Theory JMM_Type
section ‹JMM heap implementation 1›
theory JMM_Type
imports
"../Common/ExternalCallWF"
"../Common/ConformThreaded"
JMM_Heap
begin
subsection ‹Definitions›
text ‹
The JMM heap only stores type information.
›
type_synonym 'addr JMM_heap = "'addr ⇀ htype"
translations (type) "'addr JMM_heap" <= (type) "'addr ⇒ htype option"
abbreviation jmm_empty :: "'addr JMM_heap" where "jmm_empty == Map.empty"
definition jmm_allocate :: "'addr JMM_heap ⇒ htype ⇒ ('addr JMM_heap × 'addr) set"
where "jmm_allocate h hT = (λa. (h(a ↦ hT), a)) ` {a. h a = None}"
definition jmm_typeof_addr :: "'addr JMM_heap ⇒ 'addr ⇀ htype"
where "jmm_typeof_addr h = h"
definition jmm_heap_read :: "'addr JMM_heap ⇒ 'addr ⇒ addr_loc ⇒ 'addr val ⇒ bool"
where "jmm_heap_read h a ad v = True"
context
notes [[inductive_internals]]
begin
inductive jmm_heap_write :: "'addr JMM_heap ⇒ 'addr ⇒ addr_loc ⇒ 'addr val ⇒ 'addr JMM_heap ⇒ bool"
where "jmm_heap_write h a ad v h"
end
definition jmm_hconf :: "'m prog ⇒ 'addr JMM_heap ⇒ bool" (‹_ ⊢jmm _ √› [51,51] 50)
where "P ⊢jmm h √ ⟷ ty_of_htype ` ran h ⊆ {T. is_type P T}"
definition jmm_allocated :: "'addr JMM_heap ⇒ 'addr set"
where "jmm_allocated h = dom (jmm_typeof_addr h)"
definition jmm_spurious_wakeups :: bool
where "jmm_spurious_wakeups = True"
lemmas jmm_heap_ops_defs =
jmm_allocate_def jmm_typeof_addr_def
jmm_heap_read_def jmm_heap_write_def
jmm_allocated_def jmm_spurious_wakeups_def
type_synonym 'addr thread_id = "'addr"
abbreviation (input) addr2thread_id :: "'addr ⇒ 'addr thread_id"
where "addr2thread_id ≡ λx. x"
abbreviation (input) thread_id2addr :: "'addr thread_id ⇒ 'addr"
where "thread_id2addr ≡ λx. x"