Theory NodeMonad
section‹Node›
text‹In this theory, we introduce the monadic method setup for the Node class.›
theory NodeMonad
imports
ObjectMonad
"../classes/NodeClass"
begin
type_synonym ('object_ptr, 'node_ptr, 'Object, 'Node, 'result) dom_prog
= "((_) heap, exception, 'result) prog"
register_default_tvars "('object_ptr, 'node_ptr, 'Object, 'Node, 'result) dom_prog"
global_interpretation l_ptr_kinds_M node_ptr_kinds defines node_ptr_kinds_M = a_ptr_kinds_M .
lemmas node_ptr_kinds_M_defs = a_ptr_kinds_M_def
lemma node_ptr_kinds_M_eq:
assumes "|h ⊢ object_ptr_kinds_M|⇩r = |h' ⊢ object_ptr_kinds_M|⇩r"
shows "|h ⊢ node_ptr_kinds_M|⇩r = |h' ⊢ node_ptr_kinds_M|⇩r"
using assms
by(auto simp add: node_ptr_kinds_M_defs object_ptr_kinds_M_defs node_ptr_kinds_def)
global_interpretation l_dummy defines get_M⇩N⇩o⇩d⇩e = "l_get_M.a_get_M get⇩N⇩o⇩d⇩e" .
lemma get_M_is_l_get_M: "l_get_M get⇩N⇩o⇩d⇩e type_wf node_ptr_kinds"
apply(simp add: get⇩N⇩o⇩d⇩e_type_wf l_get_M_def)
by (metis ObjectClass.a_type_wf_def ObjectClass.get⇩O⇩b⇩j⇩e⇩c⇩t_type_wf bind_eq_None_conv get⇩N⇩o⇩d⇩e_def
node_ptr_kinds_commutes option.simps(3))
lemmas get_M_defs = get_M⇩N⇩o⇩d⇩e_def[unfolded l_get_M.a_get_M_def[OF get_M_is_l_get_M]]
adhoc_overloading get_M get_M⇩N⇩o⇩d⇩e
locale l_get_M⇩N⇩o⇩d⇩e_lemmas = l_type_wf⇩N⇩o⇩d⇩e
begin
sublocale l_get_M⇩O⇩b⇩j⇩e⇩c⇩t_lemmas by unfold_locales
interpretation l_get_M get⇩N⇩o⇩d⇩e type_wf node_ptr_kinds
apply(unfold_locales)
apply (simp add: get⇩N⇩o⇩d⇩e_type_wf local.type_wf⇩N⇩o⇩d⇩e)
by (meson NodeMonad.get_M_is_l_get_M l_get_M_def)
lemmas get_M⇩N⇩o⇩d⇩e_ok = get_M_ok[folded get_M⇩N⇩o⇩d⇩e_def]
end
global_interpretation l_get_M⇩N⇩o⇩d⇩e_lemmas type_wf by unfold_locales
lemma node_ptr_kinds_M_reads:
"reads (⋃object_ptr. {preserved (get_M⇩O⇩b⇩j⇩e⇩c⇩t object_ptr RObject.nothing)}) node_ptr_kinds_M h h'"
using object_ptr_kinds_M_reads
apply (simp add: reads_def node_ptr_kinds_M_defs node_ptr_kinds_def
object_ptr_kinds_M_reads preserved_def)
by (metis (mono_tags, lifting) object_ptr_kinds_preserved_small old.unit.exhaust preserved_def)
global_interpretation l_put_M type_wf node_ptr_kinds get⇩N⇩o⇩d⇩e put⇩N⇩o⇩d⇩e
rewrites "a_get_M = get_M⇩N⇩o⇩d⇩e"
defines put_M⇩N⇩o⇩d⇩e = a_put_M
apply (simp add: get_M_is_l_get_M l_put_M_def)
by (simp add: get_M⇩N⇩o⇩d⇩e_def)
lemmas put_M_defs = a_put_M_def
adhoc_overloading put_M put_M⇩N⇩o⇩d⇩e
locale l_put_M⇩N⇩o⇩d⇩e_lemmas = l_type_wf⇩N⇩o⇩d⇩e
begin
sublocale l_put_M⇩O⇩b⇩j⇩e⇩c⇩t_lemmas by unfold_locales
interpretation l_put_M type_wf node_ptr_kinds get⇩N⇩o⇩d⇩e put⇩N⇩o⇩d⇩e
apply(unfold_locales)
apply (simp add: get⇩N⇩o⇩d⇩e_type_wf local.type_wf⇩N⇩o⇩d⇩e)
by (meson NodeMonad.get_M_is_l_get_M l_get_M_def)
lemmas put_M⇩N⇩o⇩d⇩e_ok = put_M_ok[folded put_M⇩N⇩o⇩d⇩e_def]
end
global_interpretation l_put_M⇩N⇩o⇩d⇩e_lemmas type_wf by unfold_locales
lemma get_M_Object_preserved1 [simp]:
"(⋀x. getter (cast (setter (λ_. v) x)) = getter (cast x)) ⟹ h ⊢ put_M⇩N⇩o⇩d⇩e node_ptr setter v →⇩h h'
⟹ preserved (get_M⇩O⇩b⇩j⇩e⇩c⇩t object_ptr getter) h h'"
apply(cases "cast node_ptr = object_ptr")
by(auto simp add: put_M_defs get_M_defs ObjectMonad.get_M_defs get⇩N⇩o⇩d⇩e_def preserved_def put⇩N⇩o⇩d⇩e_def
bind_eq_Some_conv
split: option.splits)
lemma get_M_Object_preserved2 [simp]:
"cast node_ptr ≠ object_ptr ⟹ h ⊢ put_M⇩N⇩o⇩d⇩e node_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⇩N⇩o⇩d⇩e_def put⇩N⇩o⇩d⇩e_def ObjectMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
lemma get_M_Object_preserved3 [simp]:
"h ⊢ put_M⇩N⇩o⇩d⇩e node_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 node_ptr ≠ object_ptr")
by(auto simp add: put_M_defs get_M_defs get⇩N⇩o⇩d⇩e_def put⇩N⇩o⇩d⇩e_def ObjectMonad.get_M_defs preserved_def
split: option.splits bind_splits dest: get_heap_E)
lemma get_M_Object_preserved4 [simp]:
"cast node_ptr ≠ object_ptr ⟹ h ⊢ put_M⇩O⇩b⇩j⇩e⇩c⇩t object_ptr setter v →⇩h h'
⟹ preserved (get_M⇩N⇩o⇩d⇩e node_ptr getter) h h'"
by(auto simp add: ObjectMonad.put_M_defs get_M_defs get⇩N⇩o⇩d⇩e_def ObjectMonad.get_M_defs preserved_def
split: option.splits dest: get_heap_E)
subsection‹Modified Heaps›
lemma get_node_ptr_simp [simp]:
"get⇩N⇩o⇩d⇩e node_ptr (put⇩O⇩b⇩j⇩e⇩c⇩t ptr obj h) = (if ptr = cast node_ptr then cast obj else get node_ptr h)"
by(auto simp add: get⇩N⇩o⇩d⇩e_def)
lemma node_ptr_kinds_simp [simp]:
"node_ptr_kinds (put⇩O⇩b⇩j⇩e⇩c⇩t ptr obj h)
= node_ptr_kinds h |∪| (if is_node_ptr_kind ptr then {|the (cast ptr)|} else {||})"
by(auto simp add: node_ptr_kinds_def)
lemma type_wf_put_I:
assumes "type_wf h"
assumes "ObjectClass.type_wf (put⇩O⇩b⇩j⇩e⇩c⇩t ptr obj h)"
assumes "is_node_ptr_kind ptr ⟹ is_node_kind obj"
shows "type_wf (put⇩O⇩b⇩j⇩e⇩c⇩t ptr obj h)"
using assms
apply(auto simp add: type_wf_defs split: option.splits)[1]
using cast⇩O⇩b⇩j⇩e⇩c⇩t⇩2⇩N⇩o⇩d⇩e_none is_node_kind_def apply blast
using cast⇩O⇩b⇩j⇩e⇩c⇩t⇩2⇩N⇩o⇩d⇩e_none is_node_kind_def apply blast
done
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!: ObjectMonad.type_wf_put_ptr_not_in_heap_E
split: option.splits if_splits)
lemma type_wf_put_ptr_in_heap_E:
assumes "type_wf (put⇩O⇩b⇩j⇩e⇩c⇩t ptr obj h)"
assumes "ptr |∈| object_ptr_kinds h"
assumes "ObjectClass.type_wf h"
assumes "is_node_ptr_kind ptr ⟹ is_node_kind (the (get ptr h))"
shows "type_wf h"
using assms
apply(auto simp add: type_wf_defs split: option.splits if_splits)[1]
by (metis ObjectClass.get⇩O⇩b⇩j⇩e⇩c⇩t_type_wf bind.bind_lunit get⇩N⇩o⇩d⇩e_def is_node_kind_def option.exhaust_sel)
subsection‹Preserving Types›
lemma node_ptr_kinds_small:
assumes "⋀object_ptr. preserved (get_M⇩O⇩b⇩j⇩e⇩c⇩t object_ptr RObject.nothing) h h'"
shows "node_ptr_kinds h = node_ptr_kinds h'"
by(simp add: node_ptr_kinds_def preserved_def object_ptr_kinds_preserved_small[OF assms])
lemma node_ptr_kinds_preserved:
assumes "writes SW setter h h'"
assumes "h ⊢ setter →⇩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 "node_ptr_kinds h = node_ptr_kinds h'"
using writes_small_big[OF assms]
apply(simp add: reflp_def transp_def preserved_def node_ptr_kinds_def)
by (metis assms object_ptr_kinds_preserved)
lemma type_wf_preserved_small:
assumes "⋀object_ptr. preserved (get_M⇩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'"
shows "type_wf h = type_wf h'"
using type_wf_preserved allI[OF assms(2), of id, simplified]
apply(auto simp add: type_wf_defs)[1]
apply(auto simp add: preserved_def get_M_defs node_ptr_kinds_small[OF assms(1)]
split: option.splits)[1]
apply (metis option.simps(3))
by(auto simp add: preserved_def get_M_defs node_ptr_kinds_small[OF assms(1)]
split: option.splits, force)[1]
lemma type_wf_preserved:
assumes "writes SW setter h h'"
assumes "h ⊢ setter →⇩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'"
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
end