Theory ShadowRootMonad
section‹Shadow Root Monad›
theory ShadowRootMonad
imports
"Core_DOM.DocumentMonad"
"../classes/ShadowRootClass"
begin
type_synonym ('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'Document, 'ShadowRoot, 'result) dom_prog
= "((_) heap, exception, 'result) prog"
register_default_tvars "('object_ptr, 'node_ptr, 'element_ptr, 'character_data_ptr, 'document_ptr,
'shadow_root_ptr, 'Object, 'Node, 'Element, 'CharacterData, 'Document, 'ShadowRoot, 'result) dom_prog"
global_interpretation l_ptr_kinds_M shadow_root_ptr_kinds defines shadow_root_ptr_kinds_M = a_ptr_kinds_M .
lemmas shadow_root_ptr_kinds_M_defs = a_ptr_kinds_M_def
lemma shadow_root_ptr_kinds_M_eq:
assumes "|h ⊢ object_ptr_kinds_M|⇩r = |h' ⊢ object_ptr_kinds_M|⇩r"
shows "|h ⊢ shadow_root_ptr_kinds_M|⇩r = |h' ⊢ shadow_root_ptr_kinds_M|⇩r"
using assms
by(auto simp add: shadow_root_ptr_kinds_M_defs object_ptr_kinds_M_defs shadow_root_ptr_kinds_def)
global_interpretation l_dummy defines get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t = "l_get_M.a_get_M get⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t" .
lemma get_M_is_l_get_M: "l_get_M get⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t type_wf shadow_root_ptr_kinds"
apply(simp add: get⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_type_wf l_get_M_def)
by (metis ObjectClass.get⇩O⇩b⇩j⇩e⇩c⇩t_type_wf ObjectClass.type_wf_defs bind_eq_None_conv
shadow_root_ptr_kinds_commutes get⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def option.simps(3))
lemmas get_M_defs = get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def[unfolded l_get_M.a_get_M_def[OF get_M_is_l_get_M]]
adhoc_overloading get_M get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t
locale l_get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_lemmas = l_type_wf⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t
begin
sublocale l_get_M⇩C⇩h⇩a⇩r⇩a⇩c⇩t⇩e⇩r⇩D⇩a⇩t⇩a_lemmas by unfold_locales
interpretation l_get_M get⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t type_wf shadow_root_ptr_kinds
apply(unfold_locales)
apply (simp add: get⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_type_wf local.type_wf⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t)
by (meson ShadowRootMonad.get_M_is_l_get_M l_get_M_def)
lemmas get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_ok = get_M_ok[folded get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def]
lemmas get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_ptr_in_heap = get_M_ptr_in_heap[folded get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def]
end
global_interpretation l_get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_lemmas type_wf by unfold_locales
global_interpretation l_put_M type_wf shadow_root_ptr_kinds get⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t put⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t rewrites
"a_get_M = get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t" defines put_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t = a_put_M
apply (simp add: get_M_is_l_get_M l_put_M_def)
by (simp add: get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def)
lemmas put_M_defs = a_put_M_def
adhoc_overloading put_M put_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t
locale l_put_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_lemmas = l_type_wf⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t
begin
sublocale l_put_M⇩C⇩h⇩a⇩r⇩a⇩c⇩t⇩e⇩r⇩D⇩a⇩t⇩a_lemmas by unfold_locales
interpretation l_put_M type_wf shadow_root_ptr_kinds get⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t put⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t
apply(unfold_locales)
apply (simp add: get⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_type_wf local.type_wf⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t)
by (meson ShadowRootMonad.get_M_is_l_get_M l_get_M_def)
lemmas put_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_ok = put_M_ok[folded put_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def]
end
global_interpretation l_put_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_lemmas type_wf by unfold_locales
lemma shadow_root_put_get [simp]: "h ⊢ put_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr setter v →⇩h h'
⟹ (⋀x. getter (setter (λ_. v) x) = v)
⟹ h' ⊢ get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr getter →⇩r v"
by(auto simp add: put_M_defs get_M_defs split: option.splits)
lemma get_M_Mshadow_root_preserved1 [simp]:
"shadow_root_ptr ≠ shadow_root_ptr'
⟹ h ⊢ put_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr setter v →⇩h h'
⟹ preserved (get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr' getter) h h'"
by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma shadow_root_put_get_preserved [simp]:
"h ⊢ put_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr setter v →⇩h h'
⟹ (⋀x. getter (setter (λ_. v) x) = getter x)
⟹ preserved (get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr' getter) h h'"
apply(cases "shadow_root_ptr = shadow_root_ptr'")
by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved2 [simp]:
"h ⊢ put_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr setter v →⇩h h' ⟹ preserved (get_M⇩N⇩o⇩d⇩e node_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs NodeMonad.get_M_defs get⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def
put⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def get⇩N⇩o⇩d⇩e_def preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved3 [simp]:
"cast shadow_root_ptr ≠ object_ptr
⟹ h ⊢ put_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr setter v →⇩h h'
⟹ preserved (get_M⇩O⇩b⇩j⇩e⇩c⇩t object_ptr getter) h h'"
by(auto simp add: put_M_defs get_M_defs get⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def put⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def ObjectMonad.get_M_defs
preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved4 [simp]:
"h ⊢ put_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr setter v →⇩h h'
⟹ (⋀x. getter (cast (setter (λ_. v) x)) = getter (cast x))
⟹ preserved (get_M⇩O⇩b⇩j⇩e⇩c⇩t object_ptr getter) h h'"
apply(cases "cast shadow_root_ptr ≠ object_ptr")[1]
by(auto simp add: put_M_defs get_M_defs get⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def put⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def
ObjectMonad.get_M_defs preserved_def
split: option.splits bind_splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved5 [simp]:
"cast shadow_root_ptr ≠ object_ptr
⟹ h ⊢ put_M⇩O⇩b⇩j⇩e⇩c⇩t object_ptr setter v →⇩h h'
⟹ preserved (get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr getter) h h'"
by(auto simp add: ObjectMonad.put_M_defs get_M_defs get⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def ObjectMonad.get_M_defs
preserved_def split: option.splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved6 [simp]:
"h ⊢ put_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr setter v →⇩h h' ⟹ preserved (get_M⇩E⇩l⇩e⇩m⇩e⇩n⇩t element_ptr getter) h h'"
by(auto simp add: put_M_defs ElementMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved7 [simp]:
"h ⊢ put_M⇩E⇩l⇩e⇩m⇩e⇩n⇩t element_ptr setter v →⇩h h' ⟹ preserved (get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr getter) h h'"
by(auto simp add: ElementMonad.put_M_defs get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved8 [simp]:
"h ⊢ put_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr setter v →⇩h h'
⟹ preserved (get_M⇩C⇩h⇩a⇩r⇩a⇩c⇩t⇩e⇩r⇩D⇩a⇩t⇩a character_data_ptr getter) h h'"
by(auto simp add: put_M_defs CharacterDataMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved9 [simp]:
"h ⊢ put_M⇩C⇩h⇩a⇩r⇩a⇩c⇩t⇩e⇩r⇩D⇩a⇩t⇩a character_data_ptr setter v →⇩h h'
⟹ preserved (get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr getter) h h'"
by(auto simp add: CharacterDataMonad.put_M_defs get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Mshadow_root_preserved10 [simp]:
"(⋀x. getter (cast (setter (λ_. v) x)) = getter (cast x))
⟹ h ⊢ put_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr setter v →⇩h h' ⟹ preserved (get_M⇩O⇩b⇩j⇩e⇩c⇩t object_ptr getter) h h'"
apply(cases "cast shadow_root_ptr = object_ptr")
by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs NodeMonad.get_M_defs get⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def
get⇩N⇩o⇩d⇩e_def preserved_def put⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def put⇩N⇩o⇩d⇩e_def bind_eq_Some_conv
split: option.splits)
lemma new_element_get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t:
"h ⊢ new_element →⇩h h' ⟹ preserved (get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t ptr getter) h h'"
by(auto simp add: new_element_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_character_data_get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t:
"h ⊢ new_character_data →⇩h h' ⟹ preserved (get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t ptr getter) h h'"
by(auto simp add: new_character_data_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_document_get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t:
"h ⊢ new_document →⇩h h' ⟹ preserved (get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t ptr getter) h h'"
by(auto simp add: new_document_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
definition delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M :: "(_) shadow_root_ptr ⇒ (_, unit) dom_prog" where
"delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M shadow_root_ptr = do {
h ← get_heap;
(case delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr h of
Some h ⇒ return_heap h |
None ⇒ error HierarchyRequestError)
}"
adhoc_overloading delete_M delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M
lemma delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_ok [simp]:
assumes "shadow_root_ptr |∈| shadow_root_ptr_kinds h"
shows "h ⊢ ok (delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M shadow_root_ptr)"
using assms
by(auto simp add: delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def delete⇩O⇩b⇩j⇩e⇩c⇩t_def split: prod.splits)
lemma delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_ptr_in_heap:
assumes "h ⊢ delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M shadow_root_ptr →⇩h h'"
shows "shadow_root_ptr |∈| shadow_root_ptr_kinds h"
using assms
by(auto simp add: delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def delete⇩O⇩b⇩j⇩e⇩c⇩t_def split: if_splits)
lemma delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_ptr_not_in_heap:
assumes "h ⊢ delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M shadow_root_ptr →⇩h h'"
shows "shadow_root_ptr |∉| shadow_root_ptr_kinds h'"
using assms
apply(auto simp add: delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def delete⇩O⇩b⇩j⇩e⇩c⇩t_def split: if_splits)[1]
by (metis comp_apply fmdom_notI fmdrop_lookup heap.sel object_ptr_kinds_def shadow_root_ptr_kinds_commutes)
lemma delete_shadow_root_pointers:
assumes "h ⊢ delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M shadow_root_ptr →⇩h h'"
shows "object_ptr_kinds h = object_ptr_kinds h' |∪| {|cast shadow_root_ptr|}"
using assms
apply(auto simp add: delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def split: option.splits)[1]
apply (metis (no_types, lifting) ObjectClass.a_type_wf_def ObjectClass.get⇩O⇩b⇩j⇩e⇩c⇩t_type_wf delete⇩O⇩b⇩j⇩e⇩c⇩t_def
delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_pointer_ptr_in_heap fmlookup_drop get⇩O⇩b⇩j⇩e⇩c⇩t_def heap.sel option.sel
shadow_root_ptr_kinds_commutes)
using delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_pointer_ptr_in_heap apply blast
by (metis (no_types, lifting) ObjectClass.a_type_wf_def ObjectClass.get⇩O⇩b⇩j⇩e⇩c⇩t_type_wf delete⇩O⇩b⇩j⇩e⇩c⇩t_def
delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_pointer_ptr_in_heap fmlookup_drop get⇩O⇩b⇩j⇩e⇩c⇩t_def heap.sel option.sel
shadow_root_ptr_kinds_commutes)
lemma delete_shadow_root_get_M⇩O⇩b⇩j⇩e⇩c⇩t: "h ⊢ delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M shadow_root_ptr →⇩h h' ⟹
ptr ≠ cast shadow_root_ptr ⟹ preserved (get_M⇩O⇩b⇩j⇩e⇩c⇩t ptr getter) h h'"
by(auto simp add: delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def delete⇩O⇩b⇩j⇩e⇩c⇩t_def ObjectMonad.get_M_defs preserved_def
split: prod.splits option.splits if_splits elim!: bind_returns_heap_E)
lemma delete_shadow_root_get_M⇩N⇩o⇩d⇩e: "h ⊢ delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M shadow_root_ptr →⇩h h' ⟹
preserved (get_M⇩N⇩o⇩d⇩e ptr getter) h h'"
by(auto simp add: delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def delete⇩O⇩b⇩j⇩e⇩c⇩t_def NodeMonad.get_M_defs ObjectMonad.get_M_defs
preserved_def split: prod.splits option.splits if_splits elim!: bind_returns_heap_E)
lemma delete_shadow_root_get_M⇩E⇩l⇩e⇩m⇩e⇩n⇩t: "h ⊢ delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M shadow_root_ptr →⇩h h' ⟹
preserved (get_M⇩E⇩l⇩e⇩m⇩e⇩n⇩t ptr getter) h h'"
by(auto simp add: delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def delete⇩O⇩b⇩j⇩e⇩c⇩t_def ElementMonad.get_M_defs NodeMonad.get_M_defs
ObjectMonad.get_M_defs preserved_def split: prod.splits option.splits if_splits elim!: bind_returns_heap_E)
lemma delete_shadow_root_get_M⇩C⇩h⇩a⇩r⇩a⇩c⇩t⇩e⇩r⇩D⇩a⇩t⇩a: "h ⊢ delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M shadow_root_ptr →⇩h h' ⟹
preserved (get_M⇩C⇩h⇩a⇩r⇩a⇩c⇩t⇩e⇩r⇩D⇩a⇩t⇩a ptr getter) h h'"
by(auto simp add: delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def delete⇩O⇩b⇩j⇩e⇩c⇩t_def CharacterDataMonad.get_M_defs
NodeMonad.get_M_defs ObjectMonad.get_M_defs preserved_def split: prod.splits option.splits if_splits
elim!: bind_returns_heap_E)
lemma delete_shadow_root_get_M⇩D⇩o⇩c⇩u⇩m⇩e⇩n⇩t: "h ⊢ delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M shadow_root_ptr →⇩h h' ⟹
preserved (get_M⇩D⇩o⇩c⇩u⇩m⇩e⇩n⇩t ptr getter) h h'"
by(auto simp add: delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def delete⇩O⇩b⇩j⇩e⇩c⇩t_def DocumentMonad.get_M_defs ObjectMonad.get_M_defs
preserved_def split: prod.splits option.splits if_splits elim!: bind_returns_heap_E)
lemma delete_shadow_root_get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t: "h ⊢ delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M shadow_root_ptr →⇩h h' ⟹
shadow_root_ptr ≠ shadow_root_ptr' ⟹ preserved (get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr' getter) h h'"
by(auto simp add: delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def delete⇩O⇩b⇩j⇩e⇩c⇩t_def get_M_defs ObjectMonad.get_M_defs
preserved_def split: prod.splits option.splits if_splits elim!: bind_returns_heap_E)
lemma shadow_root_put_get_1 [simp]: "shadow_root_ptr ≠ shadow_root_ptr' ⟹
h ⊢ put_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr setter v →⇩h h' ⟹
preserved (get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr' getter) h h'"
by(auto simp add: put_M_defs get_M_defs preserved_def split: option.splits dest: get_heap_E)
lemma shadow_root_put_get_2 [simp]: "(⋀x. getter (setter (λ_. v) x) = getter x) ⟹
h ⊢ put_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr setter v →⇩h h' ⟹
preserved (get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr' getter) h h'"
by (cases "shadow_root_ptr = shadow_root_ptr'") (auto simp add: put_M_defs get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma shadow_root_put_get_3 [simp]: "h ⊢ put_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr setter v →⇩h h' ⟹
preserved (get_M⇩E⇩l⇩e⇩m⇩e⇩n⇩t element_ptr getter) h h'"
by(auto simp add: put_M_defs ElementMonad.get_M_defs preserved_def split: option.splits
dest: get_heap_E)
lemma shadow_root_put_get_4 [simp]: "h ⊢ put_M⇩E⇩l⇩e⇩m⇩e⇩n⇩t element_ptr setter v →⇩h h' ⟹
preserved (get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr getter) h h'"
by(auto simp add: ElementMonad.put_M_defs get_M_defs preserved_def split: option.splits
dest: get_heap_E)
lemma shadow_root_put_get_5 [simp]: "h ⊢ put_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr setter v →⇩h h' ⟹
preserved (get_M⇩C⇩h⇩a⇩r⇩a⇩c⇩t⇩e⇩r⇩D⇩a⇩t⇩a character_data_ptr getter) h h'"
by(auto simp add: put_M_defs CharacterDataMonad.get_M_defs preserved_def split: option.splits
dest: get_heap_E)
lemma shadow_root_put_get_6 [simp]: "h ⊢ put_M⇩C⇩h⇩a⇩r⇩a⇩c⇩t⇩e⇩r⇩D⇩a⇩t⇩a character_data_ptr setter v →⇩h h' ⟹
preserved (get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr getter) h h'"
by(auto simp add: CharacterDataMonad.put_M_defs get_M_defs preserved_def split: option.splits
dest: get_heap_E)
lemma shadow_root_put_get_7 [simp]: "h ⊢ put_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr setter v →⇩h h' ⟹
preserved (get_M⇩D⇩o⇩c⇩u⇩m⇩e⇩n⇩t document_ptr getter) h h'"
by(auto simp add: put_M_defs DocumentMonad.get_M_defs preserved_def split: option.splits
dest: get_heap_E)
lemma shadow_root_put_get_8 [simp]: "h ⊢ put_M⇩D⇩o⇩c⇩u⇩m⇩e⇩n⇩t document_ptr setter v →⇩h h' ⟹
preserved (get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr getter) h h'"
by(auto simp add: DocumentMonad.put_M_defs get_M_defs preserved_def split: option.splits
dest: get_heap_E)
lemma shadow_root_put_get_9 [simp]: "(⋀x. getter (cast (setter (λ_. v) x)) = getter (cast x)) ⟹
h ⊢ put_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr setter v →⇩h h' ⟹ preserved (get_M⇩O⇩b⇩j⇩e⇩c⇩t object_ptr getter) h h'"
by (cases "cast shadow_root_ptr = object_ptr") (auto simp add: put_M_defs get_M_defs
ObjectMonad.get_M_defs NodeMonad.get_M_defs get⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def get⇩N⇩o⇩d⇩e_def preserved_def put⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def
put⇩N⇩o⇩d⇩e_def bind_eq_Some_conv split: option.splits)
subsection ‹new\_M›
definition new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M :: "(_, (_) shadow_root_ptr) dom_prog"
where
"new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M = do {
h ← get_heap;
(new_ptr, h') ← return (new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t h);
return_heap h';
return new_ptr
}"
lemma new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_ok [simp]:
"h ⊢ ok new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M"
by(auto simp add: new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def split: prod.splits)
lemma new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_ptr_in_heap:
assumes "h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩h h'"
and "h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩r new_shadow_root_ptr"
shows "new_shadow_root_ptr |∈| shadow_root_ptr_kinds h'"
using assms
unfolding new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def
by(auto simp add: new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def Let_def put⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_ptr_in_heap is_OK_returns_result_I
elim!: bind_returns_result_E bind_returns_heap_E)
lemma new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_ptr_not_in_heap:
assumes "h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩h h'"
and "h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩r new_shadow_root_ptr"
shows "new_shadow_root_ptr |∉| shadow_root_ptr_kinds h"
using assms new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_ptr_not_in_heap
by(auto simp add: new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def split: prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_new_ptr:
assumes "h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩h h'"
and "h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩r new_shadow_root_ptr"
shows "object_ptr_kinds h' = object_ptr_kinds h |∪| {|cast new_shadow_root_ptr|}"
using assms new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_new_ptr
by(auto simp add: new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def split: prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_is_shadow_root_ptr:
assumes "h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩r new_shadow_root_ptr"
shows "is_shadow_root_ptr new_shadow_root_ptr"
using assms new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_is_shadow_root_ptr
by(auto simp add: new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def elim!: bind_returns_result_E split: prod.splits)
lemma new_shadow_root_mode:
assumes "h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩h h'"
assumes "h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩r new_shadow_root_ptr"
shows "h' ⊢ get_M new_shadow_root_ptr mode →⇩r Open"
using assms
by(auto simp add: get_M_defs new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_shadow_root_children:
assumes "h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩h h'"
assumes "h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩r new_shadow_root_ptr"
shows "h' ⊢ get_M new_shadow_root_ptr child_nodes →⇩r []"
using assms
by(auto simp add: get_M_defs new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def Let_def
split: option.splits prod.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_shadow_root_get_M⇩O⇩b⇩j⇩e⇩c⇩t:
"h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩h h' ⟹ h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩r new_shadow_root_ptr
⟹ ptr ≠ cast new_shadow_root_ptr ⟹ preserved (get_M⇩O⇩b⇩j⇩e⇩c⇩t ptr getter) h h'"
by(auto simp add: new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def ObjectMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_shadow_root_get_M⇩N⇩o⇩d⇩e:
"h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩h h' ⟹ h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩r new_shadow_root_ptr
⟹ preserved (get_M⇩N⇩o⇩d⇩e ptr getter) h h'"
by(auto simp add: new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def NodeMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_shadow_root_get_M⇩E⇩l⇩e⇩m⇩e⇩n⇩t:
"h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩h h' ⟹ h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩r new_shadow_root_ptr
⟹ preserved (get_M⇩E⇩l⇩e⇩m⇩e⇩n⇩t ptr getter) h h'"
by(auto simp add: new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def ElementMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_shadow_root_get_M⇩C⇩h⇩a⇩r⇩a⇩c⇩t⇩e⇩r⇩D⇩a⇩t⇩a:
"h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩h h' ⟹ h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩r new_shadow_root_ptr
⟹ preserved (get_M⇩C⇩h⇩a⇩r⇩a⇩c⇩t⇩e⇩r⇩D⇩a⇩t⇩a ptr getter) h h'"
by(auto simp add: new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def CharacterDataMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_shadow_root_get_M⇩D⇩o⇩c⇩u⇩m⇩e⇩n⇩t:
"h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩h h'
⟹ h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩r new_shadow_root_ptr
⟹ preserved (get_M⇩D⇩o⇩c⇩u⇩m⇩e⇩n⇩t ptr getter) h h'"
by(auto simp add: new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def DocumentMonad.get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
lemma new_shadow_root_get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t:
"h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩h h'
⟹ h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩r new_shadow_root_ptr ⟹ ptr ≠ new_shadow_root_ptr
⟹ preserved (get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t ptr getter) h h'"
by(auto simp add: new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def get_M_defs preserved_def
split: prod.splits option.splits elim!: bind_returns_result_E bind_returns_heap_E)
subsection ‹modified heaps›
lemma shadow_root_get_put_1 [simp]: "get⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr (put⇩O⇩b⇩j⇩e⇩c⇩t ptr obj h) = (if ptr = cast shadow_root_ptr
then cast obj else get shadow_root_ptr h)"
by(auto simp add: get⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def split: option.splits Option.bind_splits)
lemma shadow_root_ptr_kinds_new[simp]: "shadow_root_ptr_kinds (put⇩O⇩b⇩j⇩e⇩c⇩t ptr obj h) = shadow_root_ptr_kinds h |∪|
(if is_shadow_root_ptr_kind ptr then {|the (cast ptr)|} else {||})"
by(auto simp add: shadow_root_ptr_kinds_def split: option.splits)
lemma type_wf_put_I:
assumes "type_wf h"
assumes "DocumentClass.type_wf (put⇩O⇩b⇩j⇩e⇩c⇩t ptr obj h)"
assumes "is_shadow_root_ptr_kind ptr ⟹ is_shadow_root_kind obj"
shows "type_wf (put⇩O⇩b⇩j⇩e⇩c⇩t ptr obj h)"
using assms
by(auto simp add: type_wf_defs is_shadow_root_kind_def split: option.splits)
lemma type_wf_put_ptr_not_in_heap_E:
assumes "type_wf (put⇩O⇩b⇩j⇩e⇩c⇩t ptr obj h)"
assumes "ptr |∉| object_ptr_kinds h"
shows "type_wf h"
using assms
by(auto simp add: type_wf_defs elim!: DocumentMonad.type_wf_put_ptr_not_in_heap_E
split: option.splits if_splits)
lemma type_wf_put_ptr_in_heap_E:
assumes "type_wf (put⇩O⇩b⇩j⇩e⇩c⇩t ptr obj h)"
assumes "ptr |∈| object_ptr_kinds h"
assumes "DocumentClass.type_wf h"
assumes "is_shadow_root_ptr_kind ptr ⟹ is_shadow_root_kind (the (get ptr h))"
shows "type_wf h"
using assms
apply(auto simp add: type_wf_defs elim!: DocumentMonad.type_wf_put_ptr_in_heap_E
split: option.splits if_splits)[1]
by (metis (no_types, opaque_lifting) ObjectClass.a_type_wf_def ObjectClass.get⇩O⇩b⇩j⇩e⇩c⇩t_type_wf
bind.bind_lunit get⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def is_shadow_root_kind_def option.exhaust_sel)
subsection ‹type\_wf›
lemma new_element_type_wf_preserved [simp]:
assumes "h ⊢ new_element →⇩h h'"
shows "type_wf h = type_wf h'"
proof -
obtain new_element_ptr where "h ⊢ new_element →⇩r new_element_ptr"
using assms
by (meson is_OK_returns_heap_I is_OK_returns_result_E)
with assms have "object_ptr_kinds h' = object_ptr_kinds h |∪| {|cast new_element_ptr|}"
using new_element_new_ptr by auto
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
unfolding shadow_root_ptr_kinds_def by auto
with assms show ?thesis
by(auto simp add: ElementMonad.new_element_def type_wf_defs Let_def elim!: bind_returns_heap_E
split: prod.splits)
qed
lemma put_M⇩E⇩l⇩e⇩m⇩e⇩n⇩t_tag_name_type_wf_preserved [simp]:
assumes "h ⊢ put_M element_ptr tag_name_update v →⇩h h'"
shows "type_wf h = type_wf h'"
proof -
have "object_ptr_kinds h = object_ptr_kinds h'"
using writes_singleton assms object_ptr_kinds_preserved unfolding all_args_def by fastforce
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
unfolding shadow_root_ptr_kinds_def by simp
with assms show ?thesis
by(auto simp add: ElementMonad.put_M_defs type_wf_defs)
qed
lemma put_M⇩E⇩l⇩e⇩m⇩e⇩n⇩t_child_nodes_type_wf_preserved [simp]:
assumes "h ⊢ put_M element_ptr RElement.child_nodes_update v →⇩h h'"
shows "type_wf h = type_wf h'"
proof -
have "object_ptr_kinds h = object_ptr_kinds h'"
using writes_singleton assms object_ptr_kinds_preserved unfolding all_args_def by fastforce
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
unfolding shadow_root_ptr_kinds_def by simp
with assms show ?thesis
by(auto simp add: ElementMonad.put_M_defs type_wf_defs)
qed
lemma put_M⇩E⇩l⇩e⇩m⇩e⇩n⇩t_attrs_type_wf_preserved [simp]:
assumes "h ⊢ put_M element_ptr attrs_update v →⇩h h'"
shows "type_wf h = type_wf h'"
proof -
have "object_ptr_kinds h = object_ptr_kinds h'"
using writes_singleton assms object_ptr_kinds_preserved unfolding all_args_def by fastforce
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
unfolding shadow_root_ptr_kinds_def by simp
with assms show ?thesis
by(auto simp add: ElementMonad.put_M_defs type_wf_defs)
qed
lemma put_M⇩E⇩l⇩e⇩m⇩e⇩n⇩t_shadow_root_opt_type_wf_preserved [simp]:
assumes "h ⊢ put_M element_ptr shadow_root_opt_update v →⇩h h'"
shows "type_wf h = type_wf h'"
proof -
have "object_ptr_kinds h = object_ptr_kinds h'"
using writes_singleton assms object_ptr_kinds_preserved unfolding all_args_def by fastforce
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
unfolding shadow_root_ptr_kinds_def by simp
with assms show ?thesis
by(auto simp add: ElementMonad.put_M_defs type_wf_defs)
qed
lemma new_character_data_type_wf_preserved [simp]:
assumes "h ⊢ new_character_data →⇩h h'"
shows "type_wf h = type_wf h'"
proof -
obtain new_character_data_ptr where "h ⊢ new_character_data →⇩r new_character_data_ptr"
using assms
by (meson is_OK_returns_heap_I is_OK_returns_result_E)
with assms have "object_ptr_kinds h' = object_ptr_kinds h |∪| {|cast new_character_data_ptr|}"
using new_character_data_new_ptr by auto
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
unfolding shadow_root_ptr_kinds_def by auto
with assms show ?thesis
by(auto simp add: CharacterDataMonad.new_character_data_def type_wf_defs Let_def
elim!: bind_returns_heap_E split: prod.splits)
qed
lemma put_M⇩C⇩h⇩a⇩r⇩a⇩c⇩t⇩e⇩r⇩D⇩a⇩t⇩a_val_type_wf_preserved [simp]:
assumes "h ⊢ put_M character_data_ptr val_update v →⇩h h'"
shows "type_wf h = type_wf h'"
proof -
have "object_ptr_kinds h = object_ptr_kinds h'"
using writes_singleton assms object_ptr_kinds_preserved unfolding all_args_def by fastforce
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
unfolding shadow_root_ptr_kinds_def by simp
with assms show ?thesis
by(auto simp add: CharacterDataMonad.put_M_defs type_wf_defs)
qed
lemma new_document_type_wf_preserved [simp]:
assumes "h ⊢ new_document →⇩h h'"
shows "type_wf h = type_wf h'"
proof -
obtain new_document_ptr where "h ⊢ new_document →⇩r new_document_ptr"
using assms
by (meson is_OK_returns_heap_I is_OK_returns_result_E)
with assms have "object_ptr_kinds h' = object_ptr_kinds h |∪| {|cast new_document_ptr|}"
using new_document_new_ptr by auto
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
unfolding shadow_root_ptr_kinds_def by auto
with assms show ?thesis
by(auto simp add: DocumentMonad.new_document_def type_wf_defs Let_def elim!: bind_returns_heap_E
split: prod.splits)
qed
lemma put_M⇩D⇩o⇩c⇩u⇩m⇩e⇩n⇩t_doctype_type_wf_preserved [simp]:
assumes "h ⊢ put_M document_ptr doctype_update v →⇩h h'"
shows "type_wf h = type_wf h'"
proof -
have "object_ptr_kinds h = object_ptr_kinds h'"
using writes_singleton assms object_ptr_kinds_preserved unfolding all_args_def by fastforce
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
unfolding shadow_root_ptr_kinds_def by simp
with assms show ?thesis
by(auto simp add: DocumentMonad.put_M_defs type_wf_defs)
qed
lemma put_M⇩D⇩o⇩c⇩u⇩m⇩e⇩n⇩t_document_element_type_wf_preserved [simp]:
assumes "h ⊢ put_M document_ptr document_element_update v →⇩h h'"
shows "type_wf h = type_wf h'"
proof -
have "object_ptr_kinds h = object_ptr_kinds h'"
using writes_singleton assms object_ptr_kinds_preserved unfolding all_args_def by fastforce
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
unfolding shadow_root_ptr_kinds_def by simp
with assms show ?thesis
by(auto simp add: DocumentMonad.put_M_defs type_wf_defs)
qed
lemma put_M⇩D⇩o⇩c⇩u⇩m⇩e⇩n⇩t_disconnected_nodes_type_wf_preserved [simp]:
assumes "h ⊢ put_M document_ptr disconnected_nodes_update v →⇩h h'"
shows "type_wf h = type_wf h'"
proof -
have "object_ptr_kinds h = object_ptr_kinds h'"
using writes_singleton assms object_ptr_kinds_preserved unfolding all_args_def by fastforce
then have "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
unfolding shadow_root_ptr_kinds_def by simp
with assms show ?thesis
by(auto simp add: DocumentMonad.put_M_defs type_wf_defs)
qed
lemma put_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_mode_type_wf_preserved [simp]:
"h ⊢ put_M shadow_root_ptr mode_update v →⇩h h' ⟹ type_wf h = type_wf h'"
by(auto simp add: get_M_defs is_shadow_root_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 DocumentClass.type_wf_defs put_M_defs
put⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I DocumentMonad.type_wf_put_I CharacterDataMonad.type_wf_put_I
ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I
split: option.splits)
lemma put_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_child_nodes_type_wf_preserved [simp]:
"h ⊢ put_M shadow_root_ptr RShadowRoot.child_nodes_update v →⇩h h' ⟹ type_wf h = type_wf h'"
by(auto simp add: get_M_defs is_shadow_root_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 DocumentClass.type_wf_defs put_M_defs
put⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def
dest!: get_heap_E
elim!: bind_returns_heap_E2
intro!: type_wf_put_I DocumentMonad.type_wf_put_I CharacterDataMonad.type_wf_put_I
ElementMonad.type_wf_put_I NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I
split: option.splits)
lemma shadow_root_ptr_kinds_small:
assumes "⋀object_ptr. preserved (get_M⇩O⇩b⇩j⇩e⇩c⇩t object_ptr RObject.nothing) h h'"
shows "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
by(simp add: shadow_root_ptr_kinds_def preserved_def object_ptr_kinds_preserved_small[OF assms])
lemma shadow_root_ptr_kinds_preserved:
assumes "writes SW setter h h'"
assumes "h ⊢ setter →⇩h h'"
assumes "⋀h h'. ∀w ∈ SW. h ⊢ w →⇩h h' ⟶
(∀object_ptr. preserved (get_M⇩O⇩b⇩j⇩e⇩c⇩t object_ptr RObject.nothing) h h')"
shows "shadow_root_ptr_kinds h = shadow_root_ptr_kinds h'"
using writes_small_big[OF assms]
apply(simp add: reflp_def transp_def preserved_def shadow_root_ptr_kinds_def)
by (metis assms object_ptr_kinds_preserved)
lemma new_shadow_root_known_ptr:
assumes "h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩r new_shadow_root_ptr"
shows "known_ptr (cast new_shadow_root_ptr)"
using assms
apply(auto simp add: new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def Let_def a_known_ptr_def
elim!: bind_returns_result_E2 split: prod.splits)[1]
using assms new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_is_shadow_root_ptr by blast
lemma new_shadow_root_type_wf_preserved [simp]: "h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩h h' ⟹ type_wf h = type_wf h'"
apply(auto simp add: new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def Let_def put⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def put⇩D⇩o⇩c⇩u⇩m⇩e⇩n⇩t_def
ShadowRootClass.type_wf⇩D⇩o⇩c⇩u⇩m⇩e⇩n⇩t ShadowRootClass.type_wf⇩C⇩h⇩a⇩r⇩a⇩c⇩t⇩e⇩r⇩D⇩a⇩t⇩a ShadowRootClass.type_wf⇩E⇩l⇩e⇩m⇩e⇩n⇩t
ShadowRootClass.type_wf⇩N⇩o⇩d⇩e ShadowRootClass.type_wf⇩O⇩b⇩j⇩e⇩c⇩t
is_node_ptr_kind_none new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_ptr_not_in_heap
elim!: bind_returns_heap_E type_wf_put_ptr_not_in_heap_E
intro!: type_wf_put_I DocumentMonad.type_wf_put_I ElementMonad.type_wf_put_I
CharacterDataMonad.type_wf_put_I
NodeMonad.type_wf_put_I ObjectMonad.type_wf_put_I
split: if_splits)[1]
by(auto simp add: type_wf_defs DocumentClass.type_wf_defs ElementClass.type_wf_defs
CharacterDataClass.type_wf_defs
NodeClass.type_wf_defs ObjectClass.type_wf_defs is_shadow_root_kind_def is_document_kind_def
split: option.splits)[1]
locale l_new_shadow_root = l_type_wf +
assumes new_shadow_root_types_preserved: "h ⊢ new⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M →⇩h h' ⟹ type_wf h = type_wf h'"
lemma new_shadow_root_is_l_new_shadow_root [instances]: "l_new_shadow_root type_wf"
using l_new_shadow_root.intro new_shadow_root_type_wf_preserved
by blast
lemma type_wf_preserved_small:
assumes "⋀object_ptr. preserved (get_M⇩O⇩b⇩j⇩e⇩c⇩t object_ptr RObject.nothing) h h'"
assumes "⋀node_ptr. preserved (get_M⇩N⇩o⇩d⇩e node_ptr RNode.nothing) h h'"
assumes "⋀element_ptr. preserved (get_M⇩E⇩l⇩e⇩m⇩e⇩n⇩t element_ptr RElement.nothing) h h'"
assumes "⋀character_data_ptr. preserved (get_M⇩C⇩h⇩a⇩r⇩a⇩c⇩t⇩e⇩r⇩D⇩a⇩t⇩a character_data_ptr RCharacterData.nothing) h h'"
assumes "⋀document_ptr. preserved (get_M⇩D⇩o⇩c⇩u⇩m⇩e⇩n⇩t document_ptr RDocument.nothing) h h'"
assumes "⋀shadow_root_ptr. preserved (get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr RShadowRoot.nothing) h h'"
shows "type_wf h = type_wf h'"
using type_wf_preserved_small[OF assms(1) assms(2) assms(3) assms(4) assms(5)]
allI[OF assms(6), of id, simplified] shadow_root_ptr_kinds_small[OF assms(1)]
apply(auto simp add: type_wf_defs preserved_def get_M_defs shadow_root_ptr_kinds_small[OF assms(1)]
split: option.splits)[1]
apply(force)
apply(force)
done
lemma new_element_is_l_new_element [instances]:
"l_new_element type_wf"
using l_new_element.intro new_element_type_wf_preserved
by blast
lemma new_character_data_is_l_new_character_data [instances]:
"l_new_character_data type_wf"
using l_new_character_data.intro new_character_data_type_wf_preserved
by blast
lemma new_document_is_l_new_document [instances]:
"l_new_document type_wf"
using l_new_document.intro new_document_type_wf_preserved
by blast
lemma type_wf_preserved:
assumes "writes SW setter h h'"
assumes "h ⊢ setter →⇩h h'"
assumes "⋀h h' w. w ∈ SW ⟹ h ⊢ w →⇩h h' ⟹
∀object_ptr. preserved (get_M⇩O⇩b⇩j⇩e⇩c⇩t object_ptr RObject.nothing) h h'"
assumes "⋀h h' w. w ∈ SW ⟹ h ⊢ w →⇩h h' ⟹
∀node_ptr. preserved (get_M⇩N⇩o⇩d⇩e node_ptr RNode.nothing) h h'"
assumes "⋀h h' w. w ∈ SW ⟹ h ⊢ w →⇩h h' ⟹
∀element_ptr. preserved (get_M⇩E⇩l⇩e⇩m⇩e⇩n⇩t element_ptr RElement.nothing) h h'"
assumes "⋀h h' w. w ∈ SW ⟹ h ⊢ w →⇩h h' ⟹
∀character_data_ptr. preserved (get_M⇩C⇩h⇩a⇩r⇩a⇩c⇩t⇩e⇩r⇩D⇩a⇩t⇩a character_data_ptr RCharacterData.nothing) h h'"
assumes "⋀h h' w. w ∈ SW ⟹ h ⊢ w →⇩h h' ⟹
∀document_ptr. preserved (get_M⇩D⇩o⇩c⇩u⇩m⇩e⇩n⇩t document_ptr RDocument.nothing) h h'"
assumes "⋀h h' w. w ∈ SW ⟹ h ⊢ w →⇩h h' ⟹
∀shadow_root_ptr. preserved (get_M⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t shadow_root_ptr RShadowRoot.nothing) h h'"
shows "type_wf h = type_wf h'"
proof -
have "⋀h h' w. w ∈ SW ⟹ h ⊢ w →⇩h h' ⟹ type_wf h = type_wf h'"
using assms type_wf_preserved_small by fast
with assms(1) assms(2) show ?thesis
apply(rule writes_small_big)
by(auto simp add: reflp_def transp_def)
qed
lemma type_wf_drop: "type_wf h ⟹ type_wf (Heap (fmdrop ptr (the_heap h)))"
apply(auto simp add: type_wf_defs)[1]
using type_wf_drop
apply blast
by (metis (no_types, opaque_lifting) DocumentClass.type_wf⇩E⇩l⇩e⇩m⇩e⇩n⇩t ElementClass.get⇩O⇩b⇩j⇩e⇩c⇩t_type_wf
ElementMonad.type_wf_drop fmlookup_drop get⇩O⇩b⇩j⇩e⇩c⇩t_def get⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def
object_ptr_kinds_code5 shadow_root_ptr_kinds_commutes)
lemma delete_shadow_root_type_wf_preserved [simp]:
assumes "h ⊢ delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M shadow_root_ptr →⇩h h'"
assumes "type_wf h"
shows "type_wf h'"
using assms
using type_wf_drop
by(auto simp add: delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_M_def delete⇩S⇩h⇩a⇩d⇩o⇩w⇩R⇩o⇩o⇩t_def delete⇩O⇩b⇩j⇩e⇩c⇩t_def split: if_splits)
end