Theory Core_DOM_Heap_WF

(***********************************************************************************
 * Copyright (c) 2016-2018 The University of Sheffield, 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‹Wellformedness›
text‹In this theory, we discuss the wellformedness of the DOM. First, we define
wellformedness and, second, we show for all functions for querying and modifying the
DOM to what extend they preserve wellformendess.›

theory Core_DOM_Heap_WF
  imports
    "Core_DOM_Functions"
begin

locale l_heap_is_wellformedCore_DOM_defs =
  l_get_child_nodes_defs get_child_nodes get_child_nodes_locs +
  l_get_disconnected_nodes_defs get_disconnected_nodes get_disconnected_nodes_locs
  for get_child_nodes :: "(_::linorder) object_ptr  ((_) heap, exception, (_) node_ptr list) prog"
    and get_child_nodes_locs :: "(_) object_ptr  ((_) heap  (_) heap  bool) set"
    and get_disconnected_nodes :: "(_) document_ptr  ((_) heap, exception, (_) node_ptr list) prog"
    and get_disconnected_nodes_locs :: "(_) document_ptr  ((_) heap  (_) heap  bool) set"
begin
definition a_owner_document_valid :: "(_) heap  bool"
  where
    "a_owner_document_valid h  (node_ptr  fset (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)))"

lemma a_owner_document_valid_code [code]: "a_owner_document_valid h  node_ptr_kinds h |⊆|
  fset_of_list (concat (map (λparent. |h  get_child_nodes parent|r)
(sorted_list_of_fset (object_ptr_kinds h)) @
map (λparent. |h  get_disconnected_nodes parent|r)
(sorted_list_of_fset (document_ptr_kinds h))))
"
  apply(auto simp add: a_owner_document_valid_def l_heap_is_wellformedCore_DOM_defs.a_owner_document_valid_def)[1]
proof -
  fix x
  assume 1: " 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)"
  assume 2: "x |∈| node_ptr_kinds h"
  assume 3: "x |∉| fset_of_list (concat (map (λparent. |h  get_disconnected_nodes parent|r)
(sorted_list_of_fset (document_ptr_kinds h))))"
  have "¬(document_ptr. document_ptr |∈| document_ptr_kinds h  x  set |h  get_disconnected_nodes document_ptr|r)"
    using 1 2 3
    by (smt (verit) UN_I fset_of_list_elem image_eqI set_concat set_map sorted_list_of_fset_simps(1))
  then
  have "(parent_ptr. parent_ptr |∈| object_ptr_kinds h  x  set |h  get_child_nodes parent_ptr|r)"
    using 1 2
    by auto
  then obtain parent_ptr where parent_ptr: "parent_ptr |∈| object_ptr_kinds h 
x  set |h  get_child_nodes parent_ptr|r"
    by auto
  moreover have "parent_ptr  set (sorted_list_of_fset (object_ptr_kinds h))"
    using parent_ptr by auto
  moreover have "|h  get_child_nodes parent_ptr|r  set (map (λparent. |h  get_child_nodes parent|r)
(sorted_list_of_fset (object_ptr_kinds h)))"
    using calculation(2) by auto
  ultimately
  show "x |∈| fset_of_list (concat (map (λparent. |h  get_child_nodes parent|r)
(sorted_list_of_fset (object_ptr_kinds h))))"
    using fset_of_list_elem by fastforce
next
  fix node_ptr
  assume 1: "node_ptr_kinds h |⊆| fset_of_list (concat (map (λparent. |h  get_child_nodes parent|r)
(sorted_list_of_fset (object_ptr_kinds h)))) |∪|
fset_of_list (concat (map (λparent. |h  get_disconnected_nodes parent|r)
(sorted_list_of_fset (document_ptr_kinds h))))"
  assume 2: "node_ptr |∈| node_ptr_kinds h"
  assume 3: "parent_ptr. parent_ptr |∈| object_ptr_kinds h  node_ptr  set |h  get_child_nodes parent_ptr|r"
  have "node_ptr  set (concat (map (λparent. |h  get_child_nodes parent|r)
(sorted_list_of_fset (object_ptr_kinds h)))) 
node_ptr  set (concat (map (λparent. |h  get_disconnected_nodes parent|r)
(sorted_list_of_fset (document_ptr_kinds h))))"
    using 1 2
    by (meson fin_mono fset_of_list_elem funion_iff)
  then
  show "document_ptr. document_ptr |∈| document_ptr_kinds h 
node_ptr  set |h  get_disconnected_nodes document_ptr|r"
    using 3
    by auto
qed

definition a_parent_child_rel :: "(_) heap  ((_) object_ptr × (_) object_ptr) set"
  where
    "a_parent_child_rel h = {(parent, child). parent |∈| object_ptr_kinds h
                             child  cast ` set |h  get_child_nodes parent|r}"

lemma a_parent_child_rel_code [code]: "a_parent_child_rel h = set (concat (map
  (λparent. map
    (λchild. (parent, castnode_ptr2object_ptr child))
    |h  get_child_nodes parent|r)
  (sorted_list_of_fset (object_ptr_kinds h)))
)"
  by(auto simp add: a_parent_child_rel_def)

definition a_acyclic_heap :: "(_) heap  bool"
  where
    "a_acyclic_heap h = acyclic (a_parent_child_rel h)"

definition a_all_ptrs_in_heap :: "(_) heap  bool"
  where
    "a_all_ptrs_in_heap h 
      (ptr  fset (object_ptr_kinds h). set |h  get_child_nodes ptr|r  fset (node_ptr_kinds h)) 
      (document_ptr  fset (document_ptr_kinds h).
set |h  get_disconnected_nodes document_ptr|r  fset (node_ptr_kinds h))"

definition a_distinct_lists :: "(_) heap  bool"
  where
    "a_distinct_lists h = distinct (concat (
      (map (λptr. |h  get_child_nodes ptr|r) |h  object_ptr_kinds_M|r)
   @ (map (λdocument_ptr. |h  get_disconnected_nodes document_ptr|r) |h  document_ptr_kinds_M|r)
    ))"

definition a_heap_is_wellformed :: "(_) heap  bool"
  where
    "a_heap_is_wellformed h 
      a_acyclic_heap h  a_all_ptrs_in_heap h  a_distinct_lists h  a_owner_document_valid h"
end

locale l_heap_is_wellformed_defs =
  fixes heap_is_wellformed :: "(_) heap  bool"
  fixes parent_child_rel :: "(_) heap  ((_) object_ptr × (_) object_ptr) set"

global_interpretation l_heap_is_wellformedCore_DOM_defs get_child_nodes get_child_nodes_locs
  get_disconnected_nodes get_disconnected_nodes_locs
  defines heap_is_wellformed = "l_heap_is_wellformedCore_DOM_defs.a_heap_is_wellformed get_child_nodes
                     get_disconnected_nodes"
    and parent_child_rel = "l_heap_is_wellformedCore_DOM_defs.a_parent_child_rel get_child_nodes"
    and acyclic_heap = a_acyclic_heap
    and all_ptrs_in_heap = a_all_ptrs_in_heap
    and distinct_lists = a_distinct_lists
    and owner_document_valid = a_owner_document_valid
  .

locale l_heap_is_wellformedCore_DOM =
  l_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs
  + l_heap_is_wellformedCore_DOM_defs get_child_nodes get_child_nodes_locs get_disconnected_nodes
  get_disconnected_nodes_locs
  + l_heap_is_wellformed_defs heap_is_wellformed parent_child_rel
  + l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs
  for known_ptr :: "(_::linorder) object_ptr  bool"
    and type_wf :: "(_) heap  bool"
    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_disconnected_nodes :: "(_) document_ptr  ((_) heap, exception, (_) node_ptr list) prog"
    and get_disconnected_nodes_locs :: "(_) document_ptr  ((_) heap  (_) heap  bool) set"
    and heap_is_wellformed :: "(_) heap  bool"
    and parent_child_rel :: "(_) heap  ((_) object_ptr × (_) object_ptr) set" +
  assumes heap_is_wellformed_impl: "heap_is_wellformed = a_heap_is_wellformed"
  assumes parent_child_rel_impl: "parent_child_rel = a_parent_child_rel"
begin
lemmas heap_is_wellformed_def = heap_is_wellformed_impl[unfolded a_heap_is_wellformed_def]
lemmas parent_child_rel_def = parent_child_rel_impl[unfolded a_parent_child_rel_def]
lemmas acyclic_heap_def = a_acyclic_heap_def[folded parent_child_rel_impl]

lemma parent_child_rel_node_ptr:
  "(parent, child)  parent_child_rel h  is_node_ptr_kind child"
  by(auto simp add: parent_child_rel_def)

lemma parent_child_rel_child_nodes:
  assumes "known_ptr parent"
    and "h  get_child_nodes parent r children"
    and "child  set children"
  shows "(parent, cast child)  parent_child_rel h"
  using assms
  apply(auto simp add: parent_child_rel_def is_OK_returns_result_I )[1]
  using get_child_nodes_ptr_in_heap by blast

lemma parent_child_rel_child_nodes2:
  assumes "known_ptr parent"
    and "h  get_child_nodes parent r children"
    and "child  set children"
    and "castnode_ptr2object_ptr child = child_obj"
  shows "(parent, child_obj)  parent_child_rel h"
  using assms parent_child_rel_child_nodes by blast


lemma parent_child_rel_finite: "finite (parent_child_rel h)"
proof -
  have "parent_child_rel h = (ptr  set |h  object_ptr_kinds_M|r.
                 (child  set |h  get_child_nodes ptr|r. {(ptr, cast child)}))"
    by(auto simp add: parent_child_rel_def)
  moreover have "finite (ptr  set |h  object_ptr_kinds_M|r.
                    (child  set |h  get_child_nodes ptr|r. {(ptr, castnode_ptr2object_ptr child)}))"
    by simp
  ultimately show ?thesis
    by simp
qed

lemma distinct_lists_no_parent:
  assumes "a_distinct_lists h"
  assumes "h  get_disconnected_nodes document_ptr r disc_nodes"
  assumes "node_ptr  set disc_nodes"
  shows "¬(parent_ptr. parent_ptr |∈| object_ptr_kinds h
                 node_ptr  set |h  get_child_nodes parent_ptr|r)"
  using assms
  apply(auto simp add: a_distinct_lists_def)[1]
proof -
  fix parent_ptr :: "(_) object_ptr"
  assume a1: "parent_ptr |∈| object_ptr_kinds h"
  assume a2: "(xfset (object_ptr_kinds h).
              set |h  get_child_nodes x|r)  (xfset (document_ptr_kinds h).
                  set |h  get_disconnected_nodes x|r) = {}"
  assume a3: "h  get_disconnected_nodes document_ptr r disc_nodes"
  assume a4: "node_ptr  set disc_nodes"
  assume a5: "node_ptr  set |h  get_child_nodes parent_ptr|r"
  have f6: "parent_ptr  fset (object_ptr_kinds h)"
    using a1 by auto
  have f7: "document_ptr  fset (document_ptr_kinds h)"
    using a3 by (meson get_disconnected_nodes_ptr_in_heap is_OK_returns_result_I)
  have "|h  get_disconnected_nodes document_ptr|r = disc_nodes"
    using a3 by simp
  then show False
    using f7 f6 a5 a4 a2 by blast
qed


lemma distinct_lists_disconnected_nodes:
  assumes "a_distinct_lists h"
    and "h  get_disconnected_nodes document_ptr r disc_nodes"
  shows "distinct disc_nodes"
proof -
  have h1: "distinct (concat (map (λdocument_ptr. |h  get_disconnected_nodes document_ptr|r)
                             |h  document_ptr_kinds_M|r))"
    using assms(1)
    by(simp add: a_distinct_lists_def)
  then show ?thesis
    using concat_map_all_distinct[OF h1] assms(2) is_OK_returns_result_I get_disconnected_nodes_ok
    by (metis (no_types, lifting) DocumentMonad.ptr_kinds_ptr_kinds_M
        l_get_disconnected_nodes.get_disconnected_nodes_ptr_in_heap
        l_get_disconnected_nodes_axioms select_result_I2)
qed

lemma distinct_lists_children:
  assumes "a_distinct_lists h"
    and "known_ptr ptr"
    and "h  get_child_nodes ptr r children"
  shows "distinct children"
proof (cases "children = []", simp)
  assume "children  []"
  have h1: "distinct (concat ((map (λptr. |h  get_child_nodes ptr|r) |h  object_ptr_kinds_M|r)))"
    using assms(1)
    by(simp add: a_distinct_lists_def)
  show ?thesis
    using concat_map_all_distinct[OF h1] assms(2) assms(3)
    by (metis (no_types, lifting) ObjectMonad.ptr_kinds_ptr_kinds_M get_child_nodes_ptr_in_heap
        is_OK_returns_result_I select_result_I2)
qed

lemma heap_is_wellformed_children_in_heap:
  assumes "heap_is_wellformed h"
  assumes "h  get_child_nodes ptr r children"
  assumes "child  set children"
  shows "child |∈| node_ptr_kinds h"
  using assms
  apply(auto simp add: heap_is_wellformed_def a_all_ptrs_in_heap_def)[1]
  by (metis (no_types, opaque_lifting) is_OK_returns_result_I
      local.get_child_nodes_ptr_in_heap select_result_I2 subsetD)

lemma heap_is_wellformed_one_parent:
  assumes "heap_is_wellformed h"
  assumes "h  get_child_nodes ptr r children"
  assumes "h  get_child_nodes ptr' r children'"
  assumes "set children  set children'  {}"
  shows "ptr = ptr'"
  using assms
proof (auto simp add: heap_is_wellformed_def a_distinct_lists_def)[1]
  fix x :: "(_) node_ptr"
  assume a1: "ptr  ptr'"
  assume a2: "h  get_child_nodes ptr r children"
  assume a3: "h  get_child_nodes ptr' r children'"
  assume a4: "distinct (concat (map (λptr. |h  get_child_nodes ptr|r)
                                     (sorted_list_of_set (fset (object_ptr_kinds h)))))"
  have f5: "|h  get_child_nodes ptr|r = children"
    using a2 by simp
  have "|h  get_child_nodes ptr'|r = children'"
    using a3 by (meson select_result_I2)
  then have "ptr  set (sorted_list_of_set (fset (object_ptr_kinds h)))
           ptr'  set (sorted_list_of_set (fset (object_ptr_kinds h)))
           set children  set children' = {}"
    using f5 a4 a1 by (meson distinct_concat_map_E(1))
  then show False
    using a3 a2 by (metis (no_types) assms(4) finite_fset is_OK_returns_result_I
        local.get_child_nodes_ptr_in_heap set_sorted_list_of_set)
qed

lemma parent_child_rel_child:
  "h  get_child_nodes ptr r children  child  set children  (ptr, cast child)  parent_child_rel h"
  by (simp add: is_OK_returns_result_I get_child_nodes_ptr_in_heap parent_child_rel_def)

lemma parent_child_rel_acyclic: "heap_is_wellformed h  acyclic (parent_child_rel h)"
  by (simp add: acyclic_heap_def local.heap_is_wellformed_def)

lemma heap_is_wellformed_disconnected_nodes_distinct:
  "heap_is_wellformed h  h  get_disconnected_nodes document_ptr r disc_nodes  distinct disc_nodes"
  using distinct_lists_disconnected_nodes local.heap_is_wellformed_def by blast

lemma parent_child_rel_parent_in_heap:
  "(parent, child_ptr)  parent_child_rel h  parent |∈| object_ptr_kinds h"
  using local.parent_child_rel_def by blast

lemma parent_child_rel_child_in_heap:
  "heap_is_wellformed h  type_wf h  known_ptr parent
     (parent, child_ptr)  parent_child_rel h  child_ptr |∈| object_ptr_kinds h"
  apply(auto simp add: heap_is_wellformed_def parent_child_rel_def a_all_ptrs_in_heap_def)[1]
  using get_child_nodes_ok
  by (meson subsetD)

lemma heap_is_wellformed_disc_nodes_in_heap:
  "heap_is_wellformed h  h  get_disconnected_nodes document_ptr r disc_nodes
    node  set disc_nodes  node |∈| node_ptr_kinds h"
  by (metis (no_types, opaque_lifting) is_OK_returns_result_I local.a_all_ptrs_in_heap_def
      local.get_disconnected_nodes_ptr_in_heap local.heap_is_wellformed_def select_result_I2 subsetD)

lemma heap_is_wellformed_one_disc_parent:
  "heap_is_wellformed h  h  get_disconnected_nodes document_ptr r disc_nodes
    h  get_disconnected_nodes document_ptr' r disc_nodes'
    set disc_nodes  set disc_nodes'  {}  document_ptr = document_ptr'"
  using DocumentMonad.ptr_kinds_ptr_kinds_M concat_append distinct_append distinct_concat_map_E(1)
    is_OK_returns_result_I local.a_distinct_lists_def local.get_disconnected_nodes_ptr_in_heap
    local.heap_is_wellformed_def select_result_I2
proof -
  assume a1: "heap_is_wellformed h"
  assume a2: "h  get_disconnected_nodes document_ptr r disc_nodes"
  assume a3: "h  get_disconnected_nodes document_ptr' r disc_nodes'"
  assume a4: "set disc_nodes  set disc_nodes'  {}"
  have f5: "|h  get_disconnected_nodes document_ptr|r = disc_nodes"
    using a2 by (meson select_result_I2)
  have f6: "|h  get_disconnected_nodes document_ptr'|r = disc_nodes'"
    using a3 by (meson select_result_I2)
  have "nss nssa. ¬ distinct (concat (nss @ nssa))  distinct (concat nssa::(_) node_ptr list)"
    by (metis (no_types) concat_append distinct_append)
  then have "distinct (concat (map (λd. |h  get_disconnected_nodes d|r) |h  document_ptr_kinds_M|r))"
    using a1 local.a_distinct_lists_def local.heap_is_wellformed_def by blast
  then show ?thesis
    using f6 f5 a4 a3 a2 by (meson DocumentMonad.ptr_kinds_ptr_kinds_M distinct_concat_map_E(1)
        is_OK_returns_result_I local.get_disconnected_nodes_ptr_in_heap)
qed

lemma heap_is_wellformed_children_disc_nodes_different:
  "heap_is_wellformed h  h  get_child_nodes ptr r children
     h  get_disconnected_nodes document_ptr r disc_nodes
    set children  set disc_nodes = {}"
  by (metis (no_types, opaque_lifting) disjoint_iff_not_equal distinct_lists_no_parent
      is_OK_returns_result_I local.get_child_nodes_ptr_in_heap
      local.heap_is_wellformed_def select_result_I2)

lemma heap_is_wellformed_children_disc_nodes:
  "heap_is_wellformed h  node_ptr |∈| node_ptr_kinds h
     ¬(parent  fset (object_ptr_kinds h). node_ptr  set |h  get_child_nodes parent|r)
    (document_ptr  fset (document_ptr_kinds h). node_ptr  set |h  get_disconnected_nodes document_ptr|r)"
  by (auto simp add: heap_is_wellformed_def a_distinct_lists_def a_owner_document_valid_def)

lemma heap_is_wellformed_children_distinct:
  "heap_is_wellformed h  h  get_child_nodes ptr r children  distinct children"
  by (metis (no_types, lifting) ObjectMonad.ptr_kinds_ptr_kinds_M concat_append distinct_append
      distinct_concat_map_E(2) is_OK_returns_result_I local.a_distinct_lists_def
      local.get_child_nodes_ptr_in_heap local.heap_is_wellformed_def
      select_result_I2)
end

locale l_heap_is_wellformed = l_type_wf + l_known_ptr + l_heap_is_wellformed_defs
  + l_get_child_nodes_defs + l_get_disconnected_nodes_defs +
  assumes heap_is_wellformed_children_in_heap:
    "heap_is_wellformed h  h  get_child_nodes ptr r children  child  set children
                         child |∈| node_ptr_kinds h"
  assumes heap_is_wellformed_disc_nodes_in_heap:
    "heap_is_wellformed h  h  get_disconnected_nodes document_ptr r disc_nodes
                         node  set disc_nodes  node |∈| node_ptr_kinds h"
  assumes heap_is_wellformed_one_parent:
    "heap_is_wellformed h  h  get_child_nodes ptr r children
                         h  get_child_nodes ptr' r children'
                         set children  set children'  {}  ptr = ptr'"
  assumes heap_is_wellformed_one_disc_parent:
    "heap_is_wellformed h  h  get_disconnected_nodes document_ptr r disc_nodes
                         h  get_disconnected_nodes document_ptr' r disc_nodes'
                         set disc_nodes  set disc_nodes'  {}  document_ptr = document_ptr'"
  assumes heap_is_wellformed_children_disc_nodes_different:
    "heap_is_wellformed h  h  get_child_nodes ptr r children
                         h  get_disconnected_nodes document_ptr r disc_nodes
                         set children  set disc_nodes = {}"
  assumes heap_is_wellformed_disconnected_nodes_distinct:
    "heap_is_wellformed h  h  get_disconnected_nodes document_ptr r disc_nodes
                         distinct disc_nodes"
  assumes heap_is_wellformed_children_distinct:
    "heap_is_wellformed h  h  get_child_nodes ptr r children  distinct children"
  assumes heap_is_wellformed_children_disc_nodes:
    "heap_is_wellformed h  node_ptr |∈| node_ptr_kinds h
     ¬(parent  fset (object_ptr_kinds h). node_ptr  set |h  get_child_nodes parent|r)
    (document_ptr  fset (document_ptr_kinds h). node_ptr  set |h  get_disconnected_nodes document_ptr|r)"
  assumes parent_child_rel_child:
    "h  get_child_nodes ptr r children
    child  set children  (ptr, cast child)  parent_child_rel h"
  assumes parent_child_rel_finite:
    "heap_is_wellformed h  finite (parent_child_rel h)"
  assumes parent_child_rel_acyclic:
    "heap_is_wellformed h  acyclic (parent_child_rel h)"
  assumes parent_child_rel_node_ptr:
    "(parent, child_ptr)  parent_child_rel h  is_node_ptr_kind child_ptr"
  assumes parent_child_rel_parent_in_heap:
    "(parent, child_ptr)  parent_child_rel h  parent |∈| object_ptr_kinds h"
  assumes parent_child_rel_child_in_heap:
    "heap_is_wellformed h  type_wf h  known_ptr parent
    (parent, child_ptr)  parent_child_rel h  child_ptr |∈| object_ptr_kinds h"

interpretation i_heap_is_wellformed?: l_heap_is_wellformedCore_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
  apply(unfold_locales)
  by(auto simp add: heap_is_wellformed_def parent_child_rel_def)
declare l_heap_is_wellformedCore_DOM_axioms[instances]


lemma heap_is_wellformed_is_l_heap_is_wellformed [instances]:
  "l_heap_is_wellformed type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes
                        get_disconnected_nodes"
  apply(auto simp add: l_heap_is_wellformed_def)[1]
  using heap_is_wellformed_children_in_heap
                apply blast
  using heap_is_wellformed_disc_nodes_in_heap
               apply blast
  using heap_is_wellformed_one_parent
              apply blast
  using heap_is_wellformed_one_disc_parent
             apply blast
  using heap_is_wellformed_children_disc_nodes_different
            apply blast
  using heap_is_wellformed_disconnected_nodes_distinct
           apply blast
  using heap_is_wellformed_children_distinct
          apply blast
  using heap_is_wellformed_children_disc_nodes
         apply blast
  using parent_child_rel_child
        apply (blast)
  using parent_child_rel_child
       apply(blast)
  using parent_child_rel_finite
      apply blast
  using parent_child_rel_acyclic
     apply blast
  using parent_child_rel_node_ptr
    apply blast
  using parent_child_rel_parent_in_heap
   apply blast
  using parent_child_rel_child_in_heap
  apply blast
  done

subsection ‹get\_parent›

locale l_get_parent_wfCore_DOM =
  l_get_parentCore_DOM
  known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs
  + l_heap_is_wellformed
  type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs
  get_disconnected_nodes get_disconnected_nodes_locs
  for known_ptr :: "(_::linorder) object_ptr  bool"
    and type_wf :: "(_) heap  bool"
    and get_child_nodes :: "(_) object_ptr  ((_) heap, exception, (_) node_ptr list) prog"
    and get_child_nodes_locs :: "(_) object_ptr  ((_) heap  (_) heap  bool) set"
    and known_ptrs :: "(_) heap  bool"
    and get_parent :: "(_) node_ptr  ((_) heap, exception, (_) object_ptr option) prog"
    and get_parent_locs :: "((_) heap  (_) heap  bool) set"
    and heap_is_wellformed :: "(_) heap  bool"
    and parent_child_rel :: "(_) heap  ((_) object_ptr × (_) object_ptr) set"
    and get_disconnected_nodes :: "(_) document_ptr  ((_) heap, exception, (_) node_ptr list) prog"
    and get_disconnected_nodes_locs :: "(_) document_ptr  ((_) heap  (_) heap  bool) set"
begin
lemma child_parent_dual:
  assumes heap_is_wellformed: "heap_is_wellformed h"
  assumes "h  get_child_nodes ptr r children"
  assumes "child  set children"
  assumes "known_ptrs h"
  assumes type_wf: "type_wf h"
  shows "h  get_parent child r Some ptr"
proof -
  obtain ptrs where ptrs: "h  object_ptr_kinds_M r ptrs"
    by(simp add: object_ptr_kinds_M_defs)
  then have h1: "ptr  set ptrs"
    using get_child_nodes_ok assms(2) is_OK_returns_result_I
    by (metis (no_types, opaque_lifting) ObjectMonad.ptr_kinds_ptr_kinds_M
        thesis. (ptrs. h  object_ptr_kinds_M r ptrs  thesis)  thesis
        get_child_nodes_ptr_in_heap returns_result_eq select_result_I2)

  let ?P = "(λptr. get_child_nodes ptr  (λchildren. return (child  set children)))"
  let ?filter = "filter_M ?P ptrs"

  have "h  ok ?filter"
    using ptrs type_wf
    using get_child_nodes_ok
    apply(auto intro!: filter_M_is_OK_I bind_is_OK_pure_I get_child_nodes_ok simp add: bind_pure_I)[1]
    using assms(4) local.known_ptrs_known_ptr by blast
  then obtain parent_ptrs where parent_ptrs: "h  ?filter r parent_ptrs"
    by auto

  have h5: "∃!x. x  set ptrs  h  Heap_Error_Monad.bind (get_child_nodes x)
                  (λchildren. return (child  set children)) r True"
    apply(auto intro!: bind_pure_returns_result_I)[1]
    using heap_is_wellformed_one_parent
  proof -
    have "h  (return (child  set children)::((_) heap, exception, bool) prog) r True"
      by (simp add: assms(3))
    then show
      "z. z  set ptrs  h  Heap_Error_Monad.bind (get_child_nodes z)
                               (λns. return (child  set ns)) r True"
      by (metis (no_types) assms(2) bind_pure_returns_result_I2 h1 is_OK_returns_result_I
          local.get_child_nodes_pure select_result_I2)
  next
    fix x y
    assume 0: "x  set ptrs"
      and 1: "h  Heap_Error_Monad.bind (get_child_nodes x)
                  (λchildren. return (child  set children)) r True"
      and 2: "y  set ptrs"
      and 3: "h  Heap_Error_Monad.bind (get_child_nodes y)
                  (λchildren. return (child  set children)) r True"
      and 4: "(h ptr children ptr' children'. heap_is_wellformed h
               h  get_child_nodes ptr r children  h  get_child_nodes ptr' r children'
               set children  set children'  {}  ptr = ptr')"
    then show "x = y"
      by (metis (no_types, lifting) bind_returns_result_E disjoint_iff_not_equal heap_is_wellformed
          return_returns_result)
  qed

  have "child |∈| node_ptr_kinds h"
    using heap_is_wellformed_children_in_heap heap_is_wellformed assms(2) assms(3)
    by fast
  moreover have "parent_ptrs = [ptr]"
    apply(rule filter_M_ex1[OF parent_ptrs h1 h5])
    using ptrs assms(2) assms(3)
    by(auto simp add: object_ptr_kinds_M_defs bind_pure_I intro!: bind_pure_returns_result_I)
  ultimately  show ?thesis
    using ptrs parent_ptrs
    by(auto simp add: bind_pure_I get_parent_def
        elim!: bind_returns_result_E2
        intro!: bind_pure_returns_result_I filter_M_pure_I) (*slow, ca 1min *)
qed

lemma parent_child_rel_parent:
  assumes "heap_is_wellformed h"
    and "h  get_parent child_node r Some parent"
  shows "(parent, cast child_node)  parent_child_rel h"
  using assms parent_child_rel_child get_parent_child_dual by auto

lemma heap_wellformed_induct [consumes 1, case_names step]:
  assumes "heap_is_wellformed h"
    and step: "parent. (children child. h  get_child_nodes parent r children
                child  set children  P (cast child))  P parent"
  shows "P ptr"
proof -
  fix ptr
  have "wf ((parent_child_rel h)¯)"
    by (simp add: assms(1) finite_acyclic_wf_converse parent_child_rel_acyclic parent_child_rel_finite)
  then show "?thesis"
  proof (induct rule: wf_induct_rule)
    case (less parent)
    then show ?case
      using assms parent_child_rel_child
      by (meson converse_iff)
  qed
qed

lemma heap_wellformed_induct2 [consumes 3, case_names not_in_heap empty_children step]:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
    and not_in_heap: "parent. parent |∉| object_ptr_kinds h  P parent"
    and empty_children: "parent. h  get_child_nodes parent r []  P parent"
    and step: "parent children child. h  get_child_nodes parent r children
                child  set children  P (cast child)  P parent"
  shows "P ptr"
proof(insert assms(1), induct rule: heap_wellformed_induct)
  case (step parent)
  then show ?case
  proof(cases "parent |∈| object_ptr_kinds h")
    case True
    then obtain children where children: "h  get_child_nodes parent r children"
      using get_child_nodes_ok assms(2) assms(3)
      by (meson is_OK_returns_result_E local.known_ptrs_known_ptr)
    then show ?thesis
    proof (cases "children = []")
      case True
      then show ?thesis
        using children empty_children
        by simp
    next
      case False
      then show ?thesis
        using assms(6) children last_in_set step.hyps by blast
    qed
  next
    case False
    then show ?thesis
      by (simp add: not_in_heap)
  qed
qed

lemma heap_wellformed_induct_rev [consumes 1, case_names step]:
  assumes "heap_is_wellformed h"
    and step: "child. (parent child_node. cast child_node = child
                h  get_parent child_node r Some parent  P parent)  P child"
  shows "P ptr"
proof -
  fix ptr
  have "wf ((parent_child_rel h))"
    by (simp add: assms(1) local.parent_child_rel_acyclic local.parent_child_rel_finite
        wf_iff_acyclic_if_finite)

  then show "?thesis"
  proof (induct rule: wf_induct_rule)
    case (less child)
    show ?case
      using assms get_parent_child_dual
      by (metis less.hyps parent_child_rel_parent)
  qed
qed
end

interpretation i_get_parent_wf?: l_get_parent_wfCore_DOM known_ptr type_wf get_child_nodes
  get_child_nodes_locs known_ptrs get_parent get_parent_locs heap_is_wellformed
  parent_child_rel get_disconnected_nodes
  using instances
  by(simp add: l_get_parent_wfCore_DOM_def)
declare l_get_parent_wfCore_DOM_axioms[instances]


locale l_get_parent_wf2Core_DOM =
  l_get_parent_wfCore_DOM
  known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs
  heap_is_wellformed parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs
  + l_heap_is_wellformedCore_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
  for known_ptr :: "(_::linorder) object_ptr  bool"
    and type_wf :: "(_) heap  bool"
    and get_child_nodes :: "(_) object_ptr  ((_) heap, exception, (_) node_ptr list) prog"
    and get_child_nodes_locs :: "(_) object_ptr  ((_) heap  (_) heap  bool) set"
    and known_ptrs :: "(_) heap  bool"
    and get_parent :: "(_) node_ptr  ((_) heap, exception, (_) object_ptr option) prog"
    and get_parent_locs :: "((_) heap  (_) heap  bool) set"
    and heap_is_wellformed :: "(_) heap  bool"
    and parent_child_rel :: "(_) heap  ((_) object_ptr × (_) object_ptr) set"
    and get_disconnected_nodes :: "(_) document_ptr  ((_) heap, exception, (_) node_ptr list) prog"
    and get_disconnected_nodes_locs :: "(_) document_ptr  ((_) heap  (_) heap  bool) set"
begin
lemma preserves_wellformedness_writes_needed:
  assumes heap_is_wellformed: "heap_is_wellformed h"
    and "h  f h h'"
    and "writes SW f h h'"
    and preserved_get_child_nodes:
    "h h' w. w  SW  h  w h h'
                       object_ptr. r  get_child_nodes_locs object_ptr. r h h'"
    and preserved_get_disconnected_nodes:
    "h h' w. w  SW  h  w h h'
                       document_ptr. r  get_disconnected_nodes_locs document_ptr. r h h'"
    and preserved_object_pointers:
    "h h' w. w  SW  h  w h h'
                      object_ptr. preserved (get_MObject object_ptr RObject.nothing) h h'"
  shows "heap_is_wellformed h'"
proof -
  have object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'"
    using assms(2) assms(3) object_ptr_kinds_preserved preserved_object_pointers by blast
  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'"
    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'"
    by auto
  have children_eq:
    "ptr children. h  get_child_nodes ptr r children = h'  get_child_nodes ptr r children"
    apply(rule reads_writes_preserved[OF get_child_nodes_reads assms(3) assms(2)])
    using preserved_get_child_nodes by fast
  then have children_eq2: "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.
        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 assms(3) assms(2)])
    using preserved_get_disconnected_nodes by fast
  then have disconnected_nodes_eq2:
    "document_ptr. |h  get_disconnected_nodes document_ptr|r
                            = |h'  get_disconnected_nodes document_ptr|r"
    using select_result_eq by force
  have get_parent_eq: "ptr parent. h  get_parent ptr r parent = h'  get_parent ptr r parent"
    apply(rule reads_writes_preserved[OF get_parent_reads assms(3) assms(2)])
    using preserved_get_child_nodes preserved_object_pointers unfolding get_parent_locs_def by fast
  have "a_acyclic_heap h"
    using heap_is_wellformed by (simp add: heap_is_wellformed_def)
  have "parent_child_rel h'  parent_child_rel h"
  proof
    fix x
    assume "x  parent_child_rel h'"
    then show "x  parent_child_rel h"
      by(simp add: parent_child_rel_def  children_eq2 object_ptr_kinds_eq3)
  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 heap_is_wellformed by (simp add: heap_is_wellformed_def)
  then have "a_all_ptrs_in_heap h'"
    by (simp add: children_eq2 disconnected_nodes_eq2 document_ptr_kinds_eq3
        l_heap_is_wellformedCore_DOM_defs.a_all_ptrs_in_heap_def node_ptr_kinds_eq3 object_ptr_kinds_eq3)

  moreover have h0: "a_distinct_lists h"
    using heap_is_wellformed by (simp add: heap_is_wellformed_def)
  have h1: "map (λptr. |h  get_child_nodes ptr|r) (sorted_list_of_set (fset (object_ptr_kinds h)))
         = map (λptr. |h'  get_child_nodes ptr|r) (sorted_list_of_set (fset (object_ptr_kinds h')))"
    by (simp add: children_eq2 object_ptr_kinds_eq3)
  have h2: "map (λdocument_ptr. |h  get_disconnected_nodes document_ptr|r)
               (sorted_list_of_set (fset (document_ptr_kinds h)))
          = map (λdocument_ptr. |h'  get_disconnected_nodes document_ptr|r)
                (sorted_list_of_set (fset (document_ptr_kinds h')))"
    using disconnected_nodes_eq document_ptr_kinds_eq2 select_result_eq by force
  have "a_distinct_lists h'"
    using h0
    by(simp add: a_distinct_lists_def h1 h2)

  moreover have "a_owner_document_valid h"
    using heap_is_wellformed by (simp add: heap_is_wellformed_def)
  then have "a_owner_document_valid h'"
    by(auto simp add: a_owner_document_valid_def children_eq2 disconnected_nodes_eq2
        object_ptr_kinds_eq3 node_ptr_kinds_eq3 document_ptr_kinds_eq3)
  ultimately show ?thesis
    by (simp add: heap_is_wellformed_def)
qed
end

interpretation i_get_parent_wf2?: l_get_parent_wf2Core_DOM known_ptr type_wf get_child_nodes
  get_child_nodes_locs known_ptrs get_parent get_parent_locs
  heap_is_wellformed parent_child_rel get_disconnected_nodes
  get_disconnected_nodes_locs
  using l_get_parent_wfCore_DOM_axioms l_heap_is_wellformedCore_DOM_axioms
  by (simp add: l_get_parent_wf2Core_DOM_def)

declare l_get_parent_wf2Core_DOM_axioms[instances]
locale l_get_parent_wf = l_type_wf + l_known_ptrs  + l_heap_is_wellformed_defs
  + l_get_child_nodes_defs + l_get_parent_defs +
  assumes child_parent_dual:
    "heap_is_wellformed h
     type_wf h
     known_ptrs h
     h  get_child_nodes ptr r children
     child  set children
     h  get_parent child r Some ptr"
  assumes heap_wellformed_induct [consumes 1, case_names step]:
    "heap_is_wellformed h
     (parent. (children child. h  get_child_nodes parent r children
     child  set children  P (cast child))  P parent)
     P ptr"
  assumes heap_wellformed_induct_rev [consumes 1, case_names step]:
    "heap_is_wellformed h
     (child. (parent child_node. cast child_node = child
     h  get_parent child_node r Some parent  P parent)  P child)
     P ptr"
  assumes parent_child_rel_parent: "heap_is_wellformed h
     h  get_parent child_node r Some parent
     (parent, cast child_node)  parent_child_rel h"

lemma get_parent_wf_is_l_get_parent_wf [instances]:
  "l_get_parent_wf type_wf known_ptr known_ptrs heap_is_wellformed parent_child_rel
  get_child_nodes get_parent"
  using known_ptrs_is_l_known_ptrs
  apply(auto simp add: l_get_parent_wf_def l_get_parent_wf_axioms_def)[1]
  using child_parent_dual heap_wellformed_induct heap_wellformed_induct_rev parent_child_rel_parent
  by metis+



subsection ‹get\_disconnected\_nodes›



subsection ‹set\_disconnected\_nodes›


subsubsection ‹get\_disconnected\_nodes›

locale l_set_disconnected_nodes_get_disconnected_nodes_wfCore_DOM =
  l_set_disconnected_nodes_get_disconnected_nodes
  type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
  set_disconnected_nodes_locs
  + l_heap_is_wellformed
  type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs
  get_disconnected_nodes get_disconnected_nodes_locs
  for known_ptr :: "(_) object_ptr  bool"
    and type_wf :: "(_) heap  bool"
    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 heap_is_wellformed :: "(_) heap  bool"
    and parent_child_rel :: "(_) heap  ((_) object_ptr × (_) object_ptr) set"
    and get_child_nodes :: "(_) object_ptr  ((_) heap, exception, (_) node_ptr list) prog"
    and get_child_nodes_locs :: "(_) object_ptr  ((_) heap  (_) heap  bool) set"
begin

lemma remove_from_disconnected_nodes_removes:
  assumes "heap_is_wellformed h"
  assumes "h  get_disconnected_nodes ptr r disc_nodes"
  assumes "h  set_disconnected_nodes ptr (remove1 node_ptr disc_nodes) h h'"
  assumes "h'  get_disconnected_nodes ptr r disc_nodes'"
  shows "node_ptr  set disc_nodes'"
  using assms
  by (metis distinct_remove1_removeAll heap_is_wellformed_disconnected_nodes_distinct
      set_disconnected_nodes_get_disconnected_nodes member_remove remove_code(1)
      returns_result_eq)
end

locale l_set_disconnected_nodes_get_disconnected_nodes_wf = l_heap_is_wellformed
  + l_set_disconnected_nodes_get_disconnected_nodes +
  assumes remove_from_disconnected_nodes_removes:
    "heap_is_wellformed h  h  get_disconnected_nodes ptr r disc_nodes
                           h  set_disconnected_nodes ptr (remove1 node_ptr disc_nodes) h h'
                           h'  get_disconnected_nodes ptr r disc_nodes'
                           node_ptr  set disc_nodes'"

interpretation i_set_disconnected_nodes_get_disconnected_nodes_wfCore_DOM?:
  l_set_disconnected_nodes_get_disconnected_nodes_wfCore_DOM  known_ptr type_wf get_disconnected_nodes
  get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs heap_is_wellformed
  parent_child_rel get_child_nodes
  using instances
  by (simp add: l_set_disconnected_nodes_get_disconnected_nodes_wfCore_DOM_def)
declare l_set_disconnected_nodes_get_disconnected_nodes_wfCore_DOM_axioms[instances]

lemma set_disconnected_nodes_get_disconnected_nodes_wf_is_l_set_disconnected_nodes_get_disconnected_nodes_wf [instances]:
  "l_set_disconnected_nodes_get_disconnected_nodes_wf type_wf known_ptr heap_is_wellformed parent_child_rel
   get_child_nodes get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
      set_disconnected_nodes_locs"
  apply(auto simp add: l_set_disconnected_nodes_get_disconnected_nodes_wf_def
      l_set_disconnected_nodes_get_disconnected_nodes_wf_axioms_def instances)[1]
  using remove_from_disconnected_nodes_removes apply fast
  done


subsection ‹get\_root\_node›

locale l_get_root_node_wfCore_DOM =
  l_heap_is_wellformed
  type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs
  get_disconnected_nodes get_disconnected_nodes_locs
  + l_get_parentCore_DOM
  known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs
  + l_get_parent_wf
  type_wf known_ptr known_ptrs heap_is_wellformed parent_child_rel get_child_nodes
  get_child_nodes_locs get_parent get_parent_locs
  + l_get_root_nodeCore_DOM
  type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs
  get_ancestors get_ancestors_locs get_root_node get_root_node_locs
  for known_ptr :: "(_::linorder) object_ptr  bool"
    and type_wf :: "(_) heap  bool"
    and known_ptrs :: "(_) heap  bool"
    and heap_is_wellformed :: "(_) heap  bool"
    and parent_child_rel :: "(_) heap  ((_) object_ptr × (_) object_ptr) 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_disconnected_nodes :: "(_) document_ptr  ((_) heap, exception, (_) node_ptr list) prog"
    and get_disconnected_nodes_locs :: "(_) document_ptr  ((_) heap  (_) heap  bool) set"
    and get_parent :: "(_) node_ptr  ((_) heap, exception, (_) object_ptr option) prog"
    and get_parent_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_root_node :: "(_) object_ptr  ((_) heap, exception, (_) object_ptr) prog"
    and get_root_node_locs :: "((_) heap  (_) heap  bool) set"

begin
lemma get_ancestors_reads:
  assumes "heap_is_wellformed h"
  shows "reads get_ancestors_locs (get_ancestors node_ptr) h h'"
proof (insert assms(1), induct rule: heap_wellformed_induct_rev)
  case (step child)
  then show ?case
    using [[simproc del: Product_Type.unit_eq]] get_parent_reads[unfolded reads_def]
    apply(simp (no_asm) add: get_ancestors_def)
    by(auto simp add: get_ancestors_locs_def reads_subset[OF return_reads] get_parent_reads_pointers
        intro!: reads_bind_pure reads_subset[OF check_in_heap_reads]
        reads_subset[OF get_parent_reads] reads_subset[OF get_child_nodes_reads]
        split: option.splits)
qed

lemma get_ancestors_ok:
  assumes "heap_is_wellformed h"
    and "ptr |∈| object_ptr_kinds h"
    and "known_ptrs h"
    and type_wf: "type_wf h"
  shows "h  ok (get_ancestors ptr)"
proof (insert assms(1) assms(2), induct rule: heap_wellformed_induct_rev)
  case (step child)
  then show ?case
    using assms(3) assms(4)
    apply(simp (no_asm) add: get_ancestors_def)
    apply(simp add: assms(1) get_parent_parent_in_heap)
    by(auto intro!: bind_is_OK_pure_I bind_pure_I get_parent_ok split: option.splits)
qed

lemma get_root_node_ptr_in_heap:
  assumes "h  ok (get_root_node ptr)"
  shows "ptr |∈| object_ptr_kinds h"
  using assms
  unfolding get_root_node_def
  using get_ancestors_ptr_in_heap
  by auto


lemma get_root_node_ok:
  assumes "heap_is_wellformed h" "known_ptrs h" "type_wf h"
    and "ptr |∈| object_ptr_kinds h"
  shows "h  ok (get_root_node ptr)"
  unfolding get_root_node_def
  using assms get_ancestors_ok
  by auto


lemma get_ancestors_parent:
  assumes "heap_is_wellformed h"
    and "h  get_parent child r Some parent"
  shows "h  get_ancestors (cast child) r (cast child) # parent # ancestors
      h  get_ancestors parent r parent # ancestors"
proof
  assume a1: "h  get_ancestors (castnode_ptr2object_ptr child) r castnode_ptr2object_ptr child # parent # ancestors"
  then have "h  Heap_Error_Monad.bind (check_in_heap (castnode_ptr2object_ptr child))
          (λ_. Heap_Error_Monad.bind (get_parent child)
              (λx. Heap_Error_Monad.bind (case x of None  return [] | Some x  get_ancestors x)
               (λancestors. return (castnode_ptr2object_ptr child # ancestors))))
        r castnode_ptr2object_ptr child # parent # ancestors"
    by(simp add: get_ancestors_def)
  then show "h  get_ancestors parent r parent # ancestors"
    using assms(2) apply(auto elim!: bind_returns_result_E2 split: option.splits)[1]
    using returns_result_eq by fastforce
next
  assume "h  get_ancestors parent r parent # ancestors"
  then show "h  get_ancestors (castnode_ptr2object_ptr child) r castnode_ptr2object_ptr child # parent # ancestors"
    using assms(2)
    apply(simp (no_asm) add: get_ancestors_def)
    apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1]
    by (metis (full_types) assms(2) check_in_heap_ptr_in_heap is_OK_returns_result_I
        local.get_parent_ptr_in_heap node_ptr_kinds_commutes old.unit.exhaust
        select_result_I)
qed


lemma get_ancestors_never_empty:
  assumes "heap_is_wellformed h"
    and "h  get_ancestors child r ancestors"
  shows "ancestors  []"
proof(insert assms(2), induct arbitrary: ancestors rule: heap_wellformed_induct_rev[OF assms(1)])
  case (1 child)
  then show ?case
  proof (induct "castobject_ptr2node_ptr child")
    case None
    then show ?case
      apply(simp add: get_ancestors_def)
      by(auto elim!: bind_returns_result_E2 split: option.splits)
  next
    case (Some child_node)
    then obtain parent_opt where parent_opt: "h  get_parent child_node r parent_opt"
      apply(simp add: get_ancestors_def)
      by(auto elim!: bind_returns_result_E2 split: option.splits)
    with Some show ?case
    proof(induct parent_opt)
      case None
      then show ?case
        apply(simp add: get_ancestors_def)
        by(auto elim!: bind_returns_result_E2 split: option.splits)
    next
      case (Some option)
      then show ?case
        apply(simp add: get_ancestors_def)
        by(auto elim!: bind_returns_result_E2 split: option.splits)
    qed
  qed
qed



lemma get_ancestors_subset:
  assumes "heap_is_wellformed h"
    and "h  get_ancestors ptr r ancestors"
    and "ancestor  set ancestors"
    and "h  get_ancestors ancestor r ancestor_ancestors"
    and type_wf: "type_wf h"
    and known_ptrs: "known_ptrs h"
  shows "set ancestor_ancestors  set ancestors"
proof (insert assms(1) assms(2) assms(3), induct ptr arbitrary: ancestors
    rule: heap_wellformed_induct_rev)
  case (step child)
  have "child |∈| object_ptr_kinds h"
    using get_ancestors_ptr_in_heap step(2) by auto
      (*  then have "h ⊢ check_in_heap child →r ()"
    using returns_result_select_result by force *)
  show ?case
  proof (induct "castobject_ptr2node_ptr child")
    case None
    then have "ancestors = [child]"
      using step(2) step(3)
      by(auto simp add: get_ancestors_def elim!: bind_returns_result_E2)
    show ?case
      using step(2) step(3)
      apply(auto simp add: ancestors = [child])[1]
      using assms(4) returns_result_eq by fastforce
  next
    case (Some child_node)
    note s1 = Some
    obtain parent_opt where parent_opt: "h  get_parent child_node r parent_opt"
      using child |∈| object_ptr_kinds h assms(1) Some[symmetric] get_parent_ok[OF type_wf known_ptrs]
      by (metis (no_types, lifting) is_OK_returns_result_E known_ptrs get_parent_ok
          l_get_parentCore_DOM_axioms node_ptr_casts_commute node_ptr_kinds_commutes)
    then show ?case
    proof (induct parent_opt)
      case None
      then have "ancestors = [child]"
        using step(2) step(3) s1
        apply(simp add: get_ancestors_def)
        by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq)
      show ?case
        using step(2) step(3)
        apply(auto simp add: ancestors = [child])[1]
        using assms(4) returns_result_eq by fastforce
    next
      case (Some parent)
      have "h  Heap_Error_Monad.bind (check_in_heap child)
          (λ_. Heap_Error_Monad.bind
                 (case castobject_ptr2node_ptr child of None  return []
                  | Some node_ptr  Heap_Error_Monad.bind (get_parent node_ptr)
                                         (λparent_ptr_opt. case parent_ptr_opt of None  return []
                                                                        | Some x  get_ancestors x))
                 (λancestors. return (child # ancestors)))
    r  ancestors"
        using step(2)
        by(simp add: get_ancestors_def)
      moreover obtain tl_ancestors where tl_ancestors: "ancestors = child # tl_ancestors"
        using calculation
        by(auto elim!: bind_returns_result_E2 split: option.splits)
      ultimately have "h  get_ancestors parent r tl_ancestors"
        using s1 Some
        by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq)
      show ?case
        using step(1)[OF s1[symmetric, simplified] Some h  get_ancestors parent r tl_ancestors]
          step(3)
        apply(auto simp add: tl_ancestors)[1]
        by (metis assms(4) insert_iff list.simps(15) local.step(2) returns_result_eq tl_ancestors)
    qed
  qed
qed

lemma get_ancestors_also_parent:
  assumes "heap_is_wellformed h"
    and "h  get_ancestors some_ptr r ancestors"
    and "cast child  set ancestors"
    and "h  get_parent child r Some parent"
    and type_wf: "type_wf h"
    and known_ptrs: "known_ptrs h"
  shows "parent  set ancestors"
proof -
  obtain child_ancestors where child_ancestors: "h  get_ancestors (cast child) r child_ancestors"
    by (meson assms(1) assms(4) get_ancestors_ok is_OK_returns_result_I known_ptrs
        local.get_parent_ptr_in_heap node_ptr_kinds_commutes returns_result_select_result
        type_wf)
  then have "parent  set child_ancestors"
    apply(simp add: get_ancestors_def)
    by(auto elim!: bind_returns_result_E2 split: option.splits dest!: returns_result_eq[OF assms(4)]
        get_ancestors_ptr)
  then show ?thesis
    using assms child_ancestors get_ancestors_subset by blast
qed

lemma get_ancestors_obtains_children:
  assumes "heap_is_wellformed h"
    and "ancestor  ptr"
    and "ancestor  set ancestors"
    and "h  get_ancestors ptr r ancestors"
    and type_wf: "type_wf h"
    and known_ptrs: "known_ptrs h"
  obtains children ancestor_child where "h  get_child_nodes ancestor r children"
    and "ancestor_child  set children" and "cast ancestor_child  set ancestors"
proof -
  assume 0: "(children ancestor_child.
        h  get_child_nodes ancestor r children 
        ancestor_child  set children  castnode_ptr2object_ptr ancestor_child  set ancestors
         thesis)"
  have "child. h  get_parent child r Some ancestor  cast child  set ancestors"
  proof (insert assms(1) assms(2) assms(3) assms(4), induct ptr arbitrary: ancestors
      rule: heap_wellformed_induct_rev)
    case (step child)
    have "child |∈| object_ptr_kinds h"
      using get_ancestors_ptr_in_heap step(4) by auto
    show ?case
    proof (induct "castobject_ptr2node_ptr child")
      case None
      then have "ancestors = [child]"
        using step(3) step(4)
        by(auto simp add: get_ancestors_def elim!: bind_returns_result_E2)
      show ?case
        using step(2) step(3) step(4)
        by(auto simp add: ancestors = [child])
    next
      case (Some child_node)
      note s1 = Some
      obtain parent_opt where parent_opt: "h  get_parent child_node r parent_opt"
        using child |∈| object_ptr_kinds h assms(1) Some[symmetric]
        using get_parent_ok known_ptrs type_wf
        by (metis (no_types, lifting) is_OK_returns_result_E node_ptr_casts_commute
            node_ptr_kinds_commutes)
      then show ?case
      proof (induct parent_opt)
        case None
        then have "ancestors = [child]"
          using step(2) step(3) step(4) s1
          apply(simp add: get_ancestors_def)
          by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq)
        show ?case
          using step(2) step(3) step(4)
          by(auto simp add: ancestors = [child])
      next
        case (Some parent)
        have "h  Heap_Error_Monad.bind (check_in_heap child)
            (λ_. Heap_Error_Monad.bind
                   (case castobject_ptr2node_ptr child of None  return []
                    | Some node_ptr  Heap_Error_Monad.bind (get_parent node_ptr)
                       (λparent_ptr_opt. case parent_ptr_opt of None  return []
                                                              | Some x  get_ancestors x))
                   (λancestors. return (child # ancestors)))
      r  ancestors"
          using step(4)
          by(simp add: get_ancestors_def)
        moreover obtain tl_ancestors where tl_ancestors: "ancestors = child # tl_ancestors"
          using calculation
          by(auto elim!: bind_returns_result_E2 split: option.splits)
        ultimately have "h  get_ancestors parent r tl_ancestors"
          using s1 Some
          by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq)
            (* have "ancestor ≠ parent" *)
        have "ancestor  set tl_ancestors"
          using tl_ancestors step(2) step(3) by auto
        show ?case
        proof (cases "ancestor  parent")
          case True
          show ?thesis
            using step(1)[OF s1[symmetric, simplified] Some True
                ancestor  set tl_ancestors h  get_ancestors parent r tl_ancestors]
            using tl_ancestors by auto
        next
          case False
          have "child  set ancestors"
            using step(4) get_ancestors_ptr by simp
          then show ?thesis
            using Some False s1[symmetric] by(auto)
        qed
      qed
    qed
  qed
  then obtain child where child: "h  get_parent child r Some ancestor"
    and in_ancestors: "cast child  set ancestors"
    by auto
  then obtain children where
    children: "h  get_child_nodes ancestor r children" and
    child_in_children: "child  set children"
    using get_parent_child_dual by blast
  show thesis
    using 0[OF children child_in_children] child assms(3) in_ancestors by blast
qed

lemma get_ancestors_parent_child_rel:
  assumes "heap_is_wellformed h"
    and "h  get_ancestors child r ancestors"
    and known_ptrs: "known_ptrs h"
    and type_wf: "type_wf h"
  shows "(ptr, child)  (parent_child_rel h)*  ptr  set ancestors"
proof (safe)
  assume 3: "(ptr, child)  (parent_child_rel h)*"
  show "ptr  set ancestors"
  proof (insert 3, induct ptr rule: heap_wellformed_induct[OF assms(1)])
    case (1 ptr)
    then show ?case
    proof (cases "ptr = child")
      case True
      then show ?thesis
        by (metis (no_types, lifting) assms(2) bind_returns_result_E get_ancestors_def
            in_set_member member_rec(1) return_returns_result)
    next
      case False
      obtain ptr_child where
        ptr_child: "(ptr, ptr_child)  (parent_child_rel h)  (ptr_child, child)  (parent_child_rel h)*"
        using converse_rtranclE[OF 1(2)] ptr  child
        by metis
      then obtain ptr_child_node
        where ptr_child_ptr_child_node: "ptr_child = castnode_ptr2object_ptr ptr_child_node"
        using ptr_child node_ptr_casts_commute3 parent_child_rel_node_ptr
        by (metis )
      then obtain children where
        children: "h  get_child_nodes ptr r children" and
        ptr_child_node: "ptr_child_node  set children"
      proof -
        assume a1: "children. h  get_child_nodes ptr r children; ptr_child_node  set children
                    thesis"

        have "ptr |∈| object_ptr_kinds h"
          using local.parent_child_rel_parent_in_heap ptr_child by blast
        moreover have "ptr_child_node  set |h  get_child_nodes ptr|r"
          by (metis calculation known_ptrs local.get_child_nodes_ok local.known_ptrs_known_ptr
              local.parent_child_rel_child ptr_child ptr_child_ptr_child_node
              returns_result_select_result type_wf)
        ultimately show ?thesis
          using a1 get_child_nodes_ok type_wf known_ptrs
          by (meson local.known_ptrs_known_ptr returns_result_select_result)
      qed
      moreover have "(castnode_ptr2object_ptr ptr_child_node, child)  (parent_child_rel h)*"
        using ptr_child ptr_child_ptr_child_node by auto
      ultimately have "castnode_ptr2object_ptr ptr_child_node  set ancestors"
        using 1 by auto
      moreover have "h  get_parent ptr_child_node r Some ptr"
        using assms(1) children ptr_child_node child_parent_dual
        using known_ptrs type_wf by blast
      ultimately show ?thesis
        using get_ancestors_also_parent assms type_wf by blast
    qed
  qed
next
  assume 3: "ptr  set ancestors"
  show "(ptr, child)  (parent_child_rel h)*"
  proof (insert 3, induct ptr rule: heap_wellformed_induct[OF assms(1)])
    case (1 ptr)
    then show ?case
    proof (cases "ptr = child")
      case True
      then show ?thesis
        by simp
    next
      case False
      then obtain children ptr_child_node where
        children: "h  get_child_nodes ptr r children" and
        ptr_child_node: "ptr_child_node  set children" and
        ptr_child_node_in_ancestors: "cast ptr_child_node  set ancestors"
        using 1(2) assms(2) get_ancestors_obtains_children assms(1)
        using known_ptrs type_wf by blast
      then have "(castnode_ptr2object_ptr ptr_child_node, child)  (parent_child_rel h)*"
        using 1(1) by blast

      moreover have "(ptr, cast ptr_child_node)  parent_child_rel h"
        using children ptr_child_node assms(1) parent_child_rel_child_nodes2
        using child_parent_dual known_ptrs parent_child_rel_parent type_wf
        by blast

      ultimately show ?thesis
        by auto
    qed
  qed
qed

lemma get_root_node_parent_child_rel:
  assumes "heap_is_wellformed h"
    and "h  get_root_node child r root"
    and known_ptrs: "known_ptrs h"
    and type_wf: "type_wf h"
  shows "(root, child)  (parent_child_rel h)*"
  using assms get_ancestors_parent_child_rel
  apply(auto simp add: get_root_node_def elim!: bind_returns_result_E2)[1]
  using get_ancestors_never_empty last_in_set by blast


lemma get_ancestors_eq:
  assumes "heap_is_wellformed h"
    and "heap_is_wellformed h'"
    and "object_ptr w. object_ptr  ptr  w  get_child_nodes_locs object_ptr  w h h'"
    and pointers_preserved: "object_ptr. preserved (get_MObject object_ptr RObject.nothing) h h'"
    and known_ptrs: "known_ptrs h"
    and known_ptrs': "known_ptrs h'"
    and "h  get_ancestors ptr r ancestors"
    and type_wf: "type_wf h"
    and type_wf': "type_wf h'"
  shows "h'  get_ancestors ptr r ancestors"
proof -
  have object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'"
    using pointers_preserved object_ptr_kinds_preserved_small by blast
  then have object_ptr_kinds_M_eq:
    "ptrs. h  object_ptr_kinds_M r ptrs = h'  object_ptr_kinds_M r ptrs"
    by(simp add: object_ptr_kinds_M_defs)
  then have object_ptr_kinds_eq: "|h  object_ptr_kinds_M|r = |h'  object_ptr_kinds_M|r"
    by(simp)
  have "h'  ok (get_ancestors ptr)"
    using get_ancestors_ok get_ancestors_ptr_in_heap object_ptr_kinds_eq3 assms(1) known_ptrs
      known_ptrs' assms(2) assms(7) type_wf'
    by blast
  then obtain ancestors' where ancestors': "h'  get_ancestors ptr r ancestors'"
    by auto

  obtain root where root: "h  get_root_node ptr r root"
  proof -
    assume 0: "(root. h  get_root_node ptr r root  thesis)"
    show thesis
      apply(rule 0)
      using assms(7)
      by(auto simp add: get_root_node_def elim!: bind_returns_result_E2 split: option.splits)
  qed

  have children_eq:
    "p children. p  ptr  h  get_child_nodes p r children = h'  get_child_nodes p r children"
    using get_child_nodes_reads assms(3)
    apply(simp add: reads_def reflp_def transp_def preserved_def)
    by blast

  have "acyclic (parent_child_rel h)"
    using assms(1) local.parent_child_rel_acyclic by auto
  have "acyclic (parent_child_rel h')"
    using assms(2) local.parent_child_rel_acyclic by blast
  have 2: "c parent_opt. castnode_ptr2object_ptr c  set ancestors  set ancestors'
           h  get_parent c r parent_opt = h'  get_parent c r parent_opt"
  proof -
    fix c parent_opt
    assume 1: " castnode_ptr2object_ptr c  set ancestors  set ancestors'"

    obtain ptrs where ptrs: "h  object_ptr_kinds_M r ptrs"
      by simp

    let ?P = "(λptr. Heap_Error_Monad.bind (get_child_nodes ptr) (λchildren. return (c  set children)))"
    have children_eq_True: "p. p  set ptrs  h  ?P p r True  h'  ?P p r True"
    proof -
      fix p
      assume "p  set ptrs"
      then show "h  ?P p r True  h'  ?P p r True"
      proof (cases "p = ptr")
        case True
        have "(castnode_ptr2object_ptr c, ptr)  (parent_child_rel h)*"
          using get_ancestors_parent_child_rel 1 assms by blast
        then have "(ptr, castnode_ptr2object_ptr c)  (parent_child_rel h)"
        proof (cases "cast c = ptr")
          case True
          then show ?thesis
            using acyclic (parent_child_rel h) by(auto simp add: acyclic_def)
        next
          case False
          then have "(ptr, castnode_ptr2object_ptr c)  (parent_child_rel h)*"
            using acyclic (parent_child_rel h) False rtrancl_eq_or_trancl rtrancl_trancl_trancl
              (castnode_ptr2object_ptr c, ptr)  (parent_child_rel h)*
            by (metis acyclic_def)
          then show ?thesis
            using r_into_rtrancl by auto
        qed
        obtain children where children: "h  get_child_nodes ptr r children"
          using type_wf
          by (metis h'  ok get_ancestors ptr assms(1) get_ancestors_ptr_in_heap get_child_nodes_ok
              heap_is_wellformed_def is_OK_returns_result_E known_ptrs local.known_ptrs_known_ptr
              object_ptr_kinds_eq3)
        then have "c  set children"
          using (ptr, castnode_ptr2object_ptr c)  (parent_child_rel h) assms(1)
          using parent_child_rel_child_nodes2
          using child_parent_dual known_ptrs parent_child_rel_parent
            type_wf by blast
        with children have "h  ?P p r False"
          by(auto simp add: True)

        moreover have "(castnode_ptr2object_ptr c, ptr)  (parent_child_rel h')*"
          using get_ancestors_parent_child_rel assms(2) ancestors' 1 known_ptrs' type_wf
            type_wf' by blast
        then have "(ptr, castnode_ptr2object_ptr c)  (parent_child_rel h')"
        proof (cases "cast c = ptr")
          case True
          then show ?thesis
            using acyclic (parent_child_rel h') by(auto simp add: acyclic_def)
        next
          case False
          then have "(ptr, castnode_ptr2object_ptr c)  (parent_child_rel h')*"
            using acyclic (parent_child_rel h') False rtrancl_eq_or_trancl rtrancl_trancl_trancl
              (castnode_ptr2object_ptr c, ptr)  (parent_child_rel h')*
            by (metis acyclic_def)
          then show ?thesis
            using r_into_rtrancl by auto
        qed
        then have "(ptr, castnode_ptr2object_ptr c)  (parent_child_rel h')"
          using r_into_rtrancl by auto
        obtain children' where children': "h'  get_child_nodes ptr r children'"
          using  type_wf  type_wf'
          by (meson h'  ok (get_ancestors ptr) assms(2) get_ancestors_ptr_in_heap
              get_child_nodes_ok is_OK_returns_result_E known_ptrs'
              local.known_ptrs_known_ptr)
        then have "c  set children'"
          using (ptr, castnode_ptr2object_ptr c)  (parent_child_rel h') assms(2)  type_wf  type_wf'
          using parent_child_rel_child_nodes2 child_parent_dual known_ptrs' parent_child_rel_parent
          by auto
        with children' have "h'  ?P p r False"
          by(auto simp add: True)

        ultimately show ?thesis
          by (metis returns_result_eq)
      next
        case False
        then show ?thesis
          using children_eq ptrs
          by (metis (no_types, lifting) bind_pure_returns_result_I bind_returns_result_E
              get_child_nodes_pure return_returns_result)
      qed
    qed
    have "pa. pa  set ptrs  h  ok (get_child_nodes pa
                      (λchildren. return (c  set children))) = h'  ok ( get_child_nodes pa
                      (λchildren. return (c  set children)))"
      using assms(1) assms(2) object_ptr_kinds_eq ptrs type_wf  type_wf'
      by (metis (no_types, lifting) ObjectMonad.ptr_kinds_ptr_kinds_M bind_is_OK_pure_I
          get_child_nodes_ok get_child_nodes_pure known_ptrs'
          local.known_ptrs_known_ptr return_ok select_result_I2)
    have children_eq_False:
      "pa. pa  set ptrs  h  get_child_nodes pa
                   (λchildren. return (c  set children)) r False = h'  get_child_nodes pa
                   (λchildren. return (c  set children)) r False"
    proof
      fix pa
      assume "pa  set ptrs"
        and "h  get_child_nodes pa  (λchildren. return (c  set children)) r False"
      have "h  ok (get_child_nodes pa  (λchildren. return (c  set children)))
            h'  ok ( get_child_nodes pa  (λchildren. return (c  set children)))"
        using pa  set ptrs pa. pa  set ptrs  h  ok (get_child_nodes pa
                     (λchildren. return (c  set children))) = h'  ok ( get_child_nodes pa
                     (λchildren. return (c  set children)))
        by auto
      moreover have "h  get_child_nodes pa  (λchildren. return (c  set children)) r False
         h'  get_child_nodes pa  (λchildren. return (c  set children)) r False"
        by (metis (mono_tags, lifting) pa. pa  set ptrs
               h  get_child_nodes pa
                   (λchildren. return (c  set children)) r True = h'  get_child_nodes pa
                   (λchildren. return (c  set children)) r True pa  set ptrs
            calculation is_OK_returns_result_I returns_result_eq returns_result_select_result)
      ultimately show "h'  get_child_nodes pa  (λchildren. return (c  set children)) r False"
        using h  get_child_nodes pa  (λchildren. return (c  set children)) r False
        by auto
    next
      fix pa
      assume "pa  set ptrs"
        and "h'  get_child_nodes pa  (λchildren. return (c  set children)) r False"
      have "h'  ok (get_child_nodes pa  (λchildren. return (c  set children)))
             h  ok ( get_child_nodes pa  (λchildren. return (c  set children)))"
        using pa  set ptrs pa. pa  set ptrs
               h  ok (get_child_nodes pa
                   (λchildren. return (c  set children))) = h'  ok ( get_child_nodes pa
                   (λchildren. return (c  set children)))
        by auto
      moreover have "h'  get_child_nodes pa  (λchildren. return (c  set children)) r False
                   h  get_child_nodes pa  (λchildren. return (c  set children)) r False"
        by (metis (mono_tags, lifting)
            pa. pa  set ptrs  h  get_child_nodes pa
                         (λchildren. return (c  set children)) r True = h'  get_child_nodes pa
                         (λchildren. return (c  set children)) r True pa  set ptrs
            calculation is_OK_returns_result_I returns_result_eq returns_result_select_result)
      ultimately show "h  get_child_nodes pa  (λchildren. return (c  set children)) r False"
        using h'  get_child_nodes pa  (λchildren. return (c  set children)) r False by blast
    qed

    have filter_eq: "xs. h  filter_M ?P ptrs r xs = h'  filter_M ?P ptrs r xs"
    proof (rule filter_M_eq)
      show
        "xs x. pure (Heap_Error_Monad.bind (get_child_nodes x) (λchildren. return (c  set children))) h"
        by(auto intro!: bind_pure_I)
    next
      show
        "xs x. pure (Heap_Error_Monad.bind (get_child_nodes x) (λchildren. return (c  set children))) h'"
        by(auto intro!: bind_pure_I)
    next
      fix xs b x
      assume 0: "x  set ptrs"
      then show "h  Heap_Error_Monad.bind (get_child_nodes x) (λchildren. return (c  set children)) r b
              = h'  Heap_Error_Monad.bind (get_child_nodes x) (λchildren. return (c  set children)) r b"
        apply(induct b)
        using children_eq_True apply blast
        using children_eq_False apply blast
        done
    qed

    show "h  get_parent c r parent_opt = h'  get_parent c r parent_opt"
      apply(simp add: get_parent_def)
      apply(rule bind_cong_2)
         apply(simp)
        apply(simp)
       apply(simp add: check_in_heap_def node_ptr_kinds_def object_ptr_kinds_eq3)
      apply(rule bind_cong_2)
         apply(auto simp add: object_ptr_kinds_M_eq object_ptr_kinds_eq3)[1]
        apply(auto simp add: object_ptr_kinds_M_eq object_ptr_kinds_eq3)[1]
       apply(auto simp add: object_ptr_kinds_M_eq object_ptr_kinds_eq3)[1]
      apply(rule bind_cong_2)
         apply(auto intro!: filter_M_pure_I bind_pure_I)[1]
        apply(auto intro!: filter_M_pure_I bind_pure_I)[1]
       apply(auto simp add: filter_eq (* dest!: returns_result_eq[OF ptrs] *))[1]
      using filter_eq ptrs apply auto[1]
      using filter_eq ptrs by auto
  qed

  have "ancestors = ancestors'"
  proof(insert assms(1) assms(7) ancestors' 2, induct ptr arbitrary: ancestors ancestors'
      rule: heap_wellformed_induct_rev)
    case (step child)
    show ?case
      using step(2) step(3) step(4)
      apply(simp add: get_ancestors_def)
      apply(auto intro!: elim!: bind_returns_result_E2 split: option.splits)[1]
      using returns_result_eq apply fastforce
       apply (meson option.simps(3) returns_result_eq)
      by (metis IntD1 IntD2 option.inject returns_result_eq step.hyps)
  qed
  then show ?thesis
    using assms(5) ancestors'
    by simp
qed

lemma get_ancestors_remains_not_in_ancestors:
  assumes "heap_is_wellformed h"
    and "heap_is_wellformed h'"
    and "h  get_ancestors ptr r ancestors"
    and "h'  get_ancestors ptr r ancestors'"
    and "p children children'. h  get_child_nodes p r children
         h'  get_child_nodes p r children'  set children'  set children"
    and "node  set ancestors"
    and object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'"
    and known_ptrs: "known_ptrs h"
    and type_wf: "type_wf h"
    and type_wf': "type_wf h'"
  shows "node  set ancestors'"
proof -
  have object_ptr_kinds_M_eq:
    "ptrs. h  object_ptr_kinds_M r ptrs = h'  object_ptr_kinds_M r ptrs"
    using object_ptr_kinds_eq3
    by(simp add: object_ptr_kinds_M_defs)
  then have object_ptr_kinds_eq: "|h  object_ptr_kinds_M|r = |h'  object_ptr_kinds_M|r"
    by(simp)

  show ?thesis
  proof (insert assms(1) assms(3) assms(4) assms(6), induct ptr arbitrary: ancestors ancestors'
      rule: heap_wellformed_induct_rev)
    case (step child)
    have 1: "p parent. h'  get_parent p r Some parent  h  get_parent p r Some parent"
    proof -
      fix p parent
      assume "h'  get_parent p r Some parent"
      then obtain children' where
        children': "h'  get_child_nodes parent r children'" and
        p_in_children': "p  set children'"
        using get_parent_child_dual by blast
      obtain children where children: "h  get_child_nodes parent r children"
        using get_child_nodes_ok assms(1) get_child_nodes_ptr_in_heap object_ptr_kinds_eq children'
          known_ptrs
        using  type_wf type_wf'
        by (metis h'  get_parent p r Some parent get_parent_parent_in_heap is_OK_returns_result_E
            local.known_ptrs_known_ptr object_ptr_kinds_eq3)
      have "p  set children"
        using assms(5) children children' p_in_children'
        by blast
      then show "h  get_parent p r Some parent"
        using child_parent_dual assms(1) children known_ptrs  type_wf  by blast
    qed
    have "node  child"
      using assms(1) get_ancestors_parent_child_rel step.prems(1) step.prems(3) known_ptrs
      using type_wf type_wf'
      by blast
    then show ?case
      using step(2) step(3)
      apply(simp add: get_ancestors_def)
      using step(4)
      apply(auto elim!: bind_returns_result_E2 split: option.splits)[1]
      using 1
       apply (meson option.distinct(1) returns_result_eq)
      by (metis "1" option.inject returns_result_eq step.hyps)
  qed
qed

lemma get_ancestors_ptrs_in_heap:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_ancestors ptr r ancestors"
  assumes "ptr'  set ancestors"
  shows "ptr' |∈| object_ptr_kinds h"
proof (insert assms(4) assms(5), induct ancestors arbitrary: ptr)
  case Nil
  then show ?case
    by(auto)
next
  case (Cons a ancestors)
  then obtain x where x: "h  get_ancestors x r a # ancestors"
    by(auto simp add: get_ancestors_def[of a] elim!: bind_returns_result_E2 split: option.splits)
  then have "x = a"
    by(auto simp add: get_ancestors_def[of x] elim!: bind_returns_result_E2 split: option.splits)
  then show ?case
    using Cons.hyps Cons.prems(2) get_ancestors_ptr_in_heap x
    by (metis assms(1) assms(2) assms(3) get_ancestors_obtains_children get_child_nodes_ptr_in_heap
        is_OK_returns_result_I)
qed


lemma get_ancestors_prefix:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_ancestors ptr r ancestors"
  assumes "ptr'  set ancestors"
  assumes "h  get_ancestors ptr' r ancestors'"
  shows "pre. ancestors = pre @ ancestors'"
proof (insert assms(1) assms(5) assms(6), induct ptr' arbitrary: ancestors'
    rule: heap_wellformed_induct)
  case (step parent)
  then show ?case
  proof (cases "parent  ptr" )
    case True

    then obtain children ancestor_child where "h  get_child_nodes parent r children"
      and "ancestor_child  set children" and "cast ancestor_child  set ancestors"
      using assms(1) assms(2) assms(3) assms(4) get_ancestors_obtains_children step.prems(1) by blast
    then have "h  get_parent ancestor_child r Some parent"
      using assms(1) assms(2) assms(3) child_parent_dual by blast
    then have "h  get_ancestors (cast ancestor_child) r cast ancestor_child # ancestors'"
      apply(simp add: get_ancestors_def)
      using h  get_ancestors parent r ancestors' get_parent_ptr_in_heap
      by(auto simp add: check_in_heap_def is_OK_returns_result_I intro!: bind_pure_returns_result_I)
    then show ?thesis
      using step(1) h  get_child_nodes parent r children ancestor_child  set children
        cast ancestor_child  set ancestors h  get_ancestors (cast ancestor_child) r cast ancestor_child # ancestors'
      by fastforce
  next
    case False
    then show ?thesis
      by (metis append_Nil assms(4) returns_result_eq step.prems(2))
  qed
qed


lemma get_ancestors_same_root_node:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_ancestors ptr r ancestors"
  assumes "ptr'  set ancestors"
  assumes "ptr''  set ancestors"
  shows "h  get_root_node ptr' r root_ptr  h  get_root_node ptr'' r root_ptr"
proof -
  have "ptr' |∈| object_ptr_kinds h"
    by (metis assms(1) assms(2) assms(3) assms(4) assms(5) get_ancestors_obtains_children
        get_ancestors_ptr_in_heap get_child_nodes_ptr_in_heap is_OK_returns_result_I)
  then obtain ancestors' where ancestors': "h  get_ancestors ptr' r ancestors'"
    by (meson assms(1) assms(2) assms(3) get_ancestors_ok is_OK_returns_result_E)
  then have "pre. ancestors = pre @ ancestors'"
    using get_ancestors_prefix assms by blast
  moreover have "ptr'' |∈| object_ptr_kinds h"
    by (metis assms(1) assms(2) assms(3) assms(4) assms(6) get_ancestors_obtains_children
        get_ancestors_ptr_in_heap get_child_nodes_ptr_in_heap is_OK_returns_result_I)
  then obtain ancestors'' where ancestors'': "h  get_ancestors ptr'' r ancestors''"
    by (meson assms(1) assms(2) assms(3) get_ancestors_ok is_OK_returns_result_E)
  then have "pre. ancestors = pre @ ancestors''"
    using get_ancestors_prefix assms by blast
  ultimately show ?thesis
    using ancestors' ancestors''
    apply(auto simp add: get_root_node_def elim!: bind_returns_result_E2
        intro!: bind_pure_returns_result_I)[1]
     apply (metis (no_types, lifting) assms(1) get_ancestors_never_empty last_appendR
        returns_result_eq)
    by (metis assms(1) get_ancestors_never_empty last_appendR returns_result_eq)
qed

lemma get_root_node_parent_same:
  assumes "h  get_parent child r Some ptr"
  shows "h  get_root_node (cast child) r root  h  get_root_node ptr r root"
proof
  assume 1: " h  get_root_node (castnode_ptr2object_ptr child) r root"
  show "h  get_root_node ptr r root"
    using 1[unfolded get_root_node_def] assms
    apply(simp add: get_ancestors_def)
    apply(auto simp add: get_root_node_def dest: returns_result_eq elim!: bind_returns_result_E2
        intro!: bind_pure_returns_result_I split: option.splits)[1]
    using returns_result_eq apply fastforce
    using get_ancestors_ptr by fastforce
next
  assume 1: " h  get_root_node ptr r root"
  show "h  get_root_node (castnode_ptr2object_ptr child) r root"
    apply(simp add: get_root_node_def)
    using assms 1
    apply(simp add: get_ancestors_def)
    apply(auto simp add: get_root_node_def dest: returns_result_eq elim!: bind_returns_result_E2
        intro!: bind_pure_returns_result_I split: option.splits)[1]
     apply (simp add: check_in_heap_def is_OK_returns_result_I)
    using get_ancestors_ptr get_parent_ptr_in_heap
    apply (simp add: is_OK_returns_result_I)
    by (meson list.distinct(1) list.set_cases local.get_ancestors_ptr)
qed

lemma get_root_node_same_no_parent:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_root_node ptr r cast child"
  shows "h  get_parent child r None"
proof (insert assms(1) assms(4), induct ptr rule: heap_wellformed_induct_rev)
  case (step c)
  then show ?case
  proof (cases "castobject_ptr2node_ptr c")
    case None
    then have "c = cast child"
      using step(2)
      by(auto simp add: get_root_node_def get_ancestors_def[of c] elim!: bind_returns_result_E2)
    then show ?thesis
      using None by auto
  next
    case (Some child_node)
    note s = this
    then obtain parent_opt where parent_opt: "h  get_parent child_node r parent_opt"
      by (metis (no_types, lifting) assms(2) assms(3) get_root_node_ptr_in_heap
          is_OK_returns_result_I local.get_parent_ok node_ptr_casts_commute
          node_ptr_kinds_commutes returns_result_select_result step.prems)
    then show ?thesis
    proof(induct parent_opt)
      case None
      then show ?case
        using Some get_root_node_no_parent returns_result_eq step.prems by fastforce
    next
      case (Some parent)
      then show ?case
        using step s
        apply(auto simp add: get_root_node_def get_ancestors_def[of c]
            elim!: bind_returns_result_E2 split: option.splits list.splits)[1]
        using get_root_node_parent_same step.hyps step.prems by auto
    qed
  qed
qed

lemma get_root_node_not_node_same:
  assumes "ptr |∈| object_ptr_kinds h"
  assumes "¬is_node_ptr_kind ptr"
  shows "h  get_root_node ptr r ptr"
  using assms
  apply(simp add: get_root_node_def get_ancestors_def)
  by(auto simp add: get_root_node_def dest: returns_result_eq elim!: bind_returns_result_E2
      intro!: bind_pure_returns_result_I split: option.splits)


lemma get_root_node_root_in_heap:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_root_node ptr r root"
  shows "root |∈| object_ptr_kinds h"
  using assms
  apply(auto simp add: get_root_node_def elim!: bind_returns_result_E2)[1]
  by (simp add: get_ancestors_never_empty get_ancestors_ptrs_in_heap)


lemma get_root_node_same_no_parent_parent_child_rel:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_root_node ptr' r ptr'"
  shows "¬(p. (p, ptr')  (parent_child_rel h))"
  by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) get_root_node_same_no_parent
      l_heap_is_wellformed.parent_child_rel_child local.child_parent_dual local.get_child_nodes_ok
      local.known_ptrs_known_ptr local.l_heap_is_wellformed_axioms local.parent_child_rel_node_ptr
      local.parent_child_rel_parent_in_heap node_ptr_casts_commute3 option.simps(3) returns_result_eq
      returns_result_select_result)

end


locale l_get_ancestors_wf = l_heap_is_wellformed_defs + l_known_ptrs + l_type_wf + l_get_ancestors_defs
  + l_get_child_nodes_defs + l_get_parent_defs +
  assumes get_ancestors_never_empty:
    "heap_is_wellformed h  h  get_ancestors child r ancestors  ancestors  []"
  assumes get_ancestors_ok:
    "heap_is_wellformed h  ptr |∈| object_ptr_kinds h  known_ptrs h  type_wf h
                           h  ok (get_ancestors ptr)"
  assumes get_ancestors_reads:
    "heap_is_wellformed h  reads get_ancestors_locs (get_ancestors node_ptr) h h'"
  assumes get_ancestors_ptrs_in_heap:
    "heap_is_wellformed h  type_wf h  known_ptrs h
                           h  get_ancestors ptr r ancestors  ptr'  set ancestors
                           ptr' |∈| object_ptr_kinds h"
  assumes get_ancestors_remains_not_in_ancestors:
    "heap_is_wellformed h  heap_is_wellformed h'  h  get_ancestors ptr r ancestors
                           h'  get_ancestors ptr r ancestors'
                           (p children children'. h  get_child_nodes p r children
                                h'  get_child_nodes p r children'
                                set children'  set children)
                           node  set ancestors
                           object_ptr_kinds h = object_ptr_kinds h'  known_ptrs h
                           type_wf h  type_wf h'  node  set ancestors'"
  assumes get_ancestors_also_parent:
    "heap_is_wellformed h  h  get_ancestors some_ptr r ancestors
                           cast child_node  set ancestors
                           h  get_parent child_node r Some parent  type_wf h
                           known_ptrs h  parent  set ancestors"
  assumes get_ancestors_obtains_children:
    "heap_is_wellformed h  ancestor  ptr  ancestor  set ancestors
                           h  get_ancestors ptr r ancestors  type_wf h  known_ptrs h
                           (children ancestor_child . h  get_child_nodes ancestor r children
                                ancestor_child  set children
                                cast ancestor_child  set ancestors
                                thesis)
                           thesis"
  assumes get_ancestors_parent_child_rel:
    "heap_is_wellformed h  h  get_ancestors child r ancestors  known_ptrs h  type_wf h
                           (ptr, child)  (parent_child_rel h)*  ptr  set ancestors"

locale l_get_root_node_wf = l_heap_is_wellformed_defs + l_get_root_node_defs + l_type_wf
  + l_known_ptrs + l_get_ancestors_defs + l_get_parent_defs +
  assumes get_root_node_ok:
    "heap_is_wellformed h  known_ptrs h  type_wf h  ptr |∈| object_ptr_kinds h
                           h  ok (get_root_node ptr)"
  assumes get_root_node_ptr_in_heap:
    "h  ok (get_root_node ptr)  ptr |∈| object_ptr_kinds h"
  assumes get_root_node_root_in_heap:
    "heap_is_wellformed h  type_wf h  known_ptrs h
                           h  get_root_node ptr r root  root |∈| object_ptr_kinds h"
  assumes get_ancestors_same_root_node:
    "heap_is_wellformed h  type_wf h  known_ptrs h
                           h  get_ancestors ptr r ancestors  ptr'  set ancestors
                           ptr''  set ancestors
                           h  get_root_node ptr' r root_ptr  h  get_root_node ptr'' r root_ptr"
  assumes get_root_node_same_no_parent:
    "heap_is_wellformed h  type_wf h  known_ptrs h
                           h  get_root_node ptr r cast child  h  get_parent child r None"
  assumes get_root_node_parent_same:
    "h  get_parent child r Some ptr
         h  get_root_node (cast child) r root  h  get_root_node ptr r root"

interpretation i_get_root_node_wf?:
  l_get_root_node_wfCore_DOM known_ptr type_wf known_ptrs heap_is_wellformed parent_child_rel
  get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs
  get_parent get_parent_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs
  using instances
  by(simp add: l_get_root_node_wfCore_DOM_def)
declare l_get_root_node_wfCore_DOM_axioms[instances]

lemma get_ancestors_wf_is_l_get_ancestors_wf [instances]:
  "l_get_ancestors_wf heap_is_wellformed parent_child_rel known_ptr known_ptrs type_wf get_ancestors
  get_ancestors_locs get_child_nodes get_parent"
  using known_ptrs_is_l_known_ptrs
  apply(auto simp add: l_get_ancestors_wf_def l_get_ancestors_wf_axioms_def)[1]
  using get_ancestors_never_empty apply blast
  using get_ancestors_ok apply blast
  using get_ancestors_reads apply blast
  using get_ancestors_ptrs_in_heap apply blast
  using get_ancestors_remains_not_in_ancestors apply blast
  using get_ancestors_also_parent apply blast
  using get_ancestors_obtains_children apply blast
  using get_ancestors_parent_child_rel apply blast
  using get_ancestors_parent_child_rel apply blast
  done

lemma get_root_node_wf_is_l_get_root_node_wf [instances]:
  "l_get_root_node_wf heap_is_wellformed get_root_node type_wf known_ptr known_ptrs
   get_ancestors get_parent"
  using known_ptrs_is_l_known_ptrs
  apply(auto simp add: l_get_root_node_wf_def l_get_root_node_wf_axioms_def)[1]
  using get_root_node_ok apply blast
  using get_root_node_ptr_in_heap apply blast
  using get_root_node_root_in_heap apply blast
  using get_ancestors_same_root_node apply(blast, blast)
  using get_root_node_same_no_parent apply blast
  using get_root_node_parent_same apply (blast, blast)
  done


subsection ‹to\_tree\_order›

locale l_to_tree_order_wfCore_DOM =
  l_to_tree_orderCore_DOM +
  l_get_parent +
  l_get_parent_wf +
  l_heap_is_wellformed
  (* l_get_parent_wfCore_DOM *)
begin

lemma to_tree_order_ptr_in_heap:
  assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
  assumes "h  ok (to_tree_order ptr)"
  shows "ptr |∈| object_ptr_kinds h"
proof(insert assms(1) assms(4), induct rule: heap_wellformed_induct)
  case (step parent)
  then show ?case
    apply(auto simp add: to_tree_order_def[of parent] map_M_pure_I elim!: bind_is_OK_E3)[1]
    using get_child_nodes_ptr_in_heap by blast
qed

lemma to_tree_order_either_ptr_or_in_children:
  assumes "h  to_tree_order ptr r nodes"
    and "node  set nodes"
    and "h  get_child_nodes ptr r children"
    and "node  ptr"
  obtains child child_to where "child  set children"
    and "h  to_tree_order (cast child) r child_to" and "node  set child_to"
proof -
  obtain treeorders where treeorders: "h  map_M to_tree_order (map cast children) r treeorders"
    using assms
    apply(auto simp add: to_tree_order_def elim!: bind_returns_result_E)[1]
    using pure_returns_heap_eq returns_result_eq by fastforce
  then have "node  set (concat treeorders)"
    using assms[simplified to_tree_order_def]
    by(auto elim!: bind_returns_result_E4 dest: pure_returns_heap_eq)
  then obtain treeorder where "treeorder  set treeorders"
    and node_in_treeorder: "node  set treeorder"
    by auto
  then obtain child where "h  to_tree_order (castnode_ptr2object_ptr child) r treeorder"
    and  "child  set children"
    using assms[simplified to_tree_order_def] treeorders
    by(auto elim!: map_M_pure_E2)
  then show ?thesis
    using node_in_treeorder returns_result_eq that by auto
qed


lemma to_tree_order_ptrs_in_heap:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  to_tree_order ptr r to"
  assumes "ptr'  set to"
  shows "ptr' |∈| object_ptr_kinds h"
proof(insert assms(1) assms(4) assms(5), induct ptr arbitrary: to rule: heap_wellformed_induct)
  case (step parent)
  have "parent |∈| object_ptr_kinds h"
    using assms(1) assms(2) assms(3) step.prems(1) to_tree_order_ptr_in_heap by blast
  then obtain children where children: "h  get_child_nodes parent r children"
    by (meson assms(2) assms(3) get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr)
  then show ?case
  proof (cases "children = []")
    case True
    then have "to = [parent]"
      using step(2) children
      apply(auto simp add: to_tree_order_def[of parent] map_M_pure_I elim!: bind_returns_result_E2)[1]
      by (metis list.distinct(1) list.map_disc_iff list.set_cases map_M_pure_E2 returns_result_eq)
    then show ?thesis
      using parent |∈| object_ptr_kinds h step.prems(2) by auto
  next
    case False
    note f = this
    then show ?thesis
      using children step to_tree_order_either_ptr_or_in_children
    proof (cases "ptr' = parent")
      case True
      then show ?thesis
        using parent |∈| object_ptr_kinds h by blast
    next
      case False
      then show ?thesis
        using children step.hyps to_tree_order_either_ptr_or_in_children
        by (metis step.prems(1) step.prems(2))
    qed
  qed
qed

lemma to_tree_order_ok:
  assumes wellformed: "heap_is_wellformed h"
    and "ptr |∈| object_ptr_kinds h"
    and "known_ptrs h"
    and type_wf: "type_wf h"
  shows "h  ok (to_tree_order ptr)"
proof(insert assms(1) assms(2), induct rule: heap_wellformed_induct)
  case (step parent)
  then show ?case
    using assms(3) type_wf
    apply(simp add: to_tree_order_def)
    apply(auto simp add: heap_is_wellformed_def intro!: map_M_ok_I bind_is_OK_pure_I map_M_pure_I)[1]
    using get_child_nodes_ok known_ptrs_known_ptr apply blast
    by (simp add: local.heap_is_wellformed_children_in_heap local.to_tree_order_def wellformed)
qed

lemma to_tree_order_child_subset:
  assumes "heap_is_wellformed h"
    and "h  to_tree_order ptr r nodes"
    and "h  get_child_nodes ptr r children"
    and "node  set children"
    and "h  to_tree_order (cast node) r nodes'"
  shows "set nodes'  set nodes"
proof
  fix x
  assume a1: "x  set nodes'"
  moreover obtain treeorders
    where treeorders: "h  map_M to_tree_order (map cast children) r treeorders"
    using assms(2) assms(3)
    apply(auto simp add: to_tree_order_def elim!: bind_returns_result_E)[1]
    using pure_returns_heap_eq returns_result_eq by fastforce
  then have "nodes'  set treeorders"
    using assms(4) assms(5)
    by(auto elim!: map_M_pure_E dest: returns_result_eq)
  moreover have "set (concat treeorders)  set nodes"
    using treeorders assms(2) assms(3)
    by(auto simp add: to_tree_order_def elim!: bind_returns_result_E4 dest: pure_returns_heap_eq)
  ultimately show "x  set nodes"
    by auto
qed

lemma to_tree_order_ptr_in_result:
  assumes "h  to_tree_order ptr r nodes"
  shows "ptr  set nodes"
  using assms
  apply(simp add: to_tree_order_def)
  by(auto elim!: bind_returns_result_E2 intro!: map_M_pure_I bind_pure_I)

lemma to_tree_order_subset:
  assumes "heap_is_wellformed h"
    and "h  to_tree_order ptr r nodes"
    and "node  set nodes"
    and "h  to_tree_order node r nodes'"
    and "known_ptrs h"
    and type_wf: "type_wf h"
  shows "set nodes'  set nodes"
proof -
  have "nodes. h  to_tree_order ptr r nodes  (node. node  set nodes
        (nodes'. h  to_tree_order node r nodes'  set nodes'  set nodes))"
  proof(insert assms(1), induct ptr rule: heap_wellformed_induct)
    case (step parent)
    then show ?case
    proof safe
      fix nodes node nodes' x
      assume 1: "(children child.
             h  get_child_nodes parent r children 
             child  set children  nodes. h  to_tree_order (cast child) r nodes
             (node. node  set nodes  (nodes'. h  to_tree_order node r nodes'
                                               set nodes'  set nodes)))"
        and 2: "h  to_tree_order parent r nodes"
        and 3: "node  set nodes"
        and "h  to_tree_order node r nodes'"
        and "x  set nodes'"
      have h1: "(children child nodes node nodes'.
             h  get_child_nodes parent r children 
             child  set children  h  to_tree_order (cast child) r nodes
              (node  set nodes  (h  to_tree_order node r nodes'  set nodes'  set nodes)))"
        using 1
        by blast
      obtain children where children: "h  get_child_nodes parent r children"
        using 2
        by(auto simp add: to_tree_order_def elim!: bind_returns_result_E)
      then have "set nodes'  set nodes"
      proof (cases "children = []")
        case True
        then show ?thesis
          by (metis "2" "3" h  to_tree_order node r nodes' children empty_iff list.set(1)
              subsetI to_tree_order_either_ptr_or_in_children)
      next
        case False
        then show ?thesis
        proof (cases "node = parent")
          case True
          then show ?thesis
            using "2" h  to_tree_order node r nodes' returns_result_eq by fastforce
        next
          case False
          then obtain child nodes_of_child where
            "child  set children" and
            "h  to_tree_order (cast child) r nodes_of_child" and
            "node  set nodes_of_child"
            using 2[simplified to_tree_order_def] 3
              to_tree_order_either_ptr_or_in_children[where node=node and ptr=parent] children
            apply(auto elim!: bind_returns_result_E2 intro: map_M_pure_I)[1]
            using is_OK_returns_result_E 2 a_all_ptrs_in_heap_def assms(1) heap_is_wellformed_def
            using "3" by blast
          then have "set nodes'  set nodes_of_child"
            using h1
            using h  to_tree_order node r nodes' children by blast
          moreover have "set nodes_of_child  set nodes"
            using "2" child  set children h  to_tree_order (cast child) r nodes_of_child
              assms children to_tree_order_child_subset by auto
          ultimately show ?thesis
            by blast
        qed
      qed
      then show "x  set nodes"
        using x  set nodes' by blast
    qed
  qed
  then show ?thesis
    using assms by blast
qed

lemma to_tree_order_parent:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  to_tree_order ptr r nodes"
  assumes "h  get_parent child r Some parent"
  assumes "parent  set nodes"
  shows "cast child  set nodes"
proof -
  obtain nodes' where nodes': "h  to_tree_order parent r nodes'"
    using assms to_tree_order_ok get_parent_parent_in_heap
    by (meson get_parent_parent_in_heap is_OK_returns_result_E)

  then have "set nodes'  set nodes"
    using to_tree_order_subset assms
    by blast
  moreover obtain children where
    children: "h  get_child_nodes parent r children" and
    child: "child  set children"
    using assms get_parent_child_dual by blast
  then obtain child_to where child_to: "h  to_tree_order (castnode_ptr2object_ptr child) r child_to"
    by (meson assms(1) assms(2) assms(3) assms(5) is_OK_returns_result_E is_OK_returns_result_I
        get_parent_ptr_in_heap node_ptr_kinds_commutes to_tree_order_ok)
  then have "cast child  set child_to"
    apply(simp add: to_tree_order_def)
    by(auto elim!: bind_returns_result_E2 map_M_pure_E
        dest!: bind_returns_result_E3[rotated, OF children, rotated]  intro!: map_M_pure_I)

  have "cast child  set nodes'"
    using nodes' child
    apply(simp add: to_tree_order_def)
    apply(auto elim!: bind_returns_result_E2 map_M_pure_E
        dest!: bind_returns_result_E3[rotated, OF children, rotated]  intro!: map_M_pure_I)[1]
    using child_to cast child  set child_to returns_result_eq by fastforce
  ultimately show ?thesis
    by auto
qed

lemma to_tree_order_child:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  to_tree_order ptr r nodes"
  assumes "h  get_child_nodes parent r children"
  assumes "cast child  ptr"
  assumes "child  set children"
  assumes "cast child  set nodes"
  shows "parent  set nodes"
proof(insert assms(1) assms(4) assms(6) assms(8), induct ptr arbitrary: nodes
    rule: heap_wellformed_induct)
  case (step p)
  have "p |∈| object_ptr_kinds h"
    using h  to_tree_order p r nodes to_tree_order_ptr_in_heap
    using assms(1) assms(2) assms(3) by blast
  then obtain children where children: "h  get_child_nodes p r children"
    by (meson assms(2) assms(3) get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr)
  then show ?case
  proof (cases "children = []")
    case True
    then show ?thesis
      using step(2) step(3) step(4) children
      by(auto simp add: to_tree_order_def[of p] map_M_pure_I  elim!: bind_returns_result_E2
          dest!: bind_returns_result_E3[rotated, OF children, rotated])
  next
    case False
    then obtain c child_to where
      child: "c  set children" and
      child_to: "h  to_tree_order (castnode_ptr2object_ptr c) r child_to" and
      "cast child  set child_to"
      using step(2) children
      apply(auto simp add: to_tree_order_def[of p] map_M_pure_I  elim!: bind_returns_result_E2
          dest!: bind_returns_result_E3[rotated, OF children, rotated])[1]
      by (metis (full_types) assms(1) assms(2) assms(3) get_parent_ptr_in_heap
          is_OK_returns_result_I l_get_parent_wfCore_DOM.child_parent_dual
          l_get_parent_wfCore_DOM_axioms node_ptr_kinds_commutes
          returns_result_select_result step.prems(1) step.prems(2) step.prems(3)
          to_tree_order_either_ptr_or_in_children to_tree_order_ok)
    then have "set child_to  set nodes"
      using assms(1) child children step.prems(1) to_tree_order_child_subset by auto

    show ?thesis
    proof (cases "c = child")
      case True
      then have "parent = p"
        using step(3) children child assms(5) assms(7)
        by (meson assms(1) assms(2) assms(3) child_parent_dual option.inject returns_result_eq)

      then show ?thesis
        using step.prems(1) to_tree_order_ptr_in_result by blast
    next
      case False
      then show ?thesis
        using step(1)[OF children child child_to] step(3) step(4)
        using set child_to  set nodes
        using castnode_ptr2object_ptr child  set child_to by auto
    qed
  qed
qed

lemma to_tree_order_node_ptrs:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  to_tree_order ptr r nodes"
  assumes "ptr'  ptr"
  assumes "ptr'  set nodes"
  shows "is_node_ptr_kind ptr'"
proof(insert assms(1) assms(4) assms(5) assms(6), induct ptr arbitrary: nodes
    rule: heap_wellformed_induct)
  case (step p)
  have "p |∈| object_ptr_kinds h"
    using h  to_tree_order p r nodes to_tree_order_ptr_in_heap
    using assms(1) assms(2) assms(3) by blast
  then obtain children where children: "h  get_child_nodes p r children"
    by (meson assms(2) assms(3) get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr)
  then show ?case
  proof (cases "children = []")
    case True
    then show ?thesis
      using step(2) step(3) step(4) children
      by(auto simp add: to_tree_order_def[of p] map_M_pure_I  elim!: bind_returns_result_E2
          dest!: bind_returns_result_E3[rotated, OF children, rotated])[1]
  next
    case False
    then obtain c child_to where
      child: "c  set children" and
      child_to: "h  to_tree_order (castnode_ptr2object_ptr c) r child_to" and
      "ptr'  set child_to"
      using step(2) children
      apply(auto simp add: to_tree_order_def[of p] map_M_pure_I  elim!: bind_returns_result_E2
          dest!: bind_returns_result_E3[rotated, OF children, rotated])[1]
      using step.prems(1) step.prems(2) step.prems(3) to_tree_order_either_ptr_or_in_children by blast
    then have "set child_to  set nodes"
      using assms(1) child children step.prems(1) to_tree_order_child_subset by auto

    show ?thesis
    proof (cases "cast c = ptr")
      case True
      then show ?thesis
        using step ptr'  set child_to assms(5) child child_to children by blast
    next
      case False
      then show ?thesis
        using ptr'  set child_to child child_to children is_node_ptr_kind_cast step.hyps by blast
    qed
  qed
qed

lemma to_tree_order_child2:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  to_tree_order ptr r nodes"
  assumes "cast child  ptr"
  assumes "cast child  set nodes"
  obtains parent where "h  get_parent child r Some parent" and "parent  set nodes"
proof -
  assume 1: "(parent. h  get_parent child r Some parent  parent  set nodes  thesis)"
  show thesis
  proof(insert assms(1) assms(4) assms(5) assms(6) 1, induct ptr arbitrary: nodes
      rule: heap_wellformed_induct)
    case (step p)
    have "p |∈| object_ptr_kinds h"
      using h  to_tree_order p r nodes to_tree_order_ptr_in_heap
      using assms(1) assms(2) assms(3) by blast
    then obtain children where children: "h  get_child_nodes p r children"
      by (meson assms(2) assms(3) get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr)
    then show ?case
    proof (cases "children = []")
      case True
      then show ?thesis
        using step(2) step(3) step(4) children
        by(auto simp add: to_tree_order_def[of p] map_M_pure_I  elim!: bind_returns_result_E2
            dest!: bind_returns_result_E3[rotated, OF children, rotated])
    next
      case False
      then obtain c child_to where
        child: "c  set children" and
        child_to: "h  to_tree_order (castnode_ptr2object_ptr c) r child_to" and
        "cast child  set child_to"
        using step(2) children
        apply(auto simp add: to_tree_order_def[of p] map_M_pure_I  elim!: bind_returns_result_E2
            dest!: bind_returns_result_E3[rotated, OF children, rotated])[1]
        using step.prems(1) step.prems(2) step.prems(3) to_tree_order_either_ptr_or_in_children
        by blast
      then have "set child_to  set nodes"
        using assms(1) child children step.prems(1) to_tree_order_child_subset by auto

      have "cast child |∈| object_ptr_kinds h"
        using assms(1) assms(2) assms(3) assms(4) assms(6) to_tree_order_ptrs_in_heap by blast
      then obtain parent_opt where parent_opt: "h  get_parent child r parent_opt"
        by (meson assms(2) assms(3) is_OK_returns_result_E get_parent_ok node_ptr_kinds_commutes)
      then show ?thesis
      proof (induct parent_opt)
        case None
        then show ?case
          by (metis castnode_ptr2object_ptr child  set child_to assms(1) assms(2) assms(3)
              castnode_ptr2object_ptr_inject child child_parent_dual child_to children
              option.distinct(1) returns_result_eq step.hyps)
      next
        case (Some option)
        then show ?case
          by (meson assms(1) assms(2) assms(3) get_parent_child_dual step.prems(1) step.prems(2)
              step.prems(3) step.prems(4) to_tree_order_child)
      qed
    qed
  qed
qed

lemma to_tree_order_parent_child_rel:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  to_tree_order ptr r to"
  shows "(ptr, child)  (parent_child_rel h)*  child  set to"
proof
  assume 3: "(ptr, child)  (parent_child_rel h)*"
  show "child  set to"
  proof (insert 3, induct child rule: heap_wellformed_induct_rev[OF assms(1)])
    case (1 child)
    then show ?case
    proof (cases "ptr = child")
      case True
      then show ?thesis
        using assms(4)
        apply(simp add: to_tree_order_def)
        by(auto simp add: map_M_pure_I elim!: bind_returns_result_E2)
    next
      case False
      obtain child_parent where
        "(ptr, child_parent)  (parent_child_rel h)*" and
        "(child_parent, child)  (parent_child_rel h)"
        using ptr  child
        by (metis "1.prems" rtranclE)
      obtain child_node where child_node: "castnode_ptr2object_ptr child_node = child"
        using (child_parent, child)  parent_child_rel h node_ptr_casts_commute3
          parent_child_rel_node_ptr
        by blast
      then have "h  get_parent child_node r Some child_parent"
        using (child_parent, child)  (parent_child_rel h)
        by (meson assms(1) assms(2) assms(3) is_OK_returns_result_E l_get_parent_wf.child_parent_dual
            l_heap_is_wellformed.parent_child_rel_child local.get_child_nodes_ok
            local.known_ptrs_known_ptr local.l_get_parent_wf_axioms
            local.l_heap_is_wellformed_axioms local.parent_child_rel_parent_in_heap)
      then show ?thesis
        using 1(1) child_node (ptr, child_parent)  (parent_child_rel h)*
        using assms(1) assms(2) assms(3) assms(4) to_tree_order_parent by blast
    qed
  qed
next
  assume "child  set to"
  then show "(ptr, child)  (parent_child_rel h)*"
  proof (induct child rule: heap_wellformed_induct_rev[OF assms(1)])
    case (1 child)
    then show ?case
    proof (cases "ptr = child")
      case True
      then show ?thesis
        by simp
    next
      case False
      then have "parent. (parent, child)  (parent_child_rel h)"
        using 1(2) assms(4) to_tree_order_child2[OF assms(1) assms(2) assms(3) assms(4)]
          to_tree_order_node_ptrs
        by (metis assms(1) assms(2) assms(3) node_ptr_casts_commute3 parent_child_rel_parent)
      then obtain child_node where child_node: "castnode_ptr2object_ptr child_node = child"
        using node_ptr_casts_commute3 parent_child_rel_node_ptr by blast
      then obtain child_parent where child_parent: "h  get_parent child_node r Some child_parent"
        using parent. (parent, child)  (parent_child_rel h)
        by (metis "1.prems" False assms(1) assms(2) assms(3) assms(4) to_tree_order_child2)
      then have "(child_parent, child)  (parent_child_rel h)"
        using assms(1) child_node parent_child_rel_parent by blast
      moreover have "child_parent  set to"
        by (metis "1.prems" False assms(1) assms(2) assms(3) assms(4) child_node child_parent
            get_parent_child_dual to_tree_order_child)
      then have "(ptr, child_parent)  (parent_child_rel h)*"
        using 1 child_node child_parent by blast
      ultimately show ?thesis
        by auto
    qed
  qed
qed
end

interpretation i_to_tree_order_wf?: l_to_tree_order_wfCore_DOM known_ptr type_wf get_child_nodes
  get_child_nodes_locs to_tree_order known_ptrs get_parent
  get_parent_locs heap_is_wellformed parent_child_rel
  get_disconnected_nodes get_disconnected_nodes_locs
  using instances
  apply(simp add: l_to_tree_order_wfCore_DOM_def)
  done
declare l_to_tree_order_wfCore_DOM_axioms [instances]

locale l_to_tree_order_wf = l_heap_is_wellformed_defs + l_type_wf + l_known_ptrs
  + l_to_tree_order_defs
  + l_get_parent_defs + l_get_child_nodes_defs +
  assumes to_tree_order_ok:
    "heap_is_wellformed h  ptr |∈| object_ptr_kinds h  known_ptrs h  type_wf h
                           h  ok (to_tree_order ptr)"
  assumes to_tree_order_ptrs_in_heap:
    "heap_is_wellformed h  type_wf h  known_ptrs h  h  to_tree_order ptr r to
                           ptr'  set to  ptr' |∈| object_ptr_kinds h"
  assumes to_tree_order_parent_child_rel:
    "heap_is_wellformed h  type_wf h  known_ptrs h  h  to_tree_order ptr r to
                           (ptr, child_ptr)  (parent_child_rel h)*  child_ptr  set to"
  assumes to_tree_order_child2:
    "heap_is_wellformed h  type_wf h  known_ptrs h  h  to_tree_order ptr r nodes
                           cast child  ptr  cast child  set nodes
                           (parent. h  get_parent child r Some parent
                                          parent  set nodes  thesis)
                           thesis"
  assumes to_tree_order_node_ptrs:
    "heap_is_wellformed h  type_wf h  known_ptrs h  h  to_tree_order ptr r nodes
                           ptr'  ptr  ptr'  set nodes  is_node_ptr_kind ptr'"
  assumes to_tree_order_child:
    "heap_is_wellformed h  type_wf h  known_ptrs h  h  to_tree_order ptr r nodes
                           h  get_child_nodes parent r children  cast child  ptr
                           child  set children  cast child  set nodes
                           parent  set nodes"
  assumes to_tree_order_ptr_in_result:
    "h  to_tree_order ptr r nodes  ptr  set nodes"
  assumes to_tree_order_parent:
    "heap_is_wellformed h  type_wf h  known_ptrs h  h  to_tree_order ptr r nodes
                           h  get_parent child r Some parent  parent  set nodes
                           cast child  set nodes"
  assumes to_tree_order_subset:
    "heap_is_wellformed h  h  to_tree_order ptr r nodes  node  set nodes
                           h  to_tree_order node r nodes'  known_ptrs h
                           type_wf h  set nodes'  set nodes"

lemma to_tree_order_wf_is_l_to_tree_order_wf [instances]:
  "l_to_tree_order_wf heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs
                      to_tree_order get_parent get_child_nodes"
  using instances
  apply(auto simp add: l_to_tree_order_wf_def l_to_tree_order_wf_axioms_def)[1]
  using to_tree_order_ok
           apply blast
  using to_tree_order_ptrs_in_heap
          apply blast
  using to_tree_order_parent_child_rel
         apply(blast, blast)
  using to_tree_order_child2
       apply blast
  using to_tree_order_node_ptrs
      apply blast
  using to_tree_order_child
     apply blast
  using to_tree_order_ptr_in_result
    apply blast
  using to_tree_order_parent
   apply blast
  using to_tree_order_subset
  apply blast
  done


subsubsection ‹get\_root\_node›

locale l_to_tree_order_wf_get_root_node_wfCore_DOM =
  l_get_root_node_wfCore_DOM
  + l_to_tree_order_wf
begin
lemma to_tree_order_get_root_node:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  to_tree_order ptr r to"
  assumes "ptr'  set to"
  assumes "h  get_root_node ptr' r root_ptr"
  assumes "ptr''  set to"
  shows "h  get_root_node ptr'' r root_ptr"
proof -
  obtain ancestors' where ancestors': "h  get_ancestors ptr' r ancestors'"
    by (meson assms(1) assms(2) assms(3) assms(4) assms(5) get_ancestors_ok is_OK_returns_result_E
        to_tree_order_ptrs_in_heap )
  moreover have "ptr  set ancestors'"
    using h  get_ancestors ptr' r ancestors'
    using assms(1) assms(2) assms(3) assms(4) assms(5) get_ancestors_parent_child_rel
      to_tree_order_parent_child_rel by blast

  ultimately have "h  get_root_node ptr r root_ptr"
    using h  get_root_node ptr' r root_ptr
    using assms(1) assms(2) assms(3) get_ancestors_ptr get_ancestors_same_root_node by blast

  obtain ancestors'' where ancestors'': "h  get_ancestors ptr'' r ancestors''"
    by (meson assms(1) assms(2) assms(3) assms(4) assms(7) get_ancestors_ok is_OK_returns_result_E
        to_tree_order_ptrs_in_heap)
  moreover have "ptr  set ancestors''"
    using h  get_ancestors ptr'' r ancestors''
    using assms(1) assms(2) assms(3) assms(4) assms(7) get_ancestors_parent_child_rel
      to_tree_order_parent_child_rel by blast
  ultimately show ?thesis
    using h  get_root_node ptr r root_ptr assms(1) assms(2) assms(3) get_ancestors_ptr
      get_ancestors_same_root_node by blast
qed

lemma to_tree_order_same_root:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_root_node ptr r root_ptr"
  assumes "h  to_tree_order root_ptr r to"
  assumes "ptr'  set to"
  shows "h  get_root_node ptr' r root_ptr"
proof (insert assms(1)(*  assms(4) assms(5) *) assms(6), induct ptr' rule: heap_wellformed_induct_rev)
  case (step child)
  then show ?case
  proof (cases "h  get_root_node child r child")
    case True
    then have "child = root_ptr"
      using  assms(1) assms(2) assms(3) assms(5) step.prems
      by (metis (no_types, lifting) get_root_node_same_no_parent node_ptr_casts_commute3
          option.simps(3) returns_result_eq to_tree_order_child2 to_tree_order_node_ptrs)
    then show ?thesis
      using True by blast
  next
    case False
    then obtain child_node parent where "cast child_node = child"
      and "h  get_parent child_node r Some parent"
      by (metis assms(1) assms(2) assms(3) assms(4) assms(5) local.get_root_node_no_parent
          local.get_root_node_not_node_same local.get_root_node_same_no_parent
          local.to_tree_order_child2 local.to_tree_order_ptrs_in_heap node_ptr_casts_commute3
          step.prems)
    then show ?thesis
    proof (cases "child = root_ptr")
      case True
      then have "h  get_root_node root_ptr r root_ptr"
        using assms(4)
        using castnode_ptr2object_ptr child_node = child assms(1) assms(2) assms(3)
          get_root_node_no_parent get_root_node_same_no_parent
        by blast
      then show ?thesis
        using step assms(4)
        using True by blast
    next
      case False
      then have "parent  set to"
        using assms(5) step(2) to_tree_order_child h  get_parent child_node r Some parent
          cast child_node = child
        by (metis False assms(1) assms(2) assms(3) get_parent_child_dual)
      then show ?thesis
        using castnode_ptr2object_ptr child_node = child h  get_parent child_node r Some parent
          get_root_node_parent_same
        using step.hyps by blast
    qed

  qed
qed
end

interpretation i_to_tree_order_wf_get_root_node_wf?: l_to_tree_order_wf_get_root_node_wfCore_DOM
  known_ptr type_wf known_ptrs heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs
  get_disconnected_nodes get_disconnected_nodes_locs get_parent get_parent_locs get_ancestors
  get_ancestors_locs get_root_node get_root_node_locs to_tree_order
  using instances
  by(simp add: l_to_tree_order_wf_get_root_node_wfCore_DOM_def)

locale l_to_tree_order_wf_get_root_node_wf = l_type_wf + l_known_ptrs + l_to_tree_order_defs
  + l_get_root_node_defs + l_heap_is_wellformed_defs +
  assumes to_tree_order_get_root_node:
    "heap_is_wellformed h  type_wf h  known_ptrs h  h  to_tree_order ptr r to
                           ptr'  set to  h  get_root_node ptr' r root_ptr
                           ptr''  set to  h  get_root_node ptr'' r root_ptr"
  assumes to_tree_order_same_root:
    "heap_is_wellformed h  type_wf h  known_ptrs h
                           h  get_root_node ptr r root_ptr
                           h  to_tree_order root_ptr r to  ptr'  set to
                           h  get_root_node ptr' r root_ptr"

lemma to_tree_order_wf_get_root_node_wf_is_l_to_tree_order_wf_get_root_node_wf [instances]:
  "l_to_tree_order_wf_get_root_node_wf type_wf known_ptr known_ptrs to_tree_order
                                       get_root_node heap_is_wellformed"
  using instances
  apply(auto simp add: l_to_tree_order_wf_get_root_node_wf_def
      l_to_tree_order_wf_get_root_node_wf_axioms_def)[1]
  using to_tree_order_get_root_node apply blast
  using to_tree_order_same_root apply blast
  done


subsection ‹get\_owner\_document›

locale l_get_owner_document_wfCore_DOM =
  l_known_ptrs
  + l_heap_is_wellformed
  + l_get_root_nodeCore_DOM
  + l_get_ancestors
  + l_get_ancestors_wf
  + l_get_parent
  + l_get_parent_wf
  + l_get_root_node_wf
  + l_get_owner_documentCore_DOM
begin

lemma get_owner_document_disconnected_nodes:
  assumes "heap_is_wellformed h"
  assumes "h  get_disconnected_nodes document_ptr r disc_nodes"
  assumes "node_ptr  set disc_nodes"
  assumes known_ptrs: "known_ptrs h"
  assumes type_wf: "type_wf h"
  shows "h  get_owner_document (cast node_ptr) r document_ptr"
proof -
  have 2: "node_ptr |∈| node_ptr_kinds h"
    using assms heap_is_wellformed_disc_nodes_in_heap
    by blast
  have 3: "document_ptr |∈| document_ptr_kinds h"
    using assms(2) get_disconnected_nodes_ptr_in_heap by blast
  have 0:
    "∃!document_ptrset |h  document_ptr_kinds_M|r. node_ptr  set |h  get_disconnected_nodes document_ptr|r"
    by (metis (no_types, lifting) "3" DocumentMonad.ptr_kinds_ptr_kinds_M assms(1) assms(2) assms(3)
        disjoint_iff_not_equal l_heap_is_wellformed.heap_is_wellformed_one_disc_parent
        local.get_disconnected_nodes_ok local.l_heap_is_wellformed_axioms
        returns_result_select_result select_result_I2 type_wf)

  have "h  get_parent node_ptr r None"
    using heap_is_wellformed_children_disc_nodes_different child_parent_dual assms
    using "2" disjoint_iff_not_equal local.get_parent_child_dual local.get_parent_ok
      returns_result_select_result split_option_ex
    by (metis (no_types, lifting))

  then have 4: "h  get_root_node (cast node_ptr) r cast node_ptr"
    using 2 get_root_node_no_parent
    by blast
  obtain document_ptrs where document_ptrs: "h  document_ptr_kinds_M r document_ptrs"
    by simp

  then
  have "h  ok (filter_M (λdocument_ptr. do {
          disconnected_nodes  get_disconnected_nodes document_ptr;
          return (((castnode_ptr2object_ptr node_ptr))  cast ` set disconnected_nodes)
        }) document_ptrs)"
    using assms(1) get_disconnected_nodes_ok  type_wf unfolding heap_is_wellformed_def
    by(auto intro!: bind_is_OK_I2 filter_M_is_OK_I bind_pure_I)
  then obtain candidates where
    candidates: "h  filter_M (λdocument_ptr. do {
          disconnected_nodes  get_disconnected_nodes document_ptr;
          return (((castnode_ptr2object_ptr node_ptr))  cast ` set disconnected_nodes)
        }) document_ptrs r candidates"
    by auto


  have eq: "document_ptr. document_ptr |∈| document_ptr_kinds h
            node_ptr  set |h  get_disconnected_nodes document_ptr|r  |h  do {
          disconnected_nodes  get_disconnected_nodes document_ptr;
          return (((castnode_ptr2object_ptr node_ptr))  cast ` set disconnected_nodes)
        }|r"
    apply(auto dest!: get_disconnected_nodes_ok[OF type_wf]
        intro!: select_result_I[where P=id, simplified] elim!: bind_returns_result_E2)[1]
    apply(drule select_result_E[where P=id, simplified])
    by(auto elim!: bind_returns_result_E2)

  have filter: "filter (λdocument_ptr. |h  do {
          disconnected_nodes  get_disconnected_nodes document_ptr;
          return (castnode_ptr2object_ptr node_ptr  cast ` set disconnected_nodes)
        }|r) document_ptrs = [document_ptr]"
    apply(rule filter_ex1)
    using 0 document_ptrs apply(simp)[1]
    using eq
    using local.get_disconnected_nodes_ok apply auto[1]
    using assms(2) assms(3)
      apply(auto intro!: intro!: select_result_I[where P=id, simplified]
        elim!: bind_returns_result_E2)[1]
    using returns_result_eq apply fastforce
    using document_ptrs  3 apply(simp)
    using document_ptrs
    by simp
  have "h  filter_M (λdocument_ptr. do {
          disconnected_nodes  get_disconnected_nodes document_ptr;
          return (((castnode_ptr2object_ptr node_ptr))  cast ` set disconnected_nodes)
        }) document_ptrs r [document_ptr]"
    apply(rule filter_M_filter2)
    using get_disconnected_nodes_ok document_ptrs 3  assms(1)  type_wf filter
    unfolding heap_is_wellformed_def
    by(auto intro: bind_pure_I bind_is_OK_I2)

  with 4 document_ptrs have "h  a_get_owner_documentnode_ptr node_ptr () r document_ptr"
    by(auto simp add: a_get_owner_documentnode_ptr_def
        intro!:  bind_pure_returns_result_I filter_M_pure_I bind_pure_I
        split: option.splits)[1]
  moreover have "known_ptr (castnode_ptr2object_ptr node_ptr)"
    using "4" assms(1)  known_ptrs type_wf known_ptrs_known_ptr "2" node_ptr_kinds_commutes by blast
  ultimately show ?thesis
    using 2
    apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def)[1]
    apply(split invoke_splits, (rule conjI | rule impI)+)+
       apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl])
       apply(drule(1)  known_ptr_not_character_data_ptr)
       apply(drule(1)  known_ptr_not_element_ptr)
       apply(simp add: NodeClass.known_ptr_defs)
    by(auto split: option.splits intro!: bind_pure_returns_result_I)
qed

lemma in_disconnected_nodes_no_parent:
  assumes "heap_is_wellformed h"
    and "h  get_parent node_ptr r None"
    and "h  get_owner_document (cast node_ptr) r owner_document"
    and "h  get_disconnected_nodes owner_document r disc_nodes"
    and known_ptrs: "known_ptrs h"
    and type_wf: "type_wf h"
  shows "node_ptr  set disc_nodes"
proof -
  have 2: "cast node_ptr |∈| object_ptr_kinds h"
    using assms(3) get_owner_document_ptr_in_heap by fast
  then have 3: "h  get_root_node (cast node_ptr) r cast node_ptr"
    using assms(2) local.get_root_node_no_parent by blast

  have "¬(parent_ptr. parent_ptr |∈| object_ptr_kinds h  node_ptr  set |h  get_child_nodes parent_ptr|r)"
    apply(auto)[1]
    using assms(2) child_parent_dual[OF assms(1)] type_wf
      assms(1) assms(5) get_child_nodes_ok known_ptrs_known_ptr option.simps(3)
      returns_result_eq returns_result_select_result
    by (metis (no_types, opaque_lifting))
  moreover have "node_ptr |∈| node_ptr_kinds h"
    using assms(2) get_parent_ptr_in_heap by blast
  ultimately
  have 0: "document_ptrset |h  document_ptr_kinds_M|r. node_ptr  set |h  get_disconnected_nodes document_ptr|r"
    by (metis DocumentMonad.ptr_kinds_ptr_kinds_M assms(1) heap_is_wellformed_children_disc_nodes)
  then  obtain document_ptr where
    document_ptr: "document_ptrset |h  document_ptr_kinds_M|r" and
    node_ptr_in_disc_nodes: "node_ptr  set |h  get_disconnected_nodes document_ptr|r"
    by auto
  then show ?thesis
    using get_owner_document_disconnected_nodes known_ptrs type_wf assms
    using DocumentMonad.ptr_kinds_ptr_kinds_M assms(1) assms(3) assms(4) get_disconnected_nodes_ok
      returns_result_select_result select_result_I2
    by (metis (no_types, opaque_lifting) )
qed

lemma get_owner_document_owner_document_in_heap:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_owner_document ptr r owner_document"
  shows "owner_document |∈| document_ptr_kinds h"
  using assms
  apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
  apply(split invoke_split_asm)+
proof -
  assume "h  invoke [] ptr () r owner_document"
  then show "owner_document |∈| document_ptr_kinds h"
    by (meson invoke_empty is_OK_returns_result_I)
next
  assume "h  Heap_Error_Monad.bind (check_in_heap ptr)
          (λ_. (local.a_get_owner_documentdocument_ptr  the  castobject_ptr2document_ptr) ptr ())
    r owner_document"
  then show "owner_document |∈| document_ptr_kinds h"
    by(auto simp add: a_get_owner_documentdocument_ptr_def elim!: bind_returns_result_E2 split: if_splits)
next
  assume 0: "heap_is_wellformed h"
    and 1: "type_wf h"
    and 2: "known_ptrs h"
    and 3: "¬ is_element_ptrobject_ptr ptr"
    and 4: "is_character_data_ptrobject_ptr ptr"
    and 5: "h  Heap_Error_Monad.bind (check_in_heap ptr)
(λ_. (local.a_get_owner_documentnode_ptr  the  castobject_ptr2node_ptr) ptr ()) r owner_document"
  then obtain root where
    root: "h  get_root_node ptr r root"
    by(auto simp add: a_get_owner_documentnode_ptr_def elim!: bind_returns_result_E2
        split: option.splits)

  then show ?thesis
  proof (cases "is_document_ptr root")
    case True
    then show ?thesis
      using 4 5 root
      apply(auto simp add: a_get_owner_documentnode_ptr_def elim!:  bind_returns_result_E2
          intro!: filter_M_pure_I bind_pure_I split: option.splits)[1]
       apply(drule(1) returns_result_eq) apply(auto)[1]
      using "0" "1" "2" document_ptr_kinds_commutes local.get_root_node_root_in_heap by blast
  next
    case False
    have "known_ptr root"
      using "0" "1" "2" local.get_root_node_root_in_heap local.known_ptrs_known_ptr root by blast
    have "root |∈| object_ptr_kinds h"
      using root
      using "0" "1" "2" local.get_root_node_root_in_heap
      by blast
    then have "is_node_ptr_kind root"
      using False known_ptr root
      apply(simp add: known_ptr_impl known_ptr_defs CharacterDataClass.known_ptr_defs
          ElementClass.known_ptr_defs NodeClass.known_ptr_defs)
      using is_node_ptr_kind_none by force
    then
    have "(document_ptr  fset (document_ptr_kinds h). root  cast ` set |h  get_disconnected_nodes document_ptr|r)"
      by (metis (no_types, opaque_lifting) "0" "1" "2" root |∈| object_ptr_kinds h local.child_parent_dual
          local.get_child_nodes_ok local.get_root_node_same_no_parent local.heap_is_wellformed_children_disc_nodes
          local.known_ptrs_known_ptr node_ptr_casts_commute3 node_ptr_inclusion node_ptr_kinds_commutes
          option.distinct(1) returns_result_eq returns_result_select_result root)
    then obtain some_owner_document where
      "some_owner_document |∈| document_ptr_kinds h" and
      "root  cast ` set |h  get_disconnected_nodes some_owner_document|r"
      by auto
    then
    obtain candidates where
      candidates: "h  filter_M
             (λdocument_ptr.
                 Heap_Error_Monad.bind (get_disconnected_nodes document_ptr)
                  (λdisconnected_nodes. return (root  castnode_ptr2object_ptr ` set disconnected_nodes)))
             (sorted_list_of_set (fset (document_ptr_kinds h)))
       r candidates"
      by (metis (no_types, lifting) "1" bind_is_OK_I2 bind_pure_I filter_M_is_OK_I finite_fset
          is_OK_returns_result_E local.get_disconnected_nodes_ok local.get_disconnected_nodes_pure
          return_ok return_pure sorted_list_of_set(1))
    then have "some_owner_document  set candidates"
      apply(rule filter_M_in_result_if_ok)
      using some_owner_document |∈| document_ptr_kinds h
        root  cast ` set |h  get_disconnected_nodes some_owner_document|r
        apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1]
       apply (simp add: some_owner_document |∈| document_ptr_kinds h)
      using "1" root  castnode_ptr2object_ptr ` set |h  get_disconnected_nodes some_owner_document|r
        some_owner_document |∈| document_ptr_kinds h
        local.get_disconnected_nodes_ok by auto
    then have "candidates  []"
      by auto
    then have "owner_document  set candidates"
      using 5 root 4
      apply(auto simp add: a_get_owner_documentnode_ptr_def elim!:  bind_returns_result_E2
          intro!: filter_M_pure_I bind_pure_I split: option.splits)[1]
       apply (metis candidates list.set_sel(1) returns_result_eq)
      by (metis is_node_ptr_kind root node_ptr_no_document_ptr_cast returns_result_eq)

    then show ?thesis
      using candidates
      by (meson bind_pure_I bind_returns_result_E2 filter_M_holds_for_result is_OK_returns_result_I
          local.get_disconnected_nodes_ptr_in_heap local.get_disconnected_nodes_pure return_pure)
  qed
next
  assume 0: "heap_is_wellformed h"
    and 1: "type_wf h"
    and 2: "known_ptrs h"
    and 3: "is_element_ptrobject_ptr ptr"
    and 4: "h  Heap_Error_Monad.bind (check_in_heap ptr) (λ_. (local.a_get_owner_documentnode_ptr  the  castobject_ptr2node_ptr) ptr ()) r owner_document"
  then obtain root where
    root: "h  get_root_node ptr r root"
    by(auto simp add: a_get_owner_documentnode_ptr_def elim!: bind_returns_result_E2  split: option.splits)

  then show ?thesis
  proof (cases "is_document_ptr root")
    case True
    then show ?thesis
      using 3 4 root
      apply(auto simp add: a_get_owner_documentnode_ptr_def elim!:  bind_returns_result_E2
          intro!: filter_M_pure_I bind_pure_I split: option.splits)[1]
       apply(drule(1) returns_result_eq) apply(auto)[1]
      using "0" "1" "2" document_ptr_kinds_commutes local.get_root_node_root_in_heap by blast
  next
    case False
    have "known_ptr root"
      using "0" "1" "2" local.get_root_node_root_in_heap local.known_ptrs_known_ptr root by blast
    have "root |∈| object_ptr_kinds h"
      using root
      using "0" "1" "2" local.get_root_node_root_in_heap
      by blast
    then have "is_node_ptr_kind root"
      using False known_ptr root
      apply(simp add: known_ptr_impl known_ptr_defs CharacterDataClass.known_ptr_defs
          ElementClass.known_ptr_defs NodeClass.known_ptr_defs)
      using is_node_ptr_kind_none by force
    then
    have "(document_ptr  fset (document_ptr_kinds h). root  cast ` set |h  get_disconnected_nodes document_ptr|r)"
      by (metis (no_types, opaque_lifting) "0" "1" "2" root |∈| object_ptr_kinds h
          local.child_parent_dual local.get_child_nodes_ok local.get_root_node_same_no_parent
          local.heap_is_wellformed_children_disc_nodes local.known_ptrs_known_ptr node_ptr_casts_commute3
          node_ptr_inclusion node_ptr_kinds_commutes option.distinct(1) returns_result_eq
          returns_result_select_result root)
    then obtain some_owner_document where
      "some_owner_document |∈| document_ptr_kinds h" and
      "root  cast ` set |h  get_disconnected_nodes some_owner_document|r"
      by auto
    then
    obtain candidates where
      candidates: "h  filter_M
             (λdocument_ptr.
                 Heap_Error_Monad.bind (get_disconnected_nodes document_ptr)
                  (λdisconnected_nodes. return (root  castnode_ptr2object_ptr ` set disconnected_nodes)))
             (sorted_list_of_set (fset (document_ptr_kinds h)))
       r candidates"
      by (metis (no_types, lifting) "1" bind_is_OK_I2 bind_pure_I filter_M_is_OK_I finite_fset
          is_OK_returns_result_E local.get_disconnected_nodes_ok local.get_disconnected_nodes_pure
          return_ok return_pure sorted_list_of_set(1))
    then have "some_owner_document  set candidates"
      apply(rule filter_M_in_result_if_ok)
      using some_owner_document |∈| document_ptr_kinds h
        root  cast ` set |h  get_disconnected_nodes some_owner_document|r
        apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1]
       apply (simp add: some_owner_document |∈| document_ptr_kinds h)
      using "1" root  castnode_ptr2object_ptr ` set |h  get_disconnected_nodes some_owner_document|r
        some_owner_document |∈| document_ptr_kinds h local.get_disconnected_nodes_ok
      by auto

    then have "candidates  []"
      by auto
    then have "owner_document  set candidates"
      using 4 root 3
      apply(auto simp add: a_get_owner_documentnode_ptr_def elim!:  bind_returns_result_E2
          intro!: filter_M_pure_I bind_pure_I split: option.splits)[1]
       apply (metis candidates list.set_sel(1) returns_result_eq)
      by (metis is_node_ptr_kind root node_ptr_no_document_ptr_cast returns_result_eq)

    then show ?thesis
      using candidates
      by (meson bind_pure_I bind_returns_result_E2 filter_M_holds_for_result is_OK_returns_result_I
          local.get_disconnected_nodes_ptr_in_heap local.get_disconnected_nodes_pure return_pure)
  qed
qed

lemma get_owner_document_ok:
  assumes "heap_is_wellformed h"  "known_ptrs h" "type_wf h"
  assumes "ptr |∈| object_ptr_kinds h"
  shows "h  ok (get_owner_document ptr)"
proof -
  have "known_ptr ptr"
    using assms(2) assms(4) local.known_ptrs_known_ptr
    by blast
  then show ?thesis
    apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
    apply(split invoke_splits, (rule conjI | rule impI)+)+
       apply(auto simp add: known_ptr_impl)[1]
    using NodeClass.a_known_ptr_def known_ptr_not_character_data_ptr known_ptr_not_document_ptr
      known_ptr_not_element_ptr
       apply blast
    using assms(4)
      apply(auto simp add: a_get_owner_documentdocument_ptr_def a_get_owner_documentnode_ptr_def
        intro!: bind_is_OK_pure_I)[1]
      apply (metis (no_types, lifting) document_ptr_casts_commute3 document_ptr_kinds_commutes
        is_document_ptr_kind_none option.case_eq_if)
    using assms(4)
     apply(auto simp add: a_get_owner_documentdocument_ptr_def a_get_owner_documentnode_ptr_def
        intro!: bind_is_OK_pure_I)[1]
      apply (metis (no_types, lifting) assms(1) assms(2) assms(3) is_node_ptr_kind_none
        local.get_root_node_ok node_ptr_casts_commute3 option.case_eq_if)
    using assms(4)
     apply(auto simp add: a_get_owner_documentdocument_ptr_def a_get_owner_documentnode_ptr_def
        intro!: bind_is_OK_pure_I)[1]
     apply(auto split: option.splits intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I
        filter_M_is_OK_I)[1]
    using assms(3) local.get_disconnected_nodes_ok
     apply blast
    apply (simp add: assms(1) assms(2) assms(3) local.get_root_node_ok)
    using assms(4)
    apply(auto simp add: a_get_owner_documentdocument_ptr_def a_get_owner_documentnode_ptr_def
        intro!: bind_is_OK_pure_I)[1]
     apply(auto split: option.splits intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I
        filter_M_is_OK_I)[1]
     apply (simp add: assms(1) assms(2) assms(3) local.get_root_node_ok)[1]
    apply(auto split: option.splits intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I
        filter_M_is_OK_I)[1]
    using assms(3) local.get_disconnected_nodes_ok by blast
qed

lemma get_owner_document_child_same:
  assumes "heap_is_wellformed h"  "known_ptrs h" "type_wf h"
  assumes "h  get_child_nodes ptr r children"
  assumes "child  set children"
  shows "h  get_owner_document ptr r owner_document  h  get_owner_document (cast child) r owner_document"
proof -
  have "ptr |∈| object_ptr_kinds h"
    by (meson assms(4) is_OK_returns_result_I local.get_child_nodes_ptr_in_heap)
  then have "known_ptr ptr"
    using assms(2) local.known_ptrs_known_ptr by blast

  have "cast child |∈| object_ptr_kinds h"
    using assms(1) assms(4) assms(5) local.heap_is_wellformed_children_in_heap node_ptr_kinds_commutes
    by blast
  then
  have "known_ptr (cast child)"
    using assms(2) local.known_ptrs_known_ptr by blast

  obtain root where root: "h  get_root_node ptr r root"
    by (meson ptr |∈| object_ptr_kinds h assms(1) assms(2) assms(3) is_OK_returns_result_E
        local.get_root_node_ok)
  then have "h  get_root_node (cast child) r root"
    using assms(1) assms(2) assms(3) assms(4) assms(5) local.child_parent_dual
      local.get_root_node_parent_same
    by blast

  have "h  get_owner_document ptr r owner_document 
h  a_get_owner_documentnode_ptr child () r owner_document"
  proof (cases "is_document_ptr ptr")
    case True
    then obtain document_ptr where document_ptr: "castdocument_ptr2object_ptr document_ptr = ptr"
      using case_optionE document_ptr_casts_commute by blast
    then have "root = cast document_ptr"
      using root
      by(auto simp add: get_root_node_def get_ancestors_def elim!: bind_returns_result_E2
          split: option.splits)

    then have "h  a_get_owner_documentdocument_ptr document_ptr () r owner_document 
h  a_get_owner_documentnode_ptr child () r owner_document"
      using document_ptr
        h  get_root_node (cast child) r root[simplified root = cast document_ptr document_ptr]
      apply(auto simp add:  a_get_owner_documentnode_ptr_def a_get_owner_documentdocument_ptr_def
          elim!: bind_returns_result_E2 dest!: bind_returns_result_E3[rotated,
            OF h  get_root_node (cast child) r root[simplified root = cast document_ptr document_ptr], rotated]
          intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: if_splits option.splits)[1]
      using ptr |∈| object_ptr_kinds h document_ptr_kinds_commutes by blast
    then show ?thesis
      using known_ptr ptr
      apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def known_ptr_impl)[1]
       apply(split invoke_splits, ((rule conjI | rule impI)+)?)+
          apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl])
          apply(drule(1)  known_ptr_not_character_data_ptr)
          apply(drule(1)  known_ptr_not_element_ptr)
          apply(simp add: NodeClass.known_ptr_defs)
      using ptr |∈| object_ptr_kinds h True
      by(auto simp add: document_ptr[symmetric] intro!: bind_pure_returns_result_I
          split: option.splits)
  next
    case False
    then obtain node_ptr where node_ptr: "castnode_ptr2object_ptr node_ptr = ptr"
      using known_ptr ptr
      by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs
          ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)
    then have "h  a_get_owner_documentnode_ptr node_ptr () r owner_document 
h  a_get_owner_documentnode_ptr child () r owner_document"
      using root h  get_root_node (cast child) r root
      unfolding a_get_owner_documentnode_ptr_def
      by (meson bind_pure_returns_result_I bind_returns_result_E3 local.get_root_node_pure)
    then show ?thesis
      using known_ptr ptr
      apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def known_ptr_impl)[1]
       apply(split invoke_splits, ((rule conjI | rule impI)+)?)+
          apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl])
          apply(drule(1)  known_ptr_not_character_data_ptr)
          apply(drule(1)  known_ptr_not_element_ptr)
          apply(simp add: NodeClass.known_ptr_defs)
      using cast child |∈| object_ptr_kinds h ptr |∈| object_ptr_kinds h False
         apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
      using cast child |∈| object_ptr_kinds h ptr |∈| object_ptr_kinds h False
        apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
      using cast child |∈| object_ptr_kinds h ptr |∈| object_ptr_kinds h False
       apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
      using cast child |∈| object_ptr_kinds h ptr |∈| object_ptr_kinds h False
      apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
      apply(split invoke_splits, ((rule conjI | rule impI)+)?)+
         apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
         apply (meson invoke_empty is_OK_returns_result_I)
        apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
       apply(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
      by(auto simp add: node_ptr[symmetric] intro!: bind_pure_returns_result_I split: )[1]
  qed
  then show ?thesis
    using known_ptr (cast child)
    apply(auto simp add: get_owner_document_def[of "castnode_ptr2object_ptr child"]
        a_get_owner_document_tups_def known_ptr_impl)[1]
     apply(split invoke_splits, ((rule conjI | rule impI)+)?)+
        apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl])
        apply(drule(1)  known_ptr_not_character_data_ptr)
        apply(drule(1)  known_ptr_not_element_ptr)
        apply(simp add: NodeClass.known_ptr_defs)
    using cast child |∈| object_ptr_kinds h ptr |∈| object_ptr_kinds h
       apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1]
    using cast child |∈| object_ptr_kinds h ptr |∈| object_ptr_kinds h
      apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1]
    using cast child |∈| object_ptr_kinds h ptr |∈| object_ptr_kinds h
     apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1]
    using cast child |∈| object_ptr_kinds h ptr |∈| object_ptr_kinds h
    apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1]
    by (smt (verit) castnode_ptr2object_ptr child |∈| object_ptr_kinds h cast_document_ptr_not_node_ptr(1)
        comp_apply invoke_empty invoke_not invoke_returns_result is_OK_returns_result_I
        node_ptr_casts_commute2 option.sel)
qed

end

locale l_get_owner_document_wf = l_heap_is_wellformed_defs + l_type_wf + l_known_ptrs
  + l_get_disconnected_nodes_defs + l_get_owner_document_defs
  + l_get_parent_defs + l_get_child_nodes_defs +
  assumes get_owner_document_disconnected_nodes:
    "heap_is_wellformed h 
     known_ptrs h 
     type_wf h 
     h  get_disconnected_nodes document_ptr r disc_nodes 
     node_ptr  set disc_nodes 
     h  get_owner_document (cast node_ptr) r document_ptr"
  assumes in_disconnected_nodes_no_parent:
    "heap_is_wellformed h 
    h  get_parent node_ptr r None
    h  get_owner_document (cast node_ptr) r owner_document 
    h  get_disconnected_nodes owner_document r disc_nodes 
    known_ptrs h 
    type_wf h
    node_ptr  set disc_nodes"
  assumes get_owner_document_owner_document_in_heap:
    "heap_is_wellformed h  type_wf h  known_ptrs h 
h  get_owner_document ptr r owner_document 
owner_document |∈| document_ptr_kinds h"
  assumes get_owner_document_ok:
    "heap_is_wellformed h  known_ptrs h  type_wf h  ptr |∈| object_ptr_kinds h
                           h  ok (get_owner_document ptr)"
  assumes get_owner_document_child_same:
    "heap_is_wellformed h  type_wf h  known_ptrs h  h  get_child_nodes ptr r children 
child  set children  h  get_owner_document ptr r owner_document 
h  get_owner_document (cast child) r owner_document"

interpretation i_get_owner_document_wf?: l_get_owner_document_wfCore_DOM
  known_ptr known_ptrs type_wf heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs
  get_disconnected_nodes get_disconnected_nodes_locs get_parent get_parent_locs get_ancestors
  get_ancestors_locs get_root_node get_root_node_locs get_owner_document
  by(auto simp add: l_get_owner_document_wfCore_DOM_def instances)
declare l_get_owner_document_wfCore_DOM_axioms [instances]

lemma get_owner_document_wf_is_l_get_owner_document_wf [instances]:
  "l_get_owner_document_wf heap_is_wellformed type_wf known_ptr known_ptrs get_disconnected_nodes
                           get_owner_document get_parent get_child_nodes"
  using known_ptrs_is_l_known_ptrs
  apply(auto simp add: l_get_owner_document_wf_def l_get_owner_document_wf_axioms_def)[1]
  using get_owner_document_disconnected_nodes apply fast
  using  in_disconnected_nodes_no_parent apply fast
  using get_owner_document_owner_document_in_heap apply fast
  using get_owner_document_ok apply fast
  using get_owner_document_child_same apply (fast, fast)
  done


subsubsection ‹get\_root\_node›

locale l_get_owner_document_wf_get_root_node_wfCore_DOM =
  l_get_root_nodeCore_DOM +
  l_get_root_node_wf +
  l_get_owner_documentCore_DOM +
  l_get_owner_document_wf
begin

lemma get_root_node_document:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_root_node ptr r root"
  assumes "is_document_ptr_kind root"
  shows "h  get_owner_document ptr r the (cast root)"
proof -
  have "ptr |∈| object_ptr_kinds h"
    using assms(4)
    by (meson is_OK_returns_result_I local.get_root_node_ptr_in_heap)
  then have "known_ptr ptr"
    using assms(3) local.known_ptrs_known_ptr by blast
  {
    assume "is_document_ptr_kind ptr"
    then have "ptr = root"
      using assms(4)
      by(auto simp add: get_root_node_def get_ancestors_def elim!: bind_returns_result_E2
          split: option.splits)
    then have ?thesis
      using is_document_ptr_kind ptr known_ptr ptr ptr |∈| object_ptr_kinds h
      apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def)[1]
      apply(split invoke_splits, (rule conjI | rule impI)+)+
         apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl])
         apply(drule(1) known_ptr_not_character_data_ptr)
         apply(drule(1) known_ptr_not_element_ptr)
         apply(simp add: NodeClass.known_ptr_defs)
      by(auto simp add: a_get_owner_documentdocument_ptr_def intro!: bind_pure_returns_result_I
          split: option.splits)
  }
  moreover
  {
    assume "is_node_ptr_kind ptr"
    then have ?thesis
      using known_ptr ptr ptr |∈| object_ptr_kinds h
      apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def)[1]
      apply(split invoke_splits, (rule conjI | rule impI)+)+
         apply(drule(1) known_ptr_not_document_ptr[folded known_ptr_impl])
         apply(drule(1) known_ptr_not_character_data_ptr)
         apply(drule(1) known_ptr_not_element_ptr)
         apply(simp add: NodeClass.known_ptr_defs)
        apply(auto split: option.splits)[1]
      using h  get_root_node ptr r root assms(5)
      by(auto simp add: a_get_owner_documentnode_ptr_def is_document_ptr_kind_def
          intro!: bind_pure_returns_result_I  split: option.splits)[2]
  }
  ultimately
  show ?thesis
    using known_ptr ptr
    by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs
        ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)
qed

lemma get_root_node_same_owner_document:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_root_node ptr r root"
  shows "h  get_owner_document ptr r owner_document  h  get_owner_document root r owner_document"
proof -
  have "ptr |∈| object_ptr_kinds h"
    by (meson assms(4) is_OK_returns_result_I local.get_root_node_ptr_in_heap)
  have "root |∈| object_ptr_kinds h"
    using assms(1) assms(2) assms(3) assms(4) local.get_root_node_root_in_heap by blast
  have "known_ptr ptr"
    using ptr |∈| object_ptr_kinds h assms(3) local.known_ptrs_known_ptr by blast
  have "known_ptr root"
    using root |∈| object_ptr_kinds h assms(3) local.known_ptrs_known_ptr by blast
  show ?thesis
  proof (cases "is_document_ptr_kind ptr")
    case True
    then
    have "ptr = root"
      using assms(4)
      apply(auto simp add: get_root_node_def elim!: bind_returns_result_E2)[1]
      by (metis document_ptr_casts_commute3 last_ConsL local.get_ancestors_not_node node_ptr_no_document_ptr_cast)
    then show ?thesis
      by auto
  next
    case False
    then have "is_node_ptr_kind ptr"
      using known_ptr ptr
      by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs
          ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)
    then obtain node_ptr where node_ptr: "ptr = castnode_ptr2object_ptr node_ptr"
      by (metis node_ptr_casts_commute3)
    show ?thesis
    proof
      assume "h  get_owner_document ptr r owner_document"
      then have "h  local.a_get_owner_documentnode_ptr node_ptr () r owner_document"
        using node_ptr
        apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
        apply(split invoke_splits)+
           apply (meson invoke_empty is_OK_returns_result_I)
        by(auto  elim!: bind_returns_result_E2 split: option.splits)

      show "h  get_owner_document root r owner_document"
      proof (cases "is_document_ptr_kind root")
        case True
        have "is_document_ptr root"
          using True known_ptr root
          by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs
              ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)
        have "root = cast owner_document"
          using True
          by (smt (verit) h  get_owner_document ptr r owner_document assms(1) assms(2) assms(3) assms(4)
              document_ptr_casts_commute3 get_root_node_document returns_result_eq)
        then show ?thesis
          apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
          apply(split invoke_splits, (rule conjI | rule impI)+)+
          using is_document_ptrobject_ptr root apply blast
          using root |∈| object_ptr_kinds h
          by(auto simp add: a_get_owner_documentdocument_ptr_def is_node_ptr_kind_none)

      next
        case False
        then have "is_node_ptr_kind root"
          using known_ptr root
          by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs
              ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)
        then obtain root_node_ptr where root_node_ptr: "root = castnode_ptr2object_ptr root_node_ptr"
          by (metis node_ptr_casts_commute3)
        then have "h  local.a_get_owner_documentnode_ptr root_node_ptr () r owner_document"
          using h  local.a_get_owner_documentnode_ptr node_ptr () r owner_document assms(4)
          apply(auto simp add: a_get_owner_documentnode_ptr_def elim!: bind_returns_result_E2
              intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: option.splits)[1]
           apply (metis assms(1) assms(2) assms(3) local.get_root_node_no_parent
              local.get_root_node_same_no_parent node_ptr returns_result_eq)
          using is_node_ptr_kind root node_ptr returns_result_eq by fastforce
        then show ?thesis
          apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
          apply(split invoke_splits, (rule conjI | rule impI)+)+
          using is_node_ptr_kind root known_ptr root
             apply(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs
              CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs
              split: option.splits)[1]
            apply(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs
              CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs
              split: option.splits)[1]
          using root |∈| object_ptr_kinds h
          by(auto simp add: root_node_ptr)
      qed
    next
      assume "h  get_owner_document root r owner_document"
      show "h  get_owner_document ptr r owner_document"
      proof (cases "is_document_ptr_kind root")
        case True
        have "root = cast owner_document"
          using h  get_owner_document root r owner_document
          apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
          apply(split invoke_splits)+
             apply (meson invoke_empty is_OK_returns_result_I)
            apply(auto simp add: True a_get_owner_documentdocument_ptr_def elim!: bind_returns_result_E2
              split: if_splits)[1]
           apply (metis True cast_document_ptr_not_node_ptr(2) is_document_ptr_kind_obtains
              is_node_ptr_kind_none node_ptr_casts_commute3 option.case_eq_if)
          by (metis True cast_document_ptr_not_node_ptr(1) document_ptr_casts_commute3
              is_node_ptr_kind_none node_ptr_casts_commute3 option.case_eq_if)
        then show ?thesis
          using assms(1) assms(2) assms(3) assms(4) get_root_node_document
          by fastforce
      next
        case False
        then have "is_node_ptr_kind root"
          using known_ptr root
          by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs
              ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)
        then obtain root_node_ptr where root_node_ptr: "root = castnode_ptr2object_ptr root_node_ptr"
          by (metis node_ptr_casts_commute3)
        then have "h  local.a_get_owner_documentnode_ptr root_node_ptr () r owner_document"
          using h  get_owner_document root r owner_document
          apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
          apply(split invoke_splits)+
             apply (meson invoke_empty is_OK_returns_result_I)
          by(auto simp add: is_document_ptr_kind_none elim!: bind_returns_result_E2)
        then have "h  local.a_get_owner_documentnode_ptr node_ptr () r owner_document"
          apply(auto simp add: a_get_owner_documentnode_ptr_def elim!: bind_returns_result_E2
              intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: option.splits)[1]
          using assms(1) assms(2) assms(3) assms(4) local.get_root_node_no_parent
            local.get_root_node_same_no_parent node_ptr returns_result_eq root_node_ptr
          by fastforce+
        then show ?thesis
          apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
          apply(split invoke_splits, (rule conjI | rule impI)+)+
          using node_ptr  known_ptr ptr ptr |∈| object_ptr_kinds h

          by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs
              CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs
              intro!: bind_pure_returns_result_I split: option.splits)
      qed
    qed
  qed
qed
end

interpretation get_owner_document_wf_get_root_node_wf?: l_get_owner_document_wf_get_root_node_wfCore_DOM
  type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs
  get_ancestors get_ancestors_locs get_root_node get_root_node_locs heap_is_wellformed
  parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs get_owner_document
  by(auto simp add: l_get_owner_document_wf_get_root_node_wfCore_DOM_def instances)
declare l_get_owner_document_wf_get_root_node_wfCore_DOM_axioms [instances]

locale l_get_owner_document_wf_get_root_node_wf = l_heap_is_wellformed_defs + l_type_wf +
  l_known_ptrs + l_get_root_node_defs + l_get_owner_document_defs +
  assumes get_root_node_document:
    "heap_is_wellformed h  type_wf h  known_ptrs h  h  get_root_node ptr r root 
is_document_ptr_kind root  h  get_owner_document ptr r the (cast root)"
  assumes get_root_node_same_owner_document:
    "heap_is_wellformed h  type_wf h  known_ptrs h  h  get_root_node ptr r root 
h  get_owner_document ptr r owner_document  h  get_owner_document root r owner_document"

lemma get_owner_document_wf_get_root_node_wf_is_l_get_owner_document_wf_get_root_node_wf [instances]:
  "l_get_owner_document_wf_get_root_node_wf heap_is_wellformed type_wf known_ptr known_ptrs
get_root_node get_owner_document"
  apply(auto simp add: l_get_owner_document_wf_get_root_node_wf_def
      l_get_owner_document_wf_get_root_node_wf_axioms_def instances)[1]
  using get_root_node_document apply blast
  using get_root_node_same_owner_document apply (blast, blast)
  done


subsection ‹Preserving heap-wellformedness›

subsection ‹set\_attribute›

locale l_set_attribute_wfCore_DOM =
  l_get_parent_wf2Core_DOM +
  l_set_attributeCore_DOM +
  l_set_attribute_get_disconnected_nodes +
  l_set_attribute_get_child_nodes
begin
lemma set_attribute_preserves_wellformedness:
  assumes "heap_is_wellformed h"
    and "h  set_attribute element_ptr k v h h'"
  shows "heap_is_wellformed h'"
  thm preserves_wellformedness_writes_needed
  apply(rule preserves_wellformedness_writes_needed[OF assms set_attribute_writes])
  using  set_attribute_get_child_nodes
    apply(fast)
  using set_attribute_get_disconnected_nodes apply(fast)
  by(auto simp add: all_args_def set_attribute_locs_def)
end


subsection ‹remove\_child›

locale l_remove_child_wfCore_DOM =
  l_remove_childCore_DOM +
  l_get_parent_wfCore_DOM +
  l_heap_is_wellformed +
  l_set_disconnected_nodes_get_child_nodes
begin
lemma remove_child_removes_parent:
  assumes wellformed: "heap_is_wellformed h"
    and remove_child: "h  remove_child ptr child h h2"
    and known_ptrs: "known_ptrs h"
    and type_wf: "type_wf h"
  shows "h2  get_parent child r None"
proof -
  obtain children where children: "h  get_child_nodes ptr r children"
    using remove_child remove_child_def by auto
  then have "child  set children"
    using remove_child remove_child_def
    by(auto elim!: bind_returns_heap_E dest: returns_result_eq split: if_splits)
  then have h1: "other_ptr other_children. other_ptr  ptr
                 h  get_child_nodes other_ptr r other_children  child  set other_children"
    using assms(1) known_ptrs  type_wf child_parent_dual
    by (meson child_parent_dual children option.inject returns_result_eq)

  have known_ptr: "known_ptr ptr"
    using known_ptrs
    by (meson is_OK_returns_heap_I l_known_ptrs.known_ptrs_known_ptr l_known_ptrs_axioms
        remove_child remove_child_ptr_in_heap)

  obtain owner_document disc_nodes h' where
    owner_document: "h  get_owner_document (cast child) r owner_document" and
    disc_nodes: "h   get_disconnected_nodes owner_document r disc_nodes" and
    h': "h   set_disconnected_nodes owner_document (child # disc_nodes) h h'" and
    h2: "h'  set_child_nodes ptr (remove1 child children) h h2"
    using assms children unfolding remove_child_def
    apply(auto split: if_splits elim!: bind_returns_heap_E)[1]
    by (metis (full_types) get_child_nodes_pure get_disconnected_nodes_pure
        get_owner_document_pure pure_returns_heap_eq returns_result_eq)

  have "object_ptr_kinds h = object_ptr_kinds h2"
    using remove_child_writes remove_child unfolding remove_child_locs_def
    apply(rule writes_small_big)
    using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved
    by(auto simp add: reflp_def transp_def)
  then have "|h  object_ptr_kinds_M|r = |h2  object_ptr_kinds_M|r"
    unfolding object_ptr_kinds_M_defs by simp

  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 type_wf
    by(auto simp add: reflp_def transp_def)
  have "type_wf h2"
    using writes_small_big[where P="λh h'. type_wf h  type_wf h'",
        OF remove_child_writes remove_child]  unfolding remove_child_locs_def
    using set_disconnected_nodes_types_preserved set_child_nodes_types_preserved type_wf
    apply(auto simp add: reflp_def transp_def)[1]
    by blast
  then obtain children' where children': "h2  get_child_nodes ptr r children'"
    using h2 set_child_nodes_get_child_nodes known_ptr
    by (metis object_ptr_kinds h = object_ptr_kinds h2 children get_child_nodes_ok
        get_child_nodes_ptr_in_heap is_OK_returns_result_E is_OK_returns_result_I)

  have "child  set children'"
    by (metis (mono_tags, lifting) type_wf h' children children' distinct_remove1_removeAll h2
        known_ptr local.heap_is_wellformed_children_distinct
        local.set_child_nodes_get_child_nodes member_remove remove_code(1) select_result_I2
        wellformed)


  moreover have "other_ptr other_children. other_ptr  ptr
                 h'  get_child_nodes other_ptr r other_children  child  set other_children"
  proof -
    fix other_ptr other_children
    assume a1: "other_ptr  ptr"  and a3: "h'  get_child_nodes other_ptr r other_children"
    have "h  get_child_nodes other_ptr r other_children"
      using get_child_nodes_reads set_disconnected_nodes_writes h' a3
      apply(rule reads_writes_separate_backwards)
      using set_disconnected_nodes_get_child_nodes by fast
    show "child  set other_children"
      using h  get_child_nodes other_ptr r other_children a1 h1 by blast
  qed
  then have "other_ptr other_children. other_ptr  ptr
              h2  get_child_nodes other_ptr r other_children  child  set other_children"
  proof -
    fix other_ptr other_children
    assume a1: "other_ptr  ptr" and a3: "h2  get_child_nodes other_ptr r other_children"
    have "h'  get_child_nodes other_ptr r other_children"
      using get_child_nodes_reads set_child_nodes_writes h2 a3
      apply(rule reads_writes_separate_backwards)
      using set_disconnected_nodes_get_child_nodes a1 set_child_nodes_get_child_nodes_different_pointers
      by metis
    then show "child  set other_children"
      using other_ptr other_children. other_ptr  ptr; h'  get_child_nodes other_ptr r other_children
              child  set other_children a1 by blast
  qed
  ultimately have ha: "other_ptr other_children. h2  get_child_nodes other_ptr r other_children
                       child  set other_children"
    by (metis (full_types) children' returns_result_eq)
  moreover obtain ptrs where ptrs: "h2  object_ptr_kinds_M r ptrs"
    by (simp add: object_ptr_kinds_M_defs)
  moreover have "ptr. ptr  set ptrs  h2  ok (get_child_nodes ptr)"
    using type_wf h2 ptrs get_child_nodes_ok known_ptr
    using object_ptr_kinds h = object_ptr_kinds h2 known_ptrs local.known_ptrs_known_ptr by auto
  ultimately show "h2  get_parent child r None"
    apply(auto simp add: get_parent_def intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I)[1]
  proof -
    have "castnode_ptr2object_ptr child |∈| object_ptr_kinds h"
      using get_owner_document_ptr_in_heap owner_document by blast
    then show "h2  check_in_heap (castnode_ptr2object_ptr child) r ()"
      by (simp add: object_ptr_kinds h = object_ptr_kinds h2 check_in_heap_def)
  next
    show "(other_ptr other_children. h2  get_child_nodes other_ptr r other_children
          child  set other_children) 
    ptrs = sorted_list_of_set (fset (object_ptr_kinds h2)) 
    (ptr. ptr |∈| object_ptr_kinds h2  h2  ok get_child_nodes ptr) 
    h2  filter_M (λptr. Heap_Error_Monad.bind (get_child_nodes ptr)
       (λchildren. return (child  set children))) (sorted_list_of_set (fset (object_ptr_kinds h2))) r []"
      by(auto intro!: filter_M_empty_I bind_pure_I)
  qed
qed
end

locale l_remove_child_wf2Core_DOM =
  l_remove_child_wfCore_DOM +
  l_heap_is_wellformedCore_DOM
begin

lemma remove_child_parent_child_rel_subset:
  assumes "heap_is_wellformed h"
    and "h  remove_child ptr child h h'"
    and "known_ptrs h"
    and type_wf: "type_wf h"
  shows "parent_child_rel h'  parent_child_rel h"
proof (standard, safe)

  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(2)
    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(2)])
    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(2)])
    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(2)])
    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_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 type_wf
    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(2) 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

  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(2)] 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 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


lemma remove_child_heap_is_wellformed_preserved:
  assumes "heap_is_wellformed h"
    and "h  remove_child ptr child h h'"
    and "known_ptrs h"
    and type_wf: "type_wf h"
  shows "type_wf h'" and "known_ptrs h'" and "heap_is_wellformed 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(2)
    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(2)])
    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(2)])
    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(2)])
    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)

  show "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 type_wf
    by(auto simp add: reflp_def transp_def)
  then show "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(2) 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(2) 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(2)] 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, opaque_lifting) type_wf h' assms(2) assms(3) local.get_child_nodes_ok
        local.known_ptrs_known_ptr local.remove_child_children_subset object_ptr_kinds_eq3
        returns_result_select_result subset_code(1) type_wf)
    apply (metis (no_types, opaque_lifting) assms(2) disconnected_nodes_eq2 disconnected_nodes_h
        disconnected_nodes_h' document_ptr_kinds_eq3 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 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" type_wf assms(3) 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" type_wf assms(3) 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" type_wf assms(3) 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: type_wf 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" type_wf assms(3) 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
      have 7: "|h'  get_disconnected_nodes xb|r = disc_nodes'"
        using disc_nodes' by simp
      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
        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, opaque_lifting) assms(3) type_wf)
        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 show "heap_is_wellformed h'"
    using heap_is_wellformed_def by blast
qed

lemma remove_heap_is_wellformed_preserved:
  assumes "heap_is_wellformed h"
    and "h  remove child h h'"
    and "known_ptrs h"
    and type_wf: "type_wf h"
  shows "type_wf h'" and "known_ptrs h'" and "heap_is_wellformed h'"
  using assms
  by(auto simp add: remove_def intro: remove_child_heap_is_wellformed_preserved
      elim!: bind_returns_heap_E2 split: option.splits)

lemma remove_child_removes_child:
  assumes wellformed: "heap_is_wellformed h"
    and remove_child: "h  remove_child ptr' child h h'"
    and children: "h'  get_child_nodes ptr r children"
    and known_ptrs: "known_ptrs h"
    and type_wf: "type_wf h"
  shows "child  set children"
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(2)
    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 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 remove_child])
    unfolding remove_child_locs_def
    using set_child_nodes_pointers_preserved set_disconnected_nodes_pointers_preserved
    by (auto simp add: reflp_def transp_def)
  moreover have "type_wf h'"
    using writes_small_big[where P="λh h'. type_wf h  type_wf h'", OF remove_child_writes assms(2)]
    using set_child_nodes_types_preserved set_disconnected_nodes_types_preserved type_wf
    unfolding remove_child_locs_def
    apply(auto simp add: reflp_def transp_def)[1]
    by blast
  ultimately show ?thesis
    using remove_child_removes_parent remove_child_heap_is_wellformed_preserved child_parent_dual
    by (meson children known_ptrs local.known_ptrs_preserved option.distinct(1) remove_child
        returns_result_eq type_wf wellformed)
qed

lemma remove_child_removes_first_child:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_child_nodes ptr r node_ptr # children"
  assumes "h  remove_child ptr node_ptr h h'"
  shows "h'  get_child_nodes ptr r children"
proof -
  obtain h2 disc_nodes owner_document where
    "h  get_owner_document (cast node_ptr) r owner_document" and
    "h  get_disconnected_nodes owner_document r disc_nodes" and
    h2: "h  set_disconnected_nodes owner_document (node_ptr # disc_nodes) h h2" and
    "h2  set_child_nodes ptr children h h'"
    using assms(5)
    apply(auto simp add: remove_child_def
        dest!: bind_returns_heap_E3[rotated, OF assms(4) get_child_nodes_pure, rotated])[1]
    by(auto elim!: bind_returns_heap_E
        bind_returns_heap_E2[rotated,OF get_owner_document_pure, rotated]
        bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated])
  have "known_ptr ptr"
    by (meson assms(3) assms(4) is_OK_returns_result_I get_child_nodes_ptr_in_heap known_ptrs_known_ptr)
  moreover have "h2  get_child_nodes ptr r node_ptr # children"
    apply(rule reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes h2 assms(4)])
    using set_disconnected_nodes_get_child_nodes set_child_nodes_get_child_nodes_different_pointers
    by fast
  moreover 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 type_wf h set_disconnected_nodes_types_preserved
    by(auto simp add:  reflp_def transp_def)
  ultimately show ?thesis
    using set_child_nodes_get_child_nodesh2  set_child_nodes ptr children h h'
    by fast
qed

lemma remove_removes_child:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_child_nodes ptr r node_ptr # children"
  assumes "h  remove node_ptr h h'"
  shows "h'  get_child_nodes ptr r children"
proof -
  have "h  get_parent node_ptr r Some ptr"
    using child_parent_dual assms by fastforce
  show ?thesis
    using assms remove_child_removes_first_child
    by(auto simp add: remove_def
        dest!: bind_returns_heap_E3[rotated, OF h  get_parent node_ptr r Some ptr, rotated]
        bind_returns_heap_E3[rotated, OF assms(4) get_child_nodes_pure, rotated])
qed

lemma remove_for_all_empty_children:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_child_nodes ptr r children"
  assumes "h  forall_M remove children h h'"
  shows "h'  get_child_nodes ptr r []"
  using assms
proof(induct children arbitrary: h h')
  case Nil
  then show ?case
    by simp
next
  case (Cons a children)
  have "h  get_parent a r Some ptr"
    using child_parent_dual Cons by fastforce
  with Cons show ?case
  proof(auto elim!: bind_returns_heap_E)[1]
    fix h2
    assume 0: "(h h'. heap_is_wellformed h  type_wf h  known_ptrs h
                    h  get_child_nodes ptr r children
                        h  forall_M remove children h h'  h'  get_child_nodes ptr r [])"
      and 1: "heap_is_wellformed h"
      and 2: "type_wf h"
      and 3: "known_ptrs h"
      and 4: "h  get_child_nodes ptr r a # children"
      and 5: "h  get_parent a r Some ptr"
      and 7: "h  remove a h h2"
      and 8: "h2  forall_M remove children h h'"
    then have "h2  get_child_nodes ptr r children"
      using remove_removes_child by blast

    moreover have "heap_is_wellformed h2"
      using 7 1 2 3 remove_child_heap_is_wellformed_preserved(3)
      by(auto simp add: remove_def
          elim!: bind_returns_heap_E
          bind_returns_heap_E2[rotated, OF get_parent_pure, rotated]
          split: option.splits)
    moreover have "type_wf h2"
      using writes_small_big[where P="λh h'. type_wf h  type_wf h'", OF remove_writes 7]
      using type_wf h remove_child_types_preserved
      by(auto simp add: a_remove_child_locs_def reflp_def transp_def)
    moreover have "object_ptr_kinds h = object_ptr_kinds h2"
      using 7
      apply(rule writes_small_big[where P="λh h'. object_ptr_kinds h = object_ptr_kinds h'",
            OF remove_writes])
      using remove_child_pointers_preserved
      by (auto simp add: reflp_def transp_def)
    then have "known_ptrs h2"
      using 3  known_ptrs_preserved by blast

    ultimately show "h'  get_child_nodes ptr r []"
      using 0 8 by fast
  qed
qed
end

locale l_remove_child_wf2 = l_type_wf + l_known_ptrs + l_remove_child_defs + l_heap_is_wellformed_defs
  + l_get_child_nodes_defs + l_remove_defs +
  assumes remove_child_preserves_type_wf:
    "heap_is_wellformed h  type_wf h  known_ptrs h  h  remove_child ptr child h h'
                           type_wf h'"
  assumes remove_child_preserves_known_ptrs:
    "heap_is_wellformed h  type_wf h  known_ptrs h  h  remove_child ptr child h h'
                           known_ptrs h'"
  assumes remove_child_heap_is_wellformed_preserved:
    "type_wf h  known_ptrs h  heap_is_wellformed h  h  remove_child ptr child h h'
                heap_is_wellformed h'"
  assumes remove_preserves_type_wf:
    "heap_is_wellformed h  type_wf h  known_ptrs h  h  remove child h h'
                           type_wf h'"
  assumes remove_preserves_known_ptrs:
    "heap_is_wellformed h  type_wf h  known_ptrs h  h  remove child h h'
                           known_ptrs h'"
  assumes remove_heap_is_wellformed_preserved:
    "type_wf h  known_ptrs h  heap_is_wellformed h  h  remove child h h'
                heap_is_wellformed h'"
  assumes remove_child_removes_child:
    "heap_is_wellformed h  h  remove_child ptr' child h h'  h'  get_child_nodes ptr r children
                           known_ptrs h  type_wf h
                           child  set children"
  assumes remove_child_removes_first_child:
    "heap_is_wellformed h  type_wf h  known_ptrs h
                           h  get_child_nodes ptr r node_ptr # children
                           h  remove_child ptr node_ptr h h'
                           h'  get_child_nodes ptr r children"
  assumes remove_removes_child:
    "heap_is_wellformed h  type_wf h  known_ptrs h
                           h  get_child_nodes ptr r node_ptr # children
                           h  remove node_ptr h h'  h'  get_child_nodes ptr r children"
  assumes remove_for_all_empty_children:
    "heap_is_wellformed h  type_wf h  known_ptrs h  h  get_child_nodes ptr r children
                           h  forall_M remove children h h'  h'  get_child_nodes ptr r []"

interpretation i_remove_child_wf2?: 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
  by unfold_locales

declare l_remove_child_wf2Core_DOM_axioms[instances]

lemma remove_child_wf2_is_l_remove_child_wf2 [instances]:
  "l_remove_child_wf2 type_wf known_ptr known_ptrs remove_child heap_is_wellformed get_child_nodes remove"
  apply(auto simp add: l_remove_child_wf2_def l_remove_child_wf2_axioms_def instances)[1]
  using remove_child_heap_is_wellformed_preserved apply(fast, fast, fast)
  using remove_heap_is_wellformed_preserved apply(fast, fast, fast)
  using remove_child_removes_child apply fast
  using remove_child_removes_first_child apply fast
  using remove_removes_child apply fast
  using remove_for_all_empty_children apply fast
  done



subsection ‹adopt\_node›

locale l_adopt_node_wfCore_DOM =
  l_adopt_nodeCore_DOM +
  l_get_parent_wf +
  l_get_owner_document_wf +
  l_remove_child_wf2 +
  l_heap_is_wellformed
begin
lemma adopt_node_removes_first_child:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  adopt_node owner_document node h h'"
  assumes "h  get_child_nodes ptr' r node # children"
  shows "h'  get_child_nodes ptr' r children"
proof -
  obtain old_document parent_opt h2 where
    old_document: "h  get_owner_document (cast node) r old_document" and
    parent_opt: "h  get_parent node r parent_opt" and
    h2: "h  (case parent_opt of Some parent  do { remove_child parent node }
                               | None  do { return ()}) h h2" and
    h': "h2  (if owner_document  old_document then do {
        old_disc_nodes  get_disconnected_nodes old_document;
        set_disconnected_nodes old_document (remove1 node old_disc_nodes);
        disc_nodes  get_disconnected_nodes owner_document;
        set_disconnected_nodes owner_document (node # disc_nodes)
      } else do { return () }) h h'"
    using assms(4)
    by(auto simp add: adopt_node_def elim!: bind_returns_heap_E
        dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure]
        pure_returns_heap_eq[rotated, OF get_parent_pure])
  have "h2  get_child_nodes ptr' r children"
    using h2 remove_child_removes_first_child assms(1) assms(2) assms(3) assms(5)
    by (metis list.set_intros(1) local.child_parent_dual option.simps(5) parent_opt returns_result_eq)
  then
  show ?thesis
    using h'
    by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]
        dest!: reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes]
        split: if_splits)
qed


lemma adopt_node_document_in_heap:
  assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
  assumes "h  ok (adopt_node owner_document node)"
  shows "owner_document |∈| document_ptr_kinds h"
proof -
  obtain old_document parent_opt h2 h' where
    old_document: "h  get_owner_document (cast node) r old_document" and
    parent_opt: "h  get_parent node r parent_opt" and
    h2: "h  (case parent_opt of Some parent  do { remove_child parent node } | None  do { return ()}) h h2"
    and
    h': "h2  (if owner_document  old_document then do {
        old_disc_nodes  get_disconnected_nodes old_document;
        set_disconnected_nodes old_document (remove1 node old_disc_nodes);
        disc_nodes  get_disconnected_nodes owner_document;
        set_disconnected_nodes owner_document (node # disc_nodes)
      } else do { return () }) h h'"
    using assms(4)
    by(auto simp add: adopt_node_def
        elim!: bind_returns_heap_E
        dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure]
        pure_returns_heap_eq[rotated, OF get_parent_pure])
  show ?thesis
  proof (cases "owner_document = old_document")
    case True
    then show ?thesis
      using old_document get_owner_document_owner_document_in_heap assms(1) assms(2) assms(3)
      by auto
  next
    case False

    then obtain h3 old_disc_nodes disc_nodes where
      old_disc_nodes: "h2  get_disconnected_nodes old_document r old_disc_nodes" and
      h3: "h2  set_disconnected_nodes old_document (remove1 node old_disc_nodes) h h3" and
      old_disc_nodes: "h3  get_disconnected_nodes owner_document r disc_nodes" and
      h': "h3  set_disconnected_nodes owner_document (node # disc_nodes) h h'"
      using h'
      by(auto elim!: bind_returns_heap_E
          bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] )
    then have "owner_document |∈| document_ptr_kinds h3"
      by (meson is_OK_returns_result_I local.get_disconnected_nodes_ptr_in_heap)

    moreover have "object_ptr_kinds h = object_ptr_kinds h2"
      using h2 apply(simp split: option.splits)
      apply(rule writes_small_big[where P="λh h'. object_ptr_kinds h = object_ptr_kinds h'",
            OF remove_child_writes])
      using remove_child_pointers_preserved
      by (auto simp add: reflp_def transp_def)
    moreover have "object_ptr_kinds h2 = 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 h3])
      using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved
      by (auto simp add: reflp_def transp_def)

    ultimately show ?thesis
      by(auto simp add: document_ptr_kinds_def)
  qed
qed
end

locale l_adopt_node_wf2Core_DOM =
  l_adopt_node_wfCore_DOM +
  l_adopt_nodeCore_DOM +
  l_get_parent_wfCore_DOM +
  l_get_root_node +
  l_get_owner_document_wf +
  l_remove_child_wf2 +
  l_heap_is_wellformedCore_DOM
begin

lemma adopt_node_removes_child_step:
  assumes wellformed: "heap_is_wellformed h"
    and adopt_node: "h  adopt_node owner_document node_ptr h h2"
    and children: "h2  get_child_nodes ptr r children"
    and known_ptrs: "known_ptrs h"
    and type_wf: "type_wf h"
  shows "node_ptr  set children"
proof -
  obtain old_document parent_opt h' where
    old_document: "h  get_owner_document (cast node_ptr) r old_document" and
    parent_opt: "h  get_parent node_ptr r parent_opt" and
    h': "h  (case parent_opt of Some parent  remove_child parent node_ptr | None  return () ) h h'"
    using adopt_node get_parent_pure
    by(auto simp add: adopt_node_def
        elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated]
        bind_returns_heap_E2[rotated, OF get_parent_pure, rotated]
        bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]
        split: if_splits)

  then have "h'  get_child_nodes ptr r children"
    using adopt_node
    apply(auto simp add: adopt_node_def
        dest!: bind_returns_heap_E3[rotated, OF old_document, rotated]
        bind_returns_heap_E3[rotated, OF parent_opt, rotated]
        elim!: bind_returns_heap_E4[rotated, OF h', rotated])[1]
    apply(auto split: if_splits
        elim!: bind_returns_heap_E
        bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated])[1]
     apply (simp add: set_disconnected_nodes_get_child_nodes children
        reads_writes_preserved[OF get_child_nodes_reads set_disconnected_nodes_writes])
    using children by blast
  show ?thesis
  proof(insert parent_opt h', induct parent_opt)
    case None
    then show ?case
      using child_parent_dual wellformed known_ptrs type_wf
        h'  get_child_nodes ptr r children returns_result_eq
      by fastforce
  next
    case (Some option)
    then show ?case
      using remove_child_removes_child h'  get_child_nodes ptr r children known_ptrs type_wf
        wellformed
      by auto
  qed
qed

lemma adopt_node_removes_child:
  assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h"
  assumes "h  adopt_node owner_document node_ptr h h'"
  shows "ptr' children'.
  h'  get_child_nodes ptr' r children'  node_ptr  set children'"
  using adopt_node_removes_child_step assms by blast

lemma adopt_node_preserves_wellformedness:
  assumes "heap_is_wellformed h"
    and "h  adopt_node document_ptr child h h'"
    and known_ptrs: "known_ptrs h"
    and type_wf: "type_wf h"
  shows "heap_is_wellformed h'" and "known_ptrs h'" and "type_wf h'"
proof -
  obtain old_document parent_opt h2 where
    old_document: "h  get_owner_document (cast child) r old_document"
    and
    parent_opt: "h  get_parent child r parent_opt"
    and
    h2: "h  (case parent_opt of Some parent  remove_child parent child | None  return ()) h h2"
    and
    h': "h2  (if document_ptr  old_document then do {
        old_disc_nodes  get_disconnected_nodes old_document;
        set_disconnected_nodes old_document (remove1 child old_disc_nodes);
        disc_nodes  get_disconnected_nodes document_ptr;
        set_disconnected_nodes document_ptr (child # disc_nodes)
      } else do {
        return ()
      }) h h'"
    using assms(2)
    by(auto simp add: adopt_node_def elim!: bind_returns_heap_E
        dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure]
        pure_returns_heap_eq[rotated, OF get_parent_pure])

  have object_ptr_kinds_h_eq3: "object_ptr_kinds h = object_ptr_kinds h2"
    using h2 apply(simp split: option.splits)
    apply(rule writes_small_big[where P="λh h'. object_ptr_kinds h = object_ptr_kinds h'",
          OF remove_child_writes])
    using remove_child_pointers_preserved
    by (auto simp add: reflp_def transp_def)
  then have object_ptr_kinds_M_eq_h:
    "ptrs. h  object_ptr_kinds_M r ptrs = h2  object_ptr_kinds_M r ptrs"
    unfolding object_ptr_kinds_M_defs by simp
  then have object_ptr_kinds_eq_h: "|h  object_ptr_kinds_M|r = |h2  object_ptr_kinds_M|r"
    by simp
  then have node_ptr_kinds_eq_h: "|h  node_ptr_kinds_M|r = |h2  node_ptr_kinds_M|r"
    using node_ptr_kinds_M_eq by blast

  have wellformed_h2: "heap_is_wellformed h2"
    using h2 remove_child_heap_is_wellformed_preserved known_ptrs type_wf
    by (metis (no_types, lifting) assms(1) option.case_eq_if pure_returns_heap_eq return_pure)
  have "type_wf h2"
    using h2 remove_child_preserves_type_wf known_ptrs type_wf
    by (metis (no_types, lifting) assms(1) option.case_eq_if pure_returns_heap_eq return_pure)
  have "known_ptrs h2"
    using h2 remove_child_preserves_known_ptrs known_ptrs type_wf
    by (metis (no_types, lifting) assms(1) option.case_eq_if pure_returns_heap_eq return_pure)
  have "heap_is_wellformed h'  known_ptrs h'  type_wf h'"
  proof(cases "document_ptr = old_document")
    case True
    then show ?thesis
      using h' wellformed_h2 type_wf h2 known_ptrs h2 by auto
  next
    case False
    then obtain h3 old_disc_nodes disc_nodes_document_ptr_h3 where
      docs_neq: "document_ptr  old_document" and
      old_disc_nodes: "h2  get_disconnected_nodes old_document r old_disc_nodes" and
      h3: "h2  set_disconnected_nodes old_document (remove1 child old_disc_nodes) h h3" and
      disc_nodes_document_ptr_h3:
      "h3  get_disconnected_nodes document_ptr r disc_nodes_document_ptr_h3" and
      h': "h3  set_disconnected_nodes document_ptr (child # disc_nodes_document_ptr_h3) h h'"
      using h'
      by(auto elim!: bind_returns_heap_E
          bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] )

    have object_ptr_kinds_h2_eq3: "object_ptr_kinds h2 = 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 h3])
      using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved
      by (auto simp add: reflp_def transp_def)
    then have object_ptr_kinds_M_eq_h2:
      "ptrs. h2  object_ptr_kinds_M r ptrs = h3  object_ptr_kinds_M r ptrs"
      by(simp add: object_ptr_kinds_M_defs)
    then have object_ptr_kinds_eq_h2: "|h2  object_ptr_kinds_M|r = |h3  object_ptr_kinds_M|r"
      by(simp)
    then have node_ptr_kinds_eq_h2: "|h2  node_ptr_kinds_M|r = |h3  node_ptr_kinds_M|r"
      using node_ptr_kinds_M_eq by blast
    then have node_ptr_kinds_eq3_h2: "node_ptr_kinds h2 = node_ptr_kinds h3"
      by auto
    have document_ptr_kinds_eq2_h2: "|h2  document_ptr_kinds_M|r = |h3  document_ptr_kinds_M|r"
      using object_ptr_kinds_eq_h2 document_ptr_kinds_M_eq by auto
    then have document_ptr_kinds_eq3_h2: "document_ptr_kinds h2 = document_ptr_kinds h3"
      using object_ptr_kinds_eq_h2 document_ptr_kinds_M_eq by auto
    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_disconnected_nodes_writes h3
      apply(rule reads_writes_preserved)
      by (simp add: set_disconnected_nodes_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 object_ptr_kinds_h3_eq3: "object_ptr_kinds h3 = object_ptr_kinds h'"
      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 set_child_nodes_pointers_preserved
      by (auto simp add: reflp_def transp_def)
    then have object_ptr_kinds_M_eq_h3:
      "ptrs. h3  object_ptr_kinds_M r ptrs = h'  object_ptr_kinds_M r ptrs"
      by(simp add: object_ptr_kinds_M_defs)
    then have object_ptr_kinds_eq_h3: "|h3  object_ptr_kinds_M|r = |h'  object_ptr_kinds_M|r"
      by(simp)
    then have node_ptr_kinds_eq_h3: "|h3  node_ptr_kinds_M|r = |h'  node_ptr_kinds_M|r"
      using node_ptr_kinds_M_eq by blast
    then have node_ptr_kinds_eq3_h3: "node_ptr_kinds h3 = node_ptr_kinds h'"
      by auto
    have document_ptr_kinds_eq2_h3: "|h3  document_ptr_kinds_M|r = |h'  document_ptr_kinds_M|r"
      using object_ptr_kinds_eq_h3 document_ptr_kinds_M_eq by auto
    then have document_ptr_kinds_eq3_h3: "document_ptr_kinds h3 = document_ptr_kinds h'"
      using object_ptr_kinds_eq_h3 document_ptr_kinds_M_eq by auto
    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 (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_h2:
      "doc_ptr disc_nodes. old_document  doc_ptr
       h2  get_disconnected_nodes doc_ptr r disc_nodes = h3  get_disconnected_nodes doc_ptr r disc_nodes"
      using get_disconnected_nodes_reads set_disconnected_nodes_writes h3
      apply(rule reads_writes_preserved)
      by (simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers)
    then have disconnected_nodes_eq2_h2:
      "doc_ptr. old_document  doc_ptr
        |h2  get_disconnected_nodes doc_ptr|r = |h3  get_disconnected_nodes doc_ptr|r"
      using select_result_eq by force
    obtain disc_nodes_old_document_h2 where disc_nodes_old_document_h2:
      "h2  get_disconnected_nodes old_document r disc_nodes_old_document_h2"
      using old_disc_nodes by blast
    then have disc_nodes_old_document_h3:
      "h3  get_disconnected_nodes old_document r remove1 child disc_nodes_old_document_h2"
      using h3 old_disc_nodes returns_result_eq set_disconnected_nodes_get_disconnected_nodes
      by fastforce
    have "distinct disc_nodes_old_document_h2"
      using disc_nodes_old_document_h2 local.heap_is_wellformed_disconnected_nodes_distinct wellformed_h2
      by blast


    have "type_wf h2"
    proof (insert h2, induct  parent_opt)
      case None
      then show ?case
        using type_wf by simp
    next
      case (Some option)
      then show ?case
        using writes_small_big[where P="λh h'. type_wf h  type_wf h'", OF remove_child_writes]
          type_wf remove_child_types_preserved
        by (simp add: reflp_def transp_def)
    qed
    then have "type_wf h3"
      using writes_small_big[where P="λh h'. type_wf h  type_wf h'", OF set_disconnected_nodes_writes h3]
      using  set_disconnected_nodes_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 "known_ptrs h3"
      using known_ptrs local.known_ptrs_preserved object_ptr_kinds_h2_eq3 object_ptr_kinds_h_eq3 by blast
    then have "known_ptrs h'"
      using local.known_ptrs_preserved object_ptr_kinds_h3_eq3 by blast

    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 (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_document_ptr_h3"
      using disconnected_nodes_eq_h2 docs_neq disc_nodes_document_ptr_h3 by auto
    have disc_nodes_document_ptr_h': "
      h'  get_disconnected_nodes document_ptr r child # disc_nodes_document_ptr_h3"
      using h' disc_nodes_document_ptr_h3
      using set_disconnected_nodes_get_disconnected_nodes by blast

    have document_ptr_in_heap: "document_ptr |∈| document_ptr_kinds h2"
      using disc_nodes_document_ptr_h3 document_ptr_kinds_eq2_h2 get_disconnected_nodes_ok assms(1)
      unfolding heap_is_wellformed_def
      using disc_nodes_document_ptr_h2 get_disconnected_nodes_ptr_in_heap by blast
    have old_document_in_heap: "old_document |∈| document_ptr_kinds h2"
      using disc_nodes_old_document_h3 document_ptr_kinds_eq2_h2 get_disconnected_nodes_ok assms(1)
      unfolding heap_is_wellformed_def
      using get_disconnected_nodes_ptr_in_heap old_disc_nodes by blast

    have "child  set disc_nodes_old_document_h2"
    proof (insert parent_opt h2, induct parent_opt)
      case None
      then have "h = h2"
        by(auto)
      moreover have "a_owner_document_valid h"
        using assms(1) heap_is_wellformed_def by(simp add: heap_is_wellformed_def)
      ultimately show ?case
        using old_document disc_nodes_old_document_h2 None(1) child_parent_dual[OF assms(1)]
          in_disconnected_nodes_no_parent assms(1) known_ptrs type_wf by blast
    next
      case (Some option)
      then show ?case
        apply(simp split: option.splits)
        using assms(1) disc_nodes_old_document_h2 old_document remove_child_in_disconnected_nodes known_ptrs
        by blast
    qed
    have "child  set (remove1 child disc_nodes_old_document_h2)"
      using disc_nodes_old_document_h3 h3 known_ptrs wellformed_h2 distinct disc_nodes_old_document_h2
      by auto
    have "child  set disc_nodes_document_ptr_h3"
    proof -
      have "a_distinct_lists h2"
        using heap_is_wellformed_def wellformed_h2 by blast
      then have 0: "distinct (concat (map (λdocument_ptr. |h2  get_disconnected_nodes document_ptr|r)
                                                          |h2  document_ptr_kinds_M|r))"
        by(simp add: a_distinct_lists_def)
      show ?thesis
        using distinct_concat_map_E(1)[OF 0] child  set disc_nodes_old_document_h2
          disc_nodes_old_document_h2 disc_nodes_document_ptr_h2
        by (meson type_wf h2 docs_neq known_ptrs local.get_owner_document_disconnected_nodes
            local.known_ptrs_preserved object_ptr_kinds_h_eq3 returns_result_eq wellformed_h2)
    qed

    have child_in_heap: "child |∈| node_ptr_kinds h"
      using get_owner_document_ptr_in_heap[OF is_OK_returns_result_I[OF old_document]]
        node_ptr_kinds_commutes by blast
    have "a_acyclic_heap h2"
      using wellformed_h2 by (simp add: heap_is_wellformed_def)
    have "parent_child_rel h'  parent_child_rel h2"
    proof
      fix x
      assume "x  parent_child_rel h'"
      then show "x  parent_child_rel h2"
        using object_ptr_kinds_h2_eq3  object_ptr_kinds_h3_eq3 children_eq2_h2 children_eq2_h3
          mem_Collect_eq object_ptr_kinds_M_eq_h3 select_result_eq split_cong
        unfolding parent_child_rel_def
        by(simp)
    qed
    then have "a_acyclic_heap h'"
      using a_acyclic_heap h2 acyclic_heap_def acyclic_subset by blast

    moreover have "a_all_ptrs_in_heap h2"
      using wellformed_h2 by (simp add: heap_is_wellformed_def)
    then have "a_all_ptrs_in_heap h3"
      apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq3_h2 children_eq_h2)[1]
       apply (simp add: children_eq2_h2 object_ptr_kinds_h2_eq3 subset_code(1))
      by (metis (no_types, lifting) child  set disc_nodes_old_document_h2 type_wf h2
          disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2
          document_ptr_kinds_eq3_h2 in_set_remove1 local.get_disconnected_nodes_ok
          local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq3_h2 returns_result_select_result
          select_result_I2 wellformed_h2)
    then have "a_all_ptrs_in_heap h'"
      apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq3_h3 children_eq_h3)[1]
       apply (simp add: children_eq2_h3 object_ptr_kinds_h3_eq3 subset_code(1))
      by (metis (no_types, opaque_lifting) child  set disc_nodes_old_document_h2 disc_nodes_document_ptr_h'
          disc_nodes_document_ptr_h2 disc_nodes_old_document_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq3_h3
          local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq3_h2 node_ptr_kinds_eq3_h3
          select_result_I2 set_ConsD subset_code(1) wellformed_h2)

    moreover have "a_owner_document_valid h2"
      using wellformed_h2 by (simp add: heap_is_wellformed_def)
    then have "a_owner_document_valid h'"
      apply(simp add: a_owner_document_valid_def node_ptr_kinds_eq_h2 node_ptr_kinds_eq3_h3
          object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 document_ptr_kinds_eq2_h2
          document_ptr_kinds_eq2_h3 children_eq2_h2 children_eq2_h3 )
      by (smt (verit) disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2
          disc_nodes_old_document_h2 disc_nodes_old_document_h3
          disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_in_heap
          document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 in_set_remove1
          list.set_intros(1) node_ptr_kinds_eq3_h2 node_ptr_kinds_eq3_h3
          object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 select_result_I2
          set_subset_Cons subset_code(1))

    have a_distinct_lists_h2: "a_distinct_lists h2"
      using wellformed_h2 by (simp add: heap_is_wellformed_def)
    then have "a_distinct_lists h'"
      apply(auto simp add: a_distinct_lists_def object_ptr_kinds_eq_h3 object_ptr_kinds_eq_h2
          children_eq2_h2 children_eq2_h3)[1]
    proof -
      assume 1: "distinct (concat (map (λptr. |h'  get_child_nodes ptr|r)
                          (sorted_list_of_set (fset (object_ptr_kinds h')))))"
        and 2: "distinct (concat (map (λdocument_ptr. |h2  get_disconnected_nodes document_ptr|r)
                                             (sorted_list_of_set (fset (document_ptr_kinds h2)))))"
        and 3: "(xfset (object_ptr_kinds h'). set |h'  get_child_nodes x|r)
                    (xfset (document_ptr_kinds h2). set |h2  get_disconnected_nodes x|r) = {}"
      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)
        show "distinct (sorted_list_of_set (fset (document_ptr_kinds h')))"
          by(auto simp add: document_ptr_kinds_M_def )
      next
        fix x
        assume a1: "x  set (sorted_list_of_set (fset (document_ptr_kinds h')))"
        have 4: "distinct |h2  get_disconnected_nodes x|r"
          using a_distinct_lists_h2 "2" a1 concat_map_all_distinct document_ptr_kinds_eq2_h2
            document_ptr_kinds_eq2_h3
          by fastforce
        then show "distinct |h'  get_disconnected_nodes x|r"
        proof (cases "old_document  x")
          case True
          then show ?thesis
          proof (cases "document_ptr  x")
            case True
            then show ?thesis
              using disconnected_nodes_eq2_h2[OF old_document  x]
                disconnected_nodes_eq2_h3[OF document_ptr  x] 4
              by(auto)
          next
            case False
            then show ?thesis
              using  disc_nodes_document_ptr_h3 disc_nodes_document_ptr_h' 4
                child  set disc_nodes_document_ptr_h3
              by(auto simp add: disconnected_nodes_eq2_h2[OF old_document  x] )
          qed
        next
          case False
          then show ?thesis
            by (metis (no_types, opaque_lifting) distinct disc_nodes_old_document_h2
                disc_nodes_old_document_h3 disconnected_nodes_eq2_h3
                distinct_remove1 docs_neq select_result_I2)
        qed
      next
        fix x y
        assume a0: "x  set (sorted_list_of_set (fset (document_ptr_kinds h')))"
          and a1: "y  set (sorted_list_of_set (fset (document_ptr_kinds h')))"
          and a2: "x  y"

        moreover have 5: "set |h2  get_disconnected_nodes x|r  set |h2  get_disconnected_nodes y|r = {}"
          using 2 calculation
          by (auto simp add: document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 dest: distinct_concat_map_E(1))
        ultimately show "set |h'  get_disconnected_nodes x|r  set |h'  get_disconnected_nodes y|r = {}"
        proof(cases "old_document = x")
          case True
          have "old_document  y"
            using x  y old_document = x by simp
          have "document_ptr  x"
            using docs_neq old_document = x by auto
          show ?thesis
          proof(cases "document_ptr = y")
            case True
            then show ?thesis
              using 5 True select_result_I2[OF disc_nodes_document_ptr_h']
                select_result_I2[OF disc_nodes_document_ptr_h2]
                select_result_I2[OF disc_nodes_old_document_h2]
                select_result_I2[OF disc_nodes_old_document_h3] old_document = x
              by (metis (no_types, lifting) child  set (remove1 child disc_nodes_old_document_h2)
                  document_ptr  x disconnected_nodes_eq2_h3 disjoint_iff_not_equal
                  notin_set_remove1 set_ConsD)
          next
            case False
            then show ?thesis
              using 5 select_result_I2[OF disc_nodes_document_ptr_h']
                select_result_I2[OF disc_nodes_document_ptr_h2]
                select_result_I2[OF disc_nodes_old_document_h2]
                select_result_I2[OF disc_nodes_old_document_h3]
                disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 old_document = x
                docs_neq old_document  y
              by (metis (no_types, lifting) disjoint_iff_not_equal notin_set_remove1)
          qed
        next
          case False
          then show ?thesis
          proof(cases "old_document = y")
            case True
            then show ?thesis
            proof(cases "document_ptr = x")
              case True
              show ?thesis
                using 5 select_result_I2[OF disc_nodes_document_ptr_h']
                  select_result_I2[OF disc_nodes_document_ptr_h2]
                  select_result_I2[OF disc_nodes_old_document_h2]
                  select_result_I2[OF disc_nodes_old_document_h3]
                  old_document  x old_document = y document_ptr = x
                apply(simp)
                by (metis (no_types, lifting) child  set (remove1 child disc_nodes_old_document_h2)
                    disconnected_nodes_eq2_h3 disjoint_iff_not_equal notin_set_remove1)
            next
              case False
              then show ?thesis
                using 5 select_result_I2[OF disc_nodes_document_ptr_h']
                  select_result_I2[OF disc_nodes_document_ptr_h2]
                  select_result_I2[OF disc_nodes_old_document_h2]
                  select_result_I2[OF disc_nodes_old_document_h3]
                  old_document  x old_document = y document_ptr  x
                by (metis (no_types, lifting) disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3
                    disjoint_iff_not_equal docs_neq notin_set_remove1)
            qed
          next
            case False
            have "set |h2  get_disconnected_nodes y|r  set disc_nodes_old_document_h2 = {}"
              by (metis DocumentMonad.ptr_kinds_M_ok DocumentMonad.ptr_kinds_M_ptr_kinds False
                  type_wf h2 a1 disc_nodes_old_document_h2 document_ptr_kinds_M_def
                  document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3
                  l_ptr_kinds_M.ptr_kinds_ptr_kinds_M local.get_disconnected_nodes_ok
                  local.heap_is_wellformed_one_disc_parent returns_result_select_result
                  wellformed_h2)
            then show ?thesis
            proof(cases "document_ptr = x")
              case True
              then have "document_ptr  y"
                using x  y by auto
              have "set |h2  get_disconnected_nodes y|r  set disc_nodes_old_document_h2 = {}"
                using set |h2  get_disconnected_nodes y|r  set disc_nodes_old_document_h2 = {}
                by blast
              then show ?thesis
                using 5 select_result_I2[OF disc_nodes_document_ptr_h']
                  select_result_I2[OF disc_nodes_document_ptr_h2]
                  select_result_I2[OF disc_nodes_old_document_h2]
                  select_result_I2[OF disc_nodes_old_document_h3]
                  old_document  x old_document  y document_ptr = x document_ptr  y
                  child  set disc_nodes_old_document_h2 disconnected_nodes_eq2_h2
                  disconnected_nodes_eq2_h3
                  set |h2  get_disconnected_nodes y|r  set disc_nodes_old_document_h2 = {}
                by(auto)
            next
              case False
              then show ?thesis
              proof(cases "document_ptr = y")
                case True
                have f1: "set |h2  get_disconnected_nodes x|r  set disc_nodes_document_ptr_h3 = {}"
                  using 2 a1 document_ptr_in_heap document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3
                    document_ptr  x select_result_I2[OF disc_nodes_document_ptr_h3, symmetric]
                    disconnected_nodes_eq2_h2[OF docs_neq[symmetric], symmetric]
                  by (simp add: "5" True)
                moreover have f1:
                  "set |h2  get_disconnected_nodes x|r  set |h2  get_disconnected_nodes old_document|r = {}"
                  using 2 a1 old_document_in_heap document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3
                    old_document  x
                  by (metis (no_types, lifting) a0 distinct_concat_map_E(1) document_ptr_kinds_eq3_h2
                      document_ptr_kinds_eq3_h3 finite_fset set_sorted_list_of_set)
                ultimately show ?thesis
                  using 5 select_result_I2[OF disc_nodes_document_ptr_h']
                    select_result_I2[OF disc_nodes_old_document_h2] old_document  x
                    document_ptr  x document_ptr = y
                    child  set disc_nodes_old_document_h2 disconnected_nodes_eq2_h2
                    disconnected_nodes_eq2_h3
                  by auto
              next
                case False
                then show ?thesis
                  using 5
                    select_result_I2[OF disc_nodes_old_document_h2] old_document  x
                    document_ptr  x document_ptr  y
                    child  set disc_nodes_old_document_h2
                    disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3
                  by (metis set |h2  get_disconnected_nodes y|r  set disc_nodes_old_document_h2 = {}
                      empty_iff inf.idem)
              qed
            qed
          qed
        qed
      qed
    next
      fix x xa xb
      assume 0: "distinct (concat (map (λptr. |h'  get_child_nodes ptr|r)
                                  (sorted_list_of_set (fset (object_ptr_kinds h')))))"
        and 1: "distinct (concat (map (λdocument_ptr. |h2  get_disconnected_nodes document_ptr|r)
                                   (sorted_list_of_set (fset (document_ptr_kinds h2)))))"
        and 2: "(xfset (object_ptr_kinds h'). set |h'  get_child_nodes x|r)
                     (xfset (document_ptr_kinds h2). set |h2  get_disconnected_nodes x|r) = {}"
        and 3: "xa |∈| object_ptr_kinds h'"
        and 4: "x  set |h'  get_child_nodes xa|r"
        and 5: "xb |∈| document_ptr_kinds h'"
        and 6: "x  set |h'  get_disconnected_nodes xb|r"
      then show False
        using child  set disc_nodes_old_document_h2 disc_nodes_document_ptr_h'
          disc_nodes_document_ptr_h2 disc_nodes_old_document_h2 disc_nodes_old_document_h3
          disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq2_h2
          document_ptr_kinds_eq2_h3  old_document_in_heap
        apply(auto)[1]
        apply(cases "xb = old_document")
      proof -
        assume a1: "xb = old_document"
        assume a2: "h2  get_disconnected_nodes old_document r disc_nodes_old_document_h2"
        assume a3: "h3  get_disconnected_nodes old_document r remove1 child disc_nodes_old_document_h2"
        assume a4: "x  set |h'  get_child_nodes xa|r"
        assume "document_ptr_kinds h2 = document_ptr_kinds h'"
        assume a5: "(xfset (object_ptr_kinds h'). set |h'  get_child_nodes x|r)
                    (xfset (document_ptr_kinds h'). set |h2  get_disconnected_nodes x|r) = {}"
        have f6: "old_document |∈| document_ptr_kinds h'"
          using a1 xb |∈| document_ptr_kinds h' by blast
        have f7: "|h2  get_disconnected_nodes old_document|r = disc_nodes_old_document_h2"
          using a2 by simp
        have "x  set disc_nodes_old_document_h2"
          using f6 a3 a1 by (metis (no_types) type_wf h' x  set |h'  get_disconnected_nodes xb|r
              disconnected_nodes_eq_h3 docs_neq get_disconnected_nodes_ok returns_result_eq
              returns_result_select_result set_remove1_subset subsetCE)
        then have "set |h'  get_child_nodes xa|r   set |h2  get_disconnected_nodes xb|r = {}"
          using f7 f6 a5 a4 xa |∈| object_ptr_kinds h'
          by fastforce
        then show ?thesis
          using x  set disc_nodes_old_document_h2 a1 a4 f7 by blast
      next
        assume a1: "xb  old_document"
        assume a2: "h2  get_disconnected_nodes document_ptr r disc_nodes_document_ptr_h3"
        assume a3: "h2  get_disconnected_nodes old_document r disc_nodes_old_document_h2"
        assume a4: "xa |∈| object_ptr_kinds h'"
        assume a5: "h'  get_disconnected_nodes document_ptr r child # disc_nodes_document_ptr_h3"
        assume a6: "old_document |∈| document_ptr_kinds h'"
        assume a7: "x  set |h'  get_disconnected_nodes xb|r"
        assume a8: "x  set |h'  get_child_nodes xa|r"
        assume a9: "document_ptr_kinds h2 = document_ptr_kinds h'"
        assume a10: "doc_ptr. old_document  doc_ptr
                |h2  get_disconnected_nodes doc_ptr|r = |h3  get_disconnected_nodes doc_ptr|r"
        assume a11: "doc_ptr. document_ptr  doc_ptr
                |h3  get_disconnected_nodes doc_ptr|r = |h'  get_disconnected_nodes doc_ptr|r"
        assume a12: "(xfset (object_ptr_kinds h'). set |h'  get_child_nodes x|r)
                     (xfset (document_ptr_kinds h'). set |h2  get_disconnected_nodes x|r) = {}"
        have f13: "d. d  set |h'  document_ptr_kinds_M|r  h2  ok get_disconnected_nodes d"
          using a9 type_wf h2 get_disconnected_nodes_ok
          by simp
        then have f14: "|h2  get_disconnected_nodes old_document|r = disc_nodes_old_document_h2"
          using a6 a3 by simp
        have "x  set |h2  get_disconnected_nodes xb|r"
          using a12 a8 a4 xb |∈| document_ptr_kinds h'
          by (meson UN_I disjoint_iff_not_equal)
        then have "x = child"
          using f13 a11 a10 a7 a5 a2 a1
          by (metis (no_types, lifting) select_result_I2 set_ConsD)
        then have "child  set disc_nodes_old_document_h2"
          using f14 a12 a8 a6 a4
          by (metis type_wf h' adopt_node_removes_child assms(1) assms(2) type_wf
              get_child_nodes_ok known_ptrs local.known_ptrs_known_ptr object_ptr_kinds_h2_eq3
              object_ptr_kinds_h3_eq3 object_ptr_kinds_h_eq3 returns_result_select_result)
        then show ?thesis
          using child  set disc_nodes_old_document_h2 by fastforce
      qed
    qed
    ultimately show ?thesis
      using type_wf h' known_ptrs h' a_owner_document_valid h' heap_is_wellformed_def by blast
  qed
  then show "heap_is_wellformed h'" and "known_ptrs h'" and "type_wf h'"
    by auto
qed

lemma adopt_node_node_in_disconnected_nodes:
  assumes wellformed: "heap_is_wellformed h"
    and adopt_node: "h  adopt_node owner_document node_ptr h h'"
    and "h'  get_disconnected_nodes owner_document r disc_nodes"
    and known_ptrs: "known_ptrs h"
    and type_wf: "type_wf h"
  shows "node_ptr  set disc_nodes"
proof -
  obtain old_document parent_opt h2 where
    old_document: "h  get_owner_document (cast node_ptr) r old_document" and
    parent_opt: "h  get_parent node_ptr r parent_opt" and
    h2: "h  (case parent_opt of Some parent  remove_child parent node_ptr | None  return ()) h h2"
    and
    h': "h2  (if owner_document  old_document then do {
        old_disc_nodes  get_disconnected_nodes old_document;
        set_disconnected_nodes old_document (remove1 node_ptr old_disc_nodes);
        disc_nodes  get_disconnected_nodes owner_document;
        set_disconnected_nodes owner_document (node_ptr # disc_nodes)
      } else do {
        return ()
      }) h h'"
    using assms(2)
    by(auto simp add: adopt_node_def elim!: bind_returns_heap_E
        dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure]
        pure_returns_heap_eq[rotated, OF get_parent_pure])

  show ?thesis
  proof (cases "owner_document = old_document")
    case True
    then show ?thesis
    proof (insert parent_opt h2, induct parent_opt)
      case None
      then have "h = h'"
        using h2 h' by(auto)
      then show ?case
        using in_disconnected_nodes_no_parent assms None old_document by blast
    next
      case (Some parent)
      then show ?case
        using remove_child_in_disconnected_nodes known_ptrs True h' assms(3) old_document by auto
    qed
  next
    case False
    then show ?thesis
      using assms(3)  h' list.set_intros(1) select_result_I2 set_disconnected_nodes_get_disconnected_nodes
      apply(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated])[1]
    proof -
      fix x  and h'a and xb
      assume a1: "h'  get_disconnected_nodes owner_document r disc_nodes"
      assume a2: "h document_ptr disc_nodes h'. h  set_disconnected_nodes document_ptr disc_nodes h h'
                  h'  get_disconnected_nodes document_ptr r disc_nodes"
      assume "h'a  set_disconnected_nodes owner_document (node_ptr # xb) h h'"
      then have "node_ptr # xb = disc_nodes"
        using a2 a1 by (meson returns_result_eq)
      then show ?thesis
        by (meson list.set_intros(1))
    qed
  qed
qed
end

interpretation i_adopt_node_wf?: l_adopt_node_wfCore_DOM get_owner_document get_parent get_parent_locs
  remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs
  set_disconnected_nodes set_disconnected_nodes_locs adopt_node adopt_node_locs known_ptr
  type_wf get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes set_child_nodes_locs
  remove heap_is_wellformed parent_child_rel
  by(simp add: l_adopt_node_wfCore_DOM_def instances)
declare l_adopt_node_wfCore_DOM_axioms[instances]

interpretation i_adopt_node_wf2?: l_adopt_node_wf2Core_DOM get_owner_document get_parent get_parent_locs
  remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs
  set_disconnected_nodes set_disconnected_nodes_locs adopt_node adopt_node_locs known_ptr
  type_wf get_child_nodes get_child_nodes_locs known_ptrs set_child_nodes set_child_nodes_locs
  remove heap_is_wellformed parent_child_rel get_root_node get_root_node_locs
  by(simp add: l_adopt_node_wf2Core_DOM_def instances)
declare l_adopt_node_wf2Core_DOM_axioms[instances]


locale l_adopt_node_wf = l_heap_is_wellformed + l_known_ptrs + l_type_wf + l_adopt_node_defs
  + l_get_child_nodes_defs + l_get_disconnected_nodes_defs +
  assumes adopt_node_preserves_wellformedness:
    "heap_is_wellformed h  h  adopt_node document_ptr child h h'  known_ptrs h
                           type_wf h  heap_is_wellformed h'"
  assumes adopt_node_removes_child:
    "heap_is_wellformed h  h  adopt_node owner_document node_ptr h h2
                           h2  get_child_nodes ptr r children  known_ptrs h
                           type_wf h  node_ptr  set children"
  assumes adopt_node_node_in_disconnected_nodes:
    "heap_is_wellformed h  h  adopt_node owner_document node_ptr h h'
                           h'  get_disconnected_nodes owner_document r disc_nodes
                           known_ptrs h  type_wf h  node_ptr  set disc_nodes"
  assumes adopt_node_removes_first_child: "heap_is_wellformed h  type_wf h  known_ptrs h
                          h  adopt_node owner_document node h h'
                          h  get_child_nodes ptr' r node # children
                          h'  get_child_nodes ptr' r children"
  assumes adopt_node_document_in_heap: "heap_is_wellformed h  known_ptrs h  type_wf h
                         h  ok (adopt_node owner_document node)
                         owner_document |∈| document_ptr_kinds h"

lemma adopt_node_wf_is_l_adopt_node_wf [instances]:
  "l_adopt_node_wf type_wf known_ptr heap_is_wellformed parent_child_rel get_child_nodes
                   get_disconnected_nodes known_ptrs adopt_node"
  using heap_is_wellformed_is_l_heap_is_wellformed known_ptrs_is_l_known_ptrs
  apply(auto simp add: l_adopt_node_wf_def l_adopt_node_wf_axioms_def)[1]
  using adopt_node_preserves_wellformedness apply blast
  using adopt_node_removes_child apply blast
  using adopt_node_node_in_disconnected_nodes apply blast
  using adopt_node_removes_first_child apply blast
  using adopt_node_document_in_heap apply blast
  done


subsection ‹insert\_before›

locale l_insert_before_wfCore_DOM =
  l_insert_beforeCore_DOM +
  l_adopt_node_wf +
  l_set_disconnected_nodes_get_child_nodes +
  l_heap_is_wellformed
begin
lemma insert_before_removes_child:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "ptr  ptr'"
  assumes "h  insert_before ptr node child h h'"
  assumes "h  get_child_nodes ptr' r node # children"
  shows "h'  get_child_nodes ptr' r children"
proof -
  obtain owner_document h2 h3 disc_nodes reference_child where
    "h  (if Some node = child then a_next_sibling node else return child) r reference_child" and
    "h  get_owner_document ptr r owner_document" and
    h2: "h  adopt_node owner_document node h h2" and
    "h2  get_disconnected_nodes owner_document r disc_nodes" and
    h3: "h2  set_disconnected_nodes owner_document (remove1 node disc_nodes) h h3" and
    h': "h3  a_insert_node ptr node reference_child h h'"
    using assms(5)
    by(auto simp add: insert_before_def a_ensure_pre_insertion_validity_def
        elim!: bind_returns_heap_E bind_returns_result_E
        bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated]
        bind_returns_heap_E2[rotated, OF get_parent_pure, rotated]
        bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated]
        bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated]
        bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated]
        bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]
        split: if_splits option.splits)

  have "h2  get_child_nodes ptr' r children"
    using h2 adopt_node_removes_first_child assms(1) assms(2) assms(3) assms(6)
    by simp
  then have "h3  get_child_nodes ptr' r children"
    using h3
    by(auto simp add: set_disconnected_nodes_get_child_nodes
        dest!: reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes])
  then show ?thesis
    using h' assms(4)
    apply(auto simp add: a_insert_node_def
        elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated])[1]
    by(auto simp add: set_child_nodes_get_child_nodes_different_pointers
        elim!: reads_writes_separate_forwards[OF get_child_nodes_reads set_child_nodes_writes])
qed
end

locale l_insert_before_wf = l_heap_is_wellformed_defs + l_type_wf + l_known_ptrs
  + l_insert_before_defs + l_get_child_nodes_defs +
  assumes insert_before_removes_child:
    "heap_is_wellformed h  type_wf h  known_ptrs h  ptr  ptr'
                         h  insert_before ptr node child h h'
                         h  get_child_nodes ptr' r node # children
                         h'  get_child_nodes ptr' r children"

interpretation i_insert_before_wf?: l_insert_before_wfCore_DOM get_parent get_parent_locs
  get_child_nodes get_child_nodes_locs set_child_nodes
  set_child_nodes_locs get_ancestors get_ancestors_locs
  adopt_node adopt_node_locs set_disconnected_nodes
  set_disconnected_nodes_locs get_disconnected_nodes
  get_disconnected_nodes_locs get_owner_document insert_before
  insert_before_locs append_child type_wf known_ptr known_ptrs
  heap_is_wellformed parent_child_rel
  by(simp add: l_insert_before_wfCore_DOM_def instances)
declare l_insert_before_wfCore_DOM_axioms [instances]

lemma insert_before_wf_is_l_insert_before_wf [instances]:
  "l_insert_before_wf heap_is_wellformed type_wf known_ptr known_ptrs insert_before get_child_nodes"
  apply(auto simp add: l_insert_before_wf_def l_insert_before_wf_axioms_def instances)[1]
  using insert_before_removes_child apply fast
  done

locale l_insert_before_wf2Core_DOM =
  l_insert_before_wfCore_DOM +
  l_set_child_nodes_get_disconnected_nodes +
  l_remove_child +
  l_get_root_node_wf +
  l_set_disconnected_nodes_get_disconnected_nodes_wf +
  l_set_disconnected_nodes_get_ancestors +
  l_get_ancestors_wf +
  l_get_owner_document +
  l_heap_is_wellformedCore_DOM
begin
lemma insert_before_heap_is_wellformed_preserved:
  assumes wellformed: "heap_is_wellformed h"
    and insert_before: "h  insert_before ptr node child h h'"
    and known_ptrs: "known_ptrs h"
    and type_wf: "type_wf h"
  shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs h'"
proof -
  obtain ancestors reference_child owner_document h2 h3 disconnected_nodes_h2 where
    ancestors: "h  get_ancestors ptr r ancestors" and
    node_not_in_ancestors: "cast node  set ancestors" and
    reference_child:
    "h  (if Some node = child then a_next_sibling node else return child) r reference_child" and
    owner_document: "h  get_owner_document ptr r owner_document" and
    h2: "h  adopt_node owner_document node h h2" and
    disconnected_nodes_h2: "h2  get_disconnected_nodes owner_document r disconnected_nodes_h2" and
    h3: "h2  set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) h h3" and
    h': "h3  a_insert_node ptr node reference_child h h'"
    using assms(2)
    by(auto simp add: insert_before_def a_ensure_pre_insertion_validity_def
        elim!: bind_returns_heap_E bind_returns_result_E
        bind_returns_heap_E2[rotated, OF get_parent_pure, rotated]
        bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated]
        bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]
        bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated]
        bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated]
        bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated]
        split: if_splits option.splits)

  have "known_ptr ptr"
    by (meson get_owner_document_ptr_in_heap is_OK_returns_result_I known_ptrs
        l_known_ptrs.known_ptrs_known_ptr l_known_ptrs_axioms owner_document)

  have "type_wf h2"
    using writes_small_big[where P="λh h'. type_wf h  type_wf h'", OF adopt_node_writes h2]
    using type_wf adopt_node_types_preserved
    by(auto simp add: a_remove_child_locs_def reflp_def transp_def)
  then have "type_wf h3"
    using writes_small_big[where P="λh h'. type_wf h  type_wf h'", OF set_disconnected_nodes_writes h3]
    using  set_disconnected_nodes_types_preserved
    by(auto simp add: reflp_def transp_def)
  then show "type_wf h'"
    using writes_small_big[where P="λh h'. type_wf h  type_wf h'", OF insert_node_writes h']
    using set_child_nodes_types_preserved
    by(auto simp add: reflp_def transp_def)

  have object_ptr_kinds_M_eq3_h: "object_ptr_kinds h = object_ptr_kinds h2"
    apply(rule writes_small_big[where P="λh h'. object_ptr_kinds h = object_ptr_kinds h'",
          OF adopt_node_writes h2])
    using adopt_node_pointers_preserved
      apply blast
    by (auto simp add: reflp_def transp_def)
  then have object_ptr_kinds_M_eq_h: "ptrs. h  object_ptr_kinds_M r ptrs = h2  object_ptr_kinds_M r ptrs"
    by(simp add: object_ptr_kinds_M_defs )
  then have object_ptr_kinds_M_eq2_h: "|h  object_ptr_kinds_M|r = |h2  object_ptr_kinds_M|r"
    by simp
  then have node_ptr_kinds_eq2_h: "|h  node_ptr_kinds_M|r = |h2  node_ptr_kinds_M|r"
    using node_ptr_kinds_M_eq by blast

  have "known_ptrs h2"
    using known_ptrs object_ptr_kinds_M_eq3_h known_ptrs_preserved by blast

  have wellformed_h2: "heap_is_wellformed h2"
    using adopt_node_preserves_wellformedness[OF wellformed h2] known_ptrs type_wf .

  have object_ptr_kinds_M_eq3_h2: "object_ptr_kinds h2 = 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 h3])
    unfolding a_remove_child_locs_def
    using set_disconnected_nodes_pointers_preserved
    by (auto simp add: reflp_def transp_def)
  then have object_ptr_kinds_M_eq_h2: "ptrs. h2  object_ptr_kinds_M r ptrs = h3  object_ptr_kinds_M r ptrs"
    by(simp add: object_ptr_kinds_M_defs)
  then have object_ptr_kinds_M_eq2_h2: "|h2  object_ptr_kinds_M|r = |h3  object_ptr_kinds_M|r"
    by simp
  then have node_ptr_kinds_eq2_h2: "|h2  node_ptr_kinds_M|r = |h3  node_ptr_kinds_M|r"
    using node_ptr_kinds_M_eq by blast
  have document_ptr_kinds_eq2_h2: "|h2  document_ptr_kinds_M|r = |h3  document_ptr_kinds_M|r"
    using object_ptr_kinds_M_eq2_h2 document_ptr_kinds_M_eq by auto

  have "known_ptrs h3"
    using object_ptr_kinds_M_eq3_h2 known_ptrs_preserved known_ptrs h2 by blast

  have object_ptr_kinds_M_eq3_h': "object_ptr_kinds h3 = object_ptr_kinds h'"
    apply(rule writes_small_big[where P="λh h'. object_ptr_kinds h = object_ptr_kinds h'",
          OF insert_node_writes h'])
    unfolding a_remove_child_locs_def
    using set_child_nodes_pointers_preserved
    by (auto simp add: reflp_def transp_def)
  then have object_ptr_kinds_M_eq_h3:
    "ptrs. h3  object_ptr_kinds_M r ptrs = h'  object_ptr_kinds_M r ptrs"
    by(simp add: object_ptr_kinds_M_defs)
  then have object_ptr_kinds_M_eq2_h3:
    "|h3  object_ptr_kinds_M|r = |h'  object_ptr_kinds_M|r"
    by simp
  then have node_ptr_kinds_eq2_h3: "|h3  node_ptr_kinds_M|r = |h'  node_ptr_kinds_M|r"
    using node_ptr_kinds_M_eq by blast
  have document_ptr_kinds_eq2_h3: "|h3  document_ptr_kinds_M|r = |h'  document_ptr_kinds_M|r"
    using object_ptr_kinds_M_eq2_h3 document_ptr_kinds_M_eq by auto

  show "known_ptrs h'"
    using object_ptr_kinds_M_eq3_h' known_ptrs_preserved known_ptrs h3 by blast

  have disconnected_nodes_eq_h2:
    "doc_ptr disc_nodes. owner_document  doc_ptr
       h2  get_disconnected_nodes doc_ptr r disc_nodes = h3  get_disconnected_nodes doc_ptr r disc_nodes"
    using get_disconnected_nodes_reads set_disconnected_nodes_writes h3
    apply(rule reads_writes_preserved)
    by (auto simp add: set_disconnected_nodes_get_disconnected_nodes_different_pointers)
  then have disconnected_nodes_eq2_h2:
    "doc_ptr. doc_ptr  owner_document
      |h2  get_disconnected_nodes doc_ptr|r = |h3  get_disconnected_nodes doc_ptr|r"
    using select_result_eq by force
  have disconnected_nodes_h3:
    "h3  get_disconnected_nodes owner_document r remove1 node disconnected_nodes_h2"
    using h3 set_disconnected_nodes_get_disconnected_nodes
    by blast

  have disconnected_nodes_eq_h3:
    "doc_ptr disc_nodes. h3  get_disconnected_nodes doc_ptr r disc_nodes
                                        = h'  get_disconnected_nodes doc_ptr r disc_nodes"
    using get_disconnected_nodes_reads insert_node_writes  h'
    apply(rule reads_writes_preserved)
    using set_child_nodes_get_disconnected_nodes by fast
  then have disconnected_nodes_eq2_h3:
    "doc_ptr. |h3  get_disconnected_nodes doc_ptr|r = |h'  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_disconnected_nodes_writes h3
    apply(rule reads_writes_preserved)
    by (auto simp add: set_disconnected_nodes_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 children_eq_h3:
    "ptr' children. ptr  ptr'
         h3  get_child_nodes ptr' r children = h'  get_child_nodes ptr' r children"
    using get_child_nodes_reads insert_node_writes h'
    apply(rule reads_writes_preserved)
    by (auto simp add: set_child_nodes_get_child_nodes_different_pointers)
  then have children_eq2_h3:
    "ptr'. ptr  ptr'  |h3  get_child_nodes ptr'|r = |h'  get_child_nodes ptr'|r"
    using select_result_eq by force
  obtain children_h3 where children_h3: "h3  get_child_nodes ptr r children_h3"
    using h' a_insert_node_def by auto
  have children_h': "h'  get_child_nodes ptr r insert_before_list node reference_child children_h3"
    using h' type_wf h3 known_ptr ptr
    by(auto simp add: a_insert_node_def elim!: bind_returns_heap_E2
        dest!: set_child_nodes_get_child_nodes returns_result_eq[OF children_h3])

  have ptr_in_heap: "ptr |∈| object_ptr_kinds h3"
    using children_h3 get_child_nodes_ptr_in_heap by blast
  have node_in_heap: "node |∈| node_ptr_kinds h"
    using h2 adopt_node_child_in_heap by fast
  have child_not_in_any_children:
    "p children. h2  get_child_nodes p r children  node  set children"
    using wellformed h2 adopt_node_removes_child type_wf h known_ptrs h by auto
  have "node  set disconnected_nodes_h2"
    using disconnected_nodes_h2 h2 adopt_node_node_in_disconnected_nodes assms(1)
      type_wf h known_ptrs h by blast
  have node_not_in_disconnected_nodes:
    "d. d |∈| document_ptr_kinds h3  node  set |h3  get_disconnected_nodes d|r"
  proof -
    fix d
    assume "d |∈| document_ptr_kinds h3"
    show "node  set |h3  get_disconnected_nodes d|r"
    proof (cases "d = owner_document")
      case True
      then show ?thesis
        using disconnected_nodes_h2 wellformed_h2 h3 remove_from_disconnected_nodes_removes
          wellformed_h2 d |∈| document_ptr_kinds h3 disconnected_nodes_h3
        by fastforce
    next
      case False
      then have
        "set |h2  get_disconnected_nodes d|r  set |h2  get_disconnected_nodes owner_document|r = {}"
        using distinct_concat_map_E(1) wellformed_h2
        by (metis (no_types, lifting) d |∈| document_ptr_kinds h3 type_wf h2
            disconnected_nodes_h2 document_ptr_kinds_M_def document_ptr_kinds_eq2_h2
            l_ptr_kinds_M.ptr_kinds_ptr_kinds_M local.get_disconnected_nodes_ok
            local.heap_is_wellformed_one_disc_parent returns_result_select_result
            select_result_I2)
      then show ?thesis
        using disconnected_nodes_eq2_h2[OF False] node  set disconnected_nodes_h2
          disconnected_nodes_h2 by fastforce
    qed
  qed

  have "cast node  ptr"
    using ancestors node_not_in_ancestors get_ancestors_ptr
    by fast

  obtain ancestors_h2 where ancestors_h2: "h2  get_ancestors ptr r ancestors_h2"
    using get_ancestors_ok object_ptr_kinds_M_eq2_h2 known_ptrs h2 type_wf h2
    by (metis is_OK_returns_result_E object_ptr_kinds_M_eq3_h2 ptr_in_heap wellformed_h2)
  have ancestors_h3: "h3  get_ancestors ptr r ancestors_h2"
    using get_ancestors_reads set_disconnected_nodes_writes h3
    apply(rule reads_writes_separate_forwards)
    using heap_is_wellformed h2 ancestors_h2
    by (auto simp add: set_disconnected_nodes_get_ancestors)
  have node_not_in_ancestors_h2: "cast node  set ancestors_h2"
    apply(rule get_ancestors_remains_not_in_ancestors[OF assms(1) wellformed_h2 ancestors ancestors_h2])
    using adopt_node_children_subset using h2 known_ptrs h type_wf h apply(blast)
    using node_not_in_ancestors apply(blast)
    using object_ptr_kinds_M_eq3_h apply(blast)
    using known_ptrs h apply(blast)
    using type_wf h apply(blast)
    using type_wf h2 by blast

  moreover have "a_acyclic_heap h'"
  proof -
    have "acyclic (parent_child_rel h2)"
      using wellformed_h2  by (simp add: heap_is_wellformed_def acyclic_heap_def)
    then have "acyclic (parent_child_rel h3)"
      by(auto simp add: parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2)
    moreover have "cast node  {x. (x, ptr)  (parent_child_rel h2)*}"
      using get_ancestors_parent_child_rel node_not_in_ancestors_h2 known_ptrs h2 type_wf h2
      using ancestors_h2 wellformed_h2 by blast
    then have "cast node  {x. (x, ptr)  (parent_child_rel h3)*}"
      by(auto simp add: parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2)
    moreover have "parent_child_rel h' = insert (ptr, cast node) ((parent_child_rel h3))"
      using  children_h3 children_h' ptr_in_heap
      apply(auto simp add: parent_child_rel_def object_ptr_kinds_M_eq3_h' children_eq2_h3
          insert_before_list_node_in_set)[1]
       apply (metis (no_types, lifting) children_eq2_h3 insert_before_list_in_set select_result_I2)
      by (metis (no_types, lifting) children_eq2_h3 imageI insert_before_list_in_set select_result_I2)
    ultimately show ?thesis
      by(auto simp add: acyclic_heap_def)
  qed


  moreover have "a_all_ptrs_in_heap h2"
    using wellformed_h2  by (simp add: heap_is_wellformed_def)
  have "a_all_ptrs_in_heap h'"
  proof -
    have "a_all_ptrs_in_heap h3"
      using a_all_ptrs_in_heap h2
      apply(auto simp add: a_all_ptrs_in_heap_def object_ptr_kinds_M_eq2_h2 node_ptr_kinds_eq2_h2
          children_eq_h2)[1]
      using disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3
      using node_ptr_kinds_eq2_h2 apply auto[1]
       apply (metis known_ptrs h2 type_wf h3 children_eq_h2 local.get_child_nodes_ok
          local.heap_is_wellformed_children_in_heap local.known_ptrs_known_ptr object_ptr_kinds_M_eq3_h2
          returns_result_select_result wellformed_h2)
      by (metis (no_types, opaque_lifting) disconnected_nodes_eq2_h2 disconnected_nodes_h2
          disconnected_nodes_h3 document_ptr_kinds_commutes node_ptr_kinds_commutes
          object_ptr_kinds_M_eq3_h2 select_result_I2 set_remove1_subset subsetD)

    have "set children_h3   set |h'  node_ptr_kinds_M|r"
      using children_h3 a_all_ptrs_in_heap h3
      apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq2_h3)[1]
      by (metis children_eq_h2 l_heap_is_wellformed.heap_is_wellformed_children_in_heap
          local.l_heap_is_wellformed_axioms node_ptr_kinds_commutes object_ptr_kinds_M_eq3_h'
          object_ptr_kinds_M_eq3_h2 wellformed_h2)

    then have "set (insert_before_list node reference_child children_h3)  set |h'  node_ptr_kinds_M|r"
      using node_in_heap
      apply(auto simp add: node_ptr_kinds_eq2_h node_ptr_kinds_eq2_h2 node_ptr_kinds_eq2_h3)[1]
      by (metis (no_types, opaque_lifting) contra_subsetD insert_before_list_in_set
          node_ptr_kinds_commutes object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h'
          object_ptr_kinds_M_eq3_h2)
    then show ?thesis
      using a_all_ptrs_in_heap h3
      apply(auto simp add: object_ptr_kinds_M_eq3_h' a_all_ptrs_in_heap_def node_ptr_kinds_def
          node_ptr_kinds_eq2_h3 disconnected_nodes_eq_h3)[1]
      using children_eq_h3 children_h'
       apply (metis (no_types, lifting) children_eq2_h3 select_result_I2 subsetD)
      by (metis (no_types) type_wf h' disconnected_nodes_eq2_h3 disconnected_nodes_eq_h3
          is_OK_returns_result_I local.get_disconnected_nodes_ok
          local.get_disconnected_nodes_ptr_in_heap returns_result_select_result subsetD)
  qed

  moreover have "a_distinct_lists h2"
    using wellformed_h2  by (simp add: heap_is_wellformed_def)
  then have "a_distinct_lists h3"
  proof(auto simp add: a_distinct_lists_def object_ptr_kinds_M_eq2_h2 document_ptr_kinds_eq2_h2
      children_eq2_h2 intro!: distinct_concat_map_I)[1]
    fix x
    assume 1: "x |∈| document_ptr_kinds h3"
      and 2: "distinct (concat (map (λdocument_ptr. |h2  get_disconnected_nodes document_ptr|r)
                        (sorted_list_of_set (fset (document_ptr_kinds h3)))))"
    show "distinct |h3  get_disconnected_nodes x|r"
      using distinct_concat_map_E(2)[OF 2] select_result_I2[OF disconnected_nodes_h3]
        disconnected_nodes_eq2_h2 select_result_I2[OF disconnected_nodes_h2] 1
      by (metis (full_types) distinct_remove1 finite_fset set_sorted_list_of_set)
  next
    fix x y xa
    assume 1: "distinct (concat (map (λdocument_ptr. |h2  get_disconnected_nodes document_ptr|r)
                          (sorted_list_of_set (fset (document_ptr_kinds h3)))))"
      and 2: "x |∈| document_ptr_kinds h3"
      and 3: "y |∈| document_ptr_kinds h3"
      and 4: "x  y"
      and 5: "xa  set |h3  get_disconnected_nodes x|r"
      and 6: "xa  set |h3  get_disconnected_nodes y|r"
    show False
    proof (cases "x = owner_document")
      case True
      then have "y  owner_document"
        using 4 by simp
      show ?thesis
        using distinct_concat_map_E(1)[OF 1]
        using 2 3 4 5 6 select_result_I2[OF disconnected_nodes_h3]  select_result_I2[OF disconnected_nodes_h2]
        apply(auto simp add: True disconnected_nodes_eq2_h2[OF y  owner_document])[1]
        by (metis (no_types, opaque_lifting) disconnected_nodes_eq2_h2 disjoint_iff_not_equal notin_set_remove1)
    next
      case False
      then show ?thesis
      proof (cases "y = owner_document")
        case True
        then show ?thesis
          using distinct_concat_map_E(1)[OF 1]
          using 2 3 4 5 6 select_result_I2[OF disconnected_nodes_h3]  select_result_I2[OF disconnected_nodes_h2]
          apply(auto simp add: True disconnected_nodes_eq2_h2[OF x  owner_document])[1]
          by (metis (no_types, opaque_lifting) disconnected_nodes_eq2_h2 disjoint_iff_not_equal notin_set_remove1)
      next
        case False
        then show ?thesis
          using distinct_concat_map_E(1)[OF 1, simplified, OF 2 3 4] 5 6
          using disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3
            disjoint_iff_not_equal notin_set_remove1 select_result_I2
          by (metis (no_types, opaque_lifting))
      qed
    qed
  next
    fix x xa xb
    assume 1: "(xfset (object_ptr_kinds h3). set |h3  get_child_nodes x|r)
                      (xfset (document_ptr_kinds h3). set |h2  get_disconnected_nodes x|r) = {}"
      and 2: "xa |∈| object_ptr_kinds h3"
      and 3: "x  set |h3  get_child_nodes xa|r"
      and 4: "xb |∈| document_ptr_kinds h3"
      and 5: "x  set |h3  get_disconnected_nodes xb|r"
    have 6: "set |h3  get_child_nodes xa|r  set |h2  get_disconnected_nodes xb|r = {}"
      using 1 2 4
      by (metis type_wf h2 children_eq2_h2 document_ptr_kinds_commutes known_ptrs
          local.get_child_nodes_ok local.get_disconnected_nodes_ok
          local.heap_is_wellformed_children_disc_nodes_different local.known_ptrs_known_ptr
          object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h2 returns_result_select_result
          wellformed_h2)
    show False
    proof (cases "xb = owner_document")
      case True
      then show ?thesis
        using select_result_I2[OF disconnected_nodes_h3,folded select_result_I2[OF disconnected_nodes_h2]]
        by (metis (no_types, lifting) "3" "5" "6" disjoint_iff_not_equal notin_set_remove1)
    next
      case False
      show ?thesis
        using 2 3 4 5 6 unfolding disconnected_nodes_eq2_h2[OF False] by auto
    qed
  qed
  then have "a_distinct_lists h'"
  proof(auto simp add: a_distinct_lists_def document_ptr_kinds_eq2_h3 object_ptr_kinds_M_eq2_h3
      disconnected_nodes_eq2_h3 intro!: distinct_concat_map_I)[1]
    fix x
    assume 1: "distinct (concat (map (λptr. |h3  get_child_nodes ptr|r)
                                      (sorted_list_of_set (fset (object_ptr_kinds h')))))" and
      2: "x |∈| object_ptr_kinds h'"
    have 3: "p. p |∈| object_ptr_kinds h'  distinct |h3  get_child_nodes p|r"
      using 1  by (auto elim: distinct_concat_map_E)
    show "distinct |h'  get_child_nodes x|r"
    proof(cases "ptr = x")
      case True
      show ?thesis
        using 3[OF 2] children_h3 children_h'
        by(auto simp add: True  insert_before_list_distinct
            dest: child_not_in_any_children[unfolded children_eq_h2])
    next
      case False
      show ?thesis
        using children_eq2_h3[OF False] 3[OF 2] by auto
    qed
  next
    fix x y xa
    assume 1: "distinct (concat (map (λptr. |h3  get_child_nodes ptr|r)
                     (sorted_list_of_set (fset (object_ptr_kinds h')))))"
      and 2: "x |∈| object_ptr_kinds h'"
      and 3: "y |∈| object_ptr_kinds h'"
      and 4: "x  y"
      and 5: "xa  set |h'  get_child_nodes x|r"
      and 6: "xa  set |h'  get_child_nodes y|r"
    have 7:"set |h3  get_child_nodes x|r  set |h3  get_child_nodes y|r = {}"
      using distinct_concat_map_E(1)[OF 1] 2 3 4 by auto
    show False
    proof (cases "ptr = x")
      case True
      then have "ptr  y"
        using 4 by simp
      then show ?thesis
        using  children_h3 children_h' child_not_in_any_children[unfolded children_eq_h2] 5 6
        apply(auto simp add:  True children_eq2_h3[OF ptr  y])[1]
        by (metis (no_types, opaque_lifting) "3" "7" type_wf h3 children_eq2_h3 disjoint_iff_not_equal
            get_child_nodes_ok insert_before_list_in_set known_ptrs local.known_ptrs_known_ptr
            object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h'
            object_ptr_kinds_M_eq3_h2 returns_result_select_result select_result_I2)
    next
      case False
      then show ?thesis
      proof (cases "ptr = y")
        case True
        then show ?thesis
          using  children_h3 children_h' child_not_in_any_children[unfolded children_eq_h2] 5 6
          apply(auto simp add:  True children_eq2_h3[OF ptr  x])[1]
          by (metis (no_types, opaque_lifting) "2" "4" "7" IntI known_ptrs h3 type_wf h'
              children_eq_h3 empty_iff insert_before_list_in_set local.get_child_nodes_ok
              local.known_ptrs_known_ptr object_ptr_kinds_M_eq3_h'
              returns_result_select_result select_result_I2)
      next
        case False
        then show ?thesis
          using children_eq2_h3[OF ptr  x] children_eq2_h3[OF ptr  y] 5 6 7 by auto
      qed
    qed
  next
    fix x xa xb
    assume 1: " (xfset (object_ptr_kinds h'). set |h3  get_child_nodes x|r)
                   (xfset (document_ptr_kinds h'). set |h'  get_disconnected_nodes x|r) = {} "
      and 2: "xa |∈| object_ptr_kinds h'"
      and 3: "x  set |h'  get_child_nodes xa|r"
      and 4: "xb |∈| document_ptr_kinds h'"
      and 5: "x  set |h'  get_disconnected_nodes xb|r"
    have 6: "set |h3  get_child_nodes xa|r  set |h'  get_disconnected_nodes xb|r = {}"
      using 1 2 3 4 5
    proof -
      have "h d. ¬ type_wf h  d |∉| document_ptr_kinds h  h  ok get_disconnected_nodes d"
        using local.get_disconnected_nodes_ok by satx
      then have "h'  ok get_disconnected_nodes xb"
        using "4" type_wf h' by fastforce
      then have f1: "h3  get_disconnected_nodes xb r |h'  get_disconnected_nodes xb|r"
        by (simp add: disconnected_nodes_eq_h3)
      have "xa |∈| object_ptr_kinds h3"
        using "2" object_ptr_kinds_M_eq3_h' by blast
      then show ?thesis
        using f1 local.a_distinct_lists h3 local.distinct_lists_no_parent by fastforce
    qed
    show False
    proof (cases "ptr = xa")
      case True
      show ?thesis
        using 6 node_not_in_disconnected_nodes  3 4 5  select_result_I2[OF children_h']
          select_result_I2[OF children_h3] True disconnected_nodes_eq2_h3
        by (metis (no_types, lifting) "2" DocumentMonad.ptr_kinds_ptr_kinds_M
            a_distinct_lists h3 type_wf h' disconnected_nodes_eq_h3
            distinct_lists_no_parent document_ptr_kinds_eq2_h3 get_disconnected_nodes_ok
            insert_before_list_in_set object_ptr_kinds_M_eq3_h' returns_result_select_result)

    next
      case False
      then show ?thesis
        using 1 2 3 4 5 children_eq2_h3[OF False] by fastforce
    qed
  qed

  moreover have "a_owner_document_valid h2"
    using wellformed_h2  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_M_eq2_h2
        object_ptr_kinds_M_eq2_h3 node_ptr_kinds_eq2_h2 node_ptr_kinds_eq2_h3
        document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 children_eq2_h2)[1]
    apply(auto simp add: document_ptr_kinds_eq2_h2[simplified]  document_ptr_kinds_eq2_h3[simplified]
        object_ptr_kinds_M_eq2_h2[simplified] object_ptr_kinds_M_eq2_h3[simplified]
        node_ptr_kinds_eq2_h2[simplified] node_ptr_kinds_eq2_h3[simplified])[1]
    apply(auto simp add: disconnected_nodes_eq2_h3[symmetric])[1]
    by (smt (verit) children_eq2_h3 children_h' children_h3 disconnected_nodes_eq2_h2 disconnected_nodes_h2
        disconnected_nodes_h3 in_set_remove1 insert_before_list_in_set
        object_ptr_kinds_M_eq3_h' ptr_in_heap select_result_I2)

  ultimately show "heap_is_wellformed h'"
    by (simp add: heap_is_wellformed_def)
qed
end

locale l_insert_before_wf2 = l_type_wf + l_known_ptrs + l_insert_before_defs
  + l_heap_is_wellformed_defs + l_get_child_nodes_defs + l_remove_defs +
  assumes insert_before_preserves_type_wf:
    "heap_is_wellformed h  type_wf h  known_ptrs h  h  insert_before ptr child ref h h'
                           type_wf h'"
  assumes insert_before_preserves_known_ptrs:
    "heap_is_wellformed h  type_wf h  known_ptrs h  h  insert_before ptr child ref h h'
                           known_ptrs h'"
  assumes insert_before_heap_is_wellformed_preserved:
    "type_wf h  known_ptrs h  heap_is_wellformed h  h  insert_before ptr child ref h h'
                heap_is_wellformed h'"

interpretation i_insert_before_wf2?: l_insert_before_wf2Core_DOM get_parent get_parent_locs
  get_child_nodes get_child_nodes_locs set_child_nodes
  set_child_nodes_locs get_ancestors get_ancestors_locs
  adopt_node adopt_node_locs set_disconnected_nodes
  set_disconnected_nodes_locs get_disconnected_nodes
  get_disconnected_nodes_locs get_owner_document insert_before
  insert_before_locs append_child type_wf known_ptr known_ptrs
  heap_is_wellformed parent_child_rel remove_child
  remove_child_locs get_root_node get_root_node_locs
  by(simp add: l_insert_before_wf2Core_DOM_def instances)
declare l_insert_before_wf2Core_DOM_axioms [instances]

lemma insert_before_wf2_is_l_insert_before_wf2 [instances]:
  "l_insert_before_wf2 type_wf known_ptr known_ptrs insert_before heap_is_wellformed"
  apply(auto simp add: l_insert_before_wf2_def l_insert_before_wf2_axioms_def instances)[1]
  using insert_before_heap_is_wellformed_preserved apply(fast, fast, fast)
  done

locale l_append_child_wfCore_DOM =
  l_adopt_nodeCore_DOM +
  l_insert_beforeCore_DOM +
  l_append_childCore_DOM +
  l_insert_before_wf +
  l_insert_before_wf2  +
  l_get_child_nodes
begin


lemma append_child_heap_is_wellformed_preserved:
  assumes wellformed: "heap_is_wellformed h"
    and append_child: "h  append_child ptr node h h'"
    and known_ptrs: "known_ptrs h"
    and type_wf: "type_wf h"
  shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs h'"
  using assms
  by(auto simp add: append_child_def intro: insert_before_preserves_type_wf
      insert_before_preserves_known_ptrs insert_before_heap_is_wellformed_preserved)

lemma append_child_children:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_child_nodes ptr r xs"
  assumes "h  append_child ptr node h h'"
  assumes "node  set xs"
  shows "h'  get_child_nodes ptr r xs @ [node]"
proof -

  obtain ancestors owner_document h2 h3 disconnected_nodes_h2 where
    ancestors: "h  get_ancestors ptr r ancestors" and
    node_not_in_ancestors: "cast node  set ancestors" and
    owner_document: "h  get_owner_document ptr r owner_document" and
    h2: "h  adopt_node owner_document node h h2" and
    disconnected_nodes_h2: "h2  get_disconnected_nodes owner_document r disconnected_nodes_h2" and
    h3: "h2  set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) h h3" and
    h': "h3  a_insert_node ptr node None h h'"
    using assms(5)
    by(auto simp add: append_child_def insert_before_def a_ensure_pre_insertion_validity_def
        elim!: bind_returns_heap_E bind_returns_result_E
        bind_returns_heap_E2[rotated, OF get_parent_pure, rotated]
        bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated]
        bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]
        bind_returns_heap_E2[rotated, OF get_ancestors_pure, rotated]
        bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated]
        bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated]
        split: if_splits option.splits)

  have "parent. |h  get_parent node|r = Some parent  parent  ptr"
    using assms(1) assms(4) assms(6)
    by (metis (no_types, lifting) assms(2) assms(3) h2 is_OK_returns_heap_I is_OK_returns_result_E
        local.adopt_node_child_in_heap local.get_parent_child_dual local.get_parent_ok
        select_result_I2)
  have "h2  get_child_nodes ptr r xs"
    using get_child_nodes_reads adopt_node_writes h2 assms(4)
    apply(rule reads_writes_separate_forwards)
    using parent. |h  get_parent node|r = Some parent  parent  ptr
    apply(auto simp add: adopt_node_locs_def remove_child_locs_def)[1]
    by (meson local.set_child_nodes_get_child_nodes_different_pointers)

  have "h3  get_child_nodes ptr r xs"
    using get_child_nodes_reads set_disconnected_nodes_writes h3 h2  get_child_nodes ptr r xs
    apply(rule reads_writes_separate_forwards)
    by(auto)

  have "ptr |∈| object_ptr_kinds h"
    by (meson ancestors is_OK_returns_result_I local.get_ancestors_ptr_in_heap)
  then
  have "known_ptr ptr"
    using assms(3)
    using 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 adopt_node_writes h2]
    using adopt_node_types_preserved  type_wf h
    by(auto simp add: adopt_node_locs_def remove_child_locs_def reflp_def transp_def split: if_splits)
  then
  have "type_wf h3"
    using writes_small_big[where P="λh h'. type_wf h  type_wf h'", OF set_disconnected_nodes_writes h3]
    using  set_disconnected_nodes_types_preserved
    by(auto simp add: reflp_def transp_def)

  show "h'  get_child_nodes ptr r xs@[node]"
    using h'
    apply(auto simp add: a_insert_node_def
        dest!: bind_returns_heap_E3[rotated, OF h3  get_child_nodes ptr r xs
          get_child_nodes_pure, rotated])[1]
    using type_wf h3 set_child_nodes_get_child_nodes known_ptr ptr
    by metis
qed

lemma append_child_for_all_on_children:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_child_nodes ptr r xs"
  assumes "h  forall_M (append_child ptr) nodes h h'"
  assumes "set nodes  set xs = {}"
  assumes "distinct nodes"
  shows "h'  get_child_nodes ptr r xs@nodes"
  using assms
  apply(induct nodes arbitrary: h xs)
   apply(simp)
proof(auto elim!: bind_returns_heap_E)[1]fix a nodes h xs h'a
  assume 0: "(h xs. heap_is_wellformed h  type_wf h  known_ptrs h
               h  get_child_nodes ptr r xs  h  forall_M (append_child ptr) nodes h h'
               set nodes  set xs = {}  h'  get_child_nodes ptr r xs @ nodes)"
    and 1: "heap_is_wellformed h"
    and 2: "type_wf h"
    and 3: "known_ptrs h"
    and 4: "h  get_child_nodes ptr r xs"
    and 5: "h  append_child ptr a r ()"
    and 6: "h  append_child ptr a h h'a"
    and 7: "h'a  forall_M (append_child ptr) nodes h h'"
    and 8: "a  set xs"
    and 9: "set nodes  set xs = {}"
    and 10: "a  set nodes"
    and 11: "distinct nodes"
  then have "h'a  get_child_nodes ptr r xs @ [a]"
    using append_child_children 6
    using "1" "2" "3" "4" "8" by blast

  moreover have "heap_is_wellformed h'a" and "type_wf h'a" and "known_ptrs h'a"
    using insert_before_heap_is_wellformed_preserved insert_before_preserves_known_ptrs
      insert_before_preserves_type_wf 1 2 3 6 append_child_def
    by metis+
  moreover have "set nodes  set (xs @ [a]) = {}"
    using 9 10
    by auto
  ultimately show "h'  get_child_nodes ptr r xs @ a # nodes"
    using 0 7
    by fastforce
qed


lemma append_child_for_all_on_no_children:
  assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
  assumes "h  get_child_nodes ptr r []"
  assumes "h  forall_M (append_child ptr) nodes h h'"
  assumes "distinct nodes"
  shows "h'  get_child_nodes ptr r nodes"
  using assms append_child_for_all_on_children
  by force
end

locale l_append_child_wf = l_type_wf + l_known_ptrs + l_append_child_defs + l_heap_is_wellformed_defs +
  assumes append_child_preserves_type_wf:
    "heap_is_wellformed h  type_wf h  known_ptrs h  h  append_child ptr child h h'
                           type_wf h'"
  assumes append_child_preserves_known_ptrs:
    "heap_is_wellformed h  type_wf h  known_ptrs h  h  append_child ptr child h h'
                           known_ptrs h'"
  assumes append_child_heap_is_wellformed_preserved:
    "type_wf h  known_ptrs h  heap_is_wellformed h  h  append_child ptr child h h'
                heap_is_wellformed h'"

interpretation i_append_child_wf?: l_append_child_wfCore_DOM get_owner_document get_parent
  get_parent_locs remove_child remove_child_locs
  get_disconnected_nodes get_disconnected_nodes_locs
  set_disconnected_nodes set_disconnected_nodes_locs
  adopt_node adopt_node_locs known_ptr type_wf get_child_nodes
  get_child_nodes_locs known_ptrs set_child_nodes
  set_child_nodes_locs remove get_ancestors get_ancestors_locs
  insert_before insert_before_locs append_child heap_is_wellformed
  parent_child_rel
  by(auto simp add: l_append_child_wfCore_DOM_def instances)

lemma append_child_wf_is_l_append_child_wf [instances]: "l_append_child_wf type_wf known_ptr
known_ptrs append_child heap_is_wellformed"
  apply(auto simp add: l_append_child_wf_def l_append_child_wf_axioms_def instances)[1]
  using append_child_heap_is_wellformed_preserved by fast+


subsection ‹create\_element›

locale l_create_element_wfCore_DOM =
  l_heap_is_wellformedCore_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 +
  l_new_element_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs +
  l_set_tag_name_get_disconnected_nodes type_wf set_tag_name set_tag_name_locs
  get_disconnected_nodes get_disconnected_nodes_locs +
  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_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_set_disconnected_nodes_get_child_nodes set_disconnected_nodes set_disconnected_nodes_locs
  get_child_nodes get_child_nodes_locs +
  l_set_disconnected_nodes type_wf set_disconnected_nodes set_disconnected_nodes_locs +
  l_set_disconnected_nodes_get_disconnected_nodes  type_wf get_disconnected_nodes
  get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs +
  l_new_element type_wf +
  l_known_ptrs known_ptr known_ptrs
  for known_ptr :: "(_::linorder) object_ptr  bool"
    and known_ptrs :: "(_) heap  bool"
    and type_wf :: "(_) heap  bool"
    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_disconnected_nodes :: "(_) document_ptr  ((_) heap, exception, (_) node_ptr list) prog"
    and get_disconnected_nodes_locs :: "(_) document_ptr  ((_) heap  (_) heap  bool) set"
    and heap_is_wellformed :: "(_) heap  bool"
    and parent_child_rel :: "(_) heap  ((_) object_ptr × (_) object_ptr) 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 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_element :: "(_) document_ptr  char list  ((_) heap, exception, (_) element_ptr) prog"
begin
lemma create_element_preserves_wellformedness:
  assumes "heap_is_wellformed h"
    and "h  create_element document_ptr tag h h'"
    and "type_wf h"
    and "known_ptrs h"
  shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs 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(2)
    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
  show "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 show "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(1)
        assms(3) 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(1) assms(3) 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 (verit) 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 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(3) assms(4) 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_remove)
      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(1) assms(3) assms(4) 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 (verit) 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 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 (verit) 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 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, opaque_lifting) "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, opaque_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 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)

  show "heap_is_wellformed h'"
    using a_acyclic_heap h' a_all_ptrs_in_heap h' a_distinct_lists h' a_owner_document_valid h'
    by(simp add: heap_is_wellformed_def)
qed
end

interpretation i_create_element_wf?: 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
  using instances
  by(auto simp add: l_create_element_wfCore_DOM_def)
declare l_create_element_wfCore_DOM_axioms [instances]


subsection ‹create\_character\_data›

locale l_create_character_data_wfCore_DOM =
  l_heap_is_wellformedCore_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
  + l_new_character_data_get_disconnected_nodes
  get_disconnected_nodes get_disconnected_nodes_locs
  + l_set_val_get_disconnected_nodes
  type_wf set_val set_val_locs get_disconnected_nodes get_disconnected_nodes_locs
  + 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_new_character_data_get_child_nodes
  type_wf known_ptr get_child_nodes get_child_nodes_locs
  + l_set_val_get_child_nodes
  type_wf set_val set_val_locs known_ptr get_child_nodes get_child_nodes_locs
  + l_set_disconnected_nodes_get_child_nodes
  set_disconnected_nodes set_disconnected_nodes_locs get_child_nodes get_child_nodes_locs
  + l_set_disconnected_nodes
  type_wf set_disconnected_nodes set_disconnected_nodes_locs
  + l_set_disconnected_nodes_get_disconnected_nodes
  type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes
  set_disconnected_nodes_locs
  + l_new_character_data
  type_wf
  + l_known_ptrs
  known_ptr known_ptrs
  for known_ptr :: "(_::linorder) object_ptr  bool"
    and type_wf :: "(_) heap  bool"
    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_disconnected_nodes :: "(_) document_ptr  ((_) heap, exception, (_) node_ptr list) prog"
    and get_disconnected_nodes_locs :: "(_) document_ptr  ((_) heap  (_) heap  bool) set"
    and heap_is_wellformed :: "(_) heap  bool"
    and parent_child_rel :: "(_) heap  ((_) object_ptr × (_) object_ptr) set"
    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 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"
    and known_ptrs :: "(_) heap  bool"
begin

lemma create_character_data_preserves_wellformedness:
  assumes "heap_is_wellformed h"
    and "h  create_character_data document_ptr text h h'"
    and "type_wf h"
    and "known_ptrs h"
  shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs 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(2)
    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 "known_ptr (cast new_character_data_ptr)"
    using h  create_character_data document_ptr text r new_character_data_ptr
      local.create_character_data_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
  show "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_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 show "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 "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
          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 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]
    using node_ptr_kinds_eq_h cast new_character_data_ptr  set |h  node_ptr_kinds_M|r
      h2  get_child_nodes (cast new_character_data_ptr) r []
     apply (metis (no_types, lifting) NodeMonad.ptr_kinds_ptr_kinds_M parent_child_rel h = parent_child_rel h2
        children_eq2_h finsert_iff funion_finsert_right local.parent_child_rel_child
        local.parent_child_rel_parent_in_heap node_ptr_kinds_commutes object_ptr_kinds_eq_h
        select_result_I2 subsetD sup_bot.right_neutral)
    by (metis assms(1) assms(3) disconnected_nodes_eq2_h document_ptr_kinds_eq_h funionI1
        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 (verit) character_data_ptr_kinds_commutes children_eq2_h3 disc_nodes_document_ptr_h2
        disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3
        h' h2 local.a_all_ptrs_in_heap_def
        local.set_disconnected_nodes_get_disconnected_nodes new_character_data_ptr
        new_character_data_ptr_in_heap node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h3
        object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1))
  have "p. p |∈| object_ptr_kinds h  cast new_character_data_ptr  set |h  get_child_nodes p|r"
    using heap_is_wellformed h cast new_character_data_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(3) assms(4) 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_character_data_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 (cast new_character_data_ptr) r [] apply auto[1]
    by (metis ObjectMonad.ptr_kinds_ptr_kinds_M cast new_character_data_ptr  set |h  object_ptr_kinds_M|r)
  then have "p. p |∈| object_ptr_kinds h3  cast new_character_data_ptr  set |h3  get_child_nodes p|r"
    using object_ptr_kinds_eq_h2 children_eq2_h2 by auto
  then have new_character_data_ptr_not_in_any_children:
    "p. p |∈| object_ptr_kinds h'  cast new_character_data_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_character_data_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_remove)
      apply(case_tac "x=cast new_character_data_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(1) assms(3) assms(4) 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)[1]
  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) cast new_character_data_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"
      by (smt (verit) NodeMonad.ptr_kinds_ptr_kinds_M castcharacter_data_ptr2node_ptr new_character_data_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 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(cases "xb = document_ptr")
       apply (metis (no_types, opaque_lifting) "3" "4" "6"
          p. p |∈| object_ptr_kinds h3  cast new_character_data_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(simp add: a_owner_document_valid_def)
    apply(simp add: object_ptr_kinds_eq_h object_ptr_kinds_eq_h3 )
    apply(simp add: object_ptr_kinds_eq_h2)
    apply(simp add: document_ptr_kinds_eq_h document_ptr_kinds_eq_h3 )
    apply(simp add: document_ptr_kinds_eq_h2)
    apply(simp add: node_ptr_kinds_eq_h node_ptr_kinds_eq_h3 )
    apply(simp add: node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h )
    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 (mono_tags, opaque_lifting) castcharacter_data_ptr2object_ptr new_character_data_ptr  set |h  object_ptr_kinds_M|r
        children_eq2_h disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h h'
        l_ptr_kinds_M.ptr_kinds_ptr_kinds_M
        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 object_ptr_kinds_M_def
        select_result_I2)


  show "heap_is_wellformed h'"
    using a_acyclic_heap h' a_all_ptrs_in_heap h' a_distinct_lists h' a_owner_document_valid h'
    by(simp add: heap_is_wellformed_def)
qed
end

interpretation i_create_character_data_wf?: 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
  using instances
  by (auto simp add: l_create_character_data_wfCore_DOM_def)
declare l_create_character_data_wfCore_DOM_axioms [instances]


subsection ‹create\_document›

locale l_create_document_wfCore_DOM =
  l_heap_is_wellformedCore_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
  + l_new_document_get_disconnected_nodes
  get_disconnected_nodes get_disconnected_nodes_locs
  + l_create_documentCore_DOM
  create_document
  + l_new_document_get_child_nodes
  type_wf known_ptr get_child_nodes get_child_nodes_locs
  + l_new_document
  type_wf
  + l_known_ptrs
  known_ptr known_ptrs
  for known_ptr :: "(_::linorder) object_ptr  bool"
    and type_wf :: "(_) heap  bool"
    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_disconnected_nodes :: "(_) document_ptr  ((_) heap, exception, (_) node_ptr list) prog"
    and get_disconnected_nodes_locs :: "(_) document_ptr  ((_) heap  (_) heap  bool) set"
    and heap_is_wellformed :: "(_) heap  bool"
    and parent_child_rel :: "(_) heap  ((_) object_ptr × (_) object_ptr) set"
    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 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_document :: "((_) heap, exception, (_) document_ptr) prog"
    and known_ptrs :: "(_) heap  bool"
begin

lemma create_document_preserves_wellformedness:
  assumes "heap_is_wellformed h"
    and "h  create_document h h'"
    and "type_wf h"
    and "known_ptrs h"
  shows "heap_is_wellformed h'"
proof -
  obtain new_document_ptr where
    new_document_ptr: "h  new_document r new_document_ptr" and
    h': "h  new_document h h'"
    using assms(2)
    apply(simp add: create_document_def)
    using new_document_ok by blast

  have "new_document_ptr  set |h  document_ptr_kinds_M|r"
    using new_document_ptr DocumentMonad.ptr_kinds_ptr_kinds_M
    using new_document_ptr_not_in_heap h' by blast
  then have "cast new_document_ptr  set |h  object_ptr_kinds_M|r"
    by simp

  have "new_document_ptr |∉| document_ptr_kinds h"
    using new_document_ptr DocumentMonad.ptr_kinds_ptr_kinds_M
    using new_document_ptr_not_in_heap h' by blast
  then have "cast new_document_ptr |∉| object_ptr_kinds h"
    by simp

  have object_ptr_kinds_eq: "object_ptr_kinds h' = object_ptr_kinds h |∪| {|cast new_document_ptr|}"
    using new_document_new_ptr h' new_document_ptr by blast
  then have node_ptr_kinds_eq: "node_ptr_kinds h' = node_ptr_kinds h"
    apply(simp add: node_ptr_kinds_def)
    by force
  then have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h' = character_data_ptr_kinds h"
    by(simp add: character_data_ptr_kinds_def)
  have element_ptr_kinds_eq_h: "element_ptr_kinds h' = element_ptr_kinds h"
    using object_ptr_kinds_eq
    by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def)
  have document_ptr_kinds_eq_h: "document_ptr_kinds h' = document_ptr_kinds h |∪| {|new_document_ptr|}"
    using object_ptr_kinds_eq
    by (auto simp add: document_ptr_kinds_def)


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


  have "h'  get_child_nodes (cast new_document_ptr) r []"
    using new_document_ptr h' new_document_ptr_in_heap[OF h' new_document_ptr]
      new_document_is_document_ptr[OF new_document_ptr] new_document_no_child_nodes
    by blast
  have disconnected_nodes_eq_h:
    "doc_ptr disc_nodes. doc_ptr  new_document_ptr
     h  get_disconnected_nodes doc_ptr r disc_nodes = h'  get_disconnected_nodes doc_ptr r disc_nodes"
    using get_disconnected_nodes_reads h' get_disconnected_nodes_new_document_different_pointers new_document_ptr
    apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1]
    by (smt (verit))+
  then have disconnected_nodes_eq2_h: "doc_ptr. doc_ptr  new_document_ptr
                  |h  get_disconnected_nodes doc_ptr|r = |h'  get_disconnected_nodes doc_ptr|r"
    using select_result_eq by force
  have "h'  get_disconnected_nodes new_document_ptr r []"
    using h' local.new_document_no_disconnected_nodes new_document_ptr by blast

  have "type_wf h'"
    using type_wf h new_document_types_preserved h' 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 h'"
  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 h'"
      by (simp add: object_ptr_kinds_eq)
  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 |h'  get_child_nodes a|r"
      by (metis ObjectMonad.ptr_kinds_ptr_kinds_M
          cast new_document_ptr  set |h  object_ptr_kinds_M|r children_eq2)
  next
    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 h"
      using object_ptr_kinds_eq h'  get_child_nodes (cast new_document_ptr) r []
      by(auto)
  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 |h  get_child_nodes a|r"
      by (metis (no_types, lifting) h'  get_child_nodes (cast new_document_ptr) r []
          children_eq2 empty_iff empty_set image_eqI select_result_I2)
  qed
  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 h'"
    apply(auto simp add: a_all_ptrs_in_heap_def)[1]
    using  ObjectMonad.ptr_kinds_ptr_kinds_M
      castdocument_ptr2object_ptr new_document_ptr  set |h  object_ptr_kinds_M|r
      parent_child_rel h = parent_child_rel h' assms(1) children_eq fset_of_list_elem
      local.heap_is_wellformed_children_in_heap local.parent_child_rel_child
      local.parent_child_rel_parent_in_heap node_ptr_kinds_eq
     apply (metis (no_types, opaque_lifting) h'  get_child_nodes (castdocument_ptr2object_ptr new_document_ptr) r []
        children_eq2 finsert_iff funion_finsert_right object_ptr_kinds_eq
        select_result_I2 subsetD sup_bot.right_neutral)
    by (metis (no_types, lifting) castdocument_ptr2object_ptr new_document_ptr |∉| object_ptr_kinds h
        h'  get_child_nodes (castdocument_ptr2object_ptr new_document_ptr) r []
        h'  get_disconnected_nodes new_document_ptr r []
        parent_child_rel h = parent_child_rel h' type_wf h' assms(1) disconnected_nodes_eq_h
        local.get_disconnected_nodes_ok
        local.heap_is_wellformed_disc_nodes_in_heap local.parent_child_rel_child
        local.parent_child_rel_parent_in_heap
        node_ptr_kinds_eq returns_result_select_result select_result_I2)
  have "a_distinct_lists h"
    using heap_is_wellformed h
    by (simp add: heap_is_wellformed_def)
  then have "a_distinct_lists h'"
    using h'  get_disconnected_nodes new_document_ptr r []
      h'  get_child_nodes (cast new_document_ptr) r []

    apply(auto simp add: children_eq2[symmetric] a_distinct_lists_def insort_split object_ptr_kinds_eq
        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_remove)

         apply(auto simp add:  dest: distinct_concat_map_E)[1]
        apply(auto simp add:  dest: distinct_concat_map_E)[1]
    using new_document_ptr |∉| document_ptr_kinds h
       apply(auto simp add: distinct_insort dest: distinct_concat_map_E)[1]
    using disconnected_nodes_eq_h
      apply (metis assms(1) assms(3) disconnected_nodes_eq2_h local.get_disconnected_nodes_ok
        local.heap_is_wellformed_disconnected_nodes_distinct
        returns_result_select_result)
  proof -
    fix x :: "(_) document_ptr" and y :: "(_) document_ptr" and xa :: "(_) node_ptr"
    assume a1: "x  y"
    assume a2: "x |∈| document_ptr_kinds h"
    assume a3: "x  new_document_ptr"
    assume a4: "y |∈| document_ptr_kinds h"
    assume a5: "y  new_document_ptr"
    assume a6: "distinct (concat (map (λdocument_ptr. |h  get_disconnected_nodes document_ptr|r)
                                              (sorted_list_of_set (fset (document_ptr_kinds h)))))"
    assume a7: "xa  set |h'  get_disconnected_nodes x|r"
    assume a8: "xa  set |h'  get_disconnected_nodes y|r"
    have f9: "xa  set |h  get_disconnected_nodes x|r"
      using a7 a3 disconnected_nodes_eq2_h by presburger
    have f10: "xa  set |h  get_disconnected_nodes y|r"
      using a8 a5 disconnected_nodes_eq2_h by presburger
    have f11: "y  set (sorted_list_of_set (fset (document_ptr_kinds h)))"
      using a4 by simp
    have "x  set (sorted_list_of_set (fset (document_ptr_kinds h)))"
      using a2 by simp
    then show False
      using f11 f10 f9 a6 a1 by (meson disjoint_iff_not_equal distinct_concat_map_E(1))
  next
    fix x xa xb
    assume 0: "h'  get_disconnected_nodes new_document_ptr r []"
      and 1: "h'  get_child_nodes (castdocument_ptr2object_ptr new_document_ptr) r []"
      and 2: "distinct (concat (map (λptr. |h  get_child_nodes ptr|r)
                                                  (sorted_list_of_set (fset (object_ptr_kinds h)))))"
      and 3: "distinct (concat (map (λdocument_ptr. |h  get_disconnected_nodes document_ptr|r)
                                              (sorted_list_of_set (fset (document_ptr_kinds h)))))"
      and 4: "(xfset (object_ptr_kinds h). set |h  get_child_nodes x|r)
                       (xfset (document_ptr_kinds h). set |h  get_disconnected_nodes x|r) = {}"
      and 5: "x  set |h  get_child_nodes xa|r"
      and 6: "x  set |h'  get_disconnected_nodes xb|r"
      and 7: "xa |∈| object_ptr_kinds h"
      and 8: "xa  castdocument_ptr2object_ptr new_document_ptr"
      and 9: "xb |∈| document_ptr_kinds h"
      and 10: "xb  new_document_ptr"
    then show "False"

      by (metis local.a_distinct_lists h assms(3) disconnected_nodes_eq2_h
          local.distinct_lists_no_parent local.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'"
    apply(auto simp add: a_owner_document_valid_def)[1]
    by (metis castdocument_ptr2object_ptr new_document_ptr |∉| object_ptr_kinds h
        children_eq2 disconnected_nodes_eq2_h document_ptr_kinds_commutes
        funion_iff node_ptr_kinds_eq object_ptr_kinds_eq)
  show "heap_is_wellformed h'"
    using a_acyclic_heap h' a_all_ptrs_in_heap h' a_distinct_lists h' a_owner_document_valid h'
    by(simp add: heap_is_wellformed_def)
qed
end

interpretation i_create_document_wf?: l_create_document_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_document known_ptrs
  using instances
  by (auto simp add: l_create_document_wfCore_DOM_def)
declare l_create_document_wfCore_DOM_axioms [instances]


end