diff --git a/Core_DOM/ROOT b/Core_DOM/ROOT index dc5212e..771bdc6 100644 --- a/Core_DOM/ROOT +++ b/Core_DOM/ROOT @@ -1,7 +1,7 @@ chapter AFP session "Core_DOM-devel" (AFP) = "HOL-Library" + - options [timeout = 1200] + options [timeout = 2400] directories "common" "common/classes" @@ -15,6 +15,6 @@ session "Core_DOM-devel" (AFP) = "HOL-Library" + theories Core_DOM Core_DOM_Tests - document_files (in "document") + document_files "root.tex" "root.bib" diff --git a/Core_DOM/common/Core_DOM_Functions.thy b/Core_DOM/common/Core_DOM_Functions.thy index deae67a..da84e24 100644 --- a/Core_DOM/common/Core_DOM_Functions.thy +++ b/Core_DOM/common/Core_DOM_Functions.thy @@ -216,7 +216,7 @@ lemma get_child_nodes_reads: "reads (get_child_nodes_locs ptr) (get_child_nodes intro!: reads_bind_pure reads_subset[OF return_reads] )[1] apply(auto simp add: 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 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) + split: option.splits)[1] done end @@ -618,7 +618,8 @@ lemma set_child_nodes_get_child_nodes_different_pointers: apply(auto)[1] apply(auto)[1] apply(rule is_element_ptr_kind_obtains) - apply(auto) + apply(auto)[1] + apply(auto)[1] done lemma set_child_nodes_element_ok [simp]: @@ -630,10 +631,12 @@ lemma set_child_nodes_element_ok [simp]: proof - have "is_element_ptr ptr" using \known_ptr ptr\ assms(4) - 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) + 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) then show ?thesis using assms - apply(auto simp add: set_child_nodes_def a_set_child_nodes_tups_def set_child_nodes\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def split: option.splits)[1] + apply(auto simp add: set_child_nodes_def a_set_child_nodes_tups_def set_child_nodes\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def + split: option.splits)[1] by (simp add: DocumentMonad.put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok local.type_wf_impl) qed @@ -647,10 +650,12 @@ lemma set_child_nodes_document1_ok [simp]: proof - have "is_document_ptr ptr" using \known_ptr ptr\ assms(4) - 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) + 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) then show ?thesis using assms - apply(auto simp add: set_child_nodes_def a_set_child_nodes_tups_def set_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 split: option.splits)[1] + apply(auto simp add: set_child_nodes_def a_set_child_nodes_tups_def set_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 + split: option.splits)[1] by (simp add: DocumentMonad.put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok local.type_wf_impl) qed @@ -665,10 +670,11 @@ lemma set_child_nodes_document2_ok [simp]: proof - have "is_document_ptr ptr" using \known_ptr ptr\ assms(4) - 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) + 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) then show ?thesis using assms - apply(auto simp add: set_child_nodes_def a_set_child_nodes_tups_def set_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) + apply(auto simp add: set_child_nodes_def a_set_child_nodes_tups_def set_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)[1] apply(split invoke_splits, rule conjI)+ apply(auto simp add: is_element_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def set_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 split: option.splits)[1] apply(auto simp add: is_element_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def set_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 split: option.splits)[1] @@ -1172,7 +1178,8 @@ begin definition a_set_disconnected_nodes :: "(_) document_ptr \ (_) node_ptr list \ (_, unit) dom_prog" where - "a_set_disconnected_nodes document_ptr disc_nodes = put_M document_ptr disconnected_nodes_update disc_nodes" + "a_set_disconnected_nodes document_ptr disc_nodes = +put_M document_ptr disconnected_nodes_update disc_nodes" lemmas set_disconnected_nodes_defs = a_set_disconnected_nodes_def definition a_set_disconnected_nodes_locs :: "(_) document_ptr \ (_, unit) dom_prog set" @@ -1196,10 +1203,13 @@ locale l_set_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\ assumes set_disconnected_nodes_locs_impl: "set_disconnected_nodes_locs = a_set_disconnected_nodes_locs" begin lemmas set_disconnected_nodes_def = set_disconnected_nodes_impl[unfolded a_set_disconnected_nodes_def] -lemmas set_disconnected_nodes_locs_def = set_disconnected_nodes_locs_impl[unfolded a_set_disconnected_nodes_locs_def] +lemmas set_disconnected_nodes_locs_def = + set_disconnected_nodes_locs_impl[unfolded a_set_disconnected_nodes_locs_def] lemma set_disconnected_nodes_ok: - "type_wf h \ document_ptr |\| document_ptr_kinds h \ h \ ok (set_disconnected_nodes document_ptr node_ptrs)" - by (simp add: type_wf_impl put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok set_disconnected_nodes_impl[unfolded a_set_disconnected_nodes_def]) + "type_wf h \ document_ptr |\| document_ptr_kinds h \ +h \ ok (set_disconnected_nodes document_ptr node_ptrs)" + by (simp add: type_wf_impl put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok + set_disconnected_nodes_impl[unfolded a_set_disconnected_nodes_def]) lemma set_disconnected_nodes_ptr_in_heap: "h \ ok (set_disconnected_nodes document_ptr disc_nodes) \ document_ptr |\| document_ptr_kinds h" @@ -1234,13 +1244,17 @@ end locale l_set_disconnected_nodes = l_type_wf + l_set_disconnected_nodes_defs + assumes set_disconnected_nodes_writes: - "writes (set_disconnected_nodes_locs document_ptr) (set_disconnected_nodes document_ptr disc_nodes) h h'" + "writes (set_disconnected_nodes_locs document_ptr) +(set_disconnected_nodes document_ptr disc_nodes) h h'" assumes set_disconnected_nodes_ok: - "type_wf h \ document_ptr |\| document_ptr_kinds h \ h \ ok (set_disconnected_nodes document_ptr disc_noded)" + "type_wf h \ document_ptr |\| document_ptr_kinds h \ +h \ ok (set_disconnected_nodes document_ptr disc_noded)" assumes set_disconnected_nodes_ptr_in_heap: - "h \ ok (set_disconnected_nodes document_ptr disc_noded) \ document_ptr |\| document_ptr_kinds h" + "h \ ok (set_disconnected_nodes document_ptr disc_noded) \ +document_ptr |\| document_ptr_kinds h" assumes set_disconnected_nodes_pointers_preserved: - "w \ set_disconnected_nodes_locs document_ptr \ h \ w \\<^sub>h h' \ object_ptr_kinds h = object_ptr_kinds h'" + "w \ set_disconnected_nodes_locs document_ptr \ h \ w \\<^sub>h h' \ +object_ptr_kinds h = object_ptr_kinds h'" assumes set_disconnected_nodes_types_preserved: "w \ set_disconnected_nodes_locs document_ptr \ h \ w \\<^sub>h h' \ type_wf h = type_wf h'" @@ -1257,7 +1271,8 @@ declare l_set_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D lemma set_disconnected_nodes_is_l_set_disconnected_nodes [instances]: "l_set_disconnected_nodes type_wf set_disconnected_nodes set_disconnected_nodes_locs" apply(simp add: l_set_disconnected_nodes_def) - using set_disconnected_nodes_ok set_disconnected_nodes_writes set_disconnected_nodes_pointers_preserved + using set_disconnected_nodes_ok set_disconnected_nodes_writes + set_disconnected_nodes_pointers_preserved set_disconnected_nodes_ptr_in_heap set_disconnected_nodes_typess_preserved by blast+ @@ -1299,7 +1314,8 @@ interpretation i_set_disconnected_nodes_get_disconnected_nodes?: by unfold_locales declare l_set_disconnected_nodes_get_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] -lemma set_disconnected_nodes_get_disconnected_nodes_is_l_set_disconnected_nodes_get_disconnected_nodes [instances]: +lemma set_disconnected_nodes_get_disconnected_nodes_is_l_set_disconnected_nodes_get_disconnected_nodes + [instances]: "l_set_disconnected_nodes_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs" using set_disconnected_nodes_is_l_set_disconnected_nodes get_disconnected_nodes_is_l_get_disconnected_nodes @@ -1346,17 +1362,17 @@ subsubsection \get\_tag\_name\ locale l_get_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin -definition a_get_tag_name :: "(_) element_ptr \ (_, tag_type) dom_prog" +definition a_get_tag_name :: "(_) element_ptr \ (_, tag_name) dom_prog" where - "a_get_tag_name element_ptr = get_M element_ptr tag_type" + "a_get_tag_name element_ptr = get_M element_ptr tag_name" definition a_get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" where - "a_get_tag_name_locs element_ptr \ {preserved (get_M element_ptr tag_type)}" + "a_get_tag_name_locs element_ptr \ {preserved (get_M element_ptr tag_name)}" end locale l_get_tag_name_defs = - fixes get_tag_name :: "(_) element_ptr \ (_, tag_type) dom_prog" + fixes get_tag_name :: "(_) element_ptr \ (_, tag_name) dom_prog" fixes get_tag_name_locs :: "(_) element_ptr \ ((_) heap \ (_) heap \ bool) set" locale l_get_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = @@ -1364,7 +1380,7 @@ locale l_get_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^ l_get_tag_name_defs get_tag_name get_tag_name_locs + l_get_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs for type_wf :: "(_) heap \ bool" - and get_tag_name :: "(_) element_ptr \ (_, tag_type) dom_prog" + 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 = DocumentClass.type_wf" assumes get_tag_name_impl: "get_tag_name = a_get_tag_name" @@ -1466,7 +1482,7 @@ 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: set_child_nodes_locs_def get_tag_name_locs_def all_args_def - intro: element_put_get_preserved[where getter=tag_type and setter=child_nodes_update]) + intro: element_put_get_preserved[where getter=tag_name and setter=child_nodes_update]) end locale l_set_child_nodes_get_tag_name = l_set_child_nodes + l_get_tag_name + @@ -1489,162 +1505,207 @@ lemma set_child_nodes_get_tag_name_is_l_set_child_nodes_get_tag_name [instances] subsubsection \set\_tag\_type\ -locale l_set_tag_type\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs +locale l_set_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs begin -definition a_set_tag_type :: "(_) element_ptr \ tag_type \ (_, unit) dom_prog" +definition a_set_tag_name :: "(_) element_ptr \ tag_name \ (_, unit) dom_prog" where - "a_set_tag_type ptr tag = do { + "a_set_tag_name ptr tag = do { m \ get_M ptr attrs; - put_M ptr tag_type_update tag + put_M ptr tag_name_update tag }" -lemmas set_tag_type_defs = a_set_tag_type_def +lemmas set_tag_name_defs = a_set_tag_name_def -definition a_set_tag_type_locs :: "(_) element_ptr \ (_, unit) dom_prog set" +definition a_set_tag_name_locs :: "(_) element_ptr \ (_, unit) dom_prog set" where - "a_set_tag_type_locs element_ptr \ all_args (put_M element_ptr tag_type_update)" + "a_set_tag_name_locs element_ptr \ all_args (put_M element_ptr tag_name_update)" end -locale l_set_tag_type_defs = - fixes set_tag_type :: "(_) element_ptr \ tag_type \ (_, unit) dom_prog" - fixes set_tag_type_locs :: "(_) element_ptr \ (_, unit) dom_prog set" +locale l_set_tag_name_defs = + fixes set_tag_name :: "(_) element_ptr \ tag_name \ (_, unit) dom_prog" + fixes set_tag_name_locs :: "(_) element_ptr \ (_, unit) dom_prog set" -locale l_set_tag_type\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = +locale l_set_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_type_wf type_wf + - l_set_tag_type_defs set_tag_type set_tag_type_locs + - l_set_tag_type\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs + l_set_tag_name_defs set_tag_name set_tag_name_locs + + l_set_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs for type_wf :: "(_) heap \ bool" - and set_tag_type :: "(_) element_ptr \ char list \ (_, unit) dom_prog" - and set_tag_type_locs :: "(_) element_ptr \ (_, unit) dom_prog set" + + and set_tag_name :: "(_) element_ptr \ char list \ (_, unit) dom_prog" + and set_tag_name_locs :: "(_) element_ptr \ (_, unit) dom_prog set" + assumes type_wf_impl: "type_wf = DocumentClass.type_wf" - assumes set_tag_type_impl: "set_tag_type = a_set_tag_type" - assumes set_tag_type_locs_impl: "set_tag_type_locs = a_set_tag_type_locs" + assumes set_tag_name_impl: "set_tag_name = a_set_tag_name" + assumes set_tag_name_locs_impl: "set_tag_name_locs = a_set_tag_name_locs" begin -lemma set_tag_type_ok: - "type_wf h \ element_ptr |\| element_ptr_kinds h \ h \ ok (set_tag_type element_ptr tag)" +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_type_impl[unfolded a_set_tag_type_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 + 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 by (metis (no_types, lifting) DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t ElementMonad.get_M_pure bind_is_OK_E bind_is_OK_pure_I is_OK_returns_result_I) -lemma set_tag_type_writes: - "writes (set_tag_type_locs element_ptr) (set_tag_type element_ptr tag) h h'" - by(auto simp add: set_tag_type_impl[unfolded a_set_tag_type_def] - set_tag_type_locs_impl[unfolded a_set_tag_type_locs_def] intro: writes_bind_pure) +lemma set_tag_name_writes: + "writes (set_tag_name_locs element_ptr) (set_tag_name element_ptr tag) h h'" + by(auto simp add: set_tag_name_impl[unfolded a_set_tag_name_def] + set_tag_name_locs_impl[unfolded a_set_tag_name_locs_def] intro: writes_bind_pure) -lemma set_tag_type_pointers_preserved: - assumes "w \ set_tag_type_locs element_ptr" +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(1) object_ptr_kinds_preserved[OF writes_singleton2 assms(2)] - by(auto simp add: all_args_def set_tag_type_locs_impl[unfolded a_set_tag_type_locs_def] + by(auto simp add: all_args_def set_tag_name_locs_impl[unfolded a_set_tag_name_locs_def] split: if_splits) -lemma set_tag_type_typess_preserved: - assumes "w \ set_tag_type_locs element_ptr" +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) using assms(1) type_wf_preserved[OF writes_singleton2 assms(2)] - by(auto simp add: all_args_def set_tag_type_locs_impl[unfolded a_set_tag_type_locs_def] + by(auto simp add: all_args_def set_tag_name_locs_impl[unfolded a_set_tag_name_locs_def] split: if_splits) end -locale l_set_tag_type = l_type_wf + l_set_tag_type_defs + - assumes set_tag_type_writes: - "writes (set_tag_type_locs element_ptr) (set_tag_type element_ptr tag) h h'" - assumes set_tag_type_ok: - "type_wf h \ element_ptr |\| element_ptr_kinds h \ h \ ok (set_tag_type element_ptr tag)" - assumes set_tag_type_pointers_preserved: - "w \ set_tag_type_locs element_ptr \ h \ w \\<^sub>h h' \ object_ptr_kinds h = object_ptr_kinds h'" - assumes set_tag_type_types_preserved: - "w \ set_tag_type_locs element_ptr \ h \ w \\<^sub>h h' \ type_wf h = type_wf h'" +locale l_set_tag_name = l_type_wf + l_set_tag_name_defs + + assumes set_tag_name_writes: + "writes (set_tag_name_locs element_ptr) (set_tag_name element_ptr tag) h h'" + assumes set_tag_name_ok: + "type_wf h \ element_ptr |\| element_ptr_kinds h \ h \ ok (set_tag_name element_ptr tag)" + assumes set_tag_name_pointers_preserved: + "w \ set_tag_name_locs element_ptr \ h \ w \\<^sub>h h' \ object_ptr_kinds h = object_ptr_kinds h'" + assumes set_tag_name_types_preserved: + "w \ set_tag_name_locs element_ptr \ h \ w \\<^sub>h h' \ type_wf h = type_wf h'" -global_interpretation l_set_tag_type\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines - set_tag_type = l_set_tag_type\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_set_tag_type and - set_tag_type_locs = l_set_tag_type\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_set_tag_type_locs . +global_interpretation l_set_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs defines + set_tag_name = l_set_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_set_tag_name and + set_tag_name_locs = l_set_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_set_tag_name_locs . interpretation - i_set_tag_type?: l_set_tag_type\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_tag_type set_tag_type_locs + i_set_tag_name?: l_set_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf set_tag_name set_tag_name_locs apply(unfold_locales) - by (auto simp add: set_tag_type_def set_tag_type_locs_def) -declare l_set_tag_type\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] + by (auto simp add: set_tag_name_def set_tag_name_locs_def) +declare l_set_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] -lemma set_tag_type_is_l_set_tag_type [instances]: - "l_set_tag_type type_wf set_tag_type set_tag_type_locs" - apply(simp add: l_set_tag_type_def) - using set_tag_type_ok set_tag_type_writes set_tag_type_pointers_preserved - set_tag_type_typess_preserved +lemma set_tag_name_is_l_set_tag_name [instances]: + "l_set_tag_name type_wf set_tag_name set_tag_name_locs" + apply(simp add: l_set_tag_name_def) + using set_tag_name_ok set_tag_name_writes set_tag_name_pointers_preserved + set_tag_name_typess_preserved by blast paragraph \get\_child\_nodes\ -locale l_set_tag_type_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = - l_set_tag_type\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + +locale l_set_tag_name_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = + l_set_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin -lemma set_tag_type_get_child_nodes: - "\w \ set_tag_type_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" - by(auto simp add: set_tag_type_locs_impl[unfolded a_set_tag_type_locs_def] +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'))" + by(auto simp add: set_tag_name_locs_impl[unfolded a_set_tag_name_locs_def] get_child_nodes_locs_impl[unfolded a_get_child_nodes_locs_def] all_args_def - intro: element_put_get_preserved[where setter=tag_type_update and getter=child_nodes]) + intro: element_put_get_preserved[where setter=tag_name_update and getter=child_nodes]) end -locale l_set_tag_type_get_child_nodes = l_set_tag_type + l_get_child_nodes + - assumes set_tag_type_get_child_nodes: - "\w \ set_tag_type_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_child_nodes_locs ptr'. r h h'))" +locale l_set_tag_name_get_child_nodes = l_set_tag_name + l_get_child_nodes + + assumes 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'))" interpretation - i_set_tag_type_get_child_nodes?: l_set_tag_type_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf - set_tag_type set_tag_type_locs known_ptr + i_set_tag_name_get_child_nodes?: l_set_tag_name_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf + set_tag_name set_tag_name_locs known_ptr get_child_nodes get_child_nodes_locs by unfold_locales -declare l_set_tag_type_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] +declare l_set_tag_name_get_child_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] -lemma set_tag_type_get_child_nodes_is_l_set_tag_type_get_child_nodes [instances]: - "l_set_tag_type_get_child_nodes type_wf set_tag_type set_tag_type_locs known_ptr get_child_nodes +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_type_is_l_set_tag_type get_child_nodes_is_l_get_child_nodes - apply(simp add: l_set_tag_type_get_child_nodes_def l_set_tag_type_get_child_nodes_axioms_def) - using set_tag_type_get_child_nodes + 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 paragraph \get\_disconnected\_nodes\ -locale l_set_tag_type_get_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = - l_set_tag_type\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + +locale l_set_tag_name_get_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = + l_set_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + l_get_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M begin -lemma set_tag_type_get_disconnected_nodes: - "\w \ set_tag_type_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" - by(auto simp add: set_tag_type_locs_impl[unfolded a_set_tag_type_locs_def] +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: set_tag_name_locs_impl[unfolded a_set_tag_name_locs_def] get_disconnected_nodes_locs_impl[unfolded a_get_disconnected_nodes_locs_def] all_args_def) end -locale l_set_tag_type_get_disconnected_nodes = l_set_tag_type + l_get_disconnected_nodes + - assumes set_tag_type_get_disconnected_nodes: - "\w \ set_tag_type_locs ptr. (h \ w \\<^sub>h h' \ (\r \ get_disconnected_nodes_locs ptr'. r h h'))" +locale l_set_tag_name_get_disconnected_nodes = l_set_tag_name + l_get_disconnected_nodes + + assumes 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'))" interpretation - i_set_tag_type_get_disconnected_nodes?: l_set_tag_type_get_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf - set_tag_type set_tag_type_locs get_disconnected_nodes + i_set_tag_name_get_disconnected_nodes?: l_set_tag_name_get_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M type_wf + set_tag_name set_tag_name_locs get_disconnected_nodes get_disconnected_nodes_locs by unfold_locales -declare l_set_tag_type_get_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] +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_type_get_disconnected_nodes_is_l_set_tag_type_get_disconnected_nodes [instances]: - "l_set_tag_type_get_disconnected_nodes type_wf set_tag_type set_tag_type_locs get_disconnected_nodes +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_type_is_l_set_tag_type get_disconnected_nodes_is_l_get_disconnected_nodes - apply(simp add: l_set_tag_type_get_disconnected_nodes_def - l_set_tag_type_get_disconnected_nodes_axioms_def) - using set_tag_type_get_disconnected_nodes + 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 +paragraph \get\_tag\_type\ + +locale l_set_tag_name_get_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = l_get_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M + + l_set_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M +begin +lemma set_tag_name_get_tag_name: + assumes "h \ a_set_tag_name element_ptr tag \\<^sub>h h'" + shows "h' \ a_get_tag_name element_ptr \\<^sub>r tag" + using assms + by(auto simp add: a_get_tag_name_def a_set_tag_name_def) + +lemma set_tag_name_get_tag_name_different_pointers: + assumes "ptr \ ptr'" + assumes "w \ a_set_tag_name_locs ptr" + assumes "h \ w \\<^sub>h h'" + assumes "r \ a_get_tag_name_locs ptr'" + shows "r h h'" + using assms + by(auto simp add: all_args_def a_set_tag_name_locs_def a_get_tag_name_locs_def + split: if_splits option.splits ) +end + +locale l_set_tag_name_get_tag_name = l_get_tag_name + l_set_tag_name + + assumes set_tag_name_get_tag_name: + "h \ set_tag_name element_ptr tag \\<^sub>h h' + \ h' \ get_tag_name element_ptr \\<^sub>r tag" + assumes set_tag_name_get_tag_name_different_pointers: + "ptr \ ptr' \ w \ set_tag_name_locs ptr \ h \ w \\<^sub>h h' + \ r \ get_tag_name_locs ptr' \ r h h'" + +interpretation i_set_tag_name_get_tag_name?: + l_set_tag_name_get_tag_name\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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>C\<^sub>o\<^sub>r\<^sub>e\<^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 \set\_val\ locale l_set_val\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs @@ -1785,7 +1846,8 @@ interpretation declare l_set_val_get_disconnected_nodes\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] lemma set_val_get_disconnected_nodes_is_l_set_val_get_disconnected_nodes [instances]: - "l_set_val_get_disconnected_nodes type_wf set_val set_val_locs 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" using set_val_is_l_set_val get_disconnected_nodes_is_l_get_disconnected_nodes apply(simp add: l_set_val_get_disconnected_nodes_def l_set_val_get_disconnected_nodes_axioms_def) using set_val_get_disconnected_nodes @@ -2419,7 +2481,9 @@ lemma remove_child_child_in_heap: assumes "h \ remove_child ptr' child \\<^sub>h h'" shows "child |\| node_ptr_kinds h" using assms - apply(auto simp add: remove_child_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] split: if_splits)[1] + apply(auto simp add: remove_child_def + elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] + split: if_splits)[1] by (meson is_OK_returns_result_I local.get_owner_document_ptr_in_heap node_ptr_kinds_commutes) @@ -2437,7 +2501,8 @@ proof - using assms(1) apply(auto simp add: remove_child_def elim!: bind_returns_heap_E - dest!: returns_result_eq[OF assms(2)] pure_returns_heap_eq[rotated, OF get_owner_document_pure] + dest!: returns_result_eq[OF assms(2)] + pure_returns_heap_eq[rotated, OF get_owner_document_pure] pure_returns_heap_eq[rotated, OF get_child_nodes_pure] split: if_splits)[1] by (metis get_disconnected_nodes_pure pure_returns_heap_eq) @@ -2460,7 +2525,8 @@ lemma remove_child_writes [simp]: intro!: writes_bind) lemma remove_writes: - "writes (remove_child_locs (the |h \ get_parent child|\<^sub>r) |h \ get_owner_document (cast child)|\<^sub>r) (remove child) h h'" + "writes (remove_child_locs (the |h \ get_parent child|\<^sub>r) |h \ get_owner_document (cast child)|\<^sub>r) +(remove child) h h'" by(auto simp add: remove_def intro!: writes_bind_pure split: option.splits) lemma remove_child_children_subset: @@ -2544,7 +2610,8 @@ end locale l_remove_child = l_type_wf + l_known_ptrs + l_remove_child_defs + l_get_owner_document_defs + l_get_child_nodes_defs + l_get_disconnected_nodes_defs + assumes remove_child_writes: - "writes (remove_child_locs object_ptr |h \ get_owner_document (cast child)|\<^sub>r) (remove_child object_ptr child) h h'" + "writes (remove_child_locs object_ptr |h \ get_owner_document (cast child)|\<^sub>r) +(remove_child object_ptr child) h h'" assumes remove_child_pointers_preserved: "w \ remove_child_locs ptr owner_document \ h \ w \\<^sub>h h' \ object_ptr_kinds h = object_ptr_kinds h'" assumes remove_child_types_preserved: @@ -2742,7 +2809,8 @@ 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" + 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; @@ -2790,7 +2858,8 @@ proof - next case (Some option) then show ?case - using assms(2) \h2 \ get_child_nodes ptr \\<^sub>r children'\ remove_child_children_subset known_ptrs type_wf + using assms(2) \h2 \ get_child_nodes ptr \\<^sub>r children'\ remove_child_children_subset known_ptrs + type_wf by simp qed qed @@ -2824,7 +2893,8 @@ lemma adopt_node_types_preserved: by (auto split: if_splits) end -locale l_adopt_node = l_type_wf + l_known_ptrs + l_get_parent_defs + l_adopt_node_defs + l_get_child_nodes_defs + l_get_owner_document_defs + +locale l_adopt_node = l_type_wf + l_known_ptrs + l_get_parent_defs + l_adopt_node_defs + + l_get_child_nodes_defs + l_get_owner_document_defs + assumes adopt_node_writes: "writes (adopt_node_locs |h \ get_parent node|\<^sub>r |h \ get_owner_document (cast node)|\<^sub>r document_ptr) (adopt_node document_ptr node) h h'" @@ -3020,7 +3090,8 @@ locale l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\< 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_owner_document :: "(_) object_ptr \ ((_) heap, exception, (_) document_ptr) prog" - and insert_before :: "(_) object_ptr \ (_) node_ptr \ (_) node_ptr option \ ((_) 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" @@ -3084,7 +3155,8 @@ lemma insert_before_ptr_in_heap: shows "ptr |\| object_ptr_kinds h" using assms apply(auto simp add: insert_before_def elim!: bind_is_OK_E)[1] - by (metis (mono_tags, lifting) ensure_pre_insertion_validity_pure is_OK_returns_result_I local.get_owner_document_ptr_in_heap next_sibling_pure pure_returns_heap_eq return_returns_heap) + by (metis (mono_tags, lifting) ensure_pre_insertion_validity_pure is_OK_returns_result_I + local.get_owner_document_ptr_in_heap next_sibling_pure pure_returns_heap_eq return_returns_heap) lemma insert_before_child_in_heap: assumes "h \ ok (insert_before ptr node reference_child)" @@ -3203,7 +3275,7 @@ subsubsection \create\_element\ locale l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs = l_get_disconnected_nodes_defs get_disconnected_nodes get_disconnected_nodes_locs + l_set_disconnected_nodes_defs set_disconnected_nodes set_disconnected_nodes_locs + - l_set_tag_type_defs set_tag_type set_tag_type_locs + l_set_tag_name_defs set_tag_name set_tag_name_locs for get_disconnected_nodes :: "(_) document_ptr \ ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes_locs :: @@ -3212,16 +3284,16 @@ locale l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\ "(_) document_ptr \ (_) node_ptr list \ ((_) heap, exception, unit) prog" and set_disconnected_nodes_locs :: "(_) document_ptr \ ((_) heap, exception, unit) prog set" - and set_tag_type :: + and set_tag_name :: "(_) element_ptr \ char list \ ((_) heap, exception, unit) prog" - and set_tag_type_locs :: + and set_tag_name_locs :: "(_) element_ptr \ ((_) heap, exception, unit) prog set" begin -definition a_create_element :: "(_) document_ptr \ tag_type \ (_, (_) element_ptr) dom_prog" +definition a_create_element :: "(_) document_ptr \ tag_name \ (_, (_) element_ptr) dom_prog" where "a_create_element document_ptr tag = do { new_element_ptr \ new_element; - set_tag_type new_element_ptr tag; + set_tag_name new_element_ptr tag; disc_nodes \ get_disconnected_nodes document_ptr; set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes); return new_element_ptr @@ -3229,28 +3301,29 @@ definition a_create_element :: "(_) document_ptr \ tag_type \ tag_type \ (_, (_) element_ptr) dom_prog" + fixes create_element :: "(_) document_ptr \ tag_name \ (_, (_) element_ptr) dom_prog" global_interpretation l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs - set_tag_type set_tag_type_locs + set_tag_name set_tag_name_locs defines create_element = "l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs.a_create_element get_disconnected_nodes - set_disconnected_nodes set_tag_type" + set_disconnected_nodes set_tag_name" . locale l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = - l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_tag_type set_tag_type_locs + + l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_disconnected_nodes get_disconnected_nodes_locs + set_disconnected_nodes set_disconnected_nodes_locs set_tag_name set_tag_name_locs + l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs + - l_set_tag_type type_wf set_tag_type set_tag_type_locs + + l_set_tag_name type_wf set_tag_name set_tag_name_locs + l_create_element_defs create_element + 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_type :: "(_) element_ptr \ char list \ ((_) heap, exception, unit) prog" - and set_tag_type_locs :: "(_) element_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" + @@ -3270,7 +3343,7 @@ 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_type new_element_ptr tag \\<^sub>h h3" 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'" by(auto simp add: create_element_def @@ -3281,8 +3354,9 @@ proof - using new_element_new_ptr h2 new_element_ptr by blast moreover 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_type_writes h3]) - using set_tag_type_pointers_preserved + 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) moreover have "document_ptr |\| document_ptr_kinds h3" by (meson disc_nodes_h3 is_OK_returns_result_I local.get_disconnected_nodes_ptr_in_heap) @@ -3301,14 +3375,17 @@ proof - using new_element_is_element_ptr by blast then show ?thesis - by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs) + by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs + ElementClass.known_ptr_defs) qed end locale l_create_element = l_create_element_defs interpretation - i_create_element?: 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_type set_tag_type_locs type_wf create_element known_ptr + i_create_element?: 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 by(auto simp add: l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def create_element_def instances) declare l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances] @@ -3347,7 +3424,8 @@ global_interpretation l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^ . locale l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = - l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs set_val set_val_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs + + l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs set_val set_val_locs get_disconnected_nodes + get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs + l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs + l_set_val type_wf set_val set_val_locs + l_create_character_data_defs create_character_data + @@ -3388,7 +3466,8 @@ proof - using new_character_data_new_ptr h2 new_character_data_ptr by blast moreover 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]) + 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) moreover have "document_ptr |\| document_ptr_kinds h3" @@ -3408,15 +3487,19 @@ proof - using new_character_data_is_character_data_ptr by blast then show ?thesis - by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs) + by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs + ElementClass.known_ptr_defs) qed end locale l_create_character_data = l_create_character_data_defs interpretation - i_create_character_data?: 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 - by(auto simp add: l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def create_character_data_def instances) + i_create_character_data?: 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 + by(auto simp add: l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def + create_character_data_def instances) declare l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances] @@ -3612,11 +3695,11 @@ definition a_get_elements_by_class_name :: "(_) object_ptr \ attr_va definition a_get_elements_by_tag_name :: "(_) object_ptr \ attr_value \ (_, (_) element_ptr list) dom_prog" where - "a_get_elements_by_tag_name ptr tag_name = to_tree_order ptr \ + "a_get_elements_by_tag_name ptr tag = to_tree_order ptr \ map_filter_M (\ptr. (case cast ptr of Some element_ptr \ do { - this_tag_name \ get_M element_ptr tag_type; - (if this_tag_name = tag_name then return (Some element_ptr) else return None) + this_tag_name \ get_M element_ptr tag_name; + (if this_tag_name = tag then return (Some element_ptr) else return None) } | _ \ return None))" end @@ -3631,7 +3714,8 @@ l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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" + 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" . @@ -3642,13 +3726,17 @@ locale l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\ l_to_tree_order to_tree_order + l_get_attribute type_wf get_attribute get_attribute_locs for to_tree_order :: "(_::linorder) object_ptr \ ((_) heap, exception, (_) object_ptr list) prog" - and first_in_tree_order :: "(_) object_ptr \ ((_) object_ptr \ ((_) heap, exception, (_) element_ptr option) prog) + and first_in_tree_order :: + "(_) object_ptr \ ((_) object_ptr \ ((_) heap, exception, (_) element_ptr option) prog) \ ((_) heap, exception, (_) element_ptr option) prog" 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_element_by_id :: "(_) object_ptr \ char list \ ((_) heap, exception, (_) element_ptr option) prog" - and get_elements_by_class_name :: "(_) object_ptr \ char list \ ((_) heap, exception, (_) element_ptr list) prog" - and get_elements_by_tag_name :: "(_) object_ptr \ char list \ ((_) heap, exception, (_) element_ptr list) prog" + and get_element_by_id :: + "(_) object_ptr \ char list \ ((_) heap, exception, (_) element_ptr option) prog" + and get_elements_by_class_name :: + "(_) object_ptr \ char list \ ((_) heap, exception, (_) element_ptr list) prog" + and get_elements_by_tag_name :: + "(_) object_ptr \ char list \ ((_) heap, exception, (_) element_ptr list) prog" and type_wf :: "(_) heap \ bool" + assumes get_element_by_id_impl: "get_element_by_id = a_get_element_by_id" assumes get_elements_by_class_name_impl: "get_elements_by_class_name = a_get_elements_by_class_name" @@ -3696,7 +3784,7 @@ lemma get_elements_by_tag_name_result_in_tree_order: intro!: map_filter_M_pure map_M_pure_I bind_pure_I split: option.splits list.splits if_splits) -lemma get_elements_by_tag_name_pure [simp]: "pure (get_elements_by_tag_name ptr tag_name) h" +lemma get_elements_by_tag_name_pure [simp]: "pure (get_elements_by_tag_name ptr tag) h" by(auto simp add: get_elements_by_tag_name_def intro!: bind_pure_I map_filter_M_pure split: option.splits) @@ -3706,7 +3794,7 @@ locale l_get_element_by = l_get_element_by_defs + l_to_tree_order_defs + assumes get_element_by_id_result_in_tree_order: "h \ get_element_by_id ptr iden \\<^sub>r Some element_ptr \ h \ to_tree_order ptr \\<^sub>r to \ cast element_ptr \ set to" - assumes get_elements_by_tag_name_pure [simp]: "pure (get_elements_by_tag_name ptr tag_name) h" + assumes get_elements_by_tag_name_pure [simp]: "pure (get_elements_by_tag_name ptr tag) h" 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 diff --git a/Core_DOM/common/classes/CharacterDataClass.thy b/Core_DOM/common/classes/CharacterDataClass.thy index f434a93..23b95e9 100644 --- a/Core_DOM/common/classes/CharacterDataClass.thy +++ b/Core_DOM/common/classes/CharacterDataClass.thy @@ -164,7 +164,8 @@ lemma get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub> \ 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 h \ None" using 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 assms apply(simp add: type_wf_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 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_def) - by (metis assms bind.bind_lzero character_data_ptr_kinds_commutes fmember.rep_eq local.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf option.exhaust option.simps(3)) + by (metis assms bind.bind_lzero character_data_ptr_kinds_commutes fmember.rep_eq + local.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf option.exhaust option.simps(3)) end global_interpretation l_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_lemmas type_wf @@ -258,14 +259,16 @@ lemma new\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub> lemma new_character_data_ptr_new: "character_data_ptr.Ref (Suc (fMax (finsert 0 (character_data_ptr.the_ref |`| character_data_ptrs h)))) |\| character_data_ptrs h" - by (metis Suc_n_not_le_n character_data_ptr.sel(1) fMax_ge fimage_finsert finsertI1 finsertI2 set_finsert) + by (metis Suc_n_not_le_n character_data_ptr.sel(1) fMax_ge fimage_finsert finsertI1 + finsertI2 set_finsert) lemma 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: assumes "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 = (new_character_data_ptr, h')" shows "new_character_data_ptr |\| character_data_ptr_kinds h" using assms unfolding 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 - by (metis Pair_inject character_data_ptrs_def fMax_finsert fempty_iff ffmember_filter fimage_is_fempty is_character_data_ptr_ref max_0L new_character_data_ptr_new) + by (metis Pair_inject character_data_ptrs_def fMax_finsert fempty_iff ffmember_filter + fimage_is_fempty is_character_data_ptr_ref max_0L new_character_data_ptr_new) lemma 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: assumes "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 = (new_character_data_ptr, h')" @@ -337,7 +340,9 @@ lemma known_ptrs_preserved: lemma known_ptrs_subset: "object_ptr_kinds h' |\| object_ptr_kinds h \ a_known_ptrs h \ a_known_ptrs h'" by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD) -lemma known_ptrs_new_ptr: "object_ptr_kinds h' = object_ptr_kinds h |\| {|new_ptr|} \ known_ptr new_ptr \ a_known_ptrs h \ a_known_ptrs h'" +lemma known_ptrs_new_ptr: + "object_ptr_kinds h' = object_ptr_kinds h |\| {|new_ptr|} \ known_ptr new_ptr \ +a_known_ptrs h \ a_known_ptrs h'" by(simp add: a_known_ptrs_def) end global_interpretation l_known_ptrs\<^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 known_ptr defines known_ptrs = a_known_ptrs . diff --git a/Core_DOM/common/classes/DocumentClass.thy b/Core_DOM/common/classes/DocumentClass.thy index 094e216..420daec 100644 --- a/Core_DOM/common/classes/DocumentClass.thy +++ b/Core_DOM/common/classes/DocumentClass.thy @@ -136,7 +136,8 @@ lemma get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_type_w shows "document_ptr |\| document_ptr_kinds h \ get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h \ None" using l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_axioms assms apply(simp add: type_wf_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def) - by (metis document_ptr_kinds_commutes fmember.rep_eq is_none_bind is_none_simps(1) is_none_simps(2) local.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf) + by (metis document_ptr_kinds_commutes fmember.rep_eq is_none_bind is_none_simps(1) + is_none_simps(2) local.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf) end global_interpretation l_get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas type_wf by unfold_locales @@ -195,12 +196,14 @@ lemma get_document_ptr_simp2 [simp]: lemma get_document_ptr_simp3 [simp]: "get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr (put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr f h) = get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h" by(auto simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def) -lemma get_document_ptr_simp4 [simp]: "get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr f h) = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h" +lemma get_document_ptr_simp4 [simp]: + "get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr f h) = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h" by(auto simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_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) lemma get_document_ptr_simp5 [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>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr f h) = 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 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 get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def) -lemma get_document_ptr_simp6 [simp]: "get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (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 character_data_ptr f h) = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h" +lemma get_document_ptr_simp6 [simp]: + "get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr (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 character_data_ptr f h) = get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h" by(auto simp add: get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_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) lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t [simp]: @@ -327,7 +330,9 @@ lemma known_ptrs_preserved: lemma known_ptrs_subset: "object_ptr_kinds h' |\| object_ptr_kinds h \ a_known_ptrs h \ a_known_ptrs h'" by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD) -lemma known_ptrs_new_ptr: "object_ptr_kinds h' = object_ptr_kinds h |\| {|new_ptr|} \ known_ptr new_ptr \ a_known_ptrs h \ a_known_ptrs h'" +lemma known_ptrs_new_ptr: + "object_ptr_kinds h' = object_ptr_kinds h |\| {|new_ptr|} \ known_ptr new_ptr \ +a_known_ptrs h \ a_known_ptrs h'" by(simp add: a_known_ptrs_def) end global_interpretation l_known_ptrs\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t known_ptr defines known_ptrs = a_known_ptrs . diff --git a/Core_DOM/common/classes/NodeClass.thy b/Core_DOM/common/classes/NodeClass.thy index fdbbff1..04ec1af 100644 --- a/Core_DOM/common/classes/NodeClass.thy +++ b/Core_DOM/common/classes/NodeClass.thy @@ -181,18 +181,23 @@ definition a_known_ptrs :: "(_) heap \ bool" lemma known_ptrs_known_ptr: "a_known_ptrs h \ ptr |\| object_ptr_kinds h \ known_ptr ptr" apply(simp add: a_known_ptrs_def) using notin_fset by fastforce -lemma known_ptrs_preserved: "object_ptr_kinds h = object_ptr_kinds h' \ a_known_ptrs h = a_known_ptrs h'" +lemma known_ptrs_preserved: + "object_ptr_kinds h = object_ptr_kinds h' \ a_known_ptrs h = a_known_ptrs h'" by(auto simp add: a_known_ptrs_def) -lemma known_ptrs_subset: "object_ptr_kinds h' |\| object_ptr_kinds h \ a_known_ptrs h \ a_known_ptrs h'" +lemma known_ptrs_subset: + "object_ptr_kinds h' |\| object_ptr_kinds h \ a_known_ptrs h \ a_known_ptrs h'" by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD) -lemma known_ptrs_new_ptr: "object_ptr_kinds h' = object_ptr_kinds h |\| {|new_ptr|} \ known_ptr new_ptr \ a_known_ptrs h \ a_known_ptrs h'" +lemma known_ptrs_new_ptr: + "object_ptr_kinds h' = object_ptr_kinds h |\| {|new_ptr|} \ known_ptr new_ptr \ +a_known_ptrs h \ a_known_ptrs h'" by(simp add: a_known_ptrs_def) end global_interpretation l_known_ptrs\<^sub>N\<^sub>o\<^sub>d\<^sub>e known_ptr defines known_ptrs = a_known_ptrs . lemmas known_ptrs_defs = a_known_ptrs_def lemma known_ptrs_is_l_known_ptrs: "l_known_ptrs known_ptr known_ptrs" - using known_ptrs_known_ptr known_ptrs_preserved l_known_ptrs_def known_ptrs_subset known_ptrs_new_ptr + using known_ptrs_known_ptr known_ptrs_preserved l_known_ptrs_def known_ptrs_subset + known_ptrs_new_ptr by blast lemma get_node_ptr_simp1 [simp]: "get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr (put\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr node h) = Some node" diff --git a/Core_DOM/common/classes/ObjectClass.thy b/Core_DOM/common/classes/ObjectClass.thy index 7461138..8886459 100644 --- a/Core_DOM/common/classes/ObjectClass.thy +++ b/Core_DOM/common/classes/ObjectClass.thy @@ -127,9 +127,13 @@ lemmas known_ptr_defs = a_known_ptr_def locale l_known_ptrs = l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \ bool" + fixes known_ptrs :: "(_) heap \ bool" assumes known_ptrs_known_ptr: "known_ptrs h \ ptr |\| object_ptr_kinds h \ known_ptr ptr" - assumes known_ptrs_preserved: "object_ptr_kinds h = object_ptr_kinds h' \ known_ptrs h = known_ptrs h'" - assumes known_ptrs_subset: "object_ptr_kinds h' |\| object_ptr_kinds h \ known_ptrs h \ known_ptrs h'" - assumes known_ptrs_new_ptr: "object_ptr_kinds h' = object_ptr_kinds h |\| {|new_ptr|} \ known_ptr new_ptr \ known_ptrs h \ known_ptrs h'" + assumes known_ptrs_preserved: + "object_ptr_kinds h = object_ptr_kinds h' \ known_ptrs h = known_ptrs h'" + assumes known_ptrs_subset: + "object_ptr_kinds h' |\| object_ptr_kinds h \ known_ptrs h \ known_ptrs h'" + assumes known_ptrs_new_ptr: + "object_ptr_kinds h' = object_ptr_kinds h |\| {|new_ptr|} \ known_ptr new_ptr \ +known_ptrs h \ known_ptrs h'" locale l_known_ptrs\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t = l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \ bool" begin @@ -142,11 +146,15 @@ lemma known_ptrs_known_ptr: apply(simp add: a_known_ptrs_def) using notin_fset by fastforce -lemma known_ptrs_preserved: "object_ptr_kinds h = object_ptr_kinds h' \ a_known_ptrs h = a_known_ptrs h'" +lemma known_ptrs_preserved: + "object_ptr_kinds h = object_ptr_kinds h' \ a_known_ptrs h = a_known_ptrs h'" by(auto simp add: a_known_ptrs_def) -lemma known_ptrs_subset: "object_ptr_kinds h' |\| object_ptr_kinds h \ a_known_ptrs h \ a_known_ptrs h'" +lemma known_ptrs_subset: + "object_ptr_kinds h' |\| object_ptr_kinds h \ a_known_ptrs h \ a_known_ptrs h'" by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD) -lemma known_ptrs_new_ptr: "object_ptr_kinds h' = object_ptr_kinds h |\| {|new_ptr|} \ known_ptr new_ptr \ a_known_ptrs h \ a_known_ptrs h'" +lemma known_ptrs_new_ptr: + "object_ptr_kinds h' = object_ptr_kinds h |\| {|new_ptr|} \ known_ptr new_ptr \ +a_known_ptrs h \ a_known_ptrs h'" by(simp add: a_known_ptrs_def) end global_interpretation l_known_ptrs\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t known_ptr defines known_ptrs = a_known_ptrs . diff --git a/Core_DOM/common/monads/BaseMonad.thy b/Core_DOM/common/monads/BaseMonad.thy index 346c768..c1e2fda 100644 --- a/Core_DOM/common/monads/BaseMonad.thy +++ b/Core_DOM/common/monads/BaseMonad.thy @@ -36,8 +36,8 @@ theory BaseMonad begin subsection\Datatypes\ -datatype exception = NotFoundError | SegmentationFault | HierarchyRequestError | AssertException - | NonTerminationException | InvokeError | TypeError | DebugException nat +datatype exception = NotFoundError | HierarchyRequestError | NotSupportedError | SegmentationFault + | AssertException | NonTerminationException | InvokeError | TypeError lemma finite_set_in [simp]: "x \ fset FS \ x |\| FS" by (meson notin_fset) @@ -149,11 +149,19 @@ proof (unfold comp_def, rule ccpo.admissibleI, clarify) have h1:"\a. Complete_Partial_Order.chain (flat_ord (Inl e)) {y. \f\A. y = f a}" by (rule chain_fun[OF 1]) show "P h h2 r" - using chain_fun[OF 1] flat_lub_in_chain[OF chain_fun[OF 1]] 2 4 unfolding execute_def fun_lub_def - by force + using CollectD Inl_Inr_False prog.sel chain_fun[OF 1] flat_lub_in_chain[OF chain_fun[OF 1]] 2 4 + unfolding execute_def fun_lub_def + proof - + assume a1: "the_prog (Prog (\x. flat_lub (Inl e) {y. \f\A. y = f x})) h = Inr (r, h2)" + assume a2: "\xa\A. \h h2 r. the_prog (Prog xa) h = Inr (r, h2) \ P h h2 r" + have "Inr (r, h2) \ {s. \f. f \ A \ s = f h} \ Inr (r, h2) = Inl e" + using a1 by (metis (lifting) \\aa a. flat_lub (Inl e) {y. \f\A. y = f aa} = a \ a = Inl e \ a \ {y. \f\A. y = f aa}\ prog.sel) + then show ?thesis + using a2 by fastforce + qed qed -lemma execute_admissible2: +lemma execute_admissible2: "ccpo.admissible (fun_lub (flat_lub (Inl (e::'e)))) (fun_ord (flat_ord (Inl e))) ((\a. \(h::'heap) h' h2 h2' (r::'result) r'. h \ a = Inr (r, h2) \ h' \ a = Inr (r', h2') \ P h h' h2 h2' r r') \ Prog)" @@ -171,11 +179,11 @@ proof (unfold comp_def, rule ccpo.admissibleI, clarify) have "h \ ?lub \ {y. \f\A. y = f h}" using flat_lub_in_chain[OF h1] 4 unfolding execute_def fun_lub_def - by auto + by (metis (mono_tags, lifting) Collect_cong Inl_Inr_False prog.sel) moreover have "h' \ ?lub \ {y. \f\A. y = f h'}" using flat_lub_in_chain[OF h1] 5 unfolding execute_def fun_lub_def - by auto + by (metis (no_types, lifting) Collect_cong Inl_Inr_False prog.sel) ultimately obtain f where "f \ A" and "h \ Prog f = Inr (r, h2)" and diff --git a/Core_DOM/common/monads/CharacterDataMonad.thy b/Core_DOM/common/monads/CharacterDataMonad.thy index b86e4c3..0038c48 100644 --- a/Core_DOM/common/monads/CharacterDataMonad.thy +++ b/Core_DOM/common/monads/CharacterDataMonad.thy @@ -308,7 +308,7 @@ lemma type_wf_put_ptr_not_in_heap_E: 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) + split: option.splits if_splits)[1] using assms(2) node_ptr_kinds_commutes by blast lemma type_wf_put_ptr_in_heap_E: @@ -319,7 +319,8 @@ lemma type_wf_put_ptr_in_heap_E: 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) + 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\ @@ -340,8 +341,8 @@ 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_type_type_wf_preserved [simp]: - "h \ put_M element_ptr tag_type_update v \\<^sub>h h' \ type_wf h = type_wf h'" +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 @@ -526,6 +527,8 @@ lemma type_wf_drop: "type_wf h \ type_wf (Heap (fmdrop ptr (the_ 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) + 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/Core_DOM/common/monads/DocumentMonad.thy b/Core_DOM/common/monads/DocumentMonad.thy index 9b9958c..0878c7c 100644 --- a/Core_DOM/common/monads/DocumentMonad.thy +++ b/Core_DOM/common/monads/DocumentMonad.thy @@ -322,7 +322,8 @@ lemma type_wf_put_ptr_in_heap_E: using assms apply(auto simp add: type_wf_defs elim!: CharacterDataMonad.type_wf_put_ptr_in_heap_E split: option.splits if_splits)[1] - by (metis (no_types, lifting) CharacterDataClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf bind.bind_lunit get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def is_document_kind_def notin_fset option.exhaust_sel) + by (metis (no_types, lifting) CharacterDataClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf bind.bind_lunit get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def + is_document_kind_def notin_fset option.exhaust_sel) @@ -347,8 +348,8 @@ lemma new_element_is_l_new_element [instances]: 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_type_type_wf_preserved [simp]: - "h \ put_M element_ptr tag_type_update v \\<^sub>h h' \ type_wf h = type_wf h'" +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 DocumentClass.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 DocumentClass.type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t DocumentClass.type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e DocumentClass.type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t @@ -360,7 +361,9 @@ lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_tag_type_typ apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs CharacterDataClass.type_wf_defs split: option.splits)[1] - apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse option.distinct(1) option.simps(3)) + apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def + bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse + option.distinct(1) option.simps(3)) by (metis fmember.rep_eq) lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_child_nodes_type_wf_preserved [simp]: @@ -376,7 +379,9 @@ lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_child_nodes_ apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs CharacterDataClass.type_wf_defs split: option.splits)[1] - apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse option.distinct(1) option.simps(3)) + apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def + bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse + option.distinct(1) option.simps(3)) by (metis fmember.rep_eq) lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_attrs_type_wf_preserved [simp]: @@ -392,7 +397,9 @@ lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_attrs_type_w apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs CharacterDataClass.type_wf_defs split: option.splits)[1] - apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse option.distinct(1) option.simps(3)) + apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def + bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse + option.distinct(1) option.simps(3)) by (metis fmember.rep_eq) lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_shadow_root_opt_type_wf_preserved [simp]: @@ -408,7 +415,9 @@ lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_shadow_root_ apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs CharacterDataClass.type_wf_defs split: option.splits)[1] - apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse option.distinct(1) option.simps(3)) + apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def + bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse + option.distinct(1) option.simps(3)) by (metis fmember.rep_eq) lemma new_character_data_type_wf_preserved [simp]: @@ -458,7 +467,8 @@ lemma new_document_type_wf_preserved [simp]: "h \ new_document \D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_doct apply(auto simp add: is_document_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs CharacterDataClass.type_wf_defs split: option.splits)[1] - apply(auto simp add: get_M_defs) + apply(auto simp add: get_M_defs)[1] by (metis (mono_tags) error_returns_result finite_set_in option.exhaust_sel option.simps(4)) lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_document_element_type_wf_preserved [simp]: @@ -599,5 +609,6 @@ lemma type_wf_drop: "type_wf h \ type_wf (Heap (fmdrop ptr (the_ apply(auto simp add: type_wf_defs)[1] using type_wf_drop apply blast - by (metis (no_types, lifting) CharacterDataClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf CharacterDataMonad.type_wf_drop document_ptr_kinds_commutes finite_set_in fmlookup_drop get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def heap.sel) + by (metis (no_types, lifting) CharacterDataClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf CharacterDataMonad.type_wf_drop + document_ptr_kinds_commutes finite_set_in fmlookup_drop get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def heap.sel) end diff --git a/Core_DOM/common/monads/ElementMonad.thy b/Core_DOM/common/monads/ElementMonad.thy index 2103139..b3cd318 100644 --- a/Core_DOM/common/monads/ElementMonad.thy +++ b/Core_DOM/common/monads/ElementMonad.thy @@ -233,10 +233,10 @@ lemma new_element_child_nodes: by(auto simp add: get_M_defs new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E) -lemma new_element_tag_type: +lemma new_element_tag_name: assumes "h \ new_element \\<^sub>h h'" assumes "h \ new_element \\<^sub>r new_element_ptr" - shows "h' \ get_M new_element_ptr tag_type \\<^sub>r ''''" + shows "h' \ get_M new_element_ptr tag_name \\<^sub>r ''''" using assms by(auto simp add: get_M_defs new_element_def new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E) @@ -338,8 +338,8 @@ 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_type_type_wf_preserved [simp]: - "h \ put_M element_ptr tag_type_update v \\<^sub>h h' \ type_wf h = type_wf h'" +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: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs Let_def put_M_defs get_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 put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_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 diff --git a/Core_DOM/common/monads/NodeMonad.thy b/Core_DOM/common/monads/NodeMonad.thy index b17a231..317a1e1 100644 --- a/Core_DOM/common/monads/NodeMonad.thy +++ b/Core_DOM/common/monads/NodeMonad.thy @@ -165,7 +165,7 @@ lemma type_wf_put_ptr_in_heap_E: 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) + 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) @@ -192,7 +192,7 @@ lemma type_wf_preserved_small: 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) + 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)) diff --git a/Core_DOM/common/preliminaries/Heap_Error_Monad.thy b/Core_DOM/common/preliminaries/Heap_Error_Monad.thy index 5a4a0b4..acd5cf1 100644 --- a/Core_DOM/common/preliminaries/Heap_Error_Monad.thy +++ b/Core_DOM/common/preliminaries/Heap_Error_Monad.thy @@ -93,7 +93,8 @@ definition where "returns_result_heap h p r h' \ h \ p \\<^sub>r r \ h \ p \\<^sub>h h'" -lemma return_result_heap_code [code]: "returns_result_heap h p r h' \ (case h \ p of Inr (r', h'') \ r = r' \ h' = h'' | Inl _ \ False)" +lemma return_result_heap_code [code]: + "returns_result_heap h p r h' \ (case h \ p of Inr (r', h'') \ r = r' \ h' = h'' | Inl _ \ False)" by(auto simp add: returns_result_heap_def returns_result_def returns_heap_def split: sum.splits) fun select_result_heap ("|(_)|\<^sub>r\<^sub>h") @@ -452,32 +453,12 @@ fun forall_M :: "('y \ ('heap, 'e, 'result) prog) \ 'y l P x; forall_M P xs }" - (* -lemma forall_M_elim: - assumes "h \ forall_M P xs \\<^sub>r True" and "\x h. x \ set xs \ pure (P x) h" - shows "\x \ set xs. h \ P x \\<^sub>r True" - apply(insert assms, induct xs) - apply(simp) - apply(auto elim!: bind_returns_result_E)[1] - by (metis (full_types) pure_returns_heap_eq) *) + lemma pure_forall_M_I: "(\x. x \ set xs \ pure (P x) h) \ pure (forall_M P xs) h" apply(induct xs) by(auto intro!: bind_pure_I) - (* -lemma forall_M_pure_I: - assumes "\x. x \ set xs \ h \ P x \\<^sub>r True" and "\x h. x \ set xs \ pure (P x)h" - shows "h \ forall_M P xs \\<^sub>r True" - apply(insert assms, induct xs) - apply(simp) - by(fastforce) - -lemma forall_M_pure_eq: - assumes "\x. x \ set xs \ h \ P x \\<^sub>r True \ h' \ P x \\<^sub>r True" - and "\x h. x \ set xs \ pure (P x) h" - shows "(h \ forall_M P xs \\<^sub>r True) \ h' \ forall_M P xs \\<^sub>r True" - using assms - by(auto intro!: forall_M_pure_I dest!: forall_M_elim) *) + subsection \Fold\ @@ -506,7 +487,8 @@ lemma filter_M_pure_I [intro]: "(\x. x \ set xs \ pure apply(induct xs) by(auto intro!: bind_pure_I) -lemma filter_M_is_OK_I [intro]: "(\x. x \ set xs \ h \ ok (P x)) \ (\x. x \ set xs \ pure (P x) h) \ h \ ok (filter_M P xs)" +lemma filter_M_is_OK_I [intro]: + "(\x. x \ set xs \ h \ ok (P x)) \ (\x. x \ set xs \ pure (P x) h) \ h \ ok (filter_M P xs)" apply(induct xs) apply(simp) by(auto intro!: bind_is_OK_pure_I) @@ -518,7 +500,8 @@ lemma filter_M_not_more_elements: by(auto elim!: bind_returns_result_E2 split: if_splits intro!: set_ConsD) lemma filter_M_in_result_if_ok: - assumes "h \ filter_M P xs \\<^sub>r ys" and "\h x. x \ set xs \ pure (P x) h" and "x \ set xs" and "h \ P x \\<^sub>r True" + assumes "h \ filter_M P xs \\<^sub>r ys" and "\h x. x \ set xs \ pure (P x) h" and "x \ set xs" and + "h \ P x \\<^sub>r True" shows "x \ set ys" apply(insert assms, induct xs arbitrary: ys) apply(simp) @@ -730,7 +713,8 @@ definition preserved :: "('heap, 'e, 'result) prog \ 'heap \ (\x. h \ f \\<^sub>r x \ h' \ f \\<^sub>r x)" -lemma preserved_code [code]: "preserved f h h' = (((h \ ok f) \ (h' \ ok f) \ |h \ f|\<^sub>r = |h' \ f|\<^sub>r) \ ((\h \ ok f) \ (\h' \ ok f)))" +lemma preserved_code [code]: + "preserved f h h' = (((h \ ok f) \ (h' \ ok f) \ |h \ f|\<^sub>r = |h' \ f|\<^sub>r) \ ((\h \ ok f) \ (\h' \ ok f)))" apply(auto simp add: preserved_def)[1] apply (meson is_OK_returns_result_E is_OK_returns_result_I)+ done @@ -768,13 +752,16 @@ lemma reads_bind_pure: dest: pure_returns_heap_eq elim!: bind_returns_result_E) -lemma reads_insert_writes_set_left: "\P \ S. reflp P \ transp P \ reads {getter} f h h' \ reads (insert getter S) f h h'" +lemma reads_insert_writes_set_left: + "\P \ S. reflp P \ transp P \ reads {getter} f h h' \ reads (insert getter S) f h h'" unfolding reads_def by simp -lemma reads_insert_writes_set_right: "reflp getter \ transp getter \ reads S f h h' \ reads (insert getter S) f h h'" +lemma reads_insert_writes_set_right: + "reflp getter \ transp getter \ reads S f h h' \ reads (insert getter S) f h h'" unfolding reads_def by blast -lemma reads_subset: "reads S f h h' \ \P \ S' - S. reflp P \ transp P \ S \ S' \ reads S' f h h'" +lemma reads_subset: + "reads S f h h' \ \P \ S' - S. reflp P \ transp P \ S \ S' \ reads S' f h h'" by(auto simp add: reads_def) lemma return_reads [simp]: "reads {} (return x) h h'" diff --git a/Core_DOM/common/preliminaries/Testing_Utils.thy b/Core_DOM/common/preliminaries/Testing_Utils.thy index a8811e7..f87a276 100644 --- a/Core_DOM/common/preliminaries/Testing_Utils.thy +++ b/Core_DOM/common/preliminaries/Testing_Utils.thy @@ -75,10 +75,12 @@ val _ = Theory.setup handle Timeout.TIMEOUT _ => NONE; val t2 = Time.now() - start2; in - if length (Seq.list_of result) > 0 then (Output.information ("eval took " ^ (Time.toString t)); File.append (Path.explode "/tmp/isabellebench") (Time.toString t ^ ",")) else (); + if length (Seq.list_of result) > 0 then (Output.information ("eval took " ^ (Time.toString t)); +File.append (Path.explode "/tmp/isabellebench") (Time.toString t ^ ",")) else (); (case result2_opt of SOME result2 => - (if length (Seq.list_of result2) > 0 then (Output.information ("code_simp took " ^ (Time.toString t2)); File.append (Path.explode "/tmp/isabellebench") (Time.toString t2 ^ "\n")) else ()) + (if length (Seq.list_of result2) > 0 then (Output.information ("code_simp took " ^ (Time.toString t2)); +File.append (Path.explode "/tmp/isabellebench") (Time.toString t2 ^ "\n")) else ()) | NONE => (Output.information "code_simp timed out after 600s"; File.append (Path.explode "/tmp/isabellebench") (">600.000\n"))); result end))) diff --git a/Core_DOM/common/tests/Core_DOM_BaseTest.thy b/Core_DOM/common/tests/Core_DOM_BaseTest.thy index 62f5703..4ebeacd 100644 --- a/Core_DOM/common/tests/Core_DOM_BaseTest.thy +++ b/Core_DOM/common/tests/Core_DOM_BaseTest.thy @@ -142,19 +142,23 @@ fun get_element_by_id_with_null :: "((_::linorder) object_ptr option) \ string \ (_, ((_) object_ptr option) list) dom_prog" +fun get_elements_by_class_name_with_null :: +"((_::linorder) object_ptr option) \ string \ (_, ((_) object_ptr option) list) dom_prog" where "get_elements_by_class_name_with_null (Some ptr) class_name = get_elements_by_class_name ptr class_name \ map_M (return \ Some \ 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)" notation get_elements_by_class_name_with_null ("_ . getElementsByClassName'(_')") -fun get_elements_by_tag_name_with_null :: "((_::linorder) object_ptr option) \ string \ (_, ((_) object_ptr option) list) dom_prog" +fun get_elements_by_tag_name_with_null :: +"((_::linorder) object_ptr option) \ string \ (_, ((_) object_ptr option) list) dom_prog" where - "get_elements_by_tag_name_with_null (Some ptr) tag_name = - get_elements_by_tag_name ptr tag_name \ map_M (return \ Some \ 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)" + "get_elements_by_tag_name_with_null (Some ptr) tag = + get_elements_by_tag_name ptr tag \ map_M (return \ Some \ 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)" notation get_elements_by_tag_name_with_null ("_ . getElementsByTagName'(_')") -fun insert_before_with_null :: "((_::linorder) object_ptr option) \ ((_) object_ptr option) \ ((_) object_ptr option) \ (_, ((_) object_ptr option)) dom_prog" +fun insert_before_with_null :: +"((_::linorder) object_ptr option) \ ((_) object_ptr option) \ ((_) object_ptr option) \ +(_, ((_) object_ptr option)) dom_prog" where "insert_before_with_null (Some ptr) (Some child_obj) ref_child_obj_opt = (case cast child_obj of Some child \ do { @@ -165,7 +169,8 @@ fun insert_before_with_null :: "((_::linorder) object_ptr option) \ | None \ error HierarchyRequestError)" notation insert_before_with_null ("_ . insertBefore'(_, _')") -fun append_child_with_null :: "((_::linorder) object_ptr option) \ ((_) object_ptr option) \ (_, unit) dom_prog" +fun append_child_with_null :: "((_::linorder) object_ptr option) \ ((_) object_ptr option) \ +(_, unit) dom_prog" where "append_child_with_null (Some ptr) (Some child_obj) = (case cast child_obj of Some child \ append_child ptr child @@ -180,7 +185,8 @@ fun get_body :: "((_::linorder) object_ptr option) \ (_, ((_) object }" notation get_body ("_ . body") -fun get_document_element_with_null :: "((_::linorder) object_ptr option) \ (_, ((_) object_ptr option)) dom_prog" +fun get_document_element_with_null :: "((_::linorder) object_ptr option) \ +(_, ((_) object_ptr option)) dom_prog" where "get_document_element_with_null (Some ptr) = (case 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 of Some document_ptr \ do { @@ -190,14 +196,16 @@ fun get_document_element_with_null :: "((_::linorder) object_ptr option) \ None)})" notation get_document_element_with_null ("_ . documentElement") -fun get_owner_document_with_null :: "((_::linorder) object_ptr option) \ (_, ((_) object_ptr option)) dom_prog" +fun get_owner_document_with_null :: "((_::linorder) object_ptr option) \ +(_, ((_) object_ptr option)) dom_prog" where "get_owner_document_with_null (Some ptr) = (do { document_ptr \ get_owner_document ptr; return (Some (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))})" notation get_owner_document_with_null ("_ . ownerDocument") -fun remove_with_null :: "((_::linorder) object_ptr option) \ ((_) object_ptr option) \ (_, ((_) object_ptr option)) dom_prog" +fun remove_with_null :: "((_::linorder) object_ptr option) \ ((_) object_ptr option) \ +(_, ((_) object_ptr option)) dom_prog" where "remove_with_null (Some ptr) (Some child) = (case cast child of Some child_node \ do { @@ -208,7 +216,8 @@ fun remove_with_null :: "((_::linorder) object_ptr option) \ ((_) ob | "remove_with_null _ None = error TypeError" notation remove_with_null ("_ . remove'(')") -fun remove_child_with_null :: "((_::linorder) object_ptr option) \ ((_) object_ptr option) \ (_, ((_) object_ptr option)) dom_prog" +fun remove_child_with_null :: "((_::linorder) object_ptr option) \ ((_) object_ptr option) \ +(_, ((_) object_ptr option)) dom_prog" where "remove_child_with_null (Some ptr) (Some child) = (case cast child of Some child_node \ do { @@ -222,7 +231,7 @@ notation remove_child_with_null ("_ . removeChild") fun get_tag_name_with_null :: "((_) object_ptr option) \ (_, attr_value) dom_prog" where "get_tag_name_with_null (Some ptr) = (case cast ptr of - Some element_ptr \ get_M element_ptr tag_type)" + Some element_ptr \ get_M element_ptr tag_name)" notation get_tag_name_with_null ("_ . tagName") abbreviation "remove_attribute_with_null ptr k \ set_attribute_with_null2 ptr k None" @@ -256,7 +265,8 @@ fun first_child_with_null :: "((_) object_ptr option) \ (_, ((_) obj | None \ None)}" notation first_child_with_null ("_ . firstChild") -fun adopt_node_with_null :: "((_::linorder) object_ptr option) \ ((_) object_ptr option) \ (_, ((_) object_ptr option)) dom_prog" +fun adopt_node_with_null :: +"((_::linorder) object_ptr option) \ ((_) object_ptr option) \(_, ((_) object_ptr option)) dom_prog" where "adopt_node_with_null (Some ptr) (Some child) = (case cast ptr of Some document_ptr \ (case cast child of @@ -266,7 +276,8 @@ fun adopt_node_with_null :: "((_::linorder) object_ptr option) \ ((_ notation adopt_node_with_null ("_ . adoptNode'(_')") -definition createTestTree :: "((_::linorder) object_ptr option) \ (_, (string \ (_, ((_) object_ptr option)) dom_prog)) dom_prog" +definition createTestTree :: +"((_::linorder) object_ptr option) \ (_, (string \ (_, ((_) object_ptr option)) dom_prog)) dom_prog" where "createTestTree ref = return (\id. get_element_by_id_with_null ref id)" diff --git a/Core_DOM/document/root.tex b/Core_DOM/document/root.tex index 445b9bb..58b8f5e 100644 --- a/Core_DOM/document/root.tex +++ b/Core_DOM/document/root.tex @@ -10,7 +10,7 @@ \DeclareOldFontCommand{\bf}{\normalfont\bfseries}{\mathbf} \DeclareOldFontCommand{\it}{\normalfont\itshape}{\mathit} %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%% -\usepackage[USenglish]{babel} +\usepackage[english]{babel} \usepackage[numbers, sort&compress]{natbib} \usepackage{isabelle,isabellesym} \usepackage{booktabs} @@ -50,8 +50,6 @@ \begingroup% \def\isacharunderscore{\textunderscore}% \section{#1 (\thy)}% - \def\isacharunderscore{-}% - \expandafter\label{sec:\isabellecontext}% \endgroup% } diff --git a/Core_DOM/standard/Core_DOM_Heap_WF.thy b/Core_DOM/standard/Core_DOM_Heap_WF.thy index 55e2271..69c7e10 100644 --- a/Core_DOM/standard/Core_DOM_Heap_WF.thy +++ b/Core_DOM/standard/Core_DOM_Heap_WF.thy @@ -23,80 +23,98 @@ * 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 +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" + 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" + 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 + ((\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 + \ (\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)))) + 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] + 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 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)" + assume 3: "x |\| fset_of_list (concat (map (\parent. |h \ get_disconnected_nodes parent|\<^sub>r) +(sorted_list_of_fset (document_ptr_kinds h))))" + have "\(\document_ptr. document_ptr |\| document_ptr_kinds h \ +x \ set |h \ get_disconnected_nodes document_ptr|\<^sub>r)" using 1 2 3 by (smt UN_I fset_of_list_elem image_eqI notin_fset set_concat set_map sorted_list_of_fset_simps(1)) then have "(\parent_ptr. parent_ptr |\| object_ptr_kinds h \ x \ set |h \ get_child_nodes parent_ptr|\<^sub>r)" using 1 2 by auto - then obtain parent_ptr where parent_ptr: "parent_ptr |\| object_ptr_kinds h \ x \ set |h \ get_child_nodes parent_ptr|\<^sub>r" + 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)))" + 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))))" + 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 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))))" + 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" + 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 + "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 @@ -115,12 +133,13 @@ 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))" + (\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 + where "a_distinct_lists h = distinct (concat ( - (map (\ptr. |h \ get_child_nodes ptr|\<^sub>r) |h \ object_ptr_kinds_M|\<^sub>r) + (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) ))" @@ -134,31 +153,31 @@ 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 +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 + 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\<^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" + + 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 @@ -190,10 +209,10 @@ lemma parent_child_rel_child_nodes2: 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. + 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. + 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 @@ -204,15 +223,15 @@ 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 + 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). + 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" @@ -233,15 +252,15 @@ lemma distinct_lists_disconnected_nodes: 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) + 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) + 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: @@ -256,8 +275,8 @@ proof (cases "children = []", simp) 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) + 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: @@ -267,7 +286,8 @@ lemma heap_is_wellformed_children_in_heap: 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) + 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" @@ -281,55 +301,58 @@ proof (auto simp add: heap_is_wellformed_def a_distinct_lists_def)[1] 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) + 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))) + 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) + 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" +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" +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: +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 +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 +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) + 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' +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 + 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" @@ -344,117 +367,117 @@ proof - 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) + 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 +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, hide_lams) 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) + by (metis (no_types, hide_lams) 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) +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: +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) + 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 +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 + 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' + 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' + 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 + 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 + 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) + 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 + 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 + 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 +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 + "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 + using heap_is_wellformed_children_in_heap apply blast - using heap_is_wellformed_disc_nodes_in_heap + using heap_is_wellformed_disc_nodes_in_heap apply blast - using heap_is_wellformed_one_parent + using heap_is_wellformed_one_parent apply blast - using heap_is_wellformed_one_disc_parent + using heap_is_wellformed_one_disc_parent apply blast - using heap_is_wellformed_children_disc_nodes_different + using heap_is_wellformed_children_disc_nodes_different apply blast - using heap_is_wellformed_disconnected_nodes_distinct + using heap_is_wellformed_disconnected_nodes_distinct apply blast - using heap_is_wellformed_children_distinct + using heap_is_wellformed_children_distinct apply blast - using heap_is_wellformed_children_disc_nodes + using heap_is_wellformed_children_disc_nodes apply blast - using parent_child_rel_child + using parent_child_rel_child apply (blast) - using parent_child_rel_child + using parent_child_rel_child apply(blast) - using parent_child_rel_finite + using parent_child_rel_finite apply blast - using parent_child_rel_acyclic + using parent_child_rel_acyclic apply blast - using parent_child_rel_node_ptr + using parent_child_rel_node_ptr apply blast - using parent_child_rel_parent_in_heap + using parent_child_rel_parent_in_heap apply blast - using parent_child_rel_child_in_heap + using parent_child_rel_child_in_heap apply blast done @@ -462,21 +485,21 @@ 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 + 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 + 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" + 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" @@ -489,10 +512,10 @@ 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, hide_lams) 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) + using get_child_nodes_ok assms(2) is_OK_returns_result_I + by (metis (no_types, hide_lams) 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" @@ -501,50 +524,50 @@ proof - 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 + 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) + 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) + 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) + 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) + 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) + 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' + 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) + 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 + by fast moreover have "parent_ptrs = [ptr]" apply(rule filter_M_ex1[OF parent_ptrs h1 h5]) - using ptrs assms(2) assms(3) + 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 *) + 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: @@ -555,7 +578,7 @@ lemma parent_child_rel_parent: 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 + 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 - @@ -575,7 +598,7 @@ lemma heap_wellformed_induct2 [consumes 3, case_names not_in_heap empty_children 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 + 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) @@ -584,7 +607,7 @@ proof(insert assms(1), induct rule: heap_wellformed_induct) 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) + 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 = []") @@ -599,21 +622,21 @@ proof(insert assms(1), induct rule: heap_wellformed_induct) qed next case False - then show ?thesis + 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 + 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) + 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) @@ -625,9 +648,9 @@ proof - 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 +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] @@ -635,43 +658,43 @@ declare l_get_parent_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\ 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 + 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 + 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" + 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' + 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' + 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' + 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'" + 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_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" @@ -681,21 +704,21 @@ proof - 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: + 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 + 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 + 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" @@ -716,42 +739,43 @@ proof - 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) + 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))) + 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) + 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 + 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) + 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 +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 + +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 @@ -761,25 +785,25 @@ locale l_get_parent_wf = l_type_wf + l_known_ptrs + l_heap_is_wellformed_defs \ 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 + \ (\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 + \ (\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 + 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 +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+ + by metis+ @@ -794,21 +818,21 @@ 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 + 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 + 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" + 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: @@ -818,33 +842,33 @@ lemma remove_from_disconnected_nodes_removes: 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) + 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 + + + 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' + "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 + 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 + "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] + 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 @@ -853,31 +877,31 @@ 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 + 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 + 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 + 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 + 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" + 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: @@ -888,10 +912,10 @@ proof (insert assms(1), induct rule: heap_wellformed_induct_rev) 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) + 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: @@ -930,13 +954,14 @@ lemma get_root_node_ok: 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 + 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" + 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) + (\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) @@ -946,12 +971,12 @@ proof 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) + 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) + 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 @@ -975,12 +1000,12 @@ proof(insert assms(2), induct arbitrary: ancestors rule: heap_wellformed_induct_ with Some show ?case proof(induct parent_opt) case None - then show ?case + 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 + then show ?case apply(simp add: get_ancestors_def) by(auto elim!: bind_returns_result_E2 split: option.splits) qed @@ -994,21 +1019,21 @@ lemma get_ancestors_subset: 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" + 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) +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 ()" + (* 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) + using step(2) step(3) by(auto simp add: get_ancestors_def elim!: bind_returns_result_E2) show ?case using step(2) step(3) @@ -1018,9 +1043,10 @@ proof (insert assms(1) assms(2) assms(3), induct ptr arbitrary: ancestors 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) + 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 @@ -1037,8 +1063,8 @@ proof (insert assms(1) assms(2) assms(3), induct ptr arbitrary: ancestors 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 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" @@ -1051,8 +1077,8 @@ proof (insert assms(1) assms(2) assms(3), induct ptr arbitrary: 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) + 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 @@ -1069,13 +1095,13 @@ lemma get_ancestors_also_parent: 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) + 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) + 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 @@ -1087,16 +1113,16 @@ lemma get_ancestors_obtains_children: 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" + 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 + 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) + 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 @@ -1115,8 +1141,8 @@ proof - 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) + 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 @@ -1132,8 +1158,8 @@ proof - 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 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" @@ -1145,15 +1171,15 @@ proof - 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 \ 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\] + 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 @@ -1165,8 +1191,8 @@ proof - qed qed qed - then obtain child where child: "h \ get_parent child \\<^sub>r Some ancestor" - and in_ancestors: "cast child \ set ancestors" + 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 @@ -1181,7 +1207,7 @@ lemma get_ancestors_parent_child_rel: 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" + 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" @@ -1192,33 +1218,33 @@ proof (safe) 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) + 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" + 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\ + 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) + 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) + 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 @@ -1226,12 +1252,12 @@ proof (safe) 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 + using known_ptrs type_wf by blast ultimately show ?thesis using get_ancestors_also_parent assms type_wf by blast qed qed - next +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)]) @@ -1251,12 +1277,12 @@ proof (safe) 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 @@ -1288,15 +1314,15 @@ lemma get_ancestors_eq: 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" + 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 + known_ptrs' assms(2) assms(7) type_wf' + by blast then obtain ancestors' where ancestors': "h' \ get_ancestors ptr \\<^sub>r ancestors'" by auto @@ -1309,8 +1335,8 @@ proof - 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" + 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 @@ -1319,7 +1345,7 @@ proof - 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' + 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 @@ -1337,47 +1363,47 @@ proof - 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 + 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 + 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>*\ + 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) + 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 + 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 + 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 + 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>*\ + \(cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>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 @@ -1385,12 +1411,12 @@ proof - 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) + 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 \(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" @@ -1402,76 +1428,76 @@ proof - 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) + 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 + 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 + 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))) + 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)))\ + 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 + 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) + 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" + 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))) + 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)))\ + 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 + 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) + 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" + 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'" + 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 + 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 @@ -1492,21 +1518,21 @@ proof - 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] *)) + 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) + 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) + apply (meson option.simps(3) returns_result_eq) by (metis IntD1 IntD2 option.inject returns_result_eq step.hyps) qed then show ?thesis @@ -1519,7 +1545,7 @@ lemma get_ancestors_remains_not_in_ancestors: 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 + 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'" @@ -1528,7 +1554,7 @@ lemma get_ancestors_remains_not_in_ancestors: and type_wf': "type_wf h'" shows "node \ set ancestors'" proof - - have object_ptr_kinds_M_eq: + 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) @@ -1536,8 +1562,8 @@ proof - by(simp) show ?thesis - proof (insert assms(1) assms(3) assms(4) assms(6), induct ptr arbitrary: ancestors ancestors' - rule: heap_wellformed_induct_rev) + 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 - @@ -1548,11 +1574,11 @@ proof - 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 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) + 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 @@ -1560,13 +1586,13 @@ proof - 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 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) + using step(4) apply(auto elim!: bind_returns_result_E2 split: option.splits)[1] using 1 apply (meson option.distinct(1) returns_result_eq) @@ -1591,8 +1617,8 @@ next 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) + by (metis assms(1) assms(2) assms(3) get_ancestors_obtains_children get_child_nodes_ptr_in_heap + is_OK_returns_result_I) qed @@ -1602,26 +1628,27 @@ lemma get_ancestors_prefix: 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) +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 + 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 @@ -1638,25 +1665,25 @@ lemma get_ancestors_same_root_node: 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) + 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) + 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 + 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) + 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 @@ -1668,7 +1695,7 @@ proof 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 + 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 @@ -1678,8 +1705,8 @@ next 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(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) @@ -1704,9 +1731,9 @@ proof (insert assms(1) assms(4), induct ptr rule: heap_wellformed_induct_rev) 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) + 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 @@ -1716,8 +1743,8 @@ proof (insert assms(1) assms(4), induct ptr rule: heap_wellformed_induct_rev) 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] + 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 @@ -1729,8 +1756,8 @@ lemma get_root_node_not_node_same: 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) + 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: @@ -1746,86 +1773,86 @@ 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) + 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 + +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 + "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 + 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 + "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 + "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) + "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 + "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 +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: + 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 + 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 + 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 + 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 + 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 + 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 + "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] @@ -1841,7 +1868,7 @@ lemma get_ancestors_wf_is_l_get_ancestors_wf [instances]: 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 + "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] @@ -1856,7 +1883,7 @@ lemma get_root_node_wf_is_l_get_root_node_wf [instances]: 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 = +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 + @@ -1880,7 +1907,7 @@ lemma to_tree_order_either_ptr_or_in_children: and "node \ set nodes" and "h \ get_child_nodes ptr \\<^sub>r children" and "node \ ptr" - obtains child child_to where "child \ set children" + 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" @@ -1890,11 +1917,11 @@ proof - 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" + 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" + 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 @@ -1911,7 +1938,7 @@ proof(insert assms(1) assms(4) assms(5), induct ptr arbitrary: to rule: heap_wel 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" + 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 = []") @@ -1920,7 +1947,7 @@ proof(insert assms(1) assms(4) assms(5), induct ptr arbitrary: to rule: heap_wel 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 + then show ?thesis using \parent |\| object_ptr_kinds h\ step.prems(2) by auto next case False @@ -1933,9 +1960,9 @@ proof(insert assms(1) assms(4) assms(5), induct ptr arbitrary: to rule: heap_wel 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)) + 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 @@ -1958,16 +1985,16 @@ 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" + 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" + 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 @@ -1997,7 +2024,7 @@ lemma to_tree_order_subset: 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 + 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) @@ -2006,8 +2033,8 @@ proof - 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' + 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" @@ -2015,7 +2042,7 @@ proof - 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 + 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 @@ -2026,8 +2053,8 @@ proof - 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) + 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 @@ -2041,8 +2068,8 @@ proof - "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 + 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 @@ -2050,8 +2077,8 @@ proof - 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 + 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 @@ -2083,18 +2110,18 @@ proof - 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) + 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) - + 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] + 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 @@ -2107,9 +2134,9 @@ lemma to_tree_order_child: 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) + 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 @@ -2121,22 +2148,22 @@ proof(insert assms(1) assms(4) assms(6) assms(8), induct ptr arbitrary: nodes 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]) + 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(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) + 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 @@ -2146,16 +2173,16 @@ proof(insert assms(1) assms(4) assms(6) assms(8), induct ptr arbitrary: nodes 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 + 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 @@ -2165,8 +2192,8 @@ lemma to_tree_order_node_ptrs: 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) +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 @@ -2178,18 +2205,18 @@ proof(insert assms(1) assms(4) assms(5) assms(6), induct ptr arbitrary: nodes 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] + 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 + 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 @@ -2216,7 +2243,7 @@ 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) + 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 @@ -2228,18 +2255,18 @@ proof - 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]) + 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 + 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 @@ -2252,14 +2279,14 @@ proof - 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) + 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) + 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 @@ -2289,15 +2316,15 @@ proof 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 + 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) + 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 @@ -2316,8 +2343,8 @@ next 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 + 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 @@ -2327,8 +2354,8 @@ next 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) + 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 @@ -2338,60 +2365,60 @@ next 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 +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 +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 + 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 + "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) + 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: + 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 + 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: + 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 + 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 + "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 +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 + using to_tree_order_ok apply blast - using to_tree_order_ptrs_in_heap + using to_tree_order_ptrs_in_heap apply blast using to_tree_order_parent_child_rel apply(blast, blast) @@ -2399,7 +2426,7 @@ lemma to_tree_order_wf_is_l_to_tree_order_wf [instances]: apply blast using to_tree_order_node_ptrs apply blast - using to_tree_order_child + using to_tree_order_child apply blast using to_tree_order_ptr_in_result apply blast @@ -2425,27 +2452,27 @@ lemma to_tree_order_get_root_node: 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 ) + 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 + 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) + 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 + 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 + 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: @@ -2461,25 +2488,25 @@ proof (insert assms(1)(* assms(4) assms(5) *) assms(6), induct ptr' rule: heap_ 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) + 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 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 + 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) @@ -2487,41 +2514,43 @@ proof (insert assms(1)(* assms(4) assms(5) *) assms(6), induct ptr' rule: heap_ 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\ + 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 \cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>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 +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 +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 + 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 + "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 + 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 @@ -2529,7 +2558,7 @@ lemma to_tree_order_wf_get_root_node_wf_is_l_to_tree_order_wf_get_root_node_wf [ 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 @@ -2555,17 +2584,17 @@ proof - 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 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 + 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" @@ -2573,7 +2602,7 @@ proof - 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; @@ -2581,7 +2610,7 @@ proof - }) 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 + 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) @@ -2589,13 +2618,13 @@ proof - by auto - have eq: "\document_ptr. document_ptr |\| document_ptr_kinds h + 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(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) @@ -2608,8 +2637,8 @@ proof - 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] + 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 @@ -2619,22 +2648,22 @@ proof - 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 + 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] + 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_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) @@ -2654,11 +2683,12 @@ proof - 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)" + 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 + 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, hide_lams)) moreover have "node_ptr |\| node_ptr_kinds h" using assms(2) get_parent_ptr_in_heap by blast @@ -2671,12 +2701,12 @@ proof - 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 + 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, hide_lams) ) qed -lemma get_owner_document_owner_document_in_heap: +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" @@ -2699,17 +2729,20 @@ next 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" + 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) + 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] + 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 @@ -2722,11 +2755,16 @@ next 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) + 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) + 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" @@ -2739,42 +2777,51 @@ next (\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)) + 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\ + 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 + 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(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) + 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" + 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) + 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] + 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 @@ -2787,11 +2834,17 @@ next 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) + 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) + 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" @@ -2804,28 +2857,39 @@ next (\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)) + 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) + 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(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) + 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: +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)" @@ -2837,25 +2901,35 @@ proof - 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 + 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) + 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) + 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(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) + 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(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] + 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 @@ -2866,34 +2940,47 @@ lemma get_owner_document_child_same: 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) + 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 + 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 + 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 + 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) + 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 + 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 + 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 + 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] @@ -2903,19 +2990,24 @@ proof - 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) + 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" + 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) + 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) @@ -2933,12 +3025,13 @@ proof - 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] + 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(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) @@ -2952,14 +3045,16 @@ proof - apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1] using \cast child |\| object_ptr_kinds h\ \ptr |\| object_ptr_kinds h\ apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1] - by (smt \cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child |\| object_ptr_kinds h\ cast_document_ptr_not_node_ptr(1) comp_apply invoke_empty invoke_not invoke_returns_result is_OK_returns_result_I node_ptr_casts_commute2 option.sel) + by (smt \cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r child |\| object_ptr_kinds h\ cast_document_ptr_not_node_ptr(1) + comp_apply invoke_empty invoke_not invoke_returns_result is_OK_returns_result_I + node_ptr_casts_commute2 option.sel) qed end -locale l_get_owner_document_wf = l_heap_is_wellformed_defs + l_type_wf + l_known_ptrs - + l_get_disconnected_nodes_defs + l_get_owner_document_defs - + l_get_parent_defs + +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 \ @@ -2968,26 +3063,30 @@ locale l_get_owner_document_wf = l_heap_is_wellformed_defs + l_type_wf + l_known 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 \ + "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 + 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 + 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 + "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] @@ -3022,7 +3121,8 @@ proof - 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) + 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] @@ -3031,14 +3131,15 @@ proof - 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) + 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) + 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) @@ -3046,12 +3147,16 @@ proof - 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] + 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 + 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) + 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: @@ -3074,14 +3179,17 @@ proof - 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) + 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) + 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 @@ -3099,12 +3207,14 @@ proof - 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) + by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs + ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits) have "root = cast owner_document" - using True - by (smt \h \ get_owner_document ptr \\<^sub>r owner_document\ assms(1) assms(2) assms(3) assms(4) document_ptr_casts_commute3 get_root_node_document returns_result_eq) + using True + by (smt \h \ get_owner_document ptr \\<^sub>r owner_document\ assms(1) assms(2) assms(3) assms(4) + document_ptr_casts_commute3 get_root_node_document returns_result_eq) then show ?thesis - apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def) + 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\ @@ -3114,19 +3224,27 @@ proof - 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) + 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) - 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) + 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) + 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)[2] + 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 @@ -3137,12 +3255,15 @@ proof - 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) + 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) - 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) + 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 @@ -3150,42 +3271,58 @@ proof - 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) + 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) + 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) - 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 + 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) + 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) + 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 +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" +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) + "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 @@ -3197,8 +3334,8 @@ 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\<^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: @@ -3208,7 +3345,7 @@ lemma set_attribute_preserves_wellformedness: thm preserves_wellformedness_writes_needed apply(rule preserves_wellformedness_writes_needed[OF assms set_attribute_writes]) using set_attribute_get_child_nodes - apply(fast) + apply(fast) using set_attribute_get_disconnected_nodes apply(fast) by(auto simp add: all_args_def set_attribute_locs_def) end @@ -3234,25 +3371,25 @@ proof - 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 + 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) + 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 + 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) + 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 @@ -3263,41 +3400,41 @@ proof - 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 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 + 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) + 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) + 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 + 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 + 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 + 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 @@ -3305,21 +3442,21 @@ proof - 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 + 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\ + 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 + 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" + 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" @@ -3327,11 +3464,11 @@ proof - 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 + 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) + 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 @@ -3353,22 +3490,22 @@ 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 + 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] + 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)]) + 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 + 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 @@ -3382,42 +3519,45 @@ proof (standard, safe) 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)]) + 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 + using set_disconnected_nodes_get_child_nodes set_child_nodes_get_child_nodes_different_pointers by fast - then have children_eq2: + 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 + 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)]) + 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 + 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] ) + 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 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 + 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" @@ -3445,8 +3585,8 @@ proof (standard, safe) 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 + 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 @@ -3467,22 +3607,22 @@ lemma remove_child_heap_is_wellformed_preserved: 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 + 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] + 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)]) + OF remove_child_writes assms(2)]) unfolding remove_child_locs_def - using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved + 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 @@ -3496,29 +3636,31 @@ proof - 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)]) + 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 + 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" + 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 + 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)]) + 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 + 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] ) + 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'" @@ -3527,13 +3669,14 @@ proof - 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] + 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 + 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" @@ -3576,8 +3719,8 @@ have "type_wf h2" 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 + 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 @@ -3594,43 +3737,50 @@ have "type_wf h2" 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)) + 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] + 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' + 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 + 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) + 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' + 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 + 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) + 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 @@ -3640,9 +3790,9 @@ assume 0: "\node_ptr\fset (node_ptr_kinds h'). (\document_pt 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 + 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 + using ha1 ha2 ha3 apply(simp) using IntI by fastforce moreover have "distinct |h \ object_ptr_kinds_M|\<^sub>r" @@ -3660,28 +3810,28 @@ assume 0: "\node_ptr\fset (node_ptr_kinds h'). (\document_pt 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) + 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" + 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 + 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 + then show ?thesis using children distinct_children children_eq[OF False] using children' distinct_lists_children h0 using select_result_I2 by fastforce @@ -3693,11 +3843,11 @@ assume 0: "\node_ptr\fset (node_ptr_kinds h'). (\document_pt 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) + 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) + 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'" @@ -3746,12 +3896,12 @@ assume 0: "\node_ptr\fset (node_ptr_kinds h'). (\document_pt 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) + 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) + 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 @@ -3769,7 +3919,7 @@ assume 0: "\node_ptr\fset (node_ptr_kinds h'). (\document_pt by(simp) ultimately show ?thesis using 5 unfolding True - by simp + by simp next case False show ?thesis @@ -3791,11 +3941,11 @@ assume 0: "\node_ptr\fset (node_ptr_kinds h'). (\document_pt 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 + 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 + 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' = {}" @@ -3804,10 +3954,11 @@ assume 0: "\node_ptr\fset (node_ptr_kinds h'). (\document_pt 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' + 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 + 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 @@ -3821,9 +3972,11 @@ assume 0: "\node_ptr\fset (node_ptr_kinds h'). (\document_pt 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 + 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 + 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 @@ -3834,10 +3987,12 @@ assume 0: "\node_ptr\fset (node_ptr_kinds h'). (\document_pt 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 + 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 + 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 @@ -3846,27 +4001,27 @@ assume 0: "\node_ptr\fset (node_ptr_kinds h'). (\document_pt 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" + 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) + 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 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) + 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'" @@ -3885,29 +4040,29 @@ assume 1: "xa \ fset (object_ptr_kinds 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) + 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 + 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 + using "2" "4" "5" "6" "7" False \children' = children\ assms(1) child_in_children_h + child_parent_dual children children_h disc_nodes' get_child_nodes_ptr_in_heap + list.set_cases list.simps(3) option.simps(1) returns_result_eq set_ConsD by (metis (no_types, hide_lams) assms(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 + using "2" "4" "5" "6" "7" \children' = children\ disc_nodes disc_nodes' + disconnected_nodes_eq returns_result_eq by metis qed qed @@ -3927,33 +4082,34 @@ lemma remove_heap_is_wellformed_preserved: 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) + 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" + 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 + 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] + 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]) + 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) @@ -3965,8 +4121,8 @@ proof - 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) + 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: @@ -3981,16 +4137,16 @@ proof - 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] + 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]) + 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 + 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] @@ -4011,9 +4167,9 @@ proof - 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]) + 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: @@ -4024,7 +4180,7 @@ lemma remove_for_all_empty_children: using assms proof(induct children arbitrary: h h') case Nil - then show ?case + then show ?case by simp next case (Cons a children) @@ -4033,8 +4189,8 @@ next 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 + 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" @@ -4049,17 +4205,17 @@ next 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) + 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]) + 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" @@ -4071,22 +4227,22 @@ next 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' +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' + 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' + 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' + 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' @@ -4095,27 +4251,27 @@ locale l_remove_child_wf2 = l_type_wf + l_known_ptrs + l_remove_child_defs + l_h "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' + 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 + 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 + 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 +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]: +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) @@ -4125,11 +4281,11 @@ lemma remove_child_wf2_is_l_remove_child_wf2 [instances]: 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 + @@ -4146,7 +4302,7 @@ 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 } + 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; @@ -4155,22 +4311,22 @@ proof - 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]) + 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) + 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: +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" @@ -4178,7 +4334,7 @@ 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" + 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; @@ -4187,10 +4343,10 @@ proof - 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]) + 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 @@ -4206,21 +4362,21 @@ proof - 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] ) + 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'", + 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'", + 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 + using set_disconnected_nodes_pointers_preserved set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) ultimately show ?thesis @@ -4228,7 +4384,7 @@ proof - 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 + @@ -4239,7 +4395,7 @@ locale l_adopt_node_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\ 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: +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" @@ -4251,47 +4407,47 @@ proof - 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 + 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) + 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 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 + 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 + 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_thesis: +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'. + shows "\ptr' children'. h' \ get_child_nodes ptr' \\<^sub>r children' \ node_ptr \ set children'" - using adopt_node_removes_child assms by blast + using adopt_node_removes_child_step assms by blast lemma adopt_node_preserves_wellformedness: assumes "heap_is_wellformed h" @@ -4301,12 +4457,12 @@ lemma adopt_node_preserves_wellformedness: 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 + 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); @@ -4316,18 +4472,18 @@ proof - 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]) + 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]) + 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" + 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 @@ -4336,13 +4492,13 @@ proof - 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) + 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) + 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) + 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 @@ -4354,20 +4510,20 @@ proof - 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 + 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] ) + 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 + 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" + 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) @@ -4379,7 +4535,7 @@ proof - 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: + 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) @@ -4388,11 +4544,11 @@ proof - 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 + 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: + 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" @@ -4405,7 +4561,7 @@ proof - 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: + 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) @@ -4413,25 +4569,25 @@ proof - 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 + 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 + 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: + 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 + 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 + using disc_nodes_old_document_h2 local.heap_is_wellformed_disconnected_nodes_distinct wellformed_h2 by blast @@ -4443,35 +4599,38 @@ proof - 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 + 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 + 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 + 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 + 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 + 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 + 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: + 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': " @@ -4482,11 +4641,11 @@ proof - 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 + 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 + 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) @@ -4495,36 +4654,37 @@ proof - 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 + 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 + 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\ + 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) + 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) + 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 + 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" @@ -4532,8 +4692,8 @@ 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 + 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 @@ -4545,39 +4705,47 @@ proof - 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'" + 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) + apply (simp add: children_eq2_h3 object_ptr_kinds_h3_eq3 subset_code(1)) + by (metis (no_types, lifting) \child \ set disc_nodes_old_document_h2\ + disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2 disc_nodes_old_document_h2 + disconnected_nodes_eq2_h3 document_ptr_kinds_eq3_h3 finite_set_in + local.heap_is_wellformed_disc_nodes_in_heap node_ptr_kinds_eq3_h2 node_ptr_kinds_eq3_h3 + select_result_I2 set_ConsD subset_code(1) wellformed_h2) moreover have "a_owner_document_valid h2" using wellformed_h2 by (simp add: heap_is_wellformed_def) then have "a_owner_document_valid h'" - apply(simp add: a_owner_document_valid_def node_ptr_kinds_eq_h2 node_ptr_kinds_eq3_h3 - object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 document_ptr_kinds_eq2_h2 - document_ptr_kinds_eq2_h3 children_eq2_h2 children_eq2_h3 ) - by (smt disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2 - disc_nodes_old_document_h2 disc_nodes_old_document_h3 - disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_in_heap - document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 in_set_remove1 - list.set_intros(1) list.set_intros(2) node_ptr_kinds_eq3_h2 - node_ptr_kinds_eq3_h3 object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 - select_result_I2) + apply(simp add: a_owner_document_valid_def node_ptr_kinds_eq_h2 node_ptr_kinds_eq3_h3 + object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 document_ptr_kinds_eq2_h2 + document_ptr_kinds_eq2_h3 children_eq2_h2 children_eq2_h3 ) + by (smt disc_nodes_document_ptr_h' disc_nodes_document_ptr_h2 + disc_nodes_old_document_h2 disc_nodes_old_document_h3 + disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_in_heap + document_ptr_kinds_eq3_h2 document_ptr_kinds_eq3_h3 in_set_remove1 + list.set_intros(1) list.set_intros(2) node_ptr_kinds_eq3_h2 + node_ptr_kinds_eq3_h3 object_ptr_kinds_h2_eq3 object_ptr_kinds_h3_eq3 + select_result_I2) have a_distinct_lists_h2: "a_distinct_lists h2" using wellformed_h2 by (simp add: heap_is_wellformed_def) then have "a_distinct_lists h'" - apply(auto simp add: a_distinct_lists_def object_ptr_kinds_eq_h3 object_ptr_kinds_eq_h2 - children_eq2_h2 children_eq2_h3)[1] + 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) + 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) + 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) + 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')))" @@ -4586,32 +4754,32 @@ proof - 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 + 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 + 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 + 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 + then show ?thesis using disc_nodes_document_ptr_h3 disc_nodes_document_ptr_h' 4 - \child \ set disc_nodes_document_ptr_h3\ + \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, hide_lams) \distinct disc_nodes_old_document_h2\ - disc_nodes_old_document_h3 disconnected_nodes_eq2_h3 - distinct_remove1 docs_neq select_result_I2) + by (metis (no_types, hide_lams) \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 @@ -4620,7 +4788,7 @@ proof - 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 + 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") @@ -4633,21 +4801,21 @@ proof - proof(cases "document_ptr = y") case True then show ?thesis - using 5 True select_result_I2[OF disc_nodes_document_ptr_h'] + 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) + \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'] + 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_h2] select_result_I2[OF disc_nodes_old_document_h3] - disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 \old_document = x\ + 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 @@ -4660,49 +4828,49 @@ proof - 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) + 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) + 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) + \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 = {}\ + 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'] + 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] + 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 + \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) @@ -4712,33 +4880,33 @@ proof - 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] + 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) + 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\ + 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 + \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\ + 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\ + \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) + by (metis \set |h2 \ get_disconnected_nodes y|\<^sub>r \ set disc_nodes_old_document_h2 = {}\ + empty_iff inf.idem) qed qed qed @@ -4746,21 +4914,21 @@ proof - qed next fix x xa xb - assume 0: "distinct (concat (map (\ptr. |h' \ get_child_nodes ptr|\<^sub>r) + 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) + 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) + 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" + 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 + 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 - @@ -4769,19 +4937,19 @@ proof - 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) + 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) + 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 + by fastforce then show ?thesis using \x \ set disc_nodes_old_document_h2\ a1 a4 f7 by blast next @@ -4794,11 +4962,11 @@ proof - 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 + 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 + 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) + 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 @@ -4810,12 +4978,12 @@ proof - 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) + 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) + 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 @@ -4838,7 +5006,7 @@ 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" + 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; @@ -4849,9 +5017,9 @@ proof - 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]) + 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") @@ -4874,9 +5042,9 @@ proof - 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 + 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' + 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" @@ -4888,52 +5056,52 @@ proof - 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 +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 +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 + +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 + "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 + "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 + "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 + 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 + "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 + "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]: +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 @@ -4971,52 +5139,52 @@ proof - 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) + 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]) + 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]) + 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 +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 +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]: +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 @@ -5035,40 +5203,40 @@ locale l_insert_before_wf2\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub l_get_owner_document_wf begin -lemma insert_before_preserves_acyclitity_thesis: +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')" + 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 + 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 + 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 + 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) + 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) + 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] @@ -5076,18 +5244,18 @@ proof - 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 + 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 + 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]) + 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 + 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 ) @@ -5103,10 +5271,10 @@ proof - 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]) + 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 + 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) @@ -5119,17 +5287,17 @@ proof - 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']) + 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 + using set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) - then have object_ptr_kinds_M_eq_h3: + 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: + 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" @@ -5140,68 +5308,68 @@ proof - 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 + 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 + 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" + 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 + 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: + 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" + 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" + 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' + 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" + 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]) + 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: + 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" + 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" @@ -5209,26 +5377,26 @@ proof - 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 + 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 + 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) + 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 + using disconnected_nodes_eq2_h2[OF False] \node \ set disconnected_nodes_h2\ + disconnected_nodes_h2 by fastforce qed qed - have "cast node \ ptr" + have "cast node \ ptr" using ancestors node_not_in_ancestors get_ancestors_ptr by fast @@ -5238,7 +5406,7 @@ proof - 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 + 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]) @@ -5249,24 +5417,25 @@ proof - 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 "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 + 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(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')" @@ -5284,26 +5453,26 @@ proof - 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 + "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) + 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) + 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] @@ -5311,18 +5480,18 @@ proof - 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 + 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 + 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]) + 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 + 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 ) @@ -5338,10 +5507,10 @@ proof - 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]) + 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 + 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) @@ -5354,17 +5523,17 @@ proof - 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']) + 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 + using set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) - then have object_ptr_kinds_M_eq_h3: + 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: + 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" @@ -5375,68 +5544,68 @@ proof - 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 + 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 + 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" + 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 + 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: + 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" + 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" + 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' + 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" + 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]) + 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: + 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" + 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" @@ -5444,26 +5613,26 @@ proof - 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 + 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 + 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) + 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 + using disconnected_nodes_eq2_h2[OF False] \node \ set disconnected_nodes_h2\ + disconnected_nodes_h2 by fastforce qed qed - have "cast node \ ptr" + have "cast node \ ptr" using ancestors node_not_in_ancestors get_ancestors_ptr by fast @@ -5473,7 +5642,7 @@ proof - 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 + 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]) @@ -5497,14 +5666,14 @@ proof - 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) + 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) @@ -5512,49 +5681,57 @@ proof - 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] + 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) + 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\ + 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) + 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, hide_lams) 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) + by (metis (no_types, hide_lams) 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] + 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) + 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] + 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) + 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 + 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 + next fix x y xa - assume 1: "distinct (concat (map (\document_ptr. |h2 \ get_disconnected_nodes document_ptr|\<^sub>r) + 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" @@ -5566,8 +5743,8 @@ proof - case True then have "y \ owner_document" using 4 by simp - show ?thesis - using distinct_concat_map_E(1)[OF 1] + 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, hide_lams) disconnected_nodes_eq2_h2 disjoint_iff_not_equal notin_set_remove1) @@ -5576,18 +5753,18 @@ proof - 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, hide_lams) disconnected_nodes_eq2_h2 disjoint_iff_not_equal notin_set_remove1) + 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, hide_lams) disconnected_nodes_eq2_h2 disjoint_iff_not_equal notin_set_remove1) next case False - then show ?thesis + 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 + 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 @@ -5601,15 +5778,15 @@ proof - 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) + 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 + 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 @@ -5619,29 +5796,29 @@ proof - 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] + 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'" + 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]) + 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 + 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) + 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'" @@ -5658,23 +5835,23 @@ proof - 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, hide_lams) "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) + by (metis (no_types, hide_lams) "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 + 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, hide_lams) "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 + 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 @@ -5682,7 +5859,7 @@ proof - qed next fix x xa xb - assume 1: " (\x\fset (object_ptr_kinds h'). set |h3 \ get_child_nodes x|\<^sub>r) + 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" @@ -5698,7 +5875,7 @@ proof - 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 + 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 @@ -5706,17 +5883,17 @@ proof - 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) + 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 + next case False then show ?thesis - using 1 2 3 4 5 children_eq2_h3[OF False] by fastforce + using 1 2 3 4 5 children_eq2_h3[OF False] by fastforce qed qed @@ -5724,33 +5901,36 @@ proof - 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] + object_ptr_kinds_M_eq2_h3 node_ptr_kinds_eq2_h2 node_ptr_kinds_eq2_h3 + document_ptr_kinds_eq2_h2 document_ptr_kinds_eq2_h3 children_eq2_h2)[1] + apply(auto simp add: document_ptr_kinds_eq2_h2[simplified] document_ptr_kinds_eq2_h3[simplified] + object_ptr_kinds_M_eq2_h2[simplified] object_ptr_kinds_M_eq2_h3[simplified] + node_ptr_kinds_eq2_h2[simplified] node_ptr_kinds_eq2_h3[simplified])[1] apply(auto simp add: disconnected_nodes_eq2_h3[symmetric])[1] - by (smt children_eq2_h3 children_h' children_h3 disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 finite_set_in in_set_remove1 insert_before_list_in_set object_ptr_kinds_M_eq3_h' ptr_in_heap select_result_I2) + by (smt children_eq2_h3 children_h' children_h3 disconnected_nodes_eq2_h2 disconnected_nodes_h2 + disconnected_nodes_h3 finite_set_in in_set_remove1 insert_before_list_in_set object_ptr_kinds_M_eq3_h' + ptr_in_heap select_result_I2) ultimately show "heap_is_wellformed h'" by (simp add: heap_is_wellformed_def) qed -lemma adopt_node_children_remain_distinct_thesis: +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'. + 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 + 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_thesis: +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'. + shows "\children'. h' \ get_child_nodes ptr \\<^sub>r children' \ distinct children'" proof - fix children' @@ -5758,7 +5938,8 @@ proof - 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) + 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 @@ -5768,34 +5949,34 @@ proof - using assms(5) assms(6) insert_before_list_distinct returns_result_eq by fastforce qed -lemma insert_before_children_remain_distinct_thesis: +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'. + 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 + "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) + 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_thesis + 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 @@ -5810,11 +5991,12 @@ proof - 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) + 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_thesis: +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'" @@ -5826,26 +6008,26 @@ proof - 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 + "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) + 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) + 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] @@ -5853,18 +6035,18 @@ proof - 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 + 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 + 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]) + 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 + 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 ) @@ -5880,12 +6062,13 @@ proof - 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]) + 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 + 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" + 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 @@ -5896,17 +6079,17 @@ proof - 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']) + 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 + using set_child_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) - then have object_ptr_kinds_M_eq_h3: + 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: + 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" @@ -5917,61 +6100,62 @@ proof - 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" + 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 + 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" + 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 + 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: + 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" + 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" + 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' + 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" + 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]) + 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: + 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'" @@ -5998,8 +6182,9 @@ proof - } \\<^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 - + 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 { @@ -6015,7 +6200,8 @@ proof - 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 assms(5) assms(7) bind_pure_returns_result_I2 calculation(1) is_OK_returns_result_I + local.get_child_nodes_pure local.get_parent_child_dual returns_result_eq) moreover have "h \ do { (if is_character_data_ptr node \ is_document_ptr parent @@ -6029,51 +6215,51 @@ proof - 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 + using assms(1) assms(2) assms(3) assms(7) local.get_ancestors_ok local.get_parent_parent_in_heap apply blast apply auto[1] apply auto[1] using assms(6) - apply auto[1] - using assms(1) assms(2) assms(3) assms(7) local.get_ancestors_ok local.get_parent_parent_in_heap - apply auto[1] - apply (smt bind_returns_heap_E is_OK_returns_heap_E local.get_parent_pure pure_def - pure_returns_heap_eq return_returns_heap returns_result_eq) + apply auto[1] + using assms(1) assms(2) assms(3) assms(7) local.get_ancestors_ok local.get_parent_parent_in_heap + apply auto[1] + apply (smt bind_returns_heap_E is_OK_returns_heap_E local.get_parent_pure pure_def + pure_returns_heap_eq return_returns_heap returns_result_eq) apply(blast) - using local.get_child_nodes_pure - apply blast - apply (meson assms(7) is_OK_returns_result_I local.get_parent_child_dual) + using local.get_child_nodes_pure + apply blast + apply (meson assms(7) is_OK_returns_result_I local.get_parent_child_dual) apply (simp) - apply (smt assms(5) assms(8) is_OK_returns_result_I returns_result_eq) - by(auto) + apply (smt assms(5) assms(8) is_OK_returns_result_I returns_result_eq) + by(auto) qed end -locale l_insert_before_wf2 = l_type_wf + l_known_ptrs + l_insert_before_defs - + l_heap_is_wellformed_defs + l_get_child_nodes_defs + l_remove_defs + - assumes insert_before_preserves_type_wf: - "heap_is_wellformed h \ type_wf h \ known_ptrs h \ h \ insert_before ptr child ref \\<^sub>h h' +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' + 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 +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]: +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) @@ -6096,9 +6282,9 @@ 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) - using get_child_nodes_ok local.get_parent_parent_in_heap local.known_ptrs_known_ptr by blast + 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: @@ -6123,45 +6309,49 @@ proof - then show False using assms returns_result_eq by fastforce - qed + 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" + 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) - apply (meson invoke_empty is_OK_returns_result_I) - apply (meson invoke_empty is_OK_returns_result_I) + 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) + 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) + 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) + 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 + using \ptr |\| object_ptr_kinds h\ by blast - have "type_wf h2" + 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]) + 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) @@ -6169,25 +6359,33 @@ proof - 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 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) + 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) + 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) + 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) - 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) - using \ptr |\| object_ptr_kinds h\ \is_document_ptr_kind ptr\ \known_ptr ptr\ \object_ptr_kinds h = object_ptr_kinds h2\ \type_wf h2\ assms(4) local.set_child_nodes_document1_ok apply blast - using \ptr |\| object_ptr_kinds h\ \is_document_ptr_kind ptr\ \known_ptr ptr\ \object_ptr_kinds h = object_ptr_kinds h2\ \type_wf h2\ assms(4) local.set_child_nodes_document1_ok apply blast - using \ptr |\| object_ptr_kinds h\ \is_document_ptr_kind ptr\ \known_ptr ptr\ \object_ptr_kinds h = object_ptr_kinds h2\ \type_wf h2\ assms(4) is_element_ptr_kind_cast local.set_child_nodes_document2_ok by blast + 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 @@ -6197,13 +6395,16 @@ proof - 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] + 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] -) + 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 @@ -6214,15 +6415,16 @@ lemma adopt_node_ok: 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) + 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) + 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 @@ -6232,35 +6434,37 @@ proof - 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" + 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]) + 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 + 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) + by(auto split: option.splits) have "type_wf h2" using h2 remove_child_preserves_type_wf assms - by(auto split: option.splits) + by(auto split: option.splits) have "known_ptrs h2" using h2 remove_child_preserves_known_ptrs assms - by(auto split: option.splits) + 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]) + 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" @@ -6294,12 +6498,12 @@ proof - 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 + 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" + 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) @@ -6311,7 +6515,7 @@ proof - 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: + 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) @@ -6322,7 +6526,7 @@ proof - 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 + 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 @@ -6345,7 +6549,7 @@ proof - 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] ) + 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 @@ -6368,7 +6572,7 @@ proof - 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]) + 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 @@ -6381,37 +6585,47 @@ lemma insert_node_ok: 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) +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 + 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) + 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) + 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) + 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) + 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) + 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 @@ -6429,7 +6643,7 @@ 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 + 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 @@ -6445,19 +6659,21 @@ proof - 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 + using assms(1) assms(2) assms(3) local.get_owner_document_owner_document_in_heap owner_document + by blast - obtain h2 where + 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) + 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]) + 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 + 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) @@ -6470,23 +6686,27 @@ proof - 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) - + 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) + 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 + 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]) + 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 + using set_disconnected_nodes_pointers_preserved by (auto simp add: reflp_def transp_def) have "parent |\| object_ptr_kinds h3" @@ -6496,31 +6716,38 @@ proof - 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) + 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' + 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_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!: 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 + 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] @@ -6542,7 +6769,8 @@ lemma append_child_heap_is_wellformed_preserved: 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) + 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" @@ -6561,21 +6789,21 @@ proof - 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) + 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) + 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) @@ -6602,14 +6830,14 @@ proof - 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 + 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 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 @@ -6623,10 +6851,10 @@ lemma append_child_for_all_on_children: shows "h' \ get_child_nodes ptr \\<^sub>r xs@nodes" using assms apply(induct nodes arbitrary: h xs) - apply(simp) + 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' + 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" @@ -6644,8 +6872,8 @@ proof(auto elim!: bind_returns_heap_E)[1]fix a nodes h xs h'a 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 + 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 @@ -6667,28 +6895,29 @@ lemma append_child_for_all_on_no_children: 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' + 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' + 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 +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" +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+ @@ -6696,38 +6925,38 @@ lemma append_child_wf_is_l_append_child_wf [instances]: "l_append_child_wf type_ 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_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_type_get_disconnected_nodes type_wf set_tag_type set_tag_type_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_type set_tag_type_locs type_wf create_element known_ptr + + 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_type_get_child_nodes type_wf set_tag_type set_tag_type_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_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_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_type :: "(_) element_ptr \ char list \ ((_) heap, exception, unit) prog" - and set_tag_type_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 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" @@ -6739,22 +6968,23 @@ 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_type new_element_ptr tag \\<^sub>h h3" 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) + 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] ) + 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) + 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 + 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" @@ -6776,18 +7006,19 @@ proof - 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_type_writes h3]) - using set_tag_type_pointers_preserved + 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 + 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']) + 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" @@ -6797,7 +7028,8 @@ proof - 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 + 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 @@ -6811,92 +7043,92 @@ proof - 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 + 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 + 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 + 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 + 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 + 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: + 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: + 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_type_writes h3 + using get_child_nodes_reads set_tag_name_writes h3 apply(rule reads_writes_preserved) - by(auto simp add: set_tag_type_get_child_nodes) + 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 + 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_type_writes h3 + using get_disconnected_nodes_reads set_tag_name_writes h3 apply(rule reads_writes_preserved) - by(auto simp add: set_tag_type_get_disconnected_nodes) - then have disconnected_nodes_eq2_h2: + 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_type_writes h3] - using set_tag_type_types_preserved + 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 + using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) - have children_eq_h3: + 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 + 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 + 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 \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\ + 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] @@ -6910,12 +7142,12 @@ proof - 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) + 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" + 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) @@ -6924,9 +7156,9 @@ proof - 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) + 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) @@ -6939,69 +7171,79 @@ proof - 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) + 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) + by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 + local.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2) then have "a_all_ptrs_in_heap h'" - by (smt \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ children_eq2_h3 disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 finite_set_in h' is_OK_returns_result_I l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes local.a_all_ptrs_in_heap_def local.get_child_nodes_ptr_in_heap local.l_set_disconnected_nodes_get_disconnected_nodes_axioms node_ptr_kinds_commutes object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1)) + by (smt \h2 \ get_child_nodes (cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr) \\<^sub>r []\ children_eq2_h3 + disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 + finite_set_in h' is_OK_returns_result_I set_disconnected_nodes_get_disconnected_nodes + local.a_all_ptrs_in_heap_def local.get_child_nodes_ptr_in_heap node_ptr_kinds_commutes + object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1)) have "\p. p |\| object_ptr_kinds h \ cast new_element_ptr \ set |h \ get_child_nodes p|\<^sub>r" - using \heap_is_wellformed h\ \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ - heap_is_wellformed_children_in_heap - by (meson NodeMonad.ptr_kinds_ptr_kinds_M a_all_ptrs_in_heap_def assms(3) assms(4) fset_mp - fset_of_list_elem get_child_nodes_ok known_ptrs_known_ptr returns_result_select_result) + 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\) + 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: + 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\ + 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 + apply(auto simp add: a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1] apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert) apply(case_tac "x=cast new_element_ptr") apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] apply(auto simp add: children_eq2_h[symmetric] insort_split dest: distinct_concat_map_E(2))[1] - apply (metis IntI assms(1) assms(3) assms(4) empty_iff local.get_child_nodes_ok + 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 + 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) + 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] + 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) + 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) + 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) + 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" @@ -7013,37 +7255,39 @@ proof - ultimately show "False" apply(-) apply(cases "x = document_ptr") - apply (smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \local.a_all_ptrs_in_heap h\ - disc_nodes_h3 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 - disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' - l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes - local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms - select_result_I2 set_ConsD subsetD) - by (smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ \local.a_all_ptrs_in_heap h\ - disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 - disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' - l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes - local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms - select_result_I2 set_ConsD subsetD) + apply (smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ + \local.a_all_ptrs_in_heap h\ + disc_nodes_h3 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 + disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' + l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes + local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms + select_result_I2 set_ConsD subsetD) + by (smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ + \local.a_all_ptrs_in_heap h\ + disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 + disjoint_iff_not_equal document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' + l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes + local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms + select_result_I2 set_ConsD subsetD) next fix x xa xb - assume 2: "(\x\fset (object_ptr_kinds h3). set |h' \ get_child_nodes x|\<^sub>r) + 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" + show "False" using disc_nodes_document_ptr_h disconnected_nodes_eq2_h3 apply - apply(cases "xb = document_ptr") - apply (metis (no_types, hide_lams) "3" "4" "6" - \\p. p |\| object_ptr_kinds h3 - \ 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) + apply (metis (no_types, hide_lams) "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" @@ -7057,13 +7301,17 @@ proof - 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 + 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 (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) + by (smt ObjectMonad.ptr_kinds_ptr_kinds_M + \cast\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_element_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h + children_eq2_h2 children_eq2_h3 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 + disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h finite_set_in h' list.set_intros(2) + local.set_disconnected_nodes_get_disconnected_nodes select_result_I2) show "heap_is_wellformed h'" using \a_acyclic_heap h'\ \a_all_ptrs_in_heap h'\ \a_distinct_lists h'\ \a_owner_document_valid h'\ @@ -7071,11 +7319,11 @@ proof - 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_type set_tag_type_locs - set_disconnected_nodes set_disconnected_nodes_locs create_element +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] @@ -7085,46 +7333,46 @@ 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 + 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 + 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 + 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 + 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 + 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 + 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 + 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 + 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 + type_wf get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes + set_disconnected_nodes_locs + l_new_character_data - type_wf + type_wf + l_known_ptrs - known_ptr 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" + 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: @@ -7140,21 +7388,22 @@ proof - 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] ) + 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) - 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) + 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 + 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" @@ -7162,14 +7411,14 @@ proof - - have object_ptr_kinds_eq_h: + 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: + 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: + 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 @@ -7181,19 +7430,19 @@ proof - 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]) + 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 + 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']) + 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" @@ -7204,27 +7453,28 @@ proof - 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 + 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 + 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] + 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 + 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: + 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: + 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: + 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 @@ -7236,19 +7486,19 @@ proof - 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]) + 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 + 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']) + 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" @@ -7259,50 +7509,50 @@ proof - 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 + 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 + 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 + 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 + 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 + 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] + 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: + 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: + 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: + 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 + 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: + 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 @@ -7310,42 +7560,42 @@ proof - 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 + 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 + using set_disconnected_nodes_types_preserved by(auto simp add: reflp_def transp_def) - have children_eq_h3: + 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: + 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 + 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 + 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 + 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\ + 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] @@ -7359,12 +7609,12 @@ proof - 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) + 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" + 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) @@ -7373,8 +7623,8 @@ proof - 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) + 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) @@ -7387,30 +7637,30 @@ proof - 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) + 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) + by (simp add: children_eq2_h2 disconnected_nodes_eq2_h2 document_ptr_kinds_eq_h2 + local.a_all_ptrs_in_heap_def node_ptr_kinds_eq_h2 object_ptr_kinds_eq_h2) then have "a_all_ptrs_in_heap h'" - by (smt character_data_ptr_kinds_commutes children_eq2_h3 disc_nodes_document_ptr_h2 - disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 - finite_set_in h' h2 local.a_all_ptrs_in_heap_def - local.set_disconnected_nodes_get_disconnected_nodes new_character_data_ptr - new_character_data_ptr_in_heap node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h3 - object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1)) + by (smt character_data_ptr_kinds_commutes children_eq2_h3 disc_nodes_document_ptr_h2 + disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h3 + finite_set_in h' h2 local.a_all_ptrs_in_heap_def + local.set_disconnected_nodes_get_disconnected_nodes new_character_data_ptr + new_character_data_ptr_in_heap node_ptr_kinds_eq_h2 node_ptr_kinds_eq_h3 + object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1)) have "\p. p |\| object_ptr_kinds h \ cast new_character_data_ptr \ set |h \ get_child_nodes p|\<^sub>r" using \heap_is_wellformed h\ \cast new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ - heap_is_wellformed_children_in_heap - by (meson NodeMonad.ptr_kinds_ptr_kinds_M a_all_ptrs_in_heap_def assms(3) assms(4) fset_mp - fset_of_list_elem get_child_nodes_ok known_ptrs_known_ptr returns_result_select_result) + 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] @@ -7418,44 +7668,44 @@ proof - 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: + 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\ + 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 + apply(auto simp add: a_distinct_lists_def object_ptr_kinds_eq_h document_ptr_kinds_eq_h disconnected_nodes_eq2_h intro!: distinct_concat_map_I)[1] apply (metis distinct_sorted_list_of_set finite_fset sorted_list_of_set_insert) apply(case_tac "x=cast new_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 + 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 + 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] + 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] + 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) + 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) + 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) @@ -7468,16 +7718,17 @@ proof - moreover have "set |h3 \ get_disconnected_nodes x|\<^sub>r \ set |h3 \ get_disconnected_nodes y|\<^sub>r = {}" using calculation by(auto dest: distinct_concat_map_E(1)) ultimately show "False" - by (smt NodeMonad.ptr_kinds_ptr_kinds_M \cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ - \local.a_all_ptrs_in_heap h\ disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h - disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal - document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' - l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes - local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms - select_result_I2 set_ConsD subsetD) + by (smt NodeMonad.ptr_kinds_ptr_kinds_M + \cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr \ set |h \ node_ptr_kinds_M|\<^sub>r\ + \local.a_all_ptrs_in_heap h\ disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h + disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 disjoint_iff_not_equal + document_ptr_kinds_eq_h document_ptr_kinds_eq_h2 finite_set_in h' + l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes + local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms + select_result_I2 set_ConsD subsetD) next fix x xa xb - assume 2: "(\x\fset (object_ptr_kinds h3). set |h' \ get_child_nodes x|\<^sub>r) + 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" @@ -7487,11 +7738,11 @@ proof - using disc_nodes_document_ptr_h disconnected_nodes_eq2_h3 apply(cases "xb = document_ptr") apply (metis (no_types, hide_lams) "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) + \\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" @@ -7505,20 +7756,19 @@ proof - 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(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) + by (smt ObjectMonad.ptr_kinds_ptr_kinds_M + \cast\<^sub>c\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>_\<^sub>d\<^sub>a\<^sub>t\<^sub>a\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r new_character_data_ptr \ set |h \ object_ptr_kinds_M|\<^sub>r\ children_eq2_h + disconnected_nodes_eq2_h3 document_ptr_kinds_eq_h finite_set_in h' list.set_intros(2) + local.set_disconnected_nodes_get_disconnected_nodes select_result_I2) have "known_ptr (cast new_character_data_ptr)" - using \h \ create_character_data document_ptr text \\<^sub>r new_character_data_ptr\ local.create_character_data_known_ptr by blast + 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 @@ -7536,10 +7786,10 @@ proof - 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 +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] @@ -7549,32 +7799,32 @@ 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 + 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 + 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 + create_document + l_new_document_get_child_nodes - type_wf known_ptr get_child_nodes get_child_nodes_locs - + l_new_document - type_wf + type_wf known_ptr get_child_nodes get_child_nodes_locs + + l_new_document + type_wf + l_known_ptrs - known_ptr 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" + 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: @@ -7587,7 +7837,7 @@ 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) + using assms(2) apply(simp add: create_document_def) using new_document_ok by blast @@ -7619,30 +7869,30 @@ proof - 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 + 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 + 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 + 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 + 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 + 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 []" @@ -7652,7 +7902,7 @@ proof - using \type_wf h\ new_document_types_preserved h' by blast have "acyclic (parent_child_rel h)" - using \heap_is_wellformed 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] @@ -7666,51 +7916,55 @@ proof - 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) + 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" + 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 + 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) + 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'" + 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) + 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\ + 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 []\ + 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(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) apply(auto simp add: dest: distinct_concat_map_E)[1] @@ -7718,9 +7972,9 @@ proof - 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) + 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" @@ -7728,7 +7982,7 @@ proof - 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) + 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" @@ -7746,11 +8000,11 @@ proof - 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) + 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) + 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) + 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" @@ -7760,28 +8014,29 @@ proof - 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) + 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) + 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 +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] diff --git a/Core_DOM/standard/classes/ElementClass.thy b/Core_DOM/standard/classes/ElementClass.thy index 8128245..b4f5222 100644 --- a/Core_DOM/standard/classes/ElementClass.thy +++ b/Core_DOM/standard/classes/ElementClass.thy @@ -39,43 +39,52 @@ text\The type @{type "DOMString"} is a type synonym for @{type "string"}, type_synonym attr_key = DOMString type_synonym attr_value = DOMString type_synonym attrs = "(attr_key, attr_value) fmap" -type_synonym tag_type = DOMString +type_synonym tag_name = DOMString record ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr) RElement = RNode + nothing :: unit - tag_type :: tag_type + tag_name :: tag_name child_nodes :: "('node_ptr, 'element_ptr, 'character_data_ptr) node_ptr list" attrs :: attrs shadow_root_opt :: "'shadow_root_ptr shadow_root_ptr option" type_synonym ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element) Element - = "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) RElement_scheme" + = "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) +RElement_scheme" register_default_tvars "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element) Element" type_synonym ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Node, 'Element) Node - = "(('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) RElement_ext + 'Node) Node" + = "(('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) RElement_ext ++ 'Node) Node" register_default_tvars "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Node, 'Element) Node" type_synonym ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element) Object - = "('Object, ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) RElement_ext + 'Node) Object" + = "('Object, ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) +RElement_ext + 'Node) Object" register_default_tvars "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element) Object" type_synonym - ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element) heap - = "('document_ptr document_ptr + 'shadow_root_ptr shadow_root_ptr + 'object_ptr, 'element_ptr element_ptr + 'character_data_ptr character_data_ptr + 'node_ptr, 'Object, - ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) RElement_ext + 'Node) heap" + ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr, + 'Object, 'Node, 'Element) heap + = "('document_ptr document_ptr + 'shadow_root_ptr shadow_root_ptr + 'object_ptr, +'element_ptr element_ptr + 'character_data_ptr character_data_ptr + 'node_ptr, 'Object, +('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) RElement_ext + +'Node) heap" register_default_tvars - "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element) heap" + "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr, +'Object, 'Node, 'Element) heap" type_synonym heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l = "(unit, unit, unit, unit, unit, unit, unit, unit, unit) heap" definition element_ptr_kinds :: "(_) heap \ (_) element_ptr fset" where - "element_ptr_kinds heap = the |`| (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r |`| (ffilter is_element_ptr_kind (node_ptr_kinds heap)))" + "element_ptr_kinds heap = +the |`| (cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r |`| (ffilter is_element_ptr_kind (node_ptr_kinds heap)))" lemma element_ptr_kinds_simp [simp]: - "element_ptr_kinds (Heap (fmupd (cast element_ptr) element (the_heap h))) = {|element_ptr|} |\| element_ptr_kinds h" + "element_ptr_kinds (Heap (fmupd (cast element_ptr) element (the_heap h))) = +{|element_ptr|} |\| element_ptr_kinds h" apply(auto simp add: element_ptr_kinds_def)[1] by force @@ -85,7 +94,8 @@ definition element_ptrs :: "(_) heap \ (_) element_ptr fset" definition cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) Node \ (_) Element option" where - "cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t node = (case RNode.more node of Inl element \ Some (RNode.extend (RNode.truncate node) element) | _ \ None)" + "cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t node = +(case RNode.more node of Inl element \ Some (RNode.extend (RNode.truncate node) element) | _ \ None)" adhoc_overloading cast cast\<^sub>N\<^sub>o\<^sub>d\<^sub>e\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t abbreviation cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) Object \ (_) Element option" @@ -210,9 +220,9 @@ lemma get_elment_ptr_simp2 [simp]: by(auto simp add: get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def) -abbreviation "create_element_obj tag_type_arg child_nodes_arg attrs_arg shadow_root_opt_arg +abbreviation "create_element_obj tag_name_arg child_nodes_arg attrs_arg shadow_root_opt_arg \ \ RObject.nothing = (), RNode.nothing = (), RElement.nothing = (), - tag_type = tag_type_arg, Element.child_nodes = child_nodes_arg, attrs = attrs_arg, + tag_name = tag_name_arg, Element.child_nodes = child_nodes_arg, attrs = attrs_arg, shadow_root_opt = shadow_root_opt_arg, \ = None \" definition new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_) heap \ ((_) element_ptr \ (_) heap)" @@ -298,11 +308,15 @@ lemma known_ptrs_known_ptr: apply(simp add: a_known_ptrs_def) using notin_fset by fastforce -lemma known_ptrs_preserved: "object_ptr_kinds h = object_ptr_kinds h' \ a_known_ptrs h = a_known_ptrs h'" +lemma known_ptrs_preserved: + "object_ptr_kinds h = object_ptr_kinds h' \ a_known_ptrs h = a_known_ptrs h'" by(auto simp add: a_known_ptrs_def) -lemma known_ptrs_subset: "object_ptr_kinds h' |\| object_ptr_kinds h \ a_known_ptrs h \ a_known_ptrs h'" +lemma known_ptrs_subset: + "object_ptr_kinds h' |\| object_ptr_kinds h \ a_known_ptrs h \ a_known_ptrs h'" by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD) -lemma known_ptrs_new_ptr: "object_ptr_kinds h' = object_ptr_kinds h |\| {|new_ptr|} \ known_ptr new_ptr \ a_known_ptrs h \ a_known_ptrs h'" +lemma known_ptrs_new_ptr: + "object_ptr_kinds h' = object_ptr_kinds h |\| {|new_ptr|} \ known_ptr new_ptr \ +a_known_ptrs h \ a_known_ptrs h'" by(simp add: a_known_ptrs_def) end global_interpretation l_known_ptrs\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t known_ptr defines known_ptrs = a_known_ptrs .