Updated files to match development version of AFP.

This commit is contained in:
Achim D. Brucker 2020-05-23 11:53:43 +01:00
parent d966581670
commit fb49067a8d
38 changed files with 2459 additions and 578 deletions

View File

@ -3,14 +3,18 @@ chapter AFP
session "Core_DOM-devel" (AFP) = "HOL-Library" +
options [timeout = 1200]
directories
"classes"
"monads"
"pointers"
"preliminaries"
"tests"
"common"
"common/classes"
"common/monads"
"common/pointers"
"common/preliminaries"
"common/tests"
"standard"
"standard/classes"
"standard/pointers"
theories
Core_DOM
Core_DOM_Tests
document_files
document_files (in "document")
"root.tex"
"root.bib"

View File

@ -620,6 +620,62 @@ lemma set_child_nodes_get_child_nodes_different_pointers:
apply(rule is_element_ptr_kind_obtains)
apply(auto)
done
lemma set_child_nodes_element_ok [simp]:
assumes "known_ptr ptr"
assumes "type_wf h"
assumes "ptr |\<in>| object_ptr_kinds h"
assumes "is_element_ptr_kind ptr"
shows "h \<turnstile> ok (set_child_nodes ptr children)"
proof -
have "is_element_ptr ptr"
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)
then show ?thesis
using assms
apply(auto simp add: set_child_nodes_def a_set_child_nodes_tups_def set_child_nodes\<^sub>e\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def split: option.splits)[1]
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
lemma set_child_nodes_document1_ok [simp]:
assumes "known_ptr ptr"
assumes "type_wf h"
assumes "ptr |\<in>| object_ptr_kinds h"
assumes "is_document_ptr_kind ptr"
assumes "children = []"
shows "h \<turnstile> ok (set_child_nodes ptr children)"
proof -
have "is_document_ptr ptr"
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)
then show ?thesis
using assms
apply(auto simp add: set_child_nodes_def a_set_child_nodes_tups_def set_child_nodes\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def split: option.splits)[1]
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
lemma set_child_nodes_document2_ok [simp]:
assumes "known_ptr ptr"
assumes "type_wf h"
assumes "ptr |\<in>| object_ptr_kinds h"
assumes "is_document_ptr_kind ptr"
assumes "children = [child]"
assumes "is_element_ptr_kind child"
shows "h \<turnstile> ok (set_child_nodes ptr children)"
proof -
have "is_document_ptr ptr"
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)
then show ?thesis
using assms
apply(auto simp add: set_child_nodes_def a_set_child_nodes_tups_def set_child_nodes\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def)
apply(split invoke_splits, rule conjI)+
apply(auto simp add: is_element_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def set_child_nodes\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def split: option.splits)[1]
apply(auto simp add: is_element_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def set_child_nodes\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def split: option.splits)[1]
apply (simp add: local.type_wf_impl put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ok)
apply(auto simp add: is_element_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def set_child_nodes\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def split: option.splits)[1]
by(auto simp add: is_element_ptr_kind\<^sub>n\<^sub>o\<^sub>d\<^sub>e\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def set_child_nodes\<^sub>d\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r_def split: option.splits)[1]
qed
end
locale l_set_child_nodes_get_child_nodes = l_get_child_nodes + l_set_child_nodes +
@ -2358,6 +2414,15 @@ proof -
using assms(1) get_child_nodes_ptr_in_heap by blast
qed
lemma remove_child_child_in_heap:
assumes "h \<turnstile> remove_child ptr' child \<rightarrow>\<^sub>h h'"
shows "child |\<in>| node_ptr_kinds h"
using assms
apply(auto simp add: remove_child_def elim!: bind_returns_heap_E bind_returns_heap_E2[rotated, OF get_child_nodes_pure, rotated] split: if_splits)[1]
by (meson is_OK_returns_result_I local.get_owner_document_ptr_in_heap node_ptr_kinds_commutes)
lemma remove_child_in_disconnected_nodes:
(* assumes "known_ptrs h" *)
assumes "h \<turnstile> remove_child ptr child \<rightarrow>\<^sub>h h'"
@ -2490,6 +2555,7 @@ locale l_remove_child = l_type_wf + l_known_ptrs + l_remove_child_defs + l_get_o
\<Longrightarrow> h' \<turnstile> get_disconnected_nodes owner_document \<rightarrow>\<^sub>r disc_nodes
\<Longrightarrow> child \<in> set disc_nodes"
assumes remove_child_ptr_in_heap: "h \<turnstile> ok (remove_child ptr child) \<Longrightarrow> ptr |\<in>| object_ptr_kinds h"
assumes remove_child_child_in_heap: "h \<turnstile> remove_child ptr' child \<rightarrow>\<^sub>h h' \<Longrightarrow> child |\<in>| node_ptr_kinds h"
assumes remove_child_children_subset:
"known_ptrs h \<Longrightarrow> type_wf h \<Longrightarrow> h \<turnstile> remove_child parent child \<rightarrow>\<^sub>h h'
\<Longrightarrow> h \<turnstile> get_child_nodes ptr \<rightarrow>\<^sub>r children
@ -2535,6 +2601,7 @@ lemma remove_child_is_l_remove_child [instances]:
using remove_child_types_preserved apply(blast)
using remove_child_in_disconnected_nodes apply(blast)
using remove_child_ptr_in_heap apply(blast)
using remove_child_child_in_heap apply(blast)
using remove_child_children_subset apply(blast)
done
@ -2975,8 +3042,8 @@ lemma insert_before_list_in_set: "x \<in> set (insert_before_list v ref xs) \<lo
by(auto)
lemma insert_before_list_distinct: "x \<notin> set xs \<Longrightarrow> distinct xs \<Longrightarrow> distinct (insert_before_list x ref xs)"
by (induct x ref xs rule: insert_before_list.induct)
(auto simp add: insert_before_list_in_set)
apply(induct x ref xs rule: insert_before_list.induct)
by(auto simp add: insert_before_list_in_set)
lemma insert_before_list_subset: "set xs \<subseteq> set (insert_before_list x ref xs)"
apply(induct x ref xs rule: insert_before_list.induct)
@ -3012,6 +3079,13 @@ proof -
unfolding insert_before_def by auto
qed
lemma insert_before_ptr_in_heap:
assumes "h \<turnstile> ok (insert_before ptr node reference_child)"
shows "ptr |\<in>| object_ptr_kinds h"
using assms
apply(auto simp add: insert_before_def elim!: bind_is_OK_E)[1]
by (metis (mono_tags, lifting) ensure_pre_insertion_validity_pure is_OK_returns_result_I local.get_owner_document_ptr_in_heap next_sibling_pure pure_returns_heap_eq return_returns_heap)
lemma insert_before_child_in_heap:
assumes "h \<turnstile> ok (insert_before ptr node reference_child)"
shows "node |\<in>| node_ptr_kinds h"
@ -3166,20 +3240,76 @@ 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 =
l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs +
l_create_element_defs +
l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_tag_type set_tag_type_locs +
l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs +
l_set_tag_type type_wf set_tag_type set_tag_type_locs +
l_create_element_defs create_element +
l_known_ptr known_ptr
for 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 set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and set_tag_type :: "(_) element_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_tag_type_locs :: "(_) element_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and type_wf :: "(_) heap \<Rightarrow> bool"
and create_element :: "(_) document_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) element_ptr) prog"
and known_ptr :: "(_) object_ptr \<Rightarrow> bool" +
assumes known_ptr_impl: "known_ptr = a_known_ptr"
assumes create_element_impl: "create_element = a_create_element"
begin
lemmas create_element_def = a_create_element_def[folded create_element_impl]
lemma create_element_document_in_heap:
assumes "h \<turnstile> ok (create_element document_ptr tag)"
shows "document_ptr |\<in>| document_ptr_kinds h"
proof -
obtain h' where "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>h h'"
using assms(1)
by auto
then
obtain new_element_ptr h2 h3 disc_nodes_h3 where
new_element_ptr: "h \<turnstile> new_element \<rightarrow>\<^sub>r new_element_ptr" and
h2: "h \<turnstile> new_element \<rightarrow>\<^sub>h h2" and
h3: "h2 \<turnstile> set_tag_type new_element_ptr tag \<rightarrow>\<^sub>h h3" and
disc_nodes_h3: "h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (cast new_element_ptr # disc_nodes_h3) \<rightarrow>\<^sub>h h'"
by(auto simp add: create_element_def
elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] )
have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\<union>| {|cast new_element_ptr|}"
using new_element_new_ptr h2 new_element_ptr by blast
moreover have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_tag_type_writes h3])
using set_tag_type_pointers_preserved
by (auto simp add: reflp_def transp_def)
moreover have "document_ptr |\<in>| document_ptr_kinds h3"
by (meson disc_nodes_h3 is_OK_returns_result_I local.get_disconnected_nodes_ptr_in_heap)
ultimately show ?thesis
by (auto simp add: document_ptr_kinds_def)
qed
lemma create_element_known_ptr:
assumes "h \<turnstile> create_element document_ptr tag \<rightarrow>\<^sub>r new_element_ptr"
shows "known_ptr (cast new_element_ptr)"
proof -
have "is_element_ptr new_element_ptr"
using assms
apply(auto simp add: create_element_def elim!: bind_returns_result_E)[1]
using new_element_is_element_ptr
by blast
then show ?thesis
by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs)
qed
end
locale l_create_element = l_create_element_defs
interpretation
i_create_element?: l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs
set_disconnected_nodes set_disconnected_nodes_locs set_tag_type
set_tag_type_locs create_element
by unfold_locales (simp add: create_element_def)
i_create_element?: l_create_element\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_tag_type set_tag_type_locs type_wf create_element known_ptr
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]
@ -3217,20 +3347,76 @@ global_interpretation l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^
.
locale l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M =
l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs +
l_create_character_data_defs +
l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_defs set_val set_val_locs get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs +
l_get_disconnected_nodes type_wf get_disconnected_nodes get_disconnected_nodes_locs +
l_set_val type_wf set_val set_val_locs +
l_create_character_data_defs create_character_data +
l_known_ptr known_ptr
for 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 set_disconnected_nodes :: "(_) document_ptr \<Rightarrow> (_) node_ptr list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_disconnected_nodes_locs :: "(_) document_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and set_val :: "(_) character_data_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, unit) prog"
and set_val_locs :: "(_) character_data_ptr \<Rightarrow> ((_) heap, exception, unit) prog set"
and type_wf :: "(_) heap \<Rightarrow> bool"
and create_character_data :: "(_) document_ptr \<Rightarrow> char list \<Rightarrow> ((_) heap, exception, (_) character_data_ptr) prog"
and known_ptr :: "(_) object_ptr \<Rightarrow> bool" +
assumes known_ptr_impl: "known_ptr = a_known_ptr"
assumes create_character_data_impl: "create_character_data = a_create_character_data"
begin
lemmas create_character_data_def = a_create_character_data_def[folded create_character_data_impl]
lemma create_character_data_document_in_heap:
assumes "h \<turnstile> ok (create_character_data document_ptr text)"
shows "document_ptr |\<in>| document_ptr_kinds h"
proof -
obtain h' where "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>h h'"
using assms(1)
by auto
then
obtain new_character_data_ptr h2 h3 disc_nodes_h3 where
new_character_data_ptr: "h \<turnstile> new_character_data \<rightarrow>\<^sub>r new_character_data_ptr" and
h2: "h \<turnstile> new_character_data \<rightarrow>\<^sub>h h2" and
h3: "h2 \<turnstile> set_val new_character_data_ptr text \<rightarrow>\<^sub>h h3" and
disc_nodes_h3: "h3 \<turnstile> get_disconnected_nodes document_ptr \<rightarrow>\<^sub>r disc_nodes_h3" and
h': "h3 \<turnstile> set_disconnected_nodes document_ptr (cast new_character_data_ptr # disc_nodes_h3) \<rightarrow>\<^sub>h h'"
by(auto simp add: create_character_data_def
elim!: bind_returns_heap_E
bind_returns_heap_E2[rotated, OF get_disconnected_nodes_pure, rotated] )
have object_ptr_kinds_eq_h: "object_ptr_kinds h2 = object_ptr_kinds h |\<union>| {|cast new_character_data_ptr|}"
using new_character_data_new_ptr h2 new_character_data_ptr by blast
moreover have object_ptr_kinds_eq_h2: "object_ptr_kinds h3 = object_ptr_kinds h2"
apply(rule writes_small_big[where P="\<lambda>h h'. object_ptr_kinds h' = object_ptr_kinds h", OF set_val_writes h3])
using set_val_pointers_preserved
by (auto simp add: reflp_def transp_def)
moreover have "document_ptr |\<in>| document_ptr_kinds h3"
by (meson disc_nodes_h3 is_OK_returns_result_I local.get_disconnected_nodes_ptr_in_heap)
ultimately show ?thesis
by (auto simp add: document_ptr_kinds_def)
qed
lemma create_character_data_known_ptr:
assumes "h \<turnstile> create_character_data document_ptr text \<rightarrow>\<^sub>r new_character_data_ptr"
shows "known_ptr (cast new_character_data_ptr)"
proof -
have "is_character_data_ptr new_character_data_ptr"
using assms
apply(auto simp add: create_character_data_def elim!: bind_returns_result_E)[1]
using new_character_data_is_character_data_ptr
by blast
then show ?thesis
by(auto simp add: known_ptr_impl DocumentClass.known_ptr_defs CharacterDataClass.known_ptr_defs ElementClass.known_ptr_defs)
qed
end
locale l_create_character_data = l_create_character_data_defs
interpretation
i_create_character_data?: l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M set_val set_val_locs get_disconnected_nodes
get_disconnected_nodes_locs set_disconnected_nodes
set_disconnected_nodes_locs create_character_data
by unfold_locales (simp add: create_character_data_def)
i_create_character_data?: l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M get_disconnected_nodes get_disconnected_nodes_locs set_disconnected_nodes set_disconnected_nodes_locs set_val set_val_locs type_wf create_character_data known_ptr
by(auto simp add: l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_def l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms_def create_character_data_def instances)
declare l_create_character_data\<^sub>C\<^sub>o\<^sub>r\<^sub>e\<^sub>_\<^sub>D\<^sub>O\<^sub>M_axioms [instances]
@ -3486,6 +3672,30 @@ lemma get_element_by_id_result_in_tree_order:
intro!: map_filter_M_pure map_M_pure_I bind_pure_I
split: option.splits list.splits if_splits)
lemma get_elements_by_class_name_result_in_tree_order:
assumes "h \<turnstile> get_elements_by_class_name ptr name \<rightarrow>\<^sub>r results"
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to"
assumes "element_ptr \<in> set results"
shows "cast element_ptr \<in> set to"
using assms
by(auto simp add: get_elements_by_class_name_def first_in_tree_order_def
elim!: map_filter_M_pure_E[where y=element_ptr] bind_returns_result_E2
dest!: bind_returns_result_E3[rotated, OF assms(2), rotated]
intro!: map_filter_M_pure map_M_pure_I bind_pure_I
split: option.splits list.splits if_splits)
lemma get_elements_by_tag_name_result_in_tree_order:
assumes "h \<turnstile> get_elements_by_tag_name ptr name \<rightarrow>\<^sub>r results"
assumes "h \<turnstile> to_tree_order ptr \<rightarrow>\<^sub>r to"
assumes "element_ptr \<in> set results"
shows "cast element_ptr \<in> set to"
using assms
by(auto simp add: get_elements_by_tag_name_def first_in_tree_order_def
elim!: map_filter_M_pure_E[where y=element_ptr] bind_returns_result_E2
dest!: bind_returns_result_E3[rotated, OF assms(2), rotated]
intro!: map_filter_M_pure map_M_pure_I bind_pure_I
split: option.splits list.splits if_splits)
lemma get_elements_by_tag_name_pure [simp]: "pure (get_elements_by_tag_name ptr tag_name) h"
by(auto simp add: get_elements_by_tag_name_def
intro!: bind_pure_I map_filter_M_pure

View File

@ -65,6 +65,7 @@ type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'docume
'Object, 'CharacterData option RCharacterData_ext + 'Node, 'Element) heap"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData) heap"
type_synonym heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l = "(unit, unit, unit, unit, unit, unit, unit, unit, unit, unit) heap"
definition character_data_ptr_kinds :: "(_) heap \<Rightarrow> (_) character_data_ptr fset"
@ -139,8 +140,8 @@ begin
definition a_type_wf :: "(_) heap \<Rightarrow> bool"
where
"a_type_wf h = (ElementClass.type_wf h
\<and> (\<forall>character_data_ptr. character_data_ptr |\<in>| character_data_ptr_kinds h
\<longrightarrow> 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))"
\<and> (\<forall>character_data_ptr \<in> fset (character_data_ptr_kinds 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 \<noteq> None))"
end
global_interpretation 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 defines type_wf = a_type_wf .
lemmas type_wf_defs = a_type_wf_def
@ -163,8 +164,7 @@ 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"
using l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_axioms assms
apply(simp add: type_wf_defs get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def l_type_wf\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def)
by (metis NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf bind_eq_None_conv character_data_ptr_kinds_commutes
l_type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def local.l_type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e_axioms option.distinct(1))
by (metis assms bind.bind_lzero character_data_ptr_kinds_commutes fmember.rep_eq local.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf option.exhaust option.simps(3))
end
global_interpretation l_get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_lemmas type_wf
@ -325,23 +325,26 @@ locale l_known_ptrs\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^su
begin
definition a_known_ptrs :: "(_) heap \<Rightarrow> bool"
where
"a_known_ptrs h = (\<forall>ptr. ptr |\<in>| object_ptr_kinds h \<longrightarrow> known_ptr ptr)"
"a_known_ptrs h = (\<forall>ptr \<in> fset (object_ptr_kinds h). known_ptr ptr)"
lemma known_ptrs_known_ptr: "a_known_ptrs h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptr ptr"
by(simp add: a_known_ptrs_def)
apply(simp add: a_known_ptrs_def)
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'"
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'"
by(auto simp add: a_known_ptrs_def)
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'"
by(simp add: a_known_ptrs_def)
end
global_interpretation l_known_ptrs\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a known_ptr defines known_ptrs = a_known_ptrs .
lemmas known_ptrs_defs = a_known_ptrs_def
lemma known_ptrs_is_l_known_ptrs: "l_known_ptrs known_ptr known_ptrs"
using known_ptrs_known_ptr known_ptrs_preserved l_known_ptrs_def known_ptrs_subset
using known_ptrs_known_ptr known_ptrs_preserved known_ptrs_subset known_ptrs_new_ptr l_known_ptrs_def
by blast
end

View File

@ -65,6 +65,7 @@ type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'docume
register_default_tvars
"('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'Document) heap"
type_synonym heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l = "(unit, unit, unit, unit, unit, unit, unit, unit, unit, unit, unit) heap"
definition document_ptr_kinds :: "(_) heap \<Rightarrow> (_) document_ptr fset"
@ -115,7 +116,7 @@ begin
definition a_type_wf :: "(_) heap \<Rightarrow> bool"
where
"a_type_wf h = (CharacterDataClass.type_wf h \<and>
(\<forall>document_ptr. document_ptr |\<in>| document_ptr_kinds h \<longrightarrow> get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h \<noteq> None))"
(\<forall>document_ptr \<in> fset (document_ptr_kinds h). get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t document_ptr h \<noteq> None))"
end
global_interpretation l_type_wf_def\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t defines type_wf = a_type_wf .
lemmas type_wf_defs = a_type_wf_def
@ -135,8 +136,7 @@ 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"
using l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_axioms assms
apply(simp add: type_wf_defs get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
by (metis bind_eq_None_conv document_ptr_kinds_commutes local.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf
option.distinct(1))
by (metis document_ptr_kinds_commutes fmember.rep_eq is_none_bind is_none_simps(1) is_none_simps(2) local.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf)
end
global_interpretation l_get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_lemmas type_wf by unfold_locales
@ -228,7 +228,7 @@ abbreviation
definition new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t :: "(_)heap \<Rightarrow> ((_) document_ptr \<times> (_) heap)"
where
"new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t h =
(let new_document_ptr = document_ptr.Ref (Suc (fMax (document_ptr.the_ref |`| (document_ptrs h))))
(let new_document_ptr = document_ptr.Ref (Suc (fMax (finsert 0 (document_ptr.the_ref |`| (document_ptrs h)))))
in
(new_document_ptr, put new_document_ptr (create_document_obj '''' None []) h))"
@ -315,22 +315,26 @@ locale l_known_ptrs\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^su
begin
definition a_known_ptrs :: "(_) heap \<Rightarrow> bool"
where
"a_known_ptrs h = (\<forall>ptr. ptr |\<in>| object_ptr_kinds h \<longrightarrow> known_ptr ptr)"
"a_known_ptrs h = (\<forall>ptr \<in> fset (object_ptr_kinds h). known_ptr ptr)"
lemma known_ptrs_known_ptr: "a_known_ptrs h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptr ptr"
by(simp add: a_known_ptrs_def)
apply(simp add: a_known_ptrs_def)
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'"
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'"
by(auto simp add: a_known_ptrs_def)
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'"
by(simp add: a_known_ptrs_def)
end
global_interpretation l_known_ptrs\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t known_ptr defines known_ptrs = a_known_ptrs .
lemmas known_ptrs_defs = a_known_ptrs_def
lemma known_ptrs_is_l_known_ptrs [instances]: "l_known_ptrs known_ptr known_ptrs"
using known_ptrs_known_ptr known_ptrs_preserved l_known_ptrs_def known_ptrs_subset by blast
using known_ptrs_known_ptr known_ptrs_preserved l_known_ptrs_def known_ptrs_subset known_ptrs_new_ptr
by blast
end

View File

@ -51,6 +51,7 @@ type_synonym ('object_ptr, 'node_ptr, 'Object, 'Node) heap
= "('node_ptr node_ptr + 'object_ptr, 'Node RNode_ext + 'Object) heap"
register_default_tvars
"('object_ptr, 'node_ptr, 'Object, 'Node) heap"
type_synonym heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l = "(unit, unit, unit, unit) heap"
definition node_ptr_kinds :: "(_) heap \<Rightarrow> (_) node_ptr fset"
@ -89,8 +90,7 @@ begin
definition a_type_wf :: "(_) heap \<Rightarrow> bool"
where
"a_type_wf h = (ObjectClass.type_wf h
\<and> (\<forall>node_ptr. node_ptr |\<in>| node_ptr_kinds h
\<longrightarrow> get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr h \<noteq> None))"
\<and> (\<forall>node_ptr \<in> fset( node_ptr_kinds h). get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr h \<noteq> None))"
end
global_interpretation l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e defines type_wf = a_type_wf .
lemmas type_wf_defs = a_type_wf_def
@ -110,9 +110,8 @@ lemma get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf:
shows "node_ptr |\<in>| node_ptr_kinds h \<longleftrightarrow> get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr h \<noteq> None"
using l_type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e_axioms assms
apply(simp add: type_wf_defs get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def l_type_wf\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)
by (metis (mono_tags, lifting) bind_eq_None_conv ffmember_filter fimage_eqI
is_node_ptr_kind_cast get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf local.l_type_wf\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_axioms
node_ptr_casts_commute2 node_ptr_kinds_def option.sel option.simps(3))
by (metis bind_eq_None_conv ffmember_filter fimage_eqI fmember.rep_eq is_node_ptr_kind_cast
get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf node_ptr_casts_commute2 node_ptr_kinds_def option.sel option.simps(3))
end
global_interpretation l_get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_lemmas type_wf
@ -177,21 +176,24 @@ locale l_known_ptrs\<^sub>N\<^sub>o\<^sub>d\<^sub>e = l_known_ptr known_ptr for
begin
definition a_known_ptrs :: "(_) heap \<Rightarrow> bool"
where
"a_known_ptrs h = (\<forall>ptr. ptr |\<in>| object_ptr_kinds h \<longrightarrow> known_ptr ptr)"
"a_known_ptrs h = (\<forall>ptr \<in> fset (object_ptr_kinds h). known_ptr ptr)"
lemma known_ptrs_known_ptr: "a_known_ptrs h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptr ptr"
by(simp add: a_known_ptrs_def)
apply(simp add: a_known_ptrs_def)
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'"
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'"
by(auto simp add: a_known_ptrs_def)
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'"
by(simp add: a_known_ptrs_def)
end
global_interpretation l_known_ptrs\<^sub>N\<^sub>o\<^sub>d\<^sub>e known_ptr defines known_ptrs = a_known_ptrs .
lemmas known_ptrs_defs = a_known_ptrs_def
lemma known_ptrs_is_l_known_ptrs: "l_known_ptrs known_ptr known_ptrs"
using known_ptrs_known_ptr known_ptrs_preserved l_known_ptrs_def known_ptrs_subset by blast
using known_ptrs_known_ptr known_ptrs_preserved l_known_ptrs_def known_ptrs_subset known_ptrs_new_ptr
by blast
lemma get_node_ptr_simp1 [simp]: "get\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr (put\<^sub>N\<^sub>o\<^sub>d\<^sub>e node_ptr node h) = Some node"
by(auto simp add: get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def)

View File

@ -45,6 +45,7 @@ register_default_tvars "'Object Object"
datatype ('object_ptr, 'Object) heap = Heap (the_heap: "((_) object_ptr, (_) Object) fmap")
register_default_tvars "('object_ptr, 'Object) heap"
type_synonym heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l = "(unit, unit) heap"
definition object_ptr_kinds :: "(_) heap \<Rightarrow> (_) object_ptr fset"
where
@ -128,27 +129,32 @@ locale l_known_ptrs = l_known_ptr known_ptr for known_ptr :: "(_) object_ptr \<R
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_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"
begin
definition a_known_ptrs :: "(_) heap \<Rightarrow> bool"
where
"a_known_ptrs h = (\<forall>ptr. ptr |\<in>| object_ptr_kinds h \<longrightarrow> known_ptr ptr)"
"a_known_ptrs h = (\<forall>ptr \<in> fset (object_ptr_kinds h). known_ptr ptr)"
lemma known_ptrs_known_ptr:
"a_known_ptrs h \<Longrightarrow> ptr |\<in>| object_ptr_kinds h \<Longrightarrow> known_ptr ptr"
by(simp add: a_known_ptrs_def)
apply(simp add: a_known_ptrs_def)
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'"
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'"
by(auto simp add: a_known_ptrs_def)
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'"
by(simp add: a_known_ptrs_def)
end
global_interpretation l_known_ptrs\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t known_ptr defines known_ptrs = a_known_ptrs .
lemmas known_ptrs_defs = a_known_ptrs_def
lemma known_ptrs_is_l_known_ptrs: "l_known_ptrs known_ptr known_ptrs"
using known_ptrs_known_ptr known_ptrs_preserved l_known_ptrs_def known_ptrs_subset by blast
using known_ptrs_known_ptr known_ptrs_preserved l_known_ptrs_def known_ptrs_subset known_ptrs_new_ptr
by blast
lemma get_object_ptr_simp1 [simp]: "get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr object h) = Some object"
@ -188,4 +194,24 @@ lemma delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_ok:
using assms
by(auto simp add: delete\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def object_ptr_kinds_def split: if_splits)
subsection \<open>Code Generator Setup\<close>
definition "create_heap xs = Heap (fmap_of_list xs)"
code_datatype ObjectClass.heap.Heap create_heap
lemma object_ptr_kinds_code3 [code]:
"fmlookup (the_heap (create_heap xs)) x = map_of xs x"
by(auto simp add: create_heap_def fmlookup_of_list)
lemma object_ptr_kinds_code4 [code]:
"the_heap (create_heap xs) = fmap_of_list xs"
by(simp add: create_heap_def)
lemma object_ptr_kinds_code5 [code]:
"the_heap (Heap x) = x"
by simp
end

View File

@ -58,9 +58,8 @@ lemma character_data_ptr_kinds_M_reads:
"reads (\<Union>node_ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t node_ptr RObject.nothing)}) character_data_ptr_kinds_M h h'"
using node_ptr_kinds_M_reads
apply (simp add: reads_def node_ptr_kinds_M_defs character_data_ptr_kinds_M_defs
character_data_ptr_kinds_def preserved_def cong del: image_cong_simp)
apply (metis (mono_tags, hide_lams) node_ptr_kinds_small old.unit.exhaust preserved_def)
done
character_data_ptr_kinds_def preserved_def)
by (smt node_ptr_kinds_small preserved_def unit_all_impI)
global_interpretation l_dummy defines get_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a = "l_get_M.a_get_M get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a" .
lemma get_M_is_l_get_M: "l_get_M get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a type_wf character_data_ptr_kinds"
@ -308,8 +307,9 @@ lemma type_wf_put_ptr_not_in_heap_E:
assumes "ptr |\<notin>| object_ptr_kinds h"
shows "type_wf h"
using assms
by(auto simp add: type_wf_defs elim!: ElementMonad.type_wf_put_ptr_not_in_heap_E
apply(auto simp add: type_wf_defs elim!: ElementMonad.type_wf_put_ptr_not_in_heap_E
split: option.splits if_splits)
using assms(2) node_ptr_kinds_commutes by blast
lemma type_wf_put_ptr_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
@ -319,14 +319,7 @@ lemma type_wf_put_ptr_in_heap_E:
shows "type_wf h"
using assms
apply(auto simp add: type_wf_defs split: option.splits if_splits)[1]
apply(case_tac "x2 = cast character_data_ptr")
apply(auto)[1]
apply(drule_tac x=character_data_ptr in allE)
apply(simp)
apply (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 option.exhaust_sel)
by(blast)
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>
@ -365,8 +358,9 @@ lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_tag_type_typ
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs split: option.splits)[1]
using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast
apply (metis (mono_tags, lifting) bind_eq_Some_conv get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def option.exhaust_sel)
by (metis (no_types, lifting) option.exhaust_sel )
apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
apply (metis finite_set_in)
done
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_child_nodes_type_wf_preserved [simp]:
@ -391,9 +385,9 @@ lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_child_nodes_
NodeClass.type_wf_defs ElementMonad.get_M_defs
split: option.splits)[1]
using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast
apply (metis (mono_tags, lifting) ElementMonad.a_get_M_def bind_eq_Some_conv error_returns_result
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get_heap_returns_result option.exhaust_sel option.simps(4))
by (metis (no_types, lifting) option.exhaust_sel)
apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
apply (metis finite_set_in)
done
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_attrs_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr attrs_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
@ -413,8 +407,9 @@ lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_attrs_type_w
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs split: option.splits)[1]
using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast
apply (metis (mono_tags, lifting) ElementMonad.a_get_M_def bind_eq_Some_conv error_returns_result get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get_heap_returns_result option.exhaust_sel option.simps(4))
by (metis (no_types, lifting) option.exhaust_sel)
apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
apply (metis finite_set_in)
done
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_shadow_root_opt_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr shadow_root_opt_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
@ -434,8 +429,9 @@ lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_shadow_root_
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs split: option.splits)[1]
using ObjectMonad.type_wf_put_ptr_in_heap_E ObjectMonad.type_wf_put_ptr_not_in_heap_E apply blast
apply (metis (mono_tags, lifting) bind_eq_Some_conv get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
by (metis (no_types, lifting) option.exhaust_sel)
apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
apply (metis finite_set_in)
done
lemma new_character_data_type_wf_preserved [simp]:
@ -470,8 +466,9 @@ lemma put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^su
NodeClass.type_wf_defs CharacterDataMonad.get_M_defs
ObjectClass.a_type_wf_def
split: option.splits)[1]
apply (metis (no_types, lifting) bind_eq_Some_conv 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)
by metis
apply (metis (no_types, lifting) bind_eq_Some_conv finite_set_in get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def)
apply (metis finite_set_in)
done
lemma character_data_ptr_kinds_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
@ -529,8 +526,6 @@ 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
l_type_wf_def\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a.a_type_wf_def)[1]
using type_wf_drop
by (metis (no_types, lifting) ElementClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf character_data_ptr_kinds_commutes
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 heap.sel
node_ptr_kinds_commutes)
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,8 +322,7 @@ lemma type_wf_put_ptr_in_heap_E:
using assms
apply(auto simp add: type_wf_defs elim!: CharacterDataMonad.type_wf_put_ptr_in_heap_E
split: option.splits if_splits)[1]
by (metis (no_types, lifting) CharacterDataClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf bind.bind_lunit get\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
is_document_kind_def option.collapse)
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)
@ -361,9 +360,8 @@ lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_tag_type_typ
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs NodeClass.type_wf_defs
ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply (metis (no_types, lifting) Option.bind_cong bind_rzero
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def option.distinct(1))
by metis
apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse option.distinct(1) option.simps(3))
by (metis fmember.rep_eq)
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_child_nodes_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr child_nodes_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
@ -378,8 +376,8 @@ lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_child_nodes_
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply (metis (no_types, lifting) Option.bind_cong bind_rzero get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def option.distinct(1))
by metis
apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse option.distinct(1) option.simps(3))
by (metis fmember.rep_eq)
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_attrs_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr attrs_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
@ -394,8 +392,8 @@ lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_attrs_type_w
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply (metis (no_types, lifting) Option.bind_cong bind_rzero get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def option.distinct(1))
by metis
apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse option.distinct(1) option.simps(3))
by (metis fmember.rep_eq)
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_shadow_root_opt_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr shadow_root_opt_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
@ -410,8 +408,8 @@ lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_shadow_root_
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply (metis (no_types, lifting) Option.bind_cong bind_rzero get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def option.distinct(1))
by metis
apply (metis NodeClass.a_type_wf_def NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf ObjectClass.a_type_wf_def bind.bind_lzero finite_set_in get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf_def\<^sub>N\<^sub>o\<^sub>d\<^sub>e.a_type_wf_def option.collapse option.distinct(1) option.simps(3))
by (metis fmember.rep_eq)
lemma new_character_data_type_wf_preserved [simp]:
"h \<turnstile> new_character_data \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
@ -440,13 +438,11 @@ lemma put_M\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^su
intro!: type_wf_put_I CharacterDataMonad.type_wf_put_I ElementMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I)[1]
apply(auto simp add: is_node_kind_def type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
NodeClass.type_wf_defs CharacterDataMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply (metis (no_types, lifting) CharacterDataMonad.a_get_M_def bind_eq_None_conv
error_returns_result 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_heap_returns_result option.exhaust_sel
option.simps(4))
by (metis (no_types, lifting) CharacterDataMonad.a_get_M_def error_returns_result
get_heap_returns_result option.exhaust_sel option.simps(4))
apply (metis bind.bind_lzero finite_set_in get\<^sub>C\<^sub>h\<^sub>a\<^sub>r\<^sub>a\<^sub>c\<^sub>t\<^sub>e\<^sub>r\<^sub>D\<^sub>a\<^sub>t\<^sub>a_def option.distinct(1) option.exhaust_sel)
by (metis finite_set_in)
lemma new_document_type_wf_preserved [simp]: "h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
apply(auto simp add: new_document_def new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def put\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
@ -460,7 +456,10 @@ lemma new_document_type_wf_preserved [simp]: "h \<turnstile> new_document \<righ
apply(auto simp add: type_wf_defs ElementClass.type_wf_defs CharacterDataClass.type_wf_defs
NodeClass.type_wf_defs ObjectClass.type_wf_defs is_document_kind_def
split: option.splits)[1]
by (meson new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def new\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_not_in_heap)
using document_ptrs_def apply fastforce
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)
done
locale l_new_document = l_type_wf +
assumes new_document_types_preserved: "h \<turnstile> new_document \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
@ -500,7 +499,7 @@ lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_doct
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
apply(auto simp add: get_M_defs)
by (metis (no_types, lifting) error_returns_result option.exhaust_sel option.simps(4))
by (metis (mono_tags) error_returns_result finite_set_in option.exhaust_sel option.simps(4))
lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_document_element_type_wf_preserved [simp]:
"h \<turnstile> put_M document_ptr document_element_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
@ -519,7 +518,7 @@ lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_docu
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs
split: option.splits)[1]
by (metis)
by (metis finite_set_in)
lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_disconnected_nodes_type_wf_preserved [simp]:
"h \<turnstile> put_M document_ptr disconnected_nodes_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
@ -538,7 +537,7 @@ lemma put_M\<^sub>D\<^sub>o\<^sub>c\<^sub>u\<^sub>m\<^sub>e\<^sub>n\<^sub>t_disc
apply(auto simp add: is_document_kind_def get_M_defs type_wf_defs ElementClass.type_wf_defs
NodeClass.type_wf_defs ElementMonad.get_M_defs ObjectClass.type_wf_defs
CharacterDataClass.type_wf_defs split: option.splits)[1]
by (metis)
by (metis finite_set_in)
lemma document_ptr_kinds_small:
assumes "\<And>object_ptr. preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing) h h'"
@ -600,6 +599,5 @@ lemma type_wf_drop: "type_wf h \<Longrightarrow> type_wf (Heap (fmdrop ptr (the_
apply(auto simp add: type_wf_defs)[1]
using type_wf_drop
apply blast
by (metis (mono_tags, lifting) comp_apply document_ptr_kinds_commutes ffmember_filter fmdom_filter
fmfilter_alt_defs(1) 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 object_ptr_kinds_def)
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

View File

@ -32,7 +32,7 @@ text\<open>In this theory, we introduce the monadic method setup for the Element
theory ElementMonad
imports
NodeMonad
"../classes/ElementClass"
"ElementClass"
begin
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
@ -299,8 +299,9 @@ lemma type_wf_put_ptr_not_in_heap_E:
assumes "ptr |\<notin>| object_ptr_kinds h"
shows "type_wf h"
using assms
by(auto simp add: type_wf_defs elim!: NodeMonad.type_wf_put_ptr_not_in_heap_E
split: option.splits if_splits)
apply(auto simp add: type_wf_defs elim!: NodeMonad.type_wf_put_ptr_not_in_heap_E
split: option.splits if_splits)[1]
using assms(2) node_ptr_kinds_commutes by blast
lemma type_wf_put_ptr_in_heap_E:
assumes "type_wf (put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr obj h)"
@ -310,12 +311,9 @@ lemma type_wf_put_ptr_in_heap_E:
shows "type_wf h"
using assms
apply(auto simp add: type_wf_defs split: option.splits if_splits)[1]
apply(case_tac "x2 = cast element_ptr")
apply(drule_tac x=element_ptr in allE)
apply(auto)[1]
apply(metis (no_types, lifting) NodeClass.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>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_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>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def option.exhaust_sel)
by(auto)
by (metis (no_types, lifting) NodeClass.l_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas_axioms assms(2) bind.bind_lunit
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_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 finite_set_in 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
l_get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_lemmas.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf option.collapse)
subsection\<open>Preserving Types\<close>
@ -324,13 +322,14 @@ lemma new_element_type_wf_preserved [simp]: "h \<turnstile> new_element \<righta
new_element_def Let_def put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
split: prod.splits if_splits elim!: bind_returns_heap_E)[1]
apply (metis element_ptr_kinds_commutes element_ptrs_def fempty_iff ffmember_filter
apply (metis element_ptr_kinds_commutes element_ptrs_def fempty_iff ffmember_filter finite_set_in
is_element_ptr_ref)
using element_ptrs_def apply fastforce
apply (metis (mono_tags, hide_lams) Suc_n_not_le_n element_ptr.sel(1) element_ptr_kinds_commutes
element_ptrs_def fMax_ge ffmember_filter fimageI is_element_ptr_ref)
by (metis (no_types, lifting) fMax_finsert fempty_iff fimage_is_fempty max_0L new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_ptr_not_in_heap)
apply (metis element_ptrs_def fempty_iff ffmember_filter finite_set_in is_element_ptr_ref)
apply (metis (no_types, lifting) Suc_n_not_le_n element_ptr.sel(1) element_ptr_kinds_commutes
element_ptrs_def fMax_ge ffmember_filter fimage_eqI is_element_ptr_ref notin_fset)
apply (metis (no_types, lifting) Suc_n_not_le_n element_ptr.sel(1) element_ptrs_def
fMax_ge ffmember_filter fimage_eqI finite_set_in is_element_ptr_ref)
done
locale l_new_element = l_type_wf +
assumes new_element_types_preserved: "h \<turnstile> new_element \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
@ -345,12 +344,9 @@ lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_tag_type_typ
Let_def put_M_defs get_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
split: prod.splits option.splits Option.bind_splits elim!: bind_returns_heap_E)[1]
apply (metis option.distinct(1))
apply (metis bind.bind_lunit cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_none)
apply (metis option.distinct(1))
apply (metis bind.bind_lunit cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_none)
by (metis bind.bind_lunit 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_none cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv)
apply (metis finite_set_in option.inject)
apply (metis cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv finite_set_in option.sel)
done
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_child_nodes_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr child_nodes_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
@ -358,11 +354,9 @@ lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_child_nodes_
Let_def put_M_defs get_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
split: prod.splits option.splits Option.bind_splits elim!: bind_returns_heap_E)[1]
apply (metis option.distinct(1))
apply (metis bind.bind_lunit cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_none)
apply (metis option.distinct(1))
apply (metis bind.bind_lunit cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_none)
by (metis bind.bind_lunit 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_none cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv)
apply (metis finite_set_in option.inject)
apply (metis cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv finite_set_in option.sel)
done
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_attrs_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr attrs_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
@ -370,11 +364,9 @@ lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_attrs_type_w
put_M_defs get_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def
get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
split: prod.splits option.splits Option.bind_splits elim!: bind_returns_heap_E)[1]
apply (metis option.distinct(1))
apply (metis bind.bind_lunit cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_none)
apply (metis option.distinct(1))
apply (metis bind.bind_lunit cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_none)
by (metis bind.bind_lunit 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_none cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv)
apply (metis finite_set_in option.inject)
apply (metis cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv finite_set_in option.sel)
done
lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_shadow_root_opt_type_wf_preserved [simp]:
"h \<turnstile> put_M element_ptr shadow_root_opt_update v \<rightarrow>\<^sub>h h' \<Longrightarrow> type_wf h = type_wf h'"
@ -382,11 +374,9 @@ lemma put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_shadow_root_
Let_def put_M_defs get_M_defs put\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def put\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def put\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def
split: prod.splits option.splits Option.bind_splits elim!: bind_returns_heap_E)[1]
apply (metis option.distinct(1))
apply (metis bind.bind_lunit cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_none)
apply (metis option.distinct(1))
apply (metis bind.bind_lunit cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_none)
by (metis bind.bind_lunit 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_none cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv)
apply (metis finite_set_in option.inject)
apply (metis cast\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>2\<^sub>N\<^sub>o\<^sub>d\<^sub>e_inv finite_set_in option.sel)
done
lemma put_M_pointers_preserved:
assumes "h \<turnstile> put_M\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr setter v \<rightarrow>\<^sub>h h'"
@ -447,11 +437,9 @@ lemma type_wf_drop: "type_wf h \<Longrightarrow> type_wf (Heap (fmdrop ptr (the_
apply(auto simp add: type_wf_defs NodeClass.type_wf_defs ObjectClass.type_wf_defs
node_ptr_kinds_def object_ptr_kinds_def is_node_ptr_kind_def
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_def)[1]
apply (metis (mono_tags, lifting) comp_apply ffmember_filter fimage_eqI
is_node_ptr_kind_cast node_ptr_casts_commute2 option.sel)
apply (metis (no_types, lifting) comp_apply element_ptr_kinds_commutes ffmember_filter
fmdom_filter fmfilter_alt_defs(1) heap.sel node_ptr_kinds_commutes object_ptr_kinds_def)
by (metis comp_eq_dest_lhs element_ptr_kinds_commutes fmdom_notI fmdrop_lookup heap.sel
node_ptr_kinds_commutes object_ptr_kinds_def)
apply (metis (no_types, lifting) element_ptr_kinds_commutes finite_set_in fmdom_notD fmdom_notI
fmlookup_drop heap.sel node_ptr_kinds_commutes o_apply object_ptr_kinds_def)
by (metis element_ptr_kinds_commutes fmdom_notI fmdrop_lookup heap.sel node_ptr_kinds_commutes
o_apply object_ptr_kinds_def)
end

View File

@ -76,9 +76,8 @@ lemma node_ptr_kinds_M_reads:
"reads (\<Union>object_ptr. {preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t object_ptr RObject.nothing)}) node_ptr_kinds_M h h'"
using object_ptr_kinds_M_reads
apply (simp add: reads_def node_ptr_kinds_M_defs node_ptr_kinds_def
object_ptr_kinds_M_reads preserved_def cong del: image_cong_simp)
apply (metis (mono_tags, hide_lams) object_ptr_kinds_preserved_small old.unit.exhaust preserved_def)
done
object_ptr_kinds_M_reads preserved_def)
by (smt object_ptr_kinds_preserved_small preserved_def unit_all_impI)
global_interpretation l_put_M type_wf node_ptr_kinds get\<^sub>N\<^sub>o\<^sub>d\<^sub>e put\<^sub>N\<^sub>o\<^sub>d\<^sub>e
rewrites "a_get_M = get_M\<^sub>N\<^sub>o\<^sub>d\<^sub>e"
@ -167,7 +166,7 @@ lemma type_wf_put_ptr_in_heap_E:
shows "type_wf h"
using assms
apply(auto simp add: type_wf_defs split: option.splits if_splits)
by (metis ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf bind.bind_lunit get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def is_node_kind_def option.collapse)
by (metis ObjectClass.get\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t_type_wf bind.bind_lunit finite_set_in get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_def is_node_kind_def option.exhaust_sel)
subsection\<open>Preserving Types\<close>
@ -195,7 +194,8 @@ lemma type_wf_preserved_small:
using type_wf_preserved allI[OF assms(2), of id, simplified]
apply(auto simp add: type_wf_defs)
apply(auto simp add: preserved_def get_M_defs node_ptr_kinds_small[OF assms(1)]
split: option.splits, force)[1]
split: option.splits)[1]
apply (metis notin_fset option.simps(3))
by(auto simp add: preserved_def get_M_defs node_ptr_kinds_small[OF assms(1)]
split: option.splits, force)[1]

View File

@ -244,4 +244,15 @@ proof -
using object_ptr_kinds_preserved_small by blast
qed
lemma reads_writes_preserved2:
assumes "writes SW setter h h'"
assumes "h \<turnstile> setter \<rightarrow>\<^sub>h h'"
assumes "\<And>h h' x. \<forall>w \<in> SW. h \<turnstile> w \<rightarrow>\<^sub>h h' \<longrightarrow> preserved (get_M\<^sub>O\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t ptr getter) h h'"
shows "preserved (get_M ptr getter) h h'"
apply(clarsimp simp add: preserved_def)
using reads_singleton assms(1) assms(2)
apply(rule reads_writes_preserved)
using assms(3)
by(auto simp add: preserved_def)
end

View File

@ -79,9 +79,28 @@ definition
where
"returns_heap h p h' \<longleftrightarrow> (case h \<turnstile> p of Inr (_ , h'') \<Rightarrow> h' = h'' | Inl _ \<Rightarrow> False)"
fun select_heap ("|(_)|\<^sub>h")
where
"select_heap (Inr ( _, h)) = h"
| "select_heap (Inl _) = undefined"
lemma returns_heap_eq [elim]: "h \<turnstile> f \<rightarrow>\<^sub>h h' \<Longrightarrow> h \<turnstile> f \<rightarrow>\<^sub>h h'' \<Longrightarrow> h' = h''"
by(auto simp add: returns_heap_def split: sum.splits)
definition
returns_result_heap :: "'heap \<Rightarrow> ('heap, 'e, 'result) prog \<Rightarrow> 'result \<Rightarrow> 'heap \<Rightarrow> bool"
("((_)/ \<turnstile> (_)/ \<rightarrow>\<^sub>r (_) \<rightarrow>\<^sub>h (_))" [60, 35, 61, 62] 65)
where
"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)"
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")
where
"select_result_heap (Inr (r, h)) = (r, h)"
| "select_result_heap (Inl _) = undefined"
definition
returns_error :: "'heap \<Rightarrow> ('heap, 'e, 'result) prog \<Rightarrow> 'e \<Rightarrow> bool"
("((_)/ \<turnstile> (_)/ \<rightarrow>\<^sub>e (_))" [60, 35, 61] 65)
@ -711,6 +730,11 @@ definition preserved :: "('heap, 'e, 'result) prog \<Rightarrow> 'heap \<Rightar
where
"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)))"
apply(auto simp add: preserved_def)[1]
apply (meson is_OK_returns_result_E is_OK_returns_result_I)+
done
lemma reflp_preserved_f [simp]: "reflp (preserved f)"
by(auto simp add: preserved_def reflp_def)
lemma transp_preserved_f [simp]: "transp (preserved f)"

View File

@ -0,0 +1,92 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
theory Testing_Utils
imports Main
begin
ML \<open>
val _ = Theory.setup
(Method.setup @{binding timed_code_simp}
(Scan.succeed (SIMPLE_METHOD' o (CHANGED_PROP oo (fn a => fn b => fn tac =>
let
val start = Time.now ();
val result = Code_Simp.dynamic_tac a b tac;
val t = Time.now() - start;
in
(if length (Seq.list_of result) > 0 then Output.information ("Took " ^ (Time.toString t)) else ());
result
end))))
"timed simplification with code equations");
val _ = Theory.setup
(Method.setup @{binding timed_eval}
(Scan.succeed (SIMPLE_METHOD' o (fn a => fn b => fn tac =>
let
val eval = CONVERSION (Conv.params_conv ~1 (K (Conv.concl_conv ~1 (Code_Runtime.dynamic_holds_conv a))) a) THEN'
resolve_tac a [TrueI];
val start = Time.now ();
val result = eval b tac
val t = Time.now() - start;
in
(if length (Seq.list_of result) > 0 then Output.information ("Took " ^ (Time.toString t)) else ());
result
end)))
"timed evaluation");
val _ = Theory.setup
(Method.setup @{binding timed_eval_and_code_simp}
(Scan.succeed (SIMPLE_METHOD' o (fn a => fn b => fn tac =>
let
val eval = CONVERSION (Conv.params_conv ~1 (K (Conv.concl_conv ~1 (Code_Runtime.dynamic_holds_conv a))) a) THEN'
resolve_tac a [TrueI];
val start = Time.now ();
val result = eval b tac
val t = Time.now() - start;
val start2 = Time.now ();
val result2_opt =
Timeout.apply (seconds 600.0) (fn _ => SOME (Code_Simp.dynamic_tac a b tac)) ()
handle Timeout.TIMEOUT _ => NONE;
val t2 = Time.now() - start2;
in
if length (Seq.list_of result) > 0 then (Output.information ("eval took " ^ (Time.toString t)); File.append (Path.explode "/tmp/isabellebench") (Time.toString t ^ ",")) else ();
(case result2_opt of
SOME result2 =>
(if length (Seq.list_of result2) > 0 then (Output.information ("code_simp took " ^ (Time.toString t2)); File.append (Path.explode "/tmp/isabellebench") (Time.toString t2 ^ "\n")) else ())
| NONE => (Output.information "code_simp timed out after 600s"; File.append (Path.explode "/tmp/isabellebench") (">600.000\n")));
result
end)))
"timed evaluation and simplification with code equations with file output");
\<close>
(* To run the DOM test cases with timing information output, simply replace the use *)
(* of "eval" with either "timed_code_simp", "timed_eval", or, to run both and write the results *)
(* to /tmp/isabellebench, "timed_eval_and_code_simp". *)
end

View File

@ -79,39 +79,6 @@ definition removeWhiteSpaceOnlyTextNodes :: "((_) object_ptr option) \<Rightarro
"removeWhiteSpaceOnlyTextNodes _ = return ()"
subsection \<open>create\_heap\<close>
(* We use this construction because partially applied functions such as "map_of xs" don't play
well together with the code generator. *)
definition "create_heap xs = Heap (fmap_of_list xs)"
code_datatype ObjectClass.heap.Heap create_heap
lemma object_ptr_kinds_code1 [code]:
"object_ptr_kinds (Heap (fmap_of_list xs)) = object_ptr_kinds (create_heap xs)"
by(simp add: create_heap_def)
lemma object_ptr_kinds_code2 [code]:
"object_ptr_kinds (create_heap xs) = fset_of_list (map fst xs)"
by (simp add: object_ptr_kinds_def create_heap_def dom_map_of_conv_image_fst)
lemma object_ptr_kinds_code3 [code]:
"fmlookup (the_heap (create_heap xs)) x = map_of xs x"
by(auto simp add: create_heap_def fmlookup_of_list)
lemma object_ptr_kinds_code4 [code]:
"the_heap (create_heap xs) = fmap_of_list xs"
by(simp add: create_heap_def)
lemma object_ptr_kinds_code5 [code]:
"the_heap (Heap x) = x"
by simp
lemma object_ptr_kinds_code6 [code]:
"noop = return ()"
by(simp add: noop_def)
subsection \<open>Making the functions under test compatible with untyped languages such as JavaScript\<close>
fun set_attribute_with_null :: "((_) object_ptr option) \<Rightarrow> attr_key \<Rightarrow> attr_value \<Rightarrow> (_, unit) dom_prog"

View File

@ -1,5 +1,5 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
* Copyright (c) 2016-2019 The University of Sheffield, UK
*
* All rights reserved.
*
@ -27,15 +27,17 @@
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Testing adoptNode\<close>
text\<open>This theory contains the test cases for adoptNode.\<close>
(* This file is automatically generated, please do not modify! *)
section\<open>Testing Document\_adoptNode\<close>
text\<open>This theory contains the test cases for Document\_adoptNode.\<close>
theory Document_adoptNode
imports
"Core_DOM_BaseTest"
begin
definition Document_adoptNode_heap :: "(unit, unit, unit, unit, unit, unit, unit, unit, unit, unit, unit) heap" where
definition Document_adoptNode_heap :: heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l where
"Document_adoptNode_heap = create_heap [(cast (document_ptr.Ref 1), cast (create_document_obj html (Some (cast (element_ptr.Ref 1))) [])),
(cast (element_ptr.Ref 1), cast (create_element_obj ''html'' [cast (element_ptr.Ref 2), cast (element_ptr.Ref 8)] fmempty None)),
(cast (element_ptr.Ref 2), cast (create_element_obj ''head'' [cast (element_ptr.Ref 3), cast (element_ptr.Ref 4), cast (element_ptr.Ref 5), cast (element_ptr.Ref 6), cast (element_ptr.Ref 7)] fmempty None)),
@ -52,30 +54,30 @@ definition Document_adoptNode_heap :: "(unit, unit, unit, unit, unit, unit, unit
(cast (element_ptr.Ref 11), cast (create_element_obj ''script'' [cast (character_data_ptr.Ref 3)] fmempty None)),
(cast (character_data_ptr.Ref 3), cast (create_character_data_obj ''%3C%3Cscript%3E%3E''))]"
definition document :: "(unit, unit, unit, unit, unit, unit) object_ptr option" where "document = Some (cast (document_ptr.Ref 1))"
definition Document_adoptNode_document :: "(unit, unit, unit, unit, unit, unit) object_ptr option" where "Document_adoptNode_document = Some (cast (document_ptr.Ref 1))"
text \<open>"Adopting an Element called 'x<' should work."\<close>
lemma "test (do {
tmp0 \<leftarrow> document . getElementsByTagName(''x<'');
tmp0 \<leftarrow> Document_adoptNode_document . getElementsByTagName(''x<'');
y \<leftarrow> return (tmp0 ! 0);
child \<leftarrow> y . firstChild;
tmp1 \<leftarrow> y . parentNode;
tmp2 \<leftarrow> document . body;
tmp2 \<leftarrow> Document_adoptNode_document . body;
assert_equals(tmp1, tmp2);
tmp3 \<leftarrow> y . ownerDocument;
assert_equals(tmp3, document);
tmp4 \<leftarrow> document . adoptNode(y);
assert_equals(tmp3, Document_adoptNode_document);
tmp4 \<leftarrow> Document_adoptNode_document . adoptNode(y);
assert_equals(tmp4, y);
tmp5 \<leftarrow> y . parentNode;
assert_equals(tmp5, None);
tmp6 \<leftarrow> y . firstChild;
assert_equals(tmp6, child);
tmp7 \<leftarrow> y . ownerDocument;
assert_equals(tmp7, document);
assert_equals(tmp7, Document_adoptNode_document);
tmp8 \<leftarrow> child . ownerDocument;
assert_equals(tmp8, document);
assert_equals(tmp8, Document_adoptNode_document);
doc \<leftarrow> createDocument(None, None, None);
tmp9 \<leftarrow> doc . adoptNode(y);
assert_equals(tmp9, y);
@ -94,8 +96,8 @@ lemma "test (do {
text \<open>"Adopting an Element called ':good:times:' should work."\<close>
lemma "test (do {
x \<leftarrow> document . createElement('':good:times:'');
tmp0 \<leftarrow> document . adoptNode(x);
x \<leftarrow> Document_adoptNode_document . createElement('':good:times:'');
tmp0 \<leftarrow> Document_adoptNode_document . adoptNode(x);
assert_equals(tmp0, x);
doc \<leftarrow> createDocument(None, None, None);
tmp1 \<leftarrow> doc . adoptNode(x);

View File

@ -1,5 +1,5 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
* Copyright (c) 2016-2019 The University of Sheffield, UK
*
* All rights reserved.
*
@ -27,15 +27,17 @@
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Testing getElementById\<close>
text\<open>This theory contains the test cases for getElementById.\<close>
(* This file is automatically generated, please do not modify! *)
section\<open>Testing Document\_getElementById\<close>
text\<open>This theory contains the test cases for Document\_getElementById.\<close>
theory Document_getElementById
imports
"Core_DOM_BaseTest"
begin
definition Document_getElementById_heap :: "(unit, unit, unit, unit, unit, unit, unit, unit, unit, unit, unit) heap" where
definition Document_getElementById_heap :: heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l where
"Document_getElementById_heap = create_heap [(cast (document_ptr.Ref 1), cast (create_document_obj html (Some (cast (element_ptr.Ref 1))) [])),
(cast (element_ptr.Ref 1), cast (create_element_obj ''html'' [cast (element_ptr.Ref 2), cast (element_ptr.Ref 9)] fmempty None)),
(cast (element_ptr.Ref 2), cast (create_element_obj ''head'' [cast (element_ptr.Ref 3), cast (element_ptr.Ref 4), cast (element_ptr.Ref 5), cast (element_ptr.Ref 6), cast (element_ptr.Ref 7), cast (element_ptr.Ref 8)] fmempty None)),
@ -60,23 +62,23 @@ definition Document_getElementById_heap :: "(unit, unit, unit, unit, unit, unit,
(cast (element_ptr.Ref 19), cast (create_element_obj ''script'' [cast (character_data_ptr.Ref 3)] fmempty None)),
(cast (character_data_ptr.Ref 3), cast (create_character_data_obj ''%3C%3Cscript%3E%3E''))]"
definition document :: "(unit, unit, unit, unit, unit, unit) object_ptr option" where "document = Some (cast (document_ptr.Ref 1))"
definition Document_getElementById_document :: "(unit, unit, unit, unit, unit, unit) object_ptr option" where "Document_getElementById_document = Some (cast (document_ptr.Ref 1))"
text \<open>"Document.getElementById with a script-inserted element"\<close>
lemma "test (do {
gBody \<leftarrow> document . body;
gBody \<leftarrow> Document_getElementById_document . body;
TEST_ID \<leftarrow> return ''test2'';
test \<leftarrow> document . createElement(''div'');
test \<leftarrow> Document_getElementById_document . createElement(''div'');
test . setAttribute(''id'', TEST_ID);
gBody . appendChild(test);
result \<leftarrow> document . getElementById(TEST_ID);
result \<leftarrow> Document_getElementById_document . getElementById(TEST_ID);
assert_not_equals(result, None, ''should not be null.'');
tmp0 \<leftarrow> result . tagName;
assert_equals(tmp0, ''div'', ''should have appended element's tag name'');
gBody . removeChild(test);
removed \<leftarrow> document . getElementById(TEST_ID);
removed \<leftarrow> Document_getElementById_document . getElementById(TEST_ID);
assert_equals(removed, None, ''should not get removed element.'')
}) Document_getElementById_heap"
by eval
@ -85,19 +87,19 @@ lemma "test (do {
text \<open>"update `id` attribute via setAttribute/removeAttribute"\<close>
lemma "test (do {
gBody \<leftarrow> document . body;
gBody \<leftarrow> Document_getElementById_document . body;
TEST_ID \<leftarrow> return ''test3'';
test \<leftarrow> document . createElement(''div'');
test \<leftarrow> Document_getElementById_document . createElement(''div'');
test . setAttribute(''id'', TEST_ID);
gBody . appendChild(test);
UPDATED_ID \<leftarrow> return ''test3-updated'';
test . setAttribute(''id'', UPDATED_ID);
e \<leftarrow> document . getElementById(UPDATED_ID);
e \<leftarrow> Document_getElementById_document . getElementById(UPDATED_ID);
assert_equals(e, test, ''should get the element with id.'');
old \<leftarrow> document . getElementById(TEST_ID);
old \<leftarrow> Document_getElementById_document . getElementById(TEST_ID);
assert_equals(old, None, ''shouldn't get the element by the old id.'');
test . removeAttribute(''id'');
e2 \<leftarrow> document . getElementById(UPDATED_ID);
e2 \<leftarrow> Document_getElementById_document . getElementById(UPDATED_ID);
assert_equals(e2, None, ''should return null when the passed id is none in document.'')
}) Document_getElementById_heap"
by eval
@ -107,13 +109,13 @@ text \<open>"Ensure that the id attribute only affects elements present in a doc
lemma "test (do {
TEST_ID \<leftarrow> return ''test4-should-not-exist'';
e \<leftarrow> document . createElement(''div'');
e \<leftarrow> Document_getElementById_document . createElement(''div'');
e . setAttribute(''id'', TEST_ID);
tmp0 \<leftarrow> document . getElementById(TEST_ID);
tmp0 \<leftarrow> Document_getElementById_document . getElementById(TEST_ID);
assert_equals(tmp0, None, ''should be null'');
tmp1 \<leftarrow> document . body;
tmp1 \<leftarrow> Document_getElementById_document . body;
tmp1 . appendChild(e);
tmp2 \<leftarrow> document . getElementById(TEST_ID);
tmp2 \<leftarrow> Document_getElementById_document . getElementById(TEST_ID);
assert_equals(tmp2, e, ''should be the appended element'')
}) Document_getElementById_heap"
by eval
@ -122,23 +124,23 @@ lemma "test (do {
text \<open>"in tree order, within the context object's tree"\<close>
lemma "test (do {
gBody \<leftarrow> document . body;
gBody \<leftarrow> Document_getElementById_document . body;
TEST_ID \<leftarrow> return ''test5'';
target \<leftarrow> document . getElementById(TEST_ID);
target \<leftarrow> Document_getElementById_document . getElementById(TEST_ID);
assert_not_equals(target, None, ''should not be null'');
tmp0 \<leftarrow> target . getAttribute(''data-name'');
assert_equals(tmp0, ''1st'', ''should return the 1st'');
element4 \<leftarrow> document . createElement(''div'');
element4 \<leftarrow> Document_getElementById_document . createElement(''div'');
element4 . setAttribute(''id'', TEST_ID);
element4 . setAttribute(''data-name'', ''4th'');
gBody . appendChild(element4);
target2 \<leftarrow> document . getElementById(TEST_ID);
target2 \<leftarrow> Document_getElementById_document . getElementById(TEST_ID);
assert_not_equals(target2, None, ''should not be null'');
tmp1 \<leftarrow> target2 . getAttribute(''data-name'');
assert_equals(tmp1, ''1st'', ''should be the 1st'');
tmp2 \<leftarrow> target2 . parentNode;
tmp2 . removeChild(target2);
target3 \<leftarrow> document . getElementById(TEST_ID);
target3 \<leftarrow> Document_getElementById_document . getElementById(TEST_ID);
assert_not_equals(target3, None, ''should not be null'');
tmp3 \<leftarrow> target3 . getAttribute(''data-name'');
assert_equals(tmp3, ''4th'', ''should be the 4th'')
@ -146,17 +148,15 @@ lemma "test (do {
by eval
text \<open>"Modern browsers optimize this method with using internal id cache.
This test checks that their optimization should effect only append to
`Document`, not append to `Node`."\<close>
text \<open>"Modern browsers optimize this method with using internal id cache. This test checks that their optimization should effect only append to `Document`, not append to `Node`."\<close>
lemma "test (do {
TEST_ID \<leftarrow> return ''test6'';
s \<leftarrow> document . createElement(''div'');
s \<leftarrow> Document_getElementById_document . createElement(''div'');
s . setAttribute(''id'', TEST_ID);
tmp0 \<leftarrow> document . createElement(''div'');
tmp0 \<leftarrow> Document_getElementById_document . createElement(''div'');
tmp0 . appendChild(s);
tmp1 \<leftarrow> document . getElementById(TEST_ID);
tmp1 \<leftarrow> Document_getElementById_document . getElementById(TEST_ID);
assert_equals(tmp1, None, ''should be null'')
}) Document_getElementById_heap"
by eval
@ -165,17 +165,17 @@ lemma "test (do {
text \<open>"changing attribute's value via `Attr` gotten from `Element.attribute`."\<close>
lemma "test (do {
gBody \<leftarrow> document . body;
gBody \<leftarrow> Document_getElementById_document . body;
TEST_ID \<leftarrow> return ''test7'';
element \<leftarrow> document . createElement(''div'');
element \<leftarrow> Document_getElementById_document . createElement(''div'');
element . setAttribute(''id'', TEST_ID);
gBody . appendChild(element);
target \<leftarrow> document . getElementById(TEST_ID);
target \<leftarrow> Document_getElementById_document . getElementById(TEST_ID);
assert_equals(target, element, ''should return the element before changing the value'');
element . setAttribute(''id'', (TEST_ID @ ''-updated''));
target2 \<leftarrow> document . getElementById(TEST_ID);
target2 \<leftarrow> Document_getElementById_document . getElementById(TEST_ID);
assert_equals(target2, None, ''should return null after updated id via Attr.value'');
target3 \<leftarrow> document . getElementById((TEST_ID @ ''-updated''));
target3 \<leftarrow> Document_getElementById_document . getElementById((TEST_ID @ ''-updated''));
assert_equals(target3, element, ''should be equal to the updated element.'')
}) Document_getElementById_heap"
by eval
@ -184,19 +184,19 @@ lemma "test (do {
text \<open>"update `id` attribute via element.id"\<close>
lemma "test (do {
gBody \<leftarrow> document . body;
gBody \<leftarrow> Document_getElementById_document . body;
TEST_ID \<leftarrow> return ''test12'';
test \<leftarrow> document . createElement(''div'');
test \<leftarrow> Document_getElementById_document . createElement(''div'');
test . setAttribute(''id'', TEST_ID);
gBody . appendChild(test);
UPDATED_ID \<leftarrow> return (TEST_ID @ ''-updated'');
test . setAttribute(''id'', UPDATED_ID);
e \<leftarrow> document . getElementById(UPDATED_ID);
e \<leftarrow> Document_getElementById_document . getElementById(UPDATED_ID);
assert_equals(e, test, ''should get the element with id.'');
old \<leftarrow> document . getElementById(TEST_ID);
old \<leftarrow> Document_getElementById_document . getElementById(TEST_ID);
assert_equals(old, None, ''shouldn't get the element by the old id.'');
test . setAttribute(''id'', '''');
e2 \<leftarrow> document . getElementById(UPDATED_ID);
e2 \<leftarrow> Document_getElementById_document . getElementById(UPDATED_ID);
assert_equals(e2, None, ''should return null when the passed id is none in document.'')
}) Document_getElementById_heap"
by eval
@ -205,33 +205,33 @@ lemma "test (do {
text \<open>"where insertion order and tree order don't match"\<close>
lemma "test (do {
gBody \<leftarrow> document . body;
gBody \<leftarrow> Document_getElementById_document . body;
TEST_ID \<leftarrow> return ''test13'';
container \<leftarrow> document . createElement(''div'');
container \<leftarrow> Document_getElementById_document . createElement(''div'');
container . setAttribute(''id'', (TEST_ID @ ''-fixture''));
gBody . appendChild(container);
element1 \<leftarrow> document . createElement(''div'');
element1 \<leftarrow> Document_getElementById_document . createElement(''div'');
element1 . setAttribute(''id'', TEST_ID);
element2 \<leftarrow> document . createElement(''div'');
element2 \<leftarrow> Document_getElementById_document . createElement(''div'');
element2 . setAttribute(''id'', TEST_ID);
element3 \<leftarrow> document . createElement(''div'');
element3 \<leftarrow> Document_getElementById_document . createElement(''div'');
element3 . setAttribute(''id'', TEST_ID);
element4 \<leftarrow> document . createElement(''div'');
element4 \<leftarrow> Document_getElementById_document . createElement(''div'');
element4 . setAttribute(''id'', TEST_ID);
container . appendChild(element2);
container . appendChild(element4);
container . insertBefore(element3, element4);
container . insertBefore(element1, element2);
test \<leftarrow> document . getElementById(TEST_ID);
test \<leftarrow> Document_getElementById_document . getElementById(TEST_ID);
assert_equals(test, element1, ''should return 1st element'');
container . removeChild(element1);
test \<leftarrow> document . getElementById(TEST_ID);
test \<leftarrow> Document_getElementById_document . getElementById(TEST_ID);
assert_equals(test, element2, ''should return 2nd element'');
container . removeChild(element2);
test \<leftarrow> document . getElementById(TEST_ID);
test \<leftarrow> Document_getElementById_document . getElementById(TEST_ID);
assert_equals(test, element3, ''should return 3rd element'');
container . removeChild(element3);
test \<leftarrow> document . getElementById(TEST_ID);
test \<leftarrow> Document_getElementById_document . getElementById(TEST_ID);
assert_equals(test, element4, ''should return 4th element'');
container . removeChild(element4)
}) Document_getElementById_heap"
@ -241,16 +241,16 @@ lemma "test (do {
text \<open>"Inserting an id by inserting its parent node"\<close>
lemma "test (do {
gBody \<leftarrow> document . body;
gBody \<leftarrow> Document_getElementById_document . body;
TEST_ID \<leftarrow> return ''test14'';
a \<leftarrow> document . createElement(''a'');
b \<leftarrow> document . createElement(''b'');
a \<leftarrow> Document_getElementById_document . createElement(''a'');
b \<leftarrow> Document_getElementById_document . createElement(''b'');
a . appendChild(b);
b . setAttribute(''id'', TEST_ID);
tmp0 \<leftarrow> document . getElementById(TEST_ID);
tmp0 \<leftarrow> Document_getElementById_document . getElementById(TEST_ID);
assert_equals(tmp0, None);
gBody . appendChild(a);
tmp1 \<leftarrow> document . getElementById(TEST_ID);
tmp1 \<leftarrow> Document_getElementById_document . getElementById(TEST_ID);
assert_equals(tmp1, b)
}) Document_getElementById_heap"
by eval
@ -260,15 +260,15 @@ text \<open>"Document.getElementById must not return nodes not present in docume
lemma "test (do {
TEST_ID \<leftarrow> return ''test15'';
outer \<leftarrow> document . getElementById(''outer'');
middle \<leftarrow> document . getElementById(''middle'');
inner \<leftarrow> document . getElementById(''inner'');
tmp0 \<leftarrow> document . getElementById(''middle'');
outer \<leftarrow> Document_getElementById_document . getElementById(''outer'');
middle \<leftarrow> Document_getElementById_document . getElementById(''middle'');
inner \<leftarrow> Document_getElementById_document . getElementById(''inner'');
tmp0 \<leftarrow> Document_getElementById_document . getElementById(''middle'');
outer . removeChild(tmp0);
new_el \<leftarrow> document . createElement(''h1'');
new_el \<leftarrow> Document_getElementById_document . createElement(''h1'');
new_el . setAttribute(''id'', ''heading'');
inner . appendChild(new_el);
tmp1 \<leftarrow> document . getElementById(''heading'');
tmp1 \<leftarrow> Document_getElementById_document . getElementById(''heading'');
assert_equals(tmp1, None)
}) Document_getElementById_heap"
by eval

View File

@ -1,5 +1,5 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
* Copyright (c) 2016-2019 The University of Sheffield, UK
*
* All rights reserved.
*
@ -27,15 +27,17 @@
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Testing insertBefore\<close>
text\<open>This theory contains the test cases for insertBefore.\<close>
(* This file is automatically generated, please do not modify! *)
section\<open>Testing Node\_insertBefore\<close>
text\<open>This theory contains the test cases for Node\_insertBefore.\<close>
theory Node_insertBefore
imports
"Core_DOM_BaseTest"
begin
definition Node_insertBefore_heap :: "(unit, unit, unit, unit, unit, unit, unit, unit, unit, unit, unit) heap" where
definition Node_insertBefore_heap :: heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l where
"Node_insertBefore_heap = create_heap [(cast (document_ptr.Ref 1), cast (create_document_obj html (Some (cast (element_ptr.Ref 1))) [])),
(cast (element_ptr.Ref 1), cast (create_element_obj ''html'' [cast (element_ptr.Ref 2), cast (element_ptr.Ref 6)] fmempty None)),
(cast (element_ptr.Ref 2), cast (create_element_obj ''head'' [cast (element_ptr.Ref 3), cast (element_ptr.Ref 4), cast (element_ptr.Ref 5)] fmempty None)),
@ -48,14 +50,14 @@ definition Node_insertBefore_heap :: "(unit, unit, unit, unit, unit, unit, unit,
(cast (element_ptr.Ref 8), cast (create_element_obj ''script'' [cast (character_data_ptr.Ref 2)] fmempty None)),
(cast (character_data_ptr.Ref 2), cast (create_character_data_obj ''%3C%3Cscript%3E%3E''))]"
definition document :: "(unit, unit, unit, unit, unit, unit) object_ptr option" where "document = Some (cast (document_ptr.Ref 1))"
definition Node_insertBefore_document :: "(unit, unit, unit, unit, unit, unit) object_ptr option" where "Node_insertBefore_document = Some (cast (document_ptr.Ref 1))"
text \<open>"Calling insertBefore an a leaf node Text must throw HIERARCHY\_REQUEST\_ERR."\<close>
lemma "test (do {
node \<leftarrow> document . createTextNode(''Foo'');
tmp0 \<leftarrow> document . createTextNode(''fail'');
node \<leftarrow> Node_insertBefore_document . createTextNode(''Foo'');
tmp0 \<leftarrow> Node_insertBefore_document . createTextNode(''fail'');
assert_throws(HierarchyRequestError, node . insertBefore(tmp0, None))
}) Node_insertBefore_heap"
by eval
@ -64,13 +66,13 @@ lemma "test (do {
text \<open>"Calling insertBefore with an inclusive ancestor of the context object must throw HIERARCHY\_REQUEST\_ERR."\<close>
lemma "test (do {
tmp1 \<leftarrow> document . body;
tmp2 \<leftarrow> document . getElementById(''log'');
tmp0 \<leftarrow> document . body;
tmp1 \<leftarrow> Node_insertBefore_document . body;
tmp2 \<leftarrow> Node_insertBefore_document . getElementById(''log'');
tmp0 \<leftarrow> Node_insertBefore_document . body;
assert_throws(HierarchyRequestError, tmp0 . insertBefore(tmp1, tmp2));
tmp4 \<leftarrow> document . documentElement;
tmp5 \<leftarrow> document . getElementById(''log'');
tmp3 \<leftarrow> document . body;
tmp4 \<leftarrow> Node_insertBefore_document . documentElement;
tmp5 \<leftarrow> Node_insertBefore_document . getElementById(''log'');
tmp3 \<leftarrow> Node_insertBefore_document . body;
assert_throws(HierarchyRequestError, tmp3 . insertBefore(tmp4, tmp5))
}) Node_insertBefore_heap"
by eval
@ -79,9 +81,9 @@ lemma "test (do {
text \<open>"Calling insertBefore with a reference child whose parent is not the context node must throw a NotFoundError."\<close>
lemma "test (do {
a \<leftarrow> document . createElement(''div'');
b \<leftarrow> document . createElement(''div'');
c \<leftarrow> document . createElement(''div'');
a \<leftarrow> Node_insertBefore_document . createElement(''div'');
b \<leftarrow> Node_insertBefore_document . createElement(''div'');
c \<leftarrow> Node_insertBefore_document . createElement(''div'');
assert_throws(NotFoundError, a . insertBefore(b, c))
}) Node_insertBefore_heap"
by eval
@ -104,9 +106,9 @@ lemma "test (do {
text \<open>"Inserting a node before itself should not move the node"\<close>
lemma "test (do {
a \<leftarrow> document . createElement(''div'');
b \<leftarrow> document . createElement(''div'');
c \<leftarrow> document . createElement(''div'');
a \<leftarrow> Node_insertBefore_document . createElement(''div'');
b \<leftarrow> Node_insertBefore_document . createElement(''div'');
c \<leftarrow> Node_insertBefore_document . createElement(''div'');
a . appendChild(b);
a . appendChild(c);
tmp0 \<leftarrow> a . childNodes;

View File

@ -1,5 +1,5 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
* Copyright (c) 2016-2019 The University of Sheffield, UK
*
* All rights reserved.
*
@ -27,15 +27,17 @@
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
section\<open>Testing removeChild\<close>
text\<open>This theory contains the test cases for removeChild.\<close>
(* This file is automatically generated, please do not modify! *)
section\<open>Testing Node\_removeChild\<close>
text\<open>This theory contains the test cases for Node\_removeChild.\<close>
theory Node_removeChild
imports
"Core_DOM_BaseTest"
begin
definition Node_removeChild_heap :: "(unit, unit, unit, unit, unit, unit, unit, unit, unit, unit, unit) heap" where
definition Node_removeChild_heap :: heap\<^sub>f\<^sub>i\<^sub>n\<^sub>a\<^sub>l where
"Node_removeChild_heap = create_heap [(cast (document_ptr.Ref 1), cast (create_document_obj html (Some (cast (element_ptr.Ref 1))) [])),
(cast (element_ptr.Ref 1), cast (create_element_obj ''html'' [cast (element_ptr.Ref 2), cast (element_ptr.Ref 7)] fmempty None)),
(cast (element_ptr.Ref 2), cast (create_element_obj ''head'' [cast (element_ptr.Ref 3), cast (element_ptr.Ref 4), cast (element_ptr.Ref 5), cast (element_ptr.Ref 6)] fmempty None)),
@ -50,17 +52,17 @@ definition Node_removeChild_heap :: "(unit, unit, unit, unit, unit, unit, unit,
(cast (element_ptr.Ref 10), cast (create_element_obj ''script'' [cast (character_data_ptr.Ref 2)] fmempty None)),
(cast (character_data_ptr.Ref 2), cast (create_character_data_obj ''%3C%3Cscript%3E%3E''))]"
definition document :: "(unit, unit, unit, unit, unit, unit) object_ptr option" where "document = Some (cast (document_ptr.Ref 1))"
definition Node_removeChild_document :: "(unit, unit, unit, unit, unit, unit) object_ptr option" where "Node_removeChild_document = Some (cast (document_ptr.Ref 1))"
text \<open>"Passing a detached Element to removeChild should not affect it."\<close>
lemma "test (do {
doc \<leftarrow> return document;
doc \<leftarrow> return Node_removeChild_document;
s \<leftarrow> doc . createElement(''div'');
tmp0 \<leftarrow> s . ownerDocument;
assert_equals(tmp0, doc);
tmp1 \<leftarrow> document . body;
tmp1 \<leftarrow> Node_removeChild_document . body;
assert_throws(NotFoundError, tmp1 . removeChild(s));
tmp2 \<leftarrow> s . ownerDocument;
assert_equals(tmp2, doc)
@ -71,13 +73,13 @@ lemma "test (do {
text \<open>"Passing a non-detached Element to removeChild should not affect it."\<close>
lemma "test (do {
doc \<leftarrow> return document;
doc \<leftarrow> return Node_removeChild_document;
s \<leftarrow> doc . createElement(''div'');
tmp0 \<leftarrow> doc . documentElement;
tmp0 . appendChild(s);
tmp1 \<leftarrow> s . ownerDocument;
assert_equals(tmp1, doc);
tmp2 \<leftarrow> document . body;
tmp2 \<leftarrow> Node_removeChild_document . body;
assert_throws(NotFoundError, tmp2 . removeChild(s));
tmp3 \<leftarrow> s . ownerDocument;
assert_equals(tmp3, doc)
@ -88,7 +90,7 @@ lemma "test (do {
text \<open>"Calling removeChild on an Element with no children should throw NOT\_FOUND\_ERR."\<close>
lemma "test (do {
doc \<leftarrow> return document;
doc \<leftarrow> return Node_removeChild_document;
s \<leftarrow> doc . createElement(''div'');
tmp0 \<leftarrow> doc . body;
tmp0 . appendChild(s);
@ -106,7 +108,7 @@ lemma "test (do {
s \<leftarrow> doc . createElement(''div'');
tmp0 \<leftarrow> s . ownerDocument;
assert_equals(tmp0, doc);
tmp1 \<leftarrow> document . body;
tmp1 \<leftarrow> Node_removeChild_document . body;
assert_throws(NotFoundError, tmp1 . removeChild(s));
tmp2 \<leftarrow> s . ownerDocument;
assert_equals(tmp2, doc)
@ -123,7 +125,7 @@ lemma "test (do {
tmp0 . appendChild(s);
tmp1 \<leftarrow> s . ownerDocument;
assert_equals(tmp1, doc);
tmp2 \<leftarrow> document . body;
tmp2 \<leftarrow> Node_removeChild_document . body;
assert_throws(NotFoundError, tmp2 . removeChild(s));
tmp3 \<leftarrow> s . ownerDocument;
assert_equals(tmp3, doc)
@ -148,7 +150,7 @@ lemma "test (do {
text \<open>"Passing a value that is not a Node reference to removeChild should throw TypeError."\<close>
lemma "test (do {
tmp0 \<leftarrow> document . body;
tmp0 \<leftarrow> Node_removeChild_document . body;
assert_throws(TypeError, tmp0 . removeChild(None))
}) Node_removeChild_heap"
by eval

View File

@ -1,39 +0,0 @@
(***********************************************************************************
* Copyright (c) 2016-2018 The University of Sheffield, UK
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are met:
*
* * Redistributions of source code must retain the above copyright notice, this
* list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above copyright notice,
* this list of conditions and the following disclaimer in the documentation
* and/or other materials provided with the distribution.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS "AS IS"
* AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT LIMITED TO, THE
* IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR PURPOSE ARE
* DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT HOLDER OR CONTRIBUTORS BE LIABLE
* FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL, EXEMPLARY, OR CONSEQUENTIAL
* DAMAGES (INCLUDING, BUT NOT LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR
* SERVICES; LOSS OF USE, DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER
* CAUSED AND ON ANY THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY,
* OR TORT (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
*
* SPDX-License-Identifier: BSD-2-Clause
***********************************************************************************)
theory Testing_Utils
imports Main
begin
ML \<open>
val _ = Theory.setup
(Method.setup @{binding timed_code_simp}
(Scan.succeed (SIMPLE_METHOD' o (CHANGED_PROP oo (fn a => fn b => Timeout.apply (seconds 3600.0) (Code_Simp.dynamic_tac a b)))))
"simplification with code equations, aborts after 1 hour");
\<close>
end

View File

@ -31,8 +31,8 @@ section\<open>Element\<close>
text\<open>In this theory, we introduce the types for the Element class.\<close>
theory ElementClass
imports
NodeClass
"../pointers/ShadowRootPointer"
"NodeClass"
"ShadowRootPointer"
begin
text\<open>The type @{type "DOMString"} is a type synonym for @{type "string"}, define
in \autoref{sec:Core_DOM_Basic_Datatypes}.\<close>
@ -68,6 +68,7 @@ type_synonym
('node_ptr, 'element_ptr, 'character_data_ptr, 'shadow_root_ptr, 'Element option) RElement_ext + 'Node) heap"
register_default_tvars
"('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr, 'shadow_root_ptr, 'Object, 'Node, 'Element) heap"
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"
where
@ -131,8 +132,8 @@ locale l_type_wf_def\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
begin
definition a_type_wf :: "(_) heap \<Rightarrow> bool"
where
"a_type_wf h = (NodeClass.type_wf h \<and> (\<forall>element_ptr. element_ptr |\<in>| element_ptr_kinds h
\<longrightarrow> get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h \<noteq> None))"
"a_type_wf h = (NodeClass.type_wf h \<and> (\<forall>element_ptr \<in> fset (element_ptr_kinds h).
get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h \<noteq> None))"
end
global_interpretation l_type_wf_def\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t defines type_wf = a_type_wf .
lemmas type_wf_defs = a_type_wf_def
@ -154,7 +155,7 @@ lemma get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_type_wf:
shows "element_ptr |\<in>| element_ptr_kinds h \<longleftrightarrow> get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t element_ptr h \<noteq> None"
using l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_axioms assms
apply(simp add: type_wf_defs get\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def l_type_wf\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def)
by (metis NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf bind_eq_None_conv element_ptr_kinds_commutes
by (metis NodeClass.get\<^sub>N\<^sub>o\<^sub>d\<^sub>e_type_wf bind_eq_None_conv element_ptr_kinds_commutes notin_fset
option.distinct(1))
end
@ -273,7 +274,6 @@ lemma new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_get\<^sub>E\<^
using assms
by(auto simp add: new\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t_def Let_def)
locale l_known_ptr\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t
begin
definition a_known_ptr :: "(_) object_ptr \<Rightarrow> bool"
@ -291,21 +291,24 @@ locale l_known_ptrs\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t = l_
begin
definition a_known_ptrs :: "(_) heap \<Rightarrow> bool"
where
"a_known_ptrs h = (\<forall>ptr. ptr |\<in>| object_ptr_kinds h \<longrightarrow> known_ptr ptr)"
"a_known_ptrs h = (\<forall>ptr \<in> fset (object_ptr_kinds h). known_ptr ptr)"
lemma known_ptrs_known_ptr:
"ptr |\<in>| object_ptr_kinds h \<Longrightarrow> a_known_ptrs h \<Longrightarrow> known_ptr ptr"
by(simp add: a_known_ptrs_def)
apply(simp add: a_known_ptrs_def)
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'"
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'"
by(auto simp add: a_known_ptrs_def)
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'"
by(simp add: a_known_ptrs_def)
end
global_interpretation l_known_ptrs\<^sub>E\<^sub>l\<^sub>e\<^sub>m\<^sub>e\<^sub>n\<^sub>t known_ptr defines known_ptrs = a_known_ptrs .
lemmas known_ptrs_defs = a_known_ptrs_def
lemma known_ptrs_is_l_known_ptrs: "l_known_ptrs known_ptr known_ptrs"
using known_ptrs_known_ptr known_ptrs_preserved l_known_ptrs_def known_ptrs_subset by blast
using known_ptrs_known_ptr known_ptrs_preserved known_ptrs_subset known_ptrs_new_ptr l_known_ptrs_def by blast
end

View File

@ -34,7 +34,7 @@ We only include them here, as they are required for future work and they cannot
following the object-oriented extensibility of our data model.\<close>
theory ShadowRootPointer
imports
DocumentPointer
"DocumentPointer"
begin
datatype 'shadow_root_ptr shadow_root_ptr = Ref (the_ref: ref) | Ext 'shadow_root_ptr
@ -46,6 +46,10 @@ type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr,
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr,
'document_ptr, 'shadow_root_ptr) object_ptr"
definition 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>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_) shadow_root_ptr \<Rightarrow> (_) shadow_root_ptr"
where
"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>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 = id"
definition cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r :: "(_)shadow_root_ptr \<Rightarrow> (_) object_ptr"
where
"cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r ptr = object_ptr.Ext (Inr (Inr (Inl ptr)))"
@ -56,7 +60,7 @@ definition cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\
object_ptr.Ext (Inr (Inr (Inl shadow_root_ptr))) \<Rightarrow> Some shadow_root_ptr
| _ \<Rightarrow> None)"
adhoc_overloading cast cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r
adhoc_overloading cast cast\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r cast\<^sub>o\<^sub>b\<^sub>j\<^sub>e\<^sub>c\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r\<^sub>2\<^sub>s\<^sub>h\<^sub>a\<^sub>d\<^sub>o\<^sub>w\<^sub>_\<^sub>r\<^sub>o\<^sub>o\<^sub>t\<^sub>_\<^sub>p\<^sub>t\<^sub>r 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>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
definition is_shadow_root_ptr_kind :: "(_) object_ptr \<Rightarrow> bool"