diff --git a/thys/ConcurrentGC/Handshakes.thy b/thys/ConcurrentGC/Handshakes.thy --- a/thys/ConcurrentGC/Handshakes.thy +++ b/thys/ConcurrentGC/Handshakes.thy @@ -1,585 +1,587 @@ (*<*) (* * Copyright 2015, NICTA * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * @TAG(NICTA_BSD) *) theory Handshakes imports TSO begin +declare subst_all [simp del] [[simproc del: defined_all]] + (*>*) subsection\Handshake phases\ text\ The mutators can be at most one step behind the garbage collector (and system). If any mutator is behind then the GC is stalled on a pending handshake. Unfortunately this is a complicated by needing to consider the handshake type due to \get_work\. This relation is very precise. \ definition hp_step_rel :: "(bool \ handshake_type \ handshake_phase \ handshake_phase) set" where "hp_step_rel \ { True } \ ({ (ht_NOOP, hp, hp) |hp. hp \ {hp_Idle, hp_IdleInit, hp_InitMark, hp_Mark} } \ { (ht_GetRoots, hp_IdleMarkSweep, hp_IdleMarkSweep) , (ht_GetWork, hp_IdleMarkSweep, hp_IdleMarkSweep) }) \ { False } \ { (ht_NOOP, hp_Idle, hp_IdleMarkSweep) , (ht_NOOP, hp_IdleInit, hp_Idle) , (ht_NOOP, hp_InitMark, hp_IdleInit) , (ht_NOOP, hp_Mark, hp_InitMark) , (ht_GetRoots, hp_IdleMarkSweep, hp_Mark) , (ht_GetWork, hp_IdleMarkSweep, hp_IdleMarkSweep) }" definition handshake_phase_inv :: "('field, 'mut, 'ref) lsts_pred" where "handshake_phase_inv = (\<^bold>\m. (sys_ghost_handshake_in_sync m \<^bold>\ sys_handshake_type \<^bold>\ sys_ghost_handshake_phase \<^bold>\ mut_m.mut_ghost_handshake_phase m) \<^bold>\ \hp_step_rel\ \<^bold>\ (sys_handshake_pending m \<^bold>\ \<^bold>\(sys_ghost_handshake_in_sync m)))" (*<*) (* Sanity *) lemma handshake_step_inv: "hp' = handshake_step hp \ \in' ht. (in', ht, hp', hp) \ hp_step_rel" by (cases hp) (auto simp: hp_step_rel_def) (* Sanity *) lemma handshake_step_invD: "(False, ht, hp', hp) \ hp_step_rel \ hp' = hp_step ht hp" by (cases ht) (auto simp: hp_step_rel_def) lemma (in mut_m) handshake_phase_invD: "handshake_phase_inv s \ (sys_ghost_handshake_in_sync m s, sys_handshake_type s, sys_ghost_handshake_phase s, mut_ghost_handshake_phase s) \ hp_step_rel \ (sys_handshake_pending m s \ \sys_ghost_handshake_in_sync m s)" by (simp add: handshake_phase_inv_def) lemma handshake_in_syncD: "\ All (ghost_handshake_in_sync (s sys)); handshake_phase_inv s \ \ \m'. mut_m.mut_ghost_handshake_phase m' s = sys_ghost_handshake_phase s" by clarsimp (auto simp: hp_step_rel_def dest!: mut_m.handshake_phase_invD) lemma (in sys) handshake_phase_inv[intro]: "\ LSTP handshake_phase_inv \ sys" by (vcg_jackhammer simp: handshake_phase_inv_def) (*>*) text\ Connect @{const "sys_ghost_handshake_phase"} with locations in the GC. \ locset_definition "hp_Idle_locs = (prefixed ''idle_noop'' - { ''idle_noop_mfence'', ''idle_noop_init_type'' }) \ { ''idle_read_fM'', ''idle_invert_fM'', ''idle_write_fM'', ''idle_flip_noop_mfence'', ''idle_flip_noop_init_type'' }" locset_definition "hp_IdleInit_locs = (prefixed ''idle_flip_noop'' - { ''idle_flip_noop_mfence'', ''idle_flip_noop_init_type'' }) \ { ''idle_phase_init'', ''init_noop_mfence'', ''init_noop_init_type'' }" locset_definition "hp_InitMark_locs = (prefixed ''init_noop'' - { ''init_noop_mfence'', ''init_noop_init_type'' }) \ { ''init_phase_mark'', ''mark_read_fM'', ''mark_write_fA'', ''mark_noop_mfence'', ''mark_noop_init_type'' }" locset_definition "hp_IdleMarkSweep_locs = { ''idle_noop_mfence'', ''idle_noop_init_type'', ''mark_end'' } \ sweep_locs \ (mark_loop_locs - { ''mark_loop_get_roots_init_type'' })" locset_definition "hp_Mark_locs = (prefixed ''mark_noop'' - { ''mark_noop_mfence'', ''mark_noop_init_type'' }) \ { ''mark_loop_get_roots_init_type'' }" abbreviation "hs_noop_prefixes \ {''idle_noop'', ''idle_flip_noop'', ''init_noop'', ''mark_noop'' }" locset_definition "hs_noop_locs = (\l \ hs_noop_prefixes. prefixed l - (suffixed ''_noop_mfence'' \ suffixed ''_noop_init_type''))" locset_definition "hs_get_roots_locs = prefixed ''mark_loop_get_roots'' - {''mark_loop_get_roots_init_type''}" locset_definition "hs_get_work_locs = prefixed ''mark_loop_get_work'' - {''mark_loop_get_work_init_type''}" abbreviation "hs_prefixes \ hs_noop_prefixes \ { ''mark_loop_get_roots'', ''mark_loop_get_work'' }" locset_definition "hs_init_loop_locs = (\l \ hs_prefixes. prefixed (l @ ''_init_loop''))" locset_definition "hs_done_loop_locs = (\l \ hs_prefixes. prefixed (l @ ''_done_loop''))" locset_definition "hs_done_locs = (\l \ hs_prefixes. prefixed (l @ ''_done''))" locset_definition "hs_none_pending_locs = - (hs_init_loop_locs \ hs_done_locs)" locset_definition "hs_in_sync_locs = (- ( (\l \ hs_prefixes. prefixed (l @ ''_init'')) \ hs_done_locs )) \ (\l \ hs_prefixes. {l @ ''_init_type''})" locset_definition "hs_out_of_sync_locs = (\l \ hs_prefixes. {l @ ''_init_muts''})" locset_definition "hs_mut_in_muts_locs = (\l \ hs_prefixes. {l @ ''_init_loop_set_pending'', l @ ''_init_loop_done''})" locset_definition "hs_init_loop_done_locs = (\l \ hs_prefixes. {l @ ''_init_loop_done''})" locset_definition "hs_init_loop_not_done_locs = (hs_init_loop_locs - (\l \ hs_prefixes. {l @ ''_init_loop_done''}))" inv_definition (in gc) handshake_invL :: "('field, 'mut, 'ref) gc_pred" where "handshake_invL = (atS_gc hs_noop_locs (sys_handshake_type \<^bold>= \ht_NOOP\) \<^bold>\ atS_gc hs_get_roots_locs (sys_handshake_type \<^bold>= \ht_GetRoots\) \<^bold>\ atS_gc hs_get_work_locs (sys_handshake_type \<^bold>= \ht_GetWork\) \<^bold>\ atS_gc hs_mut_in_muts_locs (gc_mut \<^bold>\ gc_muts) \<^bold>\ atS_gc hs_init_loop_locs (\<^bold>\m. \<^bold>\(\m\ \<^bold>\ gc_muts) \<^bold>\ sys_handshake_pending m \<^bold>\ sys_ghost_handshake_in_sync m) \<^bold>\ atS_gc hs_init_loop_not_done_locs (\<^bold>\m. \m\ \<^bold>\ gc_muts \<^bold>\ \<^bold>\(sys_handshake_pending m) \<^bold>\ \<^bold>\(sys_ghost_handshake_in_sync m)) \<^bold>\ atS_gc hs_init_loop_done_locs ( (sys_handshake_pending \<^bold>$ gc_mut \<^bold>\ sys_ghost_handshake_in_sync \<^bold>$ gc_mut) \<^bold>\ (\<^bold>\m. \m\ \<^bold>\ gc_muts \<^bold>\ \m\ \<^bold>\ gc_mut \<^bold>\ \<^bold>\(sys_handshake_pending m) \<^bold>\ \<^bold>\(sys_ghost_handshake_in_sync m)) ) \<^bold>\ atS_gc hs_done_locs (\<^bold>\m. sys_handshake_pending m \<^bold>\ sys_ghost_handshake_in_sync m) \<^bold>\ atS_gc hs_done_loop_locs (\<^bold>\m. \<^bold>\(\m\ \<^bold>\ gc_muts) \<^bold>\ \<^bold>\(sys_handshake_pending m)) \<^bold>\ atS_gc hs_none_pending_locs (\<^bold>\m. \<^bold>\(sys_handshake_pending m)) \<^bold>\ atS_gc hs_in_sync_locs (\<^bold>\m. sys_ghost_handshake_in_sync m) \<^bold>\ atS_gc hs_out_of_sync_locs (\<^bold>\m. \<^bold>\(sys_handshake_pending m) \<^bold>\ \<^bold>\(sys_ghost_handshake_in_sync m)) \<^bold>\ atS_gc hp_Idle_locs (sys_ghost_handshake_phase \<^bold>= \hp_Idle\) \<^bold>\ atS_gc hp_IdleInit_locs (sys_ghost_handshake_phase \<^bold>= \hp_IdleInit\) \<^bold>\ atS_gc hp_InitMark_locs (sys_ghost_handshake_phase \<^bold>= \hp_InitMark\) \<^bold>\ atS_gc hp_IdleMarkSweep_locs (sys_ghost_handshake_phase \<^bold>= \hp_IdleMarkSweep\) \<^bold>\ atS_gc hp_Mark_locs (sys_ghost_handshake_phase \<^bold>= \hp_Mark\))" (*<*) lemma hs_get_roots_locs_subseteq_hp_IdleMarkSweep_locs: "hs_get_roots_locs \ hp_IdleMarkSweep_locs" by (auto simp: hs_get_roots_locs_def hp_IdleMarkSweep_locs_def mark_loop_locs_def intro: append_prefixD) lemma hs_get_work_locs_subseteq_hp_IdleMarkSweep_locs: "hs_get_work_locs \ hp_IdleMarkSweep_locs" apply (simp add: hs_get_work_locs_def hp_IdleMarkSweep_locs_def mark_loop_locs_def) apply clarsimp apply (drule mp) apply (auto intro: append_prefixD)[1] apply auto done lemma gc_handshake_invL_eq_imp: "eq_imp (\(_::unit) s. (AT s gc, s\ gc, sys_ghost_handshake_phase s\, handshake_pending (s\ sys), ghost_handshake_in_sync (s\ sys), sys_handshake_type s\)) gc.handshake_invL" by (simp add: gc.handshake_invL_def eq_imp_def) lemmas gc_handshake_invL_niE[nie] = iffD1[OF gc_handshake_invL_eq_imp[simplified eq_imp_simps, rule_format, unfolded conj_explode], rotated -1] lemma (in gc) handshake_invL[intro]: "\ handshake_invL \ gc" by vcg_jackhammer_ff lemma (in sys) gc_handshake_invL[intro]: notes gc.handshake_invL_def[inv] shows "\ gc.handshake_invL \ sys" by vcg_ni lemma (in gc) handshake_phase_inv[intro]: "\ handshake_invL \<^bold>\ LSTP handshake_phase_inv \ gc \ LSTP handshake_phase_inv \" by (vcg_jackhammer_ff simp: handshake_phase_inv_def hp_step_rel_def) (*>*) text\ Local handshake phase invariant for the mutators. \ locset_definition "mut_no_pending_mutates_locs = (prefixed ''hs_noop'' - { ''hs_noop'', ''hs_noop_mfence'' }) \ (prefixed ''hs_get_roots'' - { ''hs_get_roots'', ''hs_get_roots_mfence'' }) \ (prefixed ''hs_get_work'' - { ''hs_get_work'', ''hs_get_work_mfence'' })" inv_definition (in mut_m) handshake_invL :: "('field, 'mut, 'ref) gc_pred" where "handshake_invL = (atS_mut (prefixed ''hs_noop_'') (sys_handshake_type \<^bold>= \ht_NOOP\ \<^bold>\ sys_handshake_pending m) \<^bold>\ atS_mut (prefixed ''hs_get_roots_'') (sys_handshake_type \<^bold>= \ht_GetRoots\ \<^bold>\ sys_handshake_pending m) \<^bold>\ atS_mut (prefixed ''hs_get_work_'') (sys_handshake_type \<^bold>= \ht_GetWork\ \<^bold>\ sys_handshake_pending m) \<^bold>\ atS_mut mut_no_pending_mutates_locs (LIST_NULL (tso_pending_mutate (mutator m))))" (*<*) lemma (in mut_m) handshake_invL_eq_imp: "eq_imp (\(_::unit) s. (AT s (mutator m), s\ (mutator m), sys_handshake_type s\, handshake_pending (s\ sys), mem_write_buffers (s\ sys) (mutator m))) handshake_invL" by (simp add: eq_imp_def handshake_invL_def) lemmas mut_m_handshake_invL_niE[nie] = iffD1[OF mut_m.handshake_invL_eq_imp[simplified eq_imp_simps, rule_format, unfolded conj_explode], rotated -1] lemma (in mut_m) handshake_invL[intro]: "\ handshake_invL \ mutator m" by vcg_jackhammer lemma (in mut_m') handshake_invL[intro]: "\ handshake_invL \ mutator m'" by vcg_nihe vcg_ni+ lemma (in gc) mut_handshake_invL[intro]: notes mut_m.handshake_invL_def[inv] shows "\ handshake_invL \<^bold>\ mut_m.handshake_invL m \ gc \ mut_m.handshake_invL m \" by vcg_nihe vcg_ni+ lemma (in sys) mut_handshake_invL[intro]: notes mut_m.handshake_invL_def[inv] shows "\ mut_m.handshake_invL m \ sys" by (vcg_ni split: if_splits) lemma (in mut_m) gc_handshake_invL[intro]: notes gc.handshake_invL_def[inv] shows "\ handshake_invL \<^bold>\ gc.handshake_invL \ mutator m \ gc.handshake_invL \" by vcg_nihe vcg_ni+ lemma (in mut_m) handshake_phase_inv[intro]: "\ handshake_invL \<^bold>\ LSTP handshake_phase_inv \ mutator m \ LSTP handshake_phase_inv \" by (vcg_jackhammer simp: handshake_phase_inv_def) (auto simp: hp_step_rel_def) (*>*) text\ Relate @{const "sys_ghost_handshake_phase"}, @{const "gc_phase"}, @{const "sys_phase"} and writes to the phase in the GC's TSO buffer. The first relation treats the case when the GC's TSO buffer does not contain any writes to the phase. The second relation exhibits the data race on the phase variable: we need to precisely track the possible states of the GC's TSO buffer. \ definition handshake_phase_rel :: "handshake_phase \ bool \ gc_phase \ bool" where "handshake_phase_rel hp in_sync ph \ case hp of hp_Idle \ ph = ph_Idle | hp_IdleInit \ ph = ph_Idle \ (in_sync \ ph = ph_Init) | hp_InitMark \ ph = ph_Init \ (in_sync \ ph = ph_Mark) | hp_Mark \ ph = ph_Mark | hp_IdleMarkSweep \ ph = ph_Mark \ (in_sync \ ph \ { ph_Idle, ph_Sweep })" definition phase_rel :: "(bool \ handshake_phase \ gc_phase \ gc_phase \ ('field, 'ref) mem_write_action list) set" where "phase_rel \ { (in_sync, hp, ph, ph, []) |in_sync hp ph. handshake_phase_rel hp in_sync ph } \ ({True} \ { (hp_IdleInit, ph_Init, ph_Idle, [mw_Phase ph_Init]), (hp_InitMark, ph_Mark, ph_Init, [mw_Phase ph_Mark]), (hp_IdleMarkSweep, ph_Sweep, ph_Mark, [mw_Phase ph_Sweep]), (hp_IdleMarkSweep, ph_Idle, ph_Mark, [mw_Phase ph_Sweep, mw_Phase ph_Idle]), (hp_IdleMarkSweep, ph_Idle, ph_Sweep, [mw_Phase ph_Idle]) })" definition phase_rel_inv :: "('field, 'mut, 'ref) lsts_pred" where "phase_rel_inv = ((\<^bold>\m. sys_ghost_handshake_in_sync m) \<^bold>\ sys_ghost_handshake_phase \<^bold>\ gc_phase \<^bold>\ sys_phase \<^bold>\ tso_pending_phase gc \<^bold>\ \phase_rel\)" (*<*) simps_of_case handshake_phase_rel_simps[simp]: handshake_phase_rel_def (splits: handshake_phase.split) lemma phase_rel_invD: "phase_rel_inv s \ (\m. sys_ghost_handshake_in_sync m s, sys_ghost_handshake_phase s, gc_phase s, sys_phase s, tso_pending_phase gc s) \ phase_rel" by (simp add: phase_rel_inv_def) lemma phases_rel_Id[iff]: "(\m. sys_ghost_handshake_in_sync m s, sys_ghost_handshake_phase s, gc_phase s, sys_phase s, tso_pending_phase gc s) \ phase_rel \ (\ph. mw_Phase ph \ set (sys_mem_write_buffers gc s)) \ sys_phase s = gc_phase s" by (auto simp: phase_rel_def filter_empty_conv filter_eq_Cons_iff) (*>*) text\ Tie the garbage collector's control location to the value of @{const "gc_phase"}. \ locset_definition no_pending_phase_locs :: "location set" where "no_pending_phase_locs \ (idle_locs - { ''idle_noop_mfence'' }) \ (init_locs - { ''init_noop_mfence'' }) \ (mark_locs - { ''mark_read_fM'', ''mark_write_fA'', ''mark_noop_mfence'' })" inv_definition (in gc) phase_invL :: "('field, 'mut, 'ref) gc_pred" where "phase_invL = (atS_gc idle_locs (gc_phase \<^bold>= \ph_Idle\) \<^bold>\ atS_gc init_locs (gc_phase \<^bold>= \ph_Init\) \<^bold>\ atS_gc mark_locs (gc_phase \<^bold>= \ph_Mark\) \<^bold>\ atS_gc sweep_locs (gc_phase \<^bold>= \ph_Sweep\) \<^bold>\ atS_gc no_pending_phase_locs (LIST_NULL (tso_pending_phase gc)))" (*<*) lemma (in gc) phase_invL_eq_imp: "eq_imp (\r (s :: ('field, 'mut, 'ref) gc_pred_state). (AT s gc, s\ gc, tso_pending_phase gc s\)) phase_invL" by (clarsimp simp: eq_imp_def inv) lemmas gc_phase_invL_niE[nie] = iffD1[OF gc.phase_invL_eq_imp[simplified eq_imp_simps, rule_format, unfolded conj_explode], rotated -1] lemma (in gc) phase_invL[intro]: "\ phase_invL \<^bold>\ LSTP phase_rel_inv \ gc \ phase_invL \" by (vcg_jackhammer dest!: phase_rel_invD simp: phase_rel_def) lemma (in sys) gc_phase_invL[intro]: notes gc.phase_invL_def[inv] shows "\ gc.phase_invL \<^bold>\ LSTP tso_writes_inv \ sys \ gc.phase_invL \" by (vcg_ni split: if_splits) lemma (in mut_m) gc_phase_invL[intro]: notes gc.phase_invL_def[inv] shows "\ gc.phase_invL \ mutator m" by vcg_nihe lemma (in gc) phase_rel_inv[intro]: notes phase_rel_inv_def[inv] shows "\ handshake_invL \<^bold>\ phase_invL \<^bold>\ LSTP phase_rel_inv \ gc \ LSTP phase_rel_inv \" by (vcg_jackhammer_ff dest!: phase_rel_invD simp: phase_rel_def) lemma (in sys) phase_rel_inv[intro]: notes gc.phase_invL_def[inv] phase_rel_inv_def[inv] shows "\ LSTP (phase_rel_inv \<^bold>\ tso_writes_inv) \ sys \ LSTP phase_rel_inv \" apply vcg_jackhammer apply (simp add: phase_rel_def p_not_sys split: if_splits) apply (elim disjE; auto split: if_splits) done lemma (in mut_m) phase_rel_inv[intro]: "\ handshake_invL \<^bold>\ LSTP (handshake_phase_inv \<^bold>\ phase_rel_inv) \ mutator m \ LSTP phase_rel_inv \" apply (vcg_jackhammer simp: phase_rel_inv_def) subgoal by (auto dest!: handshake_phase_invD simp: handshake_phase_rel_def phase_rel_def hp_step_rel_def split: handshake_phase.splits) subgoal by (auto dest!: handshake_phase_invD simp: handshake_phase_rel_def phase_rel_def hp_step_rel_def split: handshake_phase.splits) subgoal by (auto dest!: handshake_phase_invD simp: handshake_phase_rel_def phase_rel_def hp_step_rel_def split: handshake_phase.splits) done (*>*) text\ Validity of @{const "sys_fM"} wrt @{const "gc_fM"} and the handshake phase. Effectively we use @{const "gc_fM"} as ghost state. We also include the TSO lock to rule out the GC having any pending marks during the @{const "hp_Idle"} handshake phase. \ definition fM_rel :: "(bool \ handshake_phase \ gc_mark \ gc_mark \ ('field, 'ref) mem_write_action list \ bool) set" where "fM_rel = { (in_sync, hp, fM, fM, [], l) |fM hp in_sync l. hp = hp_Idle \ \in_sync } \ { (in_sync, hp_Idle, fM, fM', [], l) |fM fM' in_sync l. in_sync } \ { (in_sync, hp_Idle, \fM, fM, [mw_fM (\fM)], False) |fM in_sync. in_sync }" definition fM_rel_inv :: "('field, 'mut, 'ref) lsts_pred" where "fM_rel_inv = ((\<^bold>\m. sys_ghost_handshake_in_sync m) \<^bold>\ sys_ghost_handshake_phase \<^bold>\ gc_fM \<^bold>\ sys_fM \<^bold>\ tso_pending_fM gc \<^bold>\ (sys_mem_lock \<^bold>= \Some gc\) \<^bold>\ \fM_rel\)" definition fA_rel :: "(bool \ handshake_phase \ gc_mark \ gc_mark \ ('field, 'ref) mem_write_action list) set" where "fA_rel = { (in_sync, hp_Idle, fA, fM, []) |fA fM in_sync. \in_sync \ fA = fM } \ { (in_sync, hp_IdleInit, fA, \fA, []) |fA in_sync. True } \ { (in_sync, hp_InitMark, fA, \fA, [mw_fA (\fA)]) |fA in_sync. in_sync } \ { (in_sync, hp_InitMark, fA, fM, []) |fA fM in_sync. \in_sync \ fA \ fM } \ { (in_sync, hp_Mark, fA, fA, []) |fA in_sync. True } \ { (in_sync, hp_IdleMarkSweep, fA, fA, []) |fA in_sync. True }" definition fA_rel_inv :: "('field, 'mut, 'ref) lsts_pred" where "fA_rel_inv = ((\<^bold>\m. sys_ghost_handshake_in_sync m) \<^bold>\ sys_ghost_handshake_phase \<^bold>\ sys_fA \<^bold>\ gc_fM \<^bold>\ tso_pending_fA gc \<^bold>\ \fA_rel\)" locset_definition "fM_eq_locs = (- { ''idle_write_fM'', ''idle_flip_noop_mfence'' })" locset_definition "fM_tso_empty_locs = (- { ''idle_flip_noop_mfence'' })" locset_definition "fA_tso_empty_locs = (- { ''mark_noop_mfence'' })" locset_definition "fA_eq_locs \ { ''idle_read_fM'', ''idle_invert_fM'' } \ prefixed ''idle_noop'' \ (mark_locs - { ''mark_read_fM'', ''mark_write_fA'', ''mark_noop_mfence'' }) \ sweep_locs" locset_definition "fA_neq_locs \ { ''idle_phase_init'', ''idle_write_fM'', ''mark_read_fM'', ''mark_write_fA'' } \ prefixed ''idle_flip_noop'' \ init_locs" inv_definition (in gc) fM_fA_invL :: "('field, 'mut, 'ref) gc_pred" where "fM_fA_invL = (atS_gc fM_eq_locs (sys_fM \<^bold>= gc_fM) \<^bold>\ at_gc ''idle_write_fM'' (sys_fM \<^bold>\ gc_fM) \<^bold>\ at_gc ''idle_flip_noop_mfence'' (sys_fM \<^bold>\ gc_fM \<^bold>\ (\<^bold>\(LIST_NULL (tso_pending_fM gc)))) \<^bold>\ atS_gc fM_tso_empty_locs (LIST_NULL (tso_pending_fM gc)) \<^bold>\ atS_gc fA_eq_locs (sys_fA \<^bold>= gc_fM) \<^bold>\ atS_gc fA_neq_locs (sys_fA \<^bold>\ gc_fM) \<^bold>\ at_gc ''mark_noop_mfence'' (sys_fA \<^bold>\ gc_fM \<^bold>\ (\<^bold>\(LIST_NULL (tso_pending_fA gc)))) \<^bold>\ atS_gc fA_tso_empty_locs (LIST_NULL (tso_pending_fA gc)))" (*<*) lemmas fM_rel_invD = iffD1[OF fun_cong[OF fM_rel_inv_def[simplified atomize_eq]]] lemma do_write_action_fM[simp]: "\ sys_mem_write_buffers p s = w # ws; tso_writes_inv s; handshake_phase_inv s; fM_rel_inv s; ghost_handshake_phase (s (mutator m)) \ hp_Idle \ (sys_ghost_handshake_phase s = hp_IdleMarkSweep \ All (ghost_handshake_in_sync (s sys))); p \ sys \ \ fM (do_write_action w (s sys)) = fM (s sys)" apply (drule mut_m.handshake_phase_invD[where m=m]) apply (clarsimp simp: do_write_action_def p_not_sys split: mem_write_action.splits) apply (simp add: hp_step_rel_def) apply (elim disjE, auto simp: fM_rel_inv_def fM_rel_def) done lemmas fA_rel_invD = iffD1[OF fun_cong[OF fA_rel_inv_def[simplified atomize_eq]]] lemma gc_fM_fA_invL_eq_imp: "eq_imp (\(_::unit) s. (AT s gc, s\ gc, sys_fA s\, sys_fM s\, sys_mem_write_buffers gc s\)) gc.fM_fA_invL" by (simp add: gc.fM_fA_invL_def eq_imp_def) lemmas gc_fM_fA_invL_niE[nie] = iffD1[OF gc_fM_fA_invL_eq_imp[simplified eq_imp_simps, rule_format, unfolded conj_explode], rotated -1] lemma (in gc) fM_fA_invL[intro]: "\ fM_fA_invL \ gc" by vcg_jackhammer lemma (in gc) fM_rel_inv[intro]: "\ fM_fA_invL \<^bold>\ handshake_invL \<^bold>\ tso_lock_invL \<^bold>\ LSTP fM_rel_inv \ gc \ LSTP fM_rel_inv \" by (vcg_jackhammer simp: fM_rel_inv_def fM_rel_def) lemma (in gc) fA_rel_inv[intro]: "\ fM_fA_invL \<^bold>\ handshake_invL \<^bold>\ LSTP fA_rel_inv \ gc \ LSTP fA_rel_inv \" by (vcg_jackhammer simp: fA_rel_inv_def; auto simp: fA_rel_def) lemma (in mut_m) gc_fM_fA_invL[intro]: notes gc.fM_fA_invL_def[inv] shows "\ gc.fM_fA_invL \ mutator m" by vcg_nihe lemma (in mut_m) fM_rel_inv[intro]: "\ LSTP fM_rel_inv \ mutator m" by (vcg_jackhammer simp: fM_rel_inv_def fM_rel_def; elim disjE; auto split: if_splits) (* FIXME trivial but eta reduction plays merry hell *) lemma (in mut_m) fA_rel_inv[intro]: "\ LSTP fA_rel_inv \ mutator m" by (vcg_jackhammer simp: fA_rel_inv_def; simp add: fA_rel_def; elim disjE; auto split: if_splits) lemma fA_neq_locs_diff_fA_tso_empty_locs[simp]: "fA_neq_locs - fA_tso_empty_locs = {}" by (simp add: fA_neq_locs_def fA_tso_empty_locs_def loc) lemma (in sys) gc_fM_fA_invL[intro]: notes gc.fM_fA_invL_def[inv] shows "\ gc.fM_fA_invL \<^bold>\ LSTP (fA_rel_inv \<^bold>\ fM_rel_inv \<^bold>\ tso_writes_inv) \ sys \ gc.fM_fA_invL \" apply (vcg_ni simp: p_not_sys) apply (erule disjE) apply (clarsimp simp: fM_rel_inv_def fM_rel_def split: if_splits) apply (clarsimp simp: fM_rel_inv_def fM_rel_def filter_empty_conv) apply (erule disjE; clarsimp) apply (rule conjI; clarsimp simp: fM_rel_inv_def fM_rel_def split: if_splits) apply (clarsimp split: if_splits) apply (erule disjE) apply (clarsimp simp: fA_rel_inv_def fA_rel_def) apply clarsimp apply (erule disjE) apply (clarsimp simp: fA_rel_inv_def fA_rel_def fM_rel_inv_def fM_rel_def) apply (drule (1) atS_dests(3), fastforce simp: atS_simps) apply clarsimp apply (rule conjI) apply (clarsimp simp: fA_rel_inv_def fA_rel_def split: if_splits) apply clarsimp apply (clarsimp split: if_splits) done lemma (in sys) fM_rel_inv[intro]: "\ LSTP (fM_rel_inv \<^bold>\ tso_writes_inv) \ sys \ LSTP fM_rel_inv \" apply (vcg_ni simp: p_not_sys) apply (fastforce simp: do_write_action_def fM_rel_inv_def fM_rel_def split: mem_write_action.splits) done lemma (in sys) fA_rel_inv[intro]: "\ LSTP (fA_rel_inv \<^bold>\ tso_writes_inv) \ sys \ LSTP fA_rel_inv \" apply (vcg_ni simp: p_not_sys) apply (fastforce simp: do_write_action_def fA_rel_inv_def fA_rel_def split: mem_write_action.splits) done lemma mut_m_get_roots_no_fM_write: "\ atS (mutator m) (prefixed ''hs_get_roots_'') s; mut_m.handshake_invL m s; handshake_phase_inv s\; fM_rel_inv s\; tso_writes_inv s\; p \ sys \ \ \sys_mem_write_buffers p s\ = mw_fM fl # ws" unfolding mut_m.handshake_invL_def apply (elim conjE) apply (drule mut_m.handshake_phase_invD[where m=m]) apply (drule fM_rel_invD) apply (fastforce simp: hp_step_rel_def fM_rel_def loc filter_empty_conv p_not_sys) done lemma mut_m_get_roots_no_phase_write: "\ atS (mutator m) (prefixed ''hs_get_roots_'') s; mut_m.handshake_invL m s; handshake_phase_inv s\; phase_rel_inv s\; tso_writes_inv s\; p \ sys \ \ \sys_mem_write_buffers p s\ = mw_Phase ph # ws" unfolding mut_m.handshake_invL_def apply (elim conjE) apply (drule mut_m.handshake_phase_invD[where m=m]) apply (drule phase_rel_invD) apply (fastforce simp: hp_step_rel_def phase_rel_def loc filter_empty_conv p_not_sys) done lemma mut_m_not_idle_no_fM_write: "\ ghost_handshake_phase (s (mutator m)) \ hp_Idle; fM_rel_inv s; handshake_phase_inv s; tso_writes_inv s; p \ sys \ \ \sys_mem_write_buffers p s = mw_fM fl # ws" apply (drule mut_m.handshake_phase_invD[where m=m]) apply (drule fM_rel_invD) apply (fastforce simp: hp_step_rel_def fM_rel_def filter_empty_conv p_not_sys) done lemma (in mut_m) mut_ghost_handshake_phase_idle: "\ mut_ghost_handshake_phase s = hp_Idle; handshake_phase_inv s; phase_rel_inv s \ \ sys_phase s = ph_Idle" apply (drule phase_rel_invD) apply (drule handshake_phase_invD) apply (auto simp: phase_rel_def hp_step_rel_def) done (*>*) (*<*) end (*>*) diff --git a/thys/ConcurrentGC/MarkObject.thy b/thys/ConcurrentGC/MarkObject.thy --- a/thys/ConcurrentGC/MarkObject.thy +++ b/thys/ConcurrentGC/MarkObject.thy @@ -1,1102 +1,1104 @@ (*<*) (* * Copyright 2015, NICTA * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * @TAG(NICTA_BSD) *) theory MarkObject imports Handshakes begin +declare subst_all [simp del] [[simproc del: defined_all]] + (*>*) subsection\Object colours, reference validity, worklist validity\ text\ We adopt the classical tricolour scheme for object colours due to @{cite [cite_macro=citet] "DBLP:journals/cacm/DijkstraLMSS78"}, but tweak it somewhat in the presence of worklists and TSO. Intuitively: \begin{description} \item[White] potential garbage, not yet reached \item[Grey] reached, presumed live, a source of possible new references (work) \item[Black] reached, presumed live, not a source of new references \end{description} In this particular setting we use the following interpretation: \begin{description} \item[White:] not marked \item[Grey:] on a worklist \item[Black:] marked and not on a worklist \end{description} Note that this allows the colours to overlap: an object being marked may be white (on the heap) and in @{const "ghost_honorary_grey"} for some process, i.e. grey. \ abbreviation marked :: "'ref \ ('field, 'mut, 'ref) lsts_pred" where "marked r s \ obj_at (\obj. obj_mark obj = sys_fM s) r s" abbreviation white :: "'ref \ ('field, 'mut, 'ref) lsts_pred" where "white r s \ obj_at (\obj. obj_mark obj = (\sys_fM s)) r s" definition WL :: "'mut process_name \ ('field, 'mut, 'ref) lsts \ 'ref set" where "WL p \ \s. W (s p) \ ghost_honorary_grey (s p)" definition grey :: "'ref \ ('field, 'mut, 'ref) lsts_pred" where "grey r \ \<^bold>\p. \r\ \<^bold>\ WL p" definition black :: "'ref \ ('field, 'mut, 'ref) lsts_pred" where "black r \ marked r \<^bold>\ \<^bold>\(grey r)" text\ We show that if a mutator can load a reference into its roots (its working set of references), then there is an object in the heap at that reference. In this particular collector, we can think of grey references and pending TSO heap mutations as extra mutator roots; in particular the GC holds no roots itself but marks everything reachable from its worklist, and so we need to know these objects exist. By the strong tricolour invariant (\S\ref{sec:strong-tricolour-invariant}), black objects point to black or grey objects, and so we do not need to treat these specially. \ abbreviation write_refs :: "('field, 'ref) mem_write_action \ 'ref set" where "write_refs w \ case w of mw_Mutate r f r' \ {r} \ Option.set_option r' | _ \ {}" definition (in mut_m) tso_write_refs :: "('field, 'mut, 'ref) lsts \ 'ref set" where "tso_write_refs = (\s. \w \ set (sys_mem_write_buffers (mutator m) s). write_refs w)" definition (in mut_m) reachable :: "'ref \ ('field, 'mut, 'ref) lsts_pred" where "reachable y = (\<^bold>\x. \x\ \<^bold>\ mut_roots \<^bold>\ mut_ghost_honorary_root \<^bold>\ tso_write_refs \<^bold>\ x reaches y)" definition grey_reachable :: "'ref \ ('field, 'mut, 'ref) lsts_pred" where "grey_reachable y = (\<^bold>\g. grey g \<^bold>\ g reaches y)" definition valid_refs_inv :: "('field, 'mut, 'ref) lsts_pred" where "valid_refs_inv = (\<^bold>\x. ((\<^bold>\m. mut_m.reachable m x) \<^bold>\ grey_reachable x) \<^bold>\ valid_ref x)" text\ \label{def:valid_W_inv} The worklists track the grey objects. The following invariant asserts that grey objects are marked on the heap except for a few steps near the end of @{const "mark_object_fn"}, the processes' worklists and @{const "ghost_honorary_grey"}s are disjoint, and that pending marks are sensible. The safety of the collector does not to depend on disjointness; we include it as proof that the single-threading of grey objects in the implementation is sound. \ definition valid_W_inv :: "('field, 'mut, 'ref) lsts_pred" where "valid_W_inv = (\<^bold>\p q r fl. (r in_W p \<^bold>\ (sys_mem_lock \<^bold>\ \Some p\ \<^bold>\ r in_ghost_honorary_grey p) \<^bold>\ marked r) \<^bold>\ (\p \ q\ \<^bold>\ \<^bold>\(\r\ \<^bold>\ WL p \<^bold>\ \r\ \<^bold>\ WL q)) \<^bold>\ (\<^bold>\(r in_ghost_honorary_grey p \<^bold>\ r in_W q)) \<^bold>\ (EMPTY sys_ghost_honorary_grey) \<^bold>\ (tso_pending_write p (mw_Mark r fl) \<^bold>\ ( \fl\ \<^bold>= sys_fM \<^bold>\ r in_ghost_honorary_grey p \<^bold>\ tso_locked_by p \<^bold>\ white r \<^bold>\ tso_pending_mark p \<^bold>= \[mw_Mark r fl]\ )))" (*<*) lemma obj_at_mark_dequeue[simp]: "obj_at P r (s(sys := s sys\ heap := (sys_heap s)(r' := Option.map_option (obj_mark_update (\_. fl)) (sys_heap s r')), mem_write_buffers := wb' \)) \ (r = r' \ obj_at (\obj. (P (obj\ obj_mark := fl \))) r s) \ (r \ r' \ obj_at P r s)" by (clarsimp split: obj_at_splits) lemma marked_not_white: "white r s \ \marked r s" by (clarsimp split: obj_at_splits) lemma valid_ref_valid_null_ref_simps[simp]: "valid_ref r (s(sys := do_write_action w (s sys)\mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\)) \ valid_ref r s" "valid_null_ref r' (s(sys := do_write_action w (s sys)\mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\)) \ valid_null_ref r' s" "valid_null_ref r' (s(mutator m := mut_s', sys := (s sys)\ heap := (heap (s sys))(r'' \ obj) \)) \ valid_null_ref r' s \ r' = Some r''" by (auto simp: do_write_action_def valid_null_ref_def split: mem_write_action.splits obj_at_splits option.splits) text\points to, reaches, reachable mut, reachable grey\ lemma reaches_fields: assumes "(x reaches y) s'" assumes "\r. (Option.set_option (sys_heap s' r) \ ran \ obj_fields) = (Option.set_option (sys_heap s r) \ ran \ obj_fields)" shows "(x reaches y) s" using assms proof induct case (step y z) then have "(y points_to z) s" by (cases "sys_heap s y") (auto simp: ran_def obj_at_def split: option.splits dest!: spec[where x=y]) with step show ?case by (blast intro: rtranclp.intros(2)) qed simp lemma reaches_eq_imp: "eq_imp (\r s. Option.set_option (sys_heap s r) \ ran \ obj_fields) (x reaches y)" by (auto simp: eq_imp_def elim: reaches_fields) lemmas reaches_fun_upd[simp] = eq_imp_fun_upd[OF reaches_eq_imp, simplified eq_imp_simps, rule_format] lemma (in mut_m) reachable_eq_imp: "eq_imp (\r'. mut_roots \<^bold>\ mut_ghost_honorary_root \<^bold>\ (\s. Option.set_option (sys_heap s r') \ ran \ obj_fields) \<^bold>\ tso_pending_mutate (mutator m)) (reachable r)" apply (clarsimp simp: eq_imp_def reachable_def tso_write_refs_def ex_disj_distrib) apply (rename_tac s s') apply (subgoal_tac "\r'. (\w\set (sys_mem_write_buffers (mutator m) s). r' \ write_refs w) \ (\w\set (sys_mem_write_buffers (mutator m) s'). r' \ write_refs w)") apply (subgoal_tac "\x. (x reaches r) s \ (x reaches r) s'") apply clarsimp apply (auto simp: reaches_fields)[1] apply (drule arg_cong[where f=set]) apply (clarsimp simp: set_eq_iff) apply (rule iffI) apply clarsimp apply (rename_tac s s' r' w) apply (drule_tac x=w in spec) apply (rule_tac x=w in bexI) apply clarsimp apply (case_tac w, simp_all) apply clarsimp apply (rename_tac s s' r' w) apply (drule_tac x=w in spec) apply (rule_tac x=w in bexI) apply clarsimp apply (case_tac w, simp_all) done lemmas reachable_fun_upd[simp] = eq_imp_fun_upd[OF mut_m.reachable_eq_imp, simplified eq_imp_simps, rule_format] lemma reachableI[intro]: "x \ mut_m.mut_roots m s \ mut_m.reachable m x s" "x \ mut_m.tso_write_refs m s \ mut_m.reachable m x s" by (auto simp: mut_m.reachable_def) lemma reachableE: "\ (x points_to y) s; mut_m.reachable m x s \ \ mut_m.reachable m y s" by (auto simp: mut_m.reachable_def elim: rtranclp.intros(2)) lemma (in mut_m) reachable_induct[consumes 1, case_names root ghost_honorary_root tso_root reaches]: assumes r: "reachable y s" assumes root: "\x. \ x \ mut_roots s \ \ P x" assumes ghost_honorary_root: "\x. \ x \ mut_ghost_honorary_root s \ \ P x" assumes tso_root: "\x. \ x \ tso_write_refs s \ \ P x" assumes reaches: "\x y. \ reachable x s; (x points_to y) s; P x \ \ P y" shows "P y" using r unfolding reachable_def proof(clarify) fix x assume xy: "(x reaches y) s" and xr: "x \ mut_roots s \ mut_ghost_honorary_root s \ tso_write_refs s" then show "P y" proof induct case base with root ghost_honorary_root tso_root show ?case by blast next case (step y z) with reaches show ?case unfolding reachable_def by blast qed qed lemmas reachable_induct = mut_m.reachable_induct[consumes 1, case_names root ghost_honorary_root tso_root reaches] lemma (in mut_m) mut_reachableE[consumes 1, case_names mut_root tso_write_refs]: "\ reachable y s; \x. \ (x reaches y) s; x \ mut_roots s \ \ Q; \x. \ (x reaches y) s; x \ mut_ghost_honorary_root s \ \ Q; \x. \ (x reaches y) s; x \ tso_write_refs s \ \ Q \ \ Q" by (auto simp: reachable_def) lemma grey_reachable_eq_imp: "eq_imp (\r'. (\s. \p. WL p s) \<^bold>\ (\s. Option.set_option (sys_heap s r') \ ran \ obj_fields)) (grey_reachable r)" by (auto simp: eq_imp_def grey_reachable_def grey_def set_eq_iff reaches_fields) lemmas grey_reachable_fun_upd[simp] = eq_imp_fun_upd[OF grey_reachable_eq_imp, simplified eq_imp_simps, rule_format] lemma grey_reachableI[intro]: "grey g s \ grey_reachable g s" by (auto simp: grey_reachable_def) lemma grey_reachableE: "\ (g points_to y) s; grey_reachable g s \ \ grey_reachable y s" by (auto simp: grey_reachable_def elim: rtranclp.intros(2)) text\colours and work lists\ lemma black_eq_imp: "eq_imp (\_::unit. (\s. r \ (\p. WL p s)) \<^bold>\ sys_fM \<^bold>\ (\s. Option.map_option obj_mark (sys_heap s r))) (black r)" by (auto simp add: eq_imp_def black_def grey_def obj_at_def split: option.splits) lemma white_eq_imp: "eq_imp (\_::unit. sys_fM \<^bold>\ (\s. Option.map_option obj_mark (sys_heap s r))) (white r)" by (auto simp add: eq_imp_def obj_at_def split: option.splits) lemma grey_eq_imp: "eq_imp (\_::unit. (\s. r \ (\p. WL p s))) (grey r)" by (auto simp add: eq_imp_def grey_def) lemmas black_fun_upd[simp] = eq_imp_fun_upd[OF black_eq_imp, simplified eq_imp_simps, rule_format] lemmas grey_fun_upd[simp] = eq_imp_fun_upd[OF grey_eq_imp, simplified eq_imp_simps, rule_format] lemmas white_fun_upd[simp] = eq_imp_fun_upd[OF white_eq_imp, simplified eq_imp_simps, rule_format] text\ These demonstrate the overlap in colours. \ lemma colours_distinct[dest]: "black r s \ \grey r s" "black r s \ \white r s" "grey r s \ \black r s" "white r s \ \black r s" by (auto simp: black_def split: obj_at_splits) lemma marked_imp_black_or_grey: "marked r s \ black r s \ grey r s" "\ white r s \ \ valid_ref r s \ black r s \ grey r s" by (auto simp: black_def grey_def split: obj_at_splits) lemma blackD[dest]: "black r s \ marked r s" "black r s \ r \ WL p s" by (simp_all add: black_def grey_def) text\valid refs inv\ lemma valid_refs_inv_eq_imp: "eq_imp (\(m', r'). (\s. roots (s (mutator m'))) \<^bold>\ (\s. ghost_honorary_root (s (mutator m'))) \<^bold>\ (\s. Option.map_option obj_fields (sys_heap s r')) \<^bold>\ tso_pending_mutate (mutator m') \<^bold>\ (\s. \p. WL p s)) valid_refs_inv" apply (clarsimp simp: eq_imp_def valid_refs_inv_def grey_reachable_def all_conj_distrib) apply (rename_tac s s') apply (subgoal_tac "\r. valid_ref r s \ valid_ref r s'") apply (subgoal_tac "\x. Option.set_option (sys_heap s x) \ ran \ obj_fields = Option.set_option (sys_heap s' x) \ ran \ obj_fields") apply (subst eq_impD[OF mut_m.reachable_eq_imp]) defer apply (subst eq_impD[OF grey_eq_imp]) defer apply (subst eq_impD[OF reaches_eq_imp]) defer apply force apply clarsimp apply (rename_tac x) apply (drule_tac x=x in spec) apply (clarsimp simp: set_eq_iff ran_def) apply (case_tac "sys_heap s x", simp_all)[1] apply (metis (hide_lams, no_types) elem_set not_Some_eq option.inject map_option_eq_Some) apply (clarsimp split: obj_at_splits) apply (rule conjI) apply (metis map_option_is_None) apply (metis map_option_eq_Some) apply clarsimp apply clarsimp apply clarsimp done lemmas valid_refs_inv_fun_upd[simp] = eq_imp_fun_upd[OF valid_refs_inv_eq_imp, simplified eq_imp_simps, rule_format] lemma valid_refs_invD[elim]: "\ x \ mut_m.mut_roots m s; (x reaches y) s; valid_refs_inv s \ \ valid_ref y s" "\ x \ mut_m.mut_roots m s; (x reaches y) s; valid_refs_inv s \ \ \obj. sys_heap s y = Some obj" "\ x \ mut_m.tso_write_refs m s; (x reaches y) s; valid_refs_inv s \ \ valid_ref y s" "\ x \ mut_m.tso_write_refs m s; (x reaches y) s; valid_refs_inv s \ \ \obj. sys_heap s y = Some obj" "\ w \ set (sys_mem_write_buffers (mutator m) s); x \ write_refs w; (x reaches y) s; valid_refs_inv s \ \ valid_ref y s" "\ w \ set (sys_mem_write_buffers (mutator m) s); x \ write_refs w; (x reaches y) s; valid_refs_inv s \ \ \obj. sys_heap s y = Some obj" "\ grey x s; (x reaches y) s; valid_refs_inv s \ \ valid_ref y s" "\ mut_m.reachable m x s; valid_refs_inv s \ \ valid_ref x s" "\ mut_m.reachable m x s; valid_refs_inv s \ \ \obj. sys_heap s x = Some obj" "\ x \ mut_m.mut_ghost_honorary_root m s; (x reaches y) s; valid_refs_inv s \ \ valid_ref y s" "\ x \ mut_m.mut_ghost_honorary_root m s; (x reaches y) s; valid_refs_inv s \ \ \obj. sys_heap s y = Some obj" by (fastforce simp: valid_refs_inv_def grey_reachable_def mut_m.reachable_def mut_m.tso_write_refs_def split: obj_at_splits)+ lemma valid_refs_invD2[elim]: "\ mut_m.reachable m x s; valid_refs_inv s; (x reaches y) s \ \ valid_ref y s" apply (clarsimp simp: valid_refs_inv_def mut_m.reachable_def) apply (frule (1) rtranclp_trans) apply auto done lemma valid_refs_invD3: "\ sys_mem_write_buffers (mutator m) s = mw_Mutate r f opt_r' # ws; (r reaches y) s; valid_refs_inv s \ \ valid_ref y s" apply (clarsimp simp: valid_refs_inv_def mut_m.reachable_def mut_m.tso_write_refs_def) apply (fastforce dest: spec[where x=y] spec[where x=m]) done text\WL\ lemma WLI[intro]: "r \ W (s p) \ r \ WL p s" "r \ ghost_honorary_grey (s p) \ r \ WL p s" by (simp_all add: WL_def) lemma WL_eq_imp: "eq_imp (\(_::unit) s. (ghost_honorary_grey (s p), W (s p))) (WL p)" by (clarsimp simp: eq_imp_def WL_def) lemmas WL_fun_upd[simp] = eq_imp_fun_upd[OF WL_eq_imp, simplified eq_imp_simps, rule_format] lemma greyI[intro]: "r \ ghost_honorary_grey (s p) \ grey r s" "r \ W (s p) \ grey r s" "r \ WL p s \ grey r s" by (case_tac [!] p) (auto simp: grey_def WL_def) text\@{const "valid_W_inv"}\ lemma valid_W_inv_eq_imp: "eq_imp (\(p, r). (\s. W (s p)) \<^bold>\ (\s. ghost_honorary_grey (s p)) \<^bold>\ sys_fM \<^bold>\ (\s. Option.map_option obj_mark (sys_heap s r)) \<^bold>\ sys_mem_lock \<^bold>\ tso_pending_mark p) valid_W_inv" apply (clarsimp simp: eq_imp_def valid_W_inv_def fun_eq_iff all_conj_distrib) apply (rename_tac s s') apply (subgoal_tac "\p. WL p s = WL p s'") apply (subgoal_tac "\x. obj_at (\obj. obj_mark obj = sys_fM s') x s \ obj_at (\obj. obj_mark obj = sys_fM s') x s'") apply (subgoal_tac "\x. obj_at (\obj. obj_mark obj = (\sys_fM s')) x s \ obj_at (\obj. obj_mark obj = (\sys_fM s')) x s'") apply (subgoal_tac "\x xa xb. mw_Mark xa xb \ set (sys_mem_write_buffers x s) \ mw_Mark xa xb \ set (sys_mem_write_buffers x s')") apply simp apply clarsimp apply (rename_tac x xa xb) apply (drule_tac x=x in spec, drule arg_cong[where f=set], fastforce) apply (clarsimp split: obj_at_splits) apply (rename_tac x) apply ( (drule_tac x=x in spec)+ )[1] apply (case_tac "sys_heap s x", simp_all) apply (case_tac "sys_heap s' x", auto)[1] apply (clarsimp split: obj_at_splits) apply (rename_tac x) apply (drule_tac x=x in spec) apply (case_tac "sys_heap s x", simp_all) apply (case_tac "sys_heap s' x", simp_all) apply (simp add: WL_def) done lemmas valid_W_inv_fun_upd[simp] = eq_imp_fun_upd[OF valid_W_inv_eq_imp, simplified eq_imp_simps, rule_format] lemma valid_W_invD[dest]: "\ r \ W (s p); valid_W_inv s \ \ marked r s" "\ r \ W (s p); valid_W_inv s \ \ valid_ref r s" "\ r \ WL p s; valid_W_inv s; p \ q \ \ r \ WL q s" "\ r \ W (s p); valid_W_inv s; p \ q \ \ r \ WL q s" "\ r \ W (s p); valid_W_inv s \ \ r \ ghost_honorary_grey (s q)" "\ r \ ghost_honorary_grey (s p); valid_W_inv s \ \ r \ W (s q)" "\ r \ ghost_honorary_grey (s p); valid_W_inv s; p \ q \ \ r \ WL q s" by (auto simp: valid_W_inv_def WL_def split: obj_at_splits) (* FIXME horrible but effective (?) *) lemma valid_W_invD2[dest]: "\ sys_mem_write_buffers p s = mw_Mark r fl # ws; valid_W_inv s \ \ fl = sys_fM s \ r \ ghost_honorary_grey (s p) \ tso_locked_by p s \ white r s \ filter is_mw_Mark ws = []" "\ mw_Mark r fl \ set (sys_mem_write_buffers p s); valid_W_inv s \ \ fl = sys_fM s \ r \ ghost_honorary_grey (s p) \ tso_locked_by p s \ white r s \ filter is_mw_Mark (sys_mem_write_buffers p s) = [mw_Mark r fl]" by (clarsimp simp: valid_W_inv_def dest!: spec[where x=p], blast)+ lemma valid_W_invD3[elim]: "\ mw_Mark r fl \ set (sys_mem_write_buffers p s); valid_W_inv s \ \ r \ ghost_honorary_grey (s p)" "\ r \ ghost_honorary_grey (s p); sys_mem_lock s \ Some p; valid_W_inv s \ \ marked r s" "\ r \ ghost_honorary_grey (s p); sys_mem_lock s \ Some p; valid_W_inv s \ \ valid_ref r s" apply (simp_all add: valid_W_inv_def) apply (clarsimp split: obj_at_splits) apply blast done lemma valid_W_invD4: "\ sys_mem_write_buffers p s = mw_Mutate r' f r'' # ws; mw_Mark r fl \ set ws; valid_W_inv s \ \ fl = sys_fM s \ r \ ghost_honorary_grey (s p) \ tso_locked_by p s \ white r s \ filter is_mw_Mark ws = [mw_Mark r fl]" "\ sys_mem_write_buffers p s = mw_fA fl' # ws; mw_Mark r fl \ set ws; valid_W_inv s \ \ fl = sys_fM s \ r \ ghost_honorary_grey (s p) \ tso_locked_by p s \ white r s \ filter is_mw_Mark ws = [mw_Mark r fl]" "\ sys_mem_write_buffers p s = mw_fM fl' # ws; mw_Mark r fl \ set ws; valid_W_inv s \ \ fl = sys_fM s \ r \ ghost_honorary_grey (s p) \ tso_locked_by p s \ white r s \ filter is_mw_Mark ws = [mw_Mark r fl]" "\ sys_mem_write_buffers p s = mw_Phase ph # ws; mw_Mark r fl \ set ws; valid_W_inv s \ \ fl = sys_fM s \ r \ ghost_honorary_grey (s p) \ tso_locked_by p s \ white r s \ filter is_mw_Mark ws = [mw_Mark r fl]" by (clarsimp simp: valid_W_inv_def dest!: spec[where x=p], blast)+ lemma valid_W_iff[iff]: "valid_W_inv s \ sys_ghost_honorary_grey s = {}" by (simp add: valid_W_inv_def) lemma valid_W_inv_no_mark_writes_invD: "\ sys_mem_lock s \ Some p; valid_W_inv s \ \ tso_pending p is_mw_Mark s = []" by (auto intro!: filter_False) lemma valid_W_inv_sys_read[simp]: "\ sys_mem_lock s \ Some p; valid_W_inv s \ \ sys_read p (mr_Mark r) (s sys) = mv_Mark (Option.map_option obj_mark (sys_heap s r))" apply (clarsimp simp: sys_read_def fold_writes_def) apply (rule fold_invariant[where P="\fr. Option.map_option obj_mark (heap (fr (s sys)) r) = Option.map_option obj_mark (sys_heap s r)" and Q="\w. \r fl. w \ mw_Mark r fl"]) apply (auto simp: Option.map_option_case do_write_action_def filter_empty_conv split: mem_write_action.splits option.splits) done (*>*) subsection\Mark Object\ text\ Local invariants for @{const "mark_object_fn"}. Invoking this code in phases where @{const "sys_fM"} is constant marks the reference in @{const "ref"}. When @{const "sys_fM"} could vary this code is not called. The two cases are distinguished by @{term "p_ph_enabled"}. Each use needs to provide extra facts to justify validity of references, etc. We do not include a post-condition for @{const "mark_object_fn"} here as it is different at each call site. \ locale mark_object = fixes p :: "'mut process_name" fixes l :: "location" fixes p_ph_enabled :: "('field, 'mut, 'ref) lsts_pred" assumes p_ph_enabled_eq_imp: "eq_imp (\(_::unit) s. s p) p_ph_enabled" begin abbreviation (input) "p_cas_mark s \ cas_mark (s p)" abbreviation (input) "p_mark s \ mark (s p)" abbreviation (input) "p_fM s \ fM (s p)" abbreviation (input) "p_ghost_handshake_phase s \ ghost_handshake_phase (s p)" abbreviation (input) "p_ghost_honorary_grey s \ ghost_honorary_grey (s p)" abbreviation (input) "p_ghost_handshake_in_sync s \ ghost_handshake_in_sync (s p)" abbreviation (input) "p_phase s \ phase (s p)" abbreviation (input) "p_ref s \ ref (s p)" abbreviation (input) "p_the_ref \ the \ p_ref" abbreviation (input) "p_W s \ W (s p)" abbreviation at_p :: "location \ ('field, 'mut, 'ref) lsts_pred \ ('field, 'mut, 'ref) gc_pred" where "at_p l' P \ at p (l @ l') \<^bold>\ LSTP P" abbreviation (input) "p_en_cond P \ p_ph_enabled \<^bold>\ P" abbreviation (input) "p_valid_ref \ \<^bold>\(NULL p_ref) \<^bold>\ valid_ref \<^bold>$ p_the_ref" abbreviation (input) "p_tso_no_pending_mark \ LIST_NULL (tso_pending_mark p)" abbreviation (input) "p_tso_no_pending_mutate \ LIST_NULL (tso_pending_mutate p)" abbreviation (input) "p_valid_W_inv \ ((p_cas_mark \<^bold>\ p_mark \<^bold>\ p_tso_no_pending_mark) \<^bold>\ marked \<^bold>$ p_the_ref) \<^bold>\ (tso_pending_mark p \<^bold>\ (\s. {[], [mw_Mark (p_the_ref s) (p_fM s)]}) )" abbreviation (input) "p_mark_inv \ \<^bold>\(NULL p_mark) \<^bold>\ ((\s. obj_at (\obj. Some (obj_mark obj) = p_mark s) (p_the_ref s) s) \<^bold>\ marked \<^bold>$ p_the_ref)" abbreviation (input) "p_cas_mark_inv \ (\s. obj_at (\obj. Some (obj_mark obj) = p_cas_mark s) (p_the_ref s) s)" abbreviation (input) "p_valid_fM \ p_fM \<^bold>= sys_fM" abbreviation (input) "p_ghg_eq_ref \ p_ghost_honorary_grey \<^bold>= pred_singleton (the \ p_ref)" abbreviation (input) "p_ghg_inv \ If p_cas_mark \<^bold>= p_mark Then p_ghg_eq_ref Else EMPTY p_ghost_honorary_grey" definition mark_object_invL :: "('field, 'mut, 'ref) gc_pred" where "mark_object_invL = (at_p ''_mo_null'' \True\ \<^bold>\ at_p ''_mo_mark'' (p_valid_ref) \<^bold>\ at_p ''_mo_fM'' (p_valid_ref \<^bold>\ p_en_cond (p_mark_inv)) \<^bold>\ at_p ''_mo_mtest'' (p_valid_ref \<^bold>\ p_en_cond (p_mark_inv \<^bold>\ p_valid_fM)) \<^bold>\ at_p ''_mo_phase'' (p_valid_ref \<^bold>\ p_mark \<^bold>\ Some \ p_fM \<^bold>\ p_en_cond (p_mark_inv \<^bold>\ p_valid_fM)) \<^bold>\ at_p ''_mo_ptest'' (p_valid_ref \<^bold>\ p_mark \<^bold>\ Some \ p_fM \<^bold>\ p_en_cond (p_mark_inv \<^bold>\ p_valid_fM)) \<^bold>\ at_p ''_mo_co_lock'' (p_valid_ref \<^bold>\ p_mark_inv \<^bold>\ p_valid_fM \<^bold>\ p_mark \<^bold>\ Some \ p_fM \<^bold>\ p_tso_no_pending_mark) \<^bold>\ at_p ''_mo_co_cmark'' (p_valid_ref \<^bold>\ p_mark_inv \<^bold>\ p_valid_fM \<^bold>\ p_mark \<^bold>\ Some \ p_fM \<^bold>\ p_tso_no_pending_mark) \<^bold>\ at_p ''_mo_co_ctest'' (p_valid_ref \<^bold>\ p_mark_inv \<^bold>\ p_valid_fM \<^bold>\ p_mark \<^bold>\ Some \ p_fM \<^bold>\ p_cas_mark_inv \<^bold>\ p_tso_no_pending_mark) \<^bold>\ at_p ''_mo_co_mark'' (p_cas_mark \<^bold>= p_mark \<^bold>\ p_valid_ref \<^bold>\ p_valid_fM \<^bold>\ white \<^bold>$ p_the_ref \<^bold>\ p_tso_no_pending_mark) \<^bold>\ at_p ''_mo_co_unlock'' (p_ghg_inv \<^bold>\ p_valid_ref \<^bold>\ p_valid_fM \<^bold>\ p_valid_W_inv) \<^bold>\ at_p ''_mo_co_won'' (p_ghg_inv \<^bold>\ p_valid_ref \<^bold>\ p_valid_fM \<^bold>\ marked \<^bold>$ p_the_ref \<^bold>\ p_tso_no_pending_mutate) \<^bold>\ at_p ''_mo_co_W'' (p_ghg_eq_ref \<^bold>\ p_valid_ref \<^bold>\ p_valid_fM \<^bold>\ marked \<^bold>$ p_the_ref \<^bold>\ p_tso_no_pending_mutate))" (*<*) lemma mark_object_invL_eq_imp: "eq_imp (\(_::unit) s. (AT s p, s\ p, sys_heap s\, sys_fM s\, sys_mem_write_buffers p s\)) mark_object_invL" apply (clarsimp simp: eq_imp_def) apply (rename_tac s s') apply (cut_tac s="s\" and s'="s'\" in eq_impD[OF p_ph_enabled_eq_imp], simp) apply (clarsimp simp: mark_object_invL_def obj_at_def cong: option.case_cong) done lemmas mark_object_invL_niE[nie] = iffD1[OF mark_object_invL_eq_imp[simplified eq_imp_simps, rule_format, unfolded conj_explode], rotated -1] (*>*) end text\ The uses of @{const "mark_object_fn"} in the GC and during the root marking are straightforward. \ (* FIXME we'd like: sublocale mut < get_roots: mark_object "mutator mut" "''hs_get_roots_loop''" . but this doesn't seem to get promoted to the top-level, so we can't use it in the other process locales. This interpretation promotes the [inv] attribute to the top-level. *) interpretation gc_mark: mark_object "gc" "''mark_loop''" "\True\" by standard (simp add: eq_imp_def) lemmas gc_mark_mark_object_invL_def2[inv] = gc_mark.mark_object_invL_def[simplified] interpretation mut_get_roots: mark_object "mutator m" "''hs_get_roots_loop''" "\True\" for m by standard (simp add: eq_imp_def) lemmas mut_get_roots_mark_object_invL_def2[inv] = mut_get_roots.mark_object_invL_def[simplified] text\ The most interesting cases are the two asynchronous uses of @{const "mark_object_fn"} in the mutators: we need something that holds even before we read the phase. In particular we need to avoid interference by an @{const "fM"} flip. \ interpretation mut_store_del: mark_object "mutator m" "''store_del''" "mut_m.mut_ghost_handshake_phase m \<^bold>\ \hp_Idle\" for m by standard (simp add: eq_imp_def) lemmas mut_store_del_mark_object_invL_def2[inv] = mut_store_del.mark_object_invL_def[simplified] interpretation mut_store_ins: mark_object "mutator m" "''store_ins''" "mut_m.mut_ghost_handshake_phase m \<^bold>\ \hp_Idle\" for m by standard (simp add: eq_imp_def) lemmas mut_store_ins_mark_object_invL_def2[inv] = mut_store_ins.mark_object_invL_def[simplified] text\ Local invariant for the mutator's uses of @{term "mark_object"}. \ locset_definition "mut_hs_get_roots_loop_locs = prefixed ''hs_get_roots_loop''" locset_definition "mut_hs_get_roots_loop_mo_locs = prefixed ''hs_get_roots_loop_mo'' \ {''hs_get_roots_loop_done''}" abbreviation "mut_async_mark_object_prefixes \ { ''store_del'', ''store_ins'' }" locset_definition "mut_hs_not_hp_Idle_locs = (\pref\mut_async_mark_object_prefixes. \l\{''mo_co_lock'', ''mo_co_cmark'', ''mo_co_ctest'', ''mo_co_mark'', ''mo_co_unlock'', ''mo_co_won'', ''mo_co_W''}. {pref @ ''_'' @ l})" locset_definition "mut_async_mo_ptest_locs = (\pref\mut_async_mark_object_prefixes. {pref @ ''_mo_ptest''})" locset_definition "mut_mo_ptest_locs = (\pref\mut_async_mark_object_prefixes. {pref @ ''_mo_ptest''})" locset_definition "mut_mo_valid_ref_locs = (prefixed ''store_del'' \ prefixed ''store_ins'' \ { ''deref_del'', ''lop_store_ins''})" (*<*) lemma mut_hs_get_roots_loop_locs_subseteq_hs_get_roots: "mut_hs_get_roots_loop_locs \ prefixed ''hs_get_roots_''" by (auto simp: mut_hs_get_roots_loop_locs_def intro: append_prefixD) lemma mut_m_ghost_handshake_phase_not_hp_Idle: "\ atS (mutator m) mut_hs_get_roots_loop_locs s; mut_m.handshake_invL m s; handshake_phase_inv s\ \ \ ghost_handshake_phase (s\ (mutator m)) \ hp_Idle" unfolding mut_m.handshake_invL_def apply (elim conjE) apply (drule mut_m.handshake_phase_invD[where m=m]) apply (drule mp, erule atS_mono[OF _ mut_hs_get_roots_loop_locs_subseteq_hs_get_roots]) apply (clarsimp simp: hp_step_rel_def) done (*>*) text\ This local invariant for the mutators illustrates the handshake structure: we can rely on the insertion barrier earlier than on the deletion barrier. Both need to be installed before \get_roots\ to ensure we preserve the strong tricolour invariant. All black objects at that point are allocated: we need to know that the insertion barrier is installed to preserve it. This limits when \fA\ can be set. It is interesting to contrast the two barriers. Intuitively a mutator can locally guarantee that it, in the relevant phases, will insert only marked references. Less often can it be sure that the reference it is overwriting is marked. We also need to consider writes pending in TSO buffers: it is key that after the \''init_noop''\ handshake there are no pending white insertions (mutations that insert unmarked references). This ensures the deletion barrier does its job. \ locset_definition "ghost_honorary_grey_empty_locs \ - (\pref\{ ''mark_loop'', ''hs_get_roots_loop'', ''store_del'', ''store_ins'' }. \l\{ ''mo_co_unlock'', ''mo_co_won'', ''mo_co_W'' }. {pref @ ''_'' @ l})" locset_definition "ghost_honorary_root_empty_locs \ - (prefixed ''store_del'' \ {''lop_store_ins''} \ prefixed ''store_ins'')" inv_definition (in mut_m) mark_object_invL :: "('field, 'mut, 'ref) gc_pred" where "mark_object_invL = (atS_mut mut_hs_get_roots_loop_locs (mut_refs \<^bold>\ mut_roots \<^bold>\ (\<^bold>\r. \r\ \<^bold>\ mut_roots \<^bold>- mut_refs \<^bold>\ marked r)) \<^bold>\ atS_mut mut_hs_get_roots_loop_mo_locs (\<^bold>\(NULL mut_ref) \<^bold>\ mut_the_ref \<^bold>\ mut_roots) \<^bold>\ at_mut ''hs_get_roots_loop_done'' (marked \<^bold>$ mut_the_ref) \<^bold>\ at_mut ''hs_get_roots_loop_mo_ptest'' (mut_phase \<^bold>\ \ph_Idle\) \<^bold>\ at_mut ''hs_get_roots_done'' (\<^bold>\r. \r\ \<^bold>\ mut_roots \<^bold>\ marked r) \<^bold>\ atS_mut mut_mo_valid_ref_locs ( (\<^bold>\(NULL mut_new_ref) \<^bold>\ mut_the_new_ref \<^bold>\ mut_roots) \<^bold>\ (mut_tmp_ref \<^bold>\ mut_roots) ) \<^bold>\ at_mut ''store_del_mo_null'' (\<^bold>\(NULL mut_ref) \<^bold>\ mut_the_ref \<^bold>\ mut_ghost_honorary_root) \<^bold>\ atS_mut (prefixed ''store_del'' - {''store_del_mo_null''}) (mut_the_ref \<^bold>\ mut_ghost_honorary_root) \<^bold>\ atS_mut (prefixed ''store_ins'') (mut_ref \<^bold>= mut_new_ref) \<^bold>\ atS_mut (suffixed ''_mo_ptest'') (mut_phase \<^bold>\ \ph_Idle\ \<^bold>\ mut_ghost_handshake_phase \<^bold>\ \hp_Idle\) \<^bold>\ atS_mut mut_hs_not_hp_Idle_locs (mut_ghost_handshake_phase \<^bold>\ \hp_Idle\) \<^bold>\ atS_mut mut_mo_ptest_locs (mut_phase \<^bold>= \ph_Idle\ \<^bold>\ (mut_ghost_handshake_phase \<^bold>\ \{hp_Idle, hp_IdleInit}\ \<^bold>\ (mut_ghost_handshake_phase \<^bold>= \hp_IdleMarkSweep\ \<^bold>\ sys_phase \<^bold>= \ph_Idle\))) \<^bold>\ atS_mut ghost_honorary_grey_empty_locs (EMPTY mut_ghost_honorary_grey) \ \insertion barrier\ \<^bold>\ at_mut ''store_ins'' ( (mut_ghost_handshake_phase \<^bold>\ \{hp_InitMark, hp_Mark}\ \<^bold>\ (mut_ghost_handshake_phase \<^bold>= \hp_IdleMarkSweep\ \<^bold>\ sys_phase \<^bold>\ \ph_Idle\)) \<^bold>\ \<^bold>\(NULL mut_new_ref) \<^bold>\ marked \<^bold>$ mut_the_new_ref ) \ \deletion barrier\ \<^bold>\ atS_mut (prefixed ''store_del_mo'' \ {''lop_store_ins''}) ( (mut_ghost_handshake_phase \<^bold>= \hp_Mark\ \<^bold>\ (mut_ghost_handshake_phase \<^bold>= \hp_IdleMarkSweep\ \<^bold>\ sys_phase \<^bold>\ \ph_Idle\)) \<^bold>\ (\s. \opt_r'. \tso_pending_write (mutator m) (mw_Mutate (mut_tmp_ref s) (mut_field s) opt_r') s) \<^bold>\ (\s. obj_at_field_on_heap (\r. mut_ref s = Some r \ marked r s) (mut_tmp_ref s) (mut_field s) s)) \<^bold>\ at_mut ''lop_store_ins'' ( (mut_ghost_handshake_phase \<^bold>= \hp_Mark\ \<^bold>\ (mut_ghost_handshake_phase \<^bold>= \hp_IdleMarkSweep\ \<^bold>\ sys_phase \<^bold>\ \ph_Idle\)) \<^bold>\ \<^bold>\(NULL mut_ref) \<^bold>\ marked \<^bold>$ mut_the_ref ) \<^bold>\ atS_mut (prefixed ''store_ins'') ( (mut_ghost_handshake_phase \<^bold>= \hp_Mark\ \<^bold>\ (mut_ghost_handshake_phase \<^bold>= \hp_IdleMarkSweep\ \<^bold>\ sys_phase \<^bold>\ \ph_Idle\)) \<^bold>\ (\s. \opt_r'. \tso_pending_write (mutator m) (mw_Mutate (mut_tmp_ref s) (mut_field s) opt_r') s) \<^bold>\ (\s. obj_at_field_on_heap (\r'. marked r' s) (mut_tmp_ref s) (mut_field s) s) ) \\after \''init_noop''\\ \<^bold>\ at_mut ''load'' (mut_tmp_ref \<^bold>\ mut_roots) \<^bold>\ at_mut ''hs_noop_done'' (LIST_NULL (tso_pending_mutate (mutator m))) \\ key: no pending white insertions \ \<^bold>\ atS_mut ghost_honorary_root_empty_locs (EMPTY mut_ghost_honorary_root) )" (*<*) lemma get_roots_get_work_subseteq_ghost_honorary_grey_empty_locs: "hs_get_roots_locs \ hs_get_work_locs \ ghost_honorary_grey_empty_locs" unfolding ghost_honorary_grey_empty_locs_def hs_get_roots_locs_def hs_get_work_locs_def apply (clarsimp simp: subset_eq) apply (intro conjI conjI; force) done lemma (in mut_m) mark_object_invL_eq_imp: "eq_imp (\r s. (AT s (mutator m), s\ (mutator m), sys_heap s\ r, sys_fM s\, sys_phase s\, tso_pending_mutate (mutator m) s\)) mark_object_invL" apply (clarsimp simp: eq_imp_def mark_object_invL_def fun_eq_iff[symmetric] obj_at_field_on_heap_def cong: option.case_cong) apply (rename_tac s s') apply (subgoal_tac "\r. marked r s\ \ marked r s'\") apply (subgoal_tac "\r. valid_null_ref r s\ \ valid_null_ref r s'\") apply (subgoal_tac "\r f opt_r'. mw_Mutate r f opt_r' \ set (sys_mem_write_buffers (mutator m) s\) \ mw_Mutate r f opt_r' \ set (sys_mem_write_buffers (mutator m) s'\)") apply (clarsimp cong: option.case_cong) apply (drule arg_cong[where f=set]) apply auto[1] apply (clarsimp simp: obj_at_def valid_null_ref_def split: option.splits) apply (clarsimp simp: obj_at_def valid_null_ref_def split: option.splits) done lemmas mut_m_mark_object_invL_niE[nie] = iffD1[OF mut_m.mark_object_invL_eq_imp[simplified eq_imp_simps, rule_format, unfolded conj_explode], rotated -1] lemma (in mut_m) mark_object_invL[intro]: "\ handshake_invL \<^bold>\ mark_object_invL \<^bold>\ mut_get_roots.mark_object_invL m \<^bold>\ mut_store_del.mark_object_invL m \<^bold>\ mut_store_ins.mark_object_invL m \<^bold>\ LSTP (phase_rel_inv \<^bold>\ handshake_phase_inv \<^bold>\ phase_rel_inv \<^bold>\ tso_writes_inv \<^bold>\ valid_refs_inv) \ mutator m \ mark_object_invL \" apply vcg_jackhammer (* store_ins_mo_ptest *) subgoal by (elim disjE; fastforce) (* store_ins_mo_ptest *) subgoal for s s' y apply (drule handshake_phase_invD) apply (drule phase_rel_invD) apply (clarsimp simp: phase_rel_def) apply (case_tac "sys_ghost_handshake_phase s\"; simp add: hp_step_rel_def; elim disjE; simp; force) done subgoal by fastforce subgoal by fastforce subgoal by fastforce subgoal for s s' y apply (drule handshake_phase_invD) apply (drule phase_rel_invD) apply (clarsimp simp: phase_rel_def) apply (case_tac "sys_ghost_handshake_phase s\"; simp add: hp_step_rel_def; elim disjE; simp; force) done subgoal by fastforce subgoal by (auto dest: obj_at_field_on_heap_no_pending_writes) (* hs get roots loop done *) subgoal by force (* hs get roots loop mo phase *) subgoal apply (drule handshake_phase_invD) apply (drule phase_rel_invD) apply (clarsimp simp: phase_rel_def hp_step_rel_def) done (* hs get roots loop choose ref *) subgoal by blast done lemma (in mut_m') mut_mark_object_invL[intro]: "\ mark_object_invL \ mutator m'" apply vcg_nihe apply vcg_ni apply (clarsimp simp: obj_at_field_on_heap_def split: option.splits) apply (clarsimp simp: obj_at_field_on_heap_def split: option.splits) done lemma (in mut_m) mut_store_ins_mark_object_invL[intro]: "\ mut_store_ins.mark_object_invL m \<^bold>\ mark_object_invL \<^bold>\ handshake_invL \<^bold>\ tso_lock_invL \<^bold>\ LSTP (handshake_phase_inv \<^bold>\ valid_W_inv \<^bold>\ tso_writes_inv \<^bold>\ valid_refs_inv) \ mutator m \ mut_store_ins.mark_object_invL m \" apply vcg_jackhammer apply (auto dest: valid_refs_invD valid_W_inv_no_mark_writes_invD split: obj_at_splits) done lemma mut_m_not_idle_no_fM_writeD: "\ sys_mem_write_buffers p s = mw_fM fl # ws; ghost_handshake_phase (s (mutator m)) \ hp_Idle; fM_rel_inv s; handshake_phase_inv s; tso_writes_inv s; p \ sys \ \ False" apply (drule mut_m.handshake_phase_invD[where m=m]) apply (drule fM_rel_invD) apply (fastforce simp: hp_step_rel_def fM_rel_def filter_empty_conv p_not_sys) done lemma (in sys) mut_store_ins_mark_object_invL[intro]: notes mut_m.mark_object_invL_def[inv] notes mut_m.tso_lock_invL_def[inv] notes mut_m_not_idle_no_fM_writeD[where m=m, dest!] notes map_option.compositionality[simp] o_def[simp] shows "\ mut_m.tso_lock_invL m \<^bold>\ mut_m.mark_object_invL m \<^bold>\ mut_store_ins.mark_object_invL m \<^bold>\ LSTP (fM_rel_inv \<^bold>\ handshake_phase_inv \<^bold>\ valid_W_inv \<^bold>\ tso_writes_inv) \ sys \ mut_store_ins.mark_object_invL m \" apply (vcg_ni simp: not_blocked_def) subgoal by (auto simp: do_write_action_def split: mem_write_action.splits obj_at_splits) subgoal by (auto simp: do_write_action_def split: mem_write_action.splits obj_at_splits) subgoal by (auto simp: do_write_action_def split: mem_write_action.splits obj_at_splits) subgoal by (auto simp: do_write_action_def split: mem_write_action.splits obj_at_splits) subgoal by (auto simp: do_write_action_def split: mem_write_action.splits obj_at_splits) subgoal by (auto simp: do_write_action_def split: mem_write_action.splits obj_at_splits) subgoal by (auto simp: do_write_action_def split: mem_write_action.splits obj_at_splits) subgoal by (auto simp: do_write_action_def split: mem_write_action.splits obj_at_splits) subgoal by (auto simp: do_write_action_def split: mem_write_action.splits obj_at_splits) subgoal by (auto simp: do_write_action_def split: mem_write_action.splits obj_at_splits) subgoal by (auto simp: do_write_action_def split: mem_write_action.splits obj_at_splits) done lemma (in sys) mut_get_roots_mark_object_invL[intro]: notes mut_m.handshake_invL_def[inv] notes mut_m.tso_lock_invL_def[inv] notes map_option.compositionality[simp] o_def[simp] shows "\ mut_m.tso_lock_invL m \<^bold>\ mut_m.handshake_invL m \<^bold>\ mut_get_roots.mark_object_invL m \<^bold>\ LSTP (fM_rel_inv \<^bold>\ handshake_phase_inv \<^bold>\ valid_W_inv \<^bold>\ tso_writes_inv) \ sys \ mut_get_roots.mark_object_invL m \" apply (vcg_ni simp: not_blocked_def p_not_sys dest!: mut_m.handshake_phase_invD[where m=m]) subgoal by (auto simp: do_write_action_def fM_rel_inv_def fM_rel_def hp_step_rel_def filter_empty_conv split: mem_write_action.splits if_splits obj_at_splits) subgoal by (auto simp: do_write_action_def fM_rel_inv_def fM_rel_def hp_step_rel_def filter_empty_conv split: mem_write_action.splits if_splits obj_at_splits) subgoal by (auto simp: do_write_action_def fM_rel_inv_def fM_rel_def hp_step_rel_def filter_empty_conv split: mem_write_action.splits if_splits obj_at_splits) subgoal by (auto simp: do_write_action_def fM_rel_inv_def fM_rel_def hp_step_rel_def filter_empty_conv split: mem_write_action.splits if_splits obj_at_splits) subgoal by (auto simp: do_write_action_def fM_rel_inv_def fM_rel_def hp_step_rel_def filter_empty_conv split: mem_write_action.splits if_splits obj_at_splits) subgoal by (auto simp: do_write_action_def fM_rel_inv_def fM_rel_def hp_step_rel_def filter_empty_conv split: mem_write_action.splits if_splits obj_at_splits) subgoal by (auto simp: do_write_action_def fM_rel_inv_def fM_rel_def hp_step_rel_def filter_empty_conv split: mem_write_action.splits if_splits obj_at_splits) subgoal by (auto simp: do_write_action_def fM_rel_inv_def fM_rel_def hp_step_rel_def filter_empty_conv split: mem_write_action.splits if_splits obj_at_splits) subgoal by (auto simp: do_write_action_def fM_rel_inv_def fM_rel_def hp_step_rel_def filter_empty_conv split: mem_write_action.splits if_splits obj_at_splits) subgoal by (auto simp: do_write_action_def fM_rel_inv_def fM_rel_def hp_step_rel_def filter_empty_conv split: mem_write_action.splits if_splits obj_at_splits) subgoal by (auto simp: do_write_action_def fM_rel_inv_def fM_rel_def hp_step_rel_def filter_empty_conv split: mem_write_action.splits if_splits obj_at_splits) done lemma (in sys) mut_store_del_mark_object_invL[intro]: notes mut_m.mark_object_invL_def[inv] notes mut_m.tso_lock_invL_def[inv] notes mut_m_not_idle_no_fM_writeD[where m=m, dest!] notes map_option.compositionality[simp] o_def[simp] shows "\ mut_m.tso_lock_invL m \<^bold>\ mut_m.mark_object_invL m \<^bold>\ mut_store_del.mark_object_invL m \<^bold>\ LSTP (fM_rel_inv \<^bold>\ handshake_phase_inv \<^bold>\ valid_W_inv \<^bold>\ tso_writes_inv) \ sys \ mut_store_del.mark_object_invL m \" apply (vcg_ni simp: not_blocked_def) subgoal by (auto simp: do_write_action_def split: mem_write_action.splits obj_at_splits) subgoal by (auto simp: do_write_action_def split: mem_write_action.splits obj_at_splits) subgoal by (auto simp: do_write_action_def split: mem_write_action.splits obj_at_splits) subgoal by (auto simp: do_write_action_def split: mem_write_action.splits obj_at_splits) subgoal by (fastforce simp: do_write_action_def split: mem_write_action.splits obj_at_splits) subgoal by (auto simp: do_write_action_def split: mem_write_action.splits obj_at_splits) subgoal by (auto simp: do_write_action_def split: mem_write_action.splits obj_at_splits) subgoal by (auto simp: do_write_action_def split: mem_write_action.splits obj_at_splits) subgoal by (auto simp: do_write_action_def split: mem_write_action.splits obj_at_splits) subgoal by (auto simp: do_write_action_def split: mem_write_action.splits obj_at_splits) subgoal by (auto simp: do_write_action_def split: mem_write_action.splits obj_at_splits) done lemma (in mut_m) mut_store_del_mark_object_invL[intro]: "\ mut_store_del.mark_object_invL m \<^bold>\ mark_object_invL \<^bold>\ handshake_invL \<^bold>\ tso_lock_invL \<^bold>\ LSTP (handshake_phase_inv \<^bold>\ valid_W_inv \<^bold>\ tso_writes_inv \<^bold>\ valid_refs_inv) \ mutator m \ mut_store_del.mark_object_invL m \" by vcg_jackhammer (auto dest: valid_refs_invD valid_W_inv_no_mark_writes_invD split: obj_at_splits) lemma (in mut_m) mut_get_roots_mark_object_invL[intro]: "\ mut_get_roots.mark_object_invL m \<^bold>\ mark_object_invL \<^bold>\ handshake_invL \<^bold>\ tso_lock_invL \<^bold>\ LSTP (handshake_phase_inv \<^bold>\ valid_W_inv \<^bold>\ tso_writes_inv \<^bold>\ valid_refs_inv) \ mutator m \ mut_get_roots.mark_object_invL m \" by vcg_jackhammer (auto dest: valid_W_inv_no_mark_writes_invD split: obj_at_splits) lemma (in mut_m') mut_get_roots_mark_object_invL[intro]: "\ mut_get_roots.mark_object_invL m \ mutator m'" by vcg_nihe vcg_ni lemma (in mut_m') mut_store_ins_mark_object_invL[intro]: "\ mut_store_ins.mark_object_invL m \ mutator m'" by vcg_nihe vcg_ni lemma (in mut_m') mut_store_del_mark_object_invL[intro]: "\ mut_store_del.mark_object_invL m \ mutator m'" by vcg_nihe vcg_ni lemma (in gc) mut_get_roots_mark_object_invL[intro]: notes mut_m.handshake_invL_def[inv] shows "\ handshake_invL \<^bold>\ mut_m.handshake_invL m \<^bold>\ mut_get_roots.mark_object_invL m \ gc \ mut_get_roots.mark_object_invL m \" by vcg_nihe vcg_ni (*>*) text\ We now show that the GC's use of @{const "mark_object_fn"} is correct. When we take grey @{const "tmp_ref"} to black, all of the objects it points to are marked, ergo the new black does not point to white, and so we preserve the strong tricolour invariant. \ definition (in gc) obj_fields_marked_inv :: "('field, 'mut, 'ref) lsts_pred" where "obj_fields_marked_inv = (\<^bold>\f. \f\ \<^bold>\ (- gc_field_set) \<^bold>\ (\s. obj_at_field_on_heap (\r. marked r s) (gc_tmp_ref s) f s))" (*<*) lemma (in gc) obj_fields_marked_inv_eq_imp: "eq_imp (\_::unit. gc_field_set \<^bold>\ gc_tmp_ref \<^bold>\ sys_heap \<^bold>\ sys_fM \<^bold>\ tso_pending_mutate gc) obj_fields_marked_inv" by (clarsimp simp: eq_imp_def obj_fields_marked_inv_def obj_at_field_on_heap_def obj_at_def cong: option.case_cong) lemmas gc_obj_fields_marked_inv_fun_upd[simp] = eq_imp_fun_upd[OF gc.obj_fields_marked_inv_eq_imp, simplified eq_imp_simps, rule_format] lemma (in gc) obj_fields_marked_inv_UNIV[iff]: "obj_fields_marked_inv (s(gc := (s gc)\ field_set := UNIV \))" by (simp_all add: obj_fields_marked_inv_def) lemma (in gc) obj_fields_marked_inv_mark_field_done[iff]: "\ obj_at_field_on_heap (\r. marked r s) (gc_tmp_ref s) (gc_field s) s; obj_fields_marked_inv s \ \ obj_fields_marked_inv (s(gc := (s gc)\field_set := gc_field_set s - {gc_field s}\))" by (force simp: obj_fields_marked_inv_def obj_at_field_on_heap_def split: option.splits obj_at_splits) (*>*) text\\ locset_definition "obj_fields_marked_locs \ { ''mark_loop_mark_object_loop'', ''mark_loop_mark_choose_field'', ''mark_loop_mark_deref'', ''mark_loop_mark_field_done'', ''mark_loop_blacken'' } \ prefixed ''mark_loop_mo''" inv_definition (in gc) obj_fields_marked_invL :: "('field, 'mut, 'ref) gc_pred" where "obj_fields_marked_invL \ (atS_gc obj_fields_marked_locs (obj_fields_marked_inv \<^bold>\ gc_tmp_ref \<^bold>\ gc_W) \<^bold>\ atS_gc (prefixed ''mark_loop_mo'' \ { ''mark_loop_mark_field_done'' }) (\s. obj_at_field_on_heap (\r. gc_ref s = Some r \ marked r s) (gc_tmp_ref s) (gc_field s) s) \<^bold>\ atS_gc (prefixed ''mark_loop_mo'') (\<^bold>\y. \<^bold>\(NULL gc_ref) \<^bold>\ (\s. ((gc_the_ref s) reaches y) s) \<^bold>\ valid_ref y) \<^bold>\ at_gc ''mark_loop_fields'' (gc_tmp_ref \<^bold>\ gc_W) \<^bold>\ at_gc ''mark_loop_mark_field_done'' (\<^bold>\(NULL gc_ref) \<^bold>\ marked \<^bold>$ gc_the_ref) \<^bold>\ at_gc ''mark_loop_blacken'' (EMPTY gc_field_set) \<^bold>\ atS_gc ghost_honorary_grey_empty_locs (EMPTY gc_ghost_honorary_grey))" (*<*) lemma (in gc) obj_fields_marked_invL_eq_imp: "eq_imp (\(_::unit) (s :: ('field, 'mut, 'ref) gc_pred_state). (AT s gc, s\ gc, sys_heap s\, sys_fM s\, sys_W s\, tso_pending_mutate gc s\)) obj_fields_marked_invL" by (clarsimp simp: eq_imp_def inv obj_at_def obj_fields_marked_inv_def obj_at_field_on_heap_def cong: option.case_cong) lemmas gc_obj_fields_marked_invL_niE[nie] = iffD1[OF gc.obj_fields_marked_invL_eq_imp[simplified eq_imp_simps, rule_format, unfolded conj_explode], rotated -1] lemma (in gc) gc_mark_mark_object_invL[intro]: "\ fM_fA_invL \<^bold>\ gc_mark.mark_object_invL \<^bold>\ obj_fields_marked_invL \<^bold>\ tso_lock_invL \<^bold>\ LSTP valid_W_inv \ gc \ gc_mark.mark_object_invL \" by vcg_jackhammer (auto dest: valid_W_inv_no_mark_writes_invD split: obj_at_splits) lemma (in gc) obj_fields_marked_invL[intro]: "\ fM_fA_invL \<^bold>\ phase_invL \<^bold>\ obj_fields_marked_invL \<^bold>\ gc_mark.mark_object_invL \<^bold>\ LSTP (tso_writes_inv \<^bold>\ valid_W_inv \<^bold>\ valid_refs_inv) \ gc \ obj_fields_marked_invL \" apply vcg_jackhammer (* mark_loop_mark_field_done *) apply (rule obj_fields_marked_inv_mark_field_done, auto)[1] (* mark_loop_mark_deref *) apply (rename_tac s s') apply (subgoal_tac "grey (gc_tmp_ref s\) s\") (* FIXME rule *) apply (clarsimp simp: obj_at_field_on_heap_def split: option.splits) apply (frule valid_refs_invD, fastforce, fastforce) apply (rule conjI) apply (clarsimp split: obj_at_splits) apply clarsimp apply (rename_tac x) apply (subgoal_tac "(gc_tmp_ref s\ reaches x) s\") apply (erule valid_refs_invD, fastforce, fastforce) apply (fastforce elim: converse_rtranclp_into_rtranclp[rotated] simp: ran_def split: obj_at_splits) apply blast done lemma (in mut_m) gc_obj_fields_marked_invL[intro]: notes gc.obj_fields_marked_invL_def[inv] notes gc.handshake_invL_def[inv] shows "\ handshake_invL \<^bold>\ gc.handshake_invL \<^bold>\ gc.obj_fields_marked_invL \<^bold>\ LSTP (tso_writes_inv \<^bold>\ valid_refs_inv) \ mutator m \ gc.obj_fields_marked_invL \" apply vcg_nihe apply vcg_ni (* FIXME rules *) apply (clarsimp simp: gc.obj_fields_marked_inv_def) apply (rename_tac s s' ra x) apply (drule_tac x=x in spec) apply clarsimp apply (erule obj_at_field_on_heapE) apply (subgoal_tac "grey (gc_tmp_ref s\) s\") apply (drule_tac y="gc_tmp_ref s\" in valid_refs_invD(7), simp+) apply (clarsimp split: obj_at_splits) apply (erule greyI) apply (clarsimp split: obj_at_splits) apply (clarsimp simp: obj_at_field_on_heap_def split: option.splits) apply vcg_ni+ done lemma (in mut_m) gc_mark_mark_object_invL[intro]: "\ gc_mark.mark_object_invL \ mutator m" by vcg_nihe vcg_ni lemma (in sys) gc_mark_mark_object_invL[intro]: notes gc.tso_lock_invL_def[inv] notes gc.phase_invL_def[inv] notes gc.fM_fA_invL_def[inv] notes gc.handshake_invL_def[inv] notes map_option.compositionality[simp] o_def[simp] shows "\ gc.fM_fA_invL \<^bold>\ gc.handshake_invL \<^bold>\ gc.phase_invL \<^bold>\ gc_mark.mark_object_invL \<^bold>\ gc.tso_lock_invL \<^bold>\ LSTP (handshake_phase_inv \<^bold>\ phase_rel_inv \<^bold>\ valid_W_inv \<^bold>\ tso_writes_inv) \ sys \ gc_mark.mark_object_invL \" apply vcg_ni subgoal by (auto dest!: valid_W_invD2 simp: do_write_action_def not_blocked_def fM_rel_def filter_empty_conv p_not_sys split: mem_write_action.split if_splits) subgoal by (auto dest!: valid_W_invD2 simp: do_write_action_def not_blocked_def fM_rel_def filter_empty_conv p_not_sys split: mem_write_action.split if_splits) subgoal by (auto dest!: valid_W_invD2 simp: do_write_action_def not_blocked_def fM_rel_def filter_empty_conv p_not_sys split: mem_write_action.split if_splits) subgoal by (auto dest!: valid_W_invD2 simp: do_write_action_def not_blocked_def fM_rel_def filter_empty_conv p_not_sys split: mem_write_action.split if_splits) subgoal by (auto dest!: valid_W_invD2 simp: do_write_action_def not_blocked_def fM_rel_def filter_empty_conv p_not_sys split: mem_write_action.split if_splits) subgoal by (auto dest!: valid_W_invD2 simp: do_write_action_def not_blocked_def fM_rel_def filter_empty_conv p_not_sys split: mem_write_action.split if_splits) subgoal by (auto dest!: valid_W_invD2 simp: do_write_action_def not_blocked_def fM_rel_def filter_empty_conv p_not_sys split: mem_write_action.split if_splits) subgoal by (auto dest!: valid_W_invD2 simp: do_write_action_def not_blocked_def fM_rel_def filter_empty_conv p_not_sys split: mem_write_action.split if_splits) subgoal by (auto dest!: valid_W_invD2 simp: do_write_action_def not_blocked_def fM_rel_def filter_empty_conv p_not_sys split: mem_write_action.split if_splits) subgoal by (auto dest!: valid_W_invD2 simp: do_write_action_def not_blocked_def fM_rel_def filter_empty_conv p_not_sys split: mem_write_action.split if_splits) subgoal by (auto dest!: valid_W_invD2 simp: do_write_action_def not_blocked_def fM_rel_def filter_empty_conv p_not_sys split: mem_write_action.split if_splits) done (*>*) (*<*) end (*>*) diff --git a/thys/ConcurrentGC/Model.thy b/thys/ConcurrentGC/Model.thy --- a/thys/ConcurrentGC/Model.thy +++ b/thys/ConcurrentGC/Model.thy @@ -1,1006 +1,1008 @@ (*<*) (* * Copyright 2015, NICTA * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * @TAG(NICTA_BSD) *) theory Model imports ConcurrentIMP.CIMP "HOL-Library.Monad_Syntax" begin +declare subst_all [simp del] [[simproc del: defined_all]] + (*>*) section\The Schism garbage collector \label{sec:gc-model}\ text \ The following formalises Figures~2.8 (\mark_object_fn\), 2.9 (load and store but not alloc), and 2.15 (garbage collector) of @{cite [cite_macro=citet] "Pizlo201xPhd"}; see also @{cite [cite_macro=citet] "Pizlo+2010PLDI"}. We additionally need to model TSO memory, the handshakes and compare-and-swap (\texttt{CAS}). We closely model things where interference is possible and abstract everything else. \textbf{\emph{NOTE}: this model is for TSO \emph{only}. We elide any details irrelevant for that memory model.} We begin by defining the types of the various parts. Our program locations are labelled with strings for readability. We enumerate the names of the processes in our system. The safety proof treats an arbitary (unbounded) number of mutators. \ type_synonym location = "char list" datatype 'mut process_name = mutator 'mut | gc | sys text \ The garbage collection process can be in one of the following phases. \ datatype gc_phase = ph_Idle | ph_Init | ph_Mark | ph_Sweep text \ The garbage collector instructs mutators to perform certain actions, and blocks until the mutators signal these actions are done. The mutators always respond with their work list (a set of references). The handshake can be of one of the specified types. \ datatype handshake_type = ht_NOOP | ht_GetRoots | ht_GetWork text\ We track how many \texttt{noop} and \texttt{get\_roots} handshakes each process has participated in as ghost state. See \S\ref{sec:gc_ragged_safepoints}. \ datatype handshake_phase = hp_Idle \\ done 1 noop \ | hp_IdleInit | hp_InitMark | hp_Mark \\ done 4 noops \ | hp_IdleMarkSweep \\ done get roots \ definition handshake_step :: "handshake_phase \ handshake_phase" where "handshake_step ph \ case ph of hp_Idle \ hp_IdleInit | hp_IdleInit \ hp_InitMark | hp_InitMark \ hp_Mark | hp_Mark \ hp_IdleMarkSweep | hp_IdleMarkSweep \ hp_Idle" text \ An object consists of a garbage collection mark and a function that maps its fields to values. A value is either a reference or \texttt{NULL}. @{typ "'field"} is the abstract type of fields. @{typ "'ref"} is the abstract type of object references. @{typ "'mut"} is the abstract type of the mutators' names. For simplicity we assume all objects define all fields and ignore all non-reference payload in objects. \ type_synonym gc_mark = bool record ('field, 'ref) object = obj_mark :: "gc_mark" obj_fields :: "'field \ 'ref option" text\ The TSO store buffers track write actions, represented by \('field, 'ref) mem_write_action\. \ datatype ('field, 'ref) mem_write_action = mw_Mark 'ref gc_mark | mw_Mutate 'ref 'field "'ref option" | mw_fA gc_mark | mw_fM gc_mark | mw_Phase gc_phase text\ The following record is the type of all processes's local states. For the mutators and the garbage collector, consider these to be local variables or registers. The system's \fA\, \fM\, \phase\ and \heap\ variables are subject to the TSO memory model, as are all heap operations. \ record ('field, 'mut, 'ref) local_state = \ \System-specific fields\ heap :: "'ref \ ('field, 'ref) object option" \ \TSO memory state\ mem_write_buffers :: "'mut process_name \ ('field, 'ref) mem_write_action list" mem_lock :: "'mut process_name option" \ \The state of the handshakes\ handshake_type :: "handshake_type" handshake_pending :: "'mut \ bool" \ \Ghost state\ ghost_handshake_in_sync :: "'mut \ bool" ghost_handshake_phase :: "handshake_phase" \ \Mutator-specific temporaries\ new_ref :: "'ref option" roots :: "'ref set" ghost_honorary_root :: "'ref set" \ \Garbage collector-specific temporaries\ field_set :: "'field set" mut :: "'mut" muts :: "'mut set" \ \Local variables used by multiple processes\ fA :: "gc_mark" fM :: "gc_mark" cas_mark :: "gc_mark option" field :: "'field" mark :: "gc_mark option" phase :: "gc_phase" tmp_ref :: "'ref" ref :: "'ref option" refs :: "'ref set" W :: "'ref set" \ \Ghost state\ ghost_honorary_grey :: "'ref set" text\ An action is a request by a mutator or the garbage collector to the system. \ datatype ('field, 'ref) mem_read_action = mr_Ref 'ref 'field | mr_Mark 'ref | mr_Phase | mr_fM | mr_fA datatype ('field, 'mut, 'ref) request_op = ro_MFENCE | ro_Read "('field, 'ref) mem_read_action" | ro_Write "('field, 'ref) mem_write_action" | ro_Lock | ro_Unlock | ro_Alloc | ro_Free 'ref | ro_hs_gc_set_type handshake_type | ro_hs_gc_set_pending 'mut | ro_hs_gc_read_pending 'mut | ro_hs_gc_load_W | ro_hs_mut_read_type handshake_type | ro_hs_mut_done "'ref set" abbreviation "ReadfM \ ro_Read mr_fM" abbreviation "ReadMark r \ ro_Read (mr_Mark r)" abbreviation "ReadPhase \ ro_Read mr_Phase" abbreviation "ReadRef r f \ ro_Read (mr_Ref r f)" abbreviation "WritefA m \ ro_Write (mw_fA m)" abbreviation "WritefM m \ ro_Write (mw_fM m)" abbreviation "WriteMark r m \ ro_Write (mw_Mark r m)" abbreviation "WritePhase ph \ ro_Write (mw_Phase ph)" abbreviation "WriteRef r f r' \ ro_Write (mw_Mutate r f r')" type_synonym ('field, 'mut, 'ref) request = "'mut process_name \ ('field, 'mut, 'ref) request_op" datatype ('field, 'ref) response = mv_Bool "bool" | mv_Mark "gc_mark option" | mv_Phase "gc_phase" | mv_Ref "'ref option" | mv_Refs "'ref set" | mv_Void text\We instantiate CIMP's types as follows:\ type_synonym ('field, 'mut, 'ref) gc_com = "(('field, 'ref) response, location, ('field, 'mut, 'ref) request, ('field, 'mut, 'ref) local_state) com" type_synonym ('field, 'mut, 'ref) gc_loc_comp = "(('field, 'ref) response, location, ('field, 'mut, 'ref) request, ('field, 'mut, 'ref) local_state) loc_comp" type_synonym ('field, 'mut, 'ref) gc_pred_state = "(('field, 'ref) response, location, 'mut process_name, ('field, 'mut, 'ref) request, ('field, 'mut, 'ref) local_state) pred_state" type_synonym ('field, 'mut, 'ref) gc_pred = "(('field, 'ref) response, location, 'mut process_name, ('field, 'mut, 'ref) request, ('field, 'mut, 'ref) local_state) pred" type_synonym ('field, 'mut, 'ref) gc_system = "(('field, 'ref) response, location, 'mut process_name, ('field, 'mut, 'ref) request, ('field, 'mut, 'ref) local_state) system" type_synonym ('field, 'mut, 'ref) gc_event = "('field, 'mut, 'ref) request \ ('field, 'ref) response" type_synonym ('field, 'mut, 'ref) gc_history = "('field, 'mut, 'ref) gc_event list" type_synonym ('field, 'mut, 'ref) lst_pred = "('field, 'mut, 'ref) local_state \ bool" type_synonym ('field, 'mut, 'ref) lsts = "'mut process_name \ ('field, 'mut, 'ref) local_state" type_synonym ('field, 'mut, 'ref) lsts_pred = "('field, 'mut, 'ref) lsts \ bool" text\ We use one locale per process to define a namespace for definitions local to these processes. Mutator definitions are parametrised by the mutator's identifier \m\. We never interpret these locales; we use their contents typically by prefixing their names the locale name. This might be considered an abuse. The attributes depend on locale scoping somewhat, which is a mixed blessing. If we have more than one mutator then we need to show that mutators do not mutually interfere. To that end we define an extra locale that contains these proofs. \ locale mut_m = fixes m :: "'mut" locale mut_m' = mut_m + fixes m' :: "'mut" assumes mm'[iff]: "m \ m'" locale gc locale sys subsection\Object marking \label{sec:gc-marking}\ text\ Both the mutators and the garbage collector mark references, which indicates that a reference is live in the current round of collection. This operation is defined in @{cite [cite_macro=citet] \Figure~2.8\ "Pizlo201xPhd"}. These definitions are parameterised by the name of the process. \ context fixes p :: "'mut process_name" begin abbreviation lock :: "location \ ('field, 'mut, 'ref) gc_com" where "lock l \ \l\ Request (\s. (p, ro_Lock)) (\_ s. {s})" notation lock ("\_\ lock") abbreviation unlock :: "location \ ('field, 'mut, 'ref) gc_com" where "unlock l \ \l\ Request (\s. (p, ro_Unlock)) (\_ s. {s})" notation unlock ("\_\ unlock") abbreviation read_mark :: "location \ (('field, 'mut, 'ref) local_state \ 'ref) \ ((gc_mark option \ gc_mark option) \ ('field, 'mut, 'ref) local_state \ ('field, 'mut, 'ref) local_state) \ ('field, 'mut, 'ref) gc_com" where "read_mark l r upd \ \l\ Request (\s. (p, ReadMark (r s))) (\mv s. { upd \m\ s |m. mv = mv_Mark m })" notation read_mark ("\_\ read'_mark") abbreviation read_fM :: "location \ ('field, 'mut, 'ref) gc_com" where "read_fM l \ \l\ Request (\s. (p, ro_Read mr_fM)) (\mv s. { s\fM := m\ |m. mv = mv_Mark (Some m) })" notation read_fM ("\_\ read'_fM") abbreviation read_phase :: "location \ ('field, 'mut, 'ref) gc_com" where "read_phase l \ \l\ Request (\s. (p, ReadPhase)) (\mv s. { s\phase := ph\ |ph. mv = mv_Phase ph })" notation read_phase ("\_\ read'_phase") abbreviation write_mark :: "location \ (('field, 'mut, 'ref) local_state \ 'ref) \ (('field, 'mut, 'ref) local_state \ bool) \ ('field, 'mut, 'ref) gc_com" where "write_mark l r fl \ \l\ Request (\s. (p, WriteMark (r s) (fl s))) (\_ s. { s\ ghost_honorary_grey := {r s} \ })" notation write_mark ("\_\ write'_mark") abbreviation add_to_W :: "location \ (('field, 'mut, 'ref) local_state \ 'ref) \ ('field, 'mut, 'ref) gc_com" where "add_to_W l r \ \l\ \\s. s\ W := W s \ {r s}, ghost_honorary_grey := {} \\" notation add_to_W ("\_\ add'_to'_W") text\ The reference we're marking is given in @{const "ref"}. If the current process wins the \texttt{CAS} race then the reference is marked and added to the local work list @{const "W"}. TSO means we cannot avoid having the mark write pending in a store buffer; in other words, we cannot have objects atomically transition from white to grey. The following scheme blackens a white object, and then reverts it to grey. The @{const "ghost_honorary_grey"} variable is used to track objects undergoing this transition. As CIMP provides no support for function calls, we prefix each statement's label with a string from its callsite. \ definition mark_object_fn :: "location \ ('field, 'mut, 'ref) gc_com" where "mark_object_fn l \ \l @ ''_mo_null''\ IF \<^bold>\ (NULL ref) THEN \l @ ''_mo_mark''\ read_mark (the \ ref) mark_update ;; \l @ ''_mo_fM''\ read_fM ;; \l @ ''_mo_mtest''\ IF mark \<^bold>\ Some \ fM THEN \l @ ''_mo_phase''\ read_phase ;; \l @ ''_mo_ptest''\ IF phase \<^bold>\ \ph_Idle\ THEN \ \CAS: claim object\ \l @ ''_mo_co_lock''\ lock ;; \l @ ''_mo_co_cmark''\ read_mark (the \ ref) cas_mark_update ;; \l @ ''_mo_co_ctest''\ IF cas_mark \<^bold>= mark THEN \l @ ''_mo_co_mark''\ write_mark (the \ ref) fM FI ;; \l @ ''_mo_co_unlock''\ unlock ;; \l @ ''_mo_co_won''\ IF cas_mark \<^bold>= mark THEN \l @ ''_mo_co_W''\ add_to_W (the \ ref) FI FI FI FI" end text\ The worklists (field @{term "W"}) are not subject to TSO. As we later show (\S\ref{def:valid_W_inv}), these are disjoint and hence operations on these are private to each process, with the sole exception of when the GC requests them from the mutators. We describe that mechanism next. \ subsection\Handshakes \label{sec:gc_ragged_safepoints}\ text\ The garbage collector needs to synchronise with the mutators. In practice this is implemented with some thread synchronisation primitives that include memory fences. The scheme we adopt here has the GC busy waiting. It sets a \pending\ flag for each mutator and then waits for each to respond. The system side of the interface collects the responses from the mutators into a single worklist, which acts as a proxy for the garbage collector's local worklist during \get_roots\ and \get_work\ handshakes. In practise this involves a \texttt{CAS} operation. We carefully model the effect these handshakes have on the process's TSO buffers. The system and mutators track handshake phases using ghost state. \ abbreviation hp_step :: "handshake_type \ handshake_phase \ handshake_phase" where "hp_step ht \ case ht of ht_NOOP \ handshake_step | ht_GetRoots \ handshake_step | ht_GetWork \ id" context sys begin definition handshake :: "('field, 'mut, 'ref) gc_com" where "handshake \ \''sys_hs_gc_set_type''\ Response (\req s. { (s\ handshake_type := ht, ghost_handshake_in_sync := \False\, ghost_handshake_phase := hp_step ht (ghost_handshake_phase s) \, mv_Void) |ht. req = (gc, ro_hs_gc_set_type ht) }) \ \''sys_hs_gc_mut_reqs''\ Response (\req s. { (s\ handshake_pending := (handshake_pending s)(m := True) \, mv_Void) |m. req = (gc, ro_hs_gc_set_pending m) }) \ \''sys_hs_gc_done''\ Response (\req s. { (s, mv_Bool (\handshake_pending s m)) |m. req = (gc, ro_hs_gc_read_pending m) }) \ \''sys_hs_gc_load_W''\ Response (\req s. { (s\ W := {} \, mv_Refs (W s)) |_::unit. req = (gc, ro_hs_gc_load_W) }) \ \''sys_hs_mut''\ Response (\req s. { (s, mv_Void) |m. req = (mutator m, ro_hs_mut_read_type (handshake_type s)) \ handshake_pending s m }) \ \''sys_hs_mut_done''\ Response (\req s. { (s\ handshake_pending := (handshake_pending s)(m := False), W := W s \ W', ghost_handshake_in_sync := (ghost_handshake_in_sync s)(m := True) \, mv_Void) |m W'. req = (mutator m, ro_hs_mut_done W') })" end text\ The mutator's side of the interface. Also updates the ghost state tracking the handshake state for @{const "ht_NOOP"} and @{const "ht_GetRoots"} but not @{const "ht_GetWork"}. \ context mut_m begin abbreviation mark_object :: "location \ ('field, 'mut, 'ref) gc_com" ("\_\ mark'_object") where "\l\ mark_object \ mark_object_fn (mutator m) l" abbreviation mfence :: "location \ ('field, 'mut, 'ref) gc_com" ("\_\ MFENCE") where "\l\ MFENCE \ \l\ Request (\s. (mutator m, ro_MFENCE)) (\_ s. {s})" abbreviation hs_read_type :: "location \ handshake_type \ ('field, 'mut, 'ref) gc_com" ("\_\ hs'_read'_type") where "\l\ hs_read_type ht \ \l\ Request (\s. (mutator m, ro_hs_mut_read_type ht)) (\_ s. {s})" abbreviation hs_noop_done :: "location \ ('field, 'mut, 'ref) gc_com" ("\_\ hs'_noop'_done") where "\l\ hs_noop_done \ \l\ Request (\s. (mutator m, ro_hs_mut_done {})) (\_ s. {s\ ghost_handshake_phase := handshake_step (ghost_handshake_phase s) \})" abbreviation hs_get_roots_done :: "location \ (('field, 'mut, 'ref) local_state \ 'ref set) \ ('field, 'mut, 'ref) gc_com" ("\_\ hs'_get'_roots'_done") where "\l\ hs_get_roots_done wl \ \l\ Request (\s. (mutator m, ro_hs_mut_done (wl s))) (\_ s. {s\ W := {}, ghost_handshake_phase := handshake_step (ghost_handshake_phase s) \})" abbreviation hs_get_work_done :: "location \ (('field, 'mut, 'ref) local_state \ 'ref set) \ ('field, 'mut, 'ref) gc_com" ("\_\ hs'_get'_work'_done") where "\l\ hs_get_work_done wl \ \l\ Request (\s. (mutator m, ro_hs_mut_done (wl s))) (\_ s. {s\ W := {} \})" definition handshake :: "('field, 'mut, 'ref) gc_com" where "handshake \ \''hs_noop''\ hs_read_type ht_NOOP ;; \''hs_noop_mfence''\ MFENCE ;; \''hs_noop_done''\ hs_noop_done \ \''hs_get_roots''\ hs_read_type ht_GetRoots ;; \''hs_get_roots_mfence''\ MFENCE ;; \''hs_get_roots_refs''\ \refs := \roots ;; \''hs_get_roots_loop''\ WHILE \<^bold>\(EMPTY refs) DO \''hs_get_roots_loop_choose_ref''\ \ref :\ Some ` \refs ;; \''hs_get_roots_loop''\ mark_object ;; \''hs_get_roots_loop_done''\ \refs := (\refs - {the \ref}) OD ;; \''hs_get_roots_done''\ hs_get_roots_done W \ \''hs_get_work''\ hs_read_type ht_GetWork ;; \''hs_get_work_mfence''\ MFENCE ;; \''hs_get_work_done''\ hs_get_work_done W" end text\ The garbage collector's side of the interface. \ context gc begin abbreviation set_hs_type :: "location \ handshake_type \ ('field, 'mut, 'ref) gc_com" ("\_\ set'_hs'_type") where "\l\ set_hs_type ht \ \l\ Request (\s. (gc, ro_hs_gc_set_type ht)) (\_ s. {s})" abbreviation set_hs_pending :: "location \ (('field, 'mut, 'ref) local_state \ 'mut) \ ('field, 'mut, 'ref) gc_com" ("\_\ set'_hs'_pending") where "\l\ set_hs_pending m \ \l\ Request (\s. (gc, ro_hs_gc_set_pending (m s))) (\_ s. {s})" definition handshake_init :: "location \ handshake_type \ ('field, 'mut, 'ref) gc_com" ("\_\ handshake'_init") where "\l\ handshake_init req \ \l @ ''_init_type''\ set_hs_type req ;; \l @ ''_init_muts''\ \muts := UNIV ;; \l @ ''_init_loop''\ WHILE \<^bold>\ (EMPTY muts) DO \l @ ''_init_loop_choose_mut''\ \mut :\ \muts ;; \l @ ''_init_loop_set_pending''\ set_hs_pending mut ;; \l @ ''_init_loop_done''\ \muts := (\muts - {\mut}) OD" definition handshake_done :: "location \ ('field, 'mut, 'ref) gc_com" ("\_\ handshake'_done") where "\l\ handshake_done \ \l @ ''_done_muts''\ \muts := UNIV ;; \l @ ''_done_loop''\ WHILE \<^bold>\(EMPTY muts) DO \l @ ''_done_loop_choose_mut''\ \mut :\ \muts ;; \l @ ''_done_loop_rendezvous''\ Request (\s. (gc, ro_hs_gc_read_pending (mut s))) (\mv s. { s\ muts := muts s - { mut s |done. mv = mv_Bool done \ done } \}) OD" abbreviation load_W :: "location \ ('field, 'mut, 'ref) gc_com" ("\_\ load'_W") where "\l\ load_W \ \l @ ''_load_W''\ Request (\s. (gc, ro_hs_gc_load_W)) (\resp s. {s\W := W'\ |W'. resp = mv_Refs W'})" abbreviation mfence :: "location \ ('field, 'mut, 'ref) gc_com" ("\_\ MFENCE") where "\l\ MFENCE \ \l\ Request (\s. (gc, ro_MFENCE)) (\_ s. {s})" definition handshake_noop :: "location \ ('field, 'mut, 'ref) gc_com" ("\_\ handshake'_noop") where "\l\ handshake_noop \ \l @ ''_mfence''\ MFENCE ;; \l\ handshake_init ht_NOOP ;; \l\ handshake_done" definition handshake_get_roots :: "location \ ('field, 'mut, 'ref) gc_com" ("\_\ handshake'_get'_roots") where "\l\ handshake_get_roots \ \l\ handshake_init ht_GetRoots ;; \l\ handshake_done ;; \l\ load_W" definition handshake_get_work :: "location \ ('field, 'mut, 'ref) gc_com" ("\_\ handshake'_get'_work") where "\l\ handshake_get_work \ \l\ handshake_init ht_GetWork ;; \l\ handshake_done ;; \l\ load_W" end subsection\The system process\ text \ The system process models the environment in which the garbage collector and mutators execute. We translate the x86-TSO memory model due to @{cite [cite_macro=citet] "DBLP:journals/cacm/SewellSONM10"} into a CIMP process. It is a reactive system: it receives requests and returns values, but initiates no communication itself. It can, however, autonomously commit a write pending in a TSO store buffer. The memory bus can be locked by atomic compare-and-swap (\texttt{CAS}) instructions (and others in general). A processor is not blocked (i.e., it can read from memory) when it holds the lock, or no-one does. \ definition not_blocked :: "('field, 'mut, 'ref) local_state \ 'mut process_name \ bool" where "not_blocked s p \ case mem_lock s of None \ True | Some p' \ p = p'" text\ We compute the view a processor has of memory by applying all its pending writes. \ definition do_write_action :: "('field, 'ref) mem_write_action \ ('field, 'mut, 'ref) local_state \ ('field, 'mut, 'ref) local_state" where "do_write_action wact \ \s. case wact of mw_Mark r gc_mark \ s\heap := (heap s)(r := map_option (\obj. obj\obj_mark := gc_mark\) (heap s r))\ | mw_Mutate r f new_r \ s\heap := (heap s)(r := map_option (\obj. obj\obj_fields := (obj_fields obj)(f := new_r) \) (heap s r))\ | mw_fM gc_mark \ s\fM := gc_mark\ | mw_fA gc_mark \ s\fA := gc_mark\ | mw_Phase gc_phase \ s\phase := gc_phase\" definition fold_writes :: "('field, 'ref) mem_write_action list \ ('field, 'mut, 'ref) local_state \ ('field, 'mut, 'ref) local_state" where "fold_writes ws \ fold (\w. (\) (do_write_action w)) ws id" abbreviation processors_view_of_memory :: "'mut process_name \ ('field, 'mut, 'ref) local_state \ ('field, 'mut, 'ref) local_state" where "processors_view_of_memory p s \ fold_writes (mem_write_buffers s p) s" definition do_read_action :: "('field, 'ref) mem_read_action \ ('field, 'mut, 'ref) local_state \ ('field, 'ref) response" where "do_read_action ract \ \s. case ract of mr_Ref r f \ mv_Ref (heap s r \ (\obj. obj_fields obj f)) | mr_Mark r \ mv_Mark (map_option obj_mark (heap s r)) | mr_Phase \ mv_Phase (phase s) | mr_fM \ mv_Mark (Some (fM s)) | mr_fA \ mv_Mark (Some (fA s))" definition sys_read :: "'mut process_name \ ('field, 'ref) mem_read_action \ ('field, 'mut, 'ref) local_state \ ('field, 'ref) response" where "sys_read p ract \ do_read_action ract \ processors_view_of_memory p" context sys begin text\ The semantics of TSO memory following @{cite [cite_macro=citet] \\S3\ "DBLP:journals/cacm/SewellSONM10"}. This differs from the earlier @{cite [cite_macro=citet] "DBLP:conf/tphol/OwensSS09"} by allowing the TSO lock to be taken by a process with a non-empty write buffer. We omit their treatment of registers; these are handled by the local states of the other processes. The system can autonomously take the oldest write in the write buffer for processor \p\ and commit it to memory, provided \p\ either holds the lock or no processor does. \ definition mem_TSO :: "('field, 'mut, 'ref) gc_com" where "mem_TSO \ \''sys_read''\ Response (\req s. { (s, sys_read p mr s) |p mr. req = (p, ro_Read mr) \ not_blocked s p }) \ \''sys_write''\ Response (\req s. { (s\ mem_write_buffers := (mem_write_buffers s)(p := mem_write_buffers s p @ [w]) \, mv_Void) |p w. req = (p, ro_Write w) }) \ \''sys_mfence''\ Response (\req s. { (s, mv_Void) |p. req = (p, ro_MFENCE) \ mem_write_buffers s p = [] }) \ \''sys_lock''\ Response (\req s. { (s\ mem_lock := Some p \, mv_Void) |p. req = (p, ro_Lock) \ mem_lock s = None }) \ \''sys_unlock''\ Response (\req s. { (s\ mem_lock := None \, mv_Void) |p. req = (p, ro_Unlock) \ mem_lock s = Some p \ mem_write_buffers s p = [] }) \ \''sys_dequeue_write_buffer''\ LocalOp (\s. { (do_write_action w s)\ mem_write_buffers := (mem_write_buffers s)(p := ws) \ | p w ws. mem_write_buffers s p = w # ws \ not_blocked s p \ p \ sys })" text\ We track which references are allocated using the domain of @{const "heap"}. \label{sec:sys_alloc} For now we assume that the system process magically allocates and deallocates references. To model this more closely we would need to take care of the underlying machine addresses. We should be able to separate out those issues from GC correctness: the latter should imply that only alloc and free can interfere with each other. We also arrange for the object to be marked atomically (see \S\ref{sec:mut_alloc}) which morally should be done by the mutator. In practice allocation pools enable this kind of atomicity (wrt the sweep loop in the GC described in \S\ref{sec:gc-model-gc}). Note that the \texttt{abort} in @{cite [cite_macro=citet] \Figure~2.9: Alloc\ "Pizlo201xPhd"} means the atomic fails and the mutator can revert to activity outside of \texttt{Alloc}, avoiding deadlock. \ definition alloc :: "('field, 'mut, 'ref) gc_com" where "alloc \ \''sys_alloc''\ Response (\req s. { ( s\ heap := (heap s)(r := Some \ obj_mark = fA s, obj_fields = \None\ \) \, mv_Ref (Some r) ) |r. r \ dom (heap s) \ snd req = ro_Alloc })" text\ References are freed by removing them from @{const "heap"}. \ definition free :: "('field, 'mut, 'ref) gc_com" where "free \ \''sys_free''\ Response (\req s. { (s\heap := (heap s)(r := None)\, mv_Void) |r. snd req = ro_Free r })" text\ The top-level system process. \ definition com :: "('field, 'mut, 'ref) gc_com" where "com \ LOOP DO mem_TSO \ alloc \ free \ handshake OD" end subsection\Mutators\ text\ The mutators need to cooperate with the garbage collector. In particular, when the garbage collector is not idle the mutators use a \emph{write barrier} (see \S\ref{sec:gc-marking}). The local state for each mutator tracks a working set of references, which abstracts from how the process's registers and stack are traversed to discover roots. \ context mut_m begin text\ \label{sec:mut_alloc} Allocation is defined in @{cite [cite_macro=citet] \Figure~2.9\ "Pizlo201xPhd"}. See \S\ref{sec:sys_alloc} for how we abstract it. \ abbreviation (in -) mut_alloc :: "'mut \ ('field, 'mut, 'ref) gc_com" where "mut_alloc m \ \''alloc''\ Request (\s. (mutator m, ro_Alloc)) (\mv s. { s\ roots := roots s \ {r} \ |r. mv = mv_Ref (Some r) })" abbreviation alloc :: "('field, 'mut, 'ref) gc_com" where "alloc \ mut_alloc m" text\ The mutator can always discard any references it holds. \ abbreviation discard :: "('field, 'mut, 'ref) gc_com" where "discard \ \''discard_refs''\ LocalOp (\s. { s\ roots := roots' \ |roots'. roots' \ roots s })" text\ Load and store are defined in @{cite [cite_macro=citet] \Figure~2.9\ "Pizlo201xPhd"}. Dereferencing a reference can increase the set of mutator roots. \ abbreviation load :: "('field, 'mut, 'ref) gc_com" where "load \ \''load_choose''\ LocalOp (\s. { s\ tmp_ref := r, field := f \ |r f. r \ roots s }) ;; \''load''\ Request (\s. (mutator m, ReadRef (tmp_ref s) (field s))) (\mv s. { s\ roots := roots s \ set_option r \ |r. mv = mv_Ref r })" text\ \label{sec:write-barriers} Storing a reference involves marking both the old and new references, i.e., both \emph{insertion} and \emph{deletion} barriers are installed. The deletion barrier preserves the \emph{weak tricolour invariant}, and the insertion barrier preserves the \emph{strong tricolour invariant}; see \S\ref{sec:strong-tricolour-invariant} for further discussion. Note that the the mutator reads the overwritten reference but does not store it in its roots. \ abbreviation mut_deref :: "location \ (('field, 'mut, 'ref) local_state \ 'ref) \ (('field, 'mut, 'ref) local_state \ 'field) \ (('ref option \ 'ref option) \ ('field, 'mut, 'ref) local_state \ ('field, 'mut, 'ref) local_state) \ ('field, 'mut, 'ref) gc_com" ("\_\ deref") where "\l\ deref r f upd \ \l\ Request (\s. (mutator m, ReadRef (r s) (f s))) (\mv s. { upd \opt_r'\ (s\ghost_honorary_root := set_option opt_r'\) |opt_r'. mv = mv_Ref opt_r' })" (* Does not work in local theory mode: syntax "_mut_fassign" :: "location \ idt \ 'ref \ 'field \ ('field, 'mut, 'ref) gc_com" ("\_\ \_ := \_ \ _" [0, 0, 70] 71) translations "\l\ \q := \r\f" => "CONST mut_deref l r \f\ (_update_name q)" \ref := \tmp_ref\\field ;; *) abbreviation write_ref :: "location \ (('field, 'mut, 'ref) local_state \ 'ref) \ (('field, 'mut, 'ref) local_state \ 'field) \ (('field, 'mut, 'ref) local_state \ 'ref option) \ ('field, 'mut, 'ref) gc_com" ("\_\ write'_ref") where "\l\ write_ref r f r' \ \l\ Request (\s. (mutator m, WriteRef (r s) (f s) (r' s))) (\_ s. {s\ghost_honorary_root := {}\})" definition store :: "('field, 'mut, 'ref) gc_com" where "store \ \ \Choose vars for: \ref\field := new_ref\\ \''store_choose''\ LocalOp (\s. { s\ tmp_ref := r, field := f, new_ref := r' \ |r f r'. r \ roots s \ r' \ Some ` roots s \ {None} }) ;; \ \Mark the reference we're about to overwrite. Does not update roots.\ \''deref_del''\ deref tmp_ref field ref_update ;; \''store_del''\ mark_object ;; \ \Mark the reference we're about to insert.\ \''lop_store_ins''\ \ref := \new_ref ;; \''store_ins''\ mark_object ;; \''store_ins''\ write_ref tmp_ref field new_ref" text\ A mutator makes a non-deterministic choice amongst its possible actions. For completeness we allow mutators to issue \texttt{MFENCE} instructions. We leave \texttt{CAS} (etc) to future work. Neither has a significant impact on the rest of the development. \ definition com :: "('field, 'mut, 'ref) gc_com" where "com \ LOOP DO \''mut local computation''\ SKIP \ alloc \ discard \ load \ store \ \''mut mfence''\ MFENCE \ handshake OD" end subsection \Garbage collector \label{sec:gc-model-gc}\ text\ We abstract the primitive actions of the garbage collector thread. \ abbreviation gc_deref :: "location \ (('field, 'mut, 'ref) local_state \ 'ref) \ (('field, 'mut, 'ref) local_state \ 'field) \ (('ref option \ 'ref option) \ ('field, 'mut, 'ref) local_state \ ('field, 'mut, 'ref) local_state) \ ('field, 'mut, 'ref) gc_com" where "gc_deref l r f upd \ \l\ Request (\s. (gc, ReadRef (r s) (f s))) (\mv s. { upd \r'\ s |r'. mv = mv_Ref r' })" abbreviation gc_read_mark :: "location \ (('field, 'mut, 'ref) local_state \ 'ref) \ ((gc_mark option \ gc_mark option) \ ('field, 'mut, 'ref) local_state \ ('field, 'mut, 'ref) local_state) \ ('field, 'mut, 'ref) gc_com" where "gc_read_mark l r upd \ \l\ Request (\s. (gc, ReadMark (r s))) (\mv s. { upd \m\ s |m. mv = mv_Mark m })" syntax "_gc_fassign" :: "location \ idt \ 'ref \ 'field \ ('field, 'mut, 'ref) gc_com" ("\_\ \_ := \_ \ _" [0, 0, 70] 71) "_gc_massign" :: "location \ idt \ 'ref \ ('field, 'mut, 'ref) gc_com" ("\_\ \_ := \_ \ flag" [0, 0] 71) translations "\l\ \q := \r\f" => "CONST gc_deref l r \f\ (_update_name q)" "\l\ \m := \r\flag" => "CONST gc_read_mark l r (_update_name m)" context gc begin abbreviation write_fA :: "location \ (('field, 'mut, 'ref) local_state \ gc_mark) \ ('field, 'mut, 'ref) gc_com" ("\_\ write'_fA") where "\l\ write_fA f \ \l\ Request (\s. (gc, WritefA (f s))) (\_ s. {s})" abbreviation read_fM :: "location \ ('field, 'mut, 'ref) gc_com" ("\_\ read'_fM") where "\l\ read_fM \ \l\ Request (\s. (gc, ReadfM)) (\mv s. { s\fM := m\ |m. mv = mv_Mark (Some m) })" abbreviation write_fM :: "location \ ('field, 'mut, 'ref) gc_com" ("\_\ write'_fM") where "\l\ write_fM \ \l\ Request (\s. (gc, WritefM (fM s))) (\_ s. {s})" abbreviation write_phase :: "location \ gc_phase \ ('field, 'mut, 'ref) gc_com" ("\_\ write'_phase") where "\l\ write_phase ph \ \l\ Request (\s. (gc, WritePhase ph)) (\_ s. {s\ phase := ph \})" abbreviation mark_object :: "location \ ('field, 'mut, 'ref) gc_com" ("\_\ mark'_object") where "\l\ mark_object \ mark_object_fn gc l" abbreviation free :: "location \ (('field, 'mut, 'ref) local_state \ 'ref) \ ('field, 'mut, 'ref) gc_com" ("\_\ free") where "\l\ free r \ \l\ Request (\s. (gc, ro_Free (r s))) (\_ s. {s})" text\ The following CIMP program encodes the garbage collector algorithm proposed in Figure~2.15 of @{cite [cite_macro=citet] "Pizlo201xPhd"}. \ definition (in gc) com :: "('field, 'mut, 'ref) gc_com" where "com \ LOOP DO \''idle_noop''\ handshake_noop ;; \ \\hp_Idle\\ \''idle_read_fM''\ read_fM ;; \''idle_invert_fM''\ \fM := (\ \fM) ;; \''idle_write_fM''\ write_fM ;; \''idle_flip_noop''\ handshake_noop ;; \ \\hp_IdleInit\\ \''idle_phase_init''\ write_phase ph_Init ;; \''init_noop''\ handshake_noop ;; \ \\hp_InitMark\\ \''init_phase_mark''\ write_phase ph_Mark ;; \''mark_read_fM''\ read_fM ;; \''mark_write_fA''\ write_fA fM ;; \''mark_noop''\ handshake_noop ;; \ \\hp_Mark\\ \''mark_loop_get_roots''\ handshake_get_roots ;; \ \\hp_IdleMarkSweep\\ \''mark_loop''\ WHILE \<^bold>\(EMPTY W) DO \''mark_loop_inner''\ WHILE \<^bold>\(EMPTY W) DO \''mark_loop_choose_ref''\ \tmp_ref :\ \W ;; \''mark_loop_fields''\ \field_set := UNIV ;; \''mark_loop_mark_object_loop''\ WHILE \<^bold>\(EMPTY field_set) DO \''mark_loop_mark_choose_field''\ \field :\ \field_set ;; \''mark_loop_mark_deref''\ \ref := \tmp_ref\\field ;; \''mark_loop''\ mark_object ;; \''mark_loop_mark_field_done''\ \field_set := (\field_set - {\field}) OD ;; \''mark_loop_blacken''\ \W := (\W - {\tmp_ref}) OD ;; \''mark_loop_get_work''\ handshake_get_work OD ;; \ \sweep\ \''mark_end''\ write_phase ph_Sweep ;; \''sweep_read_fM''\ read_fM ;; \''sweep_refs''\ \refs := UNIV ;; \''sweep_loop''\ WHILE \<^bold>\(EMPTY refs) DO \''sweep_loop_choose_ref''\ \tmp_ref :\ \refs ;; \''sweep_loop_read_mark''\ \mark := \tmp_ref\flag ;; \''sweep_loop_check''\ IF \<^bold>\(NULL mark) \<^bold>\ the \ mark \<^bold>\ fM THEN \''sweep_loop_free''\ free tmp_ref FI ;; \''sweep_loop_ref_done''\ \refs := (\refs - {\tmp_ref}) OD ;; \''sweep_idle''\ write_phase ph_Idle OD" end primrec gc_pgms :: "'mut process_name \ ('field, 'mut, 'ref) gc_com" where "gc_pgms (mutator m) = mut_m.com m" | "gc_pgms gc = gc.com" | "gc_pgms sys = sys.com" (*<*) end (*>*) diff --git a/thys/ConcurrentGC/Proofs.thy b/thys/ConcurrentGC/Proofs.thy --- a/thys/ConcurrentGC/Proofs.thy +++ b/thys/ConcurrentGC/Proofs.thy @@ -1,331 +1,333 @@ (*<*) (* * Copyright 2015, NICTA * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * @TAG(NICTA_BSD) *) theory Proofs imports StrongTricolour begin +declare subst_all [simp del] [[simproc del: defined_all]] + (*>*) section\Top-level safety \label{sec:top-level-correctness}\ subsection\Invariants\ definition (in gc) invsL :: "('field, 'mut, 'ref) gc_pred" where "invsL \ fM_fA_invL \<^bold>\ gc_mark.mark_object_invL \<^bold>\ gc_W_empty_invL \<^bold>\ handshake_invL \<^bold>\ obj_fields_marked_invL \<^bold>\ phase_invL \<^bold>\ sweep_loop_invL \<^bold>\ tso_lock_invL \<^bold>\ LSTP (fA_rel_inv \<^bold>\ fM_rel_inv)" definition (in mut_m) invsL :: "('field, 'mut, 'ref) gc_pred" where "invsL \ mark_object_invL \<^bold>\ mut_get_roots.mark_object_invL m \<^bold>\ mut_store_ins.mark_object_invL m \<^bold>\ mut_store_del.mark_object_invL m \<^bold>\ handshake_invL \<^bold>\ tso_lock_invL \<^bold>\ LSTP mutator_phase_inv" definition invs :: "('field, 'mut, 'ref) lsts_pred" where "invs \ handshake_phase_inv \<^bold>\ phase_rel_inv \<^bold>\ strong_tricolour_inv \<^bold>\ sys_phase_inv \<^bold>\ tso_writes_inv \<^bold>\ valid_refs_inv \<^bold>\ valid_W_inv" definition I :: "('field, 'mut, 'ref) gc_pred" where "I \ gc.invsL \<^bold>\ (\<^bold>\m. mut_m.invsL m) \<^bold>\ LSTP invs" (*<*) lemmas I_defs = gc.invsL_def mut_m.invsL_def invs_def I_def lemma (in gc) I: "\ I \ gc" apply (simp add: I_defs) apply (rule valid_pre) apply ( rule valid_conj_lift valid_all_lift | fastforce )+ done lemma (in sys) I: "\ I \ sys" apply (simp add: I_defs) apply (rule valid_pre) apply ( rule valid_conj_lift valid_all_lift | fastforce )+ done text\ We need to separately treat the two cases of a single mutator and multiple mutators. In the latter case we have the additional obligation of showing mutual non-interference amongst mutators. \ lemma mut_invsL[intro]: "\I\ mutator m \mut_m.invsL m'\" proof(cases "m = m'") case True interpret mut_m m' by unfold_locales from True show ?thesis apply (simp add: I_defs) apply (rule valid_pre) apply ( rule valid_conj_lift valid_all_lift | fastforce )+ done next case False then interpret mut_m' m' m by unfold_locales blast from False show ?thesis apply (simp add: I_defs) apply (rule valid_pre) apply ( rule valid_conj_lift valid_all_lift | fastforce )+ done qed lemma (in mut_m) I: "\ I \ mutator m" apply (simp add: I_def gc.invsL_def invs_def) apply (rule valid_pre) apply ( rule valid_conj_lift valid_all_lift | fastforce )+ apply (simp add: I_defs) done (*>*) subsection\Initial conditions\ text\ We ask that the GC and system initially agree on some things: \begin{itemize} \item All objects on the heap are marked (have their flags equal to @{const "sys_fM"}, and there are no grey references, i.e. the heap is uniformly black. \item The GC and system have the same values for @{term "fA"}, @{term "fM"}, etc. and the phase is @{term "Idle"}. \item No process holds the TSO lock and all write buffers are empty. \item All root-reachable references are backed by objects. \end{itemize} Note that these are merely sufficient initial conditions and can be weakened. \ locale gc_system = fixes initial_mark :: gc_mark begin definition gc_initial_state :: "('field, 'mut, 'ref) lst_pred" where "gc_initial_state s \ fM s = initial_mark \ phase s = ph_Idle \ ghost_honorary_grey s = {} \ W s = {}" definition mut_initial_state :: "('field, 'mut, 'ref) lst_pred" where "mut_initial_state s \ ghost_handshake_phase s = hp_IdleMarkSweep \ ghost_honorary_grey s = {} \ ghost_honorary_root s = {} \ W s = {}" definition sys_initial_state :: "('field, 'mut, 'ref) lst_pred" where "sys_initial_state s \ (\m. \handshake_pending s m \ ghost_handshake_in_sync s m) \ ghost_handshake_phase s = hp_IdleMarkSweep \ handshake_type s = ht_GetRoots \ obj_mark ` ran (heap s) \ {initial_mark} \ fA s = initial_mark \ fM s = initial_mark \ phase s = ph_Idle \ ghost_honorary_grey s = {} \ W s = {} \ (\p. mem_write_buffers s p = []) \ mem_lock s = None" abbreviation "root_reachable y \ \<^bold>\m x. \x\ \<^bold>\ mut_m.mut_roots m \<^bold>\ x reaches y" definition valid_refs :: "('field, 'mut, 'ref) lsts_pred" where "valid_refs \ \<^bold>\y. root_reachable y \<^bold>\ valid_ref y" definition gc_system_init :: "('field, 'mut, 'ref) lsts_pred" where "gc_system_init \ (\s. gc_initial_state (s gc)) \<^bold>\ (\s. \m. mut_initial_state (s (mutator m))) \<^bold>\ (\s. sys_initial_state (s sys)) \<^bold>\ valid_refs" text\ The system consists of the programs and these constraints on the initial state. \ abbreviation gc_system :: "('field, 'mut, 'ref) gc_system" where "gc_system \ (gc_pgms, gc_system_init)" (*<*) lemma init_strong_tricolour_inv: "\ obj_mark ` ran (sys_heap (mkP (s, []))\) \ {gc_fM (mkP (s, []))\}; sys_fM (mkP (s, []))\ = gc_fM (mkP (s, []))\ \ \ strong_tricolour_inv (mkP (s, []))\" by (auto simp: strong_tricolour_inv_def ran_def split: obj_at_splits) lemma init_no_grey_refs: "\ gc_W (mkP (s, []))\ = {}; \m. W ((mkP (s, []))\ (mutator m)) = {}; sys_W (mkP (s, []))\ = {}; gc_ghost_honorary_grey (mkP (s, []))\ = {}; \m. ghost_honorary_grey ((mkP (s, []))\ (mutator m)) = {}; sys_ghost_honorary_grey (mkP (s, []))\ = {} \ \ no_grey_refs (mkP (s, []))\" apply (clarsimp simp: no_grey_refs_def grey_def) apply (rename_tac x xa) apply (case_tac xa) apply (auto simp: WL_def) done lemma valid_refs_imp_valid_refs_inv: "\ valid_refs s; no_grey_refs s; \p. sys_mem_write_buffers p s = []; \m. ghost_honorary_root (s (mutator m)) = {} \ \ valid_refs_inv s" by (auto simp: valid_refs_def valid_refs_inv_def mut_m.reachable_def mut_m.tso_write_refs_def dest: no_grey_refs_not_grey_reachableD) lemma no_grey_refs_imp_valid_W_inv: "\ no_grey_refs s; \p. sys_mem_write_buffers p s = [] \ \ valid_W_inv s" unfolding valid_W_inv_def by (auto simp: no_grey_refs_def grey_def WL_def) lemma init_inv_sys: "\s\initial_states gc_system. invs (mkP (s, []))\" apply (clarsimp dest!: initial_statesD simp: gc_system_init_def invs_def gc_initial_state_def mut_initial_state_def sys_initial_state_def inv handshake_phase_rel_def handshake_phase_inv_def hp_step_rel_def phase_rel_inv_def phase_rel_def tso_writes_inv_def init_no_grey_refs init_strong_tricolour_inv no_grey_refs_imp_valid_W_inv valid_refs_imp_valid_refs_inv all_conj_distrib) done lemma init_inv_gc: "\s\initial_states gc_system. gc.invsL (mkP (s, []))" apply (clarsimp dest!: initial_statesD) apply (drule fun_cong[where x=gc]) (* hacky *) apply (simp add: com) apply (clarsimp simp: gc_system_init_def gc_initial_state_def mut_initial_state_def sys_initial_state_def gc.invsL_def inv gc.fM_fA_invL_def fA_rel_inv_def fA_rel_def fM_rel_inv_def fM_rel_def gc.handshake_invL_def gc.obj_fields_marked_invL_def gc.phase_invL_def gc.sweep_loop_invL_def gc.tso_lock_invL_def gc.gc_W_empty_invL_def init_no_grey_refs loc atS_simps all_conj_distrib) apply (auto simp: ran_def image_subset_iff split: obj_at_splits) done lemma valid_refs_imp_reachable_snapshot_inv: "\ valid_refs s; obj_mark ` ran (sys_heap s) \ {sys_fM s}; \p. sys_mem_write_buffers p s = []; \m. ghost_honorary_root (s (mutator m)) = {} \ \ mut_m.reachable_snapshot_inv m s" apply (clarsimp simp: mut_m.reachable_snapshot_inv_def in_snapshot_def valid_refs_def) apply (auto simp: image_subset_iff ran_def black_def mut_m.reachable_def mut_m.tso_write_refs_def split: obj_at_splits) done lemma init_inv_mut: "\s\initial_states gc_system. mut_m.invsL m (mkP (s, []))" apply (clarsimp dest!: initial_statesD) apply (drule fun_cong[where x="mutator m"]) (* hacky *) apply (simp add: com) apply (clarsimp simp: gc_system_init_def mut_initial_state_def sys_initial_state_def valid_refs_imp_reachable_snapshot_inv mut_m.invsL_def inv mut_m.mark_object_invL_def mut_m.handshake_invL_def mut_m.tso_lock_invL_def mut_m.marked_deletions_def mut_m.marked_insertions_def loc atS_simps) done lemma init_inv: "\s\initial_states gc_system. I (mkP (s, []))" by (simp add: I_def init_inv_sys init_inv_gc init_inv_mut) (*>*) text\\ theorem inv: "s \ reachable_states gc_system \ I (mkP s)" (*<*) apply (erule VCG) apply (rule init_inv) apply (rename_tac p) apply (case_tac p, simp_all) apply (rule mut_m.I[unfolded valid_proc_def, simplified]) apply (rule gc.I[unfolded valid_proc_def, simplified]) apply (rule sys.I[unfolded valid_proc_def, simplified]) done (*>*) text\ Our headline safety result follows directly. \ corollary safety: "s \ reachable_states gc_system \ valid_refs (mkP s)\" (*<*) apply (drule inv) apply (clarsimp simp: I_def invs_def valid_refs_inv_def valid_refs_def) apply (rename_tac x xa xb) apply (drule_tac x=x in spec) apply clarsimp apply (drule_tac x=xa in spec, fastforce simp: mut_m.reachable_def) done (*>*) text\\ end text\ The GC is correct for the remaining fixed-but-arbitrary initial conditions. \ interpretation gc_system_interpretation: gc_system undefined . subsection\A concrete system state\ text\ We demonstrate that our definitions are not vacuous by exhibiting a concrete initial state that satisfies the initial conditions. We use Isabelle's notation for types of a given size. \ (*<*) end (*>*) diff --git a/thys/ConcurrentGC/Proofs_basis.thy b/thys/ConcurrentGC/Proofs_basis.thy --- a/thys/ConcurrentGC/Proofs_basis.thy +++ b/thys/ConcurrentGC/Proofs_basis.thy @@ -1,349 +1,351 @@ (*<*) (* * Copyright 2015, NICTA * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * @TAG(NICTA_BSD) *) theory Proofs_basis imports Tactics "HOL-Library.Simps_Case_Conv" begin +declare subst_all [simp del] [[simproc del: defined_all]] + (*>*) subsection\Functions and predicates\ text\ We define a pile of predicates and accessor functions for the process's local states. One might hope that a more sophisticated approach would automate all of this (cf @{cite [cite_macro=citet] "DBLP:journals/entcs/SchirmerW09"}). \ abbreviation "is_mw_Mark w \ \r fl. w = mw_Mark r fl" abbreviation "is_mw_Mutate w \ \r f r'. w = mw_Mutate r f r'" abbreviation "is_mw_fA w \ \fl. w = mw_fA fl" abbreviation "is_mw_fM w \ \fl. w = mw_fM fl" abbreviation "is_mw_Phase w \ \ph. w = mw_Phase ph" abbreviation (input) pred_in_W :: "'ref \ 'mut process_name \ ('field, 'mut, 'ref) lsts_pred" (infix "in'_W" 50) where "r in_W p \ \s. r \ W (s p)" abbreviation (input) pred_in_ghost_honorary_grey :: "'ref \ 'mut process_name \ ('field, 'mut, 'ref) lsts_pred" (infix "in'_ghost'_honorary'_grey" 50) where "r in_ghost_honorary_grey p \ \s. r \ ghost_honorary_grey (s p)" context gc begin abbreviation valid_gc_syn :: "('field, 'mut, 'ref) gc_loc_comp \ ('field, 'mut, 'ref) gc_pred \ ('field, 'mut, 'ref) gc_com \ ('field, 'mut, 'ref) gc_pred \ bool" ("_ \ \_\/ _/ \_\") where "afts \ \P\ c \Q\ \ gc_pgms, gc, afts \ \P\ c \Q\" abbreviation valid_gc_inv_syn :: "('field, 'mut, 'ref) gc_loc_comp \ ('field, 'mut, 'ref) gc_pred \ ('field, 'mut, 'ref) gc_com \ bool" ("_ \ \_\/ _") where "afts \ \P\ c \ afts \ \P\ c \P\" end abbreviation "gc_cas_mark s \ cas_mark (s gc)" abbreviation "gc_fM s \ fM (s gc)" abbreviation "gc_field s \ field (s gc)" abbreviation "gc_field_set s \ field_set (s gc)" abbreviation "gc_mark s \ mark (s gc)" abbreviation "gc_mut s \ mut (s gc)" abbreviation "gc_muts s \ muts (s gc)" abbreviation "gc_phase s \ phase (s gc)" abbreviation "gc_tmp_ref s \ tmp_ref (s gc)" abbreviation "gc_ghost_honorary_grey s \ ghost_honorary_grey (s gc)" abbreviation "gc_ref s \ ref (s gc)" abbreviation "gc_refs s \ refs (s gc)" abbreviation "gc_the_ref \ the \ gc_ref" abbreviation "gc_W s \ W (s gc)" abbreviation at_gc :: "location \ ('field, 'mut, 'ref) lsts_pred \ ('field, 'mut, 'ref) gc_pred" where "at_gc l P \ at gc l \<^bold>\ LSTP P" abbreviation atS_gc :: "location set \ ('field, 'mut, 'ref) lsts_pred \ ('field, 'mut, 'ref) gc_pred" where "atS_gc ls P \ atS gc ls \<^bold>\ LSTP P" context mut_m begin abbreviation valid_mut_syn :: "('field, 'mut, 'ref) gc_loc_comp \ ('field, 'mut, 'ref) gc_pred \ ('field, 'mut, 'ref) gc_com \ ('field, 'mut, 'ref) gc_pred \ bool" ("_ \ \_\/ _/ \_\") where "afts \ \P\ c \Q\ \ gc_pgms, mutator m, afts \ \P\ c \Q\" abbreviation valid_mut_inv_syn :: "('field, 'mut, 'ref) gc_loc_comp \ ('field, 'mut, 'ref) gc_pred \ ('field, 'mut, 'ref) gc_com \ bool" ("_ \ \_\/ _") where "afts \ \P\ c \ afts \ \P\ c \P\" abbreviation at_mut :: "location \ ('field, 'mut, 'ref) lsts_pred \ ('field, 'mut, 'ref) gc_pred" where "at_mut l P \ at (mutator m) l \<^bold>\ LSTP P" abbreviation atS_mut :: "location set \ ('field, 'mut, 'ref) lsts_pred \ ('field, 'mut, 'ref) gc_pred" where "atS_mut ls P \ atS (mutator m) ls \<^bold>\ LSTP P" abbreviation "mut_cas_mark s \ cas_mark (s (mutator m))" abbreviation "mut_field s \ field (s (mutator m))" abbreviation "mut_fM s \ fM (s (mutator m))" abbreviation "mut_ghost_honorary_grey s \ ghost_honorary_grey (s (mutator m))" abbreviation "mut_ghost_handshake_phase s \ ghost_handshake_phase (s (mutator m))" abbreviation "mut_ghost_honorary_root s \ ghost_honorary_root (s (mutator m))" abbreviation "mut_mark s \ mark (s (mutator m))" abbreviation "mut_new_ref s \ new_ref (s (mutator m))" abbreviation "mut_phase s \ phase (s (mutator m))" abbreviation "mut_ref s \ ref (s (mutator m))" abbreviation "mut_tmp_ref s \ tmp_ref (s (mutator m))" abbreviation "mut_the_new_ref \ the \ mut_new_ref" abbreviation "mut_the_ref \ the \ mut_ref" abbreviation "mut_refs s \ refs (s (mutator m))" abbreviation "mut_roots s \ roots (s (mutator m))" abbreviation "mut_W s \ W (s (mutator m))" end context sys begin abbreviation valid_sys_syn :: "('field, 'mut, 'ref) gc_loc_comp \ ('field, 'mut, 'ref) gc_pred \ ('field, 'mut, 'ref) gc_com \ ('field, 'mut, 'ref) gc_pred \ bool" ("_ \ \_\/ _/ \_\") where "afts \ \P\ c \Q\ \ gc_pgms, sys, afts \ \P\ c \Q\" abbreviation valid_sys_inv_syn :: "('field, 'mut, 'ref) gc_loc_comp \ ('field, 'mut, 'ref) gc_pred \ ('field, 'mut, 'ref) gc_com \ bool" ("_ \ \_\/ _") where "afts \ \P\ c \ afts \ \P\ c \P\" end abbreviation sys_heap :: "('field, 'mut, 'ref) lsts \ 'ref \ ('field, 'ref) object option" where "sys_heap s \ heap (s sys)" abbreviation "sys_fA s \ fA (s sys)" abbreviation "sys_fM s \ fM (s sys)" abbreviation "sys_ghost_honorary_grey s \ ghost_honorary_grey (s sys)" abbreviation "sys_ghost_handshake_in_sync m s \ ghost_handshake_in_sync (s sys) m" abbreviation "sys_ghost_handshake_phase s \ ghost_handshake_phase (s sys)" abbreviation "sys_handshake_pending m s \ handshake_pending (s sys) m" abbreviation "sys_handshake_type s \ handshake_type (s sys)" abbreviation "sys_mem_write_buffers p s \ mem_write_buffers (s sys) p" abbreviation "sys_mem_lock s \ mem_lock (s sys)" abbreviation "sys_phase s \ phase (s sys)" abbreviation "sys_W s \ W (s sys)" abbreviation atS_sys :: "location set \ ('field, 'mut, 'ref) lsts_pred \ ('field, 'mut, 'ref) gc_pred" where "atS_sys ls P \ atS sys ls \<^bold>\ LSTP P" text\Projections on TSO buffers.\ abbreviation (input) "tso_unlocked s \ mem_lock (s sys) = None" abbreviation (input) "tso_locked_by p s \ mem_lock (s sys) = Some p" abbreviation (input) "tso_pending p P s \ filter P (mem_write_buffers (s sys) p)" abbreviation (input) "tso_pending_write p w s \ w \ set (mem_write_buffers (s sys) p)" abbreviation (input) "tso_pending_fA p \ tso_pending p is_mw_fA" abbreviation (input) "tso_pending_fM p \ tso_pending p is_mw_fM" abbreviation (input) "tso_pending_mark p \ tso_pending p is_mw_Mark" abbreviation (input) "tso_pending_mutate p \ tso_pending p is_mw_Mutate" abbreviation (input) "tso_pending_phase p \ tso_pending p is_mw_Phase" abbreviation (input) "tso_no_pending_marks \ \<^bold>\p. LIST_NULL (tso_pending_mark p)" text\ A somewhat-useful abstraction of the heap, following l4.verified, which asserts that there is an object at the given reference with the given property. \ definition obj_at :: "(('field, 'ref) object \ bool) \ 'ref \ ('field, 'mut, 'ref) lsts_pred" where "obj_at P r \ \s. case sys_heap s r of None \ False | Some obj \ P obj" abbreviation (input) valid_ref :: "'ref \ ('field, 'mut, 'ref) lsts_pred" where "valid_ref r \ obj_at \True\ r" definition valid_null_ref :: "'ref option \ ('field, 'mut, 'ref) lsts_pred" where "valid_null_ref r \ case r of None \ \True\ | Some r' \ valid_ref r'" abbreviation pred_points_to :: "'ref \ 'ref \ ('field, 'mut, 'ref) lsts_pred" (infix "points'_to" 51) where "x points_to y \ \s. obj_at (\obj. y \ ran (obj_fields obj)) x s" text\ We use Isabelle's standard transitive-reflexive closure to define reachability through the heap. \ abbreviation pred_reaches :: "'ref \ 'ref \ ('field, 'mut, 'ref) lsts_pred" (infix "reaches" 51) where "x reaches y \ \s. (\x y. (x points_to y) s)\<^sup>*\<^sup>* x y" text\ The predicate \obj_at_field_on_heap\ asserts that @{term \valid_ref r\} and if \f\ is a field of the object referred to by \r\ then it it satisfies \P\. \ (* FIXME rename *) definition obj_at_field_on_heap :: "('ref \ bool) \ 'ref \ 'field \ ('field, 'mut, 'ref) lsts_pred" where "obj_at_field_on_heap P r f \ \s. case Option.map_option obj_fields (sys_heap s r) of None \ False | Some fs \ (case fs f of None \ True | Some r' \ P r')" subsection\ Garbage collector locations \ locset_definition "idle_locs = prefixed ''idle''" locset_definition "init_locs = prefixed ''init''" locset_definition "mark_locs = prefixed ''mark''" locset_definition "mark_loop_locs = prefixed ''mark_loop''" locset_definition "sweep_locs = prefixed ''sweep''" (*<*) subsection\ Lemma bucket \ lemma obj_at_split: "Q (obj_at P r s) = ((sys_heap s r = None \ Q False) \ (\obj. sys_heap s r = Some obj \ Q (P obj)))" by (simp add: obj_at_def split: option.splits) lemma obj_at_split_asm: "Q (obj_at P r s) = (\ ((sys_heap s r = None \ \Q False) \ (\obj. sys_heap s r = Some obj \ \ Q (P obj))))" by (simp add: obj_at_def split: option.splits) lemmas obj_at_splits = obj_at_split obj_at_split_asm lemma p_not_sys: "p \ sys \ p = gc \ (\m. p = mutator m)" by (cases p) simp_all lemma (in mut_m') m'm[iff]: "m' \ m" using mm' by blast lemma obj_at_eq_imp: "eq_imp (\(_::unit) s. map_option P (sys_heap s r)) (obj_at P r)" by (simp add: eq_imp_def obj_at_def split: option.splits) lemmas obj_at_fun_upd[simp] = eq_imp_fun_upd[OF obj_at_eq_imp, simplified eq_imp_simps] simps_of_case handshake_step_simps[simp]: handshake_step_def (splits: handshake_phase.split) (* Looks like a good idea but seems very weak. The split rules do a better job. *) lemma obj_at_weakenE[elim]: "\ obj_at P r s; \s. P s \ P' s \ \ obj_at P' r s" by (simp add: obj_at_def split: option.splits) lemma obj_at_ref_sweep_loop_free[simp]: "obj_at P r (s(sys := (s sys)\heap := (sys_heap s)(r' := None)\)) \ obj_at P r s \ r \ r'" by (clarsimp split: obj_at_splits) lemma obj_at_alloc[simp]: "sys_heap s r' = None \ obj_at P r (s(Mut m := mut_m_s', sys := (s sys)\ heap := sys_heap s(r' \ obj) \)) \ (obj_at P r s \ (r = r' \ P obj))" by (simp add: ran_def split: obj_at_splits) simps_of_case do_read_action_simps[simp]: fun_cong[OF do_read_action_def[simplified atomize_eq]] (splits: mem_read_action.split) simps_of_case do_write_action_simps[simp]: fun_cong[OF do_write_action_def[simplified atomize_eq]] (splits: mem_write_action.split) (* This gives some indication of how much we're cheating on the TSO front. *) lemma do_write_action_prj_simps[simp]: "fM (do_write_action w s) = fl \ (fM s = fl \ w \ mw_fM (\fM s)) \ w = mw_fM fl" "fl = fM (do_write_action w s) \ (fl = fM s \ w \ mw_fM (\fM s)) \ w = mw_fM fl" "fA (do_write_action w s) = fl \ (fA s = fl \ w \ mw_fA (\fA s)) \ w = mw_fA fl" "fl = fA (do_write_action w s) \ (fl = fA s \ w \ mw_fA (\fA s)) \ w = mw_fA fl" "ghost_handshake_in_sync (do_write_action w s) = ghost_handshake_in_sync s" "ghost_handshake_phase (do_write_action w s) = ghost_handshake_phase s" "ghost_honorary_grey (do_write_action w s) = ghost_honorary_grey s" "handshake_pending (do_write_action w s) = handshake_pending s" "handshake_type (do_write_action w s) = handshake_type s" "heap (do_write_action w s) r = None \ heap s r = None" "mem_lock (do_write_action w s) = mem_lock s" "phase (do_write_action w s) = ph \ (phase s = ph \ (\ph'. w \ mw_Phase ph') \ w = mw_Phase ph)" "ph = phase (do_write_action w s) \ (ph = phase s \ (\ph'. w \ mw_Phase ph') \ w = mw_Phase ph)" "W (do_write_action w s) = W s" by (auto simp: do_write_action_def split: mem_write_action.splits obj_at_splits) lemma valid_null_ref_eq_imp: "eq_imp (\(_::unit) s. r \ (Option.map_option \True\ \ sys_heap s)) (valid_null_ref r)" by (simp add: eq_imp_def obj_at_def valid_null_ref_def split: option.splits) lemmas valid_null_ref_fun_upd[simp] = eq_imp_fun_upd[OF valid_null_ref_eq_imp, simplified] lemma valid_null_ref_simps[simp]: "valid_null_ref None s" "valid_null_ref (Some r) s \ valid_ref r s" by (simp_all add: valid_null_ref_def) lemma obj_at_field_on_heap_eq_imp: "eq_imp (\(_::unit) s. sys_heap s r) (obj_at_field_on_heap P r f)" by (simp add: eq_imp_def obj_at_field_on_heap_def) lemmas obj_at_field_on_heap_fun_upd[simp] = eq_imp_fun_upd[OF obj_at_field_on_heap_eq_imp, simplified eq_imp_simps] lemma obj_at_field_on_heapE[elim]: "\ obj_at_field_on_heap P r f s; sys_heap s' r = sys_heap s r; \r'. P r' \ P' r' \ \ obj_at_field_on_heap P' r f s'" by (simp add: obj_at_field_on_heap_def split: option.splits) lemma obj_at_field_on_heap_mw_simps[simp]: "obj_at_field_on_heap P r0 f0 (s(sys := (s sys)\ heap := (sys_heap s)(r := Option.map_option (\obj :: ('field, 'ref) object. obj\obj_fields := (obj_fields obj)(f := opt_r')\) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s Sys))(p := ws) \)) \ ( (r \ r0 \ f \ f0) \ obj_at_field_on_heap P r0 f0 s ) \ (r = r0 \ f = f0 \ valid_ref r s \ (case opt_r' of Some r'' \ P r'' | _ \ True))" "obj_at_field_on_heap P r f (s(sys := s sys\heap := (sys_heap s)(r' := Option.map_option (obj_mark_update (\_. fl)) (sys_heap s r')), mem_write_buffers := sb'\)) \ obj_at_field_on_heap P r f s" by (auto simp: obj_at_field_on_heap_def split: option.splits obj_at_splits) lemma obj_at_field_on_heap_no_pending_writes: "\ sys_read (mutator m) (mr_Ref r f) (s sys) = mv_Ref opt_r'; \opt_r'. mw_Mutate r f opt_r' \ set (sys_mem_write_buffers (mutator m) s); valid_ref r s \ \ obj_at_field_on_heap (\r. opt_r' = Some r) r f s" apply (clarsimp simp: sys_read_def fold_writes_def) apply (rule fold_invariant[where P="\fr. obj_at_field_on_heap (\r'. heap (fr (s sys)) r \ (\obj. obj_fields obj f) = Some r') r f s" and Q="\w. w \ set (sys_mem_write_buffers (mutator m) s)"]) apply auto[1] apply (clarsimp simp: obj_at_field_on_heap_def split: option.splits obj_at_splits) apply (auto simp: obj_at_field_on_heap_def do_write_action_def Option.map_option_case split: option.splits obj_at_splits mem_write_action.splits) done lemma obj_at_field_on_heap_imp_valid_ref[elim]: "obj_at_field_on_heap P r f s \ valid_ref r s" "obj_at_field_on_heap P r f s \ valid_null_ref (Some r) s" by (auto simp: obj_at_field_on_heap_def valid_null_ref_def split: obj_at_splits option.splits) lemma obj_at_field_on_heap_weakenE[elim]: "\ obj_at_field_on_heap P r f s; \s. P s \ P' s\ \ obj_at_field_on_heap P' r f s" by (clarsimp simp: obj_at_field_on_heap_def split: option.splits) lemma Set_bind_insert[simp]: "insert a A \ B = B a \ (A \ B)" by (auto simp: Set.bind_def) lemma option_bind_invE[elim]: "\ f \ g = None; \a. \ f = Some a; g a = None \ \ Q; f = None \ Q \ \ Q" "\ f \ g = Some x; \a. \ f = Some a; g a = Some x \ \ Q \ \ Q" by (case_tac [!] f) simp_all (*>*) (*<*) end (*>*) diff --git a/thys/ConcurrentGC/StrongTricolour.thy b/thys/ConcurrentGC/StrongTricolour.thy --- a/thys/ConcurrentGC/StrongTricolour.thy +++ b/thys/ConcurrentGC/StrongTricolour.thy @@ -1,3400 +1,3402 @@ (*<*) (* * Copyright 2015, NICTA * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * @TAG(NICTA_BSD) *) theory StrongTricolour imports MarkObject begin +declare subst_all [simp del] [[simproc del: defined_all]] + (*>*) subsection\The strong-tricolour invariant \label{sec:strong-tricolour-invariant} \ text\ As the GC algorithm uses both insertion and deletion barriers, it preserves the \emph{strong tricolour-invariant}: \ abbreviation points_to_white :: "'ref \ 'ref \ ('field, 'mut, 'ref) lsts_pred" (infix "points'_to'_white" 51) where "x points_to_white y \ x points_to y \<^bold>\ white y" definition strong_tricolour_inv :: "('field, 'mut, 'ref) lsts_pred" where "strong_tricolour_inv = (\<^bold>\b w. black b \<^bold>\ \<^bold>\(b points_to_white w))" text\ Intuitively this invariant says that there are no pointers from completely processed objects to the unexplored space; i.e., the grey references properly separate the two. In contrast the weak tricolour invariant allows such pointers, provided there is a grey reference that protects the unexplored object. \ definition has_white_path_to :: "'ref \ 'ref \ ('field, 'mut, 'ref) lsts_pred" (infix "has'_white'_path'_to" 51) where "x has_white_path_to y \ \s. (\x y. (x points_to_white y) s)\<^sup>*\<^sup>* x y" definition grey_protects_white :: "'ref \ 'ref \ ('field, 'mut, 'ref) lsts_pred" (infix "grey'_protects'_white" 51) where "g grey_protects_white w = (grey g \<^bold>\ g has_white_path_to w)" definition weak_tricolour_inv :: "('field, 'mut, 'ref) lsts_pred" where "weak_tricolour_inv = (\<^bold>\b w. black b \<^bold>\ b points_to_white w \<^bold>\ (\<^bold>\g. g grey_protects_white w))" lemma "strong_tricolour_inv s \ weak_tricolour_inv s" (*<*) by (clarsimp simp: strong_tricolour_inv_def weak_tricolour_inv_def grey_protects_white_def) (*>*) text\ The key invariant that the mutators establish as they perform \get_roots\: they protect their white-reachable references with grey objects. \ definition in_snapshot :: "'ref \ ('field, 'mut, 'ref) lsts_pred" where "in_snapshot r = (black r \<^bold>\ (\<^bold>\g. g grey_protects_white r))" definition (in mut_m) reachable_snapshot_inv :: "('field, 'mut, 'ref) lsts_pred" where "reachable_snapshot_inv = (\<^bold>\r. reachable r \<^bold>\ in_snapshot r)" text\ Note that it is not easy to specify precisely when the snapshot (of objects the GC will retain) is taken due to the raggedness of the initialisation. In some phases we need to know that the insertion and deletion barriers are installed, in order to preserve the snapshot. These can ignore TSO effects as marks hit system memory in a timely way. \ abbreviation marked_insertion :: "('field, 'ref) mem_write_action \ ('field, 'mut, 'ref) lsts_pred" where "marked_insertion w \ \s. case w of mw_Mutate r f (Some r') \ marked r' s | _ \ True" definition (in mut_m) marked_insertions :: "('field, 'mut, 'ref) lsts_pred" where "marked_insertions = (\<^bold>\w. tso_pending_write (mutator m) w \<^bold>\ marked_insertion w)" abbreviation marked_deletion :: "('field, 'ref) mem_write_action \ ('field, 'mut, 'ref) lsts_pred" where "marked_deletion w \ \s. case w of mw_Mutate r f opt_r' \ obj_at_field_on_heap (\r'. marked r' s) r f s | _ \ True" definition (in mut_m) marked_deletions :: "('field, 'mut, 'ref) lsts_pred" where "marked_deletions = (\<^bold>\w. tso_pending_write (mutator m) w \<^bold>\ marked_deletion w)" text\ Finally, in some phases the heap is somewhat monochrome. \ definition black_heap :: "('field, 'mut, 'ref) lsts_pred" where "black_heap = (\<^bold>\r. valid_ref r \<^bold>\ black r)" definition white_heap :: "('field, 'mut, 'ref) lsts_pred" where "white_heap = (\<^bold>\r. valid_ref r \<^bold>\ white r)" definition no_black_refs :: "('field, 'mut, 'ref) lsts_pred" where "no_black_refs = (\<^bold>\r. \<^bold>\(black r))" definition no_grey_refs :: "('field, 'mut, 'ref) lsts_pred" where "no_grey_refs = (\<^bold>\r. \<^bold>\(grey r))" (*<*) lemma no_black_refsD: "no_black_refs s \ \black r s" unfolding no_black_refs_def by simp lemma has_white_path_to_eq_imp: "eq_imp (\(_::unit). sys_fM \<^bold>\ sys_heap) (x has_white_path_to y)" by (clarsimp simp: eq_imp_def has_white_path_to_def obj_at_def cong: option.case_cong) lemmas has_white_path_to_fun_upd[simp] = eq_imp_fun_upd[OF has_white_path_to_eq_imp, simplified eq_imp_simps, rule_format] lemma has_white_path_toD[dest]: "(x has_white_path_to y) s \ white y s \ x = y" unfolding has_white_path_to_def by (fastforce elim: rtranclp.cases) lemma has_white_path_toI[iff]: "(x has_white_path_to x) s" by (simp add: has_white_path_to_def) lemma has_white_path_toE[elim!]: "\ (x points_to y) s; white y s \ \ (x has_white_path_to y) s" "\ (x has_white_path_to y) s; (y points_to z) s; white z s \ \ (x has_white_path_to z) s" by (auto simp: has_white_path_to_def elim: rtranclp.intros(2)) lemma has_white_path_to_reaches[elim]: "(x has_white_path_to y) s \ (x reaches y) s" unfolding has_white_path_to_def by (induct rule: rtranclp.induct) (auto intro: rtranclp.intros(2)) lemma has_white_path_to_blacken[simp]: "(x has_white_path_to w) (s(gc := s gc\ W := gc_W s - rs \)) \ (x has_white_path_to w) s" by (simp add: has_white_path_to_def) text\WL\ lemma WL_mo_co_mark[simp]: "ghost_honorary_grey (s p) = {} \ WL p' (s(p := s p\ ghost_honorary_grey := rs \)) = WL p' s \ { r |r. p' = p \ r \ rs}" by (auto simp: WL_def) lemma WL_blacken[simp]: "gc_ghost_honorary_grey s = {} \ WL p (s(gc := s gc\ W := gc_W s - rs \)) = WL p s - { r |r. p = gc \ r \ rs }" by (auto simp: WL_def) lemma WL_hs_done[simp]: "ghost_honorary_grey (s (mutator m)) = {} \ WL p (s(mutator m := s (mutator m)\ W := {}, ghost_handshake_phase := hp' \, sys := s sys\ handshake_pending := hsp', W := sys_W s \ W (s (mutator m)), ghost_handshake_in_sync := in' \)) = (case p of gc \ WL gc s | mutator m' \ (if m' = m then {} else WL (mutator m') s) | sys \ WL sys s \ WL (mutator m) s)" "ghost_honorary_grey (s (mutator m)) = {} \ WL p (s(mutator m := s (mutator m)\ W := {} \, sys := s sys\ handshake_pending := hsp', W := sys_W s \ W (s (mutator m)), ghost_handshake_in_sync := in' \)) = (case p of gc \ WL gc s | mutator m' \ (if m' = m then {} else WL (mutator m') s) | sys \ WL sys s \ WL (mutator m) s)" by (auto simp: WL_def split: process_name.splits) lemma colours_load_W[iff]: "gc_W s = {} \ black r (s(gc := (s gc)\W := W (s sys)\, sys := (s sys)\W := {}\)) \ black r s" "gc_W s = {} \ grey r (s(gc := (s gc)\W := W (s sys)\, sys := (s sys)\W := {}\)) \ grey r s" apply (simp_all add: black_def grey_def WL_def) apply safe apply (case_tac [!] x) apply blast+ done lemma colours_sweep_loop_free[iff]: "black r (s(sys := s sys\heap := (heap (s sys))(r' := None)\)) \ (black r s \ r \ r')" "grey r (s(sys := s sys\heap := (heap (s sys))(r' := None)\)) \ (grey r s)" "white r (s(sys := s sys\heap := (heap (s sys))(r' := None)\)) \ (white r s \ r \ r')" by (auto simp: black_def grey_def split: obj_at_splits) lemma colours_get_work_done[simp]: "black r (s(mutator m := (s (mutator m))\W := {}\, sys := (s sys)\ handshake_pending := hp', W := W (s sys) \ W (s (mutator m)), ghost_handshake_in_sync := his' \)) \ black r s" "grey r (s(mutator m := (s (mutator m))\W := {}\, sys := (s sys)\ handshake_pending := hp', W := W (s sys) \ W (s (mutator m)), ghost_handshake_in_sync := his' \)) \ grey r s" "white r (s(mutator m := (s (mutator m))\W := {}\, sys := (s sys)\ handshake_pending := hp', W := W (s sys) \ W (s (mutator m)), ghost_handshake_in_sync := his' \)) \ white r s" apply (simp_all add: black_def grey_def WL_def split: obj_at_splits) apply auto apply (metis (full_types) process_name.distinct(3)) by metis lemma colours_get_roots_done[simp]: "black r (s(mutator m := (s (mutator m))\ W := {}, ghost_handshake_phase := hs' \, sys := (s sys)\ handshake_pending := hp', W := W (s sys) \ W (s (mutator m)), ghost_handshake_in_sync := his' \)) \ black r s" "grey r (s(mutator m := (s (mutator m))\ W := {}, ghost_handshake_phase := hs' \, sys := (s sys)\ handshake_pending := hp', W := W (s sys) \ W (s (mutator m)), ghost_handshake_in_sync := his' \)) \ grey r s" "white r (s(mutator m := (s (mutator m))\ W := {}, ghost_handshake_phase := hs' \, sys := (s sys)\ handshake_pending := hp', W := W (s sys) \ W (s (mutator m)), ghost_handshake_in_sync := his' \)) \ white r s" apply (simp_all add: black_def grey_def WL_def split: obj_at_splits) apply auto apply (metis process_name.distinct(3)) by metis lemma colours_mark[simp]: "\ ghost_honorary_grey (s p) = {} \ \ black b (s(p := s p\ghost_honorary_grey := {r}\)) \ black b s \ b \ r" "\ ghost_honorary_grey (s p) = {} \ \ grey g (s(p := (s p)\ghost_honorary_grey := {r}\)) \ grey g s \ g = r" "\ ghost_honorary_grey (s p) = {} \ \ white w (s(p := s p\ghost_honorary_grey := {r}\)) \ white w s" by (auto simp: black_def grey_def) lemma colours_flip_fM[simp]: "fl \ sys_fM s \ black b (s(sys := (s sys)\fM := fl, mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\)) \ white b s \ \grey b s" by (simp_all add: black_def) lemma colours_alloc[simp]: "heap (s sys) r' = None \ black r (s(mutator m := (s (mutator m))\ roots := roots' \, sys := (s sys)\heap := heap (s sys)(r' \ \obj_mark = fl, obj_fields = Map.empty\)\)) \ black r s \ (r' = r \ fl = sys_fM s \ \grey r' s)" "grey r (s(mutator m := (s (mutator m))\ roots := roots' \, sys := (s sys)\heap := heap (s sys)(r' \ \obj_mark = fl, obj_fields = Map.empty\)\)) \ grey r s" "heap (s sys) r' = None \ white r (s(mutator m := (s (mutator m))\ roots := roots' \, sys := (s sys)\heap := heap (s sys)(r' \ \obj_mark = fl, obj_fields = Map.empty\)\)) \ white r s \ (r' = r \ fl \ sys_fM s)" by (auto simp: black_def split: obj_at_splits dest!: valid_refs_invD) lemma colours_blacken[simp]: "valid_W_inv s \ black b (s(gc := s gc\W := gc_W s - {r}\)) \ black b s \ (r \ gc_W s \ b = r)" "\ r \ gc_W s; valid_W_inv s \ \ grey g (s(gc := s gc\W := gc_W s - {r}\)) \ (grey g s \ g \ r)" "white w (s(gc := s gc\W := gc_W s - {r}\)) \ white w s" apply (auto simp: black_def grey_def WL_def split: obj_at_splits) apply metis+ done lemma colours_dequeue[simp]: "black b (s(sys := (s sys)\ heap := (sys_heap s)(r := Option.map_option (obj_mark_update (\_. fl)) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(p := ws) \)) \ (black b s \ b \ r) \ (b = r \ fl = sys_fM s \ valid_ref r s \ \grey r s)" by (auto simp: black_def split: obj_at_splits) lemma W_load_W[simp]: "gc_W s = {} \ (\p. W (if p = sys then (s sys)\W := {}\ else if p = gc then (s gc)\W := sys_W s\ else s p)) = (\p. W (s p))" apply (rule equalityI) apply (fastforce split: if_splits) apply clarsimp apply (rename_tac x p) apply (case_tac p) apply force+ done lemma WL_load_W[simp]: "gc_W s = {} \ (WL p (s(gc := (s gc)\W := sys_W s\, sys := (s sys)\W := {}\))) = (case p of gc \ WL gc s \ sys_W s | mutator m \ WL (mutator m) s | sys \ sys_ghost_honorary_grey s)" by (auto simp: WL_def split: process_name.splits) lemma colours_mo_co_W[simp]: "ghost_honorary_grey (s p') = {r} \ (WL p (s(p' := (s p')\W := insert r (W (s p')), ghost_honorary_grey := {}\))) = (WL p s)" "ghost_honorary_grey (s p') = {r} \ grey g (s(p' := (s p')\W := insert r (W (s p')), ghost_honorary_grey := {}\)) \ grey g s" by (force simp: grey_def WL_def split: process_name.splits if_splits)+ lemma no_grey_refs_eq_imp: "eq_imp (\(_::unit). (\s. \p. WL p s)) no_grey_refs" by (auto simp add: eq_imp_def grey_def no_grey_refs_def set_eq_iff) lemmas no_grey_refs_fun_upd[simp] = eq_imp_fun_upd[OF no_grey_refs_eq_imp, simplified eq_imp_simps, rule_format] lemma no_grey_refs_no_pending_marks: "\ no_grey_refs s; valid_W_inv s \ \ tso_no_pending_marks s" by (auto intro!: filter_False simp: no_grey_refs_def) lemma (in mut_m) reachable_blackD: "\ no_grey_refs s; reachable_snapshot_inv s; reachable r s \ \ black r s" by (simp add: no_grey_refs_def reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def) lemma (in mut_m) no_grey_refs_not_reachable: "\ no_grey_refs s; reachable_snapshot_inv s; white r s \ \ \reachable r s" by (fastforce simp: no_grey_refs_def reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def split: obj_at_splits) lemma (in mut_m) no_grey_refs_not_rootD: "\ no_grey_refs s; reachable_snapshot_inv s; white r s \ \ r \ mut_roots s \ r \ mut_ghost_honorary_root s \ r \ tso_write_refs s" apply (drule (2) no_grey_refs_not_reachable) apply (force simp: reachable_def) done lemma no_grey_refs_not_grey_reachableD: "no_grey_refs s \ \grey_reachable x s" by (clarsimp simp: no_grey_refs_def grey_reachable_def) lemma (in mut_m) reachable_snapshot_inv_white_root: "\ white w s; w \ mut_roots s \ w \ mut_ghost_honorary_root s; reachable_snapshot_inv s \ \ \g. (g grey_protects_white w) s" by (auto simp: reachable_snapshot_inv_def in_snapshot_def reachable_def grey_protects_white_def) lemma no_grey_refsD: "no_grey_refs s \ r \ W (s p)" "no_grey_refs s \ r \ WL p s" "no_grey_refs s \ r \ ghost_honorary_grey (s p)" by (auto simp: no_grey_refs_def) lemma no_grey_refs_marked[dest]: "\ marked r s; no_grey_refs s \ \ black r s" by (auto simp: no_grey_refs_def black_def) lemma no_grey_refs_bwD[dest]: "\ heap (s sys) r = Some obj; no_grey_refs s \ \ black r s \ white r s" by (clarsimp simp: black_def grey_def no_grey_refs_def split: obj_at_splits) text\tso write refs\ lemma tso_write_refs_simps[simp]: "mut_m.tso_write_refs m (s(mutator m' := s (mutator m')\roots := roots'\)) = mut_m.tso_write_refs m s" "mut_m.tso_write_refs m (s(mutator m' := s (mutator m')\ghost_honorary_root := {}\, sys := s sys\mem_write_buffers := (mem_write_buffers (s sys))(mutator m' := sys_mem_write_buffers (mutator m') s @ [mw_Mutate r f opt_r'])\)) = mut_m.tso_write_refs m s \ (if m' = m then write_refs (mw_Mutate r f opt_r') else {})" "mut_m.tso_write_refs m (s(sys := s sys\heap := (sys_heap s)(r' := None)\)) = mut_m.tso_write_refs m s" "mut_m.tso_write_refs m (s(mutator m' := s (mutator m')\roots := insert r (roots (s (mutator m')))\, sys := s sys\heap := sys_heap s(r \ obj)\)) = mut_m.tso_write_refs m s" "mut_m.tso_write_refs m (s(mutator m' := s (mutator m')\ghost_honorary_root := Option.set_option opt_r', ref := opt_r'\)) = mut_m.tso_write_refs m s" "mut_m.tso_write_refs m (s(sys := s sys\heap := (sys_heap s)(r := Option.map_option (\obj. obj\obj_fields := (obj_fields obj)(f := opt_r')\) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\)) = (if p = mutator m then \w \ set ws. write_refs w else mut_m.tso_write_refs m s)" "sys_mem_write_buffers p s = mw_Mark r fl # ws \ mut_m.tso_write_refs m (s(sys := s sys\heap := (sys_heap s)(r := Option.map_option (obj_mark_update (\_. fl)) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\)) = mut_m.tso_write_refs m s" by (auto simp: mut_m.tso_write_refs_def) lemma mutator_reachable_tso: "sys_mem_write_buffers (mutator m) s = mw_Mutate r f opt_r' # ws \ mut_m.reachable m r s \ (\r'. opt_r' = Some r' \ mut_m.reachable m r' s)" by (auto simp: mut_m.tso_write_refs_def) text\coloured heaps\ lemma black_heap_eq_imp: "eq_imp (\(_::unit). (\s. \p. WL p s) \<^bold>\ sys_fM \<^bold>\ (\s. Option.map_option obj_mark \ sys_heap s)) black_heap" apply (clarsimp simp add: eq_imp_def black_heap_def black_def grey_def all_conj_distrib fun_eq_iff split: option.splits) apply (rename_tac s s') apply (subgoal_tac "\x. marked x s \ marked x s'") apply (subgoal_tac "\x. valid_ref x s \ valid_ref x s'") apply (subgoal_tac "\x. (\p. x \ WL p s) \ (\p. x \ WL p s')") apply clarsimp apply (auto simp: set_eq_iff)[1] apply clarsimp apply (rename_tac x) apply (rule eq_impD[OF obj_at_eq_imp]) apply (drule_tac x=x in spec) apply (drule_tac f="Option.map_option \True\" in arg_cong) apply (clarsimp simp: map_option.compositionality o_def) apply (clarsimp simp: map_option.compositionality o_def) apply (rule eq_impD[OF obj_at_eq_imp]) apply clarsimp apply (rename_tac x) apply (drule_tac x=x in spec) apply (drule_tac f="Option.map_option (\fl. fl = sys_fM s)" in arg_cong) apply (simp add: map_option.compositionality o_def) done lemma white_heap_eq_imp: "eq_imp (\(_::unit). sys_fM \<^bold>\ (\s. Option.map_option obj_mark \ sys_heap s)) white_heap" apply (clarsimp simp: all_conj_distrib eq_imp_def white_heap_def obj_at_def fun_eq_iff split: option.splits) apply (rule iffI) apply (metis (hide_lams, no_types) map_option_eq_Some)+ done lemma no_black_refs_eq_imp: "eq_imp (\(_::unit). (\s. (\p. WL p s)) \<^bold>\ sys_fM \<^bold>\ (\s. Option.map_option obj_mark \ sys_heap s)) no_black_refs" apply (clarsimp simp add: eq_imp_def no_black_refs_def black_def grey_def all_conj_distrib fun_eq_iff set_eq_iff split: option.splits) apply (rename_tac s s') apply (subgoal_tac "\x. marked x s \ marked x s'") apply clarsimp apply clarsimp apply (rename_tac x) apply ( (drule_tac x=x in spec)+ )[1] apply (auto split: obj_at_splits) done lemmas black_heap_fun_upd[simp] = eq_imp_fun_upd[OF black_heap_eq_imp, simplified eq_imp_simps, rule_format] lemmas white_heap_fun_upd[simp] = eq_imp_fun_upd[OF white_heap_eq_imp, simplified eq_imp_simps, rule_format] lemmas no_black_refs_fun_upd[simp] = eq_imp_fun_upd[OF no_black_refs_eq_imp, simplified eq_imp_simps, rule_format] lemma white_heap_imp_no_black_refs[elim!]: "white_heap s \ no_black_refs s" apply (clarsimp simp: white_heap_def no_black_refs_def black_def) apply (rename_tac x) apply (drule_tac x=x in spec) apply (clarsimp split: obj_at_splits) done lemma (in sys) no_black_refs_dequeue[simp]: "\ sys_mem_write_buffers p s = mw_Mark r fl # ws; no_black_refs s; valid_W_inv s \ \ no_black_refs (s(sys := s sys\heap := (sys_heap s)(r := Option.map_option (obj_mark_update (\_. fl)) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\))" "\ sys_mem_write_buffers p s = mw_Mutate r f r' # ws; no_black_refs s \ \ no_black_refs (s(sys := s sys\heap := (sys_heap s)(r := Option.map_option (\obj. obj\obj_fields := (obj_fields obj)(f := r')\) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\))" by (auto simp: no_black_refs_def map_option.compositionality o_def) lemma black_heap_no_greys[elim]: "\ black_heap s; valid_refs_inv s \ \ no_grey_refs s" "\ no_grey_refs s; \r. marked r s \ \valid_ref r s \ \ black_heap s" by (auto simp: black_def black_heap_def no_grey_refs_def dest: valid_refs_invD) lemma heap_colours_alloc[simp]: "\ heap (s sys) r' = None; valid_refs_inv s \ \ black_heap (s(mutator m := s (mutator m)\roots := insert r' (roots (s (mutator m)))\, sys := s sys\heap := sys_heap s(r' \ \obj_mark = fl, obj_fields = Map.empty\)\)) \ black_heap s \ fl = sys_fM s" "heap (s sys) r' = None \ white_heap (s(mutator m := s (mutator m)\roots := insert r' (roots (s (mutator m)))\, sys := s sys\heap := sys_heap s(r' \ \obj_mark = fl, obj_fields = Map.empty\)\)) \ white_heap s \ fl \ sys_fM s" apply (simp_all add: black_heap_def white_heap_def split: obj_at_splits) apply (rule iffI) apply (intro allI conjI impI) apply (rename_tac x) apply (drule_tac x=x in spec) apply clarsimp apply (drule spec[where x=r'], auto dest!: valid_refs_invD split: obj_at_splits)[2] apply (rule iffI) apply (intro allI conjI impI) apply (rename_tac x obj) apply (drule_tac x=x in spec) apply clarsimp apply (drule spec[where x=r'], auto dest!: valid_refs_invD split: obj_at_splits)[2] done lemma heap_colours_colours: "black_heap s \ \white r s" "white_heap s \ \black r s" by (auto simp: black_heap_def white_heap_def dest!: spec[where x=r] split: obj_at_splits) lemma heap_colours_marked: "\ black_heap s; obj_at P r s \ \ marked r s" "\ white_heap s; obj_at P r s \ \ white r s" by (auto simp: black_heap_def white_heap_def dest!: spec[where x=r] split: obj_at_splits) lemma (in gc) obj_fields_marked_inv_blacken: "\ gc_field_set s = {}; obj_fields_marked_inv s; (gc_tmp_ref s points_to w) s; white w s; tso_writes_inv s \ \ False" by (simp add: obj_fields_marked_inv_def obj_at_field_on_heap_def ran_def split: option.splits obj_at_splits) lemma (in gc) obj_fields_marked_inv_has_white_path_to_blacken: "\ gc_field_set s = {}; gc_tmp_ref s \ gc_W s; (gc_tmp_ref s has_white_path_to w) s; obj_fields_marked_inv s; valid_W_inv s; tso_writes_inv s \ \ w = gc_tmp_ref s" apply (drule (1) valid_W_invD) apply (clarsimp simp: has_white_path_to_def) apply (erule converse_rtranclpE) apply (clarsimp split: obj_at_splits) apply clarsimp apply (simp add: obj_fields_marked_inv_def obj_at_field_on_heap_def ran_def split: option.splits obj_at_splits) done lemma fold_writes_points_to[rule_format, simplified conj_explode]: "heap (fold_writes ws (s sys)) r = Some obj \ obj_fields obj f = Some r' \ (r points_to r') s \ (\w \ set ws. r' \ write_refs w)" (is "?P (fold_writes ws) obj") unfolding fold_writes_def apply (rule spec[OF fold_invariant[where P="\fr. \obj. ?P fr obj" and Q="\w. w \ set ws"]]) apply simp apply (fastforce simp: ran_def split: obj_at_splits) apply clarsimp apply (drule (1) bspec) apply (clarsimp split: mem_write_action.split_asm if_splits) done lemma (in mut_m) reachable_load[simp]: assumes "sys_read (mutator m) (mr_Ref r f) (s sys) = mv_Ref r'" assumes "r \ mut_roots s" shows "mut_m.reachable m' y (s(mutator m := s (mutator m)\ roots := mut_roots s \ Option.set_option r' \)) \ mut_m.reachable m' y s" (is "?lhs = ?rhs") proof(cases "m' = m") case True show ?thesis proof(rule iffI) assume ?lhs with assms True show ?rhs apply (clarsimp simp: sys_read_def) apply (clarsimp simp: reachable_def tso_write_refs_def) apply (clarsimp simp: sys_read_def fold_writes_def) apply (elim disjE) apply blast defer apply blast apply blast apply (fold fold_writes_def) apply clarsimp apply (drule (1) fold_writes_points_to) apply (erule disjE) apply (fastforce elim!: converse_rtranclp_into_rtranclp[rotated] split: obj_at_splits intro!: ranI) apply clarsimp apply (case_tac w, simp_all) apply (erule disjE) apply (rule_tac x=x in exI) apply (fastforce elim!: converse_rtranclp_into_rtranclp[rotated] split: obj_at_splits intro!: ranI) apply (rule_tac x=x in exI) apply (fastforce elim!: converse_rtranclp_into_rtranclp[rotated] split: obj_at_splits intro!: ranI) done (* filthy *) next assume ?rhs with True show ?lhs by (fastforce simp: reachable_def) qed qed simp lemma (in mut_m) reachable_deref_del[simp]: "\ sys_read (mutator m) (mr_Ref r f) (s sys) = mv_Ref opt_r'; r \ mut_roots s; mut_ghost_honorary_root s = {} \ \ mut_m.reachable m' y (s(mutator m := s (mutator m)\ ghost_honorary_root := Option.set_option opt_r', ref := opt_r' \)) \ mut_m.reachable m' y s" apply (clarsimp simp: mut_m.reachable_def sys_read_def) apply (rule iffI) apply clarsimp apply (elim disjE) apply metis apply (erule option_bind_invE; auto dest!: fold_writes_points_to) apply (auto elim!: converse_rtranclp_into_rtranclp[rotated] simp: tso_write_refs_def) done text\strong tricolour\ lemma strong_tricolour_invD: "\ black x s; (x points_to y) s; valid_ref y s; strong_tricolour_inv s \ \ marked y s" apply (clarsimp simp: strong_tricolour_inv_def) apply (drule spec[where x=x]) apply (auto split: obj_at_splits) done text\grey protects white\ lemma grey_protects_whiteD[dest]: "(g grey_protects_white w) s \ grey g s \ (g = w \ white w s)" by (auto simp: grey_protects_white_def) lemma grey_protects_whiteI[iff]: "grey g s \ (g grey_protects_white g) s" by (simp add: grey_protects_white_def) lemma grey_protects_whiteE[elim!]: "\ (g points_to w) s; grey g s; white w s \ \ (g grey_protects_white w) s" "\ (g grey_protects_white y) s; (y points_to w) s; white w s \ \ (g grey_protects_white w) s" by (auto simp: grey_protects_white_def) lemma grey_protects_white_reaches[elim]: "(g grey_protects_white w) s \ (g reaches w) s" by (auto simp: grey_protects_white_def) lemma grey_protects_white_hs_done[simp]: "(g grey_protects_white w) (s(mutator m := s (mutator m)\ W := {}, ghost_handshake_phase := hs' \, sys := s sys\ handshake_pending := hp', W := sys_W s \ W (s (mutator m)), ghost_handshake_in_sync := his' \)) \ (g grey_protects_white w) s" by (simp add: grey_protects_white_def) lemma grey_protects_white_alloc[simp]: "\ fl = sys_fM s; sys_heap s r = None \ \ (g grey_protects_white w) (s(mutator m := s (mutator m)\roots := insert r (roots (s (mutator m)))\, sys := s sys\heap := sys_heap s(r \ \obj_mark = fl, obj_fields = Map.empty\)\)) \ (g grey_protects_white w) s" by (clarsimp simp: grey_protects_white_def has_white_path_to_def) text\reachable snapshot inv\ lemma (in mut_m) reachable_snapshot_invI[intro]: "(\y. reachable y s \ in_snapshot y s) \ reachable_snapshot_inv s" by (simp add: reachable_snapshot_inv_def) lemma (in mut_m) reachable_snapshot_inv_eq_imp: "eq_imp (\r. mut_roots \<^bold>\ mut_ghost_honorary_root \<^bold>\ (\s. r \ (\p. WL p s)) \<^bold>\ sys_fM \<^bold>\ sys_heap \<^bold>\ tso_pending_mutate (mutator m)) reachable_snapshot_inv" apply (clarsimp simp: eq_imp_def mut_m.reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def) apply (rename_tac s s') apply (clarsimp simp: black_def grey_def obj_at_def cong: option.case_cong) apply (subst eq_impD[OF has_white_path_to_eq_imp]) defer apply (subst eq_impD[OF reachable_eq_imp]) defer apply (subgoal_tac "\x. (\p. x \ WL p s) \ (\p. x \ WL p s')") apply force apply blast apply simp apply simp done lemmas reachable_snapshot_fun_upd[simp] = eq_imp_fun_upd[OF mut_m.reachable_snapshot_inv_eq_imp, simplified eq_imp_simps, rule_format] lemma in_snapshotI[intro]: "black r s \ in_snapshot r s" "grey r s \ in_snapshot r s" "\ white w s; (g grey_protects_white w) s \ \ in_snapshot w s" by (auto simp: in_snapshot_def) lemma in_snapshot_colours: "in_snapshot r s \ black r s \ grey r s \ white r s" by (auto simp: in_snapshot_def) lemma in_snapshot_valid_ref: "\ in_snapshot r s; valid_refs_inv s \ \ valid_ref r s" by (fastforce dest: in_snapshot_colours) lemma (in mut_m) reachable_snapshot_inv_sweep_loop_free: fixes s :: "('field, 'mut, 'ref) lsts" assumes nmr: "white r s" assumes ngs: "no_grey_refs s" assumes rsi: "reachable_snapshot_inv s" shows "reachable_snapshot_inv (s(sys := (s sys)\heap := (heap (s sys))(r := None)\))" (is "reachable_snapshot_inv ?s'") proof fix y :: "'ref" assume rx: "reachable y ?s'" then have "black y s \ y \ r" proof(induct rule: reachable_induct) case (root x) with ngs nmr rsi show ?case by (auto dest: reachable_blackD) next case (ghost_honorary_root x) with ngs nmr rsi show ?case apply - apply (frule (1) reachable_blackD[where r=x]) apply (auto simp: reachable_def) done next case (tso_root x) with ngs nmr rsi show ?case apply - apply (frule (1) reachable_blackD[where r=x]) apply (auto simp: reachable_def tso_write_refs_def) done next case (reaches x y) with ngs nmr rsi show ?case apply (clarsimp simp: reachable_def) apply (drule predicate2D[OF rtranclp_mono[where s="\x y. (x points_to y) s", OF predicate2I], rotated]) apply (clarsimp split: obj_at_splits if_splits) apply (rule conjI) apply (rule reachable_blackD, assumption, assumption) apply (simp add: reachable_def) apply (blast intro: rtranclp.intros(2)) apply clarsimp apply (frule (1) reachable_blackD[where r=r]) apply (simp add: reachable_def) apply (blast intro: rtranclp.intros(2)) apply auto done qed then show "in_snapshot y ?s'" unfolding in_snapshot_def by simp qed lemma reachableI2[intro]: "x \ mut_m.mut_ghost_honorary_root m s \ mut_m.reachable m x s" by (auto simp: mut_m.reachable_def) lemma reachable_alloc[simp]: assumes rn: "sys_heap s r = None" shows "mut_m.reachable m r' (s(mutator m' := (s (mutator m'))\roots := insert r (roots (s (mutator m')))\, sys := (s sys)\heap := (sys_heap s)(r \ \obj_mark = fl, obj_fields = Map.empty\)\)) \ mut_m.reachable m r' s \ (m' = m \ r' = r)" (is "?lhs \ ?rhs") proof(rule iffI) assume ?lhs from this assms show ?rhs proof(induct rule: reachable_induct) case (reaches x y) then show ?case by clarsimp (fastforce simp: mut_m.reachable_def elim: rtranclp.intros(2) split: obj_at_splits) qed (auto split: if_splits) next assume ?rhs then show ?lhs proof(rule disjE) assume "mut_m.reachable m r' s" then show ?thesis proof(induct rule: reachable_induct) case (tso_root x) then show ?case by (fastforce simp: mut_m.reachable_def simp del: fun_upd_apply) next case (reaches x y) with rn show ?case by (fastforce simp: mut_m.reachable_def simp del: fun_upd_apply elim: rtranclp.intros(2)) qed auto next assume "m' = m \ r' = r" with rn show ?thesis by (fastforce simp: mut_m.reachable_def) qed qed lemma (in mut_m) reachable_snapshot_inv_alloc[simp]: fixes s :: "('field, 'mut, 'ref) lsts" assumes rn: "sys_heap s r = None" assumes fl: "fl = sys_fM s" assumes vri: "valid_refs_inv s" assumes rsi: "reachable_snapshot_inv s" shows "reachable_snapshot_inv (s(mutator m' := (s (mutator m'))\roots := insert r (roots (s (mutator m')))\, sys := (s sys)\heap := (sys_heap s)(r \ \obj_mark = fl, obj_fields = Map.empty\)\))" (is "reachable_snapshot_inv ?s'") using assms unfolding reachable_snapshot_inv_def in_snapshot_def by (auto simp del: reachable_fun_upd) lemma (in mut_m) reachable_snapshot_inv_discard_roots[simp]: "\ reachable_snapshot_inv s; roots' \ roots (s (mutator m)) \ \ reachable_snapshot_inv (s(mutator m := (s (mutator m))\roots := roots'\))" unfolding reachable_snapshot_inv_def reachable_def in_snapshot_def grey_protects_white_def by auto lemma grey_protects_white_mark[simp]: assumes ghg: "ghost_honorary_grey (s p) = {}" shows "(\g. (g grey_protects_white w) (s(p := s p\ ghost_honorary_grey := {r} \))) \ (\g'. (g' grey_protects_white w) s) \ (r has_white_path_to w) s" (is "?lhs \ ?rhs") proof assume ?lhs then obtain g where "(g grey_protects_white w) (s(p := s p\ghost_honorary_grey := {r}\))" by blast from this ghg show ?rhs apply (clarsimp simp: grey_protects_white_def has_white_path_to_def if_distrib cong: if_cong) apply (rotate_tac 2) apply (induct rule: rtranclp.induct) apply (auto intro: rtranclp.intros(2)) done next assume ?rhs then show ?lhs proof(safe) fix g assume "(g grey_protects_white w) s" with ghg show ?thesis apply (clarsimp simp: grey_protects_white_def has_white_path_to_def) apply (rotate_tac 2) apply (induct rule: rtranclp.induct) apply (auto intro: rtranclp.intros(2)) done next assume "(r has_white_path_to w) s" with ghg show ?thesis by (auto simp: grey_protects_white_def has_white_path_to_def) qed qed (*>*) subsection\Invariants\ text (in mut_m) \ We need phase invariants in terms of both @{const "mut_ghost_handshake_phase"} and @{const "sys_ghost_handshake_phase"} which respectively track what the mutators and GC know by virtue of the synchronisation structure of the system. Read the following as ``when mutator \m\ is past the specified handshake, and has yet to reach the next one, ... holds.'' \ primrec (in mut_m) mutator_phase_inv_aux :: "handshake_phase \ ('field, 'mut, 'ref) lsts_pred" where "mutator_phase_inv_aux hp_Idle = \True\" | "mutator_phase_inv_aux hp_IdleInit = no_black_refs" | "mutator_phase_inv_aux hp_InitMark = marked_insertions" | "mutator_phase_inv_aux hp_Mark = (marked_insertions \<^bold>\ marked_deletions)" | "mutator_phase_inv_aux hp_IdleMarkSweep = (marked_insertions \<^bold>\ marked_deletions \<^bold>\ reachable_snapshot_inv)" abbreviation (in mut_m) mutator_phase_inv :: "('field, 'mut, 'ref) lsts_pred" where "mutator_phase_inv s \ mutator_phase_inv_aux (mut_ghost_handshake_phase s) s" abbreviation mutators_phase_inv :: "('field, 'mut, 'ref) lsts_pred" where "mutators_phase_inv \ (\<^bold>\m. mut_m.mutator_phase_inv m)" text\ This is what the GC guarantees. Read this as ``when the GC is at or past the specified handshake, ... holds.'' \ primrec sys_phase_inv_aux :: "handshake_phase \ ('field, 'mut, 'ref) lsts_pred" where "sys_phase_inv_aux hp_Idle = ( (If sys_fA \<^bold>= sys_fM Then black_heap Else white_heap) \<^bold>\ no_grey_refs )" | "sys_phase_inv_aux hp_IdleInit = no_black_refs" | "sys_phase_inv_aux hp_InitMark = (sys_fA \<^bold>\ sys_fM \<^bold>\ no_black_refs)" | "sys_phase_inv_aux hp_Mark = \True\" | "sys_phase_inv_aux hp_IdleMarkSweep = ( (sys_phase \<^bold>= \ph_Idle\ \<^bold>\ tso_pending_write gc (mw_Phase ph_Idle)) \<^bold>\ no_grey_refs )" abbreviation sys_phase_inv :: "('field, 'mut, 'ref) lsts_pred" where "sys_phase_inv s \ sys_phase_inv_aux (sys_ghost_handshake_phase s) s" (*<*) declare mut_m.mutator_phase_inv_aux.simps[simp] case_of_simps mutator_phase_inv_aux_case: mut_m.mutator_phase_inv_aux.simps case_of_simps sys_phase_inv_aux_case: sys_phase_inv_aux.simps lemma tso_pending_mutate_cong: "\ filter is_mw_Mutate (sys_mem_write_buffers p s) = filter is_mw_Mutate (sys_mem_write_buffers p s'); \r f r'. P r f r' \ Q r f r' \ \ (\r f r'. mw_Mutate r f r' \ set (sys_mem_write_buffers p s) \ P r f r') = (\r f r'. mw_Mutate r f r' \ set (sys_mem_write_buffers p s') \ Q r f r')" apply (intro iff_allI) apply (auto dest!: arg_cong[where f=set]) done lemma (in mut_m) marked_insertions_eq_imp: "eq_imp (\(_::unit). sys_fM \<^bold>\ (\s. Option.map_option obj_mark \ sys_heap s) \<^bold>\ tso_pending_mutate (mutator m)) marked_insertions" apply (clarsimp simp: eq_imp_def marked_insertions_def obj_at_def fun_eq_iff split: mem_write_action.splits) apply (erule tso_pending_mutate_cong) apply (clarsimp split: option.splits obj_at_splits) apply (rename_tac s s' opt x) apply (drule_tac x=x in spec) apply auto done lemmas marked_insertions_fun_upd[simp] = eq_imp_fun_upd[OF mut_m.marked_insertions_eq_imp, simplified eq_imp_simps, rule_format] lemma marked_insertionD[elim!]: "\ sys_mem_write_buffers (mutator m) s = mw_Mutate r f (Some r') # ws; mut_m.marked_insertions m s \ \ marked r' s" by (auto simp: mut_m.marked_insertions_def) lemma (in mut_m) marked_insertions_store_buffer_empty[intro]: "tso_pending_mutate (mutator m) s = [] \ marked_insertions s" by (auto simp: marked_insertions_def filter_empty_conv split: mem_write_action.splits option.splits) lemma (in mut_m) marked_insertions_blacken[simp]: "marked_insertions (s(gc := (s gc)\ W := gc_W s - {r} \)) \ marked_insertions s" by (clarsimp simp: marked_insertions_def split: mem_write_action.splits option.splits) lemma (in mut_m) marked_insertions_alloc[simp]: "\ heap (s sys) r' = None; valid_refs_inv s \ \ marked_insertions (s(mutator m' := s (mutator m')\roots := roots'\, sys := s sys\heap := sys_heap s(r' \ obj')\)) \ marked_insertions s" apply (clarsimp simp: mut_m.marked_insertions_def split: mem_write_action.splits option.splits) apply (rule iffI) apply clarsimp apply (rename_tac ref field x) apply (drule_tac x=ref in spec, drule_tac x=field in spec, drule_tac x=x in spec, clarsimp) apply (drule valid_refs_invD(6)[where x=r' and y=r'], simp_all) done lemma marked_insertions_free[simp]: "\ mut_m.marked_insertions m s; white r s \ \ mut_m.marked_insertions m (s(sys := (s sys)\heap := (heap (s sys))(r := None)\))" by (fastforce simp: mut_m.marked_insertions_def split: mem_write_action.splits obj_at_splits option.splits) lemma (in mut_m) marked_deletions_eq_imp: "eq_imp (\(_::unit). sys_fM \<^bold>\ sys_heap \<^bold>\ tso_pending_mutate (mutator m)) marked_deletions" apply (clarsimp simp: eq_imp_def fun_eq_iff marked_deletions_def obj_at_field_on_heap_def) apply (rule iffI) apply (clarsimp split: mem_write_action.splits) apply (rename_tac s s' ref field option) apply (drule_tac x="mw_Mutate ref field option" in spec) apply (drule mp) apply (metis (lifting, full_types) filter_set member_filter) apply clarsimp apply (subst eq_impD[OF obj_at_eq_imp]) prefer 2 apply (fastforce cong: option.case_cong) apply clarsimp (* opposite dir *) apply (clarsimp split: mem_write_action.splits) apply (rename_tac s s' ref field option) apply (drule_tac x="mw_Mutate ref field option" in spec) apply (drule mp) apply (metis (lifting, full_types) filter_set member_filter) apply clarsimp apply (subst eq_impD[OF obj_at_eq_imp]) prefer 2 apply (fastforce cong: option.case_cong) apply clarsimp done lemmas marked_deletions_fun_upd[simp] = eq_imp_fun_upd[OF mut_m.marked_deletions_eq_imp, simplified eq_imp_simps, rule_format] lemma (in mut_m) marked_deletions_store_buffer_empty[intro]: "tso_pending_mutate (mutator m) s = [] \ marked_deletions s" by (auto simp: marked_deletions_def filter_empty_conv split: mem_write_action.splits) lemma (in mut_m) marked_deletions_alloc[simp]: "\ marked_deletions s; heap (s sys) r' = None; valid_refs_inv s \ \ marked_deletions (s(mutator m' := s (mutator m')\roots := roots'\, sys := s sys\heap := sys_heap s(r' \ obj')\))" apply (clarsimp simp: marked_deletions_def split: mem_write_action.splits) apply (rename_tac ref field option) apply (drule_tac x="mw_Mutate ref field option" in spec) apply clarsimp apply (case_tac "ref = r'") apply (auto simp: obj_at_field_on_heap_def split: option.splits) done (*>*) subsection\ Mutator proofs \ lemma (in mut_m) reachable_snapshot_inv_hs_get_roots_done[simp]: assumes sti: "strong_tricolour_inv s" assumes m: "\r \ mut_roots s. marked r s" assumes ghr: "mut_ghost_honorary_root s = {}" assumes t: "tso_pending_mutate (mutator m) s = []" assumes vri: "valid_refs_inv s" shows "reachable_snapshot_inv (s(mutator m := s (mutator m)\W := {}, ghost_handshake_phase := ghp'\, sys := s sys\handshake_pending := hp', W := sys_W s \ mut_W s, ghost_handshake_in_sync := in'\))" (is "reachable_snapshot_inv ?s'") proof(rule, clarsimp) fix r assume "reachable r s" then show "in_snapshot r ?s'" proof (induct rule: reachable_induct) case (root x) with m show ?case apply (clarsimp simp: in_snapshot_def simp del: fun_upd_apply) (* FIXME intro rules *) apply (auto dest: marked_imp_black_or_grey) done next case (ghost_honorary_root x) with ghr show ?case by simp next case (tso_root x) with t show ?case apply (clarsimp simp: filter_empty_conv tso_write_refs_def) apply (force split: mem_write_action.splits) done next case (reaches x y) from reaches vri have "valid_ref x s" "valid_ref y s" by auto with reaches sti vri show ?case apply (clarsimp simp: in_snapshot_def simp del: fun_upd_apply) apply (elim disjE) apply (clarsimp simp: strong_tricolour_inv_def) apply (drule spec[where x=x]) apply clarsimp apply (auto dest!: marked_imp_black_or_grey)[1] apply (cases "white y s") apply (auto dest: grey_protects_whiteE dest!: marked_imp_black_or_grey) done qed qed lemma (in mut_m) reachable_snapshot_inv_hs_get_work_done[simp]: "reachable_snapshot_inv s \ reachable_snapshot_inv (s(mutator m := s (mutator m)\W := {}\, sys := s sys\handshake_pending := (handshake_pending (s sys))(m := False), W := sys_W s \ mut_W s, ghost_handshake_in_sync := (ghost_handshake_in_sync (s sys))(m := True)\))" by (simp add: reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def) lemma (in mut_m) reachable_write_enqueue[simp]: "reachable r (s(sys := s sys\mem_write_buffers := (mem_write_buffers (s sys))(mutator m := sys_mem_write_buffers (mutator m) s @ [w])\)) \ reachable r s \ (\x. x \ write_refs w \ (x reaches r) s)" by (auto simp: reachable_def tso_write_refs_def) lemma no_black_refs_mo_co_mark[simp]: "\ ghost_honorary_grey (s p) = {}; white r s \ \ no_black_refs (s(p := s p\ghost_honorary_grey := {r}\)) \ no_black_refs s" by (auto simp: no_black_refs_def) lemma (in mut_m) reachable_snapshot_inv_mo_co_mark[simp]: "\ ghost_honorary_grey (s p) = {}; reachable_snapshot_inv s \ \ reachable_snapshot_inv (s(p := s p\ ghost_honorary_grey := {r} \))" by (auto simp: in_snapshot_def reachable_snapshot_inv_def) lemma (in mut_m) no_black_refs_alloc[simp]: "\ heap (s sys) r' = None; no_black_refs s \ \ no_black_refs (s(mutator m' := s (mutator m')\roots := roots'\, sys := s sys\heap := sys_heap s(r' \ \obj_mark = fl, obj_fields = Map.empty\)\)) \ fl \ sys_fM s \ grey r' s" by (auto simp: no_black_refs_def) lemma (in mut_m) reachable_snapshot_inv_load[simp]: "\ reachable_snapshot_inv s; sys_read (mutator m) (mr_Ref r f) (s sys) = mv_Ref r'; r \ mut_roots s \ \ reachable_snapshot_inv (s(mutator m := s (mutator m)\ roots := mut_roots s \ Option.set_option r' \))" by (simp add: reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def) lemma (in mut_m) reachable_snapshot_inv_store_ins[simp]: "\ reachable_snapshot_inv s; r \ mut_roots s; (\r'. opt_r' = Some r') \ the opt_r' \ mut_roots s \ \ reachable_snapshot_inv (s(mutator m := s (mutator m)\ghost_honorary_root := {}\, sys := s sys\ mem_write_buffers := (mem_write_buffers (s sys))(mutator m := sys_mem_write_buffers (mutator m) s @ [mw_Mutate r f opt_r']) \))" apply (clarsimp simp: reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def) apply (drule_tac x=x in spec) apply (auto simp: reachable_def) done lemma (in mut_m) marked_insertions_store_ins[simp]: "\ marked_insertions s; (\r'. opt_r' = Some r') \ marked (the opt_r') s \ \ marked_insertions (s(mutator m := s (mutator m)\ghost_honorary_root := {}\, sys := s sys \mem_write_buffers := (mem_write_buffers (s sys))(mutator m := sys_mem_write_buffers (mutator m) s @ [mw_Mutate r f opt_r'])\))" by (auto simp: marked_insertions_def split: mem_write_action.splits option.splits) lemma (in mut_m) marked_deletions_store_ins[simp]: "\ marked_deletions s; obj_at_field_on_heap (\r'. marked r' s) r f s \ \ marked_deletions (s(mutator m := s (mutator m)\ghost_honorary_root := {}\, sys := s sys \mem_write_buffers := (mem_write_buffers (s sys))(mutator m := sys_mem_write_buffers (mutator m) s @ [mw_Mutate r f opt_r'])\))" by (auto simp: marked_deletions_def split: mem_write_action.splits option.splits) lemma (in mut_m) reachable_snapshot_inv_deref_del[simp]: "\ reachable_snapshot_inv s; sys_read (mutator m) (mr_Ref r f) (s sys) = mv_Ref opt_r'; r \ mut_roots s; mut_ghost_honorary_root s = {} \ \ reachable_snapshot_inv (s(mutator m := s (mutator m)\ghost_honorary_root := Option.set_option opt_r', ref := opt_r'\))" by (clarsimp simp: reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def) lemma (in mut_m) mutator_phase_inv[intro]: "\ handshake_invL \<^bold>\ mark_object_invL \<^bold>\ mut_get_roots.mark_object_invL m \<^bold>\ mut_store_del.mark_object_invL m \<^bold>\ mut_store_ins.mark_object_invL m \<^bold>\ LSTP (fA_rel_inv \<^bold>\ fM_rel_inv \<^bold>\ handshake_phase_inv \<^bold>\ mutators_phase_inv \<^bold>\ phase_rel_inv \<^bold>\ strong_tricolour_inv \<^bold>\ sys_phase_inv \<^bold>\ valid_refs_inv \<^bold>\ valid_W_inv) \ mutator m \ LSTP mutator_phase_inv \" apply (vcg_jackhammer dest!: handshake_phase_invD simp: fA_rel_inv_def fM_rel_inv_def; simp add: mutator_phase_inv_aux_case split: handshake_phase.splits if_splits) subgoal apply (drule_tac x=m in spec) apply (clarsimp simp: fM_rel_def hp_step_rel_def) apply (intro conjI impI; simp) apply (elim disjE; force simp: fA_rel_def) (* FIXME annoying: unfolding fA_rel early leads to non-termination *) apply (rule reachable_snapshot_inv_alloc, simp_all) apply (elim disjE; force simp: fA_rel_def) (* FIXME annoying: unfolding fA_rel early leads to non-termination *) done subgoal for s s' apply (drule_tac x=m in spec) apply (intro conjI impI) apply clarsimp apply (rule marked_deletions_store_ins, assumption) (* FIXME shuffle the following into this lemma *) apply (cases "(\opt_r'. mw_Mutate (mut_tmp_ref s\) (mut_field s\) opt_r' \ set (sys_mem_write_buffers (mutator m) s\))") apply force apply (force simp: marked_deletions_def) apply clarsimp apply (erule marked_insertions_store_ins) apply (drule phase_rel_invD) apply (clarsimp simp: phase_rel_def hp_step_rel_def; elim disjE; fastforce dest: reachable_blackD elim: blackD) apply clarsimp apply (rule marked_deletions_store_ins, assumption) (* FIXME as above *) apply clarsimp apply (erule disjE) apply (drule phase_rel_invD) apply (clarsimp simp: phase_rel_def) apply (elim disjE, simp_all)[1] apply (clarsimp simp: hp_step_rel_def) apply (clarsimp simp: hp_step_rel_def) apply (case_tac "sys_ghost_handshake_phase s\", simp_all)[1] (* FIXME invert handshake_phase_rel *) apply (clarsimp simp: fA_rel_def fM_rel_def) apply (elim disjE, simp_all)[1] apply (clarsimp simp: obj_at_field_on_heap_def split: option.splits) apply (rule conjI) apply fast apply clarsimp apply (frule_tac r=x2a in blackD(1)[OF reachable_blackD], simp_all)[1] apply (rule_tac x="mut_tmp_ref s\" in reachableE; auto simp: ran_def split: obj_at_splits; fail) apply (clarsimp simp: obj_at_field_on_heap_def split: option.splits) apply (rule conjI) apply fast apply clarsimp apply (frule_tac r=x2a in blackD(1)[OF reachable_blackD], simp_all)[1] apply (rule_tac x="mut_tmp_ref s\" in reachableE; auto simp: ran_def split: obj_at_splits; fail) apply (force simp: marked_deletions_def) done (* hs_noop_done *) subgoal for s s' apply (drule_tac x=m in spec) apply (simp add: fA_rel_def fM_rel_def hp_step_rel_def) apply (cases "mut_ghost_handshake_phase s\", simp_all)[1] (* FIXME invert handshake_step *) apply (erule marked_insertions_store_buffer_empty) (* FIXME simp? *) apply (erule marked_deletions_store_buffer_empty) (* FIXME simp? *) done (* hs_get_roots_done *) subgoal apply (drule_tac x=m in spec) apply (simp add: fA_rel_def fM_rel_def hp_step_rel_def) done (* hs_get_work_done *) subgoal apply (drule_tac x=m in spec) apply (simp add: fA_rel_def fM_rel_def hp_step_rel_def) done done lemma (in mut_m') mutator_phase_inv[intro]: notes mut_m.mark_object_invL_def[inv] notes mut_m.handshake_invL_def[inv] shows "\ handshake_invL \<^bold>\ mut_m.handshake_invL m' \<^bold>\ mut_m.mark_object_invL m' \<^bold>\ mut_get_roots.mark_object_invL m' \<^bold>\ mut_store_del.mark_object_invL m' \<^bold>\ mut_store_ins.mark_object_invL m' \<^bold>\ LSTP (fA_rel_inv \<^bold>\ fM_rel_inv \<^bold>\ handshake_phase_inv \<^bold>\ mutators_phase_inv \<^bold>\ valid_refs_inv) \ mutator m' \ LSTP mutator_phase_inv \" apply (vcg_jackhammer simp: fA_rel_inv_def fM_rel_inv_def dest!: handshake_phase_invD) apply (simp_all add: mutator_phase_inv_aux_case split: handshake_phase.splits) apply (drule spec[where x=m]) apply (intro conjI impI) apply (clarsimp simp: fA_rel_def fM_rel_def hp_step_rel_def) apply (elim disjE, auto)[1] apply (rule reachable_snapshot_inv_alloc, simp_all)[1] apply (clarsimp simp: fA_rel_def fM_rel_def hp_step_rel_def) apply (elim disjE, auto)[1] apply (drule spec[where x=m]) apply (clarsimp simp: no_black_refs_def) apply (clarsimp simp: reachable_snapshot_inv_def in_snapshot_def) apply (drule spec[where x=m]) apply (clarsimp simp: no_black_refs_def reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def) done lemma (in sys) grey_protects_white_dequeue_mark[simp]: assumes fl: "fl = sys_fM s" assumes "r \ ghost_honorary_grey (s p)" shows "(\g. (g grey_protects_white w) (s(sys := s sys\heap := (sys_heap s)(r := Option.map_option (obj_mark_update (\_. fl)) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\))) \ (\g. (g grey_protects_white w) s)" (is "(\g. (g grey_protects_white w) ?s') \ ?rhs") proof assume "\g. (g grey_protects_white w) ?s'" then obtain g where "(g grey_protects_white w) ?s'" by blast with assms show ?rhs apply (clarsimp simp: grey_protects_white_def has_white_path_to_def) apply (rotate_tac -1) apply (induct rule: rtranclp.induct) apply fastforce apply clarsimp apply (rename_tac a b c g) apply (case_tac "white c s") apply (rule_tac x=g in exI) apply clarsimp apply (erule rtranclp.intros) apply clarsimp apply (auto split: obj_at_splits if_splits) done next assume ?rhs then obtain g' where "(g' grey_protects_white w) s" by blast with assms show "\g. (g grey_protects_white w) ?s'" apply (clarsimp simp: grey_protects_white_def has_white_path_to_def) apply (rotate_tac -1) apply (induct rule: rtranclp.induct) apply (fastforce simp: grey_def) apply clarsimp apply (rename_tac a b c g) apply (case_tac "c = r") apply (clarsimp simp: grey_def) apply blast apply (rule_tac x=g in exI) apply clarsimp apply (erule rtranclp.intros) apply clarsimp done qed lemma (in sys) reachable_snapshot_inv_dequeue_mark[simp]: "\ sys_mem_write_buffers p s = mw_Mark r fl # ws; mut_m.reachable_snapshot_inv m s; valid_W_inv s \ \ mut_m.reachable_snapshot_inv m (s(sys := s sys\heap := (sys_heap s)(r := Option.map_option (obj_mark_update (\_. fl)) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\))" apply (clarsimp simp: mut_m.reachable_snapshot_inv_def in_snapshot_def) apply (rename_tac x) apply (subst (asm) reachable_fun_upd, simp_all)[1] apply auto[1] apply (drule_tac x=x in spec) apply clarsimp apply (subst (asm) arg_cong[where f=Not, OF grey_protects_white_dequeue_mark, simplified], simp_all) apply blast apply blast apply blast done lemma (in sys) marked_insertions_dequeue_mark[simp]: "\ sys_mem_write_buffers p s = mw_Mark r fl # ws; mut_m.marked_insertions m s; tso_writes_inv s; valid_W_inv s \ \ mut_m.marked_insertions m (s(sys := s sys\heap := (sys_heap s)(r := Option.map_option (obj_mark_update (\_. fl)) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\))" apply (clarsimp simp: mut_m.marked_insertions_def) apply (cases "mutator m = p") apply clarsimp apply (rename_tac x) apply (drule_tac x=x in spec) apply (auto split: mem_write_action.splits option.splits obj_at_splits)[1] apply clarsimp apply (rename_tac x) apply (drule_tac x=x in spec) apply (auto split: mem_write_action.splits option.splits obj_at_splits)[1] done lemma (in sys) marked_insertions_dequeue_ref[simp]: "\ sys_mem_write_buffers p s = mw_Mutate r f r' # ws; mut_m.marked_insertions m s \ \ mut_m.marked_insertions m (s(sys := s sys\heap := (sys_heap s)(r := Option.map_option (\obj. obj\obj_fields := (obj_fields obj)(f := r')\) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\))" apply (clarsimp simp: mut_m.marked_insertions_def) apply (cases "mutator m = p") apply clarsimp apply (rename_tac x) apply (drule_tac x=x in spec) apply (auto split: mem_write_action.splits option.splits obj_at_splits)[1] apply clarsimp apply (rename_tac x) apply (drule_tac x=x in spec) apply (auto split: mem_write_action.splits option.splits obj_at_splits)[1] done (* redundant to fit with other rules. Perhaps want points_to with explicit witness for f? *) lemma points_to_mw_Mutate: "(x points_to y) (s(sys := (s sys)\ heap := (sys_heap s)(r := Option.map_option (\obj :: ('field, 'ref) object. obj\obj_fields := (obj_fields obj)(f := opt_r')\) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(p := ws) \)) \ (r \ x \ (x points_to y) s) \ (r = x \ valid_ref r s \ (opt_r' = Some y \ ( (x points_to y) s \ obj_at (\obj. \f'. obj_fields obj f' = Some y \ f \ f') r s)))" by (auto simp: ran_def split: obj_at_splits) (* shows the snapshot is preserved by the two marks. *) lemma (in sys) grey_protects_white_dequeue_ref[simp]: assumes sb: "sys_mem_write_buffers (mutator m) s = mw_Mutate r f opt_r' # ws" assumes mi: "mut_m.marked_insertions m s" assumes md: "mut_m.marked_deletions m s" notes map_option.compositionality[simp] o_def[simp] shows "(\g. (g grey_protects_white w) (s(sys := s sys\heap := (sys_heap s)(r := Option.map_option (\obj. obj\obj_fields := (obj_fields obj)(f := opt_r')\) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(mutator m := ws)\))) \ (\g. (g grey_protects_white w) s)" (is "(\g. (g grey_protects_white w) ?s') \ ?rhs") proof assume "(\g. (g grey_protects_white w) ?s')" then obtain g where "(g grey_protects_white w) ?s'" by blast with mi sb show ?rhs apply (clarsimp simp: grey_protects_white_def has_white_path_to_def) apply (rotate_tac -1) apply (induct rule: rtranclp.induct) apply fastforce apply (auto simp: points_to_mw_Mutate elim: rtranclp.intros(2)) apply (rename_tac a c g) apply (clarsimp simp: mut_m.marked_insertions_def) (* FIXME rule *) apply (drule_tac x="mw_Mutate r f (Some c)" in spec) apply (simp split: obj_at_splits) done next assume ?rhs then show "(\g. (g grey_protects_white w) ?s')" proof(clarsimp) fix g assume "(g grey_protects_white w) s" with md sb show ?thesis apply (clarsimp simp: grey_protects_white_def has_white_path_to_def) apply (rotate_tac -1) apply (induct rule: rtranclp.induct) apply fastforce apply clarsimp apply (rename_tac a b c g) apply (case_tac "b = r") defer apply (auto simp: points_to_mw_Mutate elim: rtranclp.intros(2))[1] apply clarsimp apply (subst (asm) (3) obj_at_def) (* FIXME rule: witness field for r points_to c *) apply (clarsimp simp: ran_def split: option.splits) apply (rename_tac a c g x2 aa) apply (case_tac "aa = f") defer apply (rule_tac x=g in exI) apply clarsimp apply (erule rtranclp.intros) apply (auto split: obj_at_splits)[1] apply clarsimp apply (clarsimp simp: mut_m.marked_deletions_def) (* FIXME rule *) apply (drule spec[where x="mw_Mutate r f opt_r'"]) apply (clarsimp simp: obj_at_field_on_heap_def) apply (simp split: obj_at_splits) done qed qed (* write barrier installed but not all mutators are necessarily past get_roots *) lemma (in sys) reachable_snapshot_inv_dequeue_ref[simp]: fixes s :: "('field, 'mut, 'ref) lsts" assumes sb: "sys_mem_write_buffers (mutator m') s = mw_Mutate r f opt_r' # ws" assumes mi: "mut_m.marked_insertions m' s" assumes md: "mut_m.marked_deletions m' s" assumes rsi: "mut_m.reachable_snapshot_inv m s" assumes sti: "strong_tricolour_inv s" assumes vri: "valid_refs_inv s" notes map_option.compositionality[simp] o_def[simp] shows "mut_m.reachable_snapshot_inv m (s(sys := s sys\heap := (sys_heap s)(r := Option.map_option (\obj. obj\obj_fields := (obj_fields obj)(f := opt_r')\) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(mutator m' := ws)\))" (is "mut_m.reachable_snapshot_inv m ?s'") proof(rule mut_m.reachable_snapshot_invI) fix y assume y: "mut_m.reachable m y ?s'" then have "(mut_m.reachable m y s \ mut_m.reachable m' y s) \ in_snapshot y ?s'" proof(induct rule: reachable_induct) case (root x) with mi md rsi sb show ?case apply (clarsimp simp: mut_m.reachable_snapshot_inv_def in_snapshot_def simp del: fun_upd_apply) apply auto done next case (ghost_honorary_root x) with mi md rsi sb show ?case apply (clarsimp simp: mut_m.reachable_snapshot_inv_def in_snapshot_def simp del: fun_upd_apply) apply auto done next case (tso_root x) with mi md rsi sb show ?case apply (clarsimp simp: mut_m.reachable_snapshot_inv_def in_snapshot_def simp del: fun_upd_apply) apply (clarsimp split: if_splits) apply (rename_tac xa) apply (case_tac xa, simp_all)[1] apply (rename_tac ref field option) apply (clarsimp simp: mut_m.marked_deletions_def mut_m.marked_insertions_def) apply (drule_tac x="mw_Mutate ref field option" in spec) apply (drule_tac x="mw_Mutate ref field option" in spec) apply clarsimp apply (frule spec[where x=x]) apply (subgoal_tac "mut_m.reachable m x s") apply force apply (rule reachableI(2)) apply (force simp: mut_m.tso_write_refs_def) apply auto done next case (reaches x y) from reaches sb have y: "mut_m.reachable m y s \ mut_m.reachable m' y s" apply (clarsimp simp: points_to_mw_Mutate mut_m.reachable_snapshot_inv_def in_snapshot_def) apply (elim disjE, (force dest!: reachableE mutator_reachable_tso)+)[1] done moreover from y vri have "valid_ref y s" by auto with reaches mi md rsi sb sti y have "(black y s \ (\x. (x grey_protects_white y) s))" apply (clarsimp simp: mut_m.reachable_snapshot_inv_def in_snapshot_def simp del: fun_upd_apply) apply clarsimp apply (drule spec[where x=y]) apply (clarsimp simp: points_to_mw_Mutate mut_m.marked_insertions_def mut_m.marked_deletions_def) (* FIXME lemmas *) apply (drule spec[where x="mw_Mutate r f opt_r'"])+ apply clarsimp apply (elim disjE, simp_all) (* FIXME probably want points_to_mw_Mutate as an elim rule to make this robust, reduce duplication *) apply (force dest!: reachableE) apply (force dest!: reachableE) apply clarsimp apply (drule (3) strong_tricolour_invD) apply (metis (no_types) grey_protects_whiteI marked_imp_black_or_grey(1)) apply clarsimp apply (cases "white y s") (* FIXME lemma *) apply (drule (2) grey_protects_whiteE) apply blast apply (force simp: black_def split: obj_at_splits) apply clarsimp apply (elim disjE, simp_all) apply (force simp: black_def) apply clarsimp apply (drule (3) strong_tricolour_invD) apply (force simp: black_def) apply clarsimp apply (elim disjE, simp_all) apply (force simp: black_def) apply clarsimp apply (cases "white y s") (* FIXME lemma *) apply (drule (2) grey_protects_whiteE) apply blast apply (force simp: black_def split: obj_at_splits) apply clarsimp apply (elim disjE, simp_all) apply (force simp: black_def) apply clarsimp apply (drule (3) strong_tricolour_invD) apply (force simp: black_def) apply clarsimp apply (elim disjE, simp_all) apply (force simp: black_def) apply clarsimp apply (cases "white y s") (* FIXME lemma *) apply (drule (2) grey_protects_whiteE) apply blast apply (force simp: black_def split: obj_at_splits) done moreover note mi md rsi sb ultimately show ?case apply (clarsimp simp: mut_m.reachable_snapshot_inv_def in_snapshot_def simp del: fun_upd_apply) apply clarsimp done qed then show "in_snapshot y ?s'" by blast qed lemma valid_refs_invI: "\ \m x y. \ (x reaches y) s; x \ mut_m.mut_roots m s \ x \ mut_m.mut_ghost_honorary_root m s \ x \ mut_m.tso_write_refs m s \ grey x s \ \ valid_ref y s \ \ valid_refs_inv s" by (auto simp: valid_refs_inv_def mut_m.reachable_def grey_reachable_def) lemma black_heap_reachable: assumes "mut_m.reachable m y s" assumes bh: "black_heap s" assumes vri: "valid_refs_inv s" shows "black y s" using assms apply (induct rule: reachable_induct) apply (simp_all add: black_heap_def valid_refs_invD) apply (metis obj_at_weakenE reachableE valid_refs_inv_def) done lemma valid_refs_inv_dequeue_ref[simp]: notes map_option.compositionality[simp] o_def[simp] fixes s :: "('field, 'mut, 'ref) lsts" assumes vri: "valid_refs_inv s" assumes sb: "sys_mem_write_buffers (mutator m') s = mw_Mutate r f opt_r' # ws" shows "valid_refs_inv (s(sys := s sys\heap := (sys_heap s)(r := Option.map_option (\obj. obj\obj_fields := (obj_fields obj)(f := opt_r')\) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(mutator m' := ws)\))" (is "valid_refs_inv ?s'") proof(rule valid_refs_invI) fix m let ?root = "\m x (s::('field, 'mut, 'ref) lsts). x \ mut_m.mut_roots m s \ x \ mut_m.mut_ghost_honorary_root m s \ x \ mut_m.tso_write_refs m s \ grey x s" fix x y assume xy: "(x reaches y) ?s'" and x: "?root m x ?s'" from xy have "(\m x. ?root m x s \ (x reaches y) s) \ valid_ref y ?s'" proof induct case base with x sb vri show ?case apply - apply (subst obj_at_fun_upd) apply (auto simp: mut_m.tso_write_refs_def split: if_splits intro: valid_refs_invD(5)[where m=m]) apply (metis list.set_intros(2) rtranclp.rtrancl_refl) done (* FIXME rules *) next case (step y z) with sb vri show ?case apply - apply (subst obj_at_fun_upd, clarsimp) apply (subst (asm) obj_at_fun_upd, fastforce) apply (clarsimp simp: points_to_mw_Mutate simp del: fun_upd_apply) apply (fastforce elim: rtranclp.intros(2) simp: mut_m.tso_write_refs_def intro: exI[where x=m'] valid_refs_invD(5)[where m=m']) done qed then show "valid_ref y ?s'" by blast qed declare map_option.compositionality[simp] o_def[simp] lemma (in sys) reachable_snapshot_inv_black_heap_no_grey_refs_dequeue_ref[simp]: assumes sb: "sys_mem_write_buffers (mutator m') s = mw_Mutate r f opt_r' # ws" assumes bh: "black_heap s" assumes ngr: "no_grey_refs s" assumes vri: "valid_refs_inv s" shows "mut_m.reachable_snapshot_inv m (s(sys := s sys\heap := (sys_heap s)(r := Option.map_option (\obj. obj\obj_fields := (obj_fields obj)(f := opt_r')\) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(mutator m' := ws)\))" (is "mut_m.reachable_snapshot_inv m ?s'") apply (rule mut_m.reachable_snapshot_invI) apply (rule in_snapshotI) apply (erule black_heap_reachable) using bh vri apply (simp add: black_heap_def) using bh ngr sb vri apply (subst valid_refs_inv_def) apply clarsimp apply (simp add: no_grey_refs_def grey_reachable_def) apply clarsimp apply (drule black_heap_reachable) apply (simp add: black_heap_def) apply clarsimp apply auto done lemma (in sys) marked_deletions_dequeue_mark[simp]: "\ sys_mem_write_buffers p s = mw_Mark r fl # ws; mut_m.marked_deletions m s; tso_writes_inv s; valid_W_inv s \ \ mut_m.marked_deletions m (s(sys := s sys\heap := (sys_heap s)(r := Option.map_option (obj_mark_update (\_. fl)) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\))" apply (clarsimp simp: mut_m.marked_deletions_def) apply (cases "mutator m = p") apply clarsimp apply (rename_tac x) apply (drule_tac x=x in spec) apply (clarsimp split: mem_write_action.splits) apply (simp add: obj_at_field_on_heap_def cong: option.case_cong) apply (auto split: option.splits)[1] apply clarsimp apply (rename_tac x) apply (drule_tac x=x in spec) apply (clarsimp split: mem_write_action.splits) apply (simp add: obj_at_field_on_heap_def cong: option.case_cong) apply (auto split: option.splits)[1] done lemma (in sys) marked_deletions_dequeue_ref[simp]: "\ sys_mem_write_buffers (mutator m') s = mw_Mutate r f opt_r' # ws; mut_m.marked_deletions m s; mut_m.marked_insertions m' s \ \ mut_m.marked_deletions m (s(sys := s sys\heap := (sys_heap s)(r := Option.map_option (\obj. obj\obj_fields := (obj_fields obj)(f := opt_r')\) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))((mutator m') := ws)\))" supply [[simproc del: defined_all]] apply (clarsimp simp: mut_m.marked_deletions_def) apply (cases "m = m'") apply clarsimp apply (rename_tac x) apply (drule_tac x=x in spec) apply (fastforce simp: obj_at_field_on_heap_def mut_m.marked_insertions_def split: mem_write_action.splits obj_at_splits option.splits) apply clarsimp apply (rename_tac x) apply (drule_tac x=x in spec) apply (fastforce simp: obj_at_field_on_heap_def mut_m.marked_insertions_def split: mem_write_action.splits obj_at_splits option.splits) done lemma (in sys) black_heap_marked_insertions_dequeue[simp]: "\ black_heap s; valid_refs_inv s \ \ mut_m.marked_insertions m s" by (auto simp: mut_m.marked_insertions_def black_heap_def black_def split: mem_write_action.splits option.splits dest: valid_refs_invD) lemma (in sys) mutator_phase_inv[intro]: "\ LSTP (fA_rel_inv \<^bold>\ fM_rel_inv \<^bold>\ handshake_phase_inv \<^bold>\ mutators_phase_inv \<^bold>\ strong_tricolour_inv \<^bold>\ sys_phase_inv \<^bold>\ tso_writes_inv \<^bold>\ valid_refs_inv \<^bold>\ valid_W_inv) \ sys \ LSTP (mut_m.mutator_phase_inv m) \" apply vcg_nihe apply vcg_ni apply (clarsimp simp: fA_rel_inv_def fM_rel_inv_def mutator_phase_inv_aux_case p_not_sys hp_step_rel_def do_write_action_def split: handshake_phase.splits mem_write_action.splits) (* fM *) prefer 2 apply (drule mut_m.handshake_phase_invD[where m=m]) apply (erule disjE) apply (clarsimp simp: fM_rel_def hp_step_rel_def) apply clarsimp (* FIXME mess *) apply (frule mut_m.handshake_phase_invD[where m=m]) apply (intro allI conjI impI) apply (erule disjE) apply auto[1] apply clarsimp apply (rename_tac s s' ws ref field option ma) apply (rule marked_deletions_dequeue_ref, simp_all) apply (drule_tac x=ma in spec) apply (frule_tac m=ma in mut_m.handshake_phase_invD, clarsimp simp: hp_step_rel_def) apply (elim disjE, simp_all)[1] apply (erule disjE) apply auto[1] apply clarsimp apply (rule marked_deletions_dequeue_ref, simp_all) apply (drule_tac x=ma in spec) apply (frule_tac m=ma in mut_m.handshake_phase_invD, clarsimp simp: hp_step_rel_def) apply (elim disjE, simp_all)[1] apply (clarsimp simp: fA_rel_def fM_rel_def if_splits) apply (elim disjE, simp_all)[1] apply (erule disjE, simp) apply (clarsimp simp: hp_step_rel_def) apply (frule_tac m=ma in mut_m.handshake_phase_invD, clarsimp simp: hp_step_rel_def) apply (elim disjE, simp_all split: if_splits)[1] apply (clarsimp simp: fA_rel_def fM_rel_def, blast) done (*>*) subsection\The infamous termination argument.\ text\ We need to know that if the GC does not receive any further work to do at \get_roots\ and \get_work\, then there are no grey objects left. Essentially this encodes the stability property that grey objects must exist for mutators to create grey objects. Note that this is not invariant across the scan: it is possible for the GC to hold all the grey references. The two handshakes transform the GC's local knowledge that it has no more work to do into a global property, or gives it more work. \ definition (in mut_m) gc_W_empty_mut_inv :: "('field, 'mut, 'ref) lsts_pred" where "gc_W_empty_mut_inv = ((EMPTY sys_W \<^bold>\ sys_ghost_handshake_in_sync m \<^bold>\ \<^bold>\(EMPTY (WL (mutator m)))) \<^bold>\ (\<^bold>\m'. \<^bold>\(sys_ghost_handshake_in_sync m') \<^bold>\ \<^bold>\(EMPTY (WL (mutator m')))))" locset_definition (in -) gc_W_empty_locs :: "location set" where "gc_W_empty_locs \ idle_locs \ init_locs \ sweep_locs \ { ''mark_read_fM'', ''mark_write_fA'', ''mark_end'' } \ prefixed ''mark_noop'' \ prefixed ''mark_loop_get_roots'' \ prefixed ''mark_loop_get_work''" locset_definition "black_heap_locs = { ''sweep_idle'', ''idle_noop_mfence'', ''idle_noop_init_type'' }" locset_definition "no_grey_refs_locs = black_heap_locs \ sweep_locs \ {''mark_end''}" inv_definition (in gc) gc_W_empty_invL :: "('field, 'mut, 'ref) gc_pred" where "gc_W_empty_invL = (atS_gc (hs_get_roots_locs \ hs_get_work_locs) (\<^bold>\m. mut_m.gc_W_empty_mut_inv m) \<^bold>\ at_gc ''mark_loop_get_roots_load_W'' (EMPTY sys_W \<^bold>\ no_grey_refs) \<^bold>\ at_gc ''mark_loop_get_work_load_W'' (EMPTY sys_W \<^bold>\ no_grey_refs) \<^bold>\ at_gc ''mark_loop'' (EMPTY gc_W \<^bold>\ no_grey_refs) \<^bold>\ atS_gc no_grey_refs_locs no_grey_refs \<^bold>\ atS_gc gc_W_empty_locs (EMPTY gc_W))" (*<*) lemma (in mut_m) gc_W_empty_mut_inv_eq_imp: "eq_imp (\m'. sys_W \<^bold>\ WL (mutator m') \<^bold>\ sys_ghost_handshake_in_sync m') gc_W_empty_mut_inv" by (simp add: eq_imp_def gc_W_empty_mut_inv_def) lemmas gc_W_empty_mut_inv_fun_upd[simp] = eq_imp_fun_upd[OF mut_m.gc_W_empty_mut_inv_eq_imp, simplified eq_imp_simps, rule_format] lemma (in gc) gc_W_empty_invL_eq_imp: "eq_imp (\(m', p) s. (AT s gc, s\ gc, sys_W s\, WL p s\, sys_ghost_handshake_in_sync m' s\)) gc_W_empty_invL" by (simp add: eq_imp_def gc_W_empty_invL_def mut_m.gc_W_empty_mut_inv_def no_grey_refs_def grey_def) lemmas gc_W_empty_invL_niE[nie] = iffD1[OF gc.gc_W_empty_invL_eq_imp[simplified eq_imp_simps, rule_format, unfolded conj_explode, rule_format], rotated -1] lemma get_roots_get_work_subseteq_gc_W_empty_locs: "hs_get_roots_locs \ hs_get_work_locs \ gc_W_empty_locs" by (auto simp: hs_get_roots_locs_def hs_get_work_locs_def gc_W_empty_locs_def) lemma (in gc) gc_W_empty_mut_inv_handshake_init[iff]: "mut_m.gc_W_empty_mut_inv m (s(sys := s sys\handshake_type := ht, ghost_handshake_in_sync := \False\\))" "mut_m.gc_W_empty_mut_inv m (s(sys := s sys\handshake_type := ht, ghost_handshake_in_sync := \False\, ghost_handshake_phase := hp' \))" by (simp_all add: mut_m.gc_W_empty_mut_inv_def) lemma gc_W_empty_mut_inv_load_W: "\ \m. mut_m.gc_W_empty_mut_inv m s; \m. sys_ghost_handshake_in_sync m s; WL gc s = {}; WL sys s = {} \ \ no_grey_refs s" apply (clarsimp simp: mut_m.gc_W_empty_mut_inv_def no_grey_refs_def grey_def) apply (rename_tac x xa) apply (case_tac xa) apply (simp_all add: WL_def) done lemma (in gc) gc_W_empty_invL[intro]: "\ handshake_invL \<^bold>\ obj_fields_marked_invL \<^bold>\ gc_W_empty_invL \<^bold>\ LSTP valid_W_inv \ gc \ gc_W_empty_invL \" apply vcg_jackhammer apply (auto elim: gc_W_empty_mut_inv_load_W simp: WL_def) done lemma (in sys) gc_gc_W_empty_invL[intro]: "\ gc.gc_W_empty_invL \ sys" by vcg_nihe lemma (in gc) handshake_get_rootsD: "\ atS gc hs_get_roots_locs s; handshake_invL s \ \ sys_ghost_handshake_phase s\ = hp_IdleMarkSweep \ sys_handshake_type s\ = ht_GetRoots" apply (simp add: handshake_invL_def) apply (elim conjE) apply (drule mp, erule atS_mono[OF _ hs_get_roots_locs_subseteq_hp_IdleMarkSweep_locs]) apply simp done lemma (in gc) handshake_get_workD: "\ atS gc hs_get_work_locs s; handshake_invL s \ \ sys_ghost_handshake_phase s\ = hp_IdleMarkSweep \ sys_handshake_type s\ = ht_GetWork" apply (simp add: handshake_invL_def) apply (elim conjE) apply (drule mp, erule atS_mono[OF _ hs_get_work_locs_subseteq_hp_IdleMarkSweep_locs]) apply simp done lemma (in gc) handshake_get_roots_get_workD: "\ atS gc (hs_get_roots_locs \ hs_get_work_locs) s; handshake_invL s \ \ sys_ghost_handshake_phase s\ = hp_IdleMarkSweep \ sys_handshake_type s\ \ {ht_GetWork, ht_GetRoots}" apply (simp add: handshake_invL_def) apply (elim conjE) apply (drule mp, rule atS_mono[OF _ iffD2[OF Un_subset_iff, unfolded conj_explode, OF hs_get_roots_locs_subseteq_hp_IdleMarkSweep_locs hs_get_work_locs_subseteq_hp_IdleMarkSweep_locs]], assumption) apply (auto simp: atS_un) done lemma no_grey_refs_locs_subseteq_hs_in_sync_locs: "no_grey_refs_locs \ hs_in_sync_locs" by (auto simp: no_grey_refs_locs_def black_heap_locs_def hs_in_sync_locs_def hs_done_locs_def sweep_locs_def dest: prefix_same_cases) lemma (in mut_m) handshake_sweep_mark_endD: "\ atS gc no_grey_refs_locs s; gc.handshake_invL s; handshake_phase_inv s\ \ \ mut_ghost_handshake_phase s\ = hp_IdleMarkSweep \ All (ghost_handshake_in_sync (s\ sys))" apply (simp add: gc.handshake_invL_def) apply (elim conjE) apply (drule mp, erule atS_mono[OF _ no_grey_refs_locs_subseteq_hs_in_sync_locs]) apply (drule handshake_phase_invD) apply (simp only: no_grey_refs_locs_def cong del: atS_state_cong) apply (clarsimp simp: atS_un) apply (elim disjE) apply (drule mp, erule atS_mono[where ls'="hp_IdleMarkSweep_locs"]) apply (clarsimp simp: black_heap_locs_def loc) apply (clarsimp simp: hp_step_rel_def) apply blast apply (drule mp, erule atS_mono[where ls'="hp_IdleMarkSweep_locs"]) apply (clarsimp simp: hp_IdleMarkSweep_locs_def hp_step_rel_def) apply (clarsimp simp: hp_step_rel_def) apply blast apply (clarsimp simp: atS_simps loc hp_step_rel_def) apply blast done lemma empty_WL_GC: "\ atS gc (hs_get_roots_locs \ hs_get_work_locs) s; gc.obj_fields_marked_invL s \ \ gc_ghost_honorary_grey s\ = {}" by (clarsimp simp: gc.obj_fields_marked_invL_def dest!: atS_mono[OF _ get_roots_get_work_subseteq_ghost_honorary_grey_empty_locs]) (* think about showing gc_W_empty_invL instead *) lemma (in mut_m) gc_W_empty_mut_mo_co_mark[simp]: "\ \x. mut_m.gc_W_empty_mut_inv x s\; mutators_phase_inv s\; mut_ghost_honorary_grey s\ = {}; r \ mut_roots s\ \ mut_ghost_honorary_root s\; white r s\; atS gc (hs_get_roots_locs \ hs_get_work_locs) s; gc.handshake_invL s; gc.obj_fields_marked_invL s; atS gc gc_W_empty_locs s \ gc_W s\ = {}; handshake_phase_inv s\; valid_W_inv s\ \ \ mut_m.gc_W_empty_mut_inv m' (s\(mutator m := s\ (mutator m)\ghost_honorary_grey := {r}\))" apply (frule (1) gc.handshake_get_roots_get_workD) apply (frule handshake_phase_invD) apply (clarsimp simp: hp_step_rel_def simp del: Un_iff) apply (elim disjE, simp_all) (* before get work *) apply (clarsimp simp: mut_m.gc_W_empty_mut_inv_def) apply blast (* past get work *) apply (clarsimp simp: mut_m.gc_W_empty_mut_inv_def) apply (frule spec[where x=m], clarsimp) apply (frule (2) reachable_snapshot_inv_white_root) apply clarsimp apply (drule grey_protects_whiteD) apply (clarsimp simp: grey_def) apply (rename_tac g x) apply (case_tac x, simp_all) (* Can't be the GC *) prefer 2 apply (frule (1) empty_WL_GC) apply (drule mp, erule atS_mono[OF _ get_roots_get_work_subseteq_gc_W_empty_locs]) apply (clarsimp simp: WL_def) (* Can't be sys *) prefer 2 apply (clarsimp simp: WL_def) apply clarsimp apply (rename_tac g mut) apply (case_tac "sys_ghost_handshake_in_sync mut s\") apply (drule mp, rule_tac x=mut in exI, clarsimp) apply blast apply (rule_tac x=mut in exI) apply clarsimp (* before get roots *) apply (clarsimp simp: mut_m.gc_W_empty_mut_inv_def) apply blast (* after get roots *) apply (clarsimp simp: mut_m.gc_W_empty_mut_inv_def) apply (frule spec[where x=m], clarsimp) apply (frule (2) reachable_snapshot_inv_white_root) apply clarsimp apply (drule grey_protects_whiteD) apply (clarsimp simp: grey_def) apply (rename_tac g x) apply (case_tac x, simp_all) (* Can't be the GC *) prefer 2 apply (frule (1) empty_WL_GC) apply (drule mp, erule atS_mono[OF _ get_roots_get_work_subseteq_gc_W_empty_locs]) apply (clarsimp simp: WL_def) (* Can't be sys *) prefer 2 apply (clarsimp simp: WL_def) apply clarsimp apply (rename_tac g mut) apply (case_tac "sys_ghost_handshake_in_sync mut s\") apply (drule mp, rule_tac x=mut in exI, clarsimp) apply blast apply (rule_tac x=mut in exI) apply clarsimp done (* FIXME common up *) lemma (in mut_m) no_grey_refs_mo_co_mark[simp]: "\ mutators_phase_inv s\; no_grey_refs s\; gc.handshake_invL s; at gc ''mark_loop'' s \ at gc ''mark_loop_get_roots_load_W'' s \ at gc ''mark_loop_get_work_load_W'' s \ atS gc no_grey_refs_locs s; r \ mut_roots s\ \ mut_ghost_honorary_root s\; white r s\; handshake_phase_inv s\ \ \ no_grey_refs (s\(mutator m := s\ (mutator m)\ghost_honorary_grey := {r}\))" apply (elim disjE) apply (clarsimp simp: atS_simps gc.handshake_invL_def loc) apply (frule handshake_phase_invD) apply (clarsimp simp: hp_step_rel_def) apply (drule spec[where x=m]) apply (clarsimp simp: conj_disj_distribR[symmetric]) apply (drule (2) no_grey_refs_not_rootD) apply blast apply (clarsimp simp: atS_simps gc.handshake_invL_def loc) apply (frule handshake_phase_invD) apply (clarsimp simp: hp_step_rel_def) apply (drule spec[where x=m]) apply (clarsimp simp: conj_disj_distribR[symmetric]) apply (drule (2) no_grey_refs_not_rootD) apply blast apply (clarsimp simp: atS_simps gc.handshake_invL_def loc) apply (frule handshake_phase_invD) apply (clarsimp simp: hp_step_rel_def) apply (drule spec[where x=m]) apply clarsimp apply (drule (2) no_grey_refs_not_rootD) apply blast apply (frule (2) handshake_sweep_mark_endD) apply (drule spec[where x=m]) apply clarsimp apply (drule (2) no_grey_refs_not_rootD) apply blast done lemma (in mut_m) gc_W_empty_invL[intro]: notes gc.gc_W_empty_invL_def[inv] shows "\ handshake_invL \<^bold>\ mark_object_invL \<^bold>\ tso_lock_invL \<^bold>\ mut_get_roots.mark_object_invL m \<^bold>\ mut_store_del.mark_object_invL m \<^bold>\ mut_store_ins.mark_object_invL m \<^bold>\ gc.handshake_invL \<^bold>\ gc.obj_fields_marked_invL \<^bold>\ gc.gc_W_empty_invL \<^bold>\ LSTP (handshake_phase_inv \<^bold>\ mutators_phase_inv \<^bold>\ valid_W_inv) \ mutator m \ gc.gc_W_empty_invL \" apply vcg_nihe (* apply vcg_ni -- very slow *) apply(tactic \ let val ctxt = @{context} in TRY (HEADGOAL (vcg_clarsimp_tac ctxt)) THEN PARALLEL_ALLGOALS ( vcg_sem_tac ctxt THEN' (TRY o SELECT_GOAL (Local_Defs.unfold_tac ctxt (Named_Theorems.get ctxt @{named_theorems inv}))) THEN' (TRY o REPEAT_ALL_NEW (Tactic.match_tac ctxt @{thms conjI})) (* expose the location predicates, do not split the consequents *) THEN_ALL_NEW (TRY o REPEAT_ALL_NEW (Tactic.match_tac ctxt @{thms impI})) (* Preserve the label sets in atS but normalise the label in at; turn s' into s *) THEN_ALL_NEW full_simp_tac ctxt (* FIXME vcg_ni uses asm_full_simp_tac here *) THEN_ALL_NEW (TRY o REPEAT_ALL_NEW (Tactic.ematch_tac ctxt @{thms conjE})) (* The effect of vcg_pre: should be cheap *) THEN_ALL_NEW (TRY o REPEAT_ALL_NEW (Tactic.ematch_tac ctxt @{thms thin_locs} THEN' REPEAT1 o assume_tac ctxt)) THEN_ALL_NEW asm_full_simp_tac (ss_only (@{thms loc_simps} @ Named_Theorems.get ctxt @{named_theorems loc}) ctxt) THEN_ALL_NEW (TRY o REPEAT_ALL_NEW (Rule_Insts.thin_tac ctxt "True" [])) THEN_ALL_NEW clarsimp_tac ctxt) end \) (* hs_noop_done *) apply (clarsimp simp: atS_un gc.handshake_invL_def) (* hs_get_roots_done: gc_W_empty *) apply (clarsimp simp: mut_m.gc_W_empty_mut_inv_def) apply (rule conjI) apply (clarsimp simp: WL_def) apply clarsimp apply (drule mp) apply blast apply (clarsimp simp: WL_def) apply (rename_tac xa) apply (rule_tac x=xa in exI) apply clarsimp (* hs_get_roots_done: no_grey_refs *) apply (simp add: no_grey_refs_def) apply (simp add: no_grey_refs_def) (* hs_get_work_done: gc_W_empty *) apply (clarsimp simp: mut_m.gc_W_empty_mut_inv_def) apply (rule conjI) apply (clarsimp simp: WL_def) apply clarsimp apply (drule mp) apply blast apply (clarsimp simp: WL_def) apply (rename_tac xa) apply (rule_tac x=xa in exI) apply clarsimp (* hs_get_work_done: no_grey_refs *) apply (simp add: no_grey_refs_def) apply (simp add: no_grey_refs_def) done (*>*) subsection\Sweep loop invariants\ locset_definition "sweep_loop_locs = prefixed ''sweep_loop''" inv_definition (in gc) sweep_loop_invL :: "('field, 'mut, 'ref) gc_pred" where "sweep_loop_invL = (at_gc ''sweep_loop_check'' ( (\<^bold>\(NULL gc_mark) \<^bold>\ (\s. obj_at (\obj. Some (obj_mark obj) = gc_mark s) (gc_tmp_ref s) s)) \<^bold>\ ( NULL gc_mark \<^bold>\ valid_ref \<^bold>$ gc_tmp_ref \<^bold>\ marked \<^bold>$ gc_tmp_ref ) ) \<^bold>\ at_gc ''sweep_loop_free'' ( \<^bold>\(NULL gc_mark) \<^bold>\ the \ gc_mark \<^bold>\ gc_fM \<^bold>\ (\s. obj_at (\obj. Some (obj_mark obj) = gc_mark s) (gc_tmp_ref s) s) ) \<^bold>\ at_gc ''sweep_loop_ref_done'' (valid_ref \<^bold>$ gc_tmp_ref \<^bold>\ marked \<^bold>$ gc_tmp_ref) \<^bold>\ atS_gc sweep_loop_locs (\<^bold>\r. \<^bold>\(\r\ \<^bold>\ gc_refs) \<^bold>\ valid_ref r \<^bold>\ marked r) \<^bold>\ atS_gc black_heap_locs (\<^bold>\r. valid_ref r \<^bold>\ marked r) \<^bold>\ atS_gc (prefixed ''sweep_loop_'' - { ''sweep_loop_choose_ref'' }) (gc_tmp_ref \<^bold>\ gc_refs))" (*<*) lemma (in gc) sweep_loop_invL_eq_imp: "eq_imp (\(_::unit) (s :: ('field, 'mut, 'ref) gc_pred_state). (AT s gc, s\ gc, sys_fM s\, Option.map_option obj_mark \ sys_heap s\)) sweep_loop_invL" apply (clarsimp simp: eq_imp_def inv) apply (rename_tac s s') apply (subgoal_tac "\r. valid_ref r s\ \ valid_ref r s'\") apply (subgoal_tac "\P r. obj_at (\obj. P (obj_mark obj)) r s\ \ obj_at (\obj. P (obj_mark obj)) r s'\") apply (frule_tac x="\mark. Some mark = gc_mark s'\" in spec) apply (frule_tac x="\mark. mark = sys_fM s'\" in spec) apply clarsimp apply (clarsimp simp: fun_eq_iff split: obj_at_splits) apply (rename_tac r) apply ( (drule_tac x=r in spec)+, auto)[1] apply (clarsimp simp: fun_eq_iff split: obj_at_splits) apply (rename_tac r) apply (drule_tac x=r in spec, auto)[1] apply (metis map_option_eq_Some)+ done lemmas gc_sweep_loop_invL_niE[nie] = iffD1[OF gc.sweep_loop_invL_eq_imp[simplified eq_imp_simps, rule_format, unfolded conj_explode, rule_format], rotated -1] lemma (in gc) sweep_loop_invL[intro]: "\ fM_fA_invL \<^bold>\ phase_invL \<^bold>\ sweep_loop_invL \<^bold>\ tso_lock_invL \<^bold>\ LSTP (phase_rel_inv \<^bold>\ mutators_phase_inv \<^bold>\ valid_W_inv) \ gc \ sweep_loop_invL \" apply (vcg_jackhammer simp: no_grey_refs_def phase_rel_inv_def phase_rel_def) apply (rename_tac s s' x) apply (case_tac "x \ gc_refs s\") apply force apply blast apply (clarsimp split: obj_at_splits) done lemma sweep_loop_sweep_locs[iff]: "sweep_loop_locs \ sweep_locs" by (auto simp: sweep_loop_locs_def sweep_locs_def intro: append_prefixD) lemma sweep_locs_subseteq_fM_tso_empty_locs[iff]: "sweep_locs \ fM_tso_empty_locs" by (auto simp: sweep_locs_def fM_tso_empty_locs_def) lemma sweep_loop_locs_fM_eq_locs: "sweep_loop_locs \ fM_eq_locs" by (auto simp: sweep_loop_locs_def fM_eq_locs_def sweep_locs_def) lemma sweep_loop_locs_fA_eq_locs: "sweep_loop_locs \ fA_eq_locs" apply (simp add: sweep_loop_locs_def fA_eq_locs_def sweep_locs_def) apply (intro subset_insertI2) apply (auto intro: append_prefixD) done lemma black_heap_locs_subseteq_fM_tso_empty_locs[iff]: "black_heap_locs \ fM_tso_empty_locs" by (auto simp: black_heap_locs_def fM_tso_empty_locs_def) lemma black_heap_locs_fM_eq_locs: "black_heap_locs \ fM_eq_locs" by (simp add: black_heap_locs_def fM_eq_locs_def) lemma black_heap_locs_fA_eq_locs: "black_heap_locs \ fA_eq_locs" by (simp add: black_heap_locs_def fA_eq_locs_def sweep_locs_def) lemma (in gc) fM_invL_tso_emptyD: "\ atS gc ls s; fM_fA_invL s; ls \ fM_tso_empty_locs \ \ tso_pending_fM gc s\ = []" by (auto simp: fM_fA_invL_def dest: atS_mono) lemma gc_sweep_loop_invL_locsE[rule_format]: "(atS gc (sweep_locs \ black_heap_locs) s \ False) \ gc.sweep_loop_invL s" apply (simp add: gc.sweep_loop_invL_def atS_un) apply (auto simp: loc atS_simps dest: atS_mono) apply (clarsimp simp: atS_def) apply (rename_tac x) apply (drule_tac x=x in bspec) apply (auto simp: sweep_locs_def intro: append_prefixD) done lemma (in sys) gc_sweep_loop_invL[intro]: "\ gc.fM_fA_invL \<^bold>\ gc.gc_W_empty_invL \<^bold>\ gc.handshake_invL \<^bold>\ gc.phase_invL \<^bold>\ gc.sweep_loop_invL \<^bold>\ LSTP (mutators_phase_inv \<^bold>\ tso_writes_inv \<^bold>\ valid_W_inv) \ sys \ gc.sweep_loop_invL \" apply vcg_nihe apply vcg_ni apply (clarsimp simp: do_write_action_def split: mem_write_action.splits elim: gc_sweep_loop_invL_niE) (* FIXME elimination rule using tso_writes_inv *) (* mw_Mark *) apply (rename_tac s s' p ws ref bool) apply (rule gc_sweep_loop_invL_locsE) apply (simp only: gc.gc_W_empty_invL_def no_grey_refs_locs_def cong del: atS_state_cong) apply (clarsimp simp: atS_un) apply (erule disjE) apply clarsimp apply (drule (1) no_grey_refs_no_pending_marks) apply (clarsimp simp: filter_empty_conv) apply (drule_tac x=p in spec) apply fastforce apply clarsimp apply (drule (1) no_grey_refs_no_pending_marks) apply (clarsimp simp: filter_empty_conv) apply (drule_tac x=p in spec) apply fastforce (* mw_Mutate *) apply (erule gc_sweep_loop_invL_niE, simp_all add: fun_eq_iff)[1] (* FIXME should be automatic *) (* mw_fA *) apply (erule gc_sweep_loop_invL_niE, simp_all add: fun_eq_iff)[1] (* FIXME should be automatic *) (* mw_fM *) apply (rename_tac s s' p ws bool) apply (rule gc_sweep_loop_invL_locsE) apply (case_tac p, clarsimp+) apply (drule (1) gc.fM_invL_tso_emptyD) apply simp_all (* mv_Phase *) apply (erule gc_sweep_loop_invL_niE, simp_all add: fun_eq_iff)[1] (* FIXME should be automatic *) done (* FIXME weird: expect more aggressive use of gc_sweep_loop_invL_niE by clarsimp *) lemma (in mut_m) gc_sweep_loop_invL[intro]: "\ gc.fM_fA_invL \<^bold>\ gc.handshake_invL \<^bold>\ gc.sweep_loop_invL \<^bold>\ LSTP (mutators_phase_inv \<^bold>\ valid_refs_inv) \ mutator m \ gc.sweep_loop_invL \" apply vcg_nihe apply vcg_ni apply (clarsimp simp: inv gc.sweep_loop_invL_def gc.fM_fA_invL_def) apply (intro allI conjI impI) (* four subgoals *) apply ((erule (1) thin_locs)+)[1] apply (force simp: loc) apply ((erule (1) thin_locs)+)[1] apply (force simp: loc) apply (rename_tac s s' ra) apply (drule mp, erule atS_mono[OF _ sweep_loop_locs_fA_eq_locs]) apply (drule mp, erule atS_mono[OF _ sweep_loop_locs_fM_eq_locs]) apply force apply (rename_tac s s' ra) apply (drule mp, erule atS_mono[OF _ black_heap_locs_fA_eq_locs]) apply (drule mp, erule atS_mono[OF _ black_heap_locs_fM_eq_locs]) apply force done (* FIXME crappy split *) lemma (in gc) sys_phase_inv[intro]: "\ fM_fA_invL \<^bold>\ gc_W_empty_invL \<^bold>\ handshake_invL \<^bold>\ obj_fields_marked_invL \<^bold>\ phase_invL \<^bold>\ sweep_loop_invL \<^bold>\ LSTP (phase_rel_inv \<^bold>\ sys_phase_inv \<^bold>\ valid_W_inv \<^bold>\ tso_writes_inv) \ gc \ LSTP sys_phase_inv \" apply vcg_jackhammer apply ( (erule black_heap_no_greys, simp) | fastforce dest!: phase_rel_invD simp: phase_rel_def filter_empty_conv )+ done lemma (in gc) no_black_refs_sweep_loop_free[simp]: "no_black_refs s \ no_black_refs (s(sys := s sys\heap := (sys_heap s)(gc_tmp_ref s := None)\))" by (simp add: no_black_refs_def) lemma (in gc) no_black_refs_load_W[simp]: "\ no_black_refs s; gc_W s = {} \ \ no_black_refs (s(gc := s gc\W := sys_W s\, sys := s sys\W := {}\))" by (simp add: no_black_refs_def) lemma marked_deletions_sweep_loop_free[simp]: "\ mut_m.marked_deletions m s; mut_m.reachable_snapshot_inv m s; no_grey_refs s; white r s \ \ mut_m.marked_deletions m (s(sys := s sys\heap := (sys_heap s)(r := None)\))" apply (clarsimp simp: mut_m.marked_deletions_def split: mem_write_action.splits) apply (rename_tac ref field option) apply (drule_tac x="mw_Mutate ref field option" in spec) apply clarsimp apply (clarsimp simp: obj_at_field_on_heap_def split: option.splits) apply (rule conjI) apply (clarsimp simp: mut_m.reachable_snapshot_inv_def) apply (drule spec[where x=r], clarsimp simp: in_snapshot_def) apply (drule mp, auto simp: mut_m.reachable_def mut_m.tso_write_refs_def split: mem_write_action.splits)[1] (* FIXME rule *) apply (drule grey_protects_whiteD) apply (clarsimp simp: no_grey_refs_def) apply clarsimp apply (rule conjI) apply clarsimp apply clarsimp apply (rule conjI) apply (clarsimp simp: mut_m.reachable_snapshot_inv_def) apply (drule spec[where x=r], clarsimp simp: in_snapshot_def) apply (drule mp, auto simp: mut_m.reachable_def mut_m.tso_write_refs_def split: mem_write_action.splits)[1] (* FIXME rule *) apply (drule grey_protects_whiteD) apply (clarsimp simp: no_grey_refs_def) apply (clarsimp split: obj_at_splits) done lemma (in gc) mutator_phase_inv[intro]: "\ fM_fA_invL \<^bold>\ gc_W_empty_invL \<^bold>\ handshake_invL \<^bold>\ obj_fields_marked_invL \<^bold>\ sweep_loop_invL \<^bold>\ gc_mark.mark_object_invL \<^bold>\ LSTP (handshake_phase_inv \<^bold>\ mutators_phase_inv \<^bold>\ tso_writes_inv \<^bold>\ valid_refs_inv \<^bold>\ valid_W_inv) \ gc \ LSTP (mut_m.mutator_phase_inv m) \" apply vcg_jackhammer apply (simp_all add: mutator_phase_inv_aux_case split: handshake_phase.splits) (* sweep_loop_free *) apply (intro allI conjI impI) apply (drule mut_m.handshake_phase_invD[where m=m], clarsimp simp: hp_step_rel_def) apply (rule mut_m.reachable_snapshot_inv_sweep_loop_free, simp_all) (* ''mark_loop_get_work_load_W'' *) apply clarsimp apply (drule spec[where x=m]) apply (clarsimp simp: mut_m.reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def) (* FIXME rule *) (* mark_loop_blacken *) apply (drule spec[where x=m]) apply clarsimp apply (intro allI conjI impI) apply clarsimp apply (drule mut_m.handshake_phase_invD[where m=m], clarsimp simp: hp_step_rel_def) apply (clarsimp simp: mut_m.reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def) apply (rename_tac s s' x) apply (drule_tac x=x in spec) apply clarsimp apply (rename_tac s s' x xa) apply (drule_tac x=xa in spec) apply clarsimp apply (rule obj_fields_marked_inv_has_white_path_to_blacken, simp_all)[1] (* mark_loop_mo_co_mark *) apply clarsimp apply (erule mut_m.reachable_snapshot_inv_mo_co_mark) (* FIXME hoist to the top level *) apply blast (* mark_loop_get_roots_load_W *) apply clarsimp apply (drule spec[where x=m]) apply (clarsimp simp: mut_m.reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def) (* FIXME rule *) done lemma (in mut_m) sys_phase_inv[intro]: "\ handshake_invL \<^bold>\ mark_object_invL \<^bold>\ mut_get_roots.mark_object_invL m \<^bold>\ mut_store_del.mark_object_invL m \<^bold>\ mut_store_ins.mark_object_invL m \<^bold>\ LSTP (fA_rel_inv \<^bold>\ fM_rel_inv \<^bold>\ handshake_phase_inv \<^bold>\ mutators_phase_inv \<^bold>\ phase_rel_inv \<^bold>\ sys_phase_inv \<^bold>\ valid_refs_inv) \ mutator m \ LSTP sys_phase_inv \" apply (vcg_jackhammer simp: fA_rel_inv_def fM_rel_inv_def) apply (simp_all add: sys_phase_inv_aux_case heap_colours_colours split: handshake_phase.splits if_splits) (* alloc *) subgoal by (clarsimp simp: fA_rel_def fM_rel_def no_black_refs_def dest!: handshake_phase_invD phase_rel_invD split: handshake_phase.splits) (* store_ins_mo_co_mark *) subgoal by (fastforce simp: fA_rel_def fM_rel_def hp_step_rel_def dest!: handshake_phase_invD phase_rel_invD) subgoal apply (drule spec[where x=m]) apply (rule conjI) apply (clarsimp simp: hp_step_rel_def phase_rel_def conj_disj_distribR[symmetric] dest!: handshake_phase_invD phase_rel_invD) apply (elim disjE, simp_all add: no_grey_refs_not_rootD)[1] apply (clarsimp simp: hp_step_rel_def phase_rel_def dest!: handshake_phase_invD phase_rel_invD) apply (elim disjE, simp_all add: no_grey_refs_not_rootD)[1] apply clarsimp apply (elim disjE, simp_all add: no_grey_refs_not_rootD filter_empty_conv)[1] apply fastforce done (* store_del_mo_co_unlock *) subgoal by (fastforce simp: fA_rel_def fM_rel_def hp_step_rel_def dest!: handshake_phase_invD phase_rel_invD) subgoal apply (drule spec[where x=m]) apply (rule conjI) apply (clarsimp simp: hp_step_rel_def phase_rel_def conj_disj_distribR[symmetric] no_grey_refs_not_rootD dest!: handshake_phase_invD phase_rel_invD) apply (clarsimp simp: hp_step_rel_def phase_rel_def dest!: handshake_phase_invD phase_rel_invD) apply (elim disjE, simp_all add: no_grey_refs_not_rootD)[1] apply clarsimp apply (elim disjE, simp_all add: no_grey_refs_not_rootD filter_empty_conv)[1] apply fastforce done (* hs_get_roots_done *) subgoal by (clarsimp simp: hp_step_rel_def dest!: handshake_phase_invD phase_rel_invD) subgoal by (clarsimp simp: hp_step_rel_def dest!: handshake_phase_invD phase_rel_invD) subgoal by (clarsimp simp: hp_step_rel_def dest!: handshake_phase_invD phase_rel_invD) subgoal by (clarsimp simp: hp_step_rel_def dest!: handshake_phase_invD phase_rel_invD) subgoal apply (clarsimp simp: hp_step_rel_def phase_rel_def filter_empty_conv dest!: handshake_phase_invD phase_rel_invD) apply auto done (* hs_get_roots_loop_mo_co_mark *) subgoal by (clarsimp simp: hp_step_rel_def dest!: handshake_phase_invD phase_rel_invD) subgoal apply (clarsimp simp: hp_step_rel_def phase_rel_def filter_empty_conv dest!: handshake_phase_invD phase_rel_invD) apply auto done (* hs_get_work_done *) subgoal by (clarsimp simp: hp_step_rel_def dest!: handshake_phase_invD phase_rel_invD) subgoal by (clarsimp simp: hp_step_rel_def dest!: handshake_phase_invD phase_rel_invD) subgoal by (clarsimp simp: hp_step_rel_def dest!: handshake_phase_invD phase_rel_invD) subgoal by (clarsimp simp: hp_step_rel_def dest!: handshake_phase_invD phase_rel_invD) subgoal apply (clarsimp simp: hp_step_rel_def phase_rel_def filter_empty_conv dest!: handshake_phase_invD phase_rel_invD) apply auto done done lemma no_grey_refs_no_marks[simp]: "\ no_grey_refs s; valid_W_inv s \ \ \sys_mem_write_buffers p s = mw_Mark r fl # ws" by (auto simp: no_grey_refs_def) context sys begin lemma black_heap_dequeue_mark[iff]: "\ sys_mem_write_buffers p s = mw_Mark r fl # ws; black_heap s; valid_W_inv s \ \ black_heap (s(sys := s sys\heap := (sys_heap s)(r := Option.map_option (obj_mark_update (\_. fl)) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\))" by (auto simp: black_heap_def) lemma black_heap_dequeue_ref[iff]: "\ sys_mem_write_buffers p s = mw_Mutate r f r' # ws; black_heap s \ \ black_heap (s(sys := s sys\heap := (sys_heap s)(r := Option.map_option (\obj. obj\obj_fields := (obj_fields obj)(f := r')\) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\))" by (simp add: black_heap_def black_def) lemma white_heap_dequeue_fM[iff]: "black_heap s\ \ white_heap (s\(sys := s\ sys\fM := \ sys_fM s\, mem_write_buffers := (mem_write_buffers (s\ sys))(gc := ws)\))" by (auto simp: black_heap_def white_heap_def) lemma black_heap_dequeue_fM[iff]: "\ white_heap s\; no_grey_refs s\ \ \ black_heap (s\(sys := s\ sys\fM := \ sys_fM s\, mem_write_buffers := (mem_write_buffers (s\ sys))(gc := ws)\))" by (auto simp: black_heap_def white_heap_def no_grey_refs_def) lemma white_heap_dequeue_ref[iff]: "\ sys_mem_write_buffers p s = mw_Mutate r f r' # ws; white_heap s \ \ white_heap (s(sys := s sys\heap := (sys_heap s)(r := Option.map_option (\obj. obj\obj_fields := (obj_fields obj)(f := r')\) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\))" by (simp add: white_heap_def) lemma (in sys) sys_phase_inv[intro]: "\ LSTP (fA_rel_inv \<^bold>\ fM_rel_inv \<^bold>\ handshake_phase_inv \<^bold>\ mutators_phase_inv \<^bold>\ phase_rel_inv \<^bold>\ sys_phase_inv \<^bold>\ tso_writes_inv \<^bold>\ valid_W_inv) \ sys \ LSTP sys_phase_inv \" by (vcg_jackhammer simp: fA_rel_inv_def fM_rel_inv_def p_not_sys) (clarsimp simp: do_write_action_def sys_phase_inv_aux_case split: mem_write_action.splits handshake_phase.splits if_splits; erule disjE; clarsimp simp: fA_rel_def fM_rel_def; fail) end lemma valid_W_inv_unlockE[iff]: "\ sys_mem_lock s = Some p; sys_mem_write_buffers p s = []; \r. r \ ghost_honorary_grey (s p) \ marked r s; valid_W_inv s \ \ valid_W_inv (s(sys := mem_lock_update Map.empty (s sys)))" unfolding valid_W_inv_def by clarsimp (metis emptyE empty_set) lemma valid_W_inv_mark: "\ sys_mem_lock s = Some p; white w s; valid_W_inv s \ \ w \ ghost_honorary_grey (s p) \ (\q. w \ WL q s)" by (metis Un_iff WL_def marked_not_white option.inject valid_W_invD(1) valid_W_invD3(2)) lemma (in gc) valid_W_inv[intro]: notes valid_W_invD2[dest!, simp] notes valid_W_invD3[dest, simp] shows "\ fM_invL \<^bold>\ gc_mark.mark_object_invL \<^bold>\ gc_W_empty_invL \<^bold>\ obj_fields_marked_invL \<^bold>\ sweep_loop_invL \<^bold>\ tso_lock_invL \<^bold>\ LSTP (tso_writes_inv \<^bold>\ valid_W_inv) \ gc \ LSTP valid_W_inv \" apply (vcg_jackhammer simp: fM_rel_def) (* sweep loop free: what's with the case splitting? *) subgoal for s s' apply (subst valid_W_inv_def) apply (intro allI conjI impI; clarsimp simp: p_not_sys split: if_splits) apply blast apply blast apply blast apply blast apply (rename_tac p x) apply (case_tac p; auto) apply (rename_tac p) apply (case_tac p; force simp: no_grey_refs_def) done (* mark_loop_get_work_load_W *) subgoal by (subst valid_W_inv_def) (auto simp: all_conj_distrib split: process_name.splits) (* mark_loop_blacken *) subgoal by (subst valid_W_inv_def) (auto simp: all_conj_distrib) (* mark_loop_mo_co_W *) subgoal apply (subst valid_W_inv_def) apply clarsimp apply blast done (* mark_loop_mo_co_unlock *) subgoal apply (subst valid_W_inv_def) apply (clarsimp simp: all_conj_distrib) apply (intro allI conjI impI) apply blast apply blast apply (auto iff: p_not_sys split: if_splits)[1] apply blast apply blast apply blast apply force done (* mark_loop_mo_co_mark *) subgoal apply (subst valid_W_inv_def) apply (clarsimp simp: all_conj_distrib) apply (intro allI conjI impI) apply blast apply blast apply blast apply (frule (2) valid_W_inv_mark; auto)[1] apply (clarsimp dest!: valid_W_invD(1) split: obj_at_splits) (* FIXME want a cheaper contradiction between white and marked *) apply (clarsimp dest!: valid_W_invD(1) split: obj_at_splits) apply (clarsimp dest!: valid_W_invD(1) split: obj_at_splits) apply blast apply blast apply blast apply blast apply blast apply blast apply force done (* mark_loop_mo_co_lock *) subgoal apply (subst valid_W_inv_def) apply (auto simp: all_conj_distrib) done (* ''mark_loop_get_roots_load_W'' *) subgoal apply (subst valid_W_inv_def) apply (auto simp: all_conj_distrib split: process_name.splits) done done lemma (in mut_m) valid_W_inv[intro]: notes valid_W_invD2[dest!, simp] notes valid_W_invD3[dest, simp] shows "\ handshake_invL \<^bold>\ mark_object_invL \<^bold>\ tso_lock_invL \<^bold>\ mut_get_roots.mark_object_invL m \<^bold>\ mut_store_del.mark_object_invL m \<^bold>\ mut_store_ins.mark_object_invL m \<^bold>\ LSTP (fM_rel_inv \<^bold>\ sys_phase_inv \<^bold>\ tso_writes_inv \<^bold>\ valid_refs_inv \<^bold>\ valid_W_inv) \ mutator m \ LSTP valid_W_inv \" apply (vcg_jackhammer simp: fM_rel_inv_def fM_rel_def) (* alloc *) subgoal apply (subst valid_W_inv_def) apply (clarsimp simp: all_conj_distrib) apply (intro allI conjI impI; auto) done (* store ins mo co W *) subgoal apply (subst valid_W_inv_def) apply clarsimp apply blast done (* store ins mo co unlock *) subgoal for s s' y apply (subst valid_W_inv_def) apply (clarsimp simp: all_conj_distrib) apply (intro allI conjI impI) subgoal by blast subgoal by blast subgoal for p x by (case_tac "p = mutator m"; force split: if_splits) subgoal by blast subgoal by blast subgoal by blast subgoal by force done (* store ins mo co mark *) subgoal apply (subst valid_W_inv_def) apply (clarsimp simp: all_conj_distrib) apply (intro allI conjI impI) subgoal by blast subgoal by blast subgoal by blast subgoal by (blast dest: valid_W_inv_mark) subgoal by (clarsimp dest!: valid_W_invD(1) split: obj_at_splits) (* FIXME want a cheaper contradiction between white and marked *) subgoal by (clarsimp dest!: valid_W_invD(1) split: obj_at_splits) subgoal by (clarsimp dest!: valid_W_invD(1) split: obj_at_splits) subgoal by blast subgoal by blast subgoal by blast subgoal by blast subgoal by blast subgoal by blast subgoal by force done (* store ins mo co lock *) subgoal apply (subst valid_W_inv_def) apply (clarsimp simp: all_conj_distrib) apply (intro allI conjI impI) subgoal by blast subgoal by blast subgoal by force subgoal by blast subgoal by blast subgoal by blast subgoal by force subgoal by blast subgoal by force done (* store del mo co W *) subgoal apply (subst valid_W_inv_def) apply clarsimp apply blast done (* store del mo co unlock *) subgoal for s s' y apply (subst valid_W_inv_def) apply (clarsimp simp: all_conj_distrib) apply (intro allI conjI impI) subgoal by blast subgoal by blast subgoal for p x by (cases "p = mutator m") (auto split: if_splits) subgoal by blast subgoal by blast subgoal by blast subgoal by force done (* store del mo co mark *) subgoal apply (subst valid_W_inv_def) apply (clarsimp simp: all_conj_distrib) apply (intro allI conjI impI) subgoal by blast subgoal by blast subgoal by blast subgoal by (frule (2) valid_W_inv_mark; auto) subgoal by (clarsimp dest!: valid_W_invD(1) split: obj_at_splits) (* FIXME want a cheaper contradiction between white and marked *) subgoal by (clarsimp dest!: valid_W_invD(1) split: obj_at_splits) subgoal by (clarsimp dest!: valid_W_invD(1) split: obj_at_splits) subgoal by blast subgoal by blast subgoal by blast subgoal by blast subgoal by blast subgoal by blast subgoal by force done (* store del mo co lock *) subgoal apply (subst valid_W_inv_def) apply (clarsimp simp: all_conj_distrib) apply (intro allI conjI impI) subgoal by blast subgoal by blast subgoal by force subgoal by blast subgoal by blast subgoal by blast subgoal by force subgoal by blast subgoal by force done (* get roots done *) subgoal apply (subst valid_W_inv_def) apply (clarsimp simp: all_conj_distrib) apply (intro allI conjI impI) subgoal by blast subgoal by blast subgoal by blast subgoal by blast subgoal by (auto split: if_splits obj_at_splits process_name.splits) subgoal by blast subgoal by blast subgoal by blast subgoal by (auto split: if_splits obj_at_splits process_name.splits) subgoal by blast subgoal by blast done (* hs get roots loop mo co W *) subgoal apply (subst valid_W_inv_def) apply clarsimp apply blast done (* hs get roots loop mo co unlock *) subgoal for s s' y apply (subst valid_W_inv_def) apply (clarsimp simp: all_conj_distrib) apply (intro allI conjI impI) subgoal by blast subgoal by blast subgoal for p x by (cases "p = mutator m") (auto split: if_splits) subgoal by blast subgoal by blast subgoal by blast subgoal by force done (* hs get roots loop mo co mark *) subgoal apply (subst valid_W_inv_def) apply (clarsimp simp: all_conj_distrib) apply (intro allI conjI impI) subgoal by blast subgoal by blast subgoal by blast subgoal by (frule (2) valid_W_inv_mark; auto) subgoal by (clarsimp dest!: valid_W_invD(1) split: obj_at_splits) (* FIXME want a cheaper contradiction between white and marked *) subgoal by (clarsimp dest!: valid_W_invD(1) split: obj_at_splits) subgoal by (clarsimp dest!: valid_W_invD(1) split: obj_at_splits) subgoal by blast subgoal by blast subgoal by blast subgoal by blast subgoal by blast subgoal by blast subgoal by force done (* hs get roots loop mo co lock *) subgoal apply (subst valid_W_inv_def) apply (clarsimp simp: all_conj_distrib) apply (intro allI conjI impI) subgoal by blast subgoal by blast subgoal by force subgoal by blast subgoal by blast subgoal by blast subgoal by force subgoal by blast subgoal by force done (* hs get work done *) subgoal apply (subst valid_W_inv_def) apply (clarsimp simp: all_conj_distrib) apply (intro allI conjI impI) subgoal by blast subgoal by blast subgoal by blast subgoal by blast subgoal by (auto split: if_splits obj_at_splits process_name.splits) subgoal by blast subgoal by blast subgoal by blast subgoal by force subgoal by blast subgoal by blast done done lemma (in sys) valid_W_inv[intro]: notes valid_W_invD2[dest!, simp] notes valid_W_invD3[dest, simp] notes valid_W_invD4[dest!, simp] notes o_def[simp] shows "\ LSTP (fM_rel_inv \<^bold>\ sys_phase_inv \<^bold>\ tso_writes_inv \<^bold>\ valid_refs_inv \<^bold>\ valid_W_inv) \ sys \ LSTP valid_W_inv \" apply vcg_jackhammer apply (subst valid_W_inv_def) apply (clarsimp simp: do_write_action_def all_conj_distrib fM_rel_inv_def split: mem_write_action.splits) (* mw_Mark *) subgoal apply (intro allI conjI impI) apply blast apply blast apply blast apply blast apply blast apply force apply blast apply blast apply blast apply (force simp: filter_empty_conv) apply force apply blast done (* mw_Mutate, mw_fA *) subgoal by (intro allI conjI impI; auto) subgoal by (intro allI conjI impI; auto) (* mw_fM *) subgoal for s s' p ws bool apply (clarsimp simp: fM_rel_def) apply (rule disjE[OF iffD1[OF p_not_sys]], assumption) prefer 2 apply clarsimp apply (case_tac "sys_ghost_handshake_phase s\ = hp_Idle"; clarsimp) apply (intro allI conjI impI) subgoal by blast subgoal by blast subgoal by blast subgoal by blast subgoal by blast subgoal by force subgoal by (force dest: no_grey_refs_no_pending_marks) subgoal by blast subgoal by force subgoal by (fastforce dest: no_grey_refs_no_pending_marks) subgoal by (fastforce dest: no_grey_refs_no_pending_marks) done (* mw_Phase *) apply (intro allI conjI impI; auto) done lemma (in gc) strong_tricolour_inv[intro]: "\ fM_fA_invL \<^bold>\ gc_W_empty_invL \<^bold>\ gc_mark.mark_object_invL \<^bold>\ obj_fields_marked_invL \<^bold>\ sweep_loop_invL \<^bold>\ LSTP (strong_tricolour_inv \<^bold>\ tso_writes_inv \<^bold>\ valid_W_inv) \ gc \ LSTP strong_tricolour_inv \" apply (vcg_jackhammer simp: strong_tricolour_inv_def) apply (fastforce elim!: obj_fields_marked_inv_blacken) done lemma (in mut_m) strong_tricolour[intro]: "\ mark_object_invL \<^bold>\ mut_get_roots.mark_object_invL m \<^bold>\ mut_store_del.mark_object_invL m \<^bold>\ mut_store_ins.mark_object_invL m \<^bold>\ LSTP (fA_rel_inv \<^bold>\ fM_rel_inv \<^bold>\ handshake_phase_inv \<^bold>\ mutators_phase_inv \<^bold>\ strong_tricolour_inv \<^bold>\ sys_phase_inv \<^bold>\ valid_refs_inv) \ mutator m \ LSTP strong_tricolour_inv \" apply (vcg_jackhammer simp: strong_tricolour_inv_def fA_rel_inv_def fM_rel_inv_def) apply (drule handshake_phase_invD) apply (clarsimp simp: sys_phase_inv_aux_case split: handshake_phase.splits if_splits) apply (blast dest: heap_colours_colours) apply (blast dest: heap_colours_colours) apply (clarsimp simp: fA_rel_def fM_rel_def no_black_refs_def) (* FIXME rule *) apply (drule_tac x=m in spec) apply (clarsimp simp: hp_step_rel_def) apply (elim disjE, simp_all) apply (auto dest: no_black_refsD)[1] apply (auto dest: no_black_refsD)[1] apply (clarsimp simp: fA_rel_def fM_rel_def) apply (clarsimp split: obj_at_splits) apply (drule spec[where x=m]) apply (clarsimp simp: hp_step_rel_def) apply (elim disjE; force simp: fA_rel_def fM_rel_def split: obj_at_splits) apply (drule spec[where x=m]) apply (clarsimp simp: hp_step_rel_def) apply (elim disjE; force simp: fA_rel_def fM_rel_def split: obj_at_splits) done lemma (in sys) strong_tricolour_inv[intro]: "\ LSTP (fM_rel_inv \<^bold>\ handshake_phase_inv \<^bold>\ mutators_phase_inv \<^bold>\ strong_tricolour_inv \<^bold>\ sys_phase_inv \<^bold>\ tso_writes_inv \<^bold>\ valid_W_inv) \ sys \ LSTP strong_tricolour_inv \" supply [[simproc del: defined_all]] apply (vcg_jackhammer simp: strong_tricolour_inv_def p_not_sys) apply (clarsimp simp: do_write_action_def fM_rel_inv_def split: mem_write_action.splits) apply (rename_tac ref field) (* mw_Mark *) subgoal for s s' p ws x xa ref field apply (frule (1) valid_W_invD2) apply clarsimp apply (case_tac "x = ref", simp_all) apply (clarsimp simp: grey_def WL_def) done (* mw_Mutate *) apply (elim disjE, simp_all) apply (clarsimp simp: points_to_mw_Mutate) apply (elim disjE, simp_all) apply (rename_tac s s' ws x xa ref field option m) apply (drule_tac m=m in mut_m.handshake_phase_invD) apply (frule_tac x=m in spec) apply (clarsimp simp: mutator_phase_inv_aux_case hp_step_rel_def split: handshake_phase.splits) apply (elim disjE, simp_all split: if_splits) apply (blast dest!: heap_colours_colours) apply (blast dest!: heap_colours_colours) apply (drule_tac x=m in spec) apply (clarsimp simp: no_black_refs_def) apply (drule_tac x=m in spec) apply (clarsimp simp: no_black_refsD) (* FIXME split less *) apply (drule_tac x=m in spec) apply (clarsimp simp: mut_m.marked_insertions_def) apply (drule_tac x="mw_Mutate x field (Some xa)" in spec) apply (clarsimp dest!: marked_not_white) apply (drule_tac x=m in spec) apply (clarsimp simp: mut_m.marked_insertions_def) apply (drule_tac x="mw_Mutate x field (Some xa)" in spec) apply (clarsimp dest!: marked_not_white) apply (drule_tac x=m in spec) apply (clarsimp simp: mut_m.marked_insertions_def) apply (drule_tac x="mw_Mutate x field (Some xa)" in spec) apply (clarsimp dest!: marked_not_white) apply (drule_tac x=m in spec) apply (clarsimp simp: mut_m.marked_insertions_def) apply (drule_tac x="mw_Mutate x field (Some xa)" in spec) apply (clarsimp dest!: marked_not_white) apply (drule_tac x=m in spec) apply (clarsimp simp: mut_m.marked_insertions_def) apply (drule_tac x="mw_Mutate x field (Some xa)" in spec) apply (clarsimp dest!: marked_not_white) apply clarsimp apply (elim disjE, simp_all split: if_splits) apply (blast dest!: heap_colours_colours) apply (blast dest!: heap_colours_colours) apply (drule_tac x=m in spec) apply (clarsimp simp: no_black_refs_def) apply (drule_tac x=m in spec) apply (clarsimp simp: mut_m.marked_insertions_def) apply (drule_tac x="mw_Mutate x field (Some xa)" in spec) apply (clarsimp dest!: marked_not_white) apply (drule_tac x=m in spec) apply (clarsimp simp: mut_m.marked_insertions_def) apply (drule_tac x="mw_Mutate x field (Some xa)" in spec) apply (clarsimp dest!: marked_not_white) (* mw_fM *) apply (rename_tac s s' p ws x xa bool) apply (erule disjE) apply (clarsimp simp: fM_rel_def black_heap_def split: if_splits) apply ( (drule_tac x=x in spec)+ )[1] apply (frule colours_distinct) apply (clarsimp split: obj_at_splits) apply (clarsimp simp: white_heap_def) apply ( (drule_tac x=xa in spec)+ )[1] apply (clarsimp split: obj_at_splits) apply fastforce done text\Remaining non-interference proofs.\ lemma marked_insertionsD: "\ sys_mem_write_buffers (mutator m) s = mw_Mutate r f (Some r') # ws; mut_m.marked_insertions m s \ \ marked r' s" by (clarsimp simp: mut_m.marked_insertions_def) lemma mut_hs_get_roots_loop_locs_subseteq_hs_get_roots: "mut_hs_get_roots_loop_locs \ prefixed ''hs_get_roots_''" by (auto simp: mut_hs_get_roots_loop_locs_def intro: append_prefixD) lemma mut_m_handshake_invL_get_roots: "\ mut_m.handshake_invL m s; atS (mutator m) mut_hs_get_roots_loop_locs s \ \ sys_handshake_type s\ = ht_GetRoots \ sys_handshake_pending m s\" unfolding mut_m.handshake_invL_def apply (elim conjE) apply (drule mp, erule (1) atS_mono[OF _ mut_hs_get_roots_loop_locs_subseteq_hs_get_roots]) done lemma subseteq_mut_mo_valid_ref_locs: "prefixed ''store_del_mo'' \ {''lop_store_ins''} \ mut_mo_valid_ref_locs" by (auto simp: mut_mo_valid_ref_locs_def intro: append_prefixD) lemma subseteq_mut_mo_valid_ref_locs2: "prefixed ''store_ins'' \ mut_mo_valid_ref_locs" by (clarsimp simp: mut_mo_valid_ref_locs_def) lemma mut_phase_inv: "\ ghost_handshake_phase (s (mutator m)) = hp_Mark \ ghost_handshake_phase (s (mutator m)) = hp_IdleMarkSweep \ sys_phase s \ ph_Idle; mut_m.mutator_phase_inv_aux m (ghost_handshake_phase (s (mutator m))) s \ \ mut_m.marked_insertions m s \ mut_m.marked_deletions m s" by (cases "ghost_handshake_phase (s (mutator m))", simp_all) lemma (in sys) mut_mark_object_invL[intro]: notes mut_m.mark_object_invL_def[inv] notes do_write_action_fM[where m=m, simp] notes do_write_action_prj_simps(1)[simp del] do_write_action_prj_simps(2)[simp del] notes mut_m_get_roots_no_fM_write[where m=m, simp] notes mut_m_get_roots_no_phase_write[where m=m, simp] notes mut_m_ghost_handshake_phase_not_hp_Idle[where m=m, simp] notes atS_simps[simp] filter_empty_conv[simp] shows "\ mut_m.handshake_invL m \<^bold>\ mut_m.mark_object_invL m \<^bold>\ LSTP (fA_rel_inv \<^bold>\ fM_rel_inv \<^bold>\ handshake_phase_inv \<^bold>\ mutators_phase_inv \<^bold>\ phase_rel_inv \<^bold>\ valid_refs_inv \<^bold>\ valid_W_inv \<^bold>\ tso_writes_inv) \ sys \ mut_m.mark_object_invL m \" apply vcg_nihe apply vcg_ni subgoal by (clarsimp simp: do_write_action_def filter_empty_conv p_not_sys dest!: valid_W_invD2 elim!: obj_at_weakenE split: mem_write_action.splits) subgoal by (clarsimp simp: do_write_action_def filter_empty_conv p_not_sys dest!: valid_W_invD2 elim!: obj_at_weakenE split: mem_write_action.splits) subgoal by (clarsimp simp: do_write_action_def filter_empty_conv p_not_sys loc dest!: valid_W_invD2 elim!: obj_at_weakenE split: mem_write_action.splits) subgoal apply (drule phase_rel_invD) apply (drule mut_m.handshake_phase_invD[where m=m]) apply (clarsimp simp: p_not_sys) apply (erule disjE) apply (clarsimp simp: hp_step_rel_def) apply (elim disjE, simp_all add: phase_rel_def)[1] apply force done subgoal apply (clarsimp simp: do_write_action_def filter_empty_conv p_not_sys loc fM_rel_inv_def dest!: valid_W_invD2 elim!: obj_at_weakenE split: mem_write_action.splits) apply (rule conjI) apply (clarsimp elim!: obj_at_weakenE) apply clarsimp apply (drule mut_m.handshake_phase_invD[where m=m]) apply (erule disjE) apply (auto simp: fM_rel_def hp_step_rel_def)[1] apply force apply (erule disjE) apply (drule mut_m.handshake_phase_invD[where m=m]) apply (drule phase_rel_invD) apply (clarsimp simp: hp_step_rel_def) apply (auto simp: phase_rel_def)[1] apply force done subgoal for s s' p w ws apply (drule mp, erule atS_mono[OF _ subseteq_mut_mo_valid_ref_locs]) apply ((thin_tac "atS p ls s \ Q" for p ls s Q)+)[1] apply ((thin_tac "at p ls s \ Q" for p ls s Q)+)[1] apply (clarsimp simp: do_write_action_def p_not_sys loc split: mem_write_action.splits if_splits) apply (drule (1) valid_W_invD2) apply (erule obj_at_field_on_heap_weakenE) apply (fastforce split: obj_at_splits) apply (drule (1) valid_W_invD2) apply (erule obj_at_field_on_heap_weakenE) apply (fastforce split: obj_at_splits) apply (drule spec[where x=m]) apply (frule (1) mut_phase_inv) apply (rename_tac refa fielda option) apply (frule_tac y="refa" in valid_refs_invD3, simp_all)[1] apply (frule_tac y="tmp_ref (s\ (mutator m))" in valid_refs_invD(2), simp_all)[1] apply (clarsimp split: obj_at_splits option.splits) apply force apply (frule (1) marked_insertionsD) apply (auto split: obj_at_splits; fail)[1] apply (erule disjE) (* super messy case *) apply force apply (rule conjI) apply (erule obj_at_field_on_heap_imp_valid_ref) apply (clarsimp split: option.splits) apply (drule phase_rel_invD) apply (frule mut_m.handshake_phase_invD[where m=m]) apply (rename_tac ma x2) apply (drule_tac m=ma in mut_m.handshake_phase_invD) apply (frule spec[where x=m]) apply (drule_tac x=ma in spec) apply (clarsimp simp: hp_step_rel_def) apply (elim disjE, auto simp: phase_rel_def dest: marked_insertionsD; fail)[1] apply (erule disjE; clarsimp) apply (drule mut_m.handshake_phase_invD[where m=m]) apply (clarsimp simp: fM_rel_inv_def fM_rel_def hp_step_rel_def) apply (metis (no_types, lifting) handshake_phase.distinct(5) handshake_phase.distinct(7)) apply (erule disjE; clarsimp) apply (drule mut_m.handshake_phase_invD[where m=m]) apply (drule phase_rel_invD) apply (clarsimp simp: hp_step_rel_def phase_rel_def) done subgoal for s s' p w ws y apply (clarsimp simp: do_write_action_def p_not_sys split: mem_write_action.splits if_splits) apply (auto split: obj_at_splits; fail)[1] apply (erule disjE; clarsimp) apply (drule mut_m.handshake_phase_invD[where m=m]) apply (clarsimp simp: fM_rel_inv_def fM_rel_def hp_step_rel_def) apply (metis (no_types, lifting) handshake_phase.distinct(5) handshake_phase.distinct(7)) apply (erule disjE; clarsimp) apply (drule mut_m.handshake_phase_invD[where m=m]) apply (drule phase_rel_invD) apply (clarsimp simp: hp_step_rel_def phase_rel_def) done subgoal for s s' p w ws apply (drule mp, erule atS_mono[OF _ subseteq_mut_mo_valid_ref_locs2]) apply ((thin_tac "atS p ls s \ Q" for p ls s Q)+)[1] apply ((thin_tac "at p ls s \ Q" for p ls s Q)+)[1] apply (subst do_write_action_fM[where m=m], simp_all)[1] apply (elim disjE, simp_all)[1] apply (clarsimp simp: do_write_action_def p_not_sys split: mem_write_action.splits if_splits) apply (erule obj_at_field_on_heap_weakenE, auto)[1] apply (erule obj_at_field_on_heap_weakenE, auto split: obj_at_splits)[1] apply (drule spec[where x=m]) apply (frule (1) mut_phase_inv) apply (rename_tac refa fielda option) apply (frule_tac y="refa" in valid_refs_invD3, simp_all)[1] apply (frule_tac y="tmp_ref (s\ (mutator m))" in valid_refs_invD(2), simp_all)[1] apply (clarsimp split: obj_at_splits option.splits) apply auto[1] apply (frule (1) marked_insertionsD) apply (auto split: obj_at_splits)[1] apply (erule disjE) (* super messy case *) apply force apply (rule conjI) apply (erule obj_at_field_on_heap_imp_valid_ref) apply (clarsimp split: option.splits) apply (drule phase_rel_invD) apply (frule mut_m.handshake_phase_invD[where m=m]) apply (rename_tac ma x2) apply (drule_tac m=ma in mut_m.handshake_phase_invD) apply (frule spec[where x=m]) apply (drule_tac x=ma in spec) apply (clarsimp simp: hp_step_rel_def) apply (elim disjE, auto simp: phase_rel_def dest: marked_insertionsD)[1] apply (erule disjE) apply (drule mut_m.handshake_phase_invD[where m=m]) apply (drule phase_rel_invD) apply (clarsimp simp: hp_step_rel_def phase_rel_def) apply force done done lemma (in gc) mut_mark_object_invL[intro]: notes mut_m.mark_object_invL_def[inv] notes atS_simps[simp] shows "\ fM_fA_invL \<^bold>\ gc_W_empty_invL \<^bold>\ handshake_invL \<^bold>\ sweep_loop_invL \<^bold>\ mut_m.handshake_invL m \<^bold>\ mut_m.mark_object_invL m \<^bold>\ LSTP (fM_rel_inv \<^bold>\ handshake_phase_inv \<^bold>\ mutators_phase_inv \<^bold>\ sys_phase_inv) \ gc \ mut_m.mark_object_invL m \" apply vcg_nihe apply vcg_ni subgoal apply (drule (1) mut_m_handshake_invL_get_roots) apply clarsimp done subgoal by (simp add: mut_m.handshake_invL_def) subgoal by (fastforce simp: fM_rel_inv_def fM_rel_def hp_step_rel_def split: obj_at_splits) subgoal apply (drule mut_m.handshake_phase_invD[where m=m]) apply (drule spec[where x=m]) apply (clarsimp simp: valid_null_ref_def hp_step_rel_def conj_disj_distribR[symmetric] split: option.splits) apply (drule (1) mut_m.reachable_blackD) apply blast apply (clarsimp split: obj_at_splits) done subgoal apply (drule mp, erule atS_mono[OF _ subseteq_mut_mo_valid_ref_locs]) apply ((thin_tac "atS p ls s \ Q" for p ls s Q)+)[1] apply ((thin_tac "at p ls s \ Q" for p ls s Q)+)[1] apply (drule mut_m.handshake_phase_invD[where m=m]) apply (drule spec[where x=m]) apply (clarsimp simp: valid_null_ref_def hp_step_rel_def conj_disj_distribR[symmetric] split: option.splits) apply (drule (1) mut_m.reachable_blackD) apply blast apply (auto simp: obj_at_field_on_heap_def black_def split: obj_at_splits option.splits)[1] done subgoal by (clarsimp split: obj_at_splits) subgoal apply (drule mp, erule atS_mono[OF _ subseteq_mut_mo_valid_ref_locs2]) apply ((thin_tac "atS p ls s \ Q" for p ls s Q)+)[1] apply ((thin_tac "at p ls s \ Q" for p ls s Q)+)[1] apply (drule mut_m.handshake_phase_invD[where m=m]) apply (drule spec[where x=m]) apply (clarsimp simp: valid_null_ref_def hp_step_rel_def conj_disj_distribR[symmetric] split: option.splits) apply (drule (1) mut_m.reachable_blackD) apply blast apply (auto simp: obj_at_field_on_heap_def black_def split: obj_at_splits option.splits)[1] done done lemma (in gc) mut_store_old_mark_object_invL[intro]: notes mut_m.mark_object_invL_def[inv] shows "\ fM_fA_invL \<^bold>\ handshake_invL \<^bold>\ sweep_loop_invL \<^bold>\ gc_W_empty_invL \<^bold>\ mut_m.mark_object_invL m \<^bold>\ mut_store_del.mark_object_invL m \<^bold>\ LSTP (handshake_phase_inv \<^bold>\ mut_m.mutator_phase_inv m) \ gc \ mut_store_del.mark_object_invL m \" apply vcg_nihe apply vcg_ni apply ( (clarsimp dest!: mut_m.handshake_phase_invD[where m=m] simp: hp_step_rel_def conj_disj_distribR[symmetric] , drule_tac r="gc_tmp_ref s\" in mut_m.no_grey_refs_not_rootD[where m=m] , auto split: obj_at_splits if_splits)[1] )+ done lemma (in gc) mut_store_ins_mark_object_invL[intro]: notes mut_m.mark_object_invL_def[inv] shows "\ fM_fA_invL \<^bold>\ handshake_invL \<^bold>\ sweep_loop_invL \<^bold>\ gc_W_empty_invL \<^bold>\ mut_m.mark_object_invL m \<^bold>\ mut_store_ins.mark_object_invL m \<^bold>\ LSTP (handshake_phase_inv \<^bold>\ mut_m.mutator_phase_inv m) \ gc \ mut_store_ins.mark_object_invL m \" apply vcg_nihe apply vcg_ni apply ( (clarsimp dest!: mut_m.handshake_phase_invD[where m=m] simp: hp_step_rel_def conj_disj_distribR[symmetric] , drule_tac r="gc_tmp_ref s\" in mut_m.no_grey_refs_not_rootD[where m=m] , auto split: obj_at_splits if_splits)[1] )+ done lemma obj_fields_marked_locs_subseteq_hp_IdleMarkSweep_locs: "obj_fields_marked_locs \ hp_IdleMarkSweep_locs" apply (clarsimp simp: obj_fields_marked_locs_def hp_IdleMarkSweep_locs_def mark_loop_locs_def) apply (drule mp) apply (auto intro: append_prefixD) done lemma obj_fields_marked_locs_subseteq_hs_in_sync_locs: "obj_fields_marked_locs \ hs_in_sync_locs" by (auto simp: obj_fields_marked_locs_def hs_in_sync_locs_def hs_done_locs_def dest: prefix_same_cases) lemma handshake_obj_fields_markedD: "\ atS gc obj_fields_marked_locs s; gc.handshake_invL s \ \ sys_ghost_handshake_phase s\ = hp_IdleMarkSweep \ All (ghost_handshake_in_sync (s\ sys))" apply (simp add: gc.handshake_invL_def) apply (elim conjE) apply (drule mp, erule atS_mono[OF _ obj_fields_marked_locs_subseteq_hp_IdleMarkSweep_locs]) apply (drule mp, erule atS_mono[OF _ obj_fields_marked_locs_subseteq_hs_in_sync_locs]) apply clarsimp done lemma mark_loop_mo_mark_loop_field_done_subseteq_hp_IdleMarkSweep_locs: "prefixed ''mark_loop_mo'' \ {''mark_loop_mark_field_done''} \ hp_IdleMarkSweep_locs" apply (clarsimp simp: hp_IdleMarkSweep_locs_def mark_loop_locs_def) apply (drule mp) apply (auto intro: append_prefixD) done lemma mark_loop_mo_mark_loop_field_done_subseteq_hs_in_sync_locs: "prefixed ''mark_loop_mo'' \ {''mark_loop_mark_field_done''} \ hs_in_sync_locs" by (auto simp: hs_in_sync_locs_def hs_done_locs_def dest: prefix_same_cases) lemma mark_loop_mo_mark_loop_field_done_hp_phaseD: "\ atS gc (prefixed ''mark_loop_mo'' \ {''mark_loop_mark_field_done''}) s; gc.handshake_invL s \ \ sys_ghost_handshake_phase s\ = hp_IdleMarkSweep \ All (ghost_handshake_in_sync (s\ sys))" apply (simp add: gc.handshake_invL_def) apply (elim conjE) apply (drule mp, erule atS_mono[OF _ mark_loop_mo_mark_loop_field_done_subseteq_hp_IdleMarkSweep_locs]) apply (drule mp, erule atS_mono[OF _ mark_loop_mo_mark_loop_field_done_subseteq_hs_in_sync_locs]) apply clarsimp done (* an alternative is some kind of ghost_honorary_XXX for the GC while marking *) lemma gc_marking_reaches_mw_Mutate: assumes xys: "\y. (x reaches y) s \ valid_ref y s" assumes xy: "(x reaches y) (s(sys := s sys\heap := (sys_heap s)(r := Option.map_option (\obj. obj\obj_fields := (obj_fields obj)(f := opt_r')\) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\))" assumes sb: "sys_mem_write_buffers (mutator m) s = mw_Mutate r f opt_r' # ws" assumes vri: "valid_refs_inv s" shows "valid_ref y s" proof - from xy xys have "\z. z \ {x} \ mut_m.tso_write_refs m s \ (z reaches y) s \ valid_ref y s" proof(induct rule: rtranclp.induct) case (rtrancl_refl x) then show ?case by auto next case (rtrancl_into_rtrancl x y z) with sb vri show ?case apply (clarsimp simp: points_to_mw_Mutate) apply (elim disjE) apply (force dest: rtranclp.intros(2)) apply (force dest: rtranclp.intros(2)) apply clarsimp apply (elim disjE) apply (rule exI[where x=z]) apply (clarsimp simp: mut_m.tso_write_refs_def) apply (rule valid_refs_invD(3)[where m=m and x=z], auto simp: mut_m.tso_write_refs_def)[1] apply (force dest: rtranclp.intros(2)) apply clarsimp apply (elim disjE) apply (rule exI[where x=z]) apply (clarsimp simp: mut_m.tso_write_refs_def) apply (rule valid_refs_invD(3)[where m=m and x=z], auto simp: mut_m.tso_write_refs_def)[1] apply (force dest: rtranclp.intros(2)) done qed then show ?thesis by blast qed lemma (in sys) gc_obj_fields_marked_invL[intro]: notes gc.obj_fields_marked_invL_def[inv] shows "\ gc.fM_fA_invL \<^bold>\ gc.handshake_invL \<^bold>\ gc.obj_fields_marked_invL \<^bold>\ LSTP (fM_rel_inv \<^bold>\ handshake_phase_inv \<^bold>\ mutators_phase_inv \<^bold>\ tso_writes_inv \<^bold>\ valid_refs_inv \<^bold>\ valid_W_inv) \ sys \ gc.obj_fields_marked_invL \" apply vcg_nihe apply (vcg_ni simp: p_not_sys fM_rel_inv_def) apply (clarsimp simp: gc.obj_fields_marked_inv_def) apply (frule (1) handshake_obj_fields_markedD) apply (clarsimp simp: do_write_action_def split: mem_write_action.splits) (* mark *) subgoal for s s' p ws x ref bool apply (frule (1) valid_W_invD2) apply (drule_tac x=x in spec) apply clarsimp apply (erule obj_at_field_on_heap_weakenE) apply (clarsimp split: obj_at_splits) done (* ref *) subgoal for s s' p ws x x23 apply (erule disjE; clarsimp) apply (rename_tac m) apply (drule_tac m=m in mut_m.handshake_phase_invD; clarsimp simp: hp_step_rel_def conj_disj_distribR[symmetric]) apply (drule_tac x=m in spec; clarsimp) apply (drule_tac x=x in spec; clarsimp) apply (auto split: option.splits) done (* fM *) subgoal for s s' p ws x x4 apply (erule disjE) apply (auto simp: fM_rel_def filter_empty_conv)[1] apply clarsimp done subgoal for s s' p w ws apply (clarsimp simp: gc.obj_fields_marked_inv_def) apply (frule (1) mark_loop_mo_mark_loop_field_done_hp_phaseD) apply (clarsimp simp: do_write_action_def split: mem_write_action.splits) (* mark *) apply (frule (1) valid_W_invD2) apply (erule obj_at_field_on_heap_weakenE) apply (clarsimp split: obj_at_splits) (* ref *) apply (erule disjE, clarsimp+)[1] apply (rename_tac option m) apply (drule_tac m=m in mut_m.handshake_phase_invD, clarsimp simp: hp_step_rel_def conj_disj_distribR[symmetric]) apply (drule_tac x=m in spec, clarsimp) apply (rule conjI) apply (rule_tac x="gc_tmp_ref s\" and m=m in valid_refs_invD(3), auto simp: mut_m.tso_write_refs_def)[1] apply (clarsimp split: option.splits dest!: marked_insertionD) (* fM *) apply (erule disjE) apply (auto simp: fM_rel_def filter_empty_conv)[1] apply clarsimp done subgoal for s s' p w ws x y apply (clarsimp simp: do_write_action_def split: mem_write_action.splits) (* mark *) subgoal apply (drule_tac x=x in spec) apply (drule mp, erule predicate2D[OF rtranclp_mono[OF predicate2I], rotated]) apply clarsimp apply assumption done (* ref *) subgoal apply (clarsimp simp: atS_un) apply (erule disjE; clarsimp) apply (erule gc_marking_reaches_mw_Mutate; blast) done done (* mark loop mark field done *) subgoal apply (clarsimp simp: do_write_action_def split: mem_write_action.splits) (* mark *) apply fast (* fM *) apply (clarsimp simp: gc.handshake_invL_def loc atS_simps) apply (erule disjE) apply (auto simp: fM_rel_def filter_empty_conv)[1] apply clarsimp done done lemma reachable_sweep_loop_free: "mut_m.reachable m r (s(sys := s sys\heap := (sys_heap s)(r' := None)\)) \ mut_m.reachable m r s" apply (clarsimp simp: mut_m.reachable_def) apply (rename_tac x) apply (rule_tac x=x in exI, simp) apply (erule predicate2D[OF rtranclp_mono[OF predicate2I], rotated], clarsimp split: if_splits obj_at_splits) done lemma valid_refs_inv_sweep_loop_free[simp]: assumes "valid_refs_inv s" assumes ngr: "no_grey_refs s" assumes rsi: "\m'. mut_m.reachable_snapshot_inv m' s" assumes "white r' s" shows "valid_refs_inv (s(sys := s sys\heap := (sys_heap s)(r' := None)\))" using assms apply (clarsimp simp: valid_refs_inv_def grey_reachable_def no_grey_refs_def dest!: reachable_sweep_loop_free) apply (drule mut_m.reachable_blackD[OF ngr spec[OF rsi]]) apply (auto split: obj_at_splits) done lemma (in gc) valid_refs_inv[intro]: "\ fM_fA_invL \<^bold>\ handshake_invL \<^bold>\ gc_W_empty_invL \<^bold>\ gc_mark.mark_object_invL \<^bold>\ obj_fields_marked_invL \<^bold>\ phase_invL \<^bold>\ sweep_loop_invL \<^bold>\ LSTP (handshake_phase_inv \<^bold>\ mutators_phase_inv \<^bold>\ sys_phase_inv \<^bold>\ valid_refs_inv \<^bold>\ valid_W_inv) \ gc \ LSTP valid_refs_inv \" apply vcg_jackhammer (* sweep loop free *) apply (drule (1) handshake_in_syncD) apply clarsimp apply (auto simp: valid_refs_inv_def grey_reachable_def) done lemma (in sys) valid_refs_inv[intro]: "\ LSTP (valid_refs_inv \<^bold>\ tso_writes_inv) \ sys \ LSTP valid_refs_inv \" apply vcg_jackhammer apply (auto simp: do_write_action_def p_not_sys split: mem_write_action.splits) done lemma (in mut_m) valid_refs_inv_discard_roots[simp]: "\ valid_refs_inv s; roots' \ mut_roots s \ \ valid_refs_inv (s(mutator m := s (mutator m)\roots := roots'\))" by (auto simp: valid_refs_inv_def mut_m.reachable_def split: if_splits) lemma (in mut_m) valid_refs_inv_load[simp]: "\ valid_refs_inv s; sys_read (mutator m) (mr_Ref r f) (s sys) = mv_Ref r'; r \ mut_roots s \ \ valid_refs_inv (s(mutator m := s (mutator m)\roots := mut_roots s \ Option.set_option r'\))" by (auto simp: valid_refs_inv_def) lemma (in mut_m) valid_refs_inv_alloc[simp]: "\ valid_refs_inv s; sys_heap s r' = None \ \ valid_refs_inv (s(mutator m := s (mutator m)\roots := insert r' (mut_roots s)\, sys := s sys\heap := sys_heap s(r' \ \obj_mark = fl, obj_fields = Map.empty\)\))" apply (clarsimp simp: valid_refs_inv_def) apply (clarsimp simp: mut_m.reachable_def split: if_splits) apply (auto elim: converse_rtranclpE split: obj_at_splits) done lemma (in mut_m) valid_refs_inv_store_ins[simp]: "\ valid_refs_inv s; r \ mut_roots s; (\r'. opt_r' = Some r') \ the opt_r' \ mut_roots s \ \ valid_refs_inv (s(mutator m := s (mutator m)\ ghost_honorary_root := {} \, sys := s sys\ mem_write_buffers := (mem_write_buffers (s sys))(mutator m := sys_mem_write_buffers (mutator m) s @ [mw_Mutate r f opt_r']) \))" apply (subst valid_refs_inv_def) apply (clarsimp simp: grey_reachable_def mut_m.reachable_def) apply (rule conjI) apply clarsimp apply (rename_tac x xa) apply (case_tac "xa = m") apply clarsimp apply fastforce apply clarsimp apply (fastforce dest: valid_refs_invD) apply fastforce done lemma (in mut_m) valid_refs_inv_deref_del[simp]: "\ valid_refs_inv s; sys_read (mutator m) (mr_Ref r f) (s sys) = mv_Ref opt_r'; r \ mut_roots s; mut_ghost_honorary_root s = {} \ \ valid_refs_inv (s(mutator m := s (mutator m)\ghost_honorary_root := Option.set_option opt_r', ref := opt_r'\))" by (clarsimp simp: valid_refs_inv_def) lemma (in mut_m) valid_refs_inv[intro]: "\ mark_object_invL \<^bold>\ mut_get_roots.mark_object_invL m \<^bold>\ mut_store_del.mark_object_invL m \<^bold>\ mut_store_ins.mark_object_invL m \<^bold>\ LSTP valid_refs_inv \ mutator m \ LSTP valid_refs_inv \" apply vcg_jackhammer (* store ins mo co mark - FIXME some elim/dest rule really gets in the way here *) subgoal apply (subst valid_refs_inv_def) apply clarsimp apply (rule conjI) apply clarsimp apply (erule (1) valid_refs_invD) apply (clarsimp simp: grey_reachable_def) apply (erule disjE) apply (erule (1) valid_refs_invD, simp) apply clarsimp apply (erule (1) valid_refs_invD, simp) done (* store del mo co mark *) subgoal apply (subst valid_refs_inv_def) apply clarsimp apply (rule conjI) apply clarsimp apply (erule (1) valid_refs_invD) apply (clarsimp simp: grey_reachable_def) apply (erule disjE) apply (erule (1) valid_refs_invD, simp) apply clarsimp apply (erule (1) valid_refs_invD, simp) done (* get roots done *) subgoal by (clarsimp simp: valid_refs_inv_def grey_reachable_def) (* get roots loop mo co mark *) subgoal by (auto simp: valid_refs_inv_def grey_reachable_def) (* get work done *) subgoal by (clarsimp simp: valid_refs_inv_def grey_reachable_def) done (*>*) (*<*) end (*>*) diff --git a/thys/ConcurrentGC/TSO.thy b/thys/ConcurrentGC/TSO.thy --- a/thys/ConcurrentGC/TSO.thy +++ b/thys/ConcurrentGC/TSO.thy @@ -1,201 +1,203 @@ (*<*) (* * Copyright 2015, NICTA * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * @TAG(NICTA_BSD) *) theory TSO imports Proofs_basis begin +declare subst_all [simp del] [[simproc del: defined_all]] + (*>*) subsection\Coarse TSO invariants\ text\ Very coarse invariants about what processes write, and when they hold the TSO lock. \ abbreviation gc_writes :: "('field, 'ref) mem_write_action \ bool" where "gc_writes w \ case w of mw_Mark _ _ \ True | mw_Phase _ \ True | mw_fM _ \ True | mw_fA _ \ True | _ \ False" abbreviation mut_writes :: "('field, 'ref) mem_write_action \ bool" where "mut_writes w \ case w of mw_Mutate _ _ _ \ True | mw_Mark _ _ \ True | _ \ False" definition tso_writes_inv :: "('field, 'mut, 'ref) lsts_pred" where "tso_writes_inv \ (\<^bold>\w. tso_pending_write gc w \<^bold>\ \gc_writes w\) \<^bold>\ (\<^bold>\m w. tso_pending_write (mutator m) w \<^bold>\ \mut_writes w\)" (*<*) lemma tso_writes_inv_eq_imp: "eq_imp (\p s. mem_write_buffers (s sys) p) tso_writes_inv" by (simp add: eq_imp_def tso_writes_inv_def) lemmas tso_writes_inv_fun_upd[simp] = eq_imp_fun_upd[OF tso_writes_inv_eq_imp, simplified eq_imp_simps, rule_format] lemma tso_writes_invD[simp]: "tso_writes_inv s \ \sys_mem_write_buffers gc s = mw_Mutate r f r' # ws" "tso_writes_inv s \ \sys_mem_write_buffers (mutator m) s = mw_fA fl # ws" "tso_writes_inv s \ \sys_mem_write_buffers (mutator m) s = mw_fM fl # ws" "tso_writes_inv s \ \sys_mem_write_buffers (mutator m) s = mw_Phase ph # ws" by (auto simp: tso_writes_inv_def dest!: spec[where x=m]) lemma mut_do_write_action[simp]: "\ sys_mem_write_buffers (mutator m) s = w # ws; tso_writes_inv s \ \ fA (do_write_action w (s sys)) = sys_fA s" "\ sys_mem_write_buffers (mutator m) s = w # ws; tso_writes_inv s \ \ fM (do_write_action w (s sys)) = sys_fM s" "\ sys_mem_write_buffers (mutator m) s = w # ws; tso_writes_inv s \ \ phase (do_write_action w (s sys)) = sys_phase s" by (auto simp: do_write_action_def split: mem_write_action.splits) lemma tso_writes_inv_sys_read_Mut[simp]: assumes "tso_writes_inv s" assumes "(ract, v) \ { (mr_fM, mv_Mark (Some (sys_fM s))), (mr_fA, mv_Mark (Some (sys_fA s))), (mr_Phase, mv_Phase (sys_phase s)) }" shows "sys_read (mutator m) ract (s sys) = v" using assms apply (clarsimp simp: sys_read_def fold_writes_def) apply (rule fold_invariant[where P="\fr. do_read_action ract (fr (s sys)) = v" and Q="mut_writes"]) apply (fastforce simp: tso_writes_inv_def) apply (auto simp: do_read_action_def split: mem_write_action.splits) done lemma tso_writes_inv_sys_read_GC[simp]: assumes "tso_writes_inv s" shows "sys_read gc (mr_Ref r f) (s sys) = mv_Ref (sys_heap s r \ (\obj. obj_fields obj f))" (is "?lhs = mv_Ref ?rhs") using assms apply (simp add: sys_read_def fold_writes_def) apply (rule fold_invariant[where P="\fr. heap (fr (s sys)) r \ (\obj. obj_fields obj f) = ?rhs" and Q="\w. \r f r'. w \ mw_Mutate r f r'"]) apply (force simp: tso_writes_inv_def) apply (auto simp: do_write_action_def Option.map_option_case split: mem_write_action.splits option.splits) done lemma tso_no_pending_marksD[simp]: assumes "tso_pending_mark p s = []" shows "sys_read p (mr_Mark r) (s sys) = mv_Mark (Option.map_option obj_mark (sys_heap s r))" using assms apply (clarsimp simp: sys_read_def fold_writes_def) apply (rule fold_invariant[where P="\fr. Option.map_option obj_mark (heap (fr (s sys)) r) = Option.map_option obj_mark (sys_heap s r)" and Q="\w. \fl. w \ mw_Mark r fl"]) apply (auto simp: Option.map_option_case do_write_action_def filter_empty_conv split: mem_write_action.splits option.splits) done lemma no_pending_phase_sys_read[simp]: assumes "tso_pending_phase p s = []" shows "sys_read p mr_Phase (s sys) = mv_Phase (sys_phase s)" using assms apply (clarsimp simp: sys_read_def fold_writes_def) apply (rule fold_invariant[where P="\fr. phase (fr (s sys)) = sys_phase s" and Q="\w. \ph. w \ mw_Phase ph"]) apply (auto simp: do_write_action_def filter_empty_conv split: mem_write_action.splits) done lemma gc_no_pending_fM_write[simp]: assumes "tso_pending_fM gc s = []" shows "sys_read gc mr_fM (s sys) = mv_Mark (Some (sys_fM s))" using assms apply (clarsimp simp: sys_read_def fold_writes_def) apply (rule fold_invariant[where P="\fr. fM (fr (s sys)) = sys_fM s" and Q="\w. \fl. w \ mw_fM fl"]) apply (auto simp: do_write_action_def filter_empty_conv split: mem_write_action.splits) done lemma (in sys) tso_gc_writes_inv[intro]: "\ LSTP tso_writes_inv \ sys" apply (vcg_jackhammer simp: tso_writes_inv_def) apply (metis (no_types) list.set_intros(2)) done lemma (in gc) tso_writes_inv[intro]: "\ LSTP tso_writes_inv \ gc" by (vcg_jackhammer simp: tso_writes_inv_def) lemma (in mut_m) tso_writes_inv[intro]: "\ LSTP tso_writes_inv \ mutator m" by (vcg_jackhammer simp: tso_writes_inv_def) (*>*) subsubsection\Locations where the TSO lock is held\ text (in gc) \ The GC holds the TSO lock only during the \texttt{CAS} in @{const "mark_object"}. \ locset_definition gc_tso_lock_locs :: "location set" where "gc_tso_lock_locs \ \l\{ ''mo_co_cmark'', ''mo_co_ctest'', ''mo_co_mark'', ''mo_co_unlock'' }. suffixed l" inv_definition (in gc) tso_lock_invL :: "('field, 'mut, 'ref) gc_pred" where "tso_lock_invL \ atS_gc gc_tso_lock_locs (tso_locked_by gc) \<^bold>\ atS_gc (- gc_tso_lock_locs) (\<^bold>\ (tso_locked_by gc))" (*<*) lemma (in gc) tso_lock_invL[intro]: "\ tso_lock_invL \ gc" by vcg_jackhammer lemma (in sys) gc_tso_lock_invL[intro]: notes gc.tso_lock_invL_def[inv] shows "\ gc.tso_lock_invL \ sys" by vcg_ni lemma (in mut_m) gc_tso_lock_invL[intro]: notes gc.tso_lock_invL_def[inv] shows "\ gc.tso_lock_invL \ mutator m" by vcg_ni (*>*) text (in mut_m) \ A mutator holds the TSO lock only during the \texttt{CAS}s in @{const "mark_object"}. \ locset_definition "mut_tso_lock_locs = (\l\{ ''mo_co_cmark'', ''mo_co_ctest'', ''mo_co_mark'', ''mo_co_unlock'' }. suffixed l)" inv_definition (in mut_m) tso_lock_invL :: "('field, 'mut, 'ref) gc_pred" where "tso_lock_invL = (atS_mut mut_tso_lock_locs (tso_locked_by (mutator m)) \<^bold>\ atS_mut (- mut_tso_lock_locs) (\<^bold>\(tso_locked_by (mutator m))))" (*<*) lemma (in mut_m) tso_lock_invL[intro]: "\ tso_lock_invL \ mutator m" by vcg_jackhammer lemma (in gc) mut_tso_lock_invL[intro]: notes mut_m.tso_lock_invL_def[inv] shows "\ mut_m.tso_lock_invL m \ gc" by vcg_ni lemma (in sys) mut_tso_lock_invL[intro]: notes mut_m.tso_lock_invL_def[inv] shows "\ mut_m.tso_lock_invL m \ sys" by vcg_ni lemma (in mut_m') tso_lock_invL[intro]: "\ tso_lock_invL \ mutator m'" by vcg_ni (*>*) (*<*) end (*>*) diff --git a/thys/ConcurrentGC/Tactics.thy b/thys/ConcurrentGC/Tactics.thy --- a/thys/ConcurrentGC/Tactics.thy +++ b/thys/ConcurrentGC/Tactics.thy @@ -1,307 +1,309 @@ (*<*) (* * Copyright 2015, NICTA * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * @TAG(NICTA_BSD) *) theory Tactics imports Model "HOL-Library.Sublist" begin +declare subst_all [simp del] [[simproc del: defined_all]] + (*>*) section\Invariants and Proofs \label{sec:gc-invs}\ subsection\Constructors for sets of locations.\ abbreviation prefixed :: "location \ location set" where "prefixed p \ { l . prefix p l }" abbreviation suffixed :: "location \ location set" where "suffixed p \ { l . suffix p l }" subsection\Hoare triples\ text\ Specialise CIMP's pre/post validity to our system. \ definition valid_proc :: "('field, 'mut, 'ref) gc_pred \ 'mut process_name \ ('field, 'mut, 'ref) gc_pred \ bool" ("\_\ _ \_\") where "\P\ p \Q\ \ \(c, afts) \ vcg_fragments (gc_pgms p). (gc_pgms, p, afts \ \P\ c \Q\)" abbreviation valid_proc_inv_syn :: "('field, 'mut, 'ref) gc_pred \ 'mut process_name \ bool" ("\_\ _" [100,0] 100) where "\P\ p \ \P\ p \P\" (*<*) lemma valid_pre: assumes "\Q\ p \R\" assumes "\s. P s \ Q s" shows "\P\ p \R\" using assms apply (clarsimp simp: valid_proc_def) apply (drule (1) bspec) apply (auto elim: vcg_pre) done lemma valid_conj_lift: assumes x: "\P\ p \Q\" assumes y: "\P'\ p \Q'\" shows "\P \<^bold>\ P'\ p \Q \<^bold>\ Q'\" apply (clarsimp simp: valid_proc_def) apply (rule vcg_conj) apply (rule vcg_pre[OF spec[OF spec[OF x[unfolded Ball_def valid_proc_def split_paired_All]], simplified, rule_format]], simp, simp) apply (rule vcg_pre[OF spec[OF spec[OF y[unfolded Ball_def valid_proc_def split_paired_All]], simplified, rule_format]], simp, simp) done lemma valid_all_lift: assumes "\x. \P x\ p \Q x\" shows "\\s. \x. P x s\ p \\s. \x. Q x s\" using assms by (fastforce simp: valid_proc_def intro: vcg_all_lift) (*>*) text\ As we elide formal proofs in this document, we also omit our specialised proof tactics. These support essentially traditional local correctness and non-interference proofs. Their most interesting aspect is the use of Isabelle's parallelism to greatly reduce system latency. \ (*<*) subsection\Tactics\ named_theorems nie "Non-interference elimination rules" text\ Collect the component definitions. Inline everything. \ lemmas gc_defs = (* gc.com_def *) gc.handshake_done_def gc.handshake_init_def gc.handshake_noop_def gc.handshake_get_roots_def gc.handshake_get_work_def mark_object_fn_def lemmas mut_defs = (* mut.com_def *) mut_m.handshake_def mut_m.store_def mark_object_fn_def lemmas sys_defs = (* sys.com_def *) sys.alloc_def sys.free_def sys.mem_TSO_def sys.handshake_def lemmas gc_all_defs[com] = gc.com_def[simplified gc_defs append.simps if_True if_False] mut_m.com_def[simplified mut_defs append.simps if_True if_False] sys.com_def[simplified sys_defs append.simps if_True if_False] (* FIXME not automation friendly. *) lemma atS_dests: "\ atS p ls s; atS p ls' s \ \ atS p (ls \ ls') s" "\ \atS p ls s; \atS p ls' s \ \ \atS p (ls \ ls') s" "\ \atS p ls s; atS p ls' s \ \ atS p (ls' - ls) s" "\ \atS p ls s; at p l s \ \ atS p ({l} - ls) s" by (auto simp: atS_def) (* FIXME these leave \at ... assumptions when P is easy to contradict. *) lemma thin_locs: "\ at p l' s \ P; at p l s; l' \ l \ P \ Q \ \ Q" "\ atS p ls s \ P; at p l s; l \ ls \ P \ Q \ \ Q" (* "\ at p l' s \ P; atS p ls s; ... FIXME strategy: reduce atS to at and use the first of thin_locs. *) (* "\ atS p ls' s \ P; atS p ls s; \atS p (ls' \ ls) s \ P \ Q \ \ Q" *) by (auto simp: atS_def) text\ The following is unfortunately overspecialised to the GC. One might hope for general tactics that work on all CIMP programs. The system responds to all requests. The schematic variable is instantiated with the semantics of the responses. Thanks to Thomas Sewell for the hackery. \ schematic_goal system_responds_actionE: "\ (\l\ Response action, afts) \ fragments (gc_pgms p) \False\; v \ action x s; \ p = sys; ?P \ \ Q \ \ Q" apply (cases p) apply (simp_all add: gc_all_defs) apply atomize apply (drule_tac P="x \ y" and Q="v \ action p k" for x y p k in conjI, assumption) apply (thin_tac "v \ action p k" for p k) apply (simp only: conj_disj_distribR conj_assoc mem_Collect_eq cong: conj_cong) apply (erule mp) apply (thin_tac "p = sys") apply (assumption) done lemma triv: "P \ P" by simp schematic_goal system_responds_action_caseE: "\ (\l\ Response action, afts) \ fragments (gc_pgms p) \False\; v \ action (pname, req) s; \ p = sys; case_request_op ?P1 ?P2 ?P3 ?P4 ?P5 ?P6 ?P7 ?P8 ?P9 ?P10 ?P11 ?P12 ?P13 req \ \ Q \ \ Q" apply (erule(1) system_responds_actionE) apply (cases req, simp_all only: request_op.simps prod.inject simp_thms fst_conv snd_conv) apply (drule meta_mp[OF _ TrueI], erule meta_mp, erule_tac P="A \ B" for A B in triv)+ done schematic_goal system_responds_action_specE: "\ (\l\ Response action, afts) \ fragments (gc_pgms p) \False\; v \ action x s; \ p = sys; case_request_op ?P1 ?P2 ?P3 ?P4 ?P5 ?P6 ?P7 ?P8 ?P9 ?P10 ?P11 ?P12 ?P13 (snd x) \ \ Q \ \ Q" apply (erule system_responds_action_caseE[where pname="fst x" and req="snd x"]) apply simp apply assumption done (* Simplification rules for locations. *) lemmas loc_simps = bex_simps append.simps list.simps rev.simps (* evaluate string equality *) char.inject cong_exp_iff_simps (* evaluate character equality *) prefix_code suffix_to_prefix mem_Collect_eq Un_iff UNION_eq Compl_iff insertI1 insertI2 singleton_iff Diff_iff UNIV_I if_True if_False fun_upd_same fun_upd_other process_name.simps refl simp_thms atS_simps lemmas vcg_fragments'_simps = valid_proc_def gc_pgms.simps vcg_fragments'.simps atC.simps ball_Un bool_simps if_False if_True (* equations for deriving useful things from eq_imp facts. *) lemmas eq_imp_simps = eq_imp_def all_conj_distrib split_paired_All split_def fst_conv snd_conv prod_eq_iff conj_explode simp_thms (* Tweak the default simpset: - "not in dom" as a premise negates the goal - we always want to execute suffix - we may as well simplify under our non-recursive datatypes. *) declare dom_def[simp] declare suffix_to_prefix[simp] declare gc_phase.case_cong[cong] declare mem_write_action.case_cong[cong] declare process_name.case_cong[cong] declare handshake_phase.case_cong[cong] ML \ (* Thanks BNF people. *) fun ss_only thms ctxt = clear_simpset (put_simpset HOL_basic_ss ctxt) addsimps thms; (* Thanks BNF people. Actually use HOL_ss rather than HOL_basic_ss *) fun HOL_ss_only thms ctxt = clear_simpset (put_simpset HOL_ss ctxt) addsimps thms; fun vcg_clarsimp_tac ctxt = simp_tac (ss_only (@{thms vcg_fragments'_simps} @ Named_Theorems.get ctxt @{named_theorems com}) ctxt) THEN' SELECT_GOAL (safe_tac ctxt) val _ = Theory.setup (Method.setup @{binding vcg_clarsimp} (Method.sections clasimp_modifiers >> K (SIMPLE_METHOD' o vcg_clarsimp_tac)) "unfold com defs, execute vcg_fragments' and split goals") fun vcg_sem_tac ctxt = Tactic.match_tac ctxt @{thms vcg.intros} THEN' (TRY o (Tactic.ematch_tac ctxt @{thms system_responds_action_specE} THEN' assume_tac ctxt)) THEN' Rule_Insts.thin_tac ctxt "hist s = h" [(@{binding s}, NONE, NoSyn), (@{binding h}, NONE, NoSyn)] (* FIXME discard history: we don't use it here *) THEN' clarsimp_tac ctxt val _ = Theory.setup (Method.setup @{binding vcg_sem} (Method.sections clasimp_modifiers >> K (SIMPLE_METHOD' o vcg_sem_tac)) "turn VCG goal into semantics and clarsimp") fun vcg_jackhammer_gen_tac terminal_tac ctxt = SELECT_GOAL ( HEADGOAL (vcg_clarsimp_tac ctxt) THEN PARALLEL_ALLGOALS ( vcg_sem_tac ctxt THEN_ALL_NEW (full_simp_tac (Splitter.add_split @{thm lcond_split_asm} (ctxt addsimps Named_Theorems.get ctxt @{named_theorems inv}))) THEN_ALL_NEW ( (TRY o REPEAT_ALL_NEW (Tactic.ematch_tac ctxt @{thms conjE})) THEN' (TRY o REPEAT_ALL_NEW (Tactic.ematch_tac ctxt @{thms thin_locs} THEN' REPEAT1 o assume_tac ctxt)) THEN' asm_full_simp_tac (ss_only (@{thms loc_simps} @ Named_Theorems.get ctxt @{named_theorems loc}) ctxt) THEN_ALL_NEW (TRY o REPEAT_ALL_NEW (Rule_Insts.thin_tac ctxt "True" [])) (* FIXME weird, must be a standard way to do this. Leaving them in can cause simp to diverge ?? *) THEN_ALL_NEW clarsimp_tac (ctxt addsimps (Named_Theorems.get ctxt @{named_theorems loc} @ @{thms atS_simps})) (* FIXME smelly *) THEN_ALL_NEW Rule_Insts.thin_tac ctxt "AT _ = _" [] (* FIXME discard \AT s = s'(funupd)\ fact *) THEN_ALL_NEW TRY o terminal_tac ctxt))) val _ = Theory.setup (Method.setup @{binding vcg_jackhammer} (Method.sections clasimp_modifiers >> K (SIMPLE_METHOD' o vcg_jackhammer_gen_tac (fn _ => SELECT_GOAL all_tac))) "VCG supertactic, no terminal method") val _ = Theory.setup (Method.setup @{binding vcg_jackhammer_ff} (Method.sections clasimp_modifiers >> K (SIMPLE_METHOD' o vcg_jackhammer_gen_tac fast_force_tac)) "VCG supertactic, fastforce the survivors") fun vcg_ni_tac ctxt = SELECT_GOAL ( HEADGOAL (TRY o vcg_clarsimp_tac ctxt) THEN PARALLEL_ALLGOALS ( vcg_sem_tac ctxt THEN' (TRY o SELECT_GOAL (Local_Defs.unfold_tac ctxt (Named_Theorems.get ctxt @{named_theorems inv}))) THEN' (TRY o REPEAT_ALL_NEW (Tactic.match_tac ctxt @{thms conjI})) (* expose the location predicates, do not split the consequents *) THEN_ALL_NEW (TRY o REPEAT_ALL_NEW (Tactic.match_tac ctxt @{thms impI})) (* Preserve the label sets in atS but normalise the label in at; turn s' into s *) THEN_ALL_NEW asm_full_simp_tac ctxt (* FIXME diverges on some invariants *) THEN_ALL_NEW (TRY o REPEAT_ALL_NEW (Tactic.ematch_tac ctxt @{thms conjE})) (* The effect of vcg_pre: should be cheap *) THEN_ALL_NEW (TRY o REPEAT_ALL_NEW (Tactic.ematch_tac ctxt @{thms thin_locs} THEN' REPEAT1 o assume_tac ctxt)) THEN_ALL_NEW asm_full_simp_tac (ss_only (@{thms loc_simps} @ Named_Theorems.get ctxt @{named_theorems loc}) ctxt) THEN_ALL_NEW (TRY o REPEAT_ALL_NEW (Rule_Insts.thin_tac ctxt "True" [])) (* FIXME weird, must be a standard way to do this. Leaving them in can cause simp to diverge ?? *) (* THEN_ALL_NEW Rule_Insts.thin_tac ctxt "AT _ = _" [] (* FIXME discard \AT s = s'(funupd)\ fact *) doesn't work when processes communicate! see gc_sweep_loop_invL *) THEN_ALL_NEW clarsimp_tac ctxt)) fun vcg_nihe_tac ctxt = SELECT_GOAL ( HEADGOAL (vcg_clarsimp_tac ctxt) THEN PARALLEL_ALLGOALS ( (vcg_sem_tac ctxt THEN_ALL_NEW (Tactic.ematch_tac ctxt (Named_Theorems.get ctxt @{named_theorems nie}) THEN_ALL_NEW clarsimp_tac ctxt THEN_ALL_NEW SELECT_GOAL no_tac)) ORELSE' SELECT_GOAL all_tac)) (* FIXME perhaps replace with vcg_ni? but less diagnosable then *) val _ = Theory.setup (Method.setup @{binding vcg_ni} (Method.sections clasimp_modifiers >> K (SIMPLE_METHOD' o vcg_ni_tac)) "VCG non-interference supertactic, no terminal method") val _ = Theory.setup (Method.setup @{binding vcg_nihe} (Method.sections clasimp_modifiers >> K (SIMPLE_METHOD' o vcg_nihe_tac)) "cheap VCG non-interference tactic: apply non-interference Hoare and elimination rules, leaving remaining goals as Hoare triples") \ (*>*) (*<*) end (*>*) diff --git a/thys/ConcurrentGC/concrete/Concrete.thy b/thys/ConcurrentGC/concrete/Concrete.thy --- a/thys/ConcurrentGC/concrete/Concrete.thy +++ b/thys/ConcurrentGC/concrete/Concrete.thy @@ -1,77 +1,79 @@ (*<*) (* * Copyright 2015, NICTA * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * @TAG(NICTA_BSD) *) theory Concrete imports Concrete_heap begin +declare subst_all [simp del] [[simproc del: defined_all]] + (*>*) text\\ context gc_system begin abbreviation sys_init_state :: concrete_local_state where "sys_init_state \ undefined\ fA := initial_mark , fM := initial_mark , heap := sys_init_heap , handshake_pending := \False\ , handshake_type := ht_GetRoots , mem_lock := None , mem_write_buffers := \[]\ , phase := ph_Idle , W := {} , ghost_honorary_grey := {} , ghost_handshake_in_sync := \True\ , ghost_handshake_phase := hp_IdleMarkSweep \" abbreviation gc_init_state :: concrete_local_state where "gc_init_state \ undefined\ fM := initial_mark , fA := initial_mark , phase := ph_Idle , W := {} , ghost_honorary_grey := {} \" primrec lookup :: "('k \ 'v) list \ 'v \ 'k \ 'v" where "lookup [] v0 k = v0" | "lookup (kv # kvs) v0 k = (if fst kv = k then snd kv else lookup kvs v0 k)" abbreviation muts_init_states :: "(mut \ concrete_local_state) list" where "muts_init_states \ [ (0, mut_init_state0), (1, mut_init_state1), (2, mut_init_state2) ]" abbreviation init_state :: clsts where "init_state \ \p. case p of gc \ gc_init_state | sys \ sys_init_state | mutator m \ lookup muts_init_states mut_common_init_state m" lemma "gc_system_init init_state" (*<*) apply (clarsimp simp: gc_system_init_def gc_initial_state_def mut_initial_state_def sys_initial_state_def) apply (auto simp: ran_def) apply (auto simp: valid_refs_def) apply (erule rtranclp.cases; auto simp: ran_def split: if_splits obj_at_splits)+ done (*>*) text\\ end (*<*) end (*>*)