(* Title: JinjaThreads/MM/SC_Collections.thy Author: Andreas Lochbihler *) section ‹Sequential consistency with efficient data structures› theory SC_Collections imports "../Common/Conform" (*"../../Collections/impl/RBTMapImpl" "../../Collections/impl/TrieMapImpl" "../../Collections/impl/ListMapImpl"*) "../Basic/JT_ICF" MM begin hide_const (open) new_Addr hide_fact (open) new_Addr_SomeD new_Addr_SomeI subsection‹Objects and Arrays› type_synonym fields = "(char, (cname, addr val) lm) tm" type_synonym array_cells = "(nat, addr val) rbt" type_synonym array_fields = "(vname, addr val) lm" datatype heapobj = Obj cname fields ― ‹class instance with class name and fields› | Arr ty nat array_fields array_cells ― ‹element type, size, fields and cell contents› lemma rec_heapobj [simp]: "rec_heapobj = case_heapobj" by(auto intro!: ext split: heapobj.split) primrec obj_ty :: "heapobj ⇒ htype" where "obj_ty (Obj c f) = Class_type c" | "obj_ty (Arr t si f e) = Array_type t si" fun is_Arr :: "heapobj ⇒ bool" where "is_Arr (Obj C fs) = False" | "is_Arr (Arr T f si el) = True" lemma is_Arr_conv: "is_Arr arrobj = (∃T si f el. arrobj = Arr T si f el)" by(cases arrobj, auto) lemma is_ArrE: "⟦ is_Arr arrobj; ⋀T si f el. arrobj = Arr T si f el ⟹ thesis ⟧ ⟹ thesis" "⟦ ¬ is_Arr arrobj; ⋀C fs. arrobj = Obj C fs ⟹ thesis ⟧ ⟹ thesis" by(cases arrobj, auto)+ definition init_fields :: "((vname × cname) × ty) list ⇒ fields" where "init_fields FDTs ≡ foldr (λ((F, D), T) fields. let F' = String.explode F in tm_update F' (lm_update D (default_val T) (case tm_lookup F' fields of None ⇒ lm_empty () | Some lm ⇒ lm)) fields) FDTs (tm_empty ())" definition init_fields_array :: "(vname × ty) list ⇒ array_fields" where "init_fields_array ≡ lm.to_map ∘ map (λ(F, T). (F, default_val T))" definition init_cells :: "ty ⇒ nat ⇒ array_cells" where "init_cells T n = foldl (λcells i. rm_update i (default_val T) cells) (rm_empty ()) [0..<n]" primrec ― ‹a new, blank object with default values in all fields:› blank :: "'m prog ⇒ htype ⇒ heapobj" where "blank P (Class_type C) = Obj C (init_fields (map (λ(FD, (T, fm)). (FD, T)) (TypeRel.fields P C)))" | "blank P (Array_type T n) = Arr T n (init_fields_array (map (λ((F, D), (T, fm)). (F, T)) (TypeRel.fields P Object))) (init_cells T n)" lemma obj_ty_blank [iff]: "obj_ty (blank P hT) = hT" by(cases hT) simp_all subsection‹Heap› type_synonym heap = "(addr, heapobj) rbt" translations (type) "heap" <= (type) "(nat, heapobj) rbt" abbreviation sc_empty :: heap where "sc_empty ≡ rm_empty ()" fun the_obj :: "heapobj ⇒ cname × fields" where "the_obj (Obj C fs) = (C, fs)" fun the_arr :: "heapobj ⇒ ty × nat × array_fields × array_cells" where "the_arr (Arr T si f el) = (T, si, f, el)" abbreviation cname_of :: "heap ⇒ addr ⇒ cname" where "cname_of hp a == fst (the_obj (the (rm_lookup a hp)))" definition new_Addr :: "heap ⇒ addr option" where "new_Addr h = Some (case rm_max h (λ_. True) of None ⇒ 0 | Some (a, _) ⇒ a + 1)" definition sc_allocate :: "'m prog ⇒ heap ⇒ htype ⇒ (heap × addr) set" where "sc_allocate P h hT = (case new_Addr h of None ⇒ {} | Some a ⇒ {(rm_update a (blank P hT) h, a)})" definition sc_typeof_addr :: "heap ⇒ addr ⇒ htype option" where "sc_typeof_addr h a = map_option obj_ty (rm_lookup a h)" inductive sc_heap_read :: "heap ⇒ addr ⇒ addr_loc ⇒ addr val ⇒ bool" for h :: heap and a :: addr where Obj: "⟦ rm_lookup a h = ⌊Obj C fs⌋; tm_lookup (String.explode F) fs = ⌊fs'⌋; lm_lookup D fs' = ⌊v⌋ ⟧ ⟹ sc_heap_read h a (CField D F) v" | Arr: "⟦ rm_lookup a h = ⌊Arr T si f el⌋; n < si ⟧ ⟹ sc_heap_read h a (ACell n) (the (rm_lookup n el))" | ArrObj: "⟦ rm_lookup a h = ⌊Arr T si f el⌋; lm_lookup F f = ⌊v⌋ ⟧ ⟹ sc_heap_read h a (CField Object F) v" hide_fact (open) Obj Arr ArrObj inductive_cases sc_heap_read_cases [elim!]: "sc_heap_read h a (CField C F) v" "sc_heap_read h a (ACell n) v" inductive sc_heap_write :: "heap ⇒ addr ⇒ addr_loc ⇒ addr val ⇒ heap ⇒ bool" for h :: heap and a :: addr where Obj: "⟦ rm_lookup a h = ⌊Obj C fs⌋; F' = String.explode F; h' = rm_update a (Obj C (tm_update F' (lm_update D v (case tm_lookup (String.explode F) fs of None ⇒ lm_empty () | Some fs' ⇒ fs')) fs)) h ⟧ ⟹ sc_heap_write h a (CField D F) v h'" | Arr: "⟦ rm_lookup a h = ⌊Arr T si f el⌋; h' = rm_update a (Arr T si f (rm_update n v el)) h ⟧ ⟹ sc_heap_write h a (ACell n) v h'" | ArrObj: "⟦ rm_lookup a h = ⌊Arr T si f el⌋; h' = rm_update a (Arr T si (lm_update F v f) el) h ⟧ ⟹ sc_heap_write h a (CField Object F) v h'" hide_fact (open) Obj Arr ArrObj inductive_cases sc_heap_write_cases [elim!]: "sc_heap_write h a (CField C F) v h'" "sc_heap_write h a (ACell n) v h'" consts sc_spurious_wakeups :: bool lemma new_Addr_SomeD: "new_Addr h = ⌊a⌋ ⟹ rm_lookup a h = None" apply(simp add: new_Addr_def) apply(drule rm.max_None[OF rm.invar]) apply(simp add: rm.lookup_correct rel_of_def) apply(clarsimp simp add: rm.lookup_correct) apply(frule rm.max_Some[OF rm.invar]) apply(clarsimp simp add: rel_of_def) apply(hypsubst_thin) apply(rule ccontr) apply(clarsimp) apply(drule_tac k'="Suc a" in rm.max_Some(2)[OF rm.invar]) apply(auto simp add: rel_of_def) done