Theory Mask
subsection ‹Permission masks: Maps from heap locations to permission amounts›
theory Mask
imports PosRat
begin
subsubsection ‹Definitions›
type_synonym field = string
type_synonym address = nat
type_synonym heap_loc = "address × field"
type_synonym mask = "heap_loc ⇒ prat"
type_synonym bmask = "heap_loc ⇒ bool"
definition null where "null = 0"
definition full_mask :: "mask" where
"full_mask hl = (if fst hl = null then pnone else pwrite)"
definition multiply_mask :: "prat ⇒ mask ⇒ mask" where
"multiply_mask p π hl = pmult p (π hl)"
fun empty_mask where
"empty_mask hl = pnone"
fun empty_bmask where
"empty_bmask hl = False"
fun add_acc where "add_acc π hl p = π(hl := padd (π hl) p)"
inductive rm_acc where
"π hl = padd p r ⟹ rm_acc π hl p (π(hl := r))"
fun add_masks where
"add_masks π' π hl = padd (π' hl) (π hl)"
definition greater_mask where
"greater_mask π' π ⟷ (∃r. π' = add_masks π r)"
fun uni_mask where
"uni_mask hl p = empty_mask(hl := p)"
fun valid_mask :: "mask ⇒ bool" where
"valid_mask π ⟷ (∀hl. pgte pwrite (π hl)) ∧ (∀f. π (null, f) = pnone)"
definition valid_null :: "mask ⇒ bool" where
"valid_null π ⟷ (∀f. π (null, f) = pnone)"
definition equal_on_mask where
"equal_on_mask π h h' ⟷ (∀hl. ppos (π hl) ⟶ h hl = h' hl)"
definition equal_on_bmask where
"equal_on_bmask π h h' ⟷ (∀hl. π hl ⟶ h hl = h' hl)"
definition big_add_masks where
"big_add_masks Π Π' h = add_masks (Π h) (Π' h)"
definition big_greater_mask where
"big_greater_mask Π Π' ⟷ (∀h. greater_mask (Π h) (Π' h))"
definition greater_bmask where
"greater_bmask H H' ⟷ (∀h. H' h ⟶ H h)"
definition update_dm where
"update_dm dm π π' hl ⟷ (dm hl ∨ pgt (π hl) (π' hl))"
fun pre_get_m where "pre_get_m φ = fst φ"
fun pre_get_h where "pre_get_h φ = snd φ"
fun srm_acc where "srm_acc φ hl p = (rm_acc (pre_get_m φ) hl p, pre_get_h φ)"