diff --git a/thys/Core_DOM/standard/Core_DOM_Heap_WF.thy b/thys/Core_DOM/standard/Core_DOM_Heap_WF.thy --- a/thys/Core_DOM/standard/Core_DOM_Heap_WF.thy +++ b/thys/Core_DOM/standard/Core_DOM_Heap_WF.thy @@ -1,8045 +1,8045 @@ (*********************************************************************************** * 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_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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|\<^sub>r) \ (\parent_ptr. parent_ptr |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>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|\<^sub>r) (sorted_list_of_fset (object_ptr_kinds h)) @ map (\parent. |h \ get_disconnected_nodes parent|\<^sub>r) (sorted_list_of_fset (document_ptr_kinds h)))) " apply(auto simp add: a_owner_document_valid_def l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_owner_document_valid_def)[1] proof - fix x assume 1: " \node_ptr\fset (node_ptr_kinds h). (\document_ptr. document_ptr |\| document_ptr_kinds h \ node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r) \ (\parent_ptr. parent_ptr |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>r)" assume 2: "x |\| node_ptr_kinds h" assume 3: "x |\| fset_of_list (concat (map (\parent. |h \ get_disconnected_nodes parent|\<^sub>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|\<^sub>r)" using 1 2 3 by (smt UN_I fset_of_list_elem image_eqI notin_fset 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|\<^sub>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|\<^sub>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|\<^sub>r \ set (map (\parent. |h \ get_child_nodes parent|\<^sub>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|\<^sub>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|\<^sub>r) (sorted_list_of_fset (object_ptr_kinds h)))) |\| fset_of_list (concat (map (\parent. |h \ get_disconnected_nodes parent|\<^sub>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|\<^sub>r" have "node_ptr \ set (concat (map (\parent. |h \ get_child_nodes parent|\<^sub>r) (sorted_list_of_fset (object_ptr_kinds h)))) \ node_ptr \ set (concat (map (\parent. |h \ get_disconnected_nodes parent|\<^sub>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|\<^sub>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|\<^sub>r}" lemma a_parent_child_rel_code [code]: "a_parent_child_rel h = set (concat (map (\parent. map (\child. (parent, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child)) |h \ get_child_nodes parent|\<^sub>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|\<^sub>r \ fset (node_ptr_kinds h)) \ (\document_ptr \ fset (document_ptr_kinds h). set |h \ get_disconnected_nodes document_ptr|\<^sub>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|\<^sub>r) |h \ object_ptr_kinds_M|\<^sub>r) @ (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) |h \ document_ptr_kinds_M|\<^sub>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_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs defines heap_is_wellformed = "l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_heap_is_wellformed get_child_nodes get_disconnected_nodes" and parent_child_rel = "l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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 \\<^sub>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 \\<^sub>r children" and "child \ set children" and "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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|\<^sub>r. (\child \ set |h \ get_child_nodes ptr|\<^sub>r. {(ptr, cast child)}))" by(auto simp add: parent_child_rel_def) moreover have "finite (\ptr \ set |h \ object_ptr_kinds_M|\<^sub>r. (\child \ set |h \ get_child_nodes ptr|\<^sub>r. {(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>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|\<^sub>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: "(\x\fset (object_ptr_kinds h). set |h \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h). set |h \ get_disconnected_nodes x|\<^sub>r) = {}" assume a3: "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" assume a4: "node_ptr \ set disc_nodes" assume a5: "node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>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 fmember.rep_eq get_disconnected_nodes_ptr_in_heap is_OK_returns_result_I) have "|h \ get_disconnected_nodes document_ptr|\<^sub>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 \\<^sub>r disc_nodes" shows "distinct disc_nodes" proof - have h1: "distinct (concat (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) |h \ document_ptr_kinds_M|\<^sub>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 \\<^sub>r children" shows "distinct children" proof (cases "children = []", simp) assume "children \ []" have h1: "distinct (concat ((map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) |h \ object_ptr_kinds_M|\<^sub>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 \\<^sub>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, lifting) finite_set_in 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 \\<^sub>r children" assumes "h \ get_child_nodes ptr' \\<^sub>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 \\<^sub>r children" assume a3: "h \ get_child_nodes ptr' \\<^sub>r children'" assume a4: "distinct (concat (map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h)))))" have f5: "|h \ get_child_nodes ptr|\<^sub>r = children" using a2 by simp have "|h \ get_child_nodes ptr'|\<^sub>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 fmember.rep_eq 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 \\<^sub>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 \\<^sub>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 finite_set_in subsetD) lemma heap_is_wellformed_disc_nodes_in_heap: "heap_is_wellformed h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ node \ set disc_nodes \ node |\| node_ptr_kinds h" by (metis (no_types, lifting) finite_set_in 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 \\<^sub>r disc_nodes \ h \ get_disconnected_nodes document_ptr' \\<^sub>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 \\<^sub>r disc_nodes" assume a3: "h \ get_disconnected_nodes document_ptr' \\<^sub>r disc_nodes'" assume a4: "set disc_nodes \ set disc_nodes' \ {}" have f5: "|h \ get_disconnected_nodes document_ptr|\<^sub>r = disc_nodes" using a2 by (meson select_result_I2) have f6: "|h \ get_disconnected_nodes document_ptr'|\<^sub>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|\<^sub>r) |h \ document_ptr_kinds_M|\<^sub>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 \\<^sub>r children \ h \ get_disconnected_nodes document_ptr \\<^sub>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|\<^sub>r) \ (\document_ptr \ fset (document_ptr_kinds h). node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" apply(auto simp add: heap_is_wellformed_def a_distinct_lists_def a_owner_document_valid_def)[1] by (meson fmember.rep_eq) lemma heap_is_wellformed_children_distinct: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r children \ h \ get_child_nodes ptr' \\<^sub>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 \\<^sub>r disc_nodes \ h \ get_disconnected_nodes document_ptr' \\<^sub>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 \\<^sub>r children \ h \ get_disconnected_nodes document_ptr \\<^sub>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 \\<^sub>r disc_nodes \ distinct disc_nodes" assumes heap_is_wellformed_children_distinct: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>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|\<^sub>r) \ (\document_ptr \ fset (document_ptr_kinds h). node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" assumes parent_child_rel_child: "h \ get_child_nodes ptr \\<^sub>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_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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 \\<^sub>r children" assumes "child \ set children" assumes "known_ptrs h" assumes type_wf: "type_wf h" shows "h \ get_parent child \\<^sub>r Some ptr" proof - obtain ptrs where ptrs: "h \ object_ptr_kinds_M \\<^sub>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 \\<^sub>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 \\<^sub>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)) \\<^sub>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) \\<^sub>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)) \\<^sub>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)) \\<^sub>r True" and 2: "y \ set ptrs" and 3: "h \ Heap_Error_Monad.bind (get_child_nodes y) (\children. return (child \ set children)) \\<^sub>r True" and 4: "(\h ptr children ptr' children'. heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ get_child_nodes ptr' \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r [] \ P parent" and step: "\parent children child. h \ get_child_nodes parent \\<^sub>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 \\<^sub>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 \\<^sub>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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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 \\<^sub>h h'" and "writes SW f h h'" and preserved_get_child_nodes: "\h h' w. w \ SW \ h \ w \\<^sub>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 \\<^sub>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 \\<^sub>h h' \ \object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t 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 \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" unfolding object_ptr_kinds_M_defs by simp then have object_ptr_kinds_eq2: "|h \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" using select_result_eq by force then have node_ptr_kinds_eq2: "|h \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>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|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>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 \\<^sub>r children = h' \ get_child_nodes ptr \\<^sub>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|\<^sub>r = |h' \ get_child_nodes ptr|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq: "\document_ptr disconnected_nodes. h \ get_disconnected_nodes document_ptr \\<^sub>r disconnected_nodes = h' \ get_disconnected_nodes document_ptr \\<^sub>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|\<^sub>r = |h' \ get_disconnected_nodes document_ptr|\<^sub>r" using select_result_eq by force have get_parent_eq: "\ptr parent. h \ get_parent ptr \\<^sub>r parent = h' \ get_parent ptr \\<^sub>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_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h))) = map (\ptr. |h' \ get_child_nodes ptr|\<^sub>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|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h))) = map (\document_ptr. |h' \ get_disconnected_nodes document_ptr|\<^sub>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_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms by (simp add: l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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 \\<^sub>r children \ child \ set children \ h \ get_parent child \\<^sub>r Some ptr" assumes heap_wellformed_induct [consumes 1, case_names step]: "heap_is_wellformed h \ (\parent. (\children child. h \ get_child_nodes parent \\<^sub>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 \\<^sub>r Some parent \ P parent) \ P child) \ P ptr" assumes parent_child_rel_parent: "heap_is_wellformed h \ h \ get_parent child_node \\<^sub>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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = 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 \\<^sub>r disc_nodes" assumes "h \ set_disconnected_nodes ptr (remove1 node_ptr disc_nodes) \\<^sub>h h'" assumes "h' \ get_disconnected_nodes ptr \\<^sub>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 \\<^sub>r disc_nodes \ h \ set_disconnected_nodes ptr (remove1 node_ptr disc_nodes) \\<^sub>h h' \ h' \ get_disconnected_nodes ptr \\<^sub>r disc_nodes' \ node_ptr \ set disc_nodes'" interpretation i_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M?: l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = 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_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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 \\<^sub>r Some parent" shows "h \ get_ancestors (cast child) \\<^sub>r (cast child) # parent # ancestors \ h \ get_ancestors parent \\<^sub>r parent # ancestors" proof assume a1: "h \ get_ancestors (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child # parent # ancestors" then have "h \ Heap_Error_Monad.bind (check_in_heap (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child # ancestors)))) \\<^sub>r cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child # parent # ancestors" by(simp add: get_ancestors_def) then show "h \ get_ancestors parent \\<^sub>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 \\<^sub>r parent # ancestors" then show "h \ get_ancestors (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>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 "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>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 \\<^sub>r ancestors" and "ancestor \ set ancestors" and "h \ get_ancestors ancestor \\<^sub>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 \\<^sub>r ()" using returns_result_select_result by force *) show ?case proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>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_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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 cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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))) \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r ancestors" and "cast child \ set ancestors" and "h \ get_parent child \\<^sub>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) \\<^sub>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 \\<^sub>r ancestors" and type_wf: "type_wf h" and known_ptrs: "known_ptrs h" obtains children ancestor_child where "h \ get_child_nodes ancestor \\<^sub>r children" and "ancestor_child \ set children" and "cast ancestor_child \ set ancestors" proof - assume 0: "(\children ancestor_child. h \ get_child_nodes ancestor \\<^sub>r children \ ancestor_child \ set children \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_child \ set ancestors \ thesis)" have "\child. h \ get_parent child \\<^sub>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 "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>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 cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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))) \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r Some ancestor" and in_ancestors: "cast child \ set ancestors" by auto then obtain children where children: "h \ get_child_nodes ancestor \\<^sub>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 \\<^sub>r ancestors" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "(ptr, child) \ (parent_child_rel h)\<^sup>* \ ptr \ set ancestors" proof (safe) assume 3: "(ptr, child) \ (parent_child_rel h)\<^sup>*" 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)\<^sup>*" using converse_rtranclE[OF 1(2)] \ptr \ child\ by metis then obtain ptr_child_node where ptr_child_ptr_child_node: "ptr_child = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>r children" and ptr_child_node: "ptr_child_node \ set children" proof - assume a1: "\children. \h \ get_child_nodes ptr \\<^sub>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|\<^sub>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 "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node, child) \ (parent_child_rel h)\<^sup>*" using ptr_child ptr_child_ptr_child_node by auto ultimately have "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node \ set ancestors" using 1 by auto moreover have "h \ get_parent ptr_child_node \\<^sub>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)\<^sup>*" 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 \\<^sub>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 "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node, child) \ (parent_child_rel h)\<^sup>*" 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 \\<^sub>r root" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "(root, child) \ (parent_child_rel h)\<^sup>*" 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_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'" and known_ptrs: "known_ptrs h" and known_ptrs': "known_ptrs h'" and "h \ get_ancestors ptr \\<^sub>r ancestors" and type_wf: "type_wf h" and type_wf': "type_wf h'" shows "h' \ get_ancestors ptr \\<^sub>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 \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq: "|h \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>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 \\<^sub>r ancestors'" by auto obtain root where root: "h \ get_root_node ptr \\<^sub>r root" proof - assume 0: "(\root. h \ get_root_node ptr \\<^sub>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 \\<^sub>r children = h' \ get_child_nodes p \\<^sub>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. cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c \ set ancestors \ set ancestors' \ h \ get_parent c \\<^sub>r parent_opt = h' \ get_parent c \\<^sub>r parent_opt" proof - fix c parent_opt assume 1: " cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c \ set ancestors \ set ancestors'" obtain ptrs where ptrs: "h \ object_ptr_kinds_M \\<^sub>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 \\<^sub>r True \ h' \ ?P p \\<^sub>r True" proof - fix p assume "p \ set ptrs" then show "h \ ?P p \\<^sub>r True \ h' \ ?P p \\<^sub>r True" proof (cases "p = ptr") case True have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c, ptr) \ (parent_child_rel h)\<^sup>*" using get_ancestors_parent_child_rel 1 assms by blast then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \ (parent_child_rel h)\<^sup>*" using \acyclic (parent_child_rel h)\ False rtrancl_eq_or_trancl rtrancl_trancl_trancl \(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c, ptr) \ (parent_child_rel h)\<^sup>*\ by (metis acyclic_def) then show ?thesis using r_into_rtrancl by auto qed obtain children where children: "h \ get_child_nodes ptr \\<^sub>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, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>r False" by(auto simp add: True) moreover have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c, ptr) \ (parent_child_rel h')\<^sup>*" using get_ancestors_parent_child_rel assms(2) ancestors' 1 known_ptrs' type_wf type_wf' by blast then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \ (parent_child_rel h')\<^sup>*" using \acyclic (parent_child_rel h')\ False rtrancl_eq_or_trancl rtrancl_trancl_trancl \(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c, ptr) \ (parent_child_rel h')\<^sup>*\ by (metis acyclic_def) then show ?thesis using r_into_rtrancl by auto qed then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \ (parent_child_rel h')" using r_into_rtrancl by auto obtain children' where children': "h' \ get_child_nodes ptr \\<^sub>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, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>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)) \\<^sub>r False = h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False" proof fix pa assume "pa \ set ptrs" and "h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>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)) \\<^sub>r False \ h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False" by (metis (mono_tags, lifting) \\pa. pa \ set ptrs \ h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r True = h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>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)) \\<^sub>r False" using \h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False\ by auto next fix pa assume "pa \ set ptrs" and "h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>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)) \\<^sub>r False \ h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False" by (metis (mono_tags, lifting) \\pa. pa \ set ptrs \ h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r True = h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>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)) \\<^sub>r False" using \h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False\ by blast qed have filter_eq: "\xs. h \ filter_M ?P ptrs \\<^sub>r xs = h' \ filter_M ?P ptrs \\<^sub>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)) \\<^sub>r b = h' \ Heap_Error_Monad.bind (get_child_nodes x) (\children. return (c \ set children)) \\<^sub>r b" apply(induct b) using children_eq_True apply blast using children_eq_False apply blast done qed show "h \ get_parent c \\<^sub>r parent_opt = h' \ get_parent c \\<^sub>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 \\<^sub>r ancestors" and "h' \ get_ancestors ptr \\<^sub>r ancestors'" and "\p children children'. h \ get_child_nodes p \\<^sub>r children \ h' \ get_child_nodes p \\<^sub>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 \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>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|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>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 \\<^sub>r Some parent \ h \ get_parent p \\<^sub>r Some parent" proof - fix p parent assume "h' \ get_parent p \\<^sub>r Some parent" then obtain children' where children': "h' \ get_child_nodes parent \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r ancestors" assumes "ptr' \ set ancestors" assumes "h \ get_ancestors ptr' \\<^sub>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 \\<^sub>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 \\<^sub>r Some parent" using assms(1) assms(2) assms(3) child_parent_dual by blast then have "h \ get_ancestors (cast ancestor_child) \\<^sub>r cast ancestor_child # ancestors'" apply(simp add: get_ancestors_def) using \h \ get_ancestors parent \\<^sub>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 \\<^sub>r children\ \ancestor_child \ set children\ \cast ancestor_child \ set ancestors\ \h \ get_ancestors (cast ancestor_child) \\<^sub>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 \\<^sub>r ancestors" assumes "ptr' \ set ancestors" assumes "ptr'' \ set ancestors" shows "h \ get_root_node ptr' \\<^sub>r root_ptr \ h \ get_root_node ptr'' \\<^sub>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' \\<^sub>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'' \\<^sub>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 \\<^sub>r Some ptr" shows "h \ get_root_node (cast child) \\<^sub>r root \ h \ get_root_node ptr \\<^sub>r root" proof assume 1: " h \ get_root_node (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r root" show "h \ get_root_node ptr \\<^sub>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 \\<^sub>r root" show "h \ get_root_node (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>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 \\<^sub>r cast child" shows "h \ get_parent child \\<^sub>r None" proof (insert assms(1) assms(4), induct ptr rule: heap_wellformed_induct_rev) case (step c) then show ?case proof (cases "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>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 \\<^sub>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 \\<^sub>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' \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r ancestors \ h' \ get_ancestors ptr \\<^sub>r ancestors' \ (\p children children'. h \ get_child_nodes p \\<^sub>r children \ h' \ get_child_nodes p \\<^sub>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 \\<^sub>r ancestors \ cast child_node \ set ancestors \ h \ get_parent child_node \\<^sub>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 \\<^sub>r ancestors \ type_wf h \ known_ptrs h \ (\children ancestor_child . h \ get_child_nodes ancestor \\<^sub>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 \\<^sub>r ancestors \ known_ptrs h \ type_wf h \ (ptr, child) \ (parent_child_rel h)\<^sup>* \ 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 \\<^sub>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 \\<^sub>r ancestors \ ptr' \ set ancestors \ ptr'' \ set ancestors \ h \ get_root_node ptr' \\<^sub>r root_ptr \ h \ get_root_node ptr'' \\<^sub>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 \\<^sub>r cast child \ h \ get_parent child \\<^sub>r None" assumes get_root_node_parent_same: "h \ get_parent child \\<^sub>r Some ptr \ h \ get_root_node (cast child) \\<^sub>r root \ h \ get_root_node ptr \\<^sub>r root" interpretation i_get_root_node_wf?: l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent + l_get_parent_wf + l_heap_is_wellformed (* l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M *) 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 \\<^sub>r nodes" and "node \ set nodes" and "h \ get_child_nodes ptr \\<^sub>r children" and "node \ ptr" obtains child child_to where "child \ set children" and "h \ to_tree_order (cast child) \\<^sub>r child_to" and "node \ set child_to" proof - obtain treeorders where treeorders: "h \ map_M to_tree_order (map cast children) \\<^sub>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 (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r nodes" and "h \ get_child_nodes ptr \\<^sub>r children" and "node \ set children" and "h \ to_tree_order (cast node) \\<^sub>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) \\<^sub>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 \\<^sub>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 \\<^sub>r nodes" and "node \ set nodes" and "h \ to_tree_order node \\<^sub>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 \\<^sub>r nodes \ (\node. node \ set nodes \ (\nodes'. h \ to_tree_order node \\<^sub>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 \\<^sub>r children \ child \ set children \ \nodes. h \ to_tree_order (cast child) \\<^sub>r nodes \ (\node. node \ set nodes \ (\nodes'. h \ to_tree_order node \\<^sub>r nodes' \ set nodes' \ set nodes)))" and 2: "h \ to_tree_order parent \\<^sub>r nodes" and 3: "node \ set nodes" and "h \ to_tree_order node \\<^sub>r nodes'" and "x \ set nodes'" have h1: "(\children child nodes node nodes'. h \ get_child_nodes parent \\<^sub>r children \ child \ set children \ h \ to_tree_order (cast child) \\<^sub>r nodes \ (node \ set nodes \ (h \ to_tree_order node \\<^sub>r nodes' \ set nodes' \ set nodes)))" using 1 by blast obtain children where children: "h \ get_child_nodes parent \\<^sub>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 \\<^sub>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 \\<^sub>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) \\<^sub>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 \\<^sub>r nodes'\ children by blast moreover have "set nodes_of_child \ set nodes" using "2" \child \ set children\ \h \ to_tree_order (cast child) \\<^sub>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 \\<^sub>r nodes" assumes "h \ get_parent child \\<^sub>r Some parent" assumes "parent \ set nodes" shows "cast child \ set nodes" proof - obtain nodes' where nodes': "h \ to_tree_order parent \\<^sub>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 \\<^sub>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 (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>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 \\<^sub>r nodes" assumes "h \ get_child_nodes parent \\<^sub>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 \\<^sub>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 \\<^sub>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 (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \\<^sub>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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.child_parent_dual l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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 \cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>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 \\<^sub>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 \\<^sub>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 (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \\<^sub>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 \\<^sub>r nodes" assumes "cast child \ ptr" assumes "cast child \ set nodes" obtains parent where "h \ get_parent child \\<^sub>r Some parent" and "parent \ set nodes" proof - assume 1: "(\parent. h \ get_parent child \\<^sub>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 \\<^sub>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 \\<^sub>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 (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \\<^sub>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 \\<^sub>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 \cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child \ set child_to\ assms(1) assms(2) assms(3) cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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 \\<^sub>r to" shows "(ptr, child) \ (parent_child_rel h)\<^sup>* \ child \ set to" proof assume 3: "(ptr, child) \ (parent_child_rel h)\<^sup>*" 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)\<^sup>*" and "(child_parent, child) \ (parent_child_rel h)" using \ptr \ child\ by (metis "1.prems" rtranclE) obtain child_node where child_node: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>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)\<^sup>*\ 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)\<^sup>*" 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: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>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)\<^sup>*" 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) done declare l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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 \\<^sub>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 \\<^sub>r to \ (ptr, child_ptr) \ (parent_child_rel h)\<^sup>* \ child_ptr \ set to" assumes to_tree_order_child2: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ to_tree_order ptr \\<^sub>r nodes \ cast child \ ptr \ cast child \ set nodes \ (\parent. h \ get_parent child \\<^sub>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 \\<^sub>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 \\<^sub>r nodes \ h \ get_child_nodes parent \\<^sub>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 \\<^sub>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 \\<^sub>r nodes \ h \ get_parent child \\<^sub>r Some parent \ parent \ set nodes \ cast child \ set nodes" assumes to_tree_order_subset: "heap_is_wellformed h \ h \ to_tree_order ptr \\<^sub>r nodes \ node \ set nodes \ h \ to_tree_order node \\<^sub>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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + 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 \\<^sub>r to" assumes "ptr' \ set to" assumes "h \ get_root_node ptr' \\<^sub>r root_ptr" assumes "ptr'' \ set to" shows "h \ get_root_node ptr'' \\<^sub>r root_ptr" proof - obtain ancestors' where ancestors': "h \ get_ancestors ptr' \\<^sub>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' \\<^sub>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 \\<^sub>r root_ptr" using \h \ get_root_node ptr' \\<^sub>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'' \\<^sub>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'' \\<^sub>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 \\<^sub>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 \\<^sub>r root_ptr" assumes "h \ to_tree_order root_ptr \\<^sub>r to" assumes "ptr' \ set to" shows "h \ get_root_node ptr' \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r root_ptr" using assms(4) using \cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>r Some parent\ \cast child_node = child\ by (metis False assms(1) assms(2) assms(3) get_parent_child_dual) then show ?thesis using \cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node = child\ \h \ get_parent child_node \\<^sub>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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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 \\<^sub>r to \ ptr' \ set to \ h \ get_root_node ptr' \\<^sub>r root_ptr \ ptr'' \ set to \ h \ get_root_node ptr'' \\<^sub>r root_ptr" assumes to_tree_order_same_root: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_root_node ptr \\<^sub>r root_ptr \ h \ to_tree_order root_ptr \\<^sub>r to \ ptr' \ set to \ h \ get_root_node ptr' \\<^sub>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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_known_ptrs + l_heap_is_wellformed + l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_ancestors + l_get_ancestors_wf + l_get_parent + l_get_parent_wf + l_get_root_node_wf + l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_owner_document_disconnected_nodes: assumes "heap_is_wellformed h" assumes "h \ get_disconnected_nodes document_ptr \\<^sub>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) \\<^sub>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_ptr\set |h \ document_ptr_kinds_M|\<^sub>r. node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>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 \\<^sub>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) \\<^sub>r cast node_ptr" using 2 get_root_node_no_parent by blast obtain document_ptrs where document_ptrs: "h \ document_ptr_kinds_M \\<^sub>r document_ptrs" by simp then have "h \ ok (filter_M (\document_ptr. do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) \ cast ` set disconnected_nodes) }) document_ptrs \\<^sub>r candidates" by auto have eq: "\document_ptr. document_ptr |\| document_ptr_kinds h \ node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r \ |h \ do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) \ cast ` set disconnected_nodes) }|\<^sub>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 (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr \ cast ` set disconnected_nodes) }|\<^sub>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 (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) \ cast ` set disconnected_nodes) }) document_ptrs \\<^sub>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_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r document_ptr" by(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: option.splits)[1] moreover have "known_ptr (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>r None" and "h \ get_owner_document (cast node_ptr) \\<^sub>r owner_document" and "h \ get_disconnected_nodes owner_document \\<^sub>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) \\<^sub>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|\<^sub>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_ptr\set |h \ document_ptr_kinds_M|\<^sub>r. node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r" by (metis DocumentMonad.ptr_kinds_ptr_kinds_M assms(1) finite_set_in heap_is_wellformed_children_disc_nodes) then obtain document_ptr where document_ptr: "document_ptr\set |h \ document_ptr_kinds_M|\<^sub>r" and node_ptr_in_disc_nodes: "node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>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 \\<^sub>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 () \\<^sub>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_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then show "owner_document |\| document_ptr_kinds h" by(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 4: "is_character_data_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 5: "h \ Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then obtain root where root: "h \ get_root_node ptr \\<^sub>r root" by(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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|\<^sub>r)" by (metis (no_types, 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 notin_fset 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|\<^sub>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 \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ` set disconnected_nodes))) (sorted_list_of_set (fset (document_ptr_kinds h))) \\<^sub>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 notin_fset 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|\<^sub>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 \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>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_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 4: "h \ Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then obtain root where root: "h \ get_root_node ptr \\<^sub>r root" by(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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|\<^sub>r)" by (metis (no_types, 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 notin_fset 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|\<^sub>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 \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ` set disconnected_nodes))) (sorted_list_of_set (fset (document_ptr_kinds h))) \\<^sub>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 notin_fset 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|\<^sub>r\ apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1] using \some_owner_document |\| document_ptr_kinds h\ \root \ cast ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r\ apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1] using \some_owner_document |\| document_ptr_kinds h\ \root \ cast ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>r\ apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1] by (simp add: "1" local.get_disconnected_nodes_ok) then have "candidates \ []" by auto then have "owner_document \ set candidates" using 4 root 3 apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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 \\<^sub>r children" assumes "child \ set children" shows "h \ get_owner_document ptr \\<^sub>r owner_document \ h \ get_owner_document (cast child) \\<^sub>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 \\<^sub>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) \\<^sub>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 \\<^sub>r owner_document \ h \ a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \\<^sub>r owner_document" proof (cases "is_document_ptr ptr") case True then obtain document_ptr where document_ptr: "cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr () \\<^sub>r owner_document \ h \ a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \\<^sub>r owner_document" using document_ptr \h \ get_root_node (cast child) \\<^sub>r root\[simplified \root = cast document_ptr\ document_ptr] apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 dest!: bind_returns_result_E3[rotated, OF \h \ get_root_node (cast child) \\<^sub>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: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document \ h \ a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \\<^sub>r owner_document" using root \h \ get_root_node (cast child) \\<^sub>r root\ unfolding a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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 "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 + assumes get_owner_document_disconnected_nodes: "heap_is_wellformed h \ known_ptrs h \ type_wf h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ node_ptr \ set disc_nodes \ h \ get_owner_document (cast node_ptr) \\<^sub>r document_ptr" assumes in_disconnected_nodes_no_parent: "heap_is_wellformed h \ h \ get_parent node_ptr \\<^sub>r None\ h \ get_owner_document (cast node_ptr) \\<^sub>r owner_document \ h \ get_disconnected_nodes owner_document \\<^sub>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 \\<^sub>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)" interpretation i_get_owner_document_wf?: l_get_owner_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_owner_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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" 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 done subsubsection \get\_root\_node\ locale l_get_owner_document_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_root_node_wf + l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + 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 \\<^sub>r root" assumes "is_document_ptr_kind root" shows "h \ get_owner_document ptr \\<^sub>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_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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 \\<^sub>r root\ assms(5) by(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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 \\<^sub>r root" shows "h \ get_owner_document ptr \\<^sub>r owner_document \ h \ get_owner_document root \\<^sub>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 = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr" by (metis node_ptr_casts_commute3) show ?thesis proof assume "h \ get_owner_document ptr \\<^sub>r owner_document" then have "h \ local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>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 \\<^sub>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 \h \ get_owner_document ptr \\<^sub>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_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root\ apply blast using \root |\| object_ptr_kinds h\ by(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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 = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr" by (metis node_ptr_casts_commute3) then have "h \ local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr () \\<^sub>r owner_document" using \h \ local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document\ assms(4) apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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 \\<^sub>r owner_document" show "h \ get_owner_document ptr \\<^sub>r owner_document" proof (cases "is_document_ptr_kind root") case True have "root = cast owner_document" using \h \ get_owner_document root \\<^sub>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_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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 = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr" by (metis node_ptr_casts_commute3) then have "h \ local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr () \\<^sub>r owner_document" using \h \ get_owner_document root \\<^sub>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_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document" apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_owner_document_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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 \\<^sub>r root \ is_document_ptr_kind root \ h \ get_owner_document ptr \\<^sub>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 \\<^sub>r root \ h \ get_owner_document ptr \\<^sub>r owner_document \ h \ get_owner_document root \\<^sub>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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_attribute\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + 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 \\<^sub>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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + 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 \\<^sub>h h2" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "h2 \ get_parent child \\<^sub>r None" proof - obtain children where children: "h \ get_child_nodes ptr \\<^sub>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 \\<^sub>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) \\<^sub>r owner_document" and disc_nodes: "h \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" and h': "h \ set_disconnected_nodes owner_document (child # disc_nodes) \\<^sub>h h'" and h2: "h' \ set_child_nodes ptr (remove1 child children) \\<^sub>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|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r other_children" have "h \ get_child_nodes other_ptr \\<^sub>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 \\<^sub>r other_children\ a1 h1 by blast qed then have "\other_ptr other_children. other_ptr \ ptr \ h2 \ get_child_nodes other_ptr \\<^sub>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 \\<^sub>r other_children" have "h' \ get_child_nodes other_ptr \\<^sub>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 \\<^sub>r other_children\ \ child \ set other_children\ a1 by blast qed ultimately have ha: "\other_ptr other_children. h2 \ get_child_nodes other_ptr \\<^sub>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 \\<^sub>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 \\<^sub>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 "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child |\| object_ptr_kinds h" using get_owner_document_ptr_in_heap owner_document by blast then show "h2 \ check_in_heap (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>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 \\<^sub>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))) \\<^sub>r []" by(auto intro!: filter_M_empty_I bind_pure_I) qed qed end locale l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_remove_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma remove_child_parent_child_rel_subset: assumes "heap_is_wellformed h" and "h \ remove_child ptr child \\<^sub>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 (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r owner_document" and children_h: "h \ get_child_nodes ptr \\<^sub>r children_h" and child_in_children_h: "child \ set children_h" and disconnected_nodes_h: "h \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h" and h2: "h \ set_disconnected_nodes owner_document (child # disconnected_nodes_h) \\<^sub>h h2" and h': "h2 \ set_child_nodes ptr (remove1 child children_h) \\<^sub>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 \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" unfolding object_ptr_kinds_M_defs by simp then have object_ptr_kinds_eq2: "|h \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" using select_result_eq by force then have node_ptr_kinds_eq2: "|h \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>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|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>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' \\<^sub>r children =h' \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>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 \\<^sub>r disconnected_nodes = h' \ get_disconnected_nodes document_ptr \\<^sub>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|\<^sub>r = |h' \ get_disconnected_nodes document_ptr|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes ptr \\<^sub>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 \\<^sub>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 \\<^sub>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 (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r owner_document" and children_h: "h \ get_child_nodes ptr \\<^sub>r children_h" and child_in_children_h: "child \ set children_h" and disconnected_nodes_h: "h \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h" and h2: "h \ set_disconnected_nodes owner_document (child # disconnected_nodes_h) \\<^sub>h h2" and h': "h2 \ set_child_nodes ptr (remove1 child children_h) \\<^sub>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 \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" unfolding object_ptr_kinds_M_defs by simp then have object_ptr_kinds_eq2: "|h \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" using select_result_eq by force then have node_ptr_kinds_eq2: "|h \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>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|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>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' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>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 \\<^sub>r disconnected_nodes = h' \ get_disconnected_nodes document_ptr \\<^sub>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|\<^sub>r = |h' \ get_disconnected_nodes document_ptr|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes ptr \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>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, lifting) \type_wf h'\ assms(2) assms(3) local.get_child_nodes_ok local.known_ptrs_known_ptr local.remove_child_children_subset notin_fset object_ptr_kinds_eq3 returns_result_select_result subset_code(1) type_wf) apply (metis (no_types, lifting) assms(2) disconnected_nodes_eq2 disconnected_nodes_h disconnected_nodes_h' document_ptr_kinds_eq3 finite_set_in local.remove_child_child_in_heap node_ptr_kinds_eq3 select_result_I2 set_ConsD subset_code(1)) done moreover have "a_owner_document_valid h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_owner_document_valid h'" apply(auto simp add: a_owner_document_valid_def object_ptr_kinds_eq3 document_ptr_kinds_eq3 node_ptr_kinds_eq3)[1] proof - fix node_ptr assume 0: "\node_ptr\fset (node_ptr_kinds h'). (\document_ptr. document_ptr |\| document_ptr_kinds h' \ node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r) \ (\parent_ptr. parent_ptr |\| object_ptr_kinds h' \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>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|\<^sub>r" then show "\document_ptr. document_ptr |\| document_ptr_kinds h' \ node_ptr \ set |h' \ get_disconnected_nodes document_ptr|\<^sub>r" proof (cases "node_ptr = child") case True show ?thesis apply(rule exI[where x=owner_document]) using children_eq2 disconnected_nodes_eq2 children_h children_h' disconnected_nodes_h' True by (metis (no_types, lifting) get_disconnected_nodes_ptr_in_heap is_OK_returns_result_I list.set_intros(1) select_result_I2) next case False then show ?thesis using 0 1 2 children_eq2 children_h children_h' disconnected_nodes_eq2 disconnected_nodes_h disconnected_nodes_h' apply(auto simp add: children_eq2 disconnected_nodes_eq2 dest!: select_result_I2)[1] by (metis children_eq2 disconnected_nodes_eq2 finite_set_in in_set_remove1 list.set_intros(2)) qed qed moreover { have h0: "a_distinct_lists h" using assms(1) by (simp add: heap_is_wellformed_def) moreover have ha1: "(\x\set |h \ object_ptr_kinds_M|\<^sub>r. set |h \ get_child_nodes x|\<^sub>r) \ (\x\set |h \ document_ptr_kinds_M|\<^sub>r. set |h \ get_disconnected_nodes x|\<^sub>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|\<^sub>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|\<^sub>r" using ha1 ha2 ha3 apply(simp) using IntI by fastforce moreover have "distinct |h \ object_ptr_kinds_M|\<^sub>r" apply(rule select_result_I) by(auto simp add: object_ptr_kinds_M_defs) moreover have "distinct |h \ document_ptr_kinds_M|\<^sub>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|\<^sub>r" assume 1: "a_distinct_lists h" and 3: "distinct |h \ object_ptr_kinds_M|\<^sub>r" have 4: "distinct (concat ((map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) |h \ object_ptr_kinds_M|\<^sub>r)))" using 1 by(auto simp add: a_distinct_lists_def) show "distinct (concat (map (\ptr. |h' \ get_child_nodes ptr|\<^sub>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|\<^sub>r" using 4 distinct_concat_map_E object_ptr_kinds_eq2 by fastforce obtain children where children: "h \ get_child_nodes x \\<^sub>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 \\<^sub>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|\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r children_x'" using children_eq children_h' children_x by fastforce obtain children_y' where children_y': "h' \ get_child_nodes y \\<^sub>r children_y'" using children_eq children_h' children_y by fastforce have "distinct (concat (map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) |h \ object_ptr_kinds_M|\<^sub>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|\<^sub>r \ set |h' \ get_child_nodes y|\<^sub>r = {}" using children_x' children_y' by (metis (no_types, lifting) select_result_I2) qed next assume 2: "distinct |h \ document_ptr_kinds_M|\<^sub>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|\<^sub>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|\<^sub>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|\<^sub>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|\<^sub>r" proof (cases "x = owner_document") case True have "child \ set |h \ get_disconnected_nodes x|\<^sub>r" using child_not_in document_ptr_kinds_eq2 "4" by fastforce moreover have "|h' \ get_disconnected_nodes x|\<^sub>r = child # |h \ get_disconnected_nodes x|\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>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|\<^sub>r) |h \ document_ptr_kinds_M|\<^sub>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|\<^sub>r \ set |h' \ get_disconnected_nodes y|\<^sub>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|\<^sub>r" and 3: "xb \ fset (document_ptr_kinds h')" and 4: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" obtain disc_nodes where disc_nodes: "h \ get_disconnected_nodes xb \\<^sub>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 \\<^sub>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 \\<^sub>r children" by (metis "1" type_wf assms(3) finite_set_in get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr object_ptr_kinds_eq3) obtain children' where children': "h' \ get_child_nodes xa \\<^sub>r children'" using children children_eq children_h' by fastforce have "\x. x \ set |h \ get_child_nodes xa|\<^sub>r \ x \ set |h \ get_disconnected_nodes xb|\<^sub>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|\<^sub>r = children'" using children' by simp have 7: "|h' \ get_disconnected_nodes xb|\<^sub>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|\<^sub>r = remove1 child children_h" using children_h' - bysimp + 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|\<^sub>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 \\<^sub>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 \\<^sub>h h'" and children: "h' \ get_child_nodes ptr \\<^sub>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 (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r owner_document" and children_h: "h \ get_child_nodes ptr' \\<^sub>r children_h" and child_in_children_h: "child \ set children_h" and disconnected_nodes_h: "h \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h" and h2: "h \ set_disconnected_nodes owner_document (child # disconnected_nodes_h) \\<^sub>h h2" and h': "h2 \ set_child_nodes ptr' (remove1 child children_h) \\<^sub>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 \\<^sub>r node_ptr # children" assumes "h \ remove_child ptr node_ptr \\<^sub>h h'" shows "h' \ get_child_nodes ptr \\<^sub>r children" proof - obtain h2 disc_nodes owner_document where "h \ get_owner_document (cast node_ptr) \\<^sub>r owner_document" and "h \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" and h2: "h \ set_disconnected_nodes owner_document (node_ptr # disc_nodes) \\<^sub>h h2" and "h2 \ set_child_nodes ptr children \\<^sub>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 \\<^sub>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_nodes\h2 \ set_child_nodes ptr children \\<^sub>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 \\<^sub>r node_ptr # children" assumes "h \ remove node_ptr \\<^sub>h h'" shows "h' \ get_child_nodes ptr \\<^sub>r children" proof - have "h \ get_parent node_ptr \\<^sub>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 \\<^sub>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 \\<^sub>r children" assumes "h \ forall_M remove children \\<^sub>h h'" shows "h' \ get_child_nodes ptr \\<^sub>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 \\<^sub>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 \\<^sub>r children \ h \ forall_M remove children \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r [])" and 1: "heap_is_wellformed h" and 2: "type_wf h" and 3: "known_ptrs h" and 4: "h \ get_child_nodes ptr \\<^sub>r a # children" and 5: "h \ get_parent a \\<^sub>r Some ptr" and 7: "h \ remove a \\<^sub>h h2" and 8: "h2 \ forall_M remove children \\<^sub>h h'" then have "h2 \ get_child_nodes ptr \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>h h' \ heap_is_wellformed h'" assumes remove_preserves_type_wf: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ remove child \\<^sub>h h' \ type_wf h'" assumes remove_preserves_known_ptrs: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ remove child \\<^sub>h h' \ known_ptrs h'" assumes remove_heap_is_wellformed_preserved: "type_wf h \ known_ptrs h \ heap_is_wellformed h \ h \ remove child \\<^sub>h h' \ heap_is_wellformed h'" assumes remove_child_removes_child: "heap_is_wellformed h \ h \ remove_child ptr' child \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>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 \\<^sub>r node_ptr # children \ h \ remove_child ptr node_ptr \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r children" assumes remove_removes_child: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r node_ptr # children \ h \ remove node_ptr \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r children" assumes remove_for_all_empty_children: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ forall_M remove children \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r []" interpretation i_remove_child_wf2?: l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + 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 \\<^sub>h h'" assumes "h \ get_child_nodes ptr' \\<^sub>r node # children" shows "h' \ get_child_nodes ptr' \\<^sub>r children" proof - obtain old_document parent_opt h2 where old_document: "h \ get_owner_document (cast node) \\<^sub>r old_document" and parent_opt: "h \ get_parent node \\<^sub>r parent_opt" and h2: "h \ (case parent_opt of Some parent \ do { remove_child parent node } | None \ do { return ()}) \\<^sub>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 () }) \\<^sub>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' \\<^sub>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) \\<^sub>r old_document" and parent_opt: "h \ get_parent node \\<^sub>r parent_opt" and h2: "h \ (case parent_opt of Some parent \ do { remove_child parent node } | None \ do { return ()}) \\<^sub>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 () }) \\<^sub>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 \\<^sub>r old_disc_nodes" and h3: "h2 \ set_disconnected_nodes old_document (remove1 node old_disc_nodes) \\<^sub>h h3" and old_disc_nodes: "h3 \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" and h': "h3 \ set_disconnected_nodes owner_document (node # disc_nodes) \\<^sub>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_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_root_node + l_get_owner_document_wf + l_remove_child_wf2 + l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma adopt_node_removes_child_step: assumes wellformed: "heap_is_wellformed h" and adopt_node: "h \ adopt_node owner_document node_ptr \\<^sub>h h2" and children: "h2 \ get_child_nodes ptr \\<^sub>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) \\<^sub>r old_document" and parent_opt: "h \ get_parent node_ptr \\<^sub>r parent_opt" and h': "h \ (case parent_opt of Some parent \ remove_child parent node_ptr | None \ return () ) \\<^sub>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 \\<^sub>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 \\<^sub>r children\ returns_result_eq by fastforce next case (Some option) then show ?case using remove_child_removes_child \h' \ get_child_nodes ptr \\<^sub>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 \\<^sub>h h'" shows "\ptr' children'. h' \ get_child_nodes ptr' \\<^sub>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 \\<^sub>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) \\<^sub>r old_document" and parent_opt: "h \ get_parent child \\<^sub>r parent_opt" and h2: "h \ (case parent_opt of Some parent \ remove_child parent child | None \ return ()) \\<^sub>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 () }) \\<^sub>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 \\<^sub>r ptrs = h2 \ object_ptr_kinds_M \\<^sub>r ptrs" unfolding object_ptr_kinds_M_defs by simp then have object_ptr_kinds_eq_h: "|h \ object_ptr_kinds_M|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq_h: "|h \ node_ptr_kinds_M|\<^sub>r = |h2 \ node_ptr_kinds_M|\<^sub>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 \\<^sub>r old_disc_nodes" and h3: "h2 \ set_disconnected_nodes old_document (remove1 child old_disc_nodes) \\<^sub>h h3" and disc_nodes_document_ptr_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_document_ptr_h3" and h': "h3 \ set_disconnected_nodes document_ptr (child # disc_nodes_document_ptr_h3) \\<^sub>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 \\<^sub>r ptrs = h3 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq_h2: "|h2 \ object_ptr_kinds_M|\<^sub>r = |h3 \ object_ptr_kinds_M|\<^sub>r" by(simp) then have node_ptr_kinds_eq_h2: "|h2 \ node_ptr_kinds_M|\<^sub>r = |h3 \ node_ptr_kinds_M|\<^sub>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|\<^sub>r = |h3 \ document_ptr_kinds_M|\<^sub>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 \\<^sub>r children = h3 \ get_child_nodes ptr \\<^sub>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|\<^sub>r = |h3 \ get_child_nodes ptr|\<^sub>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 \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq_h3: "|h3 \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by(simp) then have node_ptr_kinds_eq_h3: "|h3 \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>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|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>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 \\<^sub>r children = h' \ get_child_nodes ptr \\<^sub>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|\<^sub>r = |h' \ get_child_nodes ptr|\<^sub>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 \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>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 \\<^sub>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|\<^sub>r) |h2 \ document_ptr_kinds_M|\<^sub>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, 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 finite_set_in 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 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) list.set_intros(2) node_ptr_kinds_eq3_h2 node_ptr_kinds_eq3_h3 object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 select_result_I2) 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|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 2: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h2)))))" and 3: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h2). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" show "distinct (concat (map (\document_ptr. |h' \ get_disconnected_nodes document_ptr|\<^sub>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|\<^sub>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|\<^sub>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|\<^sub>r \ set |h2 \ get_disconnected_nodes y|\<^sub>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|\<^sub>r \ set |h' \ get_disconnected_nodes y|\<^sub>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|\<^sub>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|\<^sub>r \ set disc_nodes_old_document_h2 = {}" using \set |h2 \ get_disconnected_nodes y|\<^sub>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|\<^sub>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|\<^sub>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|\<^sub>r \ set |h2 \ get_disconnected_nodes old_document|\<^sub>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 fmember.rep_eq 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|\<^sub>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|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 1: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h2)))))" and 2: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h2). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" and 3: "xa |\| object_ptr_kinds h'" and 4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 5: "xb |\| document_ptr_kinds h'" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>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 \\<^sub>r disc_nodes_old_document_h2" assume a3: "h3 \ get_disconnected_nodes old_document \\<^sub>r remove1 child disc_nodes_old_document_h2" assume a4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" assume "document_ptr_kinds h2 = document_ptr_kinds h'" assume a5: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h'). set |h2 \ get_disconnected_nodes x|\<^sub>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|\<^sub>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|\<^sub>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|\<^sub>r \ set |h2 \ get_disconnected_nodes xb|\<^sub>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 \\<^sub>r disc_nodes_document_ptr_h3" assume a3: "h2 \ get_disconnected_nodes old_document \\<^sub>r disc_nodes_old_document_h2" assume a4: "xa |\| object_ptr_kinds h'" assume a5: "h' \ get_disconnected_nodes document_ptr \\<^sub>r child # disc_nodes_document_ptr_h3" assume a6: "old_document |\| document_ptr_kinds h'" assume a7: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" assume a8: "x \ set |h' \ get_child_nodes xa|\<^sub>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|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" assume a11: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" assume a12: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h'). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" have f13: "\d. d \ set |h' \ document_ptr_kinds_M|\<^sub>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|\<^sub>r = disc_nodes_old_document_h2" using a6 a3 by simp have "x \ set |h2 \ get_disconnected_nodes xb|\<^sub>r" using a12 a8 a4 \xb |\| document_ptr_kinds h'\ by (meson UN_I disjoint_iff_not_equal fmember.rep_eq) 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 \\<^sub>h h'" and "h' \ get_disconnected_nodes owner_document \\<^sub>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) \\<^sub>r old_document" and parent_opt: "h \ get_parent node_ptr \\<^sub>r parent_opt" and h2: "h \ (case parent_opt of Some parent \ remove_child parent node_ptr | None \ return ()) \\<^sub>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 () }) \\<^sub>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 \\<^sub>r disc_nodes" assume a2: "\h document_ptr disc_nodes h'. h \ set_disconnected_nodes document_ptr disc_nodes \\<^sub>h h' \ h' \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" assume "h'a \ set_disconnected_nodes owner_document (node_ptr # xb) \\<^sub>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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] interpretation i_adopt_node_wf2?: l_adopt_node_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_adopt_node_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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 \\<^sub>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 \\<^sub>h h2 \ h2 \ get_child_nodes ptr \\<^sub>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 \\<^sub>h h' \ h' \ get_disconnected_nodes owner_document \\<^sub>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 \\<^sub>h h' \ h \ get_child_nodes ptr' \\<^sub>r node # children \ h' \ get_child_nodes ptr' \\<^sub>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" assumes adopt_node_preserves_type_wf: "heap_is_wellformed h \ h \ adopt_node document_ptr child \\<^sub>h h' \ known_ptrs h \ type_wf h \ type_wf h'" assumes adopt_node_preserves_known_ptrs: "heap_is_wellformed h \ h \ adopt_node document_ptr child \\<^sub>h h' \ known_ptrs h \ type_wf h \ known_ptrs 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 using adopt_node_preserves_wellformedness apply blast using adopt_node_preserves_wellformedness apply blast done subsection \insert\_before\ locale l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + 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 \\<^sub>h h'" assumes "h \ get_child_nodes ptr' \\<^sub>r node # children" shows "h' \ get_child_nodes ptr' \\<^sub>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) \\<^sub>r reference_child" and "h \ get_owner_document ptr \\<^sub>r owner_document" and h2: "h \ adopt_node owner_document node \\<^sub>h h2" and "h2 \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" and h3: "h2 \ set_disconnected_nodes owner_document (remove1 node disc_nodes) \\<^sub>h h3" and h': "h3 \ a_insert_node ptr node reference_child \\<^sub>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' \\<^sub>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' \\<^sub>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 \\<^sub>h h' \ h \ get_child_nodes ptr' \\<^sub>r node # children \ h' \ get_child_nodes ptr' \\<^sub>r children" interpretation i_insert_before_wf?: l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + 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_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_owner_document_wf begin lemma insert_before_preserves_acyclitity: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ insert_before ptr node child \\<^sub>h h'" shows "acyclic (parent_child_rel h')" proof - obtain ancestors reference_child owner_document h2 h3 disconnected_nodes_h2 where ancestors: "h \ get_ancestors ptr \\<^sub>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) \\<^sub>r reference_child" and owner_document: "h \ get_owner_document ptr \\<^sub>r owner_document" and h2: "h \ adopt_node owner_document node \\<^sub>h h2" and disconnected_nodes_h2: "h2 \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h2" and h3: "h2 \ set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \\<^sub>h h3" and h': "h3 \ a_insert_node ptr node reference_child \\<^sub>h h'" using assms(4) 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 assms 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 assms 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 have "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 \\<^sub>r ptrs = h2 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs ) then have object_ptr_kinds_M_eq2_h: "|h \ object_ptr_kinds_M|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h: "|h \ node_ptr_kinds_M|\<^sub>r = |h2 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have "known_ptrs h2" using assms object_ptr_kinds_M_eq3_h known_ptrs_preserved by blast have wellformed_h2: "heap_is_wellformed h2" using adopt_node_preserves_wellformedness[OF assms(1) h2] assms by simp 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 \\<^sub>r ptrs = h3 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_M_eq2_h2: "|h2 \ object_ptr_kinds_M|\<^sub>r = |h3 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h2: "|h2 \ node_ptr_kinds_M|\<^sub>r = |h3 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have document_ptr_kinds_eq2_h2: "|h2 \ document_ptr_kinds_M|\<^sub>r = |h3 \ document_ptr_kinds_M|\<^sub>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 \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_M_eq2_h3: "|h3 \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h3: "|h3 \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have document_ptr_kinds_eq2_h3: "|h3 \ document_ptr_kinds_M|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_M_eq2_h3 document_ptr_kinds_M_eq by auto have "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 \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disconnected_nodes_h3: "h3 \ get_disconnected_nodes owner_document \\<^sub>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 \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have children_eq_h3: "\ptr' children. ptr \ ptr' \ h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force obtain children_h3 where children_h3: "h3 \ get_child_nodes ptr \\<^sub>r children_h3" using h' a_insert_node_def by auto have children_h': "h' \ get_child_nodes ptr \\<^sub>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 \\<^sub>r children \ node \ set children" using assms h2 adopt_node_removes_child 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|\<^sub>r" proof - fix d assume "d |\| document_ptr_kinds h3" show "node \ set |h3 \ get_disconnected_nodes d|\<^sub>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|\<^sub>r \ set |h2 \ get_disconnected_nodes owner_document|\<^sub>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 \\<^sub>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 \\<^sub>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 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)\<^sup>*}" using adopt_node_removes_child using ancestors node_not_in_ancestors using \known_ptrs h2\ \type_wf h2\ ancestors_h2 local.get_ancestors_parent_child_rel node_not_in_ancestors_h2 wellformed_h2 by blast then have "cast node \ {x. (x, ptr) \ (parent_child_rel h3)\<^sup>*}" 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 "acyclic (parent_child_rel h')" by (auto simp add: heap_is_wellformed_def) qed lemma insert_before_heap_is_wellformed_preserved: assumes wellformed: "heap_is_wellformed h" and insert_before: "h \ insert_before ptr node child \\<^sub>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 \\<^sub>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) \\<^sub>r reference_child" and owner_document: "h \ get_owner_document ptr \\<^sub>r owner_document" and h2: "h \ adopt_node owner_document node \\<^sub>h h2" and disconnected_nodes_h2: "h2 \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h2" and h3: "h2 \ set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \\<^sub>h h3" and h': "h3 \ a_insert_node ptr node reference_child \\<^sub>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 \\<^sub>r ptrs = h2 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs ) then have object_ptr_kinds_M_eq2_h: "|h \ object_ptr_kinds_M|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h: "|h \ node_ptr_kinds_M|\<^sub>r = |h2 \ node_ptr_kinds_M|\<^sub>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 \\<^sub>r ptrs = h3 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_M_eq2_h2: "|h2 \ object_ptr_kinds_M|\<^sub>r = |h3 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h2: "|h2 \ node_ptr_kinds_M|\<^sub>r = |h3 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have document_ptr_kinds_eq2_h2: "|h2 \ document_ptr_kinds_M|\<^sub>r = |h3 \ document_ptr_kinds_M|\<^sub>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 \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_M_eq2_h3: "|h3 \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h3: "|h3 \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have document_ptr_kinds_eq2_h3: "|h3 \ document_ptr_kinds_M|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>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 \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disconnected_nodes_h3: "h3 \ get_disconnected_nodes owner_document \\<^sub>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 \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have children_eq_h3: "\ptr' children. ptr \ ptr' \ h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force obtain children_h3 where children_h3: "h3 \ get_child_nodes ptr \\<^sub>r children_h3" using h' a_insert_node_def by auto have children_h': "h' \ get_child_nodes ptr \\<^sub>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 \\<^sub>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|\<^sub>r" proof - fix d assume "d |\| document_ptr_kinds h3" show "node \ set |h3 \ get_disconnected_nodes d|\<^sub>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|\<^sub>r \ set |h2 \ get_disconnected_nodes owner_document|\<^sub>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 \\<^sub>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 \\<^sub>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)\<^sup>*}" 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)\<^sup>*}" 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, lifting) disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 document_ptr_kinds_commutes finite_set_in 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|\<^sub>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|\<^sub>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 finite_set_in 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 finite_set_in select_result_I2 subsetD) by (metis (no_types) \type_wf h'\ disconnected_nodes_eq2_h3 disconnected_nodes_eq_h3 finite_set_in 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|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" show "distinct |h3 \ get_disconnected_nodes x|\<^sub>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 fmember.rep_eq set_sorted_list_of_set) next fix x y xa assume 1: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>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|\<^sub>r" and 6: "xa \ set |h3 \ get_disconnected_nodes y|\<^sub>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 finite_fset fmember.rep_eq notin_set_remove1 select_result_I2 set_sorted_list_of_set by (metis (no_types, lifting)) qed qed next fix x xa xb assume 1: "(\x\fset (object_ptr_kinds h3). set |h3 \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h3). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" and 2: "xa |\| object_ptr_kinds h3" and 3: "x \ set |h3 \ get_child_nodes xa|\<^sub>r" and 4: "xb |\| document_ptr_kinds h3" and 5: "x \ set |h3 \ get_disconnected_nodes xb|\<^sub>r" have 6: "set |h3 \ get_child_nodes xa|\<^sub>r \ set |h2 \ get_disconnected_nodes xb|\<^sub>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|\<^sub>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|\<^sub>r" using 1 by (auto elim: distinct_concat_map_E) show "distinct |h' \ get_child_nodes x|\<^sub>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|\<^sub>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|\<^sub>r" and 6: "xa \ set |h' \ get_child_nodes y|\<^sub>r" have 7:"set |h3 \ get_child_nodes x|\<^sub>r \ set |h3 \ get_child_nodes y|\<^sub>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: " (\x\fset (object_ptr_kinds h'). set |h3 \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h'). set |h' \ get_disconnected_nodes x|\<^sub>r) = {} " and 2: "xa |\| object_ptr_kinds h'" and 3: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 4: "xb |\| document_ptr_kinds h'" and 5: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" have 6: "set |h3 \ get_child_nodes xa|\<^sub>r \ set |h' \ get_disconnected_nodes xb|\<^sub>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 \\<^sub>r |h' \ get_disconnected_nodes xb|\<^sub>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 children_eq2_h3 children_h' children_h3 disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 finite_set_in 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 lemma adopt_node_children_remain_distinct: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ adopt_node owner_document node_ptr \\<^sub>h h'" shows "\ptr' children'. h' \ get_child_nodes ptr' \\<^sub>r children' \ distinct children'" using assms(1) assms(2) assms(3) assms(4) local.adopt_node_preserves_wellformedness local.heap_is_wellformed_children_distinct by blast lemma insert_node_children_remain_distinct: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ a_insert_node ptr new_child reference_child_opt \\<^sub>h h'" assumes "h \ get_child_nodes ptr \\<^sub>r children" assumes "new_child \ set children" shows "\children'. h' \ get_child_nodes ptr \\<^sub>r children' \ distinct children'" proof - fix children' assume a1: "h' \ get_child_nodes ptr \\<^sub>r children'" have "h' \ get_child_nodes ptr \\<^sub>r (insert_before_list new_child reference_child_opt children)" using assms(4) assms(5) apply(auto simp add: a_insert_node_def elim!: bind_returns_heap_E)[1] using returns_result_eq set_child_nodes_get_child_nodes assms(2) assms(3) by (metis is_OK_returns_result_I local.get_child_nodes_ptr_in_heap local.get_child_nodes_pure local.known_ptrs_known_ptr pure_returns_heap_eq) moreover have "a_distinct_lists h" using assms local.heap_is_wellformed_def by blast then have "\children. h \ get_child_nodes ptr \\<^sub>r children \ distinct children" using assms local.heap_is_wellformed_children_distinct by blast ultimately show "h' \ get_child_nodes ptr \\<^sub>r children' \ distinct children'" using assms(5) assms(6) insert_before_list_distinct returns_result_eq by fastforce qed lemma insert_before_children_remain_distinct: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ insert_before ptr new_child child_opt \\<^sub>h h'" shows "\ptr' children'. h' \ get_child_nodes ptr' \\<^sub>r children' \ distinct children'" proof - obtain reference_child owner_document h2 h3 disconnected_nodes_h2 where reference_child: "h \ (if Some new_child = child_opt then a_next_sibling new_child else return child_opt) \\<^sub>r reference_child" and owner_document: "h \ get_owner_document ptr \\<^sub>r owner_document" and h2: "h \ adopt_node owner_document new_child \\<^sub>h h2" and disconnected_nodes_h2: "h2 \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h2" and h3: "h2 \ set_disconnected_nodes owner_document (remove1 new_child disconnected_nodes_h2) \\<^sub>h h3" and h': "h3 \ a_insert_node ptr new_child reference_child \\<^sub>h h'" using assms(4) 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 "\ptr children. h2 \ get_child_nodes ptr \\<^sub>r children \ distinct children" using adopt_node_children_remain_distinct using assms(1) assms(2) assms(3) h2 by blast moreover have "\ptr children. h2 \ get_child_nodes ptr \\<^sub>r children \ new_child \ set children" using adopt_node_removes_child using assms(1) assms(2) assms(3) h2 by blast moreover have "\ptr children. h2 \ get_child_nodes ptr \\<^sub>r children = h3 \ get_child_nodes ptr \\<^sub>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) ultimately show "\ptr children. h' \ get_child_nodes ptr \\<^sub>r children \ distinct children" using insert_node_children_remain_distinct by (meson assms(1) assms(2) assms(3) assms(4) insert_before_heap_is_wellformed_preserved(1) local.heap_is_wellformed_children_distinct) qed lemma insert_before_removes_child: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ insert_before ptr node child \\<^sub>h h'" assumes "ptr \ ptr'" shows "\children'. h' \ get_child_nodes ptr' \\<^sub>r children' \ node \ set children'" proof - fix children' assume a1: "h' \ get_child_nodes ptr' \\<^sub>r children'" obtain ancestors reference_child owner_document h2 h3 disconnected_nodes_h2 where ancestors: "h \ get_ancestors ptr \\<^sub>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) \\<^sub>r reference_child" and owner_document: "h \ get_owner_document ptr \\<^sub>r owner_document" and h2: "h \ adopt_node owner_document node \\<^sub>h h2" and disconnected_nodes_h2: "h2 \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h2" and h3: "h2 \ set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \\<^sub>h h3" and h': "h3 \ a_insert_node ptr node reference_child \\<^sub>h h'" using assms(4) 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 assms(2) 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 assms(3) 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 have "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 \\<^sub>r ptrs = h2 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs ) then have object_ptr_kinds_M_eq2_h: "|h \ object_ptr_kinds_M|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h: "|h \ node_ptr_kinds_M|\<^sub>r = |h2 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have "known_ptrs h2" using assms object_ptr_kinds_M_eq3_h known_ptrs_preserved by blast have wellformed_h2: "heap_is_wellformed h2" using adopt_node_preserves_wellformedness[OF assms(1) h2] assms by simp 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 \\<^sub>r ptrs = h3 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_M_eq2_h2: "|h2 \ object_ptr_kinds_M|\<^sub>r = |h3 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h2: "|h2 \ node_ptr_kinds_M|\<^sub>r = |h3 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have document_ptr_kinds_eq2_h2: "|h2 \ document_ptr_kinds_M|\<^sub>r = |h3 \ document_ptr_kinds_M|\<^sub>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 \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_M_eq2_h3: "|h3 \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h3: "|h3 \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have document_ptr_kinds_eq2_h3: "|h3 \ document_ptr_kinds_M|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_M_eq2_h3 document_ptr_kinds_M_eq by auto have "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 \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disconnected_nodes_h3: "h3 \ get_disconnected_nodes owner_document \\<^sub>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 \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have children_eq_h3: "\ptr' children. ptr \ ptr' \ h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force obtain children_h3 where children_h3: "h3 \ get_child_nodes ptr \\<^sub>r children_h3" using h' a_insert_node_def by auto have children_h': "h' \ get_child_nodes ptr \\<^sub>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 \\<^sub>r children \ node \ set children" using assms(1) assms(2) assms(3) h2 local.adopt_node_removes_child by blast show "node \ set children'" using a1 assms(5) child_not_in_any_children children_eq_h2 children_eq_h3 by blast qed lemma ensure_pre_insertion_validity_ok: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "ptr |\| object_ptr_kinds h" assumes "\is_character_data_ptr_kind parent" assumes "cast node \ set |h \ get_ancestors parent|\<^sub>r" assumes "h \ get_parent ref \\<^sub>r Some parent" assumes "is_document_ptr parent \ h \ get_child_nodes parent \\<^sub>r []" assumes "is_document_ptr parent \ \is_character_data_ptr_kind node" shows "h \ ok (a_ensure_pre_insertion_validity node parent (Some ref))" proof - have "h \ (if is_character_data_ptr_kind parent then error HierarchyRequestError else return ()) \\<^sub>r ()" using assms by (simp add: assms(4)) moreover have "h \ do { ancestors \ get_ancestors parent; (if cast node \ set ancestors then error HierarchyRequestError else return ()) } \\<^sub>r ()" using assms(6) apply(auto intro!: bind_pure_returns_result_I)[1] using assms(1) assms(2) assms(3) assms(7) local.get_ancestors_ok local.get_parent_parent_in_heap by auto moreover have "h \ do { (case Some ref of Some child \ do { child_parent \ get_parent child; (if child_parent \ Some parent then error NotFoundError else return ())} | None \ return ()) } \\<^sub>r ()" using assms(7) by(auto split: option.splits) moreover have "h \ do { children \ get_child_nodes parent; (if children \ [] \ is_document_ptr parent then error HierarchyRequestError else return ()) } \\<^sub>r ()" using assms(8) by (smt assms(5) assms(7) bind_pure_returns_result_I2 calculation(1) is_OK_returns_result_I local.get_child_nodes_pure local.get_parent_child_dual returns_result_eq) moreover have "h \ do { (if is_character_data_ptr node \ is_document_ptr parent then error HierarchyRequestError else return ()) } \\<^sub>r ()" using assms using is_character_data_ptr_kind_none by force ultimately show ?thesis unfolding a_ensure_pre_insertion_validity_def apply(intro bind_is_OK_pure_I) apply auto[1] apply auto[1] apply auto[1] using assms(1) assms(2) assms(3) assms(7) local.get_ancestors_ok local.get_parent_parent_in_heap apply blast apply auto[1] apply auto[1] using assms(6) apply auto[1] using assms(1) assms(2) assms(3) assms(7) local.get_ancestors_ok local.get_parent_parent_in_heap apply auto[1] apply (smt bind_returns_heap_E is_OK_returns_heap_E local.get_parent_pure pure_def pure_returns_heap_eq return_returns_heap returns_result_eq) apply(blast) using local.get_child_nodes_pure apply blast apply (meson assms(7) is_OK_returns_result_I local.get_parent_child_dual) apply (simp) apply (smt assms(5) assms(8) is_OK_returns_result_I returns_result_eq) by(auto) 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 \\<^sub>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 \\<^sub>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 \\<^sub>h h' \ heap_is_wellformed h'" interpretation i_insert_before_wf2?: l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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_insert_before_wf3\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_child_nodes_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_remove_child_wf2 begin lemma next_sibling_ok: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "node_ptr |\| node_ptr_kinds h" shows "h \ ok (a_next_sibling node_ptr)" proof - have "known_ptr (cast node_ptr)" using assms(2) assms(4) local.known_ptrs_known_ptr node_ptr_kinds_commutes by blast then show ?thesis using assms apply(auto simp add: a_next_sibling_def intro!: bind_is_OK_pure_I split: option.splits list.splits)[1] using get_child_nodes_ok local.get_parent_parent_in_heap local.known_ptrs_known_ptr by blast qed lemma remove_child_ok: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ get_child_nodes ptr \\<^sub>r children" assumes "child \ set children" shows "h \ ok (remove_child ptr child)" proof - have "ptr |\| object_ptr_kinds h" using assms(4) local.get_child_nodes_ptr_in_heap by blast have "child |\| node_ptr_kinds h" using assms(1) assms(4) assms(5) local.heap_is_wellformed_children_in_heap by blast have "\is_character_data_ptr ptr" proof (rule ccontr, simp) assume "is_character_data_ptr ptr" then have "h \ get_child_nodes ptr \\<^sub>r []" using \ptr |\| object_ptr_kinds h\ apply(simp add: get_child_nodes_def a_get_child_nodes_tups_def) apply(split invoke_splits)+ by(auto simp add: get_child_nodes\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_returns_result_I split: option.splits) then show False using assms returns_result_eq by fastforce qed have "is_character_data_ptr child \ \is_document_ptr_kind ptr" proof (rule ccontr, simp) assume "is_character_data_ptr\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child" and "is_document_ptr_kind ptr" then show False using assms using \ptr |\| object_ptr_kinds h\ apply(simp add: get_child_nodes_def a_get_child_nodes_tups_def) apply(split invoke_splits)+ apply(auto split: option.splits)[1] apply (meson invoke_empty is_OK_returns_result_I) apply (meson invoke_empty is_OK_returns_result_I) by(auto simp add: get_child_nodes\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: option.splits) qed obtain owner_document where owner_document: "h \ get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r owner_document" by (meson \child |\| node_ptr_kinds h\ assms(1) assms(2) assms(3) is_OK_returns_result_E local.get_owner_document_ok node_ptr_kinds_commutes) obtain disconnected_nodes_h where disconnected_nodes_h: "h \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h" by (meson assms(1) assms(2) assms(3) is_OK_returns_result_E local.get_disconnected_nodes_ok local.get_owner_document_owner_document_in_heap owner_document) obtain h2 where h2: "h \ set_disconnected_nodes owner_document (child # disconnected_nodes_h) \\<^sub>h h2" by (meson assms(1) assms(2) assms(3) is_OK_returns_heap_E l_set_disconnected_nodes.set_disconnected_nodes_ok local.get_owner_document_owner_document_in_heap local.l_set_disconnected_nodes_axioms owner_document) have "known_ptr ptr" using assms(2) assms(4) local.known_ptrs_known_ptr using \ptr |\| object_ptr_kinds h\ by blast have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h2] using set_disconnected_nodes_types_preserved assms(3) by(auto simp add: reflp_def transp_def) have "object_ptr_kinds h = object_ptr_kinds h2" using h2 apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF set_disconnected_nodes_writes]) using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) have "h2 \ ok (set_child_nodes ptr (remove1 child children))" proof (cases "is_element_ptr_kind ptr") case True then show ?thesis using set_child_nodes_element_ok \known_ptr ptr\ \object_ptr_kinds h = object_ptr_kinds h2\ \type_wf h2\ assms(4) using \ptr |\| object_ptr_kinds h\ by blast next case False then have "is_document_ptr_kind ptr" using \known_ptr ptr\ \ptr |\| object_ptr_kinds h\ \\is_character_data_ptr ptr\ by(auto simp add:known_ptr_impl known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) moreover have "is_document_ptr ptr" using \known_ptr ptr\ \ptr |\| object_ptr_kinds h\ False \\is_character_data_ptr ptr\ by(auto simp add: known_ptr_impl known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) ultimately show ?thesis using assms(4) apply(auto simp add: get_child_nodes_def a_get_child_nodes_tups_def)[1] apply(split invoke_splits)+ apply(auto elim!: bind_returns_result_E2 split: option.splits)[1] apply(auto simp add: get_child_nodes\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: option.splits)[1] using assms(5) apply auto[1] using \is_document_ptr_kind ptr\ \known_ptr ptr\ \object_ptr_kinds h = object_ptr_kinds h2\ \ptr |\| object_ptr_kinds h\ \type_wf h2\ local.set_child_nodes_document1_ok apply blast using \is_document_ptr_kind ptr\ \known_ptr ptr\ \object_ptr_kinds h = object_ptr_kinds h2\ \ptr |\| object_ptr_kinds h\ \type_wf h2\ is_element_ptr_kind_cast local.set_child_nodes_document2_ok apply blast using \\ is_character_data_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr\ apply blast by (metis False is_element_ptr_implies_kind option.case_eq_if) qed then obtain h' where h': "h2 \ set_child_nodes ptr (remove1 child children) \\<^sub>h h'" by auto show ?thesis using assms apply(auto simp add: remove_child_def simp add: is_OK_returns_heap_I[OF h2] is_OK_returns_heap_I[OF h'] is_OK_returns_result_I[OF assms(4)] is_OK_returns_result_I[OF owner_document] is_OK_returns_result_I[OF disconnected_nodes_h] intro!: bind_is_OK_pure_I[OF get_owner_document_pure] bind_is_OK_pure_I[OF get_child_nodes_pure] bind_is_OK_pure_I[OF get_disconnected_nodes_pure] bind_is_OK_I[rotated, OF h2] dest!: returns_result_eq[OF assms(4)] returns_result_eq[OF owner_document] returns_result_eq[OF disconnected_nodes_h] )[1] using h2 returns_result_select_result by force qed lemma adopt_node_ok: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "document_ptr |\| document_ptr_kinds h" assumes "child |\| node_ptr_kinds h" shows "h \ ok (adopt_node document_ptr child)" proof - obtain old_document where old_document: "h \ get_owner_document (cast child) \\<^sub>r old_document" by (meson assms(1) assms(2) assms(3) assms(5) is_OK_returns_result_E local.get_owner_document_ok node_ptr_kinds_commutes) then have "h \ ok (get_owner_document (cast child))" by auto obtain parent_opt where parent_opt: "h \ get_parent child \\<^sub>r parent_opt" by (meson assms(2) assms(3) is_OK_returns_result_I l_get_owner_document.get_owner_document_ptr_in_heap local.get_parent_ok local.l_get_owner_document_axioms node_ptr_kinds_commutes old_document returns_result_select_result) then have "h \ ok (get_parent child)" by auto have "h \ ok (case parent_opt of Some parent \ remove_child parent child | None \ return ())" apply(auto split: option.splits)[1] using remove_child_ok by (metis assms(1) assms(2) assms(3) local.get_parent_child_dual parent_opt) then obtain h2 where h2: "h \ (case parent_opt of Some parent \ remove_child parent child | None \ return ()) \\<^sub>h h2" by auto 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) then have "old_document |\| document_ptr_kinds h2" using assms(1) assms(2) assms(3) document_ptr_kinds_commutes local.get_owner_document_owner_document_in_heap old_document by blast have wellformed_h2: "heap_is_wellformed h2" using h2 remove_child_heap_is_wellformed_preserved assms by(auto split: option.splits) have "type_wf h2" using h2 remove_child_preserves_type_wf assms by(auto split: option.splits) have "known_ptrs h2" using h2 remove_child_preserves_known_ptrs assms by(auto split: option.splits) 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) then have "document_ptr_kinds h = document_ptr_kinds h2" by(auto simp add: document_ptr_kinds_def) have "h2 \ ok (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 () })" proof(cases "document_ptr = old_document") case True then show ?thesis by simp next case False then have "h2 \ ok (get_disconnected_nodes old_document)" by (simp add: \old_document |\| document_ptr_kinds h2\ \type_wf h2\ local.get_disconnected_nodes_ok) then obtain old_disc_nodes where old_disc_nodes: "h2 \ get_disconnected_nodes old_document \\<^sub>r old_disc_nodes" by auto have "h2 \ ok (set_disconnected_nodes old_document (remove1 child old_disc_nodes))" by (simp add: \old_document |\| document_ptr_kinds h2\ \type_wf h2\ local.set_disconnected_nodes_ok) then obtain h3 where h3: "h2 \ set_disconnected_nodes old_document (remove1 child old_disc_nodes) \\<^sub>h h3" by auto 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 \\<^sub>r ptrs = h3 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq_h2: "|h2 \ object_ptr_kinds_M|\<^sub>r = |h3 \ object_ptr_kinds_M|\<^sub>r" by(simp) then have node_ptr_kinds_eq_h2: "|h2 \ node_ptr_kinds_M|\<^sub>r = |h3 \ node_ptr_kinds_M|\<^sub>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|\<^sub>r = |h3 \ document_ptr_kinds_M|\<^sub>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 \\<^sub>r children = h3 \ get_child_nodes ptr \\<^sub>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|\<^sub>r = |h3 \ get_child_nodes ptr|\<^sub>r" using select_result_eq by force have "type_wf h3" using \type_wf h2\ 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) moreover have "document_ptr |\| document_ptr_kinds h3" using \document_ptr_kinds h = document_ptr_kinds h2\ assms(4) document_ptr_kinds_eq3_h2 by auto ultimately have "h3 \ ok (get_disconnected_nodes document_ptr)" by (simp add: local.get_disconnected_nodes_ok) then obtain disc_nodes where disc_nodes: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" by auto have "h3 \ ok (set_disconnected_nodes document_ptr (child # disc_nodes))" using \document_ptr |\| document_ptr_kinds h3\ \type_wf h3\ local.set_disconnected_nodes_ok by auto then obtain h' where h': "h3 \ set_disconnected_nodes document_ptr (child # disc_nodes) \\<^sub>h h'" by auto then show ?thesis using False using \h2 \ ok get_disconnected_nodes old_document\ using \h3 \ ok get_disconnected_nodes document_ptr\ apply(auto dest!: returns_result_eq[OF old_disc_nodes] returns_result_eq[OF disc_nodes] intro!: bind_is_OK_I[rotated, OF h3] bind_is_OK_pure_I[OF get_disconnected_nodes_pure] )[1] using \h2 \ ok set_disconnected_nodes old_document (remove1 child old_disc_nodes)\ by auto qed then obtain h' where 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 () }) \\<^sub>h h'" by auto show ?thesis using \h \ ok (get_owner_document (cast child))\ using \h \ ok (get_parent child)\ using h2 h' apply(auto simp add: adopt_node_def simp add: is_OK_returns_heap_I[OF h2] intro!: bind_is_OK_pure_I[OF get_owner_document_pure] bind_is_OK_pure_I[OF get_parent_pure] bind_is_OK_I[rotated, OF h2] dest!: returns_result_eq[OF parent_opt] returns_result_eq[OF old_document])[1] using \h \ ok (case parent_opt of None \ return () | Some parent \ remove_child parent child)\ by auto qed lemma insert_node_ok: assumes "known_ptr parent" and "type_wf h" assumes "parent |\| object_ptr_kinds h" assumes "\is_character_data_ptr_kind parent" assumes "is_document_ptr parent \ h \ get_child_nodes parent \\<^sub>r []" assumes "is_document_ptr parent \ \is_character_data_ptr_kind node" assumes "known_ptr (cast node)" shows "h \ ok (a_insert_node parent node ref)" proof(auto simp add: a_insert_node_def get_child_nodes_ok[OF assms(1) assms(2) assms(3)] intro!: bind_is_OK_pure_I) fix children' assume "h \ get_child_nodes parent \\<^sub>r children'" show "h \ ok set_child_nodes parent (insert_before_list node ref children')" proof (cases "is_element_ptr_kind parent") case True then show ?thesis using set_child_nodes_element_ok using assms(1) assms(2) assms(3) by blast next case False then have "is_document_ptr_kind parent" using assms(4) assms(1) 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 "is_document_ptr parent" using assms(1) 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 children where children: "h \ get_child_nodes parent \\<^sub>r children" and "children = []" using assms(5) by blast have "insert_before_list node ref children' = [node]" by (metis \children = []\ \h \ get_child_nodes parent \\<^sub>r children'\ append.left_neutral children insert_Nil l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.insert_before_list.elims l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.insert_before_list.simps(3) neq_Nil_conv returns_result_eq) moreover have "\is_character_data_ptr_kind node" using \is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r parent\ assms(6) by blast then have "is_element_ptr_kind node" by (metis (no_types, lifting) CharacterDataClass.a_known_ptr_def DocumentClass.a_known_ptr_def ElementClass.a_known_ptr_def NodeClass.a_known_ptr_def assms(7) cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_inject document_ptr_no_node_ptr_cast is_character_data_ptr_kind_none is_document_ptr_kind_none is_element_ptr_implies_kind is_node_ptr_kind_cast local.known_ptr_impl node_ptr_casts_commute3 option.case_eq_if) ultimately show ?thesis using set_child_nodes_document2_ok by (metis \is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r parent\ assms(1) assms(2) assms(3) assms(5) is_document_ptr_kind_none option.case_eq_if) qed qed lemma insert_before_ok: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "parent |\| object_ptr_kinds h" assumes "node |\| node_ptr_kinds h" assumes "\is_character_data_ptr_kind parent" assumes "cast node \ set |h \ get_ancestors parent|\<^sub>r" assumes "h \ get_parent ref \\<^sub>r Some parent" assumes "is_document_ptr parent \ h \ get_child_nodes parent \\<^sub>r []" assumes "is_document_ptr parent \ \is_character_data_ptr_kind node" shows "h \ ok (insert_before parent node (Some ref))" proof - have "h \ ok (a_ensure_pre_insertion_validity node parent (Some ref))" using assms ensure_pre_insertion_validity_ok by blast have "h \ ok (if Some node = Some ref then a_next_sibling node else return (Some ref))" (is "h \ ok ?P") apply(auto split: if_splits)[1] using assms(1) assms(2) assms(3) assms(5) next_sibling_ok by blast then obtain reference_child where reference_child: "h \ ?P \\<^sub>r reference_child" by auto obtain owner_document where owner_document: "h \ get_owner_document parent \\<^sub>r owner_document" using assms get_owner_document_ok by (meson returns_result_select_result) then have "h \ ok (get_owner_document parent)" by auto have "owner_document |\| document_ptr_kinds h" using assms(1) assms(2) assms(3) local.get_owner_document_owner_document_in_heap owner_document by blast obtain h2 where h2: "h \ adopt_node owner_document node \\<^sub>h h2" by (meson assms(1) assms(2) assms(3) assms(5) is_OK_returns_heap_E adopt_node_ok l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms local.get_owner_document_owner_document_in_heap owner_document) then have "h \ ok (adopt_node owner_document node)" by auto have "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 "document_ptr_kinds h = document_ptr_kinds h2" by(auto simp add: document_ptr_kinds_def) have "heap_is_wellformed h2" using h2 adopt_node_preserves_wellformedness assms by blast have "known_ptrs h2" using h2 adopt_node_preserves_known_ptrs assms by blast have "type_wf h2" using h2 adopt_node_preserves_type_wf assms by blast obtain disconnected_nodes_h2 where disconnected_nodes_h2: "h2 \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h2" by (metis \document_ptr_kinds h = document_ptr_kinds h2\ \type_wf h2\ assms(1) assms(2) assms(3) is_OK_returns_result_E local.get_disconnected_nodes_ok local.get_owner_document_owner_document_in_heap owner_document) obtain h3 where h3: "h2 \ set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \\<^sub>h h3" by (metis \document_ptr_kinds h = document_ptr_kinds h2\ \owner_document |\| document_ptr_kinds h\ \type_wf h2\ document_ptr_kinds_def is_OK_returns_heap_E l_set_disconnected_nodes.set_disconnected_nodes_ok local.l_set_disconnected_nodes_axioms) have "type_wf h3" using \type_wf h2\ 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) 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) have "parent |\| object_ptr_kinds h3" using \object_ptr_kinds h = object_ptr_kinds h2\ assms(4) object_ptr_kinds_M_eq3_h2 by blast moreover have "known_ptr parent" using assms(2) assms(4) local.known_ptrs_known_ptr by blast moreover have "known_ptr (cast node)" using assms(2) assms(5) local.known_ptrs_known_ptr node_ptr_kinds_commutes by blast moreover have "is_document_ptr parent \ h3 \ get_child_nodes parent \\<^sub>r []" by (metis assms(8) assms(9) distinct.simps(2) distinct_singleton local.get_parent_child_dual returns_result_eq) ultimately obtain h' where h': "h3 \ a_insert_node parent node reference_child \\<^sub>h h'" using insert_node_ok \type_wf h3\ assms by blast show ?thesis using \h \ ok (a_ensure_pre_insertion_validity node parent (Some ref))\ using reference_child \h \ ok (get_owner_document parent)\ \h \ ok (adopt_node owner_document node)\ h3 h' apply(auto simp add: insert_before_def simp add: is_OK_returns_result_I[OF disconnected_nodes_h2] simp add: is_OK_returns_heap_I[OF h3] is_OK_returns_heap_I[OF h'] intro!: bind_is_OK_I2 bind_is_OK_pure_I[OF ensure_pre_insertion_validity_pure] bind_is_OK_pure_I[OF next_sibling_pure] bind_is_OK_pure_I[OF get_owner_document_pure] bind_is_OK_pure_I[OF get_disconnected_nodes_pure] dest!: returns_result_eq[OF owner_document] returns_result_eq[OF disconnected_nodes_h2] returns_heap_eq[OF h2] returns_heap_eq[OF h3] dest!: sym[of node ref] )[1] using returns_result_eq by fastforce qed end interpretation i_insert_before_wf3?: l_insert_before_wf3\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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 remove by(auto simp add: l_insert_before_wf3\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_insert_before_wf3\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_append_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + 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 \\<^sub>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 \\<^sub>r xs" assumes "h \ append_child ptr node \\<^sub>h h'" assumes "node \ set xs" shows "h' \ get_child_nodes ptr \\<^sub>r xs @ [node]" proof - obtain ancestors owner_document h2 h3 disconnected_nodes_h2 where ancestors: "h \ get_ancestors ptr \\<^sub>r ancestors" and node_not_in_ancestors: "cast node \ set ancestors" and owner_document: "h \ get_owner_document ptr \\<^sub>r owner_document" and h2: "h \ adopt_node owner_document node \\<^sub>h h2" and disconnected_nodes_h2: "h2 \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h2" and h3: "h2 \ set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \\<^sub>h h3" and h': "h3 \ a_insert_node ptr node None \\<^sub>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|\<^sub>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 \\<^sub>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|\<^sub>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 \\<^sub>r xs" using get_child_nodes_reads set_disconnected_nodes_writes h3 \h2 \ get_child_nodes ptr \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r xs" assumes "h \ forall_M (append_child ptr) nodes \\<^sub>h h'" assumes "set nodes \ set xs = {}" assumes "distinct nodes" shows "h' \ get_child_nodes ptr \\<^sub>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 \\<^sub>r xs \ h \ forall_M (append_child ptr) nodes \\<^sub>h h' \ set nodes \ set xs = {} \ h' \ get_child_nodes ptr \\<^sub>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 \\<^sub>r xs" and 5: "h \ append_child ptr a \\<^sub>r ()" and 6: "h \ append_child ptr a \\<^sub>h h'a" and 7: "h'a \ forall_M (append_child ptr) nodes \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r []" assumes "h \ forall_M (append_child ptr) nodes \\<^sub>h h'" assumes "distinct nodes" shows "h' \ get_child_nodes ptr \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>h h' \ heap_is_wellformed h'" interpretation i_append_child_wf?: l_append_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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 \\<^sub>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 \\<^sub>r new_element_ptr" and h2: "h \ new_element \\<^sub>h h2" and h3: "h2 \ set_tag_name new_element_ptr tag \\<^sub>h h3" and disc_nodes_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" and h': "h3 \ set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes_h3) \\<^sub>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 \\<^sub>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|\<^sub>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|\<^sub>r" by simp then have "cast new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>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 \\<^sub>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' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes (cast new_element_ptr) \\<^sub>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 \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>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' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>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 \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>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 \\<^sub>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 \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>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|\<^sub>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|\<^sub>r" then show "x \ set |h2 \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq_h \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>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 \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ children_eq2_h3 disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 finite_set_in h' is_OK_returns_result_I set_disconnected_nodes_get_disconnected_nodes local.a_all_ptrs_in_heap_def local.get_child_nodes_ptr_in_heap 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|\<^sub>r" using \heap_is_wellformed h\ \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>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|\<^sub>r" using children_eq2_h apply(auto simp add: object_ptr_kinds_eq_h)[1] using \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ apply auto[1] by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\) then have "\p. p |\| object_ptr_kinds h3 \ cast new_element_ptr \ set |h3 \ get_child_nodes p|\<^sub>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|\<^sub>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) \\<^sub>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|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" then show "distinct |h' \ get_disconnected_nodes x|\<^sub>r" using document_ptr_kinds_eq_h3 disconnected_nodes_eq_h3 h' set_disconnected_nodes_get_disconnected_nodes by (metis (no_types, lifting) \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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|\<^sub>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|\<^sub>r" and "xa \ set |h' \ get_disconnected_nodes y|\<^sub>r" moreover have "set |h3 \ get_disconnected_nodes x|\<^sub>r \ set |h3 \ get_disconnected_nodes y|\<^sub>r = {}" using calculation by(auto dest: distinct_concat_map_E(1)) ultimately show "False" apply(-) apply(cases "x = document_ptr") apply (smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \local.a_all_ptrs_in_heap h\ disc_nodes_h3 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms select_result_I2 set_ConsD subsetD) by (smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \local.a_all_ptrs_in_heap h\ disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms select_result_I2 set_ConsD subsetD) next fix x xa xb assume 2: "(\x\fset (object_ptr_kinds h3). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h3). set |h3 \ get_disconnected_nodes x|\<^sub>r) = {}" and 3: "xa |\| object_ptr_kinds h3" and 4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 5: "xb |\| document_ptr_kinds h3" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>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 \ cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h3 \ get_child_nodes p|\<^sub>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 (smt ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h children_eq2_h2 children_eq2_h3 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h finite_set_in h' list.set_intros(2) local.set_disconnected_nodes_get_disconnected_nodes 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsection \create\_character\_data\ locale l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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 \\<^sub>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 \\<^sub>r new_character_data_ptr" and h2: "h \ new_character_data \\<^sub>h h2" and h3: "h2 \ set_val new_character_data_ptr text \\<^sub>h h3" and disc_nodes_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" and h': "h3 \ set_disconnected_nodes document_ptr (cast new_character_data_ptr # disc_nodes_h3) \\<^sub>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 \\<^sub>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|\<^sub>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|\<^sub>r" by simp then have "cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r" by simp have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_character_data_ptr|}" using new_character_data_new_ptr h2 new_character_data_ptr by blast then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = node_ptr_kinds h |\| {|cast new_character_data_ptr|}" apply(simp add: node_ptr_kinds_def) by force then have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h2 = character_data_ptr_kinds h |\| {|new_character_data_ptr|}" apply(simp add: character_data_ptr_kinds_def) by force have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: document_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_val_writes h3]) using set_val_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) have "document_ptr |\| document_ptr_kinds h" using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2 get_disconnected_nodes_ptr_in_heap \type_wf h\ document_ptr_kinds_def by (metis is_OK_returns_result_I) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_character_data_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>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' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>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 \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>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' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>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 \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>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 \\<^sub>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|\<^sub>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|\<^sub>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|\<^sub>r" then show "x \ set |h2 \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq_h \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>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|\<^sub>r\ \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ apply (metis (no_types, lifting) NodeMonad.ptr_kinds_ptr_kinds_M \parent_child_rel h = parent_child_rel h2\ children_eq2_h finite_set_in 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 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 finite_set_in 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|\<^sub>r" using \heap_is_wellformed h\ \cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>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|\<^sub>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) \\<^sub>r []\ apply auto[1] by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\) then have "\p. p |\| object_ptr_kinds h3 \ cast new_character_data_ptr \ set |h3 \ get_child_nodes p|\<^sub>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|\<^sub>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) \\<^sub>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|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" then show "distinct |h' \ get_disconnected_nodes x|\<^sub>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|\<^sub>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|\<^sub>r" and "xa \ set |h' \ get_disconnected_nodes y|\<^sub>r" moreover have "set |h3 \ get_disconnected_nodes x|\<^sub>r \ set |h3 \ get_disconnected_nodes y|\<^sub>r = {}" using calculation by(auto dest: distinct_concat_map_E(1)) ultimately show "False" by (smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \local.a_all_ptrs_in_heap h\ disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms select_result_I2 set_ConsD subsetD) next fix x xa xb assume 2: "(\x\fset (object_ptr_kinds h3). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h3). set |h3 \ get_disconnected_nodes x|\<^sub>r) = {}" and 3: "xa |\| object_ptr_kinds h3" and 4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 5: "xb |\| document_ptr_kinds h3" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>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|\<^sub>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 (smt ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h finite_set_in h' list.set_intros(2) local.set_disconnected_nodes_get_disconnected_nodes select_result_I2) have "known_ptr (cast new_character_data_ptr)" using \h \ create_character_data document_ptr text \\<^sub>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 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsection \create\_document\ locale l_create_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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 \\<^sub>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 \\<^sub>r new_document_ptr" and h': "h \ new_document \\<^sub>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|\<^sub>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|\<^sub>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 apply(auto simp add: document_ptr_kinds_def)[1] by (metis (no_types, lifting) document_ptr_kinds_commutes document_ptr_kinds_def finsertI1 fset.map_comp) have children_eq: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_document_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h' \ get_child_nodes (cast new_document_ptr) \\<^sub>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 \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>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 (metis(full_types) \\thesis. (\new_document_ptr. \h \ new_document \\<^sub>r new_document_ptr; h \ new_document \\<^sub>h h'\ \ thesis) \ thesis\ local.get_disconnected_nodes_new_document_different_pointers new_document_ptr)+ then have disconnected_nodes_eq2_h: "\doc_ptr. doc_ptr \ new_document_ptr \ |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have "h' \ get_disconnected_nodes new_document_ptr \\<^sub>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|\<^sub>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|\<^sub>r" then show "x \ set |h' \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_document_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2) next fix a x assume 0: "a |\| object_ptr_kinds h'" and 1: "x \ set |h' \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq \h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h'" and 1: "x \ set |h' \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h' \ get_child_nodes (cast new_document_ptr) \\<^sub>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 \cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr \ set |h \ object_ptr_kinds_M|\<^sub>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, lifting) \h' \ get_child_nodes (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr) \\<^sub>r []\ children_eq2 finite_set_in finsert_iff funion_finsert_right object_ptr_kinds_eq select_result_I2 subsetD sup_bot.right_neutral) by (metis (no_types, lifting) \cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr |\| object_ptr_kinds h\ \h' \ get_child_nodes (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr) \\<^sub>r []\ \h' \ get_disconnected_nodes new_document_ptr \\<^sub>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 \\<^sub>r []\ \h' \ get_child_nodes (cast new_document_ptr) \\<^sub>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|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h)))))" assume a7: "xa \ set |h' \ get_disconnected_nodes x|\<^sub>r" assume a8: "xa \ set |h' \ get_disconnected_nodes y|\<^sub>r" have f9: "xa \ set |h \ get_disconnected_nodes x|\<^sub>r" using a7 a3 disconnected_nodes_eq2_h by presburger have f10: "xa \ set |h \ get_disconnected_nodes y|\<^sub>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 \\<^sub>r []" and 1: "h' \ get_child_nodes (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr) \\<^sub>r []" and 2: "distinct (concat (map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h)))))" and 3: "distinct (concat (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h)))))" and 4: "(\x\fset (object_ptr_kinds h). set |h \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h). set |h \ get_disconnected_nodes x|\<^sub>r) = {}" and 5: "x \ set |h \ get_child_nodes xa|\<^sub>r" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" and 7: "xa |\| object_ptr_kinds h" and 8: "xa \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr |\| object_ptr_kinds h\ children_eq2 disconnected_nodes_eq2_h document_ptr_kinds_commutes finite_set_in 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_create_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] end diff --git a/thys/Core_SC_DOM/safely_composable/Core_DOM_Heap_WF.thy b/thys/Core_SC_DOM/safely_composable/Core_DOM_Heap_WF.thy --- a/thys/Core_SC_DOM/safely_composable/Core_DOM_Heap_WF.thy +++ b/thys/Core_SC_DOM/safely_composable/Core_DOM_Heap_WF.thy @@ -1,6965 +1,6965 @@ (*********************************************************************************** * 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_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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|\<^sub>r) \ (\parent_ptr. parent_ptr |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>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|\<^sub>r) (sorted_list_of_fset (object_ptr_kinds h)) @ map (\parent. |h \ get_disconnected_nodes parent|\<^sub>r) (sorted_list_of_fset (document_ptr_kinds h)))) " apply(auto simp add: a_owner_document_valid_def l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_owner_document_valid_def)[1] proof - fix x assume 1: " \node_ptr\fset (node_ptr_kinds h). (\document_ptr. document_ptr |\| document_ptr_kinds h \ node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r) \ (\parent_ptr. parent_ptr |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>r)" assume 2: "x |\| node_ptr_kinds h" assume 3: "x |\| fset_of_list (concat (map (\parent. |h \ get_disconnected_nodes parent|\<^sub>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|\<^sub>r)" using 1 2 3 by (smt UN_I fset_of_list_elem image_eqI notin_fset 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|\<^sub>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|\<^sub>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|\<^sub>r \ set (map (\parent. |h \ get_child_nodes parent|\<^sub>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|\<^sub>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|\<^sub>r) (sorted_list_of_fset (object_ptr_kinds h)))) |\| fset_of_list (concat (map (\parent. |h \ get_disconnected_nodes parent|\<^sub>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|\<^sub>r" have "node_ptr \ set (concat (map (\parent. |h \ get_child_nodes parent|\<^sub>r) (sorted_list_of_fset (object_ptr_kinds h)))) \ node_ptr \ set (concat (map (\parent. |h \ get_disconnected_nodes parent|\<^sub>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|\<^sub>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|\<^sub>r}" lemma a_parent_child_rel_code [code]: "a_parent_child_rel h = set (concat (map (\parent. map (\child. (parent, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child)) |h \ get_child_nodes parent|\<^sub>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|\<^sub>r \ fset (node_ptr_kinds h)) \ (\document_ptr \ fset (document_ptr_kinds h). set |h \ get_disconnected_nodes document_ptr|\<^sub>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|\<^sub>r) |h \ object_ptr_kinds_M|\<^sub>r) @ (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) |h \ document_ptr_kinds_M|\<^sub>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_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs defines heap_is_wellformed = "l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_heap_is_wellformed get_child_nodes get_disconnected_nodes" and parent_child_rel = "l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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 \\<^sub>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 \\<^sub>r children" and "child \ set children" and "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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|\<^sub>r. (\child \ set |h \ get_child_nodes ptr|\<^sub>r. {(ptr, cast child)}))" by(auto simp add: parent_child_rel_def) moreover have "finite (\ptr \ set |h \ object_ptr_kinds_M|\<^sub>r. (\child \ set |h \ get_child_nodes ptr|\<^sub>r. {(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>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|\<^sub>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: "(\x\fset (object_ptr_kinds h). set |h \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h). set |h \ get_disconnected_nodes x|\<^sub>r) = {}" assume a3: "h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" assume a4: "node_ptr \ set disc_nodes" assume a5: "node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>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 fmember.rep_eq get_disconnected_nodes_ptr_in_heap is_OK_returns_result_I) have "|h \ get_disconnected_nodes document_ptr|\<^sub>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 \\<^sub>r disc_nodes" shows "distinct disc_nodes" proof - have h1: "distinct (concat (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) |h \ document_ptr_kinds_M|\<^sub>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 \\<^sub>r children" shows "distinct children" proof (cases "children = []", simp) assume "children \ []" have h1: "distinct (concat ((map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) |h \ object_ptr_kinds_M|\<^sub>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 \\<^sub>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, lifting) finite_set_in 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 \\<^sub>r children" assumes "h \ get_child_nodes ptr' \\<^sub>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 \\<^sub>r children" assume a3: "h \ get_child_nodes ptr' \\<^sub>r children'" assume a4: "distinct (concat (map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h)))))" have f5: "|h \ get_child_nodes ptr|\<^sub>r = children" using a2 by simp have "|h \ get_child_nodes ptr'|\<^sub>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 fmember.rep_eq 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 \\<^sub>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 \\<^sub>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 finite_set_in subsetD) lemma heap_is_wellformed_disc_nodes_in_heap: "heap_is_wellformed h \ h \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes \ node \ set disc_nodes \ node |\| node_ptr_kinds h" by (metis (no_types, lifting) finite_set_in 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 \\<^sub>r disc_nodes \ h \ get_disconnected_nodes document_ptr' \\<^sub>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 \\<^sub>r disc_nodes" assume a3: "h \ get_disconnected_nodes document_ptr' \\<^sub>r disc_nodes'" assume a4: "set disc_nodes \ set disc_nodes' \ {}" have f5: "|h \ get_disconnected_nodes document_ptr|\<^sub>r = disc_nodes" using a2 by (meson select_result_I2) have f6: "|h \ get_disconnected_nodes document_ptr'|\<^sub>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|\<^sub>r) |h \ document_ptr_kinds_M|\<^sub>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 \\<^sub>r children \ h \ get_disconnected_nodes document_ptr \\<^sub>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|\<^sub>r) \ (\document_ptr \ fset (document_ptr_kinds h). node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" apply(auto simp add: heap_is_wellformed_def a_distinct_lists_def a_owner_document_valid_def)[1] by (meson fmember.rep_eq) lemma heap_is_wellformed_children_distinct: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r children \ h \ get_child_nodes ptr' \\<^sub>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 \\<^sub>r disc_nodes \ h \ get_disconnected_nodes document_ptr' \\<^sub>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 \\<^sub>r children \ h \ get_disconnected_nodes document_ptr \\<^sub>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 \\<^sub>r disc_nodes \ distinct disc_nodes" assumes heap_is_wellformed_children_distinct: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>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|\<^sub>r) \ (\document_ptr \ fset (document_ptr_kinds h). node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" assumes parent_child_rel_child: "h \ get_child_nodes ptr \\<^sub>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_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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 \\<^sub>r children" assumes "child \ set children" assumes "known_ptrs h" assumes type_wf: "type_wf h" shows "h \ get_parent child \\<^sub>r Some ptr" proof - obtain ptrs where ptrs: "h \ object_ptr_kinds_M \\<^sub>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 \\<^sub>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 \\<^sub>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)) \\<^sub>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) \\<^sub>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)) \\<^sub>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)) \\<^sub>r True" and 2: "y \ set ptrs" and 3: "h \ Heap_Error_Monad.bind (get_child_nodes y) (\children. return (child \ set children)) \\<^sub>r True" and 4: "(\h ptr children ptr' children'. heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ get_child_nodes ptr' \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r [] \ P parent" and step: "\parent children child. h \ get_child_nodes parent \\<^sub>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 \\<^sub>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 \\<^sub>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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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 \\<^sub>h h'" and "writes SW f h h'" and preserved_get_child_nodes: "\h h' w. w \ SW \ h \ w \\<^sub>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 \\<^sub>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 \\<^sub>h h' \ \object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t 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 \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" unfolding object_ptr_kinds_M_defs by simp then have object_ptr_kinds_eq2: "|h \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" using select_result_eq by force then have node_ptr_kinds_eq2: "|h \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>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|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>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 \\<^sub>r children = h' \ get_child_nodes ptr \\<^sub>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|\<^sub>r = |h' \ get_child_nodes ptr|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq: "\document_ptr disconnected_nodes. h \ get_disconnected_nodes document_ptr \\<^sub>r disconnected_nodes = h' \ get_disconnected_nodes document_ptr \\<^sub>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|\<^sub>r = |h' \ get_disconnected_nodes document_ptr|\<^sub>r" using select_result_eq by force have get_parent_eq: "\ptr parent. h \ get_parent ptr \\<^sub>r parent = h' \ get_parent ptr \\<^sub>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_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h))) = map (\ptr. |h' \ get_child_nodes ptr|\<^sub>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|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h))) = map (\document_ptr. |h' \ get_disconnected_nodes document_ptr|\<^sub>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_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms by (simp add: l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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 \\<^sub>r children \ child \ set children \ h \ get_parent child \\<^sub>r Some ptr" assumes heap_wellformed_induct [consumes 1, case_names step]: "heap_is_wellformed h \ (\parent. (\children child. h \ get_child_nodes parent \\<^sub>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 \\<^sub>r Some parent \ P parent) \ P child) \ P ptr" assumes parent_child_rel_parent: "heap_is_wellformed h \ h \ get_parent child_node \\<^sub>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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = 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 \\<^sub>r disc_nodes" assumes "h \ set_disconnected_nodes ptr (remove1 node_ptr disc_nodes) \\<^sub>h h'" assumes "h' \ get_disconnected_nodes ptr \\<^sub>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 \\<^sub>r disc_nodes \ h \ set_disconnected_nodes ptr (remove1 node_ptr disc_nodes) \\<^sub>h h' \ h' \ get_disconnected_nodes ptr \\<^sub>r disc_nodes' \ node_ptr \ set disc_nodes'" interpretation i_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M?: l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = 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_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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 \\<^sub>r Some parent" shows "h \ get_ancestors (cast child) \\<^sub>r (cast child) # parent # ancestors \ h \ get_ancestors parent \\<^sub>r parent # ancestors" proof assume a1: "h \ get_ancestors (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child # parent # ancestors" then have "h \ Heap_Error_Monad.bind (check_in_heap (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child # ancestors)))) \\<^sub>r cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child # parent # ancestors" by(simp add: get_ancestors_def) then show "h \ get_ancestors parent \\<^sub>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 \\<^sub>r parent # ancestors" then show "h \ get_ancestors (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>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 "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>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 \\<^sub>r ancestors" and "ancestor \ set ancestors" and "h \ get_ancestors ancestor \\<^sub>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 \\<^sub>r ()" using returns_result_select_result by force *) show ?case proof (induct "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>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_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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 cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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))) \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r ancestors" and "cast child \ set ancestors" and "h \ get_parent child \\<^sub>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) \\<^sub>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 \\<^sub>r ancestors" and type_wf: "type_wf h" and known_ptrs: "known_ptrs h" obtains children ancestor_child where "h \ get_child_nodes ancestor \\<^sub>r children" and "ancestor_child \ set children" and "cast ancestor_child \ set ancestors" proof - assume 0: "(\children ancestor_child. h \ get_child_nodes ancestor \\<^sub>r children \ ancestor_child \ set children \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ancestor_child \ set ancestors \ thesis)" have "\child. h \ get_parent child \\<^sub>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 "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>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 cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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))) \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r Some ancestor" and in_ancestors: "cast child \ set ancestors" by auto then obtain children where children: "h \ get_child_nodes ancestor \\<^sub>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 \\<^sub>r ancestors" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "(ptr, child) \ (parent_child_rel h)\<^sup>* \ ptr \ set ancestors" proof (safe) assume 3: "(ptr, child) \ (parent_child_rel h)\<^sup>*" 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)\<^sup>*" using converse_rtranclE[OF 1(2)] \ptr \ child\ by metis then obtain ptr_child_node where ptr_child_ptr_child_node: "ptr_child = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>r children" and ptr_child_node: "ptr_child_node \ set children" proof - assume a1: "\children. \h \ get_child_nodes ptr \\<^sub>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|\<^sub>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 "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node, child) \ (parent_child_rel h)\<^sup>*" using ptr_child ptr_child_ptr_child_node by auto ultimately have "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node \ set ancestors" using 1 by auto moreover have "h \ get_parent ptr_child_node \\<^sub>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)\<^sup>*" 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 \\<^sub>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 "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr_child_node, child) \ (parent_child_rel h)\<^sup>*" 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 \\<^sub>r root" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "(root, child) \ (parent_child_rel h)\<^sup>*" 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_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'" and known_ptrs: "known_ptrs h" and known_ptrs': "known_ptrs h'" and "h \ get_ancestors ptr \\<^sub>r ancestors" and type_wf: "type_wf h" and type_wf': "type_wf h'" shows "h' \ get_ancestors ptr \\<^sub>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 \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq: "|h \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>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 \\<^sub>r ancestors'" by auto obtain root where root: "h \ get_root_node ptr \\<^sub>r root" proof - assume 0: "(\root. h \ get_root_node ptr \\<^sub>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 \\<^sub>r children = h' \ get_child_nodes p \\<^sub>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. cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c \ set ancestors \ set ancestors' \ h \ get_parent c \\<^sub>r parent_opt = h' \ get_parent c \\<^sub>r parent_opt" proof - fix c parent_opt assume 1: " cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c \ set ancestors \ set ancestors'" obtain ptrs where ptrs: "h \ object_ptr_kinds_M \\<^sub>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 \\<^sub>r True \ h' \ ?P p \\<^sub>r True" proof - fix p assume "p \ set ptrs" then show "h \ ?P p \\<^sub>r True \ h' \ ?P p \\<^sub>r True" proof (cases "p = ptr") case True have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c, ptr) \ (parent_child_rel h)\<^sup>*" using get_ancestors_parent_child_rel 1 assms by blast then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \ (parent_child_rel h)\<^sup>*" using \acyclic (parent_child_rel h)\ False rtrancl_eq_or_trancl rtrancl_trancl_trancl \(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c, ptr) \ (parent_child_rel h)\<^sup>*\ by (metis acyclic_def) then show ?thesis using r_into_rtrancl by auto qed obtain children where children: "h \ get_child_nodes ptr \\<^sub>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, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>r False" by(auto simp add: True) moreover have "(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c, ptr) \ (parent_child_rel h')\<^sup>*" using get_ancestors_parent_child_rel assms(2) ancestors' 1 known_ptrs' type_wf type_wf' by blast then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \ (parent_child_rel h')\<^sup>*" using \acyclic (parent_child_rel h')\ False rtrancl_eq_or_trancl rtrancl_trancl_trancl \(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c, ptr) \ (parent_child_rel h')\<^sup>*\ by (metis acyclic_def) then show ?thesis using r_into_rtrancl by auto qed then have "(ptr, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \ (parent_child_rel h')" using r_into_rtrancl by auto obtain children' where children': "h' \ get_child_nodes ptr \\<^sub>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, cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>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)) \\<^sub>r False = h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False" proof fix pa assume "pa \ set ptrs" and "h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>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)) \\<^sub>r False \ h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False" by (metis (mono_tags, lifting) \\pa. pa \ set ptrs \ h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r True = h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>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)) \\<^sub>r False" using \h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False\ by auto next fix pa assume "pa \ set ptrs" and "h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>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)) \\<^sub>r False \ h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False" by (metis (mono_tags, lifting) \\pa. pa \ set ptrs \ h \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r True = h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>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)) \\<^sub>r False" using \h' \ get_child_nodes pa \ (\children. return (c \ set children)) \\<^sub>r False\ by blast qed have filter_eq: "\xs. h \ filter_M ?P ptrs \\<^sub>r xs = h' \ filter_M ?P ptrs \\<^sub>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)) \\<^sub>r b = h' \ Heap_Error_Monad.bind (get_child_nodes x) (\children. return (c \ set children)) \\<^sub>r b" apply(induct b) using children_eq_True apply blast using children_eq_False apply blast done qed show "h \ get_parent c \\<^sub>r parent_opt = h' \ get_parent c \\<^sub>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 \\<^sub>r ancestors" and "h' \ get_ancestors ptr \\<^sub>r ancestors'" and "\p children children'. h \ get_child_nodes p \\<^sub>r children \ h' \ get_child_nodes p \\<^sub>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 \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>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|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>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 \\<^sub>r Some parent \ h \ get_parent p \\<^sub>r Some parent" proof - fix p parent assume "h' \ get_parent p \\<^sub>r Some parent" then obtain children' where children': "h' \ get_child_nodes parent \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r ancestors" assumes "ptr' \ set ancestors" assumes "h \ get_ancestors ptr' \\<^sub>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 \\<^sub>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 \\<^sub>r Some parent" using assms(1) assms(2) assms(3) child_parent_dual by blast then have "h \ get_ancestors (cast ancestor_child) \\<^sub>r cast ancestor_child # ancestors'" apply(simp add: get_ancestors_def) using \h \ get_ancestors parent \\<^sub>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 \\<^sub>r children\ \ancestor_child \ set children\ \cast ancestor_child \ set ancestors\ \h \ get_ancestors (cast ancestor_child) \\<^sub>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 \\<^sub>r ancestors" assumes "ptr' \ set ancestors" assumes "ptr'' \ set ancestors" shows "h \ get_root_node ptr' \\<^sub>r root_ptr \ h \ get_root_node ptr'' \\<^sub>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' \\<^sub>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'' \\<^sub>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 \\<^sub>r Some ptr" shows "h \ get_root_node (cast child) \\<^sub>r root \ h \ get_root_node ptr \\<^sub>r root" proof assume 1: " h \ get_root_node (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r root" show "h \ get_root_node ptr \\<^sub>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 \\<^sub>r root" show "h \ get_root_node (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>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 \\<^sub>r cast child" shows "h \ get_parent child \\<^sub>r None" proof (insert assms(1) assms(4), induct ptr rule: heap_wellformed_induct_rev) case (step c) then show ?case proof (cases "cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>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 \\<^sub>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 \\<^sub>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' \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r ancestors \ h' \ get_ancestors ptr \\<^sub>r ancestors' \ (\p children children'. h \ get_child_nodes p \\<^sub>r children \ h' \ get_child_nodes p \\<^sub>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 \\<^sub>r ancestors \ cast child_node \ set ancestors \ h \ get_parent child_node \\<^sub>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 \\<^sub>r ancestors \ type_wf h \ known_ptrs h \ (\children ancestor_child . h \ get_child_nodes ancestor \\<^sub>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 \\<^sub>r ancestors \ known_ptrs h \ type_wf h \ (ptr, child) \ (parent_child_rel h)\<^sup>* \ 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 \\<^sub>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 \\<^sub>r ancestors \ ptr' \ set ancestors \ ptr'' \ set ancestors \ h \ get_root_node ptr' \\<^sub>r root_ptr \ h \ get_root_node ptr'' \\<^sub>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 \\<^sub>r cast child \ h \ get_parent child \\<^sub>r None" assumes get_root_node_parent_same: "h \ get_parent child \\<^sub>r Some ptr \ h \ get_root_node (cast child) \\<^sub>r root \ h \ get_root_node ptr \\<^sub>r root" interpretation i_get_root_node_wf?: l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent + l_get_parent_wf + l_heap_is_wellformed (* l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M *) 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 \\<^sub>r nodes" and "node \ set nodes" and "h \ get_child_nodes ptr \\<^sub>r children" and "node \ ptr" obtains child child_to where "child \ set children" and "h \ to_tree_order (cast child) \\<^sub>r child_to" and "node \ set child_to" proof - obtain treeorders where treeorders: "h \ map_M to_tree_order (map cast children) \\<^sub>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 (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r nodes" and "h \ get_child_nodes ptr \\<^sub>r children" and "node \ set children" and "h \ to_tree_order (cast node) \\<^sub>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) \\<^sub>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 \\<^sub>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 \\<^sub>r nodes" and "node \ set nodes" and "h \ to_tree_order node \\<^sub>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 \\<^sub>r nodes \ (\node. node \ set nodes \ (\nodes'. h \ to_tree_order node \\<^sub>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 \\<^sub>r children \ child \ set children \ \nodes. h \ to_tree_order (cast child) \\<^sub>r nodes \ (\node. node \ set nodes \ (\nodes'. h \ to_tree_order node \\<^sub>r nodes' \ set nodes' \ set nodes)))" and 2: "h \ to_tree_order parent \\<^sub>r nodes" and 3: "node \ set nodes" and "h \ to_tree_order node \\<^sub>r nodes'" and "x \ set nodes'" have h1: "(\children child nodes node nodes'. h \ get_child_nodes parent \\<^sub>r children \ child \ set children \ h \ to_tree_order (cast child) \\<^sub>r nodes \ (node \ set nodes \ (h \ to_tree_order node \\<^sub>r nodes' \ set nodes' \ set nodes)))" using 1 by blast obtain children where children: "h \ get_child_nodes parent \\<^sub>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 \\<^sub>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 \\<^sub>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) \\<^sub>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 \\<^sub>r nodes'\ children by blast moreover have "set nodes_of_child \ set nodes" using "2" \child \ set children\ \h \ to_tree_order (cast child) \\<^sub>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 \\<^sub>r nodes" assumes "h \ get_parent child \\<^sub>r Some parent" assumes "parent \ set nodes" shows "cast child \ set nodes" proof - obtain nodes' where nodes': "h \ to_tree_order parent \\<^sub>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 \\<^sub>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 (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>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 \\<^sub>r nodes" assumes "h \ get_child_nodes parent \\<^sub>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 \\<^sub>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 \\<^sub>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 (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \\<^sub>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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.child_parent_dual l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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 \cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>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 \\<^sub>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 \\<^sub>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 (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \\<^sub>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 \\<^sub>r nodes" assumes "cast child \ ptr" assumes "cast child \ set nodes" obtains parent where "h \ get_parent child \\<^sub>r Some parent" and "parent \ set nodes" proof - assume 1: "(\parent. h \ get_parent child \\<^sub>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 \\<^sub>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 \\<^sub>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 (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r c) \\<^sub>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 \\<^sub>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 \cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child \ set child_to\ assms(1) assms(2) assms(3) cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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 \\<^sub>r to" shows "(ptr, child) \ (parent_child_rel h)\<^sup>* \ child \ set to" proof assume 3: "(ptr, child) \ (parent_child_rel h)\<^sup>*" 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)\<^sup>*" and "(child_parent, child) \ (parent_child_rel h)" using \ptr \ child\ by (metis "1.prems" rtranclE) obtain child_node where child_node: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>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)\<^sup>*\ 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)\<^sup>*" 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: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>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)\<^sup>*" 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) done declare l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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 \\<^sub>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 \\<^sub>r to \ (ptr, child_ptr) \ (parent_child_rel h)\<^sup>* \ child_ptr \ set to" assumes to_tree_order_child2: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ to_tree_order ptr \\<^sub>r nodes \ cast child \ ptr \ cast child \ set nodes \ (\parent. h \ get_parent child \\<^sub>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 \\<^sub>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 \\<^sub>r nodes \ h \ get_child_nodes parent \\<^sub>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 \\<^sub>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 \\<^sub>r nodes \ h \ get_parent child \\<^sub>r Some parent \ parent \ set nodes \ cast child \ set nodes" assumes to_tree_order_subset: "heap_is_wellformed h \ h \ to_tree_order ptr \\<^sub>r nodes \ node \ set nodes \ h \ to_tree_order node \\<^sub>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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + 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 \\<^sub>r to" assumes "ptr' \ set to" assumes "h \ get_root_node ptr' \\<^sub>r root_ptr" assumes "ptr'' \ set to" shows "h \ get_root_node ptr'' \\<^sub>r root_ptr" proof - obtain ancestors' where ancestors': "h \ get_ancestors ptr' \\<^sub>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' \\<^sub>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 \\<^sub>r root_ptr" using \h \ get_root_node ptr' \\<^sub>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'' \\<^sub>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'' \\<^sub>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 \\<^sub>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 \\<^sub>r root_ptr" assumes "h \ to_tree_order root_ptr \\<^sub>r to" assumes "ptr' \ set to" shows "h \ get_root_node ptr' \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r root_ptr" using assms(4) using \cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>r Some parent\ \cast child_node = child\ by (metis False assms(1) assms(2) assms(3) get_parent_child_dual) then show ?thesis using \cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child_node = child\ \h \ get_parent child_node \\<^sub>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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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 \\<^sub>r to \ ptr' \ set to \ h \ get_root_node ptr' \\<^sub>r root_ptr \ ptr'' \ set to \ h \ get_root_node ptr'' \\<^sub>r root_ptr" assumes to_tree_order_same_root: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_root_node ptr \\<^sub>r root_ptr \ h \ to_tree_order root_ptr \\<^sub>r to \ ptr' \ set to \ h \ get_root_node ptr' \\<^sub>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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_known_ptrs + l_heap_is_wellformed + l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_ancestors + l_get_ancestors_wf + l_get_parent + l_get_parent_wf + l_get_root_node_wf + l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_owner_document_disconnected_nodes: assumes "heap_is_wellformed h" assumes "h \ get_disconnected_nodes document_ptr \\<^sub>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) \\<^sub>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_ptr\set |h \ document_ptr_kinds_M|\<^sub>r. node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>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 \\<^sub>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) \\<^sub>r cast node_ptr" using 2 get_root_node_no_parent by blast obtain document_ptrs where document_ptrs: "h \ document_ptr_kinds_M \\<^sub>r document_ptrs" by simp then have "h \ ok (filter_M (\document_ptr. do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) \ cast ` set disconnected_nodes) }) document_ptrs \\<^sub>r candidates" by auto have eq: "\document_ptr. document_ptr |\| document_ptr_kinds h \ node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r \ |h \ do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) \ cast ` set disconnected_nodes) }|\<^sub>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 (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr \ cast ` set disconnected_nodes) }|\<^sub>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 (((cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)) \ cast ` set disconnected_nodes) }) document_ptrs \\<^sub>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_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r document_ptr" by(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: option.splits)[1] moreover have "known_ptr (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>r None" and "h \ get_owner_document (cast node_ptr) \\<^sub>r owner_document" and "h \ get_disconnected_nodes owner_document \\<^sub>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) \\<^sub>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|\<^sub>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_ptr\set |h \ document_ptr_kinds_M|\<^sub>r. node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r" by (metis DocumentMonad.ptr_kinds_ptr_kinds_M assms(1) finite_set_in heap_is_wellformed_children_disc_nodes) then obtain document_ptr where document_ptr: "document_ptr\set |h \ document_ptr_kinds_M|\<^sub>r" and node_ptr_in_disc_nodes: "node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>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 \\<^sub>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 () \\<^sub>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_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then show "owner_document |\| document_ptr_kinds h" by(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 4: "is_character_data_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 5: "h \ Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then obtain root where root: "h \ get_root_node ptr \\<^sub>r root" by(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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|\<^sub>r)" by (metis (no_types, 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 notin_fset 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|\<^sub>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 \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ` set disconnected_nodes))) (sorted_list_of_set (fset (document_ptr_kinds h))) \\<^sub>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 notin_fset 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|\<^sub>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 \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>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_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 4: "h \ Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then obtain root where root: "h \ get_root_node ptr \\<^sub>r root" by(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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|\<^sub>r)" by (metis (no_types, 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 notin_fset 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|\<^sub>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 \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ` set disconnected_nodes))) (sorted_list_of_set (fset (document_ptr_kinds h))) \\<^sub>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 notin_fset 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|\<^sub>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 \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ` set |h \ get_disconnected_nodes some_owner_document|\<^sub>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_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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 \\<^sub>r children" assumes "child \ set children" shows "h \ get_owner_document ptr \\<^sub>r owner_document \ h \ get_owner_document (cast child) \\<^sub>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 \\<^sub>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) \\<^sub>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 \\<^sub>r owner_document \ h \ a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \\<^sub>r owner_document" proof (cases "is_document_ptr ptr") case True then obtain document_ptr where document_ptr: "cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr () \\<^sub>r owner_document \ h \ a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \\<^sub>r owner_document" using document_ptr \h \ get_root_node (cast child) \\<^sub>r root\[simplified \root = cast document_ptr\ document_ptr] apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 dest!: bind_returns_result_E3[rotated, OF \h \ get_root_node (cast child) \\<^sub>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: "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document \ h \ a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \\<^sub>r owner_document" using root \h \ get_root_node (cast child) \\<^sub>r root\ unfolding a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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 "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \\<^sub>r disc_nodes \ node_ptr \ set disc_nodes \ h \ get_owner_document (cast node_ptr) \\<^sub>r document_ptr" assumes in_disconnected_nodes_no_parent: "heap_is_wellformed h \ h \ get_parent node_ptr \\<^sub>r None\ h \ get_owner_document (cast node_ptr) \\<^sub>r owner_document \ h \ get_disconnected_nodes owner_document \\<^sub>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 \\<^sub>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 \\<^sub>r children \ child \ set children \ h \ get_owner_document ptr \\<^sub>r owner_document \ h \ get_owner_document (cast child) \\<^sub>r owner_document" interpretation i_get_owner_document_wf?: l_get_owner_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_owner_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_root_node_wf + l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + 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 \\<^sub>r root" assumes "is_document_ptr_kind root" shows "h \ get_owner_document ptr \\<^sub>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_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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 \\<^sub>r root\ assms(5) by(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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 \\<^sub>r root" shows "h \ get_owner_document ptr \\<^sub>r owner_document \ h \ get_owner_document root \\<^sub>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 = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr" by (metis node_ptr_casts_commute3) show ?thesis proof assume "h \ get_owner_document ptr \\<^sub>r owner_document" then have "h \ local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>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 \\<^sub>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 \h \ get_owner_document ptr \\<^sub>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_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root\ apply blast using \root |\| object_ptr_kinds h\ by(auto simp add: a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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 = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr" by (metis node_ptr_casts_commute3) then have "h \ local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr () \\<^sub>r owner_document" using \h \ local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document\ assms(4) apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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 \\<^sub>r owner_document" show "h \ get_owner_document ptr \\<^sub>r owner_document" proof (cases "is_document_ptr_kind root") case True have "root = cast owner_document" using \h \ get_owner_document root \\<^sub>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_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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 = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr" by (metis node_ptr_casts_commute3) then have "h \ local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr () \\<^sub>r owner_document" using \h \ get_owner_document root \\<^sub>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_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \\<^sub>r owner_document" apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_owner_document_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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 \\<^sub>r root \ is_document_ptr_kind root \ h \ get_owner_document ptr \\<^sub>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 \\<^sub>r root \ h \ get_owner_document ptr \\<^sub>r owner_document \ h \ get_owner_document root \\<^sub>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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_attribute\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + 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 \\<^sub>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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + 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 \\<^sub>h h2" and known_ptrs: "known_ptrs h" and type_wf: "type_wf h" shows "h2 \ get_parent child \\<^sub>r None" proof - obtain children where children: "h \ get_child_nodes ptr \\<^sub>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 \\<^sub>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) \\<^sub>r owner_document" and disc_nodes: "h \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" and h': "h \ set_disconnected_nodes owner_document (child # disc_nodes) \\<^sub>h h'" and h2: "h' \ set_child_nodes ptr (remove1 child children) \\<^sub>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|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r other_children" have "h \ get_child_nodes other_ptr \\<^sub>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 \\<^sub>r other_children\ a1 h1 by blast qed then have "\other_ptr other_children. other_ptr \ ptr \ h2 \ get_child_nodes other_ptr \\<^sub>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 \\<^sub>r other_children" have "h' \ get_child_nodes other_ptr \\<^sub>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 \\<^sub>r other_children\ \ child \ set other_children\ a1 by blast qed ultimately have ha: "\other_ptr other_children. h2 \ get_child_nodes other_ptr \\<^sub>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 \\<^sub>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 \\<^sub>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 "cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child |\| object_ptr_kinds h" using get_owner_document_ptr_in_heap owner_document by blast then show "h2 \ check_in_heap (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>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 \\<^sub>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))) \\<^sub>r []" by(auto intro!: filter_M_empty_I bind_pure_I) qed qed end locale l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_remove_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma remove_child_parent_child_rel_subset: assumes "heap_is_wellformed h" and "h \ remove_child ptr child \\<^sub>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 (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r owner_document" and children_h: "h \ get_child_nodes ptr \\<^sub>r children_h" and child_in_children_h: "child \ set children_h" and disconnected_nodes_h: "h \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h" and h2: "h \ set_disconnected_nodes owner_document (child # disconnected_nodes_h) \\<^sub>h h2" and h': "h2 \ set_child_nodes ptr (remove1 child children_h) \\<^sub>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 \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" unfolding object_ptr_kinds_M_defs by simp then have object_ptr_kinds_eq2: "|h \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" using select_result_eq by force then have node_ptr_kinds_eq2: "|h \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>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|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>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' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>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 \\<^sub>r disconnected_nodes = h' \ get_disconnected_nodes document_ptr \\<^sub>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|\<^sub>r = |h' \ get_disconnected_nodes document_ptr|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes ptr \\<^sub>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 \\<^sub>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 \\<^sub>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 (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r owner_document" and children_h: "h \ get_child_nodes ptr \\<^sub>r children_h" and child_in_children_h: "child \ set children_h" and disconnected_nodes_h: "h \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h" and h2: "h \ set_disconnected_nodes owner_document (child # disconnected_nodes_h) \\<^sub>h h2" and h': "h2 \ set_child_nodes ptr (remove1 child children_h) \\<^sub>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 \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" unfolding object_ptr_kinds_M_defs by simp then have object_ptr_kinds_eq2: "|h \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" using select_result_eq by force then have node_ptr_kinds_eq2: "|h \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>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|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>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' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>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 \\<^sub>r disconnected_nodes = h' \ get_disconnected_nodes document_ptr \\<^sub>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|\<^sub>r = |h' \ get_disconnected_nodes document_ptr|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes ptr \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>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, lifting) \type_wf h'\ assms(2) assms(3) local.get_child_nodes_ok local.known_ptrs_known_ptr local.remove_child_children_subset notin_fset object_ptr_kinds_eq3 returns_result_select_result subset_code(1) type_wf) apply (metis (no_types, lifting) assms(2) disconnected_nodes_eq2 disconnected_nodes_h disconnected_nodes_h' document_ptr_kinds_eq3 finite_set_in local.remove_child_child_in_heap node_ptr_kinds_eq3 select_result_I2 set_ConsD subset_code(1)) done moreover have "a_owner_document_valid h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_owner_document_valid h'" apply(auto simp add: a_owner_document_valid_def object_ptr_kinds_eq3 document_ptr_kinds_eq3 node_ptr_kinds_eq3)[1] proof - fix node_ptr assume 0: "\node_ptr\fset (node_ptr_kinds h'). (\document_ptr. document_ptr |\| document_ptr_kinds h' \ node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r) \ (\parent_ptr. parent_ptr |\| object_ptr_kinds h' \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>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|\<^sub>r" then show "\document_ptr. document_ptr |\| document_ptr_kinds h' \ node_ptr \ set |h' \ get_disconnected_nodes document_ptr|\<^sub>r" proof (cases "node_ptr = child") case True show ?thesis apply(rule exI[where x=owner_document]) using children_eq2 disconnected_nodes_eq2 children_h children_h' disconnected_nodes_h' True by (metis (no_types, lifting) get_disconnected_nodes_ptr_in_heap is_OK_returns_result_I list.set_intros(1) select_result_I2) next case False then show ?thesis using 0 1 2 children_eq2 children_h children_h' disconnected_nodes_eq2 disconnected_nodes_h disconnected_nodes_h' apply(auto simp add: children_eq2 disconnected_nodes_eq2 dest!: select_result_I2)[1] by (metis children_eq2 disconnected_nodes_eq2 finite_set_in in_set_remove1 list.set_intros(2)) qed qed moreover { have h0: "a_distinct_lists h" using assms(1) by (simp add: heap_is_wellformed_def) moreover have ha1: "(\x\set |h \ object_ptr_kinds_M|\<^sub>r. set |h \ get_child_nodes x|\<^sub>r) \ (\x\set |h \ document_ptr_kinds_M|\<^sub>r. set |h \ get_disconnected_nodes x|\<^sub>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|\<^sub>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|\<^sub>r" using ha1 ha2 ha3 apply(simp) using IntI by fastforce moreover have "distinct |h \ object_ptr_kinds_M|\<^sub>r" apply(rule select_result_I) by(auto simp add: object_ptr_kinds_M_defs) moreover have "distinct |h \ document_ptr_kinds_M|\<^sub>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|\<^sub>r" assume 1: "a_distinct_lists h" and 3: "distinct |h \ object_ptr_kinds_M|\<^sub>r" have 4: "distinct (concat ((map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) |h \ object_ptr_kinds_M|\<^sub>r)))" using 1 by(auto simp add: a_distinct_lists_def) show "distinct (concat (map (\ptr. |h' \ get_child_nodes ptr|\<^sub>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|\<^sub>r" using 4 distinct_concat_map_E object_ptr_kinds_eq2 by fastforce obtain children where children: "h \ get_child_nodes x \\<^sub>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 \\<^sub>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|\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r children_x'" using children_eq children_h' children_x by fastforce obtain children_y' where children_y': "h' \ get_child_nodes y \\<^sub>r children_y'" using children_eq children_h' children_y by fastforce have "distinct (concat (map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) |h \ object_ptr_kinds_M|\<^sub>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|\<^sub>r \ set |h' \ get_child_nodes y|\<^sub>r = {}" using children_x' children_y' by (metis (no_types, lifting) select_result_I2) qed next assume 2: "distinct |h \ document_ptr_kinds_M|\<^sub>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|\<^sub>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|\<^sub>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|\<^sub>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|\<^sub>r" proof (cases "x = owner_document") case True have "child \ set |h \ get_disconnected_nodes x|\<^sub>r" using child_not_in document_ptr_kinds_eq2 "4" by fastforce moreover have "|h' \ get_disconnected_nodes x|\<^sub>r = child # |h \ get_disconnected_nodes x|\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>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|\<^sub>r) |h \ document_ptr_kinds_M|\<^sub>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|\<^sub>r \ set |h' \ get_disconnected_nodes y|\<^sub>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|\<^sub>r" and 3: "xb \ fset (document_ptr_kinds h')" and 4: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" obtain disc_nodes where disc_nodes: "h \ get_disconnected_nodes xb \\<^sub>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 \\<^sub>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 \\<^sub>r children" by (metis "1" type_wf assms(3) finite_set_in get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr object_ptr_kinds_eq3) obtain children' where children': "h' \ get_child_nodes xa \\<^sub>r children'" using children children_eq children_h' by fastforce have "\x. x \ set |h \ get_child_nodes xa|\<^sub>r \ x \ set |h \ get_disconnected_nodes xb|\<^sub>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|\<^sub>r = children'" using children' by simp have 7: "|h' \ get_disconnected_nodes xb|\<^sub>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|\<^sub>r = remove1 child children_h" using children_h' - bysimp + 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|\<^sub>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 \\<^sub>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 \\<^sub>h h'" and children: "h' \ get_child_nodes ptr \\<^sub>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 (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r owner_document" and children_h: "h \ get_child_nodes ptr' \\<^sub>r children_h" and child_in_children_h: "child \ set children_h" and disconnected_nodes_h: "h \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h" and h2: "h \ set_disconnected_nodes owner_document (child # disconnected_nodes_h) \\<^sub>h h2" and h': "h2 \ set_child_nodes ptr' (remove1 child children_h) \\<^sub>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 \\<^sub>r node_ptr # children" assumes "h \ remove_child ptr node_ptr \\<^sub>h h'" shows "h' \ get_child_nodes ptr \\<^sub>r children" proof - obtain h2 disc_nodes owner_document where "h \ get_owner_document (cast node_ptr) \\<^sub>r owner_document" and "h \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" and h2: "h \ set_disconnected_nodes owner_document (node_ptr # disc_nodes) \\<^sub>h h2" and "h2 \ set_child_nodes ptr children \\<^sub>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 \\<^sub>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_nodes\h2 \ set_child_nodes ptr children \\<^sub>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 \\<^sub>r node_ptr # children" assumes "h \ remove node_ptr \\<^sub>h h'" shows "h' \ get_child_nodes ptr \\<^sub>r children" proof - have "h \ get_parent node_ptr \\<^sub>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 \\<^sub>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 \\<^sub>r children" assumes "h \ forall_M remove children \\<^sub>h h'" shows "h' \ get_child_nodes ptr \\<^sub>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 \\<^sub>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 \\<^sub>r children \ h \ forall_M remove children \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r [])" and 1: "heap_is_wellformed h" and 2: "type_wf h" and 3: "known_ptrs h" and 4: "h \ get_child_nodes ptr \\<^sub>r a # children" and 5: "h \ get_parent a \\<^sub>r Some ptr" and 7: "h \ remove a \\<^sub>h h2" and 8: "h2 \ forall_M remove children \\<^sub>h h'" then have "h2 \ get_child_nodes ptr \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>h h' \ heap_is_wellformed h'" assumes remove_preserves_type_wf: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ remove child \\<^sub>h h' \ type_wf h'" assumes remove_preserves_known_ptrs: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ remove child \\<^sub>h h' \ known_ptrs h'" assumes remove_heap_is_wellformed_preserved: "type_wf h \ known_ptrs h \ heap_is_wellformed h \ h \ remove child \\<^sub>h h' \ heap_is_wellformed h'" assumes remove_child_removes_child: "heap_is_wellformed h \ h \ remove_child ptr' child \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>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 \\<^sub>r node_ptr # children \ h \ remove_child ptr node_ptr \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r children" assumes remove_removes_child: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r node_ptr # children \ h \ remove node_ptr \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r children" assumes remove_for_all_empty_children: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_child_nodes ptr \\<^sub>r children \ h \ forall_M remove children \\<^sub>h h' \ h' \ get_child_nodes ptr \\<^sub>r []" interpretation i_remove_child_wf2?: l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + 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 \\<^sub>h h'" assumes "h \ get_child_nodes ptr' \\<^sub>r node # children" shows "h' \ get_child_nodes ptr' \\<^sub>r children" proof - obtain old_document parent_opt h2 where old_document: "h \ get_owner_document (cast node) \\<^sub>r old_document" and parent_opt: "h \ get_parent node \\<^sub>r parent_opt" and h2: "h \ (case parent_opt of Some parent \ do { remove_child parent node } | None \ do { return ()}) \\<^sub>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 () }) \\<^sub>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' \\<^sub>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) \\<^sub>r old_document" and parent_opt: "h \ get_parent node \\<^sub>r parent_opt" and h2: "h \ (case parent_opt of Some parent \ do { remove_child parent node } | None \ do { return ()}) \\<^sub>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 () }) \\<^sub>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 \\<^sub>r old_disc_nodes" and h3: "h2 \ set_disconnected_nodes old_document (remove1 node old_disc_nodes) \\<^sub>h h3" and old_disc_nodes: "h3 \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" and h': "h3 \ set_disconnected_nodes owner_document (node # disc_nodes) \\<^sub>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_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_root_node + l_get_owner_document_wf + l_remove_child_wf2 + l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma adopt_node_removes_child_step: assumes wellformed: "heap_is_wellformed h" and adopt_node: "h \ adopt_node owner_document node_ptr \\<^sub>h h2" and children: "h2 \ get_child_nodes ptr \\<^sub>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) \\<^sub>r old_document" and parent_opt: "h \ get_parent node_ptr \\<^sub>r parent_opt" and h': "h \ (case parent_opt of Some parent \ remove_child parent node_ptr | None \ return () ) \\<^sub>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 \\<^sub>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 \\<^sub>r children\ returns_result_eq by fastforce next case (Some option) then show ?case using remove_child_removes_child \h' \ get_child_nodes ptr \\<^sub>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 \\<^sub>h h'" shows "\ptr' children'. h' \ get_child_nodes ptr' \\<^sub>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 \\<^sub>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) \\<^sub>r old_document" and parent_opt: "h \ get_parent child \\<^sub>r parent_opt" and h2: "h \ (case parent_opt of Some parent \ remove_child parent child | None \ return ()) \\<^sub>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 () }) \\<^sub>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 \\<^sub>r ptrs = h2 \ object_ptr_kinds_M \\<^sub>r ptrs" unfolding object_ptr_kinds_M_defs by simp then have object_ptr_kinds_eq_h: "|h \ object_ptr_kinds_M|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq_h: "|h \ node_ptr_kinds_M|\<^sub>r = |h2 \ node_ptr_kinds_M|\<^sub>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 \\<^sub>r old_disc_nodes" and h3: "h2 \ set_disconnected_nodes old_document (remove1 child old_disc_nodes) \\<^sub>h h3" and disc_nodes_document_ptr_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_document_ptr_h3" and h': "h3 \ set_disconnected_nodes document_ptr (child # disc_nodes_document_ptr_h3) \\<^sub>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 \\<^sub>r ptrs = h3 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq_h2: "|h2 \ object_ptr_kinds_M|\<^sub>r = |h3 \ object_ptr_kinds_M|\<^sub>r" by(simp) then have node_ptr_kinds_eq_h2: "|h2 \ node_ptr_kinds_M|\<^sub>r = |h3 \ node_ptr_kinds_M|\<^sub>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|\<^sub>r = |h3 \ document_ptr_kinds_M|\<^sub>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 \\<^sub>r children = h3 \ get_child_nodes ptr \\<^sub>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|\<^sub>r = |h3 \ get_child_nodes ptr|\<^sub>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 \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq_h3: "|h3 \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by(simp) then have node_ptr_kinds_eq_h3: "|h3 \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>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|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>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 \\<^sub>r children = h' \ get_child_nodes ptr \\<^sub>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|\<^sub>r = |h' \ get_child_nodes ptr|\<^sub>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 \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>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 \\<^sub>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|\<^sub>r) |h2 \ document_ptr_kinds_M|\<^sub>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, 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 finite_set_in 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 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|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 2: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h2)))))" and 3: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h2). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" show "distinct (concat (map (\document_ptr. |h' \ get_disconnected_nodes document_ptr|\<^sub>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|\<^sub>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|\<^sub>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|\<^sub>r \ set |h2 \ get_disconnected_nodes y|\<^sub>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|\<^sub>r \ set |h' \ get_disconnected_nodes y|\<^sub>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|\<^sub>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|\<^sub>r \ set disc_nodes_old_document_h2 = {}" using \set |h2 \ get_disconnected_nodes y|\<^sub>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|\<^sub>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|\<^sub>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|\<^sub>r \ set |h2 \ get_disconnected_nodes old_document|\<^sub>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 fmember.rep_eq 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|\<^sub>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|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 1: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h2)))))" and 2: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h2). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" and 3: "xa |\| object_ptr_kinds h'" and 4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 5: "xb |\| document_ptr_kinds h'" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>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 \\<^sub>r disc_nodes_old_document_h2" assume a3: "h3 \ get_disconnected_nodes old_document \\<^sub>r remove1 child disc_nodes_old_document_h2" assume a4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" assume "document_ptr_kinds h2 = document_ptr_kinds h'" assume a5: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h'). set |h2 \ get_disconnected_nodes x|\<^sub>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|\<^sub>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|\<^sub>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|\<^sub>r \ set |h2 \ get_disconnected_nodes xb|\<^sub>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 \\<^sub>r disc_nodes_document_ptr_h3" assume a3: "h2 \ get_disconnected_nodes old_document \\<^sub>r disc_nodes_old_document_h2" assume a4: "xa |\| object_ptr_kinds h'" assume a5: "h' \ get_disconnected_nodes document_ptr \\<^sub>r child # disc_nodes_document_ptr_h3" assume a6: "old_document |\| document_ptr_kinds h'" assume a7: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" assume a8: "x \ set |h' \ get_child_nodes xa|\<^sub>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|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" assume a11: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" assume a12: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h'). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" have f13: "\d. d \ set |h' \ document_ptr_kinds_M|\<^sub>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|\<^sub>r = disc_nodes_old_document_h2" using a6 a3 by simp have "x \ set |h2 \ get_disconnected_nodes xb|\<^sub>r" using a12 a8 a4 \xb |\| document_ptr_kinds h'\ by (meson UN_I disjoint_iff_not_equal fmember.rep_eq) 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 \\<^sub>h h'" and "h' \ get_disconnected_nodes owner_document \\<^sub>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) \\<^sub>r old_document" and parent_opt: "h \ get_parent node_ptr \\<^sub>r parent_opt" and h2: "h \ (case parent_opt of Some parent \ remove_child parent node_ptr | None \ return ()) \\<^sub>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 () }) \\<^sub>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 \\<^sub>r disc_nodes" assume a2: "\h document_ptr disc_nodes h'. h \ set_disconnected_nodes document_ptr disc_nodes \\<^sub>h h' \ h' \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" assume "h'a \ set_disconnected_nodes owner_document (node_ptr # xb) \\<^sub>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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] interpretation i_adopt_node_wf2?: l_adopt_node_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_adopt_node_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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 \\<^sub>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 \\<^sub>h h2 \ h2 \ get_child_nodes ptr \\<^sub>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 \\<^sub>h h' \ h' \ get_disconnected_nodes owner_document \\<^sub>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 \\<^sub>h h' \ h \ get_child_nodes ptr' \\<^sub>r node # children \ h' \ get_child_nodes ptr' \\<^sub>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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + 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 \\<^sub>h h'" assumes "h \ get_child_nodes ptr' \\<^sub>r node # children" shows "h' \ get_child_nodes ptr' \\<^sub>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) \\<^sub>r reference_child" and "h \ get_owner_document ptr \\<^sub>r owner_document" and h2: "h \ adopt_node owner_document node \\<^sub>h h2" and "h2 \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" and h3: "h2 \ set_disconnected_nodes owner_document (remove1 node disc_nodes) \\<^sub>h h3" and h': "h3 \ a_insert_node ptr node reference_child \\<^sub>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' \\<^sub>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' \\<^sub>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 \\<^sub>h h' \ h \ get_child_nodes ptr' \\<^sub>r node # children \ h' \ get_child_nodes ptr' \\<^sub>r children" interpretation i_insert_before_wf?: l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + 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_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma insert_before_heap_is_wellformed_preserved: assumes wellformed: "heap_is_wellformed h" and insert_before: "h \ insert_before ptr node child \\<^sub>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 \\<^sub>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) \\<^sub>r reference_child" and owner_document: "h \ get_owner_document ptr \\<^sub>r owner_document" and h2: "h \ adopt_node owner_document node \\<^sub>h h2" and disconnected_nodes_h2: "h2 \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h2" and h3: "h2 \ set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \\<^sub>h h3" and h': "h3 \ a_insert_node ptr node reference_child \\<^sub>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 \\<^sub>r ptrs = h2 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs ) then have object_ptr_kinds_M_eq2_h: "|h \ object_ptr_kinds_M|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h: "|h \ node_ptr_kinds_M|\<^sub>r = |h2 \ node_ptr_kinds_M|\<^sub>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 \\<^sub>r ptrs = h3 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_M_eq2_h2: "|h2 \ object_ptr_kinds_M|\<^sub>r = |h3 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h2: "|h2 \ node_ptr_kinds_M|\<^sub>r = |h3 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have document_ptr_kinds_eq2_h2: "|h2 \ document_ptr_kinds_M|\<^sub>r = |h3 \ document_ptr_kinds_M|\<^sub>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 \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_M_eq2_h3: "|h3 \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h3: "|h3 \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have document_ptr_kinds_eq2_h3: "|h3 \ document_ptr_kinds_M|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>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 \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disconnected_nodes_h3: "h3 \ get_disconnected_nodes owner_document \\<^sub>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 \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have children_eq_h3: "\ptr' children. ptr \ ptr' \ h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force obtain children_h3 where children_h3: "h3 \ get_child_nodes ptr \\<^sub>r children_h3" using h' a_insert_node_def by auto have children_h': "h' \ get_child_nodes ptr \\<^sub>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 \\<^sub>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|\<^sub>r" proof - fix d assume "d |\| document_ptr_kinds h3" show "node \ set |h3 \ get_disconnected_nodes d|\<^sub>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|\<^sub>r \ set |h2 \ get_disconnected_nodes owner_document|\<^sub>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 \\<^sub>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 \\<^sub>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)\<^sup>*}" 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)\<^sup>*}" 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, lifting) disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 document_ptr_kinds_commutes finite_set_in 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|\<^sub>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|\<^sub>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 finite_set_in 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 finite_set_in select_result_I2 subsetD) by (metis (no_types) \type_wf h'\ disconnected_nodes_eq2_h3 disconnected_nodes_eq_h3 finite_set_in 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|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" show "distinct |h3 \ get_disconnected_nodes x|\<^sub>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 fmember.rep_eq set_sorted_list_of_set) next fix x y xa assume 1: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>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|\<^sub>r" and 6: "xa \ set |h3 \ get_disconnected_nodes y|\<^sub>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 finite_fset fmember.rep_eq notin_set_remove1 select_result_I2 set_sorted_list_of_set by (metis (no_types, lifting)) qed qed next fix x xa xb assume 1: "(\x\fset (object_ptr_kinds h3). set |h3 \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h3). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" and 2: "xa |\| object_ptr_kinds h3" and 3: "x \ set |h3 \ get_child_nodes xa|\<^sub>r" and 4: "xb |\| document_ptr_kinds h3" and 5: "x \ set |h3 \ get_disconnected_nodes xb|\<^sub>r" have 6: "set |h3 \ get_child_nodes xa|\<^sub>r \ set |h2 \ get_disconnected_nodes xb|\<^sub>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|\<^sub>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|\<^sub>r" using 1 by (auto elim: distinct_concat_map_E) show "distinct |h' \ get_child_nodes x|\<^sub>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|\<^sub>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|\<^sub>r" and 6: "xa \ set |h' \ get_child_nodes y|\<^sub>r" have 7:"set |h3 \ get_child_nodes x|\<^sub>r \ set |h3 \ get_child_nodes y|\<^sub>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: " (\x\fset (object_ptr_kinds h'). set |h3 \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h'). set |h' \ get_disconnected_nodes x|\<^sub>r) = {} " and 2: "xa |\| object_ptr_kinds h'" and 3: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 4: "xb |\| document_ptr_kinds h'" and 5: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" have 6: "set |h3 \ get_child_nodes xa|\<^sub>r \ set |h' \ get_disconnected_nodes xb|\<^sub>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 \\<^sub>r |h' \ get_disconnected_nodes xb|\<^sub>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 children_eq2_h3 children_h' children_h3 disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 finite_set_in 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 \\<^sub>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 \\<^sub>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 \\<^sub>h h' \ heap_is_wellformed h'" interpretation i_insert_before_wf2?: l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + 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 \\<^sub>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 \\<^sub>r xs" assumes "h \ append_child ptr node \\<^sub>h h'" assumes "node \ set xs" shows "h' \ get_child_nodes ptr \\<^sub>r xs @ [node]" proof - obtain ancestors owner_document h2 h3 disconnected_nodes_h2 where ancestors: "h \ get_ancestors ptr \\<^sub>r ancestors" and node_not_in_ancestors: "cast node \ set ancestors" and owner_document: "h \ get_owner_document ptr \\<^sub>r owner_document" and h2: "h \ adopt_node owner_document node \\<^sub>h h2" and disconnected_nodes_h2: "h2 \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h2" and h3: "h2 \ set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \\<^sub>h h3" and h': "h3 \ a_insert_node ptr node None \\<^sub>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|\<^sub>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 \\<^sub>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|\<^sub>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 \\<^sub>r xs" using get_child_nodes_reads set_disconnected_nodes_writes h3 \h2 \ get_child_nodes ptr \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r xs" assumes "h \ forall_M (append_child ptr) nodes \\<^sub>h h'" assumes "set nodes \ set xs = {}" assumes "distinct nodes" shows "h' \ get_child_nodes ptr \\<^sub>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 \\<^sub>r xs \ h \ forall_M (append_child ptr) nodes \\<^sub>h h' \ set nodes \ set xs = {} \ h' \ get_child_nodes ptr \\<^sub>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 \\<^sub>r xs" and 5: "h \ append_child ptr a \\<^sub>r ()" and 6: "h \ append_child ptr a \\<^sub>h h'a" and 7: "h'a \ forall_M (append_child ptr) nodes \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r []" assumes "h \ forall_M (append_child ptr) nodes \\<^sub>h h'" assumes "distinct nodes" shows "h' \ get_child_nodes ptr \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>h h' \ heap_is_wellformed h'" interpretation i_append_child_wf?: l_append_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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 \\<^sub>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 \\<^sub>r new_element_ptr" and h2: "h \ new_element \\<^sub>h h2" and h3: "h2 \ set_tag_name new_element_ptr tag \\<^sub>h h3" and disc_nodes_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" and h': "h3 \ set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes_h3) \\<^sub>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 \\<^sub>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|\<^sub>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|\<^sub>r" by simp then have "cast new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>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 \\<^sub>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' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes (cast new_element_ptr) \\<^sub>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 \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>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' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>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 \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>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 \\<^sub>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 \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>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|\<^sub>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|\<^sub>r" then show "x \ set |h2 \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq_h \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>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 \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ children_eq2_h3 disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 finite_set_in h' is_OK_returns_result_I l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes local.a_all_ptrs_in_heap_def local.get_child_nodes_ptr_in_heap local.l_set_disconnected_nodes_get_disconnected_nodes_axioms node_ptr_kinds_commutes object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1)) have "\p. p |\| object_ptr_kinds h \ cast new_element_ptr \ set |h \ get_child_nodes p|\<^sub>r" using \heap_is_wellformed h\ \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>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|\<^sub>r" using children_eq2_h apply(auto simp add: object_ptr_kinds_eq_h)[1] using \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ apply auto[1] by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\) then have "\p. p |\| object_ptr_kinds h3 \ cast new_element_ptr \ set |h3 \ get_child_nodes p|\<^sub>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|\<^sub>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) \\<^sub>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|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" then show "distinct |h' \ get_disconnected_nodes x|\<^sub>r" using document_ptr_kinds_eq_h3 disconnected_nodes_eq_h3 h' set_disconnected_nodes_get_disconnected_nodes by (metis (no_types, lifting) \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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|\<^sub>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|\<^sub>r" and "xa \ set |h' \ get_disconnected_nodes y|\<^sub>r" moreover have "set |h3 \ get_disconnected_nodes x|\<^sub>r \ set |h3 \ get_disconnected_nodes y|\<^sub>r = {}" using calculation by(auto dest: distinct_concat_map_E(1)) ultimately show "False" apply(-) apply(cases "x = document_ptr") apply (smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \local.a_all_ptrs_in_heap h\ disc_nodes_h3 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms select_result_I2 set_ConsD subsetD) by (smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \local.a_all_ptrs_in_heap h\ disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms select_result_I2 set_ConsD subsetD) next fix x xa xb assume 2: "(\x\fset (object_ptr_kinds h3). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h3). set |h3 \ get_disconnected_nodes x|\<^sub>r) = {}" and 3: "xa |\| object_ptr_kinds h3" and 4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 5: "xb |\| document_ptr_kinds h3" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>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 \ cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h3 \ get_child_nodes p|\<^sub>r\ \a_distinct_lists h3\ children_eq2_h3 disc_nodes_h3 distinct_lists_no_parent h' select_result_I2 set_ConsD set_disconnected_nodes_get_disconnected_nodes) by (metis "3" "4" "5" "6" \a_distinct_lists h3\ \type_wf h3\ children_eq2_h3 distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result) qed have "a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) then have "a_owner_document_valid h'" using disc_nodes_h3 \document_ptr |\| document_ptr_kinds h\ apply(auto simp add: a_owner_document_valid_def)[1] apply(auto simp add: object_ptr_kinds_eq_h object_ptr_kinds_eq_h3 )[1] apply(auto simp add: object_ptr_kinds_eq_h2)[1] apply(auto simp add: document_ptr_kinds_eq_h document_ptr_kinds_eq_h3 )[1] apply(auto simp add: document_ptr_kinds_eq_h2)[1] apply(auto simp add: node_ptr_kinds_eq_h node_ptr_kinds_eq_h3 )[1] apply(auto simp add: node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h )[1] apply(auto simp add: children_eq2_h2[symmetric] children_eq2_h3[symmetric] disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3)[1] apply (metis (no_types, lifting) document_ptr_kinds_eq_h h' list.set_intros(1) local.set_disconnected_nodes_get_disconnected_nodes select_result_I2) apply(simp add: object_ptr_kinds_eq_h) by(metis (no_types, lifting) NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ children_eq2_h children_eq2_h2 children_eq2_h3 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes list.set_intros(2) local.l_set_disconnected_nodes_get_disconnected_nodes_axioms node_ptr_kinds_commutes select_result_I2) 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsection \create\_character\_data\ locale l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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 \\<^sub>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 \\<^sub>r new_character_data_ptr" and h2: "h \ new_character_data \\<^sub>h h2" and h3: "h2 \ set_val new_character_data_ptr text \\<^sub>h h3" and disc_nodes_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" and h': "h3 \ set_disconnected_nodes document_ptr (cast new_character_data_ptr # disc_nodes_h3) \\<^sub>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 \\<^sub>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|\<^sub>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|\<^sub>r" by simp then have "cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>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 \\<^sub>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' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>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' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>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 \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>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' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>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 \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>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 \\<^sub>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|\<^sub>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|\<^sub>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|\<^sub>r" then show "x \ set |h2 \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq_h \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>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|\<^sub>r\ \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ apply (metis (no_types, lifting) NodeMonad.ptr_kinds_ptr_kinds_M \parent_child_rel h = parent_child_rel h2\ children_eq2_h finite_set_in 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 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 finite_set_in 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|\<^sub>r" using \heap_is_wellformed h\ \cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>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|\<^sub>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) \\<^sub>r []\ apply auto[1] by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\) then have "\p. p |\| object_ptr_kinds h3 \ cast new_character_data_ptr \ set |h3 \ get_child_nodes p|\<^sub>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|\<^sub>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) \\<^sub>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|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h3)))))" and "x |\| document_ptr_kinds h3" then show "distinct |h' \ get_disconnected_nodes x|\<^sub>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|\<^sub>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|\<^sub>r" and "xa \ set |h' \ get_disconnected_nodes y|\<^sub>r" moreover have "set |h3 \ get_disconnected_nodes x|\<^sub>r \ set |h3 \ get_disconnected_nodes y|\<^sub>r = {}" using calculation by(auto dest: distinct_concat_map_E(1)) ultimately show "False" by (smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \local.a_all_ptrs_in_heap h\ disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms select_result_I2 set_ConsD subsetD) next fix x xa xb assume 2: "(\x\fset (object_ptr_kinds h3). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h3). set |h3 \ get_disconnected_nodes x|\<^sub>r) = {}" and 3: "xa |\| object_ptr_kinds h3" and 4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 5: "xb |\| document_ptr_kinds h3" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>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|\<^sub>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, lifting) \cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h finite_set_in 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsection \create\_document\ locale l_create_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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 \\<^sub>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 \\<^sub>r new_document_ptr" and h': "h \ new_document \\<^sub>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|\<^sub>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|\<^sub>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 apply(auto simp add: document_ptr_kinds_def)[1] by (metis (no_types, lifting) document_ptr_kinds_commutes document_ptr_kinds_def finsertI1 fset.map_comp) have children_eq: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_document_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h' \ get_child_nodes (cast new_document_ptr) \\<^sub>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 \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>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 (metis(full_types) \\thesis. (\new_document_ptr. \h \ new_document \\<^sub>r new_document_ptr; h \ new_document \\<^sub>h h'\ \ thesis) \ thesis\ local.get_disconnected_nodes_new_document_different_pointers new_document_ptr)+ then have disconnected_nodes_eq2_h: "\doc_ptr. doc_ptr \ new_document_ptr \ |h \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have "h' \ get_disconnected_nodes new_document_ptr \\<^sub>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|\<^sub>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|\<^sub>r" then show "x \ set |h' \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_document_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2) next fix a x assume 0: "a |\| object_ptr_kinds h'" and 1: "x \ set |h' \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq \h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h'" and 1: "x \ set |h' \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h' \ get_child_nodes (cast new_document_ptr) \\<^sub>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 \cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr \ set |h \ object_ptr_kinds_M|\<^sub>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, lifting) \h' \ get_child_nodes (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr) \\<^sub>r []\ children_eq2 finite_set_in finsert_iff funion_finsert_right object_ptr_kinds_eq select_result_I2 subsetD sup_bot.right_neutral) by (metis (no_types, lifting) \cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr |\| object_ptr_kinds h\ \h' \ get_child_nodes (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr) \\<^sub>r []\ \h' \ get_disconnected_nodes new_document_ptr \\<^sub>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 \\<^sub>r []\ \h' \ get_child_nodes (cast new_document_ptr) \\<^sub>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|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h)))))" assume a7: "xa \ set |h' \ get_disconnected_nodes x|\<^sub>r" assume a8: "xa \ set |h' \ get_disconnected_nodes y|\<^sub>r" have f9: "xa \ set |h \ get_disconnected_nodes x|\<^sub>r" using a7 a3 disconnected_nodes_eq2_h by presburger have f10: "xa \ set |h \ get_disconnected_nodes y|\<^sub>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 \\<^sub>r []" and 1: "h' \ get_child_nodes (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr) \\<^sub>r []" and 2: "distinct (concat (map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h)))))" and 3: "distinct (concat (map (\document_ptr. |h \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h)))))" and 4: "(\x\fset (object_ptr_kinds h). set |h \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h). set |h \ get_disconnected_nodes x|\<^sub>r) = {}" and 5: "x \ set |h \ get_child_nodes xa|\<^sub>r" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" and 7: "xa |\| object_ptr_kinds h" and 8: "xa \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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 \cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_document_ptr |\| object_ptr_kinds h\ children_eq2 disconnected_nodes_eq2_h document_ptr_kinds_commutes finite_set_in 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_create_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] end diff --git a/thys/List_Update/List_Factoring.thy b/thys/List_Update/List_Factoring.thy --- a/thys/List_Update/List_Factoring.thy +++ b/thys/List_Update/List_Factoring.thy @@ -1,2389 +1,2389 @@ (* Title: List Factoring Author: Max Haslbeck *) section "List factoring technique" theory List_Factoring imports Partial_Cost_Model MTF2_Effects begin hide_const config compet subsection "Helper functions" subsubsection "Helper lemmas" lemma befaf: assumes "q\set s" "distinct s" shows "before q s \ {q} \ after q s = set s" proof - have "before q s \ {y. index s y = index s q \ q \ set s} = {y. index s y \ index s q \ q \ set s}" unfolding before_in_def apply(auto) by (simp add: le_neq_implies_less) also have "\ = {y. index s y \ index s q \ y\ set s \ q \ set s}" apply(auto) by (metis index_conv_size_if_notin index_less_size_conv not_less) also with \q \ set s\ have "\ = {y. index s y \ index s q \ y\ set s}" by auto finally have "before q s \ {y. index s y = index s q \ q \ set s} \ after q s = {y. index s y \ index s q \ y\ set s} \ {y. index s y > index s q \ y \ set s}" unfolding before_in_def by simp also have "\ = set s" by auto finally show ?thesis using assms by simp qed lemma index_sum: assumes "distinct s" "q\set s" shows "index s q = (\e\set s. if e < q in s then 1 else 0)" proof - from assms have bia_empty: "before q s \ ({q} \ after q s) = {}" by(auto simp: before_in_def) from befaf[OF assms(2) assms(1)] have "(\e\set s. if e < q in s then 1::nat else 0) = (\e\(before q s \ {q} \ after q s). if e < q in s then 1 else 0)" by auto also have "\ = (\e\before q s. if e < q in s then 1 else 0) + (\e\{q}. if e < q in s then 1 else 0) + (\e\after q s. if e < q in s then 1 else 0)" proof - have "(\e\(before q s \ {q} \ after q s). if e < q in s then 1::nat else 0) = (\e\(before q s \ ({q} \ after q s)). if e < q in s then 1::nat else 0)" by simp also have "\ = (\e\before q s. if e < q in s then 1 else 0) + (\e\({q} \ after q s). if e < q in s then 1 else 0) - (\e\(before q s \ ({q} \ after q s)). if e < q in s then 1 else 0)" apply(rule sum_Un_nat) by(simp_all) also have "\ = (\e\before q s. if e < q in s then 1 else 0) + (\e\({q} \ after q s). if e < q in s then 1 else 0)" using bia_empty by auto also have "\ = (\e\before q s. if e < q in s then 1 else 0) + (\e\{q}. if e < q in s then 1 else 0) + (\e\after q s. if e < q in s then 1 else 0)" by (simp add: before_in_def) finally show ?thesis . qed also have "\ = (\e\before q s. 1) + (\e\({q} \ after q s). 0)" apply(auto) unfolding before_in_def by auto also have "\ = card (before q s)" by auto also have "\ = card (set (take (index s q) s))" using before_conv_take[OF assms(2)] by simp also have "\ = length (take (index s q) s)" using distinct_card assms(1) distinct_take by metis also have "\ = min (length s) (index s q)" by simp also have "\ = index s q" using index_le_size[of s q] by(auto) finally show ?thesis by simp qed subsubsection "ALG" fun ALG :: "'a \ 'a list \ nat \ ('a list * 'is) \ nat" where "ALG x qs i s = (if x < (qs!i) in fst s then 1::nat else 0)" (* no paid exchanges, requested items in state (nice, quickcheck is awesome!) *) lemma t\<^sub>p_sumofALG: "distinct (fst s) \ snd a = [] \ (qs!i)\set (fst s) \ t\<^sub>p (fst s) (qs!i) a = (\e\set (fst s). ALG e qs i s)" unfolding t\<^sub>p_def apply(simp add: split_def ) using index_sum by metis lemma t\<^sub>p_sumofALGreal: assumes "distinct (fst s)" "snd a = []" "qs!i \ set(fst s)" shows "real(t\<^sub>p (fst s) (qs!i) a) = (\e\set (fst s). real(ALG e qs i s))" proof - from assms have "real(t\<^sub>p (fst s) (qs!i) a) = real(\e\set (fst s). ALG e qs i s)" using t\<^sub>p_sumofALG by metis also have "\ = (\e\set (fst s). real (ALG e qs i s))" by auto finally show ?thesis . qed subsubsection "The function steps'" fun steps' where "steps' s _ _ 0 = s" | "steps' s [] [] (Suc n) = s" | "steps' s (q#qs) (a#as) (Suc n) = steps' (step s q a) qs as n" lemma steps'_steps: "length as = length qs \ steps' s as qs (length as) = steps s as qs" by(induct arbitrary: s rule: list_induct2, simp_all) lemma steps'_length: "length qs = length as \ n \ length as \ length (steps' s qs as n) = length s" apply(induct qs as arbitrary: s n rule: list_induct2) apply(simp) apply(case_tac n) by (auto) lemma steps'_set: "length qs = length as \ n \ length as \ set (steps' s qs as n) = set s" apply(induct qs as arbitrary: s n rule: list_induct2) apply(simp) apply(case_tac n) by(auto simp: set_step) lemma steps'_distinct2: "length qs = length as \ n \ length as \ distinct s \ distinct (steps' s qs as n)" apply(induct qs as arbitrary: s n rule: list_induct2) apply(simp) apply(case_tac n) by(auto simp: distinct_step) lemma steps'_distinct: "length qs = length as \ length as = n \ distinct (steps' s qs as n) = distinct s" by (induct qs as arbitrary: s n rule: list_induct2) (fastforce simp add: distinct_step)+ lemma steps'_dist_perm: "length qs = length as \ length as = n \ dist_perm s s \ dist_perm (steps' s qs as n) (steps' s qs as n)" using steps'_set steps'_distinct by blast lemma steps'_rests: "length qs = length as \ n \ length as \ steps' s qs as n = steps' s (qs@r1) (as@r2) n" apply(induct qs as arbitrary: s n rule: list_induct2) apply(simp) apply(case_tac n) by auto lemma steps'_append: "length qs = length as \ length qs = n \ steps' s (qs@[q]) (as@[a]) (Suc n) = step (steps' s qs as n) q a" apply(induct qs as arbitrary: s n rule: list_induct2) by auto subsubsection "\ALG'_det\" definition "ALG'_det Strat qs init i x = ALG x qs i (swaps (snd (Strat!i)) (steps' init qs Strat i),())" lemma ALG'_det_append: "n < length Strat \ n < length qs \ ALG'_det Strat (qs@a) init n x = ALG'_det Strat qs init n x" proof - assume qs: "n < length qs" assume S: "n < length Strat" have tt: "(qs @ a) ! n = qs ! n" using qs by (simp add: nth_append) have "steps' init (take n qs) (take n Strat) n = steps' init ((take n qs) @ drop n qs) ((take n Strat) @ (drop n Strat)) n" apply(rule steps'_rests) using S qs by auto then have A: "steps' init (take n qs) (take n Strat) n = steps' init qs Strat n" by auto have "steps' init (take n qs) (take n Strat) n = steps' init ((take n qs) @ ((drop n qs)@a)) ((take n Strat) @((drop n Strat)@[])) n" apply(rule steps'_rests) using S qs by auto then have B: "steps' init (take n qs) (take n Strat) n = steps' init (qs@a) (Strat@[]) n" by (metis append_assoc List.append_take_drop_id) from A B have "steps' init qs Strat n = steps' init (qs@a) (Strat@[]) n" by auto then have C: "steps' init qs Strat n = steps' init (qs@a) Strat n" by auto show ?thesis unfolding ALG'_det_def C unfolding ALG.simps tt by auto qed subsubsection "ALG'" abbreviation "config'' A qs init n == config_rand A init (take n qs)" definition "ALG' A qs init i x = E( map_pmf (ALG x qs i) (config'' A qs init i))" lemma ALG'_refl: "qs!i = x \ ALG' A qs init i x = 0" unfolding ALG'_def by(simp add: split_def before_in_def) subsubsection "\ALGxy_det\" definition ALGxy_det where "ALGxy_det A qs init x y = (\i\{.. {y,x}) then ALG'_det A qs init i y + ALG'_det A qs init i x else 0::nat))" lemma ALGxy_det_alternativ: "ALGxy_det A qs init x y = (\i\{i. i (qs!i \ {y,x})}. ALG'_det A qs init i y + ALG'_det A qs init i x)" proof - have f: "{i. i (qs!i \ {y,x})} = {i. i {i. (qs!i \ {y,x})}" by auto have "(\i\{i. i (qs!i \ {y,x})}. ALG'_det A qs init i y + ALG'_det A qs init i x) = (\i\{i. i {i. (qs!i \ {y,x})}. ALG'_det A qs init i y + ALG'_det A qs init i x)" unfolding e by simp also have "\ = (\i\{i. i {i. (qs!i \ {y,x})} then ALG'_det A qs init i y + ALG'_det A qs init i x else 0))" apply(rule sum.inter_restrict) by auto also have "\ = (\i\{.. {i. (qs!i \ {y,x})} then ALG'_det A qs init i y + ALG'_det A qs init i x else 0))" unfolding f by auto also have "\ = ALGxy_det A qs init x y" unfolding ALGxy_det_def by auto finally show ?thesis by simp qed subsubsection "ALGxy" definition ALGxy where "ALGxy A qs init x y = (\i\{.. {i. (qs!i \ {y,x})}. ALG' A qs init i y + ALG' A qs init i x)" lemma ALGxy_def2: "ALGxy A qs init x y = (\i\{i. i (qs!i \ {y,x})}. ALG' A qs init i y + ALG' A qs init i x)" proof - have a: "{i. i (qs!i \ {y,x})} = {.. {i. (qs!i \ {y,x})}" by auto show ?thesis unfolding ALGxy_def a by simp qed lemma ALGxy_append: "ALGxy A (rs@[r]) init x y = ALGxy A rs init x y + (if (r \ {y,x}) then ALG' A (rs@[r]) init (length rs) y + ALG' A (rs@[r]) init (length rs) x else 0 )" proof - have "ALGxy A (rs@[r]) init x y = (\i\{..<(Suc (length rs))} \ {i. (rs @ [r]) ! i \ {y, x}}. ALG' A (rs @ [r]) init i y + ALG' A (rs @ [r]) init i x)" unfolding ALGxy_def by(simp) also have "\ = (\i\{..<(Suc (length rs))}. (if i\{i. (rs @ [r]) ! i \ {y, x}} then ALG' A (rs @ [r]) init i y + ALG' A (rs @ [r]) init i x else 0) )" apply(rule sum.inter_restrict) by simp also have "\ = (\i\{..{i. (rs @ [r]) ! i \ {y, x}} then ALG' A (rs @ [r]) init i y + ALG' A (rs @ [r]) init i x else 0) ) + (if length rs\{i. (rs @ [r]) ! i \ {y, x}} then ALG' A (rs @ [r]) init (length rs) y + ALG' A (rs @ [r]) init(length rs) x else 0) " by simp also have "\ = ALGxy A rs init x y + (if r \ {y, x} then ALG' A (rs @ [r]) init (length rs) y + ALG' A (rs @ [r]) init(length rs) x else 0)" apply(simp add: ALGxy_def sum.inter_restrict nth_append) unfolding ALG'_def apply(rule sum.cong) apply(simp) by(auto simp: nth_append) finally show ?thesis . qed lemma ALGxy_wholerange: "ALGxy A qs init x y = (\i<(length qs). (if qs ! i \ {y, x} then ALG' A qs init i y + ALG' A qs init i x else 0 ))" proof - have "ALGxy A qs init x y = (\i\ {i. i < length qs} \ {i. qs ! i \ {y, x}}. ALG' A qs init i y + ALG' A qs init i x)" unfolding ALGxy_def apply(rule sum.cong) apply(simp) apply(blast) by simp also have "\ = (\i\{i. i < length qs}. if i \ {i. qs ! i \ {y, x}} then ALG' A qs init i y + ALG' A qs init i x else 0)" by(rule sum.inter_restrict) simp also have "\ = (\i<(length qs). (if qs ! i \ {y, x} then ALG' A qs init i y + ALG' A qs init i x else 0 ))" apply(rule sum.cong) by(auto) finally show ?thesis . qed subsection "Transformation to Blocking Cost" lemma umformung: fixes A :: "(('a::linorder) list,'is,'a,(nat * nat list)) alg_on_rand" assumes no_paid: "\is s q. \((free,paid),_) \ (snd A (s,is) q). paid=[]" assumes inlist: "set qs \ set init" assumes dist: "distinct init" assumes "\x. x < length qs \ finite (set_pmf (config'' A qs init x))" shows "T\<^sub>p_on_rand A init qs = (\(x,y)\{(x,y). x \ set init \ y\set init \ xn. \xa \ set_pmf (config'' A qs init n). distinct (fst xa)" using dist config_rand_distinct by metis have E0: "T\<^sub>p_on_rand A init qs = (\i\{..p_on_rand_n A init qs i)" unfolding T_on_rand_as_sum by auto also have "\ = (\is. bind_pmf (snd A s (qs ! i)) (\(a, nis). return_pmf (real (\x\set init. ALG x qs i s))))))" apply(rule sum.cong) apply(simp) apply(simp add: bind_return_pmf bind_assoc_pmf) apply(rule arg_cong[where f=E]) apply(rule bind_pmf_cong) apply(simp) apply(rule bind_pmf_cong) apply(simp) apply(simp add: split_def) apply(subst t\<^sub>p_sumofALGreal) proof (goal_cases) case 1 then show ?case using config_dist by(metis) next case (2 a b c) then show ?case using no_paid[of "fst b" "snd b"] by(auto simp add: split_def) next case (3 a b c) with config_rand_set have a: "set (fst b) = set init" by metis with inlist have " set qs \ set (fst b)" by auto with 3 show ?case by auto next case (4 a b c) with config_rand_set have a: "set (fst b) = set init" by metis then show ?case by(simp) qed (* hier erst s, dann init *) also have "\ = (\i(is, s). (real (\x\set init. ALG x qs i (is,s)))) (config'' A qs init i)))" apply(simp only: map_pmf_def split_def) by simp also have E1: "\ = (\ix\set init. ALG' A qs init i x))" apply(rule sum.cong) apply(simp) apply(simp add: split_def ALG'_def) apply(rule E_linear_sum_allg) by(rule assms(4)) also have E2: "\ = (\x\set init. (\i = (\x\set init. (\y\set init. (\i\{i. i qs!i=y}. ALG' A qs init i x)))" proof (rule sum.cong, goal_cases) case (2 x) have "(\i = sum (%i. ALG' A qs init i x) (\y\{y. y \ set init}. {i. i < length qs \ qs ! i = y})" apply(rule sum.cong) apply(auto) using inlist by auto also have "\ = sum (%t. sum (%i. ALG' A qs init i x) {i. i qs ! i = t}) {y. y\ set init}" apply(rule sum.UNION_disjoint) apply(simp_all) by force also have "\ = (\y\set init. \i | i < length qs \ qs ! i = y. ALG' A qs init i x)" by auto finally show ?case . qed (simp) also have "\ = (\(x,y)\ (set init \ set init). (\i\{i. i qs!i=y}. ALG' A qs init i x))" by (rule sum.cartesian_product) also have "\ = (\(x,y)\ {(x,y). x\set init \ y\ set init}. (\i\{i. i qs!i=y}. ALG' A qs init i x))" by simp also have E4: "\ = (\(x,y)\{(x,y). x\set init \ y\ set init \ x\y}. (\i\{i. i qs!i=y}. ALG' A qs init i x))" (is "(\(x,y)\ ?L. ?f x y) = (\(x,y)\ ?R. ?f x y)") proof - let ?M = "{(x,y). x\set init \ y\ set init \ x=y}" have A: "?L = ?R \ ?M" by auto have B: "{} = ?R \ ?M" by auto have "(\(x,y)\ ?L. ?f x y) = (\(x,y)\ ?R \ ?M. ?f x y)" by(simp only: A) also have "\ = (\(x,y)\ ?R. ?f x y) + (\(x,y)\ ?M. ?f x y)" apply(rule sum.union_disjoint) apply(rule finite_subset[where B="set init \ set init"]) apply(auto) apply(rule finite_subset[where B="set init \ set init"]) by(auto) also have "(\(x,y)\ ?M. ?f x y) = 0" apply(rule sum.neutral) by (auto simp add: ALG'_refl) finally show ?thesis by simp qed also have "\ = (\(x,y)\{(x,y). x \ set init \ y\set init \ xi\{i. i qs!i=y}. ALG' A qs init i x) + (\i\{i. i qs!i=x}. ALG' A qs init i y) )" (is "(\(x,y)\ ?L. ?f x y) = (\(x,y)\ ?R. ?f x y + ?f y x)") proof - let ?R' = "{(x,y). x \ set init \ y\set init \ y ?R'" by auto have "{} = ?R \ ?R'" by auto have C: "?R' = (%(x,y). (y, x)) ` ?R" by auto have D: "(\(x,y)\ ?R'. ?f x y) = (\(x,y)\ ?R. ?f y x)" proof - have "(\(x,y)\ ?R'. ?f x y) = (\(x,y)\ (%(x,y). (y, x)) ` ?R. ?f x y)" by(simp only: C) also have "(\z\ (%(x,y). (y, x)) ` ?R. (%(x,y). ?f x y) z) = (\z\?R. ((%(x,y). ?f x y) \ (%(x,y). (y, x))) z)" apply(rule sum.reindex) by(fact swap_inj_on) also have "\ = (\z\?R. (%(x,y). ?f y x) z)" apply(rule sum.cong) by(auto) finally show ?thesis . qed have "(\(x,y)\ ?L. ?f x y) = (\(x,y)\ ?R \ ?R'. ?f x y)" by(simp only: A) also have "\ = (\(x,y)\ ?R. ?f x y) + (\(x,y)\ ?R'. ?f x y)" apply(rule sum.union_disjoint) apply(rule finite_subset[where B="set init \ set init"]) apply(auto) apply(rule finite_subset[where B="set init \ set init"]) by(auto) also have "\ = (\(x,y)\ ?R. ?f x y) + (\(x,y)\ ?R. ?f y x)" by(simp only: D) also have "\ = (\(x,y)\ ?R. ?f x y + ?f y x)" by(simp add: split_def sum.distrib[symmetric]) finally show ?thesis . qed also have E5: "\ = (\(x,y)\{(x,y). x \ set init \ y\set init \ xi\{i. i (qs!i=y \ qs!i=x)}. ALG' A qs init i y + ALG' A qs init i x))" apply(rule sum.cong) apply(simp) proof goal_cases case (1 x) then obtain a b where x: "x=(a,b)" and a: "a \ set init" "b \ set init" "a < b" by auto then have "a\b" by simp then have disj: "{i. i < length qs \ qs ! i = b} \ {i. i < length qs \ qs ! i = a} = {}" by auto have unio: "{i. i < length qs \ (qs ! i = b \ qs ! i = a)} = {i. i < length qs \ qs ! i = b} \ {i. i < length qs \ qs ! i = a}" by auto have "(\i\{i. i < length qs \ qs ! i = b} \ {i. i < length qs \ qs ! i = a}. ALG' A qs init i b + ALG' A qs init i a) = (\i\{i. i < length qs \ qs ! i = b}. ALG' A qs init i b + ALG' A qs init i a) + (\i\ {i. i < length qs \ qs ! i = a}. ALG' A qs init i b + ALG' A qs init i a) - (\i\{i. i < length qs \ qs ! i = b} \ {i. i < length qs \ qs ! i = a}. ALG' A qs init i b + ALG' A qs init i a) " apply(rule sum_Un) by(auto) also have "\ = (\i\{i. i < length qs \ qs ! i = b}. ALG' A qs init i b + ALG' A qs init i a) + (\i\ {i. i < length qs \ qs ! i = a}. ALG' A qs init i b + ALG' A qs init i a)" using disj by auto also have "\ = (\i\{i. i < length qs \ qs ! i = b}. ALG' A qs init i a) + (\i\{i. i < length qs \ qs ! i = a}. ALG' A qs init i b)" by (auto simp: ALG'_refl) finally show ?case unfolding x apply(simp add: split_def) unfolding unio by simp qed also have E6: "\ = (\(x,y)\{(x,y). x \ set init \ y\set init \ xy" shows "(if (x < y in l) then 0 else 1) = index l x" unfolding before_in_def proof (auto, goal_cases) (* bad style! *) case 1 from assms(1) have "index l y < length l" by simp with assms(2) 1(1) show "index l x = 0" by auto next case 2 from assms(1) have a: "index l x < length l" by simp from assms(1,3) have "index l y \ index l x" by simp with assms(2) 2(1) a show "Suc 0 = index l x" by simp qed (simp add: assms) lemma before_in_index2: fixes l assumes "set l = {x,y}" and "length l = 2" and "x\y" shows "(if (x < y in l) then 1 else 0) = index l y" unfolding before_in_def proof (auto, goal_cases) (* bad style! *) case 2 from assms(1,3) have a: "index l y \ index l x" by simp from assms(1) have "index l x < length l" by simp with assms(2) a 2(1) show "index l y = 0" by auto next case 1 from assms(1) have a: "index l y < length l" by simp from assms(1,3) have "index l y \ index l x" by simp with assms(2) 1(1) a show "Suc 0 = index l y" by simp qed (simp add: assms) lemma before_in_index: fixes l assumes "set l = {x,y}" and "length l = 2" and "x\y" shows "(x < y in l) = (index l x = 0)" unfolding before_in_def proof (safe, goal_cases) case 1 from assms(1) have "index l y < length l" by simp with assms(2) 1(1) show "index l x = 0" by auto next case 2 from assms(1,3) have "index l y \ index l x" by simp with 2(1) show "index l x < index l y" by simp qed (simp add: assms) subsection "The pairwise property" definition pairwise where "pairwise A = (\init. distinct init \ (\qs\{xs. set xs \ set init}. \(x::('a::linorder),y)\{(x,y). x \ set init \ y\set init \ xp_on_rand A (Lxy init {x,y}) (Lxy qs {x,y}) = ALGxy A qs init x y))" definition "Pbefore_in x y A qs init = map_pmf (\p. x < y in fst p) (config_rand A init qs)" lemma T_on_n_no_paid: assumes nopaid: "\s n. map_pmf (\x. snd (fst x)) (snd A s n) = return_pmf []" shows "T_on_rand_n A init qs i = E (config'' A qs init i \ (\p. return_pmf (real(index (fst p) (qs ! i)))))" proof - have "(\s. snd A s (qs ! i) \ (\(a, is'). return_pmf (real (t\<^sub>p (fst s) (qs ! i) a)))) = (\s. (snd A s (qs ! i) \ (\x. return_pmf (snd (fst x)))) \ (\p. return_pmf (real (index (swaps p (fst s)) (qs ! i)) + real (length p))))" by(simp add: t\<^sub>p_def split_def bind_return_pmf bind_assoc_pmf) also have "\ = (\p. return_pmf (real (index (fst p) (qs ! i))))" using nopaid[unfolded map_pmf_def] by(simp add: split_def bind_return_pmf) finally show ?thesis by simp qed lemma pairwise_property_lemma: assumes relativeorder: "(\init qs. distinct init \ qs \ {xs. set xs \ set init} \ (\x y. (x,y)\ {(x,y). x \ set init \ y\set init \ x\y} \ x \ y \ Pbefore_in x y A qs init = Pbefore_in x y A (Lxy qs {x,y}) (Lxy init {x,y}) ))" and nopaid: "\xa r. \z\ set_pmf(snd A xa r). snd(fst z) = []" shows "pairwise A" unfolding pairwise_def proof (clarify, goal_cases) case (1 init rs x y) then have xny: "x\y" by auto note dinit=1(1) then have dLyx: "distinct (Lxy init {y,x})" by(rule Lxy_distinct) from dinit have dLxy: "distinct (Lxy init {x,y})" by(rule Lxy_distinct) have setLxy: "set (Lxy init {x, y}) = {x,y}" apply(subst Lxy_set_filter) using 1 by auto have setLyx: "set (Lxy init {y, x}) = {x,y}" apply(subst Lxy_set_filter) using 1 by auto have lengthLyx:" length (Lxy init {y, x}) = 2" using setLyx distinct_card[OF dLyx] xny by simp have lengthLxy:" length (Lxy init {x, y}) = 2" using setLxy distinct_card[OF dLxy] xny by simp have aee: "{x,y} = {y,x}" by auto from 1(2) show ?case proof(induct rs rule: rev_induct) case (snoc r rs) have b: "Pbefore_in x y A rs init = Pbefore_in x y A (Lxy rs {x,y}) (Lxy init {x,y})" apply(rule relativeorder) using snoc 1 xny by(simp_all) show ?case (is "?L (rs @ [r]) = ?R (rs @ [r])") proof(cases "r\{x,y}") case True note xyrequest=this let ?expr = "E (Partial_Cost_Model.config'_rand A (fst A (Lxy init {x, y}) \ (\is. return_pmf (Lxy init {x, y}, is))) (Lxy rs {x, y}) \ (\s. snd A s r \ (\(a, is'). return_pmf (real (t\<^sub>p (fst s) r a)))))" let ?expr2 = "ALG' A (rs @ [r]) init (length rs) y + ALG' A (rs @ [r]) init (length rs) x" from xyrequest have "?L (rs @ [r]) = ?L rs + ?expr" by(simp add: Lxy_snoc T_on_rand'_append) also have "\ = ?L rs + ?expr2" proof(cases "r=x") case True let ?projS ="config'_rand A (fst A (Lxy init {x, y}) \ (\is. return_pmf (Lxy init {x, y}, is))) (Lxy rs {x, y})" let ?S = "(config'_rand A (fst A init \ (\is. return_pmf (init, is))) rs)" have "?projS \ (\s. snd A s r \ (\(a, is'). return_pmf (real (t\<^sub>p (fst s) r a)))) = ?projS \ (\s. return_pmf (real (index (fst s) r)))" proof (rule bind_pmf_cong, goal_cases) case (2 z) have "snd A z r \ (\(a, is'). return_pmf (real (t\<^sub>p (fst z) r a))) = snd A z r \ (\x. return_pmf (real (index (fst z) r)))" apply(rule bind_pmf_cong) apply(simp) using nopaid[of z r] by(simp add: split_def t\<^sub>p_def) then show ?case by(simp add: bind_return_pmf) qed simp also have "\ = map_pmf (%b. (if b then 0::real else 1)) (Pbefore_in x y A (Lxy rs {x,y}) (Lxy init {x,y}))" unfolding Pbefore_in_def map_pmf_def apply(simp add: bind_return_pmf bind_assoc_pmf) apply(rule bind_pmf_cong) apply(simp add: aee) proof goal_cases case (1 z) have " (if x < y in fst z then 0 else 1) = (index (fst z) x)" apply(rule before_in_index1) using 1 config_rand_set setLxy apply fast using 1 config_rand_length lengthLxy apply metis using xny by simp with True show ?case by(auto) qed also have "\ = map_pmf (%b. (if b then 0::real else 1)) (Pbefore_in x y A rs init)" by(simp add: b) also have "\ = map_pmf (\xa. real (if y < x in fst xa then 1 else 0)) ?S" apply(simp add: Pbefore_in_def map_pmf_comp) proof (rule map_pmf_cong, goal_cases) case (2 z) then have set_z: "set (fst z) = set init" using config_rand_set by fast have "(\ x < y in fst z) = y < x in fst z" apply(subst not_before_in) using set_z 1(3,4) xny by(simp_all) - then show ?case bysimp + then show ?case by simp qed simp finally have a: "?projS \ (\s. snd A s x \ (\(a, is'). return_pmf (real (t\<^sub>p (fst s) x a)))) = map_pmf (\xa. real (if y < x in fst xa then 1 else 0)) ?S" using True by simp from True show ?thesis apply(simp add: ALG'_refl nth_append) unfolding ALG'_def by(simp add: a) next case False with xyrequest have request: "r=y" by blast let ?projS ="config'_rand A (fst A (Lxy init {x, y}) \ (\is. return_pmf (Lxy init {x, y}, is))) (Lxy rs {x, y})" let ?S = "(config'_rand A (fst A init \ (\is. return_pmf (init, is))) rs)" have "?projS \ (\s. snd A s r \ (\(a, is'). return_pmf (real (t\<^sub>p (fst s) r a)))) = ?projS \ (\s. return_pmf (real (index (fst s) r)))" proof (rule bind_pmf_cong, goal_cases) case (2 z) have "snd A z r \ (\(a, is'). return_pmf (real (t\<^sub>p (fst z) r a))) = snd A z r \ (\x. return_pmf (real (index (fst z) r)))" apply(rule bind_pmf_cong) apply(simp) using nopaid[of z r] by(simp add: split_def t\<^sub>p_def) then show ?case by(simp add: bind_return_pmf) qed simp also have "\ = map_pmf (%b. (if b then 1::real else 0)) (Pbefore_in x y A (Lxy rs {x,y}) (Lxy init {x,y}))" unfolding Pbefore_in_def map_pmf_def apply(simp add: bind_return_pmf bind_assoc_pmf) apply(rule bind_pmf_cong) apply(simp add: aee) proof goal_cases case (1 z) have " (if x < y in fst z then 1 else 0) = (index (fst z) y)" apply(rule before_in_index2) using 1 config_rand_set setLxy apply fast using 1 config_rand_length lengthLxy apply metis using xny by simp with request show ?case by(auto) qed also have "\ = map_pmf (%b. (if b then 1::real else 0)) (Pbefore_in x y A rs init)" by(simp add: b) also have "\ = map_pmf (\xa. real (if x < y in fst xa then 1 else 0)) ?S" apply(simp add: Pbefore_in_def map_pmf_comp) apply (rule map_pmf_cong) by simp_all finally have a: "?projS \ (\s. snd A s y \ (\(a, is'). return_pmf (real (t\<^sub>p (fst s) y a)))) = map_pmf (\xa. real (if x < y in fst xa then 1 else 0)) ?S" using request by simp from request show ?thesis apply(simp add: ALG'_refl nth_append) unfolding ALG'_def by(simp add: a) qed also have "\ = ?R rs + ?expr2" using snoc by simp also from True have "\ = ?R (rs@[r])" apply(subst ALGxy_append) by(auto) finally show ?thesis . next case False then have "?L (rs @ [r]) = ?L rs" apply(subst Lxy_snoc) by simp also have "\ = ?R rs" using snoc by(simp) also have "\ = ?R (rs @ [r])" apply(subst ALGxy_append) using False by(simp) finally show ?thesis . qed qed (simp add: ALGxy_def) qed lemma umf_pair: assumes 0: "pairwise A" assumes 1: "\is s q. \((free,paid),_) \ (snd A (s, is) q). paid=[]" assumes 2: "set qs \ set init" assumes 3: "distinct init" assumes 4: "\x. x finite (set_pmf (config'' A qs init x))" shows "T\<^sub>p_on_rand A init qs = (\(x,y)\{(x, y). x \ set init \ y \ set init \ x < y}. T\<^sub>p_on_rand A (Lxy init {x,y}) (Lxy qs {x,y}))" proof - have " T\<^sub>p_on_rand A init qs = (\(x,y)\{(x, y). x \ set init \ y \ set init \ x < y}. ALGxy A qs init x y)" by(simp only: umformung[OF 1 2 3 4]) also have "\ = (\(x,y)\{(x, y). x \ set init \ y \ set init \ x < y}. T\<^sub>p_on_rand A (Lxy init {x,y}) (Lxy qs {x,y}))" apply(rule sum.cong) apply(simp) using 0[unfolded pairwise_def] 2 3 by auto finally show ?thesis . qed subsection "List Factoring for OPT" (* calculates given a list of swaps, elements x and y and a current state how many swaps between x and y there are *) fun ALG_P :: "nat list \ 'a \ 'a \ 'a list \ nat" where "ALG_P [] x y xs = (0::nat)" | "ALG_P (s#ss) x y xs = (if Suc s < length (swaps ss xs) then (if ((swaps ss xs)!s=x \ (swaps ss xs)!(Suc s)=y) \ ((swaps ss xs)!s=y \ (swaps ss xs)!(Suc s)=x) then 1 else 0) else 0) + ALG_P ss x y xs" (* nat list ersetzen durch (a::ordered) list *) lemma ALG_P_erwischt_alle: assumes dinit: "distinct init" shows "\l< length sws. Suc (sws!l) < length init \ length sws = (\(x,y)\{(x,y). x \ set (init::('a::linorder) list) \ y\set init \ xl set ?co'" by (rule nth_mem) then have a: "?co'!s \ set init" by auto from isT7 have "?co'! (Suc s) \ set ?co'" by (rule nth_mem) then have b: "?co'!(Suc s) \ set init" by auto have "{(x,y). x \ set init \ y\set init \ x {(x,y). ?expr3 x y} = {(x,y). x \ set init \ y\set init \ x (?co'!s=x \ ?co'!(Suc s)=y \ ?co'!s=y \ ?co'!(Suc s)=x)}" by auto also have "\ = {(x,y). x \ set init \ y\set init \ x ?co'!s=x \ ?co'!(Suc s)=y } \ {(x,y). x \ set init \ y\set init \ x ?co'!s=y \ ?co'!(Suc s)=x}" by auto also have "\ = {(x,y). x ?co'!s=x \ ?co'!(Suc s)=y} \ {(x,y). x ?co'!s=y \ ?co'!(Suc s)=x}" using a b by(auto) finally have c1: "{(x,y). x \ set init \ y\set init \ x {(x,y). ?expr3 x y} = {(x,y). x ?co'!s=x \ ?co'!(Suc s)=y} \ {(x,y). x ?co'!s=y \ ?co'!(Suc s)=x}" . have c2: "card ({(x,y). x ?co'!s=x \ ?co'!(Suc s)=y} \ {(x,y). x ?co'!s=y \ ?co'!(Suc s)=x}) = 1" (is "card (?A \ ?B) = 1") proof (cases "?co'!s ?B = { (?co'!s, ?co'!(Suc s)) }" apply(simp only: a b) by simp have "card (?A \ ?B) = 1" unfolding c by auto then show ?thesis . next case False then have a: "?A = {}" by auto have b: "?B = { (?co'!(Suc s), ?co'!s) } " proof - from dco distinct_conv_nth[of "?co'"] have "swaps ss init ! s \ swaps ss init ! (Suc s)" using isT2 isT3 by simp with False show ?thesis by auto qed have c: "?A \ ?B = { (?co'!(Suc s), ?co'!s) }" apply(simp only: a b) by simp have "card (?A \ ?B) = 1" unfolding c by auto then show ?thesis . qed have yeah: "(\(x,y)\{(x,y). x \ set init \ y\set init \ x(x,y)\{(x,y). x \ set init \ y\set init \ x(x,y)\{(x,y). x \ set init \ y\set init \ x = (\z\{(x,y). x \ set init \ y\set init \ x = (\z\{(x,y). x \ set init \ y\set init \ x = (\z\{(x,y). x \ set init \ y\set init \ x{(x,y). ?expr3 x y} . 1)" apply(rule sum.inter_restrict[symmetric]) apply(rule finite_subset[where B="set init \ set init"]) by(auto) also have "\ = card ({(x,y). x \ set init \ y\set init \ x {(x,y). ?expr3 x y})" by auto also have "\ = card ({(x,y). x ?co'!s=x \ ?co'!(Suc s)=y} \ {(x,y). x ?co'!s=y \ ?co'!(Suc s)=x})" by(simp only: c1) also have "\ = (1::nat)" using c2 by auto finally show ?thesis . qed have "length (s # ss) = 1 + length ss" by auto also have "\ = 1 + (\(x,y)\{(x,y). x \ set init \ y\set init \ x = (\(x,y)\{(x,y). x \ set init \ y\set init \ x(x,y)\{(x,y). x \ set init \ y\set init \ x = (\(x,y)\{(x,y). x \ set init \ y\set init \ x = (\(x,y)\{(x,y). x \ set init \ y\set init \ xp_sumofALGALGP: assumes "distinct s" "(qs!i)\set s" and "\l< length (snd a). Suc ((snd a)!l) < length s" shows "t\<^sub>p s (qs!i) a = (\e\set s. ALG e qs i (swaps (snd a) s,())) + (\(x,y)\{(x::('a::linorder),y). x \ set s \ y\set s \ x(x,y)\{(x,y). x \ set s \ y\set s \ xe\set s. ALG e qs i (swaps (snd a) s,()))" proof - have "index (swaps (snd a) s) (qs ! i) = (\e\set (swaps (snd a) s). if e < (qs ! i) in (swaps (snd a) s) then 1 else 0)" apply(rule index_sum) using assms by(simp_all) also have "\ = (\e\set s. ALG e qs i (swaps (snd a) s,()))" by auto finally show ?thesis . qed show ?thesis unfolding t\<^sub>p_def apply (simp add: split_def) unfolding ac pe by (simp add: split_def) qed (* given a Strategy Strat to serve request sequence qs on initial list init how many swaps between elements x and y occur during the ith step *) definition "ALG_P' Strat qs init i x y = ALG_P (snd (Strat!i)) x y (steps' init qs Strat i)" (* if n is in bound, Strat may be too long, that does not matter *) lemma ALG_P'_rest: "n < length qs \ n < length Strat \ ALG_P' Strat (take n qs @ [qs ! n]) init n x y = ALG_P' (take n Strat @ [Strat ! n]) (take n qs @ [qs ! n]) init n x y" proof - assume qs: "n < length qs" assume S: "n < length Strat" then have lS: "length (take n Strat) = n" by auto have "(take n Strat @ [Strat ! n]) ! n = (take n Strat @ (Strat ! n) # []) ! length (take n Strat)" using lS by auto also have "\ = Strat ! n" by(rule nth_append_length) finally have tt: "(take n Strat @ [Strat ! n]) ! n = Strat ! n" . obtain rest where rest: "Strat = (take n Strat @ [Strat ! n] @ rest)" using S apply(auto) using id_take_nth_drop by blast have "steps' init (take n qs @ [qs ! n]) (take n Strat @ [Strat ! n]) n = steps' init (take n qs) (take n Strat) n" apply(rule steps'_rests[symmetric]) using S qs by auto also have "\ = steps' init (take n qs @ [qs ! n]) (take n Strat @ ([Strat ! n] @ rest)) n" apply(rule steps'_rests) using S qs by auto finally show ?thesis unfolding ALG_P'_def tt using rest by auto qed (* verallgemeinert ALG_P'_rest, sollte mergen! *) lemma ALG_P'_rest2: "n < length qs \ n < length Strat \ ALG_P' Strat qs init n x y = ALG_P' (Strat@r1) (qs@r2) init n x y" proof - assume qs: "n < length qs" assume S: "n < length Strat" have tt: "Strat ! n = (Strat @ r1) ! n" using S by (simp add: nth_append) have "steps' init (take n qs) (take n Strat) n = steps' init ((take n qs) @ drop n qs) ((take n Strat) @ (drop n Strat)) n" apply(rule steps'_rests) using S qs by auto then have A: "steps' init (take n qs) (take n Strat) n = steps' init qs Strat n" by auto have "steps' init (take n qs) (take n Strat) n = steps' init ((take n qs) @ ((drop n qs)@r2)) ((take n Strat) @((drop n Strat)@r1)) n" apply(rule steps'_rests) using S qs by auto then have B: "steps' init (take n qs) (take n Strat) n = steps' init (qs@r2) (Strat@r1) n" by (metis append_assoc List.append_take_drop_id) from A B have C: "steps' init qs Strat n = steps' init (qs@r2) (Strat@r1) n" by auto show ?thesis unfolding ALG_P'_def tt using C by auto qed (* total number of swaps of elements x and y during execution of Strategy Strat *) definition ALG_Pxy where "ALG_Pxy Strat qs init x y = (\i b \ {x,y} \ ALGxy_det Strat (A @ [b]) init x y = ALGxy_det Strat A init x y" proof - assume bn: "b \ {x,y}" have "(A @ [b]) ! (length A) = b" by auto assume l: "length A < length Strat" term "%i. ALG'_det Strat (A @ [b]) init i y" have e: "\i. i (A @ [b]) ! i = A ! i" by(auto simp: nth_append) have "(\i\ {..< length (A @ [b])}. if (A @ [b]) ! i \ {y, x} then ALG'_det Strat (A @ [b]) init i y + ALG'_det Strat (A @ [b]) init i x else 0) = (\i\ {..< Suc(length (A))}. if (A @ [b]) ! i \ {y, x} then ALG'_det Strat (A @ [b]) init i y + ALG'_det Strat (A @ [b]) init i x else 0)" by auto also have "\ = (\i\ {..< (length (A))}. if (A @ [b]) ! i \ {y, x} then ALG'_det Strat (A @ [b]) init i y + ALG'_det Strat (A @ [b]) init i x else 0) + ( if (A @ [b]) ! (length A) \ {y, x} then ALG'_det Strat (A @ [b]) init (length A) y + ALG'_det Strat (A @ [b]) init (length A) x else 0) " by simp (* abspalten des letzten glieds *) also have "\ = (\i\ {..< (length (A))}. if (A @ [b]) ! i \ {y, x} then ALG'_det Strat (A @ [b]) init i y + ALG'_det Strat (A @ [b]) init i x else 0)" using bn by auto also have "\ = (\i\ {..< (length (A))}. if A ! i \ {y, x} then ALG'_det Strat A init i y + ALG'_det Strat A init i x else 0)" apply(rule sum.cong) apply(simp) using l ALG'_det_append[where qs=A] e by(simp) finally show ?thesis unfolding ALGxy_det_def by simp qed lemma ALG_P_split: "length qs < length Strat \ ALG_Pxy Strat (qs@[q]) init x y = ALG_Pxy Strat qs init x y + ALG_P' Strat (qs@[q]) init (length qs) x y " unfolding ALG_Pxy_def apply(auto) apply(rule sum.cong) apply(simp) using ALG_P'_rest2[symmetric, of _ qs Strat "[]" "[q]"] by(simp) lemma swap0in2: assumes "set l = {x,y}" "x\y" "length l = 2" "dist_perm l l" shows "x < y in (swap 0) l = (~ x < y in l)" proof (cases "x < y in l") case True then have a: "index l x < index l y" unfolding before_in_def by simp from assms(1) have drin: "x\set l" "y\set l" by auto from assms(1,3) have b: "index l y < 2" by simp from a b have k: "index l x = 0" "index l y = 1" by auto have g: "x = l ! 0" "y = l ! 1" using k nth_index assms(1) by force+ have "x < y in swap 0 l = (x < y in l \ \ (x = l ! 0 \ y = l ! Suc 0) \ x = l ! Suc 0 \ y = l ! 0)" apply(rule before_in_swap) apply(fact assms(4)) using assms(3) by simp also have "\ = (\ (x = l ! 0 \ y = l ! Suc 0) \ x = l ! Suc 0 \ y = l ! 0)" using True by simp also have "\ = False" using g assms(2) by auto finally have "~ x < y in (swap 0) l" by simp then show ?thesis using True by auto next case False from assms(1,2) have "index l y \ index l x" by simp with False assms(1,2) have a: "index l y < index l x" by (metis before_in_def insert_iff linorder_neqE_nat) from assms(1) have drin: "x\set l" "y\set l" by auto from assms(1,3) have b: "index l x < 2" by simp from a b have k: "index l x = 1" "index l y = 0" by auto then have g: "x = l ! 1" "y = l ! 0" using k nth_index assms(1) by force+ have "x < y in swap 0 l = (x < y in l \ \ (x = l ! 0 \ y = l ! Suc 0) \ x = l ! Suc 0 \ y = l ! 0)" apply(rule before_in_swap) apply(fact assms(4)) using assms(3) by simp also have "\ = (x = l ! Suc 0 \ y = l ! 0)" using False by simp also have "\ = True" using g by auto finally have "x < y in (swap 0) l" by simp then show ?thesis using False by auto qed lemma before_in_swap2: "dist_perm xs ys \ Suc n < size xs \ x\y \ x < y in (swap n xs) \ (~ x < y in xs \ (y = xs!n \ x = xs!Suc n) \ x < y in xs \ ~(y = xs!Suc n \ x = xs!n))" apply(simp add:before_in_def index_swap_distinct) by (metis Suc_lessD Suc_lessI index_nth_id less_Suc_eq nth_mem yes) lemma projected_paid_same_effect: assumes d: "dist_perm s1 s1" and ee: "x\y" and f: "set s2 = {x, y}" and g: "length s2 = 2" and h: "dist_perm s2 s2" shows "x < y in s1 = x < y in s2 \ x < y in swaps acs s1 = x < y in (swap 0 ^^ ALG_P acs x y s1) s2" proof (induct acs) case Nil then show ?case by auto next case (Cons s ss) from d have dd: "dist_perm (swaps ss s1) (swaps ss s1)" by simp from f have ff: "set ((swap 0 ^^ ALG_P ss x y s1) s2) = {x, y}" by (metis foldr_replicate swaps_inv) from g have gg: "length ((swap 0 ^^ ALG_P ss x y s1) s2) = 2" by (metis foldr_replicate swaps_inv) from h have hh: "dist_perm ((swap 0 ^^ ALG_P ss x y s1) s2) ((swap 0 ^^ ALG_P ss x y s1) s2)" by (metis foldr_replicate swaps_inv) show ?case (is "?LHS = ?RHS") proof (cases "Suc s < length (swaps ss s1) \ (((swaps ss s1)!s=x \ (swaps ss s1)!(Suc s)=y) \ ((swaps ss s1)!s=y \ (swaps ss s1)!(Suc s)=x))") case True from True have 1:" Suc s < length (swaps ss s1)" and 2: "(swaps ss s1 ! s = x \ swaps ss s1 ! Suc s = y \ swaps ss s1 ! s = y \ swaps ss s1 ! Suc s = x)" by auto from True have "ALG_P (s # ss) x y s1 = 1 + ALG_P ss x y s1" by auto then have "?RHS = x < y in (swap 0) ((swap 0 ^^ ALG_P ss x y s1) s2)" by auto also have "\ = (~ x < y in ((swap 0 ^^ ALG_P ss x y s1) s2))" apply(rule swap0in2) by(fact)+ also have "\ = (~ x < y in swaps ss s1)" using Cons by auto also have "\ = x < y in (swap s) (swaps ss s1)" using 1 2 before_in_swap by (metis Suc_lessD before_id dd lessI no_before_inI) (* bad *) also have "\ = ?LHS" by auto finally show ?thesis by simp next case False note F=this then have "ALG_P (s # ss) x y s1 = ALG_P ss x y s1" by auto then have "?RHS = x < y in ((swap 0 ^^ ALG_P ss x y s1) s2)" by auto also have "\ = x < y in swaps ss s1" using Cons by auto also have "\ = x < y in (swap s) (swaps ss s1)" proof (cases "Suc s < length (swaps ss s1)") case True with F have g: "swaps ss s1 ! s \ x \ swaps ss s1 ! Suc s \ y" and h: "swaps ss s1 ! s \ y \ swaps ss s1 ! Suc s \ x" by auto show ?thesis unfolding before_in_swap[OF dd True, of x y] apply(simp) using g h by auto next case False then show ?thesis unfolding swap_def by(simp) qed also have "\ = ?LHS" by auto finally show ?thesis by simp qed qed lemma steps_steps': "length qs = length as \ steps s qs as = steps' s qs as (length as)" by (induct qs as arbitrary: s rule: list_induct2) (auto) lemma T1_7': "T\<^sub>p init qs Strat = T\<^sub>p_opt init qs \ length Strat = length qs \ n\length qs \ x\(y::('a::linorder)) \ x\ set init \ y \ set init \ distinct init \ set qs \ set init \ (\Strat2 sws. \<^cancel>\T\<^sub>p_opt (Lxy init {x,y}) (Lxy (take n qs) {x,y}) \ T\<^sub>p (Lxy init {x,y}) (Lxy (take n qs) {x,y}) Strat2 \\ length Strat2 = length (Lxy (take n qs) {x,y}) \ (x < y in (steps' init (take n qs) (take n Strat) n)) = (x < y in (swaps sws (steps' (Lxy init {x,y}) (Lxy (take n qs) {x,y}) Strat2 (length Strat2)))) \ T\<^sub>p (Lxy init {x,y}) (Lxy (take n qs) {x,y}) Strat2 + length sws = ALGxy_det Strat (take n qs) init x y + ALG_Pxy Strat (take n qs) init x y)" proof(induct n) case (Suc n) from Suc(3,4) have ns: "n < length qs" by simp then have n: "n \ length qs" by simp from Suc(1)[OF Suc(2) Suc(3) n Suc(5) Suc(6) Suc(7) Suc(8) Suc(9) ] obtain Strat2 sws where (*S2: "T\<^sub>p_opt (Lxy init {x,y}) (Lxy (take n qs) {x, y}) \ T\<^sub>p (Lxy init {x,y}) (Lxy (take n qs) {x, y}) Strat2" and *) len: "length Strat2 = length (Lxy (take n qs) {x, y})" and iff: "x < y in steps' init (take n qs) (take n Strat) n = x < y in swaps sws (steps' (Lxy init {x,y}) (Lxy (take n qs) {x, y}) Strat2 (length Strat2))" and T_Strat2: "T\<^sub>p (Lxy init {x,y}) (Lxy (take n qs) {x, y}) Strat2 + length sws = ALGxy_det Strat (take n qs) init x y + ALG_Pxy Strat (take n qs) init x y " by (auto) from Suc(3-4) have nStrat: "n < length Strat" by auto from take_Suc_conv_app_nth[OF this] have tak2: "take (Suc n) Strat = take n Strat @ [Strat ! n]" by auto from take_Suc_conv_app_nth[OF ns] have tak: "take (Suc n) qs = take n qs @ [qs ! n]" by auto have aS: "length (take n Strat) = n" using Suc(3,4) by auto have aQ: "length (take n qs) = n" using Suc(4) by auto from aS aQ have qQS: "length (take n qs) = length (take n Strat)" by auto have xyininit: "x\ set init" "y : set init" by fact+ then have xysubs: "{x,y} \ set init" by auto have dI: "distinct init" by fact have "set qs \ set init" by fact then have qsnset: "qs ! n \ set init" using ns by auto from xyininit have ahjer: "set (Lxy init {x, y}) = {x,y}" using xysubs by (simp add: Lxy_set_filter) with Suc(5) have ah: "card (set (Lxy init {x, y})) = 2" by simp have ahjer3: "distinct (Lxy init {x,y})" apply(rule Lxy_distinct) by fact from ah have ahjer2: "length (Lxy init {x,y}) = 2" using distinct_card[OF ahjer3] by simp show ?case proof (cases "qs ! n \ {x,y}") case False with tak have nixzutun: "Lxy (take (Suc n) qs) {x,y} = Lxy (take n qs) {x,y}" unfolding Lxy_def by simp let ?m="ALG_P' (take n Strat @ [Strat ! n]) (take n qs @ [qs ! n]) init n x y" let ?L="replicate ?m 0 @ sws" { fix xs::"('a::linorder) list" fix m::nat fix q::'a assume "q \ {x,y}" then have 5: "y \ q" by auto assume 1: "q \ set xs" assume 2: "distinct xs" assume 3: "x \ set xs" assume 4: "y \ set xs" have "(x < y in xs) = (x < y in (mtf2 m q xs))" by (metis "1" "2" "3" "4" \q \ {x, y}\ insertCI not_before_in set_mtf2 swapped_by_mtf2) } note f=this (* taktik, erstmal das mtf weg bekommen, dann induct über snd (Strat!n) *) have "(x < y in steps' init (take (Suc n) qs) (take (Suc n) Strat) (Suc n)) = (x < y in mtf2 (fst (Strat ! n)) (qs ! n) (swaps (snd (Strat ! n)) (steps' init (take n qs) (take n Strat) n)))" unfolding tak2 tak apply(simp only: steps'_append[OF qQS aQ] ) by (simp add: step_def split_def) also have "\ = (x < y in (swaps (snd (Strat ! n)) (steps' init (take n qs) (take n Strat) n)))" apply(rule f[symmetric]) apply(fact) using qsnset steps'_set[OF qQS] aS apply(simp) using steps'_distinct[OF qQS] aS dI apply(simp) using steps'_set[OF qQS] aS xyininit by simp_all also have "\ = x < y in (swap 0 ^^ ALG_P (snd (Strat ! n)) x y (steps' init (take n qs) (take n Strat) n)) (swaps sws (steps' (Lxy init {x, y}) (Lxy (take n qs) {x, y}) Strat2 (length Strat2)))" apply(rule projected_paid_same_effect) apply(rule steps'_dist_perm) apply(fact qQS) apply(fact aS) using dI apply(simp) apply(fact Suc(5)) apply(simp) apply(rule steps'_set[where s="Lxy init {x,y}", unfolded ahjer]) using len apply(simp) apply(simp) apply(simp) apply(rule steps'_length[where s="Lxy init {x,y}", unfolded ahjer2]) using len apply(simp) apply(simp) apply(simp) apply(rule steps'_distinct2[where s="Lxy init {x,y}"]) using len apply(simp) apply(simp) apply(fact) using iff by auto finally have umfa: "x < y in steps' init (take (Suc n) qs) (take (Suc n) Strat) (Suc n) = x < y in (swap 0 ^^ ALG_P (snd (Strat ! n)) x y (steps' init (take n qs) (take n Strat) n)) (swaps sws (steps' (Lxy init {x, y}) (Lxy (take n qs) {x, y}) Strat2 (length Strat2)))" . from Suc(3,4) have lS: "length (take n Strat) = n" by auto have "(take n Strat @ [Strat ! n]) ! n = (take n Strat @ (Strat ! n) # []) ! length (take n Strat)" using lS by auto also have "\ = Strat ! n" by(rule nth_append_length) finally have tt: "(take n Strat @ [Strat ! n]) ! n = Strat ! n" . show ?thesis apply(rule exI[where x="Strat2"]) apply(rule exI[where x="?L"]) unfolding nixzutun apply(safe) apply(fact) proof goal_cases case 1 show ?case unfolding tak2 tak apply(simp add: step_def split_def) unfolding ALG_P'_def unfolding tt using aS apply(simp only: steps'_rests[OF qQS, symmetric]) using 1(1) umfa by auto next case 2 then show ?case apply(simp add: step_def split_def) unfolding ALG_P'_def unfolding tt using aS apply(simp only: steps'_rests[OF qQS, symmetric]) using umfa[symmetric] by auto next case 3 have ns2: "n < length (take n qs @ [qs ! n])" using ns by auto have er: "length (take n qs) < length Strat" using Suc.prems(2) aQ ns by linarith have "T\<^sub>p (Lxy init {x,y}) (Lxy (take n qs) {x, y}) Strat2 + length (replicate (ALG_P' Strat (take n qs @ [qs ! n]) init n x y) 0 @ sws) = ( T\<^sub>p (Lxy init {x,y}) (Lxy (take n qs) {x, y}) Strat2 + length sws) + ALG_P' Strat (take n qs @ [qs ! n]) init n x y" by simp also have "\ = ALGxy_det Strat (take n qs) init x y + ALG_Pxy Strat (take n qs) init x y + ALG_P' Strat (take n qs @ [qs ! n]) init n x y" unfolding T_Strat2 by simp also have "\ = ALGxy_det Strat (take (Suc n) qs) init x y + ALG_Pxy Strat (take (Suc n) qs) init x y" unfolding tak unfolding wegdamit[OF er False] apply(simp) unfolding ALG_P_split[of "take n qs" Strat "qs ! n" init x y, unfolded aQ, OF nStrat] by(simp) finally show ?case unfolding tak using ALG_P'_rest[OF ns nStrat] by auto qed next case True note qsinxy=this then have yeh: "Lxy (take (Suc n) qs) {x, y} = Lxy (take n qs) {x,y} @ [qs!n]" unfolding tak Lxy_def by auto from True have garar: "(take n qs @ [qs ! n]) ! n \ {y, x}" using tak[symmetric] by(auto) have aer: "\i {y, x}) = (take n qs ! i \ {y, x})" using ns by (metis less_SucI nth_take tak) (* erst definiere ich die zwischenzeitlichen Configurationen ?xs \ ?xs' \ ?xs'' und ?ys \ ?ys' \ ?ys'' \ ?ys''' und einige Eigenschaften über sie *) (* what is the mtf action taken by Strat? *) let ?Strat_mft = "fst (Strat ! n)" let ?Strat_sws = "snd (Strat ! n)" (* what is the configuration before the step? *) let ?xs = "steps' init (take n qs) (take n Strat) n" (* what is the configuration before the mtf *) let ?xs' = "(swaps (snd (Strat!n)) ?xs)" let ?xs'' = "steps' init (take (Suc n) qs) (take (Suc n) Strat) (Suc n)" let ?xs''2 = "mtf2 ?Strat_mft (qs!n) ?xs'" (* position of requested element *) let ?no_swap_occurs = "(x < y in ?xs') = (x < y in ?xs''2)" let ?mtf="(if ?no_swap_occurs then 0 else 1::nat)" let ?m="ALG_P' Strat (take n qs @ [qs ! n]) init n x y" let ?L="replicate ?m 0 @ sws" let ?newStrat="Strat2@[(?mtf,?L)]" have "?xs'' = step ?xs (qs!n) (Strat!n)" unfolding tak tak2 apply(rule steps'_append) by fact+ also have "\ = mtf2 (fst (Strat!n)) (qs!n) (swaps (snd (Strat!n)) ?xs)" unfolding step_def by (auto simp: split_def) finally have A: "?xs'' = mtf2 (fst (Strat!n)) (qs!n) ?xs'" . let ?ys = "(steps' (Lxy init {x, y}) (Lxy (take n qs) {x, y}) Strat2 (length Strat2))" let ?ys' = "( swaps sws (steps' (Lxy init {x, y}) (Lxy (take n qs) {x, y}) Strat2 (length Strat2)))" let ?ys'' = " (swap 0 ^^ ALG_P (snd (Strat!n)) x y ?xs) ?ys'" let ?ys''' = "(steps' (Lxy init {x, y}) (Lxy (take (Suc n) qs) {x, y}) ?newStrat (length ?newStrat))" have gr: "Lxy (take n qs @ [qs ! n]) {x, y} = Lxy (take n qs) {x, y} @ [qs ! n]" unfolding Lxy_def using True by(simp) have "steps' init (take n qs @ [qs ! n]) Strat n = steps' init (take n qs @ [qs ! n]) (take n Strat @ drop n Strat) n" by simp also have "\ = steps' init (take n qs) (take n Strat) n" apply(subst steps'_rests[symmetric]) using aS qQS by(simp_all) finally have t: "steps' init (take n qs @ [qs ! n]) Strat n = steps' init (take n qs) (take n Strat) n" . have gge: "swaps (replicate ?m 0) ?ys' = (swap 0 ^^ ALG_P (snd (Strat!n)) x y ?xs) ?ys'" unfolding ALG_P'_def t by simp have gg: "length ?newStrat = Suc (length Strat2)" by auto have "?ys''' = step ?ys (qs!n) (?mtf,?L)" unfolding tak gr unfolding gg apply(rule steps'_append) using len by auto also have "\ = mtf2 ?mtf (qs!n) (swaps ?L ?ys)" unfolding step_def by (simp add: split_def) also have "\ = mtf2 ?mtf (qs!n) (swaps (replicate ?m 0) ?ys')" by (simp) also have "\ = mtf2 ?mtf (qs!n) ?ys''" using gge by (simp) finally have B: "?ys''' = mtf2 ?mtf (qs!n) ?ys''" . have 3: "set ?ys' = {x,y}" apply(simp add: swaps_inv) apply(subst steps'_set) using ahjer len by(simp_all) have k: "?ys'' = swaps (replicate (ALG_P (snd (Strat!n)) x y ?xs) 0) ?ys'" by (auto) have 6: "set ?ys'' = {x,y}" unfolding k using 3 swaps_inv by metis have 7: "set ?ys''' = {x,y}" unfolding B using set_mtf2 6 by metis have 22: "x \ set ?ys''" "y \ set ?ys''" using 6 by auto have 23: "x \ set ?ys'''" "y \ set ?ys'''" using 7 by auto have 26: "(qs!n) \ set ?ys''" using 6 True by auto have "distinct ?ys" apply(rule steps'_distinct2) using len ahjer3 by(simp)+ then have 9: "distinct ?ys'" using swaps_inv by metis then have 27: "distinct ?ys''" unfolding k using swaps_inv by metis from 3 Suc(5) have "card (set ?ys') = 2" by auto then have 4: "length ?ys' = 2" using distinct_card[OF 9] by simp have "length ?ys'' = 2" unfolding k using 4 swaps_inv by metis have 5: "dist_perm ?ys' ?ys'" using 9 by auto have sxs: "set ?xs = set init" apply(rule steps'_set) using qQS n Suc(3) by(auto) have sxs': "set ?xs' = set ?xs" using swaps_inv by metis have sxs'': "set ?xs'' = set ?xs'" unfolding A using set_mtf2 by metis have 24: "x \ set ?xs'" "y\set ?xs'" "(qs!n) \ set ?xs'" using xysubs True sxs sxs' by auto have 28: "x \ set ?xs''" "y\set ?xs''" "(qs!n) \ set ?xs''" using xysubs True sxs sxs' sxs'' by auto have 0: "dist_perm init init" using dI by auto have 1: "dist_perm ?xs ?xs" apply(rule steps'_dist_perm) by fact+ then have 25: "distinct ?xs'" using swaps_inv by metis (* aus der Induktionsvorraussetzung (iff) weiß ich bereits dass die Ordnung erhalten wird bis zum nten Schritt, mit Theorem "projected_paid_same_effect" kann ich auch die paid exchanges abarbeiten ...*) from projected_paid_same_effect[OF 1 Suc(5) 3 4 5, OF iff, where acs="snd (Strat ! n)"] have aaa: "x < y in ?xs' = x < y in ?ys''" . (* ... was nun noch fehlt ist, dass die moveToFront anweisungen von Strat und Strat2 sich in gleicher Art auf die Ordnung von x und y auswirken *) have t: "?mtf = (if (xa b. a\{x,y} \ b\{x,y} \ set ?ys'' = {x,y} \ a\b \ distinct ?ys'' \ a ~a set ?ys''" by auto from d e have p: "mtf2 1 b ?ys'' = swap 0 ?ys''" unfolding mtf2_def by auto have q: "a < b in swap 0 ?ys'' = (\ a < b in ?ys'')" apply(rule swap0in2) by(fact)+ from 1(6) p q show ?case by metis qed show ?thesis proof (cases "xi\{.. {y, x} then ALG'_det Strat (take n qs @ [qs ! n]) init i y + ALG'_det Strat (take n qs @ [qs ! n]) init i x else 0)" unfolding ALGxy_det_def tak by auto also have "\ = (\i\{.. {y, x} then ALG'_det Strat (take n qs @ [qs ! n]) init i y + ALG'_det Strat (take n qs @ [qs ! n]) init i x else 0)" using ns by simp also have "\ = (\i\{.. {y, x} then ALG'_det Strat (take n qs @ [qs ! n]) init i y + ALG'_det Strat (take n qs @ [qs ! n]) init i x else 0) + (if (take n qs @ [qs ! n]) ! n \ {y, x} then ALG'_det Strat (take n qs @ [qs ! n]) init n y + ALG'_det Strat (take n qs @ [qs ! n]) init n x else 0)" by simp also have "\ = (\i\{..< n}. if take n qs ! i \ {y, x} then ALG'_det Strat (take n qs @ [qs ! n]) init i y + ALG'_det Strat (take n qs @ [qs ! n]) init i x else 0) + ALG'_det Strat (take n qs @ [qs ! n]) init n y + ALG'_det Strat (take n qs @ [qs ! n]) init n x " using aer using garar by simp also have "\ = (\i\{..< n}. if take n qs ! i \ {y, x} then ALG'_det Strat (take n qs @ [qs ! n]) init i y + ALG'_det Strat (take n qs @ [qs ! n]) init i x else 0) + ALG'_det Strat qs init n y + ALG'_det Strat qs init n x" proof - have "ALG'_det Strat qs init n y = ALG'_det Strat ((take n qs @ [qs ! n]) @ drop (Suc n) qs) init n y" unfolding tak[symmetric] by auto also have "\ = ALG'_det Strat (take n qs @ [qs ! n]) init n y " apply(rule ALG'_det_append) using nStrat ns by(auto) finally have 1: "ALG'_det Strat qs init n y = ALG'_det Strat (take n qs @ [qs ! n]) init n y" . have "ALG'_det Strat qs init n x = ALG'_det Strat ((take n qs @ [qs ! n]) @ drop (Suc n) qs) init n x" unfolding tak[symmetric] by auto also have "\ = ALG'_det Strat (take n qs @ [qs ! n]) init n x " apply(rule ALG'_det_append) using nStrat ns by(auto) finally have 2: "ALG'_det Strat qs init n x = ALG'_det Strat (take n qs @ [qs ! n]) init n x" . from 1 2 show ?thesis by auto qed also have "\ = (\i\{..< n}. if take n qs ! i \ {y, x} then ALG'_det Strat (take n qs) init i y + ALG'_det Strat (take n qs) init i x else 0) + ALG'_det Strat qs init n y + ALG'_det Strat qs init n x" apply(simp) apply(rule sum.cong) apply(simp) apply(simp) using ALG'_det_append[where qs="take n qs"] Suc.prems(2) ns by auto also have "\ = (\i\{..< length(take n qs)}. if take n qs ! i \ {y, x} then ALG'_det Strat (take n qs) init i y + ALG'_det Strat (take n qs) init i x else 0) + ALG'_det Strat qs init n y + ALG'_det Strat qs init n x" using aQ by auto also have "\ = ALGxy_det Strat (take n qs) init x y + (ALG'_det Strat qs init n y + ALG'_det Strat qs init n x)" unfolding ALGxy_det_def by(simp) finally show ?thesis . qed (* aaa: x < y in ?xs' = x < y in ?ys'' central: x < y in ?xs'' = x < y in ?ys''' *) have list: "?ys' = swaps sws (steps (Lxy init {x, y}) (Lxy (take n qs) {x, y}) Strat2)" unfolding steps_steps'[OF len[symmetric], of "(Lxy init {x, y})"] by simp have j2: "steps' init (take n qs @ [qs ! n]) Strat n = steps' init (take n qs) (take n Strat) n" proof - have "steps' init (take n qs @ [qs ! n]) Strat n = steps' init (take n qs @ [qs ! n]) (take n Strat @ drop n Strat) n" by auto also have "\ = steps' init (take n qs) (take n Strat) n" apply(rule steps'_rests[symmetric]) apply fact using aS by simp finally show ?thesis . qed have arghschonwieder: "steps' init (take n qs) (take n Strat) n = steps' init qs Strat n" proof - have "steps' init qs Strat n = steps' init (take n qs @ drop n qs) (take n Strat @ drop n Strat) n" by auto also have "\ = steps' init (take n qs) (take n Strat) n" apply(rule steps'_rests[symmetric]) apply fact using aS by simp finally show ?thesis by simp qed have indexe: "((swap 0 ^^ ?m) (swaps sws (steps (Lxy init {x,y}) (Lxy (take n qs) {x, y}) Strat2))) = ?ys''" unfolding ALG_P'_def unfolding list using j2 by auto have blocky: "ALG'_det Strat qs init n y = (if y < qs ! n in ?xs' then 1 else 0)" unfolding ALG'_det_def ALG.simps by(auto simp: arghschonwieder split_def) have blockx: "ALG'_det Strat qs init n x = (if x < qs ! n in ?xs' then 1 else 0)" unfolding ALG'_det_def ALG.simps by(auto simp: arghschonwieder split_def) have index_is_blocking_cost: "index ((swap 0 ^^ ?m) (swaps sws (steps (Lxy init {x,y}) (Lxy (take n qs) {x, y}) Strat2))) (qs ! n) = ALG'_det Strat qs init n y + ALG'_det Strat qs init n x" proof (cases "x= qs!n") case True then have "ALG'_det Strat qs init n x = 0" unfolding blockx apply(simp) using before_in_irefl by metis then have "ALG'_det Strat qs init n y + ALG'_det Strat qs init n x = (if y < x in ?xs' then 1 else 0)" unfolding blocky using True by simp also have "\ = (if ~y < x in ?xs' then 0 else 1)" by auto also have "\ = (if x < y in ?xs' then 0 else 1)" apply(simp) by (meson 24 Suc.prems(4) not_before_in) also have "\ = (if x < y in ?ys'' then 0 else 1)" using aaa by simp also have "\ = index ?ys'' x" apply(rule before_in_index1) by(fact)+ finally show ?thesis unfolding indexe using True by auto next case False then have q: "y = qs!n" using qsinxy by auto then have "ALG'_det Strat qs init n y = 0" unfolding blocky apply(simp) using before_in_irefl by metis then have "ALG'_det Strat qs init n y + ALG'_det Strat qs init n x = (if x < y in ?xs' then 1 else 0)" unfolding blockx using q by simp also have "\ = (if x < y in ?ys'' then 1 else 0)" using aaa by simp also have "\ = index ?ys'' y" apply(rule before_in_index2) by(fact)+ finally show ?thesis unfolding indexe using q by auto qed have jj: "ALG_Pxy Strat (take (Suc n) qs) init x y = ALG_Pxy Strat (take n qs) init x y + ALG_P' Strat (take n qs @ [qs ! n]) init n x y" proof - have "ALG_Pxy Strat (take (Suc n) qs) init x y = (\i = (\i< Suc n. ALG_P' Strat (take (Suc n) qs) init i x y)" unfolding tak using ns by simp also have "\ = (\i = (\i = (\i = ALG_Pxy Strat (take n qs) init x y + ALG_P' Strat (take n qs @ [qs ! n]) init n x y" unfolding ALG_Pxy_def by auto finally show ?thesis . qed have tw: "length (Lxy (take n qs) {x, y}) = length Strat2" using len by auto have "T\<^sub>p (Lxy init {x,y}) (Lxy (take (Suc n) qs) {x, y}) ?newStrat + length [] = T\<^sub>p (Lxy init {x,y}) (Lxy (take n qs) {x, y}) Strat2 + t\<^sub>p (steps (Lxy init {x, y}) (Lxy (take n qs) {x, y}) Strat2) (qs ! n) (?mtf,?L)" unfolding yeh by(simp add: T_append[OF tw, of "(Lxy init) {x,y}"]) also have "\ = T\<^sub>p (Lxy init {x,y}) (Lxy (take n qs) {x, y}) Strat2 + length sws + index ((swap 0 ^^ ?m) (swaps sws (steps (Lxy init {x,y}) (Lxy (take n qs) {x, y}) Strat2))) (qs ! n) + ALG_P' Strat (take n qs @ [qs ! n]) init n x y" by(simp add: t\<^sub>p_def) (* now use iH *) also have "\ = (ALGxy_det Strat (take n qs) init x y + index ((swap 0 ^^ ?m) (swaps sws (steps (Lxy init {x,y}) (Lxy (take n qs) {x, y}) Strat2))) (qs ! n)) + (ALG_Pxy Strat (take n qs) init x y + ALG_P' Strat (take n qs @ [qs ! n]) init n x y)" by (simp only: T_Strat2) (* the current cost are equal to the blocking costs: *) also from index_is_blocking_cost have "\ = (ALGxy_det Strat (take n qs) init x y + ALG'_det Strat qs init n y + ALG'_det Strat qs init n x) + (ALG_Pxy Strat (take n qs) init x y + ALG_P' Strat (take n qs @ [qs ! n]) init n x y)" by auto also have "\ = ALGxy_det Strat (take (Suc n) qs) init x y + (ALG_Pxy Strat (take n qs) init x y + ALG_P' Strat (take n qs @ [qs ! n]) init n x y)" using j by auto also have "\ = ALGxy_det Strat (take (Suc n) qs) init x y + ALG_Pxy Strat (take (Suc n) qs) init x y" using jj by auto finally show ?case . qed qed qed next case 0 then show ?case apply (simp add: Lxy_def ALGxy_det_def ALG_Pxy_def T_opt_def) proof goal_cases case 1 show ?case apply(rule Lxy_mono[unfolded Lxy_def, simplified]) using 1 by auto qed qed lemma T1_7: assumes "T\<^sub>p init qs Strat = T\<^sub>p_opt init qs" "length Strat = length qs" "x \ (y::('a::linorder))" "x\ set init" "y \ set init" "distinct init" "set qs \ set init" shows "T\<^sub>p_opt (Lxy init {x,y}) (Lxy qs {x,y}) \ ALGxy_det Strat qs init x y + ALG_Pxy Strat qs init x y" proof - have A:"length qs \ length qs" by auto have B:" x \ y " using assms by auto from T1_7'[OF assms(1,2), of "length qs" x y, OF A B assms(4-7)] obtain Strat2 sws where len: "length Strat2 = length (Lxy qs {x, y})" and "x < y in steps' init qs (take (length qs) Strat) (length qs) = x < y in swaps sws (steps' (Lxy init {x,y}) (Lxy qs {x, y}) Strat2 (length Strat2))" and Tp: "T\<^sub>p (Lxy init {x,y}) (Lxy qs {x, y}) Strat2 + length sws = ALGxy_det Strat qs init x y + ALG_Pxy Strat qs init x y" by auto have "T\<^sub>p_opt (Lxy init {x,y}) (Lxy qs {x,y}) \ T\<^sub>p (Lxy init {x,y}) (Lxy qs {x, y}) Strat2" unfolding T_opt_def apply(rule cInf_lower) using len by auto also have "\ \ ALGxy_det Strat qs init x y + ALG_Pxy Strat qs init x y" using Tp by auto finally show ?thesis . qed lemma T_snoc: "length rs = length as \ T init (rs@[r]) (as@[a]) = T init rs as + t\<^sub>p (steps' init rs as (length rs)) r a" apply(induct rs as arbitrary: init rule: list_induct2) by simp_all lemma steps'_snoc: "length rs = length as \ n = (length as) \ steps' init (rs@[r]) (as@[a]) (Suc n) = step (steps' init rs as n) r a" apply(induct rs as arbitrary: init n r a rule: list_induct2) by (simp_all) lemma steps'_take: assumes "n = steps' init (take n qs) (take n Strat) n" apply(subst steps'_rests[symmetric]) using assms by auto finally show ?thesis by simp qed lemma Tp_darstellung: "length qs = length Strat \ T\<^sub>p init qs Strat = (\i\{..p (steps' init qs Strat i) (qs!i) (Strat!i))" proof - assume a[simp]: "length qs = length Strat" {fix n have "n\length qs \ T\<^sub>p init (take n qs) (take n Strat) = (\i\{..p (steps' init qs Strat i) (qs!i) (Strat!i))" apply(induct n) apply(simp) apply(simp add: take_Suc_conv_app_nth) apply(subst T_snoc) apply(simp) by(simp add: min_def steps'_take) } from a this[of "length qs"] show ?thesis by auto qed (* Gleichung 1.8 in Borodin *) lemma umformung_OPT': assumes inlist: "set qs \ set init" assumes dist: "distinct init" assumes qsStrat: "length qs = length Strat" assumes noStupid: "\x l. x l< length (snd (Strat ! x)) \ Suc ((snd (Strat ! x))!l) < length init" shows "T\<^sub>p init qs Strat = (\(x,y)\{(x,y::('a::linorder)). x \ set init \ y\set init \ xn. \xa \ set_pmf (config\<^sub>p (I, S) qs init n). distinct (snd xa)" using dist config_rand_distinct by metis *) (* ersten Teil umformen: *) have "(\i\{..(x,y)\{(x,y). x \ set init \ y\set init \ xi\{..z\{(x,y). x \ set init \ y\set init \ x = (\z\{(x,y). x \ set init \ y\set init \ xi\{.. = (\(x,y)\{(x,y). x \ set init \ y\set init \ xi\{.. = (\(x,y)\{(x,y). x \ set init \ y\set init \ xi\{..(x,y)\{(x,y). x \ set init \ y\set init \ x(x,y)\{(x,y). x \ set init \ y\set init \ xi\{..e\set init. ALG e qs i (?config i, ()))) = (\e\set init. (\i\{.. = (\e\set init. (\y\set init. (\i\{i. i qs!i=y}. ALG e qs i (?config i,()))))" proof (rule sum.cong, goal_cases) case (2 x) have "(\i = sum (%i. ALG x qs i (?config i, ())) (\y\{y. y \ set init}. {i. i < length qs \ qs ! i = y})" apply(rule sum.cong) proof goal_cases case 1 show ?case apply(auto) using inlist by auto qed simp also have "\ = sum (%t. sum (%i. ALG x qs i (?config i, ())) {i. i qs ! i = t}) {y. y\ set init}" apply(rule sum.UNION_disjoint) apply(simp_all) by force also have "\ = (\y\set init. \i | i < length qs \ qs ! i = y. ALG x qs i (?config i, ()))" by auto finally show ?case . qed (simp) also have "\ = (\(x,y)\ (set init \ set init). (\i\{i. i qs!i=y}. ALG x qs i (?config i, ())))" by (rule sum.cartesian_product) also have "\ = (\(x,y)\ {(x,y). x\set init \ y\ set init}. (\i\{i. i qs!i=y}. ALG x qs i (?config i, ())))" by simp also have E4: "\ = (\(x,y)\{(x,y). x\set init \ y\ set init \ x\y}. (\i\{i. i qs!i=y}. ALG x qs i (?config i, ())))" (is "(\(x,y)\ ?L. ?f x y) = (\(x,y)\ ?R. ?f x y)") proof goal_cases case 1 let ?M = "{(x,y). x\set init \ y\ set init \ x=y}" have A: "?L = ?R \ ?M" by auto have B: "{} = ?R \ ?M" by auto have "(\(x,y)\ ?L. ?f x y) = (\(x,y)\ ?R \ ?M. ?f x y)" by(simp only: A) also have "\ = (\(x,y)\ ?R. ?f x y) + (\(x,y)\ ?M. ?f x y)" apply(rule sum.union_disjoint) apply(rule finite_subset[where B="set init \ set init"]) apply(auto) apply(rule finite_subset[where B="set init \ set init"]) by(auto) also have "(\(x,y)\ ?M. ?f x y) = 0" apply(rule sum.neutral) by (auto simp add: split_def before_in_def) finally show ?case by simp qed also have "\ = (\(x,y)\{(x,y). x \ set init \ y\set init \ xi\{i. i qs!i=y}. ALG x qs i (?config i, ())) + (\i\{i. i qs!i=x}. ALG y qs i (?config i, ())) )" (is "(\(x,y)\ ?L. ?f x y) = (\(x,y)\ ?R. ?f x y + ?f y x)") proof - let ?R' = "{(x,y). x \ set init \ y\set init \ y ?R'" by auto have "{} = ?R \ ?R'" by auto have C: "?R' = (%(x,y). (y, x)) ` ?R" by auto have D: "(\(x,y)\ ?R'. ?f x y) = (\(x,y)\ ?R. ?f y x)" proof - have "(\(x,y)\ ?R'. ?f x y) = (\(x,y)\ (%(x,y). (y, x)) ` ?R. ?f x y)" by(simp only: C) also have "(\z\ (%(x,y). (y, x)) ` ?R. (%(x,y). ?f x y) z) = (\z\?R. ((%(x,y). ?f x y) \ (%(x,y). (y, x))) z)" apply(rule sum.reindex) by(fact swap_inj_on) also have "\ = (\z\?R. (%(x,y). ?f y x) z)" apply(rule sum.cong) by(auto) finally show ?thesis . qed have "(\(x,y)\ ?L. ?f x y) = (\(x,y)\ ?R \ ?R'. ?f x y)" by(simp only: A) also have "\ = (\(x,y)\ ?R. ?f x y) + (\(x,y)\ ?R'. ?f x y)" apply(rule sum.union_disjoint) apply(rule finite_subset[where B="set init \ set init"]) apply(auto) apply(rule finite_subset[where B="set init \ set init"]) by(auto) also have "\ = (\(x,y)\ ?R. ?f x y) + (\(x,y)\ ?R. ?f y x)" by(simp only: D) also have "\ = (\(x,y)\ ?R. ?f x y + ?f y x)" by(simp add: split_def sum.distrib[symmetric]) finally show ?thesis . qed also have E5: "\ = (\(x,y)\{(x,y). x \ set init \ y\set init \ xi\{i. i (qs!i=y \ qs!i=x)}. ALG y qs i (?config i, ()) + ALG x qs i (?config i, ())))" apply(rule sum.cong) apply(simp) proof goal_cases case (1 x) then obtain a b where x: "x=(a,b)" and a: "a \ set init" "b \ set init" "a < b" by auto then have "a\b" by simp then have disj: "{i. i < length qs \ qs ! i = b} \ {i. i < length qs \ qs ! i = a} = {}" by auto have unio: "{i. i < length qs \ (qs ! i = b \ qs ! i = a)} = {i. i < length qs \ qs ! i = b} \ {i. i < length qs \ qs ! i = a}" by auto let ?f="%i. ALG b qs i (?config i, ()) + ALG a qs i (?config i, ())" let ?B="{i. i < length qs \ qs ! i = b}" let ?A="{i. i < length qs \ qs ! i = a}" have "(\i\?B \ ?A. ?f i) = (\i\?B. ?f i) + (\i\?A. ?f i) - (\i\?B \ ?A. ?f i) " apply(rule sum_Un_nat) by auto also have "\ = (\i\?B. ALG b qs i (?config i, ()) + ALG a qs i (?config i, ())) + (\i\?A. ALG b qs i (?config i, ()) + ALG a qs i (?config i, ()))" using disj by auto also have "\ = (\i\?B. ALG a qs i (?config i, ())) + (\i\?A. ALG b qs i (?config i, ()))" by (auto simp: split_def before_in_def) finally show ?case unfolding x apply(simp add: split_def) unfolding unio by simp qed also have E6: "\ = (\(x,y)\{(x,y). x \ set init \ y\set init \ xie\set init. ALG e qs i (?config i, ())) = (\(x,y)\{(x,y). x \ set init \ y\set init \ xp init qs Strat = (\i\{..p (steps' init qs Strat i) (qs!i) (Strat!i))" by auto also have "\ = (\i\{..e\set (steps' init qs Strat i). ALG e qs i (swaps (snd (Strat!i)) (steps' init qs Strat i),())) + (\(x,y)\{(x,(y::('a::linorder))). x \ set (steps' init qs Strat i) \ y\set (steps' init qs Strat i) \ xp_sumofALGALGP) apply(rule steps'_distinct2) using dist qsStrat apply(simp_all) apply(subst steps'_set) using dist qsStrat inlist apply(simp_all) apply fastforce apply(subst steps'_length) apply(simp_all) using noStupid by auto also have "\ = (\i\{..e\set init. ALG e qs i (swaps (snd (Strat!i)) (steps' init qs Strat i),())) + (\(x,y)\{(x,y). x \ set init \ y\set init \ x = (\i\{..e\set init. ALG e qs i (swaps (snd (Strat!i)) (steps' init qs Strat i), ()))) + (\i\{..(x,y)\{(x,y). x \ set init \ y\set init \ x = (\(x,y)\{(x,y). x \ set init \ y\set init \ xi\{..(x,y)\{(x,y). x \ set init \ y\set init \ x = (\(x,y)\{(x,y). x \ set init \ y\set init \ x(x,y)\{(x,y). x \ set init \ y\set init \ x = (\(x,y)\{(x,y). x \ set init \ y\set init \ x {}" shows "Inf S \ S" using assms Inf_nat_def LeastI by force lemma steps_length: "length qs = length as \ length (steps s qs as) = length s" apply(induct qs as arbitrary: s rule: list_induct2) by simp_all (* shows that OPT does not use paid exchanges that do not have an effect *) lemma OPT_noStupid: fixes Strat assumes [simp]: "length Strat = length qs" assumes opt: "T\<^sub>p init qs Strat = T\<^sub>p_opt init qs" assumes init_nempty: "init\[]" shows "\x l. x < length Strat \ l < length (snd (Strat ! x)) \ Suc ((snd (Strat ! x))!l) < length init" proof (rule ccontr, goal_cases) case (1 x l) (* construct a Stratgy that leaves out that paid exchange *) let ?sws' = "take l (snd (Strat!x)) @ drop (Suc l) (snd (Strat!x))" let ?Strat' = "take x Strat @ (fst (Strat!x),?sws') # drop (Suc x) Strat" from 1(1) have valid: "length ?Strat' = length qs" by simp from valid have isin: "T\<^sub>p init qs ?Strat' \ {T\<^sub>p init qs as |as. length as = length qs}" by blast from 1(1,2) have lsws': "length (snd (Strat!x)) = length ?sws' + 1" by (simp) have a: "(take x ?Strat') = (take x Strat)" using 1(1) by(auto simp add: min_def take_Suc_conv_app_nth) have b: "(drop (Suc x) Strat) = (drop (Suc x) ?Strat')" using 1(1) by(auto simp add: min_def take_Suc_conv_app_nth) have aa: "(take l (snd (Strat!x))) = (take l (snd (?Strat'!x)))" using 1(1,2) by(auto simp add: min_def take_Suc_conv_app_nth nth_append) have bb: "(drop (Suc l) (snd (Strat!x))) = (drop l (snd (?Strat'!x)))" using 1(1,2) by(auto simp add: min_def take_Suc_conv_app_nth nth_append ) have "(swaps (snd (Strat ! x)) (steps init (take x qs) (take x Strat))) = (swaps (take l (snd (Strat ! x)) @ (snd (Strat ! x))!l # drop (Suc l) (snd (Strat ! x))) (steps init (take x qs) (take x Strat)))" unfolding id_take_nth_drop[OF 1(2), symmetric] by simp also have "... = (swaps (take l (snd (Strat ! x)) @ drop (Suc l) (snd (Strat ! x))) (steps init (take x qs) (take x Strat)))" using 1(3) by(simp add: swap_def steps_length) finally have noeffect: "(swaps (snd (Strat ! x)) (steps init (take x qs) (take x Strat))) = (swaps (take l (snd (Strat ! x)) @ drop (Suc l) (snd (Strat ! x))) (steps init (take x qs) (take x Strat)))" . have c: "t\<^sub>p (steps init (take x qs) (take x Strat)) (qs ! x) (Strat ! x) = t\<^sub>p (steps init (take x qs) (take x ?Strat')) (qs ! x) (?Strat' ! x) + 1" unfolding a t\<^sub>p_def using 1(1,2) apply(simp add: min_def split_def nth_append) unfolding noeffect by(simp) have "T\<^sub>p init (take (Suc x) qs) (take (Suc x) Strat) = T\<^sub>p init (take x qs) (take x ?Strat') + t\<^sub>p (steps init (take x qs) (take x Strat)) (qs ! x) (Strat ! x)" using 1(1) a by(simp add: take_Suc_conv_app_nth T_append) also have "\ = T\<^sub>p init (take x qs) (take x ?Strat') + t\<^sub>p (steps init (take x qs) (take x ?Strat')) (qs ! x) (?Strat' ! x) + 1" unfolding c by(simp) also have "\ = T\<^sub>p init (take (Suc x) qs) (take (Suc x) ?Strat') + 1" using 1(1) a by(simp add: min_def take_Suc_conv_app_nth T_append nth_append) finally have bef: "T\<^sub>p init (take (Suc x) qs) (take (Suc x) Strat) = T\<^sub>p init (take (Suc x) qs) (take (Suc x) ?Strat') + 1" . let ?interstate = "(steps init (take (Suc x) qs) (take (Suc x) Strat))" let ?interstate' = "(steps init (take (Suc x) qs) (take (Suc x) ?Strat'))" have state: "?interstate' = ?interstate" using 1(1) apply(simp add: take_Suc_conv_app_nth min_def) apply(simp add: steps_append step_def split_def) using noeffect by simp have "T\<^sub>p init qs Strat = T\<^sub>p init (take (Suc x) qs @ drop (Suc x) qs) (take (Suc x) Strat @ drop (Suc x) Strat)" by simp also have "\ = T\<^sub>p init (take (Suc x) qs) (take (Suc x) Strat) + T\<^sub>p ?interstate (drop (Suc x) qs) (drop (Suc x) Strat)" apply(subst T_append2) by(simp_all) also have "\ = T\<^sub>p init (take (Suc x) qs) (take (Suc x) ?Strat') + T\<^sub>p ?interstate' (drop (Suc x) qs) (drop (Suc x) ?Strat') + 1" unfolding bef state using 1(1) by(simp add: min_def nth_append) also have "\ = T\<^sub>p init (take (Suc x) qs @ drop (Suc x) qs) (take (Suc x) ?Strat' @ drop (Suc x) ?Strat') + 1" apply(subst T_append2) using 1(1) by(simp_all add: min_def) also have "\ = T\<^sub>p init qs ?Strat' + 1" by simp finally have better: "T\<^sub>p init qs ?Strat' + 1 = T\<^sub>p init qs Strat" by simp have "T\<^sub>p init qs ?Strat' + 1 = T\<^sub>p init qs Strat" by (fact better) also have "\ = T\<^sub>p_opt init qs" by (fact opt) also from cInf_lower[OF isin] have "\ \ T\<^sub>p init qs ?Strat'" unfolding T_opt_def by simp finally show "False" using init_nempty by auto qed (* Gleichung 1.8 in Borodin *) lemma umformung_OPT: assumes inlist: "set qs \ set init" assumes dist: "distinct init" assumes a: "T\<^sub>p_opt init qs = T\<^sub>p init qs Strat" assumes b: " length qs = length Strat" assumes c: "init\[]" shows "T\<^sub>p_opt init qs = (\(x,y)\{(x,y::('a::linorder)). x \ set init \ y\set init \ xp_opt init qs = T\<^sub>p init qs Strat" by(fact a) also have "\ = (\(x,y)\{(x,y::('a::linorder)). x \ set init \ y\set init \ x[]" and setqsinit: "set qs \ set init" shows "(\(x,y)\{(x,y::('a::linorder)). x \ set init \ y\set init \ xp_opt (Lxy init {x,y}) (Lxy qs {x,y}))) \ T\<^sub>p_opt init qs" proof - have "T\<^sub>p_opt init qs \ {T\<^sub>p init qs as |as. length as = length qs}" unfolding T_opt_def apply(rule nn_contains_Inf) apply(auto) by (rule Ex_list_of_length) then obtain Strat where a: "T\<^sub>p init qs Strat = T\<^sub>p_opt init qs" and b: "length Strat = length qs" unfolding T_opt_def by auto have "(\(x,y)\{(x,y). x \ set init \ y\set init \ xp_opt (Lxy init {x,y}) (Lxy qs {x, y})) \ (\(x,y)\{(x,y). x \ set init \ y\set init \ xb" by auto show ?case apply(rule T1_7[OF a b]) by(fact)+ qed also from umformung_OPT[OF setqsinit dist] a b c have "\ = T\<^sub>p init qs Strat" by auto also from a have "\ = T\<^sub>p_opt init qs" by simp finally show ?thesis . qed subsection "Factoring Lemma" lemma cardofpairs: "S \ [] \ sorted S \ distinct S \ card {(x,y). x \ set S \ y\set S \ x set ss \ y \ set ss \ x < y} = (length ss * (length ss-1)) / 2" by auto from cons have sss: "s \ set ss" by auto from cons have tt: "(\y\set (s#ss). s \ y)" by auto with cons have tt': "(\y\set ss. s < y)" proof - from sss have "(\y\set ss. s \ y)" by auto with tt show ?thesis by fastforce qed then have "{(x, y) . x = s \ y \ set ss \ x < y} = {(x, y) . x = s \ y \ set ss}" by auto also have "\ = {s}\(set ss)" by auto finally have "{(x, y) . x = s \ y \ set ss \ x < y} = {s}\(set ss)" . then have "card {(x, y) . x = s \ y \ set ss \ x < y} = card (set ss)" by(auto) also from cons distinct_card have "\ = length ss" by auto finally have step: "card {(x, y) . x = s \ y \ set ss \ x < y} = length ss" . have uni: "{(x, y) . x \ set (s # ss) \ y \ set (s # ss) \ x < y} = {(x, y) . x \ set ss \ y \ set ss \ x < y} \ {(x, y) . x = s \ y \ set ss \ x < y}" using tt by auto have disj: "{(x, y) . x \ set ss \ y \ set ss \ x < y} \ {(x, y) . x = s \ y \ set ss \ x < y} = {}" using sss by(auto) have "card {(x, y) . x \ set (s # ss) \ y \ set (s # ss) \ x < y} = card ({(x, y) . x \ set ss \ y \ set ss \ x < y} \ {(x, y) . x = s \ y \ set ss \ x < y})" using uni by auto also have "\ = card {(x, y) . x \ set ss \ y \ set ss \ x < y} + card {(x, y) . x = s \ y \ set ss \ x < y}" apply(rule card_Un_disjoint) apply(rule finite_subset[where B="(set ss) \ (set ss)"]) apply(force) apply(simp) apply(rule finite_subset[where B="{s} \ (set ss)"]) apply(force) apply(simp) using disj apply(simp) done also have "\ = (length ss * (length ss-1)) / 2 + length ss" using iH step by auto also have "\ = (length ss * (length ss-1) + 2*length ss) / 2" by auto also have "\ = (length ss * (length ss-1) + length ss * 2) / 2" by auto also have "\ = (length ss * (length ss-1+2)) / 2" by simp also have "\ = (length ss * (length ss+1)) / 2" using cons(1) by simp also have "\ = ((length ss+1) * length ss) / 2" by auto also have "\ = (length (s#ss) * (length (s#ss)-1)) / 2" by auto finally show ?case by auto next case single thus ?case by(simp cong: conj_cong) qed (* factoring lemma *) lemma factoringlemma_withconstant: fixes A and b::real and c::real assumes c: "c \ 1" assumes dist: "\e\S0. distinct e" assumes notempty: "\e\S0. length e > 0" (* A has pairwise property *) assumes pw: "pairwise A" (* A is c-competitive on list of length 2 *) assumes on2: "\s0\S0. \b\0. \qs\{x. set x \ set s0}. \(x,y)\{(x,y). x \ set s0 \ y\set s0 \ xp_on_rand A (Lxy s0 {x,y}) (Lxy qs {x,y}) \ c * (T\<^sub>p_opt (Lxy s0 {x,y}) (Lxy qs {x,y})) + b" assumes nopaid: "\is s q. \((free,paid),_) \ (snd A (s, is) q). paid=[]" assumes 4: "\init qs. distinct init \ set qs \ set init \ (\x. x finite (set_pmf (config'' A qs init x)))" (* then A is c-competitive on arbitrary list lengths *) shows "\s0\S0. \b\0. \qs\{x. set x \ set s0}. T\<^sub>p_on_rand A s0 qs \ c * real (T\<^sub>p_opt s0 qs) + b" proof (standard, goal_cases) case (1 init) have d: "distinct init" using dist 1 by auto have d2: "init \ []" using notempty 1 by auto obtain b where on3: "\qs\{x. set x \ set init}. \(x,y)\{(x,y). x \ set init \ y\set init \ xp_on_rand A (Lxy init {x,y}) (Lxy qs {x,y}) \ c * (T\<^sub>p_opt (Lxy init {x,y}) (Lxy qs {x,y})) + b" and b: "b\0" using on2 1 by auto { fix qs assume drin: "set qs \ set init" have "T\<^sub>p_on_rand A init qs = (\(x,y)\{(x, y) . x \ set init \ y \ set init \ x < y}. T\<^sub>p_on_rand A (Lxy init {x,y}) (Lxy qs {x, y})) " apply(rule umf_pair) apply(fact)+ using 4[of init qs] drin d by(simp add: split_def) (* 1.4 *) also have "\ \ (\(x,y)\{(x,y). x \ set init \ y\set init \ xp_opt (Lxy init {x,y}) (Lxy qs {x,y})) + b)" apply(rule sum_mono) using on3 drin by(simp add: split_def) also have "\ = c * (\(x,y)\{(x,y). x \ set init \ y\set init \ xp_opt (Lxy init {x,y}) (Lxy qs {x,y})) + b*(((length init)*(length init-1)) / 2)" proof - { fix S::"'a list" assume dis: "distinct S" assume d2: "S \ []" then have d3: "sort S \ []" by (metis length_0_conv length_sort) have "card {(x,y). x \ set S \ y\set S \ x set (sort S) \ y\set (sort S) \ x = (length (sort S) * (length (sort S) - 1)) / 2" apply(rule cardofpairs) using dis d2 d3 by (simp_all) finally have "card {(x, y) . x \ set S \ y \ set S \ x < y} = (length (sort S) * (length (sort S) - 1)) / 2 " . } with d d2 have e: "card {(x,y). x \ set init \ y\set init \ x(x,y)\?S. c * (?T x y) + b) = c * ?R + b*?T2") proof - have "(\(x,y)\?S. c * (?T x y) + b) = c * (\(x,y)\?S. (?T x y)) + (\(x,y)\?S. b)" by(simp add: split_def sum.distrib sum_distrib_left) also have "\ = c * (\(x,y)\?S. (?T x y)) + b*?T2" using e by(simp add: split_def) finally show ?thesis by(simp add: split_def) qed qed also have "\ \ c * T\<^sub>p_opt init qs + (b*((length init)*(length init-1)) / 2)" proof - have "(\(x, y)\{(x, y) . x \ set init \ y \ set init \ x < y}. T\<^sub>p_opt (Lxy init {x,y}) (Lxy qs {x, y})) \ T\<^sub>p_opt init qs" using OPT_zerlegen drin d d2 by auto then have " real (\(x, y)\{(x, y) . x \ set init \ y \ set init \ x < y}. T\<^sub>p_opt (Lxy init {x,y}) (Lxy qs {x, y})) \ (T\<^sub>p_opt init qs)" by linarith with c show ?thesis by(auto simp: split_def) qed finally have f: "T\<^sub>p_on_rand A init qs \ c * real (T\<^sub>p_opt init qs) + (b*((length init)*(length init-1)) / 2)" . } note all=this show ?case unfolding compet_def apply(auto) apply(rule exI[where x="(b*((length init)*(length init-1)) / 2)"]) apply(safe) using notempty 1 b apply simp using all b by simp qed lemma factoringlemma_withconstant': fixes A and b::real and c::real assumes c: "c \ 1" assumes dist: "\e\S0. distinct e" assumes notempty: "\e\S0. length e > 0" (* A has pairwise property *) assumes pw: "pairwise A" (* A is c-competitive on list of length 2 *) assumes on2: "\s0\S0. \b\0. \qs\{x. set x \ set s0}. \(x,y)\{(x,y). x \ set s0 \ y\set s0 \ xp_on_rand A (Lxy s0 {x,y}) (Lxy qs {x,y}) \ c * (T\<^sub>p_opt (Lxy s0 {x,y}) (Lxy qs {x,y})) + b" assumes nopaid: "\is s q. \((free,paid),_) \ (snd A (s, is) q). paid=[]" assumes 4: "\init qs. distinct init \ set qs \ set init \ (\x. x finite (set_pmf (config'' A qs init x)))" (* then A is c-competitive on arbitrary list lengths *) shows "compet_rand A c S0" unfolding compet_rand_def static_def using factoringlemma_withconstant[OF assms] by simp end diff --git a/thys/MFMC_Countable/MFMC_Flow_Attainability.thy b/thys/MFMC_Countable/MFMC_Flow_Attainability.thy --- a/thys/MFMC_Countable/MFMC_Flow_Attainability.thy +++ b/thys/MFMC_Countable/MFMC_Flow_Attainability.thy @@ -1,2166 +1,2166 @@ theory MFMC_Flow_Attainability imports MFMC_Network begin section \Attainability of flows in networks\ subsection \Cleaning up flows\ text \If there is a flow along antiparallel edges, it suffices to consider the difference.\ definition cleanup :: "'a flow \ 'a flow" where "cleanup f = (\(a, b). if f (a, b) > f (b, a) then f (a, b) - f (b, a) else 0)" lemma cleanup_simps [simp]: "cleanup f (a, b) = (if f (a, b) > f (b, a) then f (a, b) - f (b, a) else 0)" by(simp add: cleanup_def) lemma value_flow_cleanup: assumes [simp]: "\x. f (x, source \) = 0" shows "value_flow \ (cleanup f) = value_flow \ f" unfolding d_OUT_def by (auto simp add: not_less intro!: nn_integral_cong intro: antisym) lemma KIR_cleanup: assumes KIR: "KIR f x" and finite_IN: "d_IN f x \ \" shows "KIR (cleanup f) x" proof - from finite_IN KIR have finite_OUT: "d_OUT f x \ \" by simp have finite_IN: "(\\<^sup>+ y\A. f (y, x)) \ \" for A using finite_IN by(rule neq_top_trans)(auto simp add: d_IN_def nn_integral_count_space_indicator intro!: nn_integral_mono split: split_indicator) have finite_OUT: "(\\<^sup>+ y\A. f (x, y)) \ \" for A using finite_OUT by(rule neq_top_trans)(auto simp add: d_OUT_def nn_integral_count_space_indicator intro!: nn_integral_mono split: split_indicator) have finite_in: "f (x, y) \ \" for y using \d_OUT f x \ \\ by(rule neq_top_trans) (rule d_OUT_ge_point) let ?M = "{y. f (x, y) > f (y, x)}" have "d_OUT (cleanup f) x = (\\<^sup>+ y\?M. f (x, y) - f (y, x))" by(auto simp add: d_OUT_def nn_integral_count_space_indicator intro!: nn_integral_cong) also have "\ = (\\<^sup>+ y\?M. f (x, y)) - (\\<^sup>+ y\?M. f (y, x))" using finite_IN by(subst nn_integral_diff)(auto simp add: AE_count_space) also have "\ = (d_OUT f x - (\\<^sup>+ y\{y. f (x, y) \ f (y, x)}. f (x, y))) - (\\<^sup>+ y\?M. f (y, x))" unfolding d_OUT_def d_IN_def using finite_IN finite_OUT apply(simp add: nn_integral_count_space_indicator) apply(subst (2) nn_integral_diff[symmetric]) apply(auto simp add: AE_count_space finite_in split: split_indicator intro!: arg_cong2[where f="(-)"] intro!: nn_integral_cong) done also have "\ = (d_IN f x - (\\<^sup>+ y\?M. f (y, x))) - (\\<^sup>+ y\{y. f (x, y) \ f (y, x)}. f (x, y))" using KIR by(simp add: diff_diff_commute_ennreal) also have "\ = (\\<^sup>+ y\{y. f (x, y) \ f (y, x)}. f (y, x)) - (\\<^sup>+ y\{y. f (x, y) \ f (y, x)}. f (x, y))" using finite_IN finite_IN[of "{ _ }"] apply(simp add: d_IN_def nn_integral_count_space_indicator) apply(subst nn_integral_diff[symmetric]) apply(auto simp add: d_IN_def AE_count_space split: split_indicator intro!: arg_cong2[where f="(-)"] intro!: nn_integral_cong) done also have "\ = (\\<^sup>+ y\{y. f (x, y) \ f (y, x)}. f (y, x) - f (x, y))" using finite_OUT by(subst nn_integral_diff)(auto simp add: AE_count_space) also have "\ = d_IN (cleanup f) x" using finite_in by(auto simp add: d_IN_def nn_integral_count_space_indicator intro!: ennreal_diff_self nn_integral_cong split: split_indicator) finally show "KIR (cleanup f) x" . qed locale flow_attainability = countable_network \ for \ :: "('v, 'more) network_scheme" (structure) + assumes finite_capacity: "\x. x \ sink \ \ d_IN (capacity \) x \ \ \ d_OUT (capacity \) x \ \" and no_loop: "\x. \ edge \ x x" and source_in: "\x. \ edge \ x (source \)" begin lemma source_in_not_cycle: assumes "cycle \ p" shows "(x, source \) \ set (cycle_edges p)" using cycle_edges_edges[OF assms] source_in[of x] by(auto) lemma source_out_not_cycle: "cycle \ p \ (source \, x) \ set (cycle_edges p)" by(auto dest: cycle_leave_ex_enter source_in_not_cycle) lemma flowD_source_IN: assumes "flow \ f" shows "d_IN f (source \) = 0" proof - have "d_IN f (source \) = (\\<^sup>+ y\\<^bold>I\<^bold>N (source \). f (y, source \))" by(rule d_IN_alt_def)(simp add: flowD_outside[OF assms]) also have "\ = (\\<^sup>+ y\\<^bold>I\<^bold>N (source \). 0)" by(rule nn_integral_cong)(simp add: source_in incoming_def) finally show ?thesis by simp qed lemma flowD_finite_IN: assumes f: "flow \ f" and x: "x \ sink \" shows "d_IN f x \ top" proof(cases "x = source \") case True thus ?thesis by(simp add: flowD_source_IN[OF f]) next case False from finite_capacity[OF x] show ?thesis proof assume *: "d_IN (capacity \) x \ \" from flowD_capacity[OF f] have "d_IN f x \ d_IN (capacity \) x" by(rule d_IN_mono) also have "\ < \" using * by (simp add: less_top) finally show ?thesis by simp next assume *: "d_OUT (capacity \) x \ \" have "d_IN f x = d_OUT f x" using flowD_KIR[OF f False x] by simp also have "\ \ d_OUT (capacity \) x" using flowD_capacity[OF f] by(rule d_OUT_mono) also have "\ < \" using * by (simp add: less_top) finally show ?thesis by simp qed qed lemma flowD_finite_OUT: assumes "flow \ f" "x \ source \" "x \ sink \" shows "d_OUT f x \ \" using flowD_KIR[OF assms] assms by(simp add: flowD_finite_IN) end locale flow_network = flow_attainability + fixes g :: "'v flow" assumes g: "flow \ g" and g_finite: "value_flow \ g \ \" and nontrivial: "\<^bold>V - {source \, sink \} \ {}" begin lemma g_outside: "e \ \<^bold>E \ g e = 0" by(rule flowD_outside)(rule g) lemma g_loop [simp]: "g (x, x) = 0" by(rule g_outside)(simp add: no_loop) lemma finite_IN_g: "x \ sink \ \ d_IN g x \ top" by(rule flowD_finite_IN[OF g]) lemma finite_OUT_g: assumes "x \ sink \" shows "d_OUT g x \ top" proof(cases "x = source \") case True with g_finite show ?thesis by simp next case False with g have "KIR g x" using assms by(auto dest: flowD_KIR) with finite_IN_g[of x] False assms show ?thesis by(simp) qed lemma g_source_in [simp]: "g (x, source \) = 0" by(rule g_outside)(simp add: source_in) lemma finite_g [simp]: "g e \ top" by(rule flowD_finite[OF g]) definition enum_v :: "nat \ 'v" where "enum_v n = from_nat_into (\<^bold>V - {source \, sink \}) (fst (prod_decode n))" lemma range_enum_v: "range enum_v \ \<^bold>V - {source \, sink \}" using from_nat_into[OF nontrivial] by(auto simp add: enum_v_def) lemma enum_v_repeat: assumes x: "x \ \<^bold>V" "x \ source \" "x \ sink \" shows "\i'>i. enum_v i' = x" proof - let ?V = "\<^bold>V - {source \, sink \}" let ?n = "to_nat_on ?V x" let ?A = "{?n} \ (UNIV :: nat set)" from x have x': "x \ \<^bold>V - {source \, sink \}" by simp have "infinite ?A" by(auto dest: finite_cartesian_productD2) hence "infinite (prod_encode ` ?A)" by(auto dest: finite_imageD simp add: inj_prod_encode) then obtain i' where "i' > i" "i' \ prod_encode ` ?A" unfolding infinite_nat_iff_unbounded by blast from this(2) have "enum_v i' = x" using x by(clarsimp simp add: enum_v_def) with \i' > i\ show ?thesis by blast qed fun h_plus :: "nat \ 'v edge \ ennreal" where "h_plus 0 (x, y) = (if x = source \ then g (x, y) else 0)" | "h_plus (Suc i) (x, y) = (if enum_v (Suc i) = x \ d_OUT (h_plus i) x < d_IN (h_plus i) x then let total = d_IN (h_plus i) x - d_OUT (h_plus i) x; share = g (x, y) - h_plus i (x, y); shares = d_OUT g x - d_OUT (h_plus i) x in h_plus i (x, y) + share * total / shares else h_plus i (x, y))" lemma h_plus_le_g: "h_plus i e \ g e" proof(induction i arbitrary: e and e) case 0 thus ?case by(cases e) simp next case (Suc i) { fix x y assume enum: "x = enum_v (Suc i)" assume less: "d_OUT (h_plus i) x < d_IN (h_plus i) x" from enum have x: "x \ source \" "x \ sink \" using range_enum_v by(auto dest: sym intro: rev_image_eqI) define share where "share = g (x, y) - h_plus i (x, y)" define shares where "shares = d_OUT g x - d_OUT (h_plus i) x" define total where "total = d_IN (h_plus i) x - d_OUT (h_plus i) x" let ?h = "h_plus i (x, y) + share * total / shares" have "d_OUT (h_plus i) x \ d_OUT g x" by(rule d_OUT_mono)(rule Suc.IH) also have "\ < top" using finite_OUT_g[of x] x by (simp add: less_top) finally have "d_OUT (h_plus i) x \ \" by simp then have shares_eq: "shares = (\\<^sup>+ y. g (x, y) - h_plus i (x, y))" unfolding shares_def d_OUT_def by(subst nn_integral_diff)(simp_all add: AE_count_space Suc.IH) have *: "share / shares \ 1" proof (cases "share = 0") case True thus ?thesis by(simp) next case False hence "share > 0" using \h_plus i (x, y) \ g _\ by(simp add: share_def dual_order.strict_iff_order) moreover have "share \ shares" unfolding share_def shares_eq by(rule nn_integral_ge_point)simp ultimately show ?thesis by(simp add: divide_le_posI_ennreal) qed note shares_def also have "d_OUT g x = d_IN g x" by(rule flowD_KIR[OF g x]) also have "d_IN (h_plus i) x \ d_IN g x" by(rule d_IN_mono)(rule Suc.IH) ultimately have *: "total \ shares" unfolding total_def by(simp add: ennreal_minus_mono) moreover have "total > 0" unfolding total_def using less by (clarsimp simp add: diff_gr0_ennreal) ultimately have "total / shares \ 1" by(intro divide_le_posI_ennreal)(simp_all) hence "share * (total / shares) \ share * 1" by(rule mult_left_mono) simp hence "?h \ h_plus i (x, y) + share" by(simp add: ennreal_times_divide add_mono) also have "\ = g (x, y)" unfolding share_def using \h_plus i (x, y) \ g _\ finite_g[of "(x, y)"] by simp moreover note calculation } note * = this show ?case using Suc.IH * by(cases e) clarsimp qed lemma h_plus_outside: "e \ \<^bold>E \ h_plus i e = 0" by (metis g_outside h_plus_le_g le_zero_eq) lemma h_plus_not_infty [simp]: "h_plus i e \ top" using h_plus_le_g[of i e] by (auto simp: top_unique) lemma h_plus_mono: "h_plus i e \ h_plus (Suc i) e" proof(cases e) case [simp]: (Pair x y) { assume "d_OUT (h_plus i) x < d_IN (h_plus i) x" hence "h_plus i (x, y) + 0 \ h_plus i (x, y) + (g (x, y) - h_plus i (x, y)) * (d_IN (h_plus i) x - d_OUT (h_plus i) x) / (d_OUT g x - d_OUT (h_plus i) x)" by(intro add_left_mono d_OUT_mono le_funI) (simp_all add: h_plus_le_g) } then show ?thesis by clarsimp qed lemma h_plus_mono': "i \ j \ h_plus i e \ h_plus j e" by(induction rule: dec_induct)(auto intro: h_plus_mono order_trans) lemma d_OUT_h_plus_not_infty': "x \ sink \ \ d_OUT (h_plus i) x \ top" using d_OUT_mono[of "h_plus i" x g, OF h_plus_le_g] finite_OUT_g[of x] by (auto simp: top_unique) lemma h_plus_OUT_le_IN: assumes "x \ source \" shows "d_OUT (h_plus i) x \ d_IN (h_plus i) x" proof(induction i) case 0 thus ?case using assms by(simp add: d_OUT_def) next case (Suc i) have "d_OUT (h_plus (Suc i)) x \ d_IN (h_plus i) x" proof(cases "enum_v (Suc i) = x \ d_OUT (h_plus i) x < d_IN (h_plus i) x") case False thus ?thesis using Suc.IH by(simp add: d_OUT_def cong: conj_cong) next case True hence x: "x \ sink \" and le: "d_OUT (h_plus i) x < d_IN (h_plus i) x" using range_enum_v by auto let ?r = "\y. (g (x, y) - h_plus i (x, y)) * (d_IN (h_plus i) x - d_OUT (h_plus i) x) / (d_OUT g x - d_OUT (h_plus i) x)" have "d_OUT (h_plus (Suc i)) x = d_OUT (h_plus i) x + (\\<^sup>+ y. ?r y)" using True unfolding d_OUT_def h_plus.simps by(simp add: AE_count_space nn_integral_add) also from True have "x \ source \" "x \ sink \" using range_enum_v by auto from flowD_KIR[OF g this] le d_IN_mono[of "h_plus i" x g, OF h_plus_le_g] have le': "d_OUT (h_plus i) x < d_OUT g x" by(simp) then have "(\\<^sup>+ y. ?r y) = (d_IN (h_plus i) x - d_OUT (h_plus i) x) * ((\\<^sup>+ y. g (x, y) - h_plus i (x, y)) / (d_OUT g x - d_OUT (h_plus i) x))" by(subst mult.commute, subst ennreal_times_divide[symmetric]) (simp add: nn_integral_cmult nn_integral_divide Suc.IH diff_gr0_ennreal) also have "(\\<^sup>+ y. g (x, y) - h_plus i (x, y)) = d_OUT g x - d_OUT (h_plus i) x" using x by(subst nn_integral_diff)(simp_all add: d_OUT_def[symmetric] h_plus_le_g d_OUT_h_plus_not_infty') also have "\ / \ = 1" using le' finite_OUT_g[of x] x by(auto intro!: ennreal_divide_self dest: diff_gr0_ennreal simp: less_top[symmetric]) also have "d_OUT (h_plus i) x + (d_IN (h_plus i) x - d_OUT (h_plus i) x) * 1 = d_IN (h_plus i) x" using x by (simp add: Suc) finally show ?thesis by simp qed also have "\ \ d_IN (h_plus (Suc i)) x" by(rule d_IN_mono)(rule h_plus_mono) finally show ?case . qed lemma h_plus_OUT_eq_IN: assumes enum: "enum_v (Suc i) = x" shows "d_OUT (h_plus (Suc i)) x = d_IN (h_plus i) x" proof(cases "d_OUT (h_plus i) x < d_IN (h_plus i) x") case False from enum have "x \ source \" using range_enum_v by auto from h_plus_OUT_le_IN[OF this, of i] False have "d_OUT (h_plus i) x = d_IN (h_plus i) x" by auto with False enum show ?thesis by(simp add: d_OUT_def) next case True from enum have x: "x \ source \" and sink: "x \ sink \" using range_enum_v by auto let ?r = "\y. (g (x, y) - h_plus i (x, y)) * (d_IN (h_plus i) x - d_OUT (h_plus i) x) / (d_OUT g x - d_OUT (h_plus i) x)" have "d_OUT (h_plus (Suc i)) x = d_OUT (h_plus i) x + (\\<^sup>+ y. ?r y)" using True enum unfolding d_OUT_def h_plus.simps by(simp add: AE_count_space nn_integral_add) also from True enum have "x \ source \" "x \ sink \" using range_enum_v by auto from flowD_KIR[OF g this] True d_IN_mono[of "h_plus i" x g, OF h_plus_le_g] have le': "d_OUT (h_plus i) x < d_OUT g x" by(simp) then have "(\\<^sup>+ y. ?r y ) = (d_IN (h_plus i) x - d_OUT (h_plus i) x) * ((\\<^sup>+ y. g (x, y) - h_plus i (x, y)) / (d_OUT g x - d_OUT (h_plus i) x))" by(subst mult.commute, subst ennreal_times_divide[symmetric]) (simp add: nn_integral_cmult nn_integral_divide h_plus_OUT_le_IN[OF x] diff_gr0_ennreal) also have "(\\<^sup>+ y. g (x, y) - h_plus i (x, y)) = d_OUT g x - d_OUT (h_plus i) x" using sink by(subst nn_integral_diff)(simp_all add: d_OUT_def[symmetric] h_plus_le_g d_OUT_h_plus_not_infty') also have "\ / \ = 1" using le' finite_OUT_g[of x] sink by(auto intro!: ennreal_divide_self dest: diff_gr0_ennreal simp: less_top[symmetric]) also have "d_OUT (h_plus i) x + (d_IN (h_plus i) x - d_OUT (h_plus i) x) * 1 = d_IN (h_plus i) x" using sink by (simp add: h_plus_OUT_le_IN x) finally show ?thesis . qed lemma h_plus_source_in [simp]: "h_plus i (x, source \) = 0" by(induction i)simp_all lemma h_plus_sum_finite: "(\\<^sup>+ e. h_plus i e) \ top" proof(induction i) case 0 have "(\\<^sup>+ e\UNIV. h_plus 0 e) = (\\<^sup>+ (x, y). h_plus 0 (x, y))" by(simp del: h_plus.simps) also have "\ = (\\<^sup>+ (x, y)\range (Pair (source \)). h_plus 0 (x, y))" by(auto simp add: nn_integral_count_space_indicator intro!: nn_integral_cong) also have "\ = value_flow \ g" by(simp add: d_OUT_def nn_integral_count_space_reindex) also have "\ < \" using g_finite by (simp add: less_top) finally show ?case by simp next case (Suc i) define xi where "xi = enum_v (Suc i)" then have xi: "xi \ source \" "xi \ sink \" using range_enum_v by auto show ?case proof(cases "d_OUT (h_plus i) xi < d_IN (h_plus i) xi") case False hence "(\\<^sup>+ e\UNIV. h_plus (Suc i) e) = (\\<^sup>+ e. h_plus i e)" by(auto intro!: nn_integral_cong simp add: xi_def) with Suc.IH show ?thesis by simp next case True have less: "d_OUT (h_plus i) xi < d_OUT g xi" using True flowD_KIR[OF g xi] d_IN_mono[of "h_plus i" xi, OF h_plus_le_g] by simp have "(\\<^sup>+ e. h_plus (Suc i) e) = (\\<^sup>+ e\UNIV. h_plus i e) + (\\<^sup>+ (x, y). ((g (x, y) - h_plus i (x, y)) * (d_IN (h_plus i) x - d_OUT (h_plus i) x) / (d_OUT g x - d_OUT (h_plus i) x)) * indicator (range (Pair xi)) (x, y))" (is "_ = ?IH + ?rest" is "_ = _ + \\<^sup>+ (x, y). ?f x y * _ \_") using xi True by(subst nn_integral_add[symmetric])(auto simp add: xi_def split_beta AE_count_space intro!: nn_integral_cong split: split_indicator intro!: h_plus_le_g h_plus_OUT_le_IN d_OUT_mono le_funI) also have "?rest = (\\<^sup>+ (x, y)\range (Pair xi). ?f x y)" by(simp add: nn_integral_count_space_indicator split_def) also have "\ = (\\<^sup>+ y. ?f xi y)" by(simp add: nn_integral_count_space_reindex) also have "\ = (\\<^sup>+ y. g (xi, y) - h_plus i (xi, y)) * ((d_IN (h_plus i) xi - d_OUT (h_plus i) xi) / (d_OUT g xi - d_OUT (h_plus i) xi))" (is "_ = ?integral * ?factor") using True less by(simp add: nn_integral_multc nn_integral_divide diff_gr0_ennreal ennreal_times_divide) also have "?integral = d_OUT g xi - d_OUT (h_plus i) xi" unfolding d_OUT_def using xi by(subst nn_integral_diff)(simp_all add: h_plus_le_g d_OUT_def[symmetric] d_OUT_h_plus_not_infty') also have "\ * ?factor = (d_IN (h_plus i) xi - d_OUT (h_plus i) xi)" using xi apply (subst ennreal_times_divide) apply (subst mult.commute) apply (subst ennreal_mult_divide_eq) apply (simp_all add: diff_gr0_ennreal finite_OUT_g less zero_less_iff_neq_zero[symmetric]) done also have "\ \ \" using h_plus_OUT_eq_IN[OF refl, of i, folded xi_def, symmetric] xi by(simp add: d_OUT_h_plus_not_infty') ultimately show ?thesis using Suc.IH by simp qed qed lemma d_OUT_h_plus_not_infty [simp]: "d_OUT (h_plus i) x \ top" proof - have "d_OUT (h_plus i) x \ (\\<^sup>+ y\UNIV. \\<^sup>+ x. h_plus i (x, y))" unfolding d_OUT_def by(rule nn_integral_mono nn_integral_ge_point)+ simp also have "\ < \" using h_plus_sum_finite by(simp add: nn_integral_snd_count_space less_top) finally show ?thesis by simp qed definition enum_cycle :: "nat \ 'v path" where "enum_cycle = from_nat_into (cycles \)" lemma cycle_enum_cycle [simp]: "cycles \ \ {} \ cycle \ (enum_cycle n)" unfolding enum_cycle_def using from_nat_into[of "cycles \" n] by simp context fixes h' :: "'v flow" assumes finite_h': "h' e \ top" begin fun h_minus_aux :: "nat \ 'v edge \ ennreal" where "h_minus_aux 0 e = 0" | "h_minus_aux (Suc j) e = (if e \ set (cycle_edges (enum_cycle j)) then h_minus_aux j e + Min {h' e' - h_minus_aux j e'|e'. e'\set (cycle_edges (enum_cycle j))} else h_minus_aux j e)" lemma h_minus_aux_le_h': "h_minus_aux j e \ h' e" proof(induction j e rule: h_minus_aux.induct) case 0: (1 e) show ?case by simp next case Suc: (2 j e) { assume e: "e \ set (cycle_edges (enum_cycle j))" then have "h_minus_aux j e + Min {h' e' - h_minus_aux j e' |e'. e' \ set (cycle_edges (enum_cycle j))} \ h_minus_aux j e + (h' e - h_minus_aux j e)" using [[simproc add: finite_Collect]] by(cases e rule: prod.exhaust)(auto intro!: add_mono Min_le) also have "\ = h' e" using e finite_h'[of e] Suc.IH(2)[of e] by(cases e rule: prod.exhaust) (auto simp add: add_diff_eq_ennreal top_unique intro!: ennreal_add_diff_cancel_left) also note calculation } then show ?case using Suc by clarsimp qed lemma h_minus_aux_finite [simp]: "h_minus_aux j e \ top" using h_minus_aux_le_h'[of j e] finite_h'[of e] by (auto simp: top_unique) lemma h_minus_aux_mono: "h_minus_aux j e \ h_minus_aux (Suc j) e" proof(cases "e \ set (cycle_edges (enum_cycle j)) = True") case True have "h_minus_aux j e + 0 \ h_minus_aux (Suc j) e" unfolding h_minus_aux.simps True if_True using True [[simproc add: finite_Collect]] by(cases e)(rule add_mono, auto intro!: Min.boundedI simp add: h_minus_aux_le_h') thus ?thesis by simp qed simp lemma d_OUT_h_minus_aux: assumes "cycles \ \ {}" shows "d_OUT (h_minus_aux j) x = d_IN (h_minus_aux j) x" proof(induction j) case 0 show ?case by simp next case (Suc j) define C where "C = enum_cycle j" define \ where "\ = Min {h' e' - h_minus_aux j e' |e'. e' \ set (cycle_edges C)}" have "d_OUT (h_minus_aux (Suc j)) x = (\\<^sup>+ y. h_minus_aux j (x, y) + (if (x, y) \ set (cycle_edges C) then \ else 0))" unfolding d_OUT_def by(simp add: if_distrib C_def \_def cong del: if_weak_cong) also have "\ = d_OUT (h_minus_aux j) x + (\\<^sup>+ y. \ * indicator (set (cycle_edges C)) (x, y))" (is "_ = _ + ?add") by(subst nn_integral_add)(auto simp add: AE_count_space d_OUT_def intro!: arg_cong2[where f="(+)"] nn_integral_cong) also have "?add = (\\<^sup>+ e\range (Pair x). \ * indicator {(x', y). (x', y) \ set (cycle_edges C) \ x' = x} e)" by(auto simp add: nn_integral_count_space_reindex intro!: nn_integral_cong split: split_indicator) also have "\ = \ * card (set (filter (\(x', y). x' = x) (cycle_edges C)))" using [[simproc add: finite_Collect]] apply(subst nn_integral_cmult_indicator; auto) apply(subst emeasure_count_space; auto simp add: split_def) done also have "card (set (filter (\(x', y). x' = x) (cycle_edges C))) = card (set (filter (\(x', y). y = x) (cycle_edges C)))" unfolding C_def by(rule cycle_enter_leave_same)(rule cycle_enum_cycle[OF assms]) also have "\ * \ = (\\<^sup>+ e\range (\x'. (x', x)). \ * indicator {(x', y). (x', y) \ set (cycle_edges C) \ y = x} e)" using [[simproc add: finite_Collect]] apply(subst nn_integral_cmult_indicator; auto) apply(subst emeasure_count_space; auto simp add: split_def) done also have "\ = (\\<^sup>+ x'. \ * indicator (set (cycle_edges C)) (x', x))" by(auto simp add: nn_integral_count_space_reindex intro!: nn_integral_cong split: split_indicator) also have "d_OUT (h_minus_aux j) x + \ = (\\<^sup>+ x'. h_minus_aux j (x', x) + \ * indicator (set (cycle_edges C)) (x', x))" unfolding Suc.IH d_IN_def by(simp add: nn_integral_add[symmetric]) also have "\ = d_IN (h_minus_aux (Suc j)) x" unfolding d_IN_def by(auto intro!: nn_integral_cong simp add: \_def C_def split: split_indicator) finally show ?case . qed lemma h_minus_aux_source: assumes "cycles \ \ {}" shows "h_minus_aux j (source \, y) = 0" proof(induction j) case 0 thus ?case by simp next case (Suc j) have "(source \, y) \ set (cycle_edges (enum_cycle j))" proof assume *: "(source \, y) \ set (cycle_edges (enum_cycle j))" have cycle: "cycle \ (enum_cycle j)" using assms by(rule cycle_enum_cycle) from cycle_leave_ex_enter[OF this *] obtain z where "(z, source \) \ set (cycle_edges (enum_cycle j))" .. with cycle_edges_edges[OF cycle] have "(z, source \) \ \<^bold>E" .. thus False using source_in[of z] by simp qed then show ?case using Suc.IH by simp qed lemma h_minus_aux_cycle: fixes j defines "C \ enum_cycle j" assumes "cycles \ \ {}" shows "\e\set (cycle_edges C). h_minus_aux (Suc j) e = h' e" proof - let ?A = "{h' e' - h_minus_aux j e'|e'. e' \ set (cycle_edges C)}" from assms have "cycle \ C" by auto from cycle_edges_not_Nil[OF this] have "Min ?A \ ?A" using [[simproc add: finite_Collect]] by(intro Min_in)(fastforce simp add: neq_Nil_conv)+ then obtain e' where e: "e' \ set (cycle_edges C)" and "Min ?A = h' e' - h_minus_aux j e'" by auto hence "h_minus_aux (Suc j) e' = h' e'" by(simp add: C_def h_minus_aux_le_h') with e show ?thesis by blast qed end fun h_minus :: "nat \ 'v edge \ ennreal" where "h_minus 0 e = 0" | "h_minus (Suc i) e = h_minus i e + (SUP j. h_minus_aux (\e'. h_plus (Suc i) e' - h_minus i e') j e)" lemma h_minus_le_h_plus: "h_minus i e \ h_plus i e" proof(induction i e rule: h_minus.induct) case 0: (1 e) show ?case by simp next case Suc: (2 i e) note IH = Suc.IH(2)[OF UNIV_I] let ?h' = "\e'. h_plus (Suc i) e' - h_minus i e'" - have h': "?h' e' \ top" for e' using IH(1)[of e'] bysimp + have h': "?h' e' \ top" for e' using IH(1)[of e'] by simp have "(\j. h_minus_aux ?h' j e) \ ?h' e" by(rule SUP_least)(rule h_minus_aux_le_h'[OF h']) hence "h_minus (Suc i) e \ h_minus i e + \" by(simp add: add_mono) also have "\ = h_plus (Suc i) e" using IH[of e] h_plus_mono[of i e] by auto finally show ?case . qed lemma finite_h': "h_plus (Suc i) e - h_minus i e \ top" by simp lemma h_minus_mono: "h_minus i e \ h_minus (Suc i) e" proof - have "h_minus i e + 0 \ h_minus (Suc i) e" unfolding h_minus.simps by(rule add_mono; simp add: SUP_upper2) thus ?thesis by simp qed lemma h_minus_finite [simp]: "h_minus i e \ \" proof - have "h_minus i e \ h_plus i e" by(rule h_minus_le_h_plus) also have "\ < \" by (simp add: less_top[symmetric]) finally show ?thesis by simp qed lemma d_OUT_h_minus: assumes cycles: "cycles \ \ {}" shows "d_OUT (h_minus i) x = d_IN (h_minus i) x" proof(induction i) case (Suc i) let ?h' = "\e. h_plus (Suc i) e - h_minus i e" have "d_OUT (\e. h_minus (Suc i) e) x = d_OUT (h_minus i) x + d_OUT (\e. SUP j. h_minus_aux ?h' j e) x" by(simp add: d_OUT_add SUP_upper2) also have "d_OUT (\e. SUP j. h_minus_aux ?h' j e) x = (SUP j. d_OUT (h_minus_aux ?h' j) x)" by(rule d_OUT_monotone_convergence_SUP incseq_SucI le_funI h_minus_aux_mono finite_h')+ also have "\ = (SUP j. d_IN (h_minus_aux ?h' j) x)" by(rule SUP_cong[OF refl])(rule d_OUT_h_minus_aux[OF finite_h' cycles]) also have "\ = d_IN (\e. SUP j. h_minus_aux ?h' j e) x" by(rule d_IN_monotone_convergence_SUP[symmetric] incseq_SucI le_funI h_minus_aux_mono finite_h')+ also have "d_OUT (h_minus i) x + \ = d_IN (\e. h_minus (Suc i) e) x" using Suc.IH by(simp add: d_IN_add SUP_upper2) finally show ?case . qed simp lemma h_minus_source: assumes "cycles \ \ {}" shows "h_minus n (source \, y) = 0" by(induction n)(simp_all add: h_minus_aux_source[OF finite_h' assms]) lemma h_minus_source_in [simp]: "h_minus i (x, source \) = 0" using h_minus_le_h_plus[of i "(x, source \)"] by simp lemma h_minus_OUT_finite [simp]: "d_OUT (h_minus i) x \ top" proof - have "d_OUT (h_minus i) x \ d_OUT (h_plus i) x" by(rule d_OUT_mono)(rule h_minus_le_h_plus) also have "\ < \" by (simp add: less_top[symmetric]) finally show ?thesis by simp qed lemma h_minus_cycle: assumes "cycle \ C" shows "\e\set (cycle_edges C). h_minus i e = h_plus i e" proof(cases i) case (Suc i) let ?h' = "\e. h_plus (Suc i) e - h_minus i e" from assms have cycles: "cycles \ \ {}" by auto with assms from_nat_into_surj[of "cycles \" C] obtain j where j: "C = enum_cycle j" by(auto simp add: enum_cycle_def) from h_minus_aux_cycle[of "?h'" j, OF finite_h' cycles] j obtain e where e: "e \ set (cycle_edges C)" and "h_minus_aux ?h' (Suc j) e = ?h' e" by(auto) then have "h_plus (Suc i) e = h_minus i e + h_minus_aux ?h' (Suc j) e" using order_trans[OF h_minus_le_h_plus h_plus_mono] by (subst eq_commute) simp also have "\ \ h_minus (Suc i) e" unfolding h_minus.simps by(intro add_mono SUP_upper; simp) finally show ?thesis using e h_minus_le_h_plus[of "Suc i" e] Suc by auto next case 0 from cycle_edges_not_Nil[OF assms] obtain x y where e: "(x, y) \ set (cycle_edges C)" by(fastforce simp add: neq_Nil_conv) then have "x \ source \" using assms by(auto dest: source_out_not_cycle) hence "h_plus 0 (x, y) = 0" by simp with e 0 show ?thesis by(auto simp del: h_plus.simps) qed abbreviation lim_h_plus :: "'v edge \ ennreal" where "lim_h_plus e \ SUP n. h_plus n e" abbreviation lim_h_minus :: "'v edge \ ennreal" where "lim_h_minus e \ SUP n. h_minus n e" lemma lim_h_plus_le_g: "lim_h_plus e \ g e" by(rule SUP_least)(rule h_plus_le_g) lemma lim_h_plus_finite [simp]: "lim_h_plus e \ top" proof - have "lim_h_plus e \ g e" by(rule lim_h_plus_le_g) also have "\ < top" by (simp add: less_top[symmetric]) finally show ?thesis unfolding less_top . qed lemma lim_h_minus_le_lim_h_plus: "lim_h_minus e \ lim_h_plus e" by(rule SUP_mono)(blast intro: h_minus_le_h_plus) lemma lim_h_minus_finite [simp]: "lim_h_minus e \ top" proof - have "lim_h_minus e \ lim_h_plus e" by(rule lim_h_minus_le_lim_h_plus) also have "\ < top" unfolding less_top[symmetric] by (rule lim_h_plus_finite) finally show ?thesis unfolding less_top[symmetric] by simp qed lemma lim_h_minus_IN_finite [simp]: assumes "x \ sink \" shows "d_IN lim_h_minus x \ top" proof - have "d_IN lim_h_minus x \ d_IN lim_h_plus x" by(intro d_IN_mono le_funI lim_h_minus_le_lim_h_plus) also have "\ \ d_IN g x" by(intro d_IN_mono le_funI lim_h_plus_le_g) also have "\ < \" using assms by(simp add: finite_IN_g less_top[symmetric]) finally show ?thesis by simp qed lemma lim_h_plus_OUT_IN: assumes "x \ source \" "x \ sink \" shows "d_OUT lim_h_plus x = d_IN lim_h_plus x" proof(cases "x \ \<^bold>V") case True have "d_OUT lim_h_plus x = (SUP n. d_OUT (h_plus n) x)" by(rule d_OUT_monotone_convergence_SUP incseq_SucI le_funI h_plus_mono)+ also have "\ = (SUP n. d_IN (h_plus n) x)" (is "?lhs = ?rhs") proof(rule antisym) show "?lhs \ ?rhs" by(rule SUP_mono)(auto intro: h_plus_OUT_le_IN[OF assms(1)]) show "?rhs \ ?lhs" proof(rule SUP_mono) fix i from enum_v_repeat[OF True assms, of i] obtain i' where "i' > i" "enum_v i' = x" by auto moreover then obtain i'' where i': "i' = Suc i''" by(cases i') auto ultimately have "d_OUT (h_plus i') x = d_IN (h_plus i'') x" using \x \ source \\ by(simp add: h_plus_OUT_eq_IN) moreover have "i \ i''" using \i < i'\ i' by simp then have "d_IN (h_plus i) x \ d_IN (h_plus i'') x" by(intro d_IN_mono h_plus_mono') ultimately have "d_IN (h_plus i) x \ d_OUT (h_plus i') x" by simp thus "\i'\UNIV. d_IN (h_plus i) x \ d_OUT (h_plus i') x" by blast qed qed also have "\ = d_IN lim_h_plus x" by(rule d_IN_monotone_convergence_SUP[symmetric] incseq_SucI le_funI h_plus_mono)+ finally show ?thesis . next case False have "(x, y) \ support_flow lim_h_plus" for y using False h_plus_outside[of "(x, y)"] by(fastforce elim!: support_flow.cases simp add: less_SUP_iff vertex_def) moreover have "(y, x) \ support_flow lim_h_plus" for y using False h_plus_outside[of "(y, x)"] by(fastforce elim!: support_flow.cases simp add: less_SUP_iff vertex_def) ultimately show ?thesis by(auto simp add: d_OUT_alt_def2 d_IN_alt_def2 AE_count_space intro!: nn_integral_cong_AE) qed lemma lim_h_minus_OUT_IN: assumes cycles: "cycles \ \ {}" shows "d_OUT lim_h_minus x = d_IN lim_h_minus x" proof - have "d_OUT lim_h_minus x = (SUP n. d_OUT (h_minus n) x)" by(rule d_OUT_monotone_convergence_SUP incseq_SucI le_funI h_minus_mono)+ also have "\ = (SUP n. d_IN (h_minus n) x)" using cycles by(simp add: d_OUT_h_minus) also have "\ = d_IN lim_h_minus x" by(rule d_IN_monotone_convergence_SUP[symmetric] incseq_SucI le_funI h_minus_mono)+ finally show ?thesis . qed definition h :: "'v edge \ ennreal" where "h e = lim_h_plus e - (if cycles \ \ {} then lim_h_minus e else 0)" lemma h_le_lim_h_plus: "h e \ lim_h_plus e" by (simp add: h_def) lemma h_le_g: "h e \ g e" using h_le_lim_h_plus[of e] lim_h_plus_le_g[of e] by simp lemma flow_h: "flow \ h" proof fix e have "h e \ lim_h_plus e" by(rule h_le_lim_h_plus) also have "\ \ g e" by(rule lim_h_plus_le_g) also have "\ \ capacity \ e" using g by(rule flowD_capacity) finally show "h e \ \" . next fix x assume "x \ source \" "x \ sink \" then show "KIR h x" by (cases "cycles \ = {}") (auto simp add: h_def[abs_def] lim_h_plus_OUT_IN d_OUT_diff d_IN_diff lim_h_minus_le_lim_h_plus lim_h_minus_OUT_IN) qed lemma value_h_plus: "value_flow \ (h_plus i) = value_flow \ g" (is "?lhs = ?rhs") proof(rule antisym) show "?lhs \ ?rhs" by(rule d_OUT_mono)(rule h_plus_le_g) have "?rhs \ value_flow \ (h_plus 0)" by(auto simp add: d_OUT_def cong: if_cong intro!: nn_integral_mono) also have "\ \ value_flow \ (h_plus i)" by(rule d_OUT_mono)(rule h_plus_mono'; simp) finally show "?rhs \ ?lhs" . qed lemma value_h: "value_flow \ h = value_flow \ g" (is "?lhs = ?rhs") proof(rule antisym) have "?lhs \ value_flow \ lim_h_plus" using ennreal_minus_mono by(fastforce simp add: h_def intro!: d_OUT_mono) also have "\ \ ?rhs" by(rule d_OUT_mono)(rule lim_h_plus_le_g) finally show "?lhs \ ?rhs" . show "?rhs \ ?lhs" by(auto simp add: d_OUT_def h_def h_minus_source cong: if_cong intro!: nn_integral_mono SUP_upper2[where i=0]) qed definition h_diff :: "nat \ 'v edge \ ennreal" where "h_diff i e = h_plus i e - (if cycles \ \ {} then h_minus i e else 0)" lemma d_IN_h_source [simp]: "d_IN (h_diff i) (source \) = 0" by(simp add: d_IN_def h_diff_def cong del: if_weak_cong) lemma h_diff_le_h_plus: "h_diff i e \ h_plus i e" by(simp add: h_diff_def) lemma h_diff_le_g: "h_diff i e \ g e" using h_diff_le_h_plus[of i e] h_plus_le_g[of i e] by simp lemma h_diff_loop [simp]: "h_diff i (x, x) = 0" using h_diff_le_g[of i "(x, x)"] by simp lemma supp_h_diff_edges: "support_flow (h_diff i) \ \<^bold>E" proof fix e assume "e \ support_flow (h_diff i)" then have "0 < h_diff i e" by(auto elim: support_flow.cases) also have "h_diff i e \ h_plus i e" by(rule h_diff_le_h_plus) finally show "e \ \<^bold>E" using h_plus_outside[of e i] by(cases "e \ \<^bold>E") auto qed lemma h_diff_OUT_le_IN: assumes "x \ source \" shows "d_OUT (h_diff i) x \ d_IN (h_diff i) x" proof(cases "cycles \ \ {}") case False thus ?thesis using assms by(simp add: h_diff_def[abs_def] h_plus_OUT_le_IN) next case cycles: True then have "d_OUT (h_diff i) x = d_OUT (h_plus i) x - d_OUT (h_minus i) x" unfolding h_diff_def[abs_def] using assms by (simp add: h_minus_le_h_plus d_OUT_diff) also have "\ \ d_IN (h_plus i) x - d_IN (h_minus i) x" using cycles assms by(intro ennreal_minus_mono h_plus_OUT_le_IN)(simp_all add: d_OUT_h_minus) also have "\ = d_IN (h_diff i) x" using cycles unfolding h_diff_def[abs_def] by(subst d_IN_diff)(simp_all add: h_minus_le_h_plus d_OUT_h_minus[symmetric]) finally show ?thesis . qed lemma h_diff_cycle: assumes "cycle \ p" shows "\e\set (cycle_edges p). h_diff i e = 0" proof - from h_minus_cycle[OF assms, of i] obtain e where e: "e \ set (cycle_edges p)" and "h_minus i e = h_plus i e" by auto hence "h_diff i e = 0" using assms by(auto simp add: h_diff_def) with e show ?thesis by blast qed lemma d_IN_h_le_value': "d_IN (h_diff i) x \ value_flow \ (h_plus i)" proof - let ?supp = "support_flow (h_diff i)" define X where "X = {y. (y, x) \ ?supp^*} - {x}" { fix x y assume x: "x \ X" and y: "y \ X" { assume yx: "(y, x) \ ?supp\<^sup>*" and neq: "y \ x" and xy: "(x, y) \ ?supp" from yx obtain p' where "rtrancl_path (\x y. (x, y) \ ?supp) y p' x" unfolding rtrancl_def rtranclp_eq_rtrancl_path by auto then obtain p where p: "rtrancl_path (\x y. (x, y) \ ?supp) y p x" and distinct: "distinct (y # p)" by(rule rtrancl_path_distinct) with neq have "p \ []" by(auto elim: rtrancl_path.cases) from xy have "(x, y) \ \<^bold>E" using supp_h_diff_edges[of i] by(auto) moreover from p have "path \ y p x" by(rule rtrancl_path_mono)(auto dest: supp_h_diff_edges[THEN subsetD]) ultimately have "path \ x (y # p) x" by(auto intro: rtrancl_path.intros) hence cycle: "cycle \ (y # p)" using _ distinct by(rule cycle) simp from h_diff_cycle[OF this, of i] obtain e where e: "e \ set (cycle_edges (y # p))" and 0: "h_diff i e = 0" by blast from e obtain n where e': "e = ((y # p) ! n, (p @ [y]) ! n)" and n: "n < Suc (length p)" by(auto simp add: cycle_edges_def set_zip) have "e \ ?supp" proof(cases "n = length p") case True with rtrancl_path_last[OF p] \p \ []\ have "(y # p) ! n = x" by(cases p)(simp_all add: last_conv_nth del: last.simps) with e' True have "e = (x, y)" by simp with xy show ?thesis by simp next case False with n have "n < length p" by simp with rtrancl_path_nth[OF p this] e' show ?thesis by(simp add: nth_append) qed with 0 have False by(simp add: support_flow.simps) } hence "(x, y) \ ?supp" using x y by(auto simp add: X_def intro: converse_rtrancl_into_rtrancl) then have "h_diff i (x, y) = 0" by(simp add: support_flow.simps) } note acyclic = this { fix y assume "y \ X" hence "(y, x) \ ?supp" by(auto simp add: X_def support_flow.simps intro: not_in_support_flowD) hence "h_diff i (y, x) = 0" by(simp add: support_flow.simps) } note in_X = this let ?diff = "\x. (\\<^sup>+ y. h_diff i (x, y) * indicator X x * indicator X y)" have finite2: "(\\<^sup>+ x. ?diff x) \ top" (is "?lhs \ _") proof - have "?lhs \ (\\<^sup>+ x\UNIV. \\<^sup>+ y. h_plus i (x, y))" by(intro nn_integral_mono)(auto simp add: h_diff_def split: split_indicator) also have "\ = (\\<^sup>+ e. h_plus i e)" by(rule nn_integral_fst_count_space) also have "\ < \" by(simp add: h_plus_sum_finite less_top[symmetric]) finally show ?thesis by simp qed have finite1: "?diff x \ top" for x using finite2 by(rule neq_top_trans)(rule nn_integral_ge_point, simp) have finite3: "(\\<^sup>+ x. d_OUT (h_diff i) x * indicator (X - {source \}) x) \ \" (is "?lhs \ _") proof - have "?lhs \ (\\<^sup>+ x\UNIV. \\<^sup>+ y. h_plus i (x, y))" unfolding d_OUT_def apply(simp add: nn_integral_multc[symmetric]) apply(intro nn_integral_mono) apply(auto simp add: h_diff_def split: split_indicator) done also have "\ = (\\<^sup>+ e. h_plus i e)" by(rule nn_integral_fst_count_space) also have "\ < \" by(simp add: h_plus_sum_finite less_top[symmetric]) finally show ?thesis by simp qed have "d_IN (h_diff i) x = (\\<^sup>+ y. h_diff i (y, x) * indicator X y)" unfolding d_IN_def by(rule nn_integral_cong)(simp add: in_X split: split_indicator) also have "\ \ (\\<^sup>+ x\- X. \\<^sup>+ y. h_diff i (y, x) * indicator X y)" by(rule nn_integral_ge_point)(simp add: X_def) also have "\ = (\\<^sup>+ x\UNIV. \\<^sup>+ y. h_diff i (y, x) * indicator X y * indicator (- X) x)" by(simp add: nn_integral_multc nn_integral_count_space_indicator) also have "\ = (\\<^sup>+ x\UNIV. \\<^sup>+ y. h_diff i (x, y) * indicator X x * indicator (- X) y)" by(subst nn_integral_snd_count_space[where f="case_prod _", simplified])(simp add: nn_integral_fst_count_space[where f="case_prod _", simplified]) also have "\ = (\\<^sup>+ x\UNIV. (\\<^sup>+ y. h_diff i (x, y) * indicator X x * indicator (- X) y) + (?diff x - ?diff x))" by(simp add: finite1) also have "\ = (\\<^sup>+ x\UNIV. (\\<^sup>+ y. h_diff i (x, y) * indicator X x * indicator (- X) y + h_diff i (x, y) * indicator X x * indicator X y) - ?diff x)" apply (subst add_diff_eq_ennreal) apply simp by(subst nn_integral_add[symmetric])(simp_all add:) also have "\ = (\\<^sup>+ x\UNIV. (\\<^sup>+ y. h_diff i (x, y) * indicator X x) - ?diff x)" by(auto intro!: nn_integral_cong arg_cong2[where f="(-)"] split: split_indicator) also have "\ = (\\<^sup>+ x\UNIV. \\<^sup>+ y\UNIV. h_diff i (x, y) * indicator X x) - (\\<^sup>+ x. ?diff x)" by(subst nn_integral_diff)(auto simp add: AE_count_space finite2 intro!: nn_integral_mono split: split_indicator) also have "(\\<^sup>+ x\UNIV. \\<^sup>+ y\UNIV. h_diff i (x, y) * indicator X x) = (\\<^sup>+ x. d_OUT (h_diff i) x * indicator X x)" unfolding d_OUT_def by(simp add: nn_integral_multc) also have "\ = (\\<^sup>+ x. d_OUT (h_diff i) x * indicator (X - {source \}) x + value_flow \ (h_diff i) * indicator X (source \) * indicator {source \} x)" by(rule nn_integral_cong)(simp split: split_indicator) also have "\ = (\\<^sup>+ x. d_OUT (h_diff i) x * indicator (X - {source \}) x) + value_flow \ (h_diff i) * indicator X (source \)" (is "_ = ?out" is "_ = _ + ?value") by(subst nn_integral_add) simp_all also have "(\\<^sup>+ x\UNIV. \\<^sup>+ y. h_diff i (x, y) * indicator X x * indicator X y) = (\\<^sup>+ x\UNIV. \\<^sup>+ y. h_diff i (x, y) * indicator X y)" using acyclic by(intro nn_integral_cong)(simp split: split_indicator) also have "\ = (\\<^sup>+ y\UNIV. \\<^sup>+ x. h_diff i (x, y) * indicator X y)" by(subst nn_integral_snd_count_space[where f="case_prod _", simplified])(simp add: nn_integral_fst_count_space[where f="case_prod _", simplified]) also have "\ = (\\<^sup>+ y. d_IN (h_diff i) y * indicator X y)" unfolding d_IN_def by(simp add: nn_integral_multc) also have "\ = (\\<^sup>+ y. d_IN (h_diff i) y * indicator (X - {source \}) y)" by(rule nn_integral_cong)(simp split: split_indicator) also have "?out - \ \ (\\<^sup>+ x. d_OUT (h_diff i) x * indicator (X - {source \}) x) - \ + ?value" by (auto simp add: add_ac intro!: add_diff_le_ennreal) also have "\ \ 0 + ?value" using h_diff_OUT_le_IN finite3 by(intro nn_integral_mono add_right_mono)(auto split: split_indicator intro!: diff_eq_0_ennreal nn_integral_mono simp add: less_top) also have "\ \ value_flow \ (h_diff i)" by(simp split: split_indicator) also have "\ \ value_flow \ (h_plus i)" by(rule d_OUT_mono le_funI h_diff_le_h_plus)+ finally show ?thesis . qed lemma d_IN_h_le_value: "d_IN h x \ value_flow \ h" (is "?lhs \ ?rhs") proof - have [tendsto_intros]: "(\i. h_plus i e) \ lim_h_plus e" for e by(rule LIMSEQ_SUP incseq_SucI h_plus_mono)+ have [tendsto_intros]: "(\i. h_minus i e) \ lim_h_minus e" for e by(rule LIMSEQ_SUP incseq_SucI h_minus_mono)+ have "(\i. h_diff i e) \ lim_h_plus e - (if cycles \ \ {} then lim_h_minus e else 0)" for e by(auto intro!: tendsto_intros tendsto_diff_ennreal simp add: h_diff_def simp del: Sup_eq_top_iff SUP_eq_top_iff) then have "d_IN h x = (\\<^sup>+ y. liminf (\i. h_diff i (y, x)))" by(simp add: d_IN_def h_def tendsto_iff_Liminf_eq_Limsup) also have "\ \ liminf (\i. d_IN (h_diff i) x)" unfolding d_IN_def by(rule nn_integral_liminf) simp_all also have "\ \ liminf (\i. value_flow \ h)" using d_IN_h_le_value'[of _ x] by(intro Liminf_mono eventually_sequentiallyI)(auto simp add: value_h_plus value_h) also have "\ = value_flow \ h" by(simp add: Liminf_const) finally show ?thesis . qed lemma flow_cleanup: \ \Lemma 5.4\ "\h \ g. flow \ h \ value_flow \ h = value_flow \ g \ (\x. d_IN h x \ value_flow \ h)" by(intro exI[where x=h] conjI strip le_funI d_IN_h_le_value flow_h value_h h_le_g) end subsection \Residual network\ context countable_network begin definition residual_network :: "'v flow \ ('v, 'more) network_scheme" where "residual_network f = \edge = \x y. edge \ x y \ edge \ y x \ y \ source \, capacity = \(x, y). if edge \ x y then capacity \ (x, y) - f (x, y) else if y = source \ then 0 else f (y, x), source = source \, sink = sink \, \ = network.more \ \" lemma residual_network_sel [simp]: "edge (residual_network f) x y \ edge \ x y \ edge \ y x \ y \ source \" "capacity (residual_network f) (x, y) = (if edge \ x y then capacity \ (x, y) - f (x, y) else if y = source \ then 0 else f (y, x))" "source (residual_network f) = source \" "sink (residual_network f) = sink \" "network.more (residual_network f) = network.more \" by(simp_all add: residual_network_def) lemma "\<^bold>E_residual_network": "\<^bold>E\<^bsub>residual_network f\<^esub> = \<^bold>E \ {(x, y). (y, x) \ \<^bold>E \ y \ source \}" by auto lemma vertices_residual_network [simp]: "vertex (residual_network f) = vertex \" by(auto simp add: vertex_def fun_eq_iff) inductive wf_residual_network :: "bool" where "\ \x y. (x, y) \ \<^bold>E \ (y, x) \ \<^bold>E; (source \, sink \) \ \<^bold>E \ \ wf_residual_network" lemma wf_residual_networkD: "\ wf_residual_network; edge \ x y \ \ \ edge \ y x" "\ wf_residual_network; e \ \<^bold>E \ \ prod.swap e \ \<^bold>E" "\ wf_residual_network; edge \ (source \) (sink \) \ \ False" by(auto simp add: wf_residual_network.simps) lemma residual_countable_network: assumes wf: "wf_residual_network" and f: "flow \ f" shows "countable_network (residual_network f)" (is "countable_network ?\") proof have "countable (converse \<^bold>E)" by simp then have "countable {(x, y). (y, x) \ \<^bold>E \ y \ source \}" by(rule countable_subset[rotated]) auto then show "countable \<^bold>E\<^bsub>?\\<^esub>" unfolding "\<^bold>E_residual_network" by simp show "source ?\ \ sink ?\" by simp show "capacity ?\ e = 0" if "e \ \<^bold>E\<^bsub>?\\<^esub>" for e using that by(cases e)(auto intro: flowD_outside[OF f]) show "capacity ?\ e \ top" for e using flowD_finite[OF f] by(cases e) auto qed end context antiparallel_edges begin interpretation \'': countable_network \'' by(rule \''_countable_network) lemma \''_flow_attainability: assumes "flow_attainability_axioms \" shows "flow_attainability \''" proof - interpret flow_attainability \ using _ assms by(rule flow_attainability.intro) unfold_locales show ?thesis proof show "d_IN (capacity \'') v \ \ \ d_OUT (capacity \'') v \ \" if "v \ sink \''" for v using that finite_capacity by(cases v)(simp_all add: max_def) show "\ edge \'' v v" for v by(auto elim: edg.cases) show "\ edge \'' v (source \'')" for v by(simp add: source_in) qed qed lemma \''_wf_residual_network: assumes no_loop: "\x. \ edge \ x x" shows "\''.wf_residual_network" by(auto simp add: \''.wf_residual_network.simps assms elim!: edg.cases) end subsection \The attainability theorem\ context flow_attainability begin lemma residual_flow_attainability: assumes wf: "wf_residual_network" and f: "flow \ f" shows "flow_attainability (residual_network f)" (is "flow_attainability ?\") proof - interpret res: countable_network "residual_network f" by(rule residual_countable_network[OF assms]) show ?thesis proof fix x assume sink: "x \ sink ?\" then consider (source) "x = source \" | (IN) "d_IN (capacity \) x \ \" | (OUT) "x \ source \" "d_OUT (capacity \) x \ \" using finite_capacity[of x] by auto then show "d_IN (capacity ?\) x \ \ \ d_OUT (capacity ?\) x \ \" proof(cases) case source hence "d_IN (capacity ?\) x = 0" by(simp add: d_IN_def source_in) thus ?thesis by simp next case IN have "d_IN (capacity ?\) x = (\\<^sup>+ y. (capacity \ (y, x) - f (y, x)) * indicator \<^bold>E (y, x) + (if x = source \ then 0 else f (x, y) * indicator \<^bold>E (x, y)))" using flowD_outside[OF f] unfolding d_IN_def by(auto intro!: nn_integral_cong split: split_indicator dest: wf_residual_networkD[OF wf]) also have "\ = (\\<^sup>+ y. (capacity \ (y, x) - f (y, x)) * indicator \<^bold>E (y, x)) + (\\<^sup>+ y. (if x = source \ then 0 else f (x, y) * indicator \<^bold>E (x, y)))" (is "_ = ?in + ?out") by(subst nn_integral_add)(auto simp add: AE_count_space split: split_indicator intro!: flowD_capacity[OF f]) also have "\ \ d_IN (capacity \) x + (if x = source \ then 0 else d_OUT f x)" (is "_ \ ?in + ?rest") unfolding d_IN_def d_OUT_def by(rule add_mono)(auto intro!: nn_integral_mono split: split_indicator simp add: nn_integral_0_iff_AE AE_count_space intro!: diff_le_self_ennreal) also consider (source) "x = source \" | (inner) "x \ source \" "x \ sink \" using sink by auto then have "?rest < \" proof cases case inner show ?thesis using inner flowD_finite_OUT[OF f inner] by (simp add: less_top) qed simp ultimately show ?thesis using IN sink by (auto simp: less_top[symmetric] top_unique) next case OUT have "d_OUT (capacity ?\) x = (\\<^sup>+ y. (capacity \ (x, y) - f (x, y)) * indicator \<^bold>E (x, y) + (if y = source \ then 0 else f (y, x) * indicator \<^bold>E (y, x)))" using flowD_outside[OF f] unfolding d_OUT_def by(auto intro!: nn_integral_cong split: split_indicator dest: wf_residual_networkD[OF wf] simp add: source_in) also have "\ = (\\<^sup>+ y. (capacity \ (x, y) - f (x, y)) * indicator \<^bold>E (x, y)) + (\\<^sup>+ y. (if y = source \ then 0 else f (y, x) * indicator \<^bold>E (y, x)))" (is "_ = ?in + ?out") by(subst nn_integral_add)(auto simp add: AE_count_space split: split_indicator intro!: flowD_capacity[OF f]) also have "\ \ d_OUT (capacity \) x + d_IN f x" (is "_ \ ?out + ?rest") unfolding d_IN_def d_OUT_def by(rule add_mono)(auto intro!: nn_integral_mono split: split_indicator simp add: nn_integral_0_iff_AE AE_count_space intro!: diff_le_self_ennreal) also have "?rest = d_OUT f x" using flowD_KIR[OF f OUT(1)] sink by simp also have "?out + \ \ ?out + ?out" by(intro add_left_mono d_OUT_mono flowD_capacity[OF f]) finally show ?thesis using OUT by (auto simp: top_unique) qed next show "\ edge ?\ x x" for x by(simp add: no_loop) show "\ edge ?\ x (source ?\)" for x by(simp add: source_in) qed qed end definition plus_flow :: "('v, 'more) graph_scheme \ 'v flow \ 'v flow \ 'v flow" (infixr "\\" 65) where "plus_flow G f g = (\(x, y). if edge G x y then f (x, y) + g (x, y) - g (y, x) else 0)" lemma plus_flow_simps [simp]: fixes G (structure) shows "(f \ g) (x, y) = (if edge G x y then f (x, y) + g (x, y) - g (y, x) else 0)" by(simp add: plus_flow_def) lemma plus_flow_outside: fixes G (structure) shows "e \ \<^bold>E \ (f \ g) e = 0" by(cases e) simp lemma fixes \ (structure) assumes f_outside: "\e. e \ \<^bold>E \ f e = 0" and g_le_f: "\x y. edge \ x y \ g (y, x) \ f (x, y)" shows OUT_plus_flow: "d_IN g x \ top \ d_OUT (f \ g) x = d_OUT f x + (\\<^sup>+ y\UNIV. g (x, y) * indicator \<^bold>E (x, y)) - (\\<^sup>+ y. g (y, x) * indicator \<^bold>E (x, y))" (is "_ \ ?OUT" is "_ \ _ = _ + ?g_out - ?g_out'") and IN_plus_flow: "d_OUT g x \ top \ d_IN (f \ g) x = d_IN f x + (\\<^sup>+ y\UNIV. g (y, x) * indicator \<^bold>E (y, x)) - (\\<^sup>+ y. g (x, y) * indicator \<^bold>E (y, x))" (is "_ \ ?IN" is "_ \ _ = _ + ?g_in - ?g_in'") proof - assume "d_IN g x \ top" then have finite1: "(\\<^sup>+ y. g (y, x) * indicator A (f y)) \ top" for A f by(rule neq_top_trans)(auto split: split_indicator simp add: d_IN_def intro!: nn_integral_mono) have "d_OUT (f \ g) x = (\\<^sup>+ y. (g (x, y) + (f (x, y) - g (y, x))) * indicator \<^bold>E (x, y))" unfolding d_OUT_def by(rule nn_integral_cong)(simp split: split_indicator add: add_diff_eq_ennreal add.commute ennreal_diff_add_assoc g_le_f) also have "\ = ?g_out + (\\<^sup>+ y. (f (x, y) - g (y, x)) * indicator \<^bold>E (x, y))" (is "_ = _ + ?rest") by(subst nn_integral_add[symmetric])(auto simp add: AE_count_space g_le_f split: split_indicator intro!: nn_integral_cong) also have "?rest = (\\<^sup>+ y. f (x, y) * indicator \<^bold>E (x, y)) - ?g_out'" (is "_ = ?f - _") apply(subst nn_integral_diff[symmetric]) apply(auto intro!: nn_integral_cong split: split_indicator simp add: AE_count_space g_le_f finite1) done also have "?f = d_OUT f x" unfolding d_OUT_def using f_outside by(auto intro!: nn_integral_cong split: split_indicator) also have "(\\<^sup>+ y. g (x, y) * indicator \<^bold>E (x, y)) + (d_OUT f x - (\\<^sup>+ y. g (y, x) * indicator \<^bold>E (x, y))) = d_OUT f x + ?g_out - ?g_out'" by (subst ennreal_diff_add_assoc[symmetric]) (auto simp: ac_simps d_OUT_def intro!: nn_integral_mono g_le_f split: split_indicator) finally show ?OUT . next assume "d_OUT g x \ top" then have finite2: "(\\<^sup>+ y. g (x, y) * indicator A (f y)) \ top" for A f by(rule neq_top_trans)(auto split: split_indicator simp add: d_OUT_def intro!: nn_integral_mono) have "d_IN (f \ g) x = (\\<^sup>+ y. (g (y, x) + (f (y, x) - g (x, y))) * indicator \<^bold>E (y, x))" unfolding d_IN_def by(rule nn_integral_cong)(simp split: split_indicator add: add_diff_eq_ennreal add.commute ennreal_diff_add_assoc g_le_f) also have "\ = (\\<^sup>+ y\UNIV. g (y, x) * indicator \<^bold>E (y, x)) + (\\<^sup>+ y. (f (y, x) - g (x, y)) * indicator \<^bold>E (y, x))" (is "_ = _ + ?rest") by(subst nn_integral_add[symmetric])(auto simp add: AE_count_space g_le_f split: split_indicator intro!: nn_integral_cong) also have "?rest = (\\<^sup>+ y. f (y, x) * indicator \<^bold>E (y, x))- ?g_in'" by(subst nn_integral_diff[symmetric])(auto intro!: nn_integral_cong split: split_indicator simp add: add_ac add_diff_eq_ennreal AE_count_space g_le_f finite2) also have "(\\<^sup>+ y. f (y, x) * indicator \<^bold>E (y, x)) = d_IN f x" unfolding d_IN_def using f_outside by(auto intro!: nn_integral_cong split: split_indicator) also have "(\\<^sup>+ y. g (y, x) * indicator \<^bold>E (y, x)) + (d_IN f x - (\\<^sup>+ y. g (x, y) * indicator \<^bold>E (y, x))) = d_IN f x + ?g_in - ?g_in'" by (subst ennreal_diff_add_assoc[symmetric]) (auto simp: ac_simps d_IN_def intro!: nn_integral_mono g_le_f split: split_indicator) finally show ?IN . qed context countable_network begin lemma d_IN_plus_flow: assumes wf: "wf_residual_network" and f: "flow \ f" and g: "flow (residual_network f) g" shows "d_IN (f \ g) x \ d_IN f x + d_IN g x" proof - have "d_IN (f \ g) x \ (\\<^sup>+ y. f (y, x) + g (y, x))" unfolding d_IN_def by(rule nn_integral_mono)(auto intro: diff_le_self_ennreal) also have "\ = d_IN f x + d_IN g x" by(subst nn_integral_add)(simp_all add: d_IN_def) finally show ?thesis . qed lemma scale_flow: assumes f: "flow \ f" and c: "c \ 1" shows "flow \ (\e. c * f e)" proof(intro flow.intros) fix e from c have "c * f e \ 1 * f e" by(rule mult_right_mono) simp also have "\ \ capacity \ e" using flowD_capacity[OF f, of e] by simp finally show "c * f e \ \" . next fix x assume x: "x \ source \" "x \ sink \" have "d_OUT (\e. c * f e) x = c * d_OUT f x" by(simp add: d_OUT_cmult) also have "d_OUT f x = d_IN f x" using f x by(rule flowD_KIR) also have "c * \ = d_IN (\e. c * f e) x" by(simp add: d_IN_cmult) finally show "KIR (\e. c * f e) x" . qed lemma value_scale_flow: "value_flow \ (\e. c * f e) = c * value_flow \ f" by(rule d_OUT_cmult) lemma value_flow: assumes f: "flow \ f" and source_out: "\y. edge \ (source \) y \ y = x" shows "value_flow \ f = f (source \, x)" proof - have "value_flow \ f = (\\<^sup>+ y\\<^bold>O\<^bold>U\<^bold>T (source \). f (source \, y))" by(rule d_OUT_alt_def)(simp add: flowD_outside[OF f]) also have "\ = (\\<^sup>+ y. f (source \, y) * indicator {x} y)" by(subst nn_integral_count_space_indicator)(auto intro!: nn_integral_cong split: split_indicator simp add: outgoing_def source_out) also have "\ = f (source \, x)" by(simp add: one_ennreal_def[symmetric] max_def) finally show ?thesis . qed end context flow_attainability begin lemma value_plus_flow: assumes wf: "wf_residual_network" and f: "flow \ f" and g: "flow (residual_network f) g" shows "value_flow \ (f \ g) = value_flow \ f + value_flow \ g" proof - interpret RES: countable_network "residual_network f" using wf f by(rule residual_countable_network) have "value_flow \ (f \ g) = (\\<^sup>+ y. f (source \, y) + g (source \, y))" unfolding d_OUT_def by(rule nn_integral_cong)(simp add: flowD_outside[OF f] RES.flowD_outside[OF g] source_in) also have "\ = value_flow \ f + value_flow \ g" unfolding d_OUT_def by(rule nn_integral_add) simp_all finally show ?thesis . qed lemma flow_residual_add: \ \Lemma 5.3\ assumes wf: "wf_residual_network" and f: "flow \ f" and g: "flow (residual_network f) g" shows "flow \ (f \ g)" proof fix e { assume e: "e \ \<^bold>E" hence "(f \ g) e = f e + g e - g (prod.swap e)" by(cases e) simp also have "\ \ f e + g e - 0" by(rule ennreal_minus_mono) simp_all also have "\ \ f e + (capacity \ e - f e)" using e flowD_capacity[OF g, of e] by(simp split: prod.split_asm add: add_mono) also have "\ = capacity \ e" using flowD_capacity[OF f, of e] by simp also note calculation } thus "(f \ g) e \ capacity \ e" by(cases e) auto next fix x assume x: "x \ source \" "x \ sink \" have g_le_f: "g (y, x) \ f (x, y)" if "edge \ x y" for x y using that flowD_capacity[OF g, of "(y, x)"] by(auto split: if_split_asm dest: wf_residual_networkD[OF wf] elim: order_trans) interpret RES: flow_attainability "residual_network f" using wf f by(rule residual_flow_attainability) have finite1: "(\\<^sup>+ y. g (y, x) * indicator A (f y)) \ \" for A f using RES.flowD_finite_IN[OF g, of x] by(rule neq_top_trans)(auto simp add: x d_IN_def split: split_indicator intro: nn_integral_mono) have finite2: "(\\<^sup>+ y. g (x, y) * indicator A (f y)) \ \" for A f using RES.flowD_finite_OUT[OF g, of x] by(rule neq_top_trans)(auto simp add: x d_OUT_def split: split_indicator intro: nn_integral_mono) have "d_OUT (f \ g) x = d_OUT f x + (\\<^sup>+ y. g (x, y) * indicator \<^bold>E (x, y)) - (\\<^sup>+ y. g (y, x) * indicator \<^bold>E (x, y))" (is "_ = ?f + ?g_out - ?g_in") using flowD_outside[OF f] g_le_f RES.flowD_finite_IN[OF g, of x] by(rule OUT_plus_flow)(simp_all add: x) also have "?f = d_IN f x" using f x by(auto dest: flowD_KIR) also have "?g_out = (\\<^sup>+ y. g (x, y) * indicator (- \<^bold>E) (y, x))" proof - have "(\\<^sup>+ y. g (x, y) * indicator (- \<^bold>E) (y, x)) = (\\<^sup>+ y. g (x, y) * indicator \<^bold>E (x, y)) + (\\<^sup>+ y. g (x, y) * indicator (- \<^bold>E) (x, y) * indicator (- \<^bold>E) (y, x))" by(subst nn_integral_add[symmetric])(auto simp add: AE_count_space dest: wf_residual_networkD[OF wf] split: split_indicator intro!: nn_integral_cong) also have "(\\<^sup>+ y. g (x, y) * indicator (- \<^bold>E) (x, y) * indicator (- \<^bold>E) (y, x)) = 0" using RES.flowD_outside[OF g] by(auto simp add: nn_integral_0_iff_AE AE_count_space split: split_indicator) finally show ?thesis by simp qed also have "\ = (\\<^sup>+ y. g (x, y) - g (x, y) * indicator \<^bold>E (y, x))" by(rule nn_integral_cong)(simp split: split_indicator add: RES.flowD_finite[OF g]) also have "\ = d_OUT g x - (\\<^sup>+ y. g (x, y) * indicator \<^bold>E (y, x))" (is "_ = _ - ?g_in_E") unfolding d_OUT_def by(subst nn_integral_diff)(simp_all add: AE_count_space finite2 split: split_indicator) also have "d_IN f x + (d_OUT g x - (\\<^sup>+ y. g (x, y) * indicator \<^bold>E (y, x))) - ?g_in = ((d_IN f x + d_OUT g x) - (\\<^sup>+ y. g (x, y) * indicator \<^bold>E (y, x))) - ?g_in" by (subst add_diff_eq_ennreal) (auto simp: d_OUT_def intro!: nn_integral_mono split: split_indicator) also have "d_OUT g x = d_IN g x" using x g by(auto dest: flowD_KIR) also have "\ = (\\<^sup>+ y\UNIV. g (y, x) * indicator (- \<^bold>E) (y, x)) + (\\<^sup>+ y. g (y, x) * indicator \<^bold>E (y, x))" (is "_ = ?x + ?g_in_E'") by(subst nn_integral_add[symmetric])(auto intro!: nn_integral_cong simp add: d_IN_def AE_count_space split: split_indicator) also have "?x = ?g_in" proof - have "?x = (\\<^sup>+ y. g (y, x) * indicator (- \<^bold>E) (x, y) * indicator (- \<^bold>E) (y, x)) + ?g_in" by(subst nn_integral_add[symmetric])(auto simp add: AE_count_space dest: wf_residual_networkD[OF wf] split: split_indicator intro!: nn_integral_cong) also have "(\\<^sup>+ y. g (y, x) * indicator (- \<^bold>E) (x, y) * indicator (- \<^bold>E) (y, x)) = 0" using RES.flowD_outside[OF g] by(auto simp add: nn_integral_0_iff_AE AE_count_space split: split_indicator) finally show ?thesis by simp qed also have "(d_IN f x + (?g_in + ?g_in_E') - ?g_in_E) - ?g_in = d_IN f x + ?g_in_E' + ?g_in - ?g_in - ?g_in_E" by (subst diff_diff_commute_ennreal) (simp add: ac_simps) also have "\ = d_IN f x + ?g_in_E' - ?g_in_E" by (subst ennreal_add_diff_cancel_right) (simp_all add: finite1) also have "\ = d_IN (f \ g) x" using flowD_outside[OF f] g_le_f RES.flowD_finite_OUT[OF g, of x] by(rule IN_plus_flow[symmetric])(simp_all add: x) finally show "KIR (f \ g) x" by simp qed definition minus_flow :: "'v flow \ 'v flow \ 'v flow" (infixl "\" 65) where "f \ g = (\(x, y). if edge \ x y then f (x, y) - g (x, y) else if edge \ y x then g (y, x) - f (y, x) else 0)" lemma minus_flow_simps [simp]: "(f \ g) (x, y) = (if edge \ x y then f (x, y) - g (x, y) else if edge \ y x then g (y, x) - f (y, x) else 0)" by(simp add: minus_flow_def) lemma minus_flow: assumes wf: "wf_residual_network" and f: "flow \ f" and g: "flow \ g" and value_le: "value_flow \ g \ value_flow \ f" and f_finite: "f (source \, x) \ \" and source_out: "\y. edge \ (source \) y \ y = x" shows "flow (residual_network g) (f \ g)" (is "flow ?\ ?f") proof show "?f e \ capacity ?\ e" for e using value_le f_finite flowD_capacity[OF g] flowD_capacity[OF f] by(cases e)(auto simp add: source_in source_out value_flow[OF f source_out] value_flow[OF g source_out] less_top intro!: diff_le_self_ennreal diff_eq_0_ennreal ennreal_minus_mono) fix x assume "x \ source ?\" "x \ sink ?\" hence x: "x \ source \" "x \ sink \" by simp_all have finite_f_in: "(\\<^sup>+ y. f (y, x) * indicator A y) \ top" for A using flowD_finite_IN[OF f, of x] by(rule neq_top_trans)(auto simp add: x d_IN_def split: split_indicator intro!: nn_integral_mono) have finite_f_out: "(\\<^sup>+ y. f (x, y) * indicator A y) \ top" for A using flowD_finite_OUT[OF f, of x] by(rule neq_top_trans)(auto simp add: x d_OUT_def split: split_indicator intro!: nn_integral_mono) have finite_f[simp]: "f (x, y) \ top" "f (y, x) \ top" for y using finite_f_in[of "{y}"] finite_f_out[of "{y}"] by auto have finite_g_in: "(\\<^sup>+ y. g (y, x) * indicator A y) \ top" for A using flowD_finite_IN[OF g, of x] by(rule neq_top_trans)(auto simp add: x d_IN_def split: split_indicator intro!: nn_integral_mono) have finite_g_out: "(\\<^sup>+ y. g (x, y) * indicator A y) \ top" for A using flowD_finite_OUT[OF g x] by(rule neq_top_trans)(auto simp add: x d_OUT_def split: split_indicator intro!: nn_integral_mono) have finite_g[simp]: "g (x, y) \ top" "g (y, x) \ top" for y using finite_g_in[of "{y}"] finite_g_out[of "{y}"] by auto have "d_OUT (f \ g) x = (\\<^sup>+ y. (f (x, y) - g (x, y)) * indicator \<^bold>E (x, y) * indicator {y. g (x, y) \ f (x, y)} y) + (\\<^sup>+ y. (g (y, x) - f (y, x)) * indicator \<^bold>E (y, x) * indicator {y. f (y, x) < g (y, x)} y)" (is "_ = ?out + ?in" is "_ = (\\<^sup>+ y\_. _ * ?f_ge_g y) + (\\<^sup>+ y\_. _ * ?g_gt_f y)") using flowD_finite[OF g] apply(subst nn_integral_add[symmetric]) apply(auto 4 4 simp add: d_OUT_def not_le less_top[symmetric] intro!: nn_integral_cong dest!: wf_residual_networkD[OF wf] split: split_indicator intro!: diff_eq_0_ennreal) done also have "?in = (\\<^sup>+ y. (g (y, x) - f (y, x)) * ?g_gt_f y)" using flowD_outside[OF f] flowD_outside[OF g] by(auto intro!: nn_integral_cong split: split_indicator) also have "\ = (\\<^sup>+ y\UNIV. g (y, x) * ?g_gt_f y) - (\\<^sup>+ y. f (y, x) * ?g_gt_f y)" (is "_ = ?g_in - ?f_in") using finite_f_in by(subst nn_integral_diff[symmetric])(auto simp add: AE_count_space split: split_indicator intro!: nn_integral_cong) also have "?out = (\\<^sup>+ y. (f (x, y) - g (x, y)) * ?f_ge_g y)" using flowD_outside[OF f] flowD_outside[OF g] by(auto intro!: nn_integral_cong split: split_indicator) also have "\ = (\\<^sup>+ y. f (x, y) * ?f_ge_g y) - (\\<^sup>+ y. g (x, y) * ?f_ge_g y)" (is "_ = ?f_out - ?g_out") using finite_g_out by(subst nn_integral_diff[symmetric])(auto simp add: AE_count_space split: split_indicator intro!: nn_integral_cong) also have "?f_out = d_OUT f x - (\\<^sup>+ y. f (x, y) * indicator {y. f (x, y) < g (x, y)} y)" (is "_ = _ - ?f_out_less") unfolding d_OUT_def using flowD_finite[OF f] using finite_f_out by(subst nn_integral_diff[symmetric])(auto split: split_indicator intro!: nn_integral_cong) also have "?g_out = d_OUT g x - (\\<^sup>+ y. g (x, y) * indicator {y. f (x, y) < g (x, y)} y)" (is "_ = _ - ?g_less_f") unfolding d_OUT_def using flowD_finite[OF g] finite_g_out by(subst nn_integral_diff[symmetric])(auto split: split_indicator intro!: nn_integral_cong) also have "d_OUT f x - ?f_out_less - (d_OUT g x - ?g_less_f) + (?g_in - ?f_in) = (?g_less_f + (d_OUT f x - ?f_out_less)) - d_OUT g x + (?g_in - ?f_in)" by (subst diff_diff_ennreal') (auto simp: ac_simps d_OUT_def nn_integral_diff[symmetric] finite_g_out finite_f_out intro!: nn_integral_mono split: split_indicator ) also have "\ = ?g_less_f + d_OUT f x - ?f_out_less - d_OUT g x + (?g_in - ?f_in)" by (subst add_diff_eq_ennreal) (auto simp: d_OUT_def intro!: nn_integral_mono split: split_indicator) also have "\ = d_OUT f x + ?g_less_f - ?f_out_less - d_OUT g x + (?g_in - ?f_in)" by (simp add: ac_simps) also have "\ = d_OUT f x + (?g_less_f - ?f_out_less) - d_OUT g x + (?g_in - ?f_in)" by (subst add_diff_eq_ennreal[symmetric]) (auto intro!: nn_integral_mono split: split_indicator) also have "\ = (?g_in - ?f_in) + ((?g_less_f - ?f_out_less) + d_OUT f x - d_OUT g x)" by (simp add: ac_simps) also have "\ = ((?g_in - ?f_in) + ((?g_less_f - ?f_out_less) + d_OUT f x)) - d_OUT g x" apply (subst add_diff_eq_ennreal) apply (simp_all add: d_OUT_def) apply (subst nn_integral_diff[symmetric]) apply (auto simp: AE_count_space finite_f_out nn_integral_add[symmetric] not_less diff_add_cancel_ennreal intro!: nn_integral_mono split: split_indicator) done also have "\ = ((?g_less_f - ?f_out_less) + (d_OUT f x + (?g_in - ?f_in))) - d_OUT g x" by (simp add: ac_simps) also have "\ = ((?g_less_f - ?f_out_less) + (d_IN f x + (?g_in - ?f_in))) - d_IN g x" unfolding flowD_KIR[OF f x] flowD_KIR[OF g x] .. also have "\ = (?g_less_f - ?f_out_less) + ((d_IN f x + (?g_in - ?f_in)) - d_IN g x)" apply (subst (2) add_diff_eq_ennreal) apply (simp_all add: d_IN_def) apply (subst nn_integral_diff[symmetric]) apply (auto simp: AE_count_space finite_f_in finite_f_out nn_integral_add[symmetric] not_less ennreal_ineq_diff_add[symmetric] intro!: nn_integral_mono split: split_indicator) done also have "\ = (?g_less_f - ?f_out_less) + (d_IN f x + ?g_in - d_IN g x - ?f_in)" by (subst (2) add_diff_eq_ennreal) (auto intro!: nn_integral_mono split: split_indicator simp: diff_diff_commute_ennreal) also have "\ = (?g_less_f - ?f_out_less) + (d_IN f x - (d_IN g x - ?g_in) - ?f_in)" apply (subst diff_diff_ennreal') apply (auto simp: d_IN_def intro!: nn_integral_mono split: split_indicator) apply (subst nn_integral_diff[symmetric]) apply (auto simp: AE_count_space finite_g_in intro!: nn_integral_mono split: split_indicator) done also have "\ =(d_IN f x - ?f_in) - (d_IN g x - ?g_in) + (?g_less_f - ?f_out_less)" by (simp add: ac_simps diff_diff_commute_ennreal) also have "?g_less_f - ?f_out_less = (\\<^sup>+ y. (g (x, y) - f (x, y)) * indicator {y. f (x, y) < g (x, y)} y)" using finite_f_out by(subst nn_integral_diff[symmetric])(auto simp add: AE_count_space split: split_indicator intro!: nn_integral_cong) also have "\ = (\\<^sup>+ y. (g (x, y) - f (x, y)) * indicator \<^bold>E (x, y) * indicator {y. f (x, y) < g (x, y)} y)" (is "_ = ?diff_out") using flowD_outside[OF f] flowD_outside[OF g] by(auto intro!: nn_integral_cong split: split_indicator) also have "d_IN f x - ?f_in = (\\<^sup>+ y. f (y, x) * indicator {y. g (y, x) \ f (y, x)} y)" unfolding d_IN_def using finite_f_in apply(subst nn_integral_diff[symmetric]) apply(auto simp add: AE_count_space split: split_indicator intro!: nn_integral_cong) done also have "d_IN g x - ?g_in = (\\<^sup>+ y. g (y, x) * indicator {y. g (y, x) \ f (y, x)} y)" unfolding d_IN_def using finite_g_in by(subst nn_integral_diff[symmetric])(auto simp add: flowD_finite[OF g] AE_count_space split: split_indicator intro!: nn_integral_cong) also have "(\\<^sup>+ y\UNIV. f (y, x) * indicator {y. g (y, x) \ f (y, x)} y) - \ = (\\<^sup>+ y. (f (y, x) - g (y, x)) * indicator {y. g (y, x) \ f (y, x)} y)" using finite_g_in by(subst nn_integral_diff[symmetric])(auto simp add: flowD_finite[OF g] AE_count_space split: split_indicator intro!: nn_integral_cong) also have "\ = (\\<^sup>+ y. (f (y, x) - g (y, x)) * indicator \<^bold>E (y, x) * indicator {y. g (y, x) \ f (y, x)} y)" using flowD_outside[OF f] flowD_outside[OF g] by(auto intro!: nn_integral_cong split: split_indicator) also have "\ + ?diff_out = d_IN ?f x" using flowD_finite[OF g] apply(subst nn_integral_add[symmetric]) apply(auto 4 4 simp add: d_IN_def not_le less_top[symmetric] intro!: nn_integral_cong dest!: wf_residual_networkD[OF wf] split: split_indicator intro: diff_eq_0_ennreal) done finally show "KIR ?f x" . qed lemma value_minus_flow: assumes f: "flow \ f" and g: "flow \ g" and value_le: "value_flow \ g \ value_flow \ f" and source_out: "\y. edge \ (source \) y \ y = x" shows "value_flow \ (f \ g) = value_flow \ f - value_flow \ g" (is "?value") proof - have "value_flow \ (f \ g) = (\\<^sup>+ y\\<^bold>O\<^bold>U\<^bold>T (source \). (f \ g) (source \, y))" by(subst d_OUT_alt_def)(auto simp add: flowD_outside[OF f] flowD_outside[OF g] source_in) also have "\ = (\\<^sup>+ y. (f (source \, y) - g (source \, y)) * indicator {x} y)" by(subst nn_integral_count_space_indicator)(auto intro!: nn_integral_cong split: split_indicator simp add: outgoing_def source_out) also have "\ = f (source \, x) - g (source \, x)" using value_le value_flow[OF f source_out] value_flow[OF g source_out] by(auto simp add: one_ennreal_def[symmetric] max_def not_le intro: antisym) also have "\ = value_flow \ f - value_flow \ g" using f g source_out by(simp add: value_flow) finally show ?value . qed context fixes \ defines "\ \ (SUP g\{g. flow \ g}. value_flow \ g)" begin lemma flow_by_value: assumes "v < \" and real[rule_format]: "\f. \ = \ \ flow \ f \ value_flow \ f < \" obtains f where "flow \ f" "value_flow \ f = v" proof - have \_pos: "\ > 0" using assms by (auto simp add: zero_less_iff_neq_zero) from \v < \\ obtain f where f: "flow \ f" and v: "v < value_flow \ f" unfolding \_def less_SUP_iff by blast have [simp]: "value_flow \ f \ \" proof assume val: "value_flow \ f = \" from f have "value_flow \ f \ \" unfolding \_def by(blast intro: SUP_upper2) with val have "\ = \" by (simp add: top_unique) from real[OF this f] val show False by simp qed let ?f = "\e. (v / value_flow \ f) * f e" note f moreover have *: "0 < value_flow \ f" using \v < value_flow \ f\ by (auto simp add: zero_less_iff_neq_zero) then have "v / value_flow \ f \ 1" using v by (auto intro!: divide_le_posI_ennreal) ultimately have "flow \ ?f" by (rule scale_flow) moreover { have "value_flow \ ?f = v * (value_flow \ f / value_flow \ f)" by(subst value_scale_flow)(simp add: divide_ennreal_def ac_simps) also have "\ = v" using * by(subst ennreal_divide_self) (auto simp: less_top[symmetric]) also note calculation } ultimately show ?thesis by(rule that) qed theorem ex_max_flow': assumes wf: "wf_residual_network" assumes source_out: "\y. edge \ (source \) y \ y = x" and nontrivial: "\<^bold>V - {source \, sink \} \ {}" and real: "\ = ennreal \'" and \'_nonneg[simp]: "0 \ \'" shows "\f. flow \ f \ value_flow \ f = \ \ (\x. d_IN f x \ value_flow \ f)" proof - have \'_not_neg[simp]: "\ \' < 0" using \'_nonneg by linarith let ?v = "\i. (1 - (1 / 2) ^ i) * \" let ?v_r = "\i. ennreal ((1 - (1 / 2) ^ i) * \')" have v_eq: "?v i = ?v_r i" for i by (auto simp: real ennreal_mult power_le_one ennreal_lessI ennreal_minus[symmetric] ennreal_power[symmetric] divide_ennreal_def) have "\f. flow \ f \ value_flow \ f = ?v i" for i :: nat proof(cases "\ = 0") case True thus ?thesis by(auto intro!: exI[where x="\_. 0"]) next case False then have "?v i < \" unfolding v_eq by (auto simp: real field_simps intro!: ennreal_lessI) (simp_all add: less_le) then obtain f where "flow \ f" and "value_flow \ f = ?v i" by(rule flow_by_value)(simp add: real) thus ?thesis by blast qed then obtain f_aux where f_aux: "\i. flow \ (f_aux i)" and value_aux: "\i. value_flow \ (f_aux i) = ?v_r i" unfolding v_eq by moura define f_i where "f_i = rec_nat (\_. 0) (\i f_i. let g = f_aux (Suc i) \ f_i; k_i = SOME k. k \ g \ flow (residual_network f_i) k \ value_flow (residual_network f_i) k = value_flow (residual_network f_i) g \ (\x. d_IN k x \ value_flow (residual_network f_i) k) in f_i \ k_i)" let ?P = "\i k. k \ f_aux (Suc i) \ f_i i \ flow (residual_network (f_i i)) k \ value_flow (residual_network (f_i i)) k = value_flow (residual_network (f_i i)) (f_aux (Suc i) \ f_i i) \ (\x. d_IN k x \ value_flow (residual_network (f_i i)) k)" define k_i where "k_i i = Eps (?P i)" for i have f_i_simps [simp]: "f_i 0 = (\_. 0)" "f_i (Suc i) = f_i i \ k_i i" for i by(simp_all add: f_i_def Let_def k_i_def) have k_i: "flow (residual_network (f_i i)) (k_i i)" (is ?k_i) and value_k_i: "value_flow (residual_network (f_i i)) (k_i i) = value_flow (residual_network (f_i i)) (f_aux (Suc i) \ f_i i)" (is "?value_k_i") and IN_k_i: "d_IN (k_i i) x \ value_flow (residual_network (f_i i)) (k_i i)" (is "?IN_k_i") and value_diff: "value_flow (residual_network (f_i i)) (f_aux (Suc i) \ f_i i) = value_flow \ (f_aux (Suc i)) - value_flow \ (f_i i)" (is "?value_diff") if "flow_network \ (f_i i)" and value_f_i: "value_flow \ (f_i i) = value_flow \ (f_aux i)" for i x proof - let ?RES = "residual_network (f_i i)" interpret fn: flow_network \ "f_i i" by(rule that) interpret RES: flow_attainability "?RES" using wf fn.g by(rule residual_flow_attainability) have le: "value_flow \ (f_i i) \ value_flow \ (f_aux (Suc i))" unfolding value_aux value_f_i unfolding v_eq by (rule ennreal_leI) (auto simp: field_simps) with wf f_aux fn.g have res_flow: "flow ?RES (f_aux (Suc i) \ f_i i)" using flowD_finite[OF f_aux] source_out by(rule minus_flow) show value': ?value_diff by(simp add: value_minus_flow[OF f_aux fn.g le source_out]) also have "\ < \" unfolding value_aux v_eq by (auto simp: less_top[symmetric]) finally have "value_flow ?RES (f_aux (Suc i) \ f_i i) \ \" by simp then have fn': "flow_network ?RES (f_aux (Suc i) \ f_i i)" using nontrivial res_flow by(unfold_locales) simp_all then interpret fn': flow_network "?RES" "f_aux (Suc i) \ f_i i" . from fn'.flow_cleanup show ?k_i ?value_k_i ?IN_k_i unfolding k_i_def by(rule someI2_ex; blast)+ qed have fn_i: "flow_network \ (f_i i)" and value_f_i: "value_flow \ (f_i i) = value_flow \ (f_aux i)" and d_IN_i: "d_IN (f_i i) x \ value_flow \ (f_i i)" for i x proof(induction i) case 0 { case 1 show ?case using nontrivial by(unfold_locales)(simp_all add: f_aux value_aux) } { case 2 show ?case by(simp add: value_aux) } { case 3 show ?case by(simp) } next case (Suc i) interpret fn: flow_network \ "f_i i" using Suc.IH(1) . let ?RES = "residual_network (f_i i)" have k_i: "flow ?RES (k_i i)" and value_k_i: "value_flow ?RES (k_i i) = value_flow ?RES (f_aux (Suc i) \ f_i i)" and d_IN_k_i: "d_IN (k_i i) x \ value_flow ?RES (k_i i)" for x using Suc.IH(1-2) by(rule k_i value_k_i IN_k_i)+ interpret RES: flow_attainability "?RES" using wf fn.g by(rule residual_flow_attainability) have le: "value_flow \ (f_i i) \ value_flow \ (f_aux (Suc i))" unfolding value_aux Suc.IH(2) v_eq using \'_nonneg by(intro ennreal_leI)(simp add: real field_simps) { case 1 show ?case unfolding f_i_simps proof show "flow \ (f_i i \ k_i i)" using wf fn.g k_i by(rule flow_residual_add) with RES.flowD_finite[OF k_i] show "value_flow \ (f_i i \ k_i i) \ \" by(simp add: value_flow[OF _ source_out]) qed(rule nontrivial) } from value_k_i have value_k: "value_flow ?RES (k_i i) = value_flow \ (f_aux (Suc i)) - value_flow \ (f_aux i)" by(simp add: value_minus_flow[OF f_aux fn.g le source_out] Suc.IH) { case 2 show ?case using value_k by(auto simp add: source_out value_plus_flow[OF wf fn.g k_i] Suc.IH value_aux field_simps intro!: ennreal_leI) } note value_f = this { case 3 have "d_IN (f_i i \ k_i i) x \ d_IN (f_i i) x + d_IN (k_i i) x" using fn.g k_i by(rule d_IN_plus_flow[OF wf]) also have "\ \ value_flow \ (f_i i) + d_IN (k_i i) x" using Suc.IH(3) by(rule add_right_mono) also have "\ \ value_flow \ (f_i i) + value_flow ?RES (k_i i)" using d_IN_k_i by(rule add_left_mono) also have "\ = value_flow \ (f_i (Suc i))" unfolding value_f Suc.IH(2) value_k by(auto simp add: value_aux field_simps intro!: ennreal_leI) finally show ?case by simp } qed interpret fn: flow_network \ "f_i i" for i by(rule fn_i) note k_i = k_i[OF fn_i value_f_i] and value_k_i = value_k_i[OF fn_i value_f_i] and IN_k_i = IN_k_i[OF fn_i value_f_i] and value_diff = value_diff[OF fn_i value_f_i] have "\x\0. f_i i e = ennreal x" for i e using fn.finite_g[of i e] by (cases "f_i i e") auto then obtain f_i' where f_i': "\i e. f_i i e = ennreal (f_i' i e)" and [simp]: "\i e. 0 \ f_i' i e" by metis { fix i e obtain x y :: 'v where e: "e = (x, y)" by(cases e) have "k_i i (x, y) \ d_IN (k_i i) y" by (rule d_IN_ge_point) also have "\ \ value_flow (residual_network (f_i i)) (k_i i)" by(rule IN_k_i) also have "\ < \" using value_k_i[of i] value_diff[of i] by(simp add: value_k_i value_f_i value_aux real less_top[symmetric]) finally have "\x\0. k_i i e = ennreal x" by(cases "k_i i e")(auto simp add: e) } then obtain k_i' where k_i': "\i e. k_i i e = ennreal (k_i' i e)" and k_i'_nonneg[simp]: "\i e. 0 \ k_i' i e" by metis have wf_k: "(x, y) \ \<^bold>E \ k_i i (y, x) \ f_i i (x, y) + k_i i (x, y)" for i x y using flowD_capacity[OF k_i, of i "(y, x)"] by (auto split: if_split_asm dest: wf_residual_networkD[OF wf] elim: order_trans) have f_i'_0[simp]: "f_i' 0 = (\_. 0)" using f_i_simps(1) by (simp del: f_i_simps add: fun_eq_iff f_i') have f_i'_Suc[simp]: "f_i' (Suc i) e = (if e \ \<^bold>E then f_i' i e + (k_i' i e - k_i' i (prod.swap e)) else 0)" for i e using f_i_simps(2)[of i, unfolded fun_eq_iff, THEN spec, of e] wf_k[of "fst e" "snd e" i] by (auto simp del: f_i_simps ennreal_plus simp add: fun_eq_iff f_i' k_i' ennreal_plus[symmetric] ennreal_minus split: if_split_asm) have k_i'_le: "k_i' i e \ \' / 2 ^ (Suc i)" for i e proof - obtain x y where e: "e = (x, y)" by(cases e) have "k_i' i (x, y) \ d_IN (k_i' i) y" by (rule d_IN_ge_point) also have "\ \ value_flow (residual_network (f_i i)) (k_i' i)" using IN_k_i[of i y] by(simp add: k_i'[abs_def]) also have "\ = \' / 2 ^ (Suc i)" using value_k_i[of i] value_diff[of i] by(simp add: value_f_i value_aux real k_i'[abs_def] field_simps ennreal_minus mult_le_cancel_left1) finally show ?thesis using e by simp qed have convergent: "convergent (\i. f_i' i e)" for e proof(cases "\' > 0") case False obtain x y where [simp]: "e = (x, y)" by(cases e) { fix i from False \'_nonneg have "\' = 0" by simp moreover have "f_i i (x, y) \ d_IN (f_i i) y" by (rule d_IN_ge_point) ultimately have "f_i i (x, y) = 0" using d_IN_i[of i y] by(simp add: value_f_i value_aux real) } thus ?thesis by(simp add: f_i' convergent_const) next case \'_pos: True show ?thesis proof(rule real_Cauchy_convergent Cauchy_real_Suc_diff)+ fix n have "\k_i' n e - k_i' n (prod.swap e)\ \ \k_i' n e\ + \k_i' n (prod.swap e)\" by (rule abs_triangle_ineq4) then have "\k_i' n e - k_i' n (prod.swap e)\ \ \' / 2 ^ n" using k_i'_le[of n e] k_i'_le[of n "prod.swap e"] by simp then have "\f_i' (Suc n) e - f_i' n e\ \ \' / 2 ^ n" using flowD_outside[OF fn.g] by (cases e) (auto simp: f_i') thus "\f_i' (Suc n) e - f_i' n e\ \ \' / 2 ^ n" by simp qed simp qed then obtain f' where f': "\e. (\i. f_i' i e) \ f' e" unfolding convergent_def by metis hence f: "\e. (\i. f_i i e) \ ennreal (f' e)" unfolding f_i' by simp have f'_nonneg: "0 \ f' e" for e by (rule LIMSEQ_le_const[OF f']) auto let ?k = "\i x y. (k_i' i (x, y) - k_i' i (y, x)) * indicator \<^bold>E (x, y)" have sum_i': "f_i' i (x, y) = (\ji. \?k i x y\)" for x y proof(rule summable_rabs_comparison_test) show "\N. \i\N. \?k i x y\ \ \' * (1 / 2) ^ i" proof (intro exI allI impI) fix i have "\?k i x y\ \ k_i' i (x, y) + k_i' i (y, x)" by (auto intro!: abs_triangle_ineq4[THEN order_trans] split: split_indicator) also have "\ \ \' * (1 / 2) ^ i" using k_i'_le[of i "(x, y)"] k_i'_le[of i "(y, x)"] \'_nonneg k_i'_nonneg[of i "(x, y)"] k_i'_nonneg[of i "(y, x)"] by(auto simp add: abs_real_def power_divide split: split_indicator) finally show "\?k i x y\ \ \' * (1 / 2) ^ i" by simp qed show "summable (\i. \' * (1 / 2) ^ i)" by(rule summable_mult complete_algebra_summable_geometric)+ simp qed hence summable_k: "summable (\i. ?k i x y)" for x y by(auto intro: summable_norm_cancel) have suminf: "(\i. (k_i' i (x, y) - k_i' i (y, x)) * indicator \<^bold>E (x, y)) = f' (x, y)" for x y by(rule LIMSEQ_unique[OF summable_LIMSEQ])(simp_all add: sum_i'[symmetric] f' summable_k) have flow: "flow \ f'" proof fix e have "f' e \ Sup {..capacity \ e}" using _ f by(rule Sup_lim)(simp add: flowD_capacity[OF fn.g]) then show "f' e \ capacity \ e" by simp next fix x assume x: "x \ source \" "x \ sink \" have integrable_f_i: "integrable (count_space UNIV) (\y. f_i' i (x, y))" for i using flowD_finite_OUT[OF fn.g x, of i] by(auto intro!: integrableI_bounded simp add: f_i' d_OUT_def less_top) have integrable_f_i': "integrable (count_space UNIV) (\y. f_i' i (y, x))" for i using flowD_finite_IN[OF fn.g, of x i] x by(auto intro!: integrableI_bounded simp add: f_i' d_IN_def less_top) have integral_k_bounded: "(\\<^sup>+ y. norm (?k i x y)) \ \' / 2 ^ i" (is ?thesis1) and integral_k'_bounded: "(\\<^sup>+ y. norm (?k i y x)) \ \' / 2 ^ i" (is ?thesis2) for i proof - define b where "b = (\\<^sup>+ y. k_i i (x, y) + k_i i (y, x))" have "b = d_OUT (k_i i) x + d_IN (k_i i) x" unfolding b_def by(subst nn_integral_add)(simp_all add: d_OUT_def d_IN_def) also have "d_OUT (k_i i) x = d_IN (k_i i) x" using k_i by(rule flowD_KIR)(simp_all add: x) also have "\ + \ \ value_flow \ (k_i i) + value_flow \ (k_i i)" using IN_k_i[of i x, simplified] by-(rule add_mono) also have "\ \ \' / 2 ^ i" using value_k_i[of i] value_diff[of i] by(simp add: value_aux value_f_i field_simps ennreal_minus_if ennreal_plus_if mult_le_cancel_left1 del: ennreal_plus) also have "(\\<^sup>+ y\UNIV. norm (?k i x y)) \ b" and "(\\<^sup>+ y. norm (?k i y x)) \ b" unfolding b_def by(rule nn_integral_mono; simp add: abs_real_def k_i' ennreal_plus_if del: ennreal_plus split: split_indicator)+ ultimately show ?thesis1 ?thesis2 by(auto) qed have integrable_k: "integrable (count_space UNIV) (\y. ?k i x y)" and integrable_k': "integrable (count_space UNIV) (\y. ?k i y x)" for i using integral_k_bounded[of i] integral_k'_bounded[of i] real by(auto intro!: integrableI_bounded simp: less_top[symmetric] top_unique ennreal_divide_eq_top_iff) have summable'_k: "summable (\i. \ y. \?k i x y\ \count_space UNIV)" proof(rule summable_comparison_test) have "\\ y. \?k i x y\ \count_space UNIV\ \ \' * (1 / 2) ^ i" for i using integral_norm_bound_ennreal[OF integrable_norm, OF integrable_k, of i] integral_k_bounded[of i] by(bestsimp simp add: real power_divide dest: order_trans) thus "\N. \i\N. norm (\ y. \?k i x y\ \count_space UNIV) \ \' * (1 / 2) ^ i" by auto show "summable (\i. \' * (1 / 2) ^ i)" by(rule summable_mult complete_algebra_summable_geometric)+ simp qed have summable'_k': "summable (\i. \ y. \?k i y x\ \count_space UNIV)" proof(rule summable_comparison_test) have "\\ y. \?k i y x\ \count_space UNIV\ \ \' * (1 / 2) ^ i" for i using integral_norm_bound_ennreal[OF integrable_norm, OF integrable_k', of i] integral_k'_bounded[of i] by(bestsimp simp add: real power_divide dest: order_trans) thus "\N. \i\N. norm (\ y. \?k i y x\ \count_space UNIV) \ \' * (1 / 2) ^ i" by auto show "summable (\i. \' * (1 / 2) ^ i)" by(rule summable_mult complete_algebra_summable_geometric)+ simp qed have "(\i. \ y. ?k i x y \count_space UNIV) sums \ y. (\i. ?k i x y) \count_space UNIV" using integrable_k by(rule sums_integral)(simp_all add: summable_nk summable'_k) also have "\ = \ y. f' (x, y) \count_space UNIV" by(rule Bochner_Integration.integral_cong[OF refl])(rule suminf) finally have "(\i. \j y. ?k j x y \count_space UNIV) \ \" unfolding sums_def . also have "(\i. \j y. ?k j x y \count_space UNIV) = (\i. \ y. f_i' i (x, y) \count_space UNIV)" unfolding sum_i' by(rule ext Bochner_Integration.integral_sum[symmetric] integrable_k)+ finally have "(\i. ennreal (\ y. f_i' i (x, y) \count_space UNIV)) \ ennreal (\ y. f' (x, y) \count_space UNIV)" by simp also have "(\i. ennreal (\ y. f_i' i (x, y) \count_space UNIV)) = (\i. d_OUT (f_i i) x)" unfolding d_OUT_def f_i' by(rule ext nn_integral_eq_integral[symmetric] integrable_f_i)+ simp also have "ennreal (\ y. f' (x, y) \count_space UNIV) = d_OUT f' x" unfolding d_OUT_def by(rule nn_integral_eq_integral[symmetric])(simp_all add: f'_nonneg, simp add: suminf[symmetric] integrable_suminf integrable_k summable_nk summable'_k) also have "(\i. d_OUT (f_i i) x) = (\i. d_IN (f_i i) x)" using flowD_KIR[OF fn.g x] by(simp) finally have *: "(\i. d_IN (f_i i) x) \ d_OUT (\x. ennreal (f' x)) x" . have "(\i. \ y. ?k i y x \count_space UNIV) sums \ y. (\i. ?k i y x) \count_space UNIV" using integrable_k' by(rule sums_integral)(simp_all add: summable_nk summable'_k') also have "\ = \ y. f' (y, x) \count_space UNIV" by(rule Bochner_Integration.integral_cong[OF refl])(rule suminf) finally have "(\i. \j y. ?k j y x \count_space UNIV) \ \" unfolding sums_def . also have "(\i. \j y. ?k j y x \count_space UNIV) = (\i. \ y. f_i' i (y, x) \count_space UNIV)" unfolding sum_i' by(rule ext Bochner_Integration.integral_sum[symmetric] integrable_k')+ finally have "(\i. ennreal (\ y. f_i' i (y, x) \count_space UNIV)) \ ennreal (\ y. f' (y, x) \count_space UNIV)" by simp also have "(\i. ennreal (\ y. f_i' i (y, x) \count_space UNIV)) = (\i. d_IN (f_i i) x)" unfolding d_IN_def f_i' by(rule ext nn_integral_eq_integral[symmetric] integrable_f_i')+ simp also have "ennreal (\ y. f' (y, x) \count_space UNIV) = d_IN f' x" unfolding d_IN_def by(rule nn_integral_eq_integral[symmetric])(simp_all add: f'_nonneg, simp add: suminf[symmetric] integrable_suminf integrable_k' summable_nk summable'_k') finally show "d_OUT f' x = d_IN f' x" using * by(blast intro: LIMSEQ_unique) qed moreover { have "incseq (\i. value_flow \ (f_i i))" by(rule incseq_SucI)(simp add: value_aux value_f_i real field_simps \'_nonneg ennreal_leI del: f_i_simps) then have "(\i. value_flow \ (f_i i)) \ (SUP i. value_flow \ (f_i i))" by(rule LIMSEQ_SUP) also have "(SUP i. value_flow \ (f_i i)) = \" proof - have "\ - (SUP i. value_flow \ (f_i i)) = (INF i. \ - value_flow \ (f_i i))" by(simp add: ennreal_SUP_const_minus real) also have "\ - value_flow \ (f_i i) = \' / 2 ^ i" for i by(simp add: value_f_i value_aux real ennreal_minus_if field_simps mult_le_cancel_left1) hence "(INF i. \ - value_flow \ (f_i i)) = (INF i. ennreal (\' / 2 ^ i))" by(auto intro: INF_cong) also have "\ = 0" proof(rule LIMSEQ_unique) show "(\i. \' / 2 ^ i) \ (INF i. ennreal (\' / 2 ^ i))" by(rule LIMSEQ_INF)(simp add: field_simps real decseq_SucI) qed(simp add: LIMSEQ_divide_realpow_zero real ennreal_0[symmetric] del: ennreal_0) finally show "(SUP i. value_flow \ (f_i i)) = \" apply (intro antisym) apply (auto simp: \_def intro!: SUP_mono fn.g) [] apply (rule ennreal_minus_eq_0) apply assumption done qed also have "(\i. value_flow \ (f_i i)) \ value_flow \ f'" by(simp add: value_flow[OF flow source_out] value_flow[OF fn.g source_out] f) ultimately have "value_flow \ f' = \" by(blast intro: LIMSEQ_unique) } note value_f = this moreover { fix x have "d_IN f' x = \\<^sup>+ y. liminf (\i. f_i i (y, x)) \count_space UNIV" unfolding d_IN_def using f by(simp add: tendsto_iff_Liminf_eq_Limsup) also have "\ \ liminf (\i. d_IN (f_i i) x)" unfolding d_IN_def by(rule nn_integral_liminf)(simp_all add:) also have "\ \ liminf (\i. \)" using d_IN_i[of _ x] fn.g by(auto intro!: Liminf_mono SUP_upper2 eventually_sequentiallyI simp add: \_def) also have "\ = value_flow \ f'" using value_f by(simp add: Liminf_const) also note calculation } ultimately show ?thesis by blast qed theorem ex_max_flow'': \ \eliminate assumption of no antiparallel edges using locale @{const wf_residual_network}\ assumes source_out: "\y. edge \ (source \) y \ y = x" and nontrivial: "\<^bold>E \ {}" and real: "\ = ennreal \'" and nn[simp]: "0 \ \'" shows "\f. flow \ f \ value_flow \ f = \ \ (\x. d_IN f x \ value_flow \ f)" proof - interpret antiparallel_edges \ .. interpret \'': flow_attainability \'' by(rule \''_flow_attainability flow_attainability.axioms(2))+unfold_locales have wf_\'': "\''.wf_residual_network" by(rule \''_wf_residual_network; simp add: no_loop) have source_out': "edge \'' (source \'') y \ y = Edge (source \) x" for y by(auto simp add: source_out) have nontrivial': "\<^bold>V\<^bsub>\''\<^esub> - {source \'', sink \''} \ {}" using nontrivial by(auto simp add: "\<^bold>V_\''") have "(SUP g \ {g. flow \'' g}. value_flow \'' g) = (SUP g \ {g. flow \ g}. value_flow \ g)" (is "?lhs = ?rhs") proof(intro antisym SUP_least; unfold mem_Collect_eq) fix g assume g: "flow \'' g" hence "value_flow \'' g = value_flow \ (collect g)" by(simp add: value_collect) also { from g have "flow \ (collect g)" by simp } then have "\ \ ?rhs" by(blast intro: SUP_upper2) finally show "value_flow \'' g \ \" . next fix g assume g: "flow \ g" hence "value_flow \ g = value_flow \'' (split g)" by simp also { from g have "flow \'' (split g)" by simp } then have "\ \ ?lhs" by(blast intro: SUP_upper2) finally show "value_flow \ g \ ?lhs" . qed with real have eq: "(SUP g \ {g. flow \'' g}. value_flow \'' g) = ennreal \'" by(simp add: \_def) from \''.ex_max_flow'[OF wf_\'' source_out' nontrivial' eq] obtain f where f: "flow \'' f" and "value_flow \'' f = \" and IN: "\x. d_IN f x \ value_flow \'' f" unfolding eq real using nn by blast hence "flow \ (collect f)" and "value_flow \ (collect f) = \" by(simp_all add: value_collect) moreover { fix x have "d_IN (collect f) x = (\\<^sup>+ y\range (\y. Edge y x). f (y, Vertex x))" by(simp add: nn_integral_count_space_reindex d_IN_def) also have "\ \ d_IN f (Vertex x)" unfolding d_IN_def by (auto intro!: nn_integral_mono simp add: nn_integral_count_space_indicator split: split_indicator) also have "\ \ value_flow \ (collect f)" using IN[of "Vertex x"] f by(simp add: value_collect) also note calculation } ultimately show ?thesis by blast qed context begin \ \We eliminate the assumption of only one edge leaving the source by introducing a new source vertex.\ private datatype (plugins del: transfer size) 'v' node = SOURCE | Inner (inner: 'v') private lemma not_Inner_conv: "x \ range Inner \ x = SOURCE" by(cases x) auto private lemma inj_on_Inner [simp]: "inj_on Inner A" by(simp add: inj_on_def) private inductive edge' :: "'v node \ 'v node \ bool" where SOURCE: "edge' SOURCE (Inner (source \))" | Inner: "edge \ x y \ edge' (Inner x) (Inner y)" private inductive_simps edge'_simps [simp]: "edge' SOURCE x" "edge' (Inner y) x" "edge' y SOURCE" "edge' y (Inner x)" private fun capacity' :: "'v node flow" where "capacity' (SOURCE, Inner x) = (if x = source \ then \ else 0)" | "capacity' (Inner x, Inner y) = capacity \ (x, y)" | "capacity' _ = 0" private lemma capacity'_source_in [simp]: "capacity' (y, Inner (source \)) = (if y = SOURCE then \ else 0)" by(cases y)(simp_all add: capacity_outside source_in) private definition \' :: "'v node network" where "\' = \edge = edge', capacity = capacity', source = SOURCE, sink = Inner (sink \)\" private lemma \'_sel [simp]: "edge \' = edge'" "capacity \' = capacity'" "source \' = SOURCE" "sink \' = Inner (sink \)" by(simp_all add: \'_def) private lemma "\<^bold>E_\'": "\<^bold>E\<^bsub>\'\<^esub> = {(SOURCE, Inner (source \))} \ (\(x, y). (Inner x, Inner y)) ` \<^bold>E" by(auto elim: edge'.cases) private lemma \'_countable_network: assumes "\ \ \" shows "countable_network \'" proof show "countable \<^bold>E\<^bsub>\'\<^esub>" unfolding "\<^bold>E_\'" by simp show "source \' \ sink \'" by simp show "capacity \' e = 0" if "e \ \<^bold>E\<^bsub>\'\<^esub>" for e using that unfolding "\<^bold>E_\'" by(cases e rule: capacity'.cases)(auto intro: capacity_outside) show "capacity \' e \ \" for e by(cases e rule: capacity'.cases)(simp_all add: assms) qed private lemma \'_flow_attainability: assumes "\ \ \" shows "flow_attainability \'" proof - interpret \': countable_network \' using assms by(rule \'_countable_network) show ?thesis proof show "d_IN (capacity \') x \ \ \ d_OUT (capacity \') x \ \" if sink: "x \ sink \'" for x proof(cases x) case (Inner x') consider (source) "x' = source \" | (IN) "x' \ source \" "d_IN (capacity \) x' \ \" | (OUT) "d_OUT (capacity \) x' \ \" using finite_capacity[of x'] sink Inner by(auto) thus ?thesis proof(cases) case source with Inner have "d_IN (capacity \') x = (\\<^sup>+ y. \ * indicator {SOURCE :: 'v node} y)" unfolding d_IN_def by(intro nn_integral_cong)(simp split: split_indicator) also have "\ = \" by(simp add: max_def) finally show ?thesis using assms by simp next case IN with Inner have "d_IN (capacity \') x = (\\<^sup>+ y\range Inner. capacity \ (node.inner y, x'))" by(auto simp add: d_IN_def nn_integral_count_space_indicator not_Inner_conv intro!: nn_integral_cong split: split_indicator) also have "\ = d_IN (capacity \) x'" unfolding d_IN_def by(simp add: nn_integral_count_space_reindex) finally show ?thesis using Inner sink IN by(simp) next case OUT from Inner have "d_OUT (capacity \') x = (\\<^sup>+ y\range Inner. capacity \ (x', node.inner y))" by(auto simp add: d_OUT_def nn_integral_count_space_indicator not_Inner_conv intro!: nn_integral_cong split: split_indicator) also have "\ = d_OUT (capacity \) x'" by(simp add: d_OUT_def nn_integral_count_space_reindex) finally show ?thesis using OUT by auto qed qed(simp add: d_IN_def) show "\ edge \' x x" for x by(cases x)(simp_all add: no_loop) show "\ edge \' x (source \')" for x by simp qed qed private fun lift :: "'v flow \ 'v node flow" where "lift f (SOURCE, Inner y) = (if y = source \ then value_flow \ f else 0)" | "lift f (Inner x, Inner y) = f (x, y)" | "lift f _ = 0" private lemma d_OUT_lift_Inner [simp]: "d_OUT (lift f) (Inner x) = d_OUT f x" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+ y\range Inner. lift f (Inner x, y))" by(auto simp add: d_OUT_def nn_integral_count_space_indicator not_Inner_conv intro!: nn_integral_cong split: split_indicator) also have "\ = ?rhs" by(simp add: nn_integral_count_space_reindex d_OUT_def) finally show ?thesis . qed private lemma d_OUT_lift_SOURCE [simp]: "d_OUT (lift f) SOURCE = value_flow \ f" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+ y. lift f (SOURCE, y) * indicator {Inner (source \)} y)" unfolding d_OUT_def by(rule nn_integral_cong)(case_tac x; simp) also have "\ = ?rhs" by(simp add: nn_integral_count_space_indicator max_def) finally show ?thesis . qed private lemma d_IN_lift_Inner [simp]: assumes "x \ source \" shows "d_IN (lift f) (Inner x) = d_IN f x" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+ y\range Inner. lift f (y, Inner x))" using assms by(auto simp add: d_IN_def nn_integral_count_space_indicator not_Inner_conv intro!: nn_integral_cong split: split_indicator) also have "\ = ?rhs" by(simp add: nn_integral_count_space_reindex d_IN_def) finally show ?thesis . qed private lemma d_IN_lift_source [simp]: "d_IN (lift f) (Inner (source \)) = value_flow \ f + d_IN f (source \)" (is "?lhs = ?rhs") proof - have "?lhs = (\\<^sup>+ y. lift f (y, Inner (source \)) * indicator {SOURCE} y) + (\\<^sup>+ y\range Inner. lift f (y, Inner (source \)))" (is "_ = ?SOURCE + ?rest") unfolding d_IN_def apply(subst nn_integral_count_space_indicator, simp) apply(subst nn_integral_add[symmetric]) apply(auto simp add: AE_count_space max_def not_Inner_conv split: split_indicator intro!: nn_integral_cong) done also have "?rest = d_IN f (source \)" by(simp add: nn_integral_count_space_reindex d_IN_def) also have "?SOURCE = value_flow \ f" by(simp add: max_def one_ennreal_def[symmetric] ) finally show ?thesis . qed private lemma flow_lift [simp]: assumes "flow \ f" shows "flow \' (lift f)" proof show "lift f e \ capacity \' e" for e by(cases e rule: capacity'.cases)(auto intro: flowD_capacity[OF assms] simp add: \_def intro: SUP_upper2 assms) fix x assume x: "x \ source \'" "x \ sink \'" then obtain x' where x': "x = Inner x'" by(cases x) auto then show "KIR (lift f) x" using x by(cases "x' = source \")(auto simp add: flowD_source_IN[OF assms] dest: flowD_KIR[OF assms]) qed private abbreviation (input) unlift :: "'v node flow \ 'v flow" where "unlift f \ (\(x, y). f (Inner x, Inner y))" private lemma flow_unlift [simp]: assumes f: "flow \' f" shows "flow \ (unlift f)" proof show "unlift f e \ capacity \ e" for e using flowD_capacity[OF f, of "map_prod Inner Inner e"] by(cases e)(simp) next fix x assume x: "x \ source \" "x \ sink \" have "d_OUT (unlift f) x = (\\<^sup>+ y\range Inner. f (Inner x, y))" by(simp add: nn_integral_count_space_reindex d_OUT_def) also have "\ = d_OUT f (Inner x)" using flowD_capacity[OF f, of "(Inner x, SOURCE)"] by(auto simp add: nn_integral_count_space_indicator d_OUT_def not_Inner_conv intro!: nn_integral_cong split: split_indicator) also have "\ = d_IN f (Inner x)" using x flowD_KIR[OF f, of "Inner x"] by(simp) also have "\ = (\\<^sup>+ y\range Inner. f (y, Inner x))" using x flowD_capacity[OF f, of "(SOURCE, Inner x)"] by(auto simp add: nn_integral_count_space_indicator d_IN_def not_Inner_conv intro!: nn_integral_cong split: split_indicator) also have "\ = d_IN (unlift f) x" by(simp add: nn_integral_count_space_reindex d_IN_def) finally show "KIR (unlift f) x" . qed private lemma value_unlift: assumes f: "flow \' f" shows "value_flow \ (unlift f) = value_flow \' f" proof - have "value_flow \ (unlift f) = (\\<^sup>+ y\range Inner. f (Inner (source \), y))" by(simp add: nn_integral_count_space_reindex d_OUT_def) also have "\ = d_OUT f (Inner (source \))" using flowD_capacity[OF f, of "(Inner (source \), SOURCE)"] by(auto simp add: nn_integral_count_space_indicator d_OUT_def not_Inner_conv intro!: nn_integral_cong split: split_indicator) also have "\ = d_IN f (Inner (source \))" using flowD_KIR[OF f, of "Inner (source \)"] by(simp) also have "\ = (\\<^sup>+ y. f (y, Inner (source \)) * indicator {SOURCE} y)" unfolding d_IN_def using flowD_capacity[OF f, of "(x, Inner (source \))" for x] by(intro nn_integral_cong)(auto intro!: antisym split: split_indicator if_split_asm elim: meta_allE) also have "\ = f (SOURCE, Inner (source \))" by simp also have "\ = (\\<^sup>+ y. f (SOURCE, y) * indicator {Inner (source \)} y)" by(simp add: one_ennreal_def[symmetric]) also have "\ = value_flow \' f" unfolding d_OUT_def unfolding d_OUT_def using flowD_capacity[OF f, of "(SOURCE, Inner x)" for x] flowD_capacity[OF f, of "(SOURCE, SOURCE)"] apply(intro nn_integral_cong) apply(case_tac x) apply(auto intro!: antisym split: split_indicator if_split_asm elim: meta_allE) done finally show ?thesis . qed theorem ex_max_flow: "\f. flow \ f \ value_flow \ f = \ \ (\x. d_IN f x \ value_flow \ f)" proof(cases "\") case (real \') hence \: "\ \ \" by simp then interpret \': flow_attainability \' by(rule \'_flow_attainability) have source_out: "edge \' (source \') y \ y = Inner (source \)" for y by(auto) have nontrivial: "\<^bold>E\<^bsub>\'\<^esub> \ {}" by(auto intro: edge'.intros) have eq: "(SUP g \ {g. flow \' g}. value_flow \' g) = (SUP g \ {g. flow \ g}. value_flow \ g)" (is "?lhs = ?rhs") proof(intro antisym SUP_least; unfold mem_Collect_eq) fix g assume g: "flow \' g" hence "value_flow \' g = value_flow \ (unlift g)" by(simp add: value_unlift) also { from g have "flow \ (unlift g)" by simp } then have "\ \ ?rhs" by(blast intro: SUP_upper2) finally show "value_flow \' g \ \" . next fix g assume g: "flow \ g" hence "value_flow \ g = value_flow \' (lift g)" by simp also { from g have "flow \' (lift g)" by simp } then have "\ \ ?lhs" by(blast intro: SUP_upper2) finally show "value_flow \ g \ ?lhs" . qed also have "\ = ennreal \'" using real by(simp add: \_def) finally obtain f where f: "flow \' f" and value_f: "value_flow \' f = (\g\{g. flow \' g}. value_flow \' g)" and IN_f: "\x. d_IN f x \ value_flow \' f" using \0 \ \'\ by(blast dest: \'.ex_max_flow''[OF source_out nontrivial]) have "flow \ (unlift f)" using f by simp moreover have "value_flow \ (unlift f) = \" using f eq value_f by(simp add: value_unlift \_def) moreover { fix x have "d_IN (unlift f) x = (\\<^sup>+ y\range Inner. f (y, Inner x))" by(simp add: nn_integral_count_space_reindex d_IN_def) also have "\ \ d_IN f (Inner x)" unfolding d_IN_def by(auto intro!: nn_integral_mono simp add: nn_integral_count_space_indicator split: split_indicator) also have "\ \ value_flow \ (unlift f)" using IN_f[of "Inner x"] f by(simp add: value_unlift) also note calculation } ultimately show ?thesis by blast next case top show ?thesis proof(cases "\f. flow \ f \ value_flow \ f = \") case True with top show ?thesis by auto next case False hence real: "\f. \ = \ \ flow \ f \ value_flow \ f < \" using top by (auto simp: less_top) { fix i have "2 * 2 ^ i < \" using top by (simp_all add: ennreal_mult_less_top power_less_top_ennreal) from flow_by_value[OF this real] have "\f. flow \ f \ value_flow \ f = 2 * 2 ^ i" by blast } then obtain f_i where f_i: "\i. flow \ (f_i i)" and value_i: "\i. value_flow \ (f_i i) = 2 * 2 ^ i" by metis define f where "f e = (\\<^sup>+ i. f_i i e / (2 * 2 ^ i))" for e have "flow \ f" proof fix e have "f e \ (\\<^sup>+ i. (SUP i. f_i i e) / (2 * 2 ^ i))" unfolding f_def by(rule nn_integral_mono)(auto intro!: divide_right_mono_ennreal SUP_upper) also have "\ = (SUP i. f_i i e) / 2 * (\\<^sup>+ i. 1 / 2 ^ i)" apply(subst nn_integral_cmult[symmetric]) apply(auto intro!: nn_integral_cong intro: SUP_upper2 simp: divide_ennreal_def ennreal_inverse_mult power_less_top_ennreal mult_ac) done also have "(\\<^sup>+ i. 1 / 2 ^ i) = (\i. ennreal ((1 / 2) ^ i))" by(simp add: nn_integral_count_space_nat power_divide divide_ennreal[symmetric] ennreal_power[symmetric]) also have "\ = ennreal (\i. (1 / 2) ^ i)" by(intro suminf_ennreal2 complete_algebra_summable_geometric) simp_all also have "\ = 2" by(subst suminf_geometric; simp) also have "(SUP i. f_i i e) / 2 * 2 = (SUP i. f_i i e)" by (simp add: ennreal_divide_times) also have "\ \ capacity \ e" by(rule SUP_least)(rule flowD_capacity[OF f_i]) finally show "f e \ capacity \ e" . fix x assume x: "x \ source \" "x \ sink \" have "d_OUT f x = (\\<^sup>+ i\UNIV. \\<^sup>+ y. f_i i (x, y) / (2 * 2 ^ i))" unfolding d_OUT_def f_def by(subst nn_integral_snd_count_space[where f="case_prod _", simplified]) (simp add: nn_integral_fst_count_space[where f="case_prod _", simplified]) also have "\ = (\\<^sup>+ i. d_OUT (f_i i) x / (2 * 2 ^ i))" unfolding d_OUT_def by(simp add: nn_integral_divide) also have "\ = (\\<^sup>+ i. d_IN (f_i i) x / (2 * 2 ^ i))" by(simp add: flowD_KIR[OF f_i, OF x]) also have "\ = (\\<^sup>+ i\UNIV. \\<^sup>+ y. f_i i (y, x) / (2 * 2 ^ i))" by(simp add: nn_integral_divide d_IN_def) also have "\ = d_IN f x" unfolding d_IN_def f_def by(subst nn_integral_snd_count_space[where f="case_prod _", simplified]) (simp add: nn_integral_fst_count_space[where f="case_prod _", simplified]) finally show "KIR f x" . qed moreover { have "value_flow \ f = (\\<^sup>+ i. value_flow \ (f_i i) / (2 * 2 ^ i))" unfolding d_OUT_def f_def by(subst nn_integral_snd_count_space[where f="case_prod _", simplified]) (simp add: nn_integral_fst_count_space[where f="case_prod _", simplified] nn_integral_divide[symmetric]) also have "\ = \" by(simp add: value_i ennreal_mult_less_top power_less_top_ennreal) finally have "value_flow \ f = \" . } ultimately show ?thesis using top by auto qed qed end end end end \ No newline at end of file diff --git a/thys/MFMC_Countable/MFMC_Unbounded.thy b/thys/MFMC_Countable/MFMC_Unbounded.thy --- a/thys/MFMC_Countable/MFMC_Unbounded.thy +++ b/thys/MFMC_Countable/MFMC_Unbounded.thy @@ -1,3467 +1,3467 @@ (* Author: Andreas Lochbihler, ETH Zurich *) section \The max-flow min-cut theorems in unbounded networks\ theory MFMC_Unbounded imports MFMC_Web MFMC_Flow_Attainability MFMC_Reduction begin subsection \More about waves\ lemma SINK_plus_current: "SINK (plus_current f g) = SINK f \ SINK g" by(auto simp add: SINK.simps set_eq_iff d_OUT_def nn_integral_0_iff emeasure_count_space_eq_0 add_eq_0_iff_both_eq_0) abbreviation plus_web :: "('v, 'more) web_scheme \ 'v current \ 'v current \ 'v current" ("_ \\ _" [66, 66] 65) where "plus_web \ f g \ plus_current f (g \ \ / f)" lemma d_OUT_plus_web: fixes \ (structure) shows "d_OUT (f \ g) x = d_OUT f x + d_OUT (g \ \ / f) x" (is "?lhs = ?rhs") proof - have "?lhs = d_OUT f x + (\\<^sup>+ y. (if x \ RF\<^sup>\ (TER f) then 0 else g (x, y) * indicator (- RF (TER f)) y))" unfolding d_OUT_def by(subst nn_integral_add[symmetric])(auto intro!: nn_integral_cong split: split_indicator) also have "\ = ?rhs" by(auto simp add: d_OUT_def intro!: arg_cong2[where f="(+)"] nn_integral_cong) finally show "?thesis" . qed lemma d_IN_plus_web: fixes \ (structure) shows "d_IN (f \ g) y = d_IN f y + d_IN (g \ \ / f) y" (is "?lhs = ?rhs") proof - have "?lhs = d_IN f y + (\\<^sup>+ x. (if y \ RF (TER f) then 0 else g (x, y) * indicator (- RF\<^sup>\ (TER f)) x))" unfolding d_IN_def by(subst nn_integral_add[symmetric])(auto intro!: nn_integral_cong split: split_indicator) also have "\ = ?rhs" by(auto simp add: d_IN_def intro!: arg_cong2[where f="(+)"] nn_integral_cong) finally show ?thesis . qed lemma plus_web_greater: "f e \ (f \\<^bsub>\\<^esub> g) e" by(cases e)(auto split: split_indicator) lemma current_plus_web: fixes \ (structure) shows "\ current \ f; wave \ f; current \ g \ \ current \ (f \ g)" by(blast intro: current_plus_current current_restrict_current) context fixes \ :: "('v, 'more) web_scheme" (structure) and f g :: "'v current" assumes f: "current \ f" and w: "wave \ f" and g: "current \ g" begin context fixes x :: "'v" assumes x: "x \ \ (TER f \ TER g)" begin qualified lemma RF_f: "x \ RF\<^sup>\ (TER f)" proof assume *: "x \ RF\<^sup>\ (TER f)" from x obtain p y where p: "path \ x p y" and y: "y \ B \" and bypass: "\z. \x \ y; z \ set p\ \ z = x \ z \ TER f \ TER g" by(rule \_E) blast from rtrancl_path_distinct[OF p] obtain p' where p: "path \ x p' y" and p': "set p' \ set p" and distinct: "distinct (x # p')" . from * have x': "x \ RF (TER f)" and \: "x \ \ (TER f)" by(auto simp add: roofed_circ_def) hence "x \ TER f" using not_essentialD[OF _ p y] p' bypass by blast with roofedD[OF x' p y] obtain z where z: "z \ set p'" "z \ TER f" by auto with p have "y \ set p'" by(auto dest!: rtrancl_path_last intro: last_in_set) with distinct have "x \ y" by auto with bypass z p' distinct show False by auto qed private lemma RF_g: "x \ RF\<^sup>\ (TER g)" proof assume *: "x \ RF\<^sup>\ (TER g)" from x obtain p y where p: "path \ x p y" and y: "y \ B \" and bypass: "\z. \x \ y; z \ set p\ \ z = x \ z \ TER f \ TER g" by(rule \_E) blast from rtrancl_path_distinct[OF p] obtain p' where p: "path \ x p' y" and p': "set p' \ set p" and distinct: "distinct (x # p')" . from * have x': "x \ RF (TER g)" and \: "x \ \ (TER g)" by(auto simp add: roofed_circ_def) hence "x \ TER g" using not_essentialD[OF _ p y] p' bypass by blast with roofedD[OF x' p y] obtain z where z: "z \ set p'" "z \ TER g" by auto with p have "y \ set p'" by(auto dest!: rtrancl_path_last intro: last_in_set) with distinct have "x \ y" by auto with bypass z p' distinct show False by auto qed lemma TER_plus_web_aux: assumes SINK: "x \ SINK (g \ \ / f)" (is "_ \ SINK ?g") shows "x \ TER (f \ g)" proof from x obtain p y where p: "path \ x p y" and y: "y \ B \" and bypass: "\z. \x \ y; z \ set p\ \ z = x \ z \ TER f \ TER g" by(rule \_E) blast from rtrancl_path_distinct[OF p] obtain p' where p: "path \ x p' y" and p': "set p' \ set p" and distinct: "distinct (x # p')" . from RF_f have "x \ SINK f" by(auto simp add: roofed_circ_def SINK.simps dest: waveD_OUT[OF w]) thus "x \ SINK (f \ g)" using SINK by(simp add: SINK.simps d_OUT_plus_web) show "x \ SAT \ (f \ g)" proof(cases "x \ TER f") case True hence "x \ SAT \ f" by simp moreover have "\ \ SAT \ (f \ g)" by(rule SAT_mono plus_web_greater)+ ultimately show ?thesis by blast next case False with x have "x \ TER g" by auto from False RF_f have "x \ RF (TER f)" by(auto simp add: roofed_circ_def) moreover { fix z assume z: "z \ RF\<^sup>\ (TER f)" have "(z, x) \ \<^bold>E" proof assume "(z, x) \ \<^bold>E" hence path': "path \ z (x # p') y" using p by(simp add: rtrancl_path.step) from z have "z \ RF (TER f)" by(simp add: roofed_circ_def) from roofedD[OF this path' y] False consider (path) z' where "z' \ set p'" "z' \ TER f" | (TER) "z \ TER f" by auto then show False proof cases { case (path z') with p distinct have "x \ y" by(auto 4 3 intro: last_in_set elim: rtrancl_path.cases dest: rtrancl_path_last[symmetric]) from bypass[OF this, of z'] path False p' show False by auto } note that = this case TER with z have "\ essential \ (B \) (TER f) z" by(simp add: roofed_circ_def) from not_essentialD[OF this path' y] False obtain z' where "z' \ set p'" "z' \ TER f" by auto thus False by(rule that) qed qed } ultimately have "d_IN ?g x = d_IN g x" unfolding d_IN_def by(intro nn_integral_cong)(clarsimp split: split_indicator simp add: currentD_outside[OF g]) hence "d_IN (f \ g) x \ d_IN g x" by(simp add: d_IN_plus_web) with \x \ TER g\ show ?thesis by(auto elim!: SAT.cases intro: SAT.intros) qed qed qualified lemma SINK_TER_in'': assumes "\x. x \ RF (TER g) \ d_OUT g x = 0" shows "x \ SINK g" using RF_g by(auto simp add: roofed_circ_def SINK.simps assms) end lemma wave_plus: "wave (quotient_web \ f) (g \ \ / f) \ wave \ (f \ g)" using f w by(rule wave_plus_current)(rule current_restrict_current[OF w g]) lemma TER_plus_web'': assumes "\x. x \ RF (TER g) \ d_OUT g x = 0" shows "\ (TER f \ TER g) \ TER (f \ g)" proof fix x assume *: "x \ \ (TER f \ TER g)" moreover have "x \ SINK (g \ \ / f)" by(rule in_SINK_restrict_current)(rule MFMC_Unbounded.SINK_TER_in''[OF f w g * assms]) ultimately show "x \ TER (f \ g)" by(rule TER_plus_web_aux) qed lemma TER_plus_web': "wave \ g \ \ (TER f \ TER g) \ TER (f \ g)" by(rule TER_plus_web'')(rule waveD_OUT) lemma wave_plus': "wave \ g \ wave \ (f \ g)" by(rule wave_plus)(rule wave_restrict_current[OF f w g]) end lemma RF_TER_plus_web: fixes \ (structure) assumes f: "current \ f" and w: "wave \ f" and g: "current \ g" and w': "wave \ g" shows "RF (TER (f \ g)) = RF (TER f \ TER g)" proof have "RF (\ (TER f \ TER g)) \ RF (TER (f \ g))" by(rule roofed_mono)(rule TER_plus_web'[OF f w g w']) also have "RF (\ (TER f \ TER g)) = RF (TER f \ TER g)" by(rule RF_essential) finally show "\ \ RF (TER (f \ g))" . next have fg: "current \ (f \ g)" using f w g by(rule current_plus_web) show "RF (TER (f \ g)) \ RF (TER f \ TER g)" proof(intro subsetI roofedI) fix x p y assume RF: "x \ RF (TER (f \ g))" and p: "path \ x p y" and y: "y \ B \" from roofedD[OF RF p y] obtain z where z: "z \ set (x # p)" and TER: "z \ TER (f \ g)" by auto from TER have SINK: "z \ SINK f" by(auto simp add: SINK.simps d_OUT_plus_web add_eq_0_iff_both_eq_0) from TER have "z \ SAT \ (f \ g)" by simp hence SAT: "z \ SAT \ f \ SAT \ g" by(cases "z \ RF (TER f)")(auto simp add: currentD_SAT[OF f] currentD_SAT[OF g] currentD_SAT[OF fg] d_IN_plus_web d_IN_restrict_current_outside restrict_current_IN_not_RF[OF g] wave_not_RF_IN_zero[OF f w]) show "(\z\set p. z \ TER f \ TER g) \ x \ TER f \ TER g" proof(cases "z \ RF (TER g)") case False hence "z \ SINK g" by(simp add: SINK.simps waveD_OUT[OF w']) with SINK SAT have "z \ TER f \ TER g" by auto thus ?thesis using z by auto next case True from split_list[OF z] obtain ys zs where split: "x # p = ys @ z # zs" by blast with p have "path \ z zs y" by(auto elim: rtrancl_path_appendE simp add: Cons_eq_append_conv) from roofedD[OF True this y] split show ?thesis by(auto simp add: Cons_eq_append_conv) qed qed qed lemma RF_TER_Sup: fixes \ (structure) assumes f: "\f. f \ Y \ current \ f" and w: "\f. f \ Y \ wave \ f" and Y: "Complete_Partial_Order.chain (\) Y" "Y \ {}" "countable (support_flow (Sup Y))" shows "RF (TER (Sup Y)) = RF (\f\Y. TER f)" proof(rule set_eqI iffI)+ fix x assume x: "x \ RF (TER (Sup Y))" have "x \ RF (RF (\f\Y. TER f))" proof fix p y assume p: "path \ x p y" and y: "y \ B \" from roofedD[OF x p y] obtain z where z: "z \ set (x # p)" and TER: "z \ TER (Sup Y)" by auto from TER have SINK: "z \ SINK f" if "f \ Y" for f using that by(auto simp add: SINK_Sup[OF Y]) from Y(2) obtain f where y: "f \ Y" by blast show "(\z\set p. z \ RF (\f\Y. TER f)) \ x \ RF (\f\Y. TER f)" proof(cases "\f\Y. z \ RF (TER f)") case True then obtain f where fY: "f \ Y" and zf: "z \ RF (TER f)" by blast from zf have "z \ RF (\f\Y. TER f)" by(rule in_roofed_mono)(auto intro: fY) with z show ?thesis by auto next case False hence *: "d_IN f z = 0" if "f \ Y" for f using that by(auto intro: wave_not_RF_IN_zero[OF f w]) hence "d_IN (Sup Y) z = 0" using Y(2) by(simp add: d_IN_Sup[OF Y]) with TER have "z \ SAT \ f" using *[OF y] by(simp add: SAT.simps) with SINK[OF y] have "z \ TER f" by simp with z y show ?thesis by(auto intro: roofed_greaterI) qed qed then show "x \ RF (\f\Y. TER f)" unfolding roofed_idem . next fix x assume x: "x \ RF (\f\Y. TER f)" have "x \ RF (RF (TER (\Y)))" proof(rule roofedI) fix p y assume p: "path \ x p y" and y: "y \ B \" from roofedD[OF x p y] obtain z f where *: "z \ set (x # p)" and **: "f \ Y" and TER: "z \ TER f" by auto have "z \ RF (TER (Sup Y))" proof(rule ccontr) assume z: "z \ RF (TER (Sup Y))" have "wave \ (Sup Y)" using Y(1-2) w Y(3) by(rule wave_lub) hence "d_OUT (Sup Y) z = 0" using z by(rule waveD_OUT) hence "z \ SINK (Sup Y)" by(simp add: SINK.simps) moreover have "z \ SAT \ (Sup Y)" using TER SAT_Sup_upper[OF **, of \] by blast ultimately have "z \ TER (Sup Y)" by simp hence "z \ RF (TER (Sup Y))" by(rule roofed_greaterI) with z show False by contradiction qed thus "(\z\set p. z \ RF (TER (Sup Y))) \ x \ RF (TER (Sup Y))" using * by auto qed then show "x \ RF (TER (\Y))" unfolding roofed_idem . qed subsection \Hindered webs with reduced weights\ context countable_bipartite_web begin context fixes u :: "'v \ ennreal" and \ defines "\ \ (\\<^sup>+ y. u y \count_space (B \))" assumes u_outside: "\x. x \ B \ \ u x = 0" and finite_\: "\ \ \" begin private lemma u_A: "x \ A \ \ u x = 0" using u_outside[of x] disjoint by auto private lemma u_finite: "u y \ \" proof(cases "y \ B \") case True have "u y \ \" unfolding \_def by(rule nn_integral_ge_point)(simp add: True) also have "\ < \" using finite_\ by (simp add: less_top[symmetric]) finally show ?thesis by simp qed(simp add: u_outside) lemma hindered_reduce: \ \Lemma 6.7\ assumes u: "u \ weight \" assumes hindered_by: "hindered_by (\\weight := weight \ - u\) \" (is "hindered_by ?\ _") shows "hindered \" proof - note [simp] = u_finite let ?TER = "TER\<^bsub>?\\<^esub>" from hindered_by obtain f where hindrance_by: "hindrance_by ?\ f \" and f: "current ?\ f" and w: "wave ?\ f" by cases from hindrance_by obtain a where a: "a \ A \" "a \ \\<^bsub>?\\<^esub> (?TER f)" and a_le: "d_OUT f a < weight \ a" and \_less: "weight \ a - d_OUT f a > \" and \_nonneg: "\ \ 0" by(auto simp add: u_A hindrance_by.simps) from f have f': "current \ f" by(rule current_weight_mono)(auto intro: diff_le_self_ennreal) write Some ("\_\") define edge' where "edge' xo yo = (case (xo, yo) of (None, Some y) \ y \ \<^bold>V \ y \ A \ | (Some x, Some y) \ edge \ x y \ edge \ y x | _ \ False)" for xo yo define cap where "cap e = (case e of (None, Some y) \ if y \ \<^bold>V then u y else 0 | (Some x, Some y) \ if edge \ x y \ x \ a then f (x, y) else if edge \ y x then max (weight \ x) (weight \ y) + 1 else 0 | _ \ 0)" for e define \ where "\ = \edge = edge', capacity = cap, source = None, sink = Some a\" have edge'_simps [simp]: "edge' None \y\ \ y \ \<^bold>V \ y \ A \" "edge' xo None \ False" "edge' \x\ \y\ \ edge \ x y \ edge \ y x" for xo x y by(simp_all add: edge'_def split: option.split) have edge_None1E [elim!]: thesis if "edge' None y" "\z. \ y = \z\; z \ \<^bold>V; z \ A \ \ \ thesis" for y thesis using that by(simp add: edge'_def split: option.split_asm sum.split_asm) have edge_Some1E [elim!]: thesis if "edge' \x\ y" "\z. \ y = \z\; edge \ x z \ edge \ z x \ \ thesis" for x y thesis using that by(simp add: edge'_def split: option.split_asm sum.split_asm) have edge_Some2E [elim!]: thesis if "edge' x \y\" "\ x = None; y \ \<^bold>V; y \ A \ \ \ thesis" "\z. \ x = \z\; edge \ z y \ edge \ y z \ \ thesis" for x y thesis using that by(simp add: edge'_def split: option.split_asm sum.split_asm) have cap_simps [simp]: "cap (None, \y\) = (if y \ \<^bold>V then u y else 0)" "cap (xo, None) = 0" "cap (\x\, \y\) = (if edge \ x y \ x \ a then f (x, y) else if edge \ y x then max (weight \ x) (weight \ y) + 1 else 0)" for xo x y by(simp_all add: cap_def split: option.split) have \_sel [simp]: "edge \ = edge'" "capacity \ = cap" "source \ = None" "sink \ = \a\" by(simp_all add: \_def) have cap_outside1: "\ vertex \ x \ cap (\x\, y) = 0" for x y by(cases y)(auto simp add: vertex_def) have capacity_A_weight: "d_OUT cap \x\ \ weight \ x" if "x \ A \" for x proof - have "d_OUT cap \x\ \ (\\<^sup>+ y\range Some. f (x, the y))" using that disjoint a(1) unfolding d_OUT_def by(auto 4 4 intro!: nn_integral_mono simp add: nn_integral_count_space_indicator notin_range_Some currentD_outside[OF f] split: split_indicator dest: edge_antiparallel bipartite_E) also have "\ = d_OUT f x" by(simp add: d_OUT_def nn_integral_count_space_reindex) also have "\ \ weight \ x" using f' by(rule currentD_weight_OUT) finally show ?thesis . qed have flow_attainability: "flow_attainability \" proof have "\<^bold>E\<^bsub>\\<^esub> = (\(x, y). (\x\, \y\)) ` \<^bold>E \ (\(x, y). (\y\, \x\)) ` \<^bold>E \ (\x. (None, \x\)) ` (\<^bold>V \ - A \)" by(auto simp add: edge'_def split: option.split_asm) thus "countable \<^bold>E\<^bsub>\\<^esub>" by simp next fix v assume "v \ sink \" consider (sink) "v = None" | (A) x where "v = \x\" "x \ A \" | (B) y where "v = \y\" "y \ A \" "y \ \<^bold>V" | (outside) x where "v = \x\" "x \ \<^bold>V" by(cases v) auto then show "d_IN (capacity \) v \ \ \ d_OUT (capacity \) v \ \" proof cases case sink thus ?thesis by(simp add: d_IN_def) next case (A x) thus ?thesis using capacity_A_weight[of x] by (auto simp: top_unique) next case (B y) have "d_IN (capacity \) v \ (\\<^sup>+ x. f (the x, y) * indicator (range Some) x + u y * indicator {None} x)" using B disjoint bipartite_V a(1) unfolding d_IN_def by(auto 4 4 intro!: nn_integral_mono simp add: nn_integral_count_space_indicator notin_range_Some currentD_outside[OF f] split: split_indicator dest: edge_antiparallel bipartite_E) also have "\ = (\\<^sup>+ x\range Some. f (the x, y)) + u y" by(subst nn_integral_add)(simp_all add: nn_integral_count_space_indicator) also have "\ = d_IN f y + u y" by(simp add: d_IN_def nn_integral_count_space_reindex) also have "d_IN f y \ weight \ y" using f' by(rule currentD_weight_IN) finally show ?thesis by(auto simp add: add_right_mono top_unique split: if_split_asm) next case outside hence "d_OUT (capacity \) v = 0" by(auto simp add: d_OUT_def nn_integral_0_iff_AE AE_count_space cap_def vertex_def split: option.split) thus ?thesis by simp qed next show "capacity \ e \ \" for e using weight_finite by(auto simp add: cap_def max_def vertex_def currentD_finite[OF f'] split: option.split prod.split simp del: weight_finite) show "capacity \ e = 0" if "e \ \<^bold>E\<^bsub>\\<^esub>" for e using that bipartite_V disjoint by(auto simp add: cap_def max_def intro: u_outside split: option.split prod.split) show "\ edge \ x (source \)" for x by simp show "\ edge \ x x" for x by(cases x)(simp_all add: no_loop) show "source \ \ sink \" by simp qed then interpret \: flow_attainability "\" . define \ where "\ = (\g\{g. flow \ g}. value_flow \ g)" have \_le: "\ \ \" proof - have "\ \ d_OUT cap None" unfolding \_def by(rule SUP_least)(auto intro!: d_OUT_mono dest: flowD_capacity) also have "\ \ \\<^sup>+ y. cap (None, y) \count_space (range Some)" unfolding d_OUT_def by(auto simp add: nn_integral_count_space_indicator notin_range_Some intro!: nn_integral_mono split: split_indicator) also have "\ \ \" unfolding \_def by (subst (2) nn_integral_count_space_indicator, auto simp add: nn_integral_count_space_reindex u_outside intro!: nn_integral_mono split: split_indicator) finally show ?thesis by simp qed then have finite_flow: "\ \ \" using finite_\ by (auto simp: top_unique) from \.ex_max_flow obtain j where j: "flow \ j" and value_j: "value_flow \ j = \" and IN_j: "\x. d_IN j x \ \" unfolding \_def by auto have j_le_f: "j (Some x, Some y) \ f (x, y)" if "edge \ x y" for x y using that flowD_capacity[OF j, of "(Some x, Some y)"] a(1) disjoint by(auto split: if_split_asm dest: bipartite_E intro: order_trans) have IN_j_finite [simp]: "d_IN j x \ \" for x using finite_flow by(rule neq_top_trans)(simp add: IN_j) have j_finite[simp]: "j (x, y) < \" for x y by (rule le_less_trans[OF d_IN_ge_point]) (simp add: IN_j_finite[of y] less_top[symmetric]) have OUT_j_finite: "d_OUT j x \ \" for x proof(cases "x = source \ \ x = sink \") case True thus ?thesis proof cases case left thus ?thesis using finite_flow value_j by simp next case right have "d_OUT (capacity \) \a\ \ \" using capacity_A_weight[of a] a(1) by(auto simp: top_unique) thus ?thesis unfolding right[simplified] by(rule neq_top_trans)(rule d_OUT_mono flowD_capacity[OF j])+ qed next case False then show ?thesis by(simp add: flowD_KIR[OF j]) qed have IN_j_le_weight: "d_IN j \x\ \ weight \ x" for x proof(cases "x \ A \") case xA: True show ?thesis proof(cases "x = a") case True have "d_IN j \x\ \ \" by(rule IN_j) also have "\ \ \" by(rule \_le) also have "\ < weight \ a" using \_less diff_le_self_ennreal less_le_trans by blast finally show ?thesis using True by(auto intro: order.strict_implies_order) next case False have "d_IN j \x\ = d_OUT j \x\" using flowD_KIR[OF j, of "Some x"] False by simp also have "\ \ d_OUT cap \x\" using flowD_capacity[OF j] by(auto intro: d_OUT_mono) also have "\ \ weight \ x" using xA by(rule capacity_A_weight) finally show ?thesis . qed next case xA: False show ?thesis proof(cases "x \ B \") case True have "d_IN j \x\ \ d_IN cap \x\" using flowD_capacity[OF j] by(auto intro: d_IN_mono) also have "\ \ (\\<^sup>+ z. f (the z, x) * indicator (range Some) z) + (\\<^sup>+ z :: 'v option. u x * indicator {None} z)" using True disjoint by(subst nn_integral_add[symmetric])(auto simp add: vertex_def currentD_outside[OF f] d_IN_def B_out intro!: nn_integral_mono split: split_indicator) also have "\ = d_IN f x + u x" by(simp add: nn_integral_count_space_indicator[symmetric] nn_integral_count_space_reindex d_IN_def) also have "\ \ weight \ x" using currentD_weight_IN[OF f, of x] u_finite[of x] using \_less u by (auto simp add: ennreal_le_minus_iff le_fun_def) finally show ?thesis . next case False with xA have "x \ \<^bold>V" using bipartite_V by blast then have "d_IN j \x\ = 0" using False by(auto simp add: d_IN_def nn_integral_0_iff emeasure_count_space_eq_0 vertex_def edge'_def split: option.split_asm intro!: \.flowD_outside[OF j]) then show ?thesis by simp qed qed let ?j = "j \ map_prod Some Some \ prod.swap" have finite_j_OUT: "(\\<^sup>+ y\\<^bold>O\<^bold>U\<^bold>T x. j (\x\, \y\)) \ \" (is "?j_OUT \ _") if "x \ A \" for x using currentD_finite_OUT[OF f', of x] by(rule neq_top_trans)(auto intro!: nn_integral_mono j_le_f simp add: d_OUT_def nn_integral_count_space_indicator outgoing_def split: split_indicator) have j_OUT_eq: "?j_OUT x = d_OUT j \x\" if "x \ A \" for x proof - have "?j_OUT x = (\\<^sup>+ y\range Some. j (Some x, y))" using that disjoint by(simp add: nn_integral_count_space_reindex)(auto 4 4 simp add: nn_integral_count_space_indicator outgoing_def intro!: nn_integral_cong \.flowD_outside[OF j] dest: bipartite_E split: split_indicator) also have "\ = d_OUT j \x\" by(auto simp add: d_OUT_def nn_integral_count_space_indicator notin_range_Some intro!: nn_integral_cong \.flowD_outside[OF j] split: split_indicator) finally show ?thesis . qed define g where "g = f \ ?j" have g_simps: "g (x, y) = (f \ ?j) (x, y)" for x y by(simp add: g_def) have OUT_g_A: "d_OUT g x = d_OUT f x + d_IN j \x\ - d_OUT j \x\" if "x \ A \" for x proof - have "d_OUT g x = (\\<^sup>+ y\\<^bold>O\<^bold>U\<^bold>T x. f (x, y) + j (\y\, \x\) - j (\x\, \y\))" by(auto simp add: d_OUT_def g_simps currentD_outside[OF f'] outgoing_def nn_integral_count_space_indicator intro!: nn_integral_cong) also have "\ = (\\<^sup>+ y\\<^bold>O\<^bold>U\<^bold>T x. f (x, y) + j (\y\, \x\)) - (\\<^sup>+ y\\<^bold>O\<^bold>U\<^bold>T x. j (\x\, \y\))" (is "_ = _ - ?j_OUT") using finite_j_OUT[OF that] by(subst nn_integral_diff)(auto simp add: AE_count_space outgoing_def intro!: order_trans[OF j_le_f]) also have "\ = (\\<^sup>+ y\\<^bold>O\<^bold>U\<^bold>T x. f (x, y)) + (\\<^sup>+ y\\<^bold>O\<^bold>U\<^bold>T x. j (Some y, Some x)) - ?j_OUT" (is "_ = ?f + ?j_IN - _") by(subst nn_integral_add) simp_all also have "?f = d_OUT f x" by(subst d_OUT_alt_def[where G=\])(simp_all add: currentD_outside[OF f]) also have "?j_OUT = d_OUT j \x\" using that by(rule j_OUT_eq) also have "?j_IN = (\\<^sup>+ y\range Some. j (y, \x\))" using that disjoint by(simp add: nn_integral_count_space_reindex)(auto 4 4 simp add: nn_integral_count_space_indicator outgoing_def intro!: nn_integral_cong \.flowD_outside[OF j] split: split_indicator dest: bipartite_E) also have "\ = d_IN j (Some x)" using that disjoint by(auto 4 3 simp add: d_IN_def nn_integral_count_space_indicator notin_range_Some intro!: nn_integral_cong \.flowD_outside[OF j] split: split_indicator) finally show ?thesis by simp qed have OUT_g_B: "d_OUT g x = 0" if "x \ A \" for x using disjoint that by(auto simp add: d_OUT_def nn_integral_0_iff_AE AE_count_space g_simps dest: bipartite_E) have OUT_g_a: "d_OUT g a < weight \ a" using a(1) proof - have "d_OUT g a = d_OUT f a + d_IN j \a\ - d_OUT j \a\" using a(1) by(rule OUT_g_A) also have "\ \ d_OUT f a + d_IN j \a\" by(rule diff_le_self_ennreal) also have "\ < weight \ a + d_IN j \a\ - \" using finite_\ \_less currentD_finite_OUT[OF f'] by (simp add: less_diff_eq_ennreal less_top ac_simps) also have "\ \ weight \ a" using IN_j[THEN order_trans, OF \_le] by (simp add: ennreal_minus_le_iff) finally show ?thesis . qed have OUT_jj: "d_OUT ?j x = d_IN j \x\ - j (None, \x\)" for x proof - have "d_OUT ?j x = (\\<^sup>+ y\range Some. j (y, \x\))" by(simp add: d_OUT_def nn_integral_count_space_reindex) also have "\ = d_IN j \x\ - (\\<^sup>+ y. j (y, \x\) * indicator {None} y)" unfolding d_IN_def by(subst nn_integral_diff[symmetric])(auto simp add: max_def \.flowD_finite[OF j] AE_count_space nn_integral_count_space_indicator split: split_indicator intro!: nn_integral_cong) also have "\ = d_IN j \x\ - j (None, \x\)" by(simp add: max_def) finally show ?thesis . qed have OUT_jj_finite [simp]: "d_OUT ?j x \ \" for x by(simp add: OUT_jj) have IN_g: "d_IN g x = d_IN f x + j (None, \x\)" for x proof(cases "x \ B \") case True have finite: "(\\<^sup>+ y\\<^bold>I\<^bold>N x. j (Some y, Some x)) \ \" using currentD_finite_IN[OF f', of x] by(rule neq_top_trans)(auto intro!: nn_integral_mono j_le_f simp add: d_IN_def nn_integral_count_space_indicator incoming_def split: split_indicator) have "d_IN g x = d_IN (f \ ?j) x" by(simp add: g_def) also have "\ = (\\<^sup>+ y\\<^bold>I\<^bold>N x. f (y, x) + j (Some x, Some y) - j (Some y, Some x))" by(auto simp add: d_IN_def currentD_outside[OF f'] incoming_def nn_integral_count_space_indicator intro!: nn_integral_cong) also have "\ = (\\<^sup>+ y\\<^bold>I\<^bold>N x. f (y, x) + j (Some x, Some y)) - (\\<^sup>+ y\\<^bold>I\<^bold>N x. j (Some y, Some x))" (is "_ = _ - ?j_IN") using finite by(subst nn_integral_diff)(auto simp add: AE_count_space incoming_def intro!: order_trans[OF j_le_f]) also have "\ = (\\<^sup>+ y\\<^bold>I\<^bold>N x. f (y, x)) + (\\<^sup>+ y\\<^bold>I\<^bold>N x. j (Some x, Some y)) - ?j_IN" (is "_ = ?f + ?j_OUT - _") by(subst nn_integral_add) simp_all also have "?f = d_IN f x" by(subst d_IN_alt_def[where G=\])(simp_all add: currentD_outside[OF f]) also have "?j_OUT = (\\<^sup>+ y\range Some. j (Some x, y))" using True disjoint by(simp add: nn_integral_count_space_reindex)(auto 4 4 simp add: nn_integral_count_space_indicator incoming_def intro!: nn_integral_cong \.flowD_outside[OF j] split: split_indicator dest: bipartite_E) also have "\ = d_OUT j (Some x)" using disjoint by(auto 4 3 simp add: d_OUT_def nn_integral_count_space_indicator notin_range_Some intro!: nn_integral_cong \.flowD_outside[OF j] split: split_indicator) also have "\ = d_IN j (Some x)" using flowD_KIR[OF j, of "Some x"] True a disjoint by auto also have "?j_IN = (\\<^sup>+ y\range Some. j (y, Some x))" using True disjoint by(simp add: nn_integral_count_space_reindex)(auto 4 4 simp add: nn_integral_count_space_indicator incoming_def intro!: nn_integral_cong \.flowD_outside[OF j] dest: bipartite_E split: split_indicator) also have "\ = d_IN j (Some x) - (\\<^sup>+ y :: 'v option. j (None, Some x) * indicator {None} y)" unfolding d_IN_def using flowD_capacity[OF j, of "(None, Some x)"] by(subst nn_integral_diff[symmetric]) (auto simp add: nn_integral_count_space_indicator AE_count_space top_unique image_iff intro!: nn_integral_cong ennreal_diff_self split: split_indicator if_split_asm) also have "d_IN f x + d_IN j (Some x) - \ = d_IN f x + j (None, Some x)" using ennreal_add_diff_cancel_right[OF IN_j_finite[of "Some x"], of "d_IN f x + j (None, Some x)"] apply(subst diff_diff_ennreal') apply(auto simp add: d_IN_def intro!: nn_integral_ge_point ennreal_diff_le_mono_left) apply(simp add: ac_simps) done finally show ?thesis . next case False hence "d_IN g x = 0" "d_IN f x = 0" "j (None, \x\) = 0" using disjoint currentD_IN[OF f', of x] bipartite_V currentD_outside_IN[OF f'] u_outside[OF False] flowD_capacity[OF j, of "(None, \x\)"] by(cases "vertex \ x"; auto simp add: d_IN_def nn_integral_0_iff_AE AE_count_space g_simps dest: bipartite_E split: if_split_asm)+ thus ?thesis by simp qed have g: "current \ g" proof show "d_OUT g x \ weight \ x" for x proof(cases "x \ A \") case False thus ?thesis by(simp add: OUT_g_B) next case True with OUT_g_a show ?thesis by(cases "x = a")(simp_all add: OUT_g_A flowD_KIR[OF j] currentD_weight_OUT[OF f']) qed show "d_IN g x \ weight \ x" for x proof(cases "x \ B \") case False hence "d_IN g x = 0" using disjoint by(auto simp add: d_IN_def nn_integral_0_iff_AE AE_count_space g_simps dest: bipartite_E) thus ?thesis by simp next case True have "d_IN g x \ (weight \ x - u x) + u x" unfolding IN_g using currentD_weight_IN[OF f, of x] flowD_capacity[OF j, of "(None, Some x)"] True bipartite_V by(intro add_mono)(simp_all split: if_split_asm) also have "\ = weight \ x" using u by (intro diff_add_cancel_ennreal) (simp add: le_fun_def) finally show ?thesis . qed show "g e = 0" if "e \ \<^bold>E" for e using that by(cases e)(auto simp add: g_simps) qed define cap' where "cap' = (\(x, y). if edge \ x y then g (x, y) else if edge \ y x then 1 else 0)" have cap'_simps [simp]: "cap' (x, y) = (if edge \ x y then g (x, y) else if edge \ y x then 1 else 0)" for x y by(simp add: cap'_def) define G where "G = \edge = \x y. cap' (x, y) > 0\" have G_sel [simp]: "edge G x y \ cap' (x, y) > 0" for x y by(simp add: G_def) define reachable where "reachable x = (edge G)\<^sup>*\<^sup>* x a" for x have reachable_alt_def: "reachable \ \x. \p. path G x p a" by(simp add: reachable_def [abs_def] rtranclp_eq_rtrancl_path) have [simp]: "reachable a" by(auto simp add: reachable_def) have AB_edge: "edge G x y" if "edge \ y x" for x y using that by(auto dest: edge_antiparallel simp add: min_def le_neq_trans add_eq_0_iff_both_eq_0) have reachable_AB: "reachable y" if "reachable x" "(x, y) \ \<^bold>E" for x y using that by(auto simp add: reachable_def simp del: G_sel dest!: AB_edge intro: rtrancl_path.step) have reachable_BA: "g (x, y) = 0" if "reachable y" "(x, y) \ \<^bold>E" "\ reachable x" for x y proof(rule ccontr) assume "g (x, y) \ 0" then have "g (x, y) > 0" by (simp add: zero_less_iff_neq_zero) hence "edge G x y" using that by simp then have "reachable x" using \reachable y\ unfolding reachable_def by(rule converse_rtranclp_into_rtranclp) with \\ reachable x\ show False by contradiction qed have reachable_V: "vertex \ x" if "reachable x" for x proof - from that obtain p where p: "path G x p a" unfolding reachable_alt_def .. then show ?thesis using rtrancl_path_nth[OF p, of 0] a(1) A_vertex by(cases "p = []")(auto 4 3 simp add: vertex_def elim: rtrancl_path.cases split: if_split_asm) qed have finite_j_IN: "(\\<^sup>+ y. j (Some y, Some x) \count_space (\<^bold>I\<^bold>N x)) \ \" for x proof - have "(\\<^sup>+ y. j (Some y, Some x) \count_space (\<^bold>I\<^bold>N x)) \ d_IN f x" by(auto intro!: nn_integral_mono j_le_f simp add: d_IN_def nn_integral_count_space_indicator incoming_def split: split_indicator) thus ?thesis using currentD_finite_IN[OF f', of x] by (auto simp: top_unique) qed have j_outside: "j (x, y) = 0" if "\ edge \ x y" for x y using that flowD_capacity[OF j, of "(x, y)"] \.capacity_outside[of "(x, y)"] by(auto) define h where "h = (\(x, y). if reachable x \ reachable y then g (x, y) else 0)" have h_simps [simp]: "h (x, y) = (if reachable x \ reachable y then g (x, y) else 0)" for x y by(simp add: h_def) have h_le_g: "h e \ g e" for e by(cases e) simp have OUT_h: "d_OUT h x = (if reachable x then d_OUT g x else 0)" for x proof - have "d_OUT h x = (\\<^sup>+ y\\<^bold>O\<^bold>U\<^bold>T x. h (x, y))" using h_le_g currentD_outside[OF g] by(intro d_OUT_alt_def) auto also have "\ = (\\<^sup>+ y\\<^bold>O\<^bold>U\<^bold>T x. if reachable x then g (x, y) else 0)" by(auto intro!: nn_integral_cong simp add: outgoing_def dest: reachable_AB) also have "\ = (if reachable x then d_OUT g x else 0)" by(auto intro!: d_OUT_alt_def[symmetric] currentD_outside[OF g]) finally show ?thesis . qed have IN_h: "d_IN h x = (if reachable x then d_IN g x else 0)" for x proof - have "d_IN h x = (\\<^sup>+ y\\<^bold>I\<^bold>N x. h (y, x))" using h_le_g currentD_outside[OF g] by(intro d_IN_alt_def) auto also have "\ = (\\<^sup>+ y\\<^bold>I\<^bold>N x. if reachable x then g (y, x) else 0)" by(auto intro!: nn_integral_cong simp add: incoming_def dest: reachable_BA) also have "\ = (if reachable x then d_IN g x else 0)" by(auto intro!: d_IN_alt_def[symmetric] currentD_outside[OF g]) finally show ?thesis . qed have h: "current \ h" using g h_le_g proof(rule current_leI) show "d_OUT h x \ d_IN h x" if "x \ A \" for x by(simp add: OUT_h IN_h currentD_OUT_IN[OF g that]) qed have reachable_full: "j (None, \y\) = u y" if reach: "reachable y" for y proof(rule ccontr) assume "j (None, \y\) \ u y" with flowD_capacity[OF j, of "(None, \y\)"] have le: "j (None, \y\) < u y" by(auto split: if_split_asm simp add: u_outside \.flowD_outside[OF j] zero_less_iff_neq_zero) then obtain y: "y \ B \" and uy: "u y > 0" using u_outside[of y] by(cases "y \ B \"; cases "u y = 0") (auto simp add: zero_less_iff_neq_zero) from reach obtain q where q: "path G y q a" and distinct: "distinct (y # q)" unfolding reachable_alt_def by(auto elim: rtrancl_path_distinct) have q_Nil: "q \ []" using q a(1) disjoint y by(auto elim!: rtrancl_path.cases) let ?E = "zip (y # q) q" define E where "E = (None, Some y) # map (map_prod Some Some) ?E" define \ where "\ = Min (insert (u y - j (None, Some y)) (cap' ` set ?E))" let ?j' = "\e. (if e \ set E then \ else 0) + j e" define j' where "j' = cleanup ?j'" have j_free: "0 < cap' e" if "e \ set ?E" for e using that unfolding E_def list.sel proof - from that obtain i where e: "e = ((y # q) ! i, q ! i)" and i: "i < length q" by(auto simp add: set_zip) have e': "edge G ((y # q) ! i) (q ! i)" using q i by(rule rtrancl_path_nth) thus ?thesis using e by(simp) qed have \_pos: "0 < \" unfolding \_def using le by(auto intro: j_free diff_gr0_ennreal) have \_le: "\ \ cap' e" if "e \ set ?E" for e using that unfolding \_def by auto have finite_\ [simplified]: "\ < \" unfolding \_def by(intro Min_less_iff[THEN iffD2])(auto simp add: less_top[symmetric]) have E_antiparallel: "(x', y') \ set ?E \ (y', x') \ set ?E" for x' y' using distinct apply(auto simp add: in_set_zip nth_Cons in_set_conv_nth) apply(auto simp add: distinct_conv_nth split: nat.split_asm) by (metis Suc_lessD less_Suc_eq less_irrefl_nat) have OUT_j': "d_OUT ?j' x' = \ * card (set [(x'', y) \ E. x'' = x']) + d_OUT j x'" for x' proof - have "d_OUT ?j' x' = d_OUT (\e. if e \ set E then \ else 0) x' + d_OUT j x'" using \_pos by(intro d_OUT_add) also have "d_OUT (\e. if e \ set E then \ else 0) x' = \\<^sup>+ y. \ * indicator (set E) (x', y) \count_space UNIV" unfolding d_OUT_def by(rule nn_integral_cong)(simp) also have "\ = (\\<^sup>+ e. \ * indicator (set E) e \embed_measure (count_space UNIV) (Pair x'))" by(simp add: measurable_embed_measure1 nn_integral_embed_measure) also have "\ = (\\<^sup>+ e. \ * indicator (set [(x'', y) \ E. x'' = x']) e \count_space UNIV)" by(auto simp add: embed_measure_count_space' nn_integral_count_space_indicator intro!: nn_integral_cong split: split_indicator) also have "\ = \ * card (set [(x'', y) \ E. x'' = x'])" using \_pos by(simp add: nn_integral_cmult) finally show ?thesis . qed have IN_j': "d_IN ?j' x' = \ * card (set [(y, x'') \ E. x'' = x']) + d_IN j x'" for x' proof - have "d_IN ?j' x' = d_IN (\e. if e \ set E then \ else 0) x' + d_IN j x'" using \_pos by(intro d_IN_add) also have "d_IN (\e. if e \ set E then \ else 0) x' = \\<^sup>+ y. \ * indicator (set E) (y, x') \count_space UNIV" unfolding d_IN_def by(rule nn_integral_cong)(simp) also have "\ = (\\<^sup>+ e. \ * indicator (set E) e \embed_measure (count_space UNIV) (\y. (y, x')))" by(simp add: measurable_embed_measure1 nn_integral_embed_measure) also have "\ = (\\<^sup>+ e. \ * indicator (set [(y, x'') \ E. x'' = x']) e \count_space UNIV)" by(auto simp add: embed_measure_count_space' nn_integral_count_space_indicator intro!: nn_integral_cong split: split_indicator) also have "\ = \ * card (set [(y, x'') \ E. x'' = x'])" using \_pos by(auto simp add: nn_integral_cmult) finally show ?thesis . qed have j': "flow \ j'" proof fix e :: "'v option edge" consider (None) "e = (None, Some y)" | (Some) x y where "e = (Some x, Some y)" "(x, y) \ set ?E" | (old) x y where "e = (Some x, Some y)" "(x, y) \ set ?E" | y' where "e = (None, Some y')" "y \ y'" | "e = (None, None)" | x where "e = (Some x, None)" by(cases e; case_tac a; case_tac b)(auto) then show "j' e \ capacity \ e" using uy \_pos flowD_capacity[OF j, of e] proof(cases) case None have "\ \ u y - j (None, Some y)" by(simp add: \_def) then have "\ + j (None, Some y) \ u y" using \_pos by (auto simp add: ennreal_le_minus_iff) thus ?thesis using reachable_V[OF reach] None \.flowD_outside[OF j, of "(Some y, None)"] uy by(auto simp add: j'_def E_def) next case (Some x' y') have e: "\ \ cap' (x', y')" using Some(2) by(rule \_le) then consider (backward) "edge \ x' y'" "x' \ a" | (forward) "edge \ y' x'" "\ edge \ x' y'" | (a') "edge \ x' y'" "x' = a" using Some \_pos by(auto split: if_split_asm) then show ?thesis proof cases case backward have "\ \ f (x', y') + j (Some y', Some x') - j (Some x', Some y')" using e backward Some(1) by(simp add: g_simps) hence "\ + j (Some x', Some y') - j (Some y', Some x') \ (f (x', y') + j (Some y', Some x') - j (Some x', Some y')) + j (Some x', Some y') - j (Some y', Some x')" by(intro ennreal_minus_mono add_right_mono) simp_all also have "\ = f (x', y')" using j_le_f[OF \edge \ x' y'\] by(simp_all add: add_increasing2 less_top diff_add_assoc2_ennreal) finally show ?thesis using Some backward by(auto simp add: j'_def E_def dest: in_set_tlD E_antiparallel) next case forward have "\ + j (Some x', Some y') - j (Some y', Some x') \ \ + j (Some x', Some y')" by(rule diff_le_self_ennreal) also have "j (Some x', Some y') \ d_IN j (Some y')" by(rule d_IN_ge_point) also have "\ \ weight \ y'" by(rule IN_j_le_weight) also have "\ \ 1" using e forward by simp finally have "\ + j (Some x', Some y') - j (Some y', Some x') \ max (weight \ x') (weight \ y') + 1" by(simp add: add_left_mono add_right_mono max_def)(metis (no_types, lifting) add.commute add_right_mono less_imp_le less_le_trans not_le) then show ?thesis using Some forward e by(auto simp add: j'_def E_def max_def dest: in_set_tlD E_antiparallel) next case a' with Some have "a \ set (map fst (zip (y # q) q))" by(auto intro: rev_image_eqI) also have "map fst (zip (y # q) q) = butlast (y # q)" by(induction q arbitrary: y) auto finally have False using rtrancl_path_last[OF q q_Nil] distinct q_Nil by(cases q rule: rev_cases) auto then show ?thesis .. qed next case (old x' y') hence "j' e \ j e" using \_pos by(auto simp add: j'_def E_def intro!: diff_le_self_ennreal) also have "j e \ capacity \ e" using j by(rule flowD_capacity) finally show ?thesis . qed(auto simp add: j'_def E_def \.flowD_outside[OF j] uy) next fix x' assume x': "x' \ source \" "x' \ sink \" then obtain x'' where x'': "x' = Some x''" by auto have "d_OUT ?j' x' = \ * card (set [(x'', y) \ E. x'' = x']) + d_OUT j x'" by(rule OUT_j') also have "card (set [(x'', y) \ E. x'' = x']) = card (set [(y, x'') \ E. x'' = x'])" (is "?lhs = ?rhs") proof - have "?lhs = length [(x'', y) \ E. x'' = x']" using distinct by(subst distinct_card)(auto simp add: E_def filter_map distinct_map inj_map_prod' distinct_zipI1) also have "\ = length [x''' \ map fst ?E. x''' = x'']" by(simp add: E_def x'' split_beta cong: filter_cong) also have "map fst ?E = butlast (y # q)" by(induction q arbitrary: y) simp_all also have "[x''' \ butlast (y # q). x''' = x''] = [x''' \ y # q. x''' = x'']" using q_Nil rtrancl_path_last[OF q q_Nil] x' x'' by(cases q rule: rev_cases) simp_all also have "q = map snd ?E" by(induction q arbitrary: y) auto also have "length [x''' \ y # \. x''' = x''] = length [x'' \ map snd E. x'' = x']" using x'' by(simp add: E_def cong: filter_cong) also have "\ = length [(y, x'') \ E. x'' = x']" by(simp cong: filter_cong add: split_beta) also have "\ = ?rhs" using distinct by(subst distinct_card)(auto simp add: E_def filter_map distinct_map inj_map_prod' distinct_zipI1) finally show ?thesis . qed also have "\ * \ + d_OUT j x' = d_IN ?j' x'" unfolding flowD_KIR[OF j x'] by(rule IN_j'[symmetric]) also have "d_IN ?j' x' \ \" using \.flowD_finite_IN[OF j x'(2)] finite_\ IN_j'[of x'] by (auto simp: top_add ennreal_mult_eq_top_iff) ultimately show "KIR j' x'" unfolding j'_def by(rule KIR_cleanup) qed hence "value_flow \ j' \ \" unfolding \_def by(auto intro: SUP_upper) moreover have "value_flow \ j' > value_flow \ j" proof - have "value_flow \ j + 0 < value_flow \ j + \ * 1" using \_pos value_j finite_flow by simp also have "[(x', y') \ E. x' = None] = [(None, Some y)]" using q_Nil by(cases q)(auto simp add: E_def filter_map cong: filter_cong split_beta) hence "\ * 1 \ \ * card (set [(x', y') \ E. x' = None])" using \_pos by(intro mult_left_mono)(auto simp add: E_def real_of_nat_ge_one_iff neq_Nil_conv card.insert_remove) also have "value_flow \ j + \ = value_flow \ ?j'" using OUT_j' by(simp add: add.commute) also have "\ = value_flow \ j'" unfolding j'_def by(subst value_flow_cleanup)(auto simp add: E_def \.flowD_outside[OF j]) finally show ?thesis by(simp add: add_left_mono) qed ultimately show False using finite_flow \_pos value_j by(cases "value_flow \ j" \ rule: ennreal2_cases) simp_all qed have sep_h: "y \ TER h" if reach: "reachable y" and y: "y \ B \" and TER: "y \ ?TER f" for y proof(rule ccontr) assume y': "y \ TER h" from y a(1) disjoint have yna: "y \ a" by auto from reach obtain p' where "path G y p' a" unfolding reachable_alt_def .. then obtain p' where p': "path G y p' a" and distinct: "distinct (y # p')" by(rule rtrancl_path_distinct) have SINK: "y \ SINK h" using y disjoint by(auto simp add: SINK.simps d_OUT_def nn_integral_0_iff emeasure_count_space_eq_0 intro: currentD_outside[OF g] dest: bipartite_E) have hg: "d_IN h y = d_IN g y" using reach by(simp add: IN_h) also have "\ = d_IN f y + j (None, Some y)" by(simp add: IN_g) also have "d_IN f y = weight \ y - u y" using currentD_weight_IN[OF f, of y] y disjoint TER by(auto elim!: SAT.cases) also have "d_IN h y < weight \ y" using y' currentD_weight_IN[OF g, of y] y disjoint SINK by(auto intro: SAT.intros) ultimately have le: "j (None, Some y) < u y" by(cases "weight \ y" "u y" "j (None, Some y)" rule: ennreal3_cases; cases "u y \ weight \ y") (auto simp: ennreal_minus ennreal_plus[symmetric] add_top ennreal_less_iff ennreal_neg simp del: ennreal_plus) moreover from reach have "j (None, \y\) = u y" by(rule reachable_full) ultimately show False by simp qed have w': "wave \ h" proof show sep: "separating \ (TER h)" proof(rule ccontr) assume "\ ?thesis" then obtain x p y where x: "x \ A \" and y: "y \ B \" and p: "path \ x p y" and x': "x \ TER h" and bypass: "\z. z \ set p \ z \ TER h" by(auto simp add: separating_gen.simps) from p disjoint x y have p_eq: "p = [y]" and edge: "(x, y) \ \<^bold>E" by -(erule rtrancl_path.cases, auto dest: bipartite_E)+ from p_eq bypass have y': "y \ TER h" by simp have "reachable x" using x' by(rule contrapos_np)(simp add: SINK.simps d_OUT_def SAT.A x) hence reach: "reachable y" using edge by(rule reachable_AB) have *: "x \ \\<^bsub>?\\<^esub> (?TER f)" using x' proof(rule contrapos_nn) assume *: "x \ \\<^bsub>?\\<^esub> (?TER f)" have "d_OUT h x \ d_OUT g x" using h_le_g by(rule d_OUT_mono) also from * have "x \ a" using a by auto then have "d_OUT j (Some x) = d_IN j (Some x)" by(auto intro: flowD_KIR[OF j]) hence "d_OUT g x \ d_OUT f x" using OUT_g_A[OF x] IN_j[of "Some x"] finite_flow by(auto split: if_split_asm) also have "\ = 0" using * by(auto elim: SINK.cases) finally have "x \ SINK h" by(simp add: SINK.simps) with x show "x \ TER h" by(simp add: SAT.A) qed from p p_eq x y have "path ?\ x [y] y" "x \ A ?\" "y \ B ?\" by simp_all from * separatingD[OF separating_essential, OF waveD_separating, OF w this] have "y \ ?TER f" by auto with reach y have "y \ TER h" by(rule sep_h) with y' show False by contradiction qed qed(rule h) have OUT_g_a: "d_OUT g a = d_OUT h a" by(simp add: OUT_h) have "a \ \ (TER h)" proof assume *: "a \ \ (TER h)" have "j (Some a, Some y) = 0" for y using flowD_capacity[OF j, of "(Some a, Some y)"] a(1) disjoint by(auto split: if_split_asm dest: bipartite_E) then have "d_OUT f a \ d_OUT g a" unfolding d_OUT_def \ \This step requires that @{term j} does not decrease the outflow of @{term a}. That's why we set the capacity of the outgoing edges from @{term "Some a"} in @{term \} to @{term "0 :: ennreal"}\ by(intro nn_integral_mono)(auto simp add: g_simps currentD_outside[OF f] intro: ) then have "a \ SINK f" using OUT_g_a * by(simp add: SINK.simps) with a(1) have "a \ ?TER f" by(auto intro: SAT.A) with a(2) have a': "\ essential \ (B \) (?TER f) a" by simp from * obtain y where ay: "edge \ a y" and y: "y \ B \" and y': "y \ TER h" using disjoint a(1) by(auto 4 4 simp add: essential_def elim: rtrancl_path.cases dest: bipartite_E) from not_essentialD[OF a' rtrancl_path.step, OF ay rtrancl_path.base y] have TER: "y \ ?TER f" by simp have "reachable y" using \reachable a\ by(rule reachable_AB)(simp add: ay) hence "y \ TER h" using y TER by(rule sep_h) with y' show False by contradiction qed with \a \ A \\ have "hindrance \ h" proof have "d_OUT h a = d_OUT g a" by(simp add: OUT_g_a) also have "\ \ d_OUT f a + \\<^sup>+ y. j (Some y, Some a) \count_space UNIV" unfolding d_OUT_def d_IN_def by(subst nn_integral_add[symmetric])(auto simp add: g_simps intro!: nn_integral_mono diff_le_self_ennreal) also have "(\\<^sup>+ y. j (Some y, Some a) \count_space UNIV) = (\\<^sup>+ y. j (y, Some a) \embed_measure (count_space UNIV) Some)" by(simp add: nn_integral_embed_measure measurable_embed_measure1) also have "\ \ d_IN j (Some a)" unfolding d_IN_def by(auto simp add: embed_measure_count_space nn_integral_count_space_indicator intro!: nn_integral_mono split: split_indicator) also have "\ \ \" by(rule IN_j) also have "\ \ \" by(rule \_le) also have "d_OUT f a + \ < d_OUT f a + (weight \ a - d_OUT f a)" using \_less using currentD_finite_OUT[OF f'] by (simp add: ennreal_add_left_cancel_less) also have "\ = weight \ a" using a_le by simp finally show "d_OUT h a < weight \ a" by(simp add: add_left_mono) qed then show ?thesis using h w' by(blast intro: hindered.intros) qed end corollary hindered_reduce_current: \ \Corollary 6.8\ fixes \ g defines "\ \ \\<^sup>+ x\B \. d_IN g x - d_OUT g x" assumes g: "current \ g" and \_finite: "\ \ \" and hindered: "hindered_by (\ \ g) \" shows "hindered \" proof - define \' where "\' = \\weight := \x. if x \ A \ then weight \ x - d_OUT g x else weight \ x\" have \'_sel [simp]: "edge \' = edge \" "A \' = A \" "B \' = B \" "weight \' x = (if x \ A \ then weight \ x - d_OUT g x else weight \ x)" "vertex \' = vertex \" "web.more \' = web.more \" for x by(simp_all add: \'_def) have "countable_bipartite_web \'" by unfold_locales(simp_all add: A_in B_out A_vertex disjoint bipartite_V no_loop weight_outside currentD_outside_OUT[OF g] currentD_weight_OUT[OF g] edge_antiparallel, rule bipartite_E) then interpret \': countable_bipartite_web \' . let ?u = "\x. (d_IN g x - d_OUT g x) * indicator (- A \) x" have "hindered \'" proof(rule \'.hindered_reduce) show "?u x = 0" if "x \ B \'" for x using that bipartite_V by(cases "vertex \' x")(auto simp add: currentD_outside_OUT[OF g] currentD_outside_IN[OF g]) have *: "(\\<^sup>+ x\B \'. ?u x) = \" using disjoint by(auto intro!: nn_integral_cong simp add: \_def nn_integral_count_space_indicator currentD_outside_OUT[OF g] currentD_outside_IN[OF g] not_vertex split: split_indicator) thus "(\\<^sup>+ x\B \'. ?u x) \ \" using \_finite by simp have **: "\'\weight := weight \' - ?u\ = \ \ g" using currentD_weight_IN[OF g] currentD_OUT_IN[OF g] currentD_IN[OF g] currentD_finite_OUT[OF g] by(intro web.equality)(simp_all add: fun_eq_iff diff_diff_ennreal' ennreal_diff_le_mono_left) show "hindered_by (\'\weight := weight \' - ?u\) (\\<^sup>+ x\B \'. ?u x)" unfolding * ** by(fact hindered) show "(\x. (d_IN g x - d_OUT g x) * indicator (- A \) x) \ weight \'" using currentD_weight_IN[OF g] by (simp add: le_fun_def ennreal_diff_le_mono_left) qed then show ?thesis by(rule hindered_mono_web[rotated -1]) simp_all qed end subsection \Reduced weight in a loose web\ definition reduce_weight :: "('v, 'more) web_scheme \ 'v \ real \ ('v, 'more) web_scheme" where "reduce_weight \ x r = \\weight := \y. weight \ y - (if x = y then r else 0)\" lemma reduce_weight_sel [simp]: "edge (reduce_weight \ x r) = edge \" "A (reduce_weight \ x r) = A \" "B (reduce_weight \ x r) = B \" "vertex (reduce_weight \ x r) = vertex \" "weight (reduce_weight \ x r) y = (if x = y then weight \ x - r else weight \ y)" "web.more (reduce_weight \ x r) = web.more \" by(simp_all add: reduce_weight_def zero_ennreal_def[symmetric] vertex_def fun_eq_iff) lemma essential_reduce_weight [simp]: "essential (reduce_weight \ x r) = essential \" by(simp add: fun_eq_iff essential_def) lemma roofed_reduce_weight [simp]: "roofed_gen (reduce_weight \ x r) = roofed_gen \" by(simp add: fun_eq_iff roofed_def) context countable_bipartite_web begin context begin private datatype (plugins del: transfer size) 'a vertex = SOURCE | SINK | Inner (inner: 'a) private lemma notin_range_Inner: "x \ range Inner \ x = SOURCE \ x = SINK" by(cases x) auto private lemma inj_Inner [simp]: "\A. inj_on Inner A" by(simp add: inj_on_def) lemma unhinder_bipartite: assumes h: "\n :: nat. current \ (h n)" and SAT: "\n. (B \ \ \<^bold>V) - {b} \ SAT \ (h n)" and b: "b \ B \" and IN: "(SUP n. d_IN (h n) b) = weight \ b" and h0_b: "\n. d_IN (h 0) b \ d_IN (h n) b" and b_V: "b \ \<^bold>V" shows "\h'. current \ h' \ wave \ h' \ B \ \ \<^bold>V \ SAT \ h'" proof - write Inner ("\_\") define edge' where "edge' xo yo = (case (xo, yo) of (\x\, \y\) \ edge \ x y \ edge \ y x | (\x\, SINK) \ x \ A \ | (SOURCE, \y\) \ y = b | (SINK, \x\) \ x \ A \ | _ \ False)" for xo yo have edge'_simps [simp]: "edge' \x\ \y\ \ edge \ x y \ edge \ y x" "edge' \x\ SINK \ x \ A \" "edge' SOURCE yo \ yo = \b\" "edge' SINK \x\ \ x \ A \" "edge' SINK SINK \ False" "edge' xo SOURCE \ False" for x y yo xo by(simp_all add: edge'_def split: vertex.split) have edge'E: "thesis" if "edge' xo yo" "\x y. \ xo = \x\; yo = \y\; edge \ x y \ edge \ y x \ \ thesis" "\x. \ xo = \x\; yo = SINK; x \ A \ \ \ thesis" "\x. \ xo = SOURCE; yo = \b\ \ \ thesis" "\y. \ xo = SINK; yo = \y\; y \ A \ \ \ thesis" for xo yo thesis using that by(auto simp add: edge'_def split: option.split_asm vertex.split_asm) have edge'_Inner1 [elim!]: "thesis" if "edge' \x\ yo" "\y. \ yo = \y\; edge \ x y \ edge \ y x \ \ thesis" "\ yo = SINK; x \ A \ \ \ thesis" for x yo thesis using that by(auto elim: edge'E) have edge'_Inner2 [elim!]: "thesis" if "edge' xo \y\" "\x. \ xo = \x\; edge \ x y \ edge \ y x \ \ thesis" "\ xo = SOURCE; y = b \ \ thesis" "\ xo = SINK; y \ A \ \ \ thesis" for xo y thesis using that by(auto elim: edge'E) have edge'_SINK1 [elim!]: "thesis" if "edge' SINK yo" "\y. \ yo = \y\; y \ A \ \ \ thesis" for yo thesis using that by(auto elim: edge'E) have edge'_SINK2 [elim!]: "thesis" if "edge' xo SINK" "\x. \ xo = \x\; x \ A \ \ \ thesis" for xo thesis using that by(auto elim: edge'E) define cap where "cap xoyo = (case xoyo of (\x\, \y\) \ if edge \ x y then h 0 (x, y) else if edge \ y x then max (weight \ x) (weight \ y) else 0 | (\x\, SINK) \ if x \ A \ then weight \ x - d_OUT (h 0) x else 0 | (SOURCE, yo) \ if yo = \b\ then weight \ b - d_IN (h 0) b else 0 | (SINK, \y\) \ if y \ A \ then weight \ y else 0 | _ \ 0)" for xoyo have cap_simps [simp]: "cap (\x\, \y\) = (if edge \ x y then h 0 (x, y) else if edge \ y x then max (weight \ x) (weight \ y) else 0)" "cap (\x\, SINK) = (if x \ A \ then weight \ x - d_OUT (h 0) x else 0)" "cap (SOURCE, yo) = (if yo = \b\ then weight \ b - d_IN (h 0) b else 0)" "cap (SINK, \y\) = (if y \ A \ then weight \ y else 0)" "cap (SINK, SINK) = 0" "cap (xo, SOURCE) = 0" for x y yo xo by(simp_all add: cap_def split: vertex.split) define \ where "\ = \edge = edge', capacity = cap, source = SOURCE, sink = SINK\" have \_sel [simp]: "edge \ = edge'" "capacity \ = cap" "source \ = SOURCE" "sink \ = SINK" by(simp_all add: \_def) have cap_outside1: "\ vertex \ x \ cap (\x\, y) = 0" for x y using A_vertex by(cases y)(auto simp add: vertex_def) have capacity_A_weight: "d_OUT cap \x\ \ 2 * weight \ x" if "x \ A \" for x proof - have "d_OUT cap \x\ \ (\\<^sup>+ y. h 0 (x, inner y) * indicator (range Inner) y + weight \ x * indicator {SINK} y)" using that disjoint unfolding d_OUT_def by(auto intro!: nn_integral_mono diff_le_self_ennreal simp add: A_in notin_range_Inner split: split_indicator) also have "\ = (\\<^sup>+ y\range Inner. h 0 (x, inner y)) + weight \ x" by(auto simp add: nn_integral_count_space_indicator nn_integral_add) also have "(\\<^sup>+ y\range Inner. h 0 (x, inner y)) = d_OUT (h 0) x" by(simp add: d_OUT_def nn_integral_count_space_reindex) also have "\ \ weight \ x" using h by(rule currentD_weight_OUT) finally show ?thesis unfolding one_add_one[symmetric] distrib_right by(simp add: add_right_mono) qed have flow_attainability: "flow_attainability \" proof have "\<^bold>E\<^bsub>\\<^esub> \ (\(x, y). (\x\, \y\)) ` \<^bold>E \ (\(x, y). (\y\, \x\)) ` \<^bold>E \ (\x. (\x\, SINK)) ` A \ \ (\x. (SINK, \x\)) ` A \ \ {(SOURCE, \b\)}" by(auto simp add: edge'_def split: vertex.split_asm) moreover have "countable (A \)" using A_vertex by(rule countable_subset) simp ultimately show "countable \<^bold>E\<^bsub>\\<^esub>" by(auto elim: countable_subset) next fix v assume "v \ sink \" then consider (source) "v = SOURCE" | (A) x where "v = \x\" "x \ A \" | (B) y where "v = \y\" "y \ A \" "y \ \<^bold>V" | (outside) x where "v = \x\" "x \ \<^bold>V" by(cases v) auto then show "d_IN (capacity \) v \ \ \ d_OUT (capacity \) v \ \" proof cases case source thus ?thesis by(simp add: d_IN_def) next case (A x) thus ?thesis using capacity_A_weight[of x] by (auto simp: top_unique ennreal_mult_eq_top_iff) next case (B y) have "d_IN (capacity \) v \ (\\<^sup>+ x. h 0 (inner x, y) * indicator (range Inner) x + weight \ b * indicator {SOURCE} x)" using B bipartite_V by(auto 4 4 intro!: nn_integral_mono simp add: diff_le_self_ennreal d_IN_def notin_range_Inner nn_integral_count_space_indicator currentD_outside[OF h] split: split_indicator dest: bipartite_E) also have "\ = (\\<^sup>+ x\range Inner. h 0 (inner x, y)) + weight \ b" by(simp add: nn_integral_add nn_integral_count_space_indicator) also have "(\\<^sup>+ x\range Inner. h 0 (inner x, y)) = d_IN (h 0) y" by(simp add: d_IN_def nn_integral_count_space_reindex) also have "d_IN (h 0) y \ weight \ y" using h by(rule currentD_weight_IN) finally show ?thesis by(auto simp add: top_unique add_right_mono split: if_split_asm) next case outside hence "d_OUT (capacity \) v = 0" using A_vertex by(auto simp add: d_OUT_def nn_integral_0_iff_AE AE_count_space cap_def vertex_def split: vertex.split) thus ?thesis by simp qed next show "capacity \ e \ \" for e by(auto simp add: cap_def max_def vertex_def currentD_finite[OF h] split: vertex.split prod.split) show "capacity \ e = 0" if "e \ \<^bold>E\<^bsub>\\<^esub>" for e using that by(auto simp add: cap_def max_def split: prod.split; split vertex.split)+ show "\ edge \ x (source \)" for x using b by(auto simp add: B_out) show "\ edge \ x x" for x by(cases x)(simp_all add: no_loop) show "source \ \ sink \" by simp qed then interpret \: flow_attainability "\" . define \ where "\ = (SUP f\{f. flow \ f}. value_flow \ f)" define f where "f n xoyo = (case xoyo of (\x\, \y\) \ if edge \ x y then h 0 (x, y) - h n (x, y) else if edge \ y x then h n (y, x) - h 0 (y, x) else 0 | (SOURCE, \y\) \ if y = b then d_IN (h n) b - d_IN (h 0) b else 0 | (\x\, SINK) \ if x \ A \ then d_OUT (h n) x - d_OUT (h 0) x else 0 | (SINK, \y\) \ if y \ A \ then d_OUT (h 0) y - d_OUT (h n) y else 0 | _ \ 0)" for n xoyo have f_cases: thesis if "\x y. e = (\x\, \y\) \ thesis" "\y. e = (SOURCE, \y\) \ thesis" "\x. e = (\x\, SINK) \ thesis" "\y. e = (SINK, \y\) \ thesis" "e = (SINK, SINK) \ thesis" "\xo. e = (xo, SOURCE) \ thesis" "e = (SOURCE, SINK) \ thesis" for e :: "'v vertex edge" and thesis using that by(cases e; cases "fst e" "snd e" rule: vertex.exhaust[case_product vertex.exhaust]) simp_all have f_simps [simp]: "f n (\x\, \y\) = (if edge \ x y then h 0 (x, y) - h n (x, y) else if edge \ y x then h n (y, x) - h 0 (y, x) else 0)" "f n (SOURCE, \y\) = (if y = b then d_IN (h n) b - d_IN (h 0) b else 0)" "f n (\x\, SINK) = (if x \ A \ then d_OUT (h n) x - d_OUT (h 0) x else 0)" "f n (SINK, \y\) = (if y \ A \ then d_OUT (h 0) y - d_OUT (h n) y else 0)" "f n (SOURCE, SINK) = 0" "f n (SINK, SINK) = 0" "f n (xo, SOURCE) = 0" for n x y xo by(simp_all add: f_def split: vertex.split) have OUT_f_SOURCE: "d_OUT (f n) SOURCE = d_IN (h n) b - d_IN (h 0) b" for n proof(rule trans) show "d_OUT (f n) SOURCE = (\\<^sup>+ y. f n (SOURCE, y) * indicator {\b\} y)" unfolding d_OUT_def apply(rule nn_integral_cong) subgoal for x by(cases x) auto done show "\ = d_IN (h n) b - d_IN (h 0) b" using h0_b[of n] by(auto simp add: max_def) qed have OUT_f_outside: "d_OUT (f n) \x\ = 0" if "x \ \<^bold>V" for x n using A_vertex that apply(clarsimp simp add: d_OUT_def nn_integral_0_iff emeasure_count_space_eq_0) subgoal for y by(cases y)(auto simp add: vertex_def) done have IN_f_outside: "d_IN (f n) \x\ = 0" if "x \ \<^bold>V" for x n using b_V that apply(clarsimp simp add: d_IN_def nn_integral_0_iff emeasure_count_space_eq_0) subgoal for y by(cases y)(auto simp add: currentD_outside_OUT[OF h] vertex_def) done have f: "flow \ (f n)" for n proof show f_le: "f n e \ capacity \ e" for e using currentD_weight_out[OF h] currentD_weight_IN[OF h] currentD_weight_OUT[OF h] by(cases e rule: f_cases) (auto dest: edge_antiparallel simp add: not_le le_max_iff_disj intro: ennreal_minus_mono ennreal_diff_le_mono_left) fix xo assume "xo \ source \" "xo \ sink \" then consider (A) x where "xo = \x\" "x \ A \" | (B) x where "xo = \x\" "x \ B \" "x \ \<^bold>V" | (outside) x where "xo = \x\" "x \ \<^bold>V" using bipartite_V by(cases xo) auto then show "KIR (f n) xo" proof cases case outside thus ?thesis by(simp add: OUT_f_outside IN_f_outside) next case A have finite1: "(\\<^sup>+ y. h n (x, y) * indicator A y) \ \" for A n using currentD_finite_OUT[OF h, of n x, unfolded d_OUT_def] by(rule neq_top_trans)(auto intro!: nn_integral_mono simp add: split: split_indicator) let ?h0_ge_hn = "{y. h 0 (x, y) \ h n (x, y)}" let ?h0_lt_hn = "{y. h 0 (x, y) < h n (x, y)}" have "d_OUT (f n) \x\ = (\\<^sup>+ y. f n (\x\, y) * indicator (range Inner) y + f n (\x\, y) * indicator {SINK} y)" unfolding d_OUT_def by(intro nn_integral_cong)(auto split: split_indicator simp add: notin_range_Inner) also have "\ = (\\<^sup>+ y\range Inner. f n (\x\, y)) + f n (\x\, SINK)" by(simp add: nn_integral_add nn_integral_count_space_indicator max.left_commute max.commute) also have "(\\<^sup>+ y\range Inner. f n (\x\, y)) = (\\<^sup>+ y. h 0 (x, y) - h n (x, y))" using A apply(simp add: nn_integral_count_space_reindex cong: nn_integral_cong_simp outgoing_def) apply(auto simp add: nn_integral_count_space_indicator outgoing_def A_in max.absorb1 currentD_outside[OF h] intro!: nn_integral_cong split: split_indicator dest: edge_antiparallel) done also have "\ = (\\<^sup>+ y. h 0 (x, y) * indicator ?h0_ge_hn y) - (\\<^sup>+ y. h n (x, y) * indicator ?h0_ge_hn y)" apply(subst nn_integral_diff[symmetric]) apply(simp_all add: AE_count_space finite1 split: split_indicator) apply(rule nn_integral_cong; auto simp add: max_def not_le split: split_indicator) by (metis diff_eq_0_ennreal le_less not_le top_greatest) also have "(\\<^sup>+ y. h n (x, y) * indicator ?h0_ge_hn y) = d_OUT (h n) x - (\\<^sup>+ y. h n (x, y) * indicator ?h0_lt_hn y)" unfolding d_OUT_def apply(subst nn_integral_diff[symmetric]) apply(auto simp add: AE_count_space finite1 currentD_finite[OF h] split: split_indicator intro!: nn_integral_cong) done also have "(\\<^sup>+ y. h 0 (x, y) * indicator ?h0_ge_hn y) - \ + f n (\x\, SINK) = (\\<^sup>+ y. h 0 (x, y) * indicator ?h0_ge_hn y) + (\\<^sup>+ y. h n (x, y) * indicator ?h0_lt_hn y) - min (d_OUT (h n) x) (d_OUT (h 0) x)" using finite1[of n "{_}"] A finite1[of n UNIV] apply (subst diff_diff_ennreal') apply (auto simp: d_OUT_def finite1 AE_count_space nn_integral_diff[symmetric] top_unique nn_integral_add[symmetric] split: split_indicator intro!: nn_integral_mono ennreal_diff_self) apply (simp add: min_def not_le diff_eq_0_ennreal finite1 less_top[symmetric]) apply (subst diff_add_assoc2_ennreal) apply (auto simp: add_diff_eq_ennreal intro!: nn_integral_mono split: split_indicator) apply (subst diff_diff_commute_ennreal) apply (simp add: ennreal_add_diff_cancel ) done also have "\ = (\\<^sup>+ y. h n (x, y) * indicator ?h0_lt_hn y) - (d_OUT (h 0) x - (\\<^sup>+ y. h 0 (x, y) * indicator ?h0_ge_hn y)) + f n (SINK, \x\)" apply(rule sym) using finite1[of 0 "{_}"] A finite1[of 0 UNIV] apply (subst diff_diff_ennreal') apply (auto simp: d_OUT_def finite1 AE_count_space nn_integral_diff[symmetric] top_unique nn_integral_add[symmetric] split: split_indicator intro!: nn_integral_mono ennreal_diff_self) apply (simp add: min_def not_le diff_eq_0_ennreal finite1 less_top[symmetric]) apply (subst diff_add_assoc2_ennreal) apply (auto simp: add_diff_eq_ennreal intro!: nn_integral_mono split: split_indicator) apply (subst diff_diff_commute_ennreal) apply (simp_all add: ennreal_add_diff_cancel ac_simps) done also have "d_OUT (h 0) x - (\\<^sup>+ y. h 0 (x, y) * indicator ?h0_ge_hn y) = (\\<^sup>+ y. h 0 (x, y) * indicator ?h0_lt_hn y)" unfolding d_OUT_def apply(subst nn_integral_diff[symmetric]) apply(auto simp add: AE_count_space finite1 currentD_finite[OF h] split: split_indicator intro!: nn_integral_cong) done also have "(\\<^sup>+ y. h n (x, y) * indicator ?h0_lt_hn y) - \ = (\\<^sup>+ y. h n (x, y) - h 0 (x, y))" apply(subst nn_integral_diff[symmetric]) apply(simp_all add: AE_count_space finite1 order.strict_implies_order split: split_indicator) apply(rule nn_integral_cong; auto simp add: currentD_finite[OF h] top_unique less_top[symmetric] not_less split: split_indicator intro!: diff_eq_0_ennreal) done also have "\ = (\\<^sup>+ y\range Inner. f n (y, \x\))" using A apply(simp add: nn_integral_count_space_reindex cong: nn_integral_cong_simp outgoing_def) apply(auto simp add: nn_integral_count_space_indicator outgoing_def A_in max.commute currentD_outside[OF h] intro!: nn_integral_cong split: split_indicator dest: edge_antiparallel) done also have "\ + f n (SINK, \x\) = (\\<^sup>+ y. f n (y, \x\) * indicator (range Inner) y + f n (y, \x\) * indicator {SINK} y)" by(simp add: nn_integral_add nn_integral_count_space_indicator) also have "\ = d_IN (f n) \x\" using A b disjoint unfolding d_IN_def by(intro nn_integral_cong)(auto split: split_indicator simp add: notin_range_Inner) finally show ?thesis using A by simp next case (B x) have finite1: "(\\<^sup>+ y. h n (y, x) * indicator A y) \ \" for A n using currentD_finite_IN[OF h, of n x, unfolded d_IN_def] by(rule neq_top_trans)(auto intro!: nn_integral_mono split: split_indicator) have finite_h[simp]: "h n (y, x) < \" for y n using finite1[of n "{y}"] by (simp add: less_top) let ?h0_gt_hn = "{y. h 0 (y, x) > h n (y, x)}" let ?h0_le_hn = "{y. h 0 (y, x) \ h n (y, x)}" have eq: "d_IN (h 0) x + f n (SOURCE, \x\) = d_IN (h n) x" proof(cases "x = b") case True with currentD_finite_IN[OF h, of _ b] show ?thesis by(simp add: add_diff_self_ennreal h0_b) next case False with B SAT have "x \ SAT \ (h n)" "x \ SAT \ (h 0)" by auto with B disjoint have "d_IN (h n) x = d_IN (h 0) x" by(auto simp add: currentD_SAT[OF h]) thus ?thesis using False by(simp add: currentD_finite_IN[OF h]) qed have "d_IN (f n) \x\ = (\\<^sup>+ y. f n (y, \x\) * indicator (range Inner) y + f n (y, \x\) * indicator {SOURCE} y)" using B disjoint unfolding d_IN_def by(intro nn_integral_cong)(auto split: split_indicator simp add: notin_range_Inner) also have "\ = (\\<^sup>+ y\range Inner. f n (y, \x\)) + f n (SOURCE, \x\)" using h0_b[of n] by(simp add: nn_integral_add nn_integral_count_space_indicator max_def) also have "(\\<^sup>+ y\range Inner. f n (y, \x\)) = (\\<^sup>+ y. h 0 (y, x) - h n (y, x))" using B disjoint apply(simp add: nn_integral_count_space_reindex cong: nn_integral_cong_simp outgoing_def) apply(auto simp add: nn_integral_count_space_indicator outgoing_def B_out max.commute currentD_outside[OF h] intro!: nn_integral_cong split: split_indicator dest: edge_antiparallel) done also have "\ = (\\<^sup>+ y. h 0 (y, x) * indicator ?h0_gt_hn y) - (\\<^sup>+ y. h n (y, x) * indicator ?h0_gt_hn y)" apply(subst nn_integral_diff[symmetric]) apply(simp_all add: AE_count_space finite1 order.strict_implies_order split: split_indicator) apply(rule nn_integral_cong; auto simp add: currentD_finite[OF h] top_unique less_top[symmetric] not_less split: split_indicator intro!: diff_eq_0_ennreal) done also have eq_h_0: "(\\<^sup>+ y. h 0 (y, x) * indicator ?h0_gt_hn y) = d_IN (h 0) x - (\\<^sup>+ y. h 0 (y, x) * indicator ?h0_le_hn y)" unfolding d_IN_def apply(subst nn_integral_diff[symmetric]) apply(auto simp add: AE_count_space finite1 currentD_finite[OF h] split: split_indicator intro!: nn_integral_cong) done also have eq_h_n: "(\\<^sup>+ y. h n (y, x) * indicator ?h0_gt_hn y) = d_IN (h n) x - (\\<^sup>+ y. h n (y, x) * indicator ?h0_le_hn y)" unfolding d_IN_def apply(subst nn_integral_diff[symmetric]) apply(auto simp add: AE_count_space finite1 currentD_finite[OF h] split: split_indicator intro!: nn_integral_cong) done also have "d_IN (h 0) x - (\\<^sup>+ y. h 0 (y, x) * indicator ?h0_le_hn y) - (d_IN (h n) x - (\\<^sup>+ y. h n (y, x) * indicator ?h0_le_hn y)) + f n (SOURCE, \x\) = (\\<^sup>+ y. h n (y, x) * indicator ?h0_le_hn y) - (\\<^sup>+ y. h 0 (y, x) * indicator ?h0_le_hn y)" apply (subst diff_add_assoc2_ennreal) subgoal by (auto simp add: eq_h_0[symmetric] eq_h_n[symmetric] split: split_indicator intro!: nn_integral_mono) apply (subst diff_add_assoc2_ennreal) subgoal by (auto simp: d_IN_def split: split_indicator intro!: nn_integral_mono) apply (subst diff_diff_commute_ennreal) apply (subst diff_diff_ennreal') subgoal by (auto simp: d_IN_def split: split_indicator intro!: nn_integral_mono) [] subgoal unfolding eq_h_n[symmetric] by (rule add_increasing2) (auto simp add: d_IN_def split: split_indicator intro!: nn_integral_mono) apply (subst diff_add_assoc2_ennreal[symmetric]) unfolding eq using currentD_finite_IN[OF h] apply simp_all done also have "(\\<^sup>+ y. h n (y, x) * indicator ?h0_le_hn y) - (\\<^sup>+ y. h 0 (y, x) * indicator ?h0_le_hn y) = (\\<^sup>+ y. h n (y, x) - h 0 (y, x))" apply(subst nn_integral_diff[symmetric]) apply(simp_all add: AE_count_space max_def finite1 split: split_indicator) apply(rule nn_integral_cong; auto simp add: not_le split: split_indicator) by (metis diff_eq_0_ennreal le_less not_le top_greatest) also have "\ = (\\<^sup>+ y\range Inner. f n (\x\, y))" using B disjoint apply(simp add: nn_integral_count_space_reindex cong: nn_integral_cong_simp outgoing_def) apply(auto simp add: B_out currentD_outside[OF h] max.commute intro!: nn_integral_cong split: split_indicator dest: edge_antiparallel) done also have "\ = (\\<^sup>+ y. f n (\x\, y) * indicator (range Inner) y)" by(simp add: nn_integral_add nn_integral_count_space_indicator max.left_commute max.commute) also have "\ = d_OUT (f n) \x\" using B disjoint unfolding d_OUT_def by(intro nn_integral_cong)(auto split: split_indicator simp add: notin_range_Inner) finally show ?thesis using B by(simp) qed qed have "weight \ b - d_IN (h 0) b = (SUP n. value_flow \ (f n))" using OUT_f_SOURCE currentD_finite_IN[OF h, of 0 b] IN by (simp add: SUP_diff_ennreal less_top) also have "(SUP n. value_flow \ (f n)) \ \" unfolding \_def apply(rule SUP_least) apply(rule SUP_upper) apply(simp add: f) done also have "\ \ weight \ b - d_IN (h 0) b" unfolding \_def proof(rule SUP_least; clarsimp) fix f assume f: "flow \ f" have "d_OUT f SOURCE = (\\<^sup>+ y. f (SOURCE, y) * indicator {\b\} y)" unfolding d_OUT_def apply(rule nn_integral_cong) subgoal for x using flowD_capacity[OF f, of "(SOURCE, x)"] by(auto split: split_indicator) done also have "\ = f (SOURCE, \b\)" by(simp add: max_def) also have "\ \ weight \ b - d_IN (h 0) b" using flowD_capacity[OF f, of "(SOURCE, \b\)"] by simp finally show "d_OUT f SOURCE \ \" . qed ultimately have \: "\ = weight \ b - d_IN (h 0) b" by(rule antisym[rotated]) hence \_finite: "\ \ \" by simp from \.ex_max_flow obtain g where g: "flow \ g" and value_g: "value_flow \ g = \" and IN_g: "\x. d_IN g x \ value_flow \ g" unfolding \_def by blast have g_le_h0: "g (\x\, \y\) \ h 0 (x, y)" if "edge \ x y" for x y using flowD_capacity[OF g, of "(\x\, \y\)"] that by simp note [simp] = \.flowD_finite[OF g] have g_SOURCE: "g (SOURCE, \x\) = (if x = b then \ else 0)" for x proof(cases "x = b") case True have "g (SOURCE, \x\) = (\\<^sup>+ y. g (SOURCE, y) * indicator {\x\} y)" by(simp add: max_def) also have "\ = value_flow \ g" unfolding d_OUT_def using True by(intro nn_integral_cong)(auto split: split_indicator intro: \.flowD_outside[OF g]) finally show ?thesis using value_g by(simp add: True) qed(simp add: \.flowD_outside[OF g]) let ?g = "\(x, y). g (\y\, \x\)" define h' where "h' = h 0 \ ?g" have h'_simps: "h' (x, y) = (if edge \ x y then h 0 (x, y) + g (\y\, \x\) - g (\x\, \y\) else 0)" for x y by(simp add: h'_def) have OUT_h'_B [simp]: "d_OUT h' x = 0" if "x \ B \" for x using that unfolding d_OUT_def by(simp add: nn_integral_0_iff emeasure_count_space_eq_0)(simp add: h'_simps B_out) have IN_h'_A [simp]: "d_IN h' x = 0" if "x \ A \" for x using that unfolding d_IN_def by(simp add: nn_integral_0_iff emeasure_count_space_eq_0)(simp add: h'_simps A_in) have h'_outside: "h' e = 0" if "e \ \<^bold>E" for e unfolding h'_def using that by(rule plus_flow_outside) have OUT_h'_outside: "d_OUT h' x = 0" and IN_h'_outside: "d_IN h' x = 0" if "x \ \<^bold>V" for x using that by(auto simp add: d_OUT_def d_IN_def nn_integral_0_iff emeasure_count_space_eq_0 vertex_def intro: h'_outside) have g_le_OUT: "g (SINK, \x\) \ d_OUT g \x\" for x by (subst flowD_KIR[OF g]) (simp_all add: d_IN_ge_point) have OUT_g_A: "d_OUT ?g x = d_OUT g \x\ - g (SINK, \x\)" if "x \ A \" for x proof - have "d_OUT ?g x = (\\<^sup>+ y\range Inner. g (y, \x\))" by(simp add: nn_integral_count_space_reindex d_OUT_def) also have "\ = d_IN g \x\ - (\\<^sup>+ y. g (y, \x\) * indicator {SINK} y)" unfolding d_IN_def using that b disjoint flowD_capacity[OF g, of "(SOURCE, \x\)"] by(subst nn_integral_diff[symmetric]) (auto simp add: nn_integral_count_space_indicator notin_range_Inner max_def intro!: nn_integral_cong split: split_indicator if_split_asm) also have "\ = d_OUT g \x\ - g (SINK, \x\)" by(simp add: flowD_KIR[OF g] max_def) finally show ?thesis . qed have IN_g_A: "d_IN ?g x = d_OUT g \x\ - g (\x\, SINK)" if "x \ A \" for x proof - have "d_IN ?g x = (\\<^sup>+ y\range Inner. g (\x\, y))" by(simp add: nn_integral_count_space_reindex d_IN_def) also have "\ = d_OUT g \x\ - (\\<^sup>+ y. g (\x\, y) * indicator {SINK} y)" unfolding d_OUT_def using that b disjoint flowD_capacity[OF g, of "(\x\, SOURCE)"] by(subst nn_integral_diff[symmetric]) (auto simp add: nn_integral_count_space_indicator notin_range_Inner max_def intro!: nn_integral_cong split: split_indicator if_split_asm) also have "\ = d_OUT g \x\ - g (\x\, SINK)" by(simp add: max_def) finally show ?thesis . qed have OUT_g_B: "d_OUT ?g x = d_IN g \x\ - g (SOURCE, \x\)" if "x \ B \" for x proof - have "d_OUT ?g x = (\\<^sup>+ y\range Inner. g (y, \x\))" by(simp add: nn_integral_count_space_reindex d_OUT_def) also have "\ = d_IN g \x\ - (\\<^sup>+ y. g (y, \x\) * indicator {SOURCE} y)" unfolding d_IN_def using that b disjoint flowD_capacity[OF g, of "(SINK, \x\)"] by(subst nn_integral_diff[symmetric]) (auto simp add: nn_integral_count_space_indicator notin_range_Inner max_def intro!: nn_integral_cong split: split_indicator if_split_asm) also have "\ = d_IN g \x\ - g (SOURCE, \x\)" by(simp add: max_def) finally show ?thesis . qed have IN_g_B: "d_IN ?g x = d_OUT g \x\" if "x \ B \" for x proof - have "d_IN ?g x = (\\<^sup>+ y\range Inner. g (\x\, y))" by(simp add: nn_integral_count_space_reindex d_IN_def) also have "\ = d_OUT g \x\" unfolding d_OUT_def using that disjoint by(auto 4 3 simp add: nn_integral_count_space_indicator notin_range_Inner intro!: nn_integral_cong \.flowD_outside[OF g] split: split_indicator) finally show ?thesis . qed have finite_g_IN: "d_IN ?g x \ \" for x using \_finite proof(rule neq_top_trans) have "d_IN ?g x = (\\<^sup>+ y\range Inner. g (\x\, y))" by(auto simp add: d_IN_def nn_integral_count_space_reindex) also have "\ \ d_OUT g \x\" unfolding d_OUT_def by(auto simp add: nn_integral_count_space_indicator intro!: nn_integral_mono split: split_indicator) also have "\ = d_IN g \x\" by(rule flowD_KIR[OF g]) simp_all also have "\ \ \" using IN_g value_g by simp finally show "d_IN ?g x \ \" . qed have OUT_h'_A: "d_OUT h' x = d_OUT (h 0) x + g (\x\, SINK) - g (SINK, \x\)" if "x \ A \" for x proof - have "d_OUT h' x = d_OUT (h 0) x + (\\<^sup>+ y. ?g (x, y) * indicator \<^bold>E (x, y)) - (\\<^sup>+ y. ?g (y, x) * indicator \<^bold>E (x, y))" unfolding h'_def apply(subst OUT_plus_flow[of \ "h 0" ?g, OF currentD_outside'[OF h]]) apply(auto simp add: g_le_h0 finite_g_IN) done also have "(\\<^sup>+ y. ?g (x, y) * indicator \<^bold>E (x, y)) = d_OUT ?g x" unfolding d_OUT_def using that by(auto simp add: A_in split: split_indicator intro!: nn_integral_cong \.flowD_outside[OF g]) also have "\ = d_OUT g \x\ - g (SINK, \x\)" using that by(rule OUT_g_A) also have "(\\<^sup>+ y. ?g (y, x) * indicator \<^bold>E (x, y)) = d_IN ?g x" using that unfolding d_IN_def by(auto simp add: A_in split: split_indicator intro!: nn_integral_cong \.flowD_outside[OF g]) also have "\ = d_OUT g \x\ - g (\x\, SINK)" using that by(rule IN_g_A) also have "d_OUT (h 0) x + (d_OUT g \x\ - g (SINK, \x\)) - \ = d_OUT (h 0) x + g (\x\, SINK) - g (SINK, \x\)" apply(simp add: g_le_OUT add_diff_eq_ennreal d_OUT_ge_point) apply(subst diff_diff_commute_ennreal) apply(simp add: add_increasing d_OUT_ge_point g_le_OUT diff_diff_ennreal') apply(subst add.assoc) apply(subst (2) add.commute) apply(subst add.assoc[symmetric]) apply(subst ennreal_add_diff_cancel_right) apply(simp_all add: \.flowD_finite_OUT[OF g]) done finally show ?thesis . qed have finite_g_OUT: "d_OUT ?g x \ \" for x using \_finite proof(rule neq_top_trans) have "d_OUT ?g x = (\\<^sup>+ y\range Inner. g (y, \x\))" by(auto simp add: d_OUT_def nn_integral_count_space_reindex) also have "\ \ d_IN g \x\" unfolding d_IN_def by(auto simp add: nn_integral_count_space_indicator intro!: nn_integral_mono split: split_indicator) also have "\ \ \" using IN_g value_g by simp finally show "d_OUT ?g x \ \" . qed have IN_h'_B: "d_IN h' x = d_IN (h 0) x + g (SOURCE, \x\)" if "x \ B \" for x proof - have g_le: "g (SOURCE, \x\) \ d_IN g \x\" by (rule d_IN_ge_point) have "d_IN h' x = d_IN (h 0) x + (\\<^sup>+ y. g (\x\, \y\) * indicator \<^bold>E (y, x)) - (\\<^sup>+ y. g (\y\, \x\) * indicator \<^bold>E (y, x))" unfolding h'_def by(subst IN_plus_flow[of \ "h 0" ?g, OF currentD_outside'[OF h]]) (auto simp add: g_le_h0 finite_g_OUT) also have "(\\<^sup>+ y. g (\x\, \y\) * indicator \<^bold>E (y, x)) = d_IN ?g x" unfolding d_IN_def using that by(intro nn_integral_cong)(auto split: split_indicator intro!: \.flowD_outside[OF g] simp add: B_out) also have "\ = d_OUT g \x\" using that by(rule IN_g_B) also have "(\\<^sup>+ y. g (\y\, \x\) * indicator \<^bold>E (y, x)) = d_OUT ?g x" unfolding d_OUT_def using that by(intro nn_integral_cong)(auto split: split_indicator intro!: \.flowD_outside[OF g] simp add: B_out) also have "\ = d_IN g \x\ - g (SOURCE, \x\)" using that by(rule OUT_g_B) also have "d_IN (h 0) x + d_OUT g \x\ - \ = d_IN (h 0) x + g (SOURCE, \x\)" using \.flowD_finite_IN[OF g] g_le by(cases "d_IN (h 0) x"; cases "d_IN g \x\"; cases "d_IN g \x\"; cases "g (SOURCE, \x\)") (auto simp: flowD_KIR[OF g] top_add ennreal_minus_if ennreal_plus_if simp del: ennreal_plus) finally show ?thesis . qed have h': "current \ h'" proof fix x consider (A) "x \ A \" | (B) "x \ B \" | (outside) "x \ \<^bold>V" using bipartite_V by auto note cases = this show "d_OUT h' x \ weight \ x" proof(cases rule: cases) case A then have "d_OUT h' x = d_OUT (h 0) x + g (\x\, SINK) - g (SINK, \x\)" by(simp add: OUT_h'_A) also have "\ \ d_OUT (h 0) x + g (\x\, SINK)" by(rule diff_le_self_ennreal) also have "g (\x\, SINK) \ weight \ x - d_OUT (h 0) x" using flowD_capacity[OF g, of "(\x\, SINK)"] A by simp also have "d_OUT (h 0) x + \ = weight \ x" by(simp add: add_diff_eq_ennreal add_diff_inverse_ennreal currentD_finite_OUT[OF h] currentD_weight_OUT[OF h]) finally show ?thesis by(simp add: add_left_mono) qed(simp_all add: OUT_h'_outside ) show "d_IN h' x \ weight \ x" proof(cases rule: cases) case B hence "d_IN h' x = d_IN (h 0) x + g (SOURCE, \x\)" by(rule IN_h'_B) also have "\ \ weight \ x" by(simp add: g_SOURCE \ currentD_weight_IN[OF h] add_diff_eq_ennreal add_diff_inverse_ennreal currentD_finite_IN[OF h]) finally show ?thesis . qed(simp_all add: IN_h'_outside) next show "h' e = 0" if "e \ \<^bold>E" for e using that by(simp split: prod.split_asm add: h'_simps) qed moreover have SAT_h': "B \ \ \<^bold>V \ SAT \ h'" proof show "x \ SAT \ h'" if "x \ B \ \ \<^bold>V" for x using that proof(cases "x = b") case True have "d_IN h' x = weight \ x" using that True by(simp add: IN_h'_B g_SOURCE \ currentD_weight_IN[OF h] add_diff_eq_ennreal add_diff_inverse_ennreal currentD_finite_IN[OF h]) thus ?thesis by(simp add: SAT.simps) next case False have "d_IN h' x = d_IN (h 0) x" using that False by(simp add: IN_h'_B g_SOURCE) also have "\ = weight \ x" using SAT[of 0, THEN subsetD, of x] False that currentD_SAT[OF h, of x 0] disjoint by auto finally show ?thesis by(simp add: SAT.simps) qed qed moreover have "wave \ h'" proof have "separating \ (B \ \ \<^bold>V)" proof fix x y p assume x: "x \ A \" and y: "y \ B \" and p: "path \ x p y" hence Nil: "p \ []" using disjoint by(auto simp add: rtrancl_path_simps) from rtrancl_path_last[OF p Nil] last_in_set[OF Nil] y rtrancl_path_Range[OF p, of y] show "(\z\set p. z \ B \ \ \<^bold>V) \ x \ B \ \ \<^bold>V" by(auto intro: vertexI2) qed moreover have TER: "B \ \ \<^bold>V \ TER h'" using SAT_h' by(auto simp add: SINK) ultimately show "separating \ (TER h')" by(rule separating_weakening) qed(rule h') ultimately show ?thesis by blast qed end lemma countable_bipartite_web_reduce_weight: assumes "weight \ x \ w" shows "countable_bipartite_web (reduce_weight \ x w)" using bipartite_V A_vertex bipartite_E disjoint assms by unfold_locales (auto 4 3 simp add: weight_outside ) lemma unhinder: \ \Lemma 6.9\ assumes loose: "loose \" and b: "b \ B \" and wb: "weight \ b > 0" and \: "\ > 0" shows "\\>0. \ < \ \ \ hindered (reduce_weight \ b \)" proof(rule ccontr) assume "\ ?thesis" hence hindered: "hindered (reduce_weight \ b \)" if "\ > 0" "\ < \" for \ using that by simp from b disjoint have bnA: "b \ A \" by blast define wb where "wb = enn2real (weight \ b)" have wb_conv: "weight \ b = ennreal wb" by(simp add: wb_def less_top[symmetric]) have wb_pos: "wb > 0" using wb by(simp add: wb_conv) define \ where "\ n = min \ wb / (n + 2)" for n :: nat have \_pos: "\ n > 0" for n using wb_pos \ by(simp add: \_def) have \_nonneg: "0 \ \ n" for n using \_pos[of n] by simp have *: "\ n \ min wb \ / 2" for n using wb_pos \ by(auto simp add: \_def field_simps min_def) have \_le: "\ n \ wb" and \_less: "\ n < wb" and \_less_\: "\ n < \" and \_le': "\ n \ wb / 2" for n using *[of n] \_pos[of n] by(auto) define \' where "\' n = reduce_weight \ b (\ n)" for n :: nat have \'_sel [simp]: "edge (\' n) = edge \" "A (\' n) = A \" "B (\' n) = B \" "weight (\' n) x = weight \ x - (if x = b then \ n else 0)" "essential (\' n) = essential \" "roofed_gen (\' n) = roofed_gen \" for n x by(simp_all add: \'_def) have vertex_\' [simp]: "vertex (\' n) = vertex \" for n by(simp add: vertex_def fun_eq_iff) from wb have "b \ \<^bold>V" using weight_outside[of b] by(auto intro: ccontr) interpret \': countable_bipartite_web "\' n" for n unfolding \'_def using wb_pos by(intro countable_bipartite_web_reduce_weight)(simp_all add: wb_conv \_le \_nonneg) obtain g where g: "\n. current (\' n) (g n)" and w: "\n. wave (\' n) (g n)" and hind: "\n. hindrance (\' n) (g n)" using hindered[OF \_pos, unfolded wb_conv ennreal_less_iff, OF \_less_\] unfolding hindered.simps \'_def by atomize_elim metis from g have g\: "current \ (g n)" for n by(rule current_weight_mono)(auto simp add: \_nonneg diff_le_self_ennreal) note [simp] = currentD_finite[OF g\] have b_TER: "b \ TER\<^bsub>\' n\<^esub> (g n)" for n proof(rule ccontr) assume b': "b \ TER\<^bsub>\' n\<^esub> (g n)" then have TER: "TER\<^bsub>\' n\<^esub> (g n) = TER (g n)" using b \_nonneg[of n] by(auto simp add: SAT.simps split: if_split_asm intro: ennreal_diff_le_mono_left) from w[of n] TER have "wave \ (g n)" by(simp add: wave.simps separating_gen.simps) moreover have "hindrance \ (g n)" using hind[of n] TER bnA b' by(auto simp add: hindrance.simps split: if_split_asm) ultimately show False using loose_unhindered[OF loose] g\[of n] by(auto intro: hindered.intros) qed have IN_g_b: "d_IN (g n) b = weight \ b - \ n" for n using b_TER[of n] bnA by(auto simp add: currentD_SAT[OF g]) define factor where "factor n = (wb - \ 0) / (wb - \ n)" for n have factor_le_1: "factor n \ 1" for n using wb_pos \ \_less[of n] by(auto simp add: factor_def field_simps \_def min_def) have factor_pos: "0 < factor n" for n using wb_pos \ * \_less by(simp add: factor_def field_simps) have factor: "(wb - \ n) * factor n = wb - \ 0" for n using \_less[of n] by(simp add: factor_def field_simps) define g' where "g' = (\n (x, y). if y = b then g n (x, y) * factor n else g n (x, y))" have g'_simps: "g' n (x, y) = (if y = b then g n (x, y) * factor n else g n (x, y))" for n x y by(simp add: g'_def) have g'_le_g: "g' n e \ g n e" for n e using factor_le_1[of n] by(cases e "g n e" rule: prod.exhaust[case_product ennreal_cases]) (auto simp add: g'_simps field_simps mult_left_le) have "4 + (n * 6 + n * (n * 2)) \ (0 :: real)" for n :: nat by(metis (mono_tags, opaque_lifting) add_is_0 of_nat_eq_0_iff of_nat_numeral zero_neq_numeral) then have IN_g': "d_IN (g' n) x = (if x = b then weight \ b - \ 0 else d_IN (g n) x)" for x n using b_TER[of n] bnA factor_pos[of n] factor[of n] wb_pos \ by(auto simp add: d_IN_def g'_simps nn_integral_divide nn_integral_cmult currentD_SAT[OF g] wb_conv \_def field_simps ennreal_minus ennreal_mult'[symmetric] intro!: arg_cong[where f=ennreal]) have OUT_g': "d_OUT (g' n) x = d_OUT (g n) x - g n (x, b) * (1 - factor n)" for n x proof - have "d_OUT (g' n) x = (\\<^sup>+ y. g n (x, y)) - (\\<^sup>+ y. (g n (x, y) * (1 - factor n)) * indicator {b} y)" using factor_le_1[of n] factor_pos[of n] apply(cases "g n (x, b)") apply(subst nn_integral_diff[symmetric]) apply(auto simp add: g'_simps nn_integral_divide d_OUT_def AE_count_space mult_left_le ennreal_mult_eq_top_iff ennreal_mult'[symmetric] ennreal_minus_if intro!: nn_integral_cong split: split_indicator) apply(simp_all add: field_simps) done also have "\ = d_OUT (g n) x - g n (x, b) * (1 - factor n)" using factor_le_1[of n] by(subst nn_integral_indicator_singleton)(simp_all add: d_OUT_def field_simps) finally show ?thesis . qed have g': "current (\' 0) (g' n)" for n proof show "d_OUT (g' n) x \ weight (\' 0) x" for x using b_TER[of n] currentD_weight_OUT[OF g, of n x] \_le[of 0] factor_le_1[of n] by(auto simp add: OUT_g' SINK.simps ennreal_diff_le_mono_left) show "d_IN (g' n) x \ weight (\' 0) x" for x using d_IN_mono[of "g' n" x, OF g'_le_g] currentD_weight_IN[OF g, of n x] b_TER[of n] b by(auto simp add: IN_g' SAT.simps wb_conv \_def) show "g' n e = 0" if "e \ \<^bold>E\<^bsub>\' 0\<^esub>" for e using that by(cases e)(clarsimp simp add: g'_simps currentD_outside[OF g]) qed have SINK_g': "SINK (g n) = SINK (g' n)" for n using factor_pos[of n] by(auto simp add: SINK.simps currentD_OUT_eq_0[OF g] currentD_OUT_eq_0[OF g'] g'_simps split: if_split_asm) have SAT_g': "SAT (\' n) (g n) = SAT (\' 0) (g' n)" for n using b_TER[of n] \_le'[of 0] by(auto simp add: SAT.simps wb_conv IN_g' IN_g_b) have TER_g': "TER\<^bsub>\' n\<^esub> (g n) = TER\<^bsub>\' 0\<^esub> (g' n)" for n using b_TER[of n] by(auto simp add: SAT.simps SINK_g' OUT_g' IN_g' wb_conv \_def) have w': "wave (\' 0) (g' n)" for n proof have "separating (\' 0) (TER\<^bsub>\' n\<^esub> (g n))" using waveD_separating[OF w, of n] by(simp add: separating_gen.simps) then show "separating (\' 0) (TER\<^bsub>\' 0\<^esub> (g' n))" unfolding TER_g' . qed(rule g') define f where "f = rec_nat (g 0) (\n rec. rec \\<^bsub>\' 0\<^esub> g' (n + 1))" have f_simps [simp]: "f 0 = g 0" "f (Suc n) = f n \\<^bsub>\' 0\<^esub> g' (n + 1)" for n by(simp_all add: f_def) have f: "current (\' 0) (f n)" and fw: "wave (\' 0) (f n)" for n proof(induction n) case (Suc n) { case 1 show ?case unfolding f_simps using Suc.IH g' by(rule current_plus_web) } { case 2 show ?case unfolding f_simps using Suc.IH g' w' by(rule wave_plus') } qed(simp_all add: g w) have f_inc: "n \ m \ f n \ f m" for n m proof(induction m rule: dec_induct) case (step k) note step.IH also have "f k \ (f k \\<^bsub>\' 0\<^esub> g' (k + 1))" by(rule le_funI plus_web_greater)+ also have "\ = f (Suc k)" by simp finally show ?case . qed simp have chain_f: "Complete_Partial_Order.chain (\) (range f)" by(rule chain_imageI[where le_a="(\)"])(simp_all add: f_inc) have "countable (support_flow (f n))" for n using current_support_flow[OF f, of n] by(rule countable_subset) simp hence supp_f: "countable (support_flow (SUP n. f n))" by(subst support_flow_Sup)simp have RF_f: "RF (TER\<^bsub>\' 0\<^esub> (f n)) = RF (\i\n. TER\<^bsub>\' 0\<^esub> (g' i))" for n proof(induction n) case 0 show ?case by(simp add: TER_g') next case (Suc n) have "RF (TER\<^bsub>\' 0\<^esub> (f (Suc n))) = RF\<^bsub>\' 0\<^esub> (TER\<^bsub>\' 0\<^esub> (f n \\<^bsub>\' 0\<^esub> g' (n + 1)))" by simp also have "\ = RF\<^bsub>\' 0\<^esub> (TER\<^bsub>\' 0\<^esub> (f n) \ TER\<^bsub>\' 0\<^esub> (g' (n + 1)))" using f fw g' w' by(rule RF_TER_plus_web) also have "\ = RF\<^bsub>\' 0\<^esub> (RF\<^bsub>\' 0\<^esub> (TER\<^bsub>\' 0\<^esub> (f n)) \ TER\<^bsub>\' 0\<^esub> (g' (n + 1)))" by(simp add: roofed_idem_Un1) also have "RF\<^bsub>\' 0\<^esub> (TER\<^bsub>\' 0\<^esub> (f n)) = RF\<^bsub>\' 0\<^esub> (\i\n. TER\<^bsub>\' 0\<^esub> (g' i))" by(simp add: Suc.IH) also have "RF\<^bsub>\' 0\<^esub> (\ \ TER\<^bsub>\' 0\<^esub> (g' (n + 1))) = RF\<^bsub>\' 0\<^esub> ((\i\n. TER\<^bsub>\' 0\<^esub> (g' i)) \ TER\<^bsub>\' 0\<^esub> (g' (n + 1)))" by(simp add: roofed_idem_Un1) also have "(\i\n. TER\<^bsub>\' 0\<^esub> (g' i)) \ TER\<^bsub>\' 0\<^esub> (g' (n + 1)) = (\i\Suc n. TER\<^bsub>\' 0\<^esub> (g' i))" unfolding atMost_Suc UN_insert by(simp add: Un_commute) finally show ?case by simp qed define g\ where "g\ = (SUP n. f n)" have g\: "current (\' 0) g\" unfolding g\_def using chain_f by(rule current_Sup)(auto simp add: f supp_f) have w\: "wave (\' 0) g\" unfolding g\_def using chain_f by(rule wave_lub)(auto simp add: fw supp_f) from g\ have g\': "current (\' n) g\" for n using wb_pos \ by(elim current_weight_mono)(auto simp add: \_le wb_conv \_def field_simps ennreal_minus_if min_le_iff_disj) have SINK_g\: "SINK g\ = (\n. SINK (f n))" unfolding g\_def by(subst SINK_Sup[OF chain_f])(simp_all add: supp_f) have SAT_g\: "SAT (\' 0) (f n) \ SAT (\' 0) g\" for n unfolding g\_def by(rule SAT_Sup_upper) simp have g_b_out: "g n (b, x) = 0" for n x using b_TER[of n] by(simp add: SINK.simps currentD_OUT_eq_0[OF g]) have g'_b_out: "g' n (b, x) = 0" for n x by(simp add: g'_simps g_b_out) have "f n (b, x) = 0" for n x by(induction n)(simp_all add: g_b_out g'_b_out) hence b_SINK_f: "b \ SINK (f n)" for n by(simp add: SINK.simps d_OUT_def) hence b_SINK_g\: "b \ SINK g\" by(simp add: SINK_g\) have RF_circ: "RF\<^sup>\\<^bsub>\' n\<^esub> (TER\<^bsub>\' 0\<^esub> (g' n)) = RF\<^sup>\\<^bsub>\' 0\<^esub> (TER\<^bsub>\' 0\<^esub> (g' n))" for n by(simp add: roofed_circ_def) have edge_restrict_\': "edge (quotient_web (\' 0) (g' n)) = edge (quotient_web (\' n) (g n))" for n by(simp add: fun_eq_iff TER_g' RF_circ) have restrict_curr_g': "f \ \' 0 / g' n = f \ \' n / g n" for n f by(simp add: restrict_current_def RF_circ TER_g') have RF_restrict: "roofed_gen (quotient_web (\' n) (g n)) = roofed_gen (quotient_web (\' 0) (g' n))" for n by(simp add: roofed_def fun_eq_iff edge_restrict_\') have g\r': "current (quotient_web (\' 0) (g' n)) (g\ \ \' 0 / g' n)" for n using w' g\ by(rule current_restrict_current) have g\r: "current (quotient_web (\' n) (g n)) (g\ \ \' n / g n)" for n using w g\' by(rule current_restrict_current) have w\r: "wave (quotient_web (\' n) (g n)) (g\ \ \' n / g n)" (is "wave ?\' ?g\") for n proof have *: "wave (quotient_web (\' 0) (g' n)) (g\ \ \' 0 / g' n)" using g' w' g\ w\ by(rule wave_restrict_current) have "d_IN (g\ \ \' n / g n) b = 0" by(rule d_IN_restrict_current_outside roofed_greaterI b_TER)+ hence SAT_subset: "SAT (quotient_web (\' 0) (g' n)) (g\ \ \' n / g n) \ SAT ?\' (g\ \ \' n / g n)" using b_TER[of n] wb_pos by(auto simp add: SAT.simps TER_g' RF_circ wb_conv \_def field_simps ennreal_minus_if split: if_split_asm) hence TER_subset: "TER\<^bsub>quotient_web (\' 0) (g' n)\<^esub> (g\ \ \' n / g n) \ TER\<^bsub>?\'\<^esub> (g\ \ \' n / g n)" using SINK_g' by(auto simp add: restrict_curr_g') show "separating ?\' (TER\<^bsub>?\'\<^esub> ?g\)" (is "separating _ ?TER") proof fix x y p assume xy: "x \ A ?\'" "y \ B ?\'" and p: "path ?\' x p y" from p have p': "path (quotient_web (\' 0) (g' n)) x p y" by(simp add: edge_restrict_\') with waveD_separating[OF *, THEN separatingD, simplified, OF p'] TER_g'[of n] SINK_g' SAT_g' restrict_curr_g' SAT_subset xy show "(\z\set p. z \ ?TER) \ x \ ?TER" by auto qed show "d_OUT (g\ \ \' n / g n) x = 0" if "x \ RF\<^bsub>?\'\<^esub> ?TER" for x unfolding restrict_curr_g'[symmetric] using TER_subset that by(intro waveD_OUT[OF *])(auto simp add: TER_g' restrict_curr_g' RF_restrict intro: in_roofed_mono) qed have RF_g\: "RF (TER\<^bsub>\' 0\<^esub> g\) = RF (\n. TER\<^bsub>\' 0\<^esub> (g' n))" proof - have "RF\<^bsub>\' 0\<^esub> (TER\<^bsub>\' 0\<^esub> g\) = RF (\i. TER\<^bsub>\' 0\<^esub> (f i))" unfolding g\_def by(subst RF_TER_Sup[OF _ _ chain_f])(auto simp add: f fw supp_f) also have "\ = RF (\i. RF (TER\<^bsub>\' 0\<^esub> (f i)))" by(simp add: roofed_UN) also have "\ = RF (\i. \j\i. TER\<^bsub>\' 0\<^esub> (g' j))" unfolding RF_f roofed_UN by simp also have "(\i. \j\i. TER\<^bsub>\' 0\<^esub> (g' j)) = (\i. TER\<^bsub>\' 0\<^esub> (g' i))" by auto finally show ?thesis by simp qed have SAT_plus_\: "SAT (\' n) (g n \\<^bsub>\' n\<^esub> g\) = SAT (\' 0) (g' n \\<^bsub>\' 0\<^esub> g\)" for n apply(intro set_eqI) apply(simp add: SAT.simps IN_plus_current[OF g w g\r] IN_plus_current[OF g' w' g\r'] TER_g') apply(cases "d_IN (g\ \ \' n / g n) b") apply(auto simp add: SAT.simps wb_conv d_IN_plus_web IN_g') apply(simp_all add: wb_conv IN_g_b restrict_curr_g' \_def field_simps) apply(metis TER_g' b_TER roofed_greaterI)+ done have SINK_plus_\: "SINK (g n \\<^bsub>\' n\<^esub> g\) = SINK (g' n \\<^bsub>\' 0\<^esub> g\)" for n apply(rule set_eqI; simp add: SINK.simps OUT_plus_current[OF g w g\r] OUT_plus_current[OF g' w'] current_restrict_current[OF w' g\]) using factor_pos[of n] by(auto simp add: RF_circ TER_g' restrict_curr_g' currentD_OUT_eq_0[OF g] currentD_OUT_eq_0[OF g'] g'_simps split: if_split_asm) have TER_plus_\: "TER\<^bsub>\' n\<^esub> (g n \\<^bsub>\' n\<^esub> g\) = TER\<^bsub>\' 0\<^esub> (g' n \\<^bsub>\' 0\<^esub> g\)" for n by(rule set_eqI iffI)+(simp_all add: SAT_plus_\ SINK_plus_\) define h where "h n = g n \\<^bsub>\' n\<^esub> g\" for n have h: "current (\' n) (h n)" for n unfolding h_def using g w by(rule current_plus_current)(rule current_restrict_current[OF w g\']) have hw: "wave (\' n) (h n)" for n unfolding h_def using g w g\' w\r by(rule wave_plus) define T where "T = TER\<^bsub>\' 0\<^esub> g\" have RF_h: "RF (TER\<^bsub>\' n\<^esub> (h n)) = RF T" for n proof - have "RF\<^bsub>\' 0\<^esub> (TER\<^bsub>\' n\<^esub> (h n)) = RF\<^bsub>\' 0\<^esub> (RF\<^bsub>\' 0\<^esub> (TER\<^bsub>\' 0\<^esub> g\) \ TER\<^bsub>\' 0\<^esub> (g' n))" unfolding h_def TER_plus_\ RF_TER_plus_web[OF g' w' g\ w\] roofed_idem_Un1 by(simp add: Un_commute) also have "\ = RF ((\n. TER\<^bsub>\' 0\<^esub> (g' n)) \ TER\<^bsub>\' 0\<^esub> (g' n))" by(simp add: RF_g\ roofed_idem_Un1) also have "\ = RF\<^bsub>\' 0\<^esub> T" unfolding T_def by(auto simp add: RF_g\ intro!: arg_cong2[where f=roofed] del: equalityI) auto finally show ?thesis by simp qed have OUT_h_nT: "d_OUT (h n) x = 0" if "x \ RF T" for n x by(rule waveD_OUT[OF hw])(simp add: RF_h that) have IN_h_nT: "d_IN (h n) x = 0" if "x \ RF T" for n x by(rule wave_not_RF_IN_zero[OF h hw])(simp add: RF_h that) have OUT_h_b: "d_OUT (h n) b = 0" for n using b_TER[of n] b_SINK_g\[THEN in_SINK_restrict_current] by(auto simp add: h_def OUT_plus_current[OF g w g\r] SINK.simps) have OUT_h_\: "d_OUT (h n) x = 0" if "x \ \ T" for x n using that apply(subst (asm) \_RF[symmetric]) apply(subst (asm) (1 2) RF_h[symmetric, of n]) apply(subst (asm) \_RF) apply(simp add: SINK.simps) done have IN_h_\: "d_IN (h n) x = weight (\' n) x" if "x \ \ T" "x \ A \" for x n using that apply(subst (asm) \_RF[symmetric]) apply(subst (asm) (1 2) RF_h[symmetric, of n]) apply(subst (asm) \_RF) apply(simp add: currentD_SAT[OF h]) done have b_SAT: "b \ SAT (\' 0) (h 0)" using b_TER[of 0] by(auto simp add: h_def SAT.simps d_IN_plus_web intro: order_trans) have b_T: "b \ T" using b_SINK_g\ b_TER by(simp add: T_def)(metis SAT_g\ subsetD f_simps(1)) have essential: "b \ \ T" proof(rule ccontr) assume "b \ \ T" hence b: "b \ \ (TER\<^bsub>\' 0\<^esub> (h 0))" proof(rule contrapos_nn) assume "b \ \ (TER\<^bsub>\' 0\<^esub> (h 0))" then obtain p y where p: "path \ b p y" and y: "y \ B \" and distinct: "distinct (b # p)" and bypass: "\z. z \ set p \ z \ RF (TER\<^bsub>\' 0\<^esub> (h 0))" by(rule \_E_RF) auto from bypass have bypass': "\z. z \ set p \ z \ T" unfolding RF_h by(auto intro: roofed_greaterI) have "essential \ (B \) T b" using p y by(rule essentialI)(auto dest: bypass') then show "b \ \ T" using b_T by simp qed have h0: "current \ (h 0)" using h[of 0] by(rule current_weight_mono)(simp_all add: wb_conv \_nonneg) moreover have "wave \ (h 0)" proof have "separating (\' 0) (\\<^bsub>\' 0\<^esub> (TER\<^bsub>\' 0\<^esub> (h 0)))" by(rule separating_essential)(rule waveD_separating[OF hw]) then have "separating \ (\ (TER\<^bsub>\' 0\<^esub> (h 0)))" by(simp add: separating_gen.simps) moreover have subset: "\ (TER\<^bsub>\' 0\<^esub> (h 0)) \ TER (h 0)" using \_nonneg[of 0] b by(auto simp add: SAT.simps wb_conv split: if_split_asm) ultimately show "separating \ (TER (h 0))" by(rule separating_weakening) qed(rule h0) ultimately have "h 0 = zero_current" by(rule looseD_wave[OF loose]) then have "d_IN (h 0) b = 0" by(simp) with b_SAT wb \b \ A \\ show False by(simp add: SAT.simps wb_conv \_def ennreal_minus_if split: if_split_asm) qed define S where "S = {x \ RF (T \ B \) \ A \. essential \ (T \ B \) (RF (T \ B \) \ A \) x}" define \_h where "\_h = \ edge = \x y. edge \ x y \ x \ S \ y \ T \ y \ B \, weight = \x. weight \ x * indicator (S \ T \ B \) x, A = S, B = T \ B \\" have \_h_sel [simp]: "edge \_h x y \ edge \ x y \ x \ S \ y \ T \ y \ B \" "A \_h = S" "B \_h = T \ B \" "weight \_h x = weight \ x * indicator (S \ T \ B \) x" for x y by(simp_all add: \_h_def) have vertex_\_hD: "x \ S \ (T \ B \)" if "vertex \_h x" for x using that by(auto simp add: vertex_def) have S_vertex: "vertex \_h x" if "x \ S" for x proof - from that have a: "x \ A \" and RF: "x \ RF (T \ B \)" and ess: "essential \ (T \ B \) (RF (T \ B \) \ A \) x" by(simp_all add: S_def) from ess obtain p y where p: "path \ x p y" and y: "y \ B \" and yT: "y \ T" and bypass: "\z. z \ set p \ z \ RF (T \ B \) \ A \" by(rule essentialE_RF)(auto intro: roofed_greaterI) from p a y disjoint have "edge \ x y" by(cases)(auto 4 3 elim: rtrancl_path.cases dest: bipartite_E) with that y yT show ?thesis by(auto intro!: vertexI1) qed have OUT_not_S: "d_OUT (h n) x = 0" if "x \ S" for x n proof(rule classical) assume *: "d_OUT (h n) x \ 0" consider (A) "x \ A \" | (B) "x \ B \" | (outside) "x \ A \" "x \ B \" by blast then show ?thesis proof cases case B with currentD_OUT[OF h, of x n] show ?thesis by simp next case outside with currentD_outside_OUT[OF h, of x n] show ?thesis by(simp add: not_vertex) next case A from * obtain y where xy: "h n (x, y) \ 0" using currentD_OUT_eq_0[OF h, of n x] by auto then have edge: "edge \ x y" using currentD_outside[OF h] by(auto) hence p: "path \ x [y] y" by(simp add: rtrancl_path_simps) from bipartite_E[OF edge] have x: "x \ A \" and y: "y \ B \" by simp_all moreover have "x \ RF (RF (T \ B \))" proof fix p y' assume p: "path \ x p y'" and y': "y' \ B \" from p x y' disjoint have py: "p = [y']" by(cases)(auto 4 3 elim: rtrancl_path.cases dest: bipartite_E) have "separating (\' 0) (RF\<^bsub>\' 0\<^esub> (TER\<^bsub>\' 0\<^esub> (h 0)))" unfolding separating_RF by(rule waveD_separating[OF hw]) from separatingD[OF this, of x p y'] py p x y' have "x \ RF T \ y' \ RF T" by(auto simp add: RF_h) thus "(\z\set p. z \ RF (T \ B \)) \ x \ RF (T \ B \)" proof cases case right with y' py show ?thesis by(simp add: RF_in_B) next case left have "x \ \ T" using OUT_h_\[of x n] xy by(auto simp add: currentD_OUT_eq_0[OF h]) with left have "x \ RF\<^sup>\ T" by(simp add: roofed_circ_def) from RF_circ_edge_forward[OF this, of y'] p py have "y' \ RF T" by(simp add: rtrancl_path_simps) with y' have "y' \ T" by(simp add: RF_in_B) with y' show ?thesis using py by(auto intro: roofed_greaterI) qed qed moreover have "y \ T" using IN_h_nT[of y n] y xy by(auto simp add: RF_in_B currentD_IN_eq_0[OF h]) with p x y disjoint have "essential \ (T \ B \) (RF (T \ B \) \ A \) x" by(auto intro!: essentialI) ultimately have "x \ S" unfolding roofed_idem by(simp add: S_def) with that show ?thesis by contradiction qed qed have B_vertex: "vertex \_h y" if T: "y \ T" and B: "y \ B \" and w: "weight \ y > 0" for y proof - from T B disjoint \_less[of 0] w have "d_IN (h 0) y > 0" using IN_h_\[of y 0] by(cases "y \ A \")(auto simp add: essential_BI wb_conv ennreal_minus_if) then obtain x where xy: "h 0 (x, y) \ 0" using currentD_IN_eq_0[OF h, of 0 y] by(auto) then have edge: "edge \ x y" using currentD_outside[OF h] by(auto) from xy have "d_OUT (h 0) x \ 0" by(auto simp add: currentD_OUT_eq_0[OF h]) hence "x \ S" using OUT_not_S[of x 0] by(auto) with edge T B show ?thesis by(simp add: vertexI2) qed have \_h: "countable_bipartite_web \_h" proof show "\<^bold>V\<^bsub>\_h\<^esub> \ A \_h \ B \_h" by(auto simp add: vertex_def) show "A \_h \ \<^bold>V\<^bsub>\_h\<^esub>" using S_vertex by auto show "x \ A \_h \ y \ B \_h" if "edge \_h x y" for x y using that by auto show "A \_h \ B \_h = {}" using disjoint by(auto simp add: S_def) have "\<^bold>E\<^bsub>\_h\<^esub> \ \<^bold>E" by auto thus "countable \<^bold>E\<^bsub>\_h\<^esub>" by(rule countable_subset) simp show "weight \_h x \ \" for x by(simp split: split_indicator) show "weight \_h x = 0" if "x \ \<^bold>V\<^bsub>\_h\<^esub>" for x using that S_vertex B_vertex[of x] by(cases "weight \_h x > 0")(auto split: split_indicator) qed then interpret \_h: countable_bipartite_web \_h . have essential_T: "essential \ (B \) T = essential \ (B \) (TER\<^bsub>\' 0\<^esub> (h 0))" proof(rule ext iffI)+ fix x assume "essential \ (B \) T x" then obtain p y where p: "path \ x p y" and y: "y \ B \" and distinct: "distinct (x # p)" and bypass: "\z. z \ set p \ z \ RF T" by(rule essentialE_RF)auto from bypass have bypass': "\z. z \ set p \ z \ TER\<^bsub>\' 0\<^esub> (h 0)" unfolding RF_h[of 0, symmetric] by(blast intro: roofed_greaterI) show "essential \ (B \) (TER\<^bsub>\' 0\<^esub> (h 0)) x" using p y by(blast intro: essentialI dest: bypass') next fix x assume "essential \ (B \) (TER\<^bsub>\' 0\<^esub> (h 0)) x" then obtain p y where p: "path \ x p y" and y: "y \ B \" and distinct: "distinct (x # p)" and bypass: "\z. z \ set p \ z \ RF (TER\<^bsub>\' 0\<^esub> (h 0))" by(rule essentialE_RF)auto from bypass have bypass': "\z. z \ set p \ z \ T" unfolding RF_h[of 0] by(blast intro: roofed_greaterI) show "essential \ (B \) T x" using p y by(blast intro: essentialI dest: bypass') qed have h': "current \_h (h n)" for n proof show "d_OUT (h n) x \ weight \_h x" for x using currentD_weight_OUT[OF h, of n x] \_nonneg[of n] \'.currentD_OUT'[OF h, of x n] OUT_not_S by(auto split: split_indicator if_split_asm elim: order_trans intro: diff_le_self_ennreal in_roofed_mono simp add: OUT_h_b roofed_circ_def) show "d_IN (h n) x \ weight \_h x" for x using currentD_weight_IN[OF h, of n x] currentD_IN[OF h, of x n] \_nonneg[of n] b_T b \'.currentD_IN'[OF h, of x n] IN_h_nT[of x n] by(cases "x \ B \")(auto 4 3 split: split_indicator split: if_split_asm elim: order_trans intro: diff_le_self_ennreal simp add: S_def roofed_circ_def RF_in_B) show "h n e = 0" if "e \ \<^bold>E\<^bsub>\_h\<^esub>" for e using that OUT_not_S[of "fst e" n] currentD_outside'[OF h, of e n] \'.currentD_IN'[OF h, of "snd e" n] disjoint apply(cases "e \ \<^bold>E") apply(auto split: prod.split_asm simp add: currentD_OUT_eq_0[OF h] currentD_IN_eq_0[OF h]) apply(cases "fst e \ S"; clarsimp simp add: S_def) apply(frule RF_circ_edge_forward[rotated]) apply(erule roofed_circI, blast) apply(drule bipartite_E) apply(simp add: RF_in_B) done qed have SAT_h': "B \_h \ \<^bold>V\<^bsub>\_h\<^esub> - {b} \ SAT \_h (h n)" for n proof fix x assume "x \ B \_h \ \<^bold>V\<^bsub>\_h\<^esub> - {b}" then have x: "x \ T" and B: "x \ B \" and b: "x \ b" and vertex: "x \ \<^bold>V\<^bsub>\_h\<^esub>" by auto from B disjoint have xnA: "x \ A \" by blast from x B have "x \ \ T" by(simp add: essential_BI) hence "d_IN (h n) x = weight (\' n) x" using xnA by(rule IN_h_\) with xnA b x B show "x \ SAT \_h (h n)" by(simp add: currentD_SAT[OF h']) qed moreover have "b \ B \_h" using b essential by simp moreover have "(\n. min \ wb * (1 / (real (n + 2)))) \ 0" apply(rule LIMSEQ_ignore_initial_segment) apply(rule tendsto_mult_right_zero) apply(rule lim_1_over_real_power[where s=1, simplified]) done then have "(INF n. ennreal (\ n)) = 0" using wb_pos \ apply(simp add: \_def) apply(rule INF_Lim) apply(rule decseq_SucI) apply(simp add: field_simps min_def) apply(simp add: add.commute ennreal_0[symmetric] del: ennreal_0) done then have "(SUP n. d_IN (h n) b) = weight \_h b" using essential b bnA wb IN_h_\[of b] by(simp add: SUP_const_minus_ennreal) moreover have "d_IN (h 0) b \ d_IN (h n) b" for n using essential b bnA wb_pos \ IN_h_\[of b] by(simp add: wb_conv \_def field_simps ennreal_minus_if min_le_iff_disj) moreover have b_V: "b \ \<^bold>V\<^bsub>\_h\<^esub>" using b wb essential by(auto dest: B_vertex) ultimately have "\h'. current \_h h' \ wave \_h h' \ B \_h \ \<^bold>V\<^bsub>\_h\<^esub> \ SAT \_h h'" by(rule \_h.unhinder_bipartite[OF h']) then obtain h' where h': "current \_h h'" and h'w: "wave \_h h'" and B_SAT': "B \_h \ \<^bold>V\<^bsub>\_h\<^esub> \ SAT \_h h'" by blast have h'': "current \ h'" proof show "d_OUT h' x \ weight \ x" for x using currentD_weight_OUT[OF h', of x] by(auto split: split_indicator_asm elim: order_trans intro: ) show "d_IN h' x \ weight \ x" for x using currentD_weight_IN[OF h', of x] by(auto split: split_indicator_asm elim: order_trans intro: ) show "h' e = 0" if "e \ \<^bold>E" for e using currentD_outside'[OF h', of e] that by auto qed moreover have "wave \ h'" proof have "separating (\' 0) T" unfolding T_def by(rule waveD_separating[OF w\]) hence "separating \ T" by(simp add: separating_gen.simps) hence *: "separating \ (\ T)" by(rule separating_essential) show "separating \ (TER h')" proof fix x p y assume x: "x \ A \" and p: "path \ x p y" and y: "y \ B \" from p x y disjoint have py: "p = [y]" by(cases)(auto 4 3 elim: rtrancl_path.cases dest: bipartite_E) from separatingD[OF * p x y] py have "x \ \ T \ y \ \ T" by auto then show "(\z\set p. z \ TER h') \ x \ TER h'" proof cases case left then have "x \ \<^bold>V\<^bsub>\_h\<^esub>" using x disjoint by(auto 4 4 dest!: vertex_\_hD simp add: S_def elim: essentialE_RF intro!: roofed_greaterI dest: roofedD) hence "d_OUT h' x = 0" by(intro currentD_outside_OUT[OF h']) with x have "x \ TER h'" by(auto simp add: SAT.A SINK.simps) thus ?thesis .. next case right have "y \ SAT \ h'" proof(cases "weight \ y > 0") case True with py x y right have "vertex \_h y" by(auto intro: B_vertex) hence "y \ SAT \_h h'" using B_SAT' right y by auto with right y disjoint show ?thesis by(auto simp add: currentD_SAT[OF h'] currentD_SAT[OF h''] S_def) qed(auto simp add: SAT.simps) with currentD_OUT[OF h', of y] y right have "y \ TER h'" by(auto simp add: SINK) thus ?thesis using py by simp qed qed qed(rule h'') ultimately have "h' = zero_current" by(rule looseD_wave[OF loose]) hence "d_IN h' b = 0" by simp moreover from essential b b_V B_SAT' have "b \ SAT \_h h'" by(auto) ultimately show False using wb b essential disjoint by(auto simp add: SAT.simps S_def) qed end subsection \Single-vertex saturation in unhindered bipartite webs\ text \ The proof of lemma 6.10 in @{cite "AharoniBergerGeorgakopoulusPerlsteinSpruessel2011JCT"} is flawed. The transfinite steps (taking the least upper bound) only preserves unhinderedness, but not looseness. However, the single steps to non-limit ordinals assumes that \\ - f\<^sub>i\ is loose in order to apply Lemma 6.9. Counterexample: The bipartite web with three nodes \a\<^sub>1\, \a\<^sub>2\, \a\<^sub>3\ in \A\ and two nodes \b\<^sub>1\, \b\<^sub>2\ in \B\ and edges \(a\<^sub>1, b\<^sub>1)\, \(a\<^sub>2, b\<^sub>1)\, \(a\<^sub>2, b\<^sub>2)\, \(a\<^sub>3, b\<^sub>2)\ and weights \a\<^sub>1 = a\<^sub>3 = 1\ and \a\<^sub>2 = 2\ and \b\<^sub>1 = 3\ and \b\<^sub>2 = 2\. Then, we can get a sequence of weight reductions on \b\<^sub>2\ from \2\ to \1.5\, \1.25\, \1.125\, etc. with limit \1\. All maximal waves in the restricted webs in the sequence are @{term [source] zero_current}, so in the limit, we get \k = 0\ and \\ = 1\ for \a\<^sub>2\ and \b\<^sub>2\. Now, the restricted web for the two is not loose because it contains the wave which assigns 1 to \(a\<^sub>3, b\<^sub>2)\. We prove a stronger version which only assumes and ensures on unhinderedness. \ context countable_bipartite_web begin lemma web_flow_iff: "web_flow \ f \ current \ f" using bipartite_V by(auto simp add: web_flow.simps) lemma countable_bipartite_web_minus_web: assumes f: "current \ f" shows "countable_bipartite_web (\ \ f)" using bipartite_V A_vertex bipartite_E disjoint currentD_finite_OUT[OF f] currentD_weight_OUT[OF f] currentD_weight_IN[OF f] currentD_outside_OUT[OF f] currentD_outside_IN[OF f] by unfold_locales (auto simp add: weight_outside) lemma current_plus_current_minus: assumes f: "current \ f" and g: "current (\ \ f) g" shows "current \ (plus_current f g)" (is "current _ ?fg") proof interpret \: countable_bipartite_web "\ \ f" using f by(rule countable_bipartite_web_minus_web) show "d_OUT ?fg x \ weight \ x" for x using currentD_weight_OUT[OF g, of x] currentD_OUT[OF g, of x] currentD_finite_OUT[OF f, of x] currentD_OUT[OF f, of x] currentD_outside_IN[OF f, of x] currentD_outside_OUT[OF f, of x] currentD_weight_OUT[OF f, of x] by(cases "x \ A \ \ x \ B \")(auto simp add: add.commute d_OUT_def nn_integral_add not_vertex ennreal_le_minus_iff split: if_split_asm) show "d_IN ?fg x \ weight \ x" for x using currentD_weight_IN[OF g, of x] currentD_IN[OF g, of x] currentD_finite_IN[OF f, of x] currentD_OUT[OF f, of x] currentD_outside_IN[OF f, of x] currentD_outside_OUT[OF f, of x] currentD_weight_IN[OF f, of x] by(cases "x \ A \ \ x \ B \")(auto simp add: add.commute d_IN_def nn_integral_add not_vertex ennreal_le_minus_iff split: if_split_asm) show "?fg e = 0" if "e \ \<^bold>E" for e using that currentD_outside'[OF f, of e] currentD_outside'[OF g, of e] by(cases e) simp qed lemma wave_plus_current_minus: assumes f: "current \ f" and w: "wave \ f" and g: "current (\ \ f) g" and w': "wave (\ \ f) g" shows "wave \ (plus_current f g)" (is "wave _ ?fg") proof show fg: "current \ ?fg" using f g by(rule current_plus_current_minus) show sep: "separating \ (TER ?fg)" proof fix x p y assume x: "x \ A \" and p: "path \ x p y" and y: "y \ B \" from p x y disjoint have py: "p = [y]" by(cases)(auto 4 3 elim: rtrancl_path.cases dest: bipartite_E) with waveD_separating[THEN separatingD, OF w p x y] have "x \ TER f \ y \ TER f" by auto thus "(\z\set p. z \ TER ?fg) \ x \ TER ?fg" proof cases case right with y disjoint have "y \ TER ?fg" using currentD_OUT[OF fg y] by(auto simp add: SAT.simps SINK.simps d_IN_def nn_integral_add not_le add_increasing2) thus ?thesis using py by simp next case x': left from p have "path (\ \ f) x p y" by simp from waveD_separating[THEN separatingD, OF w' this] x y py have "x \ TER\<^bsub>\ \ f\<^esub> g \ y \ TER\<^bsub>\ \ f\<^esub> g" by auto thus ?thesis proof cases case left hence "x \ TER ?fg" using x x' by(auto simp add: SAT.simps SINK.simps d_OUT_def nn_integral_add) thus ?thesis .. next case right hence "y \ TER ?fg" using disjoint y currentD_OUT[OF fg y] currentD_OUT[OF f y] currentD_finite_IN[OF f, of y] by(auto simp add: add.commute SINK.simps SAT.simps d_IN_def nn_integral_add ennreal_minus_le_iff split: if_split_asm) with py show ?thesis by auto qed qed qed qed lemma minus_plus_current: assumes f: "current \ f" and g: "current (\ \ f) g" shows "\ \ plus_current f g = \ \ f \ g" (is "?lhs = ?rhs") proof(rule web.equality) have "weight ?lhs x = weight ?rhs x" for x using currentD_weight_IN[OF f, of x] currentD_weight_IN[OF g, of x] by (auto simp add: d_IN_def d_OUT_def nn_integral_add diff_add_eq_diff_diff_swap_ennreal add_increasing2 diff_add_assoc2_ennreal add.assoc) thus "weight ?lhs = weight ?rhs" .. qed simp_all lemma unhindered_minus_web: assumes unhindered: "\ hindered \" and f: "current \ f" and w: "wave \ f" shows "\ hindered (\ \ f)" proof assume "hindered (\ \ f)" then obtain g where g: "current (\ \ f) g" and w': "wave (\ \ f) g" and hind: "hindrance (\ \ f) g" by cases let ?fg = "plus_current f g" have fg: "current \ ?fg" using f g by(rule current_plus_current_minus) moreover have "wave \ ?fg" using f w g w' by(rule wave_plus_current_minus) moreover from hind obtain a where a: "a \ A \" and n\: "a \ \\<^bsub>\ \ f\<^esub> (TER\<^bsub>\ \ f\<^esub> g)" and wa: "d_OUT g a < weight (\ \ f) a" by cases auto from a have "hindrance \ ?fg" proof show "a \ \ (TER ?fg)" proof assume \: "a \ \ (TER ?fg)" then obtain p y where p: "path \ a p y" and y: "y \ B \" and bypass: "\z. z \ set p \ z \ RF (TER ?fg)" by(rule \_E_RF) blast from p a y disjoint have py: "p = [y]" by(cases)(auto 4 3 elim: rtrancl_path.cases dest: bipartite_E) from bypass[of y] py have "y \ TER ?fg" by(auto intro: roofed_greaterI) with currentD_OUT[OF fg y] have "y \ SAT \ ?fg" by(auto simp add: SINK.simps) hence "y \ SAT (\ \ f) g" using y currentD_OUT[OF f y] currentD_finite_IN[OF f, of y] by(auto simp add: SAT.simps d_IN_def nn_integral_add ennreal_minus_le_iff add.commute) hence "essential (\ \ f) (B (\ \ f)) (TER\<^bsub>\ \ f\<^esub> g) a" using p py y by(auto intro!: essentialI) moreover from \ a have "a \ TER\<^bsub>\ \ f\<^esub> g" by(auto simp add: SAT.A SINK_plus_current) ultimately have "a \ \\<^bsub>\ \ f\<^esub> (TER\<^bsub>\ \ f\<^esub> g)" by blast thus False using n\ by contradiction qed show "d_OUT ?fg a < weight \ a" using a wa currentD_finite_OUT[OF f, of a] by(simp add: d_OUT_def less_diff_eq_ennreal less_top add.commute nn_integral_add) qed ultimately have "hindered \" by(blast intro: hindered.intros) with unhindered show False by contradiction qed lemma loose_minus_web: assumes unhindered: "\ hindered \" and f: "current \ f" and w: "wave \ f" and maximal: "\w. \ current \ w; wave \ w; f \ w \ \ f = w" shows "loose (\ \ f)" (is "loose ?\") proof fix g assume g: "current ?\ g" and w': "wave ?\ g" let ?g = "plus_current f g" from f g have "current \ ?g" by(rule current_plus_current_minus) moreover from f w g w' have "wave \ ?g" by(rule wave_plus_current_minus) moreover have "f \ ?g" by(clarsimp simp add: le_fun_def) ultimately have eq: "f = ?g" by(rule maximal) have "g e = 0" for e proof(cases e) case (Pair x y) have "f e \ d_OUT f x" unfolding d_OUT_def Pair by(rule nn_integral_ge_point) simp also have "\ \ weight \ x" by(rule currentD_weight_OUT[OF f]) also have "\ < \" by(simp add: less_top[symmetric]) finally show "g e = 0" using Pair eq[THEN fun_cong, of e] by(cases "f e" "g e" rule: ennreal2_cases)(simp_all add: fun_eq_iff) qed thus "g = (\_. 0)" by(simp add: fun_eq_iff) next show "\ hindrance ?\ zero_current" using unhindered proof(rule contrapos_nn) assume "hindrance ?\ zero_current" then obtain x where a: "x \ A ?\" and \: "x \ \\<^bsub>?\\<^esub> (TER\<^bsub>?\\<^esub> zero_current)" and weight: "d_OUT zero_current x < weight ?\ x" by cases have "hindrance \ f" proof show a': "x \ A \" using a by simp with weight show "d_OUT f x < weight \ x" by(simp add: less_diff_eq_ennreal less_top[symmetric] currentD_finite_OUT[OF f]) show "x \ \ (TER f)" proof assume "x \ \ (TER f)" then obtain p y where p: "path \ x p y" and y: "y \ B \" and bypass: "\z. z \ set p \ z \ RF (TER f)" by(rule \_E_RF) auto from p a' y disjoint have py: "p = [y]" by(cases)(auto 4 3 elim: rtrancl_path.cases dest: bipartite_E) hence "y \ (TER f)" using bypass[of y] by(auto intro: roofed_greaterI) hence "weight ?\ y > 0" using currentD_OUT[OF f y] disjoint y by(auto simp add: SINK.simps SAT.simps diff_gr0_ennreal) hence "y \ TER\<^bsub>?\\<^esub> zero_current" using y disjoint by(auto) hence "essential ?\ (B ?\) (TER\<^bsub>?\\<^esub> zero_current) x" using p y py by(auto intro!: essentialI) with a have "x \ \\<^bsub>?\\<^esub> (TER\<^bsub>?\\<^esub> zero_current)" by simp with \ show False by contradiction qed qed thus "hindered \" using f w .. qed qed lemma weight_minus_web: assumes f: "current \ f" shows "weight (\ \ f) x = (if x \ A \ then weight \ x - d_OUT f x else weight \ x - d_IN f x)" proof(cases "x \ B \") case True with currentD_OUT[OF f True] disjoint show ?thesis by auto next case False hence "d_IN f x = 0" "d_OUT f x = 0" if "x \ A \" using currentD_outside_OUT[OF f, of x] currentD_outside_IN[OF f, of x] bipartite_V that by auto then show ?thesis by simp qed lemma (in -) separating_minus_web [simp]: "separating_gen (G \ f) = separating_gen G" by(simp add: separating_gen.simps fun_eq_iff) lemma current_minus: assumes f: "current \ f" and g: "current \ g" and le: "\e. g e \ f e" shows "current (\ \ g) (f - g)" proof - interpret \: countable_bipartite_web "\ \ g" using g by(rule countable_bipartite_web_minus_web) note [simp del] = minus_web_sel(2) and [simp] = weight_minus_web[OF g] show ?thesis proof show "d_OUT (f - g) x \ weight (\ \ g) x" for x unfolding fun_diff_def using currentD_weight_OUT[OF f, of x] currentD_weight_IN[OF g, of x] by(subst d_OUT_diff)(simp_all add: le currentD_finite_OUT[OF g] currentD_OUT'[OF f] currentD_OUT'[OF g] ennreal_minus_mono) show "d_IN (f - g) x \ weight (\ \ g) x" for x unfolding fun_diff_def using currentD_weight_IN[OF f, of x] currentD_weight_OUT[OF g, of x] by(subst d_IN_diff)(simp_all add: le currentD_finite_IN[OF g] currentD_IN[OF f] currentD_IN[OF g] ennreal_minus_mono) show "(f - g) e = 0" if "e \ \<^bold>E\<^bsub>\ \ g\<^esub>" for e using that currentD_outside'[OF f] currentD_outside'[OF g] by simp qed qed lemma assumes w: "wave \ f" and g: "current \ g" and le: "\e. g e \ f e" shows wave_minus: "wave (\ \ g) (f - g)" and TER_minus: "TER f \ TER\<^bsub>\ \ g\<^esub> (f - g)" proof have "x \ SINK f \ x \ SINK (f - g)" for x using d_OUT_mono[of g _ f, OF le, of x] by(auto simp add: SINK.simps fun_diff_def d_OUT_diff le currentD_finite_OUT[OF g]) moreover have "x \ SAT \ f \ x \ SAT (\ \ g) (f - g)" for x by(auto simp add: SAT.simps currentD_OUT'[OF g] fun_diff_def d_IN_diff le currentD_finite_IN[OF g] ennreal_minus_mono) ultimately show TER: "TER f \ TER\<^bsub>\ \ g\<^esub> (f - g)" by(auto) from w have "separating \ (TER f)" by(rule waveD_separating) thus "separating (\ \ g) (TER\<^bsub>\ \ g\<^esub> (f - g))" using TER by(simp add: separating_weakening) fix x assume "x \ RF\<^bsub>\ \ g\<^esub> (TER\<^bsub>\ \ g\<^esub> (f - g))" hence "x \ RF (TER f)" using TER by(auto intro: in_roofed_mono) hence "d_OUT f x = 0" by(rule waveD_OUT[OF w]) moreover have "0 \ f e" for e using le[of e] by simp ultimately show "d_OUT (f - g) x = 0" unfolding d_OUT_def by(simp add: nn_integral_0_iff emeasure_count_space_eq_0) qed lemma (in -) essential_minus_web [simp]: "essential (\ \ f) = essential \" by(simp add: essential_def fun_eq_iff) lemma (in -) RF_in_essential: fixes B shows "essential \ B S x \ x \ roofed_gen \ B S \ x \ S" by(auto intro: roofed_greaterI elim!: essentialE_RF dest: roofedD) lemma (in -) d_OUT_fun_upd: assumes "f (x, y) \ \" "f (x, y) \ 0" "k \ \" "k \ 0" shows "d_OUT (f((x, y) := k)) x' = (if x = x' then d_OUT f x - f (x, y) + k else d_OUT f x')" (is "?lhs = ?rhs") proof(cases "x = x'") case True have "?lhs = (\\<^sup>+ y'. f (x, y') + k * indicator {y} y') - (\\<^sup>+ y'. f (x, y') * indicator {y} y')" unfolding d_OUT_def using assms True by(subst nn_integral_diff[symmetric]) (auto intro!: nn_integral_cong simp add: AE_count_space split: split_indicator) also have "(\\<^sup>+ y'. f (x, y') + k * indicator {y} y') = d_OUT f x + (\\<^sup>+ y'. k * indicator {y} y')" unfolding d_OUT_def using assms by(subst nn_integral_add[symmetric]) (auto intro!: nn_integral_cong split: split_indicator) also have "\ - (\\<^sup>+ y'. f (x, y') * indicator {y} y') = ?rhs" using True assms by(subst diff_add_assoc2_ennreal[symmetric])(auto simp add: d_OUT_def intro!: nn_integral_ge_point) finally show ?thesis . qed(simp add: d_OUT_def) lemma unhindered_saturate1: \ \Lemma 6.10\ assumes unhindered: "\ hindered \" and a: "a \ A \" shows "\f. current \ f \ d_OUT f a = weight \ a \ \ hindered (\ \ f)" proof - from a A_vertex have a_vertex: "vertex \ a" by auto from unhindered have "\ hindrance \ zero_current" by(auto intro!: hindered.intros simp add: ) then have a_\: "a \ \ (A \)" if "weight \ a > 0" proof(rule contrapos_np) assume "a \ \ (A \)" with a have "\ essential \ (B \) (A \) a" by simp hence "\ essential \ (B \) (A \ \ {x. weight \ x \ 0}) a" by(rule contrapos_nn)(erule essential_mono; simp) with a that show "hindrance \ zero_current" by(auto intro: hindrance) qed define F where "F = (\(\, h :: 'v current). plus_current \ h)" have F_simps: "F (\, h) = plus_current \ h" for \ h by(simp add: F_def) define Fld where "Fld = {(\, h). current \ \ \ (\x. x \ a \ d_OUT \ x = 0) \ current (\ \ \) h \ wave (\ \ \) h \ \ hindered (\ \ F (\, h))}" define leq where "leq = restrict_rel Fld {(f, f'). f \ f'}" have Fld: "Field leq = Fld" by(auto simp add: leq_def) have F_I [intro?]: "(\, h) \ Field leq" if "current \ \" and "\x. x \ a \ d_OUT \ x = 0" and "current (\ \ \) h" and "wave (\ \ \) h" and "\ hindered (\ \ F (\, h))" for \ h using that by(simp add: Fld Fld_def) have \_curr: "current \ \" if "(\, h) \ Field leq" for \ h using that by(simp add: Fld Fld_def) have OUT_\: "\x. x \ a \ d_OUT \ x = 0" if "(\, h) \ Field leq" for \ h using that by(simp add: Fld Fld_def) have h: "current (\ \ \) h" if "(\, h) \ Field leq" for \ h using that by(simp add: Fld Fld_def) have h_w: "wave (\ \ \) h" if "(\, h) \ Field leq" for \ h using that by(simp add: Fld Fld_def) have unhindered': "\ hindered (\ \ F \h)" if "\h \ Field leq" for \h using that by(simp add: Fld Fld_def split: prod.split_asm) have f: "current \ (F \h)" if "\h \ Field leq" for \h using \_curr h that by(cases \h)(simp add: F_simps current_plus_current_minus) have out_\: "\ (x, y) = 0" if "(\, h) \ Field leq" "x \ a" for \ h x y proof(rule antisym) have "\ (x, y) \ d_OUT \ x" unfolding d_OUT_def by(rule nn_integral_ge_point) simp with OUT_\[OF that] show "\ (x, y) \ 0" by simp qed simp have IN_\: "d_IN \ x = \ (a, x)" if "(\, h) \ Field leq" for \ h x proof(rule trans) show "d_IN \ x = (\\<^sup>+ y. \ (y, x) * indicator {a} y)" unfolding d_IN_def by(rule nn_integral_cong)(simp add: out_\[OF that] split: split_indicator) qed(simp add: max_def \_curr[OF that]) have leqI: "((\, h), (\', h')) \ leq" if "\ \ \'" "h \ h'" "(\, h) \ Field leq" "(\', h') \ Field leq" for \ h \' h' using that unfolding Fld by(simp add: leq_def in_restrict_rel_iff) have chain_Field: "Sup M \ Field leq" if M: "M \ Chains leq" and nempty: "M \ {}" for M unfolding Sup_prod_def proof from nempty obtain \ h where in_M: "(\, h) \ M" by auto with M have Field: "(\, h) \ Field leq" by(rule Chains_FieldD) from M have chain: "Complete_Partial_Order.chain (\\ \'. (\, \') \ leq) M" by(intro Chains_into_chain) simp hence chain': "Complete_Partial_Order.chain (\) M" by(auto simp add: chain_def leq_def in_restrict_rel_iff) hence chain1: "Complete_Partial_Order.chain (\) (fst ` M)" and chain2: "Complete_Partial_Order.chain (\) (snd ` M)" by(rule chain_imageI; auto)+ have outside1: "Sup (fst ` M) (x, y) = 0" if "\ edge \ x y" for x y using that by(auto intro!: SUP_eq_const simp add: nempty dest!: Chains_FieldD[OF M] \_curr currentD_outside) then have "support_flow (Sup (fst ` M)) \ \<^bold>E" by(auto elim!: support_flow.cases intro: ccontr) hence supp_flow1: "countable (support_flow (Sup (fst ` M)))" by(rule countable_subset) simp show SM1: "current \ (Sup (fst ` M))" by(rule current_Sup[OF chain1 _ _ supp_flow1])(auto dest: Chains_FieldD[OF M, THEN \_curr] simp add: nempty) show OUT1_na: "d_OUT (Sup (fst ` M)) x = 0" if "x \ a" for x using that by(subst d_OUT_Sup[OF chain1 _ supp_flow1])(auto simp add: nempty intro!: SUP_eq_const dest: Chains_FieldD[OF M, THEN OUT_\]) interpret SM1: countable_bipartite_web "\ \ Sup (fst ` M)" using SM1 by(rule countable_bipartite_web_minus_web) let ?h = "Sup (snd ` M)" have outside2: "?h (x, y) = 0" if "\ edge \ x y" for x y using that by(auto intro!: SUP_eq_const simp add: nempty dest!: Chains_FieldD[OF M] h currentD_outside) then have "support_flow ?h \ \<^bold>E" by(auto elim!: support_flow.cases intro: ccontr) hence supp_flow2: "countable (support_flow ?h)" by(rule countable_subset) simp have OUT1: "d_OUT (Sup (fst ` M)) x = (SUP (\, h)\M. d_OUT \ x)" for x by (subst d_OUT_Sup [OF chain1 _ supp_flow1]) (simp_all add: nempty split_beta image_comp) have OUT1': "d_OUT (Sup (fst ` M)) x = (if x = a then SUP (\, h)\M. d_OUT \ a else 0)" for x unfolding OUT1 by(auto intro!: SUP_eq_const simp add: nempty OUT_\ dest!: Chains_FieldD[OF M]) have OUT1_le: "(\\h\M. d_OUT (fst \h) x) \ weight \ x" for x using currentD_weight_OUT[OF SM1, of x] OUT1[of x] by(simp add: split_beta) have OUT1_nonneg: "0 \ (\\h\M. d_OUT (fst \h) x)" for x using in_M by(rule SUP_upper2)simp have IN1: "d_IN (Sup (fst ` M)) x = (SUP (\, h)\M. d_IN \ x)" for x by (subst d_IN_Sup [OF chain1 _ supp_flow1]) (simp_all add: nempty split_beta image_comp) have IN1_le: "(\\h\M. d_IN (fst \h) x) \ weight \ x" for x using currentD_weight_IN[OF SM1, of x] IN1[of x] by(simp add: split_beta) have IN1_nonneg: "0 \ (\\h\M. d_IN (fst \h) x)" for x using in_M by(rule SUP_upper2) simp have IN1': "d_IN (Sup (fst ` M)) x = (SUP (\, h)\M. \ (a, x))" for x unfolding IN1 by(rule SUP_cong[OF refl])(auto dest!: Chains_FieldD[OF M] IN_\) have directed: "\\k''\M. F (snd \k) + F (fst \k') \ F (snd \k'') + F (fst \k'')" if mono: "\f g. (\z. f z \ g z) \ F f \ F g" "\k \ M" "\k' \ M" for \k \k' and F :: "_ \ ennreal" using chainD[OF chain that(2-3)] proof cases case left hence "snd \k \ snd \k'" by(simp add: leq_def less_eq_prod_def in_restrict_rel_iff) hence "F (snd \k) + F (fst \k') \ F (snd \k') + F (fst \k')" by(intro add_right_mono mono)(clarsimp simp add: le_fun_def) with that show ?thesis by blast next case right hence "fst \k' \ fst \k" by(simp add: leq_def less_eq_prod_def in_restrict_rel_iff) hence "F (snd \k) + F (fst \k') \ F (snd \k) + F (fst \k)" by(intro add_left_mono mono)(clarsimp simp add: le_fun_def) with that show ?thesis by blast qed have directed_OUT: "\\k''\M. d_OUT (snd \k) x + d_OUT (fst \k') x \ d_OUT (snd \k'') x + d_OUT (fst \k'') x" if "\k \ M" "\k' \ M" for x \k \k' by(rule directed; rule d_OUT_mono that) have directed_IN: "\\k''\M. d_IN (snd \k) x + d_IN (fst \k') x \ d_IN (snd \k'') x + d_IN (fst \k'') x" if "\k \ M" "\k' \ M" for x \k \k' by(rule directed; rule d_IN_mono that) let ?\ = "\ \ Sup (fst ` M)" have hM2: "current ?\ h" if \h: "(\, h) \ M" for \ h proof from \h have Field: "(\, h) \ Field leq" by(rule Chains_FieldD[OF M]) then have H: "current (\ \ \) h" and \_curr': "current \ \" by(rule h \_curr)+ from \_curr' interpret \: countable_bipartite_web "\ \ \" by(rule countable_bipartite_web_minus_web) fix x have "d_OUT h x \ d_OUT ?h x" using \h by(intro d_OUT_mono)(auto intro: SUP_upper2) also have OUT: "\ = (SUP h\snd ` M. d_OUT h x)" using chain2 _ supp_flow2 by(rule d_OUT_Sup)(simp_all add: nempty) also have "\ = \ + (SUP \\fst ` M. d_OUT \ x) - (SUP \\fst ` M. d_OUT \ x)" using OUT1_le[of x] by (intro ennreal_add_diff_cancel_right[symmetric] neq_top_trans[OF weight_finite, of _ x]) (simp add: image_comp) also have "\ = (SUP (\, k)\M. d_OUT k x + d_OUT \ x) - (SUP \\fst ` M. d_OUT \ x)" unfolding split_def by (subst SUP_add_directed_ennreal[OF directed_OUT]) (simp_all add: image_comp) also have "(SUP (\, k)\M. d_OUT k x + d_OUT \ x) \ weight \ x" apply(clarsimp dest!: Chains_FieldD[OF M] intro!: SUP_least) subgoal premises that for \ h using currentD_weight_OUT[OF h[OF that], of x] currentD_weight_OUT[OF \_curr[OF that], of x] countable_bipartite_web_minus_web[OF \_curr, THEN countable_bipartite_web.currentD_OUT', OF that h[OF that], where x=x] by (auto simp add: ennreal_le_minus_iff split: if_split_asm) done also have "(SUP \\fst ` M. d_OUT \ x) = d_OUT (Sup (fst ` M)) x" using OUT1 by (simp add: split_beta image_comp) finally show "d_OUT h x \ weight ?\ x" using \.currentD_OUT'[OF h[OF Field], of x] currentD_weight_IN[OF SM1, of x] by(auto simp add: ennreal_minus_mono) have "d_IN h x \ d_IN ?h x" using \h by(intro d_IN_mono)(auto intro: SUP_upper2) also have IN: "\ = (SUP h\snd ` M. d_IN h x)" using chain2 _ supp_flow2 by(rule d_IN_Sup)(simp_all add: nempty) also have "\ = \ + (SUP \\fst ` M. d_IN \ x) - (SUP \\fst ` M. d_IN \ x)" using IN1_le[of x] by (intro ennreal_add_diff_cancel_right [symmetric] neq_top_trans [OF weight_finite, of _ x]) (simp add: image_comp) also have "\ = (SUP (\, k)\M. d_IN k x + d_IN \ x) - (SUP \\fst ` M. d_IN \ x)" unfolding split_def by (subst SUP_add_directed_ennreal [OF directed_IN]) (simp_all add: image_comp) also have "(SUP (\, k)\M. d_IN k x + d_IN \ x) \ weight \ x" apply(clarsimp dest!: Chains_FieldD[OF M] intro!: SUP_least) subgoal premises that for \ h using currentD_weight_OUT[OF h, OF that, where x=x] currentD_weight_IN[OF h, OF that, where x=x] countable_bipartite_web_minus_web[OF \_curr, THEN countable_bipartite_web.currentD_OUT', OF that h[OF that], where x=x] currentD_OUT'[OF \_curr, OF that, where x=x] currentD_IN[OF \_curr, OF that, of x] currentD_weight_IN[OF \_curr, OF that, where x=x] by (auto simp add: ennreal_le_minus_iff image_comp split: if_split_asm intro: add_increasing2 order_trans [rotated]) done also have "(SUP \\fst ` M. d_IN \ x) = d_IN (Sup (fst ` M)) x" using IN1 by (simp add: split_beta image_comp) finally show "d_IN h x \ weight ?\ x" using currentD_IN[OF h[OF Field], of x] currentD_weight_OUT[OF SM1, of x] by(auto simp add: ennreal_minus_mono) (auto simp add: ennreal_le_minus_iff add_increasing2) show "h e = 0" if "e \ \<^bold>E\<^bsub>?\\<^esub>" for e using currentD_outside'[OF H, of e] that by simp qed from nempty have "snd ` M \ {}" by simp from chain2 this _ supp_flow2 show current: "current ?\ ?h" by(rule current_Sup)(clarify; rule hM2; simp) have wM: "wave ?\ h" if "(\, h) \ M" for \ h proof let ?\' = "\ \ \" have subset: "TER\<^bsub>?\'\<^esub> h \ TER\<^bsub>?\\<^esub> h" using currentD_OUT'[OF SM1] currentD_OUT'[OF \_curr[OF Chains_FieldD[OF M that]]] that by(auto 4 7 elim!: SAT.cases intro: SAT.intros elim!: order_trans[rotated] intro: ennreal_minus_mono d_IN_mono intro!: SUP_upper2 split: if_split_asm) from h_w[OF Chains_FieldD[OF M], OF that] have "separating ?\' (TER\<^bsub>?\'\<^esub> h)" by(rule waveD_separating) then show "separating ?\ (TER\<^bsub>?\\<^esub> h)" using subset by(auto intro: separating_weakening) qed(rule hM2[OF that]) show wave: "wave ?\ ?h" using chain2 \snd ` M \ {}\ _ supp_flow2 by(rule wave_lub)(clarify; rule wM; simp) define f where "f = F (Sup (fst ` M), Sup (snd ` M))" have supp_flow: "countable (support_flow f)" using supp_flow1 supp_flow2 support_flow_plus_current[of "Sup (fst ` M)" ?h] unfolding f_def F_simps by(blast intro: countable_subset) have f_alt: "f = Sup ((\(\, h). plus_current \ h) ` M)" apply (simp add: fun_eq_iff split_def f_def nempty F_def image_comp) apply (subst (1 2) add.commute) apply (subst SUP_add_directed_ennreal) apply (rule directed) apply (auto dest!: Chains_FieldD [OF M]) done have f_curr: "current \ f" unfolding f_def F_simps using SM1 current by(rule current_plus_current_minus) have IN_f: "d_IN f x = d_IN (Sup (fst ` M)) x + d_IN (Sup (snd ` M)) x" for x unfolding f_def F_simps plus_current_def by(rule d_IN_add SM1 current)+ have OUT_f: "d_OUT f x = d_OUT (Sup (fst ` M)) x + d_OUT (Sup (snd ` M)) x" for x unfolding f_def F_simps plus_current_def by(rule d_OUT_add SM1 current)+ show "\ hindered (\ \ f)" (is "\ hindered ?\") \ \Assertion 6.11\ proof assume hindered: "hindered ?\" then obtain g where g: "current ?\ g" and g_w: "wave ?\ g" and hindrance: "hindrance ?\ g" by cases from hindrance obtain z where z: "z \ A \" and z\: "z \ \\<^bsub>?\\<^esub> (TER\<^bsub>?\\<^esub> g)" and OUT_z: "d_OUT g z < weight ?\ z" by cases auto define \ where "\ = weight ?\ z - d_OUT g z" have \_pos: "\ > 0" using OUT_z by(simp add: \_def diff_gr0_ennreal del: minus_web_sel) then have \_finite[simp]: "\ \ \" using z by(simp_all add: \_def) have "\(\, h) \ M. d_OUT f a < d_OUT (plus_current \ h) a + \" proof(rule ccontr) assume "\ ?thesis" hence greater: "d_OUT (plus_current \ h) a + \ \ d_OUT f a" if "(\, h) \ M" for \ h using that by auto have chain'': "Complete_Partial_Order.chain (\) ((\(\, h). plus_current \ h) ` M)" using chain' by(rule chain_imageI)(auto simp add: le_fun_def add_mono) have "d_OUT f a + 0 < d_OUT f a + \" using currentD_finite_OUT[OF f_curr, of a] by (simp add: \_pos) also have "d_OUT f a + \ = (SUP (\, h)\M. d_OUT (plus_current \ h) a) + \" using chain'' nempty supp_flow unfolding f_alt by (subst d_OUT_Sup) (simp_all add: plus_current_def [abs_def] split_def image_comp) also have "\ \ d_OUT f a" unfolding ennreal_SUP_add_left[symmetric, OF nempty] proof(rule SUP_least, clarify) show "d_OUT (plus_current \ h) a + \ \ d_OUT f a" if "(\, h) \ M" for \ h using greater[OF that] currentD_finite_OUT[OF Chains_FieldD[OF M that, THEN f], of a] by(auto simp add: ennreal_le_minus_iff add.commute F_def) qed finally show False by simp qed then obtain \ h where hM: "(\, h) \ M" and close: "d_OUT f a < d_OUT (plus_current \ h) a + \" by blast have Field: "(\, h) \ Field leq" using hM by(rule Chains_FieldD[OF M]) then have \: "current \ \" and unhindered_h: "\ hindered (\ \ F (\, h))" and h_curr: "current (\ \ \) h" and h_w: "wave (\ \ \) h" and OUT_\: "x \ a \ d_OUT \ x = 0" for x by(rule \_curr OUT_\ h h_w unhindered')+ let ?\h = "plus_current \ h" have \h_curr: "current \ ?\h" using Field unfolding F_simps[symmetric] by(rule f) interpret h: countable_bipartite_web "\ \ \" using \ by(rule countable_bipartite_web_minus_web) interpret \h: countable_bipartite_web "\ \ ?\h" using \h_curr by(rule countable_bipartite_web_minus_web) note [simp del] = minus_web_sel(2) and [simp] = weight_minus_web[OF \h_curr] weight_minus_web[OF SM1, simplified] define k where "k e = Sup (fst ` M) e - \ e" for e have k_simps: "k (x, y) = Sup (fst ` M) (x, y) - \ (x, y)" for x y by(simp add: k_def) have k_alt: "k (x, y) = (if x = a \ edge \ x y then Sup (fst ` M) (a, y) - \ (a, y) else 0)" for x y by (cases "x = a") (auto simp add: k_simps out_\ [OF Field] currentD_outside [OF \] image_comp intro!: SUP_eq_const [OF nempty] dest!: Chains_FieldD [OF M] intro: currentD_outside [OF \_curr] out_\) have OUT_k: "d_OUT k x = (if x = a then d_OUT (Sup (fst ` M)) a - d_OUT \ a else 0)" for x proof - have "d_OUT k x = (if x = a then (\\<^sup>+ y. Sup (fst ` M) (a, y) - \ (a, y)) else 0)" using currentD_outside[OF SM1] currentD_outside[OF \] by(auto simp add: k_alt d_OUT_def intro!: nn_integral_cong) also have "(\\<^sup>+ y. Sup (fst ` M) (a, y) - \ (a, y)) = d_OUT (Sup (fst `M)) a - d_OUT \ a" using currentD_finite_OUT[OF \, of a] hM unfolding d_OUT_def by(subst nn_integral_diff[symmetric])(auto simp add: AE_count_space intro!: SUP_upper2) finally show ?thesis . qed have IN_k: "d_IN k y = (if edge \ a y then Sup (fst ` M) (a, y) - \ (a, y) else 0)" for y proof - have "d_IN k y = (\\<^sup>+ x. (if edge \ x y then Sup (fst ` M) (a, y) - \ (a, y) else 0) * indicator {a} x)" unfolding d_IN_def by(rule nn_integral_cong)(auto simp add: k_alt outgoing_def split: split_indicator) also have "\ = (if edge \ a y then Sup (fst ` M) (a, y) - \ (a, y) else 0)" using hM by(auto simp add: max_def intro!: SUP_upper2) finally show ?thesis . qed have OUT_\h: "d_OUT ?\h x = d_OUT \ x + d_OUT h x" for x unfolding plus_current_def by(rule d_OUT_add)+ have IN_\h: "d_IN ?\h x = d_IN \ x + d_IN h x" for x unfolding plus_current_def by(rule d_IN_add)+ have OUT1_le': "d_OUT (Sup (fst`M)) x \ weight \ x" for x using OUT1_le[of x] unfolding OUT1 by (simp add: split_beta') have k: "current (\ \ ?\h) k" proof fix x show "d_OUT k x \ weight (\ \ ?\h) x" using a OUT1_na[of x] currentD_weight_OUT[OF hM2[OF hM], of x] currentD_weight_IN[OF \h_curr, of x] currentD_weight_IN[OF \, of x] OUT1_le'[of x] apply(auto simp add: diff_add_eq_diff_diff_swap_ennreal diff_add_assoc2_ennreal[symmetric] OUT_k OUT_\ OUT_\h IN_\h currentD_OUT'[OF \] IN_\[OF Field] h.currentD_OUT'[OF h_curr]) apply(subst diff_diff_commute_ennreal) apply(intro ennreal_minus_mono) apply(auto simp add: ennreal_le_minus_iff ac_simps less_imp_le OUT1) done have *: "(\xa\M. fst xa (a, x)) \ d_IN (Sup (fst`M)) x" unfolding IN1 by (intro SUP_subset_mono) (auto simp: split_beta' d_IN_ge_point) also have "\ \ weight \ x" using IN1_le[of x] IN1 by (simp add: split_beta') finally show "d_IN k x \ weight (\ \ ?\h) x" using currentD_weight_IN[OF \h_curr, of x] currentD_weight_OUT[OF \h_curr, of x] currentD_weight_IN[OF hM2[OF hM], of x] IN_\[OF Field, of x] * apply (auto simp add: IN_k outgoing_def IN_\h IN_\ A_in diff_add_eq_diff_diff_swap_ennreal) apply (subst diff_diff_commute_ennreal) apply (intro ennreal_minus_mono[OF _ order_refl]) apply (auto simp add: ennreal_le_minus_iff ac_simps image_comp intro: order_trans add_mono) done show "k e = 0" if "e \ \<^bold>E\<^bsub>\ \ ?\h\<^esub>" for e using that by (cases e) (simp add: k_alt) qed define q where "q = (\\<^sup>+ y\B (\ \ ?\h). d_IN k y - d_OUT k y)" have q_alt: "q = (\\<^sup>+ y\- A (\ \ ?\h). d_IN k y - d_OUT k y)" using disjoint by(auto simp add: q_def nn_integral_count_space_indicator currentD_outside_OUT[OF k] currentD_outside_IN[OF k] not_vertex split: split_indicator intro!: nn_integral_cong) have q_simps: "q = d_OUT (Sup (fst ` M)) a - d_OUT \ a" proof - have "q = (\\<^sup>+ y. d_IN k y)" using a IN1 OUT1 OUT1_na unfolding q_alt by(auto simp add: nn_integral_count_space_indicator OUT_k IN_\[OF Field] OUT_\ currentD_outside[OF \] outgoing_def no_loop A_in IN_k intro!: nn_integral_cong split: split_indicator) also have "\ = d_OUT (Sup (fst ` M)) a - d_OUT \ a" using currentD_finite_OUT[OF \, of a] hM currentD_outside[OF SM1] currentD_outside[OF \] by(subst d_OUT_diff[symmetric])(auto simp add: d_OUT_def IN_k intro!: SUP_upper2 nn_integral_cong) finally show ?thesis . qed have q_finite: "q \ \" using currentD_finite_OUT[OF SM1, of a] by(simp add: q_simps) have q_nonneg: "0 \ q" using hM by(auto simp add: q_simps intro!: d_OUT_mono SUP_upper2) have q_less_\: "q < \" using close unfolding q_simps \_def OUT_\h OUT_f proof - let ?F = "d_OUT (Sup (fst`M)) a" and ?S = "d_OUT (Sup (snd`M)) a" and ?\ = "d_OUT \ a" and ?h = "d_OUT h a" and ?w = "weight (\ \ f) z - d_OUT g z" have "?F + ?h \ ?F + ?S" using hM by (auto intro!: add_mono d_OUT_mono SUP_upper2) also assume "?F + ?S < ?\ + ?h + ?w" finally have "?h + ?F < ?h + (?w + ?\)" by (simp add: ac_simps) then show "?F - ?\ < ?w" using currentD_finite_OUT[OF \, of a] hM unfolding ennreal_add_left_cancel_less by (subst minus_less_iff_ennreal) (auto intro!: d_OUT_mono SUP_upper2 simp: less_top) qed define g' where "g' = plus_current g (Sup (snd ` M) - h)" have g'_simps: "g' e = g e + Sup (snd ` M) e - h e" for e using hM by(auto simp add: g'_def intro!: add_diff_eq_ennreal intro: SUP_upper2) have OUT_g': "d_OUT g' x = d_OUT g x + (d_OUT (Sup (snd ` M)) x - d_OUT h x)" for x unfolding g'_simps[abs_def] using \h.currentD_finite_OUT[OF k] hM h.currentD_finite_OUT[OF h_curr] hM apply(subst d_OUT_diff) apply(auto simp add: add_diff_eq_ennreal[symmetric] k_simps intro: add_increasing intro!: SUP_upper2) apply(subst d_OUT_add) apply(auto simp add: add_diff_eq_ennreal[symmetric] k_simps intro: add_increasing intro!:) apply(simp add: add_diff_eq_ennreal SUP_apply[abs_def]) apply(auto simp add: g'_def image_comp intro!: add_diff_eq_ennreal[symmetric] d_OUT_mono intro: SUP_upper2) done have IN_g': "d_IN g' x = d_IN g x + (d_IN (Sup (snd ` M)) x - d_IN h x)" for x unfolding g'_simps[abs_def] using \h.currentD_finite_IN[OF k] hM h.currentD_finite_IN[OF h_curr] hM apply(subst d_IN_diff) apply(auto simp add: add_diff_eq_ennreal[symmetric] k_simps intro: add_increasing intro!: SUP_upper2) apply(subst d_IN_add) apply(auto simp add: add_diff_eq_ennreal[symmetric] k_simps intro: add_increasing intro!: SUP_upper) apply(auto simp add: g'_def SUP_apply[abs_def] image_comp intro!: add_diff_eq_ennreal[symmetric] d_IN_mono intro: SUP_upper2) done have h': "current (\ \ Sup (fst ` M)) h" using hM by(rule hM2) let ?\ = "\ \ ?\h \ k" interpret \: web ?\ using k by(rule \h.web_minus_web) note [simp] = \h.weight_minus_web[OF k] h.weight_minus_web[OF h_curr] weight_minus_web[OF f_curr] SM1.weight_minus_web[OF h', simplified] interpret \: countable_bipartite_web "\ \ f" using f_curr by(rule countable_bipartite_web_minus_web) have *: "\ \ f = \ \ Sup (fst ` M) \ Sup (snd ` M)" unfolding f_def F_simps using SM1 current by(rule minus_plus_current) have OUT_\k: "d_OUT (Sup (fst ` M)) x = d_OUT \ x + d_OUT k x" for x using OUT1'[of x] currentD_finite_OUT[OF \] hM by(auto simp add: OUT_k OUT_\ add_diff_self_ennreal SUP_upper2) have IN_\k: "d_IN (Sup (fst ` M)) x = d_IN \ x + d_IN k x" for x using IN1'[of x] currentD_finite_IN[OF \] currentD_outside[OF \] currentD_outside[OF \_curr] by(auto simp add: IN_k IN_\[OF Field] add_diff_self_ennreal split_beta nempty image_comp dest!: Chains_FieldD[OF M] intro!: SUP_eq_const intro: SUP_upper2[OF hM]) have **: "?\ = \ \ Sup (fst ` M) \ h" proof(rule web.equality) show "weight ?\ = weight (\ \ Sup (fst ` M) \ h)" using OUT_\k OUT_\h currentD_finite_OUT[OF \] IN_\k IN_\h currentD_finite_IN[OF \] by(auto simp add: diff_add_eq_diff_diff_swap_ennreal diff_diff_commute_ennreal) qed simp_all have g'_alt: "g' = plus_current (Sup (snd ` M)) g - h" by(simp add: fun_eq_iff g'_simps add_diff_eq_ennreal add.commute) have "current (\ \ Sup (fst ` M)) (plus_current (Sup (snd ` M)) g)" using current g unfolding * by(rule SM1.current_plus_current_minus) hence g': "current ?\ g'" unfolding * ** g'_alt using hM2[OF hM] by(rule SM1.current_minus)(auto intro!: add_increasing2 SUP_upper2 hM) have "wave (\ \ Sup (fst ` M)) (plus_current (Sup (snd ` M)) g)" using current wave g g_w unfolding * by(rule SM1.wave_plus_current_minus) then have g'_w: "wave ?\ g'" unfolding * ** g'_alt using hM2[OF hM] by(rule SM1.wave_minus)(auto intro!: add_increasing2 SUP_upper2 hM) have "hindrance_by ?\ g' q" proof show "z \ A ?\" using z by simp show "z \ \\<^bsub>?\\<^esub> (TER\<^bsub>?\\<^esub> g')" proof assume "z \ \\<^bsub>?\\<^esub> (TER\<^bsub>?\\<^esub> g')" hence OUT_z: "d_OUT g' z = 0" and ess: "essential ?\ (B \) (TER\<^bsub>?\\<^esub> g') z" by(simp_all add: SINK.simps) from ess obtain p y where p: "path \ z p y" and y: "y \ B \" and bypass: "\z. z \ set p \ z \ RF (TER\<^bsub>?\\<^esub> g')" by(rule essentialE_RF) auto from y have y': "y \ A \" using disjoint by blast from p z y obtain py: "p = [y]" and edge: "edge \ z y" using disjoint by(cases)(auto 4 3 elim: rtrancl_path.cases dest: bipartite_E) hence yRF: "y \ RF (TER\<^bsub>?\\<^esub> g')" using bypass[of y] by(auto) with wave_not_RF_IN_zero[OF g' g'_w, of y] have IN_g'_y: "d_IN g' y = 0" by(auto intro: roofed_greaterI) with yRF y y' have w_y: "weight ?\ y > 0" using currentD_OUT[OF g', of y] by(auto simp add: RF_in_B currentD_SAT[OF g'] SINK.simps zero_less_iff_neq_zero) have "y \ SAT (\ \ f) g" proof assume "y \ SAT (\ \ f) g" with y disjoint have IN_g_y: "d_IN g y = weight (\ \ f) y" by(auto simp add: currentD_SAT[OF g]) have "0 < weight \ y - d_IN (\x\M. fst x) y - d_IN h y" using y' w_y unfolding ** by auto have "d_IN g' y > 0" using y' w_y hM unfolding ** apply(simp add: IN_g' IN_f IN_g_y diff_add_eq_diff_diff_swap_ennreal) apply(subst add_diff_eq_ennreal) apply(auto intro!: SUP_upper2 d_IN_mono simp: diff_add_self_ennreal diff_gt_0_iff_gt_ennreal) done with IN_g'_y show False by simp qed then have "y \ TER\<^bsub>\ \ f\<^esub> g" by simp with p y py have "essential \ (B \) (TER\<^bsub>\ \ f\<^esub> g) z" by(auto intro: essentialI) moreover with z waveD_separating[OF g_w, THEN separating_RF_A] have "z \ \\<^bsub>?\\<^esub> (TER\<^bsub>?\\<^esub> g)" by(auto simp add: RF_in_essential) with z\ show False by contradiction qed have "\ \ weight ?\ z - d_OUT g' z" unfolding ** OUT_g' using z apply (simp add: \_def OUT_f diff_add_eq_diff_diff_swap_ennreal) apply (subst (5) diff_diff_commute_ennreal) apply (rule ennreal_minus_mono[OF _ order_refl]) apply (auto simp add: ac_simps diff_add_eq_diff_diff_swap_ennreal[symmetric] add_diff_self_ennreal image_comp intro!: ennreal_minus_mono[OF order_refl] SUP_upper2[OF hM] d_OUT_mono) done then show q_z: "q < weight ?\ z - d_OUT g' z" using q_less_\ by simp then show "d_OUT g' z < weight ?\ z" using q_nonneg z by(auto simp add: less_diff_eq_ennreal less_top[symmetric] ac_simps \.currentD_finite_OUT[OF g'] intro: le_less_trans[rotated] add_increasing) qed then have hindered_by: "hindered_by (\ \ ?\h \ k) q" using g' g'_w by(rule hindered_by.intros) then have "hindered (\ \ ?\h)" using q_finite unfolding q_def by -(rule \h.hindered_reduce_current[OF k]) with unhindered_h show False unfolding F_simps by contradiction qed qed define sat where "sat = (\(\, h). let f = F (\, h); k = SOME k. current (\ \ f) k \ wave (\ \ f) k \ (\k'. current (\ \ f) k' \ wave (\ \ f) k' \ k \ k' \ k = k') in if d_OUT (plus_current f k) a < weight \ a then let \ = \ \ f \ k; y = SOME y. y \ \<^bold>O\<^bold>U\<^bold>T\<^bsub>\\<^esub> a \ weight \ y > 0; \ = SOME \. \ > 0 \ \ < enn2real (min (weight \ a) (weight \ y)) \ \ hindered (reduce_weight \ y \) in (plus_current \ (zero_current((a, y) := \)), plus_current h k) else (\, h))" have zero: "(zero_current, zero_current) \ Field leq" by(rule F_I)(simp_all add: unhindered F_def) have a_TER: "a \ TER\<^bsub>\ \ F \h\<^esub> k" if that: "\h \ Field leq" and k: "current (\ \ F \h) k" and k_w: "wave (\ \ F \h) k" and less: "d_OUT (plus_current (F \h) k) a < weight \ a" for \h k proof(rule ccontr) assume "\ ?thesis" hence \: "a \ \\<^bsub>\ \ F \h\<^esub> (TER\<^bsub>\ \ F \h\<^esub> k)" by auto from that have f: "current \ (F \h)" and unhindered: "\ hindered (\ \ F \h)" by(cases \h; simp add: f unhindered'; fail)+ from less have "d_OUT k a < weight (\ \ F \h) a" using a currentD_finite_OUT[OF f, of a] by(simp add: d_OUT_def nn_integral_add less_diff_eq_ennreal add.commute less_top[symmetric]) with _ \ have "hindrance (\ \ F \h) k" by(rule hindrance)(simp add: a) then have "hindered (\ \ F \h)" using k k_w .. with unhindered show False by contradiction qed note minus_web_sel(2)[simp del] let ?P_y = "\\h k y. y \ \<^bold>O\<^bold>U\<^bold>T\<^bsub>\ \ F \h \ k\<^esub> a \ weight (\ \ F \h \ k) y > 0" have Ex_y: "Ex (?P_y \h k)" if that: "\h \ Field leq" and k: "current (\ \ F \h) k" and k_w: "wave (\ \ F \h) k" and less: "d_OUT (plus_current (F \h) k) a < weight \ a" for \h k proof(rule ccontr) let ?\ = "\ \ F \h \ k" assume *: "\ ?thesis" interpret \: countable_bipartite_web "\ \ F \h" using f[OF that] by(rule countable_bipartite_web_minus_web) note [simp] = weight_minus_web[OF f[OF that]] \.weight_minus_web[OF k] have "hindrance ?\ zero_current" proof show "a \ A ?\" using a by simp show "a \ \\<^bsub>?\\<^esub> (TER\<^bsub>?\\<^esub> zero_current)" (is "a \ \\<^bsub>_\<^esub> ?TER") proof assume "a \ \\<^bsub>?\\<^esub> ?TER" then obtain p y where p: "path ?\ a p y" and y: "y \ B ?\" and bypass: "\z. z \ set p \ z \ RF\<^bsub>?\\<^esub> ?TER" by(rule \_E_RF)(auto) from a p y disjoint have Nil: "p \ []" by(auto simp add: rtrancl_path_simps) hence "edge ?\ a (p ! 0)" "p ! 0 \ RF\<^bsub>?\\<^esub> ?TER" using rtrancl_path_nth[OF p, of 0] bypass by auto with * show False by(auto simp add: not_less outgoing_def intro: roofed_greaterI) qed have "d_OUT (plus_current (F \h) k) x = d_OUT (F \h) x + d_OUT k x" for x by(simp add: d_OUT_def nn_integral_add) then show "d_OUT zero_current a < weight ?\ a" using less a_TER[OF that k k_w less] a by(simp add: SINK.simps diff_gr0_ennreal) qed hence "hindered ?\" by(auto intro!: hindered.intros order_trans[OF currentD_weight_OUT[OF k]] order_trans[OF currentD_weight_IN[OF k]]) moreover have "\ hindered ?\" using unhindered'[OF that] k k_w by(rule \.unhindered_minus_web) ultimately show False by contradiction qed have increasing: "\h \ sat \h \ sat \h \ Field leq" if "\h \ Field leq" for \h proof(cases \h) case (Pair \ h) with that have that: "(\, h) \ Field leq" by simp have f: "current \ (F (\, h))" and unhindered: "\ hindered (\ \ F (\, h))" and \: "current \ \" and h: "current (\ \ \) h" and h_w: "wave (\ \ \) h" and OUT_\: "x \ a \ d_OUT \ x = 0" for x using that by(rule f unhindered' \_curr OUT_\ h h_w)+ interpret \: countable_bipartite_web "\ \ F (\, h)" using f by(rule countable_bipartite_web_minus_web) note [simp] = weight_minus_web[OF f] let ?P_k = "\k. current (\ \ F (\, h)) k \ wave (\ \ F (\, h)) k \ (\k'. current (\ \ F (\, h)) k' \ wave (\ \ F (\, h)) k' \ k \ k' \ k = k')" define k where "k = Eps ?P_k" have "Ex ?P_k" by(intro ex_maximal_wave)(simp_all) hence "?P_k k" unfolding k_def by(rule someI_ex) hence k: "current (\ \ F (\, h)) k" and k_w: "wave (\ \ F (\, h)) k" and maximal: "\k'. \ current (\ \ F (\, h)) k'; wave (\ \ F (\, h)) k'; k \ k' \ \ k = k'" by blast+ note [simp] = \.weight_minus_web[OF k] let ?fk = "plus_current (F (\, h)) k" have IN_fk: "d_IN ?fk x = d_IN (F (\, h)) x + d_IN k x" for x by(simp add: d_IN_def nn_integral_add) have OUT_fk: "d_OUT ?fk x = d_OUT (F (\, h)) x + d_OUT k x" for x by(simp add: d_OUT_def nn_integral_add) have fk: "current \ ?fk" using f k by(rule current_plus_current_minus) show ?thesis proof(cases "d_OUT ?fk a < weight \ a") case less: True define \ where "\ = \ \ F (\, h) \ k" have B_\ [simp]: "B \ = B \" by(simp add: \_def) have loose: "loose \" unfolding \_def using unhindered k k_w maximal by(rule \.loose_minus_web) interpret \: countable_bipartite_web \ using k unfolding \_def by(rule \.countable_bipartite_web_minus_web) have a_\: "a \ TER\<^bsub>\ \ F (\, h)\<^esub> k" using that k k_w less by(rule a_TER) then have weight_\_a: "weight \ a = weight \ a - d_OUT (F (\, h)) a" using a disjoint by(auto simp add: roofed_circ_def \_def SINK.simps) then have weight_a: "0 < weight \ a" using less a_\ by(simp add: OUT_fk SINK.simps diff_gr0_ennreal) let ?P_y = "\y. y \ \<^bold>O\<^bold>U\<^bold>T\<^bsub>\\<^esub> a \ weight \ y > 0" define y where "y = Eps ?P_y" have "Ex ?P_y" using that k k_w less unfolding \_def by(rule Ex_y) hence "?P_y y" unfolding y_def by(rule someI_ex) hence y_OUT: "y \ \<^bold>O\<^bold>U\<^bold>T\<^bsub>\\<^esub> a" and weight_y: "weight \ y > 0" by blast+ from y_OUT have y_B: "y \ B \" by(auto simp add: outgoing_def \_def dest: bipartite_E) with weight_y have yRF: "y \ RF\<^bsub>\ \ F (\, h)\<^esub> (TER\<^bsub>\ \ F (\, h)\<^esub> k)" unfolding \_def using currentD_OUT[OF k, of y] disjoint by(auto split: if_split_asm simp add: SINK.simps currentD_SAT[OF k] roofed_circ_def RF_in_B \.currentD_finite_IN[OF k]) hence IN_k_y: "d_IN k y = 0" by(rule wave_not_RF_IN_zero[OF k k_w]) define bound where "bound = enn2real (min (weight \ a) (weight \ y))" have bound_pos: "bound > 0" using weight_y weight_a using \.weight_finite by(cases "weight \ a" "weight \ y" rule: ennreal2_cases) (simp_all add: bound_def min_def split: if_split_asm) let ?P_\ = "\\. \ > 0 \ \ < bound \ \ hindered (reduce_weight \ y \)" define \ where "\ = Eps ?P_\" let ?\ = "reduce_weight \ y \" from \.unhinder[OF loose _ weight_y bound_pos] y_B disjoint have "Ex ?P_\" by(auto simp add: \_def) hence "?P_\ \" unfolding \_def by(rule someI_ex) hence \_pos: "0 < \" and \_le_bound: "\ < bound" and unhindered': "\ hindered ?\" by blast+ from \_pos have \_nonneg: "0 \ \" by simp from \_le_bound \_pos have \_le_a: "\ < weight \ a" and \_le_y: "\ < weight \ y" by(cases "weight \ a" "weight \ y" rule: ennreal2_cases; simp add: bound_def min_def ennreal_less_iff split: if_split_asm)+ let ?\ = "\ \ ?fk" interpret \': countable_bipartite_web ?\ by(rule countable_bipartite_web_minus_web fk)+ note [simp] = weight_minus_web[OF fk] let ?g = "zero_current((a, y) := \)" have OUT_g: "d_OUT ?g x = (if x = a then \ else 0)" for x proof(rule trans) show "d_OUT ?g x = (\\<^sup>+ z. (if x = a then \ else 0) * indicator {y} z)" unfolding d_OUT_def by(rule nn_integral_cong) simp show "\ = (if x = a then \ else 0)" using \_pos by(simp add: max_def) qed have IN_g: "d_IN ?g x = (if x = y then \ else 0)" for x proof(rule trans) show "d_IN ?g x = (\\<^sup>+ z. (if x = y then \ else 0) * indicator {a} z)" unfolding d_IN_def by(rule nn_integral_cong) simp show "\ = (if x = y then \ else 0)" using \_pos by(simp add: max_def) qed have g: "current ?\ ?g" proof show "d_OUT ?g x \ weight ?\ x" for x proof(cases "x = a") case False then show ?thesis using currentD_weight_OUT[OF fk, of x] currentD_weight_IN[OF fk, of x] by(auto simp add: OUT_g zero_ennreal_def[symmetric]) next case True then show ?thesis using \_le_a a a_\ \_pos unfolding OUT_g by(simp add: OUT_g \_def SINK.simps OUT_fk split: if_split_asm) qed show "d_IN ?g x \ weight ?\ x" for x proof(cases "x = y") case False then show ?thesis using currentD_weight_OUT[OF fk, of x] currentD_weight_IN[OF fk, of x] by(auto simp add: IN_g zero_ennreal_def[symmetric]) next case True then show ?thesis using \_le_y y_B a_\ \_pos currentD_OUT[OF k, of y] IN_k_y by(simp add: OUT_g \_def SINK.simps OUT_fk IN_fk IN_g split: if_split_asm) qed show "?g e = 0" if "e \ \<^bold>E\<^bsub>?\\<^esub>" for e using y_OUT that by(auto simp add: \_def outgoing_def) qed interpret \'': web "\ \ ?fk \ ?g" using g by(rule \'.web_minus_web) let ?\' = "plus_current \ (zero_current((a, y) := \))" let ?h' = "plus_current h k" have F': "F (?\', ?h') = plus_current (plus_current (F (\, h)) k) (zero_current((a, y) := \))" (is "_ = ?f'") by(auto simp add: F_simps fun_eq_iff add_ac) have sat: "sat (\, h) = (?\', ?h')" using less by(simp add: sat_def k_def \_def Let_def y_def bound_def \_def) have le: "(\, h) \ (?\', ?h')" using \_pos by(auto simp add: le_fun_def add_increasing2 add_increasing) have "current (\ \ \) ((\_. 0)((a, y) := ennreal \))" using g by(rule current_weight_mono)(auto simp add: weight_minus_web[OF \] intro!: ennreal_minus_mono d_OUT_mono d_IN_mono, simp_all add: F_def add_increasing2) with \ have \': "current \ ?\'" by(rule current_plus_current_minus) moreover have eq_0: "d_OUT ?\' x = 0" if "x \ a" for x unfolding plus_current_def using that by(subst d_OUT_add)(simp_all add: \_nonneg d_OUT_fun_upd OUT_\) moreover from \' interpret \': countable_bipartite_web "\ \ ?\'" by(rule countable_bipartite_web_minus_web) from \ interpret \: countable_bipartite_web "\ \ \" by(rule countable_bipartite_web_minus_web) have g': "current (\ \ \) ?g" using g apply(rule current_weight_mono) apply(auto simp add: weight_minus_web[OF \] intro!: ennreal_minus_mono d_OUT_mono d_IN_mono) apply(simp_all add: F_def add_increasing2) done have k': "current (\ \ \ \ h) k" using k unfolding F_simps minus_plus_current[OF \ h] . with h have "current (\ \ \) (plus_current h k)" by(rule \.current_plus_current_minus) hence "current (\ \ \) (plus_current (plus_current h k) ?g)" using g unfolding minus_plus_current[OF f k] unfolding F_simps minus_plus_current[OF \ h] \.minus_plus_current[OF h k', symmetric] by(rule \.current_plus_current_minus) then have "current (\ \ \ \ ?g) (plus_current (plus_current h k) ?g - ?g)" using g' by(rule \.current_minus)(auto simp add: add_increasing) then have h'': "current (\ \ ?\') ?h'" by(rule arg_cong2[where f=current, THEN iffD1, rotated -1]) (simp_all add: minus_plus_current[OF \ g'] fun_eq_iff add_diff_eq_ennreal[symmetric]) moreover have "wave (\ \ ?\') ?h'" proof have "separating (\ \ \) (TER\<^bsub>\ \ \\<^esub> (plus_current h k))" using k k_w unfolding F_simps minus_plus_current[OF \ h] by(intro waveD_separating \.wave_plus_current_minus[OF h h_w]) moreover have "TER\<^bsub>\ \ \\<^esub> (plus_current h k) \ TER\<^bsub>\ \ ?\'\<^esub> (plus_current h k)" by(auto 4 4 simp add: SAT.simps weight_minus_web[OF \] weight_minus_web[OF \'] split: if_split_asm elim: order_trans[rotated] intro!: ennreal_minus_mono d_IN_mono add_increasing2 \_nonneg) ultimately show sep: "separating (\ \ ?\') (TER\<^bsub>\ \ ?\'\<^esub> ?h')" by(simp add: minus_plus_current[OF \ g'] separating_weakening) qed(rule h'') moreover have "\ hindered (\ \ F (?\', ?h'))" using unhindered' proof(rule contrapos_nn) assume "hindered (\ \ F (?\', ?h'))" thus "hindered ?\" proof(rule hindered_mono_web[rotated -1]) show "weight ?\ z = weight (\ \ F (?\', ?h')) z" if "z \ A (\ \ F (?\', ?h'))" for z using that unfolding F' apply(cases "z = y") apply(simp_all add: \_def minus_plus_current[OF fk g] \'.weight_minus_web[OF g] IN_g) apply(simp_all add: plus_current_def d_IN_add diff_add_eq_diff_diff_swap_ennreal currentD_finite_IN[OF f]) done have "y \ a" using y_B a disjoint by auto then show "weight (\ \ F (?\', ?h')) z \ weight ?\ z" if "z \ A (\ \ F (?\', ?h'))" for z using that y_B disjoint \_nonneg unfolding F' apply(cases "z = a") apply(simp_all add: \_def minus_plus_current[OF fk g] \'.weight_minus_web[OF g] OUT_g) apply(auto simp add: plus_current_def d_OUT_add diff_add_eq_diff_diff_swap_ennreal currentD_finite_OUT[OF f]) done qed(simp_all add: \_def) qed ultimately have "(?\', ?h') \ Field leq" by-(rule F_I) with Pair le sat that show ?thesis by(auto) next case False with currentD_weight_OUT[OF fk, of a] have "d_OUT ?fk a = weight \ a" by simp have "sat \h = \h" using False Pair by(simp add: sat_def k_def) thus ?thesis using that Pair by(auto) qed qed have "bourbaki_witt_fixpoint Sup leq sat" using increasing chain_Field unfolding leq_def by(intro bourbaki_witt_fixpoint_restrict_rel)(auto intro: Sup_upper Sup_least) then interpret bourbaki_witt_fixpoint Sup leq sat . define f where "f = fixp_above (zero_current, zero_current)" have Field: "f \ Field leq" using fixp_above_Field[OF zero] unfolding f_def . then have f: "current \ (F f)" and unhindered: "\ hindered (\ \ F f)" by(cases f; simp add: f unhindered'; fail)+ interpret \: countable_bipartite_web "\ \ F f" using f by(rule countable_bipartite_web_minus_web) note [simp] = weight_minus_web[OF f] have Field': "(fst f, snd f) \ Field leq" using Field by simp let ?P_k = "\k. current (\ \ F f) k \ wave (\ \ F f) k \ (\k'. current (\ \ F f) k' \ wave (\ \ F f) k' \ k \ k' \ k = k')" define k where "k = Eps ?P_k" have "Ex ?P_k" by(intro ex_maximal_wave)(simp_all) hence "?P_k k" unfolding k_def by(rule someI_ex) hence k: "current (\ \ F f) k" and k_w: "wave (\ \ F f) k" and maximal: "\k'. \ current (\ \ F f) k'; wave (\ \ F f) k'; k \ k' \ \ k = k'" by blast+ note [simp] = \.weight_minus_web[OF k] let ?fk = "plus_current (F f) k" have IN_fk: "d_IN ?fk x = d_IN (F f) x + d_IN k x" for x by(simp add: d_IN_def nn_integral_add) have OUT_fk: "d_OUT ?fk x = d_OUT (F f) x + d_OUT k x" for x by(simp add: d_OUT_def nn_integral_add) have fk: "current \ ?fk" using f k by(rule current_plus_current_minus) have "d_OUT ?fk a \ weight \ a" proof(rule ccontr) assume "\ ?thesis" hence less: "d_OUT ?fk a < weight \ a" by simp define \ where "\ = \ \ F f \ k" have B_\ [simp]: "B \ = B \" by(simp add: \_def) have loose: "loose \" unfolding \_def using unhindered k k_w maximal by(rule \.loose_minus_web) interpret \: countable_bipartite_web \ using k unfolding \_def by(rule \.countable_bipartite_web_minus_web) have a_\: "a \ TER\<^bsub>\ \ F f\<^esub> k" using Field k k_w less by(rule a_TER) then have "weight \ a = weight \ a - d_OUT (F f) a" using a disjoint by(auto simp add: roofed_circ_def \_def SINK.simps) then have weight_a: "0 < weight \ a" using less a_\ by(simp add: OUT_fk SINK.simps diff_gr0_ennreal) let ?P_y = "\y. y \ \<^bold>O\<^bold>U\<^bold>T\<^bsub>\\<^esub> a \ weight \ y > 0" define y where "y = Eps ?P_y" have "Ex ?P_y" using Field k k_w less unfolding \_def by(rule Ex_y) hence "?P_y y" unfolding y_def by(rule someI_ex) hence "y \ \<^bold>O\<^bold>U\<^bold>T\<^bsub>\\<^esub> a" and weight_y: "weight \ y > 0" by blast+ then have y_B: "y \ B \" by(auto simp add: outgoing_def \_def dest: bipartite_E) define bound where "bound = enn2real (min (weight \ a) (weight \ y))" have bound_pos: "bound > 0" using weight_y weight_a \.weight_finite by(cases "weight \ a" "weight \ y" rule: ennreal2_cases) (simp_all add: bound_def min_def split: if_split_asm) let ?P_\ = "\\. \ > 0 \ \ < bound \ \ hindered (reduce_weight \ y \)" define \ where "\ = Eps ?P_\" from \.unhinder[OF loose _ weight_y bound_pos] y_B disjoint have "Ex ?P_\" by(auto simp add: \_def) hence "?P_\ \" unfolding \_def by(rule someI_ex) hence \_pos: "0 < \" by blast+ let ?f' = "(plus_current (fst f) (zero_current((a, y) := \)), plus_current (snd f) k)" have sat: "?f' = sat f" using less by(simp add: sat_def k_def \_def Let_def y_def bound_def \_def split_def) also have "\ = f" unfolding f_def using fixp_above_unfold[OF zero] by simp finally have "fst ?f' (a, y) = fst f (a, y)" by simp hence "\ = 0" using currentD_finite[OF \_curr[OF Field']] \_pos by(cases "fst f (a, y)") simp_all with \_pos show False by simp qed with currentD_weight_OUT[OF fk, of a] have "d_OUT ?fk a = weight \ a" by simp moreover have "current \ ?fk" using f k by(rule current_plus_current_minus) moreover have "\ hindered (\ \ ?fk)" unfolding minus_plus_current[OF f k] using unhindered k k_w by(rule \.unhindered_minus_web) ultimately show ?thesis by blast qed end subsection \Linkability of unhindered bipartite webs\ context countable_bipartite_web begin theorem unhindered_linkable: assumes unhindered: "\ hindered \" shows "linkable \" proof(cases "A \ = {}") case True thus ?thesis by(auto intro!: exI[where x="zero_current"] linkage.intros simp add: web_flow_iff ) next case nempty: False let ?P = "\f a f'. current (\ \ f) f' \ d_OUT f' a = weight (\ \ f) a \ \ hindered (\ \ f \ f')" define enum where "enum = from_nat_into (A \)" have enum_A: "enum n \ A \" for n using from_nat_into[OF nempty, of n] by(simp add: enum_def) have vertex_enum [simp]: "vertex \ (enum n)" for n using enum_A[of n] A_vertex by blast define f where "f = rec_nat zero_current (\n f. let f' = SOME f'. ?P f (enum n) f' in plus_current f f')" have f_0 [simp]: "f 0 = zero_current" by(simp add: f_def) have f_Suc: "f (Suc n) = plus_current (f n) (Eps (?P (f n) (enum n)))" for n by(simp add: f_def) have f: "current \ (f n)" and sat: "\m. m < n \ d_OUT (f n) (enum m) = weight \ (enum m)" and unhindered: "\ hindered (\ \ f n)" for n proof(induction n) case 0 - { case 1 thus ?case bysimp } + { case 1 thus ?case by simp } { case 2 thus ?case by simp } { case 3 thus ?case using unhindered by simp } next case (Suc n) interpret \: countable_bipartite_web "\ \ f n" using Suc.IH(1) by(rule countable_bipartite_web_minus_web) define f' where "f' = Eps (?P (f n) (enum n))" have "Ex (?P (f n) (enum n))" using Suc.IH(3) by(rule \.unhindered_saturate1)(simp add: enum_A) hence "?P (f n) (enum n) f'" unfolding f'_def by(rule someI_ex) hence f': "current (\ \ f n) f'" and OUT: "d_OUT f' (enum n) = weight (\ \ f n) (enum n)" and unhindered': "\ hindered (\ \ f n \ f')" by blast+ have f_Suc: "f (Suc n) = plus_current (f n) f'" by(simp add: f'_def f_Suc) { case 1 show ?case unfolding f_Suc using Suc.IH(1) f' by(rule current_plus_current_minus) } note f'' = this { case (2 m) have "d_OUT (f (Suc n)) (enum m) \ weight \ (enum m)" using f'' by(rule currentD_weight_OUT) moreover have "weight \ (enum m) \ d_OUT (f (Suc n)) (enum m)" proof(cases "m = n") case True then show ?thesis unfolding f_Suc using OUT True by(simp add: d_OUT_def nn_integral_add enum_A add_diff_self_ennreal less_imp_le) next case False hence "m < n" using 2 by simp thus ?thesis using Suc.IH(2)[OF \m < n\] unfolding f_Suc by(simp add: d_OUT_def nn_integral_add add_increasing2 ) qed ultimately show ?case by(rule antisym) } { case 3 show ?case unfolding f_Suc minus_plus_current[OF Suc.IH(1) f'] by(rule unhindered') } qed interpret \: countable_bipartite_web "\ \ f n" for n using f by(rule countable_bipartite_web_minus_web) have Ex_P: "Ex (?P (f n) (enum n))" for n using unhindered by(rule \.unhindered_saturate1)(simp add: enum_A) have f_mono: "f n \ f (Suc n)" for n using someI_ex[OF Ex_P, of n] by(auto simp add: le_fun_def f_Suc enum_A intro: add_increasing2 dest: ) hence incseq: "incseq f" by(rule incseq_SucI) hence chain: "Complete_Partial_Order.chain (\) (range f)" by(rule incseq_chain_range) define g where "g = Sup (range f)" have "support_flow g \ \<^bold>E" by (auto simp add: g_def support_flow.simps currentD_outside [OF f] image_comp elim: contrapos_pp) then have countable_g: "countable (support_flow g)" by(rule countable_subset) simp with chain _ _ have g: "current \ g" unfolding g_def by(rule current_Sup)(auto simp add: f) moreover have "d_OUT g x = weight \ x" if "x \ A \" for x proof(rule antisym) show "d_OUT g x \ weight \ x" using g by(rule currentD_weight_OUT) have "countable (A \)" using A_vertex by(rule countable_subset) simp from that subset_range_from_nat_into[OF this] obtain n where "x = enum n" unfolding enum_def by blast with sat[of n "Suc n"] have "d_OUT (f (Suc n)) x \ weight \ x" by simp then show "weight \ x \ d_OUT g x" using countable_g unfolding g_def by(subst d_OUT_Sup[OF chain])(auto intro: SUP_upper2) qed ultimately show ?thesis by(auto simp add: web_flow_iff linkage.simps) qed end context countable_web begin theorem loose_linkable: \ \Theorem 6.2\ assumes "loose \" shows "linkable \" proof - interpret bi: countable_bipartite_web "bipartite_web_of \" by(rule countable_bipartite_web_of) have "\ hindered (bipartite_web_of \)" using assms by(rule unhindered_bipartite_web_of) then have "linkable (bipartite_web_of \)" by(rule bi.unhindered_linkable) then show ?thesis by(rule linkable_bipartite_web_ofD) simp qed lemma ex_orthogonal_current: \ \Lemma 4.15\ "\f S. web_flow \ f \ separating \ S \ orthogonal_current \ f S" by(rule ex_orthogonal_current')(rule countable_web.loose_linkable[OF countable_web_quotient_web]) end subsection \Glueing the reductions together\ context countable_network begin context begin qualified lemma max_flow_min_cut': assumes source_in: "\x. \ edge \ x (source \)" and sink_out: "\y. \ edge \ (sink \) y" and undead: "\x y. edge \ x y \ (\z. edge \ y z) \ (\z. edge \ z x)" and source_sink: "\ edge \ (source \) (sink \)" and no_loop: "\x. \ edge \ x x" and capacity_pos: "\e. e \ \<^bold>E \ capacity \ e > 0" shows "\f S. flow \ f \ cut \ S \ orthogonal \ f S" by(rule max_flow_min_cut')(rule countable_web.ex_orthogonal_current[OF countable_web_web_of_network], fact+) qualified lemma max_flow_min_cut'': assumes sink_out: "\y. \ edge \ (sink \) y" and source_in: "\x. \ edge \ x (source \)" and no_loop: "\x. \ edge \ x x" and capacity_pos: "\e. e \ \<^bold>E \ capacity \ e > 0" shows "\f S. flow \ f \ cut \ S \ orthogonal \ f S" proof - interpret antiparallel_edges \ .. interpret \'': countable_network \'' by(rule \''_countable_network) have "\f S. flow \'' f \ cut \'' S \ orthogonal \'' f S" by(rule \''.max_flow_min_cut')(auto simp add: sink_out source_in no_loop capacity_pos elim: edg.cases) then obtain f S where f: "flow \'' f" and cut: "cut \'' S" and ortho: "orthogonal \'' f S" by blast have "flow \ (collect f)" using f by(rule flow_collect) moreover have "cut \ (cut' S)" using cut by(rule cut_cut') moreover have "orthogonal \ (collect f) (cut' S)" using ortho f by(rule orthogonal_cut') ultimately show ?thesis by blast qed qualified lemma max_flow_min_cut''': assumes sink_out: "\y. \ edge \ (sink \) y" and source_in: "\x. \ edge \ x (source \)" and capacity_pos: "\e. e \ \<^bold>E \ capacity \ e > 0" shows "\f S. flow \ f \ cut \ S \ orthogonal \ f S" proof - interpret antiparallel_edges \ .. interpret \'': countable_network \'' by(rule \''_countable_network) have "\f S. flow \'' f \ cut \'' S \ orthogonal \'' f S" by(rule \''.max_flow_min_cut'')(auto simp add: sink_out source_in capacity_pos elim: edg.cases) then obtain f S where f: "flow \'' f" and cut: "cut \'' S" and ortho: "orthogonal \'' f S" by blast have "flow \ (collect f)" using f by(rule flow_collect) moreover have "cut \ (cut' S)" using cut by(rule cut_cut') moreover have "orthogonal \ (collect f) (cut' S)" using ortho f by(rule orthogonal_cut') ultimately show ?thesis by blast qed theorem max_flow_min_cut: "\f S. flow \ f \ cut \ S \ orthogonal \ f S" proof - interpret \''': countable_network \''' by(rule \'''_countable_network) have "\f S. flow \''' f \ cut \''' S \ orthogonal \''' f S" by(rule \'''.max_flow_min_cut''') auto then obtain f S where f: "flow \''' f" and cut: "cut \''' S" and ortho: "orthogonal \''' f S" by blast from flow_\'''[OF this] show ?thesis by blast qed end end end diff --git a/thys/MFMC_Countable/MFMC_Web.thy b/thys/MFMC_Countable/MFMC_Web.thy --- a/thys/MFMC_Countable/MFMC_Web.thy +++ b/thys/MFMC_Countable/MFMC_Web.thy @@ -1,1973 +1,1973 @@ theory MFMC_Web imports MFMC_Network begin section \Webs and currents\ record 'v web = "'v graph" + weight :: "'v \ ennreal" A :: "'v set" B :: "'v set" lemma vertex_weight_update [simp]: "vertex (weight_update f \) = vertex \" by(simp add: vertex_def fun_eq_iff) type_synonym 'v current = "'v edge \ ennreal" inductive current :: "('v, 'more) web_scheme \ 'v current \ bool" for \ f where current: "\ \x. d_OUT f x \ weight \ x; \x. d_IN f x \ weight \ x; \x. x \ A \ \ d_OUT f x \ d_IN f x; \a. a \ A \ \ d_IN f a = 0; \b. b \ B \ \ d_OUT f b = 0; \e. e \ \<^bold>E\<^bsub>\\<^esub> \ f e = 0 \ \ current \ f" lemma currentD_weight_OUT: "current \ f \ d_OUT f x \ weight \ x" by(simp add: current.simps) lemma currentD_weight_IN: "current \ f \ d_IN f x \ weight \ x" by(simp add: current.simps) lemma currentD_OUT_IN: "\ current \ f; x \ A \ \ \ d_OUT f x \ d_IN f x" by(simp add: current.simps) lemma currentD_IN: "\ current \ f; a \ A \ \ \ d_IN f a = 0" by(simp add: current.simps) lemma currentD_OUT: "\ current \ f; b \ B \ \ \ d_OUT f b = 0" by(simp add: current.simps) lemma currentD_outside: "\ current \ f; \ edge \ x y \ \ f (x, y) = 0" by(blast elim: current.cases) lemma currentD_outside': "\ current \ f; e \ \<^bold>E\<^bsub>\\<^esub> \ \ f e = 0" by(blast elim: current.cases) lemma currentD_OUT_eq_0: assumes "current \ f" shows "d_OUT f x = 0 \ (\y. f (x, y) = 0)" by(simp add: d_OUT_def nn_integral_0_iff emeasure_count_space_eq_0) lemma currentD_IN_eq_0: assumes "current \ f" shows "d_IN f x = 0 \ (\y. f (y, x) = 0)" by(simp add: d_IN_def nn_integral_0_iff emeasure_count_space_eq_0) lemma current_support_flow: fixes \ (structure) assumes "current \ f" shows "support_flow f \ \<^bold>E" using currentD_outside[OF assms] by(auto simp add: support_flow.simps intro: ccontr) lemma currentD_outside_IN: "\ current \ f; x \ \<^bold>V\<^bsub>\\<^esub> \ \ d_IN f x = 0" by(auto simp add: d_IN_def vertex_def nn_integral_0_iff AE_count_space emeasure_count_space_eq_0 dest: currentD_outside) lemma currentD_outside_OUT: "\ current \ f; x \ \<^bold>V\<^bsub>\\<^esub> \ \ d_OUT f x = 0" by(auto simp add: d_OUT_def vertex_def nn_integral_0_iff AE_count_space emeasure_count_space_eq_0 dest: currentD_outside) lemma currentD_weight_in: "current \ h \ h (x, y) \ weight \ y" by (metis order_trans d_IN_ge_point currentD_weight_IN) lemma currentD_weight_out: "current \ h \ h (x, y) \ weight \ x" by (metis order_trans d_OUT_ge_point currentD_weight_OUT) lemma current_leI: fixes \ (structure) assumes f: "current \ f" and le: "\e. g e \ f e" and OUT_IN: "\x. x \ A \ \ d_OUT g x \ d_IN g x" shows "current \ g" proof show "d_OUT g x \ weight \ x" for x using d_OUT_mono[of g x f, OF le] currentD_weight_OUT[OF f] by(rule order_trans) show "d_IN g x \ weight \ x" for x using d_IN_mono[of g x f, OF le] currentD_weight_IN[OF f] by(rule order_trans) show "d_IN g a = 0" if "a \ A \" for a using d_IN_mono[of g a f, OF le] currentD_IN[OF f that] by auto show "d_OUT g b = 0" if "b \ B \" for b using d_OUT_mono[of g b f, OF le] currentD_OUT[OF f that] by auto show "g e = 0" if "e \ \<^bold>E" for e using currentD_outside'[OF f that] le[of e] by simp qed(blast intro: OUT_IN)+ lemma current_weight_mono: "\ current \ f; edge \ = edge \'; A \ = A \'; B \ = B \'; \x. weight \ x \ weight \' x \ \ current \' f" by(auto 4 3 elim!: current.cases intro!: current.intros intro: order_trans) abbreviation (input) zero_current :: "'v current" where "zero_current \ \_. 0" lemma SINK_0 [simp]: "SINK zero_current = UNIV" by(auto simp add: SINK.simps) lemma current_0 [simp]: "current \ zero_current" by(auto simp add: current.simps) inductive web_flow :: "('v, 'more) web_scheme \ 'v current \ bool" for \ (structure) and f where web_flow: "\ current \ f; \x. \ x \ \<^bold>V; x \ A \; x \ B \ \ \ KIR f x \ \ web_flow \ f" lemma web_flowD_current: "web_flow \ f \ current \ f" by(erule web_flow.cases) lemma web_flowD_KIR: "\ web_flow \ f; x \ A \; x \ B \ \ \ KIR f x" apply(cases "x \ \<^bold>V\<^bsub>\\<^esub>") apply(fastforce elim!: web_flow.cases) apply(auto simp add: vertex_def d_OUT_def d_IN_def elim!: web_flow.cases) apply(subst (1 2) currentD_outside[of _ f]; auto) done subsection \Saturated and terminal vertices\ inductive_set SAT :: "('v, 'more) web_scheme \ 'v current \ 'v set" for \ f where A: "x \ A \ \ x \ SAT \ f" | IN: "d_IN f x \ weight \ x \ x \ SAT \ f" \ \We use @{text "\ weight"} such that @{text SAT} is monotone w.r.t. increasing currents\ lemma SAT_0 [simp]: "SAT \ zero_current = A \ \ {x. weight \ x \ 0}" by(auto simp add: SAT.simps) lemma SAT_mono: assumes "\e. f e \ g e" shows "SAT \ f \ SAT \ g" proof fix x assume "x \ SAT \ f" thus "x \ SAT \ g" proof cases case IN also have "d_IN f x \ d_IN g x" using assms by(rule d_IN_mono) finally show ?thesis .. qed(rule SAT.A) qed lemma SAT_Sup_upper: "f \ Y \ SAT \ f \ SAT \ (Sup Y)" by(rule SAT_mono)(rule Sup_upper[THEN le_funD]) lemma currentD_SAT: assumes "current \ f" shows "x \ SAT \ f \ x \ A \ \ d_IN f x = weight \ x" using currentD_weight_IN[OF assms, of x] by(auto simp add: SAT.simps) abbreviation terminal :: "('v, 'more) web_scheme \ 'v current \ 'v set" ("TER\") where "terminal \ f \ SAT \ f \ SINK f" subsection \Separation\ inductive separating_gen :: "('v, 'more) graph_scheme \ 'v set \ 'v set \ 'v set \ bool" for G A B S where separating: "(\x y p. \ x \ A; y \ B; path G x p y \ \ (\z \ set p. z \ S) \ x \ S) \ separating_gen G A B S" abbreviation separating :: "('v, 'more) web_scheme \ 'v set \ bool" where "separating \ \ separating_gen \ (A \) (B \)" abbreviation separating_network :: "('v, 'more) network_scheme \ 'v set \ bool" where "separating_network \ \ separating_gen \ {source \} {sink \}" lemma separating_networkI [intro?]: "(\p. path \ (source \) p (sink \) \ (\z \ set p. z \ S) \ source \ \ S) \ separating_network \ S" by(auto intro: separating) lemma separatingD: "\A B. \ separating_gen G A B S; path G x p y; x \ A; y \ B \ \ (\z \ set p. z \ S) \ x \ S" by(blast elim: separating_gen.cases) lemma separating_left [simp]: "\A B. A \ A' \ separating_gen \ A B A'" by(auto simp add: separating_gen.simps) lemma separating_weakening: "\A B. \ separating_gen G A B S; S \ S' \ \ separating_gen G A B S'" by(rule separating; drule (3) separatingD; blast) definition essential :: "('v, 'more) graph_scheme \ 'v set \ 'v set \ 'v \ bool" where \ \Should we allow only simple paths here?\ "\B. essential G B S x \ (\p. \y\B. path G x p y \ (x \ y \ (\z\set p. z = x \ z \ S)))" abbreviation essential_web :: "('v, 'more) web_scheme \ 'v set \ 'v set" ("\\") where "essential_web \ S \ {x\S. essential \ (B \) S x}" lemma essential_weight_update [simp]: "essential (weight_update f G) = essential G" by(simp add: essential_def fun_eq_iff) lemma not_essentialD: "\B. \ \ essential G B S x; path G x p y; y \ B \ \ x \ y \ (\z\set p. z \ x \ z \ S)" by(simp add: essential_def) lemma essentialE [elim?, consumes 1, case_names essential, cases pred: essential]: "\B. \ essential G B S x; \p y. \ path G x p y; y \ B; \z. \ x \ y; z \ set p \ \ z = x \ z \ S \ \ thesis \ \ thesis" by(auto simp add: essential_def) lemma essentialI [intro?]: "\B. \ path G x p y; y \ B; \z. \ x \ y; z \ set p \ \ z = x \ z \ S \ \ essential G B S x" by(auto simp add: essential_def) lemma essential_vertex: "\B. \ essential G B S x; x \ B \ \vertex G x" by(auto elim!: essentialE simp add: vertex_def elim: rtrancl_path.cases) lemma essential_BI: "\B. x \ B \ essential G B S x" by(auto simp add: essential_def intro: rtrancl_path.base) lemma \_E [elim?, consumes 1, case_names \, cases set: essential_web]: fixes \ (structure) assumes "x \ \ S" obtains p y where "path \ x p y" "y \ B \" "\z. \ x \ y; z \ set p \ \ z = x \ z \ S" using assms by(auto elim: essentialE) lemma essential_mono: "\B. \ essential G B S x; S' \ S \ \ essential G B S' x" by(auto simp add: essential_def) lemma separating_essential: \ \Lem. 3.4 (cf. Lem. 2.14 in [5])\ fixes G A B S assumes "separating_gen G A B S" shows "separating_gen G A B {x\S. essential G B S x}" (is "separating_gen _ _ _ ?E") proof fix x y p assume x: "x \ A" and y: "y \ B" and p: "path G x p y" from separatingD[OF assms p x y] have "\z \ set (x # p). z \ S" by auto from split_list_last_prop[OF this] obtain ys z zs where decomp: "x # p = ys @ z # zs" and z: "z \ S" and last: "\z. z \ set zs \ z \ S" by auto from decomp consider (empty) "ys = []" "x = z" "p = zs" | (Cons) ys' where "ys = x # ys'" "p = ys' @ z # zs" by(auto simp add: Cons_eq_append_conv) then show "(\z\set p. z \ ?E) \ x \ ?E" proof(cases) case empty hence "x \ ?E" using z p last y by(auto simp add: essential_def) thus ?thesis .. next case (Cons ys') from p have "path G z zs y" unfolding Cons by(rule rtrancl_path_appendE) hence "z \ ?E" using z y last by(auto simp add: essential_def) thus ?thesis using Cons by auto qed qed definition roofed_gen :: "('v, 'more) graph_scheme \ 'v set \ 'v set \ 'v set" where roofed_def: "\B. roofed_gen G B S = {x. \p. \y\B. path G x p y \ (\z\set p. z \ S) \ x \ S}" abbreviation roofed :: "('v, 'more) web_scheme \ 'v set \ 'v set" ("RF\") where "roofed \ \ roofed_gen \ (B \)" abbreviation roofed_network :: "('v, 'more) network_scheme \ 'v set \ 'v set" ("RF\<^sup>N\") where "roofed_network \ \ roofed_gen \ {sink \}" lemma roofedI [intro?]: "\B. (\p y. \ path G x p y; y \ B \ \ (\z\set p. z \ S) \ x \ S) \ x \ roofed_gen G B S" by(auto simp add: roofed_def) lemma not_roofedE: fixes B assumes "x \ roofed_gen G B S" obtains p y where "path G x p y" "y \ B" "\z. z \ set (x # p) \ z \ S" using assms by(auto simp add: roofed_def) lemma roofed_greater: "\B. S \ roofed_gen G B S" by(auto simp add: roofed_def) lemma roofed_greaterI: "\B. x \ S \ x \ roofed_gen G B S" using roofed_greater[of S G] by blast lemma roofed_mono: "\B. S \ S' \ roofed_gen G B S \ roofed_gen G B S'" by(fastforce simp add: roofed_def) lemma in_roofed_mono: "\B. \ x \ roofed_gen G B S; S \ S' \ \ x \ roofed_gen G B S'" using roofed_mono[THEN subsetD] . lemma roofedD: "\B. \ x \ roofed_gen G B S; path G x p y; y \ B \ \ (\z\set p. z \ S) \ x \ S" unfolding roofed_def by blast lemma separating_RF_A: fixes A B assumes "separating_gen G A B X" shows "A \ roofed_gen G B X" by(rule subsetI roofedI)+(erule separatingD[OF assms]) lemma roofed_idem: fixes B shows "roofed_gen G B (roofed_gen G B S) = roofed_gen G B S" proof(rule equalityI subsetI roofedI)+ fix x p y assume x: "x \ roofed_gen G B (roofed_gen G B S)" and p: "path G x p y" and y: "y \ B" from roofedD[OF x p y] obtain z where *: "z \ set (x # p)" and z: "z \ roofed_gen G B S" by auto from split_list[OF *] obtain ys zs where split: "x # p = ys @ z # zs" by blast with p have p': "path G z zs y" by(auto simp add: Cons_eq_append_conv elim: rtrancl_path_appendE) from roofedD[OF z p' y] split show "(\z\set p. z \ S) \ x \ S" by(auto simp add: Cons_eq_append_conv) qed(rule roofed_mono roofed_greater)+ lemma in_roofed_mono': "\B. \ x \ roofed_gen G B S; S \ roofed_gen G B S' \ \ x \ roofed_gen G B S'" by(subst roofed_idem[symmetric])(erule in_roofed_mono) lemma roofed_mono': "\B. S \ roofed_gen G B S' \ roofed_gen G B S \ roofed_gen G B S'" by(rule subsetI)(rule in_roofed_mono') lemma roofed_idem_Un1: fixes B shows "roofed_gen G B (roofed_gen G B S \ T) = roofed_gen G B (S \ T)" proof - have "S \ T \ roofed_gen G B S" by (metis (no_types) UnCI roofed_greater subsetCE subsetI) then have "S \ T \ T \ roofed_gen G B S \ T \ roofed_gen G B S \ roofed_gen G B (S \ T)" by (metis (no_types) Un_subset_iff Un_upper2 roofed_greater roofed_mono sup.commute) then show ?thesis by (metis (no_types) roofed_idem roofed_mono subset_antisym sup.commute) qed lemma roofed_UN: fixes A B shows "roofed_gen G B (\i\A. roofed_gen G B (X i)) = roofed_gen G B (\i\A. X i)" (is "?lhs = ?rhs") proof(rule equalityI) show "?rhs \ ?lhs" by(rule roofed_mono)(blast intro: roofed_greaterI) show "?lhs \ ?rhs" by(rule roofed_mono')(blast intro: in_roofed_mono) qed lemma RF_essential: fixes \ (structure) shows "RF (\ S) = RF S" proof(intro set_eqI iffI) fix x assume RF: "x \ RF S" show "x \ RF (\ S)" proof fix p y assume p: "path \ x p y" and y: "y \ B \" from roofedD[OF RF this] have "\z\set (x # p). z \ S" by auto from split_list_last_prop[OF this] obtain ys z zs where decomp: "x # p = ys @ z # zs" and z: "z \ S" and last: "\z. z \ set zs \ z \ S" by auto from decomp consider (empty) "ys = []" "x = z" "p = zs" | (Cons) ys' where "ys = x # ys'" "p = ys' @ z # zs" by(auto simp add: Cons_eq_append_conv) then show "(\z\set p. z \ \ S) \ x \ \ S" proof(cases) case empty hence "x \ \ S" using z p last y by(auto simp add: essential_def) thus ?thesis .. next case (Cons ys') from p have "path \ z zs y" unfolding Cons by(rule rtrancl_path_appendE) hence "z \ \ S" using z y last by(auto simp add: essential_def) thus ?thesis using Cons by auto qed qed qed(blast intro: in_roofed_mono) lemma essentialE_RF: fixes \ (structure) and B assumes "essential \ B S x" obtains p y where "path \ x p y" "y \ B" "distinct (x # p)" "\z. z \ set p \ z \ roofed_gen \ B S" proof - from assms obtain p y where p: "path \ x p y" and y: "y \ B" and bypass: "\z. \ x \ y; z \ set p \ \ z = x \ z \ S" by(rule essentialE) blast from p obtain p' where p': "path \ x p' y" and distinct: "distinct (x # p')" and subset: "set p' \ set p" by(rule rtrancl_path_distinct) { fix z assume z: "z \ set p'" hence "y \ set p'" using rtrancl_path_last[OF p', symmetric] p' by(auto elim: rtrancl_path.cases intro: last_in_set) with distinct z subset have neq: "x \ y" and "z \ set p" by(auto) from bypass[OF this] z distinct have "z \ S" by auto have "z \ roofed_gen \ B S" proof assume z': "z \ roofed_gen \ B S" from split_list[OF z] obtain ys zs where decomp: "p' = ys @ z # zs" by blast with p' have "path \ z zs y" by(auto elim: rtrancl_path_appendE) from roofedD[OF z' this y] \z \ S\ obtain z' where "z' \ set zs" "z' \ S" by auto with bypass[of z'] neq decomp subset distinct show False by auto qed } with p' y distinct show thesis .. qed lemma \_E_RF: fixes \ (structure) assumes "x \ \ S" obtains p y where "path \ x p y" "y \ B \" "distinct (x # p)" "\z. z \ set p \ z \ RF S" using assms by(auto elim: essentialE_RF) lemma in_roofed_essentialD: fixes \ (structure) assumes RF: "x \ RF S" and ess: "essential \ (B \) S x" shows "x \ S" proof - from ess obtain p y where p: "path \ x p y" and y: "y \ B \" and "distinct (x # p)" and bypass: "\z. z \ set p \ z \ S" by(rule essentialE_RF)(auto intro: roofed_greaterI) from roofedD[OF RF p y] bypass show "x \ S" by auto qed lemma separating_RF: fixes \ (structure) shows "separating \ (RF S) \ separating \ S" proof assume sep: "separating \ (RF S)" show "separating \ S" proof fix x y p assume p: "path \ x p y" and x: "x \ A \" and y: "y \ B \" from separatingD[OF sep p x y] have "\z \ set (x # p). z \ RF S" by auto from split_list_last_prop[OF this] obtain ys z zs where split: "x # p = ys @ z # zs" and z: "z \ RF S" and bypass: "\z'. z' \ set zs \ z' \ RF S" by auto from p split have "path \ z zs y" by(cases ys)(auto elim: rtrancl_path_appendE) hence "essential \ (B \) S z" using y by(rule essentialI)(auto dest: bypass intro: roofed_greaterI) with z have "z \ S" by(rule in_roofed_essentialD) with split show "(\z\set p. z \ S) \ x \ S" by(cases ys)auto qed qed(blast intro: roofed_greaterI separating_weakening) definition roofed_circ :: "('v, 'more) web_scheme \ 'v set \ 'v set" ("RF\<^sup>\\") where "roofed_circ \ S = roofed \ S - \\<^bsub>\\<^esub> S" lemma roofed_circI: fixes \ (structure) shows "\ x \ RF T; x \ T \ \ essential \ (B \) T x \ \ x \ RF\<^sup>\ T" by(simp add: roofed_circ_def) lemma roofed_circE: fixes \ (structure) assumes "x \ RF\<^sup>\ T" obtains "x \ RF T" "\ essential \ (B \) T x" using assms by(auto simp add: roofed_circ_def intro: in_roofed_essentialD) lemma \_\: fixes \ (structure) shows "\ (\ S) = \ S" by(auto intro: essential_mono) lemma roofed_circ_essential: fixes \ (structure) shows "RF\<^sup>\ (\ S) = RF\<^sup>\ S" unfolding roofed_circ_def RF_essential \_\ .. lemma essential_RF: fixes B shows "essential G B (roofed_gen G B S) = essential G B S" (is "essential _ _ ?RF = _") proof(intro ext iffI) show "essential G B S x" if "essential G B ?RF x" for x using that by(rule essential_mono)(blast intro: roofed_greaterI) show "essential G B ?RF x" if "essential G B S x" for x using that by(rule essentialE_RF)(erule (1) essentialI, blast) qed lemma \_RF: fixes \ (structure) shows "\ (RF S) = \ S" by(auto dest: in_roofed_essentialD simp add: essential_RF intro: roofed_greaterI) lemma essential_\: fixes \ (structure) shows "essential \ (B \) (\ S) = essential \ (B \) S" by(subst essential_RF[symmetric])(simp only: RF_essential essential_RF) lemma RF_in_B: fixes \ (structure) shows "x \ B \ \ x \ RF S \ x \ S" by(auto intro: roofed_greaterI dest: roofedD[OF _ rtrancl_path.base]) lemma RF_circ_edge_forward: fixes \ (structure) assumes x: "x \ RF\<^sup>\ S" and edge: "edge \ x y" shows "y \ RF S" proof fix p z assume p: "path \ y p z" and z: "z \ B \" from x have rf: "x \ RF S" and ness: "x \ \ S" by(auto elim: roofed_circE) show "(\z\set p. z \ S) \ y \ S" proof(cases "\z'\set (y # p). z' \ S") case False from edge p have p': "path \ x (y # p) z" .. from roofedD[OF rf this z] False have "x \ S" by auto moreover have "essential \ (B \) S x" using p' False z by(auto intro!: essentialI) ultimately have "x \ \ S" by simp with ness show ?thesis by contradiction qed auto qed subsection \Waves\ inductive wave :: "('v, 'more) web_scheme \ 'v current \ bool" for \ (structure) and f where wave: "\ separating \ (TER f); \x. x \ RF (TER f) \ d_OUT f x = 0 \ \ wave \ f" lemma wave_0 [simp]: "wave \ zero_current" by rule simp_all lemma waveD_separating: "wave \ f \ separating \ (TER\<^bsub>\\<^esub> f)" by(simp add: wave.simps) lemma waveD_OUT: "\ wave \ f; x \ RF\<^bsub>\\<^esub> (TER\<^bsub>\\<^esub> f) \ \ d_OUT f x = 0" by(simp add: wave.simps) lemma wave_A_in_RF: fixes \ (structure) shows "\ wave \ f; x \ A \ \ \ x \ RF (TER f)" by(auto intro!: roofedI dest!: waveD_separating separatingD) lemma wave_not_RF_IN_zero: fixes \ (structure) assumes f: "current \ f" and w: "wave \ f" and x: "x \ RF (TER f)" shows "d_IN f x = 0" proof - from x obtain p z where z: "z \ B \" and p: "path \ x p z" and bypass: "\z. z \ set p \ z \ TER f" "x \ TER f" by(clarsimp simp add: roofed_def) have "f (y, x) = 0" for y proof(cases "edge \ y x") case edge: True have "d_OUT f y = 0" proof(cases "y \ TER f") case False with z p bypass edge have "y \ RF (TER f)" by(auto simp add: roofed_def intro: rtrancl_path.step intro!: exI rev_bexI) thus "d_OUT f y = 0" by(rule waveD_OUT[OF w]) qed(auto simp add: SINK.simps) moreover have "f (y, x) \ d_OUT f y" by (rule d_OUT_ge_point) ultimately show ?thesis by simp qed(simp add: currentD_outside[OF f]) then show "d_IN f x = 0" unfolding d_IN_def by(simp add: nn_integral_0_iff emeasure_count_space_eq_0) qed lemma current_Sup: fixes \ (structure) assumes chain: "Complete_Partial_Order.chain (\) Y" and Y: "Y \ {}" and current: "\f. f \ Y \ current \ f" and countable [simp]: "countable (support_flow (Sup Y))" shows "current \ (Sup Y)" proof(rule, goal_cases) case (1 x) have "d_OUT (Sup Y) x = (SUP f\Y. d_OUT f x)" using chain Y by(simp add: d_OUT_Sup) also have "\ \ weight \ x" using 1 by(intro SUP_least)(auto dest!: current currentD_weight_OUT) finally show ?case . next case (2 x) have "d_IN (Sup Y) x = (SUP f\Y. d_IN f x)" using chain Y by(simp add: d_IN_Sup) also have "\ \ weight \ x" using 2 by(intro SUP_least)(auto dest!: current currentD_weight_IN) finally show ?case . next case (3 x) have "d_OUT (Sup Y) x = (SUP f\Y. d_OUT f x)" using chain Y by(simp add: d_OUT_Sup) also have "\ \ (SUP f\Y. d_IN f x)" using 3 by(intro SUP_mono)(auto dest: current currentD_OUT_IN) also have "\ = d_IN (Sup Y) x" using chain Y by(simp add: d_IN_Sup) finally show ?case . next case (4 a) have "d_IN (Sup Y) a = (SUP f\Y. d_IN f a)" using chain Y by(simp add: d_IN_Sup) also have "\ = (SUP f\Y. 0)" using 4 by(intro SUP_cong)(auto dest!: current currentD_IN) also have "\ = 0" using Y by simp finally show ?case . next case (5 b) have "d_OUT (Sup Y) b = (SUP f\Y. d_OUT f b)" using chain Y by(simp add: d_OUT_Sup) also have "\ = (SUP f\Y. 0)" using 5 by(intro SUP_cong)(auto dest!: current currentD_OUT) also have "\ = 0" using Y by simp finally show ?case . next fix e assume "e \ \<^bold>E" from currentD_outside'[OF current this] have "f e = 0" if "f \ Y" for f using that by simp hence "Sup Y e = (SUP _\Y. 0)" by(auto intro: SUP_cong) then show "Sup Y e = 0" using Y by(simp) qed lemma wave_lub: \ \Lemma 4.3\ fixes \ (structure) assumes chain: "Complete_Partial_Order.chain (\) Y" and Y: "Y \ {}" and wave: "\f. f \ Y \ wave \ f" and countable [simp]: "countable (support_flow (Sup Y))" shows "wave \ (Sup Y)" proof { fix x y p assume p: "path \ x p y" and y: "y \ B \" define P where "P = {x} \ set p" let ?f = "\f. SINK f \ P" have "Complete_Partial_Order.chain (\) (?f ` Y)" using chain by(rule chain_imageI)(auto dest: SINK_mono') moreover have "\ \ Pow P" by auto hence "finite (?f ` Y)" by(rule finite_subset)(simp add: P_def) ultimately have "(\(?f ` Y)) \ ?f ` Y" by(rule ccpo.in_chain_finite[OF complete_lattice_ccpo_dual])(simp add: Y) then obtain f where f: "f \ Y" and eq: "\(?f ` Y) = ?f f" by clarify hence *: "(\f\Y. SINK f) \ P = SINK f \ P" by(clarsimp simp add: prod_lub_def Y)+ { fix g assume "g \ Y" "f \ g" with * have "(\f\Y. SINK f) \ P = SINK g \ P" by(blast dest: SINK_mono') then have "TER (Sup Y) \ P \ TER g \ P" using SAT_Sup_upper[OF \g \ Y\, of \] SINK_Sup[OF chain Y countable] by blast } with f have "\f\Y. \g\Y. g \ f \ TER g \ P \ TER (Sup Y) \ P" by blast } note subset = this show "separating \ (TER (Sup Y))" proof fix x y p assume *: "path \ x p y" "y \ B \" and "x \ A \" let ?P = "{x} \ set p" from subset[OF *] obtain f where f:"f \ Y" and subset: "TER f \ ?P \ TER (Sup Y) \ ?P" by blast from wave[OF f] have "TER f \ ?P \ {}" using * \x \ A \\ by(auto simp add: wave.simps dest: separatingD) with subset show " (\z\set p. z \ TER (Sup Y)) \ x \ TER (Sup Y)" by blast qed fix x assume "x \ RF (TER (Sup Y))" then obtain p y where y: "y \ B \" and p: "path \ x p y" and ter: "TER (Sup Y) \ ({x} \ set p) = {}" by(auto simp add: roofed_def) let ?P = "{x} \ set p" from subset[OF p y] obtain f where f: "f \ Y" and subset: "\g. \ g \ Y; f \ g \ \ TER g \ ?P \ TER (Sup Y) \ ?P" by blast { fix g assume g: "g \ Y" with chain f have "f \ g \ g \ f" by(rule chainD) hence "d_OUT g x = 0" proof assume "f \ g" from subset[OF g this] ter have "TER g \ ?P = {}" by blast with p y have "x \ RF (TER g)" by(auto simp add: roofed_def) with wave[OF g] show ?thesis by(blast elim: wave.cases) next assume "g \ f" from subset ter f have "TER f \ ?P = {}" by blast with y p have "x \ RF (TER f)" by(auto simp add: roofed_def) with wave[OF f] have "d_OUT f x = 0" by(blast elim: wave.cases) moreover have "d_OUT g x \ d_OUT f x" using \g \ f\[THEN le_funD] by(rule d_OUT_mono) ultimately show ?thesis by simp qed } thus "d_OUT (Sup Y) x = 0" using chain Y by(simp add: d_OUT_Sup) qed lemma ex_maximal_wave: \ \Corollary 4.4\ fixes \ (structure) assumes countable: "countable \<^bold>E" shows "\f. current \ f \ wave \ f \ (\w. current \ w \ wave \ w \ f \ w \ f = w)" proof - define Field_r where "Field_r = {f. current \ f \ wave \ f}" define r where "r = {(f, g). f \ Field_r \ g \ Field_r \ f \ g}" have Field_r: "Field r = Field_r" by(auto simp add: Field_def r_def) have "Partial_order r" unfolding order_on_defs by(auto intro!: refl_onI transI antisymI simp add: Field_r r_def Field_def) hence "\m\Field r. \a\Field r. (m, a) \ r \ a = m" proof(rule Zorns_po_lemma) fix Y assume "Y \ Chains r" hence Y: "Complete_Partial_Order.chain (\) Y" and w: "\f. f \ Y \ wave \ f" and f: "\f. f \ Y \ current \ f" by(auto simp add: Chains_def r_def chain_def Field_r_def) show "\w \ Field r. \f \ Y. (f, w) \ r" proof(cases "Y = {}") case True have "zero_current \ Field r" by(simp add: Field_r Field_r_def) with True show ?thesis by blast next case False have "support_flow (Sup Y) \ \<^bold>E" by(auto simp add: support_flow_Sup elim!: support_flow.cases dest!: f dest: currentD_outside) hence c: "countable (support_flow (Sup Y))" using countable by(rule countable_subset) with Y False f w have "Sup Y \ Field r" unfolding Field_r Field_r_def by(blast intro: wave_lub current_Sup) moreover then have "(f, Sup Y) \ r" if "f \ Y" for f using w[OF that] f[OF that] that unfolding Field_r by(auto simp add: r_def Field_r_def intro: Sup_upper) ultimately show ?thesis by blast qed qed thus ?thesis by(simp add: Field_r Field_r_def)(auto simp add: r_def Field_r_def) qed lemma essential_leI: fixes \ (structure) assumes g: "current \ g" and w: "wave \ g" and le: "\e. f e \ g e" and x: "x \ \ (TER g)" shows "essential \ (B \) (TER f) x" proof - from x obtain p y where p: "path \ x p y" and y: "y \ B \" and distinct: "distinct (x # p)" and bypass: "\z. z \ set p \ z \ RF (TER g)" by(rule \_E_RF) blast show ?thesis using p y proof fix z assume "z \ set p" hence z: "z \ RF (TER g)" by(auto dest: bypass) with w have OUT: "d_OUT g z = 0" and IN: "d_IN g z = 0" by(rule waveD_OUT wave_not_RF_IN_zero[OF g])+ with z have "z \ A \" "weight \ z > 0" by(auto intro!: roofed_greaterI simp add: SAT.simps SINK.simps) moreover from IN d_IN_mono[of f z g, OF le] have "d_IN f z \ 0" by(simp) ultimately have "z \ TER f" by(auto simp add: SAT.simps) then show "z = x \ z \ TER f" by simp qed qed lemma essential_eq_leI: fixes \ (structure) assumes g: "current \ g" and w: "wave \ g" and le: "\e. f e \ g e" and subset: "\ (TER g) \ TER f" shows "\ (TER f) = \ (TER g)" proof show subset: "\ (TER g) \ \ (TER f)" proof fix x assume x: "x \ \ (TER g)" hence "x \ TER f" using subset by blast moreover have "essential \ (B \) (TER f) x" using g w le x by(rule essential_leI) ultimately show "x \ \ (TER f)" by simp qed show "\ \ \ (TER g)" proof fix x assume x: "x \ \ (TER f)" hence "x \ TER f" by auto hence "x \ RF (TER g)" proof(rule contrapos_pp) assume x: "x \ RF (TER g)" with w have OUT: "d_OUT g x = 0" and IN: "d_IN g x = 0" by(rule waveD_OUT wave_not_RF_IN_zero[OF g])+ with x have "x \ A \" "weight \ x > 0" by(auto intro!: roofed_greaterI simp add: SAT.simps SINK.simps) moreover from IN d_IN_mono[of f x g, OF le] have "d_IN f x \ 0" by(simp) ultimately show "x \ TER f" by(auto simp add: SAT.simps) qed moreover have "x \ RF\<^sup>\ (TER g)" proof assume "x \ RF\<^sup>\ (TER g)" hence RF: "x \ RF (\ (TER g))" and not_E: "x \ \ (TER g)" unfolding RF_essential by(simp_all add: roofed_circ_def) from x obtain p y where p: "path \ x p y" and y: "y \ B \" and distinct: "distinct (x # p)" and bypass: "\z. z \ set p \ z \ RF (TER f)" by(rule \_E_RF) blast from roofedD[OF RF p y] not_E obtain z where "z \ set p" "z \ \ (TER g)" by blast with subset bypass[of z] show False by(auto intro: roofed_greaterI) qed ultimately show "x \ \ (TER g)" by(simp add: roofed_circ_def) qed qed subsection \Hindrances and looseness\ inductive hindrance_by :: "('v, 'more) web_scheme \ 'v current \ ennreal \ bool" for \ (structure) and f and \ where hindrance_by: "\ a \ A \; a \ \ (TER f); d_OUT f a < weight \ a; \ < weight \ a - d_OUT f a \ \ hindrance_by \ f \" inductive hindrance :: "('v, 'more) web_scheme \ 'v current \ bool" for \ (structure) and f where hindrance: "\ a \ A \; a \ \ (TER f); d_OUT f a < weight \ a \ \ hindrance \ f" inductive hindered :: "('v, 'more) web_scheme \ bool" for \ (structure) where hindered: "\ hindrance \ f; current \ f; wave \ f \ \ hindered \" inductive hindered_by :: "('v, 'more) web_scheme \ ennreal \ bool" for \ (structure) and \ where hindered_by: "\ hindrance_by \ f \; current \ f; wave \ f \ \ hindered_by \ \" lemma hindrance_into_hindrance_by: assumes "hindrance \ f" shows "\\>0. hindrance_by \ f \" using assms proof cases case (hindrance a) let ?\ = "if weight \ a = \ then 1 else (weight \ a - d_OUT f a) / 2" from \d_OUT f a < weight \ a\ have "weight \ a - d_OUT f a > 0" "weight \ a \ \ \ weight \ a - d_OUT f a < \" by(simp_all add: diff_gr0_ennreal less_top diff_less_top_ennreal) from ennreal_mult_strict_left_mono[of 1 2, OF _ this] have "0 < ?\" and "?\ < weight \ a - d_OUT f a" using \d_OUT f a < weight \ a\ by(auto intro!: diff_gr0_ennreal simp: ennreal_zero_less_divide divide_less_ennreal) with hindrance show ?thesis by(auto intro!: hindrance_by.intros) qed lemma hindrance_by_into_hindrance: "hindrance_by \ f \ \ hindrance \ f" by(blast elim: hindrance_by.cases intro: hindrance.intros) lemma hindrance_conv_hindrance_by: "hindrance \ f \ (\\>0. hindrance_by \ f \)" by(blast intro: hindrance_into_hindrance_by hindrance_by_into_hindrance) lemma hindered_into_hindered_by: "hindered \ \ \\>0. hindered_by \ \" by(blast intro: hindered_by.intros elim: hindered.cases dest: hindrance_into_hindrance_by) lemma hindered_by_into_hindered: "hindered_by \ \ \ hindered \" by(blast elim: hindered_by.cases intro: hindered.intros dest: hindrance_by_into_hindrance) lemma hindered_conv_hindered_by: "hindered \ \ (\\>0. hindered_by \ \)" by(blast intro: hindered_into_hindered_by hindered_by_into_hindered) inductive loose :: "('v, 'more) web_scheme \ bool" for \ where loose: "\ \f. \ current \ f; wave \ f \ \ f = zero_current; \ hindrance \ zero_current \ \ loose \" lemma looseD_hindrance: "loose \ \ \ hindrance \ zero_current" by(simp add: loose.simps) lemma looseD_wave: "\ loose \; current \ f; wave \ f \ \ f = zero_current" by(simp add: loose.simps) lemma loose_unhindered: fixes \ (structure) assumes "loose \" shows "\ hindered \" apply auto apply(erule hindered.cases) apply(frule (1) looseD_wave[OF assms]) apply simp using looseD_hindrance[OF assms] by simp context fixes \ \' :: "('v, 'more) web_scheme" assumes [simp]: "edge \ = edge \'" "A \ = A \'" "B \ = B \'" and weight_eq: "\x. x \ A \' \ weight \ x = weight \' x" and weight_le: "\a. a \ A \' \ weight \ a \ weight \' a" begin private lemma essential_eq: "essential \ = essential \'" by(simp add: fun_eq_iff essential_def) qualified lemma TER_eq: "TER\<^bsub>\\<^esub> f = TER\<^bsub>\'\<^esub> f" apply(auto simp add: SINK.simps SAT.simps) apply(erule contrapos_np; drule weight_eq; simp)+ done qualified lemma separating_eq: "separating_gen \ = separating_gen \'" by(intro ext iffI; rule separating_gen.intros; drule separatingD; simp) qualified lemma roofed_eq: "\B. roofed_gen \ B S = roofed_gen \' B S" by(simp add: roofed_def) lemma wave_eq_web: \ \Observation 4.6\ "wave \ f \ wave \' f" by(simp add: wave.simps separating_eq TER_eq roofed_eq) lemma current_mono_web: "current \' f \ current \ f" apply(rule current, simp_all add: currentD_OUT_IN currentD_IN currentD_OUT currentD_outside') subgoal for x by(cases "x \ A \'")(auto dest!: weight_eq weight_le dest: currentD_weight_OUT intro: order_trans) subgoal for x by(cases "x \ A \'")(auto dest!: weight_eq weight_le dest: currentD_weight_IN intro: order_trans) done lemma hindrance_mono_web: "hindrance \' f \ hindrance \ f" apply(erule hindrance.cases) apply(rule hindrance) apply simp apply(unfold TER_eq, simp add: essential_eq) apply(auto dest!: weight_le) done lemma hindered_mono_web: "hindered \' \ hindered \" apply(erule hindered.cases) apply(rule hindered.intros) apply(erule hindrance_mono_web) apply(erule current_mono_web) apply(simp add: wave_eq_web) done end subsection \Linkage\ text \ The following definition of orthogonality is stronger than the original definition 3.5 in @{cite AharoniBergerGeorgakopoulusPerlsteinSpruessel2011JCT} in that the outflow from any \A\-vertices in the set must saturate the vertex; @{term "S \ SAT \ f"} is not enough. With the original definition of orthogonal current, the reduction from networks to webs fails because the induced flow need not saturate edges going out of the source. Consider the network with three nodes \s\, \x\, and \t\ and edges \(s, x)\ and \(x, t)\ with capacity \1\. Then, the corresponding web has the vertices \(s, x)\ and \(x, t)\ and one edge from \(s, x)\ to \(x, t)\. Clearly, the zero current @{term [source] zero_current} is a web-flow and \TER zero_current = {(s, x)}\, which is essential. Moreover, @{term [source] zero_current} and \{(s, x)}\ are orthogonal because @{term [source] zero_current} trivially saturates \(s, x)\ as this is a vertex in \A\. \ inductive orthogonal_current :: "('v, 'more) web_scheme \ 'v current \ 'v set \ bool" for \ (structure) and f S where orthogonal_current: "\ \x. \ x \ S; x \ A \ \ \ weight \ x \ d_IN f x; \x. \ x \ S; x \ A \; x \ B \ \ \ d_OUT f x = weight \ x; \u v. \ v \ RF S; u \ RF\<^sup>\ S \ \ f (u, v) = 0 \ \ orthogonal_current \ f S" lemma orthogonal_currentD_SAT: "\ orthogonal_current \ f S; x \ S \ \ x \ SAT \ f" by(auto elim!: orthogonal_current.cases intro: SAT.intros) lemma orthogonal_currentD_A: "\ orthogonal_current \ f S; x \ S; x \ A \; x \ B \ \ \ d_OUT f x = weight \ x" by(auto elim: orthogonal_current.cases) lemma orthogonal_currentD_in: "\ orthogonal_current \ f S; v \ RF\<^bsub>\\<^esub> S; u \ RF\<^sup>\\<^bsub>\\<^esub> S \ \ f (u, v) = 0" by(auto elim: orthogonal_current.cases) inductive linkage :: "('v, 'more) web_scheme \ 'v current \ bool" for \ f where \ \Omit the condition @{const web_flow}\ linkage: "(\x. x \ A \ \ d_OUT f x = weight \ x) \ linkage \ f" lemma linkageD: "\ linkage \ f; x \ A \ \ \ d_OUT f x = weight \ x" by(rule linkage.cases) abbreviation linkable :: "('v, 'more) web_scheme \ bool" where "linkable \ \ \f. web_flow \ f \ linkage \ f" subsection \Trimming\ context fixes \ :: "('v, 'more) web_scheme" (structure) and f :: "'v current" begin inductive trimming :: "'v current \ bool" for g where trimming: \ \omits the condition that @{term f} is a wave\ "\ current \ g; wave \ g; g \ f; \x. \ x \ RF\<^sup>\ (TER f); x \ A \ \ \ KIR g x; \ (TER g) - A \ = \ (TER f) - A \ \ \ trimming g" lemma assumes "trimming g" shows trimmingD_current: "current \ g" and trimmingD_wave: "wave \ g" and trimmingD_le: "\e. g e \ f e" and trimmingD_KIR: "\ x \ RF\<^sup>\ (TER f); x \ A \ \ \ KIR g x" and trimmingD_\: "\ (TER g) - A \ = \ (TER f) - A \" using assms by(blast elim: trimming.cases dest: le_funD)+ lemma ex_trimming: \ \Lemma 4.8\ assumes f: "current \ f" and w: "wave \ f" and countable: "countable \<^bold>E" and weight_finite: "\x. weight \ x \ \" shows "\g. trimming g" proof - define F where "F = {g. current \ g \ wave \ g \ g \ f \ \ (TER g) = \ (TER f)}" define leq where "leq = restrict_rel F {(g, g'). g' \ g}" have in_F [simp]: "g \ F \ current \ g \ wave \ g \ (\e. g e \ f e) \ \ (TER g) = \ (TER f)" for g by(simp add: F_def le_fun_def) have f_finite [simp]: "f e \ \" for e proof(cases e) case (Pair x y) have "f (x, y) \ d_IN f y" by (rule d_IN_ge_point) also have "\ \ weight \ y" by(rule currentD_weight_IN[OF f]) also have "\ < \" by(simp add: weight_finite less_top[symmetric]) finally show ?thesis using Pair by simp qed have chainD: "Inf M \ F" if M: "M \ Chains leq" and nempty: "M \ {}" for M proof - from nempty obtain g0 where g0: "g0 \ M" by auto have g0_le_f: "g0 e \ f e" and g: "current \ g0" and w0: "wave \ g0" for e using Chains_FieldD[OF M g0] by(cases e, auto simp add: leq_def) have finite_OUT: "d_OUT f x \ \" for x using weight_finite[of x] by(rule neq_top_trans)(rule currentD_weight_OUT[OF f]) have finite_IN: "d_IN f x \ \" for x using weight_finite[of x] by(rule neq_top_trans)(rule currentD_weight_IN[OF f]) from M have "M \ Chains {(g, g'). g' \ g}" by(rule mono_Chains[THEN subsetD, rotated])(auto simp add: leq_def in_restrict_rel_iff) then have chain: "Complete_Partial_Order.chain (\) M" by(rule Chains_into_chain) hence chain': "Complete_Partial_Order.chain (\) M" by(simp add: chain_dual) have countable': "countable (support_flow f)" using current_support_flow[OF f] by(rule countable_subset)(rule countable) have OUT_M: "d_OUT (Inf M) x = (INF g\M. d_OUT g x)" for x using chain' nempty countable' _ finite_OUT by(rule d_OUT_Inf)(auto dest!: Chains_FieldD[OF M] simp add: leq_def) have IN_M: "d_IN (Inf M) x = (INF g\M. d_IN g x)" for x using chain' nempty countable' _ finite_IN by(rule d_IN_Inf)(auto dest!: Chains_FieldD[OF M] simp add: leq_def) have c: "current \ (Inf M)" using g proof(rule current_leI) show "(Inf M) e \ g0 e" for e using g0 by(auto intro: INF_lower) show "d_OUT (\M) x \ d_IN (\M) x" if "x \ A \" for x by(auto 4 4 simp add: IN_M OUT_M leq_def intro!: INF_mono dest: Chains_FieldD[OF M] intro: currentD_OUT_IN[OF _ that]) qed have INF_le_f: "Inf M e \ f e" for e using g0 by(auto intro!: INF_lower2 g0_le_f) have eq: "\ (TER (Inf M)) = \ (TER f)" using f w INF_le_f proof(rule essential_eq_leI; intro subsetI) fix x assume x: "x \ \ (TER f)" hence "x \ SINK (Inf M)" using d_OUT_mono[of "Inf M" x f, OF INF_le_f] by(auto simp add: SINK.simps) moreover from x have "x \ SAT \ g" if "g \ M" for g using Chains_FieldD[OF M that] by(auto simp add: leq_def) hence "x \ SAT \ (Inf M)" by(auto simp add: SAT.simps IN_M intro!: INF_greatest) ultimately show "x \ TER (Inf M)" by auto qed have w': "wave \ (Inf M)" proof have "separating \ (\ (TER f))" by(rule separating_essential)(rule waveD_separating[OF w]) then show "separating \ (TER (Inf M))" unfolding eq[symmetric] by(rule separating_weakening) auto fix x assume "x \ RF (TER (Inf M))" hence "x \ RF (\ (TER (Inf M)))" unfolding RF_essential . hence "x \ RF (TER f)" unfolding eq RF_essential . hence "d_OUT f x = 0" by(rule waveD_OUT[OF w]) with d_OUT_mono[of _ x f, OF INF_le_f] show "d_OUT (Inf M) x = 0" by (metis le_zero_eq) qed from c w' INF_le_f eq show ?thesis by simp qed define trim1 where "trim1 g = (if trimming g then g else let z = SOME z. z \ RF\<^sup>\ (TER g) \ z \ A \ \ \ KIR g z; factor = d_OUT g z / d_IN g z in (\(y, x). (if x = z then factor else 1) * g (y, x)))" for g have increasing: "trim1 g \ g \ trim1 g \ F" if "g \ F" for g proof(cases "trimming g") case True thus ?thesis using that by(simp add: trim1_def) next case False let ?P = "\z. z \ RF\<^sup>\ (TER g) \ z \ A \ \ \ KIR g z" define z where "z = Eps ?P" from that have g: "current \ g" and w': "wave \ g" and le_f: "\e. g e \ f e" and \: "\ (TER g) = \ (TER f)" by(auto simp add: le_fun_def) { with False obtain z where z: "z \ RF\<^sup>\ (TER f)" and A: "z \ A \" and neq: "d_OUT g z \ d_IN g z" by(auto simp add: trimming.simps le_fun_def) from z have "z \ RF\<^sup>\ (\ (TER f))" unfolding roofed_circ_essential . with \ roofed_circ_essential[of \ "TER g"] have "z \ RF\<^sup>\ (TER g)" by simp with A neq have "\x. ?P x" by auto } hence "?P z" unfolding z_def by(rule someI_ex) hence RF: "z \ RF\<^sup>\ (TER g)" and A: "z \ A \" and neq: "d_OUT g z \ d_IN g z" by simp_all let ?factor = "d_OUT g z / d_IN g z" have trim1 [simp]: "trim1 g (y, x) = (if x = z then ?factor else 1) * g (y, x)" for x y using False by(auto simp add: trim1_def z_def Let_def) from currentD_OUT_IN[OF g A] neq have less: "d_OUT g z < d_IN g z" by auto hence "?factor \ 1" (is "?factor \ _") by (auto intro!: divide_le_posI_ennreal simp: zero_less_iff_neq_zero) hence le': "?factor * g (y, x) \ 1 * g (y, x)" for y x by(rule mult_right_mono) simp hence le: "trim1 g e \ g e" for e by(cases e)simp moreover { have c: "current \ (trim1 g)" using g le proof(rule current_leI) fix x assume x: "x \ A \" have "d_OUT (trim1 g) x \ d_OUT g x" unfolding d_OUT_def using le' by(auto intro: nn_integral_mono) also have "\ \ d_IN (trim1 g) x" proof(cases "x = z") case True have "d_OUT g x = d_IN (trim1 g) x" unfolding d_IN_def using True currentD_weight_IN[OF g, of x] currentD_OUT_IN[OF g x] apply (cases "d_IN g x = 0") apply(auto simp add: nn_integral_divide nn_integral_cmult d_IN_def[symmetric] ennreal_divide_times) apply (subst ennreal_divide_self) apply (auto simp: less_top[symmetric] top_unique weight_finite) done thus ?thesis by simp next case False have "d_OUT g x \ d_IN g x" using x by(rule currentD_OUT_IN[OF g]) also have "\ \ d_IN (trim1 g) x" unfolding d_IN_def using False by(auto intro!: nn_integral_mono) finally show ?thesis . qed finally show "d_OUT (trim1 g) x \ d_IN (trim1 g) x" . qed moreover have le_f: "trim1 g \ f" using le le_f by(blast intro: le_funI order_trans) moreover have eq: "\ (TER (trim1 g)) = \ (TER f)" unfolding \[symmetric] using g w' le proof(rule essential_eq_leI; intro subsetI) fix x assume x: "x \ \ (TER g)" hence "x \ SINK (trim1 g)" using d_OUT_mono[of "trim1 g" x g, OF le] by(auto simp add: SINK.simps) moreover from x have "x \ z" using RF by(auto simp add: roofed_circ_def) hence "d_IN (trim1 g) x = d_IN g x" unfolding d_IN_def by simp with \x \ \ (TER g)\ have "x \ SAT \ (trim1 g)" by(auto simp add: SAT.simps) ultimately show "x \ TER (trim1 g)" by auto qed moreover have "wave \ (trim1 g)" proof have "separating \ (\ (TER f))" by(rule separating_essential)(rule waveD_separating[OF w]) then show "separating \ (TER (trim1 g))" unfolding eq[symmetric] by(rule separating_weakening) auto fix x assume "x \ RF (TER (trim1 g))" hence "x \ RF (\ (TER (trim1 g)))" unfolding RF_essential . hence "x \ RF (TER f)" unfolding eq RF_essential . hence "d_OUT f x = 0" by(rule waveD_OUT[OF w]) with d_OUT_mono[of _ x f, OF le_f[THEN le_funD]] show "d_OUT (trim1 g) x = 0" by (metis le_zero_eq) qed ultimately have "trim1 g \ F" by(simp add: F_def) } ultimately show ?thesis using that by(simp add: le_fun_def del: trim1) qed have "bourbaki_witt_fixpoint Inf leq trim1" using chainD increasing unfolding leq_def by(intro bourbaki_witt_fixpoint_restrict_rel)(auto intro: Inf_greatest Inf_lower) then interpret bourbaki_witt_fixpoint Inf leq trim1 . have f_Field: "f \ Field leq" using f w by(simp add: leq_def) define g where "g = fixp_above f" have "g \ Field leq" using f_Field unfolding g_def by(rule fixp_above_Field) hence le_f: "g \ f" and g: "current \ g" and w': "wave \ g" and TER: "\ (TER g) = \ (TER f)" by(auto simp add: leq_def intro: le_funI) have "trimming g" proof(rule ccontr) let ?P = "\x. x \ RF\<^sup>\ (TER g) \ x \ A \ \ \ KIR g x" define x where "x = Eps ?P" assume False: "\ ?thesis" hence "\x. ?P x" using le_f g w' TER by(auto simp add: trimming.simps roofed_circ_essential[of \ "TER g", symmetric] roofed_circ_essential[of \ "TER f", symmetric]) hence "?P x" unfolding x_def by(rule someI_ex) hence x: "x \ RF\<^sup>\ (TER g)" and A: "x \ A \" and neq: "d_OUT g x \ d_IN g x" by simp_all from neq have "\y. edge \ y x \ g (y, x) > 0" proof(rule contrapos_np) assume "\ ?thesis" hence "d_IN g x = 0" using currentD_outside[OF g, of _ x] by(force simp add: d_IN_def nn_integral_0_iff_AE AE_count_space not_less) with currentD_OUT_IN[OF g A] show "KIR g x" by simp qed then obtain y where y: "edge \ y x" and gr0: "g (y, x) > 0" by blast have [simp]: "g (y, x) \ \" proof - have "g (y, x) \ d_OUT g y" by (rule d_OUT_ge_point) also have "\ \ weight \ y" by(rule currentD_weight_OUT[OF g]) also have "\ < \" by(simp add: weight_finite less_top[symmetric]) finally show ?thesis by simp qed from neq have factor: "d_OUT g x / d_IN g x \ 1" by (simp add: divide_eq_1_ennreal) have "trim1 g (y, x) = g (y, x) * (d_OUT g x / d_IN g x)" by(clarsimp simp add: False trim1_def Let_def x_def[symmetric] mult.commute) moreover have "\ \ g (y, x) * 1" unfolding ennreal_mult_cancel_left using gr0 factor by auto ultimately have "trim1 g (y, x) \ g (y, x)" by auto hence "trim1 g \ g" by(auto simp add: fun_eq_iff) moreover have "trim1 g = g" using f_Field unfolding g_def by(rule fixp_above_unfold[symmetric]) ultimately show False by contradiction qed then show ?thesis by blast qed end lemma trimming_\: fixes \ (structure) assumes w: "wave \ f" and trimming: "trimming \ f g" shows "\ (TER f) = \ (TER g)" proof(rule set_eqI) show "x \ \ (TER f) \ x \ \ (TER g)" for x proof(cases "x \ A \") case False thus ?thesis using trimmingD_\[OF trimming] by blast next case True show ?thesis proof assume x: "x \ \ (TER f)" hence "x \ TER g" using d_OUT_mono[of g x f, OF trimmingD_le[OF trimming]] True by(simp add: SINK.simps SAT.A) moreover from x have "essential \ (B \) (TER f) x" by simp then obtain p y where p: "path \ x p y" and y: "y \ B \" and bypass: "\z. z \ set p \ z \ RF (TER f)" by(rule essentialE_RF) blast from p y have "essential \ (B \) (\ (TER g)) x" proof(rule essentialI) fix z assume "z \ set p" hence z: "z \ RF (TER f)" by(rule bypass) with waveD_separating[OF w, THEN separating_RF_A] have "z \ A \" by blast with z have "z \ \ (TER g)" using trimmingD_\[OF trimming] by(auto intro: roofed_greaterI) thus "z = x \ z \ \ (TER g)" .. qed ultimately show "x \ \ (TER g)" unfolding essential_\ by simp next assume "x \ \ (TER g)" then obtain p y where p: "path \ x p y" and y: "y \ B \" and bypass: "\z. z \ set p \ z \ RF (TER g)" by(rule \_E_RF) blast have z: "z \ \ (TER f)" if "z \ set p" for z proof - from that have z: "z \ RF (TER g)" by(rule bypass) with waveD_separating[OF trimmingD_wave[OF trimming], THEN separating_RF_A] have "z \ A \" by blast with z show "z \ \ (TER f)" using trimmingD_\[OF trimming] by(auto intro: roofed_greaterI) qed then have "essential \ (B \) (\ (TER f)) x" by(intro essentialI[OF p y]) auto moreover have "x \ TER f" using waveD_separating[THEN separating_essential, THEN separatingD, OF w p True y] z by auto ultimately show "x \ \ (TER f)" unfolding essential_\ by simp qed qed qed subsection \Composition of waves via quotients\ definition quotient_web :: "('v, 'more) web_scheme \ 'v current \ ('v, 'more) web_scheme" where \ \Modifications to original Definition 4.9: No incoming edges to nodes in @{const A}, @{term "B \ - A \"} is not part of @{const A} such that @{const A} contains only vertices is disjoint from @{const B}. The weight of vertices in @{const B} saturated by @{term f} is therefore set to @{term "0 :: ennreal"}.\ "quotient_web \ f = \edge = \x y. edge \ x y \ x \ roofed_circ \ (TER\<^bsub>\\<^esub> f) \ y \ roofed \ (TER\<^bsub>\\<^esub> f), weight = \x. if x \ RF\<^sup>\\<^bsub>\\<^esub> (TER\<^bsub>\\<^esub> f) \ x \ TER\<^bsub>\\<^esub> f \ B \ then 0 else weight \ x, A = \\<^bsub>\\<^esub> (TER\<^bsub>\\<^esub> f) - (B \ - A \), B = B \, \ = web.more \\" lemma quotient_web_sel [simp]: fixes \ (structure) shows "edge (quotient_web \ f) x y \ edge \ x y \ x \ RF\<^sup>\ (TER f) \ y \ RF (TER f)" "weight (quotient_web \ f) x = (if x \ RF\<^sup>\ (TER f) \ x \ TER\<^bsub>\\<^esub> f \ B \ then 0 else weight \ x)" "A (quotient_web \ f) = \ (TER f)- (B \ - A \)" "B (quotient_web \ f) = B \" "web.more (quotient_web \ f) = web.more \" by(simp_all add: quotient_web_def) lemma vertex_quotient_webD: fixes \ (structure) shows "vertex (quotient_web \ f) x \ vertex \ x \ x \ RF\<^sup>\ (TER f)" by(auto simp add: vertex_def roofed_circ_def) lemma path_quotient_web: fixes \ (structure) assumes "path \ x p y" and "x \ RF\<^sup>\ (TER f)" and "\z. z \ set p \ z \ RF (TER f)" shows "path (quotient_web \ f) x p y" using assms by(induction)(auto intro: rtrancl_path.intros simp add: roofed_circ_def) definition restrict_current :: "('v, 'more) web_scheme \ 'v current \ 'v current \ 'v current" where "restrict_current \ f g = (\(x, y). g (x, y) * indicator (- RF\<^sup>\\<^bsub>\\<^esub> (TER\<^bsub>\\<^esub> f)) x * indicator (- RF\<^bsub>\\<^esub> (TER\<^bsub>\\<^esub> f)) y)" abbreviation restrict_curr :: "'v current \ ('v, 'more) web_scheme \ 'v current \ 'v current" ("_ \ _ '/ _" [100, 0, 100] 100) where "restrict_curr g \ f \ restrict_current \ f g" lemma restrict_current_simps [simp]: fixes \ (structure) shows "(g \ \ / f) (x, y) = (g (x, y) * indicator (- RF\<^sup>\ (TER f)) x * indicator (- RF (TER f)) y)" by(simp add: restrict_current_def) lemma d_OUT_restrict_current_outside: fixes \ (structure) shows "x \ RF\<^sup>\ (TER f) \ d_OUT (g \ \ / f) x = 0" by(simp add: d_OUT_def) lemma d_IN_restrict_current_outside: fixes \ (structure) shows "x \ RF (TER f) \ d_IN (g \ \ / f) x = 0" by(simp add: d_IN_def) lemma restrict_current_le: "(g \ \ / f) e \ g e" by(cases e)(clarsimp split: split_indicator) lemma d_OUT_restrict_current_le: "d_OUT (g \ \ / f) x \ d_OUT g x" unfolding d_OUT_def by(rule nn_integral_mono, simp split: split_indicator) lemma d_IN_restrict_current_le: "d_IN (g \ \ / f) x \ d_IN g x" unfolding d_IN_def by(rule nn_integral_mono, simp split: split_indicator) lemma restrict_current_IN_not_RF: fixes \ (structure) assumes g: "current \ g" and x: "x \ RF (TER f)" shows "d_IN (g \ \ / f) x = d_IN g x" proof - { fix y assume y: "y \ RF\<^sup>\ (TER f)" have "g (y, x) = 0" proof(cases "edge \ y x") case True from y have y': "y \ RF (TER f)" and essential: "y \ \ (TER f)" by(simp_all add: roofed_circ_def) moreover from x obtain p z where z: "z \ B \" and p: "path \ x p z" and bypass: "\z. z \ set p \ z \ TER f" "x \ TER f" by(clarsimp simp add: roofed_def) from roofedD[OF y' rtrancl_path.step, OF True p z] bypass have "x \ TER f \ y \ TER f" by auto with roofed_greater[THEN subsetD, of x "TER f" \] x have "x \ TER f" "y \ TER f" by auto with essential bypass have False by(auto dest!: not_essentialD[OF _ rtrancl_path.step, OF _ True p z]) thus ?thesis .. qed(simp add: currentD_outside[OF g]) } then show ?thesis unfolding d_IN_def using x by(auto intro!: nn_integral_cong split: split_indicator) qed lemma restrict_current_IN_A: "a \ A (quotient_web \ f) \ d_IN (g \ \ / f) a = 0" by(simp add: d_IN_restrict_current_outside roofed_greaterI) lemma restrict_current_nonneg: "0 \ g e \ 0 \ (g \ \ / f) e" by(cases e) simp lemma in_SINK_restrict_current: "x \ SINK g \ x \ SINK (g \ \ / f)" using d_OUT_restrict_current_le[of \ f g x] by(simp add: SINK.simps) lemma SAT_restrict_current: fixes \ (structure) assumes f: "current \ f" and g: "current \ g" shows "SAT (quotient_web \ f) (g \ \ / f) = RF (TER f) \ (SAT \ g - A \)" (is "SAT ?\ ?g = ?rhs") proof(intro set_eqI iffI; (elim UnE DiffE)?) show "x \ ?rhs" if "x \ SAT ?\ ?g" for x using that proof cases case IN thus ?thesis using currentD_weight_OUT[OF f, of x] by(cases "x \ RF (TER f)")(auto simp add: d_IN_restrict_current_outside roofed_circ_def restrict_current_IN_not_RF[OF g] SAT.IN currentD_IN[OF g] roofed_greaterI SAT.A SINK.simps RF_in_B split: if_split_asm intro: essentialI[OF rtrancl_path.base]) qed(simp add: roofed_greaterI) show "x \ SAT ?\ ?g" if "x \ RF (TER f)" for x using that by(auto simp add: SAT.simps roofed_circ_def d_IN_restrict_current_outside) show "x \ SAT ?\ ?g" if "x \ SAT \ g" "x \ A \" for x using that by(auto simp add: SAT.simps roofed_circ_def d_IN_restrict_current_outside restrict_current_IN_not_RF[OF g]) qed lemma current_restrict_current: fixes \ (structure) assumes w: "wave \ f" and g: "current \ g" shows "current (quotient_web \ f) (g \ \ / f)" (is "current ?\ ?g") proof show "d_OUT ?g x \ weight ?\ x" for x using d_OUT_restrict_current_le[of \ f g x] currentD_weight_OUT[OF g, of x] currentD_OUT[OF g, of x] by(auto simp add: d_OUT_restrict_current_outside) show "d_IN ?g x \ weight ?\ x" for x using d_IN_restrict_current_le[of \ f g x] currentD_weight_IN[OF g, of x] by(auto simp add: d_IN_restrict_current_outside roofed_circ_def) (subst d_IN_restrict_current_outside[of x \ f g]; simp add: roofed_greaterI) fix x assume "x \ A ?\" hence x: "x \ \ (TER f) - B \" by simp show "d_OUT ?g x \ d_IN ?g x" proof(cases "x \ RF (TER f)") case True with x have "x \ RF\<^sup>\ (TER f) \ B \" by(simp add: roofed_circ_def) with True show ?thesis using currentD_OUT[OF g, of x] d_OUT_restrict_current_le[of \ f g x] by(auto simp add: d_OUT_restrict_current_outside d_IN_restrict_current_outside) next case False then obtain p z where z: "z \ B \" and p: "path \ x p z" and bypass: "\z. z \ set p \ z \ TER f" "x \ TER f" by(clarsimp simp add: roofed_def) from g False have "d_IN ?g x = d_IN g x" by(rule restrict_current_IN_not_RF) moreover have "d_OUT ?g x \ d_OUT g x" by(rule d_OUT_mono restrict_current_le)+ moreover have "x \ A \" using separatingD[OF waveD_separating[OF w] p _ z] bypass by blast note currentD_OUT_IN[OF g this] ultimately show ?thesis by simp qed next show "d_IN ?g a = 0" if "a \ A ?\" for a using that by(rule restrict_current_IN_A) show "d_OUT ?g b = 0" if "b \ B ?\" for b using d_OUT_restrict_current_le[of \ f g b] currentD_OUT[OF g, of b] that by simp show "?g e = 0" if "e \ \<^bold>E\<^bsub>?\\<^esub>" for e using that currentD_outside'[OF g, of e] by(cases e)(auto split: split_indicator) qed lemma TER_restrict_current: fixes \ (structure) assumes f: "current \ f" and w: "wave \ f" and g: "current \ g" shows "TER g \ TER\<^bsub>quotient_web \ f\<^esub> (g \ \ / f)" (is "_ \ ?TER" is "_ \ TER\<^bsub>?\\<^esub> ?g") proof fix x assume x: "x \ TER g" hence "x \ SINK ?g" by(simp add: in_SINK_restrict_current) moreover have "x \ RF (TER f)" if "x \ A \" using waveD_separating[OF w, THEN separatingD, OF _ that] by(rule roofedI) then have "x \ SAT ?\ ?g" using SAT_restrict_current[OF f g] x by auto ultimately show "x \ ?TER" by simp qed lemma wave_restrict_current: fixes \ (structure) assumes f: "current \ f" and w: "wave \ f" and g: "current \ g" and w': "wave \ g" shows "wave (quotient_web \ f) (g \ \ / f)" (is "wave ?\ ?g") proof show "separating ?\ (TER\<^bsub>?\\<^esub> ?g)" (is "separating _ ?TER") proof fix x y p assume "x \ A ?\" "y \ B ?\" and p: "path ?\ x p y" hence x: "x \ \ (TER f)" and y: "y \ B \" and SAT: "x \ SAT ?\ ?g" by(simp_all add: SAT.A) from p have p': "path \ x p y" by(rule rtrancl_path_mono) simp { assume "x \ ?TER" hence "x \ SINK ?g" using SAT by(simp) hence "x \ SINK g" using d_OUT_restrict_current_le[of \ f g x] by(auto simp add: SINK.simps) hence "x \ RF (TER g)" using waveD_OUT[OF w'] by(auto simp add: SINK.simps) from roofedD[OF this p' y] \x \ SINK g\ have "(\z\set p. z \ ?TER)" using TER_restrict_current[OF f w g] by blast } then show "(\z\set p. z \ ?TER) \ x \ ?TER" by blast qed fix x assume "x \ RF\<^bsub>?\\<^esub> ?TER" hence "x \ RF (TER g)" proof(rule contrapos_nn) assume *: "x \ RF (TER g)" show "x \ RF\<^bsub>?\\<^esub> ?TER" proof fix p y assume "path ?\ x p y" "y \ B ?\" hence "path \ x p y" "y \ B \" by(auto elim: rtrancl_path_mono) from roofedD[OF * this] show "(\z\set p. z \ ?TER) \ x \ ?TER" using TER_restrict_current[OF f w g] by blast qed qed with w' have "d_OUT g x = 0" by(rule waveD_OUT) with d_OUT_restrict_current_le[of \ f g x] show "d_OUT ?g x = 0" by simp qed definition plus_current :: "'v current \ 'v current \ 'v current" where "plus_current f g = (\e. f e + g e)" lemma plus_current_simps [simp]: "plus_current f g e = f e + g e" by(simp add: plus_current_def) lemma plus_zero_current [simp]: "plus_current f zero_current = f" by(simp add: fun_eq_iff) lemma support_flow_plus_current: "support_flow (plus_current f g) \ support_flow f \ support_flow g" by(clarsimp simp add: support_flow.simps) context fixes \ :: "('v, 'more) web_scheme" (structure) and f g assumes f: "current \ f" and w: "wave \ f" and g: "current (quotient_web \ f) g" begin lemma OUT_plus_current: "d_OUT (plus_current f g) x = (if x \ RF\<^sup>\ (TER f) then d_OUT f x else d_OUT g x)" (is "d_OUT ?g _ = _") proof - have "d_OUT ?g x = d_OUT f x + d_OUT g x" unfolding plus_current_def by(subst d_OUT_add) simp_all also have "\ = (if x \ RF\<^sup>\ (TER f) then d_OUT f x else d_OUT g x)" proof(cases "x \ RF\<^sup>\ (TER f)") case True hence "d_OUT g x = 0" by(intro currentD_outside_OUT[OF g])(auto dest: vertex_quotient_webD) thus ?thesis using True by simp next case False hence "d_OUT f x = 0" by(auto simp add: roofed_circ_def SINK.simps dest: waveD_OUT[OF w]) with False show ?thesis by simp qed finally show ?thesis . qed lemma IN_plus_current: "d_IN (plus_current f g) x = (if x \ RF (TER f) then d_IN f x else d_IN g x)" (is "d_IN ?g _ = _") proof - have "d_IN ?g x = d_IN f x + d_IN g x" unfolding plus_current_def by(subst d_IN_add) simp_all also consider (RF) "x \ RF (TER f) - (B \ - A \)" | (B) "x \ RF (TER f)" "x \ B \ - A \" | (beyond) "x \ RF (TER f)" by blast then have "d_IN f x + d_IN g x = (if x \ RF (TER f) then d_IN f x else d_IN g x)" proof(cases) case RF hence "d_IN g x = 0" by(cases "x \ \ (TER f)")(auto intro: currentD_outside_IN[OF g] currentD_IN[OF g] dest!: vertex_quotient_webD simp add: roofed_circ_def) thus ?thesis using RF by simp next case B hence "d_IN g x = 0" using currentD_outside_IN[OF g, of x] currentD_weight_IN[OF g, of x] by(auto dest: vertex_quotient_webD simp add: roofed_circ_def) with B show ?thesis by simp next case beyond from f w beyond have "d_IN f x = 0" by(rule wave_not_RF_IN_zero) with beyond show ?thesis by simp qed finally show ?thesis . qed lemma in_TER_plus_current: assumes RF: "x \ RF\<^sup>\ (TER f)" and x: "x \ TER\<^bsub>quotient_web \ f\<^esub> g" (is "_ \ ?TER _") shows "x \ TER (plus_current f g)" (is "_ \ TER ?g") proof(cases "x \ \ (TER f) - (B \ - A \)") case True with x show ?thesis using currentD_IN[OF g, of x] by(fastforce intro: roofed_greaterI SAT.intros simp add: SINK.simps OUT_plus_current IN_plus_current elim!: SAT.cases) next case *: False have "x \ SAT \ ?g" proof(cases "x \ B \ - A \") case False with x RF * have "weight \ x \ d_IN g x" by(auto elim!: SAT.cases split: if_split_asm simp add: essential_BI) also have "\ \ d_IN ?g x" unfolding plus_current_def by(intro d_IN_mono) simp finally show ?thesis .. next case True with * x have "weight \ x \ d_IN ?g x" using currentD_OUT[OF f, of x] by(auto simp add: IN_plus_current RF_in_B SINK.simps roofed_circ_def elim!: SAT.cases split: if_split_asm) thus ?thesis .. qed moreover have "x \ SINK ?g" using x by(simp add: SINK.simps OUT_plus_current RF) ultimately show ?thesis by simp qed lemma current_plus_current: "current \ (plus_current f g)" (is "current _ ?g") proof show "d_OUT ?g x \ weight \ x" for x using currentD_weight_OUT[OF g, of x] currentD_weight_OUT[OF f, of x] by(auto simp add: OUT_plus_current split: if_split_asm elim: order_trans) show "d_IN ?g x \ weight \ x" for x using currentD_weight_IN[OF f, of x] currentD_weight_IN[OF g, of x] by(auto simp add: IN_plus_current roofed_circ_def split: if_split_asm elim: order_trans) show "d_OUT ?g x \ d_IN ?g x" if "x \ A \" for x proof(cases "x \ \ (TER f)") case False thus ?thesis using currentD_OUT_IN[OF f that] currentD_OUT_IN[OF g, of x] that by(auto simp add: OUT_plus_current IN_plus_current roofed_circ_def SINK.simps) next case True with that have "d_OUT f x = 0" "weight \ x \ d_IN f x" by(auto simp add: SINK.simps elim: SAT.cases) thus ?thesis using that True currentD_OUT_IN[OF g, of x] currentD_weight_OUT[OF g, of x] by(auto simp add: OUT_plus_current IN_plus_current roofed_circ_def intro: roofed_greaterI split: if_split_asm) qed show "d_IN ?g a = 0" if "a \ A \" for a using wave_A_in_RF[OF w that] currentD_IN[OF f that] by(simp add: IN_plus_current) show "d_OUT ?g b = 0" if "b \ B \" for b using that currentD_OUT[OF f that] currentD_OUT[OF g, of b] that by(auto simp add: OUT_plus_current SINK.simps roofed_circ_def intro: roofed_greaterI) show "?g e = 0" if "e \ \<^bold>E" for e using currentD_outside'[OF f, of e] currentD_outside'[OF g, of e] that by(cases e) auto qed context assumes w': "wave (quotient_web \ f) g" begin lemma separating_TER_plus_current: assumes x: "x \ RF (TER f)" and y: "y \ B \" and p: "path \ x p y" shows "(\z\set p. z \ TER (plus_current f g)) \ x \ TER (plus_current f g)" (is "_ \ _ \ TER ?g") proof - from x have "x \ RF (\ (TER f))" unfolding RF_essential . from roofedD[OF this p y] have "\z\set (x # p). z \ \ (TER f)" by auto from split_list_last_prop[OF this] obtain ys z zs where decomp: "x # p = ys @ z # zs" and z: "z \ \ (TER f)" and outside: "\z. z \ set zs \ z \ \ (TER f)" by auto have zs: "path \ z zs y" using decomp p by(cases ys)(auto elim: rtrancl_path_appendE) moreover have "z \ RF\<^sup>\ (TER f)" using z by(simp add: roofed_circ_def) moreover have RF: "z' \ RF (TER f)" if "z' \ set zs" for z' proof assume "z' \ RF (TER f)" hence z': "z' \ RF (\ (TER f))" by(simp only: RF_essential) from split_list[OF that] obtain ys' zs' where decomp': "zs = ys' @ z' # zs'" by blast with zs have "path \ z' zs' y" by(auto elim: rtrancl_path_appendE) from roofedD[OF z' this y] outside decomp' show False by auto qed ultimately have p': "path (quotient_web \ f) z zs y" by(rule path_quotient_web) show ?thesis proof(cases "z \ B \ - A \") case False with separatingD[OF waveD_separating[OF w'] p'] z y obtain z' where z': "z' \ set (z # zs)" and TER: "z' \ TER\<^bsub>quotient_web \ f\<^esub> g" by auto hence "z' \ TER ?g" using in_TER_plus_current[of z'] RF[of z'] \z \ RF\<^sup>\ (TER f)\ by(auto simp add: roofed_circ_def) with decomp z' show ?thesis by(cases ys) auto next case True hence "z \ TER ?g" using currentD_OUT[OF current_plus_current, of z] z by(auto simp add: SINK.simps SAT.simps IN_plus_current intro: roofed_greaterI) then show ?thesis using decomp by(cases ys) auto qed qed lemma wave_plus_current: "wave \ (plus_current f g)" (is "wave _ ?g") proof let ?\ = "quotient_web \ f" let ?TER = "TER\<^bsub>?\\<^esub>" show "separating \ (TER ?g)" using separating_TER_plus_current[OF wave_A_in_RF[OF w]] by(rule separating) fix x assume x: "x \ RF (TER ?g)" hence "x \ RF (TER f)" by(rule contrapos_nn)(rule roofedI, rule separating_TER_plus_current) hence *: "x \ RF\<^sup>\ (TER f)" by(simp add: roofed_circ_def) moreover have "x \ RF\<^bsub>?\\<^esub> (?TER g)" proof assume RF': "x \ RF\<^bsub>?\\<^esub> (?TER g)" from x obtain p y where y: "y \ B \" and p: "path \ x p y" and bypass: "\z. z \ set p \ z \ TER ?g" and x': "x \ TER ?g" by(auto simp add: roofed_def) have RF: "z \ RF (TER f)" if "z \ set p" for z proof assume z: "z \ RF (TER f)" from split_list[OF that] obtain ys' zs' where decomp: "p = ys' @ z # zs'" by blast with p have "path \ z zs' y" by(auto elim: rtrancl_path_appendE) from separating_TER_plus_current[OF z y this] decomp bypass show False by auto qed with p have "path ?\ x p y" using * by(induction)(auto intro: rtrancl_path.intros simp add: roofed_circ_def) from roofedD[OF RF' this] y consider (x) "x \ ?TER g" | (z) z where "z \ set p" "z \ ?TER g" by auto then show False proof(cases) case x with * have "x \ TER ?g" by(rule in_TER_plus_current) with x' show False by contradiction next case (z z) from z(1) have "z \ RF (TER f)" by(rule RF) hence "z \ RF\<^sup>\ (TER f)" by(simp add: roofed_circ_def) hence "z \ TER ?g" using z(2) by(rule in_TER_plus_current) moreover from z(1) have "z \ TER ?g" by(rule bypass) ultimately show False by contradiction qed qed with w' have "d_OUT g x = 0" by(rule waveD_OUT) ultimately show "d_OUT ?g x = 0" by(simp add: OUT_plus_current) qed end end lemma loose_quotient_web: fixes \ :: "('v, 'more) web_scheme" (structure) assumes weight_finite: "\x. weight \ x \ \" and f: "current \ f" and w: "wave \ f" and maximal: "\w. \ current \ w; wave \ w; f \ w \ \ f = w" shows "loose (quotient_web \ f)" (is "loose ?\") proof fix g assume g: "current ?\ g" and w': "wave ?\ g" let ?g = "plus_current f g" from f w g have "current \ ?g" "wave \ ?g" by(rule current_plus_current wave_plus_current)+ (rule w') moreover have "f \ ?g" by(clarsimp simp add: le_fun_def add_eq_0_iff_both_eq_0) ultimately have eq: "f = ?g" by(rule maximal) have "g e = 0" for e proof(cases e) case (Pair x y) have "f e \ d_OUT f x" unfolding Pair by (rule d_OUT_ge_point) also have "\ \ weight \ x" by(rule currentD_weight_OUT[OF f]) also have "\ < \" by(simp add: weight_finite less_top[symmetric]) finally show "g e = 0" using Pair eq[THEN fun_cong, of e] by(cases "f e" "g e" rule: ennreal2_cases)(simp_all add: fun_eq_iff) qed thus "g = (\_. 0)" by(simp add: fun_eq_iff) next - have 0: "current ?\ zero_current" bysimp + have 0: "current ?\ zero_current" by simp show "\ hindrance ?\ zero_current" proof assume "hindrance ?\ zero_current" then obtain x where a: "x \ A ?\" and \: "x \ \\<^bsub>?\\<^esub> (TER\<^bsub>?\\<^esub> zero_current)" and "d_OUT zero_current x < weight ?\ x" by cases from a have "x \ \ (TER f)" by simp then obtain p y where p: "path \ x p y" and y: "y \ B \" and bypass: "\z. \x \ y; z \ set p\ \ z = x \ z \ TER f" by(rule \_E) blast from p obtain p' where p': "path \ x p' y" and distinct: "distinct (x # p')" and subset: "set p' \ set p" by(auto elim: rtrancl_path_distinct) note p' moreover have RF: "z \ RF (TER f)" if "z \ set p'" for z proof assume z: "z \ RF (TER f)" from split_list[OF that] obtain ys zs where decomp: "p' = ys @ z # zs" by blast with p' have "y \ set p'" by(auto dest!: rtrancl_path_last intro: last_in_set) with distinct have neq: "x \ y" by auto from decomp p' have "path \ z zs y" by(auto elim: rtrancl_path_appendE) from roofedD[OF z this y] obtain z' where "z' \ set (z # zs)" "z' \ TER f" by auto with distinct decomp subset bypass[OF neq] show False by auto qed moreover have "x \ RF\<^sup>\ (TER f)" using \x \ \ (TER f)\ by(simp add: roofed_circ_def) ultimately have p'': "path ?\ x p' y" by(induction)(auto intro: rtrancl_path.intros simp add: roofed_circ_def) from a \ have "\ essential ?\ (B ?\) (TER\<^bsub>?\\<^esub> zero_current) x" by simp from not_essentialD[OF this p''] y obtain z where neq: "x \ y" and "z \ set p'" "z \ x" "z \ TER\<^bsub>?\\<^esub> zero_current" by auto moreover with subset RF[of z] have "z \ TER f" using currentD_weight_OUT[OF f, of z] currentD_weight_IN[OF f, of z] by(auto simp add: roofed_circ_def SINK.simps intro: SAT.IN split: if_split_asm) ultimately show False using bypass[of z] subset by auto qed qed lemma quotient_web_trimming: fixes \ (structure) assumes w: "wave \ f" and trimming: "trimming \ f g" shows "quotient_web \ f = quotient_web \ g" (is "?lhs = ?rhs") proof(rule web.equality) from trimming have \: "\ (TER g) - A \ = \ (TER f) - A \" by cases have RF: "RF (TER g) = RF (TER f)" by(subst (1 2) RF_essential[symmetric])(simp only: trimming_\[OF w trimming]) have RFc: "RF\<^sup>\ (TER g) = RF\<^sup>\ (TER f)" by(subst (1 2) roofed_circ_essential[symmetric])(simp only: trimming_\[OF w trimming]) show "edge ?lhs = edge ?rhs" by(rule ext)+(simp add: RF RFc) have "weight ?lhs = (\x. if x \ RF\<^sup>\ (TER g) \ x \ RF (TER g) \ B \ then 0 else weight \ x)" unfolding RF RFc by(auto simp add: fun_eq_iff RF_in_B) also have "\ = weight ?rhs" by(auto simp add: fun_eq_iff RF_in_B) finally show "weight ?lhs = weight ?rhs" . show "A ?lhs = A ?rhs" unfolding quotient_web_sel trimming_\[OF w trimming] .. qed simp_all subsection \Well-formed webs\ locale web = fixes \ :: "('v, 'more) web_scheme" (structure) assumes A_in: "x \ A \ \ \ edge \ y x" and B_out: "x \ B \ \ \ edge \ x y" and A_vertex: "A \ \ \<^bold>V" and disjoint: "A \ \ B \ = {}" and no_loop: "\x. \ edge \ x x" and weight_outside: "\x. x \ \<^bold>V \ weight \ x = 0" and weight_finite [simp]: "\x. weight \ x \ \" begin lemma web_weight_update: assumes "\x. \ vertex \ x \ w x = 0" and "\x. w x \ \" shows "web (\\weight := w\)" by unfold_locales(simp_all add: A_in B_out A_vertex disjoint no_loop assms) lemma currentI [intro?]: assumes "\x. d_OUT f x \ weight \ x" and "\x. d_IN f x \ weight \ x" and OUT_IN: "\x. \ x \ A \; x \ B \ \ \ d_OUT f x \ d_IN f x" and outside: "\e. e \ \<^bold>E \ f e = 0" shows "current \ f" proof show "d_IN f a = 0" if "a \ A \" for a using that by(auto simp add: d_IN_def nn_integral_0_iff emeasure_count_space_eq_0 A_in intro: outside) show "d_OUT f b = 0" if "b \ B \" for b using that by(auto simp add: d_OUT_def nn_integral_0_iff emeasure_count_space_eq_0 B_out intro: outside) then show "d_OUT f x \ d_IN f x" if "x \ A \" for x using OUT_IN[OF that] by(cases "x \ B \") auto qed(blast intro: assms)+ lemma currentD_finite_IN: assumes f: "current \ f" shows "d_IN f x \ \" proof(cases "x \ \<^bold>V") case True have "d_IN f x \ weight \ x" using f by(rule currentD_weight_IN) also have "\ < \" using True weight_finite[of x] by (simp add: less_top[symmetric]) finally show ?thesis by simp next case False then have "d_IN f x = 0" by(auto simp add: d_IN_def nn_integral_0_iff emeasure_count_space_eq_0 vertex_def intro: currentD_outside[OF f]) thus ?thesis by simp qed lemma currentD_finite_OUT: assumes f: "current \ f" shows "d_OUT f x \ \" proof(cases "x \ \<^bold>V") case True have "d_OUT f x \ weight \ x" using f by(rule currentD_weight_OUT) also have "\ < \" using True weight_finite[of x] by (simp add: less_top[symmetric]) finally show ?thesis by simp next case False then have "d_OUT f x = 0" by(auto simp add: d_OUT_def nn_integral_0_iff emeasure_count_space_eq_0 vertex_def intro: currentD_outside[OF f]) thus ?thesis by simp qed lemma currentD_finite: assumes f: "current \ f" shows "f e \ \" proof(cases e) case (Pair x y) have "f (x, y) \ d_OUT f x" by (rule d_OUT_ge_point) also have "\ < \" using currentD_finite_OUT[OF f] by (simp add: less_top[symmetric]) finally show ?thesis by(simp add: Pair) qed lemma web_quotient_web: "web (quotient_web \ f)" (is "web ?\") proof show "\ edge ?\ y x" if "x \ A ?\" for x y using that by(auto intro: roofed_greaterI) show "\ edge ?\ x y" if "x \ B ?\" for x y using that by(auto simp add: B_out) show "A ?\ \ B ?\ = {}" using disjoint by auto show "A ?\ \ \<^bold>V\<^bsub>?\\<^esub>" proof fix x assume "x \ A ?\" hence \: "x \ \ (TER f)" and x: "x \ B \" using disjoint by auto from this(1) obtain p y where p: "path \ x p y" and y: "y \ B \" and bypass: "\z. z \ set p \ z \ RF (TER f)" by(rule \_E_RF) blast from p y x have "p \ []" by(auto simp add: rtrancl_path_simps) with rtrancl_path_nth[OF p, of 0] have "edge \ x (p ! 0)" by simp moreover have "x \ RF\<^sup>\ (TER f)" using \ by(simp add: roofed_circ_def) moreover have "p ! 0 \ RF (TER f)" using bypass \p \ []\ by auto ultimately have "edge ?\ x (p ! 0)" by simp thus "x \ \<^bold>V\<^bsub>?\\<^esub>" by(auto intro: vertexI1) qed show "\ edge ?\ x x" for x by(simp add: no_loop) show "weight ?\ x = 0" if "x \ \<^bold>V\<^bsub>?\\<^esub>" for x proof(cases "x \ RF\<^sup>\ (TER f) \ x \ TER f \ B \") case True thus ?thesis by simp next case False hence RF: "x \ RF\<^sup>\ (TER f)" and B: "x \ B \ \ x \ TER f" by auto from RF obtain p y where p: "path \ x p y" and y: "y \ B \" and bypass: "\z. z \ set p \ z \ RF (TER f)" apply(cases "x \ RF (RF (TER f))") apply(auto elim!: not_roofedE)[1] apply(auto simp add: roofed_circ_def roofed_idem elim: essentialE_RF) done from that have "p = []" using p y B RF bypass by(auto 4 3 simp add: vertex_def dest!: rtrancl_path_nth[where i=0]) with p have xy: "x = y" by(simp add: rtrancl_path_simps) with B y have "x \ TER f" by simp hence RF': "x \ RF (TER f)" using xy y by(subst RF_in_B) auto have "\ vertex \ x" proof assume "vertex \ x" then obtain x' where "edge \ x' x" using xy y by(auto simp add: vertex_def B_out) moreover hence "x' \ RF\<^sup>\ (TER f)" using RF' by(auto dest: RF_circ_edge_forward) ultimately have "edge ?\ x' x" using RF' by simp hence "vertex ?\ x" by(rule vertexI2) with that show False by simp qed thus ?thesis by(simp add: weight_outside) qed show "weight ?\ x \ \" for x by simp qed end locale countable_web = web \ for \ :: "('v, 'more) web_scheme" (structure) + assumes countable [simp]: "countable \<^bold>E" begin lemma countable_V [simp]: "countable \<^bold>V" by(simp add: "\<^bold>V_def") lemma countable_web_quotient_web: "countable_web (quotient_web \ f)" (is "countable_web ?\") proof - interpret r: web ?\ by(rule web_quotient_web) show ?thesis proof have "\<^bold>E\<^bsub>?\\<^esub> \ \<^bold>E" by auto then show "countable \<^bold>E\<^bsub>?\\<^esub>" by(rule countable_subset) simp qed qed end subsection \Subtraction of a wave\ definition minus_web :: "('v, 'more) web_scheme \ 'v current \ ('v, 'more) web_scheme" (infixl "\" 65) \ \Definition 6.6\ where "\ \ f = \\weight := \x. if x \ A \ then weight \ x - d_OUT f x else weight \ x + d_OUT f x - d_IN f x\" lemma minus_web_sel [simp]: "edge (\ \ f) = edge \" "weight (\ \ f) x = (if x \ A \ then weight \ x - d_OUT f x else weight \ x + d_OUT f x - d_IN f x)" "A (\ \ f) = A \" "B (\ \ f) = B \" "\<^bold>V\<^bsub>\ \ f\<^esub> = \<^bold>V\<^bsub>\\<^esub>" "\<^bold>E\<^bsub>\ \ f\<^esub> = \<^bold>E\<^bsub>\\<^esub>" "web.more (\ \ f) = web.more \" by(auto simp add: minus_web_def vertex_def) lemma vertex_minus_web [simp]: "vertex (\ \ f) = vertex \" by(simp add: vertex_def fun_eq_iff) lemma roofed_gen_minus_web [simp]: "roofed_gen (\ \ f) = roofed_gen \" by(simp add: fun_eq_iff roofed_def) lemma minus_zero_current [simp]: "\ \ zero_current = \" by(rule web.equality)(simp_all add: fun_eq_iff) lemma (in web) web_minus_web: assumes f: "current \ f" shows "web (\ \ f)" unfolding minus_web_def apply(rule web_weight_update) apply(auto simp: weight_outside currentD_weight_IN[OF f] currentD_outside_OUT[OF f] currentD_outside_IN[OF f] currentD_weight_OUT[OF f] currentD_finite_OUT[OF f]) done subsection \Bipartite webs\ locale countable_bipartite_web = fixes \ :: "('v, 'more) web_scheme" (structure) assumes bipartite_V: "\<^bold>V \ A \ \ B \" and A_vertex: "A \ \ \<^bold>V" and bipartite_E: "edge \ x y \ x \ A \ \ y \ B \" and disjoint: "A \ \ B \ = {}" and weight_outside: "\x. x \ \<^bold>V \ weight \ x = 0" and weight_finite [simp]: "\x. weight \ x \ \" and countable_E [simp]: "countable \<^bold>E" begin lemma not_vertex: "\ x \ A \; x \ B \ \ \ \ vertex \ x" using bipartite_V by blast lemma no_loop: "\ edge \ x x" using disjoint by(auto dest: bipartite_E) lemma edge_antiparallel: "edge \ x y \ \ edge \ y x" using disjoint by(auto dest: bipartite_E) lemma A_in: "x \ A \ \ \ edge \ y x" using disjoint by(auto dest: bipartite_E) lemma B_out: "x \ B \ \ \ edge \ x y" using disjoint by(auto dest: bipartite_E) sublocale countable_web using disjoint by(unfold_locales)(auto simp add: A_in B_out A_vertex no_loop weight_outside) lemma currentD_OUT': assumes f: "current \ f" and x: "x \ A \" shows "d_OUT f x = 0" using currentD_outside_OUT[OF f, of x] x currentD_OUT[OF f, of x] bipartite_V by auto lemma currentD_IN': assumes f: "current \ f" and x: "x \ B \" shows "d_IN f x = 0" using currentD_outside_IN[OF f, of x] x currentD_IN[OF f, of x] bipartite_V by auto lemma current_bipartiteI [intro?]: assumes OUT: "\x. d_OUT f x \ weight \ x" and IN: "\x. d_IN f x \ weight \ x" and outside: "\e. e \ \<^bold>E \ f e = 0" shows "current \ f" proof fix x assume "x \ A \" "x \ B \" hence "d_OUT f x = 0" by(auto simp add: d_OUT_def nn_integral_0_iff emeasure_count_space_eq_0 intro!: outside dest: bipartite_E) then show "d_OUT f x \ d_IN f x" by simp qed(rule OUT IN outside)+ lemma wave_bipartiteI [intro?]: assumes sep: "separating \ (TER f)" and f: "current \ f" shows "wave \ f" using sep proof(rule wave.intros) fix x assume "x \ RF (TER f)" then consider "x \ \<^bold>V" | "x \ \<^bold>V" "x \ B \" using separating_RF_A[OF sep] bipartite_V by auto then show "d_OUT f x = 0" using currentD_OUT[OF f, of x] currentD_outside_OUT[OF f, of x] by cases auto qed lemma web_flow_iff: "web_flow \ f \ current \ f" using bipartite_V by(auto simp add: web_flow.simps) end end diff --git a/thys/SC_DOM_Components/Core_DOM_SC_DOM_Components.thy b/thys/SC_DOM_Components/Core_DOM_SC_DOM_Components.thy --- a/thys/SC_DOM_Components/Core_DOM_SC_DOM_Components.thy +++ b/thys/SC_DOM_Components/Core_DOM_SC_DOM_Components.thy @@ -1,3758 +1,3758 @@ (*********************************************************************************** * Copyright (c) 2016-2020 The University of Sheffield, UK * 2019-2020 University of Exeter, UK * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are met: * * * Redistributions of source code must retain the above copyright notice, this * list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above copyright notice, * this list of conditions and the following disclaimer in the documentation * and/or other materials provided with the distribution. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS" * AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE * IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE * DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE * FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL * DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR * SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER * CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, * OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. * * SPDX-License-Identifier: BSD-2-Clause ***********************************************************************************) section \Core SC DOM Components II\ theory Core_DOM_SC_DOM_Components imports Core_DOM_DOM_Components begin declare [[smt_timeout=2400]] section \Scope Components\ subsection \Definition\ locale l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_disconnected_nodes_defs get_disconnected_nodes get_disconnected_nodes_locs + l_get_owner_document_defs get_owner_document + l_to_tree_order_defs to_tree_order for get_owner_document :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and to_tree_order :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr list) prog" begin definition a_get_scdom_component :: "(_) object_ptr \ (_, (_) object_ptr list) dom_prog" where "a_get_scdom_component ptr = do { document \ get_owner_document ptr; disc_nodes \ get_disconnected_nodes document; tree_order \ to_tree_order (cast document); disconnected_tree_orders \ map_M (to_tree_order \ cast) disc_nodes; return (tree_order @ (concat disconnected_tree_orders)) }" definition a_is_strongly_scdom_component_safe :: "(_) object_ptr set \ (_) object_ptr set \ (_) heap \ (_) heap \ bool" where "a_is_strongly_scdom_component_safe S\<^sub>a\<^sub>r\<^sub>g S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t h h' = ( let removed_pointers = fset (object_ptr_kinds h) - fset (object_ptr_kinds h') in let added_pointers = fset (object_ptr_kinds h') - fset (object_ptr_kinds h) in let arg_components = (\ptr \ (\ptr \ S\<^sub>a\<^sub>r\<^sub>g. set |h \ a_get_scdom_component ptr|\<^sub>r) \ fset (object_ptr_kinds h). set |h \ a_get_scdom_component ptr|\<^sub>r) in let arg_components' = (\ptr \ (\ptr \ S\<^sub>a\<^sub>r\<^sub>g. set |h \ a_get_scdom_component ptr|\<^sub>r) \ fset (object_ptr_kinds h'). set |h' \ a_get_scdom_component ptr|\<^sub>r) in removed_pointers \ arg_components \ added_pointers \ arg_components' \ S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t \ arg_components' \ (\outside_ptr \ fset (object_ptr_kinds h) \ fset (object_ptr_kinds h') - (\ptr \ S\<^sub>a\<^sub>r\<^sub>g. set |h \ a_get_scdom_component ptr|\<^sub>r). preserved (get_M outside_ptr id) h h'))" definition a_is_weakly_scdom_component_safe :: "(_) object_ptr set \ (_) object_ptr set \ (_) heap \ (_) heap \ bool" where "a_is_weakly_scdom_component_safe S\<^sub>a\<^sub>r\<^sub>g S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t h h' = ( let removed_pointers = fset (object_ptr_kinds h) - fset (object_ptr_kinds h') in let added_pointers = fset (object_ptr_kinds h') - fset (object_ptr_kinds h) in let arg_components = (\ptr \ (\ptr \ S\<^sub>a\<^sub>r\<^sub>g. set |h \ a_get_scdom_component ptr|\<^sub>r) \ fset (object_ptr_kinds h). set |h \ a_get_scdom_component ptr|\<^sub>r) in let arg_components' = (\ptr \ (\ptr \ S\<^sub>a\<^sub>r\<^sub>g. set |h \ a_get_scdom_component ptr|\<^sub>r) \ fset (object_ptr_kinds h'). set |h' \ a_get_scdom_component ptr|\<^sub>r) in removed_pointers \ arg_components \ S\<^sub>r\<^sub>e\<^sub>s\<^sub>u\<^sub>l\<^sub>t \ arg_components' \ added_pointers \ (\outside_ptr \ fset (object_ptr_kinds h) \ fset (object_ptr_kinds h') - (\ptr \ S\<^sub>a\<^sub>r\<^sub>g. set |h \ a_get_scdom_component ptr|\<^sub>r). preserved (get_M outside_ptr id) h h'))" end global_interpretation l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_owner_document get_disconnected_nodes get_disconnected_nodes_locs to_tree_order defines get_scdom_component = "l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_scdom_component get_owner_document get_disconnected_nodes to_tree_order" and is_strongly_scdom_component_safe = a_is_strongly_scdom_component_safe and is_weakly_scdom_component_safe = a_is_weakly_scdom_component_safe . locale l_get_scdom_component_defs = fixes get_scdom_component :: "(_) object_ptr \ (_, (_) object_ptr list) dom_prog" fixes is_strongly_scdom_component_safe :: "(_) object_ptr set \ (_) object_ptr set \ (_) heap \ (_) heap \ bool" fixes is_weakly_scdom_component_safe :: "(_) object_ptr set \ (_) object_ptr set \ (_) heap \ (_) heap \ bool" locale l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_scdom_component_defs + l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + assumes get_scdom_component_impl: "get_scdom_component = a_get_scdom_component" assumes is_strongly_scdom_component_safe_impl: "is_strongly_scdom_component_safe = a_is_strongly_scdom_component_safe" assumes is_weakly_scdom_component_safe_impl: "is_weakly_scdom_component_safe = a_is_weakly_scdom_component_safe" begin lemmas get_scdom_component_def = a_get_scdom_component_def[folded get_scdom_component_impl] lemmas is_strongly_scdom_component_safe_def = a_is_strongly_scdom_component_safe_def[folded is_strongly_scdom_component_safe_impl] lemmas is_weakly_scdom_component_safe_def = a_is_weakly_scdom_component_safe_def[folded is_weakly_scdom_component_safe_impl] end interpretation i_get_scdom_component?: l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_owner_document get_disconnected_nodes get_disconnected_nodes_locs to_tree_order by(auto simp add: l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def get_scdom_component_def is_strongly_scdom_component_safe_def is_weakly_scdom_component_safe_def instances) declare l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_get_dom_component_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_heap_is_wellformed + l_get_owner_document + l_get_owner_document_wf + l_get_disconnected_nodes + l_to_tree_order + l_known_ptr + l_known_ptrs + l_get_owner_document_wf_get_root_node_wf + assumes known_ptr_impl: "known_ptr = DocumentClass.known_ptr" begin lemma known_ptr_node_or_document: "known_ptr ptr \ is_node_ptr_kind ptr \ is_document_ptr_kind ptr" by(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) lemma get_scdom_component_ptr_in_heap2: assumes "h \ ok (get_scdom_component ptr)" shows "ptr |\| object_ptr_kinds h" using assms get_root_node_ptr_in_heap apply(auto simp add: get_scdom_component_def elim!: bind_is_OK_E3 intro!: map_M_pure_I)[1] by (simp add: is_OK_returns_result_I local.get_owner_document_ptr_in_heap) lemma get_scdom_component_subset_get_dom_component: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_scdom_component ptr \\<^sub>r sc" assumes "h \ get_dom_component ptr \\<^sub>r c" shows "set c \ set sc" proof - obtain document disc_nodes tree_order disconnected_tree_orders where document: "h \ get_owner_document ptr \\<^sub>r document" and disc_nodes: "h \ get_disconnected_nodes document \\<^sub>r disc_nodes" and tree_order: "h \ to_tree_order (cast document) \\<^sub>r tree_order" and disconnected_tree_orders: "h \ map_M (to_tree_order \ cast) disc_nodes \\<^sub>r disconnected_tree_orders" and sc: "sc = tree_order @ (concat disconnected_tree_orders)" using assms(4) by(auto simp add: get_scdom_component_def elim!: bind_returns_result_E elim!: bind_returns_result_E2[rotated, OF get_owner_document_pure, rotated] elim!: bind_returns_result_E2[rotated, OF get_disconnected_nodes_pure, rotated] elim!: bind_returns_result_E2[rotated, OF to_tree_order_pure, rotated] ) obtain root_ptr where root_ptr: "h \ get_root_node ptr \\<^sub>r root_ptr" and c: "h \ to_tree_order root_ptr \\<^sub>r c" using assms(5) by(auto simp add: get_dom_component_def elim!: bind_returns_result_E2[rotated, OF get_root_node_pure, rotated]) show ?thesis proof (cases "is_document_ptr_kind root_ptr") case True then have "cast document = root_ptr" using get_root_node_document assms(1) assms(2) assms(3) root_ptr document by (metis document_ptr_casts_commute3 returns_result_eq) then have "c = tree_order" using tree_order c by auto then show ?thesis by(simp add: sc) next case False moreover have "root_ptr |\| object_ptr_kinds h" using assms(1) assms(2) assms(3) local.get_root_node_root_in_heap root_ptr by blast ultimately have "is_node_ptr_kind root_ptr" using assms(3) known_ptrs_known_ptr known_ptr_node_or_document by auto then obtain root_node_ptr where root_node_ptr: "root_ptr = cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr" by (metis node_ptr_casts_commute3) then have "h \ get_owner_document root_ptr \\<^sub>r document" using get_root_node_same_owner_document using assms(1) assms(2) assms(3) document root_ptr by blast then have "root_node_ptr \ set disc_nodes" using assms(1) assms(2) assms(3) disc_nodes in_disconnected_nodes_no_parent root_node_ptr using local.get_root_node_same_no_parent root_ptr by blast then have "c \ set disconnected_tree_orders" using c root_node_ptr using map_M_pure_E[OF disconnected_tree_orders] by (metis (mono_tags, lifting) comp_apply local.to_tree_order_pure select_result_I2) then show ?thesis by(auto simp add: sc) qed qed lemma get_scdom_component_ptrs_same_owner_document: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_scdom_component ptr \\<^sub>r sc" assumes "ptr' \ set sc" assumes "h \ get_owner_document ptr \\<^sub>r owner_document" shows "h \ get_owner_document ptr' \\<^sub>r owner_document" proof - obtain document disc_nodes tree_order disconnected_tree_orders where document: "h \ get_owner_document ptr \\<^sub>r document" and disc_nodes: "h \ get_disconnected_nodes document \\<^sub>r disc_nodes" and tree_order: "h \ to_tree_order (cast document) \\<^sub>r tree_order" and disconnected_tree_orders: "h \ map_M (to_tree_order \ cast) disc_nodes \\<^sub>r disconnected_tree_orders" and sc: "sc = tree_order @ (concat disconnected_tree_orders)" using assms(4) by(auto simp add: get_scdom_component_def elim!: bind_returns_result_E elim!: bind_returns_result_E2[rotated, OF get_owner_document_pure, rotated] elim!: bind_returns_result_E2[rotated, OF get_disconnected_nodes_pure, rotated] elim!: bind_returns_result_E2[rotated, OF to_tree_order_pure, rotated] ) show ?thesis proof (cases "ptr' \ set tree_order") case True have "owner_document = document" using assms(6) document by fastforce then show ?thesis by (metis (no_types) True assms(1) assms(2) assms(3) cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_inject document document_ptr_casts_commute3 document_ptr_document_ptr_cast document_ptr_kinds_commutes local.get_owner_document_owner_document_in_heap local.get_root_node_document local.get_root_node_not_node_same local.to_tree_order_same_root node_ptr_no_document_ptr_cast tree_order) next case False then obtain disconnected_tree_order where disconnected_tree_order: "ptr' \ set disconnected_tree_order" and "disconnected_tree_order \ set disconnected_tree_orders" using sc \ptr' \ set sc\ by auto obtain root_ptr' where root_ptr': "root_ptr' \ set disc_nodes" and "h \ to_tree_order (cast root_ptr') \\<^sub>r disconnected_tree_order" using map_M_pure_E2[OF disconnected_tree_orders \disconnected_tree_order \ set disconnected_tree_orders\] by (metis comp_apply local.to_tree_order_pure) have "\(\parent \ fset (object_ptr_kinds h). root_ptr' \ set |h \ get_child_nodes parent|\<^sub>r)" using disc_nodes by (meson assms(1) assms(2) assms(3) disjoint_iff_not_equal local.get_child_nodes_ok local.heap_is_wellformed_children_disc_nodes_different local.known_ptrs_known_ptr notin_fset returns_result_select_result root_ptr') then have "h \ get_parent root_ptr' \\<^sub>r None" using disc_nodes by (metis (no_types, lifting) assms(1) assms(2) assms(3) fmember.rep_eq local.get_parent_child_dual local.get_parent_ok local.get_parent_parent_in_heap local.heap_is_wellformed_disc_nodes_in_heap returns_result_select_result root_ptr' select_result_I2 split_option_ex) then have "h \ get_root_node ptr' \\<^sub>r cast root_ptr'" using \h \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_ptr') \\<^sub>r disconnected_tree_order\ assms(1) assms(2) assms(3) disconnected_tree_order local.get_root_node_no_parent local.to_tree_order_get_root_node local.to_tree_order_ptr_in_result by blast then have "h \ get_owner_document (cast root_ptr') \\<^sub>r document" using assms(1) assms(2) assms(3) disc_nodes local.get_owner_document_disconnected_nodes root_ptr' by blast then have "h \ get_owner_document ptr' \\<^sub>r document" using \h \ get_root_node ptr' \\<^sub>r cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_ptr'\ assms(1) assms(2) assms(3) local.get_root_node_same_owner_document by blast then show ?thesis using assms(6) document returns_result_eq by force qed qed lemma get_scdom_component_ptrs_same_scope_component: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_scdom_component ptr \\<^sub>r sc" assumes "ptr' \ set sc" shows "h \ get_scdom_component ptr' \\<^sub>r sc" proof - obtain document disc_nodes tree_order disconnected_tree_orders where document: "h \ get_owner_document ptr \\<^sub>r document" and disc_nodes: "h \ get_disconnected_nodes document \\<^sub>r disc_nodes" and tree_order: "h \ to_tree_order (cast document) \\<^sub>r tree_order" and disconnected_tree_orders: "h \ map_M (to_tree_order \ cast) disc_nodes \\<^sub>r disconnected_tree_orders" and sc: "sc = tree_order @ (concat disconnected_tree_orders)" using assms(4) by(auto simp add: get_scdom_component_def elim!: bind_returns_result_E elim!: bind_returns_result_E2[rotated, OF get_owner_document_pure, rotated] elim!: bind_returns_result_E2[rotated, OF get_disconnected_nodes_pure, rotated] elim!: bind_returns_result_E2[rotated, OF to_tree_order_pure, rotated] ) show ?thesis proof (cases "ptr' \ set tree_order") case True then have "h \ get_owner_document ptr' \\<^sub>r document" by (metis assms(1) assms(2) assms(3) cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_inject document document_ptr_casts_commute3 document_ptr_kinds_commutes known_ptr_node_or_document local.get_owner_document_owner_document_in_heap local.get_root_node_document local.get_root_node_not_node_same local.known_ptrs_known_ptr local.to_tree_order_get_root_node local.to_tree_order_ptr_in_result node_ptr_no_document_ptr_cast tree_order) then show ?thesis using disc_nodes tree_order disconnected_tree_orders sc by(auto simp add: get_scdom_component_def intro!: bind_pure_returns_result_I map_M_pure_I) next case False then obtain disconnected_tree_order where disconnected_tree_order: "ptr' \ set disconnected_tree_order" and "disconnected_tree_order \ set disconnected_tree_orders" using sc \ptr' \ set sc\ by auto obtain root_ptr' where root_ptr': "root_ptr' \ set disc_nodes" and "h \ to_tree_order (cast root_ptr') \\<^sub>r disconnected_tree_order" using map_M_pure_E2[OF disconnected_tree_orders \disconnected_tree_order \ set disconnected_tree_orders\] by (metis comp_apply local.to_tree_order_pure) have "\(\parent \ fset (object_ptr_kinds h). root_ptr' \ set |h \ get_child_nodes parent|\<^sub>r)" using disc_nodes by (meson assms(1) assms(2) assms(3) disjoint_iff_not_equal local.get_child_nodes_ok local.heap_is_wellformed_children_disc_nodes_different local.known_ptrs_known_ptr notin_fset returns_result_select_result root_ptr') then have "h \ get_parent root_ptr' \\<^sub>r None" using disc_nodes by (metis (no_types, lifting) assms(1) assms(2) assms(3) fmember.rep_eq local.get_parent_child_dual local.get_parent_ok local.get_parent_parent_in_heap local.heap_is_wellformed_disc_nodes_in_heap returns_result_select_result root_ptr' select_result_I2 split_option_ex) then have "h \ get_root_node ptr' \\<^sub>r cast root_ptr'" using \h \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_ptr') \\<^sub>r disconnected_tree_order\ assms(1) assms(2) assms(3) disconnected_tree_order local.get_root_node_no_parent local.to_tree_order_get_root_node local.to_tree_order_ptr_in_result by blast then have "h \ get_owner_document (cast root_ptr') \\<^sub>r document" using assms(1) assms(2) assms(3) disc_nodes local.get_owner_document_disconnected_nodes root_ptr' by blast then have "h \ get_owner_document ptr' \\<^sub>r document" using \h \ get_root_node ptr' \\<^sub>r cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_ptr'\ assms(1) assms(2) assms(3) local.get_root_node_same_owner_document by blast then show ?thesis using disc_nodes tree_order disconnected_tree_orders sc by(auto simp add: get_scdom_component_def intro!: bind_pure_returns_result_I map_M_pure_I) qed qed lemma get_scdom_component_ok: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "ptr |\| object_ptr_kinds h" shows "h \ ok (get_scdom_component ptr)" using assms apply(auto simp add: get_scdom_component_def intro!: bind_is_OK_pure_I map_M_pure_I map_M_ok_I)[1] using get_owner_document_ok apply blast apply (simp add: local.get_disconnected_nodes_ok local.get_owner_document_owner_document_in_heap) apply (simp add: local.get_owner_document_owner_document_in_heap local.to_tree_order_ok) using local.heap_is_wellformed_disc_nodes_in_heap local.to_tree_order_ok node_ptr_kinds_commutes by blast lemma get_scdom_component_ptr_in_heap: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_scdom_component ptr \\<^sub>r sc" assumes "ptr' \ set sc" shows "ptr' |\| object_ptr_kinds h" apply(insert assms ) apply(auto simp add: get_scdom_component_def elim!: bind_returns_result_E2 intro!: map_M_pure_I)[1] using local.to_tree_order_ptrs_in_heap apply blast by (metis (no_types, lifting) assms(4) assms(5) bind_returns_result_E get_scdom_component_ptrs_same_scope_component is_OK_returns_result_I get_scdom_component_def local.get_owner_document_ptr_in_heap) lemma get_scdom_component_contains_get_dom_component: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_scdom_component ptr \\<^sub>r sc" assumes "ptr' \ set sc" obtains c where "h \ get_dom_component ptr' \\<^sub>r c" and "set c \ set sc" proof - have "h \ get_scdom_component ptr' \\<^sub>r sc" using assms(1) assms(2) assms(3) assms(4) assms(5) get_scdom_component_ptrs_same_scope_component by blast then show ?thesis by (meson assms(1) assms(2) assms(3) assms(5) get_scdom_component_ptr_in_heap get_scdom_component_subset_get_dom_component is_OK_returns_result_E local.get_dom_component_ok that) qed lemma get_scdom_component_owner_document_same: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_scdom_component ptr \\<^sub>r sc" assumes "ptr' \ set sc" obtains owner_document where "h \ get_owner_document ptr' \\<^sub>r owner_document" and "cast owner_document \ set sc" using assms apply(auto simp add: get_scdom_component_def elim!: bind_returns_result_E2 intro!: map_M_pure_I)[1] apply (metis (no_types, lifting) assms(4) assms(5) document_ptr_casts_commute3 document_ptr_document_ptr_cast get_scdom_component_contains_get_dom_component local.get_dom_component_ptr local.get_dom_component_root_node_same local.get_dom_component_to_tree_order local.get_root_node_document local.get_root_node_not_node_same local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap node_ptr_no_document_ptr_cast) apply(rule map_M_pure_E2) apply(simp) apply(simp) apply(simp) by (smt assms(4) assms(5) comp_apply get_scdom_component_ptr_in_heap is_OK_returns_result_E local.get_owner_document_disconnected_nodes local.get_root_node_ok local.get_root_node_same_owner_document local.to_tree_order_get_root_node local.to_tree_order_ptr_in_result) lemma get_scdom_component_different_owner_documents: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_owner_document ptr \\<^sub>r owner_document" assumes "h \ get_owner_document ptr' \\<^sub>r owner_document'" assumes "owner_document \ owner_document'" shows "set |h \ get_scdom_component ptr|\<^sub>r \ set |h \ get_scdom_component ptr'|\<^sub>r = {}" using assms get_scdom_component_ptrs_same_owner_document by (smt disjoint_iff_not_equal get_scdom_component_ok is_OK_returns_result_I local.get_owner_document_ptr_in_heap returns_result_eq returns_result_select_result) lemma get_scdom_component_ptr: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_scdom_component ptr \\<^sub>r c" shows "ptr \ set c" using assms by (meson get_scdom_component_ptr_in_heap2 get_scdom_component_subset_get_dom_component is_OK_returns_result_E is_OK_returns_result_I local.get_dom_component_ok local.get_dom_component_ptr subsetD) end locale l_get_dom_component_get_scdom_component = l_get_owner_document_defs + l_heap_is_wellformed_defs + l_type_wf + l_known_ptrs + l_get_scdom_component_defs + l_get_dom_component_defs + assumes get_scdom_component_subset_get_dom_component: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_scdom_component ptr \\<^sub>r sc \ h \ get_dom_component ptr \\<^sub>r c \ set c \ set sc" assumes get_scdom_component_ptrs_same_scope_component: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_scdom_component ptr \\<^sub>r sc \ ptr' \ set sc \ h \ get_scdom_component ptr' \\<^sub>r sc" assumes get_scdom_component_ptrs_same_owner_document: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_scdom_component ptr \\<^sub>r sc \ ptr' \ set sc \ h \ get_owner_document ptr \\<^sub>r owner_document \ h \ get_owner_document ptr' \\<^sub>r owner_document" assumes get_scdom_component_ok: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ ptr |\| object_ptr_kinds h \ h \ ok (get_scdom_component ptr)" assumes get_scdom_component_ptr_in_heap: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_scdom_component ptr \\<^sub>r sc \ ptr' \ set sc \ ptr' |\| object_ptr_kinds h" assumes get_scdom_component_contains_get_dom_component: "(\c. h \ get_dom_component ptr' \\<^sub>r c \ set c \ set sc \ thesis) \ heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_scdom_component ptr \\<^sub>r sc \ ptr' \ set sc \ thesis" assumes get_scdom_component_owner_document_same: "(\owner_document. h \ get_owner_document ptr' \\<^sub>r owner_document \ cast owner_document \ set sc \ thesis) \ heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_scdom_component ptr \\<^sub>r sc \ ptr' \ set sc \ thesis" assumes get_scdom_component_different_owner_documents: "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ get_owner_document ptr \\<^sub>r owner_document \ h \ get_owner_document ptr' \\<^sub>r owner_document' \ owner_document \ owner_document' \ set |h \ get_scdom_component ptr|\<^sub>r \ set |h \ get_scdom_component ptr'|\<^sub>r = {}" interpretation i_get_dom_component_get_scdom_component?: l_get_dom_component_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_owner_document get_disconnected_nodes get_disconnected_nodes_locs to_tree_order heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name by(auto simp add: l_get_dom_component_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_dom_component_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_get_dom_component_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_dom_component_get_scdom_component_is_l_get_dom_component_get_scdom_component [instances]: "l_get_dom_component_get_scdom_component get_owner_document heap_is_wellformed type_wf known_ptr known_ptrs get_scdom_component get_dom_component" apply(auto simp add: l_get_dom_component_get_scdom_component_def l_get_dom_component_get_scdom_component_axioms_def instances)[1] using get_scdom_component_subset_get_dom_component apply fast using get_scdom_component_ptrs_same_scope_component apply fast using get_scdom_component_ptrs_same_owner_document apply fast using get_scdom_component_ok apply fast using get_scdom_component_ptr_in_heap apply fast using get_scdom_component_contains_get_dom_component apply blast using get_scdom_component_owner_document_same apply blast using get_scdom_component_different_owner_documents apply fast done subsubsection \get\_child\_nodes\ locale l_get_scdom_component_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_dom_component_get_scdom_component + l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_dom_component_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_child_nodes_is_strongly_scdom_component_safe_step: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_scdom_component ptr \\<^sub>r sc" assumes "h \ get_child_nodes ptr' \\<^sub>r children" assumes "child \ set children" shows "cast child \ set sc \ ptr' \ set sc" apply(auto)[1] apply (meson assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) contra_subsetD get_scdom_component_ptrs_same_scope_component get_scdom_component_subset_get_dom_component is_OK_returns_result_E local.get_child_nodes_is_strongly_dom_component_safe local.get_dom_component_ok local.get_dom_component_ptr local.heap_is_wellformed_children_in_heap node_ptr_kinds_commutes) by (meson assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) get_scdom_component_contains_get_dom_component is_OK_returns_result_E is_OK_returns_result_I get_child_nodes_is_strongly_dom_component_safe local.get_child_nodes_ptr_in_heap local.get_dom_component_ok local.get_dom_component_ptr set_rev_mp) lemma get_child_nodes_is_strongly_scdom_component_safe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_child_nodes ptr \\<^sub>r children" assumes "h \ get_child_nodes ptr \\<^sub>h h'" shows "is_strongly_scdom_component_safe {ptr} (cast ` set children) h h'" proof - have "h = h'" using assms(5) by (meson local.get_child_nodes_pure pure_returns_heap_eq) then show ?thesis using assms apply(auto simp add: is_strongly_scdom_component_safe_def Let_def preserved_def)[1] by (smt Int_absorb2 finite_set_in get_child_nodes_is_strongly_scdom_component_safe_step in_mono is_OK_returns_result_I local.get_child_nodes_ptr_in_heap local.get_dom_component_ok local.get_dom_component_ptr local.get_scdom_component_impl local.get_scdom_component_ok local.get_scdom_component_ptr_in_heap local.get_scdom_component_subset_get_dom_component returns_result_select_result subsetI) qed end interpretation i_get_scdom_component_get_child_nodes?: l_get_scdom_component_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_disconnected_nodes get_disconnected_nodes_locs to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs by(auto simp add: l_get_scdom_component_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_scdom_component_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_parent\ locale l_get_scdom_component_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_dom_component_get_scdom_component + l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_dom_component_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_parent_is_strongly_scdom_component_safe_step: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_scdom_component ptr \\<^sub>r sc" assumes "h \ get_parent ptr' \\<^sub>r Some parent" shows "parent \ set sc \ cast ptr' \ set sc" by (meson assms(1) assms(2) assms(3) assms(4) assms(5) contra_subsetD get_scdom_component_contains_get_dom_component local.get_dom_component_ptr local.get_parent_is_strongly_dom_component_safe_step) lemma get_parent_is_strongly_scdom_component_safe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_parent node_ptr \\<^sub>r Some parent" assumes "h \ get_parent node_ptr \\<^sub>h h'" shows "is_strongly_scdom_component_safe {cast node_ptr} {parent} h h'" proof - have "h = h'" using assms(5) by (meson local.get_parent_pure pure_returns_heap_eq) then show ?thesis using assms apply(auto simp add: is_strongly_scdom_component_safe_def Let_def preserved_def)[1] by (smt IntI finite_set_in in_mono local.get_dom_component_ok local.get_dom_component_ptr local.get_parent_is_strongly_dom_component_safe_step local.get_parent_parent_in_heap local.get_scdom_component_impl local.get_scdom_component_ok local.get_scdom_component_subset_get_dom_component local.to_tree_order_ok local.to_tree_order_parent local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap returns_result_select_result) qed end interpretation i_get_scdom_component_get_parent?: l_get_scdom_component_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_disconnected_nodes get_disconnected_nodes_locs to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs by(auto simp add: l_get_scdom_component_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_scdom_component_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_root\_node\ locale l_get_scdom_component_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_dom_component_get_scdom_component + l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_dom_component_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_root_node_is_strongly_scdom_component_safe_step: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_scdom_component ptr \\<^sub>r sc" assumes "h \ get_root_node ptr' \\<^sub>r root" shows "root \ set sc \ ptr' \ set sc" by (meson assms(1) assms(2) assms(3) assms(4) assms(5) contra_subsetD get_scdom_component_contains_get_dom_component local.get_dom_component_ptr local.get_root_node_is_strongly_dom_component_safe_step) lemma get_root_node_is_strongly_scdom_component_safe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_root_node ptr \\<^sub>r root" assumes "h \ get_root_node ptr \\<^sub>h h'" shows "is_strongly_scdom_component_safe {ptr} {root} h h'" proof - have "h = h'" using assms(5) by (meson local.get_root_node_pure pure_returns_heap_eq) then show ?thesis using assms apply(auto simp add: is_strongly_scdom_component_safe_def Let_def preserved_def)[1] by (smt Int_iff finite_set_in is_OK_returns_result_I local.get_dom_component_ok local.get_dom_component_ptr local.get_root_node_is_strongly_dom_component_safe_step local.get_root_node_ptr_in_heap local.get_scdom_component_impl local.get_scdom_component_ok local.get_scdom_component_subset_get_dom_component returns_result_select_result subset_eq) qed end interpretation i_get_scdom_component_get_root_node?: l_get_scdom_component_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_disconnected_nodes get_disconnected_nodes_locs to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs by(auto simp add: l_get_scdom_component_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_scdom_component_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_element\_by\_id\ locale l_get_scdom_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_dom_component_get_scdom_component + l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_dom_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_element_by_id_is_strongly_scdom_component_safe_step: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_scdom_component ptr \\<^sub>r sc" assumes "h \ get_element_by_id ptr' idd \\<^sub>r Some result" shows "cast result \ set sc \ ptr' \ set sc" by (meson assms(1) assms(2) assms(3) assms(4) assms(5) contra_subsetD get_element_by_id_is_strongly_dom_component_safe_step get_scdom_component_contains_get_dom_component local.get_dom_component_ptr) lemma get_element_by_id_is_strongly_scdom_component_safe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_element_by_id ptr idd \\<^sub>r Some result" assumes "h \ get_element_by_id ptr idd \\<^sub>h h'" shows "is_strongly_scdom_component_safe {ptr} {cast result} h h'" proof - have "h = h'" using assms(5) by(auto simp add: preserved_def get_element_by_id_def first_in_tree_order_def elim!: bind_returns_heap_E2 intro!: map_filter_M_pure bind_pure_I split: option.splits list.splits) have "ptr |\| object_ptr_kinds h" using assms(4) apply(auto simp add: get_element_by_id_def)[1] by (metis (no_types, lifting) assms(1) assms(2) assms(3) bind_is_OK_E is_OK_returns_result_I local.first_in_tree_order_def local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap) obtain to where to: "h \ to_tree_order ptr \\<^sub>r to" by (meson \ptr |\| object_ptr_kinds h\ assms(1) assms(2) assms(3) is_OK_returns_result_E local.to_tree_order_ok) then have "cast result \ set to" using assms(4) local.get_element_by_id_result_in_tree_order by auto obtain c where c: "h \ a_get_scdom_component ptr \\<^sub>r c" using \ptr |\| object_ptr_kinds h\ assms(1) assms(2) assms(3) local.get_scdom_component_impl local.get_scdom_component_ok by blast then show ?thesis using assms \h = h'\ apply(auto simp add: is_strongly_scdom_component_safe_def Let_def preserved_def get_element_by_id_def first_in_tree_order_def elim!: bind_returns_result_E2 intro!: map_filter_M_pure bind_pure_I split: option.splits list.splits)[1] by (smt IntI \ptr |\| object_ptr_kinds h\ assms(4) finite_set_in get_element_by_id_is_strongly_scdom_component_safe_step local.get_dom_component_ok local.get_dom_component_ptr local.get_scdom_component_impl local.get_scdom_component_subset_get_dom_component returns_result_select_result select_result_I2 subsetD) qed end interpretation i_get_scdom_component_get_element_by_id?: l_get_scdom_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_disconnected_nodes get_disconnected_nodes_locs to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs by(auto simp add: l_get_scdom_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_scdom_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_elements\_by\_class\_name\ locale l_get_scdom_component_get_elements_by_class_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_dom_component_get_scdom_component + l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_dom_component_get_elements_by_class_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_elements_by_class_name_is_strongly_scdom_component_safe_step: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_scdom_component ptr \\<^sub>r sc" assumes "h \ get_elements_by_class_name ptr' idd \\<^sub>r results" assumes "result \ set results" shows "cast result \ set sc \ ptr' \ set sc" by (meson assms local.get_dom_component_ptr local.get_elements_by_class_name_is_strongly_dom_component_safe_step local.get_scdom_component_contains_get_dom_component subsetD) lemma get_elements_by_class_name_is_strongly_scdom_component_safe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_elements_by_class_name ptr idd \\<^sub>r results" assumes "h \ get_elements_by_class_name ptr idd \\<^sub>h h'" shows "is_strongly_scdom_component_safe {ptr} (cast ` set results) h h'" proof - have "h = h'" using assms(5) by (meson local.get_elements_by_class_name_pure pure_returns_heap_eq) have "ptr |\| object_ptr_kinds h" using assms(4) apply(auto simp add: get_elements_by_class_name_def)[1] by (metis (no_types, lifting) assms(1) assms(2) assms(3) bind_is_OK_E is_OK_returns_result_I local.first_in_tree_order_def local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap) obtain to where to: "h \ to_tree_order ptr \\<^sub>r to" by (meson \ptr |\| object_ptr_kinds h\ assms(1) assms(2) assms(3) is_OK_returns_result_E local.to_tree_order_ok) then have "cast ` set results \ set to" using assms(4) local.get_elements_by_class_name_result_in_tree_order by auto obtain c where c: "h \ a_get_scdom_component ptr \\<^sub>r c" using \ptr |\| object_ptr_kinds h\ assms(1) assms(2) assms(3) local.get_scdom_component_impl local.get_scdom_component_ok by blast then show ?thesis using assms \h = h'\ apply(auto simp add: is_strongly_scdom_component_safe_def Let_def preserved_def get_element_by_id_def first_in_tree_order_def elim!: bind_returns_result_E2 intro!: map_filter_M_pure bind_pure_I split: option.splits list.splits)[1] by (smt IntI \ptr |\| object_ptr_kinds h\ finite_set_in get_elements_by_class_name_is_strongly_scdom_component_safe_step local.get_dom_component_ok local.get_dom_component_ptr local.get_scdom_component_impl local.get_scdom_component_subset_get_dom_component returns_result_select_result select_result_I2 subsetD) qed end interpretation i_get_scdom_component_get_elements_by_class_name?: l_get_scdom_component_get_elements_by_class_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_disconnected_nodes get_disconnected_nodes_locs to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs by(auto simp add: l_get_scdom_component_get_elements_by_class_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_scdom_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_elements\_by\_tag\_name\ locale l_get_scdom_component_get_elements_by_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_dom_component_get_scdom_component + l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_dom_component_get_elements_by_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_elements_by_tag_name_is_strongly_scdom_component_safe_step: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_scdom_component ptr \\<^sub>r sc" assumes "h \ get_elements_by_tag_name ptr' idd \\<^sub>r results" assumes "result \ set results" shows "cast result \ set sc \ ptr' \ set sc" by (meson assms local.get_dom_component_ptr local.get_elements_by_tag_name_is_strongly_dom_component_safe_step local.get_scdom_component_contains_get_dom_component subsetD) lemma get_elements_by_tag_name_is_strongly_scdom_component_safe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_elements_by_tag_name ptr idd \\<^sub>r results" assumes "h \ get_elements_by_tag_name ptr idd \\<^sub>h h'" shows "is_strongly_scdom_component_safe {ptr} (cast ` set results) h h'" proof - have "h = h'" using assms(5) by (meson local.get_elements_by_tag_name_pure pure_returns_heap_eq) have "ptr |\| object_ptr_kinds h" using assms(4) apply(auto simp add: get_elements_by_tag_name_def)[1] by (metis (no_types, lifting) assms(1) assms(2) assms(3) bind_is_OK_E is_OK_returns_result_I local.first_in_tree_order_def local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap) obtain to where to: "h \ to_tree_order ptr \\<^sub>r to" by (meson \ptr |\| object_ptr_kinds h\ assms(1) assms(2) assms(3) is_OK_returns_result_E local.to_tree_order_ok) then have "cast ` set results \ set to" using assms(4) local.get_elements_by_tag_name_result_in_tree_order by auto obtain c where c: "h \ a_get_scdom_component ptr \\<^sub>r c" using \ptr |\| object_ptr_kinds h\ assms(1) assms(2) assms(3) local.get_scdom_component_impl local.get_scdom_component_ok by blast then show ?thesis using assms \h = h'\ apply(auto simp add: is_strongly_scdom_component_safe_def Let_def preserved_def get_element_by_id_def first_in_tree_order_def elim!: bind_returns_result_E2 intro!: map_filter_M_pure bind_pure_I split: option.splits list.splits)[1] by (smt IntI \ptr |\| object_ptr_kinds h\ finite_set_in get_elements_by_tag_name_is_strongly_scdom_component_safe_step local.get_dom_component_ok local.get_dom_component_ptr local.get_scdom_component_impl local.get_scdom_component_subset_get_dom_component returns_result_select_result select_result_I2 subsetD) qed end interpretation i_get_scdom_component_get_elements_by_tag_name?: l_get_scdom_component_get_elements_by_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_disconnected_nodes get_disconnected_nodes_locs to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name first_in_tree_order get_attribute get_attribute_locs by(auto simp add: l_get_scdom_component_get_elements_by_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_scdom_component_get_element_by_id\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \remove\_child\ locale l_get_scdom_component_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_dom_component_get_scdom_component + l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_owner_document_wf + l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs get_parent get_parent_locs get_owner_document get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs remove_child remove_child_locs remove type_wf known_ptr known_ptrs heap_is_wellformed parent_child_rel begin lemma remove_child_is_component_unsafe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove_child ptr' child \\<^sub>h h'" assumes "ptr \ set |h \ get_dom_component ptr'|\<^sub>r" assumes "ptr \ set |h \ get_dom_component (cast |h \ get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child)|\<^sub>r)|\<^sub>r" (* assumes "ptr \ set |h \ get_dom_component (cast child)|\<^sub>r" *) shows "preserved (get_M ptr getter) h h'" proof - have "ptr \ ptr'" using assms(5) by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) is_OK_returns_heap_I is_OK_returns_result_E local.get_dom_component_ok local.get_dom_component_ptr local.remove_child_ptr_in_heap select_result_I2) obtain owner_document where owner_document: "h \ get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r owner_document" by (meson assms(1) assms(2) assms(3) assms(4) is_OK_returns_result_E local.get_owner_document_ok local.remove_child_child_in_heap node_ptr_kinds_commutes) then obtain c where "h \ get_dom_component (cast owner_document) \\<^sub>r c" using get_dom_component_ok owner_document assms(1) assms(2) assms(3) by (meson document_ptr_kinds_commutes get_owner_document_owner_document_in_heap select_result_I) then have "ptr \ cast owner_document" using assms(6) assms(1) assms(2) assms(3) local.get_dom_component_ptr owner_document by auto show ?thesis using remove_child_writes assms(4) apply(rule reads_writes_preserved2) apply(auto simp add: remove_child_locs_def set_child_nodes_locs_def set_disconnected_nodes_locs_def all_args_def split: option.splits)[1] apply (metis \ptr \ ptr'\ document_ptr_casts_commute3 get_M_Mdocument_preserved3) apply (metis (no_types, lifting) \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\ get_M_Mdocument_preserved3 owner_document select_result_I2) apply (metis (no_types, lifting) \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\ get_M_Mdocument_preserved3 owner_document select_result_I2) apply (metis \ptr \ ptr'\ element_ptr_casts_commute3 get_M_Element_preserved8) apply (metis (no_types, lifting) \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\ get_M_Mdocument_preserved3 owner_document select_result_I2) apply (metis (no_types, lifting) \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\ get_M_Mdocument_preserved3 owner_document select_result_I2) done qed lemma remove_child_is_strongly_dom_component_safe_step: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove_child ptr' child \\<^sub>h h'" assumes "ptr \ set |h \ get_scdom_component ptr'|\<^sub>r" (* assumes "ptr \ set |h \ get_dom_component (cast |h \ get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child)|\<^sub>r)|\<^sub>r" *) assumes "ptr \ set |h \ get_scdom_component (cast child)|\<^sub>r" shows "preserved (get_M ptr getter) h h'" proof - obtain sc where sc: "h \ get_scdom_component ptr' \\<^sub>r sc" using get_scdom_component_ok by (meson assms(1) assms(2) assms(3) assms(4) is_OK_returns_heap_I local.remove_child_ptr_in_heap returns_result_select_result) have "child |\| node_ptr_kinds h" using assms(4) remove_child_child_in_heap by blast then obtain child_sc where child_sc: "h \ get_scdom_component (cast child) \\<^sub>r child_sc" using get_scdom_component_ok by (meson assms(1) assms(2) assms(3) node_ptr_kinds_commutes select_result_I) then obtain owner_document where owner_document: "h \ get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r owner_document" by (meson \child |\| node_ptr_kinds h\ assms(1) assms(2) assms(3) contra_subsetD get_scdom_component_owner_document_same is_OK_returns_result_E get_scdom_component_subset_get_dom_component local.get_dom_component_ok local.get_dom_component_ptr node_ptr_kinds_commutes) then have "h \ get_scdom_component (cast owner_document) \\<^sub>r child_sc" using child_sc by (smt \child |\| node_ptr_kinds h\ assms(1) assms(2) assms(3) contra_subsetD get_scdom_component_subset_get_dom_component get_scdom_component_owner_document_same get_scdom_component_ptrs_same_scope_component local.get_dom_component_ok local.get_dom_component_ptr node_ptr_kinds_commutes returns_result_select_result select_result_I2) have "ptr \ set |h \ get_dom_component ptr'|\<^sub>r" by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) assms(5) contra_subsetD get_scdom_component_subset_get_dom_component is_OK_returns_heap_I local.get_dom_component_ok local.remove_child_ptr_in_heap returns_result_select_result sc select_result_I2) moreover have "ptr \ set |h \ get_scdom_component (cast |h \ get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child)|\<^sub>r)|\<^sub>r" using get_scdom_component_owner_document_same get_scdom_component_ptrs_same_scope_component by (metis (no_types, lifting) \h \ get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document) \\<^sub>r child_sc\ assms(6) child_sc owner_document select_result_I2) have "ptr \ set |h \ get_dom_component (cast |h \ get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child)|\<^sub>r)|\<^sub>r" using get_scdom_component_owner_document_same by (metis (no_types, lifting) \h \ get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document) \\<^sub>r child_sc\ \ptr \ set |h \ get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r |h \ get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child)|\<^sub>r)|\<^sub>r\ assms(1) assms(2) assms(3) contra_subsetD document_ptr_kinds_commutes get_scdom_component_subset_get_dom_component is_OK_returns_result_E local.get_dom_component_ok local.get_owner_document_owner_document_in_heap owner_document select_result_I2) ultimately show ?thesis using assms(1) assms(2) assms(3) assms(4) remove_child_is_component_unsafe by blast qed lemma remove_child_is_strongly_dom_component_safe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove_child ptr child \\<^sub>h h'" shows "is_strongly_scdom_component_safe {ptr, cast child} {} h h'" proof - obtain owner_document children_h h2 disconnected_nodes_h where owner_document: "h \ get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r owner_document" and children_h: "h \ get_child_nodes ptr \\<^sub>r children_h" and child_in_children_h: "child \ set children_h" and disconnected_nodes_h: "h \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h" and h2: "h \ set_disconnected_nodes owner_document (child # disconnected_nodes_h) \\<^sub>h h2" and h': "h2 \ set_child_nodes ptr (remove1 child children_h) \\<^sub>h h'" using assms(4) apply(auto simp add: remove_child_def elim!: bind_returns_heap_E dest!: pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_child_nodes_pure] split: if_splits)[1] using pure_returns_heap_eq by fastforce have object_ptr_kinds_eq3: "object_ptr_kinds h = object_ptr_kinds h'" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h = object_ptr_kinds h'", OF remove_child_writes assms(4)]) unfolding remove_child_locs_def using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have object_ptr_kinds_eq: "\ptrs. h \ object_ptr_kinds_M \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" unfolding object_ptr_kinds_M_defs by simp then have object_ptr_kinds_eq2: "|h \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" using select_result_eq by force then have node_ptr_kinds_eq2: "|h \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>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|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>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' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" apply(rule reads_writes_preserved[OF get_child_nodes_reads remove_child_writes assms(4)]) unfolding remove_child_locs_def using set_disconnected_nodes_get_child_nodes set_child_nodes_get_child_nodes_different_pointers by fast then have children_eq2: "\ptr' children. ptr \ ptr' \ |h \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>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 \\<^sub>r disconnected_nodes = h' \ get_disconnected_nodes document_ptr \\<^sub>r disconnected_nodes" apply(rule reads_writes_preserved[OF get_disconnected_nodes_reads remove_child_writes assms(4)]) unfolding remove_child_locs_def using set_child_nodes_get_disconnected_nodes set_disconnected_nodes_get_disconnected_nodes_different_pointers by (metis (no_types, lifting) Un_iff owner_document select_result_I2) then have disconnected_nodes_eq2: "\document_ptr. document_ptr \ owner_document \ |h \ get_disconnected_nodes document_ptr|\<^sub>r = |h' \ get_disconnected_nodes document_ptr|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes ptr \\<^sub>r children_h" apply(rule reads_writes_separate_forwards[OF get_child_nodes_reads set_disconnected_nodes_writes h2 children_h] ) by (simp add: set_disconnected_nodes_get_child_nodes) have "known_ptrs h'" using object_ptr_kinds_eq3 known_ptrs_preserved \known_ptrs h\ by blast have "known_ptr ptr" using assms(3) using children_h get_child_nodes_ptr_in_heap local.known_ptrs_known_ptr by blast have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h2] using set_disconnected_nodes_types_preserved assms(2) by(auto simp add: reflp_def transp_def) then have "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_child_nodes_writes h'] using set_child_nodes_types_preserved by(auto simp add: reflp_def transp_def) have children_h': "h' \ get_child_nodes ptr \\<^sub>r remove1 child children_h" using assms(4) owner_document h2 disconnected_nodes_h children_h apply(auto simp add: remove_child_def split: if_splits)[1] apply(drule bind_returns_heap_E3) apply(auto split: if_splits)[1] apply(simp) apply(auto split: if_splits)[1] apply(drule bind_returns_heap_E3) apply(auto)[1] apply(simp) apply(drule bind_returns_heap_E3) apply(auto)[1] apply(simp) apply(drule bind_returns_heap_E4) apply(auto)[1] apply simp using \type_wf h2\ set_child_nodes_get_child_nodes \known_ptr ptr\ h' by blast have disconnected_nodes_h2: "h2 \ get_disconnected_nodes owner_document \\<^sub>r child # disconnected_nodes_h" using owner_document assms(4) h2 disconnected_nodes_h apply (auto simp add: remove_child_def split: if_splits)[1] apply(drule bind_returns_heap_E2) apply(auto split: if_splits)[1] apply(simp) by(auto simp add: local.set_disconnected_nodes_get_disconnected_nodes split: if_splits) then have disconnected_nodes_h': "h' \ get_disconnected_nodes owner_document \\<^sub>r child # disconnected_nodes_h" apply(rule reads_writes_separate_forwards[OF get_disconnected_nodes_reads set_child_nodes_writes h']) by (simp add: set_child_nodes_get_disconnected_nodes) moreover have "a_acyclic_heap h" using assms(1) by (simp add: heap_is_wellformed_def) have "parent_child_rel h' \ parent_child_rel h" proof (standard, safe) fix parent child assume a1: "(parent, child) \ parent_child_rel h'" then show "(parent, child) \ parent_child_rel h" proof (cases "parent = ptr") case True then show ?thesis using a1 remove_child_removes_parent[OF assms(1) assms(4)] children_h children_h' get_child_nodes_ptr_in_heap apply(auto simp add: parent_child_rel_def object_ptr_kinds_eq )[1] by (metis imageI notin_set_remove1) next case False then show ?thesis using a1 by(auto simp add: parent_child_rel_def object_ptr_kinds_eq3 children_eq2) qed qed then have "a_acyclic_heap h'" using \a_acyclic_heap h\ acyclic_heap_def acyclic_subset by blast moreover have "a_all_ptrs_in_heap h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq3 disconnected_nodes_eq)[1] apply (metis (no_types, lifting) \type_wf h'\ assms local.get_child_nodes_ok local.known_ptrs_known_ptr local.remove_child_children_subset notin_fset object_ptr_kinds_eq3 returns_result_select_result subset_code(1)) apply (metis (no_types, lifting) assms(4) disconnected_nodes_eq2 disconnected_nodes_h disconnected_nodes_h' document_ptr_kinds_eq3 finite_set_in local.remove_child_child_in_heap node_ptr_kinds_eq3 select_result_I2 set_ConsD subset_code(1)) done moreover have "a_owner_document_valid h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_owner_document_valid h'" apply(auto simp add: a_owner_document_valid_def object_ptr_kinds_eq3 document_ptr_kinds_eq3 node_ptr_kinds_eq3)[1] proof - fix node_ptr assume 0: "\node_ptr\fset (node_ptr_kinds h'). (\document_ptr. document_ptr |\| document_ptr_kinds h' \ node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r) \ (\parent_ptr. parent_ptr |\| object_ptr_kinds h' \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>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|\<^sub>r" then show "\document_ptr. document_ptr |\| document_ptr_kinds h' \ node_ptr \ set |h' \ get_disconnected_nodes document_ptr|\<^sub>r" proof (cases "node_ptr = child") case True show ?thesis apply(rule exI[where x=owner_document]) using children_eq2 disconnected_nodes_eq2 children_h children_h' disconnected_nodes_h' True by (metis (no_types, lifting) get_disconnected_nodes_ptr_in_heap is_OK_returns_result_I list.set_intros(1) select_result_I2) next case False then show ?thesis using 0 1 2 children_eq2 children_h children_h' disconnected_nodes_eq2 disconnected_nodes_h disconnected_nodes_h' apply(auto simp add: children_eq2 disconnected_nodes_eq2 dest!: select_result_I2)[1] by (metis children_eq2 disconnected_nodes_eq2 finite_set_in in_set_remove1 list.set_intros(2)) qed qed moreover { have h0: "a_distinct_lists h" using assms(1) by (simp add: heap_is_wellformed_def) moreover have ha1: "(\x\set |h \ object_ptr_kinds_M|\<^sub>r. set |h \ get_child_nodes x|\<^sub>r) \ (\x\set |h \ document_ptr_kinds_M|\<^sub>r. set |h \ get_disconnected_nodes x|\<^sub>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|\<^sub>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|\<^sub>r" using ha1 ha2 ha3 apply(simp) using IntI by fastforce moreover have "distinct |h \ object_ptr_kinds_M|\<^sub>r" apply(rule select_result_I) by(auto simp add: object_ptr_kinds_M_defs) moreover have "distinct |h \ document_ptr_kinds_M|\<^sub>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|\<^sub>r" assume 1: "a_distinct_lists h" and 3: "distinct |h \ object_ptr_kinds_M|\<^sub>r" have 4: "distinct (concat ((map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) |h \ object_ptr_kinds_M|\<^sub>r)))" using 1 by(auto simp add: a_distinct_lists_def) show "distinct (concat (map (\ptr. |h' \ get_child_nodes ptr|\<^sub>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|\<^sub>r" using 4 distinct_concat_map_E object_ptr_kinds_eq2 by fastforce obtain children where children: "h \ get_child_nodes x \\<^sub>r children" and distinct_children: "distinct children" by (metis "5" "6" assms get_child_nodes_ok local.known_ptrs_known_ptr object_ptr_kinds_eq3 select_result_I) obtain children' where children': "h' \ get_child_nodes x \\<^sub>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|\<^sub>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 \\<^sub>r children_x" by (metis "5" assms get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr object_ptr_kinds_eq3) obtain children_y where children_y: "h \ get_child_nodes y \\<^sub>r children_y" by (metis "6" assms get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr object_ptr_kinds_eq3) obtain children_x' where children_x': "h' \ get_child_nodes x \\<^sub>r children_x'" using children_eq children_h' children_x by fastforce obtain children_y' where children_y': "h' \ get_child_nodes y \\<^sub>r children_y'" using children_eq children_h' children_y by fastforce have "distinct (concat (map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) |h \ object_ptr_kinds_M|\<^sub>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|\<^sub>r \ set |h' \ get_child_nodes y|\<^sub>r = {}" using children_x' children_y' by (metis (no_types, lifting) select_result_I2) qed next assume 2: "distinct |h \ document_ptr_kinds_M|\<^sub>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|\<^sub>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|\<^sub>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|\<^sub>r" using distinct_lists_disconnected_nodes[OF h0] 4 get_disconnected_nodes_ok by (simp add: assms document_ptr_kinds_eq3 select_result_I) show "distinct |h' \ get_disconnected_nodes x|\<^sub>r" proof (cases "x = owner_document") case True have "child \ set |h \ get_disconnected_nodes x|\<^sub>r" using child_not_in document_ptr_kinds_eq2 "4" by fastforce moreover have "|h' \ get_disconnected_nodes x|\<^sub>r = child # |h \ get_disconnected_nodes x|\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>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|\<^sub>r) |h \ document_ptr_kinds_M|\<^sub>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|\<^sub>r \ set |h' \ get_disconnected_nodes y|\<^sub>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|\<^sub>r" and 3: "xb \ fset (document_ptr_kinds h')" and 4: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" obtain disc_nodes where disc_nodes: "h \ get_disconnected_nodes xb \\<^sub>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 \\<^sub>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 \\<^sub>r children" by (metis "1" assms finite_set_in get_child_nodes_ok is_OK_returns_result_E local.known_ptrs_known_ptr object_ptr_kinds_eq3) obtain children' where children': "h' \ get_child_nodes xa \\<^sub>r children'" using children children_eq children_h' by fastforce have "\x. x \ set |h \ get_child_nodes xa|\<^sub>r \ x \ set |h \ get_disconnected_nodes xb|\<^sub>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|\<^sub>r = children'" using children' by simp have 7: "|h' \ get_disconnected_nodes xb|\<^sub>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|\<^sub>r = remove1 child children_h" using children_h' - bysimp + 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|\<^sub>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) next case False then show ?thesis using "2" "4" "5" "6" "7" \children' = children\ disc_nodes disc_nodes' disconnected_nodes_eq returns_result_eq by metis qed qed then show "x \ {}" by simp qed } ultimately have "heap_is_wellformed h'" using heap_is_wellformed_def by blast show ?thesis apply(auto simp add: is_strongly_scdom_component_safe_def Let_def object_ptr_kinds_eq3)[1] using assms(1) assms(2) assms(3) assms(4) local.get_scdom_component_impl remove_child_is_strongly_dom_component_safe_step by blast qed end interpretation i_get_scdom_component_remove_child?: l_get_scdom_component_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name set_child_nodes set_child_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs remove_child remove_child_locs remove by(auto simp add: l_get_scdom_component_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_scdom_component_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \adopt\_node\ locale l_get_scdom_component_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_dom_component_get_scdom_component + l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_adopt_node_wf + l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_owner_document_wf + l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma adopt_node_is_component_unsafe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ adopt_node document_ptr node_ptr \\<^sub>h h'" assumes "ptr \ set |h \ get_dom_component (cast document_ptr)|\<^sub>r" assumes "ptr \ set |h \ get_dom_component (cast node_ptr)|\<^sub>r" assumes "ptr \ set |h \ get_dom_component (cast |h \ get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)|\<^sub>r)|\<^sub>r" shows "preserved (get_M ptr getter) h h'" proof - obtain owner_document where owner_document: "h \ get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr) \\<^sub>r owner_document" using assms(4) local.adopt_node_def by auto then obtain c where "h \ get_dom_component (cast owner_document) \\<^sub>r c" using get_dom_component_ok assms(1) assms(2) assms(3) get_owner_document_owner_document_in_heap by (meson document_ptr_kinds_commutes select_result_I) then have "ptr \ cast owner_document" using assms(6) assms(1) assms(2) assms(3) local.get_dom_component_ptr owner_document by (metis (no_types, lifting) assms(7) select_result_I2) have "document_ptr |\| document_ptr_kinds h" using adopt_node_document_in_heap assms(1) assms(2) assms(3) assms(4) by auto then have "ptr \ cast document_ptr" using assms(5) using assms(1) assms(2) assms(3) local.get_dom_component_ptr get_dom_component_ok by (meson document_ptr_kinds_commutes returns_result_select_result) have "\parent. |h \ get_parent node_ptr|\<^sub>r = Some parent \ parent \ ptr" by (metis assms(1) assms(2) assms(3) assms(6) is_OK_returns_result_I local.get_dom_component_ok local.get_dom_component_parent_inside local.get_dom_component_ptr local.get_owner_document_ptr_in_heap local.get_parent_ok node_ptr_kinds_commutes owner_document returns_result_select_result) show ?thesis using adopt_node_writes assms(4) apply(rule reads_writes_preserved2) apply(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_locs_def set_disconnected_nodes_locs_def all_args_def)[1] apply (metis \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr\ get_M_Mdocument_preserved3) apply (metis (no_types, lifting) \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\ get_M_Mdocument_preserved3 owner_document select_result_I2) apply(drule \\parent. |h \ get_parent node_ptr|\<^sub>r = Some parent \ parent \ ptr\)[1] apply (metis element_ptr_casts_commute3 get_M_Element_preserved8 is_node_ptr_kind_none node_ptr_casts_commute3 option.case_eq_if) apply (metis (no_types, lifting) \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\ get_M_Mdocument_preserved3 owner_document select_result_I2) apply (metis \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr\ get_M_Mdocument_preserved3) apply (metis (no_types, lifting) \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\ get_M_Mdocument_preserved3 owner_document select_result_I2) apply (metis \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr\ get_M_Mdocument_preserved3) apply (metis (no_types, lifting) \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\ get_M_Mdocument_preserved3 owner_document select_result_I2) apply(drule \\parent. |h \ get_parent node_ptr|\<^sub>r = Some parent \ parent \ ptr\)[1] apply (metis document_ptr_casts_commute3 get_M_Mdocument_preserved3) apply (metis (no_types, lifting) \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\ get_M_Mdocument_preserved3 owner_document select_result_I2) apply (metis \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr\ get_M_Mdocument_preserved3) apply (metis (no_types, lifting) \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\ get_M_Mdocument_preserved3 owner_document select_result_I2) apply (metis \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr\ get_M_Mdocument_preserved3) apply (metis (no_types, lifting) \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\ get_M_Mdocument_preserved3 owner_document select_result_I2) apply(drule \\parent. |h \ get_parent node_ptr|\<^sub>r = Some parent \ parent \ ptr\)[1] apply (metis (no_types, lifting) \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\ get_M_Mdocument_preserved3 owner_document select_result_I2) apply (metis \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr\ get_M_Mdocument_preserved3) apply (metis (no_types, lifting) \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\ get_M_Mdocument_preserved3 owner_document select_result_I2) done qed lemma adopt_node_is_strongly_dom_component_safe_step: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ adopt_node document_ptr node_ptr \\<^sub>h h'" assumes "ptr \ set |h \ get_scdom_component (cast document_ptr)|\<^sub>r" assumes "ptr \ set |h \ get_scdom_component (cast node_ptr)|\<^sub>r" shows "preserved (get_M ptr getter) h h'" proof - have "document_ptr |\| document_ptr_kinds h" by (meson assms(1) assms(2) assms(3) assms(4) is_OK_returns_heap_I local.adopt_node_document_in_heap) then obtain sc where sc: "h \ get_scdom_component (cast document_ptr) \\<^sub>r sc" using get_scdom_component_ok by (meson assms(1) assms(2) assms(3) document_ptr_kinds_commutes returns_result_select_result) have "node_ptr |\| node_ptr_kinds h" using assms(4) by (meson is_OK_returns_heap_I local.adopt_node_child_in_heap) then obtain child_sc where child_sc: "h \ get_scdom_component (cast node_ptr) \\<^sub>r child_sc" using get_scdom_component_ok by (meson assms(1) assms(2) assms(3) is_OK_returns_result_E node_ptr_kinds_commutes) then obtain owner_document where owner_document: "h \ get_owner_document (cast node_ptr) \\<^sub>r owner_document" by (meson \node_ptr |\| node_ptr_kinds h\ assms(1) assms(2) assms(3) contra_subsetD get_scdom_component_owner_document_same is_OK_returns_result_E get_scdom_component_subset_get_dom_component local.get_dom_component_ok local.get_dom_component_ptr node_ptr_kinds_commutes) then have "h \ get_scdom_component (cast owner_document) \\<^sub>r child_sc" using child_sc by (metis (no_types, lifting) \node_ptr |\| node_ptr_kinds h\ assms(1) assms(2) assms(3) get_scdom_component_owner_document_same get_scdom_component_ptrs_same_scope_component get_scdom_component_subset_get_dom_component is_OK_returns_result_E local.get_dom_component_ok local.get_dom_component_ptr node_ptr_kinds_commutes select_result_I2 subset_code(1)) have "ptr \ set |h \ get_dom_component (cast document_ptr)|\<^sub>r" by (metis (no_types, lifting) \document_ptr |\| document_ptr_kinds h\ assms(1) assms(2) assms(3) assms(5) contra_subsetD document_ptr_kinds_commutes get_scdom_component_subset_get_dom_component local.get_dom_component_ok returns_result_select_result sc select_result_I2) moreover have "ptr \ set |h \ get_dom_component (cast node_ptr)|\<^sub>r" by (metis (no_types, lifting) \node_ptr |\| node_ptr_kinds h\ assms(1) assms(2) assms(3) assms(6) child_sc contra_subsetD get_scdom_component_subset_get_dom_component local.get_dom_component_ok node_ptr_kinds_commutes returns_result_select_result select_result_I2) moreover have "ptr \ set |h \ get_scdom_component (cast |h \ get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)|\<^sub>r)|\<^sub>r" using get_scdom_component_owner_document_same get_scdom_component_ptrs_same_scope_component by (metis (no_types, lifting) \h \ get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document) \\<^sub>r child_sc\ assms(6) child_sc owner_document select_result_I2) have "ptr \ set |h \ get_dom_component (cast |h \ get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)|\<^sub>r)|\<^sub>r" using get_scdom_component_owner_document_same by (metis (no_types, opaque_lifting) \\thesis. (\owner_document. h \ get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr) \\<^sub>r owner_document \ thesis) \ thesis\ \h \ get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document) \\<^sub>r child_sc\ \ptr \ set |h \ get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r |h \ get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr)|\<^sub>r)|\<^sub>r\ assms(1) assms(2) assms(3) contra_subsetD document_ptr_kinds_commutes get_scdom_component_subset_get_dom_component is_OK_returns_result_E local.get_dom_component_ok local.get_owner_document_owner_document_in_heap owner_document returns_result_eq select_result_I2) ultimately show ?thesis using assms(1) assms(2) assms(3) assms(4) adopt_node_is_component_unsafe by blast qed lemma adopt_node_is_strongly_scdom_component_safe: assumes "heap_is_wellformed h" and type_wf: "type_wf h" and known_ptrs: "known_ptrs h" assumes "h \ adopt_node document_ptr child \\<^sub>h h'" shows "is_strongly_scdom_component_safe {cast document_ptr, cast child} {} h h'" proof - obtain old_document parent_opt h2 where old_document: "h \ get_owner_document (cast child) \\<^sub>r old_document" and parent_opt: "h \ get_parent child \\<^sub>r parent_opt" and h2: "h \ (case parent_opt of Some parent \ remove_child parent child | None \ return ()) \\<^sub>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 () }) \\<^sub>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 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 \\<^sub>r ptrs = h2 \ object_ptr_kinds_M \\<^sub>r ptrs" unfolding object_ptr_kinds_M_defs by simp then have object_ptr_kinds_eq_h: "|h \ object_ptr_kinds_M|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq_h: "|h \ node_ptr_kinds_M|\<^sub>r = |h2 \ node_ptr_kinds_M|\<^sub>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 \\<^sub>r old_disc_nodes" and h3: "h2 \ set_disconnected_nodes old_document (remove1 child old_disc_nodes) \\<^sub>h h3" and disc_nodes_document_ptr_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_document_ptr_h3" and h': "h3 \ set_disconnected_nodes document_ptr (child # disc_nodes_document_ptr_h3) \\<^sub>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 \\<^sub>r ptrs = h3 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq_h2: "|h2 \ object_ptr_kinds_M|\<^sub>r = |h3 \ object_ptr_kinds_M|\<^sub>r" by(simp) then have node_ptr_kinds_eq_h2: "|h2 \ node_ptr_kinds_M|\<^sub>r = |h3 \ node_ptr_kinds_M|\<^sub>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|\<^sub>r = |h3 \ document_ptr_kinds_M|\<^sub>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 \\<^sub>r children = h3 \ get_child_nodes ptr \\<^sub>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|\<^sub>r = |h3 \ get_child_nodes ptr|\<^sub>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 \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_eq_h3: "|h3 \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by(simp) then have node_ptr_kinds_eq_h3: "|h3 \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>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|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>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 \\<^sub>r children = h' \ get_child_nodes ptr \\<^sub>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|\<^sub>r = |h' \ get_child_nodes ptr|\<^sub>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 \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>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 \\<^sub>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 \\<^sub>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 \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>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 \\<^sub>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|\<^sub>r) |h2 \ document_ptr_kinds_M|\<^sub>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, 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 finite_set_in 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 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|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 2: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h2)))))" and 3: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h2). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" show "distinct (concat (map (\document_ptr. |h' \ get_disconnected_nodes document_ptr|\<^sub>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|\<^sub>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|\<^sub>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|\<^sub>r \ set |h2 \ get_disconnected_nodes y|\<^sub>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|\<^sub>r \ set |h' \ get_disconnected_nodes y|\<^sub>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|\<^sub>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|\<^sub>r \ set disc_nodes_old_document_h2 = {}" using \set |h2 \ get_disconnected_nodes y|\<^sub>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|\<^sub>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|\<^sub>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|\<^sub>r \ set |h2 \ get_disconnected_nodes old_document|\<^sub>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 fmember.rep_eq 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|\<^sub>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|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h')))))" and 1: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) (sorted_list_of_set (fset (document_ptr_kinds h2)))))" and 2: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h2). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" and 3: "xa |\| object_ptr_kinds h'" and 4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" and 5: "xb |\| document_ptr_kinds h'" and 6: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>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 \\<^sub>r disc_nodes_old_document_h2" assume a3: "h3 \ get_disconnected_nodes old_document \\<^sub>r remove1 child disc_nodes_old_document_h2" assume a4: "x \ set |h' \ get_child_nodes xa|\<^sub>r" assume "document_ptr_kinds h2 = document_ptr_kinds h'" assume a5: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h'). set |h2 \ get_disconnected_nodes x|\<^sub>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|\<^sub>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|\<^sub>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|\<^sub>r \ set |h2 \ get_disconnected_nodes xb|\<^sub>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 \\<^sub>r disc_nodes_document_ptr_h3" assume a3: "h2 \ get_disconnected_nodes old_document \\<^sub>r disc_nodes_old_document_h2" assume a4: "xa |\| object_ptr_kinds h'" assume a5: "h' \ get_disconnected_nodes document_ptr \\<^sub>r child # disc_nodes_document_ptr_h3" assume a6: "old_document |\| document_ptr_kinds h'" assume a7: "x \ set |h' \ get_disconnected_nodes xb|\<^sub>r" assume a8: "x \ set |h' \ get_child_nodes xa|\<^sub>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|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" assume a11: "\doc_ptr. document_ptr \ doc_ptr \ |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" assume a12: "(\x\fset (object_ptr_kinds h'). set |h' \ get_child_nodes x|\<^sub>r) \ (\x\fset (document_ptr_kinds h'). set |h2 \ get_disconnected_nodes x|\<^sub>r) = {}" have f13: "\d. d \ set |h' \ document_ptr_kinds_M|\<^sub>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|\<^sub>r = disc_nodes_old_document_h2" using a6 a3 by simp have "x \ set |h2 \ get_disconnected_nodes xb|\<^sub>r" using a12 a8 a4 \xb |\| document_ptr_kinds h'\ by (meson UN_I disjoint_iff_not_equal fmember.rep_eq) 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 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 have "heap_is_wellformed h'" and "known_ptrs h'" and "type_wf h'" by auto 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 adopt_node_writes assms(4)]) unfolding adopt_node_locs_def using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved remove_child_pointers_preserved by (auto simp add: reflp_def transp_def split: if_splits) show ?thesis apply(auto simp add: is_strongly_scdom_component_safe_def Let_def object_ptr_kinds_eq3 )[1] using adopt_node_is_strongly_dom_component_safe_step get_scdom_component_impl assms by blast qed end interpretation i_get_scdom_component_adopt_node?: l_get_scdom_component_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_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 get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs remove to_tree_order get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name by(auto simp add: l_get_scdom_component_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_scdom_component_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \create\_element\ locale l_get_scdom_component_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_dom_component_get_scdom_component + l_get_dom_component_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma create_element_is_strongly_scdom_component_safe_step: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ create_element document_ptr tag \\<^sub>h h'" assumes "ptr \ set |h \ get_scdom_component (cast document_ptr)|\<^sub>r" assumes "ptr \ cast |h \ create_element document_ptr tag|\<^sub>r" shows "preserved (get_M ptr getter) h h'" proof - obtain new_element_ptr h2 h3 disc_nodes where new_element_ptr: "h \ new_element \\<^sub>r new_element_ptr" and h2: "h \ new_element \\<^sub>h h2" and h3: "h2 \set_tag_name new_element_ptr tag \\<^sub>h h3" and disc_nodes: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes" and h': "h3 \ set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes) \\<^sub>h h'" using assms(4) by(auto simp add: create_element_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated]) have 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 "heap_is_wellformed h'" using assms(4) using assms(1) assms(2) assms(3) local.create_element_preserves_wellformedness(1) by blast have "type_wf h'" using assms(1) assms(2) assms(3) assms(4) local.create_element_preserves_wellformedness(2) by blast have "known_ptrs h'" using assms(1) assms(2) assms(3) assms(4) local.create_element_preserves_wellformedness(3) by blast have "document_ptr |\| document_ptr_kinds h" by (meson assms(4) is_OK_returns_heap_I local.create_element_document_in_heap) then obtain sc where sc: "h \ get_scdom_component (cast document_ptr) \\<^sub>r sc" using get_scdom_component_ok by (meson assms(1) assms(2) assms(3) document_ptr_kinds_commutes returns_result_select_result) have "document_ptr |\| document_ptr_kinds h'" using \document_ptr |\| document_ptr_kinds h\ document_ptr_kinds_eq_h using document_ptr_kinds_eq_h2 document_ptr_kinds_eq_h3 by blast then obtain sc' where sc': "h' \ get_scdom_component (cast document_ptr) \\<^sub>r sc'" using get_scdom_component_ok by (meson \heap_is_wellformed h'\ \known_ptrs h'\ \type_wf h'\ document_ptr_kinds_commutes returns_result_select_result) obtain c where c: "h \ get_dom_component (cast document_ptr) \\<^sub>r c" by (meson \document_ptr |\| document_ptr_kinds h\ assms(1) assms(2) assms(3) document_ptr_kinds_commutes is_OK_returns_result_E local.get_dom_component_ok) have "set c \ set sc" using assms(1) assms(2) assms(3) c get_scdom_component_subset_get_dom_component sc by blast have "ptr \ set c" using \set c \ set sc\ assms(5) sc by auto then show ?thesis using create_element_is_weakly_dom_component_safe by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) assms(6) c local.create_element_is_weakly_dom_component_safe_step select_result_I2) qed lemma create_element_is_strongly_scdom_component_safe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ create_element document_ptr tag \\<^sub>r result" assumes "h \ create_element document_ptr tag \\<^sub>h h'" shows "is_strongly_scdom_component_safe {cast document_ptr} {cast result} h h'" proof - obtain new_element_ptr h2 h3 disc_nodes_h3 where new_element_ptr: "h \ new_element \\<^sub>r new_element_ptr" and h2: "h \ new_element \\<^sub>h h2" and h3: "h2 \ set_tag_name new_element_ptr tag \\<^sub>h h3" and disc_nodes_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" and h': "h3 \ set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes_h3) \\<^sub>h h'" using assms(5) by(auto simp add: create_element_def returns_result_heap_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 \\<^sub>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) then have "result = new_element_ptr" using assms(4) by auto have "new_element_ptr \ set |h \ element_ptr_kinds_M|\<^sub>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|\<^sub>r" by simp then have "cast new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>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 \\<^sub>r new_element_ptr\ local.create_element_known_ptr by blast then have "known_ptrs h2" using known_ptrs_new_ptr object_ptr_kinds_eq_h \known_ptrs h\ h2 by blast then have "known_ptrs h3" using known_ptrs_preserved object_ptr_kinds_eq_h2 by blast then have "known_ptrs h'" using known_ptrs_preserved object_ptr_kinds_eq_h3 by blast have "document_ptr |\| document_ptr_kinds h" using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2 get_disconnected_nodes_ptr_in_heap \type_wf h\ document_ptr_kinds_def by (metis is_OK_returns_result_I) have children_eq_h: "\ptr' children. ptr' \ cast new_element_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes (cast new_element_ptr) \\<^sub>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 \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have "type_wf h2" using \type_wf h\ new_element_types_preserved h2 by blast then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_tag_name_writes h3] using set_tag_name_types_preserved by(auto simp add: reflp_def transp_def) then have "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h'] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) have children_eq_h3: "\ptr' children. h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>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 \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>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 \\<^sub>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 \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ a_all_ptrs_in_heap_def heap_is_wellformed_def using NodeMonad.ptr_kinds_ptr_kinds_M local.heap_is_wellformed_disc_nodes_in_heap by blast have "parent_child_rel h = parent_child_rel h'" proof - have "parent_child_rel h = parent_child_rel h2" proof(auto simp add: parent_child_rel_def)[1] fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>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|\<^sub>r" then show "x \ set |h2 \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq_h \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ children_eq2_h empty_iff empty_set image_eqI select_result_I2) qed also have "\ = parent_child_rel h3" by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2) also have "\ = parent_child_rel h'" by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3) finally show ?thesis by simp qed have "document_ptr |\| document_ptr_kinds h'" by (simp add: \document_ptr |\| document_ptr_kinds h\ document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 document_ptr_kinds_eq_h3) have "known_ptr (cast document_ptr)" using \document_ptr |\| document_ptr_kinds h\ assms(3) document_ptr_kinds_commutes local.known_ptrs_known_ptr by blast have "h \ get_owner_document (cast document_ptr) \\<^sub>r document_ptr" using \known_ptr (cast document_ptr)\ \document_ptr |\| document_ptr_kinds h\ apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_splits, rule conjI)+ by(auto simp add: known_ptr_impl known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_returns_result_I split: option.splits) have "h' \ get_owner_document (cast document_ptr) \\<^sub>r document_ptr" using \known_ptr (cast document_ptr)\ \document_ptr |\| document_ptr_kinds h'\ apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_splits, rule conjI)+ by(auto simp add: known_ptr_impl known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_returns_result_I split: option.splits) obtain to where to: "h \ to_tree_order (cast document_ptr) \\<^sub>r to" by (meson \h \ get_owner_document (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr) \\<^sub>r document_ptr\ assms(1) assms(2) assms(3) is_OK_returns_result_E is_OK_returns_result_I local.get_owner_document_ptr_in_heap local.to_tree_order_ok) obtain to' where to': "h' \ to_tree_order (cast document_ptr) \\<^sub>r to'" by (metis \document_ptr |\| document_ptr_kinds h\ \known_ptrs h'\ \type_wf h'\ assms(1) assms(2) assms(3) assms(5) document_ptr_kinds_commutes document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 document_ptr_kinds_eq_h3 is_OK_returns_result_E local.create_element_preserves_wellformedness(1) local.to_tree_order_ok) have "set to = set to'" proof safe fix x assume "x \ set to" show "x \ set to'" using to to' using to_tree_order_parent_child_rel \parent_child_rel h = parent_child_rel h'\ by (metis \known_ptrs h'\ \type_wf h'\ \x \ set to\ assms(1) assms(2) assms(3) assms(5) local.create_element_preserves_wellformedness(1)) next fix x assume "x \ set to'" show "x \ set to" using to to' using to_tree_order_parent_child_rel \parent_child_rel h = parent_child_rel h'\ by (metis \known_ptrs h'\ \type_wf h'\ \x \ set to'\ assms(1) assms(2) assms(3) assms(5) local.create_element_preserves_wellformedness(1)) qed have "h' \ get_disconnected_nodes document_ptr \\<^sub>r cast new_element_ptr # disc_nodes_h3" using h' local.set_disconnected_nodes_get_disconnected_nodes by auto obtain disc_nodes_h' where disc_nodes_h': "h' \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h'" and "cast new_element_ptr \ set disc_nodes_h'" and "disc_nodes_h' = cast new_element_ptr # disc_nodes_h3" by (simp add: \h' \ get_disconnected_nodes document_ptr \\<^sub>r cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr # disc_nodes_h3\) have "\disc_ptr to to'. disc_ptr \ set disc_nodes_h3 \ h \ to_tree_order (cast disc_ptr) \\<^sub>r to \ h' \ to_tree_order (cast disc_ptr) \\<^sub>r to' \ set to = set to'" proof safe fix disc_ptr to to' x assume "disc_ptr \ set disc_nodes_h3" assume "h \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \\<^sub>r to" assume "h' \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \\<^sub>r to'" assume "x \ set to" show "x \ set to'" using to_tree_order_parent_child_rel \parent_child_rel h = parent_child_rel h'\ by (metis \h \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \\<^sub>r to\ \h' \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \\<^sub>r to'\ \known_ptrs h'\ \type_wf h'\ \x \ set to\ assms(1) assms(2) assms(3) assms(5) local.create_element_preserves_wellformedness(1)) next fix disc_ptr to to' x assume "disc_ptr \ set disc_nodes_h3" assume "h \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \\<^sub>r to" assume "h' \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \\<^sub>r to'" assume "x \ set to'" show "x \ set to" using to_tree_order_parent_child_rel \parent_child_rel h = parent_child_rel h'\ by (metis \h \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \\<^sub>r to\ \h' \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \\<^sub>r to'\ \known_ptrs h'\ \type_wf h'\ \x \ set to'\ assms(1) assms(2) assms(3) assms(5) local.create_element_preserves_wellformedness(1)) qed have "heap_is_wellformed h'" using assms(1) assms(2) assms(3) assms(5) local.create_element_preserves_wellformedness(1) by blast have "cast new_element_ptr |\| object_ptr_kinds h'" using \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set disc_nodes_h'\ \heap_is_wellformed h'\ disc_nodes_h' local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_commutes by blast then have "new_element_ptr |\| element_ptr_kinds h'" by simp have "\node_ptr. node_ptr \ set disc_nodes_h3 \ node_ptr |\| node_ptr_kinds h'" by (meson \heap_is_wellformed h'\ h' local.heap_is_wellformed_disc_nodes_in_heap local.set_disconnected_nodes_get_disconnected_nodes set_subset_Cons subset_code(1)) have "h \ ok (map_M (to_tree_order \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) disc_nodes_h3)" using assms(1) assms(2) assms(3) to_tree_order_ok apply(auto intro!: map_M_ok_I)[1] using disc_nodes_document_ptr_h local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_commutes by blast then obtain disc_tree_orders where disc_tree_orders: "h \ map_M (to_tree_order \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) disc_nodes_h3 \\<^sub>r disc_tree_orders" by auto have "h' \ ok (map_M (to_tree_order \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) disc_nodes_h')" apply(auto intro!: map_M_ok_I)[1] apply(simp add: \disc_nodes_h' = cast new_element_ptr # disc_nodes_h3\) using \\node_ptr. node_ptr \ set disc_nodes_h3 \ node_ptr |\| node_ptr_kinds h'\ \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set disc_nodes_h'\ \heap_is_wellformed h'\ \known_ptrs h'\ \type_wf h'\ disc_nodes_h' local.heap_is_wellformed_disc_nodes_in_heap local.to_tree_order_ok node_ptr_kinds_commutes by blast then obtain disc_tree_orders' where disc_tree_orders': "h' \ map_M (to_tree_order \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) disc_nodes_h' \\<^sub>r disc_tree_orders'" by auto have "h' \ get_child_nodes (cast new_element_ptr) \\<^sub>r []" using \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ children_eq_h2 children_eq_h3 by auto obtain new_tree_order where new_tree_order: "h' \ to_tree_order (cast new_element_ptr) \\<^sub>r new_tree_order" and "new_tree_order \ set disc_tree_orders'" using map_M_pure_E[OF disc_tree_orders' \cast new_element_ptr \ set disc_nodes_h'\] by auto then have "new_tree_order = [cast new_element_ptr]" using \h' \ get_child_nodes (cast new_element_ptr) \\<^sub>r []\ by(auto simp add: to_tree_order_def dest!: bind_returns_result_E3[rotated, OF \h' \ get_child_nodes (cast new_element_ptr) \\<^sub>r []\, rotated]) obtain foo where foo: "h' \ map_M (to_tree_order \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr # disc_nodes_h3) \\<^sub>r [cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr] # foo" apply(auto intro!: bind_pure_returns_result_I map_M_pure_I)[1] using \new_tree_order = [cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr]\ new_tree_order apply auto[1] by (smt \disc_nodes_h' = cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr # disc_nodes_h3\ bind_pure_returns_result_I bind_returns_result_E2 comp_apply disc_tree_orders' local.to_tree_order_pure map_M.simps(2) map_M_pure_I return_returns_result returns_result_eq) then have "set (concat foo) = set (concat disc_tree_orders)" apply(auto elim!: bind_returns_result_E2 intro!: map_M_pure_I)[1] apply (smt \\to' toa disc_ptr. \disc_ptr \ set disc_nodes_h3; h \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \\<^sub>r toa; h' \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \\<^sub>r to'\ \ set toa = set to'\ comp_apply disc_tree_orders local.to_tree_order_pure map_M_pure_E map_M_pure_E2) using \\to' toa disc_ptr. \disc_ptr \ set disc_nodes_h3; h \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \\<^sub>r toa; h' \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \\<^sub>r to'\ \ set toa = set to'\ comp_apply disc_tree_orders local.to_tree_order_pure map_M_pure_E map_M_pure_E2 by smt have "disc_tree_orders' = [cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr] # foo" using foo disc_tree_orders' by (simp add: \disc_nodes_h' = cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr # disc_nodes_h3\ returns_result_eq) have "set (concat disc_tree_orders') = {cast new_element_ptr} \ set (concat disc_tree_orders)" apply(auto simp add: \disc_tree_orders' = [cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr] # foo\)[1] using \set (concat foo) = set (concat disc_tree_orders)\ by auto have "h' \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr) \\<^sub>r to' @ concat disc_tree_orders'" using \h' \ get_owner_document (cast document_ptr) \\<^sub>r document_ptr\ disc_nodes_h' to' disc_tree_orders' by(auto simp add: a_get_scdom_component_def intro!: bind_pure_returns_result_I map_M_pure_I) then have "set |h' \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r = set to' \ set (concat disc_tree_orders')" by auto have "h \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr) \\<^sub>r to @ concat disc_tree_orders" using \h \ get_owner_document (cast document_ptr) \\<^sub>r document_ptr\ disc_nodes_document_ptr_h to disc_tree_orders by(auto simp add: a_get_scdom_component_def intro!: bind_pure_returns_result_I map_M_pure_I) then have "set |h \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r = set to \ set (concat disc_tree_orders)" by auto have "{cast new_element_ptr} \ set |h \ local.a_get_scdom_component (cast document_ptr)|\<^sub>r = set |h' \ local.a_get_scdom_component (cast document_ptr)|\<^sub>r" proof(safe) show "cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h' \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r" using \h' \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr) \\<^sub>r to' @ concat disc_tree_orders'\ apply(auto simp add: a_get_scdom_component_def)[1] by (meson \\thesis. (\new_tree_order. \h' \ to_tree_order (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r new_tree_order; new_tree_order \ set disc_tree_orders'\ \ thesis) \ thesis\ local.to_tree_order_ptr_in_result) next fix x assume " x \ set |h \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r" then show "x \ set |h' \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r" using \set |h \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r = set to \set (concat disc_tree_orders)\ using \set |h' \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r = set to' \ set (concat disc_tree_orders')\ using \set to = set to'\ using \set (concat disc_tree_orders') = {cast new_element_ptr} \ set (concat disc_tree_orders)\ by(auto) next fix x assume " x \ set |h' \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r" assume "x \ set |h \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r" show "x = cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr" using \set |h \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r = set to \ set (concat disc_tree_orders)\ using \set |h' \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r = set to' \ set (concat disc_tree_orders')\ using \set to = set to'\ using \set (concat disc_tree_orders') = {cast new_element_ptr} \ set (concat disc_tree_orders)\ using \x \ set |h' \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r\ \x \ set |h \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r\ by auto qed have "object_ptr_kinds h' = object_ptr_kinds h |\| {|cast new_element_ptr|}" using object_ptr_kinds_eq_h object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 by auto then show ?thesis apply(auto simp add: is_strongly_scdom_component_safe_def Let_def)[1] apply(rule bexI[where x="cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr"]) using \{cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr} \ set |h \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r = set |h' \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r\ apply auto[2] using \set to = set to'\ \set |h \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r = set to \ set (concat disc_tree_orders)\ local.to_tree_order_ptr_in_result to' apply auto[1] using \document_ptr |\| document_ptr_kinds h\ apply blast apply(rule bexI[where x="cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr"]) using \result = new_element_ptr\ \{cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr} \ set |h \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r = set |h' \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r\ apply auto[1] apply(auto)[1] using \set to = set to'\ \set |h \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r = set to \ set (concat disc_tree_orders)\ local.to_tree_order_ptr_in_result to' apply auto[1] apply (simp add: \document_ptr |\| document_ptr_kinds h\) using \\thesis. (\new_element_ptr h2 h3 disc_nodes_h3. \h \ new_element \\<^sub>r new_element_ptr; h \ new_element \\<^sub>h h2; h2 \ set_tag_name new_element_ptr tag \\<^sub>h h3; h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3; h3 \ set_disconnected_nodes document_ptr (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr # disc_nodes_h3) \\<^sub>h h'\ \ thesis) \ thesis\ new_element_ptr new_element_ptr_not_in_heap apply auto[1] using create_element_is_strongly_scdom_component_safe_step by (smt ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ \result = new_element_ptr\ assms(1) assms(2) assms(3) assms(4) assms(5) local.get_scdom_component_impl select_result_I2) qed end interpretation i_get_scdom_component_remove_child?: l_get_scdom_component_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name set_child_nodes set_child_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs remove_child remove_child_locs remove by(auto simp add: l_get_scdom_component_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_scdom_component_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \create\_character\_data\ locale l_get_scdom_component_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_dom_component_get_scdom_component + l_get_dom_component_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_create_character_data_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma create_character_data_is_strongly_dom_component_safe_step: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ create_character_data document_ptr text \\<^sub>h h'" assumes "ptr \ set |h \ get_scdom_component (cast document_ptr)|\<^sub>r" assumes "ptr \ cast |h \ create_character_data document_ptr text|\<^sub>r" shows "preserved (get_M ptr getter) h h'" proof - have "document_ptr |\| document_ptr_kinds h" by (meson assms(4) is_OK_returns_heap_I local.create_character_data_document_in_heap) then obtain sc where sc: "h \ get_scdom_component (cast document_ptr) \\<^sub>r sc" using get_scdom_component_ok by (meson assms(1) assms(2) assms(3) document_ptr_kinds_commutes returns_result_select_result) obtain c where c: "h \ get_dom_component (cast document_ptr) \\<^sub>r c" by (meson \document_ptr |\| document_ptr_kinds h\ assms(1) assms(2) assms(3) document_ptr_kinds_commutes is_OK_returns_result_E local.get_dom_component_ok) have "set c \ set sc" using assms(1) assms(2) assms(3) c get_scdom_component_subset_get_dom_component sc by blast have "ptr \ set c" using \set c \ set sc\ assms(5) sc by auto then show ?thesis by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) assms(6) c local.create_character_data_is_weakly_dom_component_safe_step select_result_I2) qed lemma create_character_data_is_strongly_dom_component_safe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ create_character_data document_ptr text \\<^sub>r result" assumes "h \ create_character_data document_ptr text \\<^sub>h h'" shows "is_strongly_scdom_component_safe {cast document_ptr} {cast result} h h'" proof - obtain new_character_data_ptr h2 h3 disc_nodes_h3 where new_character_data_ptr: "h \ new_character_data \\<^sub>r new_character_data_ptr" and h2: "h \ new_character_data \\<^sub>h h2" and h3: "h2 \ set_val new_character_data_ptr text \\<^sub>h h3" and disc_nodes_h3: "h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3" and h': "h3 \ set_disconnected_nodes document_ptr (cast new_character_data_ptr # disc_nodes_h3) \\<^sub>h h'" using assms(5) by(auto simp add: create_character_data_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] ) then have "h \ create_character_data document_ptr text \\<^sub>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) then have "result = new_character_data_ptr" using assms(4) by auto have "new_character_data_ptr \ set |h \ character_data_ptr_kinds_M|\<^sub>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|\<^sub>r" by simp then have "cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r" by simp have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_character_data_ptr|}" using new_character_data_new_ptr h2 new_character_data_ptr by blast then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = node_ptr_kinds h |\| {|cast new_character_data_ptr|}" apply(simp add: node_ptr_kinds_def) by force then have character_data_ptr_kinds_eq_h: "character_data_ptr_kinds h2 = character_data_ptr_kinds h |\| {|new_character_data_ptr|}" apply(simp add: character_data_ptr_kinds_def) by force have element_ptr_kinds_eq_h: "element_ptr_kinds h2 = element_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: node_ptr_kinds_def element_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: document_ptr_kinds_def) have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_val_writes h3]) using set_val_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h2: "node_ptr_kinds h3 = node_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def) have object_ptr_kinds_eq_h3: "object_ptr_kinds h' = object_ptr_kinds h3" apply(rule writes_small_big[where P="\h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_disconnected_nodes_writes h']) using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) then have document_ptr_kinds_eq_h3: "document_ptr_kinds h' = document_ptr_kinds h3" by (auto simp add: document_ptr_kinds_def) have node_ptr_kinds_eq_h3: "node_ptr_kinds h' = node_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: node_ptr_kinds_def) have "document_ptr |\| document_ptr_kinds h" using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2 get_disconnected_nodes_ptr_in_heap \type_wf h\ document_ptr_kinds_def by (metis is_OK_returns_result_I) have children_eq_h: "\ptr' children. ptr' \ cast new_character_data_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>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' children. ptr' \ cast new_character_data_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>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 \\<^sub>r disc_nodes = h2 \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. h2 \ get_child_nodes ptr' \\<^sub>r children = h3 \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h3 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h3 \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h3 \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have "type_wf h2" using \type_wf h\ new_character_data_types_preserved h2 by blast then have "type_wf h3" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_val_writes h3] using set_val_types_preserved by(auto simp add: reflp_def transp_def) then have "type_wf h'" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_disconnected_nodes_writes h'] using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) have children_eq_h3: "\ptr' children. h3 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>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'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>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 \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>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|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have disc_nodes_document_ptr_h2: "h2 \ get_disconnected_nodes document_ptr \\<^sub>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 \\<^sub>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|\<^sub>r\ a_all_ptrs_in_heap_def heap_is_wellformed_def using NodeMonad.ptr_kinds_ptr_kinds_M local.heap_is_wellformed_disc_nodes_in_heap by blast have "parent_child_rel h = parent_child_rel h'" proof - have "parent_child_rel h = parent_child_rel h2" proof(auto simp add: parent_child_rel_def)[1] fix a x assume 0: "a |\| object_ptr_kinds h" and 1: "x \ set |h \ get_child_nodes a|\<^sub>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|\<^sub>r" then show "x \ set |h2 \ get_child_nodes a|\<^sub>r" by (metis ObjectMonad.ptr_kinds_ptr_kinds_M \cast new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "a |\| object_ptr_kinds h" using object_ptr_kinds_eq_h \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ by(auto) next fix a x assume 0: "a |\| object_ptr_kinds h2" and 1: "x \ set |h2 \ get_child_nodes a|\<^sub>r" then show "x \ set |h \ get_child_nodes a|\<^sub>r" by (metis (no_types, lifting) \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ children_eq2_h empty_iff empty_set image_eqI select_result_I2) qed also have "\ = parent_child_rel h3" by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2) also have "\ = parent_child_rel h'" by(auto simp add: parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3) finally show ?thesis by simp qed have "known_ptr (cast new_character_data_ptr)" using \h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr\ 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 have "known_ptrs h'" using known_ptrs_preserved object_ptr_kinds_eq_h3 by blast have "document_ptr |\| document_ptr_kinds h'" by (simp add: \document_ptr |\| document_ptr_kinds h\ document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 document_ptr_kinds_eq_h3) have "known_ptr (cast document_ptr)" using \document_ptr |\| document_ptr_kinds h\ assms(3) document_ptr_kinds_commutes local.known_ptrs_known_ptr by blast have "h \ get_owner_document (cast document_ptr) \\<^sub>r document_ptr" using \known_ptr (cast document_ptr)\ \document_ptr |\| document_ptr_kinds h\ apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_splits, rule conjI)+ by(auto simp add: known_ptr_impl known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_returns_result_I split: option.splits) have "h' \ get_owner_document (cast document_ptr) \\<^sub>r document_ptr" using \known_ptr (cast document_ptr)\ \document_ptr |\| document_ptr_kinds h'\ apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(split invoke_splits, rule conjI)+ by(auto simp add: known_ptr_impl known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs NodeClass.known_ptr_defs a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_returns_result_I split: option.splits) obtain to where to: "h \ to_tree_order (cast document_ptr) \\<^sub>r to" by (meson \h \ get_owner_document (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr) \\<^sub>r document_ptr\ assms(1) assms(2) assms(3) is_OK_returns_result_E is_OK_returns_result_I local.get_owner_document_ptr_in_heap local.to_tree_order_ok) obtain to' where to': "h' \ to_tree_order (cast document_ptr) \\<^sub>r to'" by (metis \document_ptr |\| document_ptr_kinds h\ \known_ptrs h'\ \type_wf h'\ assms(1) assms(2) assms(3) assms(5) document_ptr_kinds_commutes document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 document_ptr_kinds_eq_h3 is_OK_returns_result_E local.create_character_data_preserves_wellformedness(1) local.to_tree_order_ok) have "set to = set to'" proof safe fix x assume "x \ set to" show "x \ set to'" using to to' using to_tree_order_parent_child_rel \parent_child_rel h = parent_child_rel h'\ by (metis \known_ptrs h'\ \type_wf h'\ \x \ set to\ assms(1) assms(2) assms(3) assms(5) local.create_character_data_preserves_wellformedness(1)) next fix x assume "x \ set to'" show "x \ set to" using to to' using to_tree_order_parent_child_rel \parent_child_rel h = parent_child_rel h'\ by (metis \known_ptrs h'\ \type_wf h'\ \x \ set to'\ assms(1) assms(2) assms(3) assms(5) local.create_character_data_preserves_wellformedness(1)) qed have "h' \ get_disconnected_nodes document_ptr \\<^sub>r cast new_character_data_ptr # disc_nodes_h3" using h' local.set_disconnected_nodes_get_disconnected_nodes by auto obtain disc_nodes_h' where disc_nodes_h': "h' \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h'" and "cast new_character_data_ptr \ set disc_nodes_h'" and "disc_nodes_h' = cast new_character_data_ptr # disc_nodes_h3" by (simp add: \h' \ get_disconnected_nodes document_ptr \\<^sub>r cast new_character_data_ptr # disc_nodes_h3\) have "\disc_ptr to to'. disc_ptr \ set disc_nodes_h3 \ h \ to_tree_order (cast disc_ptr) \\<^sub>r to \ h' \ to_tree_order (cast disc_ptr) \\<^sub>r to' \ set to = set to'" proof safe fix disc_ptr to to' x assume "disc_ptr \ set disc_nodes_h3" assume "h \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \\<^sub>r to" assume "h' \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \\<^sub>r to'" assume "x \ set to" show "x \ set to'" using to_tree_order_parent_child_rel \parent_child_rel h = parent_child_rel h'\ by (metis \h \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \\<^sub>r to\ \h' \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \\<^sub>r to'\ \known_ptrs h'\ \type_wf h'\ \x \ set to\ assms(1) assms(2) assms(3) assms(5) local.create_character_data_preserves_wellformedness(1)) next fix disc_ptr to to' x assume "disc_ptr \ set disc_nodes_h3" assume "h \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \\<^sub>r to" assume "h' \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \\<^sub>r to'" assume "x \ set to'" show "x \ set to" using to_tree_order_parent_child_rel \parent_child_rel h = parent_child_rel h'\ by (metis \h \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \\<^sub>r to\ \h' \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \\<^sub>r to'\ \known_ptrs h'\ \type_wf h'\ \x \ set to'\ assms(1) assms(2) assms(3) assms(5) local.create_character_data_preserves_wellformedness(1)) qed have "heap_is_wellformed h'" using assms(1) assms(2) assms(3) assms(5) local.create_character_data_preserves_wellformedness(1) by blast have "cast new_character_data_ptr |\| object_ptr_kinds h'" using \cast new_character_data_ptr \ set disc_nodes_h'\ \heap_is_wellformed h'\ disc_nodes_h' local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_commutes by blast then have "new_character_data_ptr |\| character_data_ptr_kinds h'" by simp have "\node_ptr. node_ptr \ set disc_nodes_h3 \ node_ptr |\| node_ptr_kinds h'" by (meson \heap_is_wellformed h'\ h' local.heap_is_wellformed_disc_nodes_in_heap local.set_disconnected_nodes_get_disconnected_nodes set_subset_Cons subset_code(1)) have "h \ ok (map_M (to_tree_order \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) disc_nodes_h3)" using assms(1) assms(2) assms(3) to_tree_order_ok apply(auto intro!: map_M_ok_I)[1] using disc_nodes_document_ptr_h local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_commutes by blast then obtain disc_tree_orders where disc_tree_orders: "h \ map_M (to_tree_order \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) disc_nodes_h3 \\<^sub>r disc_tree_orders" by auto have "h' \ ok (map_M (to_tree_order \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) disc_nodes_h')" apply(auto intro!: map_M_ok_I)[1] apply(simp add: \disc_nodes_h' = cast new_character_data_ptr # disc_nodes_h3\) using \\node_ptr. node_ptr \ set disc_nodes_h3 \ node_ptr |\| node_ptr_kinds h'\ \cast new_character_data_ptr \ set disc_nodes_h'\ \heap_is_wellformed h'\ \known_ptrs h'\ \type_wf h'\ disc_nodes_h' local.heap_is_wellformed_disc_nodes_in_heap local.to_tree_order_ok node_ptr_kinds_commutes by blast then obtain disc_tree_orders' where disc_tree_orders': "h' \ map_M (to_tree_order \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) disc_nodes_h' \\<^sub>r disc_tree_orders'" by auto have "h' \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []" using \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ children_eq_h2 children_eq_h3 by auto obtain new_tree_order where new_tree_order: "h' \ to_tree_order (cast new_character_data_ptr) \\<^sub>r new_tree_order" and "new_tree_order \ set disc_tree_orders'" using map_M_pure_E[OF disc_tree_orders' \cast new_character_data_ptr \ set disc_nodes_h'\] by auto then have "new_tree_order = [cast new_character_data_ptr]" using \h' \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ by(auto simp add: to_tree_order_def dest!: bind_returns_result_E3[rotated, OF \h' \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\, rotated]) obtain foo where foo: "h' \ map_M (to_tree_order \ cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) (cast new_character_data_ptr # disc_nodes_h3) \\<^sub>r [cast new_character_data_ptr] # foo" apply(auto intro!: bind_pure_returns_result_I map_M_pure_I)[1] using \new_tree_order = [cast new_character_data_ptr]\ new_tree_order apply auto[1] using \disc_nodes_h' = cast new_character_data_ptr # disc_nodes_h3\ bind_pure_returns_result_I bind_returns_result_E2 comp_apply disc_tree_orders' local.to_tree_order_pure map_M.simps(2) map_M_pure_I return_returns_result returns_result_eq apply simp by (smt \disc_nodes_h' = cast new_character_data_ptr # disc_nodes_h3\ bind_pure_returns_result_I bind_returns_result_E2 comp_apply disc_tree_orders' local.to_tree_order_pure map_M.simps(2) map_M_pure_I return_returns_result returns_result_eq) then have "set (concat foo) = set (concat disc_tree_orders)" apply(auto elim!: bind_returns_result_E2 intro!: map_M_pure_I)[1] apply (smt \\to' toa disc_ptr. \disc_ptr \ set disc_nodes_h3; h \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \\<^sub>r toa; h' \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \\<^sub>r to'\ \ set toa = set to'\ comp_apply disc_tree_orders local.to_tree_order_pure map_M_pure_E map_M_pure_E2) using \\to' toa disc_ptr. \disc_ptr \ set disc_nodes_h3; h \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \\<^sub>r toa; h' \ to_tree_order (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r disc_ptr) \\<^sub>r to'\ \ set toa = set to'\ comp_apply disc_tree_orders local.to_tree_order_pure map_M_pure_E map_M_pure_E2 by smt have "disc_tree_orders' = [cast new_character_data_ptr] # foo" using foo disc_tree_orders' by (simp add: \disc_nodes_h' = cast new_character_data_ptr # disc_nodes_h3\ returns_result_eq) have "set (concat disc_tree_orders') = {cast new_character_data_ptr} \ set (concat disc_tree_orders)" apply(auto simp add: \disc_tree_orders' = [cast new_character_data_ptr] # foo\)[1] using \set (concat foo) = set (concat disc_tree_orders)\ by auto have "h' \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr) \\<^sub>r to' @ concat disc_tree_orders'" using \h' \ get_owner_document (cast document_ptr) \\<^sub>r document_ptr\ disc_nodes_h' to' disc_tree_orders' by(auto simp add: a_get_scdom_component_def intro!: bind_pure_returns_result_I map_M_pure_I) then have "set |h' \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r = set to' \ set (concat disc_tree_orders')" by auto have "h \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr) \\<^sub>r to @ concat disc_tree_orders" using \h \ get_owner_document (cast document_ptr) \\<^sub>r document_ptr\ disc_nodes_document_ptr_h to disc_tree_orders by(auto simp add: a_get_scdom_component_def intro!: bind_pure_returns_result_I map_M_pure_I) then have "set |h \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r = set to \ set (concat disc_tree_orders)" by auto have "{cast new_character_data_ptr} \ set |h \ local.a_get_scdom_component (cast document_ptr)|\<^sub>r = set |h' \ local.a_get_scdom_component (cast document_ptr)|\<^sub>r" proof(safe) show "cast new_character_data_ptr \ set |h' \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r" using \h' \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr) \\<^sub>r to' @ concat disc_tree_orders'\ apply(auto simp add: a_get_scdom_component_def)[1] by (meson \\thesis. (\new_tree_order. \h' \ to_tree_order (cast new_character_data_ptr) \\<^sub>r new_tree_order; new_tree_order \ set disc_tree_orders'\ \ thesis) \ thesis\ local.to_tree_order_ptr_in_result) next fix x assume " x \ set |h \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r" then show "x \ set |h' \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r" using \set |h \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r = set to \ set (concat disc_tree_orders)\ using \set |h' \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r = set to' \ set (concat disc_tree_orders')\ using \set to = set to'\ using \set (concat disc_tree_orders') = {cast new_character_data_ptr} \ set (concat disc_tree_orders)\ by(auto) next fix x assume " x \ set |h' \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r" assume "x \ set |h \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r" show "x = cast new_character_data_ptr" using \set |h \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r = set to \ set (concat disc_tree_orders)\ using \set |h' \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r = set to' \ set (concat disc_tree_orders')\ using \set to = set to'\ using \set (concat disc_tree_orders') = {cast new_character_data_ptr} \ set (concat disc_tree_orders)\ using \x \ set |h' \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r\ \x \ set |h \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r\ by auto qed have "object_ptr_kinds h' = object_ptr_kinds h |\| {|cast new_character_data_ptr|}" using object_ptr_kinds_eq_h object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 by auto then show ?thesis apply(auto simp add: is_strongly_scdom_component_safe_def Let_def)[1] apply(rule bexI[where x="cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr"]) using \{cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr} \ set |h \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r = set |h' \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r\ apply auto[2] using \set to = set to'\ \set |h \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r = set to \ set (concat disc_tree_orders)\ local.to_tree_order_ptr_in_result to' apply auto[1] using \document_ptr |\| document_ptr_kinds h\ apply blast apply(rule bexI[where x="cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr"]) using \result = new_character_data_ptr\ \{cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr} \ set |h \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r = set |h' \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r\ apply auto[1] apply(auto)[1] using \set to = set to'\ \set |h \ local.a_get_scdom_component (cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr)|\<^sub>r = set to \ set (concat disc_tree_orders)\ local.to_tree_order_ptr_in_result to' apply auto[1] apply (simp add: \document_ptr |\| document_ptr_kinds h\) using \\thesis. (\new_character_data_ptr h2 h3 disc_nodes_h3. \h \ new_character_data \\<^sub>r new_character_data_ptr; h \ new_character_data \\<^sub>h h2; h2 \ set_val new_character_data_ptr text \\<^sub>h h3; h3 \ get_disconnected_nodes document_ptr \\<^sub>r disc_nodes_h3; h3 \ set_disconnected_nodes document_ptr (cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr # disc_nodes_h3) \\<^sub>h h'\ \ thesis) \ thesis\ new_character_data_ptr new_character_data_ptr_not_in_heap apply auto[1] using create_character_data_is_strongly_dom_component_safe_step by (smt ObjectMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ \result = new_character_data_ptr\ assms(1) assms(2) assms(3) assms(4) assms(5) local.get_scdom_component_impl select_result_I2) qed end interpretation i_get_scdom_component_create_character_data?: l_get_scdom_component_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name set_val set_val_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs create_character_data by(auto simp add: l_get_scdom_component_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_scdom_component_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \create\_document\ lemma create_document_not_strongly_component_safe: obtains h :: "('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" and h' and new_document_ptr where "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" and "h \ create_document \\<^sub>r new_document_ptr \\<^sub>h h'" and "\ is_strongly_scdom_component_safe {} {cast new_document_ptr} h h'" proof - let ?h0 = "Heap fmempty ::('object_ptr::{equal,linorder}, 'node_ptr::{equal,linorder}, 'element_ptr::{equal,linorder}, 'character_data_ptr::{equal,linorder}, 'document_ptr::{equal,linorder}, 'shadow_root_ptr::{equal,linorder}, 'Object::{equal,linorder}, 'Node::{equal,linorder}, 'Element::{equal,linorder}, 'CharacterData::{equal,linorder}, 'Document::{equal,linorder}) heap" let ?P = "create_document" let ?h1 = "|?h0 \ ?P|\<^sub>h" let ?document_ptr = "|?h0 \ ?P|\<^sub>r" show thesis apply(rule that[where h="?h1"]) by code_simp+ qed locale l_get_scdom_component_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_dom_component_get_scdom_component + l_get_dom_component_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma create_document_is_weakly_scdom_component_safe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ create_document \\<^sub>r result" assumes "h \ create_document \\<^sub>h h'" shows "is_weakly_scdom_component_safe {} {cast result} h h'" proof - have "object_ptr_kinds h' = {|cast result|} |\| object_ptr_kinds h" using assms(4) assms(5) local.create_document_def new_document_new_ptr by blast have "result |\| document_ptr_kinds h" using assms(4) assms(5) local.create_document_def new_document_ptr_not_in_heap by auto show ?thesis using assms apply(auto simp add: is_weakly_scdom_component_safe_def Let_def)[1] using \object_ptr_kinds h' = {|cast result|} |\| object_ptr_kinds h\ apply(auto)[1] apply (simp add: local.create_document_def new_document_ptr_in_heap) using \result |\| document_ptr_kinds h\ apply auto[1] apply (metis (no_types, lifting) \result |\| document_ptr_kinds h\ document_ptr_kinds_commutes local.create_document_is_weakly_dom_component_safe_step select_result_I2) done qed end interpretation i_get_scdom_component_create_document?: l_get_scdom_component_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_owner_document heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name create_document get_disconnected_nodes get_disconnected_nodes_locs by(auto simp add: l_get_scdom_component_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_scdom_component_create_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \insert\_before\ locale l_get_dom_component_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_owner_document_wf + l_get_dom_component_get_scdom_component + l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_insert_before_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + 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_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma insert_before_is_component_unsafe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ insert_before ptr' child ref \\<^sub>h h'" assumes "ptr \ set |h \ get_dom_component ptr'|\<^sub>r" assumes "ptr \ set |h \ get_dom_component (cast child)|\<^sub>r" assumes "ptr \ set |h \ get_dom_component (cast |h \ get_owner_document ptr'|\<^sub>r)|\<^sub>r" assumes "ptr \ set |h \ get_dom_component (cast |h \ get_owner_document (cast child)|\<^sub>r)|\<^sub>r" shows "preserved (get_M ptr getter) h h'" proof - obtain owner_document where owner_document: "h \ get_owner_document (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child) \\<^sub>r owner_document" using assms(4) by(auto simp add: local.adopt_node_def insert_before_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF ensure_pre_insertion_validity_pure, rotated] bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated] bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated] split: if_splits) then obtain c where "h \ get_dom_component (cast owner_document) \\<^sub>r c" using get_dom_component_ok assms(1) assms(2) assms(3) get_owner_document_owner_document_in_heap by (meson document_ptr_kinds_commutes select_result_I) then have "ptr \ cast owner_document" using assms(6) assms(1) assms(2) assms(3) local.get_dom_component_ptr owner_document by (metis (no_types, lifting) assms(8) select_result_I2) obtain owner_document' where owner_document': "h \ get_owner_document ptr' \\<^sub>r owner_document'" using assms(4) by(auto simp add: local.adopt_node_def insert_before_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF ensure_pre_insertion_validity_pure, rotated] bind_returns_heap_E2[rotated, OF get_owner_document_pure, rotated] bind_returns_heap_E2[rotated, OF next_sibling_pure, rotated] split: if_splits) then obtain c where "h \ get_dom_component (cast owner_document') \\<^sub>r c" using get_dom_component_ok assms(1) assms(2) assms(3) get_owner_document_owner_document_in_heap by (meson document_ptr_kinds_commutes select_result_I) then have "ptr \ cast owner_document'" using assms(1) assms(2) assms(3) assms(7) local.get_dom_component_ptr owner_document' by auto then have "ptr \ cast |h \ get_owner_document ptr'|\<^sub>r" using owner_document' by auto have "ptr \ ptr'" by (metis (mono_tags, opaque_lifting) assms(1) assms(2) assms(3) assms(5) assms(7) is_OK_returns_result_I l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_dom_component_ok l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_dom_component_ptr l_get_owner_document.get_owner_document_ptr_in_heap local.l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms local.l_get_owner_document_axioms owner_document' return_returns_result returns_result_select_result) have "\parent. h \ get_parent child \\<^sub>r Some parent \ parent \ ptr" by (meson assms(1) assms(2) assms(3) assms(6) l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_dom_component_ptr local.get_dom_component_ok local.get_dom_component_to_tree_order local.get_parent_parent_in_heap local.l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms local.to_tree_order_ok local.to_tree_order_parent local.to_tree_order_ptr_in_result local.to_tree_order_ptrs_in_heap returns_result_select_result) then have "\parent. |h \ get_parent child|\<^sub>r = Some parent \ parent \ ptr" by (metis assms(2) assms(3) assms(4) is_OK_returns_heap_I local.get_parent_ok local.insert_before_child_in_heap select_result_I) show ?thesis using insert_before_writes assms(4) apply(rule reads_writes_preserved2) apply(auto simp add: insert_before_locs_def adopt_node_locs_def all_args_def)[1] apply(auto simp add: remove_child_locs_def set_child_nodes_locs_def set_disconnected_nodes_locs_def all_args_def split: if_splits)[1] apply (metis \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r |h \ get_owner_document ptr'|\<^sub>r\ get_M_Mdocument_preserved3) apply(auto simp add: remove_child_locs_def set_child_nodes_locs_def set_disconnected_nodes_locs_def all_args_def split: if_splits)[1] apply (metis (no_types, lifting) \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\ get_M_Mdocument_preserved3 owner_document select_result_I2) apply(auto simp add: remove_child_locs_def set_child_nodes_locs_def set_disconnected_nodes_locs_def all_args_def split: if_splits)[1] apply (metis \ptr \ ptr'\ document_ptr_casts_commute3 get_M_Mdocument_preserved3) apply(auto split: option.splits)[1] apply (metis \ptr \ ptr'\ element_ptr_casts_commute3 get_M_Element_preserved8) apply(auto simp add: remove_child_locs_def set_child_nodes_locs_def set_disconnected_nodes_locs_def all_args_def split: if_splits)[1] apply (metis \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r |h \ get_owner_document ptr'|\<^sub>r\ get_M_Mdocument_preserved3) apply(auto simp add: remove_child_locs_def set_child_nodes_locs_def set_disconnected_nodes_locs_def all_args_def split: if_splits)[1] apply (metis (no_types, lifting) \\parent. |h \ get_parent child|\<^sub>r = Some parent \ parent \ ptr\ element_ptr_casts_commute3 get_M_Element_preserved8 node_ptr_casts_commute option.case_eq_if option.collapse) apply (metis (no_types, lifting) \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\ get_M_Mdocument_preserved3 owner_document select_result_I2) apply (metis \\parent. |h \ get_parent child|\<^sub>r = Some parent \ parent \ ptr\ document_ptr_casts_commute3 get_M_Mdocument_preserved3) apply (metis (no_types, lifting) \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\ get_M_Mdocument_preserved3 owner_document select_result_I2) apply (metis (no_types, lifting) \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\ get_M_Mdocument_preserved3 owner_document select_result_I2) apply(auto simp add: remove_child_locs_def set_child_nodes_locs_def set_disconnected_nodes_locs_def all_args_def split: if_splits)[1] apply (metis \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r |h \ get_owner_document ptr'|\<^sub>r\ get_M_Mdocument_preserved3) apply(auto simp add: remove_child_locs_def set_child_nodes_locs_def set_disconnected_nodes_locs_def all_args_def split: if_splits)[1] apply (metis (no_types, lifting) \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r owner_document\ get_M_Mdocument_preserved3 owner_document select_result_I2) apply(auto simp add: remove_child_locs_def set_child_nodes_locs_def set_disconnected_nodes_locs_def all_args_def split: if_splits)[1] apply (metis \ptr \ ptr'\ document_ptr_casts_commute3 get_M_Mdocument_preserved3) apply (metis (no_types, lifting) \ptr \ ptr'\ element_ptr_casts_commute3 get_M_Element_preserved8 node_ptr_casts_commute option.case_eq_if option.collapse) apply(auto simp add: remove_child_locs_def set_child_nodes_locs_def set_disconnected_nodes_locs_def all_args_def split: if_splits)[1] by (metis \ptr \ cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r |h \ get_owner_document ptr'|\<^sub>r\ get_M_Mdocument_preserved3) qed lemma insert_before_is_strongly_dom_component_safe_step: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ insert_before ptr' child ref \\<^sub>h h'" assumes "ptr \ set |h \ get_scdom_component ptr'|\<^sub>r" assumes "ptr \ set |h \ get_scdom_component (cast child)|\<^sub>r" shows "preserved (get_M ptr getter) h h'" proof - have "ptr' |\| object_ptr_kinds h" by (meson assms(4) is_OK_returns_heap_I local.insert_before_ptr_in_heap) then obtain sc' where sc': "h \ get_scdom_component ptr' \\<^sub>r sc'" by (meson assms(1) assms(2) assms(3) get_scdom_component_ok is_OK_returns_result_E) moreover obtain c' where c': "h \ get_dom_component ptr' \\<^sub>r c'" by (meson \ptr' |\| object_ptr_kinds h\ assms(1) assms(2) assms(3) is_OK_returns_result_E local.get_dom_component_ok) ultimately have "set c' \ set sc'" using assms(1) assms(2) assms(3) get_scdom_component_subset_get_dom_component by blast have "child |\| node_ptr_kinds h" by (meson assms(4) is_OK_returns_heap_I local.insert_before_child_in_heap) then obtain child_sc where child_sc: "h \ get_scdom_component (cast child) \\<^sub>r child_sc" by (meson assms(1) assms(2) assms(3) get_scdom_component_ok is_OK_returns_result_E node_ptr_kinds_commutes) moreover obtain child_c where child_c: "h \ get_dom_component (cast child) \\<^sub>r child_c" by (meson \child |\| node_ptr_kinds h\ assms(1) assms(2) assms(3) is_OK_returns_result_E local.get_dom_component_ok node_ptr_kinds_commutes) ultimately have "set child_c \ set child_sc" using assms(1) assms(2) assms(3) get_scdom_component_subset_get_dom_component by blast obtain ptr'_owner_document where ptr'_owner_document: "h \ get_owner_document ptr' \\<^sub>r ptr'_owner_document" by (meson \set c' \ set sc'\ assms(1) assms(2) assms(3) c' get_scdom_component_owner_document_same local.get_dom_component_ptr sc' subset_code(1)) then have "h \ get_scdom_component (cast ptr'_owner_document) \\<^sub>r sc'" by (metis (no_types, lifting) \set c' \ set sc'\ assms(1) assms(2) assms(3) c' get_scdom_component_owner_document_same get_scdom_component_ptrs_same_scope_component local.get_dom_component_ptr sc' select_result_I2 subset_code(1)) moreover obtain ptr'_owner_document_c where ptr'_owner_document_c: "h \ get_dom_component (cast ptr'_owner_document) \\<^sub>r ptr'_owner_document_c" by (meson assms(1) assms(2) assms(3) document_ptr_kinds_commutes is_OK_returns_result_E local.get_dom_component_ok local.get_owner_document_owner_document_in_heap ptr'_owner_document) ultimately have "set ptr'_owner_document_c \ set sc'" using assms(1) assms(2) assms(3) get_scdom_component_subset_get_dom_component by blast obtain child_owner_document where child_owner_document: "h \ get_owner_document (cast child) \\<^sub>r child_owner_document" by (meson \set child_c \ set child_sc\ assms(1) assms(2) assms(3) child_c child_sc get_scdom_component_owner_document_same local.get_dom_component_ptr subset_code(1)) have "child_owner_document |\| document_ptr_kinds h" using assms(1) assms(2) assms(3) child_owner_document local.get_owner_document_owner_document_in_heap by blast then have "h \ get_scdom_component (cast child_owner_document) \\<^sub>r child_sc" using get_scdom_component_ok assms(1) assms(2) assms(3) child_sc by (metis (no_types, lifting) \set child_c \ set child_sc\ child_c child_owner_document get_scdom_component_owner_document_same get_scdom_component_ptrs_same_scope_component local.get_dom_component_ptr returns_result_eq set_mp) moreover obtain child_owner_document_c where child_owner_document_c: "h \ get_dom_component (cast child_owner_document) \\<^sub>r child_owner_document_c" by (meson assms(1) assms(2) assms(3) child_owner_document document_ptr_kinds_commutes is_OK_returns_result_E local.get_dom_component_ok local.get_owner_document_owner_document_in_heap) ultimately have "set child_owner_document_c \ set child_sc" using assms(1) assms(2) assms(3) get_scdom_component_subset_get_dom_component by blast have "ptr \ set |h \ get_dom_component ptr'|\<^sub>r" using \set c' \ set sc'\ assms(5) c' sc' by auto moreover have "ptr \ set |h \ get_dom_component (cast child)|\<^sub>r" using \set child_c \ set child_sc\ assms(6) child_c child_sc by auto moreover have "ptr \ set |h \ get_dom_component (cast |h \ get_owner_document ptr'|\<^sub>r)|\<^sub>r" using \set ptr'_owner_document_c \ set sc'\ assms(5) ptr'_owner_document ptr'_owner_document_c sc' by auto moreover have "ptr \ set |h \ get_dom_component (cast |h \ get_owner_document (cast child)|\<^sub>r)|\<^sub>r" using \set child_owner_document_c \ set child_sc\ assms(6) child_owner_document child_owner_document_c child_sc by auto ultimately show ?thesis using assms insert_before_is_component_unsafe by blast qed lemma insert_before_is_strongly_dom_component_safe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ insert_before ptr node child \\<^sub>h h'" shows "is_strongly_scdom_component_safe ({ptr, cast node} \ (case child of Some ref \ {cast ref} | None \ {} )) {} h h'" proof - obtain ancestors reference_child owner_document h2 h3 disconnected_nodes_h2 where ancestors: "h \ get_ancestors ptr \\<^sub>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) \\<^sub>r reference_child" and owner_document: "h \ get_owner_document ptr \\<^sub>r owner_document" and h2: "h \ adopt_node owner_document node \\<^sub>h h2" and disconnected_nodes_h2: "h2 \ get_disconnected_nodes owner_document \\<^sub>r disconnected_nodes_h2" and h3: "h2 \ set_disconnected_nodes owner_document (remove1 node disconnected_nodes_h2) \\<^sub>h h3" and h': "h3 \ a_insert_node ptr node reference_child \\<^sub>h h'" using assms(4) 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 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 \\<^sub>r ptrs = h2 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs ) then have object_ptr_kinds_M_eq2_h: "|h \ object_ptr_kinds_M|\<^sub>r = |h2 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h: "|h \ node_ptr_kinds_M|\<^sub>r = |h2 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have "known_ptrs h2" using assms(3) object_ptr_kinds_M_eq3_h known_ptrs_preserved by blast have wellformed_h2: "heap_is_wellformed h2" using adopt_node_preserves_wellformedness[OF assms(1) h2] assms(3) assms(2) . 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 \\<^sub>r ptrs = h3 \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_M_eq2_h2: "|h2 \ object_ptr_kinds_M|\<^sub>r = |h3 \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h2: "|h2 \ node_ptr_kinds_M|\<^sub>r = |h3 \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have document_ptr_kinds_eq2_h2: "|h2 \ document_ptr_kinds_M|\<^sub>r = |h3 \ document_ptr_kinds_M|\<^sub>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 \\<^sub>r ptrs = h' \ object_ptr_kinds_M \\<^sub>r ptrs" by(simp add: object_ptr_kinds_M_defs) then have object_ptr_kinds_M_eq2_h3: "|h3 \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" by simp then have node_ptr_kinds_eq2_h3: "|h3 \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" using node_ptr_kinds_M_eq by blast have document_ptr_kinds_eq2_h3: "|h3 \ document_ptr_kinds_M|\<^sub>r = |h' \ document_ptr_kinds_M|\<^sub>r" using object_ptr_kinds_M_eq2_h3 document_ptr_kinds_M_eq by auto have "object_ptr_kinds h = object_ptr_kinds h'" by (simp add: object_ptr_kinds_M_eq3_h object_ptr_kinds_M_eq3_h' object_ptr_kinds_M_eq3_h2) then show ?thesis using assms apply(auto simp add: is_strongly_scdom_component_safe_def)[1] using insert_before_is_strongly_dom_component_safe_step local.get_scdom_component_impl by blast qed lemma append_child_is_strongly_dom_component_safe_step: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ append_child ptr' child \\<^sub>h h'" assumes "ptr \ set |h \ get_scdom_component ptr'|\<^sub>r" assumes "ptr \ set |h \ get_scdom_component (cast child)|\<^sub>r" shows "preserved (get_M ptr getter) h h'" by (metis assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) insert_before_is_strongly_dom_component_safe_step local.append_child_def) lemma append_child_is_strongly_dom_component_safe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ append_child ptr child \\<^sub>h h'" shows "is_strongly_scdom_component_safe {ptr, cast child} {} h h'" using assms unfolding append_child_def using insert_before_is_strongly_dom_component_safe by fastforce end interpretation i_get_dom_component_insert_before?: l_get_dom_component_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M heap_is_wellformed parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent get_parent_locs get_child_nodes get_child_nodes_locs get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_root_node get_root_node_locs get_ancestors get_ancestors_locs get_disconnected_nodes get_disconnected_nodes_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name set_child_nodes set_child_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs get_owner_document remove_child remove_child_locs remove adopt_node adopt_node_locs insert_before insert_before_locs append_child get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe by(auto simp add: l_get_dom_component_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_dom_component_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_owner\_document\ locale l_get_owner_document_scope_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_scdom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_owner_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_dom_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_dom_component_get_scdom_component + l_get_owner_document_wf_get_root_node_wf begin lemma get_owner_document_is_strongly_scdom_component_safe_step: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_scdom_component ptr \\<^sub>r sc" assumes "h \ get_owner_document ptr' \\<^sub>r owner_document" shows "cast owner_document \ set sc \ ptr' \ set sc" proof - have "h \ get_owner_document (cast owner_document) \\<^sub>r owner_document" by (metis assms(1) assms(2) assms(3) assms(5) cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_inject document_ptr_casts_commute3 document_ptr_document_ptr_cast document_ptr_kinds_commutes local.get_owner_document_owner_document_in_heap local.get_root_node_document local.get_root_node_not_node_same node_ptr_no_document_ptr_cast) then show ?thesis using assms using bind_returns_result_E contra_subsetD get_scdom_component_ok get_scdom_component_ptrs_same_scope_component get_scdom_component_subset_get_dom_component is_OK_returns_result_E is_OK_returns_result_I local.get_dom_component_ok local.get_dom_component_ptr local.get_owner_document_ptr_in_heap local.get_owner_document_pure local.get_scdom_component_def pure_returns_heap_eq returns_result_eq by (smt local.get_scdom_component_ptrs_same_owner_document subsetD) qed lemma get_owner_document_is_strongly_scdom_component_safe: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_owner_document ptr \\<^sub>r owner_document" assumes "h \ get_owner_document ptr \\<^sub>h h'" shows "is_strongly_scdom_component_safe {ptr} {cast owner_document} h h'" proof - have "h = h'" by (meson assms(5) local.get_owner_document_pure pure_returns_heap_eq) then show ?thesis using assms apply(auto simp add: is_strongly_scdom_component_safe_def Let_def preserved_def)[1] by (smt get_owner_document_is_strongly_scdom_component_safe_step inf.orderE is_OK_returns_result_I local.get_dom_component_ok local.get_dom_component_to_tree_order_subset local.get_owner_document_ptr_in_heap local.get_scdom_component_impl local.get_scdom_component_ok local.get_scdom_component_ptr_in_heap local.get_scdom_component_subset_get_dom_component local.to_tree_order_ok local.to_tree_order_ptr_in_result notin_fset returns_result_select_result subset_eq) qed end interpretation i_get_owner_document_scope_component?: l_get_owner_document_scope_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_scdom_component is_strongly_scdom_component_safe is_weakly_scdom_component_safe get_owner_document get_disconnected_nodes get_disconnected_nodes_locs to_tree_order known_ptr known_ptrs type_wf heap_is_wellformed parent_child_rel get_child_nodes get_child_nodes_locs get_parent get_parent_locs get_ancestors get_ancestors_locs get_root_node get_root_node_locs get_dom_component is_strongly_dom_component_safe is_weakly_dom_component_safe get_element_by_id get_elements_by_class_name get_elements_by_tag_name by(auto simp add: l_get_owner_document_scope_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_owner_document_scope_component\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] end diff --git a/thys/Virtual_Substitution/EliminateVariable.thy b/thys/Virtual_Substitution/EliminateVariable.thy --- a/thys/Virtual_Substitution/EliminateVariable.thy +++ b/thys/Virtual_Substitution/EliminateVariable.thy @@ -1,683 +1,683 @@ subsection "Lemmas of the elimVar function" theory EliminateVariable imports LinearCase QuadraticCase "HOL-Library.Quadratic_Discriminant" begin lemma elimVar_eq : assumes hlength : "length xs = var" assumes in_list : "Eq p \ set(L)" assumes low_pow : "MPoly_Type.degree p var = 1 \ MPoly_Type.degree p var = 2" shows "((\x. eval (list_conj (map fm.Atom L @ F)) (xs @ x # \)) = ((\x. eval (elimVar var L F (Eq p)) (xs @ x # \)))\ (\x. aEval (Eq p) (xs @ x # \)))" proof- { fix x define A where "A = (isolate_variable_sparse p var 2)" define B where "B = (isolate_variable_sparse p var 1)" define C where "C = (isolate_variable_sparse p var 0)" have freeA : "var \ vars A" unfolding A_def by (simp add: not_in_isovarspar) have freeB : "var \ vars B" unfolding B_def by (simp add: not_in_isovarspar) have freeC : "var \ vars C" unfolding C_def by (simp add: not_in_isovarspar) assume "eval (list_conj (map fm.Atom L @ F)) (xs @ x # \)" then have h : "(\a\set L. aEval a (xs @ x # \)) \ (\f\set F. eval f (xs @ x # \))" apply(simp add:eval_list_conj) by (meson Un_iff eval.simps(1) image_eqI) define X where "X=xs@x#\" have Xlength : "length X > var" using X_def hlength by auto define Aval where "Aval = insertion (nth_default 0 (list_update X var x)) A" define Bval where "Bval = insertion (nth_default 0 (list_update X var x)) B" define Cval where "Cval = insertion (nth_default 0 (list_update X var x)) C" have hinsert : "(xs @ x # \)[var := x] = (xs @ x #\)" using hlength by auto have allAval : "\x. insertion (nth_default 0 (xs @ x # \)) A = Aval" using Aval_def using not_contains_insertion[where var="var", where p = "A", OF freeA, where L = "xs @ x #\", where x="x", where val="Aval"] unfolding X_def hinsert using hlength by auto have allBval : "\x. insertion (nth_default 0 (xs @ x # \)) B = Bval" using Bval_def using not_contains_insertion[where var="var", where p = "B", OF freeB, where L = "xs @ x #\", where x="x", where val="Bval"] unfolding X_def hinsert using hlength by auto have allCval : "\x. insertion (nth_default 0 (xs @ x # \)) C = Cval" using Cval_def using not_contains_insertion[where var="var", where p = "C", OF freeC, where L = "xs @ x #\", where x="x", where val="Cval"] unfolding X_def hinsert using hlength by auto have insertion_p : "insertion (nth_default 0 X) p = 0" using in_list h aEval.simps(1) X_def by fastforce have express_p : "p = A * Var var ^ 2 + B * Var var + C" using express_poly[OF low_pow] unfolding A_def B_def C_def by fastforce have insertion_p' : "Aval *x^2+Bval *x+Cval = 0" using express_p insertion_p unfolding Aval_def Bval_def Cval_def X_def hinsert apply(simp add: insertion_add insertion_mult insertion_pow) using insertion_var by (metis X_def Xlength hinsert) have biglemma : " ((Aval = 0 \ Bval \ 0 \ (\f\set L. aEval (linear_substitution var (-C) B f) (xs @ x # \)) \ (\f\set F. eval (linear_substitution_fm var (-C) B f) (xs @ x # \)) \ Aval \ 0 \ insertion (nth_default 0 (xs @ x # \)) 4 * Aval * Cval \ (Bval)\<^sup>2 \ ((\f\set L. eval (quadratic_sub var (- B) 1 (B\<^sup>2 - 4 * A * C) (2 * A) f) (xs @ x # \))\ (\f\set F. eval (quadratic_sub_fm var (- B) 1 (B\<^sup>2 - 4 * A * C) (2 * A) f) (xs @ x # \)) \ (\f\set L. eval (quadratic_sub var (- B) (-1) (B\<^sup>2 - 4 * A * C) (2 * A) f) (xs @ x # \)) \ (\f\set F. eval (quadratic_sub_fm var (- B) (-1) (B\<^sup>2 - 4 * A * C) (2 * A) f) (xs @ x # \))) \ Aval = 0 \ Bval = 0 \ Cval = 0))" proof(cases "Aval=0") case True then have aval0 : "Aval=0" by simp show ?thesis proof(cases "Bval=0") case True then have bval0 : "Bval=0" by simp have h : "eval (list_conj (map fm.Atom L @ F)) (xs @ x # \)" using hlength h unfolding X_def using \eval (list_conj (map fm.Atom L @ F)) (xs @ x # \)\ by blast show ?thesis proof(cases "Cval=0") case True show ?thesis by(simp add:aval0 True bval0) next case False show ?thesis using insertion_p' aval0 bval0 False by(simp) qed next case False have bh : "insertion (nth_default 0 (X[var := - Cval / Bval])) B = Bval" using allBval unfolding X_def using Bval_def X_def freeB not_contains_insertion by blast have ch : "insertion (nth_default 0 (X[var := - Cval / Bval])) C = Cval" using allCval unfolding X_def using Cval_def X_def freeC not_contains_insertion by blast have xh : "x=-Cval/Bval" proof- have "Bval*x+Cval = 0" using insertion_p' aval0 by simp then show ?thesis using False by (smt nonzero_mult_div_cancel_left) qed have freecneg : "var \ vars (-C)" using freeC not_in_neg by auto have h1: "(\a\set L. aEval (linear_substitution var (-C) (B) a) (X[var := x]))" using h xh Bval_def Cval_def False LinearCase.linear[OF Xlength False freecneg freeB, of "-Cval"] freeB freeC freecneg by (metis X_def hinsert insertion_neg) have h2 : "\f\set F. eval (linear_substitution_fm var (-C) B f) (X[var := x])" using h xh Bval_def Cval_def False LinearCase.linear_fm[OF Xlength False freecneg freeB, of "-Cval"] freeB freeC by (metis X_def hinsert insertion_neg) show ?thesis using h1 h2 apply(simp add:aval0 False) using X_def hlength using hinsert by auto qed next case False then have aval0 : "Aval \0" by simp have h4 : "insertion (nth_default 0 (X[var := x])) 4 = 4" using insertion_const[where f = "(nth_default 0 (X[var := x]))", where c="4"] by (metis MPoly_Type.insertion_one insertion_add numeral_Bit0 one_add_one) show ?thesis proof(cases "4 * Aval * Cval \ Bval\<^sup>2") case True have h1a : "var\vars(-B)" by(simp add: freeB not_in_neg) have h1b : "var\vars(1::real mpoly)" using isolate_var_one not_in_isovarspar by blast have h1c : "var\vars(-1::real mpoly)" by(simp add: h1b not_in_neg) have h1d : "var\vars(4::real mpoly)" by (metis h1b not_in_add numeral_Bit0 one_add_one) have h1e : "var\vars(B^2-4*A*C)" by(simp add: freeB h1d freeA freeC not_in_mult not_in_pow not_in_sub) have h1f : "var\vars(2::real mpoly)" using h1b not_in_add by fastforce have h1g : "var\vars(2*A)" by(simp add: freeA h1f not_in_mult) have h1h : "freeIn var (quadratic_sub var (-B) (1) (B^2-4*A*C) (2*A) a)" using free_in_quad h1a h1b h1e h1g by blast have h1i : "freeIn var (quadratic_sub var (-B) (-1) (B^2-4*A*C) (2*A) a)" using free_in_quad h1a h1c h1e h1g by blast have h2 : "2*Aval \ 0" using aval0 by auto have h3 : "0 \ (Bval^2-4*Aval*Cval)" using True by auto have h4a : "var \ vars 4" by (metis monom_numeral notInKeys_notInVars not_in_add not_in_isovarspar not_in_pow one_add_one power.simps(1) rel_simps(76) vars_monom_keys) have h4 : "var \ vars (B^2-4*A*C)" by(simp add: h4a freeA freeB freeC not_in_pow not_in_mult not_in_sub) have h5 : "\x. insertion (nth_default 0 (list_update X var x)) (-B) = -Bval " using allBval apply(simp add: insertion_neg) by (simp add: B_def Bval_def insertion_isovarspars_free) have h6 : "\x. insertion (nth_default 0 (list_update X var x)) 1 = 1" by simp have h6a : "\x. insertion (nth_default 0 (list_update X var x)) (-1) = (-1)" using h6 by (simp add: insertion_neg) have h7a : "\x. insertion (nth_default 0 (list_update X var x)) 4 = 4" by (metis h6 insertion_add numeral_Bit0 one_add_one) have h7b : "var \ vars(4*A*C)" using freeA freeC by (simp add: h4a not_in_mult) have h7c : "var \ vars(B^2)" using freeB not_in_pow by auto have h7 : "\x. insertion (nth_default 0 (list_update X var x)) (B^2-4*A*C) = (Bval^2-4*Aval*Cval)" using h7a allAval allBval allCval unfolding X_def using hlength apply (simp add: insertion_mult insertion_sub power2_eq_square) by (metis A_def Aval_def Bval_def C_def Cval_def X_def freeB insertion_isovarspars_free not_contains_insertion) have h8a : "\x. insertion (nth_default 0 (list_update X var x)) 2 = 2" by (metis h6 insertion_add one_add_one) have h8 : "\x. insertion (nth_default 0 (list_update X var x)) (2*A) = (2*Aval)" apply(simp add: allAval h8a insertion_mult) by (simp add: A_def Aval_def insertion_isovarspars_free) have h1 : "- Bval\<^sup>2 + 4 * Aval * Cval \ 0" using True by simp have xh : "x = (- Bval + sqrt (Bval\<^sup>2 - 4 * Aval * Cval)) / (2 * Aval)\x=(- Bval - sqrt (Bval\<^sup>2 - 4 * Aval * Cval)) / (2 * Aval)" using insertion_p' aval0 h1 discriminant_iff unfolding discrim_def by blast have p1 : "x = (- Bval + sqrt (Bval\<^sup>2 - 4 * Aval * Cval)) / (2 * Aval) \ ((\a\ set L. eval (quadratic_sub var (-B) 1 (B^2-4*A*C) (2*A) a) (X[var := x])) \(\a\ set F. eval (quadratic_sub_fm var (-B) 1 (B^2-4*A*C) (2*A) a) (X[var := x])))" proof- assume x_def : "x = (- Bval + sqrt (Bval\<^sup>2 - 4 * Aval * Cval)) / (2 * Aval)" then have h : "(\a\set L. aEval a (X[var := (- Bval + sqrt (Bval\<^sup>2 - 4 * Aval * Cval)) / (2 * Aval)])) \ (\f\set F. eval f (X[var := (- Bval + sqrt (Bval\<^sup>2 - 4 * Aval * Cval)) / (2 * Aval)]))" using h using X_def hinsert by auto { fix a assume in_list : "a\ set L" have "eval (quadratic_sub var (- B) 1 (B\<^sup>2 - 4 * A * C) (2 * A) a) (X[var := x])" using free_in_quad[where a="-B",where b="1", where c="(B^2-4*A*C)", where d="2*A",where var="var",OF h1a h1b h1e h1g] using quadratic_sub[where a="-B",where b="1", where c="(B^2-4*A*C)", where d="2*A",where var="var", where L="X", OF Xlength, where Dv="2*Aval", OF h2, where Cv="(Bval^2-4*Aval*Cval)", OF h3, where Av="-Bval", OF h4 h5, where Bv="1", OF h6 h7 h8] h in_list using var_not_in_eval by fastforce } then have left : "(\a\set L. eval (quadratic_sub var (- B) 1 (B\<^sup>2 - 4 * A * C) (2 * A) a) (X[var := x]))" by simp { fix a assume in_list : "a\ set F" have "eval (quadratic_sub_fm var (- B) 1 (B\<^sup>2 - 4 * A * C) (2 * A) a) (X[var := x])" using free_in_quad_fm[where a="-B",where b="1", where c="(B^2-4*A*C)", where d="2*A",where var="var",OF h1a h1b h1e h1g] using quadratic_sub_fm[where a="-B",where b="1", where c="(B^2-4*A*C)", where d="2*A",where var="var", where L="X", OF Xlength, where Dv="2*Aval", OF h2, where Cv="(Bval^2-4*Aval*Cval)", OF h3, where Av="-Bval", OF h4 h5, where Bv="1", OF h6 h7 h8] h in_list using var_not_in_eval by fastforce } then have right : "(\a\set F. eval (quadratic_sub_fm var (- B) 1 (B\<^sup>2 - 4 * A * C) (2 * A) a) (X[var := x]))" by simp show ?thesis using right left by simp qed have p2 : "x = (- Bval - sqrt (Bval\<^sup>2 - 4 * Aval * Cval)) / (2 * Aval) \ ((\a\ set L. eval (quadratic_sub var (-B) (-1) (B^2-4*A*C) (2*A) a) (X[var := x])) \(\a\ set F. eval (quadratic_sub_fm var (-B) (-1) (B^2-4*A*C) (2*A) a) (X[var := x])))" proof - assume x_def : "x = (- Bval - sqrt (Bval\<^sup>2 - 4 * Aval * Cval)) / (2 * Aval)" then have h : "(\a\set L. aEval a (X[var := (- Bval - sqrt (Bval\<^sup>2 - 4 * Aval * Cval)) / (2 * Aval)])) \ (\f\set F. eval f (X[var := (- Bval - sqrt (Bval\<^sup>2 - 4 * Aval * Cval)) / (2 * Aval)]))" using h using X_def hinsert by auto then have "(\a\set L. aEval a (X[var := (- Bval - sqrt (Bval\<^sup>2 - 4 * Aval * Cval)) / (2 * Aval)])) \ (\f\set F. eval f (X[var := (- Bval - sqrt (Bval\<^sup>2 - 4 * Aval * Cval)) / (2 * Aval)]))" using h by simp { fix a assume in_list : "a\ set L" have "eval (quadratic_sub var (- B) (-1) (B\<^sup>2 - 4 * A * C) (2 * A) a) (X[var := x])" using free_in_quad[where a="-B",where b="-1", where c="(B^2-4*A*C)", where d="2*A",where var="var",OF h1a h1c h1e h1g] using quadratic_sub[where a="-B",where b="-1", where c="(B^2-4*A*C)", where d="2*A",where var="var", where L="X", OF Xlength, where Dv="2*Aval", OF h2, where Cv="(Bval^2-4*Aval*Cval)", OF h3, where Av="-Bval", OF h4 h5, where Bv="-1", OF h6a h7 h8] h in_list using var_not_in_eval by fastforce } then have left : "(\a\set L. eval (quadratic_sub var (- B) (-1) (B\<^sup>2 - 4 * A * C) (2 * A) a) (X[var := x]))" by simp { fix a assume in_list : "a\ set F" have "eval (quadratic_sub_fm var (- B) (-1) (B\<^sup>2 - 4 * A * C) (2 * A) a) (X[var := x])" using free_in_quad_fm[where a="-B",where b="-1", where c="(B^2-4*A*C)", where d="2*A",where var="var",OF h1a h1c h1e h1g] using quadratic_sub_fm[where a="-B",where b="-1", where c="(B^2-4*A*C)", where d="2*A",where var="var", where L="X", OF Xlength, where Dv="2*Aval", OF h2, where Cv="(Bval^2-4*Aval*Cval)", OF h3, where Av="-Bval", OF h4 h5, where Bv="-1", OF h6a h7 h8] h in_list using var_not_in_eval by fastforce } then have right : "(\a\set F. eval (quadratic_sub_fm var (- B) (-1) (B\<^sup>2 - 4 * A * C) (2 * A) a) (X[var := x]))" by simp show ?thesis using right left by simp qed have subst4 : "insertion (nth_default 0 (xs @ x # \)) 4 = 4" using h7a hlength X_def by auto have disj: "(\a\set L. eval (quadratic_sub var (- B) 1 (B\<^sup>2 - 4 * A * C) (2 * A) a) (xs @ x # \)) \ (\a\set F. eval (quadratic_sub_fm var (- B) 1 (B\<^sup>2 - 4 * A * C) (2 * A) a) (xs @ x # \)) \ (\a\set L. eval (quadratic_sub var (- B) (-1) (B\<^sup>2 - 4 * A * C) (2 * A) a) (xs @ x # \)) \ (\a\set F. eval (quadratic_sub_fm var (- B) (-1) (B\<^sup>2 - 4 * A * C) (2 * A) a) (xs @ x # \))" using xh p1 p2 unfolding X_def hinsert by blast show ?thesis apply(simp add: aval0 True h7a subst4) using disj unfolding X_def hinsert by auto next case False then have det : "0 < - Bval\<^sup>2 + 4 * Aval * Cval" by simp show ?thesis apply(simp add: aval0 False h4) using discriminant_negative unfolding discrim_def using insertion_p' using aval0 det by auto qed qed have "(\x. (insertion (nth_default 0 (xs @ x # \)) A = 0 \ insertion (nth_default 0 (xs @ x # \)) B \ 0 \ (\f\set L. aEval (linear_substitution var (-C) (B) f) (xs @ x # \)) \ (\f\set F. eval (linear_substitution_fm var (-C) B f) (xs @ x # \)) \ insertion (nth_default 0 (xs @ x # \)) A \ 0 \ insertion (nth_default 0 (xs @ x # \)) 4 * insertion (nth_default 0 (xs @ x # \)) A * insertion (nth_default 0 (xs @ x # \)) C \ (insertion (nth_default 0 (xs @ x # \)) B)\<^sup>2 \ ((\f\set L. eval (quadratic_sub var (- B) 1 (B\<^sup>2 - 4 * A * C) (2 * A) f) (xs @ x # \))\ (\f\set F. eval (quadratic_sub_fm var (- B) 1 (B\<^sup>2 - 4 * A * C) (2 * A) f) (xs @ x # \)) \ (\f\set L. eval (quadratic_sub var (- B) (-1) (B\<^sup>2 - 4 * A * C) (2 * A) f) (xs @ x # \)) \ (\f\set F. eval (quadratic_sub_fm var (- B) (-1) (B\<^sup>2 - 4 * A * C) (2 * A) f) (xs @ x # \)))) \ (Aval = 0 \ Bval = 0 \ Cval = 0))" apply(rule exI[where x=x]) using biglemma using allAval allBval allCval unfolding A_def B_def C_def Aval_def Bval_def Cval_def X_def hinsert by auto then obtain x where x : "(insertion (nth_default 0 (xs @ x # \)) A = 0 \ insertion (nth_default 0 (xs @ x # \)) B \ 0 \ (\f\set L. aEval (linear_substitution var (-C) (B) f) (xs @ x # \)) \ (\f\set F. eval (linear_substitution_fm var (-C) B f) (xs @ x # \)) \ insertion (nth_default 0 (xs @ x # \)) A \ 0 \ insertion (nth_default 0 (xs @ x # \)) 4 * insertion (nth_default 0 (xs @ x # \)) A * insertion (nth_default 0 (xs @ x # \)) C \ (insertion (nth_default 0 (xs @ x # \)) B)\<^sup>2 \ ((\f\set L. eval (quadratic_sub var (- B) 1 (B\<^sup>2 - 4 * A * C) (2 * A) f) (xs @ x # \))\ (\f\set F. eval (quadratic_sub_fm var (- B) 1 (B\<^sup>2 - 4 * A * C) (2 * A) f) (xs @ x # \)) \ (\f\set L. eval (quadratic_sub var (- B) (-1) (B\<^sup>2 - 4 * A * C) (2 * A) f) (xs @ x # \)) \ (\f\set F. eval (quadratic_sub_fm var (- B) (-1) (B\<^sup>2 - 4 * A * C) (2 * A) f) (xs @ x # \)))) \ (Aval = 0 \ Bval = 0 \ Cval = 0)" by auto have h : "(\x. eval (elimVar var L F (Eq p)) (xs @ x # \))\(Aval = 0 \ Bval = 0 \ Cval = 0)" proof(cases "(Aval = 0 \ Bval = 0 \ Cval = 0)") case True then show ?thesis by simp next case False have "(\x. eval (elimVar var L F (Eq p)) (xs @ x # \))" apply(rule exI[where x=x]) apply(simp add: eval_list_conj insertion_mult insertion_sub insertion_pow insertion_add del: quadratic_sub.simps linear_substitution.simps quadratic_sub_fm.simps linear_substitution_fm.simps) unfolding A_def[symmetric] B_def[symmetric] C_def[symmetric] One_nat_def[symmetric] X_def[symmetric] using hlength x by (auto simp add:False) then show ?thesis by auto qed have "(\x. eval (elimVar var L F (Eq p)) (xs @ x # \))\(\x. aEval (Eq p) (xs@ x# \))" proof(cases "(\x. eval (elimVar var L F (Eq p)) (xs @ x # \))") case True then show ?thesis by auto next case False then have "(Aval = 0 \ Bval = 0 \ Cval = 0)" using h by auto then have "(\x. aEval (Eq p) (xs @ x # \))" unfolding express_p apply(simp add:insertion_add insertion_mult insertion_pow) using allAval allBval allCval by auto then show ?thesis by auto qed } then have left : "(\x. eval (list_conj (map fm.Atom L @ F)) (xs @ x # \)) \ ((\x. eval (elimVar var L F (Eq p)) (xs @ x # \))\(\x. aEval (Eq p) (xs@ x# \)))" by blast { assume hlength : "length (xs::real list) = var" define A where "A = (isolate_variable_sparse p var 2)" define B where "B = (isolate_variable_sparse p var 1)" define C where "C = (isolate_variable_sparse p var 0)" have freeA : "var \ vars A" unfolding A_def by (simp add: not_in_isovarspar) have freeB : "var \ vars B" unfolding B_def by (simp add: not_in_isovarspar) have freeC : "var \ vars C" unfolding C_def by (simp add: not_in_isovarspar) have express_p : "p = A*(Var var)^2+B*(Var var)+C" using express_poly[OF low_pow] unfolding A_def B_def C_def by fastforce assume h : "(\x. (eval (elimVar var L F (Eq p)) (list_update (xs@x#\) var x)))" fix x define X where "X=xs@x#\" have Xlength : "length X > var" using X_def hlength by auto define Aval where "Aval = insertion (nth_default 0 (list_update X var x)) A" define Bval where "Bval = insertion (nth_default 0 (list_update X var x)) B" define Cval where "Cval = insertion (nth_default 0 (list_update X var x)) C" have allAval : "\x. insertion (nth_default 0 (list_update X var x)) A = Aval" using freeA Aval_def using not_contains_insertion by blast have allBval : "\x. insertion (nth_default 0 (list_update X var x)) B = Bval" using freeB Bval_def using not_contains_insertion by blast have allCval : "\x. insertion (nth_default 0 (list_update X var x)) C = Cval" using freeC Cval_def using not_contains_insertion by blast assume "(eval (elimVar var L F (Eq p)) (list_update (xs@x#\) var x))" then have h : "(eval (elimVar var L F (Eq p)) (list_update X var x))" unfolding X_def . have "(Aval = 0 \ Bval \ 0 \ (\f\(\a. Atom(linear_substitution var (-C) B a)) ` set L \ linear_substitution_fm var (-C) B ` set F. eval f (X[var := x])) \ Aval \ 0 \ insertion (nth_default 0 (X[var := x])) 4 * Aval * Cval \ Bval\<^sup>2 \ ((\f\(quadratic_sub var (-B) 1 (B^2-4*A*C) (2*A)) ` set L \ (quadratic_sub_fm var (-B) 1 (B^2-4*A*C) (2*A)) ` set F. eval f (X[var := x])) \(\f\(quadratic_sub var (-B) (-1) (B^2-4*A*C) (2*A)) ` set L \ (quadratic_sub_fm var (-B) (-1) (B^2-4*A*C) (2*A)) ` set F. eval f (X[var := x])) ))" unfolding Aval_def Bval_def Cval_def A_def B_def C_def using h by(simp add: eval_list_conj insertion_mult insertion_sub insertion_pow insertion_add insertion_var Xlength) then have h : "(Aval = 0 \ Bval \ 0 \ ((\a\ set L. aEval (linear_substitution var (-C) B a) (X[var := x])) \ (\a\ set F. eval (linear_substitution_fm var (-C) B a) (X[var := x]))) \ Aval \ 0 \ insertion (nth_default 0 (X[var := x])) 4 * Aval * Cval \ Bval\<^sup>2 \ (((\a\ set L. eval (quadratic_sub var (-B) 1 (B^2-4*A*C) (2*A) a) (X[var := x])) \(\a\ set F. eval (quadratic_sub_fm var (-B) 1 (B^2-4*A*C) (2*A) a) (X[var := x]))) \((\a\ set L. eval (quadratic_sub var (-B) (-1) (B^2-4*A*C) (2*A) a) (X[var := x])) \(\a\ set F. eval (quadratic_sub_fm var (-B) (-1) (B^2-4*A*C) (2*A) a) (X[var := x]))))) " apply(cases "Aval = 0 ") apply auto by (meson Un_iff eval.simps(1) imageI) have h : "(\x. ((\a\set L . aEval a ((xs@x#\)[var := x])) \ (\f\set F. eval f ((xs@x#\)[var := x]))))\(Aval=0\Bval=0\Cval=0)" proof(cases "Aval=0") case True then have aval0 : "Aval=0" by simp show ?thesis proof(cases "Bval = 0") case True then have bval0 : "Bval = 0" by simp show ?thesis proof(cases "Cval=0") case True then show ?thesis using aval0 bval0 True by auto next case False then show ?thesis using h by(simp add:aval0 bval0 False) qed next case False have hb : "insertion (nth_default 0 (X[var := - Cval / Bval])) B = Bval" using allBval by simp have hc : "insertion (nth_default 0 (X[var := - Cval / Bval])) (-C) = -Cval" using allCval by (simp add: insertion_neg) have freecneg : "var\vars(-C)" using freeC not_in_neg by auto have p1 : "(\a\set L. aEval a ((xs @ x # \)[var := - Cval / Bval]))" using h apply(simp add: False aval0) using linear[OF Xlength False freecneg freeB hc hb] list_update_length var_not_in_linear[OF freecneg freeB] unfolding X_def using hlength by (metis divide_minus_left) have p2 : "(\a\set F. eval a ((xs @ x # \)[var := - Cval / Bval]))" using h apply(simp add: False aval0) using linear_fm[OF Xlength False freecneg freeB hc hb] list_update_length var_not_in_linear_fm[OF freecneg freeB] unfolding X_def using hlength var_not_in_eval by (metis divide_minus_left linear_substitution_fm.elims linear_substitution_fm_helper.elims) show ?thesis using p1 p2 hlength by fastforce qed next case False then have aval0 : "Aval \ 0" by simp have h4 : "insertion (nth_default 0 (X[var := x])) 4 = 4" using insertion_const[where f = "(nth_default 0 (X[var := x]))", where c="4"] by (metis MPoly_Type.insertion_one insertion_add numeral_Bit0 one_add_one) show ?thesis proof(cases "4 * Aval * Cval \ Bval\<^sup>2") case True then have h1 : "- Bval\<^sup>2 + 4 * Aval * Cval \ 0" by simp have h : "(((\a\ set L. eval (quadratic_sub var (-B) 1 (B^2-4*A*C) (2*A) a) (X[var := x])) \(\a\ set F. eval (quadratic_sub_fm var (-B) 1 (B^2-4*A*C) (2*A) a) (X[var := x]))) \((\a\ set L. eval (quadratic_sub var (-B) (-1) (B^2-4*A*C) (2*A) a) (X[var := x])) \(\a\ set F. eval (quadratic_sub_fm var (-B) (-1) (B^2-4*A*C) (2*A) a) (X[var := x]))))" using h by(simp add: h1 aval0) have h1a : "var\vars(-B)" by(simp add: freeB not_in_neg) have h1b : "var\vars(1::real mpoly)" using isolate_var_one not_in_isovarspar by blast have h1c : "var\vars(-1::real mpoly)" by(simp add: h1b not_in_neg) have h1d : "var\vars(4::real mpoly)" by (metis h1b not_in_add numeral_Bit0 one_add_one) have h1e : "var\vars(B^2-4*A*C)" by(simp add: freeB h1d freeA freeC not_in_mult not_in_pow not_in_sub) have h1f : "var\vars(2::real mpoly)" using h1b not_in_add by fastforce have h1g : "var\vars(2*A)" by(simp add: freeA h1f not_in_mult) have h1h : "freeIn var (quadratic_sub var (-B) (1) (B^2-4*A*C) (2*A) a)" using free_in_quad h1a h1b h1e h1g by blast have h1i : "freeIn var (quadratic_sub var (-B) (-1) (B^2-4*A*C) (2*A) a)" using free_in_quad h1a h1c h1e h1g by blast have h2 : "2*Aval \ 0" using aval0 by auto have h3 : "0 \ (Bval^2-4*Aval*Cval)" using True by auto have h4a : "var \ vars 4" by (metis monom_numeral notInKeys_notInVars not_in_add not_in_isovarspar not_in_pow one_add_one power.simps(1) rel_simps(76) vars_monom_keys) have h4 : "var \ vars (B^2-4*A*C)" by(simp add: h4a freeA freeB freeC not_in_pow not_in_mult not_in_sub) have h5 : "\x. insertion (nth_default 0 (list_update X var x)) (-B) = -Bval " using allBval by(simp add: insertion_neg) have h6 : "\x. insertion (nth_default 0 (list_update X var x)) 1 = 1" by simp have h6a : "\x. insertion (nth_default 0 (list_update X var x)) (-1) = (-1)" using h6 by (simp add: insertion_neg) have h7a : "\x. insertion (nth_default 0 (list_update X var x)) 4 = 4" by (metis h6 insertion_add numeral_Bit0 one_add_one) have h7b : "var \ vars(4*A*C)" using freeA freeC by (simp add: h4a not_in_mult) have h7c : "var \ vars(B^2)" using freeB not_in_pow by auto have h7 : "\x. insertion (nth_default 0 (list_update X var x)) (B^2-4*A*C) = (Bval^2-4*Aval*Cval)" by (simp add: h7a allAval allBval allCval insertion_mult insertion_sub power2_eq_square) have h8a : "\x. insertion (nth_default 0 (list_update X var x)) 2 = 2" by (metis h6 insertion_add one_add_one) have h8 : "\x. insertion (nth_default 0 (list_update X var x)) (2*A) = (2*Aval)" by(simp add: allAval h8a insertion_mult) have p1 : "(\a\ set L. eval (quadratic_sub var (-B) 1 (B^2-4*A*C) (2*A) a) (X[var := x])) \(\a\ set F. eval (quadratic_sub_fm var (-B) 1 (B^2-4*A*C) (2*A) a) (X[var := x])) \ \x. length xs = var \ ((\a\set L . aEval a ((xs@x#\)[var := x])) \ (\f\set F. eval f ((xs@x#\)[var := x])))" proof- assume p1 : "(\a\ set L. eval (quadratic_sub var (-B) 1 (B^2-4*A*C) (2*A) a) (X[var := x]))" assume p2 : "(\a\ set F. eval (quadratic_sub_fm var (-B) 1 (B^2-4*A*C) (2*A) a) (X[var := x]))" show ?thesis using free_in_quad[where a="-B",where b="1", where c="(B^2-4*A*C)", where d="2*A",where var="var",OF h1a h1b h1e h1g] using quadratic_sub[where a="-B",where b="1", where c="(B^2-4*A*C)", where d="2*A",where var="var", where L="X", OF Xlength, where Dv="2*Aval", OF h2, where Cv="(Bval^2-4*Aval*Cval)", OF h3, where Av="-Bval", OF h4 h5, where Bv="1", OF h6 h7 h8] using free_in_quad_fm[where a="-B",where b="1", where c="(B^2-4*A*C)", where d="2*A",where var="var",OF h1a h1b h1e h1g] using quadratic_sub_fm[where a="-B",where b="1", where c="(B^2-4*A*C)", where d="2*A",where var="var", where L="X", OF Xlength, where Dv="2*Aval", OF h2, where Cv="(Bval^2-4*Aval*Cval)", OF h3, where Av="-Bval", OF h4 h5, where Bv="1", OF h6 h7 h8] p1 p2 using var_not_in_eval by (metis X_def hlength list_update_length) qed have p2 : "(\a\ set L. eval (quadratic_sub var (-B) (-1) (B^2-4*A*C) (2*A) a) (X[var := x])) \(\a\ set F. eval (quadratic_sub_fm var (-B) (-1) (B^2-4*A*C) (2*A) a) (X[var := x])) \\x. length xs = var \ ((\a\set L . aEval a ((xs@x#\)[var := x])) \ (\f\set F. eval f ((xs@x#\)[var := x])))" using free_in_quad[where a="-B",where b="-1", where c="(B^2-4*A*C)", where d="2*A",where var="var",OF h1a h1c h1e h1g] using quadratic_sub[where a="-B",where b="-1", where c="(B^2-4*A*C)", where d="2*A",where var="var", where L="X", OF Xlength, where Dv="2*Aval", OF h2, where Cv="(Bval^2-4*Aval*Cval)", OF h3, where Av="-Bval", OF h4 h5, where Bv="-1", OF h6a h7 h8] using free_in_quad_fm[where a="-B",where b="-1", where c="(B^2-4*A*C)", where d="2*A",where var="var",OF h1a h1c h1e h1g] using quadratic_sub_fm[where a="-B",where b="-1", where c="(B^2-4*A*C)", where d="2*A",where var="var", where L="X", OF Xlength, where Dv="2*Aval", OF h2, where Cv="(Bval^2-4*Aval*Cval)", OF h3, where Av="-Bval", OF h4 h5, where Bv="-1", OF h6a h7 h8] using var_not_in_eval by (metis X_def hlength list_update_length) then show ?thesis using h p1 p2 by blast next case False then show ?thesis using h by(simp add: aval0 False h4) qed qed have "(\x.((\a\set L . aEval a ((xs@x#\)[var := x])) \ (\f\set F. eval f ((xs@x#\)[var := x]))))\(\x. aEval (Eq p) (xs @ x#\))" proof(cases "(\x.((\a\set L . aEval a ((xs@x#\)[var := x])) \ (\f\set F. eval f ((xs@x#\)[var := x]))))") case True then show ?thesis by auto next case False then have "Aval=0\Bval=0\Cval=0" using h by auto then have "(\x. aEval (Eq p) (xs @ x # \))" unfolding express_p apply(simp add:insertion_add insertion_mult insertion_pow) using allAval allBval allCval hlength unfolding X_def by auto then show ?thesis by auto qed } then have right : "(\x. eval (elimVar var L F (Eq p)) (xs @ x # \)) \ ((\x. eval (list_conj (map fm.Atom L @ F)) (xs @ x # \))\(\x. aEval (Eq p) (xs @ x # \)))" by (smt UnE eval.simps(1) eval_list_conj hlength imageE list_update_length set_append set_map) show ?thesis using right left by blast qed text "simply states that the variable is free in the equality case of the elimVar function" lemma freeIn_elimVar_eq : "freeIn var (elimVar var L F (Eq p))" proof- have h4 : "var \ vars(4:: real mpoly)" using var_not_in_Const by (metis (full_types) isolate_var_one monom_numeral not_in_isovarspar numeral_One vars_monom_keys zero_neq_numeral) have hlinear: "\f\set(map (\a. Atom(linear_substitution var (-isolate_variable_sparse p var 0) (isolate_variable_sparse p var (Suc 0)) a)) L @ map (linear_substitution_fm var (-isolate_variable_sparse p var 0) (isolate_variable_sparse p var (Suc 0))) F). freeIn var f" using var_not_in_linear[where c="(isolate_variable_sparse p var (Suc 0))", where b="(- isolate_variable_sparse p var 0)", where var="var"] var_not_in_linear_fm[where c="(isolate_variable_sparse p var (Suc 0))", where b="(-isolate_variable_sparse p var 0)", where var="var"] not_in_isovarspar not_in_neg by auto have freeA : "var \ vars (- isolate_variable_sparse p var (Suc 0))" using not_in_isovarspar not_in_neg by auto have freeB1 : "var \ vars (1::real mpoly)" by (metis h4 monom_numeral monom_one notInKeys_notInVars vars_monom_keys zero_neq_numeral) have freeC : "var \ vars (((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0))" using not_in_isovarspar not_in_pow not_in_sub not_in_mult h4 by auto have freeD : "var \ vars ((2 * isolate_variable_sparse p var 2))" using not_in_isovarspar not_in_mult by (metis mult_2 not_in_add) have freeB2 : "var\vars (-1::real mpoly)" using freeB1 not_in_neg by auto have quadratic1 : "\f\set(map (quadratic_sub var (- isolate_variable_sparse p var (Suc 0)) 1 ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) (2 * isolate_variable_sparse p var 2)) L @ map (quadratic_sub_fm var (- isolate_variable_sparse p var (Suc 0)) 1 ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) (2 * isolate_variable_sparse p var 2)) F). freeIn var f" using free_in_quad[OF freeA freeB1 freeC freeD] free_in_quad_fm[OF freeA freeB1 freeC freeD] by auto have quadratic2 : "\f\set(map (quadratic_sub var (- isolate_variable_sparse p var (Suc 0)) (-1) ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) (2 * isolate_variable_sparse p var 2)) L @ map (quadratic_sub_fm var (- isolate_variable_sparse p var (Suc 0)) (-1) ((isolate_variable_sparse p var (Suc 0))\<^sup>2 - 4 * isolate_variable_sparse p var 2 * isolate_variable_sparse p var 0) (2 * isolate_variable_sparse p var 2)) F). freeIn var f" using free_in_quad[OF freeA freeB2 freeC freeD] free_in_quad_fm[OF freeA freeB2 freeC freeD] by auto show ?thesis using not_in_mult not_in_add h4 not_in_pow not_in_sub freeIn_list_conj not_in_isovarspar hlinear quadratic1 quadratic2 - bysimp + by simp qed text "Theorem 20.2 in the textbook" lemma elimVar_eq_2 : assumes hlength : "length xs = var" assumes in_list : "Eq p \ set(L)" assumes low_pow : "MPoly_Type.degree p var = 1 \ MPoly_Type.degree p var = 2" assumes nonzero : "\x. insertion (nth_default 0 (xs @ x # \)) (isolate_variable_sparse p var 2) \ 0 \ insertion (nth_default 0 (xs @ x # \)) (isolate_variable_sparse p var 1) \ 0 \ insertion (nth_default 0 (xs @ x # \)) (isolate_variable_sparse p var 0) \ 0" (is ?non0) shows "(\x. eval (list_conj (map fm.Atom L @ F)) (xs @ x # \)) = (\x. eval (elimVar var L F (Eq p)) (xs @ x # \))" proof- define A where "A = (isolate_variable_sparse p var 2)" define B where "B = (isolate_variable_sparse p var 1)" define C where "C = (isolate_variable_sparse p var 0)" have freeA : "var \ vars A" unfolding A_def by (simp add: not_in_isovarspar) have freeB : "var \ vars B" unfolding B_def by (simp add: not_in_isovarspar) have freeC : "var \ vars C" unfolding C_def by (simp add: not_in_isovarspar) have express_p : "p = A*(Var var)^2+B*(Var var)+C" using express_poly[OF low_pow] unfolding A_def B_def C_def by fastforce have af : "isolate_variable_sparse p var 2 = A" using A_def by auto have bf : "isolate_variable_sparse p var (Suc 0) = B" using B_def by auto have cf : "isolate_variable_sparse p var 0 = C" using C_def by auto have xlength : "\x. (insertion (nth_default 0 (xs @ x # \)) (Var var))= x" using hlength insertion_var by (metis add.commute add_strict_increasing length_append length_greater_0_conv list.distinct(1) list_update_id nth_append_length order_refl) fix x define c where "c i = (insertion (nth_default 0 (xs @ x # \)) (isolate_variable_sparse p var i))" for i have c2 : "\x. insertion (nth_default 0 (xs @ x # \)) A = c 2" using freeA apply(simp add: A_def c_def) by (simp add: hlength insertion_lowerPoly1) have c1 : "\x. insertion (nth_default 0 (xs @ x # \)) B = c 1" using freeB apply(simp add: B_def c_def) by (simp add: hlength insertion_lowerPoly1) have c0 : "\x. insertion (nth_default 0 (xs @ x # \)) C = c 0" using freeC apply(simp add: C_def c_def) by (simp add: hlength insertion_lowerPoly1) have sum : "\x. c 2 * x\<^sup>2 + c (Suc 0) * x + c 0 = (\i\2. c i * x ^ i)" by (simp add: numerals(2)) have "(\x. aEval (Eq p) (xs @ x # \)) = (\?non0)" apply(simp add : af bf cf) unfolding express_p apply(simp add:insertion_add insertion_mult insertion_pow xlength) apply(simp add:c2 c1 c0) apply(simp add: sum) using polyfun_eq_0[where c="c", where n="2"] using sum by auto then have "\(\x. aEval (Eq p) (xs @ x \))" using nonzero by auto then show ?thesis using disjE[OF elimVar_eq[OF hlength in_list, where F="F", where \="\"], where R="?thesis"] using \(\x. aEval (Eq p) (xs @ x # \)) = (\ (\x. insertion (nth_default 0 (xs @ x # \)) (isolate_variable_sparse p var 2) \ 0 \ insertion (nth_default 0 (xs @ x # \)) (isolate_variable_sparse p var 1) \ 0 \ insertion (nth_default 0 (xs @ x # \)) (isolate_variable_sparse p var 0) \ 0))\ low_pow nonzero by blast qed end \ No newline at end of file