diff --git a/thys/Core_DOM/common/monads/CharacterDataMonad.thy b/thys/Core_DOM/common/monads/CharacterDataMonad.thy --- a/thys/Core_DOM/common/monads/CharacterDataMonad.thy +++ b/thys/Core_DOM/common/monads/CharacterDataMonad.thy @@ -1,534 +1,534 @@ (*********************************************************************************** * 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\CharacterData\ text\In this theory, we introduce the monadic method setup for the CharacterData class.\ theory CharacterDataMonad imports ElementMonad "../classes/CharacterDataClass" begin type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'result) dom_prog = "((_) heap, exception, 'result) prog" register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'result) dom_prog" global_interpretation l_ptr_kinds_M character_data_ptr_kinds defines character_data_ptr_kinds_M = a_ptr_kinds_M . lemmas character_data_ptr_kinds_M_defs = a_ptr_kinds_M_def lemma character_data_ptr_kinds_M_eq: assumes "|h \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" shows "|h \ character_data_ptr_kinds_M|\<^sub>r = |h' \ character_data_ptr_kinds_M|\<^sub>r" using assms by(auto simp add: character_data_ptr_kinds_M_defs node_ptr_kinds_M_defs character_data_ptr_kinds_def) lemma character_data_ptr_kinds_M_reads: "reads (\node_ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t node_ptr RObject.nothing)}) character_data_ptr_kinds_M h h'" using node_ptr_kinds_M_reads apply (simp add: reads_def node_ptr_kinds_M_defs character_data_ptr_kinds_M_defs character_data_ptr_kinds_def preserved_def) - by (smt node_ptr_kinds_small preserved_def unit_all_impI) + by (metis (mono_tags, lifting) node_ptr_kinds_small old.unit.exhaust preserved_def) global_interpretation l_dummy defines get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a = "l_get_M.a_get_M get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a" . lemma get_M_is_l_get_M: "l_get_M get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a type_wf character_data_ptr_kinds" apply(simp add: get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_type_wf l_get_M_def) by (metis (no_types, opaque_lifting) NodeMonad.get_M_is_l_get_M bind_eq_Some_conv character_data_ptr_kinds_commutes get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def l_get_M_def option.distinct(1)) lemmas get_M_defs = get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def[unfolded l_get_M.a_get_M_def[OF get_M_is_l_get_M]] adhoc_overloading get_M get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a locale l_get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas = l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a begin sublocale l_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas by unfold_locales interpretation l_get_M get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a type_wf character_data_ptr_kinds apply(unfold_locales) apply (simp add: get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_type_wf local.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a) by (meson CharacterDataMonad.get_M_is_l_get_M l_get_M_def) lemmas get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ok = get_M_ok[folded get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def] end global_interpretation l_get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas type_wf by unfold_locales global_interpretation l_put_M type_wf character_data_ptr_kinds get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a rewrites "a_get_M = get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a" defines put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a = a_put_M apply (simp add: get_M_is_l_get_M l_put_M_def) by (simp add: get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def) lemmas put_M_defs = a_put_M_def adhoc_overloading put_M put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a locale l_put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas = l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a begin sublocale l_put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas by unfold_locales interpretation l_put_M type_wf character_data_ptr_kinds get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a apply(unfold_locales) using get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_type_wf l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a.type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a local.l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_axioms apply blast by (meson CharacterDataMonad.get_M_is_l_get_M l_get_M_def) lemmas put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ok = put_M_ok[folded put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def] end global_interpretation l_put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas type_wf by unfold_locales lemma CharacterData_simp1 [simp]: "(\x. getter (setter (\_. v) x) = v) \ h \ put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \\<^sub>h h' \ h' \ get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr getter \\<^sub>r v" by(auto simp add: put_M_defs get_M_defs split: option.splits) lemma CharacterData_simp2 [simp]: "character_data_ptr \ character_data_ptr' \ h \ put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \\<^sub>h h' \ preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr' getter) h h'" by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E) lemma CharacterData_simp3 [simp]: " (\x. getter (setter (\_. v) x) = getter x) \ h \ put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \\<^sub>h h' \ preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr' getter) h h'" apply(cases "character_data_ptr = character_data_ptr'") by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E) lemma CharacterData_simp4 [simp]: "h \ put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \\<^sub>h h' \ preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr getter) h h'" by(auto simp add: put_M_defs ElementMonad.get_M_defs preserved_def split: option.splits dest: get_heap_E) lemma CharacterData_simp5 [simp]: "h \ put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \\<^sub>h h' \ preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr getter) h h'" by(auto simp add: ElementMonad.put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E) lemma CharacterData_simp6 [simp]: "(\x. getter (cast (setter (\_. v) x)) = getter (cast x)) \ h \ put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \\<^sub>h h' \ preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'" apply (cases "cast character_data_ptr = object_ptr") by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs NodeMonad.get_M_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def bind_eq_Some_conv split: option.splits) lemma CharacterData_simp7 [simp]: "(\x. getter (cast (setter (\_. v) x)) = getter (cast x)) \ h \ put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \\<^sub>h h' \ preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'" apply(cases "cast character_data_ptr = node_ptr") by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs NodeMonad.get_M_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def bind_eq_Some_conv split: option.splits) lemma CharacterData_simp8 [simp]: "cast character_data_ptr \ node_ptr \ h \ put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \\<^sub>h h' \ preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'" by(auto simp add: put_M_defs get_M_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def NodeMonad.get_M_defs preserved_def split: option.splits dest: get_heap_E) lemma CharacterData_simp9 [simp]: "h \ put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \\<^sub>h h' \ (\x. getter (cast (setter (\_. v) x)) = getter (cast x)) \ preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'" apply(cases "cast character_data_ptr \ node_ptr") by(auto simp add: put_M_defs get_M_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def NodeMonad.get_M_defs preserved_def split: option.splits bind_splits dest: get_heap_E) lemma CharacterData_simp10 [simp]: "cast character_data_ptr \ node_ptr \ h \ put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr setter v \\<^sub>h h' \ preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr getter) h h'" by(auto simp add: NodeMonad.put_M_defs get_M_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def NodeMonad.get_M_defs preserved_def split: option.splits dest: get_heap_E) lemma CharacterData_simp11 [simp]: "cast character_data_ptr \ object_ptr \ h \ put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \\<^sub>h h' \ preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'" by(auto simp add: put_M_defs get_M_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def ObjectMonad.get_M_defs preserved_def split: option.splits dest: get_heap_E) lemma CharacterData_simp12 [simp]: "h \ put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr setter v \\<^sub>h h' \ (\x. getter (cast (setter (\_. v) x)) = getter (cast x)) \ preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'" apply(cases "cast character_data_ptr \ object_ptr") apply(auto simp add: put_M_defs get_M_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def split: option.splits bind_splits dest: get_heap_E)[1] by(auto simp add: put_M_defs get_M_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def split: option.splits bind_splits dest: get_heap_E)[1] lemma CharacterData_simp13 [simp]: "cast character_data_ptr \ object_ptr \ h \ put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr setter v \\<^sub>h h' \ preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr getter) h h'" by(auto simp add: ObjectMonad.put_M_defs get_M_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def split: option.splits dest: get_heap_E) lemma new_element_get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a: "h \ new_element \\<^sub>h h' \ preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr getter) h h'" by(auto simp add: new_element_def get_M_defs preserved_def split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E) subsection\Creating CharacterData\ definition new_character_data :: "(_, (_) character_data_ptr) dom_prog" where "new_character_data = do { h \ get_heap; (new_ptr, h') \ return (new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a h); return_heap h'; return new_ptr }" lemma new_character_data_ok [simp]: "h \ ok new_character_data" by(auto simp add: new_character_data_def split: prod.splits) lemma new_character_data_ptr_in_heap: assumes "h \ new_character_data \\<^sub>h h'" and "h \ new_character_data \\<^sub>r new_character_data_ptr" shows "new_character_data_ptr |\| character_data_ptr_kinds h'" using assms unfolding new_character_data_def by(auto simp add: new_character_data_def new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ptr_in_heap is_OK_returns_result_I elim!: bind_returns_result_E bind_returns_heap_E) lemma new_character_data_ptr_not_in_heap: assumes "h \ new_character_data \\<^sub>h h'" and "h \ new_character_data \\<^sub>r new_character_data_ptr" shows "new_character_data_ptr |\| character_data_ptr_kinds h" using assms new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ptr_not_in_heap by(auto simp add: new_character_data_def split: prod.splits elim!: bind_returns_result_E bind_returns_heap_E) lemma new_character_data_new_ptr: assumes "h \ new_character_data \\<^sub>h h'" and "h \ new_character_data \\<^sub>r new_character_data_ptr" shows "object_ptr_kinds h' = object_ptr_kinds h |\| {|cast new_character_data_ptr|}" using assms new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_new_ptr by(auto simp add: new_character_data_def split: prod.splits elim!: bind_returns_result_E bind_returns_heap_E) lemma new_character_data_is_character_data_ptr: assumes "h \ new_character_data \\<^sub>r new_character_data_ptr" shows "is_character_data_ptr new_character_data_ptr" using assms new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_is_character_data_ptr by(auto simp add: new_character_data_def elim!: bind_returns_result_E split: prod.splits) lemma new_character_data_child_nodes: assumes "h \ new_character_data \\<^sub>h h'" assumes "h \ new_character_data \\<^sub>r new_character_data_ptr" shows "h' \ get_M new_character_data_ptr val \\<^sub>r ''''" using assms by(auto simp add: get_M_defs new_character_data_def new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E) lemma new_character_data_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t: "h \ new_character_data \\<^sub>h h' \ h \ new_character_data \\<^sub>r new_character_data_ptr \ ptr \ cast new_character_data_ptr \ preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr getter) h h'" by(auto simp add: new_character_data_def ObjectMonad.get_M_defs preserved_def split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E) lemma new_character_data_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e: "h \ new_character_data \\<^sub>h h' \ h \ new_character_data \\<^sub>r new_character_data_ptr \ ptr \ cast new_character_data_ptr \ preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e ptr getter) h h'" by(auto simp add: new_character_data_def NodeMonad.get_M_defs preserved_def split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E) lemma new_character_data_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t: "h \ new_character_data \\<^sub>h h' \ h \ new_character_data \\<^sub>r new_character_data_ptr \ preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ptr getter) h h'" by(auto simp add: new_character_data_def ElementMonad.get_M_defs preserved_def split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E) lemma new_character_data_get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a: "h \ new_character_data \\<^sub>h h' \ h \ new_character_data \\<^sub>r new_character_data_ptr \ ptr \ new_character_data_ptr \ preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a ptr getter) h h'" by(auto simp add: new_character_data_def get_M_defs preserved_def split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E) subsection\Modified Heaps\ lemma get_CharacterData_ptr_simp [simp]: "get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h) = (if ptr = cast character_data_ptr then cast obj else get character_data_ptr h)" by(auto simp add: get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def split: option.splits Option.bind_splits) lemma Character_data_ptr_kinds_simp [simp]: "character_data_ptr_kinds (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h) = character_data_ptr_kinds h |\| (if is_character_data_ptr_kind ptr then {|the (cast ptr)|} else {||})" by(auto simp add: character_data_ptr_kinds_def is_node_ptr_kind_def split: option.splits) lemma type_wf_put_I: assumes "type_wf h" assumes "ElementClass.type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)" assumes "is_character_data_ptr_kind ptr \ is_character_data_kind obj" shows "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)" using assms by(auto simp add: type_wf_defs split: option.splits) lemma type_wf_put_ptr_not_in_heap_E: assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)" assumes "ptr |\| object_ptr_kinds h" shows "type_wf h" using assms apply(auto simp add: type_wf_defs elim!: ElementMonad.type_wf_put_ptr_not_in_heap_E split: option.splits if_splits)[1] using assms(2) node_ptr_kinds_commutes by blast lemma type_wf_put_ptr_in_heap_E: assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)" assumes "ptr |\| object_ptr_kinds h" assumes "ElementClass.type_wf h" assumes "is_character_data_ptr_kind ptr \ is_character_data_kind (the (get ptr h))" shows "type_wf h" using assms apply(auto simp add: type_wf_defs split: option.splits if_splits)[1] by (metis (no_types, lifting) ElementClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf assms(2) bind.bind_lunit cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_inv cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def notin_fset option.collapse) subsection\Preserving Types\ lemma new_element_type_wf_preserved [simp]: assumes "h \ new_element \\<^sub>h h'" shows "type_wf h = type_wf h'" using assms apply(auto simp add: new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def elim!: bind_returns_heap_E type_wf_put_ptr_not_in_heap_E intro!: type_wf_put_I split: if_splits)[1] using CharacterDataClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t assms new_element_type_wf_preserved apply blast using element_ptrs_def apply fastforce using CharacterDataClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t assms new_element_type_wf_preserved apply blast by (metis Suc_n_not_le_n element_ptr.sel(1) element_ptrs_def fMax_ge ffmember_filter fimage_eqI is_element_ptr_ref) lemma new_element_is_l_new_element: "l_new_element type_wf" using l_new_element.intro new_element_type_wf_preserved by blast lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_tag_name_type_wf_preserved [simp]: "h \ put_M element_ptr tag_name_update v \\<^sub>h h' \ type_wf h = type_wf h'" apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def dest!: get_heap_E elim!: bind_returns_heap_E2 intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1] apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1] apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1] apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1] apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1] apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1] using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def) apply (metis finite_set_in) done lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_child_nodes_type_wf_preserved [simp]: "h \ put_M element_ptr child_nodes_update v \\<^sub>h h' \ type_wf h = type_wf h'" apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def dest!: get_heap_E elim!: bind_returns_heap_E2 intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1] apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1] apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1] apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1] apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1] apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1] using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def) apply (metis finite_set_in) done lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_attrs_type_wf_preserved [simp]: "h \ put_M element_ptr attrs_update v \\<^sub>h h' \ type_wf h = type_wf h'" apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def dest!: get_heap_E elim!: bind_returns_heap_E2 intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1] apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1] apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1] apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1] apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1] apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1] using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def) apply (metis finite_set_in) done lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_shadow_root_opt_type_wf_preserved [simp]: "h \ put_M element_ptr shadow_root_opt_update v \\<^sub>h h' \ type_wf h = type_wf h'" apply(auto simp add: ElementMonad.put_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def dest!: get_heap_E elim!: bind_returns_heap_E2 intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1] apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1] apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1] apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1] apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1] apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1] using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def) apply (metis finite_set_in) done lemma new_character_data_type_wf_preserved [simp]: "h \ new_character_data \\<^sub>h h' \ type_wf h = type_wf h'" apply(auto simp add: new_character_data_def new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def Let_def put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def elim!: bind_returns_heap_E type_wf_put_ptr_not_in_heap_E intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I split: if_splits)[1] apply(simp_all add: type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs is_node_kind_def) by (meson new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_ptr_not_in_heap) locale l_new_character_data = l_type_wf + assumes new_character_data_types_preserved: "h \ new_character_data \\<^sub>h h' \ type_wf h = type_wf h'" lemma new_character_data_is_l_new_character_data: "l_new_character_data type_wf" using l_new_character_data.intro new_character_data_type_wf_preserved by blast lemma put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_val_type_wf_preserved [simp]: "h \ put_M character_data_ptr val_update v \\<^sub>h h' \ type_wf h = type_wf h'" apply(auto simp add: CharacterDataMonad.put_M_defs put\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def CharacterDataClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e CharacterDataClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t is_node_kind_def dest!: get_heap_E elim!: bind_returns_heap_E2 intro!: type_wf_put_I ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1] apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs CharacterDataMonad.get_M_defs split: option.splits)[1] apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs CharacterDataMonad.get_M_defs ObjectClass.a_type_wf_def split: option.splits)[1] apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def) apply (metis finite_set_in) done lemma character_data_ptr_kinds_small: assumes "\object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'" shows "character_data_ptr_kinds h = character_data_ptr_kinds h'" by(simp add: character_data_ptr_kinds_def node_ptr_kinds_def preserved_def object_ptr_kinds_preserved_small[OF assms]) lemma character_data_ptr_kinds_preserved: assumes "writes SW setter h h'" assumes "h \ setter \\<^sub>h h'" assumes "\h h'. \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 "character_data_ptr_kinds h = character_data_ptr_kinds h'" using writes_small_big[OF assms] apply(simp add: reflp_def transp_def preserved_def character_data_ptr_kinds_def) by (metis assms node_ptr_kinds_preserved) lemma type_wf_preserved_small: assumes "\object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'" assumes "\node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'" assumes "\element_ptr. preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr RElement.nothing) h h'" assumes "\character_data_ptr. preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr RCharacterData.nothing) h h'" shows "type_wf h = type_wf h'" using type_wf_preserved_small[OF assms(1) assms(2) assms(3)] allI[OF assms(4), of id, simplified] character_data_ptr_kinds_small[OF assms(1)] apply(auto simp add: type_wf_defs preserved_def get_M_defs character_data_ptr_kinds_small[OF assms(1)] split: option.splits)[1] apply(force) by force lemma type_wf_preserved: assumes "writes SW setter h h'" assumes "h \ setter \\<^sub>h h'" assumes "\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'" assumes "\h h' w. w \ SW \ h \ w \\<^sub>h h' \ \node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'" assumes "\h h' w. w \ SW \ h \ w \\<^sub>h h' \ \element_ptr. preserved (get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr RElement.nothing) h h'" assumes "\h h' w. w \ SW \ h \ w \\<^sub>h h' \ \character_data_ptr. preserved (get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a character_data_ptr RCharacterData.nothing) h h'" shows "type_wf h = type_wf h'" proof - have "\h h' w. w \ SW \ h \ w \\<^sub>h h' \ type_wf h = type_wf h'" using assms type_wf_preserved_small by fast with assms(1) assms(2) show ?thesis apply(rule writes_small_big) by(auto simp add: reflp_def transp_def) qed lemma type_wf_drop: "type_wf h \ type_wf (Heap (fmdrop ptr (the_heap h)))" apply(auto simp add: type_wf_def ElementMonad.type_wf_drop l_type_wf_def\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a.a_type_wf_def)[1] using type_wf_drop by (metis (no_types, lifting) ElementClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf character_data_ptr_kinds_commutes finite_set_in fmlookup_drop get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def node_ptr_kinds_commutes object_ptr_kinds_code5) end diff --git a/thys/Core_DOM/common/monads/NodeMonad.thy b/thys/Core_DOM/common/monads/NodeMonad.thy --- a/thys/Core_DOM/common/monads/NodeMonad.thy +++ b/thys/Core_DOM/common/monads/NodeMonad.thy @@ -1,218 +1,218 @@ (*********************************************************************************** * 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\Node\ text\In this theory, we introduce the monadic method setup for the Node class.\ theory NodeMonad imports ObjectMonad "../classes/NodeClass" begin type_synonym ('object_ptr, 'node_ptr, 'Object, 'Node, 'result) dom_prog = "((_) heap, exception, 'result) prog" register_default_tvars "('object_ptr, 'node_ptr, 'Object, 'Node, 'result) dom_prog" global_interpretation l_ptr_kinds_M node_ptr_kinds defines node_ptr_kinds_M = a_ptr_kinds_M . lemmas node_ptr_kinds_M_defs = a_ptr_kinds_M_def lemma node_ptr_kinds_M_eq: assumes "|h \ object_ptr_kinds_M|\<^sub>r = |h' \ object_ptr_kinds_M|\<^sub>r" shows "|h \ node_ptr_kinds_M|\<^sub>r = |h' \ node_ptr_kinds_M|\<^sub>r" using assms by(auto simp add: node_ptr_kinds_M_defs object_ptr_kinds_M_defs node_ptr_kinds_def) global_interpretation l_dummy defines get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e = "l_get_M.a_get_M get\<^sub>N\<^sub>o\<^sub>d\<^sub>e" . lemma get_M_is_l_get_M: "l_get_M get\<^sub>N\<^sub>o\<^sub>d\<^sub>e type_wf node_ptr_kinds" apply(simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf l_get_M_def) by (metis ObjectClass.a_type_wf_def ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf bind_eq_None_conv get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def node_ptr_kinds_commutes option.simps(3)) lemmas get_M_defs = get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def[unfolded l_get_M.a_get_M_def[OF get_M_is_l_get_M]] adhoc_overloading get_M get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e locale l_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas = l_type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e begin sublocale l_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas by unfold_locales interpretation l_get_M get\<^sub>N\<^sub>o\<^sub>d\<^sub>e type_wf node_ptr_kinds apply(unfold_locales) apply (simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf local.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e) by (meson NodeMonad.get_M_is_l_get_M l_get_M_def) lemmas get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_ok = get_M_ok[folded get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def] end global_interpretation l_get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas type_wf by unfold_locales lemma node_ptr_kinds_M_reads: "reads (\object_ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing)}) node_ptr_kinds_M h h'" using object_ptr_kinds_M_reads apply (simp add: reads_def node_ptr_kinds_M_defs node_ptr_kinds_def object_ptr_kinds_M_reads preserved_def) - by (smt object_ptr_kinds_preserved_small preserved_def unit_all_impI) + by (metis (mono_tags, lifting) object_ptr_kinds_preserved_small old.unit.exhaust preserved_def) global_interpretation l_put_M type_wf node_ptr_kinds get\<^sub>N\<^sub>o\<^sub>d\<^sub>e put\<^sub>N\<^sub>o\<^sub>d\<^sub>e rewrites "a_get_M = get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e" defines put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e = a_put_M apply (simp add: get_M_is_l_get_M l_put_M_def) by (simp add: get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def) lemmas put_M_defs = a_put_M_def adhoc_overloading put_M put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e locale l_put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas = l_type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e begin sublocale l_put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas by unfold_locales interpretation l_put_M type_wf node_ptr_kinds get\<^sub>N\<^sub>o\<^sub>d\<^sub>e put\<^sub>N\<^sub>o\<^sub>d\<^sub>e apply(unfold_locales) apply (simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf local.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e) by (meson NodeMonad.get_M_is_l_get_M l_get_M_def) lemmas put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_ok = put_M_ok[folded put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def] end global_interpretation l_put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas type_wf by unfold_locales lemma get_M_Object_preserved1 [simp]: "(\x. getter (cast (setter (\_. v) x)) = getter (cast x)) \ h \ put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr setter v \\<^sub>h h' \ preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'" apply(cases "cast node_ptr = object_ptr") by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def preserved_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def bind_eq_Some_conv split: option.splits) lemma get_M_Object_preserved2 [simp]: "cast node_ptr \ object_ptr \ h \ put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr setter v \\<^sub>h h' \ preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'" by(auto simp add: put_M_defs get_M_defs get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def split: option.splits dest: get_heap_E) lemma get_M_Object_preserved3 [simp]: "h \ put_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr setter v \\<^sub>h h' \ (\x. getter (cast (setter (\_. v) x)) = getter (cast x)) \ preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr getter) h h'" apply(cases "cast node_ptr \ object_ptr") by(auto simp add: put_M_defs get_M_defs get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def split: option.splits bind_splits dest: get_heap_E) lemma get_M_Object_preserved4 [simp]: "cast node_ptr \ object_ptr \ h \ put_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr setter v \\<^sub>h h' \ preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr getter) h h'" by(auto simp add: ObjectMonad.put_M_defs get_M_defs get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def ObjectMonad.get_M_defs preserved_def split: option.splits dest: get_heap_E) subsection\Modified Heaps\ lemma get_node_ptr_simp [simp]: "get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h) = (if ptr = cast node_ptr then cast obj else get node_ptr h)" by(auto simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def) lemma node_ptr_kinds_simp [simp]: "node_ptr_kinds (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h) = node_ptr_kinds h |\| (if is_node_ptr_kind ptr then {|the (cast ptr)|} else {||})" by(auto simp add: node_ptr_kinds_def) lemma type_wf_put_I: assumes "type_wf h" assumes "ObjectClass.type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)" assumes "is_node_ptr_kind ptr \ is_node_kind obj" shows "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)" using assms apply(auto simp add: type_wf_defs split: option.splits)[1] using cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_none is_node_kind_def apply blast using cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_none is_node_kind_def apply blast done lemma type_wf_put_ptr_not_in_heap_E: assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)" assumes "ptr |\| object_ptr_kinds h" shows "type_wf h" using assms by(auto simp add: type_wf_defs elim!: ObjectMonad.type_wf_put_ptr_not_in_heap_E split: option.splits if_splits) lemma type_wf_put_ptr_in_heap_E: assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)" assumes "ptr |\| object_ptr_kinds h" assumes "ObjectClass.type_wf h" assumes "is_node_ptr_kind ptr \ is_node_kind (the (get ptr h))" shows "type_wf h" using assms apply(auto simp add: type_wf_defs split: option.splits if_splits)[1] by (metis ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf bind.bind_lunit finite_set_in get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def is_node_kind_def option.exhaust_sel) subsection\Preserving Types\ lemma node_ptr_kinds_small: assumes "\object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'" shows "node_ptr_kinds h = node_ptr_kinds h'" by(simp add: node_ptr_kinds_def preserved_def object_ptr_kinds_preserved_small[OF assms]) lemma node_ptr_kinds_preserved: assumes "writes SW setter h h'" assumes "h \ setter \\<^sub>h h'" assumes "\h h'. \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 "node_ptr_kinds h = node_ptr_kinds h'" using writes_small_big[OF assms] apply(simp add: reflp_def transp_def preserved_def node_ptr_kinds_def) by (metis assms object_ptr_kinds_preserved) lemma type_wf_preserved_small: assumes "\object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'" assumes "\node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'" shows "type_wf h = type_wf h'" using type_wf_preserved allI[OF assms(2), of id, simplified] apply(auto simp add: type_wf_defs)[1] apply(auto simp add: preserved_def get_M_defs node_ptr_kinds_small[OF assms(1)] split: option.splits)[1] apply (metis notin_fset option.simps(3)) by(auto simp add: preserved_def get_M_defs node_ptr_kinds_small[OF assms(1)] split: option.splits, force)[1] lemma type_wf_preserved: assumes "writes SW setter h h'" assumes "h \ setter \\<^sub>h h'" assumes "\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'" assumes "\h h' w. w \ SW \ h \ w \\<^sub>h h' \ \node_ptr. preserved (get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr RNode.nothing) h h'" shows "type_wf h = type_wf h'" proof - have "\h h' w. w \ SW \ h \ w \\<^sub>h h' \ type_wf h = type_wf h'" using assms type_wf_preserved_small by fast with assms(1) assms(2) show ?thesis apply(rule writes_small_big) by(auto simp add: reflp_def transp_def) qed end 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,8044 @@ (*********************************************************************************** * 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)) + by (smt (verit) 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) + by (smt (verit) \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) + by (metis \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' 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 + by (smt (verit) disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2 disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_in_heap document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 in_set_remove1 list.set_intros(1) 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 + by (smt (verit) 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) + by (smt (verit, best) assms(5) assms(7) assms(8) 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 + apply (smt (verit) 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) + apply (smt (verit) 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)) - + by (smt (verit) children_eq2_h3 disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h2 + disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 element_ptr_kinds_commutes finite_set_in + h' h2 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 + new_element_ptr new_element_ptr_in_heap node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h3 + object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subsetD subsetI) 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\ + apply (smt (verit) 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\ + by (smt (verit) 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 + by (smt (verit) 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 + by (smt (verit) character_data_ptr_kinds_commutes children_eq2_h3 disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 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 + by (smt (verit) 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 + by (smt (verit) 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/Shadow_DOM/Shadow_DOM.thy b/thys/Shadow_DOM/Shadow_DOM.thy --- a/thys/Shadow_DOM/Shadow_DOM.thy +++ b/thys/Shadow_DOM/Shadow_DOM.thy @@ -1,10501 +1,10462 @@ (*********************************************************************************** * 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\The Shadow DOM\ theory Shadow_DOM imports "monads/ShadowRootMonad" Core_DOM.Core_DOM begin abbreviation "safe_shadow_root_element_types \ {''article'', ''aside'', ''blockquote'', ''body'', ''div'', ''footer'', ''h1'', ''h2'', ''h3'', ''h4'', ''h5'', ''h6'', ''header'', ''main'', ''nav'', ''p'', ''section'', ''span''}" subsection \Function Definitions\ subsubsection \get\_child\_nodes\ locale l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = CD: l_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin definition get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) shadow_root_ptr \ unit \ (_, (_) node_ptr list) dom_prog" where "get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr _ = get_M shadow_root_ptr RShadowRoot.child_nodes" definition a_get_child_nodes_tups :: "(((_) object_ptr \ bool) \ ((_) object_ptr \ unit \ (_, (_) node_ptr list) dom_prog)) list" where "a_get_child_nodes_tups \ [(is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r, get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast)]" definition a_get_child_nodes :: "(_) object_ptr \ (_, (_) node_ptr list) dom_prog" where "a_get_child_nodes ptr = invoke (CD.a_get_child_nodes_tups @ a_get_child_nodes_tups) ptr ()" definition a_get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" where "a_get_child_nodes_locs ptr \ (if is_shadow_root_ptr_kind ptr then {preserved (get_M (the (cast ptr)) RShadowRoot.child_nodes)} else {}) \ CD.a_get_child_nodes_locs ptr" definition first_child :: "(_) object_ptr \ (_, (_) node_ptr option) dom_prog" where "first_child ptr = do { children \ a_get_child_nodes ptr; return (case children of [] \ None | child#_ \ Some child)}" end global_interpretation l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines get_child_nodes = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_child_nodes and get_child_nodes_locs = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_child_nodes_locs . locale l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_type_wf type_wf + l_known_ptr known_ptr + l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_child_nodes_defs get_child_nodes get_child_nodes_locs + CD: l_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for type_wf :: "(_) heap \ bool" and known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and get_child_nodes :: "(_) object_ptr \ (_, (_) node_ptr list) dom_prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ (_, (_) node_ptr list) dom_prog" and get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" + assumes known_ptr_impl: "known_ptr = ShadowRootClass.known_ptr" assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes get_child_nodes_impl: "get_child_nodes = a_get_child_nodes" assumes get_child_nodes_locs_impl: "get_child_nodes_locs = a_get_child_nodes_locs" begin lemmas get_child_nodes_def = get_child_nodes_impl[unfolded a_get_child_nodes_def get_child_nodes_def] lemmas get_child_nodes_locs_def = get_child_nodes_locs_impl[unfolded a_get_child_nodes_locs_def get_child_nodes_locs_def, folded CD.get_child_nodes_locs_impl] lemma get_child_nodes_ok: assumes "known_ptr ptr" assumes "type_wf h" assumes "ptr |\| object_ptr_kinds h" shows "h \ ok (get_child_nodes ptr)" using assms[unfolded known_ptr_impl type_wf_impl] apply(auto simp add: get_child_nodes_def)[1] apply(split CD.get_child_nodes_splits, rule conjI)+ using ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t CD.get_child_nodes_ok CD.known_ptr_impl CD.type_wf_impl apply blast apply(auto simp add: CD.known_ptr_impl a_get_child_nodes_tups_def get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok dest!: known_ptr_new_shadow_root_ptr intro!: bind_is_OK_I2)[1] by (metis is_shadow_root_ptr_kind_none l_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_lemmas.get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok l_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_lemmas_axioms option.case_eq_if shadow_root_ptr_casts_commute3 shadow_root_ptr_kinds_commutes) lemma get_child_nodes_ptr_in_heap: assumes "h \ get_child_nodes ptr \\<^sub>r children" shows "ptr |\| object_ptr_kinds h" using assms by(auto simp add: get_child_nodes_def invoke_ptr_in_heap dest: is_OK_returns_result_I) lemma get_child_nodes_pure [simp]: "pure (get_child_nodes ptr) h" unfolding get_child_nodes_def a_get_child_nodes_tups_def proof(split CD.get_child_nodes_splits, rule conjI; clarify) assume "known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M ptr" then show "pure (get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M ptr) h" by simp next assume "\ known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M ptr" then show "pure (invoke [(is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r, get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^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>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r)] ptr ()) h" by(auto simp add: get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro: bind_pure_I split: invoke_splits) qed lemma get_child_nodes_reads: "reads (get_child_nodes_locs ptr) (get_child_nodes ptr) h h'" apply (simp add: get_child_nodes_def a_get_child_nodes_tups_def get_child_nodes_locs_def CD.get_child_nodes_locs_def) apply(split CD.get_child_nodes_splits, rule conjI)+ apply(auto intro!: reads_subset[OF CD.get_child_nodes_reads[unfolded CD.get_child_nodes_locs_def]] split: if_splits)[1] apply(split invoke_splits, rule conjI)+ apply(auto)[1] apply(auto simp add: get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro: reads_subset[OF reads_singleton] reads_subset[OF check_in_heap_reads] intro!: reads_bind_pure reads_subset[OF return_reads] split: option.splits)[1] done end interpretation i_get_child_nodes?: l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by(simp add: l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_child_nodes_is_l_get_child_nodes [instances]: "l_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs" apply(auto simp add: l_get_child_nodes_def instances)[1] using get_child_nodes_reads get_child_nodes_ok get_child_nodes_ptr_in_heap get_child_nodes_pure by blast+ paragraph \new\_document\ locale l_new_document_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_new_document_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for type_wf :: "(_) heap \ bool" and known_ptr :: "(_) object_ptr \ 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 type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_child_nodes_new_document: "ptr' \ cast new_document_ptr \ h \ new_document \\<^sub>r new_document_ptr \ h \ new_document \\<^sub>h h' \ r \ get_child_nodes_locs ptr' \ r h h'" apply(auto simp add: get_child_nodes_locs_def)[1] using CD.get_child_nodes_new_document apply (metis document_ptr_casts_commute3 empty_iff is_document_ptr_kind_none new_document_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t option.case_eq_if shadow_root_ptr_casts_commute3 singletonD) by (simp add: CD.get_child_nodes_new_document) lemma new_document_no_child_nodes: "h \ new_document \\<^sub>r new_document_ptr \ h \ new_document \\<^sub>h h' \ h' \ get_child_nodes (cast new_document_ptr) \\<^sub>r []" apply(auto simp add: get_child_nodes_def)[1] apply(split CD.get_child_nodes_splits, rule conjI)+ using CD.new_document_no_child_nodes apply auto[1] by(auto simp add: DocumentClass.a_known_ptr_def CD.known_ptr_impl known_ptr_def dest!: new_document_is_document_ptr) end interpretation i_new_document_get_child_nodes?: l_new_document_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr get_child_nodes get_child_nodes_locs DocumentClass.type_wf DocumentClass.known_ptr Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by(unfold_locales) declare l_new_document_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma new_document_get_child_nodes_is_l_new_document_get_child_nodes [instances]: "l_new_document_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs" using new_document_is_l_new_document get_child_nodes_is_l_get_child_nodes apply(simp add: l_new_document_get_child_nodes_def l_new_document_get_child_nodes_axioms_def) using get_child_nodes_new_document new_document_no_child_nodes by fast paragraph \new\_shadow\_root\ locale l_new_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for type_wf :: "(_) heap \ bool" and known_ptr :: "(_) object_ptr \ 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 type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_child_nodes_new_shadow_root: "ptr' \ cast new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_child_nodes_locs ptr' \ r h h'" apply(auto simp add: get_child_nodes_locs_def)[1] apply (metis document_ptr_casts_commute3 insert_absorb insert_not_empty is_document_ptr_kind_none new_shadow_root_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t option.case_eq_if shadow_root_ptr_casts_commute3 singletonD) apply(auto simp add: CD.get_child_nodes_locs_def)[1] using new_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t apply blast - apply (smt insertCI new_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t singleton_iff) - apply (metis document_ptr_casts_commute3 empty_iff new_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t singletonD) + apply (metis empty_iff insertE new_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t) + apply (metis empty_iff new_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t singletonD) done lemma new_shadow_root_no_child_nodes: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ h' \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []" apply(auto simp add: get_child_nodes_def )[1] apply(split CD.get_child_nodes_splits, rule conjI)+ apply(auto simp add: CD.get_child_nodes_def CD.a_get_child_nodes_tups_def)[1] apply(split invoke_splits, rule conjI)+ using NodeClass.a_known_ptr_def known_ptr_not_character_data_ptr known_ptr_not_document_ptr known_ptr_not_element_ptr local.CD.known_ptr_impl apply blast apply(auto simp add: is_document_ptr_def split: option.splits document_ptr.splits)[1] apply(auto simp add: is_character_data_ptr_def split: option.splits document_ptr.splits)[1] apply(auto simp add: is_element_ptr_def split: option.splits document_ptr.splits)[1] apply(auto simp add: a_get_child_nodes_tups_def)[1] apply(split invoke_splits, rule conjI)+ apply(auto simp add: is_shadow_root_ptr_def split: shadow_root_ptr.splits dest!: new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_is_shadow_root_ptr)[1] apply(auto intro!: bind_pure_returns_result_I)[1] apply(drule(1) new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_in_heap) apply(auto simp add: shadow_root_ptr_kinds_def document_ptr_kinds_def)[1] apply (metis check_in_heap_ptr_in_heap is_OK_returns_result_E old.unit.exhaust) using new_shadow_root_children by (simp add: new_shadow_root_children get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def) end interpretation i_new_shadow_root_get_child_nodes?: l_new_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr get_child_nodes get_child_nodes_locs DocumentClass.type_wf DocumentClass.known_ptr Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by(unfold_locales) declare l_new_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def[instances] locale l_new_shadow_root_get_child_nodes = l_get_child_nodes + assumes get_child_nodes_new_shadow_root: "ptr' \ cast new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_child_nodes_locs ptr' \ r h h'" assumes new_shadow_root_no_child_nodes: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ h' \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []" lemma new_shadow_root_get_child_nodes_is_l_new_shadow_root_get_child_nodes [instances]: "l_new_shadow_root_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs" apply(simp add: l_new_shadow_root_get_child_nodes_def l_new_shadow_root_get_child_nodes_axioms_def instances) using get_child_nodes_new_shadow_root new_shadow_root_no_child_nodes by fast paragraph \new\_element\ locale l_new_element_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_new_element_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_child_nodes_new_element: "ptr' \ cast new_element_ptr \ h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ r \ get_child_nodes_locs ptr' \ r h h'" by (auto simp add: get_child_nodes_locs_def CD.get_child_nodes_locs_def new_element_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_element_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t new_element_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t new_element_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) lemma new_element_no_child_nodes: "h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ h' \ get_child_nodes (cast new_element_ptr) \\<^sub>r []" apply(auto simp add: get_child_nodes_def a_get_child_nodes_tups_def split: prod.splits elim!: bind_returns_result_E bind_returns_heap_E)[1] apply(split CD.get_child_nodes_splits, rule conjI)+ using local.new_element_no_child_nodes apply auto[1] apply(auto simp add: invoke_def)[1] apply(auto simp add: new_element_ptr_in_heap get_child_nodes\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def check_in_heap_def new_element_child_nodes intro!: bind_pure_returns_result_I intro: new_element_is_element_ptr elim!: new_element_ptr_in_heap)[1] proof - assume " h \ new_element \\<^sub>r new_element_ptr" assume "h \ new_element \\<^sub>h h'" assume "\ known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>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)" moreover have "known_ptr (cast new_element_ptr)" using new_element_is_element_ptr \h \ new_element \\<^sub>r new_element_ptr\ by(auto simp add: known_ptr_impl ShadowRootClass.a_known_ptr_def DocumentClass.a_known_ptr_def CharacterDataClass.a_known_ptr_def ElementClass.a_known_ptr_def) ultimately show "False" by(simp add: known_ptr_impl CD.known_ptr_impl ShadowRootClass.a_known_ptr_def is_document_ptr_kind_none) qed end interpretation i_new_element_get_child_nodes?: l_new_element_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by(unfold_locales) declare l_new_element_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma new_element_get_child_nodes_is_l_new_element_get_child_nodes [instances]: "l_new_element_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs" using new_element_is_l_new_element get_child_nodes_is_l_get_child_nodes apply(auto simp add: l_new_element_get_child_nodes_def l_new_element_get_child_nodes_axioms_def)[1] using get_child_nodes_new_element new_element_no_child_nodes by fast+ subsubsection \delete\_shadow\_root\ locale l_delete_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_child_nodes_delete_shadow_root: "ptr' \ cast shadow_root_ptr \ h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_child_nodes_locs ptr' \ r h h'" by(auto simp add: get_child_nodes_locs_def CD.get_child_nodes_locs_def delete_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t delete_shadow_root_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t delete_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t delete_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: if_splits option.splits intro: is_shadow_root_ptr_kind_obtains) end locale l_delete_shadow_root_get_child_nodes = l_get_child_nodes_defs + assumes get_child_nodes_delete_shadow_root: "ptr' \ cast shadow_root_ptr \ h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_child_nodes_locs ptr' \ r h h'" interpretation l_delete_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by(auto simp add: l_delete_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) lemma l_delete_shadow_root_get_child_nodes_get_child_nodes_locs [instances]: "l_delete_shadow_root_get_child_nodes get_child_nodes_locs" apply(auto simp add: l_delete_shadow_root_get_child_nodes_def)[1] using get_child_nodes_delete_shadow_root apply fast done subsubsection \set\_child\_nodes\ locale l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = CD: l_set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin definition set_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) shadow_root_ptr \ (_) node_ptr list \ (_, unit) dom_prog" where "set_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr = put_M shadow_root_ptr RShadowRoot.child_nodes_update" definition a_set_child_nodes_tups :: "(((_) object_ptr \ bool) \ ((_) object_ptr \ (_) node_ptr list \ (_, unit) dom_prog)) list" where "a_set_child_nodes_tups \ [(is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r, set_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast)]" definition a_set_child_nodes :: "(_) object_ptr \ (_) node_ptr list \ (_, unit) dom_prog" where "a_set_child_nodes ptr children = invoke (CD.a_set_child_nodes_tups @ a_set_child_nodes_tups) ptr children" definition a_set_child_nodes_locs :: "(_) object_ptr \ (_, unit) dom_prog set" where "a_set_child_nodes_locs ptr \ (if is_shadow_root_ptr_kind ptr then all_args (put_M (the (cast ptr)) RShadowRoot.child_nodes_update) else {}) \ CD.a_set_child_nodes_locs ptr" end global_interpretation l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines set_child_nodes = l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_set_child_nodes and set_child_nodes_locs = l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_set_child_nodes_locs . locale l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_type_wf type_wf + l_known_ptr known_ptr + l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_set_child_nodes_defs set_child_nodes set_child_nodes_locs + CD: l_set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for type_wf :: "(_) heap \ bool" and known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and set_child_nodes :: "(_) object_ptr \ (_) node_ptr list \ (_, unit) dom_prog" and set_child_nodes_locs :: "(_) object_ptr \ (_, unit) dom_prog set" and set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ (_) node_ptr list \ (_, unit) dom_prog" and set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ (_, unit) dom_prog set" + assumes known_ptr_impl: "known_ptr = ShadowRootClass.known_ptr" assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes set_child_nodes_impl: "set_child_nodes = a_set_child_nodes" assumes set_child_nodes_locs_impl: "set_child_nodes_locs = a_set_child_nodes_locs" begin lemmas set_child_nodes_def = set_child_nodes_impl[unfolded a_set_child_nodes_def set_child_nodes_def] lemmas set_child_nodes_locs_def = set_child_nodes_locs_impl[unfolded a_set_child_nodes_locs_def set_child_nodes_locs_def, folded CD.set_child_nodes_locs_impl] lemma set_child_nodes_writes: "writes (set_child_nodes_locs ptr) (set_child_nodes ptr children) h h'" apply (simp add: set_child_nodes_def a_set_child_nodes_tups_def set_child_nodes_locs_def) apply(split CD.set_child_nodes_splits, rule conjI)+ apply (simp add: CD.set_child_nodes_writes writes_union_right_I) apply(split invoke_splits, rule conjI)+ apply(auto simp add: a_set_child_nodes_def)[1] apply(auto simp add: set_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: writes_bind_pure intro: writes_union_right_I writes_union_left_I split: list.splits)[1] by (simp add: is_shadow_root_ptr_kind_none) lemma set_child_nodes_pointers_preserved: assumes "w \ set_child_nodes_locs object_ptr" assumes "h \ w \\<^sub>h h'" shows "object_ptr_kinds h = object_ptr_kinds h'" using assms(1) object_ptr_kinds_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def set_child_nodes_locs_def CD.set_child_nodes_locs_def split: if_splits) lemma set_child_nodes_types_preserved: assumes "w \ set_child_nodes_locs object_ptr" assumes "h \ w \\<^sub>h h'" shows "type_wf h = type_wf h'" using assms(1) type_wf_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def type_wf_impl a_set_child_nodes_tups_def set_child_nodes_locs_def CD.set_child_nodes_locs_def split: if_splits option.splits) end interpretation i_set_child_nodes?: l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs apply(unfold_locales) by (auto simp add: set_child_nodes_def set_child_nodes_locs_def) declare l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_child_nodes_is_l_set_child_nodes [instances]: "l_set_child_nodes type_wf set_child_nodes set_child_nodes_locs" using instances Shadow_DOM.i_set_child_nodes.set_child_nodes_pointers_preserved Shadow_DOM.i_set_child_nodes.set_child_nodes_writes i_set_child_nodes.set_child_nodes_types_preserved unfolding l_set_child_nodes_def by blast paragraph \get\_child\_nodes\ locale l_set_child_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes set_child_nodes_locs set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + CD: l_set_child_nodes_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for type_wf :: "(_) heap \ bool" and known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ 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_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and set_child_nodes :: "(_) object_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_child_nodes_locs :: "(_) object_ptr \ ((_) heap, exception, unit) prog set" and set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, unit) prog set" begin lemma set_child_nodes_get_child_nodes: assumes "known_ptr ptr" assumes "type_wf h" assumes "h \ set_child_nodes ptr children \\<^sub>h h'" shows "h' \ get_child_nodes ptr \\<^sub>r children" proof - have "h \ check_in_heap ptr \\<^sub>r ()" using assms set_child_nodes_def invoke_ptr_in_heap by (metis (full_types) check_in_heap_ptr_in_heap is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust) then have ptr_in_h: "ptr |\| object_ptr_kinds h" by (simp add: check_in_heap_ptr_in_heap is_OK_returns_result_I) have "type_wf h'" apply(unfold type_wf_impl) apply(rule subst[where P=id, OF type_wf_preserved[OF set_child_nodes_writes assms(3), unfolded all_args_def], simplified]) by(auto simp add: all_args_def assms(2)[unfolded type_wf_impl] set_child_nodes_locs_def CD.set_child_nodes_locs_def split: if_splits) have "h' \ check_in_heap ptr \\<^sub>r ()" using check_in_heap_reads set_child_nodes_writes assms(3) \h \ check_in_heap ptr \\<^sub>r ()\ apply(rule reads_writes_separate_forwards) apply(auto simp add: all_args_def set_child_nodes_locs_def CD.set_child_nodes_locs_def)[1] done then have "ptr |\| object_ptr_kinds h'" using check_in_heap_ptr_in_heap by blast with assms ptr_in_h \type_wf h'\ show ?thesis apply(auto simp add: type_wf_impl known_ptr_impl get_child_nodes_def a_get_child_nodes_tups_def set_child_nodes_def a_set_child_nodes_tups_def del: bind_pure_returns_result_I2 intro!: bind_pure_returns_result_I2)[1] apply(split CD.get_child_nodes_splits, (rule conjI impI)+)+ apply(split CD.set_child_nodes_splits)+ apply(auto simp add: CD.set_child_nodes_get_child_nodes type_wf_impl CD.type_wf_impl dest: ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t)[1] apply(auto simp add: CD.set_child_nodes_get_child_nodes type_wf_impl CD.type_wf_impl dest: ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t)[1] apply(split CD.set_child_nodes_splits)+ by(auto simp add: known_ptr_impl CD.known_ptr_impl set_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def get_child_nodes\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def CD.type_wf_impl ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t dest: known_ptr_new_shadow_root_ptr)[2] qed lemma set_child_nodes_get_child_nodes_different_pointers: assumes "ptr \ ptr'" assumes "w \ set_child_nodes_locs ptr" assumes "h \ w \\<^sub>h h'" assumes "r \ get_child_nodes_locs ptr'" shows "r h h'" using assms apply(auto simp add: set_child_nodes_locs_def CD.set_child_nodes_locs_def get_child_nodes_locs_def CD.get_child_nodes_locs_def)[1] by(auto simp add: all_args_def elim!: is_document_ptr_kind_obtains is_shadow_root_ptr_kind_obtains is_element_ptr_kind_obtains split: if_splits option.splits) end interpretation i_set_child_nodes_get_child_nodes?: l_set_child_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs using instances by(auto simp add: l_set_child_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def ) declare l_set_child_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_child_nodes_get_child_nodes_is_l_set_child_nodes_get_child_nodes [instances]: "l_set_child_nodes_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs" apply(auto simp add: instances l_set_child_nodes_get_child_nodes_def l_set_child_nodes_get_child_nodes_axioms_def)[1] using set_child_nodes_get_child_nodes apply fast using set_child_nodes_get_child_nodes_different_pointers apply fast done subsubsection \set\_tag\_type\ locale l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_set_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_tag_name set_tag_name_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and set_tag_name :: "(_) element_ptr \ tag_name \ (_, unit) dom_prog" and set_tag_name_locs :: "(_) element_ptr \ (_, unit) dom_prog set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" begin lemmas set_tag_name_def = CD.set_tag_name_impl[unfolded CD.a_set_tag_name_def set_tag_name_def] lemmas set_tag_name_locs_def = CD.set_tag_name_locs_impl[unfolded CD.a_set_tag_name_locs_def set_tag_name_locs_def] lemma set_tag_name_ok: "type_wf h \ element_ptr |\| element_ptr_kinds h \ h \ ok (set_tag_name element_ptr tag)" apply(unfold type_wf_impl) unfolding set_tag_name_impl[unfolded a_set_tag_name_def] using get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok using CD.set_tag_name_ok CD.type_wf_impl ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t by blast lemma set_tag_name_writes: "writes (set_tag_name_locs element_ptr) (set_tag_name element_ptr tag) h h'" using CD.set_tag_name_writes . lemma set_tag_name_pointers_preserved: assumes "w \ set_tag_name_locs element_ptr" assumes "h \ w \\<^sub>h h'" shows "object_ptr_kinds h = object_ptr_kinds h'" using assms by(simp add: CD.set_tag_name_pointers_preserved) lemma set_tag_name_typess_preserved: assumes "w \ set_tag_name_locs element_ptr" assumes "h \ w \\<^sub>h h'" shows "type_wf h = type_wf h'" apply(unfold type_wf_impl) apply(rule type_wf_preserved[OF writes_singleton2 assms(2)]) using assms(1) set_tag_name_locs_def by(auto simp add: all_args_def set_tag_name_locs_def split: if_splits) end interpretation i_set_tag_name?: l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_tag_name set_tag_name_locs by(auto simp add: l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma set_tag_name_is_l_set_tag_name [instances]: "l_set_tag_name type_wf set_tag_name set_tag_name_locs" apply(auto simp add: l_set_tag_name_def)[1] using set_tag_name_writes apply fast using set_tag_name_ok apply fast using set_tag_name_pointers_preserved apply (fast, fast) using set_tag_name_typess_preserved apply (fast, fast) done paragraph \get\_child\_nodes\ locale l_set_tag_name_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + CD: l_set_tag_name_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_tag_name set_tag_name_locs known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_tag_name_get_child_nodes: "\w \ set_tag_name_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" apply(auto simp add: get_child_nodes_locs_def)[1] apply(auto simp add: set_tag_name_locs_def all_args_def)[1] using CD.set_tag_name_get_child_nodes apply(blast) using CD.set_tag_name_get_child_nodes apply(blast) done end interpretation i_set_tag_name_get_child_nodes?: l_set_tag_name_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_tag_name set_tag_name_locs known_ptr DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by unfold_locales declare l_set_tag_name_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_tag_name_get_child_nodes_is_l_set_tag_name_get_child_nodes [instances]: "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" using set_tag_name_is_l_set_tag_name get_child_nodes_is_l_get_child_nodes apply(simp add: l_set_tag_name_get_child_nodes_def l_set_tag_name_get_child_nodes_axioms_def) using set_tag_name_get_child_nodes by fast subsubsection \get\_shadow\_root\ locale l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin definition a_get_shadow_root :: "(_) element_ptr \ (_, (_) shadow_root_ptr option) dom_prog" where "a_get_shadow_root element_ptr = get_M element_ptr shadow_root_opt" definition a_get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" where "a_get_shadow_root_locs element_ptr \ {preserved (get_M element_ptr shadow_root_opt)}" end global_interpretation l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines get_shadow_root = a_get_shadow_root and get_shadow_root_locs = a_get_shadow_root_locs . locale l_get_shadow_root_defs = fixes get_shadow_root :: "(_) element_ptr \ (_, (_) shadow_root_ptr option) dom_prog" fixes get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" locale l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_shadow_root_defs get_shadow_root get_shadow_root_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes get_shadow_root_impl: "get_shadow_root = a_get_shadow_root" assumes get_shadow_root_locs_impl: "get_shadow_root_locs = a_get_shadow_root_locs" begin lemmas get_shadow_root_def = get_shadow_root_impl[unfolded get_shadow_root_def a_get_shadow_root_def] lemmas get_shadow_root_locs_def = get_shadow_root_locs_impl[unfolded get_shadow_root_locs_def a_get_shadow_root_locs_def] lemma get_shadow_root_ok: "type_wf h \ element_ptr |\| element_ptr_kinds h \ h \ ok (get_shadow_root element_ptr)" unfolding get_shadow_root_def type_wf_impl using ShadowRootMonad.get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok by blast lemma get_shadow_root_pure [simp]: "pure (get_shadow_root element_ptr) h" unfolding get_shadow_root_def by simp lemma get_shadow_root_ptr_in_heap: assumes "h \ get_shadow_root element_ptr \\<^sub>r children" shows "element_ptr |\| element_ptr_kinds h" using assms by(auto simp add: get_shadow_root_def get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap dest: is_OK_returns_result_I) lemma get_shadow_root_reads: "reads (get_shadow_root_locs element_ptr) (get_shadow_root element_ptr) h h'" by(simp add: get_shadow_root_def get_shadow_root_locs_def reads_bind_pure reads_insert_writes_set_right) end interpretation i_get_shadow_root?: l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs using instances by (auto simp add: l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_get_shadow_root = l_type_wf + l_get_shadow_root_defs + assumes get_shadow_root_reads: "reads (get_shadow_root_locs element_ptr) (get_shadow_root element_ptr) h h'" assumes get_shadow_root_ok: "type_wf h \ element_ptr |\| element_ptr_kinds h \ h \ ok (get_shadow_root element_ptr)" assumes get_shadow_root_ptr_in_heap: "h \ ok (get_shadow_root element_ptr) \ element_ptr |\| element_ptr_kinds h" assumes get_shadow_root_pure [simp]: "pure (get_shadow_root element_ptr) h" lemma get_shadow_root_is_l_get_shadow_root [instances]: "l_get_shadow_root type_wf get_shadow_root get_shadow_root_locs" using instances unfolding l_get_shadow_root_def by (metis (no_types, lifting) ElementClass.l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_axioms get_shadow_root_ok get_shadow_root_pure get_shadow_root_reads l_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas.get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_in_heap l_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas.intro l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_shadow_root_def) paragraph \set\_disconnected\_nodes\ locale l_set_disconnected_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^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 type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_disconnected_nodes set_disconnected_nodes_locs + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" 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 get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma set_disconnected_nodes_get_shadow_root: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" by(auto simp add: set_disconnected_nodes_locs_def get_shadow_root_locs_def all_args_def) end locale l_set_disconnected_nodes_get_shadow_root = l_set_disconnected_nodes_defs + l_get_shadow_root_defs + assumes set_disconnected_nodes_get_shadow_root: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" interpretation i_set_disconnected_nodes_get_shadow_root?: l_set_disconnected_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_disconnected_nodes set_disconnected_nodes_locs get_shadow_root get_shadow_root_locs by(auto simp add: l_set_disconnected_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_disconnected_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_disconnected_nodes_get_shadow_root_is_l_set_disconnected_nodes_get_shadow_root [instances]: "l_set_disconnected_nodes_get_shadow_root set_disconnected_nodes_locs get_shadow_root_locs" apply(auto simp add: l_set_disconnected_nodes_get_shadow_root_def)[1] using set_disconnected_nodes_get_shadow_root apply fast done paragraph \set\_tag\_type\ locale l_set_tag_name_get_shadow_root\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_tag_name_get_shadow_root: "\w \ set_tag_name_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" by(auto simp add: set_tag_name_locs_def get_shadow_root_locs_def all_args_def intro: element_put_get_preserved[where setter=tag_name_update and getter=shadow_root_opt]) end locale l_set_tag_name_get_shadow_root = l_set_tag_name + l_get_shadow_root + assumes set_tag_name_get_shadow_root: "\w \ set_tag_name_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" interpretation i_set_tag_name_get_shadow_root?: l_set_tag_name_get_shadow_root\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_tag_name set_tag_name_locs get_shadow_root get_shadow_root_locs apply(auto simp add: l_set_tag_name_get_shadow_root\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] using l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms by unfold_locales declare l_set_tag_name_get_shadow_root\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_tag_name_get_shadow_root_is_l_set_tag_name_get_shadow_root [instances]: "l_set_tag_name_get_shadow_root type_wf set_tag_name set_tag_name_locs get_shadow_root get_shadow_root_locs" using set_tag_name_is_l_set_tag_name get_shadow_root_is_l_get_shadow_root apply(simp add: l_set_tag_name_get_shadow_root_def l_set_tag_name_get_shadow_root_axioms_def) using set_tag_name_get_shadow_root by fast paragraph \set\_child\_nodes\ locale l_set_child_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes set_child_nodes_locs set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for type_wf :: "(_) heap \ bool" and known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" and set_child_nodes :: "(_) object_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_child_nodes_locs :: "(_) object_ptr \ ((_) heap, exception, unit) prog set" and set_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, unit) prog set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma set_child_nodes_get_shadow_root: "\w \ set_child_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" apply(auto simp add: set_child_nodes_locs_def get_shadow_root_locs_def CD.set_child_nodes_locs_def all_args_def)[1] by(auto intro!: element_put_get_preserved[where getter=shadow_root_opt and setter=RElement.child_nodes_update]) end locale l_set_child_nodes_get_shadow_root = l_set_child_nodes_defs + l_get_shadow_root_defs + assumes set_child_nodes_get_shadow_root: "\w \ set_child_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" interpretation i_set_child_nodes_get_shadow_root?: l_set_child_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs get_shadow_root get_shadow_root_locs by(auto simp add: l_set_child_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_child_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_child_nodes_get_shadow_root_is_l_set_child_nodes_get_shadow_root [instances]: "l_set_child_nodes_get_shadow_root set_child_nodes_locs get_shadow_root_locs" apply(auto simp add: l_set_child_nodes_get_shadow_root_def)[1] using set_child_nodes_get_shadow_root apply fast done paragraph \delete\_shadow\_root\ locale l_delete_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_shadow_root_delete_shadow_root: "h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" by(auto simp add: get_shadow_root_locs_def delete_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t) end locale l_delete_shadow_root_get_shadow_root = l_get_shadow_root_defs + assumes get_shadow_root_delete_shadow_root: "h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" interpretation l_delete_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs by(auto simp add: l_delete_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) lemma l_delete_shadow_root_get_shadow_root_get_shadow_root_locs [instances]: "l_delete_shadow_root_get_shadow_root get_shadow_root_locs" apply(auto simp add: l_delete_shadow_root_get_shadow_root_def)[1] using get_shadow_root_delete_shadow_root apply fast done paragraph \new\_character\_data\ locale l_new_character_data_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for type_wf :: "(_) heap \ bool" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_shadow_root_new_character_data: "h \ new_character_data \\<^sub>r new_character_data_ptr \ h \ new_character_data \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" by (auto simp add: get_shadow_root_locs_def new_character_data_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_character_data_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) end locale l_new_character_data_get_shadow_root = l_new_character_data + l_get_shadow_root + assumes get_shadow_root_new_character_data: "h \ new_character_data \\<^sub>r new_character_data_ptr \ h \ new_character_data \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" interpretation i_new_character_data_get_shadow_root?: l_new_character_data_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs by(unfold_locales) declare l_new_character_data_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma new_character_data_get_shadow_root_is_l_new_character_data_get_shadow_root [instances]: "l_new_character_data_get_shadow_root type_wf get_shadow_root get_shadow_root_locs" using new_character_data_is_l_new_character_data get_shadow_root_is_l_get_shadow_root apply(auto simp add: l_new_character_data_get_shadow_root_def l_new_character_data_get_shadow_root_axioms_def instances)[1] using get_shadow_root_new_character_data by fast paragraph \new\_document\ locale l_new_document_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for type_wf :: "(_) heap \ bool" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_shadow_root_new_document: "h \ new_document \\<^sub>r new_document_ptr \ h \ new_document \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" by (auto simp add: get_shadow_root_locs_def new_document_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_document_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) end locale l_new_document_get_shadow_root = l_new_document + l_get_shadow_root + assumes get_shadow_root_new_document: "h \ new_document \\<^sub>r new_document_ptr \ h \ new_document \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" interpretation i_new_document_get_shadow_root?: l_new_document_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs by(unfold_locales) declare l_new_document_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma new_document_get_shadow_root_is_l_new_document_get_shadow_root [instances]: "l_new_document_get_shadow_root type_wf get_shadow_root get_shadow_root_locs" using new_document_is_l_new_document get_shadow_root_is_l_get_shadow_root apply(auto simp add: l_new_document_get_shadow_root_def l_new_document_get_shadow_root_axioms_def instances)[1] using get_shadow_root_new_document by fast paragraph \new\_element\ locale l_new_element_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for type_wf :: "(_) heap \ bool" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_shadow_root_new_element: "ptr' \ new_element_ptr \ h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" by (auto simp add: get_shadow_root_locs_def new_element_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_element_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t new_element_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) lemma new_element_no_shadow_root: "h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ h' \ get_shadow_root new_element_ptr \\<^sub>r None" by(simp add: get_shadow_root_def new_element_shadow_root_opt) end locale l_new_element_get_shadow_root = l_new_element + l_get_shadow_root + assumes get_shadow_root_new_element: "ptr' \ new_element_ptr \ h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" assumes new_element_no_shadow_root: "h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ h' \ get_shadow_root new_element_ptr \\<^sub>r None" interpretation i_new_element_get_shadow_root?: l_new_element_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs by(unfold_locales) declare l_new_element_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma new_element_get_shadow_root_is_l_new_element_get_shadow_root [instances]: "l_new_element_get_shadow_root type_wf get_shadow_root get_shadow_root_locs" using new_element_is_l_new_element get_shadow_root_is_l_get_shadow_root apply(auto simp add: l_new_element_get_shadow_root_def l_new_element_get_shadow_root_axioms_def instances)[1] using get_shadow_root_new_element new_element_no_shadow_root by fast+ paragraph \new\_shadow\_root\ locale l_new_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs for type_wf :: "(_) heap \ bool" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_shadow_root_new_shadow_root: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" by (auto simp add: get_shadow_root_locs_def new_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) end locale l_new_shadow_root_get_shadow_root = l_get_shadow_root + assumes get_shadow_root_new_shadow_root: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" interpretation i_new_shadow_root_get_shadow_root?: l_new_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_locs by(unfold_locales) declare l_new_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma new_shadow_root_get_shadow_root_is_l_new_shadow_root_get_shadow_root [instances]: "l_new_shadow_root_get_shadow_root type_wf get_shadow_root get_shadow_root_locs" using get_shadow_root_is_l_get_shadow_root apply(auto simp add: l_new_shadow_root_get_shadow_root_def l_new_shadow_root_get_shadow_root_axioms_def instances)[1] using get_shadow_root_new_shadow_root by fast subsubsection \set\_shadow\_root\ locale l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin definition a_set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ (_, unit) dom_prog" where "a_set_shadow_root element_ptr = put_M element_ptr shadow_root_opt_update" definition a_set_shadow_root_locs :: "(_) element_ptr \ ((_, unit) dom_prog) set" where "a_set_shadow_root_locs element_ptr \ all_args (put_M element_ptr shadow_root_opt_update)" end global_interpretation l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines set_shadow_root = a_set_shadow_root and set_shadow_root_locs = a_set_shadow_root_locs . locale l_set_shadow_root_defs = fixes set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ (_, unit) dom_prog" fixes set_shadow_root_locs :: "(_) element_ptr \ (_, unit) dom_prog set" locale l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_type_wf type_wf + l_set_shadow_root_defs set_shadow_root set_shadow_root_locs + l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs for type_wf :: "(_) heap \ bool" and set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ (_, unit) dom_prog" and set_shadow_root_locs :: "(_) element_ptr \ (_, unit) dom_prog set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes set_shadow_root_impl: "set_shadow_root = a_set_shadow_root" assumes set_shadow_root_locs_impl: "set_shadow_root_locs = a_set_shadow_root_locs" begin lemmas set_shadow_root_def = set_shadow_root_impl[unfolded set_shadow_root_def a_set_shadow_root_def] lemmas set_shadow_root_locs_def = set_shadow_root_locs_impl[unfolded set_shadow_root_locs_def a_set_shadow_root_locs_def] lemma set_shadow_root_ok: "type_wf h \ element_ptr |\| element_ptr_kinds h \ h \ ok (set_shadow_root element_ptr tag)" apply(unfold type_wf_impl) unfolding set_shadow_root_def using get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok by (simp add: ShadowRootMonad.put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok) lemma set_shadow_root_ptr_in_heap: "h \ ok (set_shadow_root element_ptr shadow_root) \ element_ptr |\| element_ptr_kinds h" by(simp add: set_shadow_root_def ElementMonad.put_M_ptr_in_heap) lemma set_shadow_root_writes: "writes (set_shadow_root_locs element_ptr) (set_shadow_root element_ptr tag) h h'" by(auto simp add: set_shadow_root_def set_shadow_root_locs_def intro: writes_bind_pure) lemma set_shadow_root_pointers_preserved: assumes "w \ set_shadow_root_locs element_ptr" assumes "h \ w \\<^sub>h h'" shows "object_ptr_kinds h = object_ptr_kinds h'" using assms(1) object_ptr_kinds_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def set_shadow_root_locs_def split: if_splits) lemma set_shadow_root_types_preserved: assumes "w \ set_shadow_root_locs element_ptr" assumes "h \ w \\<^sub>h h'" shows "type_wf h = type_wf h'" apply(unfold type_wf_impl) using assms(1) type_wf_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def set_shadow_root_locs_def split: if_splits) end interpretation i_set_shadow_root?: l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_shadow_root set_shadow_root_locs by (auto simp add: l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_set_shadow_root = l_type_wf +l_set_shadow_root_defs + assumes set_shadow_root_writes: "writes (set_shadow_root_locs element_ptr) (set_shadow_root element_ptr disc_nodes) h h'" assumes set_shadow_root_ok: "type_wf h \ element_ptr |\| element_ptr_kinds h \ h \ ok (set_shadow_root element_ptr shadow_root)" assumes set_shadow_root_ptr_in_heap: "h \ ok (set_shadow_root element_ptr shadow_root) \ element_ptr |\| element_ptr_kinds h" assumes set_shadow_root_pointers_preserved: "w \ set_shadow_root_locs element_ptr \ h \ w \\<^sub>h h' \ object_ptr_kinds h = object_ptr_kinds h'" assumes set_shadow_root_types_preserved: "w \ set_shadow_root_locs element_ptr \ h \ w \\<^sub>h h' \ type_wf h = type_wf h'" lemma set_shadow_root_is_l_set_shadow_root [instances]: "l_set_shadow_root type_wf set_shadow_root set_shadow_root_locs" apply(auto simp add: l_set_shadow_root_def instances)[1] using set_shadow_root_writes apply blast using set_shadow_root_ok apply (blast) using set_shadow_root_ptr_in_heap apply blast using set_shadow_root_pointers_preserved apply(blast, blast) using set_shadow_root_types_preserved apply(blast, blast) done paragraph \get\_shadow\_root\ locale l_set_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_shadow_root_get_shadow_root: "type_wf h \ h \ set_shadow_root ptr shadow_root_ptr_opt \\<^sub>h h' \ h' \ get_shadow_root ptr \\<^sub>r shadow_root_ptr_opt" by(auto simp add: set_shadow_root_def get_shadow_root_def) lemma set_shadow_root_get_shadow_root_different_pointers: "ptr \ ptr' \ \w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" by(auto simp add: set_shadow_root_locs_def get_shadow_root_locs_def all_args_def) end interpretation i_set_shadow_root_get_shadow_root?: l_set_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_shadow_root set_shadow_root_locs get_shadow_root get_shadow_root_locs apply(auto simp add: l_set_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] by(unfold_locales) declare l_set_shadow_root_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_shadow_root_get_shadow_root = l_type_wf + l_set_shadow_root_defs + l_get_shadow_root_defs + assumes set_shadow_root_get_shadow_root: "type_wf h \ h \ set_shadow_root ptr shadow_root_ptr_opt \\<^sub>h h' \ h' \ get_shadow_root ptr \\<^sub>r shadow_root_ptr_opt" assumes set_shadow_root_get_shadow_root_different_pointers: "ptr \ ptr' \ w \ set_shadow_root_locs ptr \ h \ w \\<^sub>h h' \ r \ get_shadow_root_locs ptr' \ r h h'" lemma set_shadow_root_get_shadow_root_is_l_set_shadow_root_get_shadow_root [instances]: "l_set_shadow_root_get_shadow_root type_wf set_shadow_root set_shadow_root_locs get_shadow_root get_shadow_root_locs" apply(auto simp add: l_set_shadow_root_get_shadow_root_def instances)[1] using set_shadow_root_get_shadow_root apply fast using set_shadow_root_get_shadow_root_different_pointers apply fast done subsubsection \set\_mode\ locale l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin definition a_set_mode :: "(_) shadow_root_ptr \ shadow_root_mode \ (_, unit) dom_prog" where "a_set_mode shadow_root_ptr = put_M shadow_root_ptr mode_update" definition a_set_mode_locs :: "(_) shadow_root_ptr \ ((_, unit) dom_prog) set" where "a_set_mode_locs shadow_root_ptr \ all_args (put_M shadow_root_ptr mode_update)" end global_interpretation l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines set_mode = a_set_mode and set_mode_locs = a_set_mode_locs . locale l_set_mode_defs = fixes set_mode :: "(_) shadow_root_ptr \ shadow_root_mode \ (_, unit) dom_prog" fixes set_mode_locs :: "(_) shadow_root_ptr \ (_, unit) dom_prog set" locale l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_type_wf type_wf + l_set_mode_defs set_mode set_mode_locs + l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs for type_wf :: "(_) heap \ bool" and set_mode :: "(_) shadow_root_ptr \ shadow_root_mode \ (_, unit) dom_prog" and set_mode_locs :: "(_) shadow_root_ptr \ (_, unit) dom_prog set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes set_mode_impl: "set_mode = a_set_mode" assumes set_mode_locs_impl: "set_mode_locs = a_set_mode_locs" begin lemmas set_mode_def = set_mode_impl[unfolded set_mode_def a_set_mode_def] lemmas set_mode_locs_def = set_mode_locs_impl[unfolded set_mode_locs_def a_set_mode_locs_def] lemma set_mode_ok: "type_wf h \ shadow_root_ptr |\| shadow_root_ptr_kinds h \ h \ ok (set_mode shadow_root_ptr shadow_root_mode)" apply(unfold type_wf_impl) unfolding set_mode_def using get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok by (simp add: ShadowRootMonad.put_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok) lemma set_mode_ptr_in_heap: "h \ ok (set_mode shadow_root_ptr shadow_root_mode) \ shadow_root_ptr |\| shadow_root_ptr_kinds h" by(simp add: set_mode_def put_M_ptr_in_heap) lemma set_mode_writes: "writes (set_mode_locs shadow_root_ptr) (set_mode shadow_root_ptr shadow_root_mode) h h'" by(auto simp add: set_mode_def set_mode_locs_def intro: writes_bind_pure) lemma set_mode_pointers_preserved: assumes "w \ set_mode_locs element_ptr" assumes "h \ w \\<^sub>h h'" shows "object_ptr_kinds h = object_ptr_kinds h'" using assms(1) object_ptr_kinds_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def set_mode_locs_def split: if_splits) lemma set_mode_types_preserved: assumes "w \ set_mode_locs element_ptr" assumes "h \ w \\<^sub>h h'" shows "type_wf h = type_wf h'" apply(unfold type_wf_impl) using assms(1) type_wf_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def set_mode_locs_def split: if_splits) end interpretation i_set_mode?: l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_mode set_mode_locs by (auto simp add: l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_set_mode = l_type_wf +l_set_mode_defs + assumes set_mode_writes: "writes (set_mode_locs shadow_root_ptr) (set_mode shadow_root_ptr shadow_root_mode) h h'" assumes set_mode_ok: "type_wf h \ shadow_root_ptr |\| shadow_root_ptr_kinds h \ h \ ok (set_mode shadow_root_ptr shadow_root_mode)" assumes set_mode_ptr_in_heap: "h \ ok (set_mode shadow_root_ptr shadow_root_mode) \ shadow_root_ptr |\| shadow_root_ptr_kinds h" assumes set_mode_pointers_preserved: "w \ set_mode_locs shadow_root_ptr \ h \ w \\<^sub>h h' \ object_ptr_kinds h = object_ptr_kinds h'" assumes set_mode_types_preserved: "w \ set_mode_locs shadow_root_ptr \ h \ w \\<^sub>h h' \ type_wf h = type_wf h'" lemma set_mode_is_l_set_mode [instances]: "l_set_mode type_wf set_mode set_mode_locs" apply(auto simp add: l_set_mode_def instances)[1] using set_mode_writes apply blast using set_mode_ok apply (blast) using set_mode_ptr_in_heap apply blast using set_mode_pointers_preserved apply(blast, blast) using set_mode_types_preserved apply(blast, blast) done paragraph \get\_child\_nodes\ locale l_set_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_shadow_root_get_child_nodes: "\w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" by(auto simp add: get_child_nodes_locs_def set_shadow_root_locs_def CD.get_child_nodes_locs_def all_args_def intro: element_put_get_preserved[where setter=shadow_root_opt_update]) end interpretation i_set_shadow_root_get_child_nodes?: l_set_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs set_shadow_root set_shadow_root_locs by(unfold_locales) declare l_set_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_shadow_root_get_child_nodes = l_set_shadow_root + l_get_child_nodes + assumes set_shadow_root_get_child_nodes: "\w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" lemma set_shadow_root_get_child_nodes_is_l_set_shadow_root_get_child_nodes [instances]: "l_set_shadow_root_get_child_nodes type_wf set_shadow_root set_shadow_root_locs known_ptr get_child_nodes get_child_nodes_locs" apply(auto simp add: l_set_shadow_root_get_child_nodes_def l_set_shadow_root_get_child_nodes_axioms_def instances)[1] using set_shadow_root_get_child_nodes apply blast done paragraph \get\_shadow\_root\ locale l_set_mode_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_mode_get_shadow_root: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" by(auto simp add: set_mode_locs_def get_shadow_root_locs_def all_args_def) end interpretation i_set_mode_get_shadow_root?: l_set_mode_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_mode set_mode_locs get_shadow_root get_shadow_root_locs by unfold_locales declare l_set_mode_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_mode_get_shadow_root = l_set_mode + l_get_shadow_root + assumes set_mode_get_shadow_root: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" lemma set_mode_get_shadow_root_is_l_set_mode_get_shadow_root [instances]: "l_set_mode_get_shadow_root type_wf set_mode set_mode_locs get_shadow_root get_shadow_root_locs" using set_mode_is_l_set_mode get_shadow_root_is_l_get_shadow_root apply(simp add: l_set_mode_get_shadow_root_def l_set_mode_get_shadow_root_axioms_def) using set_mode_get_shadow_root by fast paragraph \get\_child\_nodes\ locale l_set_mode_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_mode_get_child_nodes: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" by(auto simp add: get_child_nodes_locs_def CD.get_child_nodes_locs_def set_mode_locs_def all_args_def)[1] end interpretation i_set_mode_get_child_nodes?: l_set_mode_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_mode set_mode_locs known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by unfold_locales declare l_set_mode_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_mode_get_child_nodes = l_set_mode + l_get_child_nodes + assumes set_mode_get_child_nodes: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" lemma set_mode_get_child_nodes_is_l_set_mode_get_child_nodes [instances]: "l_set_mode_get_child_nodes type_wf set_mode set_mode_locs known_ptr get_child_nodes get_child_nodes_locs" using set_mode_is_l_set_mode get_child_nodes_is_l_get_child_nodes apply(simp add: l_set_mode_get_child_nodes_def l_set_mode_get_child_nodes_axioms_def) using set_mode_get_child_nodes by fast subsubsection \get\_host\ locale l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_shadow_root_defs get_shadow_root get_shadow_root_locs for get_shadow_root :: "(_::linorder) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_get_host :: "(_) shadow_root_ptr \ (_, (_) element_ptr) dom_prog" where "a_get_host shadow_root_ptr = do { host_ptrs \ element_ptr_kinds_M \ filter_M (\element_ptr. do { shadow_root_opt \ get_shadow_root element_ptr; return (shadow_root_opt = Some shadow_root_ptr) }); (case host_ptrs of host_ptr#[] \ return host_ptr | _ \ error HierarchyRequestError) }" definition "a_get_host_locs \ (\element_ptr. (get_shadow_root_locs element_ptr)) \ (\ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr RObject.nothing)})" end global_interpretation l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_shadow_root get_shadow_root_locs defines get_host = "a_get_host" and get_host_locs = "a_get_host_locs" . locale l_get_host_defs = fixes get_host :: "(_) shadow_root_ptr \ (_, (_) element_ptr) dom_prog" fixes get_host_locs :: "((_) heap \ (_) heap \ bool) set" locale l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_host_defs + l_get_shadow_root + assumes get_host_impl: "get_host = a_get_host" assumes get_host_locs_impl: "get_host_locs = a_get_host_locs" begin lemmas get_host_def = get_host_impl[unfolded a_get_host_def] lemmas get_host_locs_def = get_host_locs_impl[unfolded a_get_host_locs_def] lemma get_host_pure [simp]: "pure (get_host element_ptr) h" by(auto simp add: get_host_def intro!: bind_pure_I filter_M_pure_I split: list.splits) lemma get_host_reads: "reads get_host_locs (get_host element_ptr) h h'" using get_shadow_root_reads[unfolded reads_def] by(auto simp add: get_host_def get_host_locs_def intro!: reads_bind_pure reads_subset[OF check_in_heap_reads] reads_subset[OF error_reads] reads_subset[OF get_shadow_root_reads] reads_subset[OF return_reads] reads_subset[OF element_ptr_kinds_M_reads] filter_M_reads filter_M_pure_I bind_pure_I split: list.splits) end locale l_get_host = l_get_host_defs + assumes get_host_pure [simp]: "pure (get_host element_ptr) h" assumes get_host_reads: "reads get_host_locs (get_host node_ptr) h h'" interpretation i_get_host?: l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_shadow_root get_shadow_root_locs get_host get_host_locs type_wf using instances by (simp add: l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_host_def get_host_locs_def) declare l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_host_is_l_get_host [instances]: "l_get_host get_host get_host_locs" apply(auto simp add: l_get_host_def)[1] using get_host_reads apply fast done subsubsection \get\_mode\ locale l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin definition a_get_mode :: "(_) shadow_root_ptr \ (_, shadow_root_mode) dom_prog" where "a_get_mode shadow_root_ptr = get_M shadow_root_ptr mode" definition a_get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" where "a_get_mode_locs shadow_root_ptr \ {preserved (get_M shadow_root_ptr mode)}" end global_interpretation l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines get_mode = a_get_mode and get_mode_locs = a_get_mode_locs . locale l_get_mode_defs = fixes get_mode :: "(_) shadow_root_ptr \ (_, shadow_root_mode) dom_prog" fixes get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" locale l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_mode_defs get_mode get_mode_locs + l_type_wf type_wf for get_mode :: "(_) shadow_root_ptr \ ((_) heap, exception, shadow_root_mode) prog" and get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" and type_wf :: "(_) heap \ bool" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes get_mode_impl: "get_mode = a_get_mode" assumes get_mode_locs_impl: "get_mode_locs = a_get_mode_locs" begin lemmas get_mode_def = get_mode_impl[unfolded get_mode_def a_get_mode_def] lemmas get_mode_locs_def = get_mode_locs_impl[unfolded get_mode_locs_def a_get_mode_locs_def] lemma get_mode_ok: "type_wf h \ shadow_root_ptr |\| shadow_root_ptr_kinds h \ h \ ok (get_mode shadow_root_ptr)" unfolding get_mode_def type_wf_impl using ShadowRootMonad.get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ok by blast lemma get_mode_pure [simp]: "pure (get_mode element_ptr) h" unfolding get_mode_def by simp lemma get_mode_ptr_in_heap: assumes "h \ get_mode shadow_root_ptr \\<^sub>r children" shows "shadow_root_ptr |\| shadow_root_ptr_kinds h" using assms by(auto simp add: get_mode_def get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_ptr_in_heap dest: is_OK_returns_result_I) lemma get_mode_reads: "reads (get_mode_locs element_ptr) (get_mode element_ptr) h h'" by(simp add: get_mode_def get_mode_locs_def reads_bind_pure reads_insert_writes_set_right) end interpretation i_get_mode?: l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_mode get_mode_locs type_wf using instances by (auto simp add: l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_get_mode = l_type_wf + l_get_mode_defs + assumes get_mode_reads: "reads (get_mode_locs shadow_root_ptr) (get_mode shadow_root_ptr) h h'" assumes get_mode_ok: "type_wf h \ shadow_root_ptr |\| shadow_root_ptr_kinds h \ h \ ok (get_mode shadow_root_ptr)" assumes get_mode_ptr_in_heap: "h \ ok (get_mode shadow_root_ptr) \ shadow_root_ptr |\| shadow_root_ptr_kinds h" assumes get_mode_pure [simp]: "pure (get_mode shadow_root_ptr) h" lemma get_mode_is_l_get_mode [instances]: "l_get_mode type_wf get_mode get_mode_locs" apply(auto simp add: l_get_mode_def instances)[1] using get_mode_reads apply blast using get_mode_ok apply blast using get_mode_ptr_in_heap apply blast done subsubsection \get\_shadow\_root\_safe\ locale l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_shadow_root_defs get_shadow_root get_shadow_root_locs + l_get_mode_defs get_mode get_mode_locs for get_shadow_root :: "(_) element_ptr \ (_, (_) shadow_root_ptr option) dom_prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_mode :: "(_) shadow_root_ptr \ (_, shadow_root_mode) dom_prog" and get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_get_shadow_root_safe :: "(_) element_ptr \ (_, (_) shadow_root_ptr option) dom_prog" where "a_get_shadow_root_safe element_ptr = do { shadow_root_ptr_opt \ get_shadow_root element_ptr; (case shadow_root_ptr_opt of Some shadow_root_ptr \ do { mode \ get_mode shadow_root_ptr; (if mode = Open then return (Some shadow_root_ptr) else return None ) } | None \ return None) }" definition a_get_shadow_root_safe_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" where "a_get_shadow_root_safe_locs element_ptr shadow_root_ptr \ (get_shadow_root_locs element_ptr) \ (get_mode_locs shadow_root_ptr)" end global_interpretation l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_shadow_root get_shadow_root_locs get_mode get_mode_locs defines get_shadow_root_safe = a_get_shadow_root_safe and get_shadow_root_safe_locs = a_get_shadow_root_safe_locs . locale l_get_shadow_root_safe_defs = fixes get_shadow_root_safe :: "(_) element_ptr \ (_, (_) shadow_root_ptr option) dom_prog" fixes get_shadow_root_safe_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" locale l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_shadow_root get_shadow_root_locs get_mode get_mode_locs + l_get_shadow_root_safe_defs get_shadow_root_safe get_shadow_root_safe_locs + l_get_shadow_root type_wf get_shadow_root get_shadow_root_locs + l_get_mode type_wf get_mode get_mode_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and get_shadow_root_safe :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_safe_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ (_, (_) shadow_root_ptr option) dom_prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_mode :: "(_) shadow_root_ptr \ (_, shadow_root_mode) dom_prog" and get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" assumes get_shadow_root_safe_impl: "get_shadow_root_safe = a_get_shadow_root_safe" assumes get_shadow_root_safe_locs_impl: "get_shadow_root_safe_locs = a_get_shadow_root_safe_locs" begin lemmas get_shadow_root_safe_def = get_shadow_root_safe_impl[unfolded get_shadow_root_safe_def a_get_shadow_root_safe_def] lemmas get_shadow_root_safe_locs_def = get_shadow_root_safe_locs_impl[unfolded get_shadow_root_safe_locs_def a_get_shadow_root_safe_locs_def] lemma get_shadow_root_safe_pure [simp]: "pure (get_shadow_root_safe element_ptr) h" - apply(auto simp add: get_shadow_root_safe_def)[1] - by (smt bind_returns_heap_E is_OK_returns_heap_E local.get_mode_pure local.get_shadow_root_pure - option.case_eq_if pure_def pure_returns_heap_eq return_pure) + by (auto simp add: get_shadow_root_safe_def bind_pure_I option.case_eq_if) end interpretation i_get_shadow_root_safe?: l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root_safe get_shadow_root_safe_locs get_shadow_root get_shadow_root_locs get_mode get_mode_locs using instances by (auto simp add: l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_shadow_root_safe_def get_shadow_root_safe_locs_def) declare l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_get_shadow_root_safe = l_get_shadow_root_safe_defs + assumes get_shadow_root_safe_pure [simp]: "pure (get_shadow_root_safe element_ptr) h" lemma get_shadow_root_safe_is_l_get_shadow_root_safe [instances]: "l_get_shadow_root_safe get_shadow_root_safe" using instances apply(auto simp add: l_get_shadow_root_safe_def)[1] done subsubsection \set\_disconnected\_nodes\ locale l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_set_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_disconnected_nodes set_disconnected_nodes_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" 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" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" begin lemma set_disconnected_nodes_ok: "type_wf h \ document_ptr |\| document_ptr_kinds h \ h \ ok (set_disconnected_nodes document_ptr node_ptrs)" using CD.set_disconnected_nodes_ok CD.type_wf_impl ShadowRootClass.type_wf_defs local.type_wf_impl by blast lemma set_disconnected_nodes_typess_preserved: assumes "w \ set_disconnected_nodes_locs object_ptr" assumes "h \ w \\<^sub>h h'" shows "type_wf h = type_wf h'" using assms(1) type_wf_preserved[OF writes_singleton2 assms(2)] apply(unfold type_wf_impl) by(auto simp add: all_args_def CD.set_disconnected_nodes_locs_def split: if_splits) end interpretation i_set_disconnected_nodes?: l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_disconnected_nodes set_disconnected_nodes_locs by(auto simp add: l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma set_disconnected_nodes_is_l_set_disconnected_nodes [instances]: "l_set_disconnected_nodes type_wf set_disconnected_nodes set_disconnected_nodes_locs" apply(auto simp add: l_set_disconnected_nodes_def)[1] apply (simp add: i_set_disconnected_nodes.set_disconnected_nodes_writes) using set_disconnected_nodes_ok apply blast apply (simp add: i_set_disconnected_nodes.set_disconnected_nodes_ptr_in_heap) using i_set_disconnected_nodes.set_disconnected_nodes_pointers_preserved apply (blast, blast) using set_disconnected_nodes_typess_preserved apply(blast, blast) done paragraph \get\_child\_nodes\ locale l_set_disconnected_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^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 type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_disconnected_nodes set_disconnected_nodes_locs + l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for type_wf :: "(_) heap \ bool" 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 known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ 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_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma set_disconnected_nodes_get_child_nodes: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" by(auto simp add: set_disconnected_nodes_locs_def get_child_nodes_locs_def CD.get_child_nodes_locs_def all_args_def) end interpretation i_set_disconnected_nodes_get_child_nodes?: l_set_disconnected_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_disconnected_nodes set_disconnected_nodes_locs known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs by(auto simp add: l_set_disconnected_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_disconnected_nodes_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_disconnected_nodes_get_child_nodes_is_l_set_disconnected_nodes_get_child_nodes [instances]: "l_set_disconnected_nodes_get_child_nodes set_disconnected_nodes_locs get_child_nodes_locs" apply(auto simp add: l_set_disconnected_nodes_get_child_nodes_def)[1] using set_disconnected_nodes_get_child_nodes apply fast done paragraph \get\_host\ locale l_set_disconnected_nodes_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_disconnected_nodes_get_host: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_host_locs. r h h'))" by(auto simp add: CD.set_disconnected_nodes_locs_def get_shadow_root_locs_def get_host_locs_def all_args_def) end interpretation i_set_disconnected_nodes_get_host?: l_set_disconnected_nodes_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_disconnected_nodes set_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_host get_host_locs by(auto simp add: l_set_disconnected_nodes_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_disconnected_nodes_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_set_disconnected_nodes_get_host = l_set_disconnected_nodes_defs + l_get_host_defs + assumes set_disconnected_nodes_get_host: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_host_locs. r h h'))" lemma set_disconnected_nodes_get_host_is_l_set_disconnected_nodes_get_host [instances]: "l_set_disconnected_nodes_get_host set_disconnected_nodes_locs get_host_locs" apply(auto simp add: l_set_disconnected_nodes_get_host_def instances)[1] using set_disconnected_nodes_get_host by fast subsubsection \get\_tag\_name\ locale l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_get_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_tag_name get_tag_name_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_tag_name :: "(_) element_ptr \ (_, tag_name) dom_prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" begin lemma get_tag_name_ok: "type_wf h \ element_ptr |\| element_ptr_kinds h \ h \ ok (get_tag_name element_ptr)" apply(unfold type_wf_impl get_tag_name_impl[unfolded a_get_tag_name_def]) using CD.get_tag_name_ok CD.type_wf_impl ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t by blast end interpretation i_get_tag_name?: l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs by(auto simp add: l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_tag_name_is_l_get_tag_name [instances]: "l_get_tag_name type_wf get_tag_name get_tag_name_locs" apply(auto simp add: l_get_tag_name_def)[1] using get_tag_name_reads apply fast using get_tag_name_ok apply fast done paragraph \set\_disconnected\_nodes\ locale l_set_disconnected_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_disconnected_nodes_get_tag_name: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" by(auto simp add: CD.set_disconnected_nodes_locs_def CD.get_tag_name_locs_def all_args_def) end interpretation i_set_disconnected_nodes_get_tag_name?: l_set_disconnected_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs by(auto simp add: l_set_disconnected_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_disconnected_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma set_disconnected_nodes_get_tag_name_is_l_set_disconnected_nodes_get_tag_name [instances]: "l_set_disconnected_nodes_get_tag_name type_wf set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs" apply(auto simp add: l_set_disconnected_nodes_get_tag_name_def l_set_disconnected_nodes_get_tag_name_axioms_def instances)[1] using set_disconnected_nodes_get_tag_name by fast paragraph \set\_child\_nodes\ locale l_set_child_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_child_nodes_get_tag_name: "\w \ set_child_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" by(auto simp add: CD.set_child_nodes_locs_def set_child_nodes_locs_def CD.get_tag_name_locs_def all_args_def intro: element_put_get_preserved[where getter=tag_name and setter=RElement.child_nodes_update]) end interpretation i_set_child_nodes_get_tag_name?: l_set_child_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs get_tag_name get_tag_name_locs by(auto simp add: l_set_child_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_child_nodes_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma set_child_nodes_get_tag_name_is_l_set_child_nodes_get_tag_name [instances]: "l_set_child_nodes_get_tag_name type_wf set_child_nodes set_child_nodes_locs get_tag_name get_tag_name_locs" apply(auto simp add: l_set_child_nodes_get_tag_name_def l_set_child_nodes_get_tag_name_axioms_def instances)[1] using set_child_nodes_get_tag_name by fast paragraph \delete\_shadow\_root\ locale l_delete_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_tag_name_delete_shadow_root: "h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" by(auto simp add: CD.get_tag_name_locs_def delete_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t) end locale l_delete_shadow_root_get_tag_name = l_get_tag_name_defs + assumes get_tag_name_delete_shadow_root: "h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" interpretation l_delete_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs by(auto simp add: l_delete_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) lemma l_delete_shadow_root_get_tag_name_get_tag_name_locs [instances]: "l_delete_shadow_root_get_tag_name get_tag_name_locs" apply(auto simp add: l_delete_shadow_root_get_tag_name_def)[1] using get_tag_name_delete_shadow_root apply fast done paragraph \set\_shadow\_root\ locale l_set_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_shadow_root_get_tag_name: "\w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" by(auto simp add: set_shadow_root_locs_def CD.get_tag_name_locs_def all_args_def element_put_get_preserved[where setter=shadow_root_opt_update]) end interpretation i_set_shadow_root_get_tag_name?: l_set_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_shadow_root set_shadow_root_locs DocumentClass.type_wf get_tag_name get_tag_name_locs apply(auto simp add: l_set_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] by(unfold_locales) declare l_set_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_shadow_root_get_tag_name = l_set_shadow_root_defs + l_get_tag_name_defs + assumes set_shadow_root_get_tag_name: "\w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" lemma set_shadow_root_get_tag_name_is_l_set_shadow_root_get_tag_name [instances]: "l_set_shadow_root_get_tag_name set_shadow_root_locs get_tag_name_locs" using set_shadow_root_is_l_set_shadow_root get_tag_name_is_l_get_tag_name apply(simp add: l_set_shadow_root_get_tag_name_def ) using set_shadow_root_get_tag_name by fast paragraph \new\_element\ locale l_new_element_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_tag_name get_tag_name_locs for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_tag_name :: "(_) element_ptr \ (_, tag_name) dom_prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_tag_name_new_element: "ptr' \ new_element_ptr \ h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" by (auto simp add: CD.get_tag_name_locs_def new_element_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_element_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t new_element_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) lemma new_element_empty_tag_name: "h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ h' \ get_tag_name new_element_ptr \\<^sub>r ''''" by(simp add: CD.get_tag_name_def new_element_tag_name) end locale l_new_element_get_tag_name = l_new_element + l_get_tag_name + assumes get_tag_name_new_element: "ptr' \ new_element_ptr \ h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" assumes new_element_empty_tag_name: "h \ new_element \\<^sub>r new_element_ptr \ h \ new_element \\<^sub>h h' \ h' \ get_tag_name new_element_ptr \\<^sub>r ''''" interpretation i_new_element_get_tag_name?: l_new_element_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs by(auto simp add: l_new_element_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_new_element_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma new_element_get_tag_name_is_l_new_element_get_tag_name [instances]: "l_new_element_get_tag_name type_wf get_tag_name get_tag_name_locs" using new_element_is_l_new_element get_tag_name_is_l_get_tag_name apply(auto simp add: l_new_element_get_tag_name_def l_new_element_get_tag_name_axioms_def instances)[1] using get_tag_name_new_element new_element_empty_tag_name by fast+ paragraph \get\_shadow\_root\ locale l_set_mode_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_mode_get_tag_name: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" by(auto simp add: set_mode_locs_def CD.get_tag_name_locs_def all_args_def) end interpretation i_set_mode_get_tag_name?: l_set_mode_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_mode set_mode_locs DocumentClass.type_wf get_tag_name get_tag_name_locs by unfold_locales declare l_set_mode_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_mode_get_tag_name = l_set_mode + l_get_tag_name + assumes set_mode_get_tag_name: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" lemma set_mode_get_tag_name_is_l_set_mode_get_tag_name [instances]: "l_set_mode_get_tag_name type_wf set_mode set_mode_locs get_tag_name get_tag_name_locs" using set_mode_is_l_set_mode get_tag_name_is_l_get_tag_name apply(simp add: l_set_mode_get_tag_name_def l_set_mode_get_tag_name_axioms_def) using set_mode_get_tag_name by fast paragraph \new\_document\ locale l_new_document_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_tag_name get_tag_name_locs for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, tag_name) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_tag_name_new_document: "h \ new_document \\<^sub>r new_document_ptr \ h \ new_document \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" by(auto simp add: CD.get_tag_name_locs_def new_document_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t) end locale l_new_document_get_tag_name = l_get_tag_name_defs + assumes get_tag_name_new_document: "h \ new_document \\<^sub>r new_document_ptr \ h \ new_document \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" interpretation i_new_document_get_tag_name?: l_new_document_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs by unfold_locales declare l_new_document_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def[instances] lemma new_document_get_tag_name_is_l_new_document_get_tag_name [instances]: "l_new_document_get_tag_name get_tag_name_locs" unfolding l_new_document_get_tag_name_def unfolding get_tag_name_locs_def using new_document_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t by blast paragraph \new\_shadow\_root\ locale l_new_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_tag_name_new_shadow_root: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" by (auto simp add: CD.get_tag_name_locs_def new_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t new_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t split: prod.splits if_splits option.splits elim!: bind_returns_result_E bind_returns_heap_E intro: is_element_ptr_kind_obtains) end locale l_new_shadow_root_get_tag_name = l_get_tag_name + assumes get_tag_name_new_shadow_root: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" interpretation i_new_shadow_root_get_tag_name?: l_new_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs by(unfold_locales) declare l_new_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma new_shadow_root_get_tag_name_is_l_new_shadow_root_get_tag_name [instances]: "l_new_shadow_root_get_tag_name type_wf get_tag_name get_tag_name_locs" using get_tag_name_is_l_get_tag_name apply(auto simp add: l_new_shadow_root_get_tag_name_def l_new_shadow_root_get_tag_name_axioms_def instances)[1] using get_tag_name_new_shadow_root by fast paragraph \new\_character\_data\ locale l_new_character_data_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_tag_name get_tag_name_locs for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, tag_name) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_tag_name_new_character_data: "h \ new_character_data \\<^sub>r new_character_data_ptr \ h \ new_character_data \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" by(auto simp add: CD.get_tag_name_locs_def new_character_data_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t) end locale l_new_character_data_get_tag_name = l_get_tag_name_defs + assumes get_tag_name_new_character_data: "h \ new_character_data \\<^sub>r new_character_data_ptr \ h \ new_character_data \\<^sub>h h' \ r \ get_tag_name_locs ptr' \ r h h'" interpretation i_new_character_data_get_tag_name?: l_new_character_data_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs by unfold_locales declare l_new_character_data_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def[instances] lemma new_character_data_get_tag_name_is_l_new_character_data_get_tag_name [instances]: "l_new_character_data_get_tag_name get_tag_name_locs" unfolding l_new_character_data_get_tag_name_def unfolding get_tag_name_locs_def using new_character_data_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t by blast paragraph \get\_tag\_type\ locale l_set_tag_name_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_tag_name_get_tag_name: assumes "h \ CD.a_set_tag_name element_ptr tag \\<^sub>h h'" shows "h' \ CD.a_get_tag_name element_ptr \\<^sub>r tag" using assms by(auto simp add: CD.a_get_tag_name_def CD.a_set_tag_name_def) lemma set_tag_name_get_tag_name_different_pointers: assumes "ptr \ ptr'" assumes "w \ CD.a_set_tag_name_locs ptr" assumes "h \ w \\<^sub>h h'" assumes "r \ CD.a_get_tag_name_locs ptr'" shows "r h h'" using assms by(auto simp add: all_args_def CD.a_set_tag_name_locs_def CD.a_get_tag_name_locs_def split: if_splits option.splits ) end interpretation i_set_tag_name_get_tag_name?: l_set_tag_name_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs set_tag_name set_tag_name_locs by unfold_locales declare l_set_tag_name_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_tag_name_get_tag_name_is_l_set_tag_name_get_tag_name [instances]: "l_set_tag_name_get_tag_name type_wf get_tag_name get_tag_name_locs set_tag_name set_tag_name_locs" using set_tag_name_is_l_set_tag_name get_tag_name_is_l_get_tag_name apply(simp add: l_set_tag_name_get_tag_name_def l_set_tag_name_get_tag_name_axioms_def) using set_tag_name_get_tag_name set_tag_name_get_tag_name_different_pointers by fast+ subsubsection \attach\_shadow\_root\ locale l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_set_shadow_root_defs set_shadow_root set_shadow_root_locs + l_set_mode_defs set_mode set_mode_locs + l_get_tag_name_defs get_tag_name get_tag_name_locs + l_get_shadow_root_defs get_shadow_root get_shadow_root_locs for set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ ((_) heap, exception, unit) prog" and set_shadow_root_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and set_mode :: "(_) shadow_root_ptr \ shadow_root_mode \ ((_) heap, exception, unit) prog" and set_mode_locs :: "(_) shadow_root_ptr \ ((_) heap, exception, unit) prog set" and get_tag_name :: "(_) element_ptr \ (_, char list) dom_prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_attach_shadow_root :: "(_) element_ptr \ shadow_root_mode \ (_, (_) shadow_root_ptr) dom_prog" where "a_attach_shadow_root element_ptr shadow_root_mode = do { tag \ get_tag_name element_ptr; (if tag \ safe_shadow_root_element_types then error NotSupportedError else return ()); prev_shadow_root \ get_shadow_root element_ptr; (if prev_shadow_root \ None then error NotSupportedError else return ()); new_shadow_root_ptr \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M; set_mode new_shadow_root_ptr shadow_root_mode; set_shadow_root element_ptr (Some new_shadow_root_ptr); return new_shadow_root_ptr }" end locale l_attach_shadow_root_defs = fixes attach_shadow_root :: "(_) element_ptr \ shadow_root_mode \ (_, (_) shadow_root_ptr) dom_prog" global_interpretation l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs set_shadow_root set_shadow_root_locs set_mode set_mode_locs get_tag_name get_tag_name_locs get_shadow_root get_shadow_root_locs defines attach_shadow_root = a_attach_shadow_root . locale l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs set_shadow_root set_shadow_root_locs set_mode set_mode_locs get_tag_name get_tag_name_locs get_shadow_root get_shadow_root_locs + l_attach_shadow_root_defs attach_shadow_root + l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_shadow_root set_shadow_root_locs + l_set_mode type_wf set_mode set_mode_locs + l_get_tag_name type_wf get_tag_name get_tag_name_locs + l_get_shadow_root type_wf get_shadow_root get_shadow_root_locs + l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \ bool" and set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ ((_) heap, exception, unit) prog" and set_shadow_root_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and set_mode :: "(_) shadow_root_ptr \ shadow_root_mode \ ((_) heap, exception, unit) prog" and set_mode_locs :: "(_) shadow_root_ptr \ ((_) heap, exception, unit) prog set" and attach_shadow_root :: "(_) element_ptr \ shadow_root_mode \ ((_) heap, exception, (_) shadow_root_ptr) prog" and type_wf :: "(_) heap \ bool" and get_tag_name :: "(_) element_ptr \ (_, char list) dom_prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" + assumes known_ptr_impl: "known_ptr = a_known_ptr" assumes attach_shadow_root_impl: "attach_shadow_root = a_attach_shadow_root" begin lemmas attach_shadow_root_def = a_attach_shadow_root_def[folded attach_shadow_root_impl] lemma attach_shadow_root_element_ptr_in_heap: assumes "h \ ok (attach_shadow_root element_ptr shadow_root_mode)" shows "element_ptr |\| element_ptr_kinds h" proof - obtain h' where "h \ attach_shadow_root element_ptr shadow_root_mode \\<^sub>h h'" using assms by auto then obtain h2 h3 new_shadow_root_ptr where h2: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h2" and new_shadow_root_ptr: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr" and h3: "h2 \ set_mode new_shadow_root_ptr shadow_root_mode \\<^sub>h h3" and "h3 \ set_shadow_root element_ptr (Some new_shadow_root_ptr) \\<^sub>h h'" by(auto simp add: attach_shadow_root_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_tag_name_pure, rotated] bind_returns_heap_E2[rotated, OF get_shadow_root_pure, rotated] split: if_splits) then have "element_ptr |\| element_ptr_kinds h3" using set_shadow_root_ptr_in_heap by blast moreover have "object_ptr_kinds h2 = object_ptr_kinds h |\| {|cast new_shadow_root_ptr|}" using h2 new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_new_ptr new_shadow_root_ptr by auto 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_mode_writes h3]) using set_mode_pointers_preserved apply blast by (auto simp add: reflp_def transp_def) ultimately show ?thesis by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def) qed lemma create_shadow_root_known_ptr: assumes "h \ attach_shadow_root element_ptr shadow_root_mode \\<^sub>r new_shadow_root_ptr" shows "known_ptr (cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^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_shadow_root_ptr)" using assms by(auto simp add: attach_shadow_root_def known_ptr_impl ShadowRootClass.a_known_ptr_def new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_def new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_def Let_def elim!: bind_returns_result_E) end locale l_attach_shadow_root = l_attach_shadow_root_defs interpretation i_attach_shadow_root?: l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr set_shadow_root set_shadow_root_locs set_mode set_mode_locs attach_shadow_root type_wf get_tag_name get_tag_name_locs get_shadow_root get_shadow_root_locs by(auto simp add: l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def attach_shadow_root_def instances) declare l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_parent\ global_interpretation l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs defines get_parent = "l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_parent get_child_nodes" and get_parent_locs = "l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_parent_locs get_child_nodes_locs" . interpretation i_get_parent?: 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 by(simp add: l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_parent_def get_parent_locs_def instances) declare l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_parent_is_l_get_parent [instances]: "l_get_parent type_wf known_ptr known_ptrs get_parent get_parent_locs get_child_nodes get_child_nodes_locs" apply(simp add: l_get_parent_def l_get_parent_axioms_def instances) using get_parent_reads get_parent_ok get_parent_ptr_in_heap get_parent_pure get_parent_parent_in_heap get_parent_child_dual get_parent_reads_pointers by blast paragraph \set\_disconnected\_nodes\ locale l_set_disconnected_nodes_get_parent\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes_get_child_nodes + l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^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 begin lemma set_disconnected_nodes_get_parent [simp]: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_parent_locs. r h h'))" by(auto simp add: get_parent_locs_def CD.set_disconnected_nodes_locs_def all_args_def) end interpretation i_set_disconnected_nodes_get_parent?: l_set_disconnected_nodes_get_parent\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_disconnected_nodes set_disconnected_nodes_locs get_child_nodes get_child_nodes_locs type_wf DocumentClass.type_wf known_ptr known_ptrs get_parent get_parent_locs by (simp add: l_set_disconnected_nodes_get_parent\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_disconnected_nodes_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_disconnected_nodes_get_parent_is_l_set_disconnected_nodes_get_parent [instances]: "l_set_disconnected_nodes_get_parent set_disconnected_nodes_locs get_parent_locs" by(simp add: l_set_disconnected_nodes_get_parent_def) subsubsection \get\_root\_node\ global_interpretation l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_parent get_parent_locs defines get_root_node = "l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_root_node get_parent" and get_root_node_locs = "l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_root_node_locs get_parent_locs" and get_ancestors = "l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_ancestors get_parent" and get_ancestors_locs = "l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_ancestors_locs get_parent_locs" . declare a_get_ancestors.simps [code] interpretation i_get_root_node?: 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 by(simp add: l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_root_node_def get_root_node_locs_def get_ancestors_def get_ancestors_locs_def instances) declare l_get_root_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_ancestors_is_l_get_ancestors [instances]: "l_get_ancestors get_ancestors" apply(auto simp add: l_get_ancestors_def)[1] using get_ancestors_ptr_in_heap apply fast using get_ancestors_ptr apply fast done lemma get_root_node_is_l_get_root_node [instances]: "l_get_root_node get_root_node get_parent" by (simp add: l_get_root_node_def Shadow_DOM.i_get_root_node.get_root_node_no_parent) subsubsection \get\_root\_node\_si\ locale l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_parent_defs get_parent get_parent_locs + l_get_host_defs get_host get_host_locs for get_parent :: "(_) node_ptr \ ((_) heap, exception, (_::linorder) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" begin partial_function (dom_prog) a_get_ancestors_si :: "(_::linorder) object_ptr \ (_, (_) object_ptr list) dom_prog" where "a_get_ancestors_si ptr = do { check_in_heap ptr; ancestors \ (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 ptr of Some node_ptr \ do { parent_ptr_opt \ get_parent node_ptr; (case parent_ptr_opt of Some parent_ptr \ a_get_ancestors_si parent_ptr | None \ return []) } | None \ (case cast ptr of Some shadow_root_ptr \ do { host \ get_host shadow_root_ptr; a_get_ancestors_si (cast host) } | None \ return [])); return (ptr # ancestors) }" definition "a_get_ancestors_si_locs = get_parent_locs \ get_host_locs" definition a_get_root_node_si :: "(_) object_ptr \ (_, (_) object_ptr) dom_prog" where "a_get_root_node_si ptr = do { ancestors \ a_get_ancestors_si ptr; return (last ancestors) }" definition "a_get_root_node_si_locs = a_get_ancestors_si_locs" end locale l_get_ancestors_si_defs = fixes get_ancestors_si :: "(_::linorder) object_ptr \ (_, (_) object_ptr list) dom_prog" fixes get_ancestors_si_locs :: "((_) heap \ (_) heap \ bool) set" locale l_get_root_node_si_defs = fixes get_root_node_si :: "(_) object_ptr \ (_, (_) object_ptr) dom_prog" fixes get_root_node_si_locs :: "((_) heap \ (_) heap \ bool) set" locale l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_parent + l_get_host + l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_ancestors_si_defs + l_get_root_node_si_defs + assumes get_ancestors_si_impl: "get_ancestors_si = a_get_ancestors_si" assumes get_ancestors_si_locs_impl: "get_ancestors_si_locs = a_get_ancestors_si_locs" assumes get_root_node_si_impl: "get_root_node_si = a_get_root_node_si" assumes get_root_node_si_locs_impl: "get_root_node_si_locs = a_get_root_node_si_locs" begin lemmas get_ancestors_si_def = a_get_ancestors_si.simps[folded get_ancestors_si_impl] lemmas get_ancestors_si_locs_def = a_get_ancestors_si_locs_def[folded get_ancestors_si_locs_impl] lemmas get_root_node_si_def = a_get_root_node_si_def[folded get_root_node_si_impl get_ancestors_si_impl] lemmas get_root_node_si_locs_def = a_get_root_node_si_locs_def[folded get_root_node_si_locs_impl get_ancestors_si_locs_impl] lemma get_ancestors_si_pure [simp]: "pure (get_ancestors_si ptr) h" proof - have "\ptr h h' x. h \ get_ancestors_si ptr \\<^sub>r x \ h \ get_ancestors_si ptr \\<^sub>h h' \ h = h'" proof (induct rule: a_get_ancestors_si.fixp_induct[folded get_ancestors_si_impl]) case 1 then show ?case by(rule admissible_dom_prog) next case 2 then show ?case by simp next case (3 f) then show ?case using get_parent_pure get_host_pure apply(auto simp add: pure_returns_heap_eq pure_def split: option.splits elim!: bind_returns_heap_E bind_returns_result_E dest!: pure_returns_heap_eq[rotated, OF check_in_heap_pure])[1] apply (meson option.simps(3) returns_result_eq) apply(metis get_parent_pure pure_returns_heap_eq) apply(metis get_host_pure pure_returns_heap_eq) done qed then show ?thesis by (meson pure_eq_iff) qed lemma get_root_node_si_pure [simp]: "pure (get_root_node_si ptr) h" by(auto simp add: get_root_node_si_def bind_pure_I) lemma get_ancestors_si_ptr_in_heap: assumes "h \ ok (get_ancestors_si ptr)" shows "ptr |\| object_ptr_kinds h" using assms by(auto simp add: get_ancestors_si_def check_in_heap_ptr_in_heap elim!: bind_is_OK_E dest: is_OK_returns_result_I) lemma get_ancestors_si_ptr: assumes "h \ get_ancestors_si ptr \\<^sub>r ancestors" shows "ptr \ set ancestors" using assms by(simp add: get_ancestors_si_def) (auto elim!: bind_returns_result_E2 split: option.splits intro!: bind_pure_I) lemma get_ancestors_si_never_empty: assumes "h \ get_ancestors_si child \\<^sub>r ancestors" shows "ancestors \ []" using assms apply(simp add: get_ancestors_si_def) by(auto elim!: bind_returns_result_E2 split: option.splits) lemma get_root_node_si_no_parent: "h \ get_parent node_ptr \\<^sub>r None \ h \ get_root_node_si (cast node_ptr) \\<^sub>r cast node_ptr" apply(auto simp add: check_in_heap_def get_root_node_si_def get_ancestors_si_def intro!: bind_pure_returns_result_I )[1] using get_parent_ptr_in_heap by blast lemma get_root_node_si_root_not_shadow_root: assumes "h \ get_root_node_si ptr \\<^sub>r root" shows "\ is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r root" using assms proof(auto simp add: get_root_node_si_def elim!: bind_returns_result_E2) fix y assume "h \ get_ancestors_si ptr \\<^sub>r y" and "is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (last y)" and "root = last y" then show False proof(induct y arbitrary: ptr) case Nil then show ?case using assms(1) get_ancestors_si_never_empty by blast next case (Cons a x) then show ?case apply(auto simp add: get_ancestors_si_def[of ptr] elim!: bind_returns_result_E2 split: option.splits if_splits)[1] using get_ancestors_si_never_empty apply blast using Cons.prems(2) apply auto[1] using \is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (last y)\ \root = last y\ by auto qed qed end global_interpretation l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_parent get_parent_locs get_host get_host_locs defines get_root_node_si = a_get_root_node_si and get_root_node_si_locs = a_get_root_node_si_locs and get_ancestors_si = a_get_ancestors_si and get_ancestors_si_locs = a_get_ancestors_si_locs . declare a_get_ancestors_si.simps [code] interpretation i_get_root_node_si?: l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^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_host get_host_locs get_ancestors_si get_ancestors_si_locs get_root_node_si get_root_node_si_locs apply(auto simp add: l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances)[1] by(auto simp add: get_root_node_si_def get_root_node_si_locs_def get_ancestors_si_def get_ancestors_si_locs_def) declare l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma get_ancestors_si_is_l_get_ancestors [instances]: "l_get_ancestors get_ancestors_si" unfolding l_get_ancestors_def using get_ancestors_si_pure get_ancestors_si_ptr get_ancestors_si_ptr_in_heap by blast lemma get_root_node_si_is_l_get_root_node [instances]: "l_get_root_node get_root_node_si get_parent" apply(simp add: l_get_root_node_def) using get_root_node_si_no_parent by fast paragraph \set\_disconnected\_nodes\ locale l_set_disconnected_nodes_get_ancestors_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes_get_parent + l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_disconnected_nodes_get_host begin lemma set_disconnected_nodes_get_ancestors_si: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_ancestors_si_locs. r h h'))" by(auto simp add: get_parent_locs_def set_disconnected_nodes_locs_def set_disconnected_nodes_get_host get_ancestors_si_locs_def all_args_def) end locale l_set_disconnected_nodes_get_ancestors_si = l_set_disconnected_nodes_defs + l_get_ancestors_si_defs + assumes set_disconnected_nodes_get_ancestors_si: "\w \ set_disconnected_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_ancestors_si_locs. r h h'))" interpretation i_set_disconnected_nodes_get_ancestors_si?: l_set_disconnected_nodes_get_ancestors_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_disconnected_nodes set_disconnected_nodes_locs get_parent get_parent_locs type_wf known_ptr known_ptrs get_child_nodes get_child_nodes_locs get_host get_host_locs get_ancestors_si get_ancestors_si_locs get_root_node_si get_root_node_si_locs DocumentClass.type_wf by (auto simp add: l_set_disconnected_nodes_get_ancestors_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_set_disconnected_nodes_get_ancestors_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_disconnected_nodes_get_ancestors_si_is_l_set_disconnected_nodes_get_ancestors_si [instances]: "l_set_disconnected_nodes_get_ancestors_si set_disconnected_nodes_locs get_ancestors_si_locs" using instances apply(simp add: l_set_disconnected_nodes_get_ancestors_si_def) using set_disconnected_nodes_get_ancestors_si by fast subsubsection \get\_attribute\ lemma get_attribute_is_l_get_attribute [instances]: "l_get_attribute type_wf get_attribute get_attribute_locs" apply(auto simp add: l_get_attribute_def)[1] using i_get_attribute.get_attribute_reads apply fast using type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t i_get_attribute.get_attribute_ok apply blast using i_get_attribute.get_attribute_ptr_in_heap apply fast done subsubsection \to\_tree\_order\ global_interpretation l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs defines to_tree_order = "l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_to_tree_order get_child_nodes" . declare a_to_tree_order.simps [code] interpretation i_to_tree_order?: l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M ShadowRootClass.known_ptr ShadowRootClass.type_wf Shadow_DOM.get_child_nodes Shadow_DOM.get_child_nodes_locs to_tree_order by(auto simp add: l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def to_tree_order_def instances) declare l_to_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \to\_tree\_order\_si\ locale l_to_tree_order_si\<^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_shadow_root_defs get_shadow_root get_shadow_root_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_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin partial_function (dom_prog) a_to_tree_order_si :: "(_) object_ptr \ (_, (_) object_ptr list) dom_prog" where "a_to_tree_order_si ptr = (do { children \ get_child_nodes ptr; shadow_root_part \ (case cast ptr of Some element_ptr \ do { shadow_root_opt \ get_shadow_root element_ptr; (case shadow_root_opt of Some shadow_root_ptr \ return [cast shadow_root_ptr] | None \ return []) } | None \ return []); treeorders \ map_M a_to_tree_order_si ((map (cast) children) @ shadow_root_part); return (ptr # concat treeorders) })" end locale l_to_tree_order_si_defs = fixes to_tree_order_si :: "(_) object_ptr \ (_, (_) object_ptr list) dom_prog" global_interpretation l_to_tree_order_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs defines to_tree_order_si = "a_to_tree_order_si" . declare a_to_tree_order_si.simps [code] locale l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_to_tree_order_si_defs + l_to_tree_order_si\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_child_nodes + l_get_shadow_root + assumes to_tree_order_si_impl: "to_tree_order_si = a_to_tree_order_si" begin lemmas to_tree_order_si_def = a_to_tree_order_si.simps[folded to_tree_order_si_impl] lemma to_tree_order_si_pure [simp]: "pure (to_tree_order_si ptr) h" proof - have "\ptr h h' x. h \ to_tree_order_si ptr \\<^sub>r x \ h \ to_tree_order_si ptr \\<^sub>h h' \ h = h'" proof (induct rule: a_to_tree_order_si.fixp_induct[folded to_tree_order_si_impl]) case 1 then show ?case by (rule admissible_dom_prog) next case 2 then show ?case by simp next case (3 f) then have "\x h. pure (f x) h" by (metis is_OK_returns_heap_E is_OK_returns_result_E pure_def) then have "\xs h. pure (map_M f xs) h" by(rule map_M_pure_I) then show ?case by(auto elim!: bind_returns_heap_E2 split: option.splits) qed then show ?thesis unfolding pure_def by (metis is_OK_returns_heap_E is_OK_returns_result_E) qed end interpretation i_to_tree_order_si?: l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M to_tree_order_si get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs type_wf known_ptr by(auto simp add: l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def to_tree_order_si_def instances) declare l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \first\_in\_tree\_order\ global_interpretation l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs to_tree_order defines first_in_tree_order = "l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_first_in_tree_order to_tree_order" . interpretation i_first_in_tree_order?: l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M to_tree_order first_in_tree_order by(auto simp add: l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def first_in_tree_order_def) declare l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma to_tree_order_is_l_to_tree_order [instances]: "l_to_tree_order to_tree_order" by(auto simp add: l_to_tree_order_def) subsubsection \first\_in\_tree\_order\ global_interpretation l_dummy defines first_in_tree_order_si = "l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_first_in_tree_order to_tree_order_si" . subsubsection \get\_element\_by\ global_interpretation l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs to_tree_order first_in_tree_order get_attribute get_attribute_locs defines get_element_by_id = "l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_element_by_id first_in_tree_order get_attribute" and get_elements_by_class_name = "l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_elements_by_class_name to_tree_order get_attribute" and get_elements_by_tag_name = "l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_elements_by_tag_name to_tree_order" . interpretation i_get_element_by?: l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M to_tree_order first_in_tree_order get_attribute get_attribute_locs get_element_by_id get_elements_by_class_name get_elements_by_tag_name type_wf by(auto simp add: l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_element_by_id_def get_elements_by_class_name_def get_elements_by_tag_name_def instances) declare l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma get_element_by_is_l_get_element_by [instances]: "l_get_element_by get_element_by_id get_elements_by_tag_name to_tree_order" apply(auto simp add: l_get_element_by_def)[1] using get_element_by_id_result_in_tree_order apply fast done subsubsection \get\_element\_by\_si\ global_interpretation l_dummy defines get_element_by_id_si = "l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_element_by_id first_in_tree_order_si get_attribute" and get_elements_by_class_name_si = "l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_elements_by_class_name to_tree_order_si get_attribute" and get_elements_by_tag_name_si = "l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_elements_by_tag_name to_tree_order_si" . subsubsection \find\_slot\ locale l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_parent_defs get_parent get_parent_locs + l_get_shadow_root_defs get_shadow_root get_shadow_root_locs + l_get_mode_defs get_mode get_mode_locs + l_get_attribute_defs get_attribute get_attribute_locs + l_get_tag_name_defs get_tag_name get_tag_name_locs + l_first_in_tree_order_defs first_in_tree_order for get_parent :: "(_) node_ptr \ ((_) heap, exception, (_::linorder) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_mode :: "(_) shadow_root_ptr \ ((_) heap, exception, shadow_root_mode) prog" and get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" and get_attribute :: "(_) element_ptr \ char list \ ((_) heap, exception, char list option) prog" and get_attribute_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and first_in_tree_order :: "(_) object_ptr \ ((_) object_ptr \ ((_) heap, exception, (_) element_ptr option) prog) \ ((_) heap, exception, (_) element_ptr option) prog" begin definition a_find_slot :: "bool \ (_) node_ptr \ (_, (_) element_ptr option) dom_prog" where "a_find_slot open_flag slotable = do { parent_opt \ get_parent slotable; (case parent_opt of Some parent \ if is_element_ptr_kind parent then do { shadow_root_ptr_opt \ get_shadow_root (the (cast parent)); (case shadow_root_ptr_opt of Some shadow_root_ptr \ do { shadow_root_mode \ get_mode shadow_root_ptr; if open_flag \ shadow_root_mode \ Open then return None else first_in_tree_order (cast shadow_root_ptr) (\ptr. if is_element_ptr_kind ptr then do { tag \ get_tag_name (the (cast ptr)); name_attr \ get_attribute (the (cast ptr)) ''name''; slotable_name_attr \ (if is_element_ptr_kind slotable then get_attribute (the (cast slotable)) ''slot'' else return None); (if (tag = ''slot'' \ (name_attr = slotable_name_attr \ (name_attr = None \ slotable_name_attr = Some '''') \ (name_attr = Some '''' \ slotable_name_attr = None))) then return (Some (the (cast ptr))) else return None)} else return None)} | None \ return None)} else return None | _ \ return None)}" definition a_assigned_slot :: "(_) node_ptr \ (_, (_) element_ptr option) dom_prog" where "a_assigned_slot = a_find_slot True" end global_interpretation l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_parent get_parent_locs get_shadow_root get_shadow_root_locs get_mode get_mode_locs get_attribute get_attribute_locs get_tag_name get_tag_name_locs first_in_tree_order defines find_slot = a_find_slot and assigned_slot = a_assigned_slot . locale l_find_slot_defs = fixes find_slot :: "bool \ (_) node_ptr \ (_, (_) element_ptr option) dom_prog" and assigned_slot :: "(_) node_ptr \ (_, (_) element_ptr option) dom_prog" locale l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_find_slot_defs + l_get_parent + l_get_shadow_root + l_get_mode + l_get_attribute + l_get_tag_name + l_to_tree_order + l_first_in_tree_order\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + assumes find_slot_impl: "find_slot = a_find_slot" assumes assigned_slot_impl: "assigned_slot = a_assigned_slot" begin lemmas find_slot_def = find_slot_impl[unfolded a_find_slot_def] lemmas assigned_slot_def = assigned_slot_impl[unfolded a_assigned_slot_def] lemma find_slot_ptr_in_heap: assumes "h \ find_slot open_flag slotable \\<^sub>r slot_opt" shows "slotable |\| node_ptr_kinds h" using assms apply(auto simp add: find_slot_def elim!: bind_returns_result_E2)[1] using get_parent_ptr_in_heap by blast lemma find_slot_slot_in_heap: assumes "h \ find_slot open_flag slotable \\<^sub>r Some slot" shows "slot |\| element_ptr_kinds h" using assms apply(auto simp add: find_slot_def first_in_tree_order_def elim!: bind_returns_result_E2 map_filter_M_pure_E[where y=slot] split: option.splits if_splits list.splits intro!: map_filter_M_pure bind_pure_I)[1] using get_tag_name_ptr_in_heap by blast+ lemma find_slot_pure [simp]: "pure (find_slot open_flag slotable) h" by(auto simp add: find_slot_def first_in_tree_order_def intro!: bind_pure_I map_filter_M_pure split: option.splits list.splits) end interpretation i_find_slot?: l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs get_shadow_root get_shadow_root_locs get_mode get_mode_locs get_attribute get_attribute_locs get_tag_name get_tag_name_locs first_in_tree_order find_slot assigned_slot type_wf known_ptr known_ptrs get_child_nodes get_child_nodes_locs to_tree_order by (auto simp add: find_slot_def assigned_slot_def l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_find_slot\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_find_slot = l_find_slot_defs + assumes find_slot_ptr_in_heap: "h \ find_slot open_flag slotable \\<^sub>r slot_opt \ slotable |\| node_ptr_kinds h" assumes find_slot_slot_in_heap: "h \ find_slot open_flag slotable \\<^sub>r Some slot \ slot |\| element_ptr_kinds h" assumes find_slot_pure [simp]: "pure (find_slot open_flag slotable) h" lemma find_slot_is_l_find_slot [instances]: "l_find_slot find_slot" apply(auto simp add: l_find_slot_def)[1] using find_slot_ptr_in_heap apply fast using find_slot_slot_in_heap apply fast done subsubsection \get\_disconnected\_nodes\ locale l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_get_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_disconnected_nodes :: "(_) document_ptr \ (_, (_) node_ptr list) dom_prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" begin lemma get_disconnected_nodes_ok: "type_wf h \ document_ptr |\| document_ptr_kinds h \ h \ ok (get_disconnected_nodes document_ptr)" apply(unfold type_wf_impl get_disconnected_nodes_impl[unfolded a_get_disconnected_nodes_def]) using CD.get_disconnected_nodes_ok CD.type_wf_impl ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t by blast end interpretation i_get_disconnected_nodes?: l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs by(auto simp add: l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_disconnected_nodes_is_l_get_disconnected_nodes [instances]: "l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs" apply(auto simp add: l_get_disconnected_nodes_def)[1] using i_get_disconnected_nodes.get_disconnected_nodes_reads apply fast using get_disconnected_nodes_ok apply fast using i_get_disconnected_nodes.get_disconnected_nodes_ptr_in_heap apply fast done paragraph \set\_child\_nodes\ locale l_set_child_nodes_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_child_nodes_get_disconnected_nodes: "\w \ set_child_nodes_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" by(auto simp add: set_child_nodes_locs_def CD.set_child_nodes_locs_def CD.get_disconnected_nodes_locs_def all_args_def) end interpretation i_set_child_nodes_get_disconnected_nodes?: l_set_child_nodes_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs apply(auto simp add: l_set_child_nodes_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] by(unfold_locales) declare l_set_child_nodes_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_child_nodes_get_disconnected_nodes_is_l_set_child_nodes_get_disconnected_nodes [instances]: "l_set_child_nodes_get_disconnected_nodes type_wf set_child_nodes set_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs" using set_child_nodes_is_l_set_child_nodes get_disconnected_nodes_is_l_get_disconnected_nodes apply(simp add: l_set_child_nodes_get_disconnected_nodes_def l_set_child_nodes_get_disconnected_nodes_axioms_def) using set_child_nodes_get_disconnected_nodes by fast paragraph \set\_disconnected\_nodes\ lemma set_disconnected_nodes_get_disconnected_nodes_l_set_disconnected_nodes_get_disconnected_nodes [instances]: "l_set_disconnected_nodes_get_disconnected_nodes ShadowRootClass.type_wf 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_def l_set_disconnected_nodes_get_disconnected_nodes_axioms_def instances)[1] using i_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes apply fast using i_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes_different_pointers apply fast done paragraph \delete\_shadow\_root\ locale l_delete_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma get_disconnected_nodes_delete_shadow_root: "h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_disconnected_nodes_locs ptr' \ r h h'" by(auto simp add: CD.get_disconnected_nodes_locs_def delete_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t) end locale l_delete_shadow_root_get_disconnected_nodes = l_get_disconnected_nodes_defs + assumes get_disconnected_nodes_delete_shadow_root: "h \ delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M shadow_root_ptr \\<^sub>h h' \ r \ get_disconnected_nodes_locs ptr' \ r h h'" interpretation l_delete_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs by(auto simp add: l_delete_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) lemma l_delete_shadow_root_get_disconnected_nodes_get_disconnected_nodes_locs [instances]: "l_delete_shadow_root_get_disconnected_nodes get_disconnected_nodes_locs" apply(auto simp add: l_delete_shadow_root_get_disconnected_nodes_def)[1] using get_disconnected_nodes_delete_shadow_root apply fast done paragraph \set\_shadow\_root\ locale l_set_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_shadow_root_get_disconnected_nodes: "\w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" by(auto simp add: set_shadow_root_locs_def CD.get_disconnected_nodes_locs_def all_args_def) end interpretation i_set_shadow_root_get_disconnected_nodes?: l_set_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_shadow_root set_shadow_root_locs DocumentClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs apply(auto simp add: l_set_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] by(unfold_locales) declare l_set_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_shadow_root_get_disconnected_nodes = l_set_shadow_root_defs + l_get_disconnected_nodes_defs + assumes set_shadow_root_get_disconnected_nodes: "\w \ set_shadow_root_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" lemma set_shadow_root_get_disconnected_nodes_is_l_set_shadow_root_get_disconnected_nodes [instances]: "l_set_shadow_root_get_disconnected_nodes set_shadow_root_locs get_disconnected_nodes_locs" using set_shadow_root_is_l_set_shadow_root get_disconnected_nodes_is_l_get_disconnected_nodes apply(simp add: l_set_shadow_root_get_disconnected_nodes_def ) using set_shadow_root_get_disconnected_nodes by fast paragraph \set\_mode\ locale l_set_mode_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_mode\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_mode_get_disconnected_nodes: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" by(auto simp add: set_mode_locs_def CD.get_disconnected_nodes_locs_impl[unfolded CD.a_get_disconnected_nodes_locs_def] all_args_def) end interpretation i_set_mode_get_disconnected_nodes?: l_set_mode_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_mode set_mode_locs DocumentClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs by unfold_locales declare l_set_mode_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_set_mode_get_disconnected_nodes = l_set_mode + l_get_disconnected_nodes + assumes set_mode_get_disconnected_nodes: "\w \ set_mode_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" lemma set_mode_get_disconnected_nodes_is_l_set_mode_get_disconnected_nodes [instances]: "l_set_mode_get_disconnected_nodes type_wf set_mode set_mode_locs get_disconnected_nodes get_disconnected_nodes_locs" using set_mode_is_l_set_mode get_disconnected_nodes_is_l_get_disconnected_nodes apply(simp add: l_set_mode_get_disconnected_nodes_def l_set_mode_get_disconnected_nodes_axioms_def) using set_mode_get_disconnected_nodes by fast paragraph \new\_shadow\_root\ locale l_new_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_disconnected_nodes :: "(_) document_ptr \ (_, (_) node_ptr list) dom_prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" begin lemma get_disconnected_nodes_new_shadow_root: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_disconnected_nodes_locs ptr' \ r h h'" by(auto simp add: CD.get_disconnected_nodes_locs_def new_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t) end interpretation i_new_shadow_root_get_disconnected_nodes?: l_new_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_disconnected_nodes get_disconnected_nodes_locs by unfold_locales declare l_new_shadow_root_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] locale l_new_shadow_root_get_disconnected_nodes = l_get_disconnected_nodes_defs + assumes get_disconnected_nodes_new_shadow_root: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr \ h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h' \ r \ get_disconnected_nodes_locs ptr' \ r h h'" lemma new_shadow_root_get_disconnected_nodes_is_l_new_shadow_root_get_disconnected_nodes [instances]: "l_new_shadow_root_get_disconnected_nodes get_disconnected_nodes_locs" apply (auto simp add: l_new_shadow_root_get_disconnected_nodes_def)[1] using get_disconnected_nodes_new_shadow_root apply fast done subsubsection \remove\_shadow\_root\ locale l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_child_nodes_defs get_child_nodes get_child_nodes_locs + l_get_shadow_root_defs get_shadow_root get_shadow_root_locs + l_set_shadow_root_defs set_shadow_root set_shadow_root_locs + l_get_disconnected_nodes_defs get_disconnected_nodes get_disconnected_nodes_locs for get_child_nodes :: "(_) object_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_child_nodes_locs :: "(_) object_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ ((_) heap, exception, unit) prog" and set_shadow_root_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_remove_shadow_root :: "(_) element_ptr \ (_, unit) dom_prog" where "a_remove_shadow_root element_ptr = do { shadow_root_ptr_opt \ get_shadow_root element_ptr; (case shadow_root_ptr_opt of Some shadow_root_ptr \ do { children \ get_child_nodes (cast shadow_root_ptr); (if children = [] then do { set_shadow_root element_ptr None; delete_M shadow_root_ptr } else do { error HierarchyRequestError }) } | None \ error HierarchyRequestError) }" definition a_remove_shadow_root_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_, unit) dom_prog) set" where "a_remove_shadow_root_locs element_ptr shadow_root_ptr \ set_shadow_root_locs element_ptr \ {delete_M shadow_root_ptr}" end global_interpretation l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs get_disconnected_nodes get_disconnected_nodes_locs defines remove_shadow_root = "a_remove_shadow_root" and remove_shadow_root_locs = a_remove_shadow_root_locs . locale l_remove_shadow_root_defs = fixes remove_shadow_root :: "(_) element_ptr \ (_, unit) dom_prog" fixes remove_shadow_root_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_, unit) dom_prog) set" locale l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_remove_shadow_root_defs + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_child_nodes + l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs + assumes remove_shadow_root_impl: "remove_shadow_root = a_remove_shadow_root" assumes remove_shadow_root_locs_impl: "remove_shadow_root_locs = a_remove_shadow_root_locs" begin lemmas remove_shadow_root_def = remove_shadow_root_impl[unfolded remove_shadow_root_def a_remove_shadow_root_def] lemmas remove_shadow_root_locs_def = remove_shadow_root_locs_impl[unfolded remove_shadow_root_locs_def a_remove_shadow_root_locs_def] lemma remove_shadow_root_writes: "writes (remove_shadow_root_locs element_ptr (the |h \ get_shadow_root element_ptr|\<^sub>r)) (remove_shadow_root element_ptr) h h'" apply(auto simp add: remove_shadow_root_locs_def remove_shadow_root_def all_args_def writes_union_right_I writes_union_left_I set_shadow_root_writes intro!: writes_bind writes_bind_pure[OF get_shadow_root_pure] writes_bind_pure[OF get_child_nodes_pure] intro: writes_subset[OF set_shadow_root_writes] writes_subset[OF writes_singleton2] split: option.splits)[1] using writes_union_left_I[OF set_shadow_root_writes] apply (metis inf_sup_aci(5) insert_is_Un) using writes_union_right_I[OF writes_singleton[of delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M]] - by (smt insert_is_Un writes_singleton2 writes_union_left_I) + by (smt (verit, best) insert_is_Un writes_singleton2 writes_union_left_I) end interpretation i_remove_shadow_root?: l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs get_disconnected_nodes get_disconnected_nodes_locs remove_shadow_root remove_shadow_root_locs type_wf known_ptr by(auto simp add: l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def remove_shadow_root_def remove_shadow_root_locs_def instances) declare l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] paragraph \get\_child\_nodes\ locale l_remove_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma remove_shadow_root_get_child_nodes_different_pointers: assumes "ptr \ cast shadow_root_ptr" assumes "w \ remove_shadow_root_locs element_ptr shadow_root_ptr" assumes "h \ w \\<^sub>h h'" assumes "r \ get_child_nodes_locs ptr" shows "r h h'" using assms apply(auto simp add: all_args_def get_child_nodes_locs_def CD.get_child_nodes_locs_def remove_shadow_root_locs_def set_shadow_root_locs_def delete_shadow_root_get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t delete_shadow_root_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t delete_shadow_root_get_M\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t delete_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t delete_shadow_root_get_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t[rotated] element_put_get_preserved[where setter=shadow_root_opt_update] intro: is_shadow_root_ptr_kind_obtains elim: is_document_ptr_kind_obtains is_shadow_root_ptr_kind_obtains split: if_splits option.splits)[1] done end locale l_remove_shadow_root_get_child_nodes = l_get_child_nodes_defs + l_remove_shadow_root_defs + assumes remove_shadow_root_get_child_nodes_different_pointers: "ptr \ cast shadow_root_ptr \ w \ remove_shadow_root_locs element_ptr shadow_root_ptr \ h \ w \\<^sub>h h' \ r \ get_child_nodes_locs ptr \ r h h'" interpretation i_remove_shadow_root_get_child_nodes?: l_remove_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr get_child_nodes get_child_nodes_locs Core_DOM_Functions.get_child_nodes Core_DOM_Functions.get_child_nodes_locs get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs get_disconnected_nodes get_disconnected_nodes_locs remove_shadow_root remove_shadow_root_locs by(auto simp add: l_remove_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_remove_shadow_root_get_child_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma remove_shadow_root_get_child_nodes_is_l_remove_shadow_root_get_child_nodes [instances]: "l_remove_shadow_root_get_child_nodes get_child_nodes_locs remove_shadow_root_locs" apply(auto simp add: l_remove_shadow_root_get_child_nodes_def instances )[1] using remove_shadow_root_get_child_nodes_different_pointers apply fast done paragraph \get\_tag\_name\ locale l_remove_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma remove_shadow_root_get_tag_name: assumes "w \ remove_shadow_root_locs element_ptr shadow_root_ptr" assumes "h \ w \\<^sub>h h'" assumes "r \ get_tag_name_locs ptr" shows "r h h'" using assms by(auto simp add: all_args_def remove_shadow_root_locs_def set_shadow_root_locs_def CD.get_tag_name_locs_def delete_shadow_root_get_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_put_get_preserved[where setter=shadow_root_opt_update] split: if_splits option.splits) end locale l_remove_shadow_root_get_tag_name = l_get_tag_name_defs + l_remove_shadow_root_defs + assumes remove_shadow_root_get_tag_name: "w \ remove_shadow_root_locs element_ptr shadow_root_ptr \ h \ w \\<^sub>h h' \ r \ get_tag_name_locs ptr \ r h h'" interpretation i_remove_shadow_root_get_tag_name?: l_remove_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf get_tag_name get_tag_name_locs get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs get_disconnected_nodes get_disconnected_nodes_locs remove_shadow_root remove_shadow_root_locs known_ptr by(auto simp add: l_remove_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_remove_shadow_root_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma remove_shadow_root_get_tag_name_is_l_remove_shadow_root_get_tag_name [instances]: "l_remove_shadow_root_get_tag_name get_tag_name_locs remove_shadow_root_locs" apply(auto simp add: l_remove_shadow_root_get_tag_name_def instances )[1] using remove_shadow_root_get_tag_name apply fast done subsubsection \get\_owner\_document\ locale l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_host_defs get_host get_host_locs + CD: l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_root_node get_root_node_locs get_disconnected_nodes get_disconnected_nodes_locs for get_root_node :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) object_ptr) prog" and get_root_node_locs :: "((_) 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_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" begin definition a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) shadow_root_ptr \ unit \ (_, (_) document_ptr) dom_prog" where "a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr _ = do { host \ get_host shadow_root_ptr; CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r (cast host) () }" definition a_get_owner_document_tups :: "(((_) object_ptr \ bool) \ ((_) object_ptr \ unit \ (_, (_) document_ptr) dom_prog)) list" where "a_get_owner_document_tups = [(is_shadow_root_ptr, a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r \ the \ cast)]" definition a_get_owner_document :: "(_::linorder) object_ptr \ (_, (_) document_ptr) dom_prog" where "a_get_owner_document ptr = invoke (CD.a_get_owner_document_tups @ a_get_owner_document_tups) ptr ()" end global_interpretation l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_root_node_si get_root_node_si_locs get_disconnected_nodes get_disconnected_nodes_locs get_host get_host_locs defines get_owner_document_tups = a_get_owner_document_tups and get_owner_document = a_get_owner_document and get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r = a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r and get_owner_document_tups\<^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_defs.a_get_owner_document_tups get_root_node_si get_disconnected_nodes" and get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r = "l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r get_root_node_si get_disconnected_nodes" . locale l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_known_ptr known_ptr + l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_root_node_si get_root_node_si_locs get_disconnected_nodes get_disconnected_nodes_locs get_host get_host_locs + l_get_owner_document_defs get_owner_document + l_get_host get_host get_host_locs + CD: l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_parent get_parent_locs known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_disconnected_nodes get_disconnected_nodes_locs get_root_node_si get_root_node_si_locs get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M for known_ptr :: "(_::linorder) object_ptr \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_::linorder) object_ptr \ bool" and get_parent :: "(_) node_ptr \ ((_) heap, exception, (_) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" 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 get_root_node_si :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr) prog" and get_root_node_si_locs :: "((_) heap \ (_) heap \ bool) set" and get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_owner_document :: "(_) object_ptr \ ((_) heap, exception, (_) document_ptr) prog" + assumes known_ptr_impl: "known_ptr = a_known_ptr" assumes get_owner_document_impl: "get_owner_document = a_get_owner_document" begin lemmas known_ptr_def = known_ptr_impl[unfolded a_known_ptr_def] lemmas get_owner_document_def = a_get_owner_document_def[folded get_owner_document_impl] lemma get_owner_document_pure [simp]: "pure (get_owner_document ptr) h" proof - have "\shadow_root_ptr. pure (a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr ()) h" apply(auto simp add: a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_I filter_M_pure_I split: option.splits)[1] by(auto simp add: CD.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_pure_I filter_M_pure_I split: option.splits) then show ?thesis apply(auto simp add: get_owner_document_def)[1] apply(split CD.get_owner_document_splits, rule conjI)+ apply(simp) apply(auto simp add: a_get_owner_document_tups_def)[1] apply(split invoke_splits, rule conjI)+ apply simp by(auto intro!: bind_pure_I) qed lemma get_owner_document_ptr_in_heap: assumes "h \ ok (get_owner_document ptr)" shows "ptr |\| object_ptr_kinds h" using assms by(auto simp add: get_owner_document_def invoke_ptr_in_heap dest: is_OK_returns_heap_I) end interpretation i_get_owner_document?: l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr DocumentClass.known_ptr get_parent get_parent_locs type_wf get_disconnected_nodes get_disconnected_nodes_locs get_root_node_si get_root_node_si_locs CD.a_get_owner_document get_host get_host_locs get_owner_document by(auto simp add: instances l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_owner_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_owner_document_def Core_DOM_Functions.get_owner_document_def) declare l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_owner_document_is_l_get_owner_document [instances]: "l_get_owner_document get_owner_document" apply(auto simp add: l_get_owner_document_def)[1] using get_owner_document_ptr_in_heap apply fast done subsubsection \remove\_child\ global_interpretation l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs 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 defines remove = "l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_remove get_child_nodes set_child_nodes get_parent get_owner_document get_disconnected_nodes set_disconnected_nodes" and remove_child = "l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_remove_child get_child_nodes set_child_nodes get_owner_document get_disconnected_nodes set_disconnected_nodes" and remove_child_locs = "l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_remove_child_locs set_child_nodes_locs set_disconnected_nodes_locs" . interpretation i_remove_child?: l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M Shadow_DOM.get_child_nodes Shadow_DOM.get_child_nodes_locs Shadow_DOM.set_child_nodes Shadow_DOM.set_child_nodes_locs Shadow_DOM.get_parent Shadow_DOM.get_parent_locs Shadow_DOM.get_owner_document get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs remove_child remove_child_locs remove ShadowRootClass.type_wf ShadowRootClass.known_ptr ShadowRootClass.known_ptrs by(auto simp add: l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def remove_child_def remove_child_locs_def remove_def instances) declare l_remove_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_disconnected\_document\ locale l_get_disconnected_document\<^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 for get_disconnected_nodes :: "(_::linorder) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_get_disconnected_document :: "(_) node_ptr \ (_, (_) document_ptr) dom_prog" where "a_get_disconnected_document node_ptr = do { check_in_heap (cast node_ptr); ptrs \ document_ptr_kinds_M; candidates \ filter_M (\document_ptr. do { disconnected_nodes \ get_disconnected_nodes document_ptr; return (node_ptr \ set disconnected_nodes) }) ptrs; (case candidates of Cons document_ptr [] \ return document_ptr | _ \ error HierarchyRequestError ) }" definition "a_get_disconnected_document_locs = (\document_ptr. get_disconnected_nodes_locs document_ptr) \ (\ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr RObject.nothing)})" end locale l_get_disconnected_document_defs = fixes get_disconnected_document :: "(_) node_ptr \ (_, (_::linorder) document_ptr) dom_prog" fixes get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) set" locale l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_get_disconnected_document_defs + l_get_disconnected_nodes + assumes get_disconnected_document_impl: "get_disconnected_document = a_get_disconnected_document" assumes get_disconnected_document_locs_impl: "get_disconnected_document_locs = a_get_disconnected_document_locs" begin lemmas get_disconnected_document_def = get_disconnected_document_impl[unfolded a_get_disconnected_document_def] lemmas get_disconnected_document_locs_def = get_disconnected_document_locs_impl[unfolded a_get_disconnected_document_locs_def] lemma get_disconnected_document_pure [simp]: "pure (get_disconnected_document ptr) h" using get_disconnected_nodes_pure by(auto simp add: get_disconnected_document_def intro!: bind_pure_I filter_M_pure_I split: list.splits) lemma get_disconnected_document_ptr_in_heap [simp]: "h \ ok (get_disconnected_document node_ptr) \ node_ptr |\| node_ptr_kinds h" using get_disconnected_document_def is_OK_returns_result_I check_in_heap_ptr_in_heap by (metis (no_types, lifting) bind_returns_heap_E get_disconnected_document_pure node_ptr_kinds_commutes pure_pure) lemma get_disconnected_document_disconnected_document_in_heap: assumes "h \ get_disconnected_document child_node \\<^sub>r disconnected_document" shows "disconnected_document |\| document_ptr_kinds h" using assms get_disconnected_nodes_pure by(auto simp add: get_disconnected_document_def elim!: bind_returns_result_E2 dest!: filter_M_not_more_elements[where x=disconnected_document] intro!: filter_M_pure_I bind_pure_I split: if_splits list.splits) lemma get_disconnected_document_reads: "reads get_disconnected_document_locs (get_disconnected_document node_ptr) h h'" using get_disconnected_nodes_reads[unfolded reads_def] by(auto simp add: get_disconnected_document_def get_disconnected_document_locs_def intro!: reads_bind_pure reads_subset[OF check_in_heap_reads] reads_subset[OF error_reads] reads_subset[OF get_disconnected_nodes_reads] reads_subset[OF return_reads] reads_subset[OF document_ptr_kinds_M_reads] filter_M_reads filter_M_pure_I bind_pure_I split: list.splits) end locale l_get_disconnected_document = l_get_disconnected_document_defs + assumes get_disconnected_document_reads: "reads get_disconnected_document_locs (get_disconnected_document node_ptr) h h'" assumes get_disconnected_document_ptr_in_heap: "h \ ok (get_disconnected_document node_ptr) \ node_ptr |\| node_ptr_kinds h" assumes get_disconnected_document_pure [simp]: "pure (get_disconnected_document node_ptr) h" assumes get_disconnected_document_disconnected_document_in_heap: "h \ get_disconnected_document child_node \\<^sub>r disconnected_document \ disconnected_document |\| document_ptr_kinds h" global_interpretation l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_disconnected_nodes get_disconnected_nodes_locs defines get_disconnected_document = a_get_disconnected_document and get_disconnected_document_locs = a_get_disconnected_document_locs . interpretation i_get_disconnected_document?: l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs get_disconnected_document get_disconnected_document_locs type_wf by(auto simp add: l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def get_disconnected_document_def get_disconnected_document_locs_def instances) declare l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma get_disconnected_document_is_l_get_disconnected_document [instances]: "l_get_disconnected_document get_disconnected_document get_disconnected_document_locs" apply(auto simp add: l_get_disconnected_document_def instances)[1] using get_disconnected_document_ptr_in_heap get_disconnected_document_pure get_disconnected_document_disconnected_document_in_heap get_disconnected_document_reads by blast+ paragraph \get\_disconnected\_nodes\ locale l_set_tag_name_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_tag_name_get_disconnected_nodes: "\w \ set_tag_name_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" by(auto simp add: CD.set_tag_name_locs_impl[unfolded CD.a_set_tag_name_locs_def] CD.get_disconnected_nodes_locs_impl[unfolded CD.a_get_disconnected_nodes_locs_def] all_args_def) end interpretation i_set_tag_name_get_disconnected_nodes?: l_set_tag_name_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_tag_name set_tag_name_locs get_disconnected_nodes get_disconnected_nodes_locs by unfold_locales declare l_set_tag_name_get_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_tag_name_get_disconnected_nodes_is_l_set_tag_name_get_disconnected_nodes [instances]: "l_set_tag_name_get_disconnected_nodes type_wf set_tag_name set_tag_name_locs get_disconnected_nodes get_disconnected_nodes_locs" using set_tag_name_is_l_set_tag_name get_disconnected_nodes_is_l_get_disconnected_nodes apply(simp add: l_set_tag_name_get_disconnected_nodes_def l_set_tag_name_get_disconnected_nodes_axioms_def) using set_tag_name_get_disconnected_nodes by fast subsubsection \adopt\_node\ global_interpretation l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs 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 defines adopt_node = a_adopt_node and adopt_node_locs = a_adopt_node_locs . interpretation i_adopt_node?: l_adopt_node\<^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 by(auto simp add: l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def adopt_node_def adopt_node_locs_def instances) declare l_adopt_node\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma adopt_node_is_l_adopt_node [instances]: "l_adopt_node type_wf known_ptr known_ptrs get_parent adopt_node adopt_node_locs get_child_nodes get_owner_document" apply(auto simp add: l_adopt_node_def l_adopt_node_axioms_def instances)[1] using adopt_node_writes apply fast using adopt_node_pointers_preserved apply (fast, fast) using adopt_node_types_preserved apply (fast, fast) using adopt_node_child_in_heap apply fast using adopt_node_children_subset apply fast done paragraph \get\_shadow\_root\ locale l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_child_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_set_disconnected_nodes_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^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 begin lemma adopt_node_get_shadow_root: "\w \ adopt_node_locs parent owner_document document_ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" by(auto simp add: adopt_node_locs_def remove_child_locs_def all_args_def set_disconnected_nodes_get_shadow_root set_child_nodes_get_shadow_root) end locale l_adopt_node_get_shadow_root = l_adopt_node_defs + l_get_shadow_root_defs + assumes adopt_node_get_shadow_root: "\w \ adopt_node_locs parent owner_document document_ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" interpretation i_adopt_node_get_shadow_root?: l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr DocumentClass.type_wf DocumentClass.known_ptr set_child_nodes set_child_nodes_locs Core_DOM_Functions.set_child_nodes Core_DOM_Functions.set_child_nodes_locs get_shadow_root get_shadow_root_locs set_disconnected_nodes set_disconnected_nodes_locs get_owner_document get_parent get_parent_locs remove_child remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs adopt_node adopt_node_locs get_child_nodes get_child_nodes_locs known_ptrs remove by(auto simp add: l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_adopt_node_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma adopt_node_get_shadow_root_is_l_adopt_node_get_shadow_root [instances]: "l_adopt_node_get_shadow_root adopt_node_locs get_shadow_root_locs" apply(auto simp add: l_adopt_node_get_shadow_root_def)[1] using adopt_node_get_shadow_root apply fast done subsubsection \insert\_before\ global_interpretation l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_parent get_parent_locs get_child_nodes get_child_nodes_locs set_child_nodes set_child_nodes_locs get_ancestors_si get_ancestors_si_locs adopt_node adopt_node_locs set_disconnected_nodes set_disconnected_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_owner_document defines next_sibling = a_next_sibling and insert_node = a_insert_node and ensure_pre_insertion_validity = a_ensure_pre_insertion_validity and insert_before = a_insert_before and insert_before_locs = a_insert_before_locs . global_interpretation l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs insert_before defines append_child = "l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_append_child insert_before" . interpretation i_insert_before?: l_insert_before\<^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_si get_ancestors_si_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 by(auto simp add: l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def insert_before_def insert_before_locs_def instances) declare l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] interpretation i_append_child?: l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M append_child insert_before insert_before_locs by(simp add: l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances append_child_def) declare l_append_child\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] subsubsection \get\_assigned\_nodes\ fun map_filter_M2 :: "('x \ ('heap, 'e, 'y option) prog) \ 'x list \ ('heap, 'e, 'y list) prog" where "map_filter_M2 f [] = return []" | "map_filter_M2 f (x # xs) = do { res \ f x; remainder \ map_filter_M2 f xs; return ((case res of Some r \ [r] | None \ []) @ remainder) }" lemma map_filter_M2_pure [simp]: assumes "\x. x \ set xs \ pure (f x) h" shows "pure (map_filter_M2 f xs) h" using assms apply(induct xs arbitrary: h) by(auto elim!: bind_returns_result_E2 intro!: bind_pure_I) lemma map_filter_pure_no_monad: assumes "\x. x \ set xs \ pure (f x) h" assumes "h \ map_filter_M2 f xs \\<^sub>r ys" shows "ys = map the (filter (\x. x \ None) (map (\x. |h \ f x|\<^sub>r) xs))" and "\x. x \ set xs \ h \ ok (f x)" using assms apply(induct xs arbitrary: h ys) by(auto elim!: bind_returns_result_E2) lemma map_filter_pure_foo: assumes "\x. x \ set xs \ pure (f x) h" assumes "h \ map_filter_M2 f xs \\<^sub>r ys" assumes "y \ set ys" obtains x where "h \ f x \\<^sub>r Some y" and "x \ set xs" using assms apply(induct xs arbitrary: ys) by(auto elim!: bind_returns_result_E2) lemma map_filter_M2_in_result: assumes "h \ map_filter_M2 P xs \\<^sub>r ys" assumes "a \ set xs" assumes "\x. x \ set xs \ pure (P x) h" assumes "h \ P a \\<^sub>r Some b" shows "b \ set ys" using assms apply(induct xs arbitrary: h ys) by(auto elim!: bind_returns_result_E2 ) locale l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_tag_name_defs get_tag_name get_tag_name_locs + l_get_root_node_defs get_root_node get_root_node_locs + l_get_host_defs get_host get_host_locs + l_get_child_nodes_defs get_child_nodes get_child_nodes_locs + l_find_slot_defs find_slot assigned_slot + l_remove_defs remove + l_insert_before_defs insert_before insert_before_locs + l_append_child_defs append_child + l_remove_shadow_root_defs remove_shadow_root remove_shadow_root_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_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_root_node :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr) prog" and get_root_node_locs :: "((_) heap \ (_) heap \ bool) set" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and find_slot :: "bool \ (_) node_ptr \ ((_) heap, exception, (_) element_ptr option) prog" and assigned_slot :: "(_) node_ptr \ ((_) heap, exception, (_) element_ptr option) prog" and remove :: "(_) node_ptr \ ((_) heap, exception, unit) prog" and insert_before :: "(_) object_ptr \ (_) node_ptr \ (_) node_ptr option \ ((_) heap, exception, unit) prog" and insert_before_locs :: "(_) object_ptr \ (_) object_ptr option \ (_) document_ptr \ (_) document_ptr \ (_, unit) dom_prog set" and append_child :: "(_) object_ptr \ (_) node_ptr \ ((_) heap, exception, unit) prog" and remove_shadow_root :: "(_) element_ptr \ ((_) heap, exception, unit) prog" and remove_shadow_root_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_) heap, exception, unit) prog set" begin definition a_assigned_nodes :: "(_) element_ptr \ (_, (_) node_ptr list) dom_prog" where "a_assigned_nodes slot = do { tag \ get_tag_name slot; (if tag \ ''slot'' then error HierarchyRequestError else return ()); root \ get_root_node (cast slot); if is_shadow_root_ptr_kind root then do { host \ get_host (the (cast root)); children \ get_child_nodes (cast host); filter_M (\slotable. do { found_slot \ find_slot False slotable; return (found_slot = Some slot)}) children} else return []}" partial_function (dom_prog) a_assigned_nodes_flatten :: "(_) element_ptr \ (_, (_) node_ptr list) dom_prog" where "a_assigned_nodes_flatten slot = do { tag \ get_tag_name slot; (if tag \ ''slot'' then error HierarchyRequestError else return ()); root \ get_root_node (cast slot); (if is_shadow_root_ptr_kind root then do { slotables \ a_assigned_nodes slot; slotables_or_child_nodes \ (if slotables = [] then do { get_child_nodes (cast slot) } else do { return slotables }); list_of_lists \ map_M (\node_ptr. do { (case cast node_ptr of Some element_ptr \ do { tag \ get_tag_name element_ptr; (if tag = ''slot'' then do { root \ get_root_node (cast element_ptr); (if is_shadow_root_ptr_kind root then do { a_assigned_nodes_flatten element_ptr } else do { return [node_ptr] }) } else do { return [node_ptr] }) } | None \ return [node_ptr]) }) slotables_or_child_nodes; return (concat list_of_lists) } else return []) }" definition a_flatten_dom :: "(_, unit) dom_prog" where "a_flatten_dom = do { tups \ element_ptr_kinds_M \ map_filter_M2 (\element_ptr. do { tag \ get_tag_name element_ptr; assigned_nodes \ a_assigned_nodes element_ptr; (if tag = ''slot'' \ assigned_nodes \ [] then return (Some (element_ptr, assigned_nodes)) else return None)}); forall_M (\(slot, assigned_nodes). do { get_child_nodes (cast slot) \ forall_M remove; forall_M (append_child (cast slot)) assigned_nodes }) tups; shadow_root_ptr_kinds_M \ forall_M (\shadow_root_ptr. do { host \ get_host shadow_root_ptr; get_child_nodes (cast host) \ forall_M remove; get_child_nodes (cast shadow_root_ptr) \ forall_M (append_child (cast host)); remove_shadow_root host }); return () }" end global_interpretation l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_tag_name get_tag_name_locs get_root_node get_root_node_locs get_host get_host_locs find_slot assigned_slot remove insert_before insert_before_locs append_child remove_shadow_root remove_shadow_root_locs defines assigned_nodes = a_assigned_nodes and assigned_nodes_flatten = a_assigned_nodes_flatten and flatten_dom = a_flatten_dom . declare a_assigned_nodes_flatten.simps [code] locale l_assigned_nodes_defs = fixes assigned_nodes :: "(_) element_ptr \ (_, (_) node_ptr list) dom_prog" fixes assigned_nodes_flatten :: "(_) element_ptr \ (_, (_) node_ptr list) dom_prog" fixes flatten_dom :: "(_, unit) dom_prog" locale l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_assigned_nodes_defs assigned_nodes assigned_nodes_flatten flatten_dom + l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_tag_name get_tag_name_locs get_root_node get_root_node_locs get_host get_host_locs find_slot assigned_slot remove insert_before insert_before_locs append_child remove_shadow_root remove_shadow_root_locs + l_get_shadow_root type_wf get_shadow_root get_shadow_root_locs + l_set_shadow_root type_wf set_shadow_root set_shadow_root_locs + l_remove + l_insert_before insert_before insert_before_locs + l_find_slot find_slot assigned_slot + l_get_tag_name type_wf get_tag_name get_tag_name_locs + l_get_root_node get_root_node get_root_node_locs get_parent get_parent_locs + l_get_host get_host get_host_locs + l_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_to_tree_order to_tree_order for known_ptr :: "(_::linorder) object_ptr \ bool" and assigned_nodes :: "(_) element_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and assigned_nodes_flatten :: "(_) element_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and flatten_dom :: "((_) heap, exception, unit) prog" 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_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_root_node :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr) prog" and get_root_node_locs :: "((_) heap \ (_) heap \ bool) set" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and find_slot :: "bool \ (_) node_ptr \ ((_) heap, exception, (_) element_ptr option) prog" and assigned_slot :: "(_) node_ptr \ ((_) heap, exception, (_) element_ptr option) prog" and remove :: "(_) node_ptr \ ((_) heap, exception, unit) prog" and insert_before :: "(_) object_ptr \ (_) node_ptr \ (_) node_ptr option \ ((_) heap, exception, unit) prog" and insert_before_locs :: "(_) object_ptr \ (_) object_ptr option \ (_) document_ptr \ (_) document_ptr \ (_, unit) dom_prog set" and append_child :: "(_) object_ptr \ (_) node_ptr \ ((_) heap, exception, unit) prog" and remove_shadow_root :: "(_) element_ptr \ ((_) heap, exception, unit) prog" and remove_shadow_root_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_) heap, exception, unit) prog set" and type_wf :: "(_) heap \ bool" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ ((_) heap, exception, unit) prog" and set_shadow_root_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and get_parent :: "(_) node_ptr \ ((_) heap, exception, (_) object_ptr option) prog" and get_parent_locs :: "((_) heap \ (_) heap \ bool) set" and to_tree_order :: "(_) object_ptr \ ((_) heap, exception, (_) object_ptr list) prog" + assumes assigned_nodes_impl: "assigned_nodes = a_assigned_nodes" assumes flatten_dom_impl: "flatten_dom = a_flatten_dom" begin lemmas assigned_nodes_def = assigned_nodes_impl[unfolded a_assigned_nodes_def] lemmas flatten_dom_def = flatten_dom_impl[unfolded a_flatten_dom_def, folded assigned_nodes_impl] lemma assigned_nodes_pure [simp]: "pure (assigned_nodes slot) h" by(auto simp add: assigned_nodes_def intro!: bind_pure_I filter_M_pure_I) lemma assigned_nodes_ptr_in_heap: assumes "h \ ok (assigned_nodes slot)" shows "slot |\| element_ptr_kinds h" using assms apply(auto simp add: assigned_nodes_def)[1] by (meson bind_is_OK_E is_OK_returns_result_I local.get_tag_name_ptr_in_heap) lemma assigned_nodes_slot_is_slot: assumes "h \ ok (assigned_nodes slot)" shows "h \ get_tag_name slot \\<^sub>r ''slot''" using assms by(auto simp add: assigned_nodes_def elim!: bind_is_OK_E split: if_splits) lemma assigned_nodes_different_ptr: assumes "h \ assigned_nodes slot \\<^sub>r nodes" assumes "h \ assigned_nodes slot' \\<^sub>r nodes'" assumes "slot \ slot'" shows "set nodes \ set nodes' = {}" proof (rule ccontr) assume "set nodes \ set nodes' \ {} " then obtain common_ptr where "common_ptr \ set nodes" and "common_ptr \ set nodes'" by auto have "h \ find_slot False common_ptr \\<^sub>r Some slot" using \common_ptr \ set nodes\ using assms(1) by(auto simp add: assigned_nodes_def elim!: bind_returns_result_E2 split: if_splits dest!: filter_M_holds_for_result[where x=common_ptr] intro!: bind_pure_I) moreover have "h \ find_slot False common_ptr \\<^sub>r Some slot'" using \common_ptr \ set nodes'\ using assms(2) by(auto simp add: assigned_nodes_def elim!: bind_returns_result_E2 split: if_splits dest!: filter_M_holds_for_result[where x=common_ptr] intro!: bind_pure_I) ultimately show False using assms(3) by (meson option.inject returns_result_eq) qed end interpretation i_assigned_nodes?: l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr assigned_nodes assigned_nodes_flatten flatten_dom get_child_nodes get_child_nodes_locs get_tag_name get_tag_name_locs get_root_node get_root_node_locs get_host get_host_locs find_slot assigned_slot remove insert_before insert_before_locs append_child remove_shadow_root remove_shadow_root_locs type_wf get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs get_parent get_parent_locs to_tree_order by(auto simp add: instances l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def assigned_nodes_def flatten_dom_def) declare l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_assigned_nodes = l_assigned_nodes_defs + assumes assigned_nodes_pure [simp]: "pure (assigned_nodes slot) h" assumes assigned_nodes_ptr_in_heap: "h \ ok (assigned_nodes slot) \ slot |\| element_ptr_kinds h" assumes assigned_nodes_slot_is_slot: "h \ ok (assigned_nodes slot) \ h \ get_tag_name slot \\<^sub>r ''slot''" assumes assigned_nodes_different_ptr: "h \ assigned_nodes slot \\<^sub>r nodes \ h \ assigned_nodes slot' \\<^sub>r nodes' \ slot \ slot' \ set nodes \ set nodes' = {}" lemma assigned_nodes_is_l_assigned_nodes [instances]: "l_assigned_nodes assigned_nodes" apply(auto simp add: l_assigned_nodes_def)[1] using assigned_nodes_ptr_in_heap apply fast using assigned_nodes_slot_is_slot apply fast using assigned_nodes_different_ptr apply fast done subsubsection \set\_val\ locale l_set_val\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_set_val\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_val set_val_locs + l_type_wf type_wf for type_wf :: "(_) heap \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and set_val :: "(_) character_data_ptr \ char list \ (_, unit) dom_prog" and set_val_locs :: "(_) character_data_ptr \ (_, unit) dom_prog set" + assumes type_wf_impl: "type_wf = ShadowRootClass.type_wf" begin lemma set_val_ok: "type_wf h \ character_data_ptr |\| character_data_ptr_kinds h \ h \ ok (set_val character_data_ptr tag)" using CD.set_val_ok CD.type_wf_impl ShadowRootClass.type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t local.type_wf_impl by blast lemma set_val_writes: "writes (set_val_locs character_data_ptr) (set_val character_data_ptr tag) h h'" using CD.set_val_writes . lemma set_val_pointers_preserved: assumes "w \ set_val_locs character_data_ptr" assumes "h \ w \\<^sub>h h'" shows "object_ptr_kinds h = object_ptr_kinds h'" using assms CD.set_val_pointers_preserved by simp lemma set_val_typess_preserved: assumes "w \ set_val_locs character_data_ptr" assumes "h \ w \\<^sub>h h'" shows "type_wf h = type_wf h'" apply(unfold type_wf_impl) using assms(1) type_wf_preserved[OF writes_singleton2 assms(2)] by(auto simp add: all_args_def CD.set_val_locs_impl[unfolded CD.a_set_val_locs_def] split: if_splits) end interpretation i_set_val?: l_set_val\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_val set_val_locs apply(unfold_locales) by (auto simp add: set_val_def set_val_locs_def) declare l_set_val\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_val_is_l_set_val [instances]: "l_set_val type_wf set_val set_val_locs" apply(simp add: l_set_val_def) using set_val_ok set_val_writes set_val_pointers_preserved set_val_typess_preserved by blast paragraph \get\_shadow\_root\ locale l_set_val_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_val\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_val_get_shadow_root: "\w \ set_val_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" by(auto simp add: CD.set_val_locs_impl[unfolded CD.a_set_val_locs_def] get_shadow_root_locs_def all_args_def) end locale l_set_val_get_shadow_root = l_set_val + l_get_shadow_root + assumes set_val_get_shadow_root: "\w \ set_val_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_shadow_root_locs ptr'. r h h'))" interpretation i_set_val_get_shadow_root?: l_set_val_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_val set_val_locs get_shadow_root get_shadow_root_locs apply(auto simp add: l_set_val_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] using l_set_val\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms by unfold_locales declare l_set_val_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_val_get_shadow_root_is_l_set_val_get_shadow_root [instances]: "l_set_val_get_shadow_root type_wf set_val set_val_locs get_shadow_root get_shadow_root_locs" using set_val_is_l_set_val get_shadow_root_is_l_get_shadow_root apply(simp add: l_set_val_get_shadow_root_def l_set_val_get_shadow_root_axioms_def) using set_val_get_shadow_root by fast paragraph \get\_tag\_type\ locale l_set_val_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_val\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma set_val_get_tag_name: "\w \ set_val_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" by(auto simp add: CD.set_val_locs_impl[unfolded CD.a_set_val_locs_def] CD.get_tag_name_locs_impl[unfolded CD.a_get_tag_name_locs_def] all_args_def) end locale l_set_val_get_tag_name = l_set_val + l_get_tag_name + assumes set_val_get_tag_name: "\w \ set_val_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_tag_name_locs ptr'. r h h'))" interpretation i_set_val_get_tag_name?: l_set_val_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf DocumentClass.type_wf set_val set_val_locs get_tag_name get_tag_name_locs by unfold_locales declare l_set_val_get_tag_name\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_val_get_tag_name_is_l_set_val_get_tag_name [instances]: "l_set_val_get_tag_name type_wf set_val set_val_locs get_tag_name get_tag_name_locs" using set_val_is_l_set_val get_tag_name_is_l_get_tag_name apply(simp add: l_set_val_get_tag_name_def l_set_val_get_tag_name_axioms_def) using set_val_get_tag_name by fast subsubsection \create\_character\_data\ locale l_create_character_data\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ _ _ _ _ _ type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" + assumes known_ptr_impl: "known_ptr = a_known_ptr" begin lemma create_character_data_document_in_heap: assumes "h \ ok (create_character_data document_ptr text)" shows "document_ptr |\| document_ptr_kinds h" using assms CD.create_character_data_document_in_heap by simp lemma create_character_data_known_ptr: assumes "h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr" shows "known_ptr (cast new_character_data_ptr)" using assms CD.create_character_data_known_ptr by(simp add: known_ptr_impl CD.known_ptr_impl ShadowRootClass.a_known_ptr_def) end locale l_create_character_data = l_create_character_data_defs interpretation i_create_character_data?: l_create_character_data\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^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 create_character_data known_ptr DocumentClass.type_wf DocumentClass.known_ptr by(auto simp add: l_create_character_data\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_create_character_data\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \create\_element\ locale l_create_element\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD: 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\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M create_element known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_known_ptr known_ptr for get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap \ (_) heap \ bool) set" and set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" and set_tag_name :: "(_) element_ptr \ char list \ ((_) heap, exception, unit) prog" and set_tag_name_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" and type_wf :: "(_) heap \ bool" and create_element :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) element_ptr) prog" and known_ptr :: "(_) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) object_ptr \ bool" + assumes known_ptr_impl: "known_ptr = a_known_ptr" begin lemmas create_element_def = CD.create_element_def lemma create_element_document_in_heap: assumes "h \ ok (create_element document_ptr tag)" shows "document_ptr |\| document_ptr_kinds h" using CD.create_element_document_in_heap assms . lemma create_element_known_ptr: assumes "h \ create_element document_ptr tag \\<^sub>r new_element_ptr" shows "known_ptr (cast new_element_ptr)" proof - have "is_element_ptr new_element_ptr" using assms apply(auto simp add: create_element_def elim!: bind_returns_result_E)[1] using new_element_is_element_ptr by blast then show ?thesis by(auto simp add: known_ptr_impl known_ptr_defs DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs) qed end interpretation i_create_element?: l_create_element\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^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 DocumentClass.type_wf DocumentClass.known_ptr by(auto simp add: l_create_element\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_create_element\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_create_element\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] subsection \A wellformed heap (Core DOM)\ subsubsection \wellformed\_heap\ locale l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = CD: 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_get_shadow_root_defs get_shadow_root get_shadow_root_locs + l_get_tag_name_defs get_tag_name get_tag_name_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" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" begin definition a_host_shadow_root_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" where "a_host_shadow_root_rel h = (\(x, y). (cast x, cast y)) ` {(host, shadow_root). host |\| element_ptr_kinds h \ |h \ get_shadow_root host|\<^sub>r = Some shadow_root}" lemma a_host_shadow_root_rel_code [code]: "a_host_shadow_root_rel h = set (concat (map (\host. (case |h \ get_shadow_root host|\<^sub>r of Some shadow_root \ [(cast host, cast shadow_root)] | None \ [])) (sorted_list_of_fset (element_ptr_kinds h))) )" by(auto simp add: a_host_shadow_root_rel_def) definition a_all_ptrs_in_heap :: "(_) heap \ bool" where "a_all_ptrs_in_heap h = ((\host shadow_root_ptr. (h \ get_shadow_root host \\<^sub>r Some shadow_root_ptr) \ shadow_root_ptr |\| shadow_root_ptr_kinds h))" definition a_distinct_lists :: "(_) heap \ bool" where "a_distinct_lists h = distinct (concat ( map (\element_ptr. (case |h \ get_shadow_root element_ptr|\<^sub>r of Some shadow_root_ptr \ [shadow_root_ptr] | None \ [])) |h \ element_ptr_kinds_M|\<^sub>r ))" definition a_shadow_root_valid :: "(_) heap \ bool" where "a_shadow_root_valid h = (\shadow_root_ptr \ fset (shadow_root_ptr_kinds h). (\host \ fset(element_ptr_kinds h). |h \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr))" definition a_heap_is_wellformed :: "(_) heap \ bool" where "a_heap_is_wellformed h \ CD.a_heap_is_wellformed h \ acyclic (CD.a_parent_child_rel h \ a_host_shadow_root_rel h) \ a_all_ptrs_in_heap h \ a_distinct_lists h \ a_shadow_root_valid h" end global_interpretation l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs defines heap_is_wellformed = a_heap_is_wellformed and parent_child_rel = CD.a_parent_child_rel and host_shadow_root_rel = a_host_shadow_root_rel and all_ptrs_in_heap = a_all_ptrs_in_heap and distinct_lists = a_distinct_lists and shadow_root_valid = a_shadow_root_valid and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_heap_is_wellformed and parent_child_rel\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_parent_child_rel and acyclic_heap\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_acyclic_heap and all_ptrs_in_heap\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_all_ptrs_in_heap and distinct_lists\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_distinct_lists and owner_document_valid\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = CD.a_owner_document_valid . interpretation i_heap_is_wellformed\<^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\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel by (auto simp add: l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def parent_child_rel_def instances) declare i_heap_is_wellformed\<^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_axioms[instances] lemma heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_heap_is_wellformed [instances]: "l_heap_is_wellformed type_wf known_ptr heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_child_nodes get_disconnected_nodes" apply(auto simp add: l_heap_is_wellformed_def)[1] using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_children_in_heap apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_disc_nodes_in_heap apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_one_parent apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_one_disc_parent apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_children_disc_nodes_different apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_disconnected_nodes_distinct apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_children_distinct apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_is_wellformed_children_disc_nodes apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_child apply (blast, blast) using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_finite apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_acyclic apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_node_ptr apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_parent_in_heap apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_child_in_heap apply blast done locale l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs + CD: 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\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel + l_heap_is_wellformed_defs heap_is_wellformed parent_child_rel + l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_shadow_root get_shadow_root_locs get_host get_host_locs type_wf + l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs get_disconnected_document get_disconnected_document_locs type_wf + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root get_shadow_root_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" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and known_ptr :: "(_) object_ptr \ bool" and type_wf :: "(_) heap \ bool" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_document :: "(_) node_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) set" + assumes heap_is_wellformed_impl: "heap_is_wellformed = a_heap_is_wellformed" begin lemmas heap_is_wellformed_def = heap_is_wellformed_impl[unfolded a_heap_is_wellformed_def, folded CD.heap_is_wellformed_impl CD.parent_child_rel_impl] lemma a_distinct_lists_code [code]: "a_all_ptrs_in_heap h = ((\host \ fset (element_ptr_kinds h). h \ ok (get_shadow_root host) \ (case |h \ get_shadow_root host|\<^sub>r of Some shadow_root_ptr \ shadow_root_ptr |\| shadow_root_ptr_kinds h | None \ True)))" apply(auto simp add: a_all_ptrs_in_heap_def split: option.splits)[1] by (meson is_OK_returns_result_I local.get_shadow_root_ptr_in_heap notin_fset select_result_I2) lemma get_shadow_root_shadow_root_ptr_in_heap: assumes "heap_is_wellformed h" assumes "h \ get_shadow_root host \\<^sub>r Some shadow_root_ptr" shows "shadow_root_ptr |\| shadow_root_ptr_kinds h" using assms by(auto simp add: heap_is_wellformed_def a_all_ptrs_in_heap_def) lemma get_host_ptr_in_heap: assumes "heap_is_wellformed h" assumes "h \ get_host shadow_root_ptr \\<^sub>r host" shows "shadow_root_ptr |\| shadow_root_ptr_kinds h" using assms get_shadow_root_shadow_root_ptr_in_heap by(auto simp add: get_host_def elim!: bind_returns_result_E2 dest!: filter_M_holds_for_result intro!: bind_pure_I split: list.splits) lemma shadow_root_same_host: assumes "heap_is_wellformed h" and "type_wf h" assumes "h \ get_shadow_root host \\<^sub>r Some shadow_root_ptr" assumes "h \ get_shadow_root host' \\<^sub>r Some shadow_root_ptr" shows "host = host'" proof (rule ccontr) assume " host \ host'" have "host |\| element_ptr_kinds h" using assms(3) by (meson is_OK_returns_result_I local.get_shadow_root_ptr_in_heap) moreover have "host' |\| element_ptr_kinds h" using assms(4) by (meson is_OK_returns_result_I local.get_shadow_root_ptr_in_heap) ultimately show False using assms apply(auto simp add: heap_is_wellformed_def a_distinct_lists_def)[1] apply(drule distinct_concat_map_E(1)[where x=host and y=host']) apply(simp) apply(simp) using \host \ host'\ apply(simp) apply(auto)[1] done qed lemma shadow_root_host_dual: assumes "h \ get_host shadow_root_ptr \\<^sub>r host" shows "h \ get_shadow_root host \\<^sub>r Some shadow_root_ptr" using assms by(auto simp add: get_host_def dest: filter_M_holds_for_result elim!: bind_returns_result_E2 intro!: bind_pure_I split: list.splits) lemma disc_doc_disc_node_dual: assumes "h \ get_disconnected_document disc_node \\<^sub>r disc_doc" obtains disc_nodes where "h \ get_disconnected_nodes disc_doc \\<^sub>r disc_nodes" and "disc_node \ set disc_nodes" using assms get_disconnected_nodes_pure by(auto simp add: get_disconnected_document_def bind_pure_I dest!: filter_M_holds_for_result elim!: bind_returns_result_E2 intro!: filter_M_pure_I split: if_splits list.splits) lemma get_host_valid_tag_name: assumes "heap_is_wellformed h" and "type_wf h" assumes "h \ get_host shadow_root_ptr \\<^sub>r host" assumes "h \ get_tag_name host \\<^sub>r tag" shows "tag \ safe_shadow_root_element_types" proof - obtain host' where "host' |\| element_ptr_kinds h" and "|h \ get_tag_name host'|\<^sub>r \ safe_shadow_root_element_types" and "h \ get_shadow_root host' \\<^sub>r Some shadow_root_ptr" using assms - apply(auto simp add: heap_is_wellformed_def a_shadow_root_valid_def)[1] - by (smt assms(1) finite_set_in get_host_ptr_in_heap local.get_shadow_root_ok - returns_result_select_result) + by (metis finite_set_in get_host_ptr_in_heap local.a_shadow_root_valid_def + local.get_shadow_root_ok local.heap_is_wellformed_def returns_result_select_result) then have "host = host'" by (meson assms(1) assms(2) assms(3) shadow_root_host_dual shadow_root_same_host) then show ?thesis - by (smt\\thesis. (\host'. \host' |\| element_ptr_kinds h; - |h \ get_tag_name host'|\<^sub>r \ safe_shadow_root_element_types; - h \ get_shadow_root host' \\<^sub>r Some shadow_root_ptr\ \ thesis) \ thesis\ - \h \ get_shadow_root host' \\<^sub>r Some shadow_root_ptr\ assms(1) assms(2) assms(4) select_result_I2 - shadow_root_same_host) + using \|h \ get_tag_name host'|\<^sub>r \ safe_shadow_root_element_types\ assms(4) by auto qed lemma a_host_shadow_root_rel_finite: "finite (a_host_shadow_root_rel h)" proof - have "a_host_shadow_root_rel h = (\host \ fset (element_ptr_kinds h). (case |h \ get_shadow_root host|\<^sub>r of Some shadow_root \ {(cast host, cast shadow_root)} | None \ {}))" by(auto simp add: a_host_shadow_root_rel_def split: option.splits) moreover have "finite (\host \ fset (element_ptr_kinds h). (case |h \ get_shadow_root host|\<^sub>r of Some shadow_root \ {(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 host, cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^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 shadow_root)} | None \ {}))" by(auto split: option.splits) ultimately show ?thesis by auto qed lemma 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" using CD.heap_is_wellformed_children_in_heap local.heap_is_wellformed_def by blast 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" using CD.heap_is_wellformed_disc_nodes_in_heap local.heap_is_wellformed_def by blast lemma 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'" using CD.heap_is_wellformed_one_parent local.heap_is_wellformed_def by blast 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 CD.heap_is_wellformed_one_disc_parent local.heap_is_wellformed_def by blast 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 = {}" using CD.heap_is_wellformed_children_disc_nodes_different local.heap_is_wellformed_def by blast 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 CD.heap_is_wellformed_disconnected_nodes_distinct local.heap_is_wellformed_def by blast lemma heap_is_wellformed_children_distinct: "heap_is_wellformed h \ h \ get_child_nodes ptr \\<^sub>r children \ distinct children" using CD.heap_is_wellformed_children_distinct local.heap_is_wellformed_def by blast 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)" using CD.heap_is_wellformed_children_disc_nodes local.heap_is_wellformed_def by blast lemma parent_child_rel_finite: "heap_is_wellformed h \ finite (parent_child_rel h)" using CD.parent_child_rel_finite by blast lemma parent_child_rel_acyclic: "heap_is_wellformed h \ acyclic (parent_child_rel h)" using CD.parent_child_rel_acyclic heap_is_wellformed_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" using CD.parent_child_rel_child_in_heap local.heap_is_wellformed_def by blast end interpretation i_heap_is_wellformed?: l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs by(auto simp add: l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def parent_child_rel_def heap_is_wellformed_def instances) declare l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma heap_is_wellformed_is_l_heap_is_wellformed [instances]: "l_heap_is_wellformed ShadowRootClass.type_wf ShadowRootClass.known_ptr Shadow_DOM.heap_is_wellformed Shadow_DOM.parent_child_rel Shadow_DOM.get_child_nodes get_disconnected_nodes" apply(auto simp add: l_heap_is_wellformed_def instances)[1] using heap_is_wellformed_children_in_heap apply metis using heap_is_wellformed_disc_nodes_in_heap apply metis 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 metis using heap_is_wellformed_children_distinct apply metis using heap_is_wellformed_children_disc_nodes apply metis using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_child apply(blast, blast) using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_finite apply blast using parent_child_rel_acyclic apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_node_ptr apply blast using i_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_parent_in_heap apply blast using parent_child_rel_child_in_heap apply metis done subsubsection \get\_parent\ interpretation i_get_parent_wf\<^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\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_disconnected_nodes by(simp add: l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare i_get_parent_wf\<^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_axioms[instances] interpretation i_get_parent_wf2\<^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 known_ptr type_wf get_child_nodes get_child_nodes_locs known_ptrs get_parent get_parent_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs by(auto simp add: l_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare i_get_parent_wf2\<^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_axioms[instances] lemma get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_get_parent_wf [instances]: "l_get_parent_wf type_wf known_ptr known_ptrs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_child_nodes get_parent" apply(auto simp add: l_get_parent_wf_def l_get_parent_wf_axioms_def instances)[1] using i_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.child_parent_dual apply fast using i_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_wellformed_induct apply metis using i_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.heap_wellformed_induct_rev apply metis using i_get_parent_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_parent apply fast done subsubsection \get\_disconnected\_nodes\ subsubsection \set\_disconnected\_nodes\ paragraph \get\_disconnected\_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\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_child_nodes 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 instances) declare 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_axioms[instances] lemma set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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 i_set_disconnected_nodes_get_disconnected_nodes_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_from_disconnected_nodes_removes apply fast done paragraph \get\_root\_node\ interpretation i_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 known_ptr type_wf known_ptrs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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 by(simp add: l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare i_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_axioms[instances] lemma get_ancestors_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_get_ancestors_wf [instances]: "l_get_ancestors_wf heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel known_ptr known_ptrs type_wf get_ancestors get_ancestors_locs get_child_nodes get_parent" apply(auto simp add: l_get_ancestors_wf_def l_get_ancestors_wf_axioms_def instances)[1] using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_never_empty apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_ok apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_reads apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_ptrs_in_heap apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_remains_not_in_ancestors apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_also_parent apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_obtains_children apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_parent_child_rel apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_parent_child_rel apply blast done lemma get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_get_root_node_wf [instances]: "l_get_root_node_wf heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_root_node type_wf known_ptr known_ptrs get_ancestors get_parent" apply(auto simp add: l_get_root_node_wf_def l_get_root_node_wf_axioms_def instances)[1] using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_root_node_ok apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_root_node_ptr_in_heap apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_root_node_root_in_heap apply blast using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_ancestors_same_root_node apply(blast, blast) using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_root_node_same_no_parent apply blast (* using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_root_node_not_node_same apply blast *) using i_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_root_node_parent_same apply (blast, blast) done subsubsection \to\_tree\_order\ interpretation i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: 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\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs apply(simp add: l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) done declare i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_to_tree_order_wf [instances]: "l_to_tree_order_wf heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel type_wf known_ptr known_ptrs to_tree_order get_parent get_child_nodes" apply(auto simp add: l_to_tree_order_wf_def l_to_tree_order_wf_axioms_def instances)[1] using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_ok apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_ptrs_in_heap apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_parent_child_rel apply(blast, blast) using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_child2 apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_node_ptrs apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_child apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_ptr_in_result apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_parent apply blast using i_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_subset apply blast done paragraph \get\_root\_node\ interpretation i_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: 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\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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 by(auto 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 instances) declare i_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_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\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M" 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 instances)[1] using i_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_get_root_node apply blast using i_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.to_tree_order_same_root apply blast done subsubsection \remove\_child\ interpretation i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M: 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\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M parent_child_rel by unfold_locales declare i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.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\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_is_l_remove_child_wf2 [instances]: "l_remove_child_wf2 type_wf known_ptr known_ptrs remove_child heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes remove" apply(auto simp add: l_remove_child_wf2_def l_remove_child_wf2_axioms_def instances)[1] using i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_child_heap_is_wellformed_preserved apply(fast, fast, fast) using i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_heap_is_wellformed_preserved apply(fast, fast, fast) using i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_child_removes_child apply fast using i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_child_removes_first_child apply fast using i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_removes_child apply fast using i_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.remove_for_all_empty_children apply fast done subsection \A wellformed heap\ subsubsection \get\_parent\ 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] lemma get_parent_wf_is_l_get_parent_wf [instances]: "l_get_parent_wf ShadowRootClass.type_wf ShadowRootClass.known_ptr ShadowRootClass.known_ptrs heap_is_wellformed parent_child_rel Shadow_DOM.get_child_nodes Shadow_DOM.get_parent" apply(auto simp add: l_get_parent_wf_def l_get_parent_wf_axioms_def instances)[1] using child_parent_dual apply blast using heap_wellformed_induct apply metis using heap_wellformed_induct_rev apply metis using parent_child_rel_parent apply metis done subsubsection \remove\_shadow\_root\ locale l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name + l_get_disconnected_nodes + l_set_shadow_root_get_tag_name + l_get_child_nodes + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_remove_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_delete_shadow_root_get_disconnected_nodes + l_delete_shadow_root_get_child_nodes + l_set_shadow_root_get_disconnected_nodes + l_set_shadow_root_get_child_nodes + l_delete_shadow_root_get_tag_name + l_set_shadow_root_get_shadow_root + l_delete_shadow_root_get_shadow_root + l_get_parent\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma remove_shadow_root_preserves: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove_shadow_root ptr \\<^sub>h h'" shows "known_ptrs h'" and "type_wf h'" "heap_is_wellformed h'" proof - obtain shadow_root_ptr h2 where "h \ get_shadow_root ptr \\<^sub>r Some shadow_root_ptr" and "h \ get_child_nodes (cast shadow_root_ptr) \\<^sub>r []" and h2: "h \ set_shadow_root ptr None \\<^sub>h h2" and h': "h2 \ delete_M shadow_root_ptr \\<^sub>h h'" using assms(4) by(auto simp add: remove_shadow_root_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_shadow_root_pure, rotated] bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] split: option.splits if_splits) have "type_wf h2" using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF set_shadow_root_writes h2] using \type_wf h\ set_shadow_root_types_preserved by(auto simp add: reflp_def transp_def) then show "type_wf h'" using h' delete_shadow_root_type_wf_preserved local.type_wf_impl by blast have object_ptr_kinds_eq_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 set_shadow_root_writes h2]) using set_shadow_root_pointers_preserved apply blast by (auto simp add: reflp_def transp_def) have node_ptr_kinds_eq_h: "node_ptr_kinds h = node_ptr_kinds h2" using object_ptr_kinds_eq_h by (simp add: node_ptr_kinds_def) have element_ptr_kinds_eq_h: "element_ptr_kinds h = element_ptr_kinds h2" using node_ptr_kinds_eq_h by (simp add: element_ptr_kinds_def) have document_ptr_kinds_eq_h: "document_ptr_kinds h = document_ptr_kinds h2" using object_ptr_kinds_eq_h by (simp add: document_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h: "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h2" using object_ptr_kinds_eq_h by (simp add: shadow_root_ptr_kinds_def) have "known_ptrs h2" using \known_ptrs h\ object_ptr_kinds_eq_h known_ptrs_subset by blast have object_ptr_kinds_eq_h2: "object_ptr_kinds h' |\| object_ptr_kinds h2" using h' delete_shadow_root_pointers by auto have object_ptr_kinds_eq2_h2: "object_ptr_kinds h2 = object_ptr_kinds h' |\| {|cast shadow_root_ptr|}" using h' delete_shadow_root_pointers by auto have node_ptr_kinds_eq_h2: "node_ptr_kinds h2 = node_ptr_kinds h'" using object_ptr_kinds_eq_h2 by(auto simp add: node_ptr_kinds_def delete_shadow_root_pointers[OF h']) have element_ptr_kinds_eq_h2: "element_ptr_kinds h2 = element_ptr_kinds h'" using node_ptr_kinds_eq_h2 by (simp add: element_ptr_kinds_def) have document_ptr_kinds_eq_h2: "document_ptr_kinds h2 = document_ptr_kinds h'" using object_ptr_kinds_eq_h2 by(auto simp add: document_ptr_kinds_def delete_shadow_root_pointers[OF h']) have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h' |\| shadow_root_ptr_kinds h2" using object_ptr_kinds_eq_h2 by (auto simp add: shadow_root_ptr_kinds_def) have shadow_root_ptr_kinds_eq2_h2: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h' |\| {|shadow_root_ptr|}" using object_ptr_kinds_eq2_h2 apply (auto simp add: shadow_root_ptr_kinds_def)[1] by (metis \h \ get_shadow_root ptr \\<^sub>r Some shadow_root_ptr\ assms(1) fset.map_comp local.get_shadow_root_shadow_root_ptr_in_heap object_ptr_kinds_eq_h shadow_root_ptr_kinds_def) show "known_ptrs h'" using object_ptr_kinds_eq_h2 \known_ptrs h2\ known_ptrs_subset 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 set_shadow_root_writes h2 set_shadow_root_get_disconnected_nodes by(rule reads_writes_preserved) 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 disconnected_nodes_eq_h2: "\doc_ptr disc_nodes. h2 \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes = h' \ get_disconnected_nodes doc_ptr \\<^sub>r disc_nodes" using get_disconnected_nodes_reads get_disconnected_nodes_delete_shadow_root[OF h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have disconnected_nodes_eq2_h2: "\doc_ptr. |h2 \ get_disconnected_nodes doc_ptr|\<^sub>r = |h' \ get_disconnected_nodes doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h: "\doc_ptr disc_nodes. h \ get_tag_name doc_ptr \\<^sub>r disc_nodes = h2 \ get_tag_name doc_ptr \\<^sub>r disc_nodes" using get_tag_name_reads set_shadow_root_writes h2 set_shadow_root_get_tag_name by(rule reads_writes_preserved) then have tag_name_eq2_h: "\doc_ptr. |h \ get_tag_name doc_ptr|\<^sub>r = |h2 \ get_tag_name doc_ptr|\<^sub>r" using select_result_eq by force have tag_name_eq_h2: "\doc_ptr disc_nodes. h2 \ get_tag_name doc_ptr \\<^sub>r disc_nodes = h' \ get_tag_name doc_ptr \\<^sub>r disc_nodes" using get_tag_name_reads get_tag_name_delete_shadow_root[OF h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have tag_name_eq2_h2: "\doc_ptr. |h2 \ get_tag_name doc_ptr|\<^sub>r = |h' \ get_tag_name doc_ptr|\<^sub>r" using select_result_eq by force have children_eq_h: "\ptr' children. h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads set_shadow_root_writes h2 set_shadow_root_get_child_nodes by(rule reads_writes_preserved) then have children_eq2_h: "\ptr'. |h \ get_child_nodes ptr'|\<^sub>r = |h2 \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have children_eq_h2: "\ptr' children. ptr' \ cast shadow_root_ptr \ h2 \ get_child_nodes ptr' \\<^sub>r children = h' \ get_child_nodes ptr' \\<^sub>r children" using get_child_nodes_reads h' get_child_nodes_delete_shadow_root apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have children_eq2_h2: "\ptr'. ptr' \ cast shadow_root_ptr \ |h2 \ get_child_nodes ptr'|\<^sub>r = |h' \ get_child_nodes ptr'|\<^sub>r" using select_result_eq by force have "cast shadow_root_ptr |\| object_ptr_kinds h'" using h' delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_not_in_heap by auto have get_shadow_root_eq_h: "\shadow_root_opt ptr'. ptr \ ptr' \ h \ get_shadow_root ptr' \\<^sub>r shadow_root_opt = h2 \ get_shadow_root ptr' \\<^sub>r shadow_root_opt" using get_shadow_root_reads set_shadow_root_writes h2 apply(rule reads_writes_preserved) using set_shadow_root_get_shadow_root_different_pointers by fast have get_shadow_root_eq_h2: "\shadow_root_opt ptr'. h2 \ get_shadow_root ptr' \\<^sub>r shadow_root_opt = h' \ get_shadow_root ptr' \\<^sub>r shadow_root_opt" using get_shadow_root_reads get_shadow_root_delete_shadow_root[OF h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have get_shadow_root_eq2_h2: "\ptr'. |h2 \ get_shadow_root ptr'|\<^sub>r = |h' \ get_shadow_root ptr'|\<^sub>r" using select_result_eq by force have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) moreover have "parent_child_rel h = parent_child_rel h2" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h children_eq2_h) moreover have "parent_child_rel h' \ parent_child_rel h2" using object_ptr_kinds_eq_h2 apply(auto simp add: CD.parent_child_rel_def)[1] by (metis \cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^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 shadow_root_ptr |\| object_ptr_kinds h'\ children_eq2_h2) ultimately have "CD.a_acyclic_heap h'" using acyclic_subset by (auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) moreover have "CD.a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_all_ptrs_in_heap h2" by(auto simp add: children_eq2_h disconnected_nodes_eq2_h document_ptr_kinds_eq_h CD.a_all_ptrs_in_heap_def object_ptr_kinds_eq_h node_ptr_kinds_def children_eq_h disconnected_nodes_eq_h) then have "CD.a_all_ptrs_in_heap h'" apply(auto simp add: CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 children_eq_h2 disconnected_nodes_eq_h2)[1] apply(case_tac "ptr = cast shadow_root_ptr") using object_ptr_kinds_eq_h2 children_eq_h2 apply (meson \cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^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 shadow_root_ptr |\| object_ptr_kinds h'\ is_OK_returns_result_I local.get_child_nodes_ptr_in_heap) apply (metis (no_types, lifting) children_eq2_h2 fin_mono finite_set_in object_ptr_kinds_eq_h2 subsetD) by (metis (full_types) assms(1) assms(2) disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq_h node_ptr_kinds_eq_h2 returns_result_select_result) moreover have "CD.a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_distinct_lists h2" by(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h children_eq2_h disconnected_nodes_eq2_h) then have "CD.a_distinct_lists h'" apply(auto simp add: CD.a_distinct_lists_def document_ptr_kinds_eq_h2 disconnected_nodes_eq2_h2)[1] apply(auto simp add: intro!: distinct_concat_map_I)[1] apply(case_tac "x = cast shadow_root_ptr") using \cast shadow_root_ptr |\| object_ptr_kinds h'\ apply simp using children_eq_h2 concat_map_all_distinct[of "(\ptr. |h2 \ get_child_nodes ptr|\<^sub>r)"] apply (metis (no_types, lifting) children_eq2_h2 finite_fset fmember.rep_eq fset_mp object_ptr_kinds_eq_h2 set_sorted_list_of_set) apply(case_tac "x = cast shadow_root_ptr") using \cast shadow_root_ptr |\| object_ptr_kinds h'\ apply simp apply(case_tac "y = cast shadow_root_ptr") using \cast shadow_root_ptr |\| object_ptr_kinds h'\ apply simp using children_eq_h2 distinct_concat_map_E(1)[of "(\ptr. |h2 \ get_child_nodes ptr|\<^sub>r)"] - apply (smt IntI children_eq2_h2 empty_iff finite_fset fmember.rep_eq fset_mp - object_ptr_kinds_eq_h2 set_sorted_list_of_set) + apply (metis (no_types, lifting) IntI \known_ptrs h'\ \type_wf h'\ children_eq2_h2 empty_iff + finite_fset finite_set_in is_OK_returns_result_I l_get_child_nodes.get_child_nodes_ptr_in_heap + local.get_child_nodes_ok local.known_ptrs_known_ptr local.l_get_child_nodes_axioms + returns_result_select_result sorted_list_of_set.set_sorted_key_list_of_set) apply(case_tac "xa = cast shadow_root_ptr") using \cast shadow_root_ptr |\| object_ptr_kinds h'\ apply simp - by (smt \local.CD.a_distinct_lists h2\ \type_wf h'\ children_eq2_h2 disconnected_nodes_eq_h2 - fset_mp is_OK_returns_result_E local.CD.distinct_lists_no_parent - local.get_disconnected_nodes_ok object_ptr_kinds_eq_h2 select_result_I2) - + by (metis (mono_tags, lifting) CD.distinct_lists_no_parent \known_ptrs h'\ + \local.CD.a_distinct_lists h2\ \type_wf h'\ children_eq2_h2 children_eq_h2 + disconnected_nodes_eq_h2 is_OK_returns_result_I l_get_child_nodes.get_child_nodes_ptr_in_heap + local.get_child_nodes_ok local.get_disconnected_nodes_ok local.known_ptrs_known_ptr + local.l_get_child_nodes_axioms returns_result_select_result) moreover have "CD.a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h2" by(auto simp add: CD.a_owner_document_valid_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h node_ptr_kinds_eq_h children_eq2_h disconnected_nodes_eq2_h) then have "CD.a_owner_document_valid h'" apply(auto simp add: CD.a_owner_document_valid_def document_ptr_kinds_eq_h2 node_ptr_kinds_eq_h2 disconnected_nodes_eq2_h2)[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 h2 \ node_ptr \ set |h2 \ 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 have "\parent_ptr. parent_ptr |\| object_ptr_kinds h2 \ node_ptr \ set |h2 \ get_child_nodes parent_ptr|\<^sub>r" apply(auto)[1] apply(case_tac "parent_ptr = cast shadow_root_ptr") using \h \ get_child_nodes (cast shadow_root_ptr) \\<^sub>r []\ children_eq_h apply auto[1] using children_eq2_h2 object_ptr_kinds_eq_h2 h' delete_shadow_root_pointers by (metis fempty_iff finsert_iff funionE) then show "\document_ptr. document_ptr |\| document_ptr_kinds h' \ node_ptr \ set |h' \ get_disconnected_nodes document_ptr|\<^sub>r" using 0 1 by auto qed ultimately have "heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'" by(simp add: CD.heap_is_wellformed_def) moreover have "acyclic (parent_child_rel h \ a_host_shadow_root_rel h)" using \heap_is_wellformed h\ by(simp add: heap_is_wellformed_def) then have "acyclic (parent_child_rel h2 \ a_host_shadow_root_rel h2)" proof - have "a_host_shadow_root_rel h2 \ a_host_shadow_root_rel h" apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h)[1] apply(case_tac "aa = ptr") apply(simp) apply (metis (no_types, lifting) \type_wf h2\ assms(2) h2 local.get_shadow_root_ok local.type_wf_impl option.distinct(1) returns_result_eq returns_result_select_result set_shadow_root_get_shadow_root) using get_shadow_root_eq_h by (metis (mono_tags, lifting) \type_wf h2\ image_eqI is_OK_returns_result_E local.get_shadow_root_ok mem_Collect_eq prod.simps(2) select_result_I2) then show ?thesis using \parent_child_rel h = parent_child_rel h2\ by (metis (no_types, opaque_lifting) \acyclic (parent_child_rel h \ a_host_shadow_root_rel h)\ acyclic_subset order_refl sup_mono) qed then have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h')" proof - have "a_host_shadow_root_rel h' \ a_host_shadow_root_rel h2" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 get_shadow_root_eq2_h2) then show ?thesis using \parent_child_rel h' \ parent_child_rel h2\ \acyclic (parent_child_rel h2 \ a_host_shadow_root_rel h2)\ using acyclic_subset sup_mono by (metis (no_types, opaque_lifting)) qed moreover 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 shadow_root_ptr_kinds_eq_h)[1] apply(case_tac "host = ptr") apply(simp) apply (metis assms(2) h2 local.type_wf_impl option.distinct(1) returns_result_eq set_shadow_root_get_shadow_root) using get_shadow_root_eq_h by fastforce then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def get_shadow_root_eq_h2)[1] apply(auto simp add: shadow_root_ptr_kinds_eq2_h2)[1] by (metis (no_types, lifting) \h \ get_shadow_root ptr \\<^sub>r Some shadow_root_ptr\ assms(1) assms(2) get_shadow_root_eq_h get_shadow_root_eq_h2 h2 local.shadow_root_same_host local.type_wf_impl option.distinct(1) select_result_I2 set_shadow_root_get_shadow_root) moreover have "a_distinct_lists h" using \heap_is_wellformed h\ by(simp add: heap_is_wellformed_def) then have "a_distinct_lists h2" apply(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h)[1] apply(auto intro!: distinct_concat_map_I split: option.splits)[1] apply(case_tac "x = ptr") apply(simp) apply (metis (no_types, opaque_lifting) assms(2) h2 is_OK_returns_result_I l_set_shadow_root_get_shadow_root.set_shadow_root_get_shadow_root l_set_shadow_root_get_shadow_root_axioms local.type_wf_impl option.discI returns_result_eq returns_result_select_result) apply(case_tac "y = ptr") apply(simp) apply (metis (no_types, opaque_lifting) assms(2) h2 is_OK_returns_result_I l_set_shadow_root_get_shadow_root.set_shadow_root_get_shadow_root l_set_shadow_root_get_shadow_root_axioms local.type_wf_impl option.discI returns_result_eq returns_result_select_result) by (metis \type_wf h2\ assms(1) assms(2) get_shadow_root_eq_h local.get_shadow_root_ok local.shadow_root_same_host returns_result_select_result) then have "a_distinct_lists h'" by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h2 get_shadow_root_eq2_h2) moreover have "a_shadow_root_valid h" using \heap_is_wellformed h\ by(simp add: heap_is_wellformed_def) then have "a_shadow_root_valid h'" - apply(auto simp add: a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h element_ptr_kinds_eq_h - tag_name_eq2_h)[1] - apply(simp add: shadow_root_ptr_kinds_eq2_h2 element_ptr_kinds_eq_h2 tag_name_eq2_h2) - using get_shadow_root_eq_h get_shadow_root_eq_h2 - by (smt \cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^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 shadow_root_ptr |\| object_ptr_kinds h'\ - \h \ get_shadow_root ptr \\<^sub>r Some shadow_root_ptr\ assms(2) element_ptr_kinds_eq_h - element_ptr_kinds_eq_h2 finite_set_in local.get_shadow_root_ok option.inject - returns_result_select_result select_result_I2 shadow_root_ptr_kinds_commutes) - + by (smt (verit) \h \ get_shadow_root ptr \\<^sub>r Some shadow_root_ptr\ assms(2) + delete\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_not_in_heap element_ptr_kinds_eq_h element_ptr_kinds_eq_h2 + finite_set_in finsert_iff funion_finsert_right get_shadow_root_eq2_h2 get_shadow_root_eq_h h' + local.a_shadow_root_valid_def local.get_shadow_root_ok object_ptr_kinds_eq2_h2 + object_ptr_kinds_eq_h option.sel returns_result_select_result select_result_I2 + shadow_root_ptr_kinds_commutes sup_bot.right_neutral tag_name_eq2_h tag_name_eq2_h2) ultimately show "heap_is_wellformed h'" by(simp add: heap_is_wellformed_def) qed end interpretation i_remove_shadow_root_wf?: l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_tag_name get_tag_name_locs get_disconnected_nodes get_disconnected_nodes_locs set_shadow_root set_shadow_root_locs known_ptr get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs remove_shadow_root remove_shadow_root_locs known_ptrs get_parent get_parent_locs by(auto simp add: l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_root\_node\ 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 by(simp add: l_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) 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" apply(auto simp add: l_get_ancestors_wf_def l_get_ancestors_wf_axioms_def instances)[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_not_node_same apply blast *) using get_root_node_parent_same apply (blast, blast) done subsubsection \get\_parent\_get\_host\ locale l_get_parent_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf + l_get_shadow_root + l_get_host + l_get_child_nodes begin lemma host_shadow_root_rel_finite: "finite (a_host_shadow_root_rel h)" proof - have "a_host_shadow_root_rel h = (\host \ fset (element_ptr_kinds h). (case |h \ get_shadow_root host|\<^sub>r of Some shadow_root \ {(cast host, cast shadow_root)} | None \ {}))" by(auto simp add: a_host_shadow_root_rel_def split: option.splits) moreover have "finite (\host \ fset (element_ptr_kinds h). (case |h \ get_shadow_root host|\<^sub>r of Some shadow_root \ {(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 host, cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^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 shadow_root)} | None \ {}))" by(auto split: option.splits) ultimately show ?thesis by auto qed lemma host_shadow_root_rel_shadow_root: "h \ get_shadow_root host \\<^sub>r shadow_root_option \ shadow_root_option = Some shadow_root \ ((cast host, cast shadow_root) \ a_host_shadow_root_rel h)" apply(auto simp add: a_host_shadow_root_rel_def)[1] by(metis (mono_tags, lifting) case_prodI is_OK_returns_result_I l_get_shadow_root.get_shadow_root_ptr_in_heap local.l_get_shadow_root_axioms mem_Collect_eq pair_imageI select_result_I2) lemma host_shadow_root_rel_host: "heap_is_wellformed h \ h \ get_host shadow_root \\<^sub>r host \ (cast host, cast shadow_root) \ a_host_shadow_root_rel h" apply(auto simp add: a_host_shadow_root_rel_def)[1] using shadow_root_host_dual by (metis (no_types, lifting) Collect_cong host_shadow_root_rel_shadow_root local.a_host_shadow_root_rel_def split_cong) lemma heap_wellformed_induct_si [consumes 1, case_names step]: assumes "heap_is_wellformed h" assumes "\parent. (\children child. h \ get_child_nodes parent \\<^sub>r children \ child \ set children \ P (cast child)) \ (\shadow_root host. parent = cast host \ h \ get_shadow_root host \\<^sub>r Some shadow_root \ P (cast shadow_root)) \ P parent" shows "P ptr" proof - fix ptr have "finite (parent_child_rel h \ a_host_shadow_root_rel h)" using host_shadow_root_rel_finite using local.CD.parent_child_rel_finite local.CD.parent_child_rel_impl by auto then have "wf ((parent_child_rel h \ a_host_shadow_root_rel h)\)" using assms(1) apply(simp add: heap_is_wellformed_def) by (simp add: finite_acyclic_wf_converse local.CD.parent_child_rel_impl) then show "?thesis" proof (induct rule: wf_induct_rule) case (less parent) then show ?case apply(auto)[1] using assms host_shadow_root_rel_shadow_root local.CD.parent_child_rel_child by blast qed qed lemma heap_wellformed_induct_rev_si [consumes 1, case_names step]: assumes "heap_is_wellformed h" assumes "\child. (\parent child_node. child = cast child_node \ h \ get_parent child_node \\<^sub>r Some parent \ P parent) \ (\host shadow_root. child = cast shadow_root \ h \ get_host shadow_root \\<^sub>r host \ P (cast host)) \ P child" shows "P ptr" proof - fix ptr have "finite (parent_child_rel h \ a_host_shadow_root_rel h)" using host_shadow_root_rel_finite using local.CD.parent_child_rel_finite local.CD.parent_child_rel_impl by auto then have "wf (parent_child_rel h \ a_host_shadow_root_rel h)" using assms(1) apply(simp add: heap_is_wellformed_def) by (simp add: finite_acyclic_wf) then show "?thesis" proof (induct rule: wf_induct_rule) case (less parent) then show ?case apply(auto)[1] using parent_child_rel_parent host_shadow_root_rel_host using assms(1) assms(2) by auto qed qed end interpretation i_get_parent_get_host_wf?: l_get_parent_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs known_ptrs get_parent get_parent_locs by(auto simp add: l_get_parent_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_parent_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_get_parent_get_host_wf = l_heap_is_wellformed_defs + l_get_parent_defs + l_get_shadow_root_defs + l_get_host_defs + l_get_child_nodes_defs + assumes heap_wellformed_induct_si [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)) \ (\shadow_root host. parent = cast host \ h \ get_shadow_root host \\<^sub>r Some shadow_root \ P (cast shadow_root)) \ P parent) \ P ptr" assumes heap_wellformed_induct_rev_si [consumes 1, case_names step]: "heap_is_wellformed h \ (\child. (\parent child_node. child = cast child_node \ h \ get_parent child_node \\<^sub>r Some parent \ P parent) \ (\host shadow_root. child = cast shadow_root \ h \ get_host shadow_root \\<^sub>r host \ P (cast host)) \ P child) \ P ptr" lemma l_get_parent_get_host_wf_is_get_parent_get_host_wf [instances]: "l_get_parent_get_host_wf heap_is_wellformed get_parent get_shadow_root get_host get_child_nodes" apply(auto simp add: l_get_parent_get_host_wf_def instances)[1] using heap_wellformed_induct_si apply metis using heap_wellformed_induct_rev_si apply blast done subsubsection \get\_host\ locale l_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs + l_type_wf type_wf + l_get_host\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_shadow_root get_shadow_root_locs get_host get_host_locs type_wf + l_get_shadow_root type_wf get_shadow_root get_shadow_root_locs for known_ptr :: "(_::linorder) object_ptr \ bool" and known_ptrs :: "(_) heap \ bool" and type_wf :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and 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" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" begin lemma get_host_ok [simp]: assumes "heap_is_wellformed h" assumes "type_wf h" assumes "known_ptrs h" assumes "shadow_root_ptr |\| shadow_root_ptr_kinds h" shows "h \ ok (get_host shadow_root_ptr)" proof - obtain host where host: "host |\| element_ptr_kinds h" and "|h \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types" and shadow_root: "h \ get_shadow_root host \\<^sub>r Some shadow_root_ptr" using assms(1) assms(4) get_shadow_root_ok assms(2) apply(auto simp add: heap_is_wellformed_def a_shadow_root_valid_def)[1] by (smt finite_set_in returns_result_select_result) obtain host_candidates where host_candidates: "h \ filter_M (\element_ptr. Heap_Error_Monad.bind (get_shadow_root element_ptr) (\shadow_root_opt. return (shadow_root_opt = Some shadow_root_ptr))) (sorted_list_of_set (fset (element_ptr_kinds h))) \\<^sub>r host_candidates" apply(drule is_OK_returns_result_E[rotated]) using get_shadow_root_ok assms(2) by(auto intro!: filter_M_is_OK_I bind_pure_I bind_is_OK_I2) then have "host_candidates = [host]" apply(rule filter_M_ex1) apply (simp add: host) - apply (smt assms(1) assms(2) bind_pure_returns_result_I2 bind_returns_result_E finite_set_in host - local.get_shadow_root_ok local.get_shadow_root_pure local.shadow_root_same_host - return_returns_result returns_result_eq shadow_root sorted_list_of_fset.rep_eq - sorted_list_of_fset_simps(1)) + apply (smt (verit) assms(1) assms(2) bind_pure_returns_result_I2 finite_fset finite_set_in + host local.get_shadow_root_ok local.get_shadow_root_pure local.shadow_root_same_host + return_returns_result returns_result_eq shadow_root sorted_list_of_set(1)) by (simp_all add: assms(2) bind_pure_I bind_pure_returns_result_I2 host local.get_shadow_root_ok returns_result_eq shadow_root) then show ?thesis using host_candidates host assms(1) get_shadow_root_ok apply(auto simp add: get_host_def known_ptrs_known_ptr intro!: bind_is_OK_pure_I filter_M_pure_I filter_M_is_OK_I bind_pure_I split: list.splits)[1] using assms(2) apply blast apply (meson list.distinct(1) returns_result_eq) by (meson list.distinct(1) list.inject returns_result_eq) qed end interpretation i_get_host_wf?: l_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document get_disconnected_document_locs known_ptr known_ptrs type_wf get_host get_host_locs get_shadow_root get_shadow_root_locs get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M by(auto simp add: l_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] locale l_get_host_wf = l_heap_is_wellformed_defs + l_known_ptrs + l_type_wf + l_get_host_defs + assumes get_host_ok: "heap_is_wellformed h \ known_ptrs h \ type_wf h \ shadow_root_ptr |\| shadow_root_ptr_kinds h \ h \ ok (get_host shadow_root_ptr)" lemma get_host_wf_is_l_get_host_wf [instances]: "l_get_host_wf heap_is_wellformed known_ptr known_ptrs type_wf get_host" by(auto simp add: l_get_host_wf_def l_get_host_wf_axioms_def instances) subsubsection \get\_root\_node\_si\ locale l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_root_node_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf + l_get_parent_get_host_wf + l_get_host_wf begin lemma get_root_node_si_ptr_in_heap: assumes "h \ ok (get_root_node_si ptr)" shows "ptr |\| object_ptr_kinds h" using assms unfolding get_root_node_si_def using get_ancestors_si_ptr_in_heap by auto lemma get_ancestors_si_ok: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" and "ptr |\| object_ptr_kinds h" shows "h \ ok (get_ancestors_si ptr)" proof (insert assms(1) assms(4), induct rule: heap_wellformed_induct_rev_si) case (step child) then show ?case using assms(2) assms(3) apply(auto simp add: get_ancestors_si_def[of child] assms(1) get_parent_parent_in_heap intro!: bind_is_OK_pure_I split: option.splits)[1] using local.get_parent_ok apply blast using get_host_ok assms(1) apply blast by (meson assms(1) is_OK_returns_result_I local.get_shadow_root_ptr_in_heap local.shadow_root_host_dual) qed lemma get_ancestors_si_remains_not_in_ancestors: assumes "heap_is_wellformed h" and "heap_is_wellformed h'" and "h \ get_ancestors_si ptr \\<^sub>r ancestors" and "h' \ get_ancestors_si 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 "\p shadow_root_option shadow_root_option'. h \ get_shadow_root p \\<^sub>r shadow_root_option \ h' \ get_shadow_root p \\<^sub>r shadow_root_option' \ (if shadow_root_option = None then shadow_root_option' = None else shadow_root_option' = None \ shadow_root_option' = shadow_root_option)" 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(7), induct ptr arbitrary: ancestors ancestors' rule: heap_wellformed_induct_rev_si) case (step child) obtain ancestors_remains where ancestors_remains: "ancestors = child # ancestors_remains" using \h \ get_ancestors_si child \\<^sub>r ancestors\ get_ancestors_si_never_empty by(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2 split: option.splits) obtain ancestors_remains' where ancestors_remains': "ancestors' = child # ancestors_remains'" using \h' \ get_ancestors_si child \\<^sub>r ancestors'\ get_ancestors_si_never_empty by(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2 split: option.splits) have "child |\| object_ptr_kinds h" using local.get_ancestors_si_ptr_in_heap object_ptr_kinds_eq3 step.prems(2) by fastforce have "node \ child" using ancestors_remains step.prems(3) by auto 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 2: "\p host. h' \ get_host p \\<^sub>r host \ h \ get_host p \\<^sub>r host" proof - fix p host assume "h' \ get_host p \\<^sub>r host" then have "h' \ get_shadow_root host \\<^sub>r Some p" using local.shadow_root_host_dual by blast then have "h \ get_shadow_root host \\<^sub>r Some p" by (metis assms(6) element_ptr_kinds_commutes is_OK_returns_result_I local.get_shadow_root_ok local.get_shadow_root_ptr_in_heap node_ptr_kinds_commutes object_ptr_kinds_eq3 option.distinct(1) returns_result_select_result type_wf) then show "h \ get_host p \\<^sub>r host" by (metis assms(1) is_OK_returns_result_E known_ptrs local.get_host_ok local.get_shadow_root_shadow_root_ptr_in_heap local.shadow_root_host_dual local.shadow_root_same_host type_wf) qed 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 child") case None then show ?thesis using step(3) step(4) \node \ child\ apply(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2 split: option.splits)[1] by (metis "2" assms(1) l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M.shadow_root_same_host list.set_intros(2) local.l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms local.shadow_root_host_dual step.hyps(2) step.prems(3) type_wf) next case (Some node_child) then show ?thesis using step(3) step(4) \node \ child\ apply(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2 split: option.splits)[1] apply (meson "1" option.distinct(1) returns_result_eq) by (metis "1" list.set_intros(2) option.inject returns_result_eq step.hyps(1) step.prems(3)) qed qed qed lemma get_ancestors_si_ptrs_in_heap: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_ancestors_si 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_si x \\<^sub>r a # ancestors" by(auto simp add: get_ancestors_si_def[of a] elim!: bind_returns_result_E2 split: option.splits) then have "x = a" by(auto simp add: get_ancestors_si_def[of x] elim!: bind_returns_result_E2 split: option.splits) then show ?case proof (cases "ptr' = a") case True then show ?thesis using Cons.hyps Cons.prems(2) get_ancestors_si_ptr_in_heap x using \x = a\ by blast next case False obtain ptr'' where ptr'': "h \ get_ancestors_si ptr'' \\<^sub>r ancestors" using \ h \ get_ancestors_si x \\<^sub>r a # ancestors\ Cons.prems(2) False by(auto simp add: get_ancestors_si_def[of x] elim!: bind_returns_result_E2 split: option.splits) then show ?thesis using Cons.hyps Cons.prems(2) False by auto qed qed lemma get_ancestors_si_reads: assumes "heap_is_wellformed h" shows "reads get_ancestors_si_locs (get_ancestors_si node_ptr) h h'" proof (insert assms(1), induct rule: heap_wellformed_induct_rev_si) case (step child) then show ?case using [[simproc del: Product_Type.unit_eq]] get_parent_reads[unfolded reads_def] get_host_reads[unfolded reads_def] apply(simp (no_asm) add: get_ancestors_si_def) by(auto simp add: get_ancestors_si_locs_def get_parent_reads_pointers intro!: reads_bind_pure reads_subset[OF check_in_heap_reads] reads_subset[OF return_reads] reads_subset[OF get_parent_reads] reads_subset[OF get_child_nodes_reads] reads_subset[OF get_host_reads] split: option.splits) qed lemma get_ancestors_si_subset: assumes "heap_is_wellformed h" and "h \ get_ancestors_si ptr \\<^sub>r ancestors" and "ancestor \ set ancestors" and "h \ get_ancestors_si 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_si) case (step child) have "child |\| object_ptr_kinds h" using get_ancestors_si_ptr_in_heap step(3) by auto (* then have "h \ check_in_heap child \\<^sub>r ()" using returns_result_select_result by force *) obtain tl_ancestors where tl_ancestors: "ancestors = child # tl_ancestors" using step(3) by(auto simp add: get_ancestors_si_def[of child] intro!: bind_pure_I elim!: bind_returns_result_E2 split: option.splits) 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 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>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child") case None then show ?case using step(3) \None = 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\ apply(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2)[1] by (metis (no_types, lifting) assms(4) empty_iff empty_set select_result_I2 set_ConsD step.prems(1) step.prems(2)) next case (Some shadow_root_child) then have "shadow_root_child |\| shadow_root_ptr_kinds h" using \child |\| object_ptr_kinds h\ by (metis (no_types, lifting) shadow_root_ptr_casts_commute shadow_root_ptr_kinds_commutes) obtain host where host: "h \ get_host shadow_root_child \\<^sub>r host" using get_host_ok assms by (meson \shadow_root_child |\| shadow_root_ptr_kinds h\ is_OK_returns_result_E) then have "h \ get_ancestors_si (cast host) \\<^sub>r tl_ancestors" using Some step(3) tl_ancestors None by(auto simp add: get_ancestors_si_def[of child] intro!: bind_pure_returns_result_I elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq) then show ?case using step(2) Some host step(4) tl_ancestors by (metis (no_types, lifting) assms(4) dual_order.trans eq_iff returns_result_eq set_ConsD set_subset_Cons shadow_root_ptr_casts_commute step.prems(1)) qed 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(3) s1 apply(simp add: get_ancestors_si_def) by(auto elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq) show ?case using step(3) step(4) apply(auto simp add: \ancestors = [child]\)[1] using assms(4) returns_result_eq by fastforce next case (Some parent) then have "h \ get_ancestors_si parent \\<^sub>r tl_ancestors" using s1 tl_ancestors step(3) by(auto simp add: get_ancestors_si_def[of child] elim!: bind_returns_result_E2 split: option.splits dest: returns_result_eq) show ?case by (metis (no_types, lifting) Some.prems \h \ get_ancestors_si parent \\<^sub>r tl_ancestors\ assms(4) eq_iff node_ptr_casts_commute order_trans s1 select_result_I2 set_ConsD set_subset_Cons step.hyps(1) step.prems(1) step.prems(2) tl_ancestors) qed qed qed lemma get_ancestors_si_also_parent: assumes "heap_is_wellformed h" and "h \ get_ancestors_si 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_si (cast child) \\<^sub>r child_ancestors" by (meson assms(1) assms(4) get_ancestors_si_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_si_def) by(auto elim!: bind_returns_result_E2 split: option.splits dest!: returns_result_eq[OF assms(4)] get_ancestors_si_ptr) then show ?thesis using assms child_ancestors get_ancestors_si_subset by blast qed lemma get_ancestors_si_also_host: assumes "heap_is_wellformed h" and "h \ get_ancestors_si some_ptr \\<^sub>r ancestors" and "cast shadow_root \ set ancestors" and "h \ get_host shadow_root \\<^sub>r host" and type_wf: "type_wf h" and known_ptrs: "known_ptrs h" shows "cast host \ set ancestors" proof - obtain child_ancestors where child_ancestors: "h \ get_ancestors_si (cast shadow_root) \\<^sub>r child_ancestors" by (meson assms(1) assms(2) assms(3) get_ancestors_si_ok get_ancestors_si_ptrs_in_heap is_OK_returns_result_E known_ptrs type_wf) then have "cast host \ set child_ancestors" apply(simp add: get_ancestors_si_def) by(auto elim!: bind_returns_result_E2 split: option.splits dest!: returns_result_eq[OF assms(4)] get_ancestors_si_ptr) then show ?thesis using assms child_ancestors get_ancestors_si_subset by blast qed lemma get_ancestors_si_parent_child_rel: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ get_ancestors_si child \\<^sub>r ancestors" assumes "((ptr, child) \ (parent_child_rel h)\<^sup>*)" shows "ptr \ set ancestors" proof (insert assms(5), induct ptr rule: heap_wellformed_induct_si[OF assms(1)]) case (1 ptr) then show ?case proof (cases "ptr = child") case True then show ?thesis using assms(4) local.get_ancestors_si_ptr by blast 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(3)] \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 CD.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 CD.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 h\ local.get_child_nodes_ok local.known_ptrs_known_ptr CD.parent_child_rel_child ptr_child ptr_child_ptr_child_node returns_result_select_result \type_wf h\) ultimately show ?thesis using a1 get_child_nodes_ok \type_wf h\ \known_ptrs h\ 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 h\ \type_wf h\ by blast ultimately show ?thesis using get_ancestors_si_also_parent assms \type_wf h\ by blast qed qed lemma get_ancestors_si_parent_child_host_shadow_root_rel: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ get_ancestors_si child \\<^sub>r ancestors" assumes "((ptr, child) \ (parent_child_rel h \ a_host_shadow_root_rel h)\<^sup>*)" shows "ptr \ set ancestors" proof (insert assms(5), induct ptr rule: heap_wellformed_induct_si[OF assms(1)]) case (1 ptr) then show ?case proof (cases "ptr = child") case True then show ?thesis using assms(4) local.get_ancestors_si_ptr by blast next case False obtain ptr_child where ptr_child: "(ptr, ptr_child) \ (parent_child_rel h \ local.a_host_shadow_root_rel h) \ (ptr_child, child) \ (parent_child_rel h \ local.a_host_shadow_root_rel h)\<^sup>*" using converse_rtranclE[OF 1(3)] \ptr \ child\ by metis then show ?thesis proof(cases "(ptr, ptr_child) \ parent_child_rel h") case True 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 CD.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 CD.parent_child_rel_parent_in_heap True by blast moreover have "ptr_child_node \ set |h \ get_child_nodes ptr|\<^sub>r" by (metis True assms(2) assms(3) calculation local.CD.parent_child_rel_child local.get_child_nodes_ok local.known_ptrs_known_ptr ptr_child_ptr_child_node returns_result_select_result) ultimately show ?thesis using a1 get_child_nodes_ok \type_wf h\ \known_ptrs h\ 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 \ local.a_host_shadow_root_rel h)\<^sup>*" using ptr_child True 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 h\ \type_wf h\ by blast ultimately show ?thesis using get_ancestors_si_also_parent assms \type_wf h\ by blast next case False then obtain host where host: "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 host" using ptr_child by(auto simp add: a_host_shadow_root_rel_def) then obtain shadow_root where shadow_root: "h \ get_shadow_root host \\<^sub>r Some shadow_root" and ptr_child_shadow_root: "ptr_child = cast shadow_root" using ptr_child False apply(auto simp add: a_host_shadow_root_rel_def)[1] by (metis (no_types, lifting) assms(3) local.get_shadow_root_ok select_result_I) moreover have "(cast shadow_root, child) \ (parent_child_rel h \ local.a_host_shadow_root_rel h)\<^sup>*" using ptr_child ptr_child_shadow_root by blast ultimately have "cast shadow_root \ set ancestors" using "1.hyps"(2) host by blast moreover have "h \ get_host shadow_root \\<^sub>r host" by (metis assms(1) assms(2) assms(3) is_OK_returns_result_E local.get_host_ok local.get_shadow_root_shadow_root_ptr_in_heap local.shadow_root_host_dual local.shadow_root_same_host shadow_root) ultimately show ?thesis using get_ancestors_si_also_host assms(1) assms(2) assms(3) assms(4) host by blast qed qed qed lemma get_root_node_si_ok: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" and "ptr |\| object_ptr_kinds h" shows "h \ ok (get_root_node_si ptr)" using assms get_ancestors_si_ok by(auto simp add: get_root_node_si_def) lemma get_root_node_si_root_in_heap: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_root_node_si ptr \\<^sub>r root" shows "root |\| object_ptr_kinds h" using assms apply(auto simp add: get_root_node_si_def elim!: bind_returns_result_E2)[1] by (simp add: get_ancestors_si_never_empty get_ancestors_si_ptrs_in_heap) lemma get_root_node_si_same_no_parent: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_root_node_si 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_si) 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 show ?thesis using step(3) by(auto simp add: get_root_node_si_def get_ancestors_si_def[of c] elim!: bind_returns_result_E2 split: if_splits option.splits intro!: step(2) bind_pure_returns_result_I) next case (Some child_node) note s = this then obtain parent_opt where parent_opt: "h \ get_parent child_node \\<^sub>r parent_opt" using step(3) apply(auto simp add: get_root_node_si_def get_ancestors_si_def intro!: bind_pure_I elim!: bind_returns_result_E2)[1] by(auto split: option.splits) then show ?thesis proof(induct parent_opt) case None then show ?case using Some get_root_node_si_no_parent returns_result_eq step.prems by fastforce next case (Some parent) then show ?case using step(3) s apply(auto simp add: get_root_node_si_def get_ancestors_si_def[of c] elim!: bind_returns_result_E2 split: option.splits list.splits if_splits)[1] using assms(1) get_ancestors_si_never_empty apply blast by(auto simp add: get_root_node_si_def dest: returns_result_eq intro!: step(1) bind_pure_returns_result_I) qed qed qed end interpretation i_get_root_node_si_wf?: l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^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_host get_host_locs get_ancestors_si get_ancestors_si_locs get_root_node_si get_root_node_si_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document get_disconnected_document_locs by(auto simp add: instances l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def) declare l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_disconnected\_document\ locale l_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_disconnected_document\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf + l_get_parent begin lemma get_disconnected_document_ok: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ get_parent node_ptr \\<^sub>r None" shows "h \ ok (get_disconnected_document node_ptr)" proof - have "node_ptr |\| node_ptr_kinds h" by (meson assms(4) is_OK_returns_result_I local.get_parent_ptr_in_heap) have "\(\parent \ fset (object_ptr_kinds h). node_ptr \ set |h \ get_child_nodes parent|\<^sub>r)" apply(auto)[1] using assms(4) child_parent_dual[OF assms(1)] assms(1) assms(2) assms(3) known_ptrs_known_ptr option.simps(3) returns_result_eq returns_result_select_result by (metis (no_types, lifting) CD.get_child_nodes_ok) then have "(\document_ptr \ fset (document_ptr_kinds h). node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" using heap_is_wellformed_children_disc_nodes using \node_ptr |\| node_ptr_kinds h\ assms(1) by blast then obtain some_owner_document where "some_owner_document \ set (sorted_list_of_set (fset (document_ptr_kinds h)))" and "node_ptr \ set |h \ get_disconnected_nodes some_owner_document|\<^sub>r" by auto have h5: "\!x. x \ set (sorted_list_of_set (fset (document_ptr_kinds h))) \ h \ Heap_Error_Monad.bind (get_disconnected_nodes x) (\children. return (node_ptr \ set children)) \\<^sub>r True" apply(auto intro!: bind_pure_returns_result_I)[1] - apply (smt CD.get_disconnected_nodes_ok CD.get_disconnected_nodes_pure - \\document_ptr\fset (document_ptr_kinds h). node_ptr \ - set |h \ get_disconnected_nodes document_ptr|\<^sub>r\ assms(2) - bind_pure_returns_result_I2 notin_fset return_returns_result select_result_I2) - + apply (smt (verit, best) CD.get_disconnected_nodes_ok CD.get_disconnected_nodes_pure + \\document_ptr\fset (document_ptr_kinds h). node_ptr \ set |h \ get_disconnected_nodes + document_ptr|\<^sub>r\ assms(2) bind_pure_returns_result_I finite_set_in return_returns_result + returns_result_select_result) apply(auto elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I)[1] using heap_is_wellformed_one_disc_parent assms(1) by blast let ?filter_M = "filter_M (\document_ptr. Heap_Error_Monad.bind (get_disconnected_nodes document_ptr) (\disconnected_nodes. return (node_ptr \ set disconnected_nodes))) (sorted_list_of_set (fset (document_ptr_kinds h)))" have "h \ ok (?filter_M)" using CD.get_disconnected_nodes_ok - by (smt CD.get_disconnected_nodes_pure DocumentMonad.ptr_kinds_M_ptr_kinds + by (smt (verit) CD.get_disconnected_nodes_pure DocumentMonad.ptr_kinds_M_ptr_kinds DocumentMonad.ptr_kinds_ptr_kinds_M assms(2) bind_is_OK_pure_I bind_pure_I document_ptr_kinds_M_def filter_M_is_OK_I l_ptr_kinds_M.ptr_kinds_M_ok return_ok return_pure returns_result_select_result) then obtain candidates where candidates: "h \ filter_M (\document_ptr. Heap_Error_Monad.bind (get_disconnected_nodes document_ptr) (\disconnected_nodes. return (node_ptr \ set disconnected_nodes))) (sorted_list_of_set (fset (document_ptr_kinds h))) \\<^sub>r candidates" by auto have "candidates = [some_owner_document]" apply(rule filter_M_ex1[OF candidates \some_owner_document \ set (sorted_list_of_set (fset (document_ptr_kinds h)))\ h5]) using \node_ptr \ set |h \ get_disconnected_nodes some_owner_document|\<^sub>r\ \some_owner_document \ set (sorted_list_of_set (fset (document_ptr_kinds h)))\ by(auto simp add: CD.get_disconnected_nodes_ok assms(2) intro!: bind_pure_I intro!: bind_pure_returns_result_I) then show ?thesis using candidates \node_ptr |\| node_ptr_kinds h\ apply(auto simp add: get_disconnected_document_def intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I split: list.splits)[1] apply (meson not_Cons_self2 returns_result_eq) by (meson list.distinct(1) list.inject returns_result_eq) qed end interpretation i_get_disconnected_document_wf?: l_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs known_ptrs get_parent get_parent_locs by(auto simp add: l_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_disconnected_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_owner\_document\ locale l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_nodes + l_get_child_nodes + l_get_owner_document\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_parent_wf + l_known_ptrs + l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^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 + assumes known_ptr_impl: "known_ptr = ShadowRootClass.known_ptr" 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 apply(auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.a_all_ptrs_in_heap_def)[1] using assms(1) local.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 then have 4: "\(\parent_ptr. parent_ptr |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent_ptr|\<^sub>r)" using CD.distinct_lists_no_parent assms unfolding heap_is_wellformed_def CD.heap_is_wellformed_def by simp moreover have "(\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)" using assms(1) 2 "3" assms(2) assms(3) by auto ultimately have 0: "\!document_ptr\set |h \ document_ptr_kinds_M|\<^sub>r. node_ptr \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r" - using concat_map_distinct assms(1) known_ptrs_implies - by (smt CD.heap_is_wellformed_one_disc_parent DocumentMonad.ptr_kinds_ptr_kinds_M - disjoint_iff_not_equal local.get_disconnected_nodes_ok local.heap_is_wellformed_def + by (meson DocumentMonad.ptr_kinds_ptr_kinds_M assms(1) disjoint_iff + local.get_disconnected_nodes_ok local.heap_is_wellformed_one_disc_parent returns_result_select_result type_wf) - have "h \ get_parent node_ptr \\<^sub>r None" - using 4 2 - apply(auto simp add: get_parent_def - intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I )[1] - apply(auto intro!: filter_M_empty_I bind_pure_I bind_pure_returns_result_I)[1] - using get_child_nodes_ok assms(4) type_wf - by (metis get_child_nodes_ok known_ptrs_known_ptr returns_result_select_result) - + by (metis (mono_tags, lifting) "2" "4" known_ptrs local.get_parent_child_dual + local.get_parent_ok local.get_parent_parent_in_heap returns_result_select_result + select_result_I2 split_option_ex type_wf) then have 4: "h \ get_root_node_si (cast node_ptr) \\<^sub>r cast node_ptr" using get_root_node_si_no_parent by simp 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 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 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] - apply (smt "0" "3" assms bind_is_OK_pure_I bind_pure_returns_result_I bind_pure_returns_result_I2 - bind_returns_result_E2 bind_returns_result_E3 document_ptr_kinds_M_def get_disconnected_nodes_ok - get_disconnected_nodes_pure image_eqI is_OK_returns_result_E l_ptr_kinds_M.ptr_kinds_ptr_kinds_M - return_ok return_returns_result returns_result_eq select_result_E select_result_I select_result_I2 - select_result_I2) + using 0 document_ptrs + apply (smt (verit) DocumentMonad.ptr_kinds_ptr_kinds_M bind_pure_returns_result_I2 + local.get_disconnected_nodes_ok local.get_disconnected_nodes_pure node_ptr_inclusion + return_returns_result select_result_I2 type_wf) using assms(2) assms(3) apply (metis (no_types, lifting) bind_pure_returns_result_I2 is_OK_returns_result_I local.get_disconnected_nodes_pure node_ptr_inclusion return_returns_result select_result_I2) 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 by(auto intro: bind_pure_I bind_is_OK_I2) with 4 document_ptrs have "h \ CD.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: CD.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) moreover have "known_ptr (cast node_ptr)" using known_ptrs_known_ptr[OF known_ptrs, where 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"] 2 known_ptrs_implies by(simp) ultimately show ?thesis using 2 apply(auto simp add: CD.a_get_owner_document_tups_def 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_shadow_root_ptr) apply(drule(1) known_ptr_not_document_ptr) 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" assumes "h \ get_parent node_ptr \\<^sub>r None" assumes "h \ get_owner_document (cast node_ptr) \\<^sub>r owner_document" assumes "h \ get_disconnected_nodes owner_document \\<^sub>r disc_nodes" assumes "known_ptrs h" assumes "type_wf h" shows "node_ptr \ set disc_nodes" proof - have "\parent. parent |\| object_ptr_kinds h \ node_ptr \ set |h \ get_child_nodes parent|\<^sub>r" using assms(2) by (meson get_child_nodes_ok assms(1) assms(5) assms(6) local.child_parent_dual local.known_ptrs_known_ptr option.distinct(1) returns_result_eq returns_result_select_result) then show ?thesis - by (smt assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) finite_set_in is_OK_returns_result_I - local.get_disconnected_nodes_ok local.get_owner_document_disconnected_nodes - local.get_parent_ptr_in_heap local.heap_is_wellformed_children_disc_nodes - returns_result_select_result select_result_I2) + by (metis (no_types, lifting) assms(1) assms(2) assms(3) assms(4) assms(5) assms(6) + finite_set_in is_OK_returns_result_I local.get_disconnected_nodes_ok + local.get_owner_document_disconnected_nodes local.get_parent_ptr_in_heap + local.heap_is_wellformed_children_disc_nodes returns_result_eq returns_result_select_result) qed lemma get_owner_document_owner_document_in_heap_node: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ CD.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" shows "owner_document |\| document_ptr_kinds h" proof - obtain root where root: "h \ get_root_node_si (cast node_ptr) \\<^sub>r root" using assms(4) by(auto simp add: CD.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 assms(4) root apply(auto simp add: CD.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 assms document_ptr_kinds_commutes get_root_node_si_root_in_heap by blast next case False have "known_ptr root" using assms local.get_root_node_si_root_in_heap local.known_ptrs_known_ptr root by blast have "root |\| object_ptr_kinds h" using root using assms local.get_root_node_si_root_in_heap by blast have "\is_shadow_root_ptr root" using root using local.get_root_node_si_root_not_shadow_root by blast then have "is_node_ptr_kind root" using False \known_ptr root\ \root |\| object_ptr_kinds h\ apply(simp add: known_ptr_impl known_ptr_defs DocumentClass.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)" using local.child_parent_dual local.get_child_nodes_ok local.get_root_node_si_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 by (metis (no_types, lifting) assms \root |\| object_ptr_kinds h\) 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) assms 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 simp add: assms local.get_disconnected_nodes_ok intro!: bind_pure_I bind_pure_returns_result_I)[1] done then have "candidates \ []" by auto then have "owner_document \ set candidates" using assms(4) root apply(auto simp add: CD.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_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 CD.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) (\_. (CD.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: CD.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) (\_. (CD.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 show ?thesis by (metis bind_returns_result_E2 check_in_heap_pure comp_apply get_owner_document_owner_document_in_heap_node) 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) (\_. (CD.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 show ?thesis by (metis bind_returns_result_E2 check_in_heap_pure comp_apply get_owner_document_owner_document_in_heap_node) 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: "\ is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 6: "is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 7: "h \ Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^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>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ()) \\<^sub>r owner_document" then show "owner_document |\| document_ptr_kinds h" apply(auto simp add: CD.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>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: filter_M_pure_I bind_pure_I elim!: bind_returns_result_E2 split: if_splits option.splits)[1] using get_owner_document_owner_document_in_heap_node by blast 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(simp add: get_owner_document_def a_get_owner_document_tups_def CD.a_get_owner_document_tups_def) apply(split invoke_splits, (rule conjI | rule impI)+)+ proof - assume 0: "known_ptr ptr" and 1: "\ is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 2: "\ 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 3: "\ is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 4: "\ is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" then show "h \ ok invoke [] ptr ()" using NodeClass.a_known_ptr_def known_ptr_not_character_data_ptr known_ptr_not_document_ptr known_ptr_not_shadow_root_ptr known_ptr_not_element_ptr known_ptr_impl by blast next assume 0: "known_ptr ptr" and 1: "\ is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 2: "\ 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 3: "\ is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" then show "is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \ h \ ok Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.a_get_owner_document\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^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>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r) ptr ())" using assms(1) assms(2) assms(3) assms(4) by(auto simp add: local.get_host_ok get_root_node_def CD.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 CD.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>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I get_root_node_si_ok get_disconnected_nodes_ok intro!: local.get_shadow_root_ptr_in_heap local.shadow_root_host_dual split: option.splits) next show "is_document_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \ h \ ok Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.CD.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 ())" using assms(4) by(auto simp add: CD.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 split: option.splits) next show "is_character_data_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \ h \ ok Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.CD.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 ())" using assms(1) assms(2) assms(3) assms(4) by(auto simp add: local.get_host_ok get_root_node_def CD.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 CD.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>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I get_root_node_si_ok get_disconnected_nodes_ok intro!: local.get_shadow_root_ptr_in_heap local.shadow_root_host_dual split: option.splits) next show "is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \ h \ ok Heap_Error_Monad.bind (check_in_heap ptr) (\_. (local.CD.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 ())" using assms(1) assms(2) assms(3) assms(4) by(auto simp add: local.get_host_ok get_root_node_def CD.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 CD.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>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I get_root_node_si_ok get_disconnected_nodes_ok intro!: local.get_shadow_root_ptr_in_heap local.shadow_root_host_dual split: option.splits) qed qed end interpretation i_get_owner_document_wf?: l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_disconnected_nodes get_disconnected_nodes_locs known_ptr get_child_nodes get_child_nodes_locs DocumentClass.known_ptr get_parent get_parent_locs get_root_node_si get_root_node_si_locs CD.a_get_owner_document get_host get_host_locs get_owner_document get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document get_disconnected_document_locs known_ptrs get_ancestors_si get_ancestors_si_locs by(auto simp add: l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def instances) declare l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^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" apply(auto simp add: l_get_owner_document_wf_def l_get_owner_document_wf_axioms_def instances)[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 \remove\_child\ locale l_remove_child_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_set_disconnected_nodes_get_disconnected_nodes + l_get_child_nodes + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_owner_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^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_get_shadow_root + l_set_disconnected_nodes_get_shadow_root + l_set_child_nodes_get_tag_name + l_set_disconnected_nodes_get_tag_name + CD: l_remove_child_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma remove_child_preserves_type_wf: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove_child ptr child \\<^sub>h h'" shows "type_wf h'" using CD.remove_child_heap_is_wellformed_preserved(1) assms unfolding heap_is_wellformed_def by auto lemma remove_child_preserves_known_ptrs: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove_child ptr child \\<^sub>h h'" shows "known_ptrs h'" using CD.remove_child_heap_is_wellformed_preserved(2) assms unfolding heap_is_wellformed_def by auto lemma remove_child_heap_is_wellformed_preserved: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove_child ptr child \\<^sub>h h'" shows "heap_is_wellformed h'" proof - have "heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'" using CD.remove_child_heap_is_wellformed_preserved(3) assms unfolding heap_is_wellformed_def by auto have shadow_root_eq: "\ptr' shadow_root_ptr_opt. h \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt = h' \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt" using get_shadow_root_reads remove_child_writes assms(4) apply(rule reads_writes_preserved) by(auto simp add: remove_child_locs_def set_child_nodes_get_shadow_root set_disconnected_nodes_get_shadow_root) then have shadow_root_eq2: "\ptr'. |h \ get_shadow_root ptr'|\<^sub>r = |h' \ get_shadow_root ptr'|\<^sub>r" by (meson select_result_eq) have tag_name_eq: "\ptr' tag. h \ get_tag_name ptr' \\<^sub>r tag = h' \ get_tag_name ptr' \\<^sub>r tag" using get_tag_name_reads remove_child_writes assms(4) apply(rule reads_writes_preserved) by(auto simp add: remove_child_locs_def set_child_nodes_get_tag_name set_disconnected_nodes_get_tag_name) then have tag_name_eq2: "\ptr'. |h \ get_tag_name ptr'|\<^sub>r = |h' \ get_tag_name ptr'|\<^sub>r" by (meson select_result_eq) have object_ptr_kinds_eq: "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) have shadow_root_ptr_kinds_eq: "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'" using object_ptr_kinds_eq by(auto simp add: shadow_root_ptr_kinds_def) have element_ptr_kinds_eq: "element_ptr_kinds h = element_ptr_kinds h'" using object_ptr_kinds_eq by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def) have "parent_child_rel h' \ parent_child_rel h" using \heap_is_wellformed h\ heap_is_wellformed_def using CD.remove_child_parent_child_rel_subset using \known_ptrs h\ \type_wf h\ assms(4) by simp show ?thesis using \heap_is_wellformed h\ using \heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'\ \parent_child_rel h' \ parent_child_rel h\ apply(auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def a_host_shadow_root_rel_def a_all_ptrs_in_heap_def a_distinct_lists_def a_shadow_root_valid_def object_ptr_kinds_eq element_ptr_kinds_eq shadow_root_ptr_kinds_eq shadow_root_eq shadow_root_eq2 tag_name_eq tag_name_eq2)[1] by (meson acyclic_subset order_refl sup_mono) qed lemma remove_preserves_type_wf: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove child \\<^sub>h h'" shows "type_wf h'" using CD.remove_heap_is_wellformed_preserved(1) assms unfolding heap_is_wellformed_def by auto lemma remove_preserves_known_ptrs: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove child \\<^sub>h h'" shows "known_ptrs h'" using CD.remove_heap_is_wellformed_preserved(2) assms unfolding heap_is_wellformed_def by auto lemma remove_heap_is_wellformed_preserved: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ remove child \\<^sub>h h'" shows "heap_is_wellformed h'" using assms by(auto simp add: remove_def elim!: bind_returns_heap_E2 intro: remove_child_heap_is_wellformed_preserved split: option.splits) lemma 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" using CD.remove_child_removes_child local.heap_is_wellformed_def by blast lemma 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" using CD.remove_child_removes_first_child local.heap_is_wellformed_def by blast lemma 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" using CD.remove_removes_child local.heap_is_wellformed_def by blast lemma 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 []" using CD.remove_for_all_empty_children local.heap_is_wellformed_def by blast end interpretation i_remove_child_wf2?: l_remove_child_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs known_ptr get_child_nodes get_child_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs DocumentClass.known_ptr get_parent get_parent_locs get_root_node_si get_root_node_si_locs CD.a_get_owner_document get_owner_document known_ptrs get_ancestors_si get_ancestors_si_locs set_child_nodes set_child_nodes_locs remove_child remove_child_locs remove by(auto simp add: l_remove_child_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_remove_child_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^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_preserves_type_wf apply fast using remove_child_preserves_known_ptrs apply fast using remove_child_heap_is_wellformed_preserved apply (fast) using remove_preserves_type_wf apply fast using remove_preserves_known_ptrs apply fast using remove_heap_is_wellformed_preserved apply (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 subsubsection \adopt\_node\ locale l_adopt_node_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes + l_get_disconnected_nodes + l_set_child_nodes_get_shadow_root + l_set_disconnected_nodes_get_shadow_root + l_set_child_nodes_get_tag_name + l_set_disconnected_nodes_get_tag_name + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_root_node + l_set_disconnected_nodes_get_child_nodes + l_get_owner_document_wf + l_remove_child_wf2 + 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 begin lemma adopt_node_removes_child: 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 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_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 assms by(auto split: option.splits) have "known_ptrs h2" using h2 remove_child_preserves_known_ptrs assms by(auto split: option.splits) then have "heap_is_wellformed h' \ known_ptrs h' \ type_wf h'" proof(cases "document_ptr = old_document") case True then show "heap_is_wellformed h' \ known_ptrs h' \ type_wf h'" using h' wellformed_h2 \known_ptrs h2\ \type_wf 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 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 "CD.a_owner_document_valid h" using assms(1) by(simp add: heap_is_wellformed_def CD.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 "CD.a_distinct_lists h2" using heap_is_wellformed_def CD.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: CD.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 "CD.a_acyclic_heap h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.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 CD.parent_child_rel_def by(simp) qed then have " CD.a_acyclic_heap h'" using \ CD.a_acyclic_heap h2\ CD.acyclic_heap_def acyclic_subset by blast moreover have " CD.a_all_ptrs_in_heap h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_all_ptrs_in_heap h3" apply(auto simp add: CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq3_h2 children_eq_h2)[1] apply (metis \type_wf h'\ children_eq2_h3 children_eq_h2 children_eq_h3 known_ptrs l_heap_is_wellformed.heap_is_wellformed_children_in_heap local.get_child_nodes_ok local.known_ptrs_known_ptr local.l_heap_is_wellformed_axioms node_ptr_kinds_eq3_h2 object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 object_ptr_kinds_h_eq3 returns_result_select_result wellformed_h2) by (metis (no_types, opaque_lifting) disc_nodes_old_document_h2 disc_nodes_old_document_h3 disconnected_nodes_eq2_h2 document_ptr_kinds_eq3_h2 finite_set_in select_result_I2 set_remove1_subset subsetD) then have "CD.a_all_ptrs_in_heap h'" apply(auto simp add: CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq3_h3 children_eq_h3)[1] apply (metis (no_types, opaque_lifting) children_eq2_h3 finite_set_in object_ptr_kinds_h3_eq3 subsetD) by (metis (no_types, opaque_lifting) \child \ set disc_nodes_old_document_h2\ disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2 disc_nodes_old_document_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq3_h3 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 subsetD wellformed_h2) moreover have "CD.a_owner_document_valid h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h'" apply(simp add: CD.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 (metis (no_types) 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: "CD.a_distinct_lists h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_distinct_lists h'" apply(auto simp add: CD.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 have "heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'" using \CD.a_owner_document_valid h'\ CD.heap_is_wellformed_def by simp have shadow_root_eq_h2: "\ptr' shadow_root_ptr_opt. h2 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt = h3 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt" using get_shadow_root_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_shadow_root set_disconnected_nodes_get_shadow_root) then have shadow_root_eq2_h2: "\ptr'. |h2 \ get_shadow_root ptr'|\<^sub>r = |h3 \ get_shadow_root ptr'|\<^sub>r" by (meson select_result_eq) have shadow_root_eq_h3: "\ptr' shadow_root_ptr_opt. h3 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt = h' \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt" using get_shadow_root_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_shadow_root set_disconnected_nodes_get_shadow_root) then have shadow_root_eq2_h3: "\ptr'. |h3 \ get_shadow_root ptr'|\<^sub>r = |h' \ get_shadow_root ptr'|\<^sub>r" by (meson select_result_eq) have tag_name_eq_h2: "\ptr' tag. h2 \ get_tag_name ptr' \\<^sub>r tag = h3 \ get_tag_name ptr' \\<^sub>r tag" using get_tag_name_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_tag_name set_disconnected_nodes_get_tag_name) then have tag_name_eq2_h2: "\ptr'. |h2 \ get_tag_name ptr'|\<^sub>r = |h3 \ get_tag_name ptr'|\<^sub>r" by (meson select_result_eq) have tag_name_eq_h3: "\ptr' tag. h3 \ get_tag_name ptr' \\<^sub>r tag = h' \ get_tag_name ptr' \\<^sub>r tag" using get_tag_name_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_tag_name set_disconnected_nodes_get_tag_name) then have tag_name_eq2_h3: "\ptr'. |h3 \ get_tag_name ptr'|\<^sub>r = |h' \ get_tag_name ptr'|\<^sub>r" by (meson select_result_eq) have object_ptr_kinds_eq_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 adopt_node_locs_def remove_child_locs_def using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def split: if_splits) have object_ptr_kinds_eq_h3: "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']) unfolding adopt_node_locs_def remove_child_locs_def using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def split: if_splits) have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h3" using object_ptr_kinds_eq_h2 by(auto simp add: shadow_root_ptr_kinds_def) have element_ptr_kinds_eq_h2: "element_ptr_kinds h2 = element_ptr_kinds h3" using object_ptr_kinds_eq_h2 by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h'" using object_ptr_kinds_eq_h3 by(auto simp add: shadow_root_ptr_kinds_def) have element_ptr_kinds_eq_h3: "element_ptr_kinds h3 = element_ptr_kinds h'" using object_ptr_kinds_eq_h3 by(auto simp add: element_ptr_kinds_def node_ptr_kinds_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 show "heap_is_wellformed h' \ known_ptrs h' \ type_wf h'" using \heap_is_wellformed h2\ using \heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'\ \known_ptrs h'\ \type_wf h'\ using \parent_child_rel h' \ parent_child_rel h2\ apply(auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def a_host_shadow_root_rel_def a_all_ptrs_in_heap_def a_distinct_lists_def a_shadow_root_valid_def object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 element_ptr_kinds_eq_h2 element_ptr_kinds_eq_h3 shadow_root_ptr_kinds_eq_h2 shadow_root_ptr_kinds_eq_h3 shadow_root_eq_h2 shadow_root_eq_h3 shadow_root_eq2_h2 shadow_root_eq2_h3 tag_name_eq_h2 tag_name_eq_h3 tag_name_eq2_h2 tag_name_eq2_h3 CD.parent_child_rel_def children_eq2_h2 children_eq2_h3 object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3)[1] done 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 l_adopt_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M Shadow_DOM.get_owner_document Shadow_DOM.get_parent Shadow_DOM.get_parent_locs Shadow_DOM.remove_child Shadow_DOM.remove_child_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs Shadow_DOM.adopt_node Shadow_DOM.adopt_node_locs ShadowRootClass.known_ptr ShadowRootClass.type_wf Shadow_DOM.get_child_nodes Shadow_DOM.get_child_nodes_locs ShadowRootClass.known_ptrs Shadow_DOM.set_child_nodes Shadow_DOM.set_child_nodes_locs Shadow_DOM.remove Shadow_DOM.heap_is_wellformed Shadow_DOM.parent_child_rel by(auto 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>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs set_child_nodes set_child_nodes_locs get_shadow_root get_shadow_root_locs set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs get_root_node get_root_node_locs get_parent get_parent_locs known_ptrs get_owner_document remove_child remove_child_locs remove adopt_node adopt_node_locs by(auto simp add: l_adopt_node_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_adopt_node_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] 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" apply(auto simp add: l_adopt_node_wf_def l_adopt_node_wf_axioms_def instances)[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 subsubsection \insert\_before\ locale l_insert_before_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes + l_get_disconnected_nodes + l_set_child_nodes_get_shadow_root + l_set_disconnected_nodes_get_shadow_root + l_set_child_nodes_get_tag_name + l_set_disconnected_nodes_get_tag_name + l_set_disconnected_nodes_get_disconnected_nodes + l_set_child_nodes_get_disconnected_nodes + l_set_disconnected_nodes_get_disconnected_nodes_wf + l_set_disconnected_nodes_get_ancestors_si + l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M _ _ _ _ _ _ get_ancestors_si get_ancestors_si_locs + l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_owner_document + l_adopt_node + l_adopt_node_wf + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_adopt_node_get_shadow_root begin lemma insert_before_child_preserves: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ insert_before ptr node child \\<^sub>h h'" shows "type_wf h'" and "known_ptrs h'" and "heap_is_wellformed h'" proof - obtain ancestors reference_child owner_document h2 h3 disconnected_nodes_h2 where ancestors: "h \ get_ancestors_si 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'" (* children: "h3 \ get_child_nodes ptr \\<^sub>r children" and *) (* h': "h3 \ set_child_nodes ptr (insert_before_list node reference_child children) \\<^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 "type_wf h2" using \type_wf h\ using writes_small_big[where P="\h h'. type_wf h \ type_wf h'", OF adopt_node_writes h2] using adopt_node_types_preserved by(auto simp add: 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 h = object_ptr_kinds h2" using adopt_node_writes h2 apply(rule writes_small_big) using adopt_node_pointers_preserved by(auto simp add: reflp_def transp_def) moreover have "\ = object_ptr_kinds h3" using set_disconnected_nodes_writes h3 apply(rule writes_small_big) using set_disconnected_nodes_pointers_preserved by(auto simp add: reflp_def transp_def) moreover have "\ = object_ptr_kinds h'" using insert_node_writes h' apply(rule writes_small_big) using set_child_nodes_pointers_preserved by(auto simp add: reflp_def transp_def) ultimately show "known_ptrs h'" using \known_ptrs h\ known_ptrs_preserved by blast have "known_ptrs h2" using \known_ptrs h\ known_ptrs_preserved \object_ptr_kinds h = object_ptr_kinds h2\ by blast then have "known_ptrs h3" using known_ptrs_preserved \object_ptr_kinds h2 = object_ptr_kinds h3\ by blast have "known_ptr ptr" by (meson get_owner_document_ptr_in_heap is_OK_returns_result_I \known_ptrs h\ l_known_ptrs.known_ptrs_known_ptr l_known_ptrs_axioms owner_document) 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 wellformed_h2: "heap_is_wellformed h2" using adopt_node_preserves_wellformedness[OF \heap_is_wellformed h\ h2] \known_ptrs h\ \type_wf h\ . 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 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 shadow_root_eq_h2: "\ptr' shadow_root. h \ get_shadow_root ptr' \\<^sub>r shadow_root = h2 \ get_shadow_root ptr' \\<^sub>r shadow_root" using get_shadow_root_reads adopt_node_writes h2 apply(rule reads_writes_preserved) using local.adopt_node_get_shadow_root 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 \heap_is_wellformed h\ 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_si ptr \\<^sub>r ancestors_h2" using get_ancestors_si_ok by (metis \known_ptrs h2\ \type_wf h2\ is_OK_returns_result_E object_ptr_kinds_M_eq3_h2 ptr_in_heap wellformed_h2) have ancestors_h3: "h3 \ get_ancestors_si ptr \\<^sub>r ancestors_h2" using get_ancestors_si_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_separate_forwards) using \heap_is_wellformed h2\ apply simp using ancestors_h2 apply simp apply(auto simp add: get_ancestors_si_locs_def get_parent_locs_def)[1] apply (simp add: local.get_ancestors_si_locs_def local.get_parent_reads_pointers local.set_disconnected_nodes_get_ancestors_si) using local.get_ancestors_si_locs_def local.set_disconnected_nodes_get_ancestors_si by blast have node_not_in_ancestors_h2: "cast node \ set ancestors_h2" using \heap_is_wellformed h\ \heap_is_wellformed h2\ ancestors ancestors_h2 apply(rule get_ancestors_si_remains_not_in_ancestors) using assms(2) assms(3) h2 local.adopt_node_children_subset apply blast using shadow_root_eq_h2 node_not_in_ancestors object_ptr_kinds_M_eq2_h assms(2) assms(3) \type_wf h2\ by(auto dest: returns_result_eq) moreover have "parent_child_rel h2 = parent_child_rel h3" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2) 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: CD.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) have "CD.a_acyclic_heap h'" proof - have "acyclic (parent_child_rel h2)" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) then have "acyclic (parent_child_rel h3)" by(auto simp add: CD.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_si_parent_child_rel using \known_ptrs h2\ \type_wf h2\ ancestors_h2 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: CD.parent_child_rel_def object_ptr_kinds_M_eq3_h2 children_eq2_h2) ultimately show ?thesis using \parent_child_rel h' = insert (ptr, cast node) ((parent_child_rel h3))\ by(auto simp add: CD.acyclic_heap_def) qed moreover have "CD.a_all_ptrs_in_heap h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) have "CD.a_all_ptrs_in_heap h'" proof - have "CD.a_all_ptrs_in_heap h3" using \CD.a_all_ptrs_in_heap h2\ apply(auto simp add: CD.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 (no_types, lifting) children_eq2_h2 in_mono notin_fset object_ptr_kinds_M_eq3_h2) by (metis (no_types, opaque_lifting) NodeMonad.ptr_kinds_ptr_kinds_M disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 document_ptr_kinds_commutes finite_set_in node_ptr_kinds_eq2_h2 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 \CD.a_all_ptrs_in_heap h3\ apply(auto simp add: CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq2_h3)[1] using CD.parent_child_rel_child_nodes2 \known_ptr ptr\ \parent_child_rel h2 = parent_child_rel h3\ \type_wf h2\ local.parent_child_rel_child_in_heap node_ptr_kinds_commutes object_ptr_kinds_M_eq3_h' object_ptr_kinds_M_eq3_h2 wellformed_h2 by blast 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 \CD.a_all_ptrs_in_heap h3\ apply(auto simp add: object_ptr_kinds_M_eq3_h' CD.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, lifting) DocumentMonad.ptr_kinds_ptr_kinds_M disconnected_nodes_eq2_h3 document_ptr_kinds_eq2_h3 finite_set_in subsetD) qed moreover have "CD.a_distinct_lists h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def ) then have "CD.a_distinct_lists h3" proof(auto simp add: CD.a_distinct_lists_def object_ptr_kinds_M_eq2_h2 document_ptr_kinds_eq2_h2 children_eq2_h2 intro!: distinct_concat_map_I) 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 h\ 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 "CD.a_distinct_lists h'" proof(auto simp add: CD.a_distinct_lists_def document_ptr_kinds_eq2_h3 object_ptr_kinds_M_eq2_h3 disconnected_nodes_eq2_h3 intro!: distinct_concat_map_I) 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 h\ 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.CD.a_distinct_lists h3\ CD.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 \CD.a_distinct_lists h3\ \type_wf h'\ disconnected_nodes_eq_h3 CD.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 "CD.a_owner_document_valid h2" using wellformed_h2 by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h'" apply(auto simp add: CD.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] thm children_eq2_h3 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 (metis (no_types, lifting) Core_DOM_Functions.i_insert_before.insert_before_list_in_set children_eq2_h3 children_h' children_h3 disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 finite_set_in in_set_remove1 is_OK_returns_result_I object_ptr_kinds_M_eq3_h' ptr_in_heap returns_result_eq returns_result_select_result) ultimately have "heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'" by (simp add: CD.heap_is_wellformed_def) have shadow_root_eq_h2: "\ptr' shadow_root_ptr_opt. h2 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt = h3 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt" using get_shadow_root_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_shadow_root set_disconnected_nodes_get_shadow_root) then have shadow_root_eq2_h2: "\ptr'. |h2 \ get_shadow_root ptr'|\<^sub>r = |h3 \ get_shadow_root ptr'|\<^sub>r" by (meson select_result_eq) have shadow_root_eq_h3: "\ptr' shadow_root_ptr_opt. h3 \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt = h' \ get_shadow_root ptr' \\<^sub>r shadow_root_ptr_opt" using get_shadow_root_reads insert_node_writes h' apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_shadow_root set_disconnected_nodes_get_shadow_root) then have shadow_root_eq2_h3: "\ptr'. |h3 \ get_shadow_root ptr'|\<^sub>r = |h' \ get_shadow_root ptr'|\<^sub>r" by (meson select_result_eq) have tag_name_eq_h2: "\ptr' tag. h2 \ get_tag_name ptr' \\<^sub>r tag = h3 \ get_tag_name ptr' \\<^sub>r tag" using get_tag_name_reads set_disconnected_nodes_writes h3 apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_tag_name set_disconnected_nodes_get_tag_name) then have tag_name_eq2_h2: "\ptr'. |h2 \ get_tag_name ptr'|\<^sub>r = |h3 \ get_tag_name ptr'|\<^sub>r" by (meson select_result_eq) have tag_name_eq_h3: "\ptr' tag. h3 \ get_tag_name ptr' \\<^sub>r tag = h' \ get_tag_name ptr' \\<^sub>r tag" using get_tag_name_reads insert_node_writes h' apply(rule reads_writes_preserved) by(auto simp add: adopt_node_locs_def remove_child_locs_def set_child_nodes_get_tag_name set_disconnected_nodes_get_tag_name) then have tag_name_eq2_h3: "\ptr'. |h3 \ get_tag_name ptr'|\<^sub>r = |h' \ get_tag_name ptr'|\<^sub>r" by (meson select_result_eq) have object_ptr_kinds_eq_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 adopt_node_locs_def remove_child_locs_def using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def split: if_splits) have object_ptr_kinds_eq_h3: "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']) using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def split: if_splits) have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h3" using object_ptr_kinds_eq_h2 by(auto simp add: shadow_root_ptr_kinds_def) have element_ptr_kinds_eq_h2: "element_ptr_kinds h2 = element_ptr_kinds h3" using object_ptr_kinds_eq_h2 by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h'" using object_ptr_kinds_eq_h3 by(auto simp add: shadow_root_ptr_kinds_def) have element_ptr_kinds_eq_h3: "element_ptr_kinds h3 = element_ptr_kinds h'" using object_ptr_kinds_eq_h3 by(auto simp add: element_ptr_kinds_def node_ptr_kinds_def) have "a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 shadow_root_eq2_h2) have "a_host_shadow_root_rel h3 = a_host_shadow_root_rel h'" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h3 shadow_root_eq2_h3) have "cast node \ {x. (x, ptr) \ (parent_child_rel h3 \ a_host_shadow_root_rel h3)\<^sup>*}" using get_ancestors_si_parent_child_host_shadow_root_rel using \known_ptrs h2\ \local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\ \parent_child_rel h2 = parent_child_rel h3\ \type_wf h2\ ancestors_h2 node_not_in_ancestors_h2 wellformed_h2 by auto have "acyclic (parent_child_rel h3 \ a_host_shadow_root_rel h3)" using \heap_is_wellformed h2\ by(auto simp add: heap_is_wellformed_def \parent_child_rel h2 = parent_child_rel h3\ \a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3\) then have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h')" apply(auto simp add: \a_host_shadow_root_rel h3 = a_host_shadow_root_rel h'\ \parent_child_rel h' = insert (ptr, cast node) ((parent_child_rel h3))\)[1] using \cast node \ {x. (x, ptr) \ (parent_child_rel h3 \ a_host_shadow_root_rel h3)\<^sup>*}\ by (simp add: \local.a_host_shadow_root_rel h3 = local.a_host_shadow_root_rel h'\) then show "heap_is_wellformed h'" using \heap_is_wellformed h2\ using \heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M h'\ apply(auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def a_all_ptrs_in_heap_def a_distinct_lists_def a_shadow_root_valid_def)[1] by(auto simp add: object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 element_ptr_kinds_eq_h2 element_ptr_kinds_eq_h3 shadow_root_ptr_kinds_eq_h2 shadow_root_ptr_kinds_eq_h3 shadow_root_eq_h2 shadow_root_eq_h3 shadow_root_eq2_h2 shadow_root_eq2_h3 tag_name_eq_h2 tag_name_eq_h3 tag_name_eq2_h2 tag_name_eq2_h3) qed end 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_si get_ancestors_si_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 Shadow_DOM.heap_is_wellformed ShadowRootClass.type_wf ShadowRootClass.known_ptr ShadowRootClass.known_ptrs Shadow_DOM.insert_before Shadow_DOM.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 lemma l_set_disconnected_nodes_get_disconnected_nodes_wf [instances]: "l_set_disconnected_nodes_get_disconnected_nodes_wf ShadowRootClass.type_wf ShadowRootClass.known_ptr Shadow_DOM.heap_is_wellformed Shadow_DOM.parent_child_rel Shadow_DOM.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] by (metis Diff_iff Shadow_DOM.i_heap_is_wellformed.heap_is_wellformed_disconnected_nodes_distinct Shadow_DOM.i_remove_child.set_disconnected_nodes_get_disconnected_nodes insert_iff returns_result_eq set_remove1_eq) interpretation l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^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_host get_host_locs get_ancestors_si get_ancestors_si_locs get_root_node_si get_root_node_si_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document get_disconnected_document_locs by(auto simp add: l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_get_root_node_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] interpretation i_insert_before_wf2?: l_insert_before_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs set_child_nodes set_child_nodes_locs get_shadow_root get_shadow_root_locs set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel get_ancestors_si get_ancestors_si_locs get_parent get_parent_locs adopt_node adopt_node_locs get_owner_document insert_before insert_before_locs append_child known_ptrs get_host get_host_locs get_root_node_si get_root_node_si_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document get_disconnected_document_locs by(auto simp add: l_insert_before_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_insert_before_wf2\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] lemma insert_before_wf2_is_l_insert_before_wf2 [instances]: "l_insert_before_wf2 ShadowRootClass.type_wf ShadowRootClass.known_ptr ShadowRootClass.known_ptrs Shadow_DOM.insert_before Shadow_DOM.heap_is_wellformed" apply(auto simp add: l_insert_before_wf2_def l_insert_before_wf2_axioms_def instances)[1] using insert_before_child_preserves apply(fast, fast, fast) done subsubsection \append\_child\ 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_si get_ancestors_si_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) declare l_append_child_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [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+ subsubsection \to\_tree\_order\ 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 apply(auto simp add: l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances)[1] done declare l_to_tree_order_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] 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" apply(auto simp add: l_to_tree_order_wf_def l_to_tree_order_wf_axioms_def instances)[1] using to_tree_order_ok apply fast using to_tree_order_ptrs_in_heap apply fast using to_tree_order_parent_child_rel apply(fast, fast) using to_tree_order_child2 apply blast using to_tree_order_node_ptrs apply fast using to_tree_order_child apply fast using to_tree_order_ptr_in_result apply fast using to_tree_order_parent apply fast using to_tree_order_subset apply fast done paragraph \get\_root\_node\ 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 by(auto 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 instances) declare l_to_tree_order_wf_get_root_node_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] 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 ShadowRootClass.type_wf ShadowRootClass.known_ptr ShadowRootClass.known_ptrs to_tree_order Shadow_DOM.get_root_node Shadow_DOM.heap_is_wellformed" 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 instances)[1] using to_tree_order_get_root_node apply fast using to_tree_order_same_root apply fast done subsubsection \to\_tree\_order\_si\ locale l_to_tree_order_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_child_nodes + l_get_parent_get_host_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_to_tree_order_si\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma to_tree_order_si_ok: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" and "ptr |\| object_ptr_kinds h" shows "h \ ok (to_tree_order_si ptr)" proof(insert assms(1) assms(4), induct rule: heap_wellformed_induct_si) case (step parent) have "known_ptr parent" using assms(2) local.known_ptrs_known_ptr step.prems by blast then show ?case using step using assms(1) assms(2) assms(3) using local.heap_is_wellformed_children_in_heap local.get_shadow_root_shadow_root_ptr_in_heap by(auto simp add: to_tree_order_si_def[of parent] intro: get_child_nodes_ok get_shadow_root_ok intro!: bind_is_OK_pure_I map_M_pure_I bind_pure_I map_M_ok_I split: option.splits) qed end interpretation i_to_tree_order_si_wf?: l_to_tree_order_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf known_ptr get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs known_ptrs get_parent get_parent_locs to_tree_order_si by(auto simp add: l_to_tree_order_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_to_tree_order_si_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_assigned\_nodes\ lemma forall_M_small_big: "h \ forall_M f xs \\<^sub>h h' \ P h \ (\h h' x. x \ set xs \ h \ f x \\<^sub>h h' \ P h \ P h') \ P h'" by(induct xs arbitrary: h) (auto elim!: bind_returns_heap_E) locale l_assigned_nodes_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_assigned_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_heap_is_wellformed + l_remove_child_wf2 + l_append_child_wf + l_remove_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin lemma assigned_nodes_distinct: assumes "heap_is_wellformed h" assumes "h \ assigned_nodes slot \\<^sub>r nodes" shows "distinct nodes" proof - have "\ptr children. h \ get_child_nodes ptr \\<^sub>r children \ distinct children" using assms(1) local.heap_is_wellformed_children_distinct by blast then show ?thesis using assms apply(auto simp add: assigned_nodes_def elim!: bind_returns_result_E2 split: if_splits)[1] by (simp add: filter_M_distinct) qed lemma flatten_dom_preserves: assumes "heap_is_wellformed h" and "known_ptrs h" and "type_wf h" assumes "h \ flatten_dom \\<^sub>h h'" shows "heap_is_wellformed h'" and "known_ptrs h'" and "type_wf h'" proof - obtain tups h2 element_ptrs shadow_root_ptrs where "h \ element_ptr_kinds_M \\<^sub>r element_ptrs" and tups: "h \ map_filter_M2 (\element_ptr. do { tag \ get_tag_name element_ptr; assigned_nodes \ assigned_nodes element_ptr; (if tag = ''slot'' \ assigned_nodes \ [] then return (Some (element_ptr, assigned_nodes)) else return None)}) element_ptrs \\<^sub>r tups" (is "h \ map_filter_M2 ?f element_ptrs \\<^sub>r tups") and h2: "h \ forall_M (\(slot, assigned_nodes). do { get_child_nodes (cast slot) \ forall_M remove; forall_M (append_child (cast slot)) assigned_nodes }) tups \\<^sub>h h2" and "h2 \ shadow_root_ptr_kinds_M \\<^sub>r shadow_root_ptrs" and h': "h2 \ forall_M (\shadow_root_ptr. do { host \ get_host shadow_root_ptr; get_child_nodes (cast host) \ forall_M remove; get_child_nodes (cast shadow_root_ptr) \ forall_M (append_child (cast host)); remove_shadow_root host }) shadow_root_ptrs \\<^sub>h h'" using \h \ flatten_dom \\<^sub>h h'\ apply(auto simp add: flatten_dom_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF ElementMonad.ptr_kinds_M_pure, rotated] bind_returns_heap_E2[rotated, OF ShadowRootMonad.ptr_kinds_M_pure, rotated])[1] apply(drule pure_returns_heap_eq) by(auto intro!: map_filter_M2_pure bind_pure_I) have "heap_is_wellformed h2 \ known_ptrs h2 \ type_wf h2" using h2 \heap_is_wellformed h\ \known_ptrs h\ \type_wf h\ by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] elim!: forall_M_small_big[where P = "\h. heap_is_wellformed h \ known_ptrs h \ type_wf h", simplified] intro: remove_preserves_known_ptrs remove_heap_is_wellformed_preserved remove_preserves_type_wf append_child_preserves_known_ptrs append_child_heap_is_wellformed_preserved append_child_preserves_type_wf) then show "heap_is_wellformed h'" and "known_ptrs h'" and "type_wf h'" using h' by(auto elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_host_pure, rotated] bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] dest!: forall_M_small_big[where P = "\h. heap_is_wellformed h \ known_ptrs h \ type_wf h", simplified] intro: remove_preserves_known_ptrs remove_heap_is_wellformed_preserved remove_preserves_type_wf append_child_preserves_known_ptrs append_child_heap_is_wellformed_preserved append_child_preserves_type_wf remove_shadow_root_preserves ) qed end interpretation i_assigned_nodes_wf?: l_assigned_nodes_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr assigned_nodes assigned_nodes_flatten flatten_dom get_child_nodes get_child_nodes_locs get_tag_name get_tag_name_locs get_root_node get_root_node_locs get_host get_host_locs find_slot assigned_slot remove insert_before insert_before_locs append_child remove_shadow_root remove_shadow_root_locs type_wf get_shadow_root get_shadow_root_locs set_shadow_root set_shadow_root_locs get_parent get_parent_locs to_tree_order heap_is_wellformed parent_child_rel get_disconnected_nodes get_disconnected_nodes_locs known_ptrs remove_child remove_child_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_document get_disconnected_document_locs by(auto simp add: l_assigned_nodes_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_assigned_nodes_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \get\_shadow\_root\_safe\ locale l_get_shadow_root_safe_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs + l_type_wf type_wf + l_get_shadow_root_safe\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf get_shadow_root_safe get_shadow_root_safe_locs get_shadow_root get_shadow_root_locs get_mode get_mode_locs for known_ptr :: "(_::linorder) object_ptr \ bool" and known_ptrs :: "(_) heap \ bool" and type_wf :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_shadow_root_safe :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_safe_locs :: "(_) element_ptr \ (_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" and 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" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed :: "(_) heap \ bool" and parent_child_rel :: "(_) heap \ ((_) object_ptr \ (_) object_ptr) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_mode :: "(_) shadow_root_ptr \ ((_) heap, exception, shadow_root_mode) prog" and get_mode_locs :: "(_) shadow_root_ptr \ ((_) heap \ (_) heap \ bool) set" begin end subsubsection \create\_element\ locale l_create_element_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs + l_set_tag_name type_wf set_tag_name set_tag_name_locs + l_create_element_defs create_element + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs + 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>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^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 type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + 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_tag_name_get_tag_name type_wf get_tag_name get_tag_name_locs set_tag_name set_tag_name_locs + l_new_element_get_tag_name type_wf get_tag_name get_tag_name_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_get_shadow_root set_disconnected_nodes set_disconnected_nodes_locs get_shadow_root get_shadow_root_locs + l_set_disconnected_nodes_get_tag_name type_wf set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_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_get_shadow_root type_wf get_shadow_root get_shadow_root_locs + l_set_tag_name_get_shadow_root type_wf set_tag_name set_tag_name_locs get_shadow_root get_shadow_root_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" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_document :: "(_) node_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) set" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_::linorder) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" 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 CD.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 CD.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) then have element_ptr_kinds_eq_h2: "element_ptr_kinds h3 = element_ptr_kinds h2" by(simp add: element_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) then have element_ptr_kinds_eq_h3: "element_ptr_kinds h' = element_ptr_kinds h3" by(simp add: element_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 CD.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 CD.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 tag_name_eq_h: "\ptr' disc_nodes. ptr' \ new_element_ptr \ h \ get_tag_name ptr' \\<^sub>r disc_nodes = h2 \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads h2 get_tag_name_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 tag_name_eq2_h: "\ptr'. ptr' \ new_element_ptr \ |h \ get_tag_name ptr'|\<^sub>r = |h2 \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_tag_name 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_empty_tag_name 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 CD.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 CD.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 CD.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 tag_name_eq_h2: "\ptr' disc_nodes. ptr' \ new_element_ptr \ h2 \ get_tag_name ptr' \\<^sub>r disc_nodes = h3 \ get_tag_name ptr' \\<^sub>r disc_nodes" apply(rule reads_writes_preserved[OF get_tag_name_reads set_tag_name_writes h3]) by (metis local.set_tag_name_get_tag_name_different_pointers) then have tag_name_eq2_h2: "\ptr'. ptr' \ new_element_ptr \ |h2 \ get_tag_name ptr'|\<^sub>r = |h3 \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have "h2 \ get_tag_name 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_empty_tag_name by blast 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 CD.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 CD.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 tag_name_eq_h2: "\ptr' disc_nodes. ptr' \ new_element_ptr \ h2 \ get_tag_name ptr' \\<^sub>r disc_nodes = h3 \ get_tag_name ptr' \\<^sub>r disc_nodes" apply(rule reads_writes_preserved[OF get_tag_name_reads set_tag_name_writes h3]) by (metis local.set_tag_name_get_tag_name_different_pointers) then have tag_name_eq2_h2: "\ptr'. ptr' \ new_element_ptr \ |h2 \ get_tag_name ptr'|\<^sub>r = |h3 \ get_tag_name 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 tag_name_eq_h3: "\ptr' disc_nodes. h3 \ get_tag_name ptr' \\<^sub>r disc_nodes = h' \ get_tag_name ptr' \\<^sub>r disc_nodes" apply(rule reads_writes_preserved[OF get_tag_name_reads set_disconnected_nodes_writes h']) using set_disconnected_nodes_get_tag_name by blast then have tag_name_eq2_h3: "\ptr'. |h3 \ get_tag_name ptr'|\<^sub>r = |h' \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) also have "parent_child_rel h = parent_child_rel h2" proof(auto simp add: CD.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: CD.parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2) also have "\ = parent_child_rel h'" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3) finally have "CD.a_acyclic_heap h'" by (simp add: CD.acyclic_heap_def) have "CD.a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_all_ptrs_in_heap h2" apply(auto simp add: CD.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 CD.get_child_nodes_ok local.known_ptrs_known_ptr local.parent_child_rel_child_in_heap CD.parent_child_rel_child_nodes2 node_ptr_kinds_commutes node_ptr_kinds_eq_h returns_result_select_result) by (metis (no_types, lifting) CD.get_child_nodes_ok CD.get_child_nodes_ptr_in_heap \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 []\ assms(3) assms(4) children_eq_h disconnected_nodes_eq2_h document_ptr_kinds_eq_h finite_set_in is_OK_returns_result_I local.known_ptrs_known_ptr node_ptr_kinds_commutes returns_result_select_result subsetD) then have "CD.a_all_ptrs_in_heap h3" by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2) then have "CD.a_all_ptrs_in_heap h'" - by (smt children_eq2_h3 disc_nodes_h3 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 - element_ptr_kinds_commutes h' h2 local.CD.a_all_ptrs_in_heap_def - local.set_disconnected_nodes_get_disconnected_nodes new_element_ptr new_element_ptr_in_heap - node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h3 notin_fset object_ptr_kinds_eq_h3 select_result_I2 - set_ConsD subset_code(1)) - + by (smt (verit) children_eq2_h3 disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h2 + disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 finite_set_in finsertCI funion_finsert_right + h' insert_iff list.simps(15) local.CD.a_all_ptrs_in_heap_def + local.set_disconnected_nodes_get_disconnected_nodes node_ptr_kinds_eq_h node_ptr_kinds_eq_h2 + node_ptr_kinds_eq_h3 object_ptr_kinds_eq_h3 select_result_I2 subsetD subsetI) 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 CD.a_all_ptrs_in_heap_def assms(3) assms(4) fset_mp fset_of_list_elem CD.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 "CD.a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: CD.heap_is_wellformed_def heap_is_wellformed_def) then have "CD.a_distinct_lists h2" using \h2 \ get_child_nodes (cast new_element_ptr) \\<^sub>r []\ apply(auto simp add: CD.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 CD.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 \ CD.a_distinct_lists h\ \type_wf h2\ disconnected_nodes_eq_h document_ptr_kinds_eq_h CD.distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result) then have " CD.a_distinct_lists h3" by(auto simp add: CD.a_distinct_lists_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 children_eq2_h2 object_ptr_kinds_eq_h2) then have " CD.a_distinct_lists h'" proof(auto simp add: CD.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\ \ CD.a_distinct_lists h3\ \type_wf h'\ disc_nodes_h3 distinct.simps(2) CD.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\ \CD.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' - set_disconnected_nodes_get_disconnected_nodes - CD.a_all_ptrs_in_heap_def - 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\ \CD.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' + by (smt (verit) 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.CD.a_all_ptrs_in_heap h\ + disc_nodes_document_ptr_h disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 + disconnected_nodes_eq2_h3 disjoint_iff 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 - CD.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms + local.CD.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\ \ CD.a_distinct_lists h3\ children_eq2_h3 disc_nodes_h3 CD.distinct_lists_no_parent h' select_result_I2 set_ConsD set_disconnected_nodes_get_disconnected_nodes) by (metis "3" "4" "5" "6" \ CD.a_distinct_lists h3\ \type_wf h3\ children_eq2_h3 CD.distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result) qed have "CD.a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h'" using disc_nodes_h3 \document_ptr |\| document_ptr_kinds h\ apply(auto simp add: CD.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) have "CD.a_heap_is_wellformed h'" using \CD.a_acyclic_heap h'\ \CD.a_all_ptrs_in_heap h'\ \CD.a_distinct_lists h'\ \CD.a_owner_document_valid h'\ by(simp add: CD.a_heap_is_wellformed_def) have shadow_root_ptr_kinds_eq_h: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: shadow_root_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: shadow_root_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h' = shadow_root_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: shadow_root_ptr_kinds_def) have shadow_root_eq_h: "\element_ptr shadow_root_opt. element_ptr \ new_element_ptr \ h \ get_shadow_root element_ptr \\<^sub>r shadow_root_opt = h2 \ get_shadow_root element_ptr \\<^sub>r shadow_root_opt" proof - fix element_ptr shadow_root_opt assume "element_ptr \ new_element_ptr " have "\P \ get_shadow_root_locs element_ptr. P h h2" using get_shadow_root_new_element new_element_ptr h2 using \element_ptr \ new_element_ptr\ by blast then have "preserved (get_shadow_root element_ptr) h h2" using get_shadow_root_new_element[rotated, OF new_element_ptr h2] using get_shadow_root_reads by(simp add: reads_def) then show "h \ get_shadow_root element_ptr \\<^sub>r shadow_root_opt = h2 \ get_shadow_root element_ptr \\<^sub>r shadow_root_opt" by (simp add: preserved_def) qed have shadow_root_none: "h2 \ get_shadow_root (new_element_ptr) \\<^sub>r None" 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_shadow_root by blast have shadow_root_eq_h2: "\ptr' children. h2 \ get_shadow_root ptr' \\<^sub>r children = h3 \ get_shadow_root ptr' \\<^sub>r children" using get_shadow_root_reads set_tag_name_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_tag_name_get_shadow_root) have shadow_root_eq_h3: "\ptr' children. h3 \ get_shadow_root ptr' \\<^sub>r children = h' \ get_shadow_root ptr' \\<^sub>r children" using get_shadow_root_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) using set_disconnected_nodes_get_shadow_root by(auto simp add: set_disconnected_nodes_get_shadow_root) have "a_all_ptrs_in_heap h" by (simp add: assms(1) local.a_all_ptrs_in_heap_def local.get_shadow_root_shadow_root_ptr_in_heap) then have "a_all_ptrs_in_heap h2" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h)[1] using returns_result_eq shadow_root_eq_h shadow_root_none by fastforce then have "a_all_ptrs_in_heap h3" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h2)[1] using shadow_root_eq_h2 by blast then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h3)[1] by (simp add: shadow_root_eq_h3) have "a_distinct_lists h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h2" apply(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h)[1] apply(auto simp add: distinct_insort intro!: distinct_concat_map_I split: option.splits)[1] apply(case_tac "x = new_element_ptr") using shadow_root_none apply auto[1] using shadow_root_eq_h - by (smt Diff_empty Diff_insert0 ElementMonad.ptr_kinds_M_ptr_kinds + by (smt (verit) Diff_empty Diff_insert0 ElementMonad.ptr_kinds_M_ptr_kinds ElementMonad.ptr_kinds_ptr_kinds_M assms(1) assms(3) finite_set_in h2 insort_split local.get_shadow_root_ok local.shadow_root_same_host new_element_ptr new_element_ptr_not_in_heap option.distinct(1) returns_result_select_result select_result_I2 shadow_root_none) then have "a_distinct_lists h3" by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2]) then have "a_distinct_lists h'" by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h3 select_result_eq[OF shadow_root_eq_h3]) have "a_shadow_root_valid h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_shadow_root_valid h2" proof (unfold a_shadow_root_valid_def; safe) fix shadow_root_ptr assume "\shadow_root_ptr\fset (shadow_root_ptr_kinds h). \host\fset (element_ptr_kinds h). |h \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" assume "shadow_root_ptr \ fset (shadow_root_ptr_kinds h2)" obtain previous_host where "previous_host \ fset (element_ptr_kinds h)" and "|h \ get_tag_name previous_host|\<^sub>r \ safe_shadow_root_element_types" and "|h \ get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr" by (metis \local.a_shadow_root_valid h\ \shadow_root_ptr \ fset (shadow_root_ptr_kinds h2)\ local.a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h) moreover have "previous_host \ new_element_ptr" using calculation(1) h2 new_element_ptr new_element_ptr_not_in_heap by auto ultimately have "|h2 \ get_tag_name previous_host|\<^sub>r \ safe_shadow_root_element_types" and "|h2 \ get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr" using shadow_root_eq_h apply (simp add: tag_name_eq2_h) by (metis \previous_host \ new_element_ptr\ \|h \ get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr\ select_result_eq shadow_root_eq_h) then show "\host\fset (element_ptr_kinds h2). |h2 \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h2 \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" by (meson \previous_host \ fset (element_ptr_kinds h)\ \previous_host \ new_element_ptr\ assms(3) local.get_shadow_root_ok local.get_shadow_root_ptr_in_heap notin_fset returns_result_select_result shadow_root_eq_h) qed then have "a_shadow_root_valid h3" proof (unfold a_shadow_root_valid_def; safe) fix shadow_root_ptr assume "\shadow_root_ptr\fset (shadow_root_ptr_kinds h2). \host\fset (element_ptr_kinds h2). |h2 \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h2 \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" assume "shadow_root_ptr \ fset (shadow_root_ptr_kinds h3)" obtain previous_host where "previous_host \ fset (element_ptr_kinds h2)" and "|h2 \ get_tag_name previous_host|\<^sub>r \ safe_shadow_root_element_types" and "|h2 \ get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr" by (metis \local.a_shadow_root_valid h2\ \shadow_root_ptr \ fset (shadow_root_ptr_kinds h3)\ local.a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h2) moreover have "previous_host \ new_element_ptr" using calculation(1) h3 new_element_ptr new_element_ptr_not_in_heap using calculation(3) shadow_root_none by auto ultimately have "|h2 \ get_tag_name previous_host|\<^sub>r \ safe_shadow_root_element_types" and "|h2 \ get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr" using shadow_root_eq_h2 apply (simp add: tag_name_eq2_h2) by (metis \previous_host \ new_element_ptr\ \|h2 \ get_shadow_root previous_host|\<^sub>r = Some shadow_root_ptr\ select_result_eq shadow_root_eq_h) then show "\host\fset (element_ptr_kinds h3). |h3 \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h3 \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" - by (smt \previous_host \ fset (element_ptr_kinds h2)\ \previous_host \ new_element_ptr\ - \type_wf h2\ \type_wf h3\ element_ptr_kinds_eq_h2 finite_set_in local.get_shadow_root_ok - returns_result_eq returns_result_select_result shadow_root_eq_h2 tag_name_eq2_h2) + by (metis \previous_host \ fset (element_ptr_kinds h2)\ \previous_host \ new_element_ptr\ + element_ptr_kinds_eq_h2 select_result_eq shadow_root_eq_h2 tag_name_eq2_h2) qed then have "a_shadow_root_valid h'" - apply(auto simp add: a_shadow_root_valid_def element_ptr_kinds_eq_h3 shadow_root_eq_h3 - shadow_root_ptr_kinds_eq_h3 tag_name_eq2_h3)[1] - by (smt \type_wf h3\ finite_set_in local.get_shadow_root_ok returns_result_select_result - select_result_I2 shadow_root_eq_h3) - - + by (smt (verit) \type_wf h3\ element_ptr_kinds_eq_h3 finite_set_in local.a_shadow_root_valid_def + local.get_shadow_root_ok returns_result_select_result select_result_I2 shadow_root_eq_h3 + shadow_root_ptr_kinds_eq_h3 tag_name_eq2_h3) have "a_host_shadow_root_rel h = a_host_shadow_root_rel h2" apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h shadow_root_eq_h)[1] - apply (smt assms(3) case_prod_conv h2 image_iff local.get_shadow_root_ok mem_Collect_eq - new_element_ptr new_element_ptr_not_in_heap returns_result_select_result select_result_I2 - shadow_root_eq_h) + apply (smt (verit, del_insts) assms(3) case_prodI h2 local.get_shadow_root_ok mem_Collect_eq + new_element_ptr new_element_ptr_not_in_heap pair_imageI returns_result_select_result + select_result_I2 shadow_root_eq_h) using shadow_root_none apply auto[1] - by (metis (no_types, lifting) Collect_cong assms(3) case_prodE case_prodI h2 - host_shadow_root_rel_def host_shadow_root_rel_shadow_root local.a_host_shadow_root_rel_def - local.get_shadow_root_impl local.get_shadow_root_ok local.new_element_no_shadow_root - new_element_ptr option.distinct(1) returns_result_select_result select_result_I2 shadow_root_eq_h) + by (metis (no_types, lifting) assms(3) h2 host_shadow_root_rel_def + host_shadow_root_rel_shadow_root local.a_host_shadow_root_rel_def local.get_shadow_root_impl + local.get_shadow_root_ok local.new_element_no_shadow_root new_element_ptr option.distinct(1) + returns_result_select_result select_result_I2 shadow_root_eq_h) have "a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3" apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 shadow_root_eq_h2)[1] - apply (smt Collect_cong Shadow_DOM.a_host_shadow_root_rel_def assms(3) h2 - host_shadow_root_rel_shadow_root is_OK_returns_result_E local.get_shadow_root_impl - local.get_shadow_root_ok local.new_element_types_preserved select_result_I2 shadow_root_eq_h2 - split_cong) - apply (metis (no_types, lifting) Collect_cong \type_wf h3\ element_ptr_kinds_eq_h2 - host_shadow_root_rel_def host_shadow_root_rel_shadow_root local.a_host_shadow_root_rel_def - local.get_shadow_root_impl local.get_shadow_root_ok returns_result_select_result shadow_root_eq_h2 - split_cong) + apply (metis (mono_tags, lifting) \type_wf h2\ case_prodI local.get_shadow_root_ok + mem_Collect_eq pair_imageI returns_result_select_result select_result_I2 shadow_root_eq_h2) + apply (metis (mono_tags, lifting) case_prodI mem_Collect_eq pair_imageI select_result_eq shadow_root_eq_h2) done have "a_host_shadow_root_rel h3 = a_host_shadow_root_rel h'" apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 shadow_root_eq_h2)[1] - apply (metis (no_types, lifting) Collect_cong \type_wf h3\ case_prodE case_prodI - element_ptr_kinds_eq_h2 host_shadow_root_rel_def host_shadow_root_rel_shadow_root - local.a_host_shadow_root_rel_def local.get_shadow_root_impl local.get_shadow_root_ok - returns_result_select_result shadow_root_eq_h3) - apply (smt Collect_cong \type_wf h'\ \type_wf h2\ case_prodD case_prodI2 host_shadow_root_rel_def - host_shadow_root_rel_shadow_root is_OK_returns_result_E local.a_host_shadow_root_rel_def - local.get_shadow_root_impl local.get_shadow_root_ok select_result_I2 shadow_root_eq_h2 shadow_root_eq_h3) + apply (metis (mono_tags, lifting) \type_wf h'\ case_prodI element_ptr_kinds_eq_h2 + element_ptr_kinds_eq_h3 local.get_shadow_root_ok mem_Collect_eq pair_imageI + returns_result_select_result select_result_I2 shadow_root_eq_h3) + apply (metis (mono_tags, lifting) \type_wf h'\ case_prodI element_ptr_kinds_eq_h2 + element_ptr_kinds_eq_h3 local.get_shadow_root_ok mem_Collect_eq pair_imageI + returns_result_select_result select_result_I2 shadow_root_eq_h3) done - have "acyclic (parent_child_rel h \ a_host_shadow_root_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) have "parent_child_rel h \ a_host_shadow_root_rel h = parent_child_rel h2 \ a_host_shadow_root_rel h2" using \local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h2\ \parent_child_rel h = parent_child_rel h2\ by auto have "parent_child_rel h2 \ a_host_shadow_root_rel h2 = parent_child_rel h3 \ a_host_shadow_root_rel h3" using \local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\ \parent_child_rel h2 = parent_child_rel h3\ by auto have "parent_child_rel h' \ a_host_shadow_root_rel h' = parent_child_rel h3 \ a_host_shadow_root_rel h3" by (simp add: \local.a_host_shadow_root_rel h3 = local.a_host_shadow_root_rel h'\ \parent_child_rel h3 = parent_child_rel h'\) have "acyclic (parent_child_rel h3 \ a_host_shadow_root_rel h3)" using \acyclic (parent_child_rel h \ local.a_host_shadow_root_rel h)\ \parent_child_rel h \ local.a_host_shadow_root_rel h = parent_child_rel h2 \ local.a_host_shadow_root_rel h2\ \parent_child_rel h2 \ local.a_host_shadow_root_rel h2 = parent_child_rel h3 \ local.a_host_shadow_root_rel h3\ by auto then have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h')" by(simp add: \parent_child_rel h' \ a_host_shadow_root_rel h' = parent_child_rel h3 \ a_host_shadow_root_rel h3\) show " heap_is_wellformed h' " using \acyclic (parent_child_rel h' \ local.a_host_shadow_root_rel h')\ by(simp add: heap_is_wellformed_def CD.heap_is_wellformed_impl \local.CD.a_heap_is_wellformed h'\ \local.a_all_ptrs_in_heap h'\ \local.a_distinct_lists h'\ \local.a_shadow_root_valid h'\) qed end interpretation i_create_element_wf?: l_create_element_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^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 get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs DocumentClass.known_ptr DocumentClass.type_wf by(auto simp add: l_create_element_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_create_element_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] subsubsection \create\_character\_data\ locale l_create_character_data_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs + l_create_character_data\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^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 create_character_data known_ptr type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + 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_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_set_val_get_shadow_root type_wf set_val set_val_locs get_shadow_root get_shadow_root_locs + l_set_disconnected_nodes_get_shadow_root set_disconnected_nodes set_disconnected_nodes_locs get_shadow_root get_shadow_root_locs + l_new_character_data_get_tag_name get_tag_name get_tag_name_locs + l_set_val_get_tag_name type_wf set_val set_val_locs get_tag_name get_tag_name_locs + l_get_tag_name type_wf get_tag_name get_tag_name_locs + l_set_disconnected_nodes_get_tag_name type_wf set_disconnected_nodes set_disconnected_nodes_locs get_tag_name get_tag_name_locs + l_new_character_data 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" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_document :: "(_) node_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) 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 create_character_data :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) character_data_ptr) prog" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_::linorder) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) 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: CD.create_character_data_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF CD.get_disconnected_nodes_pure, rotated] ) then have "h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr" apply(auto simp add: CD.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.CD.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 CD.set_val_writes h3]) using CD.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 CD.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 CD.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 CD.set_val_writes h3]) using CD.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) then have character_data_ptr_kinds_eq_h2: "character_data_ptr_kinds h3 = character_data_ptr_kinds h2" by(simp add: character_data_ptr_kinds_def) have element_ptr_kinds_eq_h2: "element_ptr_kinds h3 = element_ptr_kinds h2" using node_ptr_kinds_eq_h2 by(simp add: element_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) then have character_data_ptr_kinds_eq_h3: "character_data_ptr_kinds h' = character_data_ptr_kinds h3" by(simp add: character_data_ptr_kinds_def) have element_ptr_kinds_eq_h3: "element_ptr_kinds h' = element_ptr_kinds h3" using node_ptr_kinds_eq_h3 by(simp add: element_ptr_kinds_def) have "document_ptr |\| document_ptr_kinds h" using disc_nodes_h3 document_ptr_kinds_eq_h object_ptr_kinds_eq_h2 CD.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 CD.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 CD.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 tag_name_eq_h: "\ptr' disc_nodes. h \ get_tag_name ptr' \\<^sub>r disc_nodes = h2 \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads h2 get_tag_name_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 tag_name_eq2_h: "\ptr'. |h \ get_tag_name ptr'|\<^sub>r = |h2 \ get_tag_name 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 CD.get_child_nodes_reads CD.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 CD.get_disconnected_nodes_reads CD.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 tag_name_eq_h2: "\ptr' disc_nodes. h2 \ get_tag_name ptr' \\<^sub>r disc_nodes = h3 \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads CD.set_val_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_val_get_tag_name) then have tag_name_eq2_h2: "\ptr'. |h2 \ get_tag_name ptr'|\<^sub>r = |h3 \ get_tag_name 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 CD.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 CD.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 CD.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 tag_name_eq_h3: "\ptr' disc_nodes. h3 \ get_tag_name ptr' \\<^sub>r disc_nodes = h' \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_disconnected_nodes_get_tag_name) then have tag_name_eq2_h3: "\ptr'. |h3 \ get_tag_name ptr'|\<^sub>r = |h' \ get_tag_name 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 CD.heap_is_wellformed_def CD.acyclic_heap_def) also have "parent_child_rel h = parent_child_rel h2" proof(auto simp add: CD.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: CD.parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2) also have "\ = parent_child_rel h'" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3) finally have "CD.a_acyclic_heap h'" by (simp add: CD.acyclic_heap_def) have "CD.a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_all_ptrs_in_heap h2" apply(auto simp add: CD.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 CD.parent_child_rel_child CD.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 (no_types, lifting) CD.get_child_nodes_ok CD.get_child_nodes_ptr_in_heap \h2 \ get_child_nodes (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) \\<^sub>r []\ assms(3) assms(4) children_eq_h disconnected_nodes_eq2_h document_ptr_kinds_eq_h finite_set_in is_OK_returns_result_I local.known_ptrs_known_ptr node_ptr_kinds_commutes returns_result_select_result subset_code(1)) then have "CD.a_all_ptrs_in_heap h3" by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2) then have "CD.a_all_ptrs_in_heap h'" - by (smt character_data_ptr_kinds_commutes character_data_ptr_kinds_eq_h2 children_eq2_h3 - disc_nodes_h3 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 h' h2 - local.CD.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_h3 notin_fset - object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1)) - + by (smt (verit) children_eq2_h3 disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h2 + disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 finite_set_in finsertCI funion_finsert_right + h' local.CD.a_all_ptrs_in_heap_def local.set_disconnected_nodes_get_disconnected_nodes + node_ptr_kinds_eq_h node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h3 object_ptr_kinds_eq_h3 + select_result_I2 set_ConsD subsetD subsetI) 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 CD.a_all_ptrs_in_heap_def assms(3) assms(4) fset_mp fset_of_list_elem CD.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 "CD.a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_distinct_lists h2" using \h2 \ get_child_nodes (cast new_character_data_ptr) \\<^sub>r []\ apply(auto simp add: CD.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 CD.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] thm children_eq2_h using \CD.a_distinct_lists h\ \type_wf h2\ disconnected_nodes_eq_h document_ptr_kinds_eq_h CD.distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result by metis then have "CD.a_distinct_lists h3" by(auto simp add: CD.a_distinct_lists_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 children_eq2_h2 object_ptr_kinds_eq_h2)[1] then have "CD.a_distinct_lists h'" proof(auto simp add: CD.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, opaque_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>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr \ set disc_nodes_h3\ \type_wf h2\ assms(1) disc_nodes_document_ptr_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disconnected_nodes_eq_h distinct.simps(2) document_ptr_kinds_eq_h2 local.get_disconnected_nodes_ok local.heap_is_wellformed_disconnected_nodes_distinct returns_result_select_result select_result_I2) 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" using 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\ - - by (smt local.CD.a_all_ptrs_in_heap_def \CD.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 + by (smt (verit) \local.CD.a_all_ptrs_in_heap h\ disc_nodes_document_ptr_h + disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff 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 + local.CD.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 "document_ptr = xb") apply (metis (no_types, lifting) "3" "4" "5" "6" CD.distinct_lists_no_parent \local.CD.a_distinct_lists h2\ \type_wf h'\ children_eq2_h2 children_eq2_h3 disc_nodes_document_ptr_h2 document_ptr_kinds_eq_h3 h' local.get_disconnected_nodes_ok local.set_disconnected_nodes_get_disconnected_nodes new_character_data_ptr_not_in_any_children object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 returns_result_eq returns_result_select_result set_ConsD) by (metis "3" "4" "5" "6" CD.distinct_lists_no_parent \local.CD.a_distinct_lists h3\ \type_wf h3\ children_eq2_h3 local.get_disconnected_nodes_ok returns_result_select_result) qed have "CD.a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h'" using disc_nodes_h3 \document_ptr |\| document_ptr_kinds h\ apply(simp add: CD.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) have shadow_root_ptr_kinds_eq_h: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h" using object_ptr_kinds_eq_h by(auto simp add: shadow_root_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h2" using object_ptr_kinds_eq_h2 by(auto simp add: shadow_root_ptr_kinds_def) have shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h' = shadow_root_ptr_kinds h3" using object_ptr_kinds_eq_h3 by(auto simp add: shadow_root_ptr_kinds_def) have shadow_root_eq_h: "\character_data_ptr shadow_root_opt. h \ get_shadow_root character_data_ptr \\<^sub>r shadow_root_opt = h2 \ get_shadow_root character_data_ptr \\<^sub>r shadow_root_opt" using get_shadow_root_reads h2 get_shadow_root_new_character_data[rotated, OF h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] using local.get_shadow_root_locs_impl new_character_data_ptr apply blast using local.get_shadow_root_locs_impl new_character_data_ptr by blast have shadow_root_eq_h2: "\ptr' children. h2 \ get_shadow_root ptr' \\<^sub>r children = h3 \ get_shadow_root ptr' \\<^sub>r children" using get_shadow_root_reads set_val_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_val_get_shadow_root) have shadow_root_eq_h3: "\ptr' children. h3 \ get_shadow_root ptr' \\<^sub>r children = h' \ get_shadow_root ptr' \\<^sub>r children" using get_shadow_root_reads set_disconnected_nodes_writes h' apply(rule reads_writes_preserved) using set_disconnected_nodes_get_shadow_root by(auto simp add: set_disconnected_nodes_get_shadow_root) have "a_all_ptrs_in_heap h" by (simp add: assms(1) local.a_all_ptrs_in_heap_def local.get_shadow_root_shadow_root_ptr_in_heap) then have "a_all_ptrs_in_heap h2" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h)[1] using returns_result_eq shadow_root_eq_h by fastforce then have "a_all_ptrs_in_heap h3" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h2)[1] using shadow_root_eq_h2 by blast then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h3)[1] by (simp add: shadow_root_eq_h3) have "a_distinct_lists h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h2" apply(auto simp add: a_distinct_lists_def character_data_ptr_kinds_eq_h)[1] apply(auto simp add: distinct_insort intro!: distinct_concat_map_I split: option.splits)[1] by (metis \type_wf h2\ assms(1) assms(3) local.get_shadow_root_ok local.shadow_root_same_host returns_result_select_result shadow_root_eq_h) then have "a_distinct_lists h3" by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2]) then have "a_distinct_lists h'" by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h3 select_result_eq[OF shadow_root_eq_h3]) have "a_shadow_root_valid h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_shadow_root_valid h2" by(auto simp add: a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h element_ptr_kinds_eq_h select_result_eq[OF shadow_root_eq_h] tag_name_eq2_h) then have "a_shadow_root_valid h3" by(auto simp add: a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h2 element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2] tag_name_eq2_h2) then have "a_shadow_root_valid h'" by(auto simp add: a_shadow_root_valid_def shadow_root_ptr_kinds_eq_h3 element_ptr_kinds_eq_h3 select_result_eq[OF shadow_root_eq_h3] tag_name_eq2_h3) have "a_host_shadow_root_rel h = a_host_shadow_root_rel h2" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h select_result_eq[OF shadow_root_eq_h]) have "a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2]) have "a_host_shadow_root_rel h3 = a_host_shadow_root_rel h'" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h3 select_result_eq[OF shadow_root_eq_h3]) have "acyclic (parent_child_rel h \ a_host_shadow_root_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) have "parent_child_rel h \ a_host_shadow_root_rel h = parent_child_rel h2 \ a_host_shadow_root_rel h2" using \local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h2\ \parent_child_rel h = parent_child_rel h2\ by auto have "parent_child_rel h2 \ a_host_shadow_root_rel h2 = parent_child_rel h3 \ a_host_shadow_root_rel h3" using \local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\ \parent_child_rel h2 = parent_child_rel h3\ by auto have "parent_child_rel h' \ a_host_shadow_root_rel h' = parent_child_rel h3 \ a_host_shadow_root_rel h3" by (simp add: \local.a_host_shadow_root_rel h3 = local.a_host_shadow_root_rel h'\ \parent_child_rel h3 = parent_child_rel h'\) have "acyclic (parent_child_rel h3 \ a_host_shadow_root_rel h3)" using \acyclic (parent_child_rel h \ local.a_host_shadow_root_rel h)\ \parent_child_rel h \ local.a_host_shadow_root_rel h = parent_child_rel h2 \ local.a_host_shadow_root_rel h2\ \parent_child_rel h2 \ local.a_host_shadow_root_rel h2 = parent_child_rel h3 \ local.a_host_shadow_root_rel h3\ by auto then have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h')" by(simp add: \parent_child_rel h' \ a_host_shadow_root_rel h' = parent_child_rel h3 \ a_host_shadow_root_rel h3\) have "CD.a_heap_is_wellformed h'" apply(simp add: CD.a_heap_is_wellformed_def) by (simp add: \local.CD.a_acyclic_heap h'\ \local.CD.a_all_ptrs_in_heap h'\ \local.CD.a_distinct_lists h'\ \local.CD.a_owner_document_valid h'\) show "heap_is_wellformed h' " using \acyclic (parent_child_rel h' \ local.a_host_shadow_root_rel h')\ by(simp add: heap_is_wellformed_def CD.heap_is_wellformed_impl \local.CD.a_heap_is_wellformed h'\ \local.a_all_ptrs_in_heap h'\ \local.a_distinct_lists h'\ \local.a_shadow_root_valid h'\) qed end subsubsection \create\_document\ locale l_create_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs + 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_get_tag_name type_wf get_tag_name get_tag_name_locs + l_new_document_get_tag_name get_tag_name get_tag_name_locs + l_get_disconnected_nodes\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_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 type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) 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 get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_document :: "(_) node_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_disconnected_document_locs :: "((_) 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 shadow_root_ptr_kinds_eq: "shadow_root_ptr_kinds h' = shadow_root_ptr_kinds h" using object_ptr_kinds_eq apply(simp add: shadow_root_ptr_kinds_def) by force 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 CD.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 CD.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 (auto simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) also have "parent_child_rel h = parent_child_rel h'" proof(auto simp add: CD.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 "CD.a_acyclic_heap h'" by (simp add: CD.acyclic_heap_def) have "CD.a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def ) then have "CD.a_all_ptrs_in_heap h'" apply(auto simp add: CD.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 CD.parent_child_rel_child CD.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) \h' \ get_disconnected_nodes new_document_ptr \\<^sub>r []\ \type_wf h'\ assms(1) disconnected_nodes_eq_h empty_iff empty_set local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq returns_result_select_result select_result_I2) have "CD.a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.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] CD.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] apply (metis assms(1) assms(3) disconnected_nodes_eq2_h 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 \CD.a_distinct_lists h\ assms(3) disconnected_nodes_eq2_h CD.distinct_lists_no_parent get_disconnected_nodes_ok returns_result_select_result) qed have "CD.a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h'" apply(auto simp add: CD.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) have shadow_root_eq_h: "\character_data_ptr shadow_root_opt. h \ get_shadow_root character_data_ptr \\<^sub>r shadow_root_opt = h' \ get_shadow_root character_data_ptr \\<^sub>r shadow_root_opt" using get_shadow_root_reads assms(2) get_shadow_root_new_document[rotated, OF h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] using local.get_shadow_root_locs_impl new_document_ptr apply blast using local.get_shadow_root_locs_impl new_document_ptr by blast have "a_all_ptrs_in_heap h" by (simp add: assms(1) local.a_all_ptrs_in_heap_def local.get_shadow_root_shadow_root_ptr_in_heap) then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq document_ptr_kinds_eq_h)[1] using shadow_root_eq_h by fastforce have "a_distinct_lists h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h'" apply(auto simp add: a_distinct_lists_def character_data_ptr_kinds_eq_h)[1] apply(auto simp add: distinct_insort intro!: distinct_concat_map_I split: option.splits)[1] by (metis \type_wf h'\ assms(1) assms(3) local.get_shadow_root_ok local.shadow_root_same_host returns_result_select_result shadow_root_eq_h) have tag_name_eq_h: "\ptr' disc_nodes. h \ get_tag_name ptr' \\<^sub>r disc_nodes = h' \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads h' get_tag_name_new_document[OF new_document_ptr h'] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ have "a_shadow_root_valid h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_shadow_root_valid h'" using new_document_is_document_ptr[OF new_document_ptr] by(auto simp add: a_shadow_root_valid_def element_ptr_kinds_eq_h document_ptr_kinds_eq_h shadow_root_ptr_kinds_eq select_result_eq[OF shadow_root_eq_h] select_result_eq[OF tag_name_eq_h]) have "a_host_shadow_root_rel h = a_host_shadow_root_rel h'" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h select_result_eq[OF shadow_root_eq_h]) have "acyclic (parent_child_rel h \ a_host_shadow_root_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) moreover have "parent_child_rel h \ a_host_shadow_root_rel h = parent_child_rel h' \ a_host_shadow_root_rel h'" by (simp add: \local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h'\ \parent_child_rel h = parent_child_rel h'\) ultimately have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h')" by simp have "CD.a_heap_is_wellformed h'" apply(simp add: CD.a_heap_is_wellformed_def) by (simp add: \local.CD.a_acyclic_heap h'\ \local.CD.a_all_ptrs_in_heap h'\ \local.CD.a_distinct_lists h'\ \local.CD.a_owner_document_valid h'\) show "heap_is_wellformed h'" using CD.heap_is_wellformed_impl \acyclic (parent_child_rel h' \ local.a_host_shadow_root_rel h')\ \local.CD.a_heap_is_wellformed h'\ \local.a_all_ptrs_in_heap h'\ \local.a_distinct_lists h'\ \local.a_shadow_root_valid h'\ local.heap_is_wellformed_def by auto qed end interpretation l_create_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr type_wf DocumentClass.type_wf get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs heap_is_wellformed parent_child_rel set_val set_val_locs set_disconnected_nodes set_disconnected_nodes_locs create_document known_ptrs by(auto simp add: l_create_document_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) subsubsection \attach\_shadow\_root\ locale l_attach_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_child_nodes get_child_nodes_locs get_disconnected_nodes get_disconnected_nodes_locs get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs known_ptr type_wf heap_is_wellformed parent_child_rel heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs + l_attach_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M known_ptr set_shadow_root set_shadow_root_locs set_mode set_mode_locs attach_shadow_root type_wf get_tag_name get_tag_name_locs get_shadow_root get_shadow_root_locs + l_new_shadow_root_get_disconnected_nodes get_disconnected_nodes get_disconnected_nodes_locs + l_set_mode_get_disconnected_nodes type_wf set_mode set_mode_locs get_disconnected_nodes get_disconnected_nodes_locs + l_new_shadow_root_get_child_nodes type_wf known_ptr get_child_nodes get_child_nodes_locs + l_new_shadow_root_get_tag_name type_wf get_tag_name get_tag_name_locs + l_set_mode_get_child_nodes type_wf set_mode set_mode_locs known_ptr get_child_nodes get_child_nodes_locs + l_set_shadow_root_get_child_nodes type_wf set_shadow_root set_shadow_root_locs known_ptr get_child_nodes get_child_nodes_locs + l_set_shadow_root type_wf set_shadow_root set_shadow_root_locs + l_set_shadow_root_get_disconnected_nodes set_shadow_root set_shadow_root_locs get_disconnected_nodes get_disconnected_nodes_locs + l_set_mode_get_shadow_root type_wf set_mode set_mode_locs get_shadow_root get_shadow_root_locs + l_set_shadow_root_get_shadow_root type_wf set_shadow_root set_shadow_root_locs get_shadow_root get_shadow_root_locs + l_new_character_data_get_tag_name get_tag_name get_tag_name_locs + l_set_mode_get_tag_name type_wf set_mode set_mode_locs get_tag_name get_tag_name_locs + l_get_tag_name type_wf get_tag_name get_tag_name_locs + l_set_shadow_root_get_tag_name set_shadow_root set_shadow_root_locs get_tag_name get_tag_name_locs + l_new_shadow_root 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" and get_shadow_root :: "(_) element_ptr \ ((_) heap, exception, (_) shadow_root_ptr option) prog" and get_shadow_root_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and get_tag_name :: "(_) element_ptr \ ((_) heap, exception, char list) prog" and get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" and heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and get_host :: "(_) shadow_root_ptr \ ((_) heap, exception, (_) element_ptr) prog" and get_host_locs :: "((_) heap \ (_) heap \ bool) set" and get_disconnected_document :: "(_) node_ptr \ ((_) heap, exception, (_) document_ptr) prog" and get_disconnected_document_locs :: "((_) heap \ (_) heap \ bool) 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 create_character_data :: "(_) document_ptr \ char list \ ((_) heap, exception, (_) character_data_ptr) prog" and known_ptr\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_::linorder) object_ptr \ bool" and type_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M :: "(_) heap \ bool" and set_shadow_root :: "(_) element_ptr \ (_) shadow_root_ptr option \ (_, unit) dom_prog" and set_shadow_root_locs :: "(_) element_ptr \ (_, unit) dom_prog set" and set_mode :: "(_) shadow_root_ptr \ shadow_root_mode \ (_, unit) dom_prog" and set_mode_locs :: "(_) shadow_root_ptr \ (_, unit) dom_prog set" and attach_shadow_root :: "(_) element_ptr \ shadow_root_mode \ (_, (_) shadow_root_ptr) dom_prog" begin lemma attach_shadow_root_child_preserves: assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "h \ attach_shadow_root element_ptr new_mode \\<^sub>h h'" shows "type_wf h'" and "known_ptrs h'" and "heap_is_wellformed h'" proof - obtain h2 h3 new_shadow_root_ptr element_tag_name where element_tag_name: "h \ get_tag_name element_ptr \\<^sub>r element_tag_name" and "element_tag_name \ safe_shadow_root_element_types" and prev_shadow_root: "h \ get_shadow_root element_ptr \\<^sub>r None" and h2: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h2" and new_shadow_root_ptr: "h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr" and h3: "h2 \ set_mode new_shadow_root_ptr new_mode \\<^sub>h h3" and h': "h3 \ set_shadow_root element_ptr (Some new_shadow_root_ptr) \\<^sub>h h'" using assms(4) by(auto simp add: attach_shadow_root_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_tag_name_pure, rotated] bind_returns_heap_E2[rotated, OF get_shadow_root_pure, rotated] split: if_splits) have "h \ attach_shadow_root element_ptr new_mode \\<^sub>r new_shadow_root_ptr" thm bind_pure_returns_result_I[OF get_tag_name_pure] apply(unfold attach_shadow_root_def)[1] using element_tag_name apply(rule bind_pure_returns_result_I[OF get_tag_name_pure]) apply(rule bind_pure_returns_result_I) using \element_tag_name \ safe_shadow_root_element_types\ apply(simp) using \element_tag_name \ safe_shadow_root_element_types\ apply(simp) using prev_shadow_root apply(rule bind_pure_returns_result_I[OF get_shadow_root_pure]) apply(rule bind_pure_returns_result_I) apply(simp) apply(simp) using h2 new_shadow_root_ptr h3 h' by(auto intro!: bind_returns_result_I intro: is_OK_returns_result_E[OF is_OK_returns_heap_I[OF h3]] is_OK_returns_result_E[OF is_OK_returns_heap_I[OF h']]) have "new_shadow_root_ptr \ set |h \ shadow_root_ptr_kinds_M|\<^sub>r" using new_shadow_root_ptr ShadowRootMonad.ptr_kinds_ptr_kinds_M h2 using new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_not_in_heap by blast then have "cast new_shadow_root_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_shadow_root_ptr|}" using new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_new_ptr h2 new_shadow_root_ptr by blast then have document_ptr_kinds_eq_h: "document_ptr_kinds h2 = document_ptr_kinds h" apply(simp add: document_ptr_kinds_def) by force have shadow_root_ptr_kinds_eq_h: "shadow_root_ptr_kinds h2 = shadow_root_ptr_kinds h |\| {|new_shadow_root_ptr|}" using object_ptr_kinds_eq_h apply(simp add: shadow_root_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 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_mode_writes h3]) using set_mode_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 shadow_root_ptr_kinds_eq_h2: "shadow_root_ptr_kinds h3 = shadow_root_ptr_kinds h2" using object_ptr_kinds_eq_h2 by (auto simp add: shadow_root_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_shadow_root_writes h']) using set_shadow_root_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 shadow_root_ptr_kinds_eq_h3: "shadow_root_ptr_kinds h' = shadow_root_ptr_kinds h3" using object_ptr_kinds_eq_h3 by (auto simp add: shadow_root_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_shadow_root_ptr)" using \h \ attach_shadow_root element_ptr new_mode \\<^sub>r new_shadow_root_ptr\ create_shadow_root_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 "element_ptr |\| element_ptr_kinds h" by (meson \h \ attach_shadow_root element_ptr new_mode \\<^sub>r new_shadow_root_ptr\ is_OK_returns_result_I local.attach_shadow_root_element_ptr_in_heap) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_shadow_root_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads h2 get_child_nodes_new_shadow_root[rotated, OF new_shadow_root_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_shadow_root_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_shadow_root_ptr|}" using new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_new_ptr h2 new_shadow_root_ptr object_ptr_kinds_eq_h by blast then have node_ptr_kinds_eq_h: "node_ptr_kinds h2 = 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 h2 = character_data_ptr_kinds h" apply(simp add: character_data_ptr_kinds_def) done 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 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_mode_writes h3]) using set_mode_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) then have character_data_ptr_kinds_eq_h2: "character_data_ptr_kinds h3 = character_data_ptr_kinds h2" by(simp add: character_data_ptr_kinds_def) have element_ptr_kinds_eq_h2: "element_ptr_kinds h3 = element_ptr_kinds h2" using node_ptr_kinds_eq_h2 by(simp add: element_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_shadow_root_writes h']) using set_shadow_root_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) then have character_data_ptr_kinds_eq_h3: "character_data_ptr_kinds h' = character_data_ptr_kinds h3" by(simp add: character_data_ptr_kinds_def) have element_ptr_kinds_eq_h3: "element_ptr_kinds h' = element_ptr_kinds h3" using node_ptr_kinds_eq_h3 by(simp add: element_ptr_kinds_def) have children_eq_h: "\(ptr'::(_) object_ptr) children. ptr' \ cast new_shadow_root_ptr \ h \ get_child_nodes ptr' \\<^sub>r children = h2 \ get_child_nodes ptr' \\<^sub>r children" using CD.get_child_nodes_reads h2 get_child_nodes_new_shadow_root[rotated, OF new_shadow_root_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_shadow_root_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_shadow_root_ptr) \\<^sub>r []" using h2 local.new_shadow_root_no_child_nodes new_shadow_root_ptr by auto 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_shadow_root[rotated, OF h2,rotated,OF new_shadow_root_ptr] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by (metis (no_types, lifting))+ 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 tag_name_eq_h: "\ptr' disc_nodes. h \ get_tag_name ptr' \\<^sub>r disc_nodes = h2 \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads h2 get_tag_name_new_shadow_root[OF new_shadow_root_ptr h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] by blast+ then have tag_name_eq2_h: "\ptr'. |h \ get_tag_name ptr'|\<^sub>r = |h2 \ get_tag_name 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 CD.get_child_nodes_reads set_mode_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_mode_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_mode_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_mode_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 tag_name_eq_h2: "\ptr' disc_nodes. h2 \ get_tag_name ptr' \\<^sub>r disc_nodes = h3 \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads set_mode_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_mode_get_tag_name) then have tag_name_eq2_h2: "\ptr'. |h2 \ get_tag_name ptr'|\<^sub>r = |h3 \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have "type_wf h2" using \type_wf h\ new_shadow_root_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_mode_writes h3] using set_mode_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_shadow_root_writes h'] using set_shadow_root_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 CD.get_child_nodes_reads set_shadow_root_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_shadow_root_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. 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_shadow_root_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_shadow_root_get_disconnected_nodes) 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 tag_name_eq_h3: "\ptr' disc_nodes. h3 \ get_tag_name ptr' \\<^sub>r disc_nodes = h' \ get_tag_name ptr' \\<^sub>r disc_nodes" using get_tag_name_reads set_shadow_root_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_shadow_root_get_tag_name) then have tag_name_eq2_h3: "\ptr'. |h3 \ get_tag_name ptr'|\<^sub>r = |h' \ get_tag_name ptr'|\<^sub>r" using select_result_eq by force have "acyclic (parent_child_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def CD.acyclic_heap_def) also have "parent_child_rel h = parent_child_rel h2" proof(auto simp add: CD.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_shadow_root_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_shadow_root_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_shadow_root_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: CD.parent_child_rel_def object_ptr_kinds_eq_h2 children_eq2_h2) also have "\ = parent_child_rel h'" by(auto simp add: CD.parent_child_rel_def object_ptr_kinds_eq_h3 children_eq2_h3) finally have "CD.a_acyclic_heap h'" by (simp add: CD.acyclic_heap_def) have "CD.a_all_ptrs_in_heap h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_all_ptrs_in_heap h2" apply(auto simp add: CD.a_all_ptrs_in_heap_def)[1] using node_ptr_kinds_eq_h \h2 \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []\ apply (metis (no_types, lifting) CD.get_child_nodes_ok CD.l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms \known_ptrs h2\ \parent_child_rel h = parent_child_rel h2\ \type_wf h2\ assms(1) assms(2) l_heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M.parent_child_rel_child local.known_ptrs_known_ptr local.parent_child_rel_child_in_heap node_ptr_kinds_commutes returns_result_select_result) by (metis assms(1) assms(2) disconnected_nodes_eq2_h document_ptr_kinds_eq_h local.get_disconnected_nodes_ok local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq_h returns_result_select_result) then have "CD.a_all_ptrs_in_heap h3" by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2) then have "CD.a_all_ptrs_in_heap h'" by (simp add: children_eq2_h3 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 CD.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h3 object_ptr_kinds_eq_h3) have "CD.a_distinct_lists h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_distinct_lists h2" using \h2 \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []\ children_eq2_h apply(auto simp add: select_result_eq[OF disconnected_nodes_eq_h] CD.a_distinct_lists_def insort_split object_ptr_kinds_eq_h document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I dest: distinct_concat_map_E)[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(case_tac "x = cast new_shadow_root_ptr") using \h2 \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []\ children_eq2_h apply blast apply(case_tac "y = cast new_shadow_root_ptr") using \h2 \ get_child_nodes (cast new_shadow_root_ptr) \\<^sub>r []\ children_eq2_h apply blast proof - fix x y :: "(_) object_ptr" fix xa :: "(_) node_ptr" assume a1: "distinct (concat (map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) (sorted_list_of_set (fset (object_ptr_kinds h)))))" assume "x \ y" assume "xa \ set |h2 \ get_child_nodes x|\<^sub>r" assume "xa \ set |h2 \ get_child_nodes y|\<^sub>r" assume "x |\| object_ptr_kinds h" assume "x \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^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_shadow_root_ptr" assume "y |\| object_ptr_kinds h" assume "y \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^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_shadow_root_ptr" show False using distinct_concat_map_E(1)[OF a1, of x y] using \x |\| object_ptr_kinds h\ \y |\| object_ptr_kinds h\ using \xa \ set |h2 \ get_child_nodes x|\<^sub>r\ \xa \ set |h2 \ get_child_nodes y|\<^sub>r\ using \x \ y\ by(auto simp add: children_eq2_h[OF \x \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^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_shadow_root_ptr\] children_eq2_h[OF \y \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^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_shadow_root_ptr\]) next fix x :: "(_) node_ptr" fix xa :: "(_) object_ptr" fix xb :: "(_) document_ptr" assume "(\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) = {}" assume "x \ set |h2 \ get_child_nodes xa|\<^sub>r" assume "xb |\| document_ptr_kinds h" assume "x \ set |h2 \ get_disconnected_nodes xb|\<^sub>r" assume "xa |\| object_ptr_kinds h" assume "xa \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^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_shadow_root_ptr" have "set |h \ get_child_nodes xa|\<^sub>r \ set |h2 \ get_disconnected_nodes xb|\<^sub>r = {}" by (metis (no_types, lifting) CD.get_child_nodes_ok \xa |\| object_ptr_kinds h\ \xb |\| document_ptr_kinds h\ assms(1) assms(2) assms(3) disconnected_nodes_eq2_h is_OK_returns_result_E local.get_disconnected_nodes_ok local.heap_is_wellformed_children_disc_nodes_different local.known_ptrs_known_ptr select_result_I2) then show "False" using \x \ set |h2 \ get_child_nodes xa|\<^sub>r\ \x \ set |h2 \ get_disconnected_nodes xb|\<^sub>r\ \xa \ cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^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_shadow_root_ptr\ children_eq2_h by auto qed then have "CD.a_distinct_lists h3" by(auto simp add: CD.a_distinct_lists_def disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 children_eq2_h2 object_ptr_kinds_eq_h2)[1] then have "CD.a_distinct_lists h'" by(auto simp add: CD.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) have "CD.a_owner_document_valid h" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def CD.heap_is_wellformed_def) then have "CD.a_owner_document_valid h'" (* using disc_nodes_h3 \document_ptr |\| document_ptr_kinds h\ *) apply(simp add: CD.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] by (metis CD.get_child_nodes_ok CD.get_child_nodes_ptr_in_heap \cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^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_shadow_root_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ assms(2) assms(3) children_eq2_h children_eq_h document_ptr_kinds_eq_h finite_set_in is_OK_returns_result_I l_ptr_kinds_M.ptr_kinds_ptr_kinds_M local.known_ptrs_known_ptr object_ptr_kinds_M_def returns_result_select_result) have shadow_root_eq_h: "\character_data_ptr shadow_root_opt. h \ get_shadow_root character_data_ptr \\<^sub>r shadow_root_opt = h2 \ get_shadow_root character_data_ptr \\<^sub>r shadow_root_opt" using get_shadow_root_reads h2 get_shadow_root_new_shadow_root[rotated, OF h2] apply(auto simp add: reads_def reflp_def transp_def preserved_def)[1] using local.get_shadow_root_locs_impl new_shadow_root_ptr apply blast using local.get_shadow_root_locs_impl new_shadow_root_ptr by blast have shadow_root_eq_h2: "\ptr' children. h2 \ get_shadow_root ptr' \\<^sub>r children = h3 \ get_shadow_root ptr' \\<^sub>r children" using get_shadow_root_reads set_mode_writes h3 apply(rule reads_writes_preserved) by(auto simp add: set_mode_get_shadow_root) have shadow_root_eq_h3: "\ptr' children. element_ptr \ ptr' \ h3 \ get_shadow_root ptr' \\<^sub>r children = h' \ get_shadow_root ptr' \\<^sub>r children" using get_shadow_root_reads set_shadow_root_writes h' apply(rule reads_writes_preserved) by(auto simp add: set_shadow_root_get_shadow_root_different_pointers) have shadow_root_h3: "h' \ get_shadow_root element_ptr \\<^sub>r Some new_shadow_root_ptr" using \type_wf h3\ h' local.set_shadow_root_get_shadow_root by blast have "a_all_ptrs_in_heap h" by (simp add: assms(1) local.a_all_ptrs_in_heap_def local.get_shadow_root_shadow_root_ptr_in_heap) then have "a_all_ptrs_in_heap h2" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h)[1] using returns_result_eq shadow_root_eq_h by fastforce then have "a_all_ptrs_in_heap h3" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h2)[1] using shadow_root_eq_h2 by blast then have "a_all_ptrs_in_heap h'" apply(auto simp add: a_all_ptrs_in_heap_def shadow_root_ptr_kinds_eq_h3)[1] apply(case_tac "shadow_root_ptr = new_shadow_root_ptr") using h2 new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_in_heap new_shadow_root_ptr shadow_root_ptr_kinds_eq_h2 apply blast using \type_wf h3\ h' local.set_shadow_root_get_shadow_root returns_result_eq shadow_root_eq_h3 apply fastforce done have "a_distinct_lists h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h2" apply(auto simp add: a_distinct_lists_def character_data_ptr_kinds_eq_h)[1] apply(auto simp add: distinct_insort intro!: distinct_concat_map_I split: option.splits)[1] by (metis \type_wf h2\ assms(1) assms(2) local.get_shadow_root_ok local.shadow_root_same_host returns_result_select_result shadow_root_eq_h) then have "a_distinct_lists h3" by(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2]) then have "a_distinct_lists h'" apply(auto simp add: a_distinct_lists_def element_ptr_kinds_eq_h3 select_result_eq[OF shadow_root_eq_h3])[1] apply(auto simp add: distinct_insort intro!: distinct_concat_map_I split: option.splits)[1] - by (smt \type_wf h3\ assms(1) assms(2) h' h2 local.get_shadow_root_ok - local.get_shadow_root_shadow_root_ptr_in_heap local.set_shadow_root_get_shadow_root - local.shadow_root_same_host new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_not_in_heap new_shadow_root_ptr - returns_result_select_result select_result_I2 shadow_root_eq_h shadow_root_eq_h2 shadow_root_eq_h3) - + by (smt (verit, best) ShadowRootMonad.ptr_kinds_ptr_kinds_M \new_shadow_root_ptr \ set |h \ + shadow_root_ptr_kinds_M|\<^sub>r\ \type_wf h'\ assms(1) assms(2) element_ptr_kinds_eq_h3 + l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M.shadow_root_same_host local.get_shadow_root_ok + local.get_shadow_root_shadow_root_ptr_in_heap + local.l_heap_is_wellformed\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms returns_result_select_result + select_result_I2 shadow_root_eq_h shadow_root_eq_h2 shadow_root_eq_h3 shadow_root_h3) have "a_shadow_root_valid h" using assms(1) by (simp add: heap_is_wellformed_def) then have "a_shadow_root_valid h'" proof(unfold a_shadow_root_valid_def; safe) fix shadow_root_ptr assume "\shadow_root_ptr\fset (shadow_root_ptr_kinds h). \host\fset (element_ptr_kinds h). |h \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" assume "a_shadow_root_valid h" assume "shadow_root_ptr \ fset (shadow_root_ptr_kinds h')" show "\host\fset (element_ptr_kinds h'). |h' \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h' \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" proof (cases "shadow_root_ptr = new_shadow_root_ptr") case True have "element_ptr \ fset (element_ptr_kinds h')" by (simp add: \element_ptr |\| element_ptr_kinds h\ element_ptr_kinds_eq_h element_ptr_kinds_eq_h2 element_ptr_kinds_eq_h3) moreover have "|h' \ get_tag_name element_ptr|\<^sub>r \ safe_shadow_root_element_types" - by (smt \\thesis. (\h2 h3 new_shadow_root_ptr element_tag_name. -\h \ get_tag_name element_ptr \\<^sub>r element_tag_name; -element_tag_name \ safe_shadow_root_element_types; -h \ get_shadow_root element_ptr \\<^sub>r None; h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h2; -h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr; h2 \ set_mode new_shadow_root_ptr new_mode \\<^sub>h h3; -h3 \ set_shadow_root element_ptr (Some new_shadow_root_ptr) \\<^sub>h h'\ \ thesis) \ thesis\ - select_result_I2 tag_name_eq2_h tag_name_eq2_h2 tag_name_eq2_h3) + by (smt (verit, best) \\thesis. (\h2 h3 new_shadow_root_ptr element_tag_name. \h \ + get_tag_name element_ptr \\<^sub>r element_tag_name; element_tag_name \ + safe_shadow_root_element_types; h \ get_shadow_root element_ptr \\<^sub>r None; h \ + new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>h h2; h \ new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M \\<^sub>r new_shadow_root_ptr; h2 + \ set_mode new_shadow_root_ptr new_mode \\<^sub>h h3; h3 \ set_shadow_root element_ptr (Some + new_shadow_root_ptr) \\<^sub>h h'\ \ thesis) \ thesis\ select_result_I2 tag_name_eq2_h2 + tag_name_eq2_h3 tag_name_eq_h) moreover have "|h' \ get_shadow_root element_ptr|\<^sub>r = Some shadow_root_ptr" using shadow_root_h3 by (simp add: True) ultimately show ?thesis by meson next case False then obtain host where host: "host \ fset (element_ptr_kinds h)" and "|h \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types" and "|h \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr" using \shadow_root_ptr \ fset (shadow_root_ptr_kinds h')\ using \\shadow_root_ptr\fset (shadow_root_ptr_kinds h). \host\fset (element_ptr_kinds h). |h \ get_tag_name host|\<^sub>r \ safe_shadow_root_element_types \ |h \ get_shadow_root host|\<^sub>r = Some shadow_root_ptr\ apply(simp add: shadow_root_ptr_kinds_eq_h3 shadow_root_ptr_kinds_eq_h2 shadow_root_ptr_kinds_eq_h) by (meson finite_set_in) moreover have "host \ element_ptr" using calculation(3) prev_shadow_root by auto ultimately show ?thesis using element_ptr_kinds_eq_h3 element_ptr_kinds_eq_h2 element_ptr_kinds_eq_h - by (smt \type_wf h'\ assms(2) finite_set_in local.get_shadow_root_ok returns_result_eq - returns_result_select_result shadow_root_eq_h shadow_root_eq_h2 shadow_root_eq_h3 tag_name_eq2_h + by (metis (no_types, lifting) assms(2) + l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M.get_shadow_root_ok + local.l_get_shadow_root\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms notin_fset returns_result_select_result + select_result_I2 shadow_root_eq_h shadow_root_eq_h2 shadow_root_eq_h3 tag_name_eq2_h tag_name_eq2_h2 tag_name_eq2_h3) qed qed have "a_host_shadow_root_rel h = a_host_shadow_root_rel h2" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h select_result_eq[OF shadow_root_eq_h]) have "a_host_shadow_root_rel h2 = a_host_shadow_root_rel h3" by(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h2 select_result_eq[OF shadow_root_eq_h2]) have "a_host_shadow_root_rel h' = {(cast element_ptr, cast new_shadow_root_ptr)} \ a_host_shadow_root_rel h3" apply(auto simp add: a_host_shadow_root_rel_def element_ptr_kinds_eq_h3 )[1] apply(case_tac "element_ptr \ aa") using select_result_eq[OF shadow_root_eq_h3] apply (simp add: image_iff) using select_result_eq[OF shadow_root_eq_h3] apply (metis (no_types, lifting) \local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h2\ \local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\ \type_wf h3\ host_shadow_root_rel_def local.get_shadow_root_impl local.get_shadow_root_ok option.distinct(1) prev_shadow_root returns_result_select_result) apply (metis (mono_tags, lifting) \\ptr'. (\x. element_ptr \ ptr') \ |h3 \ get_shadow_root ptr'|\<^sub>r = |h' \ get_shadow_root ptr'|\<^sub>r\ case_prod_conv image_iff is_OK_returns_result_I mem_Collect_eq option.inject returns_result_eq returns_result_select_result shadow_root_h3) using element_ptr_kinds_eq_h3 local.get_shadow_root_ptr_in_heap shadow_root_h3 apply fastforce using Shadow_DOM.a_host_shadow_root_rel_def \\ptr'. (\x. element_ptr \ ptr') \ |h3 \ get_shadow_root ptr'|\<^sub>r = |h' \ get_shadow_root ptr'|\<^sub>r\ \type_wf h3\ case_prodE case_prodI host_shadow_root_rel_shadow_root image_iff local.get_shadow_root_impl local.get_shadow_root_ok mem_Collect_eq option.discI prev_shadow_root returns_result_select_result select_result_I2 shadow_root_eq_h shadow_root_eq_h2 - apply(auto)[1] - by (smt case_prodI mem_Collect_eq option.distinct(1) pair_imageI returns_result_eq returns_result_select_result) + by (smt (verit)) have "acyclic (parent_child_rel h \ a_host_shadow_root_rel h)" using \heap_is_wellformed h\ by (simp add: heap_is_wellformed_def) have "parent_child_rel h \ a_host_shadow_root_rel h = parent_child_rel h2 \ a_host_shadow_root_rel h2" using \local.a_host_shadow_root_rel h = local.a_host_shadow_root_rel h2\ \parent_child_rel h = parent_child_rel h2\ by auto have "parent_child_rel h2 \ a_host_shadow_root_rel h2 = parent_child_rel h3 \ a_host_shadow_root_rel h3" using \local.a_host_shadow_root_rel h2 = local.a_host_shadow_root_rel h3\ \parent_child_rel h2 = parent_child_rel h3\ by auto have "parent_child_rel h' \ a_host_shadow_root_rel h' = {(cast element_ptr, cast new_shadow_root_ptr)} \ parent_child_rel h3 \ a_host_shadow_root_rel h3" by (simp add: \local.a_host_shadow_root_rel 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>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r element_ptr, cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^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_shadow_root_ptr)} \ local.a_host_shadow_root_rel h3\ \parent_child_rel h3 = parent_child_rel h'\) have "\a b. (a, b) \ parent_child_rel h3 \ a \ cast new_shadow_root_ptr" using CD.parent_child_rel_parent_in_heap \parent_child_rel h = parent_child_rel h2\ \parent_child_rel h2 = parent_child_rel h3\ document_ptr_kinds_commutes by (metis h2 new\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>R\<^sub>o\<^sub>o\<^sub>t_M_ptr_not_in_heap new_shadow_root_ptr shadow_root_ptr_kinds_commutes) moreover have "\a b. (a, b) \ a_host_shadow_root_rel h3 \ a \ cast new_shadow_root_ptr" using shadow_root_eq_h2 by(auto simp add: a_host_shadow_root_rel_def) moreover have "cast new_shadow_root_ptr \ {x. (x, cast element_ptr) \ (parent_child_rel h3 \ a_host_shadow_root_rel h3)\<^sup>*}" by (metis (no_types, lifting) UnE calculation(1) calculation(2) cast_shadow_root_ptr_not_node_ptr(1) converse_rtranclE mem_Collect_eq) moreover have "acyclic (parent_child_rel h3 \ a_host_shadow_root_rel h3)" using \acyclic (parent_child_rel h \ local.a_host_shadow_root_rel h)\ \parent_child_rel h \ local.a_host_shadow_root_rel h = parent_child_rel h2 \ local.a_host_shadow_root_rel h2\ \parent_child_rel h2 \ local.a_host_shadow_root_rel h2 = parent_child_rel h3 \ local.a_host_shadow_root_rel h3\ by auto ultimately have "acyclic (parent_child_rel h' \ a_host_shadow_root_rel h')" by(simp add: \parent_child_rel h' \ a_host_shadow_root_rel h' = {(cast element_ptr, cast new_shadow_root_ptr)} \ parent_child_rel h3 \ a_host_shadow_root_rel h3\) have "CD.a_heap_is_wellformed h'" apply(simp add: CD.a_heap_is_wellformed_def) by (simp add: \local.CD.a_acyclic_heap h'\ \local.CD.a_all_ptrs_in_heap h'\ \local.CD.a_distinct_lists h'\ \local.CD.a_owner_document_valid h'\) show "heap_is_wellformed h' " using \acyclic (parent_child_rel h' \ local.a_host_shadow_root_rel h')\ by(simp add: heap_is_wellformed_def CD.heap_is_wellformed_impl \local.CD.a_heap_is_wellformed h'\ \local.a_all_ptrs_in_heap h'\ \local.a_distinct_lists h'\ \local.a_shadow_root_valid h'\) qed end interpretation l_attach_shadow_root_wf?: l_attach_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^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 get_shadow_root get_shadow_root_locs get_tag_name get_tag_name_locs heap_is_wellformed\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_host get_host_locs get_disconnected_document get_disconnected_document_locs set_val set_val_locs create_character_data DocumentClass.known_ptr DocumentClass.type_wf set_shadow_root set_shadow_root_locs set_mode set_mode_locs attach_shadow_root by(auto simp add: l_attach_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def instances) declare l_attach_shadow_root_wf\<^sub>S\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] end diff --git a/thys/SimplifiedOntologicalArgument/DisableKodkodScala.thy b/thys/SimplifiedOntologicalArgument/DisableKodkodScala.thy deleted file mode 100644 --- a/thys/SimplifiedOntologicalArgument/DisableKodkodScala.thy +++ /dev/null @@ -1,14 +0,0 @@ -theory DisableKodkodScala - imports Main -begin - -text \Some of the nitpick invocation within this AFP entry do not work, - if "Kodkod Scala" is enabled, i.e., if the box under - Plugin Options — Isabelle — General — Miscelleaneous Tools — Kodkod Scala - is activated. Therefore, in this theory we explicitly disable this configuration option.\ - -ML \ - Options.default_put_bool \<^system_option>\kodkod_scala\ false -\ - -end \ No newline at end of file diff --git a/thys/SimplifiedOntologicalArgument/SimpleVariant.thy b/thys/SimplifiedOntologicalArgument/SimpleVariant.thy --- a/thys/SimplifiedOntologicalArgument/SimpleVariant.thy +++ b/thys/SimplifiedOntologicalArgument/SimpleVariant.thy @@ -1,42 +1,41 @@ subsection\Simplified Variant (Fig.~6 in \cite{C85})\ theory SimpleVariant imports HOML MFilter BaseDefs - DisableKodkodScala begin text\Axiom's of new, simplified variant.\ axiomatization where A1': "\\<^bold>\(\

(\x.(x\<^bold>\x)))\" and A2': "\\<^bold>\X Y.(((\

X) \<^bold>\ ((X\<^bold>\Y) \<^bold>\ (X\Y))) \<^bold>\ (\

Y))\" and A3: "\\<^bold>\\.((\

\\ \) \<^bold>\ (\<^bold>\X.((X\\) \<^bold>\ (\

X))))\" lemma T2: "\\

\\" by (metis A3 G_def) \\From A3\ lemma L1: "\\

(\x.(x\<^bold>=x))\" by (metis A2' A3) text\Necessary existence of a Godlike entity.\ theorem T6: "\\<^bold>\(\<^bold>\\<^sup>E \)\" \\Proof found by sledgehammer\ proof - have T1: "\\<^bold>\X.((\

X) \<^bold>\ \<^bold>\(\<^bold>\\<^sup>E X))\" by (metis A1' A2') have T3: "\\<^bold>\(\<^bold>\\<^sup>E \)\" using T1 T2 by simp have T5: "\(\<^bold>\(\<^bold>\\<^sup>E \)) \<^bold>\ \<^bold>\(\<^bold>\\<^sup>E \)\" by (metis A1' A2' T2) thus ?thesis using T3 by blast qed lemma True nitpick[satisfy] oops \\Consistency: model found\ text\Modal collapse and monotheism: not implied.\ lemma MC: "\\<^bold>\\.(\ \<^bold>\ \<^bold>\\)\" nitpick oops \\Countermodel\ lemma MT: "\\<^bold>\x y.(((\ x) \<^bold>\ (\ y)) \<^bold>\ (x\<^bold>=y))\" nitpick oops \\Countermodel.\ text\Gödel's A1, A4, A5: not implied anymore.\ lemma A1: "\\<^bold>\X.((\<^bold>\(\

X))\<^bold>\(\

(\<^bold>\X)))\" nitpick oops \\Countermodel\ lemma A4: "\\<^bold>\X.((\

X) \<^bold>\ \<^bold>\(\

X))\" nitpick oops \\Countermodel\ lemma A5: "\\

\\\" nitpick oops \\Countermodel\ text\Checking filter and ultrafilter properties.\ theorem F1: "\Filter \

\" oops \\Proof found by sledgehammer, but reconstruction timeout\ theorem U1: "\UFilter \

\" nitpick oops \\Countermodel\ end diff --git a/thys/SimplifiedOntologicalArgument/SimpleVariantHF.thy b/thys/SimplifiedOntologicalArgument/SimpleVariantHF.thy --- a/thys/SimplifiedOntologicalArgument/SimpleVariantHF.thy +++ b/thys/SimplifiedOntologicalArgument/SimpleVariantHF.thy @@ -1,39 +1,38 @@ subsection\Hauptfiltervariant (Fig.~10 in \cite{C85})\ theory SimpleVariantHF imports HOML MFilter BaseDefs - DisableKodkodScala begin text\Definition: Set of supersets of \X\, we call this \\\ X\.\ abbreviation HF::"\\(\\\)" where "HF X \ \Y.(X\<^bold>\Y)" text\Postulate: \\\ \\ is a filter; i.e., Hauptfilter of \\\.\ axiomatization where F1: "\Filter (HF \)\" text\Necessary existence of a Godlike entity.\ theorem T6: "\\<^bold>\(\<^bold>\\<^sup>E \)\" using F1 by auto \\Proof found\ theorem T6again: "\\<^bold>\(\<^bold>\\<^sup>E \)\" proof - have T3': "\\<^bold>\\<^sup>E \\" using F1 by auto have T6: "\\<^bold>\(\<^bold>\\<^sup>E \)\" using T3' by blast thus ?thesis by simp qed text\Possible existence of Godlike entity not implied.\ lemma T3: "\\<^bold>\(\<^bold>\\<^sup>E \)\" nitpick oops \\Countermodel\ text\Axiom T enforces possible existence of Godlike entity.\ axiomatization lemma T3: assumes T: "\\<^bold>\\.((\<^bold>\\) \<^bold>\ \)\" shows "\\<^bold>\(\<^bold>\\<^sup>E \)\" using F1 T by auto lemma True nitpick[satisfy] oops \\Consistency: model found\ text\Modal collapse: not implied anymore.\ lemma MC: "\\<^bold>\\.(\ \<^bold>\ \<^bold>\\)\" nitpick oops \\Countermodel\ lemma MT: "\\<^bold>\x y.(((\ x) \<^bold>\ (\ y)) \<^bold>\ (x\<^bold>=y))\" nitpick oops \\Countermodel\ end diff --git a/thys/SimplifiedOntologicalArgument/SimpleVariantPG.thy b/thys/SimplifiedOntologicalArgument/SimpleVariantPG.thy --- a/thys/SimplifiedOntologicalArgument/SimpleVariantPG.thy +++ b/thys/SimplifiedOntologicalArgument/SimpleVariantPG.thy @@ -1,28 +1,27 @@ subsection\Simplified Variant with Axiom T2 (Fig.~7 in \cite{C85})\ theory SimpleVariantPG imports HOML MFilter BaseDefs - DisableKodkodScala begin text\Axiom's of simplified variant with A3 replaced.\ axiomatization where A1': "\\<^bold>\(\

(\x.(x\<^bold>\x)))\" and A2': "\\<^bold>\X Y.(((\

X) \<^bold>\ ((X\<^bold>\Y)\<^bold>\(X\Y))) \<^bold>\ (\

Y))\" and T2: "\\

\\" text\Necessary existence of a Godlike entity.\ theorem T6: "\\<^bold>\(\<^bold>\\<^sup>E \)\" \\Proof found by sledgehammer\ proof - have T1: "\\<^bold>\X.((\

X) \<^bold>\ \<^bold>\(\<^bold>\\<^sup>E X))\" by (metis A1' A2') have T3: "\\<^bold>\(\<^bold>\\<^sup>E \)\" using T1 T2 by simp have T5: "\(\<^bold>\(\<^bold>\\<^sup>E \)) \<^bold>\ \<^bold>\(\<^bold>\\<^sup>E \)\" by (metis A1' A2' T2) thus ?thesis using T3 by blast qed lemma True nitpick[satisfy] oops \\Consistency: model found\ text\Modal collapse and Monotheism: not implied.\ lemma MC: "\\<^bold>\\.(\ \<^bold>\ \<^bold>\\)\" nitpick oops \\Countermodel\ lemma MT: "\\<^bold>\x y.(((\ x)\<^bold>\(\ y))\<^bold>\(x\<^bold>=y))\" nitpick oops \\Countermodel\ end diff --git a/thys/SimplifiedOntologicalArgument/SimplifiedOntologicalArgument.thy b/thys/SimplifiedOntologicalArgument/SimplifiedOntologicalArgument.thy --- a/thys/SimplifiedOntologicalArgument/SimplifiedOntologicalArgument.thy +++ b/thys/SimplifiedOntologicalArgument/SimplifiedOntologicalArgument.thy @@ -1,72 +1,71 @@ section \Selected Simplified Ontological Argument \label{sec:selected}\ theory SimplifiedOntologicalArgument imports HOML - DisableKodkodScala begin text \Positive properties:\ consts posProp::"\\\" ("\

") text \An entity x is God-like if it possesses all positive properties.\ definition G ("\") where "\(x) \ \<^bold>\\.(\

(\) \<^bold>\ \(x))" text \The axiom's of the simplified variant are presented next; these axioms are further motivated in \cite{C85,J52}).\ text \Self-difference is not a positive property (possible alternative: the empty property is not a positive property).\ axiomatization where CORO1: "\\<^bold>\(\

(\x.(x\<^bold>\x)))\" text \A property entailed by a positive property is positive.\ axiomatization where CORO2: "\\<^bold>\\ \. \

(\) \<^bold>\ (\<^bold>\x. \(x) \<^bold>\ \(x)) \<^bold>\ \

(\)\" text \Being Godlike is a positive property.\ axiomatization where AXIOM3: "\\

\\" subsection\Verifying the Selected Simplified Ontological Argument (version 1)\ text \The existence of a non-exemplified positive property implies that self-difference (or, alternatively, the empty property) is a positive property.\ lemma LEMMA1: "\(\<^bold>\\.(\

(\) \<^bold>\ \<^bold>\(\<^bold>\x. \(x)))) \<^bold>\ \

(\x.(x\<^bold>\x))\" using CORO2 by meson text \A non-exemplified positive property does not exist.\ lemma LEMMA2: "\\<^bold>\(\<^bold>\\.(\

(\) \<^bold>\ \<^bold>\(\<^bold>\x. \(x))))\" using CORO1 LEMMA1 by blast text \Positive properties are exemplified.\ lemma LEMMA3: "\\<^bold>\\.(\

(\) \<^bold>\ (\<^bold>\x. \(x)))\" using LEMMA2 by blast text \There exists a God-like entity.\ theorem THEOREM3': "\\<^bold>\x. \(x)\" using AXIOM3 LEMMA3 by auto text \Necessarily, there exists a God-like entity.\ theorem THEOREM3: "\\<^bold>\(\<^bold>\x. \(x))\" using THEOREM3' by simp text \However, the possible existence of Godlike entity is not implied.\ theorem CORO: "\\<^bold>\(\<^bold>\x. \(x))\" nitpick oops (*Countermodel*) subsection\Verifying the Selected Simplified Ontological Argument (version 2)\ text \We switch to logic T.\ axiomatization where T: "\\<^bold>\\. \<^bold>\\ \<^bold>\ \\" lemma T': "\\<^bold>\\. \ \<^bold>\ \<^bold>\\\" using T by metis text \Positive properties are possibly exemplified.\ theorem THEOREM1: "\\<^bold>\\. \

(\) \<^bold>\ \<^bold>\(\<^bold>\x. \(x))\" using CORO1 CORO2 T' by metis text \Possibly there exists a God-like entity.\ theorem CORO: "\\<^bold>\(\<^bold>\x. \(x))\" using AXIOM3 THEOREM1 by auto text \The possible existence of a God-like entity impplies the necessary existence of a God-like entity.\ theorem THEOREM2: "\\<^bold>\(\<^bold>\x. \(x)) \<^bold>\ \<^bold>\(\<^bold>\x. \(x))\" using AXIOM3 CORO1 CORO2 by metis text \Necessarily, there exists a God-like entity.\ theorem THEO3: "\\<^bold>\(\<^bold>\x. \(x))\" using CORO THEOREM2 by blast text \There exists a God-like entity.\ theorem THEO3': "\\<^bold>\x. \(x)\" using T THEO3 by metis text \Modal collapse is not implied; nitpick reports a countermodel.\ lemma MC: "\\<^bold>\\. \ \<^bold>\ \<^bold>\\\" nitpick oops text \Consistency of the theory; nitpick reports a model.\ lemma True nitpick[satisfy] oops (*Model found*) end diff --git a/thys/SimplifiedOntologicalArgument/UFilterVariant.thy b/thys/SimplifiedOntologicalArgument/UFilterVariant.thy --- a/thys/SimplifiedOntologicalArgument/UFilterVariant.thy +++ b/thys/SimplifiedOntologicalArgument/UFilterVariant.thy @@ -1,46 +1,45 @@ subsection\Ultrafilter Variant (Fig.~5 in \cite{C85})\ theory UFilterVariant imports HOML MFilter BaseDefs - DisableKodkodScala begin text\Axiom's of ultrafilter variant.\ axiomatization where U1: "\UFilter \

\" and A2: "\\<^bold>\X Y.(((\

X) \<^bold>\ (X\Y)) \<^bold>\ (\

Y))\" and A3: "\\<^bold>\\.((\

\\ \) \<^bold>\ (\<^bold>\X.((X\\) \<^bold>\ (\

X))))\" text\Necessary existence of a Godlike entity.\ theorem T6: "\\<^bold>\(\<^bold>\\<^sup>E \)\" \\Proof also found by sledgehammer\ proof - have T1: "\\<^bold>\X.((\

X) \<^bold>\ \<^bold>\(\<^bold>\\<^sup>E X))\" by (metis A2 U1) have T2: "\\

\\" by (metis A3 G_def) have T3: "\\<^bold>\(\<^bold>\\<^sup>E \)\" using T1 T2 by simp have T5: "\(\<^bold>\(\<^bold>\\<^sup>E \)) \<^bold>\ \<^bold>\(\<^bold>\\<^sup>E \)\" by (metis A2 G_def T2 U1) thus ?thesis using T3 by blast qed text\Checking for consistency.\ lemma True nitpick[satisfy] oops \\Model found\ text\Checking for modal collapse.\ lemma MC: "\\<^bold>\\.(\ \<^bold>\ \<^bold>\\)\" nitpick oops \\Countermodel\ end (* definition ess ("\") where "\ Y x \ Y x \<^bold>\ (\<^bold>\Z.(Z x \<^bold>\ (Y\Z)))" definition NE ("NE") where "NE x \ \w.((\<^bold>\Y.(\ Y x \<^bold>\ \<^bold>\\<^bold>\\<^sup>E Y)) w)" consts Godlike::\ UltraFilter::"(\\\)\\" NeEx::\ axiomatization where 1: "Godlike = G" and 2: "UltraFilter = Ultrafilter" and 3: "NeEx = NE" lemma True nitpick[satisfy] oops lemma MC: "\\<^bold>\\. \ \<^bold>\ \<^bold>\\\" nitpick[format=8] oops (*Countermodel*) lemma A1: "\\<^bold>\X. \<^bold>\(\

X) \<^bold>\ \

(\<^bold>\X)\" using U1 by fastforce lemma A4: "\(\

X \<^bold>\ \<^bold>\(\

X))\" nitpick [format=4] oops (*Countermodel*) lemma A5: "\\

NE\" nitpick oops (*Countermodel*) *) \ No newline at end of file