Fixed long lines and simp lemmas without names.

This commit is contained in:
Michael Herzberg 2020-07-22 22:11:21 +01:00
parent 99a6566ed0
commit f955f2fa56
35 changed files with 6466 additions and 5933 deletions

View File

@ -631,10 +631,12 @@ lemma set_child_nodes_element_ok [simp]:
proof - proof -
have "is_element_ptr ptr" have "is_element_ptr ptr"
using \<open>known_ptr ptr\<close> assms(4) using \<open>known_ptr ptr\<close> 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 then show ?thesis
using assms 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) 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 qed
@ -648,10 +650,12 @@ lemma set_child_nodes_document1_ok [simp]:
proof - proof -
have "is_document_ptr ptr" have "is_document_ptr ptr"
using \<open>known_ptr ptr\<close> assms(4) using \<open>known_ptr ptr\<close> 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 then show ?thesis
using assms 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) 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 qed
@ -666,7 +670,8 @@ lemma set_child_nodes_document2_ok [simp]:
proof - proof -
have "is_document_ptr ptr" have "is_document_ptr ptr"
using \<open>known_ptr ptr\<close> assms(4) using \<open>known_ptr ptr\<close> 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 then show ?thesis
using assms 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)[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)[1]
@ -1173,7 +1178,8 @@ begin
definition a_set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> (_, unit) dom_prog" definition a_set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> (_, unit) dom_prog"
where 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 lemmas set_disconnected_nodes_defs = a_set_disconnected_nodes_def
definition a_set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> (_, unit) dom_prog set" definition a_set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> (_, unit) dom_prog set"
@ -1197,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" assumes set_disconnected_nodes_locs_impl: "set_disconnected_nodes_locs = a_set_disconnected_nodes_locs"
begin begin
lemmas set_disconnected_nodes_def = set_disconnected_nodes_impl[unfolded a_set_disconnected_nodes_def] 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: lemma set_disconnected_nodes_ok:
"type_wf h \<Longrightarrow> document_ptr |\<in>| document_ptr_kinds h \<Longrightarrow> h \<turnstile> ok (set_disconnected_nodes document_ptr node_ptrs)" "type_wf h \<Longrightarrow> document_ptr |\<in>| document_ptr_kinds h \<Longrightarrow>
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]) h \<turnstile> 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: lemma set_disconnected_nodes_ptr_in_heap:
"h \<turnstile> ok (set_disconnected_nodes document_ptr disc_nodes) \<Longrightarrow> document_ptr |\<in>| document_ptr_kinds h" "h \<turnstile> ok (set_disconnected_nodes document_ptr disc_nodes) \<Longrightarrow> document_ptr |\<in>| document_ptr_kinds h"
@ -1235,13 +1244,17 @@ end
locale l_set_disconnected_nodes = l_type_wf + l_set_disconnected_nodes_defs + locale l_set_disconnected_nodes = l_type_wf + l_set_disconnected_nodes_defs +
assumes set_disconnected_nodes_writes: 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: assumes set_disconnected_nodes_ok:
"type_wf h \<Longrightarrow> document_ptr |\<in>| document_ptr_kinds h \<Longrightarrow> h \<turnstile> ok (set_disconnected_nodes document_ptr disc_noded)" "type_wf h \<Longrightarrow> document_ptr |\<in>| document_ptr_kinds h \<Longrightarrow>
h \<turnstile> ok (set_disconnected_nodes document_ptr disc_noded)"
assumes set_disconnected_nodes_ptr_in_heap: assumes set_disconnected_nodes_ptr_in_heap:
"h \<turnstile> ok (set_disconnected_nodes document_ptr disc_noded) \<Longrightarrow> document_ptr |\<in>| document_ptr_kinds h" "h \<turnstile> ok (set_disconnected_nodes document_ptr disc_noded) \<Longrightarrow>
document_ptr |\<in>| document_ptr_kinds h"
assumes set_disconnected_nodes_pointers_preserved: assumes set_disconnected_nodes_pointers_preserved:
"w \<in> set_disconnected_nodes_locs document_ptr \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> object_ptr_kinds h = object_ptr_kinds h'" "w \<in> set_disconnected_nodes_locs document_ptr \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow>
object_ptr_kinds h = object_ptr_kinds h'"
assumes set_disconnected_nodes_types_preserved: assumes set_disconnected_nodes_types_preserved:
"w \<in> set_disconnected_nodes_locs document_ptr \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'" "w \<in> set_disconnected_nodes_locs document_ptr \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
@ -1258,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]: lemma set_disconnected_nodes_is_l_set_disconnected_nodes [instances]:
"l_set_disconnected_nodes type_wf set_disconnected_nodes set_disconnected_nodes_locs" "l_set_disconnected_nodes type_wf set_disconnected_nodes set_disconnected_nodes_locs"
apply(simp add: l_set_disconnected_nodes_def) 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 set_disconnected_nodes_ptr_in_heap set_disconnected_nodes_typess_preserved
by blast+ by blast+
@ -1300,7 +1314,8 @@ interpretation i_set_disconnected_nodes_get_disconnected_nodes?:
by unfold_locales 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] 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 "l_set_disconnected_nodes_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs
set_disconnected_nodes set_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 using set_disconnected_nodes_is_l_set_disconnected_nodes get_disconnected_nodes_is_l_get_disconnected_nodes
@ -1833,7 +1848,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] 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]: 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 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) apply(simp add: l_set_val_get_disconnected_nodes_def l_set_val_get_disconnected_nodes_axioms_def)
using set_val_get_disconnected_nodes using set_val_get_disconnected_nodes
@ -2467,7 +2483,9 @@ lemma remove_child_child_in_heap:
assumes "h \<turnstile> remove_child ptr' child \<rightarrow>\<^sub>h h'" assumes "h \<turnstile> remove_child ptr' child \<rightarrow>\<^sub>h h'"
shows "child |\<in>| node_ptr_kinds h" shows "child |\<in>| node_ptr_kinds h"
using assms 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) by (meson is_OK_returns_result_I local.get_owner_document_ptr_in_heap node_ptr_kinds_commutes)
@ -2485,7 +2503,8 @@ proof -
using assms(1) using assms(1)
apply(auto simp add: remove_child_def apply(auto simp add: remove_child_def
elim!: bind_returns_heap_E 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] pure_returns_heap_eq[rotated, OF get_child_nodes_pure]
split: if_splits)[1] split: if_splits)[1]
by (metis get_disconnected_nodes_pure pure_returns_heap_eq) by (metis get_disconnected_nodes_pure pure_returns_heap_eq)
@ -2508,7 +2527,8 @@ lemma remove_child_writes [simp]:
intro!: writes_bind) intro!: writes_bind)
lemma remove_writes: lemma remove_writes:
"writes (remove_child_locs (the |h \<turnstile> get_parent child|\<^sub>r) |h \<turnstile> get_owner_document (cast child)|\<^sub>r) (remove child) h h'" "writes (remove_child_locs (the |h \<turnstile> get_parent child|\<^sub>r) |h \<turnstile> get_owner_document (cast child)|\<^sub>r)
(remove child) h h'"
by(auto simp add: remove_def intro!: writes_bind_pure split: option.splits) by(auto simp add: remove_def intro!: writes_bind_pure split: option.splits)
lemma remove_child_children_subset: lemma remove_child_children_subset:
@ -2592,7 +2612,8 @@ end
locale l_remove_child = l_type_wf + l_known_ptrs + l_remove_child_defs + l_get_owner_document_defs 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 + + l_get_child_nodes_defs + l_get_disconnected_nodes_defs +
assumes remove_child_writes: assumes remove_child_writes:
"writes (remove_child_locs object_ptr |h \<turnstile> get_owner_document (cast child)|\<^sub>r) (remove_child object_ptr child) h h'" "writes (remove_child_locs object_ptr |h \<turnstile> get_owner_document (cast child)|\<^sub>r)
(remove_child object_ptr child) h h'"
assumes remove_child_pointers_preserved: assumes remove_child_pointers_preserved:
"w \<in> remove_child_locs ptr owner_document \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> object_ptr_kinds h = object_ptr_kinds h'" "w \<in> remove_child_locs ptr owner_document \<Longrightarrow> h \<turnstile> w \<rightarrow>\<^sub>h h' \<Longrightarrow> object_ptr_kinds h = object_ptr_kinds h'"
assumes remove_child_types_preserved: assumes remove_child_types_preserved:
@ -2790,7 +2811,8 @@ proof -
obtain old_document parent_opt h2 where obtain old_document parent_opt h2 where
old_document: "h \<turnstile> get_owner_document (cast node) \<rightarrow>\<^sub>r old_document" and old_document: "h \<turnstile> get_owner_document (cast node) \<rightarrow>\<^sub>r old_document" and
parent_opt: "h \<turnstile> get_parent node \<rightarrow>\<^sub>r parent_opt" and parent_opt: "h \<turnstile> get_parent node \<rightarrow>\<^sub>r parent_opt" and
h2: "h \<turnstile> (case parent_opt of Some parent \<Rightarrow> do { remove_child parent node } | None \<Rightarrow> do { return ()}) \<rightarrow>\<^sub>h h2" h2: "h \<turnstile> (case parent_opt of Some parent \<Rightarrow> do { remove_child parent node } |
None \<Rightarrow> do { return ()}) \<rightarrow>\<^sub>h h2"
and and
h': "h2 \<turnstile> (if owner_document \<noteq> old_document then do { h': "h2 \<turnstile> (if owner_document \<noteq> old_document then do {
old_disc_nodes \<leftarrow> get_disconnected_nodes old_document; old_disc_nodes \<leftarrow> get_disconnected_nodes old_document;
@ -2838,7 +2860,8 @@ proof -
next next
case (Some option) case (Some option)
then show ?case then show ?case
using assms(2) \<open>h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children'\<close> remove_child_children_subset known_ptrs type_wf using assms(2) \<open>h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children'\<close> remove_child_children_subset known_ptrs
type_wf
by simp by simp
qed qed
qed qed
@ -2872,7 +2895,8 @@ lemma adopt_node_types_preserved:
by (auto split: if_splits) by (auto split: if_splits)
end 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: assumes adopt_node_writes:
"writes (adopt_node_locs |h \<turnstile> get_parent node|\<^sub>r "writes (adopt_node_locs |h \<turnstile> get_parent node|\<^sub>r
|h \<turnstile> get_owner_document (cast node)|\<^sub>r document_ptr) (adopt_node document_ptr node) h h'" |h \<turnstile> get_owner_document (cast node)|\<^sub>r document_ptr) (adopt_node document_ptr node) h h'"
@ -3068,7 +3092,8 @@ locale l_insert_before\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<
and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog" and get_disconnected_nodes :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, (_) node_ptr list) prog"
and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set" and get_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_owner_document :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) document_ptr) prog" and get_owner_document :: "(_) object_ptr \<Rightarrow> ((_) heap, exception, (_) document_ptr) prog"
and insert_before :: "(_) object_ptr \<Rightarrow> (_) node_ptr \<Rightarrow> (_) node_ptr option \<Rightarrow> ((_) heap, exception, unit) prog" and insert_before ::
"(_) object_ptr \<Rightarrow> (_) node_ptr \<Rightarrow> (_) node_ptr option \<Rightarrow> ((_) heap, exception, unit) prog"
and insert_before_locs :: "(_) object_ptr \<Rightarrow> (_) object_ptr option \<Rightarrow> (_) document_ptr and insert_before_locs :: "(_) object_ptr \<Rightarrow> (_) object_ptr option \<Rightarrow> (_) document_ptr
\<Rightarrow> (_) document_ptr \<Rightarrow> (_, unit) dom_prog set" \<Rightarrow> (_) document_ptr \<Rightarrow> (_, unit) dom_prog set"
and append_child :: "(_) object_ptr \<Rightarrow> (_) node_ptr \<Rightarrow> ((_) heap, exception, unit) prog" and append_child :: "(_) object_ptr \<Rightarrow> (_) node_ptr \<Rightarrow> ((_) heap, exception, unit) prog"
@ -3132,7 +3157,8 @@ lemma insert_before_ptr_in_heap:
shows "ptr |\<in>| object_ptr_kinds h" shows "ptr |\<in>| object_ptr_kinds h"
using assms using assms
apply(auto simp add: insert_before_def elim!: bind_is_OK_E)[1] 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: lemma insert_before_child_in_heap:
assumes "h \<turnstile> ok (insert_before ptr node reference_child)" assumes "h \<turnstile> ok (insert_before ptr node reference_child)"
@ -3288,7 +3314,8 @@ global_interpretation l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<
. .
locale l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M = 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_name set_tag_name_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_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs +
l_set_tag_name type_wf set_tag_name set_tag_name_locs + l_set_tag_name type_wf set_tag_name set_tag_name_locs +
l_create_element_defs create_element + l_create_element_defs create_element +
@ -3329,7 +3356,8 @@ proof -
using new_element_new_ptr h2 new_element_ptr by blast 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" moreover have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_tag_name_writes h3]) apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_tag_name_writes h3])
using set_tag_name_pointers_preserved using set_tag_name_pointers_preserved
by (auto simp add: reflp_def transp_def) by (auto simp add: reflp_def transp_def)
moreover have "document_ptr |\<in>| document_ptr_kinds h3" moreover have "document_ptr |\<in>| document_ptr_kinds h3"
@ -3349,14 +3377,17 @@ proof -
using new_element_is_element_ptr using new_element_is_element_ptr
by blast by blast
then show ?thesis 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 qed
end end
locale l_create_element = l_create_element_defs locale l_create_element = l_create_element_defs
interpretation 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_name set_tag_name_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) 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] declare l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms[instances]
@ -3395,7 +3426,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 = 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_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs +
l_set_val type_wf set_val set_val_locs + l_set_val type_wf set_val set_val_locs +
l_create_character_data_defs create_character_data + l_create_character_data_defs create_character_data +
@ -3436,7 +3468,8 @@ proof -
using new_character_data_new_ptr h2 new_character_data_ptr by blast 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" moreover have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_val_writes h3]) apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_val_writes h3])
using set_val_pointers_preserved using set_val_pointers_preserved
by (auto simp add: reflp_def transp_def) by (auto simp add: reflp_def transp_def)
moreover have "document_ptr |\<in>| document_ptr_kinds h3" moreover have "document_ptr |\<in>| document_ptr_kinds h3"
@ -3456,15 +3489,19 @@ proof -
using new_character_data_is_character_data_ptr using new_character_data_is_character_data_ptr
by blast by blast
then show ?thesis 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 qed
end end
locale l_create_character_data = l_create_character_data_defs locale l_create_character_data = l_create_character_data_defs
interpretation 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 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
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) 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] declare l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
@ -3679,7 +3716,8 @@ l_get_element_by\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M
defines 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" 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 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 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" . 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" .
@ -3690,13 +3728,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_to_tree_order to_tree_order +
l_get_attribute type_wf get_attribute get_attribute_locs l_get_attribute type_wf get_attribute get_attribute_locs
for to_tree_order :: "(_::linorder) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog" for to_tree_order :: "(_::linorder) object_ptr \<Rightarrow> ((_) heap, exception, (_) object_ptr list) prog"
and first_in_tree_order :: "(_) object_ptr \<Rightarrow> ((_) object_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr option) prog) and first_in_tree_order ::
"(_) object_ptr \<Rightarrow> ((_) object_ptr \<Rightarrow> ((_) heap, exception, (_) element_ptr option) prog)
\<Rightarrow> ((_) heap, exception, (_) element_ptr option) prog" \<Rightarrow> ((_) heap, exception, (_) element_ptr option) prog"
and get_attribute :: "(_) element_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, char list option) prog" and get_attribute :: "(_) element_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, char list option) prog"
and get_attribute_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set" and get_attribute_locs :: "(_) element_ptr \<Rightarrow> ((_) heap \<Rightarrow> (_) heap \<Rightarrow> bool) set"
and get_element_by_id :: "(_) object_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr option) prog" and get_element_by_id ::
and get_elements_by_class_name :: "(_) object_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr list) prog" "(_) object_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr option) prog"
and get_elements_by_tag_name :: "(_) object_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr list) prog" and get_elements_by_class_name ::
"(_) object_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr list) prog"
and get_elements_by_tag_name ::
"(_) object_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr list) prog"
and type_wf :: "(_) heap \<Rightarrow> bool" + and type_wf :: "(_) heap \<Rightarrow> bool" +
assumes get_element_by_id_impl: "get_element_by_id = a_get_element_by_id" 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" assumes get_elements_by_class_name_impl: "get_elements_by_class_name = a_get_elements_by_class_name"

View File

@ -164,7 +164,8 @@ lemma get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>
\<longleftrightarrow> 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 \<noteq> None" \<longleftrightarrow> 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 \<noteq> 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 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) 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 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 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: 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_ptr.Ref (Suc (fMax (finsert 0 (character_data_ptr.the_ref |`| character_data_ptrs h))))
|\<notin>| character_data_ptrs h" |\<notin>| 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: 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')" 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 |\<notin>| character_data_ptr_kinds h" shows "new_character_data_ptr |\<notin>| character_data_ptr_kinds h"
using assms 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 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: 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')" 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: lemma known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'" "object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD) 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 |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'" lemma known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def) by(simp add: a_known_ptrs_def)
end 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 . 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 .

View File

@ -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 |\<in>| document_ptr_kinds h \<longleftrightarrow> get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h \<noteq> None" shows "document_ptr |\<in>| document_ptr_kinds h \<longleftrightarrow> get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h \<noteq> None"
using l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_axioms assms 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) 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 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 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]: 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" "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) 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) 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]: 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" "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) 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) 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]: 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: lemma known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'" "object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD) 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 |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'" lemma known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def) by(simp add: a_known_ptrs_def)
end 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 . 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 .

View File

@ -181,18 +181,23 @@ definition a_known_ptrs :: "(_) heap \<Rightarrow> bool"
lemma known_ptrs_known_ptr: "a_known_ptrs h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptr ptr" lemma known_ptrs_known_ptr: "a_known_ptrs h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptr ptr"
apply(simp add: a_known_ptrs_def) apply(simp add: a_known_ptrs_def)
using notin_fset by fastforce using notin_fset by fastforce
lemma known_ptrs_preserved: "object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'" lemma known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
by(auto simp add: a_known_ptrs_def) by(auto simp add: a_known_ptrs_def)
lemma known_ptrs_subset: "object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'" lemma known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD) 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 |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'" lemma known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def) by(simp add: a_known_ptrs_def)
end end
global_interpretation l_known_ptrs\<^sub>N\<^sub>o\<^sub>d\<^sub>e known_ptr defines known_ptrs = a_known_ptrs . 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 lemmas known_ptrs_defs = a_known_ptrs_def
lemma known_ptrs_is_l_known_ptrs: "l_known_ptrs known_ptr known_ptrs" 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 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" 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"

View File

@ -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 \<Rightarrow> bool" + locale l_known_ptrs = l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \<Rightarrow> bool" +
fixes known_ptrs :: "(_) heap \<Rightarrow> bool" fixes known_ptrs :: "(_) heap \<Rightarrow> bool"
assumes known_ptrs_known_ptr: "known_ptrs h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptr ptr" assumes known_ptrs_known_ptr: "known_ptrs h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptr ptr"
assumes known_ptrs_preserved: "object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> known_ptrs h = known_ptrs h'" assumes known_ptrs_preserved:
assumes known_ptrs_subset: "object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> known_ptrs h \<Longrightarrow> known_ptrs h'" "object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> known_ptrs h = known_ptrs h'"
assumes known_ptrs_new_ptr: "object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow> known_ptrs h \<Longrightarrow> known_ptrs h'" assumes known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> known_ptrs h \<Longrightarrow> known_ptrs h'"
assumes known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
known_ptrs h \<Longrightarrow> 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 \<Rightarrow> bool" 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 \<Rightarrow> bool"
begin begin
@ -142,11 +146,15 @@ lemma known_ptrs_known_ptr:
apply(simp add: a_known_ptrs_def) apply(simp add: a_known_ptrs_def)
using notin_fset by fastforce using notin_fset by fastforce
lemma known_ptrs_preserved: "object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'" lemma known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
by(auto simp add: a_known_ptrs_def) by(auto simp add: a_known_ptrs_def)
lemma known_ptrs_subset: "object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'" lemma known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD) 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 |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'" lemma known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def) by(simp add: a_known_ptrs_def)
end 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 . 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 .

View File

@ -319,7 +319,8 @@ lemma type_wf_put_ptr_in_heap_E:
shows "type_wf h" shows "type_wf h"
using assms using assms
apply(auto simp add: type_wf_defs split: option.splits if_splits)[1] 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\<open>Preserving Types\<close> subsection\<open>Preserving Types\<close>
@ -526,6 +527,8 @@ lemma type_wf_drop: "type_wf h \<Longrightarrow> type_wf (Heap (fmdrop ptr (the_
apply(auto simp add: type_wf_def ElementMonad.type_wf_drop 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] 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 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 end

View File

@ -322,7 +322,8 @@ lemma type_wf_put_ptr_in_heap_E:
using assms using assms
apply(auto simp add: type_wf_defs elim!: CharacterDataMonad.type_wf_put_ptr_in_heap_E apply(auto simp add: type_wf_defs elim!: CharacterDataMonad.type_wf_put_ptr_in_heap_E
split: option.splits if_splits)[1] 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)
@ -360,7 +361,9 @@ lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_tag_name_typ
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs 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 ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1] 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) 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]: 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 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 NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1] 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) 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]: 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 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 NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1] 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) 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]: 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 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 NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1] 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) by (metis fmember.rep_eq)
lemma new_character_data_type_wf_preserved [simp]: lemma new_character_data_type_wf_preserved [simp]:
@ -458,7 +467,8 @@ lemma new_document_type_wf_preserved [simp]: "h \<turnstile> new_document \<righ
split: option.splits)[1] split: option.splits)[1]
using document_ptrs_def apply fastforce using document_ptrs_def apply fastforce
apply (simp add: is_document_kind_def) apply (simp add: is_document_kind_def)
apply (metis Suc_n_not_le_n document_ptr.sel(1) document_ptrs_def fMax_ge ffmember_filter fimage_eqI is_document_ptr_ref) apply (metis Suc_n_not_le_n document_ptr.sel(1) document_ptrs_def fMax_ge ffmember_filter
fimage_eqI is_document_ptr_ref)
done done
locale l_new_document = l_type_wf + locale l_new_document = l_type_wf +
@ -599,5 +609,6 @@ lemma type_wf_drop: "type_wf h \<Longrightarrow> type_wf (Heap (fmdrop ptr (the_
apply(auto simp add: type_wf_defs)[1] apply(auto simp add: type_wf_defs)[1]
using type_wf_drop using type_wf_drop
apply blast 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 end

View File

@ -93,7 +93,8 @@ definition
where where
"returns_result_heap h p r h' \<longleftrightarrow> h \<turnstile> p \<rightarrow>\<^sub>r r \<and> h \<turnstile> p \<rightarrow>\<^sub>h h'" "returns_result_heap h p r h' \<longleftrightarrow> h \<turnstile> p \<rightarrow>\<^sub>r r \<and> h \<turnstile> p \<rightarrow>\<^sub>h h'"
lemma return_result_heap_code [code]: "returns_result_heap h p r h' \<longleftrightarrow> (case h \<turnstile> p of Inr (r', h'') \<Rightarrow> r = r' \<and> h' = h'' | Inl _ \<Rightarrow> False)" lemma return_result_heap_code [code]:
"returns_result_heap h p r h' \<longleftrightarrow> (case h \<turnstile> p of Inr (r', h'') \<Rightarrow> r = r' \<and> h' = h'' | Inl _ \<Rightarrow> False)"
by(auto simp add: returns_result_heap_def returns_result_def returns_heap_def split: sum.splits) 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") fun select_result_heap ("|(_)|\<^sub>r\<^sub>h")
@ -452,32 +453,10 @@ fun forall_M :: "('y \<Rightarrow> ('heap, 'e, 'result) prog) \<Rightarrow> 'y l
P x; P x;
forall_M P xs forall_M P xs
}" }"
(*
lemma forall_M_elim:
assumes "h \<turnstile> forall_M P xs \<rightarrow>\<^sub>r True" and "\<And>x h. x \<in> set xs \<Longrightarrow> pure (P x) h"
shows "\<forall>x \<in> set xs. h \<turnstile> P x \<rightarrow>\<^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: "(\<And>x. x \<in> set xs \<Longrightarrow> pure (P x) h) \<Longrightarrow> pure (forall_M P xs) h" lemma pure_forall_M_I: "(\<And>x. x \<in> set xs \<Longrightarrow> pure (P x) h) \<Longrightarrow> pure (forall_M P xs) h"
apply(induct xs) apply(induct xs)
by(auto intro!: bind_pure_I) by(auto intro!: bind_pure_I)
(*
lemma forall_M_pure_I:
assumes "\<And>x. x \<in> set xs \<Longrightarrow> h \<turnstile> P x \<rightarrow>\<^sub>r True" and "\<And>x h. x \<in> set xs \<Longrightarrow> pure (P x)h"
shows "h \<turnstile> forall_M P xs \<rightarrow>\<^sub>r True"
apply(insert assms, induct xs)
apply(simp)
by(fastforce)
lemma forall_M_pure_eq:
assumes "\<And>x. x \<in> set xs \<Longrightarrow> h \<turnstile> P x \<rightarrow>\<^sub>r True \<longleftrightarrow> h' \<turnstile> P x \<rightarrow>\<^sub>r True"
and "\<And>x h. x \<in> set xs \<Longrightarrow> pure (P x) h"
shows "(h \<turnstile> forall_M P xs \<rightarrow>\<^sub>r True) \<longleftrightarrow> h' \<turnstile> forall_M P xs \<rightarrow>\<^sub>r True"
using assms
by(auto intro!: forall_M_pure_I dest!: forall_M_elim) *)
subsection \<open>Fold\<close> subsection \<open>Fold\<close>
@ -506,7 +485,8 @@ lemma filter_M_pure_I [intro]: "(\<And>x. x \<in> set xs \<Longrightarrow> pure
apply(induct xs) apply(induct xs)
by(auto intro!: bind_pure_I) by(auto intro!: bind_pure_I)
lemma filter_M_is_OK_I [intro]: "(\<And>x. x \<in> set xs \<Longrightarrow> h \<turnstile> ok (P x)) \<Longrightarrow> (\<And>x. x \<in> set xs \<Longrightarrow> pure (P x) h) \<Longrightarrow> h \<turnstile> ok (filter_M P xs)" lemma filter_M_is_OK_I [intro]:
"(\<And>x. x \<in> set xs \<Longrightarrow> h \<turnstile> ok (P x)) \<Longrightarrow> (\<And>x. x \<in> set xs \<Longrightarrow> pure (P x) h) \<Longrightarrow> h \<turnstile> ok (filter_M P xs)"
apply(induct xs) apply(induct xs)
apply(simp) apply(simp)
by(auto intro!: bind_is_OK_pure_I) by(auto intro!: bind_is_OK_pure_I)
@ -518,7 +498,8 @@ lemma filter_M_not_more_elements:
by(auto elim!: bind_returns_result_E2 split: if_splits intro!: set_ConsD) by(auto elim!: bind_returns_result_E2 split: if_splits intro!: set_ConsD)
lemma filter_M_in_result_if_ok: lemma filter_M_in_result_if_ok:
assumes "h \<turnstile> filter_M P xs \<rightarrow>\<^sub>r ys" and "\<And>h x. x \<in> set xs \<Longrightarrow> pure (P x) h" and "x \<in> set xs" and "h \<turnstile> P x \<rightarrow>\<^sub>r True" assumes "h \<turnstile> filter_M P xs \<rightarrow>\<^sub>r ys" and "\<And>h x. x \<in> set xs \<Longrightarrow> pure (P x) h" and "x \<in> set xs" and
"h \<turnstile> P x \<rightarrow>\<^sub>r True"
shows "x \<in> set ys" shows "x \<in> set ys"
apply(insert assms, induct xs arbitrary: ys) apply(insert assms, induct xs arbitrary: ys)
apply(simp) apply(simp)
@ -730,7 +711,8 @@ definition preserved :: "('heap, 'e, 'result) prog \<Rightarrow> 'heap \<Rightar
where where
"preserved f h h' \<longleftrightarrow> (\<forall>x. h \<turnstile> f \<rightarrow>\<^sub>r x \<longleftrightarrow> h' \<turnstile> f \<rightarrow>\<^sub>r x)" "preserved f h h' \<longleftrightarrow> (\<forall>x. h \<turnstile> f \<rightarrow>\<^sub>r x \<longleftrightarrow> h' \<turnstile> f \<rightarrow>\<^sub>r x)"
lemma preserved_code [code]: "preserved f h h' = (((h \<turnstile> ok f) \<and> (h' \<turnstile> ok f) \<and> |h \<turnstile> f|\<^sub>r = |h' \<turnstile> f|\<^sub>r) \<or> ((\<not>h \<turnstile> ok f) \<and> (\<not>h' \<turnstile> ok f)))" lemma preserved_code [code]:
"preserved f h h' = (((h \<turnstile> ok f) \<and> (h' \<turnstile> ok f) \<and> |h \<turnstile> f|\<^sub>r = |h' \<turnstile> f|\<^sub>r) \<or> ((\<not>h \<turnstile> ok f) \<and> (\<not>h' \<turnstile> ok f)))"
apply(auto simp add: preserved_def)[1] apply(auto simp add: preserved_def)[1]
apply (meson is_OK_returns_result_E is_OK_returns_result_I)+ apply (meson is_OK_returns_result_E is_OK_returns_result_I)+
done done
@ -768,13 +750,16 @@ lemma reads_bind_pure:
dest: pure_returns_heap_eq dest: pure_returns_heap_eq
elim!: bind_returns_result_E) elim!: bind_returns_result_E)
lemma reads_insert_writes_set_left: "\<forall>P \<in> S. reflp P \<and> transp P \<Longrightarrow> reads {getter} f h h' \<Longrightarrow> reads (insert getter S) f h h'" lemma reads_insert_writes_set_left:
"\<forall>P \<in> S. reflp P \<and> transp P \<Longrightarrow> reads {getter} f h h' \<Longrightarrow> reads (insert getter S) f h h'"
unfolding reads_def by simp unfolding reads_def by simp
lemma reads_insert_writes_set_right: "reflp getter \<Longrightarrow> transp getter \<Longrightarrow> reads S f h h' \<Longrightarrow> reads (insert getter S) f h h'" lemma reads_insert_writes_set_right:
"reflp getter \<Longrightarrow> transp getter \<Longrightarrow> reads S f h h' \<Longrightarrow> reads (insert getter S) f h h'"
unfolding reads_def by blast unfolding reads_def by blast
lemma reads_subset: "reads S f h h' \<Longrightarrow> \<forall>P \<in> S' - S. reflp P \<and> transp P \<Longrightarrow> S \<subseteq> S' \<Longrightarrow> reads S' f h h'" lemma reads_subset:
"reads S f h h' \<Longrightarrow> \<forall>P \<in> S' - S. reflp P \<and> transp P \<Longrightarrow> S \<subseteq> S' \<Longrightarrow> reads S' f h h'"
by(auto simp add: reads_def) by(auto simp add: reads_def)
lemma return_reads [simp]: "reads {} (return x) h h'" lemma return_reads [simp]: "reads {} (return x) h h'"

View File

@ -75,10 +75,12 @@ val _ = Theory.setup
handle Timeout.TIMEOUT _ => NONE; handle Timeout.TIMEOUT _ => NONE;
val t2 = Time.now() - start2; val t2 = Time.now() - start2;
in 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 (case result2_opt of
SOME result2 => 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"))); | NONE => (Output.information "code_simp timed out after 600s"; File.append (Path.explode "/tmp/isabellebench") (">600.000\n")));
result result
end))) end)))

View File

@ -142,19 +142,23 @@ fun get_element_by_id_with_null :: "((_::linorder) object_ptr option) \<Rightarr
| "get_element_by_id_with_null _ _ = error SegmentationFault" | "get_element_by_id_with_null _ _ = error SegmentationFault"
notation get_element_by_id_with_null ("_ . getElementById'(_')") notation get_element_by_id_with_null ("_ . getElementById'(_')")
fun get_elements_by_class_name_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> string \<Rightarrow> (_, ((_) object_ptr option) list) dom_prog" fun get_elements_by_class_name_with_null ::
"((_::linorder) object_ptr option) \<Rightarrow> string \<Rightarrow> (_, ((_) object_ptr option) list) dom_prog"
where where
"get_elements_by_class_name_with_null (Some ptr) class_name = "get_elements_by_class_name_with_null (Some ptr) class_name =
get_elements_by_class_name ptr class_name \<bind> map_M (return \<circ> Some \<circ> 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_class_name ptr class_name \<bind> map_M (return \<circ> Some \<circ> 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'(_')") notation get_elements_by_class_name_with_null ("_ . getElementsByClassName'(_')")
fun get_elements_by_tag_name_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> string \<Rightarrow> (_, ((_) object_ptr option) list) dom_prog" fun get_elements_by_tag_name_with_null ::
"((_::linorder) object_ptr option) \<Rightarrow> string \<Rightarrow> (_, ((_) object_ptr option) list) dom_prog"
where where
"get_elements_by_tag_name_with_null (Some ptr) tag = "get_elements_by_tag_name_with_null (Some ptr) tag =
get_elements_by_tag_name ptr tag \<bind> map_M (return \<circ> Some \<circ> 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 ptr tag \<bind> map_M (return \<circ> Some \<circ> 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'(_')") notation get_elements_by_tag_name_with_null ("_ . getElementsByTagName'(_')")
fun insert_before_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> ((_) object_ptr option) \<Rightarrow> ((_) object_ptr option) \<Rightarrow> (_, ((_) object_ptr option)) dom_prog" fun insert_before_with_null ::
"((_::linorder) object_ptr option) \<Rightarrow> ((_) object_ptr option) \<Rightarrow> ((_) object_ptr option) \<Rightarrow>
(_, ((_) object_ptr option)) dom_prog"
where where
"insert_before_with_null (Some ptr) (Some child_obj) ref_child_obj_opt = (case cast child_obj of "insert_before_with_null (Some ptr) (Some child_obj) ref_child_obj_opt = (case cast child_obj of
Some child \<Rightarrow> do { Some child \<Rightarrow> do {
@ -165,7 +169,8 @@ fun insert_before_with_null :: "((_::linorder) object_ptr option) \<Rightarrow>
| None \<Rightarrow> error HierarchyRequestError)" | None \<Rightarrow> error HierarchyRequestError)"
notation insert_before_with_null ("_ . insertBefore'(_, _')") notation insert_before_with_null ("_ . insertBefore'(_, _')")
fun append_child_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> ((_) object_ptr option) \<Rightarrow> (_, unit) dom_prog" fun append_child_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> ((_) object_ptr option) \<Rightarrow>
(_, unit) dom_prog"
where where
"append_child_with_null (Some ptr) (Some child_obj) = (case cast child_obj of "append_child_with_null (Some ptr) (Some child_obj) = (case cast child_obj of
Some child \<Rightarrow> append_child ptr child Some child \<Rightarrow> append_child ptr child
@ -180,7 +185,8 @@ fun get_body :: "((_::linorder) object_ptr option) \<Rightarrow> (_, ((_) object
}" }"
notation get_body ("_ . body") notation get_body ("_ . body")
fun get_document_element_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> (_, ((_) object_ptr option)) dom_prog" fun get_document_element_with_null :: "((_::linorder) object_ptr option) \<Rightarrow>
(_, ((_) object_ptr option)) dom_prog"
where 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 "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 \<Rightarrow> do { Some document_ptr \<Rightarrow> do {
@ -190,14 +196,16 @@ fun get_document_element_with_null :: "((_::linorder) object_ptr option) \<Right
| None \<Rightarrow> None)})" | None \<Rightarrow> None)})"
notation get_document_element_with_null ("_ . documentElement") notation get_document_element_with_null ("_ . documentElement")
fun get_owner_document_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> (_, ((_) object_ptr option)) dom_prog" fun get_owner_document_with_null :: "((_::linorder) object_ptr option) \<Rightarrow>
(_, ((_) object_ptr option)) dom_prog"
where where
"get_owner_document_with_null (Some ptr) = (do { "get_owner_document_with_null (Some ptr) = (do {
document_ptr \<leftarrow> get_owner_document ptr; document_ptr \<leftarrow> 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))})" 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") notation get_owner_document_with_null ("_ . ownerDocument")
fun remove_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> ((_) object_ptr option) \<Rightarrow> (_, ((_) object_ptr option)) dom_prog" fun remove_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> ((_) object_ptr option) \<Rightarrow>
(_, ((_) object_ptr option)) dom_prog"
where where
"remove_with_null (Some ptr) (Some child) = (case cast child of "remove_with_null (Some ptr) (Some child) = (case cast child of
Some child_node \<Rightarrow> do { Some child_node \<Rightarrow> do {
@ -208,7 +216,8 @@ fun remove_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> ((_) ob
| "remove_with_null _ None = error TypeError" | "remove_with_null _ None = error TypeError"
notation remove_with_null ("_ . remove'(')") notation remove_with_null ("_ . remove'(')")
fun remove_child_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> ((_) object_ptr option) \<Rightarrow> (_, ((_) object_ptr option)) dom_prog" fun remove_child_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> ((_) object_ptr option) \<Rightarrow>
(_, ((_) object_ptr option)) dom_prog"
where where
"remove_child_with_null (Some ptr) (Some child) = (case cast child of "remove_child_with_null (Some ptr) (Some child) = (case cast child of
Some child_node \<Rightarrow> do { Some child_node \<Rightarrow> do {
@ -256,7 +265,8 @@ fun first_child_with_null :: "((_) object_ptr option) \<Rightarrow> (_, ((_) obj
| None \<Rightarrow> None)}" | None \<Rightarrow> None)}"
notation first_child_with_null ("_ . firstChild") notation first_child_with_null ("_ . firstChild")
fun adopt_node_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> ((_) object_ptr option) \<Rightarrow> (_, ((_) object_ptr option)) dom_prog" fun adopt_node_with_null ::
"((_::linorder) object_ptr option) \<Rightarrow> ((_) object_ptr option) \<Rightarrow>(_, ((_) object_ptr option)) dom_prog"
where where
"adopt_node_with_null (Some ptr) (Some child) = (case cast ptr of "adopt_node_with_null (Some ptr) (Some child) = (case cast ptr of
Some document_ptr \<Rightarrow> (case cast child of Some document_ptr \<Rightarrow> (case cast child of
@ -266,7 +276,8 @@ fun adopt_node_with_null :: "((_::linorder) object_ptr option) \<Rightarrow> ((_
notation adopt_node_with_null ("_ . adoptNode'(_')") notation adopt_node_with_null ("_ . adoptNode'(_')")
definition createTestTree :: "((_::linorder) object_ptr option) \<Rightarrow> (_, (string \<Rightarrow> (_, ((_) object_ptr option)) dom_prog)) dom_prog" definition createTestTree ::
"((_::linorder) object_ptr option) \<Rightarrow> (_, (string \<Rightarrow> (_, ((_) object_ptr option)) dom_prog)) dom_prog"
where where
"createTestTree ref = return (\<lambda>id. get_element_by_id_with_null ref id)" "createTestTree ref = return (\<lambda>id. get_element_by_id_with_null ref id)"

File diff suppressed because it is too large Load Diff

View File

@ -48,34 +48,43 @@ record ('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr) RElement
shadow_root_opt :: "'shadow_root_ptr shadow_root_ptr option" shadow_root_opt :: "'shadow_root_ptr shadow_root_ptr option"
type_synonym 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) 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 register_default_tvars
"('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element) Element" "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element) Element"
type_synonym 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, '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 register_default_tvars
"('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Node, 'Element) Node" "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Node, 'Element) Node"
type_synonym type_synonym
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element) Object ('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 register_default_tvars
"('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element) Object" "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element) Object"
type_synonym type_synonym
('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,
= "('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, 'Object, 'Node, 'Element) heap
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) RElement_ext + 'Node) 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 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" 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 \<Rightarrow> (_) element_ptr fset" definition element_ptr_kinds :: "(_) heap \<Rightarrow> (_) element_ptr fset"
where 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]: lemma element_ptr_kinds_simp [simp]:
"element_ptr_kinds (Heap (fmupd (cast element_ptr) element (the_heap h))) = {|element_ptr|} |\<union>| element_ptr_kinds h" "element_ptr_kinds (Heap (fmupd (cast element_ptr) element (the_heap h))) =
{|element_ptr|} |\<union>| element_ptr_kinds h"
apply(auto simp add: element_ptr_kinds_def)[1] apply(auto simp add: element_ptr_kinds_def)[1]
by force by force
@ -85,7 +94,8 @@ definition element_ptrs :: "(_) heap \<Rightarrow> (_) 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 \<Rightarrow> (_) Element option" 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 \<Rightarrow> (_) Element option"
where 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 \<Rightarrow> Some (RNode.extend (RNode.truncate node) element) | _ \<Rightarrow> 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 \<Rightarrow> Some (RNode.extend (RNode.truncate node) element) | _ \<Rightarrow> 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 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 \<Rightarrow> (_) Element option" 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 \<Rightarrow> (_) Element option"
@ -298,11 +308,15 @@ lemma known_ptrs_known_ptr:
apply(simp add: a_known_ptrs_def) apply(simp add: a_known_ptrs_def)
using notin_fset by fastforce using notin_fset by fastforce
lemma known_ptrs_preserved: "object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'" lemma known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
by(auto simp add: a_known_ptrs_def) by(auto simp add: a_known_ptrs_def)
lemma known_ptrs_subset: "object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'" lemma known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD) 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 |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'" lemma known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def) by(simp add: a_known_ptrs_def)
end 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 . 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 .

View File

@ -54,16 +54,21 @@ definition a_owner_document_valid :: "(_) heap \<Rightarrow> bool"
\<and> node_ptr \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r)))" \<and> node_ptr \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r)))"
lemma a_owner_document_valid_code [code]: "a_owner_document_valid h \<longleftrightarrow> node_ptr_kinds h |\<subseteq>| lemma a_owner_document_valid_code [code]: "a_owner_document_valid h \<longleftrightarrow> node_ptr_kinds h |\<subseteq>|
fset_of_list (concat (map (\<lambda>parent. |h \<turnstile> get_child_nodes parent|\<^sub>r) (sorted_list_of_fset (object_ptr_kinds h)) @ map (\<lambda>parent. |h \<turnstile> get_disconnected_nodes parent|\<^sub>r) (sorted_list_of_fset (document_ptr_kinds h)))) fset_of_list (concat (map (\<lambda>parent. |h \<turnstile> get_child_nodes parent|\<^sub>r)
(sorted_list_of_fset (object_ptr_kinds h)) @
map (\<lambda>parent. |h \<turnstile> 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 - proof -
fix x fix x
assume 1: " \<forall>node_ptr\<in>fset (node_ptr_kinds h). assume 1: " \<forall>node_ptr\<in>fset (node_ptr_kinds h).
(\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h \<and> node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r) \<or> (\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h \<and>
node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r) \<or>
(\<exists>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h \<and> node_ptr \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r)" (\<exists>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h \<and> node_ptr \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r)"
assume 2: "x |\<in>| node_ptr_kinds h" assume 2: "x |\<in>| node_ptr_kinds h"
assume 3: "x |\<notin>| fset_of_list (concat (map (\<lambda>parent. |h \<turnstile> get_disconnected_nodes parent|\<^sub>r) (sorted_list_of_fset (document_ptr_kinds h))))" assume 3: "x |\<notin>| fset_of_list (concat (map (\<lambda>parent. |h \<turnstile> get_disconnected_nodes parent|\<^sub>r)
(sorted_list_of_fset (document_ptr_kinds h))))"
have "\<not>(\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h \<and> x \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)" have "\<not>(\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h \<and> x \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)"
using 1 2 3 using 1 2 3
by (smt UN_I fset_of_list_elem image_eqI notin_fset set_concat set_map sorted_list_of_fset_simps(1)) by (smt UN_I fset_of_list_elem image_eqI notin_fset set_concat set_map sorted_list_of_fset_simps(1))
@ -71,25 +76,35 @@ proof -
have "(\<exists>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h \<and> x \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r)" have "(\<exists>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h \<and> x \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r)"
using 1 2 using 1 2
by auto by auto
then obtain parent_ptr where parent_ptr: "parent_ptr |\<in>| object_ptr_kinds h \<and> x \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r" then obtain parent_ptr where parent_ptr: "parent_ptr |\<in>| object_ptr_kinds h \<and>
x \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r"
by auto by auto
moreover have "parent_ptr \<in> set (sorted_list_of_fset (object_ptr_kinds h))" moreover have "parent_ptr \<in> set (sorted_list_of_fset (object_ptr_kinds h))"
using parent_ptr by auto using parent_ptr by auto
moreover have "|h \<turnstile> get_child_nodes parent_ptr|\<^sub>r \<in> set (map (\<lambda>parent. |h \<turnstile> get_child_nodes parent|\<^sub>r) (sorted_list_of_fset (object_ptr_kinds h)))" moreover have "|h \<turnstile> get_child_nodes parent_ptr|\<^sub>r \<in> set (map (\<lambda>parent. |h \<turnstile> get_child_nodes parent|\<^sub>r)
(sorted_list_of_fset (object_ptr_kinds h)))"
using calculation(2) by auto using calculation(2) by auto
ultimately ultimately
show "x |\<in>| fset_of_list (concat (map (\<lambda>parent. |h \<turnstile> get_child_nodes parent|\<^sub>r) (sorted_list_of_fset (object_ptr_kinds h))))" show "x |\<in>| fset_of_list (concat (map (\<lambda>parent. |h \<turnstile> get_child_nodes parent|\<^sub>r)
(sorted_list_of_fset (object_ptr_kinds h))))"
using fset_of_list_elem by fastforce using fset_of_list_elem by fastforce
next next
fix node_ptr fix node_ptr
assume 1: "node_ptr_kinds h |\<subseteq>| fset_of_list (concat (map (\<lambda>parent. |h \<turnstile> get_child_nodes parent|\<^sub>r) (sorted_list_of_fset (object_ptr_kinds h)))) |\<union>| fset_of_list (concat (map (\<lambda>parent. |h \<turnstile> get_disconnected_nodes parent|\<^sub>r) (sorted_list_of_fset (document_ptr_kinds h))))" assume 1: "node_ptr_kinds h |\<subseteq>| fset_of_list (concat (map (\<lambda>parent. |h \<turnstile> get_child_nodes parent|\<^sub>r)
(sorted_list_of_fset (object_ptr_kinds h)))) |\<union>|
fset_of_list (concat (map (\<lambda>parent. |h \<turnstile> get_disconnected_nodes parent|\<^sub>r)
(sorted_list_of_fset (document_ptr_kinds h))))"
assume 2: "node_ptr |\<in>| node_ptr_kinds h" assume 2: "node_ptr |\<in>| node_ptr_kinds h"
assume 3: "\<forall>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h \<longrightarrow> node_ptr \<notin> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r" assume 3: "\<forall>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h \<longrightarrow> node_ptr \<notin> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r"
have "node_ptr \<in> set (concat (map (\<lambda>parent. |h \<turnstile> get_child_nodes parent|\<^sub>r) (sorted_list_of_fset (object_ptr_kinds h)))) \<or> node_ptr \<in> set (concat (map (\<lambda>parent. |h \<turnstile> get_disconnected_nodes parent|\<^sub>r) (sorted_list_of_fset (document_ptr_kinds h))))" have "node_ptr \<in> set (concat (map (\<lambda>parent. |h \<turnstile> get_child_nodes parent|\<^sub>r)
(sorted_list_of_fset (object_ptr_kinds h)))) \<or>
node_ptr \<in> set (concat (map (\<lambda>parent. |h \<turnstile> get_disconnected_nodes parent|\<^sub>r)
(sorted_list_of_fset (document_ptr_kinds h))))"
using 1 2 using 1 2
by (meson fin_mono fset_of_list_elem funion_iff) by (meson fin_mono fset_of_list_elem funion_iff)
then then
show "\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h \<and> node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r" show "\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h \<and>
node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
using 3 using 3
by auto by auto
qed qed
@ -115,7 +130,8 @@ definition a_all_ptrs_in_heap :: "(_) heap \<Rightarrow> bool"
where where
"a_all_ptrs_in_heap h \<longleftrightarrow> "a_all_ptrs_in_heap h \<longleftrightarrow>
(\<forall>ptr \<in> fset (object_ptr_kinds h). set |h \<turnstile> get_child_nodes ptr|\<^sub>r \<subseteq> fset (node_ptr_kinds h)) \<and> (\<forall>ptr \<in> fset (object_ptr_kinds h). set |h \<turnstile> get_child_nodes ptr|\<^sub>r \<subseteq> fset (node_ptr_kinds h)) \<and>
(\<forall>document_ptr \<in> fset (document_ptr_kinds h). set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r \<subseteq> fset (node_ptr_kinds h))" (\<forall>document_ptr \<in> fset (document_ptr_kinds h).
set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r \<subseteq> fset (node_ptr_kinds h))"
definition a_distinct_lists :: "(_) heap \<Rightarrow> bool" definition a_distinct_lists :: "(_) heap \<Rightarrow> bool"
where where
@ -267,7 +283,8 @@ lemma heap_is_wellformed_children_in_heap:
shows "child |\<in>| node_ptr_kinds h" shows "child |\<in>| node_ptr_kinds h"
using assms using assms
apply(auto simp add: heap_is_wellformed_def a_all_ptrs_in_heap_def)[1] 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: lemma heap_is_wellformed_one_parent:
assumes "heap_is_wellformed h" assumes "heap_is_wellformed h"
@ -321,7 +338,8 @@ lemma parent_child_rel_child_in_heap:
lemma heap_is_wellformed_disc_nodes_in_heap: lemma heap_is_wellformed_disc_nodes_in_heap:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes "heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes
\<Longrightarrow> node \<in> set disc_nodes \<Longrightarrow> node |\<in>| node_ptr_kinds h" \<Longrightarrow> node \<in> set disc_nodes \<Longrightarrow> node |\<in>| 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: lemma heap_is_wellformed_one_disc_parent:
"heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes "heap_is_wellformed h \<Longrightarrow> h \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes
@ -716,7 +734,8 @@ proof -
moreover have "a_all_ptrs_in_heap h" moreover have "a_all_ptrs_in_heap h"
using heap_is_wellformed by (simp add: heap_is_wellformed_def) using heap_is_wellformed by (simp add: heap_is_wellformed_def)
then have "a_all_ptrs_in_heap h'" 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" moreover have h0: "a_distinct_lists h"
using heap_is_wellformed by (simp add: heap_is_wellformed_def) using heap_is_wellformed by (simp add: heap_is_wellformed_def)
@ -2501,7 +2520,9 @@ qed
end 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 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 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 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) 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)
@ -2699,17 +2720,20 @@ next
and 2: "known_ptrs h" and 2: "known_ptrs h"
and 3: "\<not> is_element_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" and 3: "\<not> 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 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 \<turnstile> Heap_Error_Monad.bind (check_in_heap ptr) (\<lambda>_. (local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \<circ> the \<circ> 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 ()) \<rightarrow>\<^sub>r owner_document" and 5: "h \<turnstile> Heap_Error_Monad.bind (check_in_heap ptr)
(\<lambda>_. (local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r \<circ> the \<circ> 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 ()) \<rightarrow>\<^sub>r owner_document"
then obtain root where then obtain root where
root: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root" root: "h \<turnstile> get_root_node ptr \<rightarrow>\<^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 then show ?thesis
proof (cases "is_document_ptr root") proof (cases "is_document_ptr root")
case True case True
then show ?thesis then show ?thesis
using 4 5 root 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(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] 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 using "0" "1" "2" document_ptr_kinds_commutes local.get_root_node_root_in_heap by blast
next next
@ -2722,11 +2746,15 @@ next
by blast by blast
then have "is_node_ptr_kind root" then have "is_node_ptr_kind root"
using False \<open>known_ptr root\<close> using False \<open>known_ptr root\<close>
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 using is_node_ptr_kind_none by force
then then
have "(\<exists>document_ptr \<in> fset (document_ptr_kinds h). root \<in> cast ` set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)" have "(\<exists>document_ptr \<in> fset (document_ptr_kinds h). root \<in> cast ` set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)"
by (metis (no_types, lifting) "0" "1" "2" \<open>root |\<in>| object_ptr_kinds h\<close> 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) by (metis (no_types, lifting) "0" "1" "2" \<open>root |\<in>| object_ptr_kinds h\<close> 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 then obtain some_owner_document where
"some_owner_document |\<in>| document_ptr_kinds h" and "some_owner_document |\<in>| document_ptr_kinds h" and
"root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r" "root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r"
@ -2739,25 +2767,31 @@ next
(\<lambda>disconnected_nodes. return (root \<in> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>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))) (\<lambda>disconnected_nodes. return (root \<in> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>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))) (sorted_list_of_set (fset (document_ptr_kinds h)))
\<rightarrow>\<^sub>r candidates" \<rightarrow>\<^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 \<in> set candidates" then have "some_owner_document \<in> set candidates"
apply(rule filter_M_in_result_if_ok) apply(rule filter_M_in_result_if_ok)
using \<open>some_owner_document |\<in>| document_ptr_kinds h\<close> \<open>root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close> using \<open>some_owner_document |\<in>| document_ptr_kinds h\<close>
\<open>root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close>
apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1] apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1]
apply (simp add: \<open>some_owner_document |\<in>| document_ptr_kinds h\<close>) apply (simp add: \<open>some_owner_document |\<in>| document_ptr_kinds h\<close>)
using "1" \<open>root \<in> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>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 \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close> \<open>some_owner_document |\<in>| document_ptr_kinds h\<close> using "1" \<open>root \<in> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>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 \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close>
\<open>some_owner_document |\<in>| document_ptr_kinds h\<close>
local.get_disconnected_nodes_ok by auto local.get_disconnected_nodes_ok by auto
then have "candidates \<noteq> []" then have "candidates \<noteq> []"
by auto by auto
then have "owner_document \<in> set candidates" then have "owner_document \<in> set candidates"
using 5 root 4 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) apply (metis candidates list.set_sel(1) returns_result_eq)
by (metis \<open>is_node_ptr_kind root\<close> node_ptr_no_document_ptr_cast returns_result_eq) by (metis \<open>is_node_ptr_kind root\<close> node_ptr_no_document_ptr_cast returns_result_eq)
then show ?thesis then show ?thesis
using candidates 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
next next
assume 0: "heap_is_wellformed h" assume 0: "heap_is_wellformed h"
@ -2774,7 +2808,8 @@ next
case True case True
then show ?thesis then show ?thesis
using 3 4 root 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(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] 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 using "0" "1" "2" document_ptr_kinds_commutes local.get_root_node_root_in_heap by blast
next next
@ -2787,11 +2822,16 @@ next
by blast by blast
then have "is_node_ptr_kind root" then have "is_node_ptr_kind root"
using False \<open>known_ptr root\<close> using False \<open>known_ptr root\<close>
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 using is_node_ptr_kind_none by force
then then
have "(\<exists>document_ptr \<in> fset (document_ptr_kinds h). root \<in> cast ` set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)" have "(\<exists>document_ptr \<in> fset (document_ptr_kinds h). root \<in> cast ` set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r)"
by (metis (no_types, lifting) "0" "1" "2" \<open>root |\<in>| object_ptr_kinds h\<close> 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) by (metis (no_types, lifting) "0" "1" "2" \<open>root |\<in>| object_ptr_kinds h\<close>
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 then obtain some_owner_document where
"some_owner_document |\<in>| document_ptr_kinds h" and "some_owner_document |\<in>| document_ptr_kinds h" and
"root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r" "root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r"
@ -2804,26 +2844,32 @@ next
(\<lambda>disconnected_nodes. return (root \<in> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>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))) (\<lambda>disconnected_nodes. return (root \<in> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>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))) (sorted_list_of_set (fset (document_ptr_kinds h)))
\<rightarrow>\<^sub>r candidates" \<rightarrow>\<^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 \<in> set candidates" then have "some_owner_document \<in> set candidates"
apply(rule filter_M_in_result_if_ok) apply(rule filter_M_in_result_if_ok)
using \<open>some_owner_document |\<in>| document_ptr_kinds h\<close> \<open>root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close> using \<open>some_owner_document |\<in>| document_ptr_kinds h\<close>
\<open>root \<in> cast ` set |h \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close>
apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1] apply(auto intro!: bind_pure_I bind_pure_returns_result_I)[1]
apply (simp add: \<open>some_owner_document |\<in>| document_ptr_kinds h\<close>) apply (simp add: \<open>some_owner_document |\<in>| document_ptr_kinds h\<close>)
using "1" \<open>root \<in> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>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 \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close> \<open>some_owner_document |\<in>| document_ptr_kinds h\<close> local.get_disconnected_nodes_ok using "1" \<open>root \<in> cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>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 \<turnstile> get_disconnected_nodes some_owner_document|\<^sub>r\<close>
\<open>some_owner_document |\<in>| document_ptr_kinds h\<close> local.get_disconnected_nodes_ok
by auto by auto
then have "candidates \<noteq> []" then have "candidates \<noteq> []"
by auto by auto
then have "owner_document \<in> set candidates" then have "owner_document \<in> set candidates"
using 4 root 3 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) apply (metis candidates list.set_sel(1) returns_result_eq)
by (metis \<open>is_node_ptr_kind root\<close> node_ptr_no_document_ptr_cast returns_result_eq) by (metis \<open>is_node_ptr_kind root\<close> node_ptr_no_document_ptr_cast returns_result_eq)
then show ?thesis then show ?thesis
using candidates 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
qed qed
@ -2839,25 +2885,35 @@ proof -
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+ apply(split invoke_splits, (rule conjI | rule impI)+)+
apply(auto simp add: known_ptr_impl)[1] 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 apply blast
using assms(4) 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 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
apply (metis (no_types, lifting) document_ptr_casts_commute3 document_ptr_kinds_commutes is_document_ptr_kind_none option.case_eq_if) 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) 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 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
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) 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) 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 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
apply(auto split: option.splits intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I)[1] 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 using assms(3) local.get_disconnected_nodes_ok
apply blast 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) 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 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
apply(auto split: option.splits intro!: bind_is_OK_pure_I filter_M_pure_I bind_pure_I filter_M_is_OK_I)[1] 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 (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 using assms(3) local.get_disconnected_nodes_ok by blast
qed qed
@ -2873,28 +2929,39 @@ proof -
using assms(2) local.known_ptrs_known_ptr by blast using assms(2) local.known_ptrs_known_ptr by blast
have "cast child |\<in>| object_ptr_kinds h" have "cast child |\<in>| 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 then
have "known_ptr (cast child)" 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 \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root" obtain root where root: "h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root"
by (meson \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(1) assms(2) assms(3) is_OK_returns_result_E local.get_root_node_ok) by (meson \<open>ptr |\<in>| object_ptr_kinds h\<close> assms(1) assms(2) assms(3) is_OK_returns_result_E
local.get_root_node_ok)
then have "h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root" then have "h \<turnstile> get_root_node (cast child) \<rightarrow>\<^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 \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<longleftrightarrow> h \<turnstile> a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \<rightarrow>\<^sub>r owner_document" have "h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<longleftrightarrow>
h \<turnstile> a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \<rightarrow>\<^sub>r owner_document"
proof (cases "is_document_ptr ptr") proof (cases "is_document_ptr ptr")
case True 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" 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" then have "root = cast document_ptr"
using root using root
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 "h \<turnstile> 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 () \<rightarrow>\<^sub>r owner_document \<longleftrightarrow> h \<turnstile> a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \<rightarrow>\<^sub>r owner_document" then have "h \<turnstile> 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 () \<rightarrow>\<^sub>r owner_document \<longleftrightarrow>
using document_ptr \<open>h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root\<close>[simplified \<open>root = cast document_ptr\<close> document_ptr] h \<turnstile> a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \<rightarrow>\<^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 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 \<open>h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root\<close>[simplified \<open>root = cast document_ptr\<close> document_ptr], rotated] intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: if_splits option.splits)[1] using document_ptr
\<open>h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root\<close>[simplified \<open>root = cast document_ptr\<close> 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 \<open>h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root\<close>[simplified \<open>root = cast document_ptr\<close> document_ptr], rotated]
intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: if_splits option.splits)[1]
using \<open>ptr |\<in>| object_ptr_kinds h\<close> document_ptr_kinds_commutes by blast using \<open>ptr |\<in>| object_ptr_kinds h\<close> document_ptr_kinds_commutes by blast
then show ?thesis then show ?thesis
using \<open>known_ptr ptr\<close> using \<open>known_ptr ptr\<close>
@ -2905,13 +2972,16 @@ proof -
apply(drule(1) known_ptr_not_element_ptr) apply(drule(1) known_ptr_not_element_ptr)
apply(simp add: NodeClass.known_ptr_defs) apply(simp add: NodeClass.known_ptr_defs)
using \<open>ptr |\<in>| object_ptr_kinds h\<close> True using \<open>ptr |\<in>| object_ptr_kinds h\<close> 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 next
case False 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" 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 \<open>known_ptr ptr\<close> using \<open>known_ptr ptr\<close>
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
then have "h \<turnstile> a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document \<longleftrightarrow> h \<turnstile> a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \<rightarrow>\<^sub>r owner_document" ElementClass.known_ptr_defs NodeClass.known_ptr_defs split: option.splits)
then have "h \<turnstile> a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document \<longleftrightarrow>
h \<turnstile> a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r child () \<rightarrow>\<^sub>r owner_document"
using root \<open>h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root\<close> using root \<open>h \<turnstile> get_root_node (cast child) \<rightarrow>\<^sub>r root\<close>
unfolding a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def 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) by (meson bind_pure_returns_result_I bind_returns_result_E3 local.get_root_node_pure)
@ -2940,7 +3010,8 @@ proof -
qed qed
then show ?thesis then show ?thesis
using \<open>known_ptr (cast child)\<close> using \<open>known_ptr (cast child)\<close>
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(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_document_ptr[folded known_ptr_impl])
apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_character_data_ptr)
@ -2954,7 +3025,9 @@ proof -
apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1] apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1]
using \<open>cast child |\<in>| object_ptr_kinds h\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close> using \<open>cast child |\<in>| object_ptr_kinds h\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close>
apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1] apply(auto intro!: bind_pure_returns_result_I split: option.splits)[1]
by (smt \<open>cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>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 |\<in>| object_ptr_kinds h\<close> 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 \<open>cast\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>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 |\<in>| object_ptr_kinds h\<close> 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 qed
end end
@ -2978,15 +3051,21 @@ locale l_get_owner_document_wf = l_heap_is_wellformed_defs + l_type_wf + l_known
type_wf h\<Longrightarrow> type_wf h\<Longrightarrow>
node_ptr \<in> set disc_nodes" node_ptr \<in> set disc_nodes"
assumes get_owner_document_owner_document_in_heap: assumes get_owner_document_owner_document_in_heap:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<Longrightarrow> owner_document |\<in>| document_ptr_kinds h" "heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow>
h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<Longrightarrow>
owner_document |\<in>| document_ptr_kinds h"
assumes get_owner_document_ok: assumes get_owner_document_ok:
"heap_is_wellformed h \<Longrightarrow> known_ptrs h \<Longrightarrow> type_wf h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h "heap_is_wellformed h \<Longrightarrow> known_ptrs h \<Longrightarrow> type_wf h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h
\<Longrightarrow> h \<turnstile> ok (get_owner_document ptr)" \<Longrightarrow> h \<turnstile> ok (get_owner_document ptr)"
assumes get_owner_document_child_same: assumes get_owner_document_child_same:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow> child \<in> set children \<Longrightarrow> h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<longleftrightarrow> h \<turnstile> get_owner_document (cast child) \<rightarrow>\<^sub>r owner_document" "heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children \<Longrightarrow>
child \<in> set children \<Longrightarrow> h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<longleftrightarrow>
h \<turnstile> get_owner_document (cast child) \<rightarrow>\<^sub>r owner_document"
interpretation i_get_owner_document_wf?: l_get_owner_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M 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) 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] declare l_get_owner_document_wf\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
@ -3027,7 +3106,8 @@ proof -
assume "is_document_ptr_kind ptr" assume "is_document_ptr_kind ptr"
then have "ptr = root" then have "ptr = root"
using assms(4) 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 then have ?thesis
using \<open>is_document_ptr_kind ptr\<close> \<open>known_ptr ptr\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close> using \<open>is_document_ptr_kind ptr\<close> \<open>known_ptr ptr\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close>
apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def)[1] apply(auto simp add: known_ptr_impl get_owner_document_def a_get_owner_document_tups_def)[1]
@ -3036,7 +3116,8 @@ proof -
apply(drule(1) known_ptr_not_character_data_ptr) apply(drule(1) known_ptr_not_character_data_ptr)
apply(drule(1) known_ptr_not_element_ptr) apply(drule(1) known_ptr_not_element_ptr)
apply(simp add: NodeClass.known_ptr_defs) 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 moreover
{ {
@ -3051,12 +3132,14 @@ proof -
apply(simp add: NodeClass.known_ptr_defs) apply(simp add: NodeClass.known_ptr_defs)
apply(auto split: option.splits)[1] apply(auto split: option.splits)[1]
using \<open>h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root\<close> assms(5) using \<open>h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root\<close> 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 show ?thesis
using \<open>known_ptr ptr\<close> using \<open>known_ptr ptr\<close>
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 qed
lemma get_root_node_same_owner_document: lemma get_root_node_same_owner_document:
@ -3086,7 +3169,8 @@ proof -
case False case False
then have "is_node_ptr_kind ptr" then have "is_node_ptr_kind ptr"
using \<open>known_ptr ptr\<close> using \<open>known_ptr ptr\<close>
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" 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) by (metis node_ptr_casts_commute3)
show ?thesis show ?thesis
@ -3104,10 +3188,12 @@ proof -
case True case True
have "is_document_ptr root" have "is_document_ptr root"
using True \<open>known_ptr root\<close> using True \<open>known_ptr root\<close>
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" have "root = cast owner_document"
using True using True
by (smt \<open>h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document\<close> assms(1) assms(2) assms(3) assms(4) document_ptr_casts_commute3 get_root_node_document returns_result_eq) by (smt \<open>h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document\<close> assms(1) assms(2) assms(3) assms(4)
document_ptr_casts_commute3 get_root_node_document returns_result_eq)
then show ?thesis then show ?thesis
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+ apply(split invoke_splits, (rule conjI | rule impI)+)+
@ -3119,20 +3205,27 @@ proof -
case False case False
then have "is_node_ptr_kind root" then have "is_node_ptr_kind root"
using \<open>known_ptr root\<close> using \<open>known_ptr root\<close>
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" 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) by (metis node_ptr_casts_commute3)
then have "h \<turnstile> local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr () \<rightarrow>\<^sub>r owner_document" then have "h \<turnstile> local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr () \<rightarrow>\<^sub>r owner_document"
using \<open>h \<turnstile> local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document\<close> assms(4) using \<open>h \<turnstile> local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document\<close> assms(4)
apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: option.splits)[1] apply(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
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) 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 \<open>is_node_ptr_kind root\<close> node_ptr returns_result_eq by fastforce using \<open>is_node_ptr_kind root\<close> node_ptr returns_result_eq by fastforce
then show ?thesis then show ?thesis
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+ apply(split invoke_splits, (rule conjI | rule impI)+)+
using \<open>is_node_ptr_kind root\<close> \<open>known_ptr root\<close> using \<open>is_node_ptr_kind root\<close> \<open>known_ptr root\<close>
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
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] 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 \<open>root |\<in>| object_ptr_kinds h\<close> using \<open>root |\<in>| object_ptr_kinds h\<close>
by(auto simp add: root_node_ptr) by(auto simp add: root_node_ptr)
qed qed
@ -3146,9 +3239,12 @@ proof -
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits)+ apply(split invoke_splits)+
apply (meson invoke_empty is_OK_returns_result_I) apply (meson invoke_empty is_OK_returns_result_I)
apply(auto simp add: True a_get_owner_document\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 split: if_splits)[1] apply(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
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) split: if_splits)[1]
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 (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 then show ?thesis
using assms(1) assms(2) assms(3) assms(4) get_root_node_document using assms(1) assms(2) assms(3) assms(4) get_root_node_document
by fastforce by fastforce
@ -3156,7 +3252,8 @@ proof -
case False case False
then have "is_node_ptr_kind root" then have "is_node_ptr_kind root"
using \<open>known_ptr root\<close> using \<open>known_ptr root\<close>
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" 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) by (metis node_ptr_casts_commute3)
then have "h \<turnstile> local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr () \<rightarrow>\<^sub>r owner_document" then have "h \<turnstile> local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r root_node_ptr () \<rightarrow>\<^sub>r owner_document"
@ -3166,32 +3263,46 @@ proof -
apply (meson invoke_empty is_OK_returns_result_I) apply (meson invoke_empty is_OK_returns_result_I)
by(auto simp add: is_document_ptr_kind_none elim!: bind_returns_result_E2) by(auto simp add: is_document_ptr_kind_none elim!: bind_returns_result_E2)
then have "h \<turnstile> local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document" then have "h \<turnstile> local.a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r node_ptr () \<rightarrow>\<^sub>r owner_document"
apply(auto simp add: a_get_owner_document\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def elim!: bind_returns_result_E2 intro!: bind_pure_returns_result_I filter_M_pure_I bind_pure_I split: option.splits)[1] 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
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 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+ by fastforce+
then show ?thesis then show ?thesis
apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1] apply(auto simp add: get_owner_document_def a_get_owner_document_tups_def)[1]
apply(split invoke_splits, (rule conjI | rule impI)+)+ apply(split invoke_splits, (rule conjI | rule impI)+)+
using node_ptr \<open>known_ptr ptr\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close> using node_ptr \<open>known_ptr ptr\<close> \<open>ptr |\<in>| object_ptr_kinds h\<close>
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
qed qed
qed qed
end 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) 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] 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 + locale l_get_owner_document_wf_get_root_node_wf = l_heap_is_wellformed_defs + l_type_wf +
assumes get_root_node_document: "heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root \<Longrightarrow> is_document_ptr_kind root \<Longrightarrow> h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r the (cast root)" l_known_ptrs + l_get_root_node_defs + l_get_owner_document_defs +
assumes get_root_node_same_owner_document: "heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root \<Longrightarrow> h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<longleftrightarrow> h \<turnstile> get_owner_document root \<rightarrow>\<^sub>r owner_document" assumes get_root_node_document:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root \<Longrightarrow>
is_document_ptr_kind root \<Longrightarrow> h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r the (cast root)"
assumes get_root_node_same_owner_document:
"heap_is_wellformed h \<Longrightarrow> type_wf h \<Longrightarrow> known_ptrs h \<Longrightarrow> h \<turnstile> get_root_node ptr \<rightarrow>\<^sub>r root \<Longrightarrow>
h \<turnstile> get_owner_document ptr \<rightarrow>\<^sub>r owner_document \<longleftrightarrow> h \<turnstile> get_owner_document root \<rightarrow>\<^sub>r owner_document"
lemma get_owner_document_wf_get_root_node_wf_is_l_get_owner_document_wf_get_root_node_wf [instances]: 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" "l_get_owner_document_wf_get_root_node_wf heap_is_wellformed type_wf known_ptr known_ptrs
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] 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_document apply blast
using get_root_node_same_owner_document apply (blast, blast) using get_root_node_same_owner_document apply (blast, blast)
done done
@ -3503,7 +3614,8 @@ proof -
then have document_ptr_kinds_eq3: "document_ptr_kinds h = document_ptr_kinds h'" then have document_ptr_kinds_eq3: "document_ptr_kinds h = document_ptr_kinds h'"
using document_ptr_kinds_M_eq by auto using document_ptr_kinds_M_eq by auto
have children_eq: have children_eq:
"\<And>ptr' children. ptr \<noteq> ptr' \<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children = h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children" "\<And>ptr' children. ptr \<noteq> ptr' \<Longrightarrow> h \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children =
h' \<turnstile> get_child_nodes ptr' \<rightarrow>\<^sub>r children"
apply(rule reads_writes_preserved[OF get_child_nodes_reads remove_child_writes assms(2)]) apply(rule reads_writes_preserved[OF get_child_nodes_reads remove_child_writes assms(2)])
unfolding remove_child_locs_def 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
@ -3516,7 +3628,8 @@ proof -
= h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disconnected_nodes" = h' \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^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 unfolding remove_child_locs_def
using set_child_nodes_get_disconnected_nodes set_disconnected_nodes_get_disconnected_nodes_different_pointers 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) by (metis (no_types, lifting) Un_iff owner_document select_result_I2)
then have disconnected_nodes_eq2: then have disconnected_nodes_eq2:
"\<And>document_ptr. document_ptr \<noteq> owner_document "\<And>document_ptr. document_ptr \<noteq> owner_document
@ -3524,7 +3637,8 @@ proof -
using select_result_eq by force using select_result_eq by force
have "h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children_h" have "h2 \<turnstile> get_child_nodes ptr \<rightarrow>\<^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) by (simp add: set_disconnected_nodes_get_child_nodes)
show "known_ptrs h'" show "known_ptrs h'"
@ -3534,7 +3648,8 @@ proof -
using assms(3) using assms(3)
using children_h get_child_nodes_ptr_in_heap local.known_ptrs_known_ptr by blast using children_h get_child_nodes_ptr_in_heap local.known_ptrs_known_ptr by blast
have "type_wf h2" have "type_wf h2"
using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'", OF set_disconnected_nodes_writes h2] using writes_small_big[where P="\<lambda>h h'. type_wf h \<longrightarrow> type_wf h'",
OF set_disconnected_nodes_writes h2]
using set_disconnected_nodes_types_preserved type_wf using set_disconnected_nodes_types_preserved type_wf
by(auto simp add: reflp_def transp_def) by(auto simp add: reflp_def transp_def)
then show "type_wf h'" then show "type_wf h'"
@ -3561,14 +3676,16 @@ have "type_wf h2"
using \<open>type_wf h2\<close> set_child_nodes_get_child_nodes \<open>known_ptr ptr\<close> h' using \<open>type_wf h2\<close> set_child_nodes_get_child_nodes \<open>known_ptr ptr\<close> h'
by blast by blast
have disconnected_nodes_h2: "h2 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r child # disconnected_nodes_h" have disconnected_nodes_h2:
"h2 \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r child # disconnected_nodes_h"
using owner_document assms(2) h2 disconnected_nodes_h using owner_document assms(2) h2 disconnected_nodes_h
apply (auto simp add: remove_child_def split: if_splits)[1] apply (auto simp add: remove_child_def split: if_splits)[1]
apply(drule bind_returns_heap_E2) apply(drule bind_returns_heap_E2)
apply(auto split: if_splits)[1] apply(auto split: if_splits)[1]
apply(simp) apply(simp)
by(auto simp add: local.set_disconnected_nodes_get_disconnected_nodes split: if_splits) by(auto simp add: local.set_disconnected_nodes_get_disconnected_nodes split: if_splits)
then have disconnected_nodes_h': "h' \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r child # disconnected_nodes_h" then have disconnected_nodes_h':
"h' \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r child # disconnected_nodes_h"
apply(rule reads_writes_separate_forwards[OF get_disconnected_nodes_reads set_child_nodes_writes h']) apply(rule reads_writes_separate_forwards[OF get_disconnected_nodes_reads set_child_nodes_writes h'])
by (simp add: set_child_nodes_get_disconnected_nodes) by (simp add: set_child_nodes_get_disconnected_nodes)
@ -3600,8 +3717,12 @@ have "type_wf h2"
using assms(1) by (simp add: heap_is_wellformed_def) using assms(1) 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 node_ptr_kinds_eq3 disconnected_nodes_eq)[1] apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq3 disconnected_nodes_eq)[1]
apply (metis (no_types, lifting) \<open>type_wf h'\<close> 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) \<open>type_wf h'\<close> assms(2) assms(3) local.get_child_nodes_ok
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)) 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 done
moreover have "a_owner_document_valid h" moreover have "a_owner_document_valid h"
using assms(1) by (simp add: heap_is_wellformed_def) using assms(1) by (simp add: heap_is_wellformed_def)
@ -3610,9 +3731,12 @@ have "type_wf h2"
node_ptr_kinds_eq3)[1] node_ptr_kinds_eq3)[1]
proof - proof -
fix node_ptr fix node_ptr
assume 0: "\<forall>node_ptr\<in>fset (node_ptr_kinds h'). (\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h' \<and> node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r) \<or> (\<exists>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h' \<and> node_ptr \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r)" assume 0: "\<forall>node_ptr\<in>fset (node_ptr_kinds h'). (\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h' \<and>
node_ptr \<in> set |h \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r) \<or>
(\<exists>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h' \<and> node_ptr \<in> set |h \<turnstile> get_child_nodes parent_ptr|\<^sub>r)"
and 1: "node_ptr |\<in>| node_ptr_kinds h'" and 1: "node_ptr |\<in>| node_ptr_kinds h'"
and 2: "\<forall>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h' \<longrightarrow> node_ptr \<notin> set |h' \<turnstile> get_child_nodes parent_ptr|\<^sub>r" and 2: "\<forall>parent_ptr. parent_ptr |\<in>| object_ptr_kinds h' \<longrightarrow>
node_ptr \<notin> set |h' \<turnstile> get_child_nodes parent_ptr|\<^sub>r"
then show "\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h' then show "\<exists>document_ptr. document_ptr |\<in>| document_ptr_kinds h'
\<and> node_ptr \<in> set |h' \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r" \<and> node_ptr \<in> set |h' \<turnstile> get_disconnected_nodes document_ptr|\<^sub>r"
proof (cases "node_ptr = child") proof (cases "node_ptr = child")
@ -3933,7 +4057,8 @@ lemma remove_heap_is_wellformed_preserved:
and type_wf: "type_wf h" and type_wf: "type_wf h"
shows "type_wf h'" and "known_ptrs h'" and "heap_is_wellformed h'" shows "type_wf h'" and "known_ptrs h'" and "heap_is_wellformed h'"
using assms 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: lemma remove_child_removes_child:
assumes wellformed: "heap_is_wellformed h" assumes wellformed: "heap_is_wellformed h"
@ -4553,11 +4678,18 @@ proof -
then have "a_all_ptrs_in_heap h3" 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(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)) apply (simp add: children_eq2_h2 object_ptr_kinds_h2_eq3 subset_code(1))
by (metis (no_types, lifting) \<open>child \<in> set disc_nodes_old_document_h2\<close> \<open>type_wf h2\<close> 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) by (metis (no_types, lifting) \<open>child \<in> set disc_nodes_old_document_h2\<close> \<open>type_wf h2\<close>
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'" 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(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)) apply (simp add: children_eq2_h3 object_ptr_kinds_h3_eq3 subset_code(1))
by (metis (no_types, lifting) \<open>child \<in> set disc_nodes_old_document_h2\<close> 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) by (metis (no_types, lifting) \<open>child \<in> set disc_nodes_old_document_h2\<close> 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" moreover have "a_owner_document_valid h2"
using wellformed_h2 by (simp add: heap_is_wellformed_def) using wellformed_h2 by (simp add: heap_is_wellformed_def)
@ -5275,13 +5407,19 @@ proof -
children_eq_h2)[1] children_eq_h2)[1]
using disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3 using disconnected_nodes_eq2_h2 disconnected_nodes_h2 disconnected_nodes_h3
using node_ptr_kinds_eq2_h2 apply auto[1] using node_ptr_kinds_eq2_h2 apply auto[1]
apply (metis \<open>known_ptrs h2\<close> \<open>type_wf h3\<close> 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) apply (metis \<open>known_ptrs h2\<close> \<open>type_wf h3\<close> children_eq_h2 local.get_child_nodes_ok
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) 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 \<subseteq> set |h' \<turnstile> node_ptr_kinds_M|\<^sub>r" have "set children_h3 \<subseteq> set |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using children_h3 \<open>a_all_ptrs_in_heap h3\<close> using children_h3 \<open>a_all_ptrs_in_heap h3\<close>
apply(auto simp add: a_all_ptrs_in_heap_def node_ptr_kinds_eq2_h3)[1] 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) \<subseteq> set |h' \<turnstile> node_ptr_kinds_M|\<^sub>r" then have "set (insert_before_list node reference_child children_h3) \<subseteq> set |h' \<turnstile> node_ptr_kinds_M|\<^sub>r"
using node_in_heap using node_in_heap
@ -5295,7 +5433,9 @@ proof -
node_ptr_kinds_eq2_h3 disconnected_nodes_eq_h3)[1] node_ptr_kinds_eq2_h3 disconnected_nodes_eq_h3)[1]
using children_eq_h3 children_h' using children_eq_h3 children_h'
apply (metis (no_types, lifting) children_eq2_h3 finite_set_in select_result_I2 subsetD) apply (metis (no_types, lifting) children_eq2_h3 finite_set_in select_result_I2 subsetD)
by (metis (no_types) \<open>type_wf h'\<close> 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) \<open>type_wf h'\<close> 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 qed
moreover have "a_distinct_lists h2" moreover have "a_distinct_lists h2"
@ -5489,7 +5629,9 @@ proof -
object_ptr_kinds_M_eq2_h2[simplified] object_ptr_kinds_M_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] node_ptr_kinds_eq2_h2[simplified] node_ptr_kinds_eq2_h3[simplified])[1]
apply(auto simp add: disconnected_nodes_eq2_h3[symmetric])[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'" ultimately show "heap_is_wellformed h'"
by (simp add: heap_is_wellformed_def) by (simp add: heap_is_wellformed_def)
@ -5543,7 +5685,8 @@ lemma append_child_heap_is_wellformed_preserved:
and type_wf: "type_wf h" and type_wf: "type_wf h"
shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs h'" shows "heap_is_wellformed h'" and "type_wf h'" and "known_ptrs h'"
using assms 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: lemma append_child_children:
assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h" assumes "heap_is_wellformed h" and "type_wf h" and "known_ptrs h"
@ -5689,7 +5832,8 @@ interpretation i_append_child_wf?: l_append_child_wf\<^sub>C\<^sub>o\<^sub>r\<^s
parent_child_rel 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) 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] 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+ using append_child_heap_is_wellformed_preserved by fast+
@ -5750,7 +5894,8 @@ proof -
then have "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr" then have "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr"
apply(auto simp add: create_element_def intro!: bind_returns_result_I)[1] 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_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_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) by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
have "new_element_ptr \<notin> set |h \<turnstile> element_ptr_kinds_M|\<^sub>r" have "new_element_ptr \<notin> set |h \<turnstile> element_ptr_kinds_M|\<^sub>r"
@ -5777,7 +5922,8 @@ proof -
by(auto simp add: document_ptr_kinds_def) by(auto simp add: document_ptr_kinds_def)
have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2" have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_tag_name_writes h3]) apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h",
OF set_tag_name_writes h3])
using set_tag_name_pointers_preserved using set_tag_name_pointers_preserved
by (auto simp add: reflp_def transp_def) by (auto simp add: reflp_def transp_def)
then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2" then have document_ptr_kinds_eq_h2: "document_ptr_kinds h3 = document_ptr_kinds h2"
@ -5798,7 +5944,8 @@ proof -
by(auto simp add: node_ptr_kinds_def) by(auto simp add: node_ptr_kinds_def)
have "known_ptr (cast new_element_ptr)" have "known_ptr (cast new_element_ptr)"
using \<open>h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr\<close> local.create_element_known_ptr by blast using \<open>h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr\<close> local.create_element_known_ptr
by blast
then then
have "known_ptrs h2" have "known_ptrs h2"
using known_ptrs_new_ptr object_ptr_kinds_eq_h \<open>known_ptrs h\<close> h2 using known_ptrs_new_ptr object_ptr_kinds_eq_h \<open>known_ptrs h\<close> h2
@ -5940,12 +6087,24 @@ proof -
using \<open>heap_is_wellformed h\<close> by (simp add: heap_is_wellformed_def) using \<open>heap_is_wellformed h\<close> by (simp add: heap_is_wellformed_def)
then have "a_all_ptrs_in_heap h2" then have "a_all_ptrs_in_heap h2"
apply(auto simp add: a_all_ptrs_in_heap_def)[1] apply(auto simp add: a_all_ptrs_in_heap_def)[1]
apply (metis \<open>known_ptrs h2\<close> \<open>parent_child_rel h = parent_child_rel h2\<close> \<open>type_wf h2\<close> 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) apply (metis \<open>known_ptrs h2\<close> \<open>parent_child_rel h = parent_child_rel h2\<close> \<open>type_wf h2\<close> assms(1)
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) 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" 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'" then have "a_all_ptrs_in_heap h'"
by (smt \<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close> 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 \<open>h2 \<turnstile> 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) \<rightarrow>\<^sub>r []\<close> children_eq2_h3
disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3
document_ptr_kinds_eq_h3 finite_set_in h' is_OK_returns_result_I
l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
local.a_all_ptrs_in_heap_def local.get_child_nodes_ptr_in_heap
local.l_set_disconnected_nodes_get_disconnected_nodes_axioms node_ptr_kinds_commutes
object_ptr_kinds_eq_h2 object_ptr_kinds_eq_h3 select_result_I2 set_ConsD subset_code(1))
have "\<And>p. p |\<in>| object_ptr_kinds h \<Longrightarrow> cast new_element_ptr \<notin> set |h \<turnstile> get_child_nodes p|\<^sub>r" have "\<And>p. p |\<in>| object_ptr_kinds h \<Longrightarrow> cast new_element_ptr \<notin> set |h \<turnstile> get_child_nodes p|\<^sub>r"
using \<open>heap_is_wellformed h\<close> \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close> using \<open>heap_is_wellformed h\<close> \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
@ -6014,14 +6173,17 @@ proof -
ultimately show "False" ultimately show "False"
apply(-) apply(-)
apply(cases "x = document_ptr") apply(cases "x = document_ptr")
apply (smt NodeMonad.ptr_kinds_ptr_kinds_M \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close> \<open>local.a_all_ptrs_in_heap h\<close> apply (smt NodeMonad.ptr_kinds_ptr_kinds_M \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close>
\<open>local.a_all_ptrs_in_heap h\<close>
disc_nodes_h3 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 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' 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 l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms
select_result_I2 set_ConsD subsetD) select_result_I2 set_ConsD subsetD)
by (smt NodeMonad.ptr_kinds_ptr_kinds_M \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close> \<open>local.a_all_ptrs_in_heap h\<close> by (smt NodeMonad.ptr_kinds_ptr_kinds_M
disc_nodes_document_ptr_h2 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3 \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close> \<open>local.a_all_ptrs_in_heap h\<close>
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' 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 l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms local.a_all_ptrs_in_heap_def local.l_set_disconnected_nodes_get_disconnected_nodes_axioms
@ -6064,7 +6226,13 @@ proof -
apply (metis (no_types, lifting) document_ptr_kinds_eq_h h' list.set_intros(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) local.set_disconnected_nodes_get_disconnected_nodes select_result_I2)
apply(simp add: object_ptr_kinds_eq_h) apply(simp add: object_ptr_kinds_eq_h)
by(metis (no_types, lifting) NodeMonad.ptr_kinds_ptr_kinds_M \<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close> 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(metis (no_types, lifting) NodeMonad.ptr_kinds_ptr_kinds_M
\<open>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 \<notin> set |h \<turnstile> node_ptr_kinds_M|\<^sub>r\<close> children_eq2_h children_eq2_h2
children_eq2_h3 disconnected_nodes_eq2_h disconnected_nodes_eq2_h2 disconnected_nodes_eq2_h3
document_ptr_kinds_eq_h finite_set_in h'
l_set_disconnected_nodes_get_disconnected_nodes.set_disconnected_nodes_get_disconnected_nodes
list.set_intros(2) local.l_set_disconnected_nodes_get_disconnected_nodes_axioms
node_ptr_kinds_commutes select_result_I2)
show "heap_is_wellformed h'" show "heap_is_wellformed h'"
using \<open>a_acyclic_heap h'\<close> \<open>a_all_ptrs_in_heap h'\<close> \<open>a_distinct_lists h'\<close> \<open>a_owner_document_valid h'\<close> using \<open>a_acyclic_heap h'\<close> \<open>a_all_ptrs_in_heap h'\<close> \<open>a_distinct_lists h'\<close> \<open>a_owner_document_valid h'\<close>
@ -6148,7 +6316,8 @@ proof -
then have "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr" then have "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr"
apply(auto simp add: create_character_data_def intro!: bind_returns_result_I)[1] 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_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_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) by (metis is_OK_returns_heap_I is_OK_returns_result_E old.unit.exhaust)
@ -6203,7 +6372,8 @@ proof -
by(auto simp add: node_ptr_kinds_def) by(auto simp add: node_ptr_kinds_def)
have "known_ptr (cast new_character_data_ptr)" have "known_ptr (cast new_character_data_ptr)"
using \<open>h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr\<close> local.create_character_data_known_ptr by blast using \<open>h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr\<close>
local.create_character_data_known_ptr by blast
then then
have "known_ptrs h2" have "known_ptrs h2"
using known_ptrs_new_ptr object_ptr_kinds_eq_h \<open>known_ptrs h\<close> h2 using known_ptrs_new_ptr object_ptr_kinds_eq_h \<open>known_ptrs h\<close> h2
@ -6696,11 +6866,15 @@ proof -
local.heap_is_wellformed_children_in_heap local.parent_child_rel_child local.heap_is_wellformed_children_in_heap local.parent_child_rel_child
local.parent_child_rel_parent_in_heap node_ptr_kinds_eq local.parent_child_rel_parent_in_heap node_ptr_kinds_eq
apply (metis (no_types, lifting) \<open>h' \<turnstile> 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) \<rightarrow>\<^sub>r []\<close> apply (metis (no_types, lifting) \<open>h' \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
children_eq2 finite_set_in finsert_iff funion_finsert_right object_ptr_kinds_eq select_result_I2 subsetD sup_bot.right_neutral) 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) \<open>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 |\<notin>| object_ptr_kinds h\<close> by (metis (no_types, lifting) \<open>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 |\<notin>| object_ptr_kinds h\<close>
\<open>h' \<turnstile> 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) \<rightarrow>\<^sub>r []\<close> \<open>h' \<turnstile> 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) \<rightarrow>\<^sub>r []\<close>
\<open>h' \<turnstile> get_disconnected_nodes new_document_ptr \<rightarrow>\<^sub>r []\<close> \<open>parent_child_rel h = parent_child_rel h'\<close> \<open>type_wf h'\<close> assms(1) disconnected_nodes_eq_h local.get_disconnected_nodes_ok \<open>h' \<turnstile> get_disconnected_nodes new_document_ptr \<rightarrow>\<^sub>r []\<close>
local.heap_is_wellformed_disc_nodes_in_heap local.parent_child_rel_child local.parent_child_rel_parent_in_heap \<open>parent_child_rel h = parent_child_rel h'\<close> \<open>type_wf h'\<close> 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) node_ptr_kinds_eq returns_result_select_result select_result_I2)
have "a_distinct_lists h" have "a_distinct_lists h"
using \<open>heap_is_wellformed h\<close> using \<open>heap_is_wellformed h\<close>
@ -6770,7 +6944,8 @@ proof -
then have "a_owner_document_valid h'" then have "a_owner_document_valid h'"
apply(auto simp add: a_owner_document_valid_def)[1] apply(auto simp add: a_owner_document_valid_def)[1]
by (metis \<open>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 |\<notin>| object_ptr_kinds h\<close> by (metis \<open>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 |\<notin>| object_ptr_kinds h\<close>
children_eq2 disconnected_nodes_eq2_h document_ptr_kinds_commutes finite_set_in funion_iff node_ptr_kinds_eq object_ptr_kinds_eq) 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'" show "heap_is_wellformed h'"
using \<open>a_acyclic_heap h'\<close> \<open>a_all_ptrs_in_heap h'\<close> \<open>a_distinct_lists h'\<close> \<open>a_owner_document_valid h'\<close> using \<open>a_acyclic_heap h'\<close> \<open>a_all_ptrs_in_heap h'\<close> \<open>a_distinct_lists h'\<close> \<open>a_owner_document_valid h'\<close>
by(simp add: heap_is_wellformed_def) by(simp add: heap_is_wellformed_def)

View File

@ -58,24 +58,31 @@ register_default_tvars
"('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Node, 'Element) Node" "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Node, 'Element) Node"
type_synonym type_synonym
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element) Object ('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 register_default_tvars
"('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element) Object" "('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element) Object"
type_synonym type_synonym
('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,
= "(('document_ptr, 'shadow_root_ptr) document_ptr + 'object_ptr, 'element_ptr element_ptr + 'character_data_ptr character_data_ptr + 'node_ptr, 'Object, 'Object, 'Node, 'Element) heap
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) RElement_ext + 'Node) heap" = "(('document_ptr, 'shadow_root_ptr) document_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 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" 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 \<Rightarrow> (_) element_ptr fset" definition element_ptr_kinds :: "(_) heap \<Rightarrow> (_) element_ptr fset"
where 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]: lemma element_ptr_kinds_simp [simp]:
"element_ptr_kinds (Heap (fmupd (cast element_ptr) element (the_heap h))) = {|element_ptr|} |\<union>| element_ptr_kinds h" "element_ptr_kinds (Heap (fmupd (cast element_ptr) element (the_heap h))) =
{|element_ptr|} |\<union>| element_ptr_kinds h"
apply(auto simp add: element_ptr_kinds_def)[1] apply(auto simp add: element_ptr_kinds_def)[1]
by force by force
@ -85,7 +92,8 @@ definition element_ptrs :: "(_) heap \<Rightarrow> (_) 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 \<Rightarrow> (_) Element option" 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 \<Rightarrow> (_) Element option"
where 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 \<Rightarrow> Some (RNode.extend (RNode.truncate node) element) | _ \<Rightarrow> 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 \<Rightarrow>
Some (RNode.extend (RNode.truncate node) element) | _ \<Rightarrow> 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 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 \<Rightarrow> (_) Element option" 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 \<Rightarrow> (_) Element option"
@ -298,17 +306,22 @@ lemma known_ptrs_known_ptr:
apply(simp add: a_known_ptrs_def) apply(simp add: a_known_ptrs_def)
using notin_fset by fastforce using notin_fset by fastforce
lemma known_ptrs_preserved: "object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'" lemma known_ptrs_preserved:
"object_ptr_kinds h = object_ptr_kinds h' \<Longrightarrow> a_known_ptrs h = a_known_ptrs h'"
by(auto simp add: a_known_ptrs_def) by(auto simp add: a_known_ptrs_def)
lemma known_ptrs_subset: "object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'" lemma known_ptrs_subset:
"object_ptr_kinds h' |\<subseteq>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def less_eq_fset.rep_eq subsetD) 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 |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow> a_known_ptrs h \<Longrightarrow> a_known_ptrs h'" lemma known_ptrs_new_ptr:
"object_ptr_kinds h' = object_ptr_kinds h |\<union>| {|new_ptr|} \<Longrightarrow> known_ptr new_ptr \<Longrightarrow>
a_known_ptrs h \<Longrightarrow> a_known_ptrs h'"
by(simp add: a_known_ptrs_def) by(simp add: a_known_ptrs_def)
end 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 . 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 .
lemmas known_ptrs_defs = a_known_ptrs_def lemmas known_ptrs_defs = a_known_ptrs_def
lemma known_ptrs_is_l_known_ptrs: "l_known_ptrs known_ptr known_ptrs" lemma known_ptrs_is_l_known_ptrs: "l_known_ptrs known_ptr known_ptrs"
using known_ptrs_known_ptr known_ptrs_preserved known_ptrs_subset known_ptrs_new_ptr l_known_ptrs_def by blast using known_ptrs_known_ptr known_ptrs_preserved known_ptrs_subset known_ptrs_new_ptr l_known_ptrs_def
by blast
end end

View File

@ -79,7 +79,8 @@ adhoc_overloading cast cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^su
consts is_shadow_root_ptr_kind :: 'a consts is_shadow_root_ptr_kind :: 'a
definition is_shadow_root_ptr_kind\<^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 \<Rightarrow> bool" definition is_shadow_root_ptr_kind\<^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 \<Rightarrow> bool"
where where
"is_shadow_root_ptr_kind\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr = (case 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>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr of Some _ \<Rightarrow> True | _ \<Rightarrow> False)" "is_shadow_root_ptr_kind\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr =
(case 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>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr of Some _ \<Rightarrow> True | _ \<Rightarrow> False)"
abbreviation is_shadow_root_ptr_kind\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) object_ptr \<Rightarrow> bool" abbreviation is_shadow_root_ptr_kind\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) object_ptr \<Rightarrow> bool"
where where
@ -107,15 +108,18 @@ abbreviation is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\
Some document_ptr \<Rightarrow> is_shadow_root_ptr\<^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 Some document_ptr \<Rightarrow> is_shadow_root_ptr\<^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
| None \<Rightarrow> False)" | None \<Rightarrow> False)"
adhoc_overloading is_shadow_root_ptr is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r is_shadow_root_ptr\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r is_shadow_root_ptr\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r adhoc_overloading is_shadow_root_ptr is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r is_shadow_root_ptr\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r
is_shadow_root_ptr\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r
lemmas is_shadow_root_ptr_def = is_shadow_root_ptr\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def lemmas is_shadow_root_ptr_def = is_shadow_root_ptr\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
consts is_shadow_root_ptr_ext :: 'a consts is_shadow_root_ptr_ext :: 'a
abbreviation "is_shadow_root_ptr_ext\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> \<not> is_shadow_root_ptr\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" abbreviation "is_shadow_root_ptr_ext\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> \<not> is_shadow_root_ptr\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr"
abbreviation "is_shadow_root_ptr_ext\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> is_shadow_root_ptr_kind ptr \<and> (\<not> is_shadow_root_ptr\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr)" abbreviation "is_shadow_root_ptr_ext\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv>
is_shadow_root_ptr_kind ptr \<and> (\<not> is_shadow_root_ptr\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr)"
abbreviation "is_shadow_root_ptr_ext\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv> is_shadow_root_ptr_kind ptr \<and> (\<not> is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr)" abbreviation "is_shadow_root_ptr_ext\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<equiv>
is_shadow_root_ptr_kind ptr \<and> (\<not> is_shadow_root_ptr\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr)"
adhoc_overloading is_shadow_root_ptr_ext is_shadow_root_ptr_ext\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r is_shadow_root_ptr_ext\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r adhoc_overloading is_shadow_root_ptr_ext is_shadow_root_ptr_ext\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r is_shadow_root_ptr_ext\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r
@ -139,7 +143,8 @@ lemma is_shadow_root_ptr_ref [simp]: "is_shadow_root_ptr (shadow_root_ptr.Ref n)
by(simp add: is_shadow_root_ptr\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def) by(simp add: is_shadow_root_ptr\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def)
lemma shadow_root_ptr_casts_commute [simp]: lemma shadow_root_ptr_casts_commute [simp]:
"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>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr = Some shadow_root_ptr \<longleftrightarrow> cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr = 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>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r document_ptr =
Some shadow_root_ptr \<longleftrightarrow> cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r shadow_root_ptr = document_ptr"
unfolding 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>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def unfolding 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>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def
by(auto split: document_ptr.splits sum.splits) by(auto split: document_ptr.splits sum.splits)
@ -178,7 +183,8 @@ lemma cast\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub
"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>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (document_ptr.Ext (Inr (Inr document_ext_ptr))) = None" "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>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r (document_ptr.Ext (Inr (Inr document_ext_ptr))) = None"
by(simp add: 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>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def) by(simp add: 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>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def)
lemma is_shadow_root_ptr_implies_kind [dest]: "is_shadow_root_ptr\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<Longrightarrow> is_shadow_root_ptr_kind\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr" lemma is_shadow_root_ptr_implies_kind [dest]:
"is_shadow_root_ptr\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr \<Longrightarrow> is_shadow_root_ptr_kind\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr"
by(auto split: option.splits) by(auto split: option.splits)
lemma is_shadow_root_ptr_kind_not_document_ptr [simp]: "\<not>is_shadow_root_ptr_kind (document_ptr.Ref x)" lemma is_shadow_root_ptr_kind_not_document_ptr [simp]: "\<not>is_shadow_root_ptr_kind (document_ptr.Ref x)"