Session SC_DOM_Components

Theory Core_DOM_DOM_Components

(***********************************************************************************
 * Copyright (c) 2016-2020 The University of Sheffield, UK
 *               2019-2020 University of Exeter, UK
 *
 * All rights reserved.
 *
 * Redistribution and use in source and binary forms, with or without
 * modification, are permitted provided that the following conditions are met:
 *
 * * Redistributions of source code must retain the above copyright notice, this
 *   list of conditions and the following disclaimer.
 *
 * * Redistributions in binary form must reproduce the above copyright notice,
 *   this list of conditions and the following disclaimer in the documentation
 *   and/or other materials provided with the distribution.
 *
 * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
 * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
 * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
 * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
 * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
 * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
 * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
 * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
 * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
 * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
 *
 * SPDX-License-Identifier: BSD-2-Clause
 ***********************************************************************************)

section ‹Core SC DOM Components›
theory Core_DOM_DOM_Components
  imports Core_SC_DOM.Core_DOM
begin

subsection ‹Components›

locale l_get_dom_componentCore_DOM_defs =
  l_get_root_node_defs get_root_node get_root_node_locs +
  l_to_tree_order_defs to_tree_order
  for get_root_node :: "(_) object_ptr  ((_) heap, exception, (_) object_ptr) prog"
    and get_root_node_locs :: "((_) heap  (_) heap  bool) set"
    and to_tree_order :: "(_) object_ptr  ((_) heap, exception, (_) object_ptr list) prog"
begin

definition a_get_dom_component :: "(_) object_ptr  (_, (_) object_ptr list) dom_prog"
  where
    "a_get_dom_component ptr = do {
      root  get_root_node ptr;
      to_tree_order root
    }"

definition a_is_strongly_dom_component_safe ::
  "(_) object_ptr set  (_) object_ptr set  (_) heap  (_) heap  bool"
  where
    "a_is_strongly_dom_component_safe Sarg Sresult h h' = (
      let removed_pointers = fset (object_ptr_kinds h) - fset (object_ptr_kinds h') in
      let added_pointers = fset (object_ptr_kinds h') - fset (object_ptr_kinds h) in
      let arg_components =
        (ptr  (ptr  Sarg. set |h  a_get_dom_component ptr|r) 
        fset (object_ptr_kinds h).  set |h  a_get_dom_component ptr|r) in
      let arg_components' =
        (ptr  (ptr  Sarg. set |h  a_get_dom_component ptr|r) 
        fset (object_ptr_kinds h').  set |h'  a_get_dom_component ptr|r) in
      removed_pointers  arg_components 
      added_pointers  arg_components' 
      Sresult  arg_components' 
      (outside_ptr  fset (object_ptr_kinds h)  fset (object_ptr_kinds h') -
        (ptr  Sarg. set |h   a_get_dom_component ptr|r). preserved (get_M outside_ptr id) h h'))"

definition a_is_weakly_dom_component_safe ::
  "(_) object_ptr set  (_) object_ptr set  (_) heap  (_) heap  bool"
  where
    "a_is_weakly_dom_component_safe Sarg Sresult h h' = (
      let removed_pointers = fset (object_ptr_kinds h) - fset (object_ptr_kinds h') in
      let added_pointers = fset (object_ptr_kinds h') - fset (object_ptr_kinds h) in
      let arg_components =
        (ptr  (ptr  Sarg. set |h  a_get_dom_component ptr|r) 
        fset (object_ptr_kinds h).  set |h  a_get_dom_component ptr|r) in
      let arg_components' =
        (ptr  (ptr  Sarg. set |h  a_get_dom_component ptr|r) 
        fset (object_ptr_kinds h').  set |h'  a_get_dom_component ptr|r) in
      removed_pointers  arg_components 
      Sresult  arg_components'  added_pointers  
      (outside_ptr  fset (object_ptr_kinds h)  fset (object_ptr_kinds h') -
        (ptr  Sarg. set |h   a_get_dom_component ptr|r). preserved (get_M outside_ptr id) h h'))"

lemma "a_is_strongly_dom_component_safe Sarg Sresult h h'  a_is_weakly_dom_component_safe Sarg Sresult h h'"
  by(auto simp add: a_is_strongly_dom_component_safe_def a_is_weakly_dom_component_safe_def Let_def)

definition is_document_component :: "(_) object_ptr list  bool"
  where
    "is_document_component c = is_document_ptr_kind (hd c)"

definition is_disconnected_component :: "(_) object_ptr list  bool"
  where
    "is_disconnected_component c = is_node_ptr_kind (hd c)"
end

global_interpretation l_get_dom_componentCore_DOM_defs get_root_node get_root_node_locs to_tree_order
  defines get_dom_component = a_get_dom_component
    and is_strongly_dom_component_safe = a_is_strongly_dom_component_safe
    and is_weakly_dom_component_safe = a_is_weakly_dom_component_safe
  .


locale l_get_dom_component_defs =
  fixes get_dom_component :: "(_) object_ptr  (_, (_) object_ptr list) dom_prog"
  fixes is_strongly_dom_component_safe ::
    "(_) object_ptr set  (_) object_ptr set  (_) heap  (_) heap  bool"
  fixes is_weakly_dom_component_safe ::
    "(_) object_ptr set  (_) object_ptr set  (_) heap  (_) heap  bool"

locale l_get_dom_componentCore_DOM =
  l_to_tree_order_wf +
  l_get_dom_component_defs +
  l_get_dom_componentCore_DOM_defs +
  l_get_ancestors +
  l_get_ancestors_wf +
  l_get_root_node +
  l_get_root_node_wfCore_DOM +
  l_get_parent +
  l_get_parent_wf +
  l_get_element_by +
  l_to_tree_order_wf_get_root_node_wf +
  (* l_to_tree_order_wfCore_DOM _ _ _ get_child_nodes +
  l_get_root_node_wfCore_DOM _ _ get_child_nodes+
  l_get_element_byCore_DOM _ _ "l_to_tree_orderCore_DOM.a_to_tree_order get_child_nodes"
  for get_child_nodes :: "(_::linorder) object_ptr ⇒ (_, (_) node_ptr list) dom_prog" *)
  assumes get_dom_component_impl: "get_dom_component = a_get_dom_component"
  assumes is_strongly_dom_component_safe_impl:
    "is_strongly_dom_component_safe = a_is_strongly_dom_component_safe"
  assumes is_weakly_dom_component_safe_impl:
    "is_weakly_dom_component_safe = a_is_weakly_dom_component_safe"
begin
lemmas get_dom_component_def = a_get_dom_component_def[folded get_dom_component_impl]
lemmas is_strongly_dom_component_safe_def =
  a_is_strongly_dom_component_safe_def[folded is_strongly_dom_component_safe_impl]
lemmas is_weakly_dom_component_safe_def =
  a_is_weakly_dom_component_safe_def[folded is_weakly_dom_component_safe_impl]

lemma get_dom_component_ptr_in_heap:
  assumes "h  ok (get_dom_component ptr)"
  shows "ptr |∈| object_ptr_kinds h"
  using assms get_root_node_ptr_in_heap
  by(auto simp add: get_dom_component_def)

lemma get_dom_component_ok:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "ptr |∈| object_ptr_kinds h"
  shows "h  ok (get_dom_component ptr)"
  using assms
  apply(auto simp add: get_dom_component_def a_get_root_node_def intro!: bind_is_OK_pure_I)[1]
  using get_root_node_ok  to_tree_order_ok get_root_node_ptr_in_heap
   apply blast
  by (simp add: local.get_root_node_root_in_heap local.to_tree_order_ok)

lemma get_dom_component_ptr:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_dom_component ptr r c"
  shows "ptr  set c"
proof(insert assms(1) assms(4), induct ptr rule: heap_wellformed_induct_rev )
  case (step child)
  then show ?case
  proof (cases "is_node_ptr_kind child")
    case True
    obtain node_ptr where
      node_ptr: "castnode_ptr2object_ptr node_ptr = child"
      using ‹is_node_ptr_kind child node_ptr_casts_commute3 by blast
    have "child |∈| object_ptr_kinds h"
      using h  get_dom_component child r c get_dom_component_ptr_in_heap by fast
    with node_ptr have "node_ptr |∈| node_ptr_kinds h"
      by auto
    then obtain parent_opt where
      parent: "h  get_parent node_ptr r parent_opt"
      using get_parent_ok type_wf h known_ptrs h
      by fast
    then show ?thesis
    proof (induct parent_opt)
      case None
      then have "h  get_root_node (cast node_ptr) r cast node_ptr"
        by (simp add: local.get_root_node_no_parent)
      then show ?case
        using type_wf h known_ptrs h node_ptr step(2)
        apply(auto simp add: get_dom_component_def a_get_root_node_def elim!: bind_returns_result_E2)[1]
        using to_tree_order_ptr_in_result returns_result_eq by fastforce
    next
      case (Some parent_ptr)
      then have "h  get_dom_component parent_ptr r c"
        using step(2) node_ptr type_wf h known_ptrs h heap_is_wellformed h
          get_root_node_parent_same
        by(auto simp add: get_dom_component_def elim!: bind_returns_result_E2
            intro!: bind_pure_returns_result_I)
      then have "parent_ptr  set c"
        using step node_ptr Some by blast
      then show ?case
        using  type_wf h known_ptrs h heap_is_wellformed h step(2) node_ptr Some
        apply(auto simp add: get_dom_component_def elim!: bind_returns_result_E2)[1]
        using to_tree_order_parent by blast
    qed
  next
    case False
    then show ?thesis
      using type_wf h known_ptrs h step(2)
      apply(auto simp add: get_dom_component_def elim!: bind_returns_result_E2)[1]
      by (metis is_OK_returns_result_I local.get_root_node_not_node_same
          local.get_root_node_ptr_in_heap local.to_tree_order_ptr_in_result returns_result_eq)
  qed
qed

lemma get_dom_component_parent_inside:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_dom_component ptr r c"
  assumes "cast node_ptr  set c"
  assumes "h  get_parent node_ptr r Some parent"
  shows "parent  set c"
proof -
  have "parent |∈| object_ptr_kinds h"
    using assms(6) get_parent_parent_in_heap by blast

  obtain root_ptr where root_ptr: "h  get_root_node ptr r root_ptr" and c: "h  to_tree_order root_ptr r c"
    using assms(4)
    by (metis (no_types, hide_lams) bind_returns_result_E2 get_dom_component_def get_root_node_pure)
  then have "h  get_root_node (cast node_ptr) r root_ptr"
    using assms(1) assms(2) assms(3) assms(5) to_tree_order_same_root by blast
  then have "h  get_root_node parent r root_ptr"
    using assms(6)  get_root_node_parent_same by blast
  then have "h  get_dom_component parent r c"
    using c get_dom_component_def by auto
  then show ?thesis
    using assms(1) assms(2) assms(3) get_dom_component_ptr by blast
qed

lemma get_dom_component_subset:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_dom_component ptr r c"
  assumes "ptr'  set c"
  shows "h  get_dom_component ptr' r c"
proof(insert assms(1) assms(5), induct ptr' rule: heap_wellformed_induct_rev )
  case (step child)
  then show ?case
  proof (cases "is_node_ptr_kind child")
    case True
    obtain node_ptr where
      node_ptr: "castnode_ptr2object_ptr node_ptr = child"
      using ‹is_node_ptr_kind child node_ptr_casts_commute3 by blast
    have "child |∈| object_ptr_kinds h"
      using to_tree_order_ptrs_in_heap assms(1) assms(2) assms(3) assms(4) step(2)
      unfolding get_dom_component_def
      by (meson bind_returns_result_E2 get_root_node_pure)
    with node_ptr have "node_ptr |∈| node_ptr_kinds h"
      by auto
    then obtain parent_opt where
      parent: "h  get_parent node_ptr r parent_opt"
      using get_parent_ok type_wf h known_ptrs h
      by fast
    then show ?thesis
    proof (induct parent_opt)
      case None
      then have "h  get_root_node child r child"
        using assms(1) get_root_node_no_parent node_ptr by blast
      then show ?case
        using type_wf h known_ptrs h node_ptr step(2) assms(4) assms(1)
        by (metis (no_types) bind_pure_returns_result_I2 bind_returns_result_E2
            get_dom_component_def get_root_node_pure is_OK_returns_result_I returns_result_eq
            to_tree_order_same_root)
    next
      case (Some parent_ptr)
      then have "h  get_dom_component parent_ptr r c"
        using step get_dom_component_parent_inside assms node_ptr by blast
      then show ?case
        using Some node_ptr
        apply(auto simp add: get_dom_component_def elim!: bind_returns_result_E2
            del: bind_pure_returns_result_I intro!: bind_pure_returns_result_I)[1]
        using get_root_node_parent_same by blast
    qed
  next
    case False
    then have "child |∈| object_ptr_kinds h"
      using assms(1) assms(4) step(2)
      by (metis (no_types, lifting) assms(2) assms(3) bind_returns_result_E2 get_root_node_pure
          get_dom_component_def to_tree_order_ptrs_in_heap)
    then have "h  get_root_node child r child"
      using assms(1) False get_root_node_not_node_same by blast
    then show ?thesis
      using assms(1) assms(2) assms(3) assms(4) step.prems
      by (metis (no_types) False child |∈| object_ptr_kinds h bind_pure_returns_result_I2
          bind_returns_result_E2 get_dom_component_def get_root_node_ok get_root_node_pure returns_result_eq
          to_tree_order_node_ptrs)
  qed
qed

lemma get_dom_component_to_tree_order_subset:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  to_tree_order ptr r nodes"
  assumes "h  get_dom_component ptr r c"
  shows "set nodes  set c"
  using assms
  apply(auto simp add: get_dom_component_def elim!: bind_returns_result_E2)[1]
  by (meson to_tree_order_subset  assms(5) contra_subsetD get_dom_component_ptr)

lemma get_dom_component_to_tree_order:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_dom_component ptr r c"
  assumes "h  to_tree_order ptr' r to"
  assumes "ptr  set to"
  shows "h  get_dom_component ptr' r c"
  by (metis (no_types, hide_lams) assms(1) assms(2) assms(3) assms(4) assms(5) assms(6)
      get_dom_component_ok get_dom_component_subset get_dom_component_to_tree_order_subset
      is_OK_returns_result_E local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap
      select_result_I2 subsetCE)

lemma get_dom_component_root_node_same:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_dom_component ptr r c"
  assumes "h  get_root_node ptr r root_ptr"
  assumes "x  set c"
  shows "h  get_root_node x r root_ptr"
proof(insert assms(1) assms(6), induct x rule: heap_wellformed_induct_rev )
  case (step child)
  then show ?case
  proof (cases "is_node_ptr_kind child")
    case True
    obtain node_ptr where
      node_ptr: "castnode_ptr2object_ptr node_ptr = child"
      using ‹is_node_ptr_kind child node_ptr_casts_commute3 by blast
    have "child |∈| object_ptr_kinds h"
      using to_tree_order_ptrs_in_heap assms(1) assms(2) assms(3) assms(4) step(2)
      unfolding get_dom_component_def
      by (meson bind_returns_result_E2 get_root_node_pure)
    with node_ptr have "node_ptr |∈| node_ptr_kinds h"
      by auto
    then obtain parent_opt where
      parent: "h  get_parent node_ptr r parent_opt"
      using get_parent_ok type_wf h known_ptrs h
      by fast
    then show ?thesis
    proof (induct parent_opt)
      case None
      then have "h  get_root_node child r child"
        using assms(1) get_root_node_no_parent  node_ptr by blast
      then show ?case
        using type_wf h known_ptrs h node_ptr step(2) assms(4) assms(1) assms(5)
        by (metis (no_types) child |∈| object_ptr_kinds h bind_pure_returns_result_I
            get_dom_component_def get_dom_component_ptr get_dom_component_subset get_root_node_pure
            is_OK_returns_result_E returns_result_eq to_tree_order_ok to_tree_order_same_root)
    next
      case (Some parent_ptr)
      then have "h  get_dom_component parent_ptr r c"
        using step get_dom_component_parent_inside assms node_ptr
        by (meson get_dom_component_subset)
      then show ?case
        using Some node_ptr
        apply(auto simp add: get_dom_component_def elim!: bind_returns_result_E2)[1]
        using get_root_node_parent_same
        using h  get_dom_component parent_ptr r c assms(1) assms(2) assms(3)
          get_dom_component_ptr step.hyps by blast
    qed
  next
    case False
    then have "child |∈| object_ptr_kinds h"
      using assms(1) assms(4) step(2)
      by (metis (no_types, lifting) assms(2) assms(3) bind_returns_result_E2 get_root_node_pure
          get_dom_component_def to_tree_order_ptrs_in_heap)
    then have "h  get_root_node child r child"
      using assms(1) False get_root_node_not_node_same by auto
    then show ?thesis
      using assms(1) assms(2) assms(3) assms(4) step.prems assms(5)
      by (metis (no_types, hide_lams) bind_returns_result_E2 get_dom_component_def
          get_root_node_pure returns_result_eq to_tree_order_same_root)
  qed
qed


lemma get_dom_component_no_overlap:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_dom_component ptr r c"
  assumes "h  get_dom_component ptr' r c'"
  shows "set c  set c' = {}  c = c'"
proof (rule ccontr, auto)
  fix x
  assume 1: "c  c'" and 2: "x  set c" and 3: "x  set c'"
  obtain root_ptr where root_ptr: "h  get_root_node ptr r root_ptr"
    using assms(4) unfolding get_dom_component_def
    by (meson bind_is_OK_E is_OK_returns_result_I)
  moreover obtain root_ptr' where root_ptr': "h  get_root_node ptr' r root_ptr'"
    using assms(5) unfolding get_dom_component_def
    by (meson bind_is_OK_E is_OK_returns_result_I)
  ultimately have "root_ptr  root_ptr'"
    using 1 assms
    unfolding get_dom_component_def
    by (meson bind_returns_result_E3 get_root_node_pure returns_result_eq)

  moreover have "h  get_root_node x r root_ptr"
    using 2 root_ptr get_dom_component_root_node_same assms by blast
  moreover have "h  get_root_node x r root_ptr'"
    using 3 root_ptr' get_dom_component_root_node_same assms by blast
  ultimately show False
    using select_result_I2 by force
qed

lemma get_dom_component_separates_tree_order:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_dom_component ptr r c"
  assumes "h  to_tree_order ptr r to"
  assumes "h  get_dom_component ptr' r c'"
  assumes "ptr'  set c"
  shows "set to  set c' = {}"
proof -
  have "c  c'"
    using assms get_dom_component_ptr by blast
  then have "set c  set c' = {}"
    using assms get_dom_component_no_overlap by blast
  moreover have "set to  set c"
    using assms get_dom_component_to_tree_order_subset by blast
  ultimately show ?thesis
    by blast
qed

lemma get_dom_component_separates_tree_order_general:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_dom_component ptr r c"
  assumes "h  to_tree_order ptr'' r to''"
  assumes "ptr''  set c"
  assumes "h  get_dom_component ptr' r c'"
  assumes "ptr'  set c"
  shows "set to''  set c' = {}"
proof -
  obtain root_ptr where root_ptr: "h  get_root_node ptr r root_ptr"
    using assms(4)
    by (metis bind_is_OK_E get_dom_component_def is_OK_returns_result_I)
  then have c: "h  to_tree_order root_ptr r c"
    using assms(4) unfolding get_dom_component_def
    by (simp add: bind_returns_result_E3)
  with root_ptr show ?thesis
    using assms get_dom_component_separates_tree_order get_dom_component_subset
    by meson
qed
end

interpretation i_get_dom_component?: l_get_dom_componentCore_DOM
  heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent
  get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe
  is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
  get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name
  get_elements_by_tag_name
  by(auto simp add: l_get_dom_componentCore_DOM_def l_get_dom_componentCore_DOM_axioms_def
      get_dom_component_def is_strongly_dom_component_safe_def is_weakly_dom_component_safe_def instances)
declare l_get_dom_componentCore_DOM_axioms [instances]


subsubsection ‹get\_child\_nodes›

locale l_get_dom_component_get_child_nodesCore_DOM =
  l_get_dom_componentCore_DOM +
  l_get_element_byCore_DOM
begin

lemma get_child_nodes_is_strongly_dom_component_safe:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_dom_component ptr r c"
  assumes "h  get_child_nodes ptr' r children"
  assumes "child  set children"
  shows "cast child  set c  ptr'  set c"
proof
  assume 1: "cast child  set c"
  obtain root_ptr where
    root_ptr: "h  get_root_node ptr r root_ptr"
    by (metis assms(4) bind_is_OK_E get_dom_component_def is_OK_returns_result_I)
  have "h  get_root_node (cast child) r root_ptr"
    using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_root_node_same root_ptr
    by blast
  then have "h  get_root_node ptr' r root_ptr"
    using assms(1) assms(2) assms(3) assms(5) assms(6) local.child_parent_dual
      local.get_root_node_parent_same by blast
  then have "h  get_dom_component ptr' r c"
    using "1" assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) local.child_parent_dual
      local.get_dom_component_parent_inside local.get_dom_component_subset by blast
  then show "ptr'  set c"
    using assms(1) assms(2) assms(3) get_dom_component_ptr by blast
next
  assume 1: "ptr'  set c"
  obtain root_ptr where
    root_ptr: "h  get_root_node ptr r root_ptr"
    by (metis assms(4) bind_is_OK_E get_dom_component_def is_OK_returns_result_I)
  have "h  get_root_node ptr' r root_ptr"
    using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_root_node_same root_ptr
    by blast
  then have "h  get_root_node (cast child) r root_ptr"
    using assms(1) assms(2) assms(3) assms(5) assms(6) local.child_parent_dual
      local.get_root_node_parent_same by blast
  then have "h  get_dom_component ptr' r c"
    using "1" assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) local.child_parent_dual
      local.get_dom_component_parent_inside local.get_dom_component_subset by blast
  then show "castnode_ptr2object_ptr child  set c"
    by (smt h  get_root_node (castnode_ptr2object_ptr child) r root_ptr assms(1) assms(2) assms(3)
        assms(5) assms(6) disjoint_iff_not_equal is_OK_returns_result_E is_OK_returns_result_I
        l_get_dom_componentCore_DOM.get_dom_component_no_overlap local.child_parent_dual local.get_dom_component_ok
        local.get_dom_component_parent_inside local.get_dom_component_ptr local.get_root_node_ptr_in_heap
        local.l_get_dom_componentCore_DOM_axioms)
qed


lemma get_child_nodes_get_dom_component:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_dom_component ptr r c"
  assumes "h  get_child_nodes ptr r children"
  shows "cast ` set children  set c"
  using assms get_child_nodes_is_strongly_dom_component_safe
  using local.get_dom_component_ptr by auto


lemma
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_child_nodes ptr r children"
  assumes "h  get_child_nodes ptr h h'"
  shows "is_strongly_dom_component_safe {ptr} (cast ` set children) h h'"
proof -
  have "h = h'"
    using assms(5)
    by (meson local.get_child_nodes_pure pure_returns_heap_eq)
  then show ?thesis
    using assms
    apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def)[1]
    by (smt IntI finite_set_in get_child_nodes_is_strongly_dom_component_safe is_OK_returns_result_E
        is_OK_returns_result_I local.get_child_nodes_ptr_in_heap local.get_dom_component_impl
        local.get_dom_component_ok local.get_dom_component_ptr select_result_I2)
qed
end

interpretation i_get_dom_component_get_child_nodes?: l_get_dom_component_get_child_nodesCore_DOM
  heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent
  get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe
  is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
  get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name
  get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs
  by(auto simp add: l_get_dom_component_get_child_nodesCore_DOM_def instances)
declare l_get_dom_component_get_child_nodesCore_DOM_axioms [instances]


subsubsection ‹get\_parent›

locale l_get_dom_component_get_parentCore_DOM =
  l_get_dom_componentCore_DOM +
  l_get_parentCore_DOM +
  l_get_element_byCore_DOM
begin

lemma get_parent_is_strongly_dom_component_safe_step:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_dom_component ptr r c"
  assumes "h  get_parent ptr' r Some parent"
  shows "parent  set c  cast ptr'  set c"
  by (meson assms(1) assms(2) assms(3) assms(4) assms(5) is_OK_returns_result_E
      l_to_tree_order_wf.to_tree_order_parent local.get_dom_component_parent_inside
      local.get_dom_component_subset local.get_dom_component_to_tree_order_subset
      local.get_parent_parent_in_heap local.l_to_tree_order_wf_axioms local.to_tree_order_ok
      local.to_tree_order_ptr_in_result subsetCE)


lemma get_parent_is_strongly_dom_component_safe:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_parent node_ptr r Some parent"
  assumes "h  get_parent node_ptr h h'"
  shows "is_strongly_dom_component_safe {cast node_ptr} {parent} h h'"
proof -
  have "h = h'"
    using assms(5)
    by (meson local.get_parent_pure pure_returns_heap_eq)
  then show ?thesis
    using assms
    apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def)[1]
    by (metis IntI finite_set_in local.get_dom_component_impl local.get_dom_component_ok
        local.get_dom_component_parent_inside local.get_dom_component_ptr local.get_parent_parent_in_heap
        local.known_ptrs_known_ptr local.parent_child_rel_child_in_heap local.parent_child_rel_parent
        returns_result_select_result)
qed
end

interpretation i_get_dom_component_get_parent?: l_get_dom_component_get_parentCore_DOM
  heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent
  get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe
  is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
  get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name
  get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs
  by(auto simp add: l_get_dom_component_get_parentCore_DOM_def instances)
declare l_get_dom_component_get_parentCore_DOM_axioms [instances]


subsubsection ‹get\_root\_node›

locale l_get_dom_component_get_root_nodeCore_DOM =
  l_get_dom_componentCore_DOM +
  l_get_root_nodeCore_DOM +
  l_get_element_byCore_DOM
begin

lemma get_root_node_is_strongly_dom_component_safe_step:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_dom_component ptr r c"
  assumes "h  get_root_node ptr' r root"
  shows "root  set c  ptr'  set c"
proof
  assume 1: "root  set c"
  obtain root_ptr where
    root_ptr: "h  get_root_node ptr r root_ptr"
    by (metis assms(4) bind_is_OK_E get_dom_component_def is_OK_returns_result_I)
  have "h  get_root_node root r root_ptr"
    using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_root_node_same root_ptr
    by blast
  moreover have "h  get_root_node ptr' r root_ptr"
    by (metis (no_types, lifting) calculation assms(1) assms(2) assms(3) assms(5)
        is_OK_returns_result_E local.get_root_node_root_in_heap local.to_tree_order_ok
        local.to_tree_order_ptr_in_result local.to_tree_order_same_root select_result_I2)
  ultimately have "h  get_dom_component ptr' r c"
    apply(auto simp add: get_dom_component_def)[1]
    by (smt "1" assms(1) assms(2) assms(3) assms(4) bind_pure_returns_result_I bind_returns_result_E3
        local.get_dom_component_def local.get_dom_component_subset local.get_root_node_pure)
  then show "ptr'  set c"
    using assms(1) assms(2) assms(3) get_dom_component_ptr by blast
next
  assume 1: "ptr'  set c"
  obtain root_ptr where
    root_ptr: "h  get_root_node ptr r root_ptr"
    by (metis assms(4) bind_is_OK_E get_dom_component_def is_OK_returns_result_I)
  have "h  get_root_node ptr' r root_ptr"
    using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_root_node_same root_ptr
    by blast
  then have "h  get_root_node root r root_ptr"
    by (metis assms(1) assms(2) assms(3) assms(5) is_OK_returns_result_E
        local.get_root_node_root_in_heap local.to_tree_order_ok local.to_tree_order_ptr_in_result
        local.to_tree_order_same_root returns_result_eq)
  then have "h  get_dom_component ptr' r c"
    using "1" assms(1) assms(2) assms(3) assms(4) local.get_dom_component_subset by blast
  then show "root  set c"
    using assms(5) bind_returns_result_E3 local.get_dom_component_def local.to_tree_order_ptr_in_result
    by fastforce
qed

lemma get_root_node_is_strongly_dom_component_safe:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_root_node ptr r root"
  assumes "h  get_root_node ptr h h'"
  shows "is_strongly_dom_component_safe {ptr} {root} h h'"
proof -
  have "h = h'"
    using assms(5)
    by (meson local.get_root_node_pure pure_returns_heap_eq)
  then show ?thesis
    using assms
    apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def)[1]
    by (metis (no_types, lifting) IntI finite_set_in get_root_node_is_strongly_dom_component_safe_step
        is_OK_returns_result_I local.get_dom_component_impl local.get_dom_component_ok
        local.get_dom_component_ptr local.get_root_node_ptr_in_heap returns_result_select_result)
qed
end

interpretation i_get_dom_component_get_root_node?: l_get_dom_component_get_root_nodeCore_DOM
  heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent
  get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe
  is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
  get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name
  get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs
  by(auto simp add: l_get_dom_component_get_root_nodeCore_DOM_def instances)
declare l_get_dom_component_get_root_nodeCore_DOM_axioms [instances]


subsubsection ‹get\_element\_by\_id›

locale l_get_dom_component_get_element_by_idCore_DOM =
  l_get_dom_componentCore_DOM +
  l_first_in_tree_orderCore_DOM +
  l_to_tree_orderCore_DOM +
  l_get_element_byCore_DOM
begin

lemma get_element_by_id_is_strongly_dom_component_safe_step:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_dom_component ptr r c"
  assumes "h  get_element_by_id ptr' idd r Some result"
  shows "cast result  set c  ptr'  set c"
proof
  assume 1: "cast result  set c"
  obtain root_ptr where
    root_ptr: "h  get_root_node ptr r root_ptr"
    by (metis assms(4) bind_is_OK_E get_dom_component_def is_OK_returns_result_I)
  then have "h  to_tree_order root_ptr r c"
    using h  get_dom_component ptr r c
    by (simp add: get_dom_component_def bind_returns_result_E3)

  obtain to' where to': "h  to_tree_order ptr' r to'"
    using h  get_element_by_id ptr' idd r Some result
    apply(simp add: get_element_by_id_def first_in_tree_order_def)
    by (meson bind_is_OK_E is_OK_returns_result_I)
  then have "cast result  set to'"
    using h  get_element_by_id ptr' idd r Some result get_element_by_id_result_in_tree_order
    by blast

  have "h  get_root_node (cast result) r root_ptr"
    using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_root_node_same root_ptr
    by blast
  then have "h  get_root_node ptr' r root_ptr"
    using ‹cast result  set to' h  to_tree_order ptr' r to'
    using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_ptr get_dom_component_root_node_same
      get_dom_component_subset get_dom_component_to_tree_order
    by blast
  then have "h  get_dom_component ptr' r c"
    using h  to_tree_order root_ptr r c
    using get_dom_component_def by auto
  then show "ptr'  set c"
    using assms(1) assms(2) assms(3) get_dom_component_ptr by blast
next
  assume "ptr'  set c"
  moreover obtain to' where to': "h  to_tree_order ptr' r to'"
    by (meson assms(1) assms(2) assms(3) assms(4) calculation get_dom_component_ptr_in_heap
        get_dom_component_subset is_OK_returns_result_E is_OK_returns_result_I to_tree_order_ok)
  ultimately have "set to'  set c"
    using assms(1) assms(2) assms(3) assms(4) get_dom_component_subset
      get_dom_component_to_tree_order_subset
    by blast
  moreover have "cast result  set to'"
    using assms(5) get_element_by_id_result_in_tree_order to' by blast
  ultimately show "cast result  set c"
    by blast
qed


lemma get_element_by_id_is_strongly_dom_component_safe:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_element_by_id ptr idd r Some result"
  assumes "h  get_element_by_id ptr idd h h'"
  shows "is_strongly_dom_component_safe {ptr} {cast result} h h'"
proof -
  have "h = h'"
    using assms(5)
    by(auto simp add: preserved_def get_element_by_id_def first_in_tree_order_def
        elim!: bind_returns_heap_E2 intro!: map_filter_M_pure bind_pure_I
        split: option.splits list.splits)
  have "ptr |∈| object_ptr_kinds h"
    using assms(4)
    apply(auto simp add: get_element_by_id_def)[1]
    by (metis (no_types, lifting) assms(1) assms(2) assms(3) bind_is_OK_E is_OK_returns_result_I
        local.first_in_tree_order_def local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap)
  obtain to where to: "h  to_tree_order ptr r to"
    by (meson ptr |∈| object_ptr_kinds h assms(1) assms(2) assms(3) is_OK_returns_result_E
        local.to_tree_order_ok)
  then have "cast result  set to"
    using assms(4) local.get_element_by_id_result_in_tree_order by auto
  obtain c where c: "h  a_get_dom_component ptr r c"
    using ptr |∈| object_ptr_kinds h assms(1) assms(2) assms(3) local.get_dom_component_impl
      local.get_dom_component_ok by blast

  then show ?thesis
    using assms h = h'
    apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def get_element_by_id_def
        first_in_tree_order_def elim!: bind_returns_result_E2 intro!: map_filter_M_pure bind_pure_I
        split: option.splits list.splits)[1]
    by (metis (no_types, lifting) Int_iff ptr |∈| object_ptr_kinds h assms(4) finite_set_in
        get_element_by_id_is_strongly_dom_component_safe_step local.get_dom_component_impl
        local.get_dom_component_ptr select_result_I2)
qed
end

interpretation i_get_dom_component_get_element_by_id?: l_get_dom_component_get_element_by_idCore_DOM
  heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent
  get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe
  is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
  get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name
  get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs
  by(auto simp add: l_get_dom_component_get_element_by_idCore_DOM_def instances)
declare l_get_dom_component_get_element_by_idCore_DOM_axioms [instances]

subsubsection ‹get\_elements\_by\_class\_name›

locale l_get_dom_component_get_elements_by_class_nameCore_DOM =
  l_get_dom_componentCore_DOM +
  l_first_in_tree_orderCore_DOM +
  l_to_tree_orderCore_DOM +
  l_get_element_byCore_DOM
begin

lemma get_elements_by_class_name_is_strongly_dom_component_safe_step:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_dom_component ptr r c"
  assumes "h  get_elements_by_class_name ptr' idd r results"
  assumes "result  set results"
  shows "cast result  set c  ptr'  set c"
proof
  assume 1: "cast result  set c"
  obtain root_ptr where
    root_ptr: "h  get_root_node ptr r root_ptr"
    by (metis assms(4) bind_is_OK_E get_dom_component_def is_OK_returns_result_I)
  then have "h  to_tree_order root_ptr r c"
    using h  get_dom_component ptr r c
    by (simp add: get_dom_component_def bind_returns_result_E3)

  obtain to' where to': "h  to_tree_order ptr' r to'"
    using assms(5)
    apply(simp add: get_elements_by_class_name_def first_in_tree_order_def)
    by (meson bind_is_OK_E is_OK_returns_result_I)
  then have "cast result  set to'"
    using assms get_elements_by_class_name_result_in_tree_order by blast

  have "h  get_root_node (cast result) r root_ptr"
    using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_root_node_same root_ptr
    by blast
  then have "h  get_root_node ptr' r root_ptr"
    using ‹cast result  set to' h  to_tree_order ptr' r to'
    using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_ptr get_dom_component_root_node_same
      get_dom_component_subset get_dom_component_to_tree_order
    by blast
  then have "h  get_dom_component ptr' r c"
    using h  to_tree_order root_ptr r c
    using get_dom_component_def by auto
  then show "ptr'  set c"
    using assms(1) assms(2) assms(3) get_dom_component_ptr by blast
next
  assume "ptr'  set c"
  moreover obtain to' where to': "h  to_tree_order ptr' r to'"
    by (meson assms(1) assms(2) assms(3) assms(4) calculation get_dom_component_ptr_in_heap
        get_dom_component_subset is_OK_returns_result_E is_OK_returns_result_I to_tree_order_ok)
  ultimately have "set to'  set c"
    using assms(1) assms(2) assms(3) assms(4) get_dom_component_subset
      get_dom_component_to_tree_order_subset
    by blast
  moreover have "cast result  set to'"
    using assms get_elements_by_class_name_result_in_tree_order to' by blast
  ultimately show "cast result  set c"
    by blast
qed

lemma get_elements_by_class_name_pure [simp]:
  "pure (get_elements_by_class_name ptr name) h"
  by(auto simp add: get_elements_by_class_name_def intro!: bind_pure_I map_filter_M_pure
      split: option.splits)

lemma get_elements_by_class_name_is_strongly_dom_component_safe:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_elements_by_class_name ptr name r results"
  assumes "h  get_elements_by_class_name ptr name h h'"
  shows "is_strongly_dom_component_safe {ptr} (cast ` set results) h h'"
proof -
  have "h = h'"
    using assms(5)
    by (meson get_elements_by_class_name_pure pure_returns_heap_eq)
  have "ptr |∈| object_ptr_kinds h"
    using assms(4)
    apply(auto simp add: get_elements_by_class_name_def)[1]
    by (metis (no_types, lifting) assms(1) assms(2) assms(3) bind_is_OK_E is_OK_returns_result_I
        local.first_in_tree_order_def local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap)
  obtain to where to: "h  to_tree_order ptr r to"
    by (meson ptr |∈| object_ptr_kinds h assms(1) assms(2) assms(3) is_OK_returns_result_E
        local.to_tree_order_ok)
  then have "cast ` set results  set to"
    using assms(4) local.get_elements_by_class_name_result_in_tree_order by auto
  obtain c where c: "h  a_get_dom_component ptr r c"
    using ptr |∈| object_ptr_kinds h assms(1) assms(2) assms(3) local.get_dom_component_impl
      local.get_dom_component_ok by blast

  then show ?thesis
    using assms h = h'
    apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def
        get_elements_by_class_name_def first_in_tree_order_def elim!: bind_returns_result_E2
        intro!: map_filter_M_pure bind_pure_I split: option.splits list.splits)[1]
    by (metis (no_types, lifting) Int_iff ptr |∈| object_ptr_kinds h assms(4) finite_set_in
        get_elements_by_class_name_is_strongly_dom_component_safe_step local.get_dom_component_impl
        local.get_dom_component_ptr select_result_I2)
qed
end

interpretation i_get_dom_component_get_elements_by_class_name?:
  l_get_dom_component_get_elements_by_class_nameCore_DOM
  heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent
  get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe
  is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
  get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name
  get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs
  by(auto simp add: l_get_dom_component_get_elements_by_class_nameCore_DOM_def instances)
declare l_get_dom_component_get_elements_by_class_nameCore_DOM_axioms [instances]


subsubsection ‹get\_elements\_by\_tag\_name›

locale l_get_dom_component_get_elements_by_tag_nameCore_DOM =
  l_get_dom_componentCore_DOM +
  l_first_in_tree_orderCore_DOM +
  l_to_tree_orderCore_DOM +
  l_get_element_byCore_DOM
begin

lemma get_elements_by_tag_name_is_strongly_dom_component_safe_step:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_dom_component ptr r c"
  assumes "h  get_elements_by_tag_name ptr' idd r results"
  assumes "result  set results"
  shows "cast result  set c  ptr'  set c"
proof
  assume 1: "cast result  set c"
  obtain root_ptr where
    root_ptr: "h  get_root_node ptr r root_ptr"
    by (metis assms(4) bind_is_OK_E get_dom_component_def is_OK_returns_result_I)
  then have "h  to_tree_order root_ptr r c"
    using h  get_dom_component ptr r c
    by (simp add: get_dom_component_def bind_returns_result_E3)

  obtain to' where to': "h  to_tree_order ptr' r to'"
    using assms(5)
    apply(simp add: get_elements_by_tag_name_def first_in_tree_order_def)
    by (meson bind_is_OK_E is_OK_returns_result_I)
  then have "cast result  set to'"
    using assms get_elements_by_tag_name_result_in_tree_order by blast

  have "h  get_root_node (cast result) r root_ptr"
    using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_root_node_same root_ptr
    by blast
  then have "h  get_root_node ptr' r root_ptr"
    using ‹cast result  set to' h  to_tree_order ptr' r to'
    using "1" assms(1) assms(2) assms(3) assms(4) get_dom_component_ptr
      get_dom_component_root_node_same get_dom_component_subset get_dom_component_to_tree_order
    by blast
  then have "h  get_dom_component ptr' r c"
    using h  to_tree_order root_ptr r c
    using get_dom_component_def by auto
  then show "ptr'  set c"
    using assms(1) assms(2) assms(3) get_dom_component_ptr by blast
next
  assume "ptr'  set c"
  moreover obtain to' where to': "h  to_tree_order ptr' r to'"
    by (meson assms(1) assms(2) assms(3) assms(4) calculation get_dom_component_ptr_in_heap
        get_dom_component_subset is_OK_returns_result_E is_OK_returns_result_I to_tree_order_ok)
  ultimately have "set to'  set c"
    using assms(1) assms(2) assms(3) assms(4) get_dom_component_subset
      get_dom_component_to_tree_order_subset
    by blast
  moreover have "cast result  set to'"
    using assms get_elements_by_tag_name_result_in_tree_order to' by blast
  ultimately show "cast result  set c"
    by blast
qed

lemma get_elements_by_tag_name_is_strongly_dom_component_safe:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_elements_by_tag_name ptr name r results"
  assumes "h  get_elements_by_tag_name ptr name h h'"
  shows "is_strongly_dom_component_safe {ptr} (cast ` set results) h h'"
proof -
  have "h = h'"
    using assms(5)
    by (meson get_elements_by_tag_name_pure pure_returns_heap_eq)
  have "ptr |∈| object_ptr_kinds h"
    using assms(4)
    apply(auto simp add: get_elements_by_tag_name_def)[1]
    by (metis (no_types, lifting) assms(1) assms(2) assms(3) bind_is_OK_E is_OK_returns_result_I
        local.first_in_tree_order_def local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap)
  obtain to where to: "h  to_tree_order ptr r to"
    by (meson ptr |∈| object_ptr_kinds h assms(1) assms(2) assms(3) is_OK_returns_result_E
        local.to_tree_order_ok)
  then have "cast ` set results  set to"
    using assms(4) local.get_elements_by_tag_name_result_in_tree_order by auto
  obtain c where c: "h  a_get_dom_component ptr r c"
    using ptr |∈| object_ptr_kinds h assms(1) assms(2) assms(3) local.get_dom_component_impl
      local.get_dom_component_ok by blast

  then show ?thesis
    using assms h = h'
    apply(auto simp add: is_strongly_dom_component_safe_def Let_def preserved_def
        get_elements_by_class_name_def first_in_tree_order_def elim!: bind_returns_result_E2
        intro!: map_filter_M_pure bind_pure_I split: option.splits list.splits)[1]
    by (metis (no_types, lifting) IntI ptr |∈| object_ptr_kinds h finite_set_in
        get_elements_by_tag_name_is_strongly_dom_component_safe_step local.get_dom_component_impl
        local.get_dom_component_ptr select_result_I2)
qed
end

interpretation i_get_dom_component_get_elements_by_tag_name?:
  l_get_dom_component_get_elements_by_tag_nameCore_DOM
  heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent
  get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe
  is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
  get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name
  get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs
  by(auto simp add: l_get_dom_component_get_elements_by_tag_nameCore_DOM_def instances)
declare l_get_dom_component_get_elements_by_tag_nameCore_DOM_axioms [instances]


subsubsection ‹remove\_child›

lemma remove_child_unsafe: "¬((h
    :: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap
  ) h' ptr child. heap_is_wellformed h  type_wf h  known_ptrs h 
h  remove_child ptr child h h'  is_weakly_dom_component_safe {ptr, cast child} {} h h')"
proof -
  obtain h document_ptr e1 e2 where h: "Inr ((document_ptr, e1, e2), h) = (Heap (fmempty)
      :: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap
    )  (do {
      document_ptr  create_document;
      e1  create_element document_ptr ''div'';
      e2  create_element document_ptr ''div'';
      append_child (cast e1) (cast e2);
      return (document_ptr, e1, e2)
  })"
    by(code_simp, auto simp add: equal_eq List.member_def)+
  then obtain h' where
    "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
    h': "h  remove_child (cast e1) (cast e2) h h'" and
    "¬(is_weakly_dom_component_safe {cast e1, cast e2} {} h h')"
    apply(code_simp)
    apply(clarify)
    by(code_simp, auto simp add: equal_eq List.member_def)+
  then show ?thesis
    by auto
qed


subsubsection ‹adopt\_node›

lemma adopt_node_unsafe: "¬((h
    :: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap
  ) h' document_ptr child. heap_is_wellformed h  type_wf h  known_ptrs h  h  adopt_node document_ptr child h h'  is_weakly_dom_component_safe {cast document_ptr, cast child} {} h h')"
proof -
  obtain h document_ptr document_ptr2 e1 where h: "Inr ((document_ptr, document_ptr2, e1), h) = (Heap (fmempty)
      :: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap
    )  (do {
      document_ptr  create_document;
      document_ptr2  create_document;
      e1  create_element document_ptr ''div'';
      adopt_node document_ptr2 (cast e1);
      return (document_ptr, document_ptr2, e1)
  })"
    by(code_simp, auto simp add: equal_eq List.member_def)+
  then obtain h' where
    "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
    h': "h  adopt_node document_ptr (cast e1) h h'" and
    "¬(is_weakly_dom_component_safe {cast document_ptr, cast e1} {} h h')"
    apply(code_simp)
    apply(clarify)
    by(code_simp, auto simp add: equal_eq List.member_def)+
  then show ?thesis
    by auto
qed


subsubsection ‹create\_element›

lemma create_element_not_strongly_dom_component_safe:
  obtains
    h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" and
    h' and document_ptr and new_element_ptr and tag where
    "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
    "h  create_element document_ptr tag r new_element_ptr h h'" and
    "¬ is_strongly_dom_component_safe {cast document_ptr} {cast new_element_ptr} h h'"
proof -
  let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder},
'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder},
'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder},
'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap"
  let ?P = "create_document"
  let ?h1 = "|?h0  ?P|h"
  let ?document_ptr = "|?h0  ?P|r"

  show thesis
    apply(rule that[where h="?h1" and document_ptr="?document_ptr"])
    by code_simp+
qed


locale l_get_dom_component_create_elementCore_DOM =
  l_get_dom_componentCore_DOM heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
  to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component
  is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs
  get_ancestors get_ancestors_locs get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id
  get_elements_by_class_name get_elements_by_tag_name +
  l_create_elementCore_DOM get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
  set_disconnected_nodes_locs set_tag_name set_tag_name_locs type_wf create_element known_ptr +
  l_get_disconnected_nodesCore_DOM type_wf get_disconnected_nodes get_disconnected_nodes_locs +
  l_set_disconnected_nodesCore_DOM type_wf set_disconnected_nodes set_disconnected_nodes_locs +
  l_set_tag_nameCore_DOM type_wf set_tag_name set_tag_name_locs +
  l_new_element_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs +
  l_new_element_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs +
  l_set_tag_name_get_child_nodes type_wf set_tag_name set_tag_name_locs known_ptr
  get_child_nodes get_child_nodes_locs +
  l_create_element_wfCore_DOM known_ptr known_ptrs type_wf get_child_nodes get_child_nodes_locs
  get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel set_tag_name
  set_tag_name_locs
  set_disconnected_nodes set_disconnected_nodes_locs create_element
  for known_ptr :: "(_::linorder) object_ptr  bool"
    and heap_is_wellformed :: "(_) heap  bool"
    and parent_child_rel :: "(_) heap  ((_) object_ptr × (_) object_ptr) set"
    and type_wf :: "(_) heap  bool"
    and known_ptrs :: "(_) heap  bool"
    and to_tree_order :: "(_) object_ptr  ((_) heap, exception, (_) object_ptr list) prog"
    and get_parent :: "(_) node_ptr  ((_) heap, exception, (_) object_ptr option) prog"
    and get_parent_locs :: "((_) heap  (_) heap  bool) set"
    and get_child_nodes :: "(_) object_ptr  ((_) heap, exception, (_) node_ptr list) prog"
    and get_child_nodes_locs :: "(_) object_ptr  ((_) heap  (_) heap  bool) set"
    and get_dom_component :: "(_) object_ptr  ((_) heap, exception, (_) object_ptr list) prog"
    and is_strongly_dom_component_safe :: "(_) object_ptr set  (_) object_ptr set  (_) heap  (_) heap  bool"
    and is_weakly_dom_component_safe :: "(_) object_ptr set  (_) object_ptr set  (_) heap  (_) heap  bool"
    and get_root_node :: "(_) object_ptr  ((_) heap, exception, (_) object_ptr) prog"
    and get_root_node_locs :: "((_) heap  (_) heap  bool) set"
    and get_ancestors :: "(_) object_ptr  ((_) heap, exception, (_) object_ptr list) prog"
    and get_ancestors_locs :: "((_) heap  (_) heap  bool) set"
    and get_element_by_id :: "(_) object_ptr  char list  ((_) heap, exception, (_) element_ptr option) prog"
    and get_elements_by_class_name :: "(_) object_ptr  char list  ((_) heap, exception, (_) element_ptr list) prog"
    and get_elements_by_tag_name :: "(_) object_ptr  char list  ((_) heap, exception, (_) element_ptr list) prog"
    and get_disconnected_nodes :: "(_) document_ptr  ((_) heap, exception, (_) node_ptr list) prog"
    and get_disconnected_nodes_locs :: "(_) document_ptr  ((_) heap  (_) heap  bool) set"
    and set_disconnected_nodes :: "(_) document_ptr  (_) node_ptr list  ((_) heap, exception, unit) prog"
    and set_disconnected_nodes_locs :: "(_) document_ptr  ((_) heap, exception, unit) prog set"
    and set_tag_name :: "(_) element_ptr  char list  ((_) heap, exception, unit) prog"
    and set_tag_name_locs :: "(_) element_ptr  ((_) heap, exception, unit) prog set"
    and create_element :: "(_) document_ptr  char list  ((_) heap, exception, (_) element_ptr) prog"
begin

lemma create_element_is_weakly_dom_component_safe_step:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  create_element document_ptr tag h h'"
  assumes "ptr  set |h  get_dom_component (cast document_ptr)|r"
  assumes "ptr  cast |h  create_element document_ptr tag|r"
  shows "preserved (get_M ptr getter) h h'"
proof -
  obtain new_element_ptr h2 h3 disc_nodes where
    new_element_ptr: "h  new_element r new_element_ptr" and
    h2: "h  new_element h h2" and
    h3: "h2 set_tag_name new_element_ptr tag h h3" and
    disc_nodes: "h3  get_disconnected_nodes document_ptr r disc_nodes" and
    h': "h3  set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes) h h'"
    using assms(4)
    by(auto simp add: create_element_def
        elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated])
  have "h  create_element document_ptr tag r new_element_ptr"
    using new_element_ptr h2 h3 disc_nodes h'
    apply(auto simp add: create_element_def intro!: bind_returns_result_I
        bind_pure_returns_result_I[OF get_disconnected_nodes_pure])[1]
     apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
    by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
  have "preserved (get_M ptr getter) h h2"
    using h2 new_element_ptr
    apply(rule new_element_get_MObject)
    using new_element_ptr assms(6) h  create_element document_ptr tag r new_element_ptr
    by simp

  have "preserved (get_M ptr getter) h2 h3"
    using set_tag_name_writes h3
    apply(rule reads_writes_preserved2)
    apply(auto simp add: set_tag_name_locs_impl a_set_tag_name_locs_def all_args_def)[1]
    by (metis (no_types, lifting) h  create_element document_ptr tag r new_element_ptr assms(6)
        get_M_Element_preserved8 select_result_I2)

  have "document_ptr |∈| document_ptr_kinds h"
    using create_element_document_in_heap
    using assms(4)
    by blast
  then
  have "ptr  (cast document_ptr)"
    using assms(5) assms(1) assms(2) assms(3) local.get_dom_component_ok local.get_dom_component_ptr
    by auto
  have "preserved (get_M ptr getter) h3 h'"
    using set_disconnected_nodes_writes h'
    apply(rule reads_writes_preserved2)
    apply(auto simp add: set_disconnected_nodes_locs_def all_args_def)[1]
    by (metis ptr  castdocument_ptr2object_ptr document_ptr get_M_Mdocument_preserved3)

  show ?thesis
    using ‹preserved (get_M ptr getter) h h2 ‹preserved (get_M ptr getter) h2 h3
      ‹preserved (get_M ptr getter) h3 h'
    by(auto simp add: preserved_def)
qed


lemma create_element_is_weakly_dom_component_safe:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  create_element document_ptr tag r result"
  assumes "h  create_element document_ptr tag h h'"
  shows "is_weakly_dom_component_safe {cast document_ptr} {cast result} h h'"
proof -

  obtain new_element_ptr h2 h3 disc_nodes_h3 where
    new_element_ptr: "h  new_element r new_element_ptr" and
    h2: "h  new_element h h2" and
    h3: "h2  set_tag_name new_element_ptr tag h h3" and
    disc_nodes_h3: "h3  get_disconnected_nodes document_ptr r disc_nodes_h3" and
    h': "h3  set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes_h3) h h'"
    using assms(5)
    by(auto simp add: create_element_def
        elim!: bind_returns_heap_E
        bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] )
  then have "h  create_element document_ptr tag r new_element_ptr"
    apply(auto simp add: create_element_def intro!: bind_returns_result_I)[1]
      apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
     apply (metis is_OK_returns_heap_E is_OK_returns_result_I local.get_disconnected_nodes_pure
        pure_returns_heap_eq)
    by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)

  have "new_element_ptr  set |h  element_ptr_kinds_M|r"
    using new_element_ptr ElementMonad.ptr_kinds_ptr_kinds_M h2
    using new_element_ptr_not_in_heap by blast
  then have "cast new_element_ptr  set |h  node_ptr_kinds_M|r"
    by simp
  then have "cast new_element_ptr  set |h  object_ptr_kinds_M|r"
    by simp

  have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |∪| {|cast new_element_ptr|}"
    using new_element_new_ptr h2 new_element_ptr by blast
  then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = node_ptr_kinds h |∪| {|cast new_element_ptr|}"
    apply(simp add: node_ptr_kinds_def)
    by force
  then have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h |∪| {|new_element_ptr|}"
    apply(simp add: element_ptr_kinds_def)
    by force
  have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h2 = character_data_ptr_kinds h"
    using object_ptr_kinds_eq_h
    by(auto simp add: node_ptr_kinds_def character_data_ptr_kinds_def)
  have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h"
    using object_ptr_kinds_eq_h
    by(auto simp add: document_ptr_kinds_def)

  have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2"
    apply(rule writes_small_big[where P="λh h'. object_ptr_kinds h' = object_ptr_kinds h",
          OF set_tag_name_writes h3])
    using set_tag_name_pointers_preserved
    by (auto simp add: reflp_def transp_def)
  then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2"
    by (auto simp add: document_ptr_kinds_def)
  have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2"
    using object_ptr_kinds_eq_h2
    by(auto simp add: node_ptr_kinds_def)

  have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3"
    apply(rule writes_small_big[where P="λh h'. object_ptr_kinds h' = object_ptr_kinds h",
          OF set_disconnected_nodes_writes h'])
    using set_disconnected_nodes_pointers_preserved
    by (auto simp add: reflp_def transp_def)
  then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3"
    by (auto simp add: document_ptr_kinds_def)
  have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3"
    using object_ptr_kinds_eq_h3
    by(auto simp add: node_ptr_kinds_def)

  have "known_ptr (cast new_element_ptr)"
    using h  create_element document_ptr tag r new_element_ptr local.create_element_known_ptr
    by blast
  then
  have "known_ptrs h2"
    using known_ptrs_new_ptr object_ptr_kinds_eq_h known_ptrs h h2
    by blast
  then
  have "known_ptrs h3"
    using known_ptrs_preserved object_ptr_kinds_eq_h2 by blast
  then
  have "known_ptrs h'"
    using known_ptrs_preserved object_ptr_kinds_eq_h3 by blast


  have "document_ptr |∈| document_ptr_kinds h"
    using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2
      get_disconnected_nodes_ptr_in_heap type_wf h document_ptr_kinds_def
    by (metis is_OK_returns_result_I)

  have children_eq_h: "(ptr'::(_) object_ptr) children. ptr'  cast new_element_ptr
               h  get_child_nodes ptr' r children = h2  get_child_nodes ptr' r children"
    using get_child_nodes_reads h2 get_child_nodes_new_element[rotated, OF new_element_ptr h2]
    apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
    by blast+
  then have children_eq2_h: "ptr'. ptr'  cast new_element_ptr
                                     |h  get_child_nodes ptr'|r = |h2  get_child_nodes ptr'|r"
    using select_result_eq by force

  have "h2  get_child_nodes (cast new_element_ptr) r []"
    using new_element_ptr h2 new_element_ptr_in_heap[OF h2 new_element_ptr]
      new_element_is_element_ptr[OF new_element_ptr] new_element_no_child_nodes
    by blast
  have disconnected_nodes_eq_h:
    "doc_ptr disc_nodes. h  get_disconnected_nodes doc_ptr r disc_nodes
                                             = h2  get_disconnected_nodes doc_ptr r disc_nodes"
    using get_disconnected_nodes_reads h2 get_disconnected_nodes_new_element[OF new_element_ptr h2]
    apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
    by blast+
  then have disconnected_nodes_eq2_h:
    "doc_ptr. |h  get_disconnected_nodes doc_ptr|r = |h2  get_disconnected_nodes doc_ptr|r"
    using select_result_eq by force

  have children_eq_h2:
    "ptr' children. h2  get_child_nodes ptr' r children = h3  get_child_nodes ptr' r children"
    using get_child_nodes_reads set_tag_name_writes h3
    apply(rule reads_writes_preserved)
    by(auto simp add: set_tag_name_get_child_nodes)
  then have children_eq2_h2: "ptr'. |h2  get_child_nodes ptr'|r = |h3  get_child_nodes ptr'|r"
    using select_result_eq by force
  have disconnected_nodes_eq_h2:
    "doc_ptr disc_nodes. h2  get_disconnected_nodes doc_ptr r disc_nodes
                                               = h3  get_disconnected_nodes doc_ptr r disc_nodes"
    using get_disconnected_nodes_reads set_tag_name_writes h3
    apply(rule reads_writes_preserved)
    by(auto simp add: set_tag_name_get_disconnected_nodes)
  then have disconnected_nodes_eq2_h2:
    "doc_ptr. |h2  get_disconnected_nodes doc_ptr|r = |h3  get_disconnected_nodes doc_ptr|r"
    using select_result_eq by force

  have "type_wf h2"
    using type_wf h new_element_types_preserved h2 by blast
  then have "type_wf h3"
    using writes_small_big[where P="λh h'. type_wf h  type_wf h'", OF set_tag_name_writes h3]
    using  set_tag_name_types_preserved
    by(auto simp add: reflp_def transp_def)
  then have "type_wf h'"
    using writes_small_big[where P="λh h'. type_wf h  type_wf h'", OF set_disconnected_nodes_writes h']
    using  set_disconnected_nodes_types_preserved
    by(auto simp add: reflp_def transp_def)

  have children_eq_h3:
    "ptr' children. h3  get_child_nodes ptr' r children = h'  get_child_nodes ptr' r children"
    using get_child_nodes_reads set_disconnected_nodes_writes h'
    apply(rule reads_writes_preserved)
    by(auto simp add: set_disconnected_nodes_get_child_nodes)
  then have children_eq2_h3: "ptr'. |h3  get_child_nodes ptr'|r = |h'  get_child_nodes ptr'|r"
    using select_result_eq by force
  have disconnected_nodes_eq_h3:
    "doc_ptr disc_nodes. document_ptr  doc_ptr
        h3  get_disconnected_nodes doc_ptr r disc_nodes
                                               = h'  get_disconnected_nodes doc_ptr r disc_nodes"
    using get_disconnected_nodes_reads set_disconnected_nodes_writes h'
    apply(rule reads_writes_preserved)
    by(auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers)
  then have disconnected_nodes_eq2_h3:
    "doc_ptr. document_ptr  doc_ptr
                 |h3  get_disconnected_nodes doc_ptr|r = |h'  get_disconnected_nodes doc_ptr|r"
    using select_result_eq by force

  have disc_nodes_document_ptr_h2: "h2  get_disconnected_nodes document_ptr r disc_nodes_h3"
    using disconnected_nodes_eq_h2 disc_nodes_h3 by auto
  then have disc_nodes_document_ptr_h: "h  get_disconnected_nodes document_ptr r disc_nodes_h3"
    using disconnected_nodes_eq_h by auto
  then have "cast new_element_ptr  set disc_nodes_h3"
    using heap_is_wellformed h
    using ‹castelement_ptr2node_ptr new_element_ptr  set |h  node_ptr_kinds_M|r
      a_all_ptrs_in_heap_def heap_is_wellformed_def
    using NodeMonad.ptr_kinds_ptr_kinds_M local.heap_is_wellformed_disc_nodes_in_heap by blast

  have "acyclic (parent_child_rel h)"
    using heap_is_wellformed h
    by (simp add: heap_is_wellformed_def acyclic_heap_def)
  also have "parent_child_rel h = parent_child_rel h2"
  proof(auto simp add: parent_child_rel_def)[1]
    fix a x
    assume 0: "a |∈| object_ptr_kinds h"
      and 1: "x  set |h  get_child_nodes a|r"
    then show "a |∈| object_ptr_kinds h2"
      by (simp add: object_ptr_kinds_eq_h)
  next
    fix a x
    assume 0: "a |∈| object_ptr_kinds h"
      and 1: "x  set |h  get_child_nodes a|r"
    then show "x  set |h2  get_child_nodes a|r"
      by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
          ‹castelement_ptr2object_ptr new_element_ptr  set |h  object_ptr_kinds_M|r children_eq2_h)
  next
    fix a x
    assume 0: "a |∈| object_ptr_kinds h2"
      and 1: "x  set |h2  get_child_nodes a|r"
    then show "a |∈| object_ptr_kinds h"
      using object_ptr_kinds_eq_h h2  get_child_nodes (castelement_ptr2object_ptr new_element_ptr) r []
      by(auto)
  next
    fix a x
    assume 0: "a |∈| object_ptr_kinds h2"
      and 1: "x  set |h2  get_child_nodes a|r"
    then show "x  set |h  get_child_nodes a|r"
      by (metis (no_types, lifting)
          h2  get_child_nodes (castelement_ptr2object_ptr new_element_ptr) r []
          children_eq2_h empty_iff empty_set image_eqI select_result_I2)
  qed
  also have " = parent_child_rel h3"
    by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2)
  also have " = parent_child_rel h'"
    by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3)
  finally have "a_acyclic_heap h'"
    by (simp add: acyclic_heap_def)

  have "a_all_ptrs_in_heap h"
    using heap_is_wellformed h  by (simp add: heap_is_wellformed_def)
  then have "a_all_ptrs_in_heap h2"
    apply(auto simp add: a_all_ptrs_in_heap_def)[1]
     apply (metis known_ptrs h2 parent_child_rel h = parent_child_rel h2 type_wf h2 assms
        funion_iff local.get_child_nodes_ok local.known_ptrs_known_ptr local.parent_child_rel_child_in_heap
        local.parent_child_rel_child_nodes2 node_ptr_kinds_commutes node_ptr_kinds_eq_h
        returns_result_select_result)
    by (metis assms disconnected_nodes_eq2_h document_ptr_kinds_eq_h funion_iff
        local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq_h
        returns_result_select_result)
  then have "a_all_ptrs_in_heap h3"
    by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2
        local.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2)
  then have "a_all_ptrs_in_heap h'"
    by (smt h2  get_child_nodes (castelement_ptr2object_ptr new_element_ptr) r [] children_eq2_h3
        disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3
        finite_set_in h' is_OK_returns_result_I
        l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
        local.a_all_ptrs_in_heap_def local.get_child_nodes_ptr_in_heap
        local.l_set_disconnected_nodes_get_disconnected_nodes_axioms node_ptr_kinds_commutes
        object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1))

  have "p. p |∈| object_ptr_kinds h  cast new_element_ptr  set |h  get_child_nodes p|r"
    using heap_is_wellformed h ‹castelement_ptr2node_ptr new_element_ptr  set |h  node_ptr_kinds_M|r
      heap_is_wellformed_children_in_heap
    by (meson NodeMonad.ptr_kinds_ptr_kinds_M a_all_ptrs_in_heap_def assms fset_mp
        fset_of_list_elem get_child_nodes_ok known_ptrs_known_ptr returns_result_select_result)
  then have "p. p |∈| object_ptr_kinds h2  cast new_element_ptr  set |h2  get_child_nodes p|r"
    using children_eq2_h
    apply(auto simp add: object_ptr_kinds_eq_h)[1]
    using h2  get_child_nodes (castelement_ptr2object_ptr new_element_ptr) r [] apply auto[1]
    by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
        ‹castelement_ptr2object_ptr new_element_ptr  set |h  object_ptr_kinds_M|r)
  then have "p. p |∈| object_ptr_kinds h3  cast new_element_ptr  set |h3  get_child_nodes p|r"
    using object_ptr_kinds_eq_h2 children_eq2_h2 by auto
  then have new_element_ptr_not_in_any_children:
    "p. p |∈| object_ptr_kinds h'  cast new_element_ptr  set |h'  get_child_nodes p|r"
    using object_ptr_kinds_eq_h3 children_eq2_h3 by auto

  have "a_distinct_lists h"
    using heap_is_wellformed h
    by (simp add: heap_is_wellformed_def)
  then have "a_distinct_lists h2"

    using h2  get_child_nodes (cast new_element_ptr) r []
    apply(auto simp add: a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h
        disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1]
       apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert)
      apply(case_tac "x=cast new_element_ptr")
       apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
      apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
     apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
     apply (metis IntI assms empty_iff local.get_child_nodes_ok
        local.heap_is_wellformed_one_parent local.known_ptrs_known_ptr returns_result_select_result)
    apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1]
    by (metis ‹local.a_distinct_lists h type_wf h2 disconnected_nodes_eq_h document_ptr_kinds_eq_h
        local.distinct_lists_no_parent local.get_disconnected_nodes_ok returns_result_select_result)

  then have "a_distinct_lists h3"
    by(auto simp add: a_distinct_lists_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2
        children_eq2_h2 object_ptr_kinds_eq_h2)
  then have "a_distinct_lists h'"
  proof(auto simp add: a_distinct_lists_def disconnected_nodes_eq2_h3 children_eq2_h3
      object_ptr_kinds_eq_h3 document_ptr_kinds_eq_h3
      intro!: distinct_concat_map_I)[1]
    fix x
    assume "distinct (concat (map (λdocument_ptr. |h3  get_disconnected_nodes document_ptr|r)
                                              (sorted_list_of_set (fset (document_ptr_kinds h3)))))"
      and "x |∈| document_ptr_kinds h3"
    then show "distinct |h'  get_disconnected_nodes x|r"
      using document_ptr_kinds_eq_h3 disconnected_nodes_eq_h3 h' set_disconnected_nodes_get_disconnected_nodes
      by (metis (no_types, lifting) ‹castelement_ptr2node_ptr new_element_ptr  set disc_nodes_h3
          ‹a_distinct_lists h3 type_wf h' disc_nodes_h3 distinct.simps(2)
          distinct_lists_disconnected_nodes get_disconnected_nodes_ok returns_result_eq
          returns_result_select_result)
  next
    fix x y xa
    assume "distinct (concat (map (λdocument_ptr. |h3  get_disconnected_nodes document_ptr|r)
                                            (sorted_list_of_set (fset (document_ptr_kinds h3)))))"
      and "x |∈| document_ptr_kinds h3"
      and "y |∈| document_ptr_kinds h3"
      and "x  y"
      and "xa  set |h'  get_disconnected_nodes x|r"
      and "xa  set |h'  get_disconnected_nodes y|r"
    moreover have "set |h3  get_disconnected_nodes x|r  set |h3  get_disconnected_nodes y|r = {}"
      using calculation by(auto dest: distinct_concat_map_E(1))
    ultimately show "False"
      apply(-)
      apply(cases "x = document_ptr")
       apply (smt NodeMonad.ptr_kinds_ptr_kinds_M
          ‹castelement_ptr2node_ptr new_element_ptr  set |h  node_ptr_kinds_M|r ‹local.a_all_ptrs_in_heap h
          disc_nodes_h3 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3
          disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h'
          l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
          local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms
          select_result_I2 set_ConsD subsetD)
      by (smt NodeMonad.ptr_kinds_ptr_kinds_M
          ‹castelement_ptr2node_ptr new_element_ptr  set |h  node_ptr_kinds_M|r ‹local.a_all_ptrs_in_heap h
          disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2
          disconnected_nodes_eq2_h3
          disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h'
          l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
          local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms
          select_result_I2 set_ConsD subsetD)
  next
    fix x xa xb
    assume 2: "(xfset (object_ptr_kinds h3). set |h'  get_child_nodes x|r)
                    (xfset (document_ptr_kinds h3). set |h3  get_disconnected_nodes x|r) = {}"
      and 3: "xa |∈| object_ptr_kinds h3"
      and 4: "x  set |h'  get_child_nodes xa|r"
      and 5: "xb |∈| document_ptr_kinds h3"
      and 6: "x  set |h'  get_disconnected_nodes xb|r"
    show "False"
      using disc_nodes_document_ptr_h disconnected_nodes_eq2_h3
      apply -
      apply(cases "xb = document_ptr")
       apply (metis (no_types, hide_lams) "3" "4" "6"
          p. p |∈| object_ptr_kinds h3
                       castelement_ptr2node_ptr new_element_ptr  set |h3  get_child_nodes p|r
          ‹a_distinct_lists h3 children_eq2_h3 disc_nodes_h3 distinct_lists_no_parent h'
          select_result_I2 set_ConsD set_disconnected_nodes_get_disconnected_nodes)
      by (metis "3" "4" "5" "6" ‹a_distinct_lists h3 type_wf h3 children_eq2_h3
          distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result)
  qed

  have "a_owner_document_valid h"
    using heap_is_wellformed h  by (simp add: heap_is_wellformed_def)
  then have "a_owner_document_valid h'"
    using disc_nodes_h3 document_ptr |∈| document_ptr_kinds h
    apply(auto simp add: a_owner_document_valid_def)[1]
    apply(auto simp add:  object_ptr_kinds_eq_h object_ptr_kinds_eq_h3 )[1]
    apply(auto simp add: object_ptr_kinds_eq_h2)[1]
    apply(auto simp add:  document_ptr_kinds_eq_h document_ptr_kinds_eq_h3 )[1]
    apply(auto simp add: document_ptr_kinds_eq_h2)[1]
    apply(auto simp add:  node_ptr_kinds_eq_h node_ptr_kinds_eq_h3 )[1]
    apply(auto simp add: node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h )[1]
     apply(auto simp add: children_eq2_h2[symmetric] children_eq2_h3[symmetric]
        disconnected_nodes_eq2_h disconnected_nodes_eq2_h2
        disconnected_nodes_eq2_h3)[1]
     apply (metis (no_types, lifting) document_ptr_kinds_eq_h h' list.set_intros(1)
        local.set_disconnected_nodes_get_disconnected_nodes select_result_I2)
    apply(simp add: object_ptr_kinds_eq_h)
    by(metis (no_types, lifting) NodeMonad.ptr_kinds_ptr_kinds_M
        ‹castelement_ptr2node_ptr new_element_ptr  set |h  node_ptr_kinds_M|r children_eq2_h
        children_eq2_h2 children_eq2_h3 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2
        disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h finite_set_in h'
        l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
        list.set_intros(2) local.l_set_disconnected_nodes_get_disconnected_nodes_axioms
        node_ptr_kinds_commutes select_result_I2)


  have "parent_child_rel h = parent_child_rel h'"
  proof -
    have "parent_child_rel h = parent_child_rel h2"
    proof(auto simp add: parent_child_rel_def)[1]
      fix a x
      assume 0: "a |∈| object_ptr_kinds h"
        and 1: "x  set |h  get_child_nodes a|r"
      then show "a |∈| object_ptr_kinds h2"
        by (simp add: object_ptr_kinds_eq_h)
    next
      fix a x
      assume 0: "a |∈| object_ptr_kinds h"
        and 1: "x  set |h  get_child_nodes a|r"
      then show "x  set |h2  get_child_nodes a|r"
        by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
            ‹cast new_element_ptr  set |h  object_ptr_kinds_M|r children_eq2_h)
    next
      fix a x
      assume 0: "a |∈| object_ptr_kinds h2"
        and 1: "x  set |h2  get_child_nodes a|r"
      then show "a |∈| object_ptr_kinds h"
        using object_ptr_kinds_eq_h h2  get_child_nodes (cast new_element_ptr) r []
        by(auto)
    next
      fix a x
      assume 0: "a |∈| object_ptr_kinds h2"
        and 1: "x  set |h2  get_child_nodes a|r"
      then show "x  set |h  get_child_nodes a|r"
        by (metis (no_types, lifting) h2  get_child_nodes (cast new_element_ptr) r []
            children_eq2_h empty_iff empty_set image_eqI select_result_I2)
    qed
    also have " = parent_child_rel h3"
      by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2)
    also have " = parent_child_rel h'"
      by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3)
    finally show ?thesis
      by simp
  qed
  have root: "h  get_root_node (cast document_ptr) r cast document_ptr"
    by (simp add: document_ptr |∈| document_ptr_kinds h local.get_root_node_not_node_same)
  then
  have root': "h'  get_root_node (cast document_ptr) r cast document_ptr"
    by (simp add: document_ptr |∈| document_ptr_kinds h document_ptr_kinds_eq_h
        local.get_root_node_not_node_same object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3)

  have "heap_is_wellformed h'" and "known_ptrs h'"
    using create_element_preserves_wellformedness assms
    by blast+

  have "cast result |∉| object_ptr_kinds h"
    using ‹cast new_element_ptr  set |h  object_ptr_kinds_M|r
    by (metis (full_types) ObjectMonad.ptr_kinds_ptr_kinds_M
        h  create_element document_ptr tag r new_element_ptr assms(4) returns_result_eq)

  obtain to where to: "h  to_tree_order (cast document_ptr) r to"
    by (meson document_ptr |∈| document_ptr_kinds h assms(1) assms(2) assms(3)
        document_ptr_kinds_commutes is_OK_returns_result_E local.to_tree_order_ok)
  then
  have "h  a_get_dom_component (cast document_ptr) r to"
    using root
    by(auto simp add: a_get_dom_component_def)
  moreover
  obtain to' where to': "h'  to_tree_order (cast document_ptr) r to'"
    by (meson heap_is_wellformed h' known_ptrs h' type_wf h' is_OK_returns_result_E
        local.get_root_node_root_in_heap local.to_tree_order_ok root')
  then
  have "h'  a_get_dom_component (cast document_ptr) r to'"
    using root'
    by(auto simp add: a_get_dom_component_def)
  moreover
  have "child. child  set to  child  set to'"
    by (metis heap_is_wellformed h' known_ptrs h' parent_child_rel h = parent_child_rel h'
        type_wf h' assms(1) assms(2) assms(3) local.to_tree_order_parent_child_rel to to')
  ultimately
  have "set |h  local.a_get_dom_component (cast document_ptr)|r =
set |h'  local.a_get_dom_component (cast document_ptr)|r"
    by(auto simp add: a_get_dom_component_def)

  show ?thesis
    apply(auto simp add: is_weakly_dom_component_safe_def Let_def)[1]
    using h2  get_child_nodes (castelement_ptr2object_ptr new_element_ptr) r [] assms(2) assms(3)
      children_eq_h local.get_child_nodes_ok local.get_child_nodes_ptr_in_heap local.known_ptrs_known_ptr
      object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 returns_result_select_result
       apply (metis is_OK_returns_result_I)
      apply (metis h  create_element document_ptr tag r new_element_ptr assms(4)
        element_ptr_kinds_commutes h2 new_element_ptr new_element_ptr_in_heap node_ptr_kinds_eq_h2
        node_ptr_kinds_eq_h3 returns_result_eq returns_result_heap_def)
    using ‹castelement_ptr2object_ptr result |∉| object_ptr_kinds h element_ptr_kinds_commutes
      node_ptr_kinds_commutes apply blast
    using assms(1) assms(2) assms(3) h  create_element document_ptr tag h h'
    apply(rule create_element_is_weakly_dom_component_safe_step)
     apply (simp add: local.get_dom_component_impl)
    using ‹castelement_ptr2object_ptr new_element_ptr  set |h  object_ptr_kinds_M|r
      h  create_element document_ptr tag r new_element_ptr
    by auto

qed
end

interpretation i_get_dom_component_create_element?: l_get_dom_component_create_elementCore_DOM
  known_ptr heap_is_wellformed parent_child_rel type_wf known_ptrs to_tree_order get_parent
  get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe
  is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
  get_element_by_id get_elements_by_class_name get_elements_by_tag_name get_disconnected_nodes
  get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_tag_name
  set_tag_name_locs create_element
  by(auto simp add: l_get_dom_component_create_elementCore_DOM_def instances)
declare l_get_dom_component_create_elementCore_DOM_axioms [instances]



subsubsection ‹create\_character\_data›

lemma create_character_data_not_strongly_dom_component_safe:
  obtains
    h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder},
'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder},
'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" and
    h' and document_ptr and create_character_datanew_character_data_ptr and tag where
    "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
    "h  create_character_data document_ptr tag r create_character_datanew_character_data_ptr h h'" and
    "¬ is_strongly_dom_component_safe {cast document_ptr} {cast create_character_datanew_character_data_ptr} h h'"
proof -
  let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder},
'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder},
'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder},
'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap"
  let ?P = "create_document"
  let ?h1 = "|?h0  ?P|h"
  let ?document_ptr = "|?h0  ?P|r"

  show thesis
    apply(rule that[where h="?h1" and document_ptr="?document_ptr"])
    by code_simp+
qed

locale l_get_dom_component_create_character_dataCore_DOM =
  l_get_dom_componentCore_DOM heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
  to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component
  is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs
  get_ancestors get_ancestors_locs get_disconnected_nodes get_disconnected_nodes_locs
  get_element_by_id get_elements_by_class_name get_elements_by_tag_name +
  l_create_character_dataCore_DOM get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
  set_disconnected_nodes_locs set_val set_val_locs type_wf create_character_data known_ptr +
  l_get_disconnected_nodesCore_DOM type_wf get_disconnected_nodes get_disconnected_nodes_locs +
  l_set_disconnected_nodesCore_DOM type_wf set_disconnected_nodes set_disconnected_nodes_locs +
  l_set_valCore_DOM type_wf set_val set_val_locs +
  l_create_character_data_wfCore_DOM known_ptr type_wf get_child_nodes get_child_nodes_locs
  get_disconnected_nodes get_disconnected_nodes_locs heap_is_wellformed parent_child_rel set_val
  set_val_locs set_disconnected_nodes set_disconnected_nodes_locs
  create_character_data known_ptrs
  for known_ptr :: "(_::linorder) object_ptr  bool"
    and heap_is_wellformed :: "(_) heap  bool"
    and parent_child_rel :: "(_) heap  ((_) object_ptr × (_) object_ptr) set"
    and type_wf :: "(_) heap  bool"
    and known_ptrs :: "(_) heap  bool"
    and to_tree_order :: "(_) object_ptr  ((_) heap, exception, (_) object_ptr list) prog"
    and get_parent :: "(_) node_ptr  ((_) heap, exception, (_) object_ptr option) prog"
    and get_parent_locs :: "((_) heap  (_) heap  bool) set"
    and get_child_nodes :: "(_) object_ptr  ((_) heap, exception, (_) node_ptr list) prog"
    and get_child_nodes_locs :: "(_) object_ptr  ((_) heap  (_) heap  bool) set"
    and get_dom_component :: "(_) object_ptr  ((_) heap, exception, (_) object_ptr list) prog"
    and is_strongly_dom_component_safe :: "(_) object_ptr set  (_) object_ptr set  (_) heap  (_) heap  bool"
    and is_weakly_dom_component_safe :: "(_) object_ptr set  (_) object_ptr set  (_) heap  (_) heap  bool"
    and get_root_node :: "(_) object_ptr  ((_) heap, exception, (_) object_ptr) prog"
    and get_root_node_locs :: "((_) heap  (_) heap  bool) set"
    and get_ancestors :: "(_) object_ptr  ((_) heap, exception, (_) object_ptr list) prog"
    and get_ancestors_locs :: "((_) heap  (_) heap  bool) set"
    and get_element_by_id :: "(_) object_ptr  char list  ((_) heap, exception, (_) element_ptr option) prog"
    and get_elements_by_class_name :: "(_) object_ptr  char list  ((_) heap, exception, (_) element_ptr list) prog"
    and get_elements_by_tag_name :: "(_) object_ptr  char list  ((_) heap, exception, (_) element_ptr list) prog"
    and set_val :: "(_) character_data_ptr  char list  ((_) heap, exception, unit) prog"
    and set_val_locs :: "(_) character_data_ptr  ((_) heap, exception, unit) prog set"
    and get_disconnected_nodes :: "(_) document_ptr  ((_) heap, exception, (_) node_ptr list) prog"
    and get_disconnected_nodes_locs :: "(_) document_ptr  ((_) heap  (_) heap  bool) set"
    and set_disconnected_nodes :: "(_) document_ptr  (_) node_ptr list  ((_) heap, exception, unit) prog"
    and set_disconnected_nodes_locs :: "(_) document_ptr  ((_) heap, exception, unit) prog set"
    and create_character_data :: "(_) document_ptr  char list  ((_) heap, exception, (_) character_data_ptr) prog"
begin

lemma create_character_data_is_weakly_dom_component_safe_step:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  create_character_data document_ptr text h h'"
  assumes "ptr  set |h  get_dom_component (cast document_ptr)|r"
  assumes "ptr  cast |h  create_character_data document_ptr text|r"
  shows "preserved (get_M ptr getter) h h'"
proof -
  obtain new_character_data_ptr h2 h3 disc_nodes where
    new_character_data_ptr: "h  new_character_data r new_character_data_ptr" and
    h2: "h  new_character_data h h2" and
    h3: "h2  set_val new_character_data_ptr text h h3" and
    disc_nodes: "h3  get_disconnected_nodes document_ptr r disc_nodes" and
    h': "h3  set_disconnected_nodes document_ptr (cast new_character_data_ptr # disc_nodes) h h'"
    using assms(4)
    by(auto simp add: create_character_data_def elim!: bind_returns_heap_E
        bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated])

  have "h  create_character_data document_ptr text r new_character_data_ptr"
    using new_character_data_ptr h2 h3 disc_nodes h'
    apply(auto simp add: create_character_data_def intro!: bind_returns_result_I
        bind_pure_returns_result_I[OF get_disconnected_nodes_pure])[1]
     apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
    by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
  have "preserved (get_M ptr getter) h h2"
    using h2 new_character_data_ptr
    apply(rule new_character_data_get_MObject)
    using new_character_data_ptr assms(6)
      h  create_character_data document_ptr text r new_character_data_ptr
    by simp

  have "preserved (get_M ptr getter) h2 h3"
    using set_val_writes h3
    apply(rule reads_writes_preserved2)
    apply(auto simp add: set_val_locs_impl a_set_val_locs_def all_args_def)[1]
    by (metis (mono_tags) CharacterData_simp11
        h  create_character_data document_ptr text r new_character_data_ptr assms(4) assms(6)
        is_OK_returns_heap_I is_OK_returns_result_E returns_result_eq select_result_I2)

  have "document_ptr |∈| document_ptr_kinds h"
    using create_character_data_document_in_heap
    using assms(4)
    by blast
  then
  have "ptr  (cast document_ptr)"
    using assms(5) assms(1) assms(2) assms(3) local.get_dom_component_ok local.get_dom_component_ptr
    by auto
  have "preserved (get_M ptr getter) h3 h'"
    using set_disconnected_nodes_writes h'
    apply(rule reads_writes_preserved2)
    apply(auto simp add: set_disconnected_nodes_locs_def all_args_def)[1]
    by (metis ptr  castdocument_ptr2object_ptr document_ptr get_M_Mdocument_preserved3)

  show ?thesis
    using ‹preserved (get_M ptr getter) h h2 ‹preserved (get_M ptr getter) h2 h3
      ‹preserved (get_M ptr getter) h3 h'
    by(auto simp add: preserved_def)
qed

lemma create_character_data_is_weakly_dom_component_safe:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  create_character_data document_ptr text r result"
  assumes "h  create_character_data document_ptr text h h'"
  shows "is_weakly_dom_component_safe {cast document_ptr} {cast result} h h'"
proof -


  obtain new_character_data_ptr h2 h3 disc_nodes_h3 where
    new_character_data_ptr: "h  new_character_data r new_character_data_ptr" and
    h2: "h  new_character_data h h2" and
    h3: "h2  set_val new_character_data_ptr text h h3" and
    disc_nodes_h3: "h3  get_disconnected_nodes document_ptr r disc_nodes_h3" and
    h': "h3  set_disconnected_nodes document_ptr (cast new_character_data_ptr # disc_nodes_h3) h h'"
    using assms(5)
    by(auto simp add: create_character_data_def
        elim!: bind_returns_heap_E
        bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] )
  then
  have "h  create_character_data document_ptr text r new_character_data_ptr"
    apply(auto simp add: create_character_data_def intro!: bind_returns_result_I)[1]
      apply (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
     apply (metis is_OK_returns_heap_E is_OK_returns_result_I local.get_disconnected_nodes_pure
        pure_returns_heap_eq)
    by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)

  have "new_character_data_ptr  set |h  character_data_ptr_kinds_M|r"
    using new_character_data_ptr CharacterDataMonad.ptr_kinds_ptr_kinds_M h2
    using new_character_data_ptr_not_in_heap by blast
  then have "cast new_character_data_ptr  set |h  node_ptr_kinds_M|r"
    by simp
  then have "cast new_character_data_ptr  set |h  object_ptr_kinds_M|r"
    by simp



  have object_ptr_kinds_eq_h:
    "object_ptr_kinds h2 = object_ptr_kinds h |∪| {|cast new_character_data_ptr|}"
    using new_character_data_new_ptr h2 new_character_data_ptr by blast
  then have node_ptr_kinds_eq_h:
    "node_ptr_kinds h2 = node_ptr_kinds h |∪| {|cast new_character_data_ptr|}"
    apply(simp add: node_ptr_kinds_def)
    by force
  then have character_data_ptr_kinds_eq_h:
    "character_data_ptr_kinds h2 = character_data_ptr_kinds h |∪| {|new_character_data_ptr|}"
    apply(simp add: character_data_ptr_kinds_def)
    by force
  have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h"
    using object_ptr_kinds_eq_h
    by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def)
  have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h"
    using object_ptr_kinds_eq_h
    by(auto simp add: document_ptr_kinds_def)

  have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2"
    apply(rule writes_small_big[where P="λh h'. object_ptr_kinds h' = object_ptr_kinds h",
          OF set_val_writes h3])
    using set_val_pointers_preserved
    by (auto simp add: reflp_def transp_def)
  then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2"
    by (auto simp add: document_ptr_kinds_def)
  have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2"
    using object_ptr_kinds_eq_h2
    by(auto simp add: node_ptr_kinds_def)

  have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3"
    apply(rule writes_small_big[where P="λh h'. object_ptr_kinds h' = object_ptr_kinds h",
          OF set_disconnected_nodes_writes h'])
    using set_disconnected_nodes_pointers_preserved
    by (auto simp add: reflp_def transp_def)
  then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3"
    by (auto simp add: document_ptr_kinds_def)
  have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3"
    using object_ptr_kinds_eq_h3
    by(auto simp add: node_ptr_kinds_def)


  have "document_ptr |∈| document_ptr_kinds h"
    using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2
      get_disconnected_nodes_ptr_in_heap type_wf h document_ptr_kinds_def
    by (metis is_OK_returns_result_I)

  have children_eq_h: "(ptr'::(_) object_ptr) children. ptr'  cast new_character_data_ptr
                   h  get_child_nodes ptr' r children = h2  get_child_nodes ptr' r children"
    using get_child_nodes_reads h2 get_child_nodes_new_character_data[rotated, OF new_character_data_ptr h2]
    apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
    by blast+
  then have children_eq2_h:
    "ptr'. ptr'  cast new_character_data_ptr
       |h  get_child_nodes ptr'|r = |h2  get_child_nodes ptr'|r"
    using select_result_eq by force
  have object_ptr_kinds_eq_h:
    "object_ptr_kinds h2 = object_ptr_kinds h |∪| {|cast new_character_data_ptr|}"
    using new_character_data_new_ptr h2 new_character_data_ptr by blast
  then have node_ptr_kinds_eq_h:
    "node_ptr_kinds h2 = node_ptr_kinds h |∪| {|cast new_character_data_ptr|}"
    apply(simp add: node_ptr_kinds_def)
    by force
  then have character_data_ptr_kinds_eq_h:
    "character_data_ptr_kinds h2 = character_data_ptr_kinds h |∪| {|new_character_data_ptr|}"
    apply(simp add: character_data_ptr_kinds_def)
    by force
  have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h"
    using object_ptr_kinds_eq_h
    by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def)
  have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h"
    using object_ptr_kinds_eq_h
    by(auto simp add: document_ptr_kinds_def)

  have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2"
    apply(rule writes_small_big[where P="λh h'. object_ptr_kinds h' = object_ptr_kinds h",
          OF set_val_writes h3])
    using set_val_pointers_preserved
    by (auto simp add: reflp_def transp_def)
  then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2"
    by (auto simp add: document_ptr_kinds_def)
  have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2"
    using object_ptr_kinds_eq_h2
    by(auto simp add: node_ptr_kinds_def)

  have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3"
    apply(rule writes_small_big[where P="λh h'. object_ptr_kinds h' = object_ptr_kinds h",
          OF set_disconnected_nodes_writes h'])
    using set_disconnected_nodes_pointers_preserved
    by (auto simp add: reflp_def transp_def)
  then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3"
    by (auto simp add: document_ptr_kinds_def)
  have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3"
    using object_ptr_kinds_eq_h3
    by(auto simp add: node_ptr_kinds_def)


  have "document_ptr |∈| document_ptr_kinds h"
    using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2
      get_disconnected_nodes_ptr_in_heap type_wf h document_ptr_kinds_def
    by (metis is_OK_returns_result_I)

  have children_eq_h: "(ptr'::(_) object_ptr) children. ptr'  cast new_character_data_ptr
                 h  get_child_nodes ptr' r children = h2  get_child_nodes ptr' r children"
    using get_child_nodes_reads h2 get_child_nodes_new_character_data[rotated, OF new_character_data_ptr h2]
    apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
    by blast+
  then have children_eq2_h: "ptr'. ptr'  cast new_character_data_ptr
                                  |h  get_child_nodes ptr'|r = |h2  get_child_nodes ptr'|r"
    using select_result_eq by force

  have "h2  get_child_nodes (cast new_character_data_ptr) r []"
    using new_character_data_ptr h2 new_character_data_ptr_in_heap[OF h2 new_character_data_ptr]
      new_character_data_is_character_data_ptr[OF new_character_data_ptr]
      new_character_data_no_child_nodes
    by blast
  have disconnected_nodes_eq_h:
    "doc_ptr disc_nodes. h  get_disconnected_nodes doc_ptr r disc_nodes
                                             = h2  get_disconnected_nodes doc_ptr r disc_nodes"
    using get_disconnected_nodes_reads h2
      get_disconnected_nodes_new_character_data[OF new_character_data_ptr h2]
    apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
    by blast+
  then have disconnected_nodes_eq2_h:
    "doc_ptr. |h  get_disconnected_nodes doc_ptr|r = |h2  get_disconnected_nodes doc_ptr|r"
    using select_result_eq by force

  have children_eq_h2:
    "ptr' children. h2  get_child_nodes ptr' r children = h3  get_child_nodes ptr' r children"
    using get_child_nodes_reads set_val_writes h3
    apply(rule reads_writes_preserved)
    by(auto simp add: set_val_get_child_nodes)
  then have children_eq2_h2:
    "ptr'. |h2  get_child_nodes ptr'|r = |h3  get_child_nodes ptr'|r"
    using select_result_eq by force
  have disconnected_nodes_eq_h2:
    "doc_ptr disc_nodes. h2  get_disconnected_nodes doc_ptr r disc_nodes
                                              = h3  get_disconnected_nodes doc_ptr r disc_nodes"
    using get_disconnected_nodes_reads set_val_writes h3
    apply(rule reads_writes_preserved)
    by(auto simp add: set_val_get_disconnected_nodes)
  then have disconnected_nodes_eq2_h2:
    "doc_ptr. |h2  get_disconnected_nodes doc_ptr|r = |h3  get_disconnected_nodes doc_ptr|r"
    using select_result_eq by force

  have "type_wf h2"
    using type_wf h new_character_data_types_preserved h2 by blast
  then have "type_wf h3"
    using writes_small_big[where P="λh h'. type_wf h  type_wf h'", OF set_val_writes h3]
    using  set_val_types_preserved
    by(auto simp add: reflp_def transp_def)
  then have "type_wf h'"
    using writes_small_big[where P="λh h'. type_wf h  type_wf h'", OF set_disconnected_nodes_writes h']
    using  set_disconnected_nodes_types_preserved
    by(auto simp add: reflp_def transp_def)

  have children_eq_h3:
    "ptr' children. h3  get_child_nodes ptr' r children = h'  get_child_nodes ptr' r children"
    using get_child_nodes_reads set_disconnected_nodes_writes h'
    apply(rule reads_writes_preserved)
    by(auto simp add: set_disconnected_nodes_get_child_nodes)
  then have children_eq2_h3:
    " ptr'. |h3  get_child_nodes ptr'|r = |h'  get_child_nodes ptr'|r"
    using select_result_eq by force
  have disconnected_nodes_eq_h3: "doc_ptr disc_nodes. document_ptr  doc_ptr
     h3  get_disconnected_nodes doc_ptr r disc_nodes
                            = h'  get_disconnected_nodes doc_ptr r disc_nodes"
    using get_disconnected_nodes_reads set_disconnected_nodes_writes h'
    apply(rule reads_writes_preserved)
    by(auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers)
  then have disconnected_nodes_eq2_h3: "doc_ptr. document_ptr  doc_ptr
              |h3  get_disconnected_nodes doc_ptr|r = |h'  get_disconnected_nodes doc_ptr|r"
    using select_result_eq by force

  have disc_nodes_document_ptr_h2: "h2  get_disconnected_nodes document_ptr r disc_nodes_h3"
    using disconnected_nodes_eq_h2 disc_nodes_h3 by auto
  then have disc_nodes_document_ptr_h: "h  get_disconnected_nodes document_ptr r disc_nodes_h3"
    using disconnected_nodes_eq_h by auto
  then have "cast new_character_data_ptr  set disc_nodes_h3"
    using heap_is_wellformed h using ‹cast new_character_data_ptr  set |h  node_ptr_kinds_M|r
      a_all_ptrs_in_heap_def heap_is_wellformed_def
    using NodeMonad.ptr_kinds_ptr_kinds_M local.heap_is_wellformed_disc_nodes_in_heap by blast

  have "parent_child_rel h = parent_child_rel h'"
  proof -
    have "parent_child_rel h = parent_child_rel h2"
    proof(auto simp add: parent_child_rel_def)[1]
      fix a x
      assume 0: "a |∈| object_ptr_kinds h"
        and 1: "x  set |h  get_child_nodes a|r"
      then show "a |∈| object_ptr_kinds h2"
        by (simp add: object_ptr_kinds_eq_h)
    next
      fix a x
      assume 0: "a |∈| object_ptr_kinds h"
        and 1: "x  set |h  get_child_nodes a|r"
      then show "x  set |h2  get_child_nodes a|r"
        by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
            ‹cast new_character_data_ptr  set |h  object_ptr_kinds_M|r children_eq2_h)
    next
      fix a x
      assume 0: "a |∈| object_ptr_kinds h2"
        and 1: "x  set |h2  get_child_nodes a|r"
      then show "a |∈| object_ptr_kinds h"
        using object_ptr_kinds_eq_h h2  get_child_nodes (cast new_character_data_ptr) r []
        by(auto)
    next
      fix a x
      assume 0: "a |∈| object_ptr_kinds h2"
        and 1: "x  set |h2  get_child_nodes a|r"
      then show "x  set |h  get_child_nodes a|r"
        by (metis (no_types, lifting) h2  get_child_nodes (cast new_character_data_ptr) r []
            children_eq2_h empty_iff empty_set image_eqI select_result_I2)
    qed
    also have " = parent_child_rel h3"
      by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2)
    also have " = parent_child_rel h'"
      by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3)
    finally show ?thesis
      by simp
  qed
  have root: "h  get_root_node (cast document_ptr) r cast document_ptr"
    by (simp add: document_ptr |∈| document_ptr_kinds h local.get_root_node_not_node_same)
  then
  have root': "h'  get_root_node (cast document_ptr) r cast document_ptr"
    by (simp add: document_ptr |∈| document_ptr_kinds h document_ptr_kinds_eq_h
        local.get_root_node_not_node_same object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3)


  have "heap_is_wellformed h'" and "known_ptrs h'"
    using create_character_data_preserves_wellformedness assms
    by blast+

  have "cast result |∉| object_ptr_kinds h"
    using ‹cast new_character_data_ptr  set |h  object_ptr_kinds_M|r
    by (metis (full_types) ObjectMonad.ptr_kinds_ptr_kinds_M
        h  create_character_data document_ptr text r new_character_data_ptr assms(4) returns_result_eq)

  obtain to where to: "h  to_tree_order (cast document_ptr) r to"
    by (meson document_ptr |∈| document_ptr_kinds h assms(1) assms(2) assms(3)
        document_ptr_kinds_commutes is_OK_returns_result_E local.to_tree_order_ok)
  then
  have "h  a_get_dom_component (cast document_ptr) r to"
    using root
    by(auto simp add: a_get_dom_component_def)
  moreover
  obtain to' where to': "h'  to_tree_order (cast document_ptr) r to'"
    by (meson heap_is_wellformed h' known_ptrs h' type_wf h' is_OK_returns_result_E
        local.get_root_node_root_in_heap local.to_tree_order_ok root')
  then
  have "h'  a_get_dom_component (cast document_ptr) r to'"
    using root'
    by(auto simp add: a_get_dom_component_def)
  moreover
  have "child. child  set to  child  set to'"
    by (metis heap_is_wellformed h' known_ptrs h' parent_child_rel h = parent_child_rel h'
        type_wf h' assms(1) assms(2) assms(3) local.to_tree_order_parent_child_rel to to')
  ultimately
  have "set |h  local.a_get_dom_component (cast document_ptr)|r =
set |h'  local.a_get_dom_component (cast document_ptr)|r"
    by(auto simp add: a_get_dom_component_def)

  show ?thesis
    apply(auto simp add: is_weakly_dom_component_safe_def Let_def)[1]
    using  assms(2) assms(3) children_eq_h local.get_child_nodes_ok
      local.get_child_nodes_ptr_in_heap local.known_ptrs_known_ptr object_ptr_kinds_eq_h2
      object_ptr_kinds_eq_h3 returns_result_select_result
       apply (metis h2  get_child_nodes (castcharacter_data_ptr2object_ptr new_character_data_ptr) r []
        is_OK_returns_result_I)

      apply (metis h  create_character_data document_ptr text r new_character_data_ptr assms(4)
        character_data_ptr_kinds_commutes h2 new_character_data_ptr new_character_data_ptr_in_heap
        node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h3 returns_result_eq)
    using h  create_character_data document_ptr text r new_character_data_ptr
      new_character_data_ptr  set |h  character_data_ptr_kinds_M|r assms(4) returns_result_eq
     apply fastforce
    using  assms(2) assms(3) children_eq_h local.get_child_nodes_ok local.get_child_nodes_ptr_in_heap
      local.known_ptrs_known_ptr object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 returns_result_select_result
    apply (smt ObjectMonad.ptr_kinds_ptr_kinds_M
        ‹castcharacter_data_ptr2object_ptr new_character_data_ptr  set |h  object_ptr_kinds_M|r
        h  create_character_data document_ptr text r new_character_data_ptr assms(1) assms(5)
        create_character_data_is_weakly_dom_component_safe_step local.get_dom_component_impl select_result_I2)
    done
qed

end

interpretation i_get_dom_component_create_character_data?: l_get_dom_component_create_character_dataCore_DOM
  known_ptr heap_is_wellformed parent_child_rel type_wf known_ptrs to_tree_order get_parent
  get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe
  is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
  get_element_by_id get_elements_by_class_name get_elements_by_tag_name set_val set_val_locs
  get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs
  create_character_data
  by(auto simp add: l_get_dom_component_create_character_dataCore_DOM_def instances)
declare l_get_dom_component_create_character_dataCore_DOM_axioms [instances]


subsubsection ‹create\_document›

lemma create_document_unsafe: "¬((h
    :: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap
  ) h' new_document_ptr. heap_is_wellformed h  type_wf h  known_ptrs h 
h  create_document r new_document_ptr  h  create_document h h' 
is_strongly_dom_component_safe {} {cast new_document_ptr} h h')"
proof -
  obtain h document_ptr where h: "Inr (document_ptr, h) = (Heap (fmempty)
      :: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap
    )  (do {
      document_ptr  create_document;
      return (document_ptr)
  })"
    by(code_simp, auto simp add: equal_eq List.member_def)+
  then obtain h' new_document_ptr where
    "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
    h': "h  create_document r new_document_ptr" and
    h': "h  create_document h h'" and
    "¬(is_strongly_dom_component_safe {} {cast new_document_ptr} h h')"
    by(code_simp, auto simp add: equal_eq List.member_def)+
  then show ?thesis
    by blast
qed

locale l_get_dom_component_create_documentCore_DOM =
  l_get_dom_componentCore_DOM heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
  to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component
  is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs
  get_ancestors get_ancestors_locs get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id
  get_elements_by_class_name get_elements_by_tag_name +
  l_create_documentCore_DOM create_document
  for known_ptr :: "(_::linorder) object_ptr  bool"
    and heap_is_wellformed :: "(_) heap  bool"
    and parent_child_rel :: "(_) heap  ((_) object_ptr × (_) object_ptr) set"
    and type_wf :: "(_) heap  bool"
    and known_ptrs :: "(_) heap  bool"
    and to_tree_order :: "(_) object_ptr  ((_) heap, exception, (_) object_ptr list) prog"
    and get_parent :: "(_) node_ptr  ((_) heap, exception, (_) object_ptr option) prog"
    and get_parent_locs :: "((_) heap  (_) heap  bool) set"
    and get_child_nodes :: "(_) object_ptr  ((_) heap, exception, (_) node_ptr list) prog"
    and get_child_nodes_locs :: "(_) object_ptr  ((_) heap  (_) heap  bool) set"
    and get_dom_component :: "(_) object_ptr  ((_) heap, exception, (_) object_ptr list) prog"
    and is_strongly_dom_component_safe ::
    "(_) object_ptr set  (_) object_ptr set  (_) heap  (_) heap  bool"
    and is_weakly_dom_component_safe ::
    "(_) object_ptr set  (_) object_ptr set  (_) heap  (_) heap  bool"
    and get_root_node :: "(_) object_ptr  ((_) heap, exception, (_) object_ptr) prog"
    and get_root_node_locs :: "((_) heap  (_) heap  bool) set"
    and get_ancestors :: "(_) object_ptr  ((_) heap, exception, (_) object_ptr list) prog"
    and get_ancestors_locs :: "((_) heap  (_) heap  bool) set"
    and get_element_by_id ::
    "(_) object_ptr  char list  ((_) heap, exception, (_) element_ptr option) prog"
    and get_elements_by_class_name ::
    "(_) object_ptr  char list  ((_) heap, exception, (_) element_ptr list) prog"
    and get_elements_by_tag_name ::
    "(_) object_ptr  char list  ((_) heap, exception, (_) element_ptr list) prog"
    and create_document :: "((_) heap, exception, (_) document_ptr) prog"
begin

lemma create_document_is_weakly_dom_component_safe_step:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  create_document h h'"
  assumes "ptr  cast |h  create_document|r"
  shows "preserved (get_MObject ptr getter) h h'"
  using assms
  apply(auto simp add: create_document_def)[1]
  by (metis assms(4) assms(5) is_OK_returns_heap_I local.create_document_def new_document_get_MObject
      select_result_I)

lemma create_document_is_weakly_dom_component_safe:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  create_document r result"
  assumes "h  create_document h h'"
  shows "is_weakly_dom_component_safe {} {cast result} h h'"
proof -
  have "object_ptr_kinds h' = object_ptr_kinds h |∪| {|cast result|}"
    using assms(4) assms(5) local.create_document_def new_document_new_ptr by auto
  moreover have "result |∉| document_ptr_kinds h"
    using assms(4) assms(5) local.create_document_def new_document_ptr_not_in_heap by auto
  ultimately show ?thesis
    using assms
    apply(auto simp add: is_weakly_dom_component_safe_def Let_def local.create_document_def
        new_document_ptr_not_in_heap)[1]
    by (metis result |∉| document_ptr_kinds h document_ptr_kinds_commutes new_document_get_MObject)
qed
end

interpretation i_get_dom_component_create_document?: l_get_dom_component_create_documentCore_DOM
  known_ptr heap_is_wellformed parent_child_rel type_wf known_ptrs to_tree_order get_parent
  get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe
  is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs
  get_element_by_id get_elements_by_class_name get_elements_by_tag_name create_document
  by(auto simp add: l_get_dom_component_create_documentCore_DOM_def instances)
declare l_get_dom_component_create_documentCore_DOM_axioms [instances]


subsubsection ‹insert\_before›

lemma insert_before_unsafe: "¬((h
    :: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap
  ) h' ptr child. heap_is_wellformed h  type_wf h  known_ptrs h 
h  insert_before ptr child None h h'  is_weakly_dom_component_safe {ptr, cast child} {} h h')"
proof -
  obtain h document_ptr e1 e2 where h: "Inr ((document_ptr, e1, e2), h) = (Heap (fmempty)
      :: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap
    )  (do {
      document_ptr  create_document;
      e1  create_element document_ptr ''div'';
      e2  create_element document_ptr ''div'';
      return (document_ptr, e1, e2)
  })"
    by(code_simp, auto simp add: equal_eq List.member_def)+
  then obtain h' where
    "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
    h': "h  insert_before (cast e1) (cast e2) None h h'" and
    "¬(is_weakly_dom_component_safe {cast e1, cast e2} {} h h')"
    by(code_simp, auto simp add: equal_eq List.member_def)+
  then show ?thesis
    by auto
qed

lemma insert_before_unsafe2: "¬((h
    :: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap
  ) h' ptr child ref. heap_is_wellformed h  type_wf h  known_ptrs h 
h  insert_before ptr child (Some ref) h h' 
is_weakly_dom_component_safe {ptr, cast child, cast ref} {} h h')"
proof -
  obtain h document_ptr e1 e2 e3 where h: "Inr ((document_ptr, e1, e2, e3), h) = (Heap (fmempty)
      :: ('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap
    )  (do {
      document_ptr  create_document;
      e1  create_element document_ptr ''div'';
      e2  create_element document_ptr ''div'';
      e3  create_element document_ptr ''div'';
      append_child (cast e1) (cast e2);
      return (document_ptr, e1, e2, e3)
  })"
    by(code_simp, auto simp add: equal_eq List.member_def)+
  then obtain h' where
    "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
    h': "h  insert_before (cast e1) (cast e3) (Some (cast e2)) h h'" and
    "¬(is_weakly_dom_component_safe {cast e1, cast e3, cast e2} {} h h')"
    apply(code_simp)
    apply(clarify)
    by(code_simp, auto simp add: equal_eq List.member_def)+
  then show ?thesis
    by fast
qed


lemma append_child_unsafe:
  obtains
    h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" and
    h' and ptr and child where
    "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
    "h  append_child ptr child h h'" and
    "¬ is_weakly_dom_component_safe {ptr, cast child} {} h h'"
proof -
  let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder},
'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder},
'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder},
'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap"
  let ?P = "do {
      document_ptr  create_document;
      e1  create_element document_ptr ''div'';
      e2  create_element document_ptr ''div'';
      return (e1, e2)
  }"
  let ?h1 = "|?h0  ?P|h"
  let ?e1 = "fst |?h0  ?P|r"
  let ?e2 = "snd |?h0  ?P|r"

  show thesis
    apply(rule that[where h="?h1" and ptr="castelement_ptr2object_ptr ?e1" and child="castelement_ptr2node_ptr ?e2"])
    by code_simp+
qed


subsubsection ‹get\_owner\_document›

lemma get_owner_document_unsafe:
  obtains
    h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder},
'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder},
'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder},
'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" and
    h' and ptr and owner_document where
    "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and
    "h  get_owner_document ptr r owner_document h h'" and
    "¬is_weakly_dom_component_safe {ptr} {cast owner_document} h h'"
proof -
  let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder},
'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder},
'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder},
'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap"
  let ?P = "do {
      document_ptr  create_document;
      e1  create_element document_ptr ''div'';
      return (document_ptr, e1)
  }"
  let ?h1 = "|?h0  ?P|h"
  let ?document_ptr = "fst |?h0  ?P|r"
  let ?e1 = "snd |?h0  ?P|r"

  show thesis
    apply(rule that[where h="?h1" and ptr="castelement_ptr2object_ptr ?e1" and owner_document="?document_ptr"])
    by code_simp+
qed

end

Theory Core_DOM_SC_DOM_Components

(***********************************************************************************
 * Copyright (c) 2016-2020 The University of Sheffield, UK
 *               2019-2020 University of Exeter, UK
 *
 * All rights reserved.
 *
 * Redistribution and use in source and binary forms, with or without
 * modification, are permitted provided that the following conditions are met:
 *
 * * Redistributions of source code must retain the above copyright notice, this
 *   list of conditions and the following disclaimer.
 *
 * * Redistributions in binary form must reproduce the above copyright notice,
 *   this list of conditions and the following disclaimer in the documentation
 *   and/or other materials provided with the distribution.
 *
 * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
 * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
 * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
 * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
 * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
 * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
 * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
 * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
 * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
 * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
 *
 * SPDX-License-Identifier: BSD-2-Clause
 ***********************************************************************************)

section ‹Core SC DOM Components II›
theory Core_DOM_SC_DOM_Components
  imports
    Core_DOM_DOM_Components
begin
declare [[smt_timeout=2400]]

section ‹Scope Components›

subsection ‹Definition›

locale l_get_scdom_componentCore_DOM_defs =
  l_get_disconnected_nodes_defs get_disconnected_nodes get_disconnected_nodes_locs +
  l_get_owner_document_defs get_owner_document +
  l_to_tree_order_defs to_tree_order
  for get_owner_document :: "(_::linorder) object_ptr  ((_) heap, exception, (_) document_ptr) prog"
    and get_disconnected_nodes :: "(_) document_ptr  ((_) heap, exception, (_) node_ptr list) prog"
    and get_disconnected_nodes_locs :: "(_) document_ptr  ((_) heap  (_) heap  bool) set"
    and to_tree_order :: "(_) object_ptr  ((_) heap, exception, (_) object_ptr list) prog"
begin
definition a_get_scdom_component :: "(_) object_ptr  (_, (_) object_ptr list) dom_prog"
  where
    "a_get_scdom_component ptr = do {
      document  get_owner_document ptr;
      disc_nodes  get_disconnected_nodes document;
      tree_order  to_tree_order (cast document);
      disconnected_tree_orders  map_M (to_tree_order  cast) disc_nodes;
      return (tree_order @ (concat disconnected_tree_orders))
    }"

definition a_is_strongly_scdom_component_safe ::
  "(_) object_ptr set  (_) object_ptr set  (_) heap  (_) heap  bool"
  where
    "a_is_strongly_scdom_component_safe Sarg Sresult h h' = (
      let removed_pointers = fset (object_ptr_kinds h) - fset (object_ptr_kinds h') in
      let added_pointers = fset (object_ptr_kinds h') - fset (object_ptr_kinds h) in
      let arg_components =
        (ptr  (ptr  Sarg. set |h  a_get_scdom_component ptr|r) 
        fset (object_ptr_kinds h).  set |h  a_get_scdom_component ptr|r) in
      let arg_components' =
        (ptr  (ptr  Sarg. set |h  a_get_scdom_component ptr|r) 
        fset (object_ptr_kinds h').  set |h'  a_get_scdom_component ptr|r) in
      removed_pointers  arg_components 
      added_pointers  arg_components' 
      Sresult  arg_components' 
      (outside_ptr  fset (object_ptr_kinds h)  fset (object_ptr_kinds h') -
        (ptr  Sarg. set |h   a_get_scdom_component ptr|r). preserved (get_M outside_ptr id) h h'))"

definition a_is_weakly_scdom_component_safe ::
  "(_) object_ptr set  (_) object_ptr set  (_) heap  (_) heap  bool"
  where
    "a_is_weakly_scdom_component_safe Sarg Sresult h h' = (
      let removed_pointers = fset (object_ptr_kinds h) - fset (object_ptr_kinds h') in
      let added_pointers = fset (object_ptr_kinds h') - fset (object_ptr_kinds h) in
      let arg_components =
        (ptr  (ptr  Sarg. set |h  a_get_scdom_component ptr|r) 
        fset (object_ptr_kinds h).  set |h  a_get_scdom_component ptr|r) in
      let arg_components' =
        (ptr  (ptr  Sarg. set |h  a_get_scdom_component ptr|r) 
        fset (object_ptr_kinds h').  set |h'  a_get_scdom_component ptr|r) in
      removed_pointers  arg_components 
      Sresult  arg_components'  added_pointers  
      (outside_ptr  fset (object_ptr_kinds h)  fset (object_ptr_kinds h') -
        (ptr  Sarg. set |h   a_get_scdom_component ptr|r). preserved (get_M outside_ptr id) h h'))"
end

global_interpretation l_get_scdom_componentCore_DOM_defs get_owner_document get_disconnected_nodes
  get_disconnected_nodes_locs to_tree_order
  defines get_scdom_component = "l_get_scdom_componentCore_DOM_defs.a_get_scdom_component
get_owner_document get_disconnected_nodes to_tree_order"
    and is_strongly_scdom_component_safe = a_is_strongly_scdom_component_safe
    and is_weakly_scdom_component_safe = a_is_weakly_scdom_component_safe
  .

locale l_get_scdom_component_defs =
  fixes get_scdom_component :: "(_) object_ptr  (_, (_) object_ptr list) dom_prog"
  fixes is_strongly_scdom_component_safe ::
    "(_) object_ptr set  (_) object_ptr set  (_) heap  (_) heap  bool"
  fixes is_weakly_scdom_component_safe ::
    "(_) object_ptr set  (_) object_ptr set  (_) heap  (_) heap  bool"


locale l_get_scdom_componentCore_DOM =
  l_get_scdom_component_defs +
  l_get_scdom_componentCore_DOM_defs +
  assumes get_scdom_component_impl: "get_scdom_component = a_get_scdom_component"
  assumes is_strongly_scdom_component_safe_impl:
    "is_strongly_scdom_component_safe = a_is_strongly_scdom_component_safe"
  assumes is_weakly_scdom_component_safe_impl:
    "is_weakly_scdom_component_safe = a_is_weakly_scdom_component_safe"
begin
lemmas get_scdom_component_def = a_get_scdom_component_def[folded get_scdom_component_impl]
lemmas is_strongly_scdom_component_safe_def =
  a_is_strongly_scdom_component_safe_def[folded is_strongly_scdom_component_safe_impl]
lemmas is_weakly_scdom_component_safe_def =
  a_is_weakly_scdom_component_safe_def[folded is_weakly_scdom_component_safe_impl]
end

interpretation i_get_scdom_component?: l_get_scdom_componentCore_DOM
  get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe
  get_owner_document get_disconnected_nodes get_disconnected_nodes_locs to_tree_order
  by(auto simp add: l_get_scdom_componentCore_DOM_def get_scdom_component_def
      is_strongly_scdom_component_safe_def is_weakly_scdom_component_safe_def instances)
declare l_get_scdom_componentCore_DOM_axioms [instances]


locale l_get_dom_component_get_scdom_componentCore_DOM =
  l_get_scdom_componentCore_DOM +
  l_get_dom_componentCore_DOM +
  l_heap_is_wellformed +
  l_get_owner_document +
  l_get_owner_document_wf +
  l_get_disconnected_nodes +
  l_to_tree_order +
  l_known_ptr +
  l_known_ptrs +
  l_get_owner_document_wf_get_root_node_wf +
  assumes known_ptr_impl: "known_ptr = DocumentClass.known_ptr"
begin

lemma known_ptr_node_or_document: "known_ptr ptr  is_node_ptr_kind ptr  is_document_ptr_kind ptr"
  by(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs
      CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs
      split: option.splits)

lemma get_scdom_component_ptr_in_heap2:
  assumes "h  ok (get_scdom_component ptr)"
  shows "ptr |∈| object_ptr_kinds h"
  using assms get_root_node_ptr_in_heap
  apply(auto simp add: get_scdom_component_def elim!: bind_is_OK_E3 intro!: map_M_pure_I)[1]
  by (simp add: is_OK_returns_result_I local.get_owner_document_ptr_in_heap)

lemma get_scdom_component_subset_get_dom_component:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_scdom_component ptr r sc"
  assumes "h  get_dom_component ptr r c"
  shows "set c  set sc"
proof -
  obtain document disc_nodes tree_order disconnected_tree_orders where
    document: "h  get_owner_document ptr r document"
    and disc_nodes: "h  get_disconnected_nodes document r disc_nodes"
    and tree_order: "h  to_tree_order (cast document) r tree_order"
    and disconnected_tree_orders: "h  map_M (to_tree_order  cast) disc_nodes r disconnected_tree_orders"
    and sc: "sc = tree_order @ (concat disconnected_tree_orders)"
    using assms(4)
    by(auto simp add: get_scdom_component_def elim!: bind_returns_result_E
        elim!: bind_returns_result_E2[rotated, OF get_owner_document_pure, rotated]
        elim!: bind_returns_result_E2[rotated, OF get_disconnected_nodes_pure, rotated]
        elim!: bind_returns_result_E2[rotated, OF to_tree_order_pure, rotated]
        )

  obtain root_ptr where root_ptr: "h  get_root_node ptr r root_ptr"
    and c: "h  to_tree_order root_ptr r c"
    using assms(5)
    by(auto simp add: get_dom_component_def elim!: bind_returns_result_E2[rotated, OF get_root_node_pure, rotated])

  show ?thesis
  proof (cases "is_document_ptr_kind root_ptr")
    case True
    then have "cast document = root_ptr"
      using get_root_node_document assms(1) assms(2) assms(3) root_ptr document
      by (metis document_ptr_casts_commute3 returns_result_eq)
    then have "c = tree_order"
      using tree_order c
      by auto
    then show ?thesis
      by(simp add: sc)
  next
    case False
    moreover have "root_ptr |∈| object_ptr_kinds h"
      using assms(1) assms(2) assms(3) local.get_root_node_root_in_heap root_ptr by blast
    ultimately have "is_node_ptr_kind root_ptr"
      using assms(3) known_ptrs_known_ptr known_ptr_node_or_document
      by auto
    then obtain root_node_ptr where root_node_ptr: "root_ptr = castnode_ptr2object_ptr root_node_ptr"
      by (metis node_ptr_casts_commute3)
    then have "h  get_owner_document root_ptr r document"
      using get_root_node_same_owner_document
      using assms(1) assms(2) assms(3) document root_ptr by blast
    then have "root_node_ptr  set disc_nodes"
      using assms(1) assms(2) assms(3) disc_nodes in_disconnected_nodes_no_parent root_node_ptr
      using local.get_root_node_same_no_parent root_ptr by blast
    then have "c  set disconnected_tree_orders"
      using c root_node_ptr
      using map_M_pure_E[OF disconnected_tree_orders]
      by (metis (mono_tags, lifting) comp_apply local.to_tree_order_pure select_result_I2)
    then show ?thesis
      by(auto simp add: sc)
  qed
qed

lemma get_scdom_component_ptrs_same_owner_document:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_scdom_component ptr r sc"
  assumes "ptr'  set sc"
  assumes "h  get_owner_document ptr r owner_document"
  shows "h  get_owner_document ptr' r owner_document"
proof -
  obtain document disc_nodes tree_order disconnected_tree_orders where
    document: "h  get_owner_document ptr r document"
    and disc_nodes: "h  get_disconnected_nodes document r disc_nodes"
    and tree_order: "h  to_tree_order (cast document) r tree_order"
    and disconnected_tree_orders: "h  map_M (to_tree_order  cast) disc_nodes r disconnected_tree_orders"
    and sc: "sc = tree_order @ (concat disconnected_tree_orders)"
    using assms(4)
    by(auto simp add: get_scdom_component_def elim!: bind_returns_result_E
        elim!: bind_returns_result_E2[rotated, OF get_owner_document_pure, rotated]
        elim!: bind_returns_result_E2[rotated, OF get_disconnected_nodes_pure, rotated]
        elim!: bind_returns_result_E2[rotated, OF to_tree_order_pure, rotated]
        )
  show ?thesis
  proof (cases "ptr'  set tree_order")
    case True
    have "owner_document = document"
      using assms(6) document by fastforce
    then show ?thesis
      by (metis (no_types) True assms(1) assms(2) assms(3) castdocument_ptr2object_ptr_inject document
          document_ptr_casts_commute3 document_ptr_document_ptr_cast document_ptr_kinds_commutes
          local.get_owner_document_owner_document_in_heap local.get_root_node_document
          local.get_root_node_not_node_same local.to_tree_order_same_root node_ptr_no_document_ptr_cast tree_order)
  next
    case False
    then obtain disconnected_tree_order where disconnected_tree_order:
      "ptr'  set disconnected_tree_order" and "disconnected_tree_order  set disconnected_tree_orders"
      using sc ptr'  set sc
      by auto
    obtain root_ptr' where
      root_ptr': "root_ptr'  set disc_nodes" and
      "h  to_tree_order (cast root_ptr') r disconnected_tree_order"
      using map_M_pure_E2[OF disconnected_tree_orders disconnected_tree_order  set disconnected_tree_orders]
      by (metis comp_apply local.to_tree_order_pure)
    have "¬(parent  fset (object_ptr_kinds h). root_ptr'  set |h  get_child_nodes parent|r)"
      using disc_nodes
      by (meson assms(1) assms(2) assms(3) disjoint_iff_not_equal local.get_child_nodes_ok
          local.heap_is_wellformed_children_disc_nodes_different local.known_ptrs_known_ptr notin_fset
          returns_result_select_result root_ptr')
    then
    have "h  get_parent root_ptr' r None"
      using disc_nodes
      by (metis (no_types, lifting) assms(1) assms(2) assms(3) fmember.rep_eq local.get_parent_child_dual
          local.get_parent_ok local.get_parent_parent_in_heap local.heap_is_wellformed_disc_nodes_in_heap
          returns_result_select_result root_ptr' select_result_I2 split_option_ex)
    then have "h  get_root_node ptr' r cast root_ptr'"
      using h  to_tree_order (castnode_ptr2object_ptr root_ptr') r disconnected_tree_order assms(1)
        assms(2) assms(3) disconnected_tree_order local.get_root_node_no_parent
        local.to_tree_order_get_root_node local.to_tree_order_ptr_in_result
      by blast
    then have "h  get_owner_document (cast root_ptr') r document"
      using assms(1) assms(2) assms(3) disc_nodes local.get_owner_document_disconnected_nodes root_ptr'
      by blast

    then have "h  get_owner_document ptr' r document"
      using h  get_root_node ptr' r castnode_ptr2object_ptr root_ptr' assms(1) assms(2) assms(3)
        local.get_root_node_same_owner_document
      by blast
    then show ?thesis
      using assms(6) document returns_result_eq by force
  qed
qed

lemma get_scdom_component_ptrs_same_scope_component:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_scdom_component ptr r sc"
  assumes "ptr'  set sc"
  shows "h  get_scdom_component ptr' r sc"
proof -
  obtain document disc_nodes tree_order disconnected_tree_orders where
    document: "h  get_owner_document ptr r document"
    and disc_nodes: "h  get_disconnected_nodes document r disc_nodes"
    and tree_order: "h  to_tree_order (cast document) r tree_order"
    and disconnected_tree_orders: "h  map_M (to_tree_order  cast) disc_nodes r disconnected_tree_orders"
    and sc: "sc = tree_order @ (concat disconnected_tree_orders)"
    using assms(4)
    by(auto simp add: get_scdom_component_def elim!: bind_returns_result_E
        elim!: bind_returns_result_E2[rotated, OF get_owner_document_pure, rotated]
        elim!: bind_returns_result_E2[rotated, OF get_disconnected_nodes_pure, rotated]
        elim!: bind_returns_result_E2[rotated, OF to_tree_order_pure, rotated]
        )
  show ?thesis
  proof (cases "ptr'  set tree_order")
    case True
    then have "h  get_owner_document ptr' r document"
      by (metis assms(1) assms(2) assms(3) castdocument_ptr2object_ptr_inject document
          document_ptr_casts_commute3 document_ptr_kinds_commutes known_ptr_node_or_document
          local.get_owner_document_owner_document_in_heap local.get_root_node_document
          local.get_root_node_not_node_same local.known_ptrs_known_ptr local.to_tree_order_get_root_node
          local.to_tree_order_ptr_in_result node_ptr_no_document_ptr_cast tree_order)
    then show ?thesis
      using disc_nodes tree_order disconnected_tree_orders sc
      by(auto simp add: get_scdom_component_def intro!: bind_pure_returns_result_I map_M_pure_I)
  next
    case False
    then obtain disconnected_tree_order where disconnected_tree_order:
      "ptr'  set disconnected_tree_order" and "disconnected_tree_order  set disconnected_tree_orders"
      using sc ptr'  set sc
      by auto
    obtain root_ptr' where
      root_ptr': "root_ptr'  set disc_nodes" and
      "h  to_tree_order (cast root_ptr') r disconnected_tree_order"
      using map_M_pure_E2[OF disconnected_tree_orders disconnected_tree_order  set disconnected_tree_orders]
      by (metis comp_apply local.to_tree_order_pure)
    have "¬(parent  fset (object_ptr_kinds h). root_ptr'  set |h  get_child_nodes parent|r)"
      using disc_nodes
      by (meson assms(1) assms(2) assms(3) disjoint_iff_not_equal local.get_child_nodes_ok
          local.heap_is_wellformed_children_disc_nodes_different local.known_ptrs_known_ptr notin_fset
          returns_result_select_result root_ptr')
    then
    have "h  get_parent root_ptr' r None"
      using disc_nodes
      by (metis (no_types, lifting) assms(1) assms(2) assms(3) fmember.rep_eq
          local.get_parent_child_dual local.get_parent_ok local.get_parent_parent_in_heap
          local.heap_is_wellformed_disc_nodes_in_heap returns_result_select_result root_ptr'
          select_result_I2 split_option_ex)
    then have "h  get_root_node ptr' r cast root_ptr'"
      using h  to_tree_order (castnode_ptr2object_ptr root_ptr') r disconnected_tree_order assms(1)
        assms(2) assms(3) disconnected_tree_order local.get_root_node_no_parent
        local.to_tree_order_get_root_node local.to_tree_order_ptr_in_result
      by blast
    then have "h  get_owner_document (cast root_ptr') r document"
      using assms(1) assms(2) assms(3) disc_nodes local.get_owner_document_disconnected_nodes root_ptr'
      by blast

    then have "h  get_owner_document ptr' r document"
      using h  get_root_node ptr' r castnode_ptr2object_ptr root_ptr' assms(1) assms(2) assms(3)
        local.get_root_node_same_owner_document
      by blast
    then show ?thesis
      using disc_nodes tree_order disconnected_tree_orders sc
      by(auto simp add: get_scdom_component_def intro!: bind_pure_returns_result_I map_M_pure_I)
  qed
qed

lemma get_scdom_component_ok:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "ptr |∈| object_ptr_kinds h"
  shows "h  ok (get_scdom_component ptr)"
  using assms
  apply(auto simp add: get_scdom_component_def intro!: bind_is_OK_pure_I map_M_pure_I map_M_ok_I)[1]
  using get_owner_document_ok
     apply blast
    apply (simp add: local.get_disconnected_nodes_ok local.get_owner_document_owner_document_in_heap)
   apply (simp add: local.get_owner_document_owner_document_in_heap local.to_tree_order_ok)
  using local.heap_is_wellformed_disc_nodes_in_heap local.to_tree_order_ok node_ptr_kinds_commutes
  by blast

lemma get_scdom_component_ptr_in_heap:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_scdom_component ptr r sc"
  assumes "ptr'  set sc"
  shows "ptr' |∈| object_ptr_kinds h"
  apply(insert assms )
  apply(auto simp add: get_scdom_component_def elim!: bind_returns_result_E2 intro!: map_M_pure_I)[1]
  using local.to_tree_order_ptrs_in_heap apply blast
  by (metis (no_types, lifting) assms(4) assms(5) bind_returns_result_E
      get_scdom_component_ptrs_same_scope_component is_OK_returns_result_I get_scdom_component_def
      local.get_owner_document_ptr_in_heap)

lemma get_scdom_component_contains_get_dom_component:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_scdom_component ptr r sc"
  assumes "ptr'  set sc"
  obtains c where "h  get_dom_component ptr' r c" and "set c  set sc"
proof -
  have "h  get_scdom_component ptr' r sc"
    using assms(1) assms(2) assms(3) assms(4) assms(5) get_scdom_component_ptrs_same_scope_component
    by blast
  then show ?thesis
    by (meson assms(1) assms(2) assms(3) assms(5) get_scdom_component_ptr_in_heap
        get_scdom_component_subset_get_dom_component is_OK_returns_result_E local.get_dom_component_ok that)
qed

lemma get_scdom_component_owner_document_same:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_scdom_component ptr r sc"
  assumes "ptr'  set sc"
  obtains owner_document where "h  get_owner_document ptr' r owner_document" and "cast owner_document  set sc"
  using assms
  apply(auto simp add: get_scdom_component_def elim!: bind_returns_result_E2 intro!: map_M_pure_I)[1]
   apply (metis (no_types, lifting) assms(4) assms(5) document_ptr_casts_commute3
      document_ptr_document_ptr_cast get_scdom_component_contains_get_dom_component
      local.get_dom_component_ptr local.get_dom_component_root_node_same local.get_dom_component_to_tree_order
      local.get_root_node_document local.get_root_node_not_node_same local.to_tree_order_ptr_in_result
      local.to_tree_order_ptrs_in_heap node_ptr_no_document_ptr_cast)
  apply(rule map_M_pure_E2)
     apply(simp)
    apply(simp)
   apply(simp)
  by (smt assms(4) assms(5) comp_apply get_scdom_component_ptr_in_heap is_OK_returns_result_E
      local.get_owner_document_disconnected_nodes local.get_root_node_ok local.get_root_node_same_owner_document
      local.to_tree_order_get_root_node local.to_tree_order_ptr_in_result)

lemma get_scdom_component_different_owner_documents:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_owner_document ptr r owner_document"
  assumes "h  get_owner_document ptr' r owner_document'"
  assumes "owner_document  owner_document'"
  shows "set |h  get_scdom_component ptr|r  set |h  get_scdom_component ptr'|r = {}"
  using assms get_scdom_component_ptrs_same_owner_document
  by (smt disjoint_iff_not_equal get_scdom_component_ok is_OK_returns_result_I
      local.get_owner_document_ptr_in_heap returns_result_eq returns_result_select_result)


lemma get_scdom_component_ptr:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_scdom_component ptr r c"
  shows "ptr  set c"
  using assms
  by (meson get_scdom_component_ptr_in_heap2 get_scdom_component_subset_get_dom_component
      is_OK_returns_result_E is_OK_returns_result_I local.get_dom_component_ok local.get_dom_component_ptr
      subsetD)
end

locale l_get_dom_component_get_scdom_component = l_get_owner_document_defs + l_heap_is_wellformed_defs +
  l_type_wf + l_known_ptrs + l_get_scdom_component_defs + l_get_dom_component_defs +
  assumes get_scdom_component_subset_get_dom_component:
    "heap_is_wellformed h  type_wf h  known_ptrs h  h  get_scdom_component ptr r sc 
h  get_dom_component ptr r c  set c  set sc"
  assumes get_scdom_component_ptrs_same_scope_component:
    "heap_is_wellformed h  type_wf h  known_ptrs h  h  get_scdom_component ptr r sc 
ptr'  set sc  h  get_scdom_component ptr' r sc"
  assumes get_scdom_component_ptrs_same_owner_document:
    "heap_is_wellformed h  type_wf h  known_ptrs h  h  get_scdom_component ptr r sc 
ptr'  set sc  h  get_owner_document ptr r owner_document  h  get_owner_document ptr' r owner_document"
  assumes get_scdom_component_ok:
    "heap_is_wellformed h  type_wf h  known_ptrs h  ptr |∈| object_ptr_kinds h 
h  ok (get_scdom_component ptr)"
  assumes get_scdom_component_ptr_in_heap:
    "heap_is_wellformed h  type_wf h  known_ptrs h  h  get_scdom_component ptr r sc 
ptr'  set sc  ptr' |∈| object_ptr_kinds h"
  assumes get_scdom_component_contains_get_dom_component:
    "(c. h  get_dom_component ptr' r c  set c  set sc  thesis)  heap_is_wellformed h 
type_wf h  known_ptrs h  h  get_scdom_component ptr r sc  ptr'  set sc  thesis"
  assumes get_scdom_component_owner_document_same:
    "(owner_document. h  get_owner_document ptr' r owner_document  cast owner_document  set sc  thesis) 
    heap_is_wellformed h  type_wf h  known_ptrs h  h  get_scdom_component ptr r sc 
ptr'  set sc  thesis"
  assumes get_scdom_component_different_owner_documents:
    "heap_is_wellformed h  type_wf h  known_ptrs h  h  get_owner_document ptr r owner_document 
h  get_owner_document ptr' r owner_document'  owner_document  owner_document' 
set |h  get_scdom_component ptr|r  set |h  get_scdom_component ptr'|r = {}"

interpretation i_get_dom_component_get_scdom_component?: l_get_dom_component_get_scdom_componentCore_DOM
  get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_owner_document
  get_disconnected_nodes get_disconnected_nodes_locs to_tree_order heap_is_wellformed parent_child_rel
  type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component
  is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors
  get_ancestors_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name
  by(auto simp add: l_get_dom_component_get_scdom_componentCore_DOM_def l_get_dom_component_get_scdom_componentCore_DOM_axioms_def instances)
declare l_get_dom_component_get_scdom_componentCore_DOM_axioms [instances]

lemma get_dom_component_get_scdom_component_is_l_get_dom_component_get_scdom_component [instances]:
  "l_get_dom_component_get_scdom_component get_owner_document heap_is_wellformed type_wf known_ptr
known_ptrs get_scdom_component get_dom_component"
  apply(auto simp add: l_get_dom_component_get_scdom_component_def l_get_dom_component_get_scdom_component_axioms_def instances)[1]
  using get_scdom_component_subset_get_dom_component apply fast
  using get_scdom_component_ptrs_same_scope_component apply fast
  using get_scdom_component_ptrs_same_owner_document apply fast
  using get_scdom_component_ok apply fast
  using get_scdom_component_ptr_in_heap apply fast
  using get_scdom_component_contains_get_dom_component apply blast
  using get_scdom_component_owner_document_same apply blast
  using get_scdom_component_different_owner_documents apply fast
  done


subsubsection ‹get\_child\_nodes›

locale l_get_scdom_component_get_child_nodesCore_DOM =
  l_get_dom_component_get_scdom_component +
  l_get_scdom_componentCore_DOM +
  l_get_dom_component_get_child_nodesCore_DOM
begin
lemma get_child_nodes_is_strongly_scdom_component_safe_step:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_scdom_component ptr r sc"
  assumes "h  get_child_nodes ptr' r children"
  assumes "child  set children"
  shows "cast child  set sc  ptr'  set sc"
  apply(auto)[1]
   apply (meson assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) contra_subsetD
      get_scdom_component_ptrs_same_scope_component get_scdom_component_subset_get_dom_component
      is_OK_returns_result_E local.get_child_nodes_is_strongly_dom_component_safe local.get_dom_component_ok
      local.get_dom_component_ptr local.heap_is_wellformed_children_in_heap node_ptr_kinds_commutes)
  by (meson assms(1) assms(2) assms(3) assms(4) assms(5) assms(6)
      get_scdom_component_contains_get_dom_component is_OK_returns_result_E is_OK_returns_result_I
      get_child_nodes_is_strongly_dom_component_safe local.get_child_nodes_ptr_in_heap
      local.get_dom_component_ok local.get_dom_component_ptr set_rev_mp)

lemma get_child_nodes_is_strongly_scdom_component_safe:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_child_nodes ptr r children"
  assumes "h  get_child_nodes ptr h h'"
  shows "is_strongly_scdom_component_safe {ptr} (cast ` set children) h h'"
proof -
  have "h = h'"
    using assms(5)
    by (meson local.get_child_nodes_pure pure_returns_heap_eq)
  then show ?thesis
    using assms
    apply(auto simp add: is_strongly_scdom_component_safe_def Let_def preserved_def)[1]
    by (smt Int_absorb2 finite_set_in get_child_nodes_is_strongly_scdom_component_safe_step in_mono
        is_OK_returns_result_I local.get_child_nodes_ptr_in_heap local.get_dom_component_ok
        local.get_dom_component_ptr local.get_scdom_component_impl local.get_scdom_component_ok
        local.get_scdom_component_ptr_in_heap local.get_scdom_component_subset_get_dom_component
        returns_result_select_result subsetI)
qed
end

interpretation i_get_scdom_component_get_child_nodes?: l_get_scdom_component_get_child_nodesCore_DOM
  get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
  get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe
  get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_disconnected_nodes
  get_disconnected_nodes_locs to_tree_order get_parent get_parent_locs get_child_nodes
  get_child_nodes_locs get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_element_by_id
  get_elements_by_class_name get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs
  by(auto simp add: l_get_scdom_component_get_child_nodesCore_DOM_def instances)
declare l_get_scdom_component_get_child_nodesCore_DOM_axioms [instances]


subsubsection ‹get\_parent›

locale l_get_scdom_component_get_parentCore_DOM =
  l_get_dom_component_get_scdom_component +
  l_get_scdom_componentCore_DOM +
  l_get_dom_component_get_parentCore_DOM
begin

lemma get_parent_is_strongly_scdom_component_safe_step:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_scdom_component ptr r sc"
  assumes "h  get_parent ptr' r Some parent"
  shows "parent  set sc  cast ptr'  set sc"
  by (meson assms(1) assms(2) assms(3) assms(4) assms(5) contra_subsetD
      get_scdom_component_contains_get_dom_component local.get_dom_component_ptr
      local.get_parent_is_strongly_dom_component_safe_step)

lemma get_parent_is_strongly_scdom_component_safe:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_parent node_ptr r Some parent"
  assumes "h  get_parent node_ptr h h'"
  shows "is_strongly_scdom_component_safe {cast node_ptr} {parent} h h'"
proof -
  have "h = h'"
    using assms(5)
    by (meson local.get_parent_pure pure_returns_heap_eq)
  then show ?thesis
    using assms
    apply(auto simp add: is_strongly_scdom_component_safe_def Let_def preserved_def)[1]
    by (smt IntI finite_set_in in_mono local.get_dom_component_ok local.get_dom_component_ptr
        local.get_parent_is_strongly_dom_component_safe_step local.get_parent_parent_in_heap
        local.get_scdom_component_impl local.get_scdom_component_ok local.get_scdom_component_subset_get_dom_component
        local.to_tree_order_ok local.to_tree_order_parent local.to_tree_order_ptr_in_result
        local.to_tree_order_ptrs_in_heap returns_result_select_result)
qed
end

interpretation i_get_scdom_component_get_parent?: l_get_scdom_component_get_parentCore_DOM
  get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
  get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component
  is_strongly_dom_component_safe is_weakly_dom_component_safe get_disconnected_nodes
  get_disconnected_nodes_locs to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs
  get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_element_by_id
  get_elements_by_class_name get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs
  by(auto simp add: l_get_scdom_component_get_parentCore_DOM_def instances)
declare l_get_scdom_component_get_parentCore_DOM_axioms [instances]


subsubsection ‹get\_root\_node›

locale l_get_scdom_component_get_root_nodeCore_DOM =
  l_get_dom_component_get_scdom_component +
  l_get_scdom_componentCore_DOM +
  l_get_dom_component_get_root_nodeCore_DOM
begin

lemma get_root_node_is_strongly_scdom_component_safe_step:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_scdom_component ptr r sc"
  assumes "h  get_root_node ptr' r root"
  shows "root  set sc  ptr'  set sc"
  by (meson assms(1) assms(2) assms(3) assms(4) assms(5) contra_subsetD
      get_scdom_component_contains_get_dom_component local.get_dom_component_ptr
      local.get_root_node_is_strongly_dom_component_safe_step)

lemma get_root_node_is_strongly_scdom_component_safe:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_root_node ptr r root"
  assumes "h  get_root_node ptr h h'"
  shows "is_strongly_scdom_component_safe {ptr} {root} h h'"
proof -
  have "h = h'"
    using assms(5)
    by (meson local.get_root_node_pure pure_returns_heap_eq)
  then show ?thesis
    using assms
    apply(auto simp add: is_strongly_scdom_component_safe_def Let_def preserved_def)[1]
    by (smt Int_iff finite_set_in is_OK_returns_result_I local.get_dom_component_ok
        local.get_dom_component_ptr local.get_root_node_is_strongly_dom_component_safe_step
        local.get_root_node_ptr_in_heap local.get_scdom_component_impl local.get_scdom_component_ok
        local.get_scdom_component_subset_get_dom_component returns_result_select_result subset_eq)
qed
end

interpretation i_get_scdom_component_get_root_node?: l_get_scdom_component_get_root_nodeCore_DOM
  get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
  get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component
  is_strongly_dom_component_safe is_weakly_dom_component_safe get_disconnected_nodes
  get_disconnected_nodes_locs to_tree_order get_parent get_parent_locs get_child_nodes
  get_child_nodes_locs get_root_node get_root_node_locs get_ancestors get_ancestors_locs
  get_element_by_id get_elements_by_class_name get_elements_by_tag_name first_in_tree_order
  get_attribute get_attribute_locs
  by(auto simp add: l_get_scdom_component_get_root_nodeCore_DOM_def instances)
declare l_get_scdom_component_get_root_nodeCore_DOM_axioms [instances]


subsubsection ‹get\_element\_by\_id›

locale l_get_scdom_component_get_element_by_idCore_DOM =
  l_get_dom_component_get_scdom_component +
  l_get_scdom_componentCore_DOM +
  l_get_dom_component_get_element_by_idCore_DOM
begin

lemma get_element_by_id_is_strongly_scdom_component_safe_step:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_scdom_component ptr r sc"
  assumes "h  get_element_by_id ptr' idd r Some result"
  shows "cast result  set sc  ptr'  set sc"
  by (meson assms(1) assms(2) assms(3) assms(4) assms(5) contra_subsetD
      get_element_by_id_is_strongly_dom_component_safe_step get_scdom_component_contains_get_dom_component
      local.get_dom_component_ptr)

lemma get_element_by_id_is_strongly_scdom_component_safe:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_element_by_id ptr idd r Some result"
  assumes "h  get_element_by_id ptr idd h h'"
  shows "is_strongly_scdom_component_safe {ptr} {cast result} h h'"
proof -
  have "h = h'"
    using assms(5)
    by(auto simp add: preserved_def get_element_by_id_def first_in_tree_order_def
        elim!: bind_returns_heap_E2 intro!: map_filter_M_pure bind_pure_I
        split: option.splits list.splits)
  have "ptr |∈| object_ptr_kinds h"
    using assms(4)
    apply(auto simp add: get_element_by_id_def)[1]
    by (metis (no_types, lifting) assms(1) assms(2) assms(3) bind_is_OK_E is_OK_returns_result_I
        local.first_in_tree_order_def local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap)
  obtain to where to: "h  to_tree_order ptr r to"
    by (meson ptr |∈| object_ptr_kinds h assms(1) assms(2) assms(3) is_OK_returns_result_E
        local.to_tree_order_ok)
  then have "cast result  set to"
    using assms(4) local.get_element_by_id_result_in_tree_order by auto
  obtain c where c: "h  a_get_scdom_component ptr r c"
    using ptr |∈| object_ptr_kinds h assms(1) assms(2) assms(3) local.get_scdom_component_impl
      local.get_scdom_component_ok
    by blast

  then show ?thesis
    using assms h = h'
    apply(auto simp add: is_strongly_scdom_component_safe_def Let_def preserved_def
        get_element_by_id_def first_in_tree_order_def elim!: bind_returns_result_E2
        intro!: map_filter_M_pure bind_pure_I split: option.splits list.splits)[1]
    by (smt IntI ptr |∈| object_ptr_kinds h assms(4) finite_set_in
        get_element_by_id_is_strongly_scdom_component_safe_step local.get_dom_component_ok
        local.get_dom_component_ptr local.get_scdom_component_impl
        local.get_scdom_component_subset_get_dom_component returns_result_select_result select_result_I2
        subsetD)
qed
end

interpretation i_get_scdom_component_get_element_by_id?: l_get_scdom_component_get_element_by_idCore_DOM
  get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
  get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component
  is_strongly_dom_component_safe is_weakly_dom_component_safe get_disconnected_nodes
  get_disconnected_nodes_locs to_tree_order get_parent get_parent_locs get_child_nodes
  get_child_nodes_locs get_root_node get_root_node_locs get_ancestors get_ancestors_locs
  get_element_by_id get_elements_by_class_name get_elements_by_tag_name first_in_tree_order
  get_attribute get_attribute_locs
  by(auto simp add: l_get_scdom_component_get_element_by_idCore_DOM_def instances)
declare l_get_scdom_component_get_element_by_idCore_DOM_axioms [instances]


subsubsection ‹get\_elements\_by\_class\_name›

locale l_get_scdom_component_get_elements_by_class_nameCore_DOM =
  l_get_dom_component_get_scdom_component +
  l_get_scdom_componentCore_DOM +
  l_get_dom_component_get_elements_by_class_nameCore_DOM
begin

lemma get_elements_by_class_name_is_strongly_scdom_component_safe_step:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_scdom_component ptr r sc"
  assumes "h  get_elements_by_class_name ptr' idd r results"
  assumes "result  set results"
  shows "cast result  set sc  ptr'  set sc"
  by (meson assms local.get_dom_component_ptr
      local.get_elements_by_class_name_is_strongly_dom_component_safe_step
      local.get_scdom_component_contains_get_dom_component subsetD)


lemma get_elements_by_class_name_is_strongly_scdom_component_safe:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_elements_by_class_name ptr idd r results"
  assumes "h  get_elements_by_class_name ptr idd h h'"
  shows "is_strongly_scdom_component_safe {ptr} (cast ` set results) h h'"
proof -
  have "h = h'"
    using assms(5)
    by (meson local.get_elements_by_class_name_pure pure_returns_heap_eq)
  have "ptr |∈| object_ptr_kinds h"
    using assms(4)
    apply(auto simp add: get_elements_by_class_name_def)[1]
    by (metis (no_types, lifting) assms(1) assms(2) assms(3) bind_is_OK_E is_OK_returns_result_I
        local.first_in_tree_order_def local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap)
  obtain to where to: "h  to_tree_order ptr r to"
    by (meson ptr |∈| object_ptr_kinds h assms(1) assms(2) assms(3) is_OK_returns_result_E
        local.to_tree_order_ok)
  then have "cast ` set results  set to"
    using assms(4) local.get_elements_by_class_name_result_in_tree_order by auto
  obtain c where c: "h  a_get_scdom_component ptr r c"
    using ptr |∈| object_ptr_kinds h assms(1) assms(2) assms(3) local.get_scdom_component_impl
      local.get_scdom_component_ok by blast

  then show ?thesis
    using assms h = h'
    apply(auto simp add: is_strongly_scdom_component_safe_def Let_def preserved_def
        get_element_by_id_def first_in_tree_order_def elim!: bind_returns_result_E2 intro!: map_filter_M_pure
        bind_pure_I split: option.splits list.splits)[1]
    by (smt IntI ptr |∈| object_ptr_kinds h finite_set_in
        get_elements_by_class_name_is_strongly_scdom_component_safe_step local.get_dom_component_ok
        local.get_dom_component_ptr local.get_scdom_component_impl
        local.get_scdom_component_subset_get_dom_component returns_result_select_result select_result_I2 subsetD)
qed
end

interpretation i_get_scdom_component_get_elements_by_class_name?: l_get_scdom_component_get_elements_by_class_nameCore_DOM
  get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
  get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component
  is_strongly_dom_component_safe is_weakly_dom_component_safe get_disconnected_nodes
  get_disconnected_nodes_locs to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs
  get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_element_by_id
  get_elements_by_class_name get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs
  by(auto simp add: l_get_scdom_component_get_elements_by_class_nameCore_DOM_def instances)
declare l_get_scdom_component_get_element_by_idCore_DOM_axioms [instances]


subsubsection ‹get\_elements\_by\_tag\_name›

locale l_get_scdom_component_get_elements_by_tag_nameCore_DOM =
  l_get_dom_component_get_scdom_component +
  l_get_scdom_componentCore_DOM +
  l_get_dom_component_get_elements_by_tag_nameCore_DOM
begin

lemma get_elements_by_tag_name_is_strongly_scdom_component_safe_step:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_scdom_component ptr r sc"
  assumes "h  get_elements_by_tag_name ptr' idd r results"
  assumes "result  set results"
  shows "cast result  set sc  ptr'  set sc"
  by (meson assms local.get_dom_component_ptr
      local.get_elements_by_tag_name_is_strongly_dom_component_safe_step
      local.get_scdom_component_contains_get_dom_component subsetD)


lemma get_elements_by_tag_name_is_strongly_scdom_component_safe:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_elements_by_tag_name ptr idd r results"
  assumes "h  get_elements_by_tag_name ptr idd h h'"
  shows "is_strongly_scdom_component_safe {ptr} (cast ` set results) h h'"
proof -
  have "h = h'"
    using assms(5)
    by (meson local.get_elements_by_tag_name_pure pure_returns_heap_eq)
  have "ptr |∈| object_ptr_kinds h"
    using assms(4)
    apply(auto simp add: get_elements_by_tag_name_def)[1]
    by (metis (no_types, lifting) assms(1) assms(2) assms(3) bind_is_OK_E is_OK_returns_result_I
        local.first_in_tree_order_def local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap)
  obtain to where to: "h  to_tree_order ptr r to"
    by (meson ptr |∈| object_ptr_kinds h assms(1) assms(2) assms(3) is_OK_returns_result_E
        local.to_tree_order_ok)
  then have "cast ` set results  set to"
    using assms(4) local.get_elements_by_tag_name_result_in_tree_order by auto
  obtain c where c: "h  a_get_scdom_component ptr r c"
    using ptr |∈| object_ptr_kinds h assms(1) assms(2) assms(3) local.get_scdom_component_impl
      local.get_scdom_component_ok by blast

  then show ?thesis
    using assms h = h'
    apply(auto simp add: is_strongly_scdom_component_safe_def Let_def preserved_def
        get_element_by_id_def first_in_tree_order_def elim!: bind_returns_result_E2 intro!:
        map_filter_M_pure bind_pure_I split: option.splits list.splits)[1]
    by (smt IntI ptr |∈| object_ptr_kinds h finite_set_in
        get_elements_by_tag_name_is_strongly_scdom_component_safe_step local.get_dom_component_ok
        local.get_dom_component_ptr local.get_scdom_component_impl
        local.get_scdom_component_subset_get_dom_component returns_result_select_result select_result_I2
        subsetD)
qed
end

interpretation i_get_scdom_component_get_elements_by_tag_name?:
  l_get_scdom_component_get_elements_by_tag_nameCore_DOM
  get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
  get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe
  get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe
  get_disconnected_nodes get_disconnected_nodes_locs to_tree_order get_parent get_parent_locs
  get_child_nodes get_child_nodes_locs get_root_node get_root_node_locs get_ancestors
  get_ancestors_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name
  first_in_tree_order get_attribute get_attribute_locs
  by(auto simp add: l_get_scdom_component_get_elements_by_tag_nameCore_DOM_def instances)
declare l_get_scdom_component_get_element_by_idCore_DOM_axioms [instances]


subsubsection ‹remove\_child›

locale l_get_scdom_component_remove_childCore_DOM =
  l_get_dom_component_get_scdom_component +
  l_get_dom_componentCore_DOM +
  l_get_scdom_componentCore_DOM +
  l_remove_childCore_DOM +
  l_set_child_nodesCore_DOM +
  l_set_disconnected_nodesCore_DOM +
  l_get_owner_document_wf +
  l_remove_child_wf2Core_DOM get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs
  get_parent get_parent_locs get_owner_document get_disconnected_nodes get_disconnected_nodes_locs
  set_disconnected_nodes set_disconnected_nodes_locs remove_child remove_child_locs remove type_wf
  known_ptr known_ptrs heap_is_wellformed parent_child_rel
begin
lemma remove_child_is_component_unsafe:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  remove_child ptr' child h h'"
  assumes "ptr  set |h  get_dom_component ptr'|r"
  assumes "ptr  set |h  get_dom_component (cast |h  get_owner_document (castnode_ptr2object_ptr child)|r)|r"
    (* assumes "ptr ∉ set |h ⊢ get_dom_component (cast child)|r" *)
  shows "preserved (get_M ptr getter) h h'"
proof -
  have "ptr  ptr'"
    using assms(5)
    by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) is_OK_returns_heap_I
        is_OK_returns_result_E local.get_dom_component_ok local.get_dom_component_ptr
        local.remove_child_ptr_in_heap select_result_I2)

  obtain owner_document where owner_document: "h  get_owner_document (castnode_ptr2object_ptr child) r owner_document"
    by (meson assms(1) assms(2) assms(3) assms(4) is_OK_returns_result_E local.get_owner_document_ok
        local.remove_child_child_in_heap node_ptr_kinds_commutes)
  then
  obtain c where "h  get_dom_component (cast owner_document) r c"
    using get_dom_component_ok owner_document assms(1) assms(2) assms(3)
    by (meson document_ptr_kinds_commutes get_owner_document_owner_document_in_heap select_result_I)
  then
  have "ptr  cast owner_document"
    using assms(6) assms(1) assms(2) assms(3) local.get_dom_component_ptr owner_document
    by auto

  show ?thesis
    using remove_child_writes assms(4)
    apply(rule reads_writes_preserved2)
    apply(auto simp add: remove_child_locs_def set_child_nodes_locs_def
        set_disconnected_nodes_locs_def all_args_def split: option.splits)[1]
         apply (metis ptr  ptr' document_ptr_casts_commute3 get_M_Mdocument_preserved3)
        apply (metis (no_types, lifting) ptr  castdocument_ptr2object_ptr owner_document
        get_M_Mdocument_preserved3 owner_document select_result_I2)
       apply (metis (no_types, lifting) ptr  castdocument_ptr2object_ptr owner_document
        get_M_Mdocument_preserved3 owner_document select_result_I2)
      apply (metis ptr  ptr' element_ptr_casts_commute3 get_M_Element_preserved8)
     apply (metis (no_types, lifting) ptr  castdocument_ptr2object_ptr owner_document
        get_M_Mdocument_preserved3 owner_document select_result_I2)
    apply (metis (no_types, lifting) ptr  castdocument_ptr2object_ptr owner_document
        get_M_Mdocument_preserved3 owner_document select_result_I2)
    done
qed

lemma remove_child_is_strongly_dom_component_safe_step:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  remove_child ptr' child h h'"
  assumes "ptr  set |h  get_scdom_component ptr'|r"
    (* assumes "ptr ∉ set |h ⊢ get_dom_component (cast |h ⊢ get_owner_document (castnode_ptr2object_ptr child)|r)|r" *)
  assumes "ptr  set |h  get_scdom_component (cast child)|r"
  shows "preserved (get_M ptr getter) h h'"
proof -
  obtain sc where sc: "h  get_scdom_component ptr' r sc"
    using get_scdom_component_ok
    by (meson assms(1) assms(2) assms(3) assms(4) is_OK_returns_heap_I local.remove_child_ptr_in_heap
        returns_result_select_result)

  have "child |∈| node_ptr_kinds h"
    using assms(4) remove_child_child_in_heap by blast
  then
  obtain child_sc where child_sc: "h  get_scdom_component (cast child) r child_sc"
    using get_scdom_component_ok
    by (meson assms(1) assms(2) assms(3) node_ptr_kinds_commutes select_result_I)
  then obtain owner_document where owner_document: "h  get_owner_document (castnode_ptr2object_ptr child) r owner_document"
    by (meson child |∈| node_ptr_kinds h assms(1) assms(2) assms(3) contra_subsetD
        get_scdom_component_owner_document_same is_OK_returns_result_E
        get_scdom_component_subset_get_dom_component local.get_dom_component_ok local.get_dom_component_ptr
        node_ptr_kinds_commutes)
  then have "h  get_scdom_component (cast owner_document) r child_sc"
    using child_sc
    by (smt child |∈| node_ptr_kinds h assms(1) assms(2) assms(3) contra_subsetD
        get_scdom_component_subset_get_dom_component get_scdom_component_owner_document_same
        get_scdom_component_ptrs_same_scope_component local.get_dom_component_ok local.get_dom_component_ptr
        node_ptr_kinds_commutes returns_result_select_result select_result_I2)

  have "ptr  set |h  get_dom_component ptr'|r"
    by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) assms(5) contra_subsetD
        get_scdom_component_subset_get_dom_component is_OK_returns_heap_I local.get_dom_component_ok
        local.remove_child_ptr_in_heap returns_result_select_result sc select_result_I2)

  moreover have "ptr  set |h  get_scdom_component (cast |h  get_owner_document (castnode_ptr2object_ptr child)|r)|r"
    using get_scdom_component_owner_document_same get_scdom_component_ptrs_same_scope_component
    by (metis (no_types, lifting)
        h  get_scdom_component (castdocument_ptr2object_ptr owner_document) r child_sc assms(6) child_sc
        owner_document select_result_I2)
  have "ptr  set |h  get_dom_component (cast |h  get_owner_document (castnode_ptr2object_ptr child)|r)|r"
    using get_scdom_component_owner_document_same
    by (metis (no_types, lifting)
        h  get_scdom_component (castdocument_ptr2object_ptr owner_document) r child_sc
        ptr  set |h  get_scdom_component (castdocument_ptr2object_ptr |h  get_owner_document (castnode_ptr2object_ptr child)|r)|r
        assms(1) assms(2) assms(3) contra_subsetD document_ptr_kinds_commutes get_scdom_component_subset_get_dom_component
        is_OK_returns_result_E local.get_dom_component_ok local.get_owner_document_owner_document_in_heap owner_document
        select_result_I2)
  ultimately show ?thesis
    using assms(1) assms(2) assms(3) assms(4) remove_child_is_component_unsafe by blast
qed



lemma remove_child_is_strongly_dom_component_safe:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  remove_child ptr child h h'"
  shows "is_strongly_scdom_component_safe {ptr, cast child} {} h h'"
proof -
  obtain owner_document children_h h2 disconnected_nodes_h where
    owner_document: "h  get_owner_document (castnode_ptr2object_ptr child) r owner_document" and
    children_h: "h  get_child_nodes ptr r children_h" and
    child_in_children_h: "child  set children_h" and
    disconnected_nodes_h: "h   get_disconnected_nodes owner_document r disconnected_nodes_h" and
    h2: "h   set_disconnected_nodes owner_document (child # disconnected_nodes_h) h h2" and
    h': "h2  set_child_nodes ptr (remove1 child children_h) h h'"
    using assms(4)
    apply(auto simp add: remove_child_def elim!: bind_returns_heap_E
        dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure]
        pure_returns_heap_eq[rotated, OF get_child_nodes_pure] split: if_splits)[1]
    using pure_returns_heap_eq by fastforce

  have object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'"
    apply(rule writes_small_big[where P="λh h'. object_ptr_kinds h = object_ptr_kinds h'",
          OF remove_child_writes assms(4)])
    unfolding remove_child_locs_def
    using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved
    by (auto simp add: reflp_def transp_def)
  then have object_ptr_kinds_eq: "ptrs. h  object_ptr_kinds_M r ptrs = h'  object_ptr_kinds_M r ptrs"
    unfolding object_ptr_kinds_M_defs by simp
  then have object_ptr_kinds_eq2: "|h  object_ptr_kinds_M|r = |h'  object_ptr_kinds_M|r"
    using select_result_eq by force
  then have node_ptr_kinds_eq2: "|h  node_ptr_kinds_M|r = |h'  node_ptr_kinds_M|r"
    using node_ptr_kinds_M_eq by auto
  then have node_ptr_kinds_eq3: "node_ptr_kinds h = node_ptr_kinds h'"
    using node_ptr_kinds_M_eq by auto
  have document_ptr_kinds_eq2: "|h  document_ptr_kinds_M|r = |h'  document_ptr_kinds_M|r"
    using object_ptr_kinds_eq2 document_ptr_kinds_M_eq by auto
  then have document_ptr_kinds_eq3: "document_ptr_kinds h = document_ptr_kinds h'"
    using document_ptr_kinds_M_eq by auto
  have children_eq:
    "ptr' children. ptr  ptr'  h  get_child_nodes ptr' r children = h'  get_child_nodes ptr' r children"
    apply(rule reads_writes_preserved[OF get_child_nodes_reads remove_child_writes assms(4)])
    unfolding remove_child_locs_def
    using set_disconnected_nodes_get_child_nodes set_child_nodes_get_child_nodes_different_pointers
    by fast
  then have children_eq2:
    "ptr' children. ptr  ptr'  |h  get_child_nodes ptr'|r = |h'  get_child_nodes ptr'|r"
    using select_result_eq by force
  have disconnected_nodes_eq: "document_ptr disconnected_nodes. document_ptr  owner_document
          h  get_disconnected_nodes document_ptr r disconnected_nodes
             = h'  get_disconnected_nodes document_ptr r disconnected_nodes"
    apply(rule reads_writes_preserved[OF get_disconnected_nodes_reads remove_child_writes assms(4)])
    unfolding remove_child_locs_def
    using set_child_nodes_get_disconnected_nodes set_disconnected_nodes_get_disconnected_nodes_different_pointers
    by (metis (no_types, lifting) Un_iff owner_document select_result_I2)
  then have disconnected_nodes_eq2:
    "document_ptr. document_ptr  owner_document
     |h  get_disconnected_nodes document_ptr|r = |h'  get_disconnected_nodes document_ptr|r"
    using select_result_eq by force

  have "h2  get_child_nodes ptr r children_h"
    apply(rule reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes h2 children_h] )
    by (simp add: set_disconnected_nodes_get_child_nodes)

  have "known_ptrs h'"
    using object_ptr_kinds_eq3 known_ptrs_preserved known_ptrs h by blast

  have "known_ptr ptr"
    using assms(3)
    using children_h get_child_nodes_ptr_in_heap local.known_ptrs_known_ptr by blast
  have "type_wf h2"
    using writes_small_big[where P="λh h'. type_wf h  type_wf h'", OF set_disconnected_nodes_writes h2]
    using set_disconnected_nodes_types_preserved assms(2)
    by(auto simp add: reflp_def transp_def)
  then have "type_wf h'"
    using writes_small_big[where P="λh h'. type_wf h  type_wf h'", OF set_child_nodes_writes h']
    using  set_child_nodes_types_preserved
    by(auto simp add: reflp_def transp_def)

  have children_h': "h'  get_child_nodes ptr r remove1 child children_h"
    using assms(4) owner_document h2 disconnected_nodes_h children_h
    apply(auto simp add: remove_child_def split: if_splits)[1]
    apply(drule bind_returns_heap_E3)
      apply(auto split: if_splits)[1]
     apply(simp)
    apply(auto split: if_splits)[1]
    apply(drule bind_returns_heap_E3)
      apply(auto)[1]
     apply(simp)
    apply(drule bind_returns_heap_E3)
      apply(auto)[1]
     apply(simp)
    apply(drule bind_returns_heap_E4)
      apply(auto)[1]
     apply simp
    using type_wf h2 set_child_nodes_get_child_nodes known_ptr ptr h'
    by blast

  have disconnected_nodes_h2: "h2  get_disconnected_nodes owner_document r child # disconnected_nodes_h"
    using owner_document assms(4) h2 disconnected_nodes_h
    apply (auto simp add: remove_child_def  split: if_splits)[1]
    apply(drule bind_returns_heap_E2)
      apply(auto split: if_splits)[1]
     apply(simp)
    by(auto simp add: local.set_disconnected_nodes_get_disconnected_nodes split: if_splits)
  then have disconnected_nodes_h': "h'  get_disconnected_nodes owner_document r child # disconnected_nodes_h"
    apply(rule reads_writes_separate_forwards[OF get_disconnected_nodes_reads set_child_nodes_writes h'])
    by (simp add: set_child_nodes_get_disconnected_nodes)

  moreover have "a_acyclic_heap h"
    using assms(1)  by (simp add: heap_is_wellformed_def)
  have "parent_child_rel h'  parent_child_rel h"
  proof (standard, safe)
    fix parent child
    assume a1: "(parent, child)  parent_child_rel h'"
    then show "(parent, child)  parent_child_rel h"
    proof (cases "parent = ptr")
      case True
      then show ?thesis
        using a1 remove_child_removes_parent[OF assms(1) assms(4)] children_h children_h'
          get_child_nodes_ptr_in_heap
        apply(auto simp add: parent_child_rel_def object_ptr_kinds_eq )[1]
        by (metis imageI notin_set_remove1)
    next
      case False
      then show ?thesis
        using a1
        by(auto simp add: parent_child_rel_def object_ptr_kinds_eq3 children_eq2)
    qed
  qed
  then have "a_acyclic_heap h'"
    using ‹a_acyclic_heap h acyclic_heap_def acyclic_subset by blast

  moreover have "a_all_ptrs_in_heap h"
    using assms(1)  by (simp add: heap_is_wellformed_def)
  then have "a_all_ptrs_in_heap h'"
    apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq3 disconnected_nodes_eq)[1]
     apply (metis (no_types, lifting) type_wf h' assms local.get_child_nodes_ok local.known_ptrs_known_ptr
        local.remove_child_children_subset notin_fset object_ptr_kinds_eq3 returns_result_select_result subset_code(1))
    apply (metis (no_types, lifting) assms(4) disconnected_nodes_eq2 disconnected_nodes_h disconnected_nodes_h'
        document_ptr_kinds_eq3 finite_set_in local.remove_child_child_in_heap node_ptr_kinds_eq3 select_result_I2
        set_ConsD subset_code(1))
    done
  moreover have "a_owner_document_valid h"
    using assms(1)  by (simp add: heap_is_wellformed_def)
  then have "a_owner_document_valid h'"
    apply(auto simp add: a_owner_document_valid_def object_ptr_kinds_eq3 document_ptr_kinds_eq3
        node_ptr_kinds_eq3)[1]
  proof -
    fix node_ptr
    assume 0: "node_ptrfset (node_ptr_kinds h'). (document_ptr. document_ptr |∈| document_ptr_kinds h' 
node_ptr  set |h  get_disconnected_nodes document_ptr|r)  (parent_ptr. parent_ptr |∈| object_ptr_kinds h' 
node_ptr  set |h  get_child_nodes parent_ptr|r)"
      and 1: "node_ptr |∈| node_ptr_kinds h'"
      and 2: "parent_ptr. parent_ptr |∈| object_ptr_kinds h'  node_ptr  set |h'  get_child_nodes parent_ptr|r"
    then show "document_ptr. document_ptr |∈| document_ptr_kinds h'
                        node_ptr  set |h'  get_disconnected_nodes document_ptr|r"
    proof (cases "node_ptr = child")
      case True
      show ?thesis
        apply(rule exI[where x=owner_document])
        using children_eq2 disconnected_nodes_eq2 children_h children_h' disconnected_nodes_h' True
        by (metis (no_types, lifting) get_disconnected_nodes_ptr_in_heap is_OK_returns_result_I
            list.set_intros(1) select_result_I2)
    next
      case False
      then show ?thesis
        using 0 1 2 children_eq2 children_h children_h' disconnected_nodes_eq2 disconnected_nodes_h
          disconnected_nodes_h'
        apply(auto simp add: children_eq2 disconnected_nodes_eq2 dest!: select_result_I2)[1]
        by (metis children_eq2 disconnected_nodes_eq2 finite_set_in in_set_remove1 list.set_intros(2))
    qed
  qed

  moreover
  {
    have h0: "a_distinct_lists h"
      using assms(1)  by (simp add: heap_is_wellformed_def)
    moreover have ha1: "(xset |h  object_ptr_kinds_M|r. set |h  get_child_nodes x|r)
                  (xset |h  document_ptr_kinds_M|r. set |h  get_disconnected_nodes x|r) = {}"
      using ‹a_distinct_lists h
      unfolding a_distinct_lists_def
      by(auto)
    have ha2: "ptr |∈| object_ptr_kinds h"
      using children_h get_child_nodes_ptr_in_heap by blast
    have ha3: "child  set |h  get_child_nodes ptr|r"
      using child_in_children_h children_h
      by(simp)
    have child_not_in: "document_ptr. document_ptr |∈| document_ptr_kinds h
                           child  set |h  get_disconnected_nodes document_ptr|r"
      using ha1 ha2 ha3
      apply(simp)
      using IntI by fastforce
    moreover have "distinct |h  object_ptr_kinds_M|r"
      apply(rule select_result_I)
      by(auto simp add: object_ptr_kinds_M_defs)
    moreover have "distinct |h  document_ptr_kinds_M|r"
      apply(rule select_result_I)
      by(auto simp add: document_ptr_kinds_M_defs)
    ultimately have "a_distinct_lists h'"
    proof(simp (no_asm) add: a_distinct_lists_def, safe)
      assume 1: "a_distinct_lists h"
        and 3: "distinct |h  object_ptr_kinds_M|r"

      assume 1: "a_distinct_lists h"
        and 3: "distinct |h  object_ptr_kinds_M|r"
      have 4: "distinct (concat ((map (λptr. |h  get_child_nodes ptr|r) |h  object_ptr_kinds_M|r)))"
        using 1 by(auto simp add: a_distinct_lists_def)
      show "distinct (concat (map (λptr. |h'  get_child_nodes ptr|r)
                     (sorted_list_of_set (fset (object_ptr_kinds h')))))"
      proof(rule distinct_concat_map_I[OF 3[unfolded object_ptr_kinds_eq2], simplified])
        fix x
        assume 5: "x |∈| object_ptr_kinds h'"
        then have 6: "distinct |h  get_child_nodes x|r"
          using 4 distinct_concat_map_E object_ptr_kinds_eq2 by fastforce
        obtain children where children: "h  get_child_nodes x r children"
          and distinct_children: "distinct children"
          by (metis "5" "6" assms get_child_nodes_ok local.known_ptrs_known_ptr
              object_ptr_kinds_eq3 select_result_I)
        obtain children' where children': "h'  get_child_nodes x r children'"
          using children children_eq children_h' by fastforce
        then have "distinct children'"
        proof (cases "ptr = x")
          case True
          then show ?thesis
            using children distinct_children children_h children_h'
            by (metis children' distinct_remove1 returns_result_eq)
        next
          case False
          then show ?thesis
            using children distinct_children children_eq[OF False]
            using children' distinct_lists_children h0
            using select_result_I2 by fastforce
        qed

        then show "distinct |h'  get_child_nodes x|r"
          using children' by(auto simp add: )
      next
        fix x y
        assume 5: "x |∈| object_ptr_kinds h'" and 6: "y |∈| object_ptr_kinds h'" and 7: "x  y"
        obtain children_x where children_x: "h  get_child_nodes x r children_x"
          by (metis "5" assms get_child_nodes_ok is_OK_returns_result_E
              local.known_ptrs_known_ptr object_ptr_kinds_eq3)
        obtain children_y where children_y: "h  get_child_nodes y r children_y"
          by (metis "6" assms get_child_nodes_ok is_OK_returns_result_E
              local.known_ptrs_known_ptr object_ptr_kinds_eq3)
        obtain children_x' where children_x': "h'  get_child_nodes x r children_x'"
          using children_eq children_h' children_x by fastforce
        obtain children_y' where children_y': "h'  get_child_nodes y r children_y'"
          using children_eq children_h' children_y by fastforce
        have "distinct (concat (map (λptr. |h  get_child_nodes ptr|r) |h  object_ptr_kinds_M|r))"
          using h0 by(auto simp add: a_distinct_lists_def)
        then have 8: "set children_x  set children_y = {}"
          using "7" assms(1) children_x children_y local.heap_is_wellformed_one_parent by blast
        have "set children_x'  set children_y' = {}"
        proof (cases "ptr = x")
          case True
          then have "ptr  y"
            by(simp add: 7)
          have "children_x' = remove1 child children_x"
            using children_h children_h' children_x children_x' True returns_result_eq by fastforce
          moreover have "children_y' = children_y"
            using children_y children_y' children_eq[OF ptr  y] by auto
          ultimately show ?thesis
            using 8 set_remove1_subset by fastforce
        next
          case False
          then show ?thesis
          proof (cases "ptr = y")
            case True
            have "children_y' = remove1 child children_y"
              using children_h children_h' children_y children_y' True returns_result_eq by fastforce
            moreover have "children_x' = children_x"
              using children_x children_x' children_eq[OF ptr  x] by auto
            ultimately show ?thesis
              using 8 set_remove1_subset by fastforce
          next
            case False
            have "children_x' = children_x"
              using children_x children_x' children_eq[OF ptr  x] by auto
            moreover have "children_y' = children_y"
              using children_y children_y' children_eq[OF ptr  y] by auto
            ultimately show ?thesis
              using 8 by simp
          qed
        qed
        then show "set |h'  get_child_nodes x|r  set |h'  get_child_nodes y|r = {}"
          using children_x' children_y'
          by (metis (no_types, lifting) select_result_I2)
      qed
    next
      assume 2: "distinct |h  document_ptr_kinds_M|r"
      then have 4: "distinct  (sorted_list_of_set (fset (document_ptr_kinds h')))"
        by simp
      have 3: "distinct (concat (map (λdocument_ptr. |h  get_disconnected_nodes document_ptr|r)
                        (sorted_list_of_set (fset (document_ptr_kinds h')))))"
        using h0
        by(simp add: a_distinct_lists_def document_ptr_kinds_eq3)

      show "distinct (concat (map (λdocument_ptr. |h'  get_disconnected_nodes document_ptr|r)
                         (sorted_list_of_set (fset (document_ptr_kinds h')))))"
      proof(rule distinct_concat_map_I[OF 4[unfolded document_ptr_kinds_eq3]])
        fix x
        assume 4: "x  set (sorted_list_of_set (fset (document_ptr_kinds h')))"
        have 5: "distinct |h  get_disconnected_nodes x|r"
          using distinct_lists_disconnected_nodes[OF h0] 4 get_disconnected_nodes_ok
          by (simp add: assms document_ptr_kinds_eq3 select_result_I)
        show "distinct |h'  get_disconnected_nodes x|r"
        proof (cases "x = owner_document")
          case True
          have "child  set |h  get_disconnected_nodes x|r"
            using child_not_in  document_ptr_kinds_eq2 "4" by fastforce
          moreover have "|h'  get_disconnected_nodes x|r = child # |h  get_disconnected_nodes x|r"
            using disconnected_nodes_h' disconnected_nodes_h unfolding True
            by(simp)
          ultimately show ?thesis
            using 5 unfolding True
            by simp
        next
          case False
          show ?thesis
            using "5" False disconnected_nodes_eq2 by auto
        qed
      next
        fix x y
        assume 4: "x  set (sorted_list_of_set (fset (document_ptr_kinds h')))"
          and 5: "y  set (sorted_list_of_set (fset (document_ptr_kinds h')))" and "x  y"
        obtain disc_nodes_x where disc_nodes_x: "h  get_disconnected_nodes x r disc_nodes_x"
          using 4 get_disconnected_nodes_ok[OF type_wf h, of x] document_ptr_kinds_eq2
          by auto
        obtain disc_nodes_y where disc_nodes_y: "h  get_disconnected_nodes y r disc_nodes_y"
          using 5 get_disconnected_nodes_ok[OF type_wf h, of y] document_ptr_kinds_eq2
          by auto
        obtain disc_nodes_x' where disc_nodes_x': "h'  get_disconnected_nodes x r disc_nodes_x'"
          using 4 get_disconnected_nodes_ok[OF type_wf h', of x] document_ptr_kinds_eq2
          by auto
        obtain disc_nodes_y' where disc_nodes_y': "h'  get_disconnected_nodes y r disc_nodes_y'"
          using 5 get_disconnected_nodes_ok[OF type_wf h', of y] document_ptr_kinds_eq2
          by auto
        have "distinct
               (concat (map (λdocument_ptr. |h  get_disconnected_nodes document_ptr|r) |h  document_ptr_kinds_M|r))"
          using h0 by (simp add: a_distinct_lists_def)
        then have 6: "set disc_nodes_x  set disc_nodes_y = {}"
          using x  y assms(1) disc_nodes_x disc_nodes_y local.heap_is_wellformed_one_disc_parent
          by blast

        have "set disc_nodes_x'  set disc_nodes_y' = {}"
        proof (cases "x = owner_document")
          case True
          then have "y  owner_document"
            using x  y by simp
          then have "disc_nodes_y' = disc_nodes_y"
            using disconnected_nodes_eq[OF y  owner_document] disc_nodes_y disc_nodes_y'
            by auto
          have "disc_nodes_x' = child # disc_nodes_x"
            using disconnected_nodes_h' disc_nodes_x disc_nodes_x' True disconnected_nodes_h returns_result_eq
            by fastforce
          have "child  set disc_nodes_y"
            using child_not_in  disc_nodes_y 5
            using document_ptr_kinds_eq2  by fastforce
          then show ?thesis
            apply(unfold disc_nodes_x' = child # disc_nodes_x disc_nodes_y' = disc_nodes_y)
            using 6 by auto
        next
          case False
          then show ?thesis
          proof (cases "y = owner_document")
            case True
            then have "disc_nodes_x' = disc_nodes_x"
              using disconnected_nodes_eq[OF x  owner_document] disc_nodes_x disc_nodes_x' by auto
            have "disc_nodes_y' = child # disc_nodes_y"
              using disconnected_nodes_h' disc_nodes_y disc_nodes_y' True disconnected_nodes_h returns_result_eq
              by fastforce
            have "child  set disc_nodes_x"
              using child_not_in  disc_nodes_x 4
              using document_ptr_kinds_eq2  by fastforce
            then show ?thesis
              apply(unfold disc_nodes_y' = child # disc_nodes_y disc_nodes_x' = disc_nodes_x)
              using 6 by auto
          next
            case False
            have "disc_nodes_x' = disc_nodes_x"
              using disconnected_nodes_eq[OF x  owner_document] disc_nodes_x disc_nodes_x' by auto
            have "disc_nodes_y' = disc_nodes_y"
              using disconnected_nodes_eq[OF y  owner_document] disc_nodes_y disc_nodes_y' by auto
            then show ?thesis
              apply(unfold disc_nodes_y' = disc_nodes_y disc_nodes_x' = disc_nodes_x)
              using 6 by auto
          qed
        qed
        then show "set |h'  get_disconnected_nodes x|r  set |h'  get_disconnected_nodes y|r = {}"
          using disc_nodes_x' disc_nodes_y' by auto
      qed
    next
      fix x xa xb
      assume 1: "xa  fset (object_ptr_kinds h')"
        and 2: "x  set |h'  get_child_nodes xa|r"
        and 3: "xb  fset (document_ptr_kinds h')"
        and 4: "x  set |h'  get_disconnected_nodes xb|r"
      obtain disc_nodes where disc_nodes: "h  get_disconnected_nodes xb r disc_nodes"
        using 3 get_disconnected_nodes_ok[OF type_wf h, of xb] document_ptr_kinds_eq2 by auto
      obtain disc_nodes' where disc_nodes': "h'  get_disconnected_nodes xb r disc_nodes'"
        using 3 get_disconnected_nodes_ok[OF type_wf h', of xb]  document_ptr_kinds_eq2 by auto

      obtain children where children: "h  get_child_nodes xa r children"
        by (metis "1" assms finite_set_in get_child_nodes_ok is_OK_returns_result_E
            local.known_ptrs_known_ptr object_ptr_kinds_eq3)
      obtain children' where children': "h'  get_child_nodes xa r children'"
        using children children_eq children_h' by fastforce
      have "x. x  set |h  get_child_nodes xa|r  x  set |h  get_disconnected_nodes xb|r  False"
        using 1 3
        apply(fold ‹ object_ptr_kinds h =  object_ptr_kinds h')
        apply(fold ‹ document_ptr_kinds h =  document_ptr_kinds h')
        using children disc_nodes h0 apply(auto simp add: a_distinct_lists_def)[1]
        by (metis (no_types, lifting) h0 local.distinct_lists_no_parent select_result_I2)
      then have 5: "x. x  set children  x  set disc_nodes  False"
        using children disc_nodes  by fastforce
      have 6: "|h'  get_child_nodes xa|r = children'"
        using children' by (simp add: )
      have 7: "|h'  get_disconnected_nodes xb|r = disc_nodes'"
        using disc_nodes' by (simp add: )
      have "False"
      proof (cases "xa = ptr")
        case True
        have "distinct children_h"
          using children_h distinct_lists_children h0 known_ptr ptr by blast
        have "|h'  get_child_nodes ptr|r = remove1 child children_h"
          using children_h'
          by(simp add: )
        have "children = children_h"
          using True children children_h by auto
        show ?thesis
          using disc_nodes' children' 5 2 4 children_h ‹distinct children_h disconnected_nodes_h'
          apply(auto simp add: 6 7
              xa = ptr |h'  get_child_nodes ptr|r = remove1 child children_h children = children_h)[1]
          by (metis (no_types, lifting) disc_nodes disconnected_nodes_eq2 disconnected_nodes_h
              select_result_I2 set_ConsD)
      next
        case False
        have "children' = children"
          using children' children children_eq[OF False[symmetric]]
          by auto
        then show ?thesis
        proof (cases "xb = owner_document")
          case True
          then show ?thesis
            using disc_nodes disconnected_nodes_h disconnected_nodes_h'
            using "2" "4" "5" "6" "7" False children' = children assms(1) child_in_children_h
              child_parent_dual children children_h disc_nodes' get_child_nodes_ptr_in_heap
              list.set_cases list.simps(3) option.simps(1) returns_result_eq set_ConsD
            by (metis (no_types, hide_lams) assms)
        next
          case False
          then show ?thesis
            using "2" "4" "5" "6" "7" children' = children disc_nodes disc_nodes'
              disconnected_nodes_eq returns_result_eq
            by metis
        qed
      qed
      then show "x  {}"
        by simp
    qed
  }

  ultimately have "heap_is_wellformed h'"
    using heap_is_wellform