diff --git a/thys/Buchi_Complementation/code/analysis.py b/thys/Buchi_Complementation/code/analysis.py --- a/thys/Buchi_Complementation/code/analysis.py +++ b/thys/Buchi_Complementation/code/analysis.py @@ -1,45 +1,75 @@ #!/usr/bin/python -import sys, random, subprocess, time +import sys, random, subprocess, time, math def read(path): f = open(path, "r") text = f.read() f.close return text def write(path, text): f = open(path, "w") f.write(text) f.close() def stats(entries, finished): entry_count = "{}".format(len(entries)) finished_ratio = "{:.2f} \\%".format(100 * len(finished) / len(entries)) if entries else "N/A" finished_time = "{:.3f} s".format(sum(finished) / len(finished)) if finished else "N/A" return "{} & {} & {}".format(entry_count, finished_ratio, finished_time) def stats_entries(entries, label, predicate): entries = list(time for (state_count, result, time) in entries if predicate(state_count, result, time)) finished = list(float(time) for time in entries if time != "T") print("{} & {}\\\\".format(label, stats(entries, finished))) if sys.argv[1] == "states": entries = list(line.split() for line in read(sys.argv[2]).splitlines()) entries = list((int(state_count_1), int(state_count_2), result, time) for [state_count_1, state_count_2, result, time] in entries) non_timeout = list(time for (state_count_1, state_count_2, result, time) in entries if time != "T") print(all(result == "True" or time == "T" for (state_count_1, state_count_2, result, time) in entries)) print("completion rate: {}".format(len(non_timeout) / len(entries))) print("{} {}".format(sum(state_count_1 for (state_count_1, state_count_2, result, time) in entries) / len(entries), sum(state_count_2 for (state_count_1, state_count_2, result, time) in entries) / len(entries))) if sys.argv[1] == "complement": entries = list(line.split() for line in read(sys.argv[2]).splitlines()) entries = list(time for [state_count, time] in entries) non_timeout = list(float(time) for time in entries if time != "T") print(stats(entries, non_timeout)) if sys.argv[1] == "equivalence": - cutoff = 20 + segments = 4 + size = 5 entries = list(line.split() for line in read(sys.argv[2]).splitlines()) entries = list((max(int(state_count_1), int(state_count_2)), result, time) for [state_count_1, state_count_2, result, time] in entries) print(all(result == "True" or time == "T" for (state_count, result, time) in entries)) - for n in range(1, cutoff + 1): stats_entries(entries, n, lambda state_count, result, time: state_count == n) - stats_entries(entries, "> {}".format(cutoff), lambda state_count, result, time: state_count > cutoff) + for n in range(segments): stats_entries(entries, "$({}, {}]$".format(n * size, (n + 1) * size), lambda state_count, result, time: n * size < state_count and state_count <= (n + 1) * size) + stats_entries(entries, "$({}, \infty)$".format(segments * size), lambda state_count, result, time: segments * size < state_count) + print("\hline") stats_entries(entries, "total", lambda state_count, result, time: True) +if sys.argv[1] == "compare": + def read_block(blocks): + if blocks[0] == 'None': return (None, blocks[1:]) + return (int(blocks[0]), int(blocks[1]), float(blocks[2])), blocks[3:] + def parse(line): + blocks = line.replace(',', '').replace('(', '').replace(')', '').split() + state_count, blocks = int(blocks[0]), blocks[1:] + r1, blocks = read_block(blocks) + r2, blocks = read_block(blocks) + r3, blocks = read_block(blocks) + r4, blocks = read_block(blocks) + return r1, r2, r3, r4 + def stats(label, entries, finished, completed): + def round_(digits, value): return float(('%.' + str(digits) + 'g') % value) + def digits(digits, value): return round(value, digits -int(math.floor(math.log10(abs(value)))) - 1) + ratio = "{:.2f} \\%".format(100 * len(finished) / len(entries)) + time = "{:.3f} s".format(sum(time for _, _, time in completed) / len(completed)) + states = "{:g}".format(digits(4, sum(states for states, _, _ in completed) / len(completed))) + return "{} & {} & {} & {}\\\\".format(label, ratio, time, states) + entries = list(parse(line) for line in read(sys.argv[2]).splitlines()) + completed = list(entry for entry in entries if all(entry)) + print("finished by all: {} of {}".format(len(completed), len(entries))) + label = ["Spot (\\texttt{-{}-complement -{}-ba})", "\\textsc{Goal} (\\texttt{rank -tr})", "\\textsc{Goal} (\\texttt{rank -rd})", "Our Tool"] + for index in range(4): + block_entries = list(entry[index] for entry in entries) + block_finished = list(entry[index] for entry in entries if entry[index] is not None) + block_completed = list(entry[index] for entry in completed) + print(stats(label[index], block_entries, block_finished, block_completed)) diff --git a/thys/Buchi_Complementation/code/benchmark.py b/thys/Buchi_Complementation/code/benchmark.py --- a/thys/Buchi_Complementation/code/benchmark.py +++ b/thys/Buchi_Complementation/code/benchmark.py @@ -1,110 +1,121 @@ -#!/usr/bin/python +#!/usr/bin/python -u import sys, random, subprocess, time -owl = "/home/Projects/owl/build/distributions/owl-minimized-20.XX-development/bin/owl" - def read(path): - f = open(path, "r") - text = f.read() - f.close - return text + with open(path, "r") as f: return f.read() def write(path, text): - f = open(path, "w") - f.write(text) - f.close() + with open(path, "w") as f: f.write(text) + +def measure(action): + start = time.time() + result = action() + return result, time.time() - start def get_state_count(path): result = subprocess.run(["autfilt", path, "--stats=%s"], stdout = subprocess.PIPE, text = True) return int(result.stdout.strip()) +def get_transition_count(path): + result = subprocess.run(["autfilt", path, "--stats=%t"], stdout = subprocess.PIPE, text = True) + return int(result.stdout.strip()) def spot_generate_automaton(state_count, proposition_count): seed = random.randrange(0x10000) result = subprocess.run(["randaut", "--seed={}".format(seed), "--ba", "--states={}".format(state_count), str(proposition_count)], stdout = subprocess.PIPE, text = True) return result.stdout def spot_generate_formula(proposition_count): seed = random.randrange(0x10000) result = subprocess.run(["randltl", "--seed={}".format(seed), str(proposition_count)], stdout = subprocess.PIPE, text = True) return result.stdout.strip() def spot_simplify_automaton(path): result = subprocess.run(["autfilt", "--ba", "--small", path], stdout = subprocess.PIPE, text = True) return result.stdout def owl_simplify_automaton(path): result = subprocess.run(["owl", "-I", path, "---", "hoa", "---", "optimize-aut", "---", "hoa", "-s"], stdout = subprocess.PIPE, text = True) return result.stdout -def spot_to_buchi_automaton(path): - result = subprocess.run(["autfilt", "--ba", "--small", path], stdout = subprocess.PIPE, text = True, timeout = 60) - return result.stdout - def spot_formula_to_nba(formula): result = subprocess.run(["ltl2tgba", "--ba", formula], stdout = subprocess.PIPE, text = True) return result.stdout def owl_formula_to_nba(formula): result = subprocess.run(["owl", "-i", formula, "---", "ltl", "---", "simplify-ltl", "---", "ltl2nba", "---", "hoa", "-s"], stdout = subprocess.PIPE, text = True) return result.stdout def owl_formula_to_nba_opt(formula): result = subprocess.run(["owl", "-i", formula, "---", "ltl", "---", "simplify-ltl", "---", "ltl2nba", "---", "optimize-aut", "---", "hoa", "-s"], stdout = subprocess.PIPE, text = True) return result.stdout def owl_formula_to_dra(formula): result = subprocess.run(["owl", "-i", formula, "---", "ltl", "---", "simplify-ltl", "---", "ltl2dra", "---", "hoa", "-s"], stdout = subprocess.PIPE, text = True) return result.stdout def owl_formula_to_dra_opt(formula): result = subprocess.run(["owl", "-i", formula, "---", "ltl", "---", "simplify-ltl", "---", "ltl2dra", "---", "optimize-aut", "---", "hoa", "-s"], stdout = subprocess.PIPE, text = True) return result.stdout -def bc_complement(path1, path2): - subprocess.run(["./Autool", "complement_quick", path1, path2], timeout = 60) -def bc_equivalence(path1, path2): - result = subprocess.run(["./Autool", "equivalence", path1, path2], timeout = 60, stdout = subprocess.PIPE, text = True) - return result.stdout.strip() == "true" +def spot_complement(path_input, path_output): + try: _, time = measure(lambda: write(path_output, subprocess.run(["autfilt", "--complement", "--ba", path_input], timeout = 60, stdout = subprocess.PIPE, text = True).stdout)) + except subprocess.TimeoutExpired: return None + state_count = get_state_count(path_output) + transition_count = get_transition_count(path_output) + return state_count, transition_count, time -def complement(path_input, path_output): - start = time.time() - try: bc_complement(path_input, path_output) - except subprocess.TimeoutExpired: duration = "T" - else: duration = str(time.time() - start) - print("{} {}".format(get_state_count(path_input), duration)) +def goal_complement(path_input, path_output, tight): + try: _, time = measure(lambda: subprocess.run(["gc", "complement", "-m", "rank", "-tr" if tight else "-rd", "1", "-o", path_output, path_input], timeout = 60)) + except subprocess.TimeoutExpired: + subprocess.run(["killall", "java"]) + return None + state_count = int(subprocess.run(["gc", "stat", "-s", path_output], stdout = subprocess.PIPE, text = True).stdout.strip()) + transition_count = int(subprocess.run(["gc", "stat", "-t", path_output], stdout = subprocess.PIPE, text = True).stdout.strip()) + return state_count, transition_count, time + +def bc_complement(path_input, path_output): + try: _, time = measure(lambda: subprocess.run(["./Autool", "complement_quick", path_input, path_output], timeout = 60)) + except subprocess.TimeoutExpired: return None + complement = read(path_output).splitlines() + state_count = int(complement[0].split()[0]) + 1 + transition_count = len(complement) + return state_count, transition_count, time +def bc_equivalence(path1, path2): + try: result, time = measure(lambda: subprocess.run(["./Autool", "equivalence", path1, path2], timeout = 60, stdout = subprocess.PIPE, text = True)) + except subprocess.TimeoutExpired: return None + return result.stdout.strip() == "true", time + +def complement(path): + r1 = spot_complement(path, path + ".spot") + r2 = goal_complement(path, path + ".goaltr", True) + r3 = goal_complement(path, path + ".goalrd", False) + r4 = bc_complement(path, path + ".bc") + print("{} {} {} {} {}".format(get_state_count(path), r1, r2, r3, r4)) def equivalence(path_1, path_2): - start = time.time() - try: eq = bc_equivalence(path_1, path_2) - except subprocess.TimeoutExpired: eq = None; duration = "T" - else: duration = str(time.time() - start) - print("{} {} {} {}".format(get_state_count(path_1), get_state_count(path_2), eq, duration)) + r = bc_equivalence(path_1, path_2) + print("{} {} {}".format(get_state_count(path_1), get_state_count(path_2), r)) + if not r: return None + eq, _ = r return eq if sys.argv[1] == "complement_automaton": while True: write("a.hoa", spot_generate_automaton(20, 4)) - complement("a.hoa", "c.txt") + complement("a.hoa") if sys.argv[1] == "complement_formula": while True: - write("a.hoa", spot_formula_to_nba(spot_generate_formula(4))) - complement("a.hoa", "c.txt") + while True: + write("a.hoa", spot_formula_to_nba(spot_generate_formula(4))) + if get_state_count("a.hoa") == 10: break + complement("a.hoa") if sys.argv[1] == "equivalence_random": while True: write("a.hoa", spot_generate_automaton(100, 4)) write("b.hoa", spot_generate_automaton(100, 4)) equivalence("a.hoa", "b.hoa") if sys.argv[1] == "equivalence_simplify": while True: - write("a.hoa", spot_generate_automaton(10, 4)) + write("a.hoa", spot_generate_automaton(20, 4)) write("b.hoa", spot_simplify_automaton("a.hoa")) equivalence("a.hoa", "b.hoa") if sys.argv[1] == "equivalence_translate": while True: formula = spot_generate_formula(4) write("a.hoa", spot_formula_to_nba(formula)) write("b.hoa", owl_formula_to_dra_opt(formula)) write("c.hoa", spot_simplify_automaton("b.hoa")) if equivalence("a.hoa", "c.hoa") == False: print(formula) -if sys.argv[1] == "equivalence_translate_collection": - for formula in read(sys.argv[2]).splitlines(): - print(formula) - try: - write("a.hoa", spot_formula_to_nba(formula)) - write("b.hoa", owl_formula_to_dra_opt(formula)) - write("c.hoa", spot_to_buchi_automaton("b.hoa")) - equivalence("a.hoa", "c.hoa") - except subprocess.TimeoutExpired: print("T") diff --git a/thys/ConcurrentGC/Global_Invariants.thy b/thys/ConcurrentGC/Global_Invariants.thy new file mode 100644 --- /dev/null +++ b/thys/ConcurrentGC/Global_Invariants.thy @@ -0,0 +1,405 @@ +(*<*) +(* + * 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 Global_Invariants +imports + Proofs_Basis +begin + +(*>*) +section\Global Invariants \label{sec:global-invariants}\ + + +subsection\The valid references invariant\ + +text\ + +The key safety property of a GC is that it does not free objects that +are reachable from mutator roots. The GC also requires that there are +objects for all references reachable from grey objects. + +\ + +definition valid_refs_inv :: "('field, 'mut, 'payload, 'ref) lsts_pred" where + "valid_refs_inv = (\<^bold>\m x. mut_m.reachable m x \<^bold>\ grey_reachable x \<^bold>\ valid_ref x)" + +text\ + +The remainder of the invariants support the inductive argument that +this one holds. + +\ + + +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, 'payload, '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, 'payload, '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, 'payload, '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, 'payload, '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, 'payload, '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) (* FIXME elide *) + +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, 'payload, '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, 'payload, 'ref) lsts_pred" where + "reachable_snapshot_inv = (\<^bold>\r. reachable r \<^bold>\ in_snapshot r)" + + +subsection\Phase invariants \label{sec:phase-invariants}\ + +text (in mut_m) \ + +The phase structure of this GC algorithm greatly complicates this +safety proof. The following assertions capture this structure in +several relations. + +We begin by relating the mutators' @{const +"mut_ghost_hs_phase"} to @{const "sys_ghost_hs_phase"}, +which tracks the GC's. Each mutator can be at most one handshake step +behind the GC. If any mutator is behind then the GC is stalled on a +pending handshake. We include the handshake type as +\get_work\ can occur any number of times. + +\ + +definition hp_step_rel :: "(bool \ hs_type \ hs_phase \ hs_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, 'payload, 'ref) lsts_pred" where + "handshake_phase_inv = (\<^bold>\m. + sys_ghost_hs_in_sync m \<^bold>\ sys_hs_type \<^bold>\ sys_ghost_hs_phase \<^bold>\ mut_m.mut_ghost_hs_phase m \<^bold>\ \hp_step_rel\ + \<^bold>\ (sys_hs_pending m \<^bold>\ \<^bold>\sys_ghost_hs_in_sync m))" + +text \ + +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 the process doing the marking holds the TSO lock +until the mark is committed to the shared memory (see +\S\ref{def:valid_W_inv}). + +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. + +Read the following as ``when mutator \m\ is past the +specified handshake, and has yet to reach the next one, ... holds.'' + +\ + +abbreviation marked_insertion :: "('field, 'payload, 'ref) mem_store_action \ ('field, 'mut, 'payload, 'ref) lsts_pred" where + "marked_insertion w \ \s. case w of mw_Mutate r f (Some r') \ marked r' s | _ \ True" + +abbreviation marked_deletion :: "('field, 'payload, 'ref) mem_store_action \ ('field, 'mut, 'payload, '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" + +context mut_m +begin + +definition marked_insertions :: "('field, 'mut, 'payload, 'ref) lsts_pred" where + "marked_insertions = (\<^bold>\w. tso_pending_store (mutator m) w \<^bold>\ marked_insertion w)" + +definition marked_deletions :: "('field, 'mut, 'payload, 'ref) lsts_pred" where + "marked_deletions = (\<^bold>\w. tso_pending_store (mutator m) w \<^bold>\ marked_deletion w)" + +primrec mutator_phase_inv_aux :: "hs_phase \ ('field, 'mut, 'payload, '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 mutator_phase_inv :: "('field, 'mut, 'payload, 'ref) lsts_pred" where + "mutator_phase_inv \ mutator_phase_inv_aux \<^bold>$ mut_ghost_hs_phase" + +end + +abbreviation mutators_phase_inv :: "('field, 'mut, 'payload, '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 :: "hs_phase \ ('field, 'mut, 'payload, '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_store gc (mw_Phase ph_Idle)) \<^bold>\ no_grey_refs )" + +abbreviation sys_phase_inv :: "('field, 'mut, 'payload, 'ref) lsts_pred" where + "sys_phase_inv \ sys_phase_inv_aux \<^bold>$ sys_ghost_hs_phase" + + +subsubsection\Writes to shared GC variables\ + +text\ + +Relate @{const "sys_ghost_hs_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 :: "hs_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 \ hs_phase \ gc_phase \ gc_phase \ ('field, 'payload, 'ref) mem_store_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, 'payload, 'ref) lsts_pred" where + "phase_rel_inv = ((\<^bold>\m. sys_ghost_hs_in_sync m) \<^bold>\ sys_ghost_hs_phase \<^bold>\ gc_phase \<^bold>\ sys_phase \<^bold>\ tso_pending_phase gc \<^bold>\ \phase_rel\)" + +text\ + +Similarly we track the validity of @{const "sys_fM"} (respectively, +@{const "sys_fA"}) wrt @{const "gc_fM"} (@{const "sys_fA"}) and the +handshake phase. 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 \ hs_phase \ gc_mark \ gc_mark \ ('field, 'payload, 'ref) mem_store_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, 'payload, 'ref) lsts_pred" where + "fM_rel_inv = ((\<^bold>\m. sys_ghost_hs_in_sync m) \<^bold>\ sys_ghost_hs_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 \ hs_phase \ gc_mark \ gc_mark \ ('field, 'payload, 'ref) mem_store_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, 'payload, 'ref) lsts_pred" where + "fA_rel_inv = ((\<^bold>\m. sys_ghost_hs_in_sync m) \<^bold>\ sys_ghost_hs_phase \<^bold>\ sys_fA \<^bold>\ gc_fM \<^bold>\ tso_pending_fA gc \<^bold>\ \fA_rel\)" + + +subsection\Worklist invariants \label{def:valid_W_inv}\ + +text\ + +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. + +Note that the phase invariants of \S\ref{sec:phase-invariants} limit +the scope of this invariant. + +\ + +definition valid_W_inv :: "('field, 'mut, 'payload, 'ref) lsts_pred" where + "valid_W_inv = + ((\<^bold>\p r. r in_W p \<^bold>\ (sys_mem_lock \<^bold>\ \Some p\ \<^bold>\ r in_ghost_honorary_grey p) \<^bold>\ marked r) + \<^bold>\ (\<^bold>\p q. \p \ q\ \<^bold>\ WL p \<^bold>\ WL q \<^bold>= \{}\) + \<^bold>\ (\<^bold>\p q r. \<^bold>\(r in_ghost_honorary_grey p \<^bold>\ r in_W q)) + \<^bold>\ (EMPTY sys_ghost_honorary_grey) + \<^bold>\ (\<^bold>\p r fl. tso_pending_store 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]\ ))" + + +subsection\Coarse invariants about the stores a process can issue\ + +abbreviation gc_writes :: "('field, 'payload, 'ref) mem_store_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, 'payload, 'ref) mem_store_action \ bool" where + "mut_writes w \ case w of mw_Mutate _ _ _ \ True | mw_Mark _ _ \ True | _ \ False" + +definition tso_store_inv :: "('field, 'mut, 'payload, 'ref) lsts_pred" where + "tso_store_inv = + ((\<^bold>\w. tso_pending_store gc w \<^bold>\ \gc_writes w\) + \<^bold>\ (\<^bold>\m w. tso_pending_store (mutator m) w \<^bold>\ \mut_writes w\))" + + +subsection\The global invariants collected\ + +definition invs :: "('field, 'mut, 'payload, 'ref) lsts_pred" where + "invs = + (handshake_phase_inv + \<^bold>\ phase_rel_inv + \<^bold>\ strong_tricolour_inv + \<^bold>\ sys_phase_inv + \<^bold>\ tso_store_inv + \<^bold>\ valid_refs_inv + \<^bold>\ valid_W_inv + \<^bold>\ mutators_phase_inv + \<^bold>\ fA_rel_inv \<^bold>\ fM_rel_inv)" + + +subsection\Initial conditions \label{sec: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, 'payload, '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, 'payload, 'ref) lst_pred" where + "mut_initial_state s = + (ghost_hs_phase s = hp_IdleMarkSweep + \ ghost_honorary_grey s = {} + \ ghost_honorary_root s = {} + \ W s = {})" + +definition sys_initial_state :: "('field, 'mut, 'payload, 'ref) lst_pred" where + "sys_initial_state s = + ((\m. \hs_pending s m \ ghost_hs_in_sync s m) + \ ghost_hs_phase s = hp_IdleMarkSweep \ hs_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_store_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, 'payload, 'ref) lsts_pred" where + "valid_refs = (\<^bold>\y. root_reachable y \<^bold>\ valid_ref y)" + +definition gc_system_init :: "('field, 'mut, 'payload, '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, 'payload, 'ref) gc_system" where + "gc_system \ \PGMs = gc_coms, INIT = gc_system_init, FAIR = \True\\" (* FIXME add fairness hypotheses *) + +end + +(*<*) + +end +(*>*) diff --git a/thys/ConcurrentGC/Global_Invariants_Lemmas.thy b/thys/ConcurrentGC/Global_Invariants_Lemmas.thy new file mode 100644 --- /dev/null +++ b/thys/ConcurrentGC/Global_Invariants_Lemmas.thy @@ -0,0 +1,1285 @@ +(*<*) +(* + * 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 Global_Invariants_Lemmas +imports + Global_Invariants +begin + +declare subst_all [simp del] [[simproc del: defined_all]] + +(*>*) +section\Global invariants lemma bucket\ + +(* FIXME this file should be about the invs. This split makes it easier to move things around, maybe. *) + +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 + + +subsection\TSO invariants\ + +lemma tso_store_inv_eq_imp: + "eq_imp (\p s. mem_store_buffers (s sys) p) + tso_store_inv" +by (simp add: eq_imp_def tso_store_inv_def) + +lemmas tso_store_inv_fun_upd[simp] = eq_imp_fun_upd[OF tso_store_inv_eq_imp, simplified eq_imp_simps, rule_format] + +lemma tso_store_invD[simp]: + "tso_store_inv s \ \sys_mem_store_buffers gc s = mw_Mutate r f r' # ws" + "tso_store_inv s \ \sys_mem_store_buffers gc s = mw_Mutate_Payload r f pl # ws" + "tso_store_inv s \ \sys_mem_store_buffers (mutator m) s = mw_fA fl # ws" + "tso_store_inv s \ \sys_mem_store_buffers (mutator m) s = mw_fM fl # ws" + "tso_store_inv s \ \sys_mem_store_buffers (mutator m) s = mw_Phase ph # ws" +by (auto simp: tso_store_inv_def dest!: spec[where x=m]) + +lemma mut_do_store_action[simp]: + "\ sys_mem_store_buffers (mutator m) s = w # ws; tso_store_inv s \ \ fA (do_store_action w (s sys)) = sys_fA s" + "\ sys_mem_store_buffers (mutator m) s = w # ws; tso_store_inv s \ \ fM (do_store_action w (s sys)) = sys_fM s" + "\ sys_mem_store_buffers (mutator m) s = w # ws; tso_store_inv s \ \ phase (do_store_action w (s sys)) = sys_phase s" +by (auto simp: do_store_action_def split: mem_store_action.splits) + +lemma tso_store_inv_sys_load_Mut[simp]: + assumes "tso_store_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_load (mutator m) ract (s sys) = v" +using assms +apply (clarsimp simp: sys_load_def fold_stores_def) +apply (rule fold_invariant[where P="\fr. do_load_action ract (fr (s sys)) = v" and Q="mut_writes"]) + apply (fastforce simp: tso_store_inv_def) + apply (auto simp: do_load_action_def split: mem_store_action.splits) +done + +lemma tso_store_inv_sys_load_GC[simp]: + assumes "tso_store_inv s" + shows "sys_load gc (mr_Ref r f) (s sys) = mv_Ref (Option.bind (sys_heap s r) (\obj. obj_fields obj f))" (is "?lhs = mv_Ref ?rhs") +using assms unfolding sys_load_def fold_stores_def +apply clarsimp +apply (rule fold_invariant[where P="\fr. Option.bind (heap (fr (s sys)) r) (\obj. obj_fields obj f) = ?rhs" + and Q="\w. \r f r'. w \ mw_Mutate r f r'"]) + apply (fastforce simp: tso_store_inv_def) + apply (auto simp: do_store_action_def map_option_case fun_upd_apply + split: mem_store_action.splits option.splits) +done + +lemma tso_no_pending_marksD[simp]: + assumes "tso_pending_mark p s = []" + shows "sys_load p (mr_Mark r) (s sys) = mv_Mark (map_option obj_mark (sys_heap s r))" +using assms unfolding sys_load_def fold_stores_def +apply clarsimp +apply (rule fold_invariant[where P="\fr. map_option obj_mark (heap (fr (s sys)) r) = map_option obj_mark (sys_heap s r)" + and Q="\w. \fl. w \ mw_Mark r fl"]) + apply (auto simp: map_option_case do_store_action_def filter_empty_conv fun_upd_apply + split: mem_store_action.splits option.splits) +done + +lemma no_pending_phase_sys_load[simp]: + assumes "tso_pending_phase p s = []" + shows "sys_load p mr_Phase (s sys) = mv_Phase (sys_phase s)" +using assms +apply (clarsimp simp: sys_load_def fold_stores_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_store_action_def filter_empty_conv + split: mem_store_action.splits) +done + +lemma gc_no_pending_fM_write[simp]: + assumes "tso_pending_fM gc s = []" + shows "sys_load gc mr_fM (s sys) = mv_Mark (Some (sys_fM s))" +using assms +apply (clarsimp simp: sys_load_def fold_stores_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_store_action_def filter_empty_conv + split: mem_store_action.splits) +done + +lemma tso_store_refs_simps[simp]: + "mut_m.tso_store_refs m (s(mutator m' := s (mutator m')\roots := roots'\)) + = mut_m.tso_store_refs m s" + "mut_m.tso_store_refs m (s(mutator m' := s (mutator m')\ghost_honorary_root := {}\, + sys := s sys\mem_store_buffers := (mem_store_buffers (s sys))(mutator m' := sys_mem_store_buffers (mutator m') s @ [mw_Mutate r f opt_r'])\)) + = mut_m.tso_store_refs m s \ (if m' = m then store_refs (mw_Mutate r f opt_r') else {})" + "mut_m.tso_store_refs m (s(sys := s sys\mem_store_buffers := (mem_store_buffers (s sys))(mutator m' := sys_mem_store_buffers (mutator m') s @ [mw_Mutate_Payload r f pl])\)) + = mut_m.tso_store_refs m s \ (if m' = m then store_refs (mw_Mutate_Payload r f pl) else {})" + "mut_m.tso_store_refs m (s(sys := s sys\heap := (sys_heap s)(r' := None)\)) + = mut_m.tso_store_refs m s" + "mut_m.tso_store_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_store_refs m s" + "mut_m.tso_store_refs m (s(mutator m' := s (mutator m')\ghost_honorary_root := Option.set_option opt_r', ref := opt_r'\)) + = mut_m.tso_store_refs m s" + "mut_m.tso_store_refs m (s(sys := s sys\heap := (sys_heap s)(r := map_option (\obj. obj\obj_fields := (obj_fields obj)(f := opt_r')\) (sys_heap s r)), + mem_store_buffers := (mem_store_buffers (s sys))(p := ws)\)) + = (if p = mutator m then \w \ set ws. store_refs w else mut_m.tso_store_refs m s)" + "mut_m.tso_store_refs m (s(sys := s sys\heap := (sys_heap s)(r := map_option (\obj. obj\obj_payload := (obj_payload obj)(f := pl)\) (sys_heap s r)), + mem_store_buffers := (mem_store_buffers (s sys))(p := ws)\)) + = (if p = mutator m then \w \ set ws. store_refs w else mut_m.tso_store_refs m s)" + "sys_mem_store_buffers p s = mw_Mark r fl # ws +\ mut_m.tso_store_refs m (s(sys := s sys\heap := (sys_heap s)(r := map_option (obj_mark_update (\_. fl)) (sys_heap s r)), mem_store_buffers := (mem_store_buffers (s sys))(p := ws)\)) + = mut_m.tso_store_refs m s" +unfolding mut_m.tso_store_refs_def by (auto simp: fun_upd_apply) + +lemma fold_stores_points_to[rule_format, simplified conj_explode]: + "heap (fold_stores ws (s sys)) r = Some obj \ obj_fields obj f = Some r' + \ (r points_to r') s \ (\w \ set ws. r' \ store_refs w)" (is "?P (fold_stores ws) obj") +unfolding fold_stores_def +apply (rule spec[OF fold_invariant[where P="\fr. \obj. ?P fr obj" and Q="\w. w \ set ws"]]) + apply fastforce + apply (fastforce simp: ran_def split: obj_at_splits) +apply clarsimp +apply (drule (1) bspec) +apply (clarsimp simp: fun_upd_apply split: mem_store_action.split_asm if_splits) +done + +lemma points_to_Mutate: + "(x points_to y) + (s(sys := (s sys)\ heap := (sys_heap s)(r := map_option (\obj. obj\obj_fields := (obj_fields obj)(f := opt_r')\) (sys_heap s r)), + mem_store_buffers := (mem_store_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)))" +unfolding ran_def by (auto simp: fun_upd_apply split: obj_at_splits) + + +subsection\ FIXME mutator handshake facts \ + +lemma \ \Sanity\ + "hp' = hs_step hp \ \in' ht. (in', ht, hp', hp) \ hp_step_rel" +by (cases hp) (auto simp: hp_step_rel_def) + +lemma \ \Sanity\ + "(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: + assumes "handshake_phase_inv s" + shows "(sys_ghost_hs_in_sync m s, sys_hs_type s, sys_ghost_hs_phase s, mut_ghost_hs_phase s) \ hp_step_rel + \ (sys_hs_pending m s \ \sys_ghost_hs_in_sync m s)" +using assms unfolding handshake_phase_inv_def by simp + +lemma handshake_in_syncD: + "\ All (ghost_hs_in_sync (s sys)); handshake_phase_inv s \ + \ \m'. mut_m.mut_ghost_hs_phase m' s = sys_ghost_hs_phase s" +by clarsimp (auto simp: hp_step_rel_def dest!: mut_m.handshake_phase_invD) + +lemmas fM_rel_invD = iffD1[OF fun_cong[OF fM_rel_inv_def[simplified atomize_eq]]] + +text\ + +Relate @{const "sys_ghost_hs_phase"}, @{const "gc_phase"}, +@{const "sys_phase"} and writes to the phase in the GC's TSO buffer. + +\ + +simps_of_case handshake_phase_rel_simps[simp]: handshake_phase_rel_def (splits: hs_phase.split) + +lemma phase_rel_invD: + assumes "phase_rel_inv s" + shows "(\m. sys_ghost_hs_in_sync m s, sys_ghost_hs_phase s, gc_phase s, sys_phase s, tso_pending_phase gc s) \ phase_rel" +using assms unfolding phase_rel_inv_def by simp + +lemma mut_m_not_idle_no_fM_write: + "\ ghost_hs_phase (s (mutator m)) \ hp_Idle; fM_rel_inv s; handshake_phase_inv s; tso_store_inv s; p \ sys \ + \ \sys_mem_store_buffers p s = mw_fM fl # ws" +apply (drule mut_m.handshake_phase_invD[where m=m]) +apply (drule fM_rel_invD) +apply (clarsimp simp: hp_step_rel_def fM_rel_def filter_empty_conv p_not_sys) +apply (metis list.set_intros(1) tso_store_invD(4)) +done + +lemma (in mut_m) mut_ghost_handshake_phase_idle: + "\ mut_ghost_hs_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 + +lemma mut_m_not_idle_no_fM_writeD: + "\ sys_mem_store_buffers p s = mw_fM fl # ws; ghost_hs_phase (s (mutator m)) \ hp_Idle; fM_rel_inv s; handshake_phase_inv s; tso_store_inv s; p \ sys \ + \ False" +apply (drule mut_m.handshake_phase_invD[where m=m]) +apply (drule fM_rel_invD) +apply (clarsimp simp: hp_step_rel_def fM_rel_def filter_empty_conv p_not_sys) +apply (metis list.set_intros(1) tso_store_invD(4)) +done + + +subsection\points to, reaches, reachable mut\ + +lemma (in mut_m) reachable_eq_imp: + "eq_imp (\r'. mut_roots \<^bold>\ mut_ghost_honorary_root \<^bold>\ (\s. \(ran ` obj_fields ` set_option (sys_heap s r'))) + \<^bold>\ tso_pending_mutate (mutator m)) + (reachable r)" +unfolding eq_imp_def reachable_def tso_store_refs_def +apply clarsimp +apply (rename_tac s s') +apply (subgoal_tac "\r'. (\w\set (sys_mem_store_buffers (mutator m) s). r' \ store_refs w) \ (\w\set (sys_mem_store_buffers (mutator m) s'). r' \ store_refs w)") + apply (subgoal_tac "\x. (x reaches r) s \ (x reaches r) s'") + apply (clarsimp; fail) + apply (auto simp: reaches_fields; fail)[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; fail) + apply (clarsimp split: mem_store_action.splits; fail) +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; fail) +apply (clarsimp split: mem_store_action.splits; fail) +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_store_refs m s \ mut_m.reachable m x s" +by (auto simp: mut_m.reachable_def reaches_def) + +lemma reachable_points_to[elim]: + "\ (x points_to y) s; mut_m.reachable m x s \ \ mut_m.reachable m y s" +by (auto simp: mut_m.reachable_def reaches_def elim: rtranclp.intros(2)) + +lemma (in mut_m) mut_reachableE[consumes 1, case_names mut_root tso_store_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_store_refs s \ \ Q \ \ Q" +by (auto simp: reachable_def) + +lemma reachable_induct[consumes 1, case_names root ghost_honorary_root tso_root reaches]: + assumes r: "mut_m.reachable m y s" + assumes root: "\x. \ x \ mut_m.mut_roots m s \ \ P x" + assumes ghost_honorary_root: "\x. \ x \ mut_m.mut_ghost_honorary_root m s \ \ P x" + assumes tso_root: "\x. x \ mut_m.tso_store_refs m s \ P x" + assumes reaches: "\x y. \ mut_m.reachable m x s; (x points_to y) s; P x \ \ P y" + shows "P y" +using r unfolding mut_m.reachable_def +proof(clarify) + fix x + assume "(x reaches y) s" and "x \ mut_m.mut_roots m s \ mut_m.mut_ghost_honorary_root m s \ mut_m.tso_store_refs m s" + then show "P y" + unfolding reaches_def 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 mut_m.reachable_def reaches_def by meson + qed +qed + +lemma mutator_reachable_tso: + "sys_mem_store_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)" + "sys_mem_store_buffers (mutator m) s = mw_Mutate_Payload r f pl # ws + \ mut_m.reachable m r s" +by (auto simp: mut_m.tso_store_refs_def) + + +subsection\ Colours \ + +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" +unfolding grey_def WL_def by (case_tac [!] p) auto + +lemma blackD[dest]: + "black r s \ marked r s" + "black r s \ r \ WL p s" +unfolding black_def grey_def by simp_all + +lemma whiteI[intro]: (* FIXME simp normal form of def *) + "obj_at (\obj. obj_mark obj = (\ sys_fM s)) r s \ white r s" +unfolding white_def by simp + +lemma marked_not_white[dest]: + "white r s \ \marked r s" +unfolding white_def by (simp_all split: obj_at_splits) + +lemma white_valid_ref[elim!]: + "white r s \ valid_ref r s" +unfolding white_def by (simp_all split: obj_at_splits) + +lemma not_white_marked[elim!]: + "\\ white r s; valid_ref r s\ \ marked r s" +unfolding white_def by (simp split: obj_at_splits) + +lemma black_eq_imp: + "eq_imp (\_::unit. (\s. r \ (\p. WL p s)) \<^bold>\ sys_fM \<^bold>\ (\s. map_option obj_mark (sys_heap s r))) + (black r)" +unfolding eq_imp_def black_def grey_def by (auto split: obj_at_splits) + +lemma grey_eq_imp: + "eq_imp (\_::unit. (\s. r \ (\p. WL p s))) + (grey r)" +unfolding eq_imp_def grey_def by auto + +lemma white_eq_imp: + "eq_imp (\_::unit. sys_fM \<^bold>\ (\s. map_option obj_mark (sys_heap s r))) + (white r)" +unfolding eq_imp_def white_def by (auto split: obj_at_splits) + +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\coloured heaps\ + +lemma black_heap_eq_imp: + "eq_imp (\r'. (\s. \p. WL p s) \<^bold>\ sys_fM \<^bold>\ (\s. map_option obj_mark (sys_heap s r'))) + black_heap" +apply (clarsimp simp: 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="map_option \True\" in arg_cong) + apply fastforce +apply clarsimp +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="map_option (\fl. fl = sys_fM s)" in arg_cong) +apply simp +done + +lemma white_heap_eq_imp: + "eq_imp (\r'. sys_fM \<^bold>\ (\s. map_option obj_mark (sys_heap s r'))) + white_heap" +apply (clarsimp simp: all_conj_distrib eq_imp_def white_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 (\r'. (\s. (\p. WL p s)) \<^bold>\ sys_fM \<^bold>\ (\s. map_option obj_mark (sys_heap s r'))) + 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 split: obj_at_splits) +apply (rename_tac x) +apply (drule_tac x=x in spec)+ +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_def 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 black_heap_no_greys[elim]: + "\ no_grey_refs s; \r. marked r s \ \valid_ref r s \ \ black_heap s" +unfolding black_def black_heap_def no_grey_refs_def by fastforce + +lemma heap_colours_colours: + "black_heap s \ \white r s" + "white_heap s \ \black r s" +by (auto simp: black_heap_def white_def white_heap_def + dest!: spec[where x=r] + split: obj_at_splits) + + +text\The strong-tricolour invariant \ + +lemma strong_tricolour_invD: + "\ black x s; (x points_to y) s; valid_ref y s; strong_tricolour_inv s \ + \ marked y s" +unfolding strong_tricolour_inv_def by fastforce + +lemma no_black_refsD: + "no_black_refs s \ \black r s" +unfolding no_black_refs_def by simp + +lemma has_white_path_to_induct[consumes 1, case_names refl step, induct set: has_white_path_to]: + assumes "(x has_white_path_to y) s" + assumes "\x. P x x" + assumes "\x y z. \(x has_white_path_to y) s; P x y; (y points_to z) s; white z s\ \ P x z" + shows "P x y" +using assms unfolding has_white_path_to_def by (rule rtranclp.induct; blast) + +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_to_refl[iff]: + "(x has_white_path_to x) s" +unfolding has_white_path_to_def by simp + +lemma has_white_path_to_step[intro]: + "\(x has_white_path_to y) s; (y points_to z) s; white z s\ \ (x has_white_path_to z) s" + "\(y has_white_path_to z) s; (x points_to y) s; white y s\ \ (x has_white_path_to z) s" +unfolding has_white_path_to_def + apply (simp add: rtranclp.rtrancl_into_rtrancl) +apply (simp add: converse_rtranclp_into_rtranclp) +done + +lemma has_white_path_toE[elim!]: + "\ (x points_to y) s; white y s \ \ (x has_white_path_to y) s" +unfolding has_white_path_to_def by (auto 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 reaches_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" +unfolding has_white_path_to_def by (simp add: fun_upd_apply) + +lemma has_white_path_to_eq_imp': \ \Complicated condition takes care of \alloc\: collapses no object and object with no fields\ + assumes "(x has_white_path_to y) s'" + assumes "\r'. \(ran ` obj_fields ` set_option (sys_heap s' r')) = \(ran ` obj_fields ` set_option (sys_heap s r'))" + assumes "\r'. map_option obj_mark (sys_heap s' r') = map_option obj_mark (sys_heap s r')" + assumes "sys_fM s' = sys_fM s" + shows "(x has_white_path_to y) s" +using assms +proof induct + case (step x y z) + then have "(y points_to z) s" + by (cases "sys_heap s y") + (auto 10 10 simp: ran_def obj_at_def split: option.splits dest!: spec[where x=y]) + with step show ?case +apply - +apply (rule has_white_path_to_step, assumption, assumption) +apply (clarsimp simp: white_def split: obj_at_splits) +apply (metis map_option_eq_Some option.sel) +done +qed simp + +lemma has_white_path_to_eq_imp: + "eq_imp (\r'. sys_fM \<^bold>\ (\s. \(ran ` obj_fields ` set_option (sys_heap s r'))) \<^bold>\ (\s. map_option obj_mark (sys_heap s r'))) + (x has_white_path_to y)" +unfolding eq_imp_def +apply (clarsimp simp: all_conj_distrib) +apply (rule iffI) + apply (erule has_white_path_to_eq_imp'; auto) +apply (erule has_white_path_to_eq_imp'; auto) +done + +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] + + +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_induct[consumes 1, case_names refl step, induct set: grey_protects_white]: + assumes "(g grey_protects_white w) s" + assumes "\x. grey x s \ P x x" + assumes "\x y z. \(x has_white_path_to y) s; P x y; (y points_to z) s; white z s\ \ P x z" + shows "P g w" +using assms unfolding grey_protects_white_def +apply - +apply (elim conjE) +apply (rotate_tac -1) +apply (induct rule: has_white_path_to_induct) +apply blast+ +done + + +subsection\ @{term "valid_W_inv"} \ + +lemma valid_W_inv_sys_ghg_empty_iff[elim!]: + "valid_W_inv s \ sys_ghost_honorary_grey s = {}" +unfolding valid_W_inv_def by simp + +lemma WLI[intro]: + "r \ W (s p) \ r \ WL p s" + "r \ ghost_honorary_grey (s p) \ r \ WL p s" +unfolding WL_def by simp_all + +lemma WL_eq_imp: + "eq_imp (\(_::unit) s. (ghost_honorary_grey (s p), W (s p))) + (WL p)" +unfolding eq_imp_def WL_def by simp + +lemmas WL_fun_upd[simp] = eq_imp_fun_upd[OF WL_eq_imp, simplified eq_imp_simps, rule_format] + +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. 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 white_def) +apply (rename_tac s s') +apply (subgoal_tac "\p. WL p s = WL p s'") + apply (subgoal_tac "\x. marked x s \ marked 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_store_buffers x s) \ mw_Mark xa xb \ set (sys_mem_store_buffers x s')") + apply (simp; fail) + 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_invE[elim!]: + "\ r \ W (s p); valid_W_inv s \ \ marked r s" + "\ r \ ghost_honorary_grey (s p); sys_mem_lock s \ Some p; valid_W_inv s \ \ marked r s" + "\ r \ W (s p); valid_W_inv s \ \ valid_ref r s" + "\ r \ ghost_honorary_grey (s p); sys_mem_lock s \ Some p; valid_W_inv s \ \ valid_ref r s" + "\ mw_Mark r fl \ set (sys_mem_store_buffers p s); valid_W_inv s \ \ r \ ghost_honorary_grey (s p)" +unfolding valid_W_inv_def +apply (simp_all add: split: obj_at_splits) +apply blast+ +done + +lemma valid_W_invD: + "\ sys_mem_store_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_store_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_store_buffers p s) = [mw_Mark r fl]" +unfolding valid_W_inv_def white_def by (clarsimp dest!: spec[where x=p], blast)+ + +lemma valid_W_inv_colours: + "\white x s; valid_W_inv s\ \ x \ W (s p)" +using marked_not_white valid_W_invE(1) by force + +lemma valid_W_inv_no_mark_stores_invD: + "\ sys_mem_lock s \ Some p; valid_W_inv s \ + \ tso_pending p is_mw_Mark s = []" +by (auto dest: valid_W_invD(2) intro!: filter_False) + +lemma valid_W_inv_sys_load[simp]: + "\ sys_mem_lock s \ Some p; valid_W_inv s \ + \ sys_load p (mr_Mark r) (s sys) = mv_Mark (map_option obj_mark (sys_heap s r))" +unfolding sys_load_def fold_stores_def +apply clarsimp +apply (rule fold_invariant[where P="\fr. map_option obj_mark (heap (fr (s sys)) r) = map_option obj_mark (sys_heap s r)" + and Q="\w. \r fl. w \ mw_Mark r fl"]) + apply (auto simp: map_option_case do_store_action_def filter_empty_conv fun_upd_apply + dest: valid_W_invD(2) + split: mem_store_action.splits option.splits) +done + + +subsection\ \grey_reachable\ \ + +lemma grey_reachable_eq_imp: + "eq_imp (\r'. (\s. \p. WL p s) \<^bold>\ (\s. Set.bind (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" +unfolding grey_reachable_def reaches_def by blast + +lemma grey_reachableE: + "\ (g points_to y) s; grey_reachable g s \ \ grey_reachable y s" +unfolding grey_reachable_def reaches_def by (auto elim: rtranclp.intros(2)) + + +subsection\valid refs inv\ + +lemma valid_refs_invI: + "\ \m x y. \ (x reaches y) s; mut_m.root m x 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 valid_refs_inv_eq_imp: + "eq_imp (\(m', r'). (\s. roots (s (mutator m'))) \<^bold>\ (\s. ghost_honorary_root (s (mutator m'))) \<^bold>\ (\s. 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 "\r'. \(ran ` obj_fields ` set_option (sys_heap s r')) = \(ran ` obj_fields ` set_option (sys_heap s' r'))") + 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 (metis option.set_map) + apply (clarsimp split: obj_at_splits) + apply (metis (no_types, hide_lams) None_eq_map_option_iff option.exhaust) + 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_store_refs m s; (x reaches y) s; valid_refs_inv s \ \ valid_ref y s" + "\ x \ mut_m.tso_store_refs m s; (x reaches y) s; valid_refs_inv s \ \ \obj. sys_heap s y = Some obj" + "\ w \ set (sys_mem_store_buffers (mutator m) s); x \ store_refs w; (x reaches y) s; valid_refs_inv s \ \ valid_ref y s" + "\ w \ set (sys_mem_store_buffers (mutator m) s); x \ store_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" +apply (simp_all add: valid_refs_inv_def grey_reachable_def mut_m.reachable_def mut_m.tso_store_refs_def + split: obj_at_splits) +apply blast+ +done + + +text\reachable snapshot inv\ + +context mut_m +begin + +lemma reachable_snapshot_invI[intro]: + "(\y. reachable y s \ in_snapshot y s) \ reachable_snapshot_inv s" +by (simp add: reachable_snapshot_inv_def) + +lemma 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>\ (\s. \(ran ` obj_fields ` set_option (sys_heap s r'))) \<^bold>\ (\s. map_option obj_mark (sys_heap s r')) + \<^bold>\ tso_pending_mutate (mutator m)) + reachable_snapshot_inv" +unfolding eq_imp_def mut_m.reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def black_def grey_def +apply (clarsimp simp: all_conj_distrib) +apply (rename_tac s s') +apply (subst (1) eq_impD[OF has_white_path_to_eq_imp]) + apply force +apply (subst eq_impD[OF reachable_eq_imp]) + apply force +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 force +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) +done + +end + +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 \ \Sanity\ + "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 (metis blackD(1) grey_protects_whiteD grey_protects_white_reaches in_snapshot_def obj_at_cong obj_at_def option.case(2) valid_refs_invD(7)) + +lemma reachableI2[intro]: + "x \ mut_m.mut_ghost_honorary_root m s \ mut_m.reachable m x s" +unfolding mut_m.reachable_def reaches_def by blast + +lemma tso_pending_mw_mutate_cong: + "\ filter is_mw_Mutate (sys_mem_store_buffers p s) = filter is_mw_Mutate (sys_mem_store_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_store_buffers p s) \ P r f r') + \ (\r f r'. mw_Mutate r f r' \ set (sys_mem_store_buffers p s') \ Q r f r')" +by (intro iff_allI) (auto dest!: arg_cong[where f=set]) + +lemma (in mut_m) marked_insertions_eq_imp: + "eq_imp (\r'. sys_fM \<^bold>\ (\s. map_option obj_mark (sys_heap s r')) \<^bold>\ tso_pending_mw_mutate (mutator m)) + marked_insertions" +unfolding eq_imp_def marked_insertions_def obj_at_def +apply (clarsimp split: mem_store_action.splits) +apply (erule tso_pending_mw_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_store_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 marked_insertions_store_buffer_empty[intro]: + "tso_pending_mutate (mutator m) s = [] \ mut_m.marked_insertions m s" +unfolding mut_m.marked_insertions_def by (auto simp: filter_empty_conv split: mem_store_action.splits) + +(* marked_deletions *) + +lemma (in mut_m) marked_deletions_eq_imp: + "eq_imp (\r'. sys_fM \<^bold>\ (\s. map_option obj_fields (sys_heap s r')) \<^bold>\ (\s. map_option obj_mark (sys_heap s r')) \<^bold>\ tso_pending_mw_mutate (mutator m)) + marked_deletions" +unfolding eq_imp_def marked_deletions_def obj_at_field_on_heap_def ran_def +apply (clarsimp simp: all_conj_distrib) +apply (drule arg_cong[where f=set]) +apply (subgoal_tac "\x. marked x s \ marked x s'") + apply (clarsimp cong: option.case_cong) + apply (rule iffI; clarsimp simp: set_eq_iff split: option.splits mem_store_action.splits; blast) +apply clarsimp +apply (rename_tac s s' x) +apply (drule_tac x=x in spec)+ +apply (force split: obj_at_splits) +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 marked_deletions_store_buffer_empty[intro]: + "tso_pending_mutate (mutator m) s = [] \ mut_m.marked_deletions m s" +unfolding mut_m.marked_deletions_def by (auto simp: filter_empty_conv split: mem_store_action.splits) + + +subsection\Location-specific simplification rules\ + +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 simp: fun_upd_apply split: obj_at_splits) + +lemma obj_at_alloc[simp]: + "sys_heap s r' = None + \ obj_at P r (s(m := mut_m_s', sys := (s sys)\ heap := sys_heap s(r' \ obj) \)) + \ (obj_at P r s \ (r = r' \ P obj))" +unfolding ran_def by (simp add: fun_upd_apply split: obj_at_splits) + +lemma valid_ref_valid_null_ref_simps[simp]: + "valid_ref r (s(sys := do_store_action w (s sys)\mem_store_buffers := (mem_store_buffers (s sys))(p := ws)\)) \ valid_ref r s" + "valid_null_ref r' (s(sys := do_store_action w (s sys)\mem_store_buffers := (mem_store_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''" +unfolding do_store_action_def valid_null_ref_def +by (auto simp: fun_upd_apply + split: mem_store_action.splits obj_at_splits option.splits) + +context mut_m +begin + +lemma reachable_load[simp]: + assumes "sys_load (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 +unfolding sys_load_def +apply clarsimp +apply (clarsimp simp: reachable_def reaches_def tso_store_refs_def sys_load_def fold_stores_def fun_upd_apply) +apply (elim disjE) + apply blast + defer + apply blast +apply blast + +apply (fold fold_stores_def) +apply clarsimp +apply (drule (1) fold_stores_points_to) +apply (erule disjE) + apply (fastforce elim!: converse_rtranclp_into_rtranclp[rotated] split: obj_at_splits intro!: ranI) +apply (clarsimp split: mem_store_action.splits) +apply meson +done + next + assume ?rhs with True show ?lhs unfolding mut_m.reachable_def by (fastforce simp: fun_upd_apply) + qed +qed (simp add: fun_upd_apply) + +end + +text\WL\ + +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 }" +unfolding WL_def by (auto simp: fun_upd_apply) + +lemma WL_hs_done[simp]: + "ghost_honorary_grey (s (mutator m)) = {} + \ WL p (s(mutator m := s (mutator m)\ W := {}, ghost_hs_phase := hp' \, + sys := s sys\ hs_pending := hsp', W := sys_W s \ W (s (mutator m)), + ghost_hs_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\ hs_pending := hsp', W := sys_W s \ W (s (mutator m)), + ghost_hs_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)" +unfolding WL_def by (auto simp: fun_upd_apply 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" +unfolding black_def grey_def WL_def +apply (simp_all add: fun_upd_apply) +apply safe +apply (case_tac [!] x) +apply blast+ +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)" +unfolding WL_def by (auto simp: fun_upd_apply split: process_name.splits) + +text\no grey refs\ + +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" +unfolding no_grey_refs_def by (auto intro!: filter_False dest: valid_W_invD(2)) + +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 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 white_def split: obj_at_splits) + +context mut_m +begin + +lemma 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 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 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_store_refs s" +apply (drule (2) no_grey_refs_not_reachable) +apply (force simp: reachable_def reaches_def) +done + +lemma 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" +unfolding reachable_snapshot_inv_def in_snapshot_def reachable_def grey_protects_white_def reaches_def by auto + +end + +(* colours *) + +lemma black_dequeue_Mark[simp]: + "black b (s(sys := (s sys)\ heap := (sys_heap s)(r := map_option (obj_mark_update (\_. fl)) (sys_heap s r)), mem_store_buffers := (mem_store_buffers (s sys))(p := ws) \)) +\ (black b s \ b \ r) \ (b = r \ fl = sys_fM s \ valid_ref r s \ \grey r s)" +unfolding black_def by (auto simp: fun_upd_apply split: obj_at_splits) + +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')" +unfolding black_def grey_def white_def by (auto simp: fun_upd_apply split: obj_at_splits) + +lemma colours_get_work_done[simp]: + "black r (s(mutator m := (s (mutator m))\W := {}\, + sys := (s sys)\ hs_pending := hp', W := W (s sys) \ W (s (mutator m)), + ghost_hs_in_sync := his' \)) \ black r s" + "grey r (s(mutator m := (s (mutator m))\W := {}\, + sys := (s sys)\ hs_pending := hp', W := W (s sys) \ W (s (mutator m)), + ghost_hs_in_sync := his' \)) \ grey r s" + "white r (s(mutator m := (s (mutator m))\W := {}\, + sys := (s sys)\ hs_pending := hp', W := W (s sys) \ W (s (mutator m)), + ghost_hs_in_sync := his' \)) \ white r s" +unfolding black_def grey_def WL_def +apply (simp_all add: fun_upd_apply split: obj_at_splits) + apply blast +apply (metis process_name.distinct(3)) +done + +lemma colours_get_roots_done[simp]: + "black r (s(mutator m := (s (mutator m))\ W := {}, ghost_hs_phase := hs' \, + sys := (s sys)\ hs_pending := hp', W := W (s sys) \ W (s (mutator m)), + ghost_hs_in_sync := his' \)) \ black r s" + "grey r (s(mutator m := (s (mutator m))\ W := {}, ghost_hs_phase := hs' \, + sys := (s sys)\ hs_pending := hp', W := W (s sys) \ W (s (mutator m)), + ghost_hs_in_sync := his' \)) \ grey r s" + "white r (s(mutator m := (s (mutator m))\ W := {}, ghost_hs_phase := hs' \, + sys := (s sys)\ hs_pending := hp', W := W (s sys) \ W (s (mutator m)), + ghost_hs_in_sync := his' \)) \ white r s" +unfolding black_def grey_def WL_def +apply (simp_all add: fun_upd_apply split: obj_at_splits) + apply blast +apply (metis process_name.distinct(3)) +done + +lemma colours_flip_fM[simp]: + "fl \ sys_fM s \ black b (s(sys := (s sys)\fM := fl, mem_store_buffers := (mem_store_buffers (s sys))(p := ws)\)) \ white b s \ \grey b s" +unfolding black_def white_def by (simp add: fun_upd_apply) + +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, obj_payload = 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, obj_payload = 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, obj_payload = Map.empty\)\)) + \ white r s \ (r' = r \ fl \ sys_fM s)" +unfolding black_def white_def by (auto simp: fun_upd_apply split: obj_at_splits) + +lemma heap_colours_alloc[simp]: + "\ heap (s sys) r' = None; valid_refs_inv s \ + \ black_heap (s(mutator m := s (mutator m)\roots := roots'\, sys := s sys\heap := sys_heap s(r' \ \obj_mark = fl, obj_fields = Map.empty, obj_payload = Map.empty\)\)) + \ black_heap s \ fl = sys_fM s" + "heap (s sys) r' = None + \ white_heap (s(mutator m := s (mutator m)\roots := roots'\, sys := s sys\heap := sys_heap s(r' \ \obj_mark = fl, obj_fields = Map.empty, obj_payload = Map.empty\)\)) + \ white_heap s \ fl \ sys_fM s" +unfolding black_heap_def white_def white_heap_def +apply (simp_all add: fun_upd_apply 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 simp: reaches_def 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 grey_protects_white_hs_done[simp]: + "(g grey_protects_white w) (s(mutator m := s (mutator m)\ W := {}, ghost_hs_phase := hs' \, + sys := s sys\ hs_pending := hp', W := sys_W s \ W (s (mutator m)), + ghost_hs_in_sync := his' \)) + \ (g grey_protects_white w) s" +unfolding grey_protects_white_def by (simp add: fun_upd_apply) + +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 := roots'\, sys := s sys\heap := sys_heap s(r \ \obj_mark = fl, obj_fields = Map.empty, obj_payload = Map.empty\)\)) + \ (g grey_protects_white w) s" +unfolding grey_protects_white_def has_white_path_to_def by simp + +lemma (in mut_m) reachable_snapshot_inv_sweep_loop_free: + fixes s :: "('field, 'mut, 'payload, '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 simp: fun_upd_apply dest: reachable_blackD) + next + case (ghost_honorary_root x) with ngs nmr rsi show ?case + unfolding reachable_def reaches_def by (auto simp: fun_upd_apply dest: reachable_blackD) + next + case (tso_root x) with ngs nmr rsi show ?case + unfolding reachable_def reaches_def by (auto simp: fun_upd_apply dest: reachable_blackD) + next + case (reaches x y) with ngs nmr rsi show ?case + unfolding reachable_def reaches_def + apply (clarsimp simp: fun_upd_apply) + 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 reaches_def) + apply (blast intro: rtranclp.intros(2)) + apply clarsimp + apply (frule (1) reachable_blackD[where r=r]) + apply (simp add: reachable_def reaches_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 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, obj_payload = 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 reaches_def elim: rtranclp.intros(2) split: obj_at_splits) + qed (auto simp: fun_upd_apply 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 + unfolding mut_m.reachable_def by fastforce + next + case (reaches x y) with rn show ?case + unfolding mut_m.reachable_def by fastforce + qed (auto simp: fun_upd_apply) + next + assume "m' = m \ r' = r" with rn show ?thesis + unfolding mut_m.reachable_def by (fastforce simp: fun_upd_apply) + qed +qed + +context mut_m +begin + +lemma reachable_snapshot_inv_alloc[simp, elim!]: + fixes s :: "('field, 'mut, 'payload, 'ref) lsts" + assumes rsi: "reachable_snapshot_inv s" + assumes rn: "sys_heap s r = None" + assumes fl: "fl = sys_fM s" + assumes vri: "valid_refs_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, obj_payload = 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 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 simp: fun_upd_apply) + +lemma reachable_snapshot_inv_load[simp]: + "\ reachable_snapshot_inv s; sys_load (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' \))" +unfolding reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def by (simp add: fun_upd_apply) + +lemma 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_store_buffers := (mem_store_buffers (s sys))(mutator m := sys_mem_store_buffers (mutator m) s @ [mw_Mutate r f opt_r']) \))" +unfolding reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def reachable_def +apply clarsimp +apply (drule_tac x=x in spec) +apply (auto simp: fun_upd_apply) +(* FIXME what's gone wrong here? *) +apply (subst (asm) tso_store_refs_simps; force)+ +done + +end + +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}" +unfolding WL_def by (simp add: fun_upd_apply) + +lemma ghost_honorary_grey_mo_co_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" +unfolding black_def grey_def by (auto simp: fun_upd_apply) + +lemma ghost_honorary_grey_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" +unfolding grey_def WL_def by (auto simp: fun_upd_apply split: process_name.splits if_splits) + +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" +unfolding mut_m.reachable_def reaches_def by (clarsimp simp: fun_upd_apply) (metis (no_types, lifting) mono_rtranclp) + +lemma reachable_deref_del[simp]: + "\ sys_load (mutator m) (mr_Ref r f) (s sys) = mv_Ref opt_r'; r \ mut_m.mut_roots m s; mut_m.mut_ghost_honorary_root m 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" +unfolding mut_m.reachable_def reaches_def sys_load_def +apply (clarsimp simp: fun_upd_apply) +apply (rule iffI) + apply clarsimp + apply (elim disjE) + apply metis + apply (erule option_bind_invE; auto dest!: fold_stores_points_to) + apply (auto elim!: converse_rtranclp_into_rtranclp[rotated] + simp: mut_m.tso_store_refs_def) +done + +lemma no_black_refs_dequeue[simp]: + "\ sys_mem_store_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 := map_option (obj_mark_update (\_. fl)) (sys_heap s r)), mem_store_buffers := (mem_store_buffers (s sys))(p := ws)\))" + "\ sys_mem_store_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 := map_option (\obj. obj\obj_fields := (obj_fields obj)(f := r')\) (sys_heap s r)), + mem_store_buffers := (mem_store_buffers (s sys))(p := ws)\))" +unfolding no_black_refs_def by (auto simp: fun_upd_apply dest: valid_W_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" is redundant *) +unfolding black_def grey_def valid_W_inv_def +apply (simp_all add: all_conj_distrib split: obj_at_splits if_splits) +apply safe +apply (simp_all add: WL_def fun_upd_apply split: if_splits) + apply (metis option.distinct(1)) +apply blast +apply blast +apply blast +apply blast +apply blast +apply blast +apply metis +done +(* FIXME +apply (auto simp: black_def grey_def WL_def split: obj_at_splits) +apply metis+ +done +*) + +lemma 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, obj_payload = Map.empty\)\)) + \ fl \ sys_fM s \ grey r' s" +unfolding no_black_refs_def by simp + +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" +unfolding no_black_refs_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 by induct (auto simp: fun_upd_apply) +next + assume ?rhs then show ?lhs + proof(safe) + fix g assume "(g grey_protects_white w) s" + from this ghg show ?thesis +apply induct + apply force +unfolding grey_protects_white_def +apply (auto simp: fun_upd_apply) +done + next + assume "(r has_white_path_to w) s" with ghg show ?thesis + unfolding grey_protects_white_def has_white_path_to_def by (auto simp: fun_upd_apply) + qed +qed + +lemma valid_refs_inv_dequeue_Mutate: + fixes s :: "('field, 'mut, 'payload, 'ref) lsts" + assumes vri: "valid_refs_inv s" + assumes sb: "sys_mem_store_buffers (mutator m') s = mw_Mutate r f opt_r' # ws" + shows "valid_refs_inv (s(sys := s sys\heap := (sys_heap s)(r := map_option (\obj. obj\obj_fields := (obj_fields obj)(f := opt_r')\) (sys_heap s r)), + mem_store_buffers := (mem_store_buffers (s sys))(mutator m' := ws)\))" (is "valid_refs_inv ?s'") +proof(rule valid_refs_invI) + fix m + let ?root = "\m x. mut_m.root m x \<^bold>\ grey x" + 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'" + unfolding reaches_def proof induct + case base with x sb vri show ?case + apply - + apply (subst obj_at_fun_upd) + apply (auto simp: mut_m.tso_store_refs_def reaches_def fun_upd_apply 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 simp: fun_upd_apply) + apply (subst (asm) obj_at_fun_upd, fastforce simp: fun_upd_apply) + apply (clarsimp simp: points_to_Mutate fun_upd_apply) + apply (fastforce elim: rtranclp.intros(2) simp: mut_m.tso_store_refs_def reaches_def fun_upd_apply intro: exI[where x=m'] valid_refs_invD(5)[where m=m']) + done + qed + then show "valid_ref y ?s'" by blast +qed + +lemma valid_refs_inv_dequeue_Mutate_Payload: + notes if_split_asm[split del] + fixes s :: "('field, 'mut, 'payload, 'ref) lsts" + assumes vri: "valid_refs_inv s" + assumes sb: "sys_mem_store_buffers (mutator m') s = mw_Mutate_Payload r f pl # ws" + shows "valid_refs_inv (s(sys := s sys\heap := (sys_heap s)(r := map_option (\obj. obj\obj_payload := (obj_payload obj)(f := pl)\) (sys_heap s r)), + mem_store_buffers := (mem_store_buffers (s sys))(mutator m := ws)\))" (is "valid_refs_inv ?s'") +apply (rule valid_refs_invI) +using assms +apply (clarsimp simp: valid_refs_invD fun_upd_apply split: obj_at_splits mem_store_action.splits) +apply auto + apply (metis (mono_tags, lifting) UN_insert Un_iff list.simps(15) mut_m.tso_store_refs_def valid_refs_invD(4)) +apply (metis case_optionE obj_at_def valid_refs_invD(7)) +done + +(*<*) + +end +(*>*) diff --git a/thys/ConcurrentGC/Global_Noninterference.thy b/thys/ConcurrentGC/Global_Noninterference.thy new file mode 100644 --- /dev/null +++ b/thys/ConcurrentGC/Global_Noninterference.thy @@ -0,0 +1,452 @@ +(*<*) +(* + * 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 Global_Noninterference +imports + Global_Invariants_Lemmas + Tactics +begin + +(*>*) +section\ Global non-interference \ + +text\ proofs that depend only on global invariants + lemmas \ + +lemma (in sys) strong_tricolour_inv[intro]: + notes fun_upd_apply[simp] + shows + "\ LSTP (fM_rel_inv \<^bold>\ handshake_phase_inv \<^bold>\ mutators_phase_inv \<^bold>\ strong_tricolour_inv \<^bold>\ sys_phase_inv \<^bold>\ tso_store_inv \<^bold>\ valid_W_inv) \ + sys + \ LSTP strong_tricolour_inv \" +unfolding strong_tricolour_inv_def +proof(vcg_jackhammer (no_thin_post_inv), vcg_name_cases) + case (tso_dequeue_store_buffer s s' p w ws x xa) then show ?case + proof(cases w) + case (mw_Mark ref field) with tso_dequeue_store_buffer show ?thesis + apply - + apply clarsimp + apply (frule (1) valid_W_invD) + apply clarsimp + apply (cases "x = ref"; clarsimp simp: grey_def white_def WL_def split: if_splits) + apply (drule_tac x=x in spec; force split: obj_at_splits) + done + next case (mw_Mutate ref field opt_r') with tso_dequeue_store_buffer show ?thesis + apply - + apply (clarsimp simp: fM_rel_inv_def p_not_sys) + apply (elim disjE; clarsimp simp: points_to_Mutate) + apply (elim disjE; clarsimp) + apply (case_tac "sys_ghost_hs_phase s\"; clarsimp simp: hp_step_rel_def heap_colours_colours no_black_refsD) + proof(goal_cases hp_InitMark hp_Mark hp_IdleMarkSweep) + case (hp_InitMark m) then show ?case + apply - + apply (drule mut_m.handshake_phase_invD[where m=m]) + apply (drule_tac x=m in spec) + apply (elim disjE; clarsimp simp: hp_step_rel_def) + apply (elim disjE; clarsimp simp: mut_m.marked_insertions_def no_black_refsD marked_not_white) + done + next case (hp_Mark m) then show ?case + apply - + apply (drule mut_m.handshake_phase_invD[where m=m]) + apply (drule_tac x=m in spec) + apply (elim disjE; clarsimp simp: hp_step_rel_def) + apply (elim disjE; clarsimp simp: mut_m.marked_insertions_def no_black_refsD) + apply blast+ + done + next case (hp_IdleMarkSweep m) then show ?case + apply - + apply (drule mut_m.handshake_phase_invD[where m=m]) + apply (drule_tac x=m in spec) + apply (elim disjE; clarsimp simp: hp_step_rel_def) + apply (elim disjE; clarsimp simp: marked_not_white mut_m.marked_insertions_def) + done + qed + next case (mw_fM fM) with tso_dequeue_store_buffer show ?thesis + apply - + apply (clarsimp simp: fM_rel_inv_def p_not_sys) + apply (erule disjE) + apply (clarsimp simp: fM_rel_def black_heap_def split: if_splits) + apply (metis colours_distinct(2) white_valid_ref) + apply (clarsimp simp: white_heap_def) + apply ( (drule_tac x=xa in spec)+ )[1] + apply (clarsimp simp: white_def split: obj_at_splits) + apply (fastforce simp: white_def) + done + qed (clarsimp simp: fM_rel_inv_def p_not_sys)+ +qed + +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 (full_types) reachable_points_to valid_refs_inv_def) +done + +lemma black_heap_valid_ref_marked_insertions: + "\ 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_store_action.splits option.splits + dest: valid_refs_invD) + +context sys +begin + +lemma reachable_snapshot_inv_black_heap_no_grey_refs_dequeue_Mutate: + assumes sb: "sys_mem_store_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 := map_option (\obj. obj\obj_fields := (obj_fields obj)(f := opt_r')\) (sys_heap s r)), + mem_store_buffers := (mem_store_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 fun_upd_apply; fail) +using bh ngr sb vri +apply (subst valid_refs_inv_def) +apply (clarsimp simp add: no_grey_refs_def grey_reachable_def fun_upd_apply) +apply (drule black_heap_reachable) + apply (simp add: black_heap_def fun_upd_apply; fail) + apply (clarsimp simp: valid_refs_inv_dequeue_Mutate; fail) +apply (clarsimp simp: in_snapshot_def in_snapshot_valid_ref fun_upd_apply) +done + +lemma marked_deletions_dequeue_Mark: + "\ sys_mem_store_buffers p s = mw_Mark r fl # ws; mut_m.marked_deletions m s; tso_store_inv s; valid_W_inv s \ + \ mut_m.marked_deletions m (s(sys := s sys\heap := (sys_heap s)(r := map_option (obj_mark_update (\_. fl)) (sys_heap s r)), mem_store_buffers := (mem_store_buffers (s sys))(p := ws)\))" +unfolding mut_m.marked_deletions_def +by (auto simp: fun_upd_apply obj_at_field_on_heap_def + split: obj_at_splits option.splits mem_store_action.splits + dest: valid_W_invD) + +lemma marked_deletions_dequeue_Mutate: + "\ sys_mem_store_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 := map_option (\obj. obj\obj_fields := (obj_fields obj)(f := opt_r')\) (sys_heap s r)), + mem_store_buffers := (mem_store_buffers (s sys))((mutator m') := ws)\))" +unfolding mut_m.marked_insertions_def mut_m.marked_deletions_def +apply (clarsimp simp: fun_upd_apply split: mem_store_action.splits option.splits) +apply (metis list.set_intros(2) obj_at_field_on_heap_imp_valid_ref(1))+ +done + +lemma grey_protects_white_dequeue_Mark: + 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 := map_option (obj_mark_update (\_. fl)) (sys_heap s r)), mem_store_buffers := (mem_store_buffers (s sys))(p := ws)\))) + \ (\g. (g grey_protects_white w) s)" (is "(\g. (g grey_protects_white w) ?s') \ ?rhs") +proof(rule iffI) + assume "\g. (g grey_protects_white w) ?s'" + then obtain g where "(g grey_protects_white w) ?s'" by blast + from this assms show ?rhs + proof induct + case (step x y z) then show ?case + apply (cases "y = r"; clarsimp simp: fun_upd_apply) + apply (metis black_dequeue_Mark colours_distinct(2) do_store_action_simps(1) greyI(1) grey_protects_whiteE(1) grey_protects_whiteI marked_imp_black_or_grey(2) valid_ref_valid_null_ref_simps(1) white_valid_ref) + apply (metis black_dequeue_Mark colours_distinct(2) do_store_action_simps(1) grey_protects_whiteE(2) grey_protects_whiteI marked_imp_black_or_grey(2) valid_ref_valid_null_ref_simps(1) white_valid_ref) + done + qed (fastforce simp: fun_upd_apply) +next + assume ?rhs + then obtain g' where "(g' grey_protects_white w) s" .. + then show "\g. (g grey_protects_white w) ?s'" + proof induct + case (refl g) with assms show ?case +apply - +apply (rule exI[where x=g]) +apply (rule grey_protects_whiteI) +apply (subst grey_fun_upd; simp add: fun_upd_apply) +done (* FIXME something eta-ish going wrong here: fun_upd_apply triggers too early, why? Maybe the WL rules are borked too *) + next + case (step x y z) with assms show ?case +apply clarsimp +apply (rename_tac g) +apply (clarsimp simp add: grey_protects_white_def) +apply (case_tac "z = r") + apply (rule exI[where x=r]) + apply (clarsimp simp add: grey_protects_white_def) + apply (subst grey_fun_upd; force simp: fun_upd_apply) +apply (rule_tac x=g in exI) +apply (fastforce elim!: has_white_path_to_step) +done + qed +qed + +lemma reachable_snapshot_inv_dequeue_Mark: + "\ sys_mem_store_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 := map_option (obj_mark_update (\_. fl)) (sys_heap s r)), mem_store_buffers := (mem_store_buffers (s sys))(p := ws)\))" +unfolding mut_m.reachable_snapshot_inv_def in_snapshot_def +apply clarsimp +apply (rename_tac x) +apply (drule_tac x=x in spec) +apply (subst (asm) arg_cong[where f=Not, OF grey_protects_white_dequeue_Mark, simplified]; simp add: colours_distinct(4) valid_W_invD(1) fun_upd_apply) +done + +lemma marked_insertions_dequeue_Mark: + "\ sys_mem_store_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 := map_option (obj_mark_update (\_. fl)) (sys_heap s r)), mem_store_buffers := (mem_store_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 simp: valid_W_invD split: mem_store_action.splits option.splits obj_at_splits; fail) +apply clarsimp +apply (rename_tac x) +apply (drule_tac x=x in spec) +apply (auto simp: valid_W_invD split: mem_store_action.splits option.splits obj_at_splits) +done + +lemma marked_insertions_dequeue_Mutate: + "\ sys_mem_store_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 := map_option (\obj. obj\obj_fields := (obj_fields obj)(f := r')\) (sys_heap s r)), + mem_store_buffers := (mem_store_buffers (s sys))(p := ws)\))" +unfolding 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 simp: fun_upd_apply split: mem_store_action.splits option.splits obj_at_splits; fail)[1] +apply clarsimp +apply (rename_tac x) +apply (drule_tac x=x in spec) +apply (auto simp: fun_upd_apply split: mem_store_action.splits option.splits obj_at_splits)[1] +done + +(* shows the snapshot is preserved by the two marks. *) +lemma grey_protects_white_dequeue_Mutate: + assumes sb: "sys_mem_store_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" + shows "(\g. (g grey_protects_white w) (s(sys := s sys\heap := (sys_heap s)(r := map_option (\obj. obj\obj_fields := (obj_fields obj)(f := opt_r')\) (sys_heap s r)), + mem_store_buffers := (mem_store_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 + from this mi sb show ?rhs + proof(induct rule: grey_protects_white_induct) + case (refl x) then show ?case by (fastforce simp: fun_upd_apply) + next case (step x y z) then show ?case + unfolding white_def + apply (clarsimp simp: points_to_Mutate grey_protects_white_def) + apply (auto dest: marked_insertionD simp: marked_not_white whiteI fun_upd_apply) + done + qed +next + assume ?rhs then show "(\g. (g grey_protects_white w) ?s')" + proof(clarsimp) + fix g assume "(g grey_protects_white w) s" + from this show ?thesis + proof(induct rule: grey_protects_white_induct) + case (refl x) then show ?case + apply - + apply (rule exI[where x=x]) + apply (clarsimp simp: grey_protects_white_def) + apply (subst grey_fun_upd; simp add: fun_upd_apply) (* FIXME *) + done + next case (step x y z) with md sb show ?case + apply clarsimp + apply (clarsimp simp: grey_protects_white_def) + apply (rename_tac g) + + apply (case_tac "y = r") + defer + apply (auto simp: points_to_Mutate fun_upd_apply elim!: has_white_path_to_step; fail)[1] + apply (clarsimp simp: ran_def fun_upd_apply split: obj_at_split_asm) (* FIXME rule: witness field for r points_to c *) + apply (rename_tac g obj aa) + apply (case_tac "aa = f") + defer + apply (rule_tac x=g in exI) + apply clarsimp + apply (clarsimp simp: has_white_path_to_def fun_upd_apply) + apply (erule rtranclp.intros) + apply (auto simp: fun_upd_apply ran_def split: obj_at_splits; fail)[1] + + apply (clarsimp simp: has_white_path_to_def) + 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 add: white_def split: obj_at_splits) + done + qed + qed +qed + +(* write barrier installed but not all mutators are necessarily past get_roots *) +lemma reachable_snapshot_inv_dequeue_Mutate: + notes grey_protects_white_dequeue_Mutate[simp] + fixes s :: "('field, 'mut, 'payload, 'ref) lsts" + assumes sb: "sys_mem_store_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" + shows "mut_m.reachable_snapshot_inv m (s(sys := s sys\heap := (sys_heap s)(r := map_option (\obj. obj\obj_fields := (obj_fields obj)(f := opt_r')\) (sys_heap s r)), + mem_store_buffers := (mem_store_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) + apply (auto simp: fun_upd_apply) + done + next + case (ghost_honorary_root x) with mi md rsi sb show ?case + unfolding mut_m.reachable_snapshot_inv_def in_snapshot_def by (auto simp: fun_upd_apply) + next + case (tso_root x) with mi md rsi sb show ?case + apply (clarsimp simp: mut_m.reachable_snapshot_inv_def in_snapshot_def) + apply (rename_tac w) + apply (case_tac w; simp) (* FIXME cut and paste here *) + apply (rename_tac ref field option) + apply (clarsimp simp: mut_m.marked_deletions_def mut_m.marked_insertions_def fun_upd_apply) + apply (drule_tac x="mw_Mutate ref field option" in spec) + apply (drule_tac x="mw_Mutate ref field option" in spec) + apply (clarsimp simp: fun_upd_apply) + apply (frule spec[where x=x]) + apply (subgoal_tac "mut_m.reachable m x s") + apply (force simp: fun_upd_apply) + apply (rule reachableI(2)) + apply (force simp: mut_m.tso_store_refs_def) + apply (rename_tac ref field pl) + apply (clarsimp simp: mut_m.marked_deletions_def mut_m.marked_insertions_def fun_upd_apply) + apply (drule_tac x="mw_Mutate_Payload x field pl" in spec) + apply (drule_tac x="mw_Mutate_Payload x field pl" in spec) + apply (clarsimp simp: fun_upd_apply) + apply (frule spec[where x=x]) + apply (subgoal_tac "mut_m.reachable m x s") + apply (force simp: fun_upd_apply) + apply (rule reachableI(2)) + apply (force simp: mut_m.tso_store_refs_def) + apply (auto simp: fun_upd_apply) + 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_Mutate mut_m.reachable_snapshot_inv_def in_snapshot_def) + apply (elim disjE, (force dest!: reachable_points_to 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) + apply (clarsimp simp: fun_upd_apply) + apply (drule spec[where x=y]) + apply (clarsimp simp: points_to_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; clarsimp simp: reachable_points_to) (* FIXME probably want points_to_Mutate as an elim rule to make this robust, reduce duplication *) + apply (drule (3) strong_tricolour_invD) + apply (metis (no_types) grey_protects_whiteI marked_imp_black_or_grey(1)) + + apply (metis (no_types) grey_protects_whiteE(2) grey_protects_whiteI marked_imp_black_or_grey(2)) + + apply (elim disjE; clarsimp simp: reachable_points_to) + apply (force simp: black_def) + + apply (elim disjE; clarsimp simp: reachable_points_to) + apply (force simp: black_def) + + apply (elim disjE; clarsimp simp: reachable_points_to) + apply (force simp: black_def) + + apply (drule (3) strong_tricolour_invD) + apply (force simp: black_def) + + apply (elim disjE; clarsimp) + apply (force simp: black_def fun_upd_apply) + apply (metis (no_types) grey_protects_whiteE(2) grey_protects_whiteI marked_imp_black_or_grey(2)) + done + moreover note mi md rsi sb + ultimately show ?case + apply (clarsimp simp: mut_m.reachable_snapshot_inv_def in_snapshot_def) + apply (clarsimp simp: fun_upd_apply) + done + qed + then show "in_snapshot y ?s'" by blast +qed + +lemma 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_store_inv \<^bold>\ valid_refs_inv \<^bold>\ valid_W_inv) \ + sys + \ LSTP (mut_m.mutator_phase_inv m) \" +proof(vcg_jackhammer (no_thin_post_inv), vcg_name_cases) + case (tso_dequeue_store_buffer s s' p w ws) show ?case + proof(cases w) + case (mw_Mark ref field) with tso_dequeue_store_buffer show ?thesis + by (clarsimp simp: mutator_phase_inv_aux_case + marked_deletions_dequeue_Mark marked_insertions_dequeue_Mark reachable_snapshot_inv_dequeue_Mark + split: hs_phase.splits) + next case (mw_Mutate ref field opt_r') show ?thesis + proof(cases "ghost_hs_phase (s\ (mutator m))") + case hp_IdleInit + with \sys_mem_store_buffers p s\ = w # ws\ spec[OF \mutators_phase_inv s\\, where x=m] mw_Mutate + show ?thesis by simp + next case hp_InitMark + with \sys_mem_store_buffers p s\ = w # ws\ spec[OF \mutators_phase_inv s\\, where x=m] mw_Mutate + show ?thesis by (simp add: marked_insertions_dequeue_Mutate) + next case hp_Mark with tso_dequeue_store_buffer mw_Mutate show ?thesis + apply - + apply (clarsimp simp: mutator_phase_inv_aux_case p_not_sys split: hs_phase.splits) + apply (erule disjE; clarsimp simp: marked_insertions_dequeue_Mutate) + apply (rename_tac m') + apply (frule mut_m.handshake_phase_invD[where m=m]) + apply (rule marked_deletions_dequeue_Mutate, simp_all) + apply (drule_tac m=m' in mut_m.handshake_phase_invD, clarsimp simp: hp_step_rel_def) + using hs_phase.distinct(11) hs_phase.distinct(15) hs_type.distinct(1) apply presburger + done + next case hp_IdleMarkSweep with tso_dequeue_store_buffer mw_Mutate show ?thesis + apply - + apply (clarsimp simp: mutator_phase_inv_aux_case p_not_sys + split: hs_phase.splits) + apply (intro allI conjI impI; erule disjE; clarsimp simp: sys.marked_insertions_dequeue_Mutate) + apply (rename_tac m') + apply (rule marked_deletions_dequeue_Mutate, simp_all)[1] + apply (drule_tac x=m' in spec) + apply (frule mut_m.handshake_phase_invD[where m=m]) + apply (drule_tac m=m' in mut_m.handshake_phase_invD, clarsimp simp: hp_step_rel_def) + apply (elim disjE; clarsimp split del: if_split_asm) + apply (clarsimp simp: fA_rel_inv_def fM_rel_inv_def fA_rel_def fM_rel_def split del: if_split_asm) + apply (meson black_heap_valid_ref_marked_insertions; fail) + apply (rename_tac m') + apply (frule_tac m=m in mut_m.handshake_phase_invD) + apply (drule_tac m=m' in mut_m.handshake_phase_invD, clarsimp simp: hp_step_rel_def) + apply (elim disjE; clarsimp simp: reachable_snapshot_inv_black_heap_no_grey_refs_dequeue_Mutate reachable_snapshot_inv_dequeue_Mutate) + apply (clarsimp simp: fA_rel_inv_def fM_rel_inv_def fA_rel_def fM_rel_def) + apply blast + done + qed simp + next case (mw_Mutate_Payload r f pl) with tso_dequeue_store_buffer show ?thesis +apply (clarsimp simp: mutator_phase_inv_aux_case fun_upd_apply split: hs_phase.splits) +apply (subst reachable_snapshot_fun_upd) +apply (simp_all add: fun_upd_apply) +apply (metis (no_types, lifting) list.set_intros(1) mem_store_action.simps(39) tso_store_inv_def) +done + next case (mw_fA mark) with tso_dequeue_store_buffer show ?thesis + by (clarsimp simp: mutator_phase_inv_aux_case fun_upd_apply split: hs_phase.splits) + next case (mw_fM mark) with tso_dequeue_store_buffer show ?thesis + using mut_m_not_idle_no_fM_writeD by fastforce + next case (mw_Phase phase) with tso_dequeue_store_buffer show ?thesis + by (clarsimp simp: mutator_phase_inv_aux_case fun_upd_apply split: hs_phase.splits) + qed +qed + +end + +(*<*) + +end +(*>*) diff --git a/thys/ConcurrentGC/Handshakes.thy b/thys/ConcurrentGC/Handshakes.thy deleted file mode 100644 --- a/thys/ConcurrentGC/Handshakes.thy +++ /dev/null @@ -1,587 +0,0 @@ -(*<*) -(* - * 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/Initial_Conditions.thy b/thys/ConcurrentGC/Initial_Conditions.thy new file mode 100644 --- /dev/null +++ b/thys/ConcurrentGC/Initial_Conditions.thy @@ -0,0 +1,112 @@ +(*<*) +(* + * 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 Initial_Conditions +imports + Local_Invariants_Lemmas + Global_Invariants_Lemmas + Tactics +begin + +(*>*) +section\Initial conditions \label{sec:initial-conditions-proofs}\ + +context gc_system +begin + +lemma init_strong_tricolour_inv: + "\ obj_mark ` ran (sys_heap \GST = s, HST = []\\) \ {gc_fM \GST = s, HST = []\\}; sys_fM \GST = s, HST = []\\ = gc_fM \GST = s, HST = []\\ \ + \ strong_tricolour_inv \GST = s, HST = []\\" +unfolding strong_tricolour_inv_def ran_def white_def by (auto split: obj_at_splits) + +lemma init_no_grey_refs: + "\ gc_W \GST = s, HST = []\\ = {}; \m. W (\GST = s, HST = []\\ (mutator m)) = {}; sys_W \GST = s, HST = []\\ = {}; + gc_ghost_honorary_grey \GST = s, HST = []\\ = {}; \m. ghost_honorary_grey (\GST = s, HST = []\\ (mutator m)) = {}; sys_ghost_honorary_grey \GST = s, HST = []\\ = {} \ + \ no_grey_refs \GST = s, HST = []\\" +unfolding no_grey_refs_def grey_def WL_def by (metis equals0D process_name.exhaust sup_bot.left_neutral) + +lemma valid_refs_imp_valid_refs_inv: + "\ valid_refs s; no_grey_refs s; \p. sys_mem_store_buffers p s = []; \m. ghost_honorary_root (s (mutator m)) = {} \ + \ valid_refs_inv s" +unfolding valid_refs_inv_def valid_refs_def mut_m.reachable_def mut_m.tso_store_refs_def +using no_grey_refs_not_grey_reachableD by fastforce + +lemma no_grey_refs_imp_valid_W_inv: + "\ no_grey_refs s; \p. sys_mem_store_buffers p s = [] \ + \ valid_W_inv s" +unfolding valid_W_inv_def no_grey_refs_def grey_def WL_def by auto + +lemma valid_refs_imp_reachable_snapshot_inv: + "\ valid_refs s; obj_mark ` ran (sys_heap s) \ {sys_fM s}; \p. sys_mem_store_buffers p s = []; \m. ghost_honorary_root (s (mutator m)) = {} \ + \ mut_m.reachable_snapshot_inv m s" +unfolding mut_m.reachable_snapshot_inv_def in_snapshot_def valid_refs_def black_def mut_m.reachable_def mut_m.tso_store_refs_def +apply clarsimp +apply (auto simp: image_subset_iff ran_def split: obj_at_splits) +done + +lemma init_inv_sys: "\s. initial_state gc_system s \ invs \GST = s, HST = []\\" +apply (clarsimp dest!: initial_stateD + 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_store_inv_def + init_no_grey_refs init_strong_tricolour_inv no_grey_refs_imp_valid_W_inv + valid_refs_imp_reachable_snapshot_inv + valid_refs_imp_valid_refs_inv + mut_m.marked_deletions_def mut_m.marked_insertions_def + fA_rel_inv_def fA_rel_def fM_rel_inv_def fM_rel_def + all_conj_distrib) +done + +lemma init_inv_mut: "\s. initial_state gc_system s \ mut_m.invsL m \GST = s, HST = []\" +apply (clarsimp dest!: initial_stateD) +apply (drule fun_cong[where x="mutator m"]) +apply (clarsimp simp: all_com_interned_defs) +unfolding mut_m.invsL_def mut_m.mut_get_roots_mark_object_invL_def2 mut_m.mut_store_del_mark_object_invL_def2 mut_m.mut_store_ins_mark_object_invL_def2 + mut_m.mark_object_invL_def mut_m.handshake_invL_def mut_m.tso_lock_invL_def + gc_system_init_def mut_initial_state_def sys_initial_state_def +apply (intro conjI; simp add: locset_cache atS_simps; simp add: mut_m.loc_defs) +done + +lemma init_inv_gc: "\s. initial_state gc_system s \ gc.invsL \GST = s, HST = []\" +apply (clarsimp dest!: initial_stateD) +apply (drule fun_cong[where x=gc]) +apply (clarsimp simp: all_com_interned_defs) +unfolding gc.invsL_def gc.fM_fA_invL_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 gc.gc_mark_mark_object_invL_def2 +apply (intro conjI; simp add: locset_cache atS_simps init_no_grey_refs; simp add: gc.loc_defs) +apply (simp_all add: gc_system_init_def gc_initial_state_def mut_initial_state_def sys_initial_state_def + gc_system.init_no_grey_refs) + apply blast +apply (clarsimp simp: image_subset_iff ranI split: obj_at_splits) +done + +end + +(* FIXME really deserves to be somewhere very public but there's no shared theory immediately above Local and Global invs by design *) + +definition I :: "('field, 'mut, 'payload, 'ref) gc_pred" where + "I = (invsL \<^bold>\ LSTP invs)" + +lemmas I_defs = gc.invsL_def mut_m.invsL_def invsL_def invs_def I_def + +context gc_system +begin + +theorem init_inv: "\s. initial_state gc_system s \ I \GST = s, HST = []\" +unfolding I_def invsL_def by (simp add: init_inv_sys init_inv_gc init_inv_mut) + +end + +(*<*) + +end +(*>*) diff --git a/thys/ConcurrentGC/Local_Invariants.thy b/thys/ConcurrentGC/Local_Invariants.thy new file mode 100644 --- /dev/null +++ b/thys/ConcurrentGC/Local_Invariants.thy @@ -0,0 +1,619 @@ +(*<*) +(* + * 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 Local_Invariants +imports + Proofs_Basis +begin + +(*>*) +section\ Local invariants \label{sec:local-invariants}\ + + +subsection\TSO invariants\ + +context gc +begin + +text \ + +The GC holds the TSO lock only during the \texttt{CAS} in \mark_object\. + +\ + +locset_definition tso_lock_locs :: "location set" where + "tso_lock_locs = (\l\{ ''mo_co_cmark'', ''mo_co_ctest'', ''mo_co_mark'', ''mo_co_unlock'' }. suffixed l)" + +definition tso_lock_invL :: "('field, 'mut, 'payload, 'ref) gc_pred" where +[inv]: "tso_lock_invL = + (atS_gc tso_lock_locs (tso_locked_by gc) + \<^bold>\ atS_gc (- tso_lock_locs) (\<^bold>\ tso_locked_by gc))" + +end + +context mut_m +begin + +text \ + +A mutator holds the TSO lock only during the \texttt{CAS}s in \mark_object\. + +\ + +locset_definition "tso_lock_locs = + (\l\{ ''mo_co_cmark'', ''mo_co_ctest'', ''mo_co_mark'', ''mo_co_unlock'' }. suffixed l)" + +definition tso_lock_invL :: "('field, 'mut, 'payload, 'ref) gc_pred" where +[inv]: "tso_lock_invL = + (atS_mut tso_lock_locs (tso_locked_by (mutator m)) + \<^bold>\ atS_mut (-tso_lock_locs) (\<^bold>\tso_locked_by (mutator m)))" + +end + + +subsection\Handshake phases \label{sec:local-handshake-phases}\ + +text\ + +Connect @{const "sys_ghost_hs_phase"} with locations in the GC. + +\ + +context gc +begin + +locset_definition "idle_locs = prefixed ''idle''" +locset_definition "init_locs = prefixed ''init''" +locset_definition "mark_locs = prefixed ''mark''" +locset_definition "sweep_locs = prefixed ''sweep''" +locset_definition "mark_loop_locs = prefixed ''mark_loop''" + +locset_definition "hp_Idle_locs = + (prefixed ''idle_noop'' - { idle_noop_mfence, idle_noop_init_type }) + \ { idle_load_fM, idle_invert_fM, idle_store_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_load_fM, mark_store_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''}))" + +definition handshake_invL :: "('field, 'mut, 'payload, 'ref) gc_pred" where +[inv]: "handshake_invL = + (atS_gc hs_noop_locs (sys_hs_type \<^bold>= \ht_NOOP\) + \<^bold>\ atS_gc hs_get_roots_locs (sys_hs_type \<^bold>= \ht_GetRoots\) + \<^bold>\ atS_gc hs_get_work_locs (sys_hs_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_hs_pending m + \<^bold>\ sys_ghost_hs_in_sync m) + \<^bold>\ atS_gc hs_init_loop_not_done_locs (\<^bold>\m. \m\ \<^bold>\ gc_muts \<^bold>\ \<^bold>\sys_hs_pending m + \<^bold>\ \<^bold>\sys_ghost_hs_in_sync m) + \<^bold>\ atS_gc hs_init_loop_done_locs ( (sys_hs_pending \<^bold>$ gc_mut + \<^bold>\ sys_ghost_hs_in_sync \<^bold>$ gc_mut) + \<^bold>\ (\<^bold>\m. \m\ \<^bold>\ gc_muts \<^bold>\ \m\ \<^bold>\ gc_mut + \<^bold>\ \<^bold>\sys_hs_pending m + \<^bold>\ \<^bold>\sys_ghost_hs_in_sync m) ) + \<^bold>\ atS_gc hs_done_locs (\<^bold>\m. sys_hs_pending m \<^bold>\ sys_ghost_hs_in_sync m) + \<^bold>\ atS_gc hs_done_loop_locs (\<^bold>\m. \<^bold>\\m\ \<^bold>\ gc_muts \<^bold>\ \<^bold>\sys_hs_pending m) + \<^bold>\ atS_gc hs_none_pending_locs (\<^bold>\m. \<^bold>\sys_hs_pending m) + \<^bold>\ atS_gc hs_in_sync_locs (\<^bold>\m. sys_ghost_hs_in_sync m) + \<^bold>\ atS_gc hs_out_of_sync_locs (\<^bold>\m. \<^bold>\sys_hs_pending m + \<^bold>\ \<^bold>\sys_ghost_hs_in_sync m) + + \<^bold>\ atS_gc hp_Idle_locs (sys_ghost_hs_phase \<^bold>= \hp_Idle\) + \<^bold>\ atS_gc hp_IdleInit_locs (sys_ghost_hs_phase \<^bold>= \hp_IdleInit\) + \<^bold>\ atS_gc hp_InitMark_locs (sys_ghost_hs_phase \<^bold>= \hp_InitMark\) + \<^bold>\ atS_gc hp_IdleMarkSweep_locs (sys_ghost_hs_phase \<^bold>= \hp_IdleMarkSweep\) + \<^bold>\ atS_gc hp_Mark_locs (sys_ghost_hs_phase \<^bold>= \hp_Mark\))" + +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_load_fM, mark_store_fA, mark_noop_mfence })" + +definition phase_invL :: "('field, 'mut, 'payload, 'ref) gc_pred" where +[inv]: "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)))" + +end + +text\ + +Local handshake phase invariant for the mutators. + +\ + +context mut_m +begin + +locset_definition "hs_noop_locs = prefixed ''hs_noop_''" +locset_definition "hs_get_roots_locs = prefixed ''hs_get_roots_''" +locset_definition "hs_get_work_locs = prefixed ''hs_get_work_''" +locset_definition "no_pending_mutations_locs = + { hs_load_ht } + \ (prefixed ''hs_noop'') + \ (prefixed ''hs_get_roots'') + \ (prefixed ''hs_get_work'')" +locset_definition "hs_pending_loaded_locs = (prefixed ''hs_'' - { hs_load_pending })" +locset_definition "hs_pending_locs = (prefixed ''hs_'' - { hs_load_pending, hs_pending })" +locset_definition "ht_loaded_locs = (prefixed ''hs_'' - { hs_load_pending, hs_pending, hs_mfence, hs_load_ht })" + +definition handshake_invL :: "('field, 'mut, 'payload, 'ref) gc_pred" where +[inv]: "handshake_invL = + (atS_mut hs_noop_locs (sys_hs_type \<^bold>= \ht_NOOP\) + \<^bold>\ atS_mut hs_get_roots_locs (sys_hs_type \<^bold>= \ht_GetRoots\) + \<^bold>\ atS_mut hs_get_work_locs (sys_hs_type \<^bold>= \ht_GetWork\) + \<^bold>\ atS_mut ht_loaded_locs (mut_hs_pending \<^bold>\ mut_hs_type \<^bold>= sys_hs_type) + \<^bold>\ atS_mut hs_pending_loaded_locs (mut_hs_pending \<^bold>\ sys_hs_pending m) + \<^bold>\ atS_mut hs_pending_locs (mut_hs_pending) + \<^bold>\ atS_mut no_pending_mutations_locs (LIST_NULL (tso_pending_mutate (mutator m))))" + +end + + +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. + +\ + +context gc +begin + +locset_definition "fM_eq_locs = (- { idle_store_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_load_fM, idle_invert_fM } + \ prefixed ''idle_noop'' + \ (mark_locs - { mark_load_fM, mark_store_fA, mark_noop_mfence }) + \ sweep_locs" + +locset_definition + "fA_neq_locs = { idle_phase_init, idle_store_fM, mark_load_fM, mark_store_fA } + \ prefixed ''idle_flip_noop'' + \ init_locs" + +definition fM_fA_invL :: "('field, 'mut, 'payload, 'ref) gc_pred" where +[inv]: "fM_fA_invL = + (atS_gc fM_eq_locs (gc_fM \<^bold>= sys_fM) + \<^bold>\ at_gc idle_store_fM (gc_fM \<^bold>\ sys_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 (gc_fM \<^bold>= sys_fA) + \<^bold>\ atS_gc fA_neq_locs (gc_fM \<^bold>\ sys_fA) + \<^bold>\ at_gc mark_noop_mfence (gc_fM \<^bold>\ sys_fA \<^bold>\ \<^bold>\LIST_NULL (tso_pending_fA gc)) + \<^bold>\ atS_gc fA_tso_empty_locs (LIST_NULL (tso_pending_fA gc)))" + +end + + +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, 'payload, '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_hs_phase s \ ghost_hs_phase (s p)" +abbreviation (input) "p_ghost_honorary_grey s \ ghost_honorary_grey (s p)" +abbreviation (input) "p_ghost_hs_in_sync s \ ghost_hs_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, 'payload, 'ref) lsts_pred \ ('field, 'mut, 'payload, '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)" + +(* FIXME rename: these are assertions? *) +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, 'payload, '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))" + +end + +text\ + +The uses of @{const "mark_object_fn"} in the GC and during the root +marking are straightforward. + +\ + +interpretation gc_mark: mark_object "gc" "gc.mark_loop" "\True\" +by standard (simp add: eq_imp_def) + +lemmas (in gc) gc_mark_mark_object_invL_def2[inv] = gc_mark.mark_object_invL_def[unfolded loc_defs, simplified, folded loc_defs] + +interpretation mut_get_roots: mark_object "mutator m" "mut_m.hs_get_roots_loop" "\True\" for m +by standard (simp add: eq_imp_def) + +lemmas (in mut_m) mut_get_roots_mark_object_invL_def2[inv] = mut_get_roots.mark_object_invL_def[unfolded loc_defs, simplified, folded loc_defs] + +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_hs_phase m \<^bold>\ \hp_Idle\" for m (* FIXME store del, why the string? *) +by standard (simp add: eq_imp_def) + +lemmas (in mut_m) mut_store_del_mark_object_invL_def2[inv] = mut_store_del.mark_object_invL_def[simplified, folded loc_defs] + +interpretation mut_store_ins: mark_object "mutator m" "mut_m.store_ins" "mut_m.mut_ghost_hs_phase m \<^bold>\ \hp_Idle\" for m +by standard (simp add: eq_imp_def) + +lemmas (in mut_m) mut_store_ins_mark_object_invL_def2[inv] = mut_store_ins.mark_object_invL_def[unfolded loc_defs, simplified, folded loc_defs] + +text\ + +Local invariant for the mutator's uses of \mark_object\. + +\ + +context mut_m +begin + +locset_definition "hs_get_roots_loop_locs = prefixed ''hs_get_roots_loop''" +locset_definition "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 "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 "async_mo_ptest_locs = + (\pref\mut_async_mark_object_prefixes. {pref @ ''_mo_ptest''})" + +locset_definition "mo_ptest_locs = + (\pref\mut_async_mark_object_prefixes. {pref @ ''_mo_ptest''})" + +locset_definition "mo_valid_ref_locs = + (prefixed ''store_del'' \ prefixed ''store_ins'' \ {deref_del, lop_store_ins})" + +(*>*) +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 stores 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\{ ''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''))" + +locset_definition "ghost_honorary_root_nonempty_locs = prefixed ''store_del'' - {store_del_mo_null}" +locset_definition "not_idle_locs = suffixed ''_mo_ptest''" +locset_definition "ins_barrier_locs = prefixed ''store_ins''" +locset_definition "del_barrier1_locs = prefixed ''store_del_mo'' \ {lop_store_ins}" + +definition mark_object_invL :: "('field, 'mut, 'payload, 'ref) gc_pred" where +[inv]: "mark_object_invL = + (atS_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 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 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 ghost_honorary_root_nonempty_locs (mut_the_ref \<^bold>\ mut_ghost_honorary_root) + + \<^bold>\ atS_mut not_idle_locs (mut_phase \<^bold>\ \ph_Idle\ \<^bold>\ mut_ghost_hs_phase \<^bold>\ \hp_Idle\) + \<^bold>\ atS_mut hs_not_hp_Idle_locs (mut_ghost_hs_phase \<^bold>\ \hp_Idle\) + + \<^bold>\ atS_mut mo_ptest_locs (mut_phase \<^bold>= \ph_Idle\ \<^bold>\ (mut_ghost_hs_phase \<^bold>\ \{hp_Idle, hp_IdleInit}\ + \<^bold>\ (mut_ghost_hs_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_hs_phase \<^bold>\ \{hp_InitMark, hp_Mark}\ + \<^bold>\ (mut_ghost_hs_phase \<^bold>= \hp_IdleMarkSweep\ \<^bold>\ sys_phase \<^bold>\ \ph_Idle\)) + \<^bold>\ \<^bold>\NULL mut_new_ref + \<^bold>\ marked \<^bold>$ mut_the_new_ref ) + \<^bold>\ atS_mut ins_barrier_locs ( ( (mut_ghost_hs_phase \<^bold>= \hp_Mark\ + \<^bold>\ (mut_ghost_hs_phase \<^bold>= \hp_IdleMarkSweep\ \<^bold>\ sys_phase \<^bold>\ \ph_Idle\)) + \<^bold>\ (\s. \opt_r'. \tso_pending_store (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) ) + \<^bold>\ (mut_ref \<^bold>= mut_new_ref) ) +\ \deletion barrier\ + \<^bold>\ atS_mut del_barrier1_locs ( (mut_ghost_hs_phase \<^bold>= \hp_Mark\ + \<^bold>\ (mut_ghost_hs_phase \<^bold>= \hp_IdleMarkSweep\ \<^bold>\ sys_phase \<^bold>\ \ph_Idle\)) + \<^bold>\ (\s. \opt_r'. \tso_pending_store (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_hs_phase \<^bold>= \hp_Mark\ + \<^bold>\ (mut_ghost_hs_phase \<^bold>= \hp_IdleMarkSweep\ \<^bold>\ sys_phase \<^bold>\ \ph_Idle\)) + \<^bold>\ \<^bold>\NULL mut_ref + \<^bold>\ marked \<^bold>$ mut_the_ref ) +\\after \init_noop\. key: no pending white insertions \at_mut hs_noop_done\ which we get from @{const \handshake_invL\}.\ + \<^bold>\ at_mut mut_load (mut_tmp_ref \<^bold>\ mut_roots) + \<^bold>\ atS_mut ghost_honorary_root_empty_locs (EMPTY mut_ghost_honorary_root) )" + +end + + +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. + +\ + +(* FIXME this is an assertion? *) +definition (in mut_m) gc_W_empty_mut_inv :: "('field, 'mut, 'payload, 'ref) lsts_pred" where + "gc_W_empty_mut_inv = + ((EMPTY sys_W \<^bold>\ sys_ghost_hs_in_sync m \<^bold>\ \<^bold>\EMPTY (WL (mutator m))) + \<^bold>\ (\<^bold>\m'. \<^bold>\sys_ghost_hs_in_sync m' \<^bold>\ \<^bold>\EMPTY (WL (mutator m'))))" + +context gc +begin + +locset_definition gc_W_empty_locs :: "location set" where + "gc_W_empty_locs = + idle_locs \ init_locs \ sweep_locs \ {mark_load_fM, mark_store_fA, mark_end} + \ prefixed ''mark_noop'' + \ prefixed ''mark_loop_get_roots'' + \ prefixed ''mark_loop_get_work''" + +locset_definition "get_roots_UN_get_work_locs = hs_get_roots_locs \ hs_get_work_locs" +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}" + +definition gc_W_empty_invL :: "('field, 'mut, 'payload, 'ref) gc_pred" where +[inv]: "gc_W_empty_invL = + (atS_gc get_roots_UN_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))" + +end + + +subsection\Sweep loop invariants\ + +context gc +begin + +locset_definition "sweep_loop_locs = prefixed ''sweep_loop''" +locset_definition "sweep_loop_not_choose_ref_locs = (prefixed ''sweep_loop_'' - {sweep_loop_choose_ref})" + +definition sweep_loop_invL :: "('field, 'mut, 'payload, 'ref) gc_pred" where +[inv]: "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 sweep_loop_not_choose_ref_locs (gc_tmp_ref \<^bold>\ gc_refs))" + +text\ + +For showing 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 obj_fields_marked :: "('field, 'mut, 'payload, 'ref) lsts_pred" where + "obj_fields_marked = + (\<^bold>\f. \f\ \<^bold>\ (- gc_field_set) \<^bold>\ (\s. obj_at_field_on_heap (\r. marked r s) (gc_tmp_ref s) f s))" + +locset_definition "mark_loop_mo_locs = prefixed ''mark_loop_mo''" +locset_definition "obj_fields_marked_good_ref_locs = mark_loop_mo_locs \ {mark_loop_mark_field_done}" + +locset_definition + "ghost_honorary_grey_empty_locs = + (- { mark_loop_mo_co_unlock, mark_loop_mo_co_won, mark_loop_mo_co_W })" + +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} + \ mark_loop_mo_locs" + +definition obj_fields_marked_invL :: "('field, 'mut, 'payload, 'ref) gc_pred" where +[inv]: "obj_fields_marked_invL = + (atS_gc obj_fields_marked_locs (obj_fields_marked \<^bold>\ gc_tmp_ref \<^bold>\ gc_W) + \<^bold>\ atS_gc obj_fields_marked_good_ref_locs (\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 mark_loop_mo_locs (\<^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))" + +end + + +subsection\ The local innvariants collected \ + +definition (in gc) invsL :: "('field, 'mut, 'payload, '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)" + +definition (in mut_m) invsL :: "('field, 'mut, 'payload, '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)" + +definition invsL :: "('field, 'mut, 'payload, 'ref) gc_pred" where + "invsL = (gc.invsL \<^bold>\ (\<^bold>\m. mut_m.invsL m))" +(*<*) + +end +(*>*) diff --git a/thys/ConcurrentGC/Local_Invariants_Lemmas.thy b/thys/ConcurrentGC/Local_Invariants_Lemmas.thy new file mode 100644 --- /dev/null +++ b/thys/ConcurrentGC/Local_Invariants_Lemmas.thy @@ -0,0 +1,245 @@ +(*<*) +(* + * 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 Local_Invariants_Lemmas +imports + Local_Invariants +begin + +declare subst_all [simp del] [[simproc del: defined_all]] + +(*>*) +section\Local invariants lemma bucket\ + + +subsection\ Location facts\ + +(* FIXME loads more in StrongTricolour. These might be mostly about non-interference, in which case it might make sense to + split those proofs off into a separate theory? *) + +context mut_m +begin + +lemma hs_get_roots_loop_locs_subseteq_hs_get_roots_locs: + "hs_get_roots_loop_locs \ hs_get_roots_locs" +unfolding hs_get_roots_loop_locs_def hs_get_roots_locs_def by (fastforce intro: append_prefixD) + +lemma hs_pending_locs_subseteq_hs_pending_loaded_locs: + "hs_pending_locs \ hs_pending_loaded_locs" +unfolding hs_pending_locs_def hs_pending_loaded_locs_def by (fastforce intro: append_prefixD) + +lemma ht_loaded_locs_subseteq_hs_pending_loaded_locs: + "ht_loaded_locs \ hs_pending_loaded_locs" +unfolding ht_loaded_locs_def hs_pending_loaded_locs_def by (fastforce intro: append_prefixD) + +lemma hs_noop_locs_subseteq_hs_pending_loaded_locs: + "hs_noop_locs \ hs_pending_loaded_locs" +unfolding hs_noop_locs_def hs_pending_loaded_locs_def loc_defs by (fastforce intro: append_prefixD) + +lemma hs_noop_locs_subseteq_hs_pending_locs: + "hs_noop_locs \ hs_pending_locs" +unfolding hs_noop_locs_def hs_pending_locs_def loc_defs by (fastforce intro: append_prefixD) + +lemma hs_noop_locs_subseteq_ht_loaded_locs: + "hs_noop_locs \ ht_loaded_locs" +unfolding hs_noop_locs_def ht_loaded_locs_def loc_defs by (fastforce intro: append_prefixD) + +lemma hs_get_roots_locs_subseteq_hs_pending_loaded_locs: + "hs_get_roots_locs \ hs_pending_loaded_locs" +unfolding hs_get_roots_locs_def hs_pending_loaded_locs_def loc_defs by (fastforce intro: append_prefixD) + +lemma hs_get_roots_locs_subseteq_hs_pending_locs: + "hs_get_roots_locs \ hs_pending_locs" +unfolding hs_get_roots_locs_def hs_pending_locs_def loc_defs by (fastforce intro: append_prefixD) + +lemma hs_get_roots_locs_subseteq_ht_loaded_locs: + "hs_get_roots_locs \ ht_loaded_locs" +unfolding hs_get_roots_locs_def ht_loaded_locs_def loc_defs by (fastforce intro: append_prefixD) + +lemma hs_get_work_locs_subseteq_hs_pending_loaded_locs: + "hs_get_work_locs \ hs_pending_loaded_locs" +unfolding hs_get_work_locs_def hs_pending_loaded_locs_def loc_defs by (fastforce intro: append_prefixD) + +lemma hs_get_work_locs_subseteq_hs_pending_locs: + "hs_get_work_locs \ hs_pending_locs" +unfolding hs_get_work_locs_def hs_pending_locs_def loc_defs by (fastforce intro: append_prefixD) + +lemma hs_get_work_locs_subseteq_ht_loaded_locs: + "hs_get_work_locs \ ht_loaded_locs" +unfolding hs_get_work_locs_def ht_loaded_locs_def loc_defs by (fastforce intro: append_prefixD) + +end + +declare + mut_m.hs_get_roots_loop_locs_subseteq_hs_get_roots_locs[locset_cache] + mut_m.hs_pending_locs_subseteq_hs_pending_loaded_locs[locset_cache] + mut_m.ht_loaded_locs_subseteq_hs_pending_loaded_locs[locset_cache] + mut_m.hs_noop_locs_subseteq_hs_pending_loaded_locs[locset_cache] + mut_m.hs_noop_locs_subseteq_hs_pending_locs[locset_cache] + mut_m.hs_noop_locs_subseteq_ht_loaded_locs[locset_cache] + mut_m.hs_get_roots_locs_subseteq_hs_pending_loaded_locs[locset_cache] + mut_m.hs_get_roots_locs_subseteq_hs_pending_locs[locset_cache] + mut_m.hs_get_roots_locs_subseteq_ht_loaded_locs[locset_cache] + mut_m.hs_get_work_locs_subseteq_hs_pending_loaded_locs[locset_cache] + mut_m.hs_get_work_locs_subseteq_hs_pending_locs[locset_cache] + mut_m.hs_get_work_locs_subseteq_ht_loaded_locs[locset_cache] + +context gc +begin + +lemma get_roots_UN_get_work_locs_subseteq_ghost_honorary_grey_empty_locs: + "get_roots_UN_get_work_locs \ ghost_honorary_grey_empty_locs" +unfolding get_roots_UN_get_work_locs_def ghost_honorary_grey_empty_locs_def hs_get_roots_locs_def hs_get_work_locs_def loc_defs +by (fastforce intro: append_prefixD) + +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 loc_defs) +apply clarsimp +apply (drule mp) + apply (auto intro: append_prefixD)[1] +apply auto +done + +end + +declare + gc.get_roots_UN_get_work_locs_subseteq_ghost_honorary_grey_empty_locs[locset_cache] + gc.hs_get_roots_locs_subseteq_hp_IdleMarkSweep_locs[locset_cache] + gc.hs_get_work_locs_subseteq_hp_IdleMarkSweep_locs[locset_cache] + + +subsection\ \obj_fields_marked_inv\ \ + +context gc +begin + +lemma obj_fields_marked_eq_imp: + "eq_imp (\r'. gc_field_set \<^bold>\ gc_tmp_ref \<^bold>\ (\s. map_option obj_fields (sys_heap s r')) \<^bold>\ (\s. map_option obj_mark (sys_heap s r')) \<^bold>\ sys_fM \<^bold>\ tso_pending_mutate gc) + obj_fields_marked" +unfolding eq_imp_def obj_fields_marked_def obj_at_field_on_heap_def obj_at_def +apply (clarsimp simp: all_conj_distrib) +apply (rule iffI; clarsimp split: option.splits) + apply (intro allI conjI impI) + apply simp_all + apply (metis (no_types, hide_lams) option.distinct(1) option.map_disc_iff) + apply (metis (no_types, lifting) option.distinct(1) option.map_sel option.sel) +apply (intro allI conjI impI) + apply simp_all + apply (metis (no_types, hide_lams) option.distinct(1) option.map_disc_iff) +apply (metis (no_types, lifting) option.distinct(1) option.map_sel option.sel) +done + +lemma obj_fields_marked_UNIV[iff]: + "obj_fields_marked (s(gc := (s gc)\ field_set := UNIV \))" +unfolding obj_fields_marked_def by (simp add: fun_upd_apply) + +lemma obj_fields_marked_invL_eq_imp: + "eq_imp (\r' s. (AT s gc, s\ gc, map_option obj_fields (sys_heap s\ r'), map_option obj_mark (sys_heap s\ r'), sys_fM s\, sys_W s\, tso_pending_mutate gc s\)) + obj_fields_marked_invL" +unfolding eq_imp_def inv obj_at_def obj_at_field_on_heap_def +apply (clarsimp simp: all_conj_distrib cong: option.case_cong) +apply (rule iffI) + apply (intro conjI impI; clarsimp) + apply (subst eq_impD[OF obj_fields_marked_eq_imp]; force) + apply (clarsimp split: option.split_asm) + apply (metis (no_types, lifting) None_eq_map_option_iff option.simps(3)) + apply (metis (no_types, lifting) option.distinct(1) option.map_sel option.sel) + apply (metis (no_types, lifting) None_eq_map_option_iff option.simps(3)) + apply (metis (no_types, lifting) option.distinct(1) option.map_sel option.sel) + apply (subst (asm) (2) eq_impD[OF reaches_eq_imp]) + prefer 2 apply (drule spec, drule mp, assumption) + apply (metis (no_types) option.disc_eq_case(2) option.map_disc_iff) + apply (metis option.set_map) + apply (clarsimp split: option.splits) + apply (metis (no_types, hide_lams) atS_simps(2) atS_un obj_fields_marked_good_ref_locs_def) + apply (metis (no_types, hide_lams) map_option_eq_Some option.inject option.simps(9)) + apply (metis (no_types, hide_lams) map_option_eq_Some option.inject option.simps(9)) + apply (metis (no_types, hide_lams) map_option_eq_Some option.inject option.simps(9)) +(* cut and paste *) +apply (intro conjI impI; clarsimp) + apply (subst eq_impD[OF obj_fields_marked_eq_imp]; force) + apply (clarsimp split: option.split_asm) + apply (metis (no_types, lifting) None_eq_map_option_iff option.simps(3)) + apply (metis (no_types, lifting) option.distinct(1) option.map_sel option.sel) + apply (metis (no_types, lifting) None_eq_map_option_iff option.simps(3)) + apply (metis (no_types, lifting) option.distinct(1) option.map_sel option.sel) + apply (subst (asm) (2) eq_impD[OF reaches_eq_imp]) + prefer 2 apply (drule spec, drule mp, assumption) + apply (metis (no_types, lifting) None_eq_map_option_iff option.case_eq_if) + apply (metis option.set_map) +apply (clarsimp split: option.splits) + apply (metis (no_types, hide_lams) atS_simps(2) atS_un obj_fields_marked_good_ref_locs_def) + apply (metis (no_types, hide_lams) map_option_eq_Some option.inject option.simps(9)) + apply (metis (no_types, hide_lams) map_option_eq_Some option.inject option.simps(9)) + apply (metis (no_types, hide_lams) map_option_eq_Some option.inject option.simps(9)) +done + +lemma obj_fields_marked_mark_field_done[iff]: + "\ obj_at_field_on_heap (\r. marked r s) (gc_tmp_ref s) (gc_field s) s; obj_fields_marked s \ + \ obj_fields_marked (s(gc := (s gc)\field_set := gc_field_set s - {gc_field s}\))" +unfolding obj_fields_marked_def obj_at_field_on_heap_def by (fastforce simp: fun_upd_apply split: option.splits obj_at_splits)+ + +end + +lemmas gc_obj_fields_marked_inv_fun_upd[simp] = eq_imp_fun_upd[OF gc.obj_fields_marked_eq_imp, simplified eq_imp_simps, rule_format] +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] + + +subsection\ mark object \ + +context mark_object +begin + +lemma mark_object_invL_eq_imp: + "eq_imp (\(_::unit) s. (AT s p, s\ p, sys_heap s\, sys_fM s\, sys_mem_store_buffers p s\)) + mark_object_invL" +unfolding eq_imp_def +apply clarsimp +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 white_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 + +lemma 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\)) + (mut_m.mark_object_invL m)" +apply (clarsimp simp: eq_imp_def mut_m.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_store_buffers (mutator m) s\) + \ mw_Mutate r f opt_r' \ set (sys_mem_store_buffers (mutator m) s'\)") + apply (clarsimp cong: option.case_cong) + apply (metis (mono_tags, lifting) filter_set member_filter) + 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] + +(*<*) + +end +(*>*) diff --git a/thys/ConcurrentGC/MarkObject.thy b/thys/ConcurrentGC/MarkObject.thy --- a/thys/ConcurrentGC/MarkObject.thy +++ b/thys/ConcurrentGC/MarkObject.thy @@ -1,1104 +1,428 @@ (*<*) (* * 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 + Global_Invariants_Lemmas + Local_Invariants_Lemmas + Tactics begin -declare subst_all [simp del] [[simproc del: defined_all]] - (*>*) -subsection\Object colours, reference validity, worklist validity\ +section\ Mark Object \ 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. +These are the most intricate proofs in this development. \ -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" +context mut_m 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]: +lemma 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) \ + \<^bold>\ LSTP (phase_rel_inv \<^bold>\ handshake_phase_inv \<^bold>\ phase_rel_inv \<^bold>\ tso_store_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 +proof (vcg_jackhammer, vcg_name_cases) + case (store_ins_mo_ptest s s' obj) then show ?case + apply - + apply (drule handshake_phase_invD) + apply (drule phase_rel_invD) + apply (clarsimp simp: phase_rel_def) + apply (cases "sys_ghost_hs_phase s\"; simp add: hp_step_rel_def; elim disjE; simp; force) + done +next case (store_ins_mo_phase s s') then show ?case + apply - + apply (drule handshake_phase_invD) + apply (drule phase_rel_invD) + apply (clarsimp simp: phase_rel_def) + apply (case_tac "sys_ghost_hs_phase s\"; simp add: hp_step_rel_def; elim disjE; simp; force) + done +next case (store_del_mo_phase s s' y) then show ?case + apply - + apply (drule handshake_phase_invD) + apply (drule phase_rel_invD) + apply (clarsimp simp: phase_rel_def) + apply (case_tac "sys_ghost_hs_phase s\"; simp add: hp_step_rel_def; elim disjE; simp; force) + done +next case (deref_del s s' opt_r') then show ?case + apply - + apply (rule obj_at_field_on_heapE[OF obj_at_field_on_heap_no_pending_stores[where m=m]]) + apply auto + done +next case (hs_get_roots_loop_mo_phase s s' y) then show ?case + apply - + apply (drule handshake_phase_invD) + apply (drule phase_rel_invD) + apply (clarsimp simp: phase_rel_def hp_step_rel_def) + done +qed 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]: +lemma 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) \ + \<^bold>\ LSTP (handshake_phase_inv \<^bold>\ valid_W_inv \<^bold>\ tso_store_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]) +proof(vcg_jackhammer, vcg_name_cases) + case store_ins_mo_null then show ?case by (metis reachableI(1) valid_refs_invD(8)) +next case store_ins_mo_mark then show ?case by (clarsimp split: obj_at_splits) +next case (store_ins_mo_ptest s s' obj) then show ?case by (simp add: valid_W_inv_no_mark_stores_invD filter_empty_conv) metis +next case store_ins_mo_co_won then show ?case by metis +next case store_ins_mo_mtest then show ?case by metis +next case store_ins_mo_co_ctest0 then show ?case by (metis whiteI) +next case (store_ins_mo_co_ctest s s' obj) then show ?case + apply (elim disjE; clarsimp split: obj_at_splits) + apply metis + done +qed -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]: +lemma 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) \ + \<^bold>\ LSTP (handshake_phase_inv \<^bold>\ valid_W_inv \<^bold>\ tso_store_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) +proof(vcg_jackhammer, vcg_name_cases) + case store_del_mo_co_ctest0 then show ?case by blast +next case store_del_mo_co_ctest then show ?case by (clarsimp split: obj_at_splits) +next case store_del_mo_ptest then show ?case by (auto dest: valid_W_inv_no_mark_stores_invD) +next case store_del_mo_mark then show ?case by (clarsimp split: obj_at_splits) +next case store_del_mo_null then show ?case by (auto dest: valid_refs_invD) +qed -lemma (in mut_m) mut_get_roots_mark_object_invL[intro]: +lemma 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) \ + \<^bold>\ LSTP (handshake_phase_inv \<^bold>\ valid_W_inv \<^bold>\ tso_store_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))" -(*<*) +proof(vcg_jackhammer, vcg_name_cases) + case hs_get_roots_loop_mo_co_ctest0 then show ?case by blast +next case hs_get_roots_loop_mo_co_ctest then show ?case by (clarsimp split: obj_at_splits) +next case hs_get_roots_loop_mo_ptest then show ?case by (auto dest: valid_W_inv_no_mark_stores_invD split: obj_at_splits) +next case hs_get_roots_loop_mo_mark then show ?case by (clarsimp split: obj_at_splits) +next case hs_get_roots_loop_mo_null then show ?case by (auto dest: valid_W_inv_no_mark_stores_invD split: obj_at_splits) +qed -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\\ +end -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''" +lemma (in mut_m') mut_mark_object_invL[intro]: + notes obj_at_field_on_heap_splits[split] + notes fun_upd_apply[simp] + shows + "\ mark_object_invL \ mutator m'" +by (vcg_chainsaw mark_object_invL_def) -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) +subsection\ @{term "obj_fields_marked_inv"} \ -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] +context gc +begin -lemma (in gc) gc_mark_mark_object_invL[intro]: +lemma 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) +by vcg_jackhammer (auto dest: valid_W_inv_no_mark_stores_invD split: obj_at_splits) -lemma (in gc) obj_fields_marked_invL[intro]: +lemma 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) \ + \<^bold>\ LSTP (tso_store_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] +proof(vcg_jackhammer, vcg_name_cases) + case (mark_loop_mark_field_done s s') then show ?case by - (rule obj_fields_marked_mark_field_done, auto) +next case (mark_loop_mark_deref s s') + then have "grey (gc_tmp_ref s\) s\" by blast + with mark_loop_mark_deref show ?case +apply (clarsimp split: obj_at_field_on_heap_splits) +apply (rule conjI) + apply (metis (no_types) case_optionE obj_at_def valid_W_invE(3)) +apply clarsimp +apply (erule valid_refs_invD; auto simp: obj_at_def ranI reaches_step(2)) +done +qed -(* 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 +end + +context sys +begin + +lemma mut_store_ins_mark_object_invL[intro]: + notes mut_m_not_idle_no_fM_writeD[where m=m, dest!] + notes not_blocked_def[simp] + notes fun_upd_apply[simp] + notes if_split_asm[split del] + 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_store_inv) \ + sys + \ mut_store_ins.mark_object_invL m \" +proof(vcg_chainsaw mut_m.mark_object_invL_def mut_m.tso_lock_invL_def mut_m.mut_store_ins_mark_object_invL_def2, vcg_name_cases "mutator m") + case (store_ins_mo_fM s s' p w ws ref fl) then show ?case + apply (clarsimp simp: do_store_action_def filter_empty_conv mut_m_not_idle_no_fM_write valid_W_invD split: mem_store_action.splits obj_at_splits if_splits) + apply (metis (no_types, lifting) valid_W_invD(1)) + done +next case (store_ins_mo_mtest s s' p w ws y ya) then show ?case + apply (clarsimp simp: do_store_action_def filter_empty_conv mut_m_not_idle_no_fM_write valid_W_invD split: mem_store_action.splits obj_at_splits if_splits) + done +next case (store_ins_mo_phase s s' p w ws y) then show ?case + apply (clarsimp simp: do_store_action_def filter_empty_conv mut_m_not_idle_no_fM_write valid_W_invD split: mem_store_action.splits obj_at_splits if_splits) + done +next case (store_ins_mo_ptest s s' p w ws y) then show ?case + apply (clarsimp simp: do_store_action_def filter_empty_conv mut_m_not_idle_no_fM_write valid_W_invD split: mem_store_action.splits obj_at_splits) + done +next case (store_ins_mo_co_lock s s' p w ws y) then show ?case + apply (intro conjI impI notI; clarsimp simp: do_store_action_def filter_empty_conv mut_m_not_idle_no_fM_write split: mem_store_action.splits obj_at_splits) + apply metis + done +next case (store_ins_mo_co_cmark s s' w ws y) then show ?case + apply (clarsimp simp: do_store_action_def filter_empty_conv mut_m_not_idle_no_fM_write split: mem_store_action.splits obj_at_splits) + done +next case (store_ins_mo_co_ctest s s' w ws y) then show ?case + apply (clarsimp simp: do_store_action_def filter_empty_conv mut_m_not_idle_no_fM_write split: mem_store_action.splits obj_at_splits) + done +next case (store_ins_mo_co_mark s s' w ws y) then show ?case + apply (clarsimp simp: do_store_action_def filter_empty_conv mut_m_not_idle_no_fM_write split: mem_store_action.splits obj_at_splits) + done +next case (store_ins_mo_co_unlock s s' w ws y) then show ?case + apply (clarsimp simp: do_store_action_def filter_empty_conv mut_m_not_idle_no_fM_write split: mem_store_action.splits obj_at_splits) + done +next case (store_ins_mo_co_won s s' p w ws y) then show ?case + apply (intro conjI impI notI; clarsimp simp: do_store_action_def filter_empty_conv mut_m_not_idle_no_fM_write valid_W_invD(1) split: mem_store_action.splits obj_at_splits) + done +next case (store_ins_mo_co_W s s' p w ws y) then show ?case + apply (clarsimp simp: do_store_action_def filter_empty_conv mut_m_not_idle_no_fM_write valid_W_invD(1) split: mem_store_action.splits obj_at_splits) + apply auto + done +qed + +lemma mut_store_del_mark_object_invL[intro]: + notes mut_m_not_idle_no_fM_writeD[where m=m, dest!] + notes not_blocked_def[simp] + notes fun_upd_apply[simp] + notes if_split_asm[split del] + 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_store_inv) \ + sys + \ mut_store_del.mark_object_invL m \" +proof(vcg_chainsaw mut_m.mark_object_invL_def mut_m.tso_lock_invL_def mut_m.mut_store_del_mark_object_invL_def2, vcg_name_cases "mutator m") + case (store_del_mo_fM s s' p w ws y ya) then show ?case + apply (clarsimp simp: do_store_action_def filter_empty_conv mut_m_not_idle_no_fM_write valid_W_invD split: mem_store_action.splits obj_at_splits if_splits) + apply (metis (no_types, lifting) valid_W_invD(1)) + done +next case (store_del_mo_mtest s s' p w ws y ya) then show ?case + apply (clarsimp simp: do_store_action_def filter_empty_conv mut_m_not_idle_no_fM_write valid_W_invD split: mem_store_action.splits obj_at_splits if_splits) + done +next case (store_del_mo_phase s s' p w ws y) then show ?case + apply (clarsimp simp: do_store_action_def filter_empty_conv mut_m_not_idle_no_fM_write valid_W_invD split: mem_store_action.splits obj_at_splits) + done +next case (store_del_mo_ptest s s' p w ws y) then show ?case + apply (clarsimp simp: do_store_action_def filter_empty_conv mut_m_not_idle_no_fM_write valid_W_invD split: mem_store_action.splits obj_at_splits) + done +next case (store_del_mo_co_lock s s' p w ws y) then show ?case + apply (intro conjI impI notI; clarsimp simp: do_store_action_def filter_empty_conv mut_m_not_idle_no_fM_write split: mem_store_action.splits obj_at_splits) + apply metis + done +next case (store_del_mo_co_cmark s s' w ws y) then show ?case + apply (clarsimp simp: do_store_action_def filter_empty_conv mut_m_not_idle_no_fM_write valid_W_invD split: mem_store_action.splits obj_at_splits) + done +next case (store_del_mo_co_ctest s s' w ws y) then show ?case + apply (clarsimp simp: do_store_action_def filter_empty_conv mut_m_not_idle_no_fM_write valid_W_invD split: mem_store_action.splits obj_at_splits) + done +next case (store_del_mo_co_mark s s' w ws y) then show ?case + apply (clarsimp simp: do_store_action_def filter_empty_conv mut_m_not_idle_no_fM_write valid_W_invD split: mem_store_action.splits obj_at_splits) + done +next case (store_del_mo_co_unlock s s' w ws y) then show ?case + apply (clarsimp simp: do_store_action_def filter_empty_conv mut_m_not_idle_no_fM_write valid_W_invD split: mem_store_action.splits obj_at_splits) + apply (metis (mono_tags, lifting) filter_empty_conv valid_W_invD(1)) + done +next case (store_del_mo_co_won s s' p w ws y) then show ?case + apply (clarsimp simp: do_store_action_def filter_empty_conv mut_m_not_idle_no_fM_write valid_W_invD split: mem_store_action.splits obj_at_splits) + apply auto + done +next case (store_del_mo_co_W s s' p w ws y) then show ?case + apply (clarsimp simp: do_store_action_def filter_empty_conv mut_m_not_idle_no_fM_write valid_W_invD split: mem_store_action.splits obj_at_splits) + apply auto + done +qed + +lemma mut_get_roots_mark_object_invL[intro]: + notes not_blocked_def[simp] + notes p_not_sys[simp] + notes mut_m.handshake_phase_invD[where m=m, dest!] + notes fun_upd_apply[simp] + notes if_split_asm[split del] + 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_store_inv) \ + sys + \ mut_get_roots.mark_object_invL m \" +proof(vcg_chainsaw mut_m.tso_lock_invL_def mut_m.handshake_invL_def mut_m.mut_get_roots_mark_object_invL_def2, vcg_name_cases "mutator m") + case (hs_get_roots_loop_mo_fM s s' p w ws y ya) then show ?case + apply (clarsimp simp: do_store_action_def valid_W_invD split: mem_store_action.splits obj_at_splits if_splits) + apply (metis (no_types, lifting) valid_W_invD(1)) + apply (force simp: fM_rel_inv_def fM_rel_def hp_step_rel_def filter_empty_conv valid_W_invD)+ + done +next case (hs_get_roots_loop_mo_mtest s s' p w ws y ya) then show ?case + apply (clarsimp simp: do_store_action_def valid_W_invD split: mem_store_action.splits obj_at_splits) + apply (force simp: fM_rel_inv_def fM_rel_def hp_step_rel_def filter_empty_conv valid_W_invD)+ + done +next case (hs_get_roots_loop_mo_phase s s' p w ws y) then show ?case + apply (clarsimp simp: do_store_action_def valid_W_invD split: mem_store_action.splits obj_at_splits) + apply (force simp: fM_rel_inv_def fM_rel_def hp_step_rel_def filter_empty_conv valid_W_invD)+ + done +next case (hs_get_roots_loop_mo_ptest s s' p w ws y) then show ?case + apply (clarsimp simp: do_store_action_def valid_W_invD split: mem_store_action.splits obj_at_splits) + apply (force simp: fM_rel_inv_def fM_rel_def hp_step_rel_def filter_empty_conv valid_W_invD)+ + done +next case (hs_get_roots_loop_mo_co_lock s s' p w ws y) then show ?case + apply (clarsimp simp: do_store_action_def valid_W_invD split: mem_store_action.splits obj_at_splits) + apply (force simp: fM_rel_inv_def fM_rel_def hp_step_rel_def filter_empty_conv valid_W_invD)+ + done +next case (hs_get_roots_loop_mo_co_cmark s s' w ws y) then show ?case + apply (clarsimp simp: do_store_action_def valid_W_invD split: mem_store_action.splits obj_at_splits) + done +next case (hs_get_roots_loop_mo_co_ctest s s' w ws y) then show ?case + apply (clarsimp simp: do_store_action_def valid_W_invD split: mem_store_action.splits obj_at_splits) + done +next case (hs_get_roots_loop_mo_co_mark s s' w ws y) then show ?case + apply (clarsimp simp: do_store_action_def valid_W_invD split: mem_store_action.splits obj_at_splits) + done +next case (hs_get_roots_loop_mo_co_unlock s s' w ws y) then show ?case + apply (clarsimp simp: do_store_action_def valid_W_invD split: mem_store_action.splits obj_at_splits) + done +next case (hs_get_roots_loop_mo_co_won s s' p w ws y) then show ?case + apply (clarsimp simp: do_store_action_def valid_W_invD split: mem_store_action.splits obj_at_splits) + apply (force simp: fM_rel_inv_def fM_rel_def hp_step_rel_def filter_empty_conv valid_W_invD)+ + done +next case (hs_get_roots_loop_mo_co_W s s' p w ws y) then show ?case + apply (clarsimp simp: do_store_action_def valid_W_invD split: mem_store_action.splits obj_at_splits) + apply (force simp: fM_rel_inv_def fM_rel_def hp_step_rel_def filter_empty_conv valid_W_invD)+ + done +qed + +lemma gc_mark_mark_object_invL[intro]: + notes fun_upd_apply[simp] + notes if_split_asm[split del] + 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_store_inv) \ + sys + \ gc_mark.mark_object_invL \" +proof(vcg_chainsaw gc.gc_mark_mark_object_invL_def2 gc.tso_lock_invL_def gc.phase_invL_def gc.fM_fA_invL_def gc.handshake_invL_def, vcg_name_cases "gc") + case (mark_loop_mo_fM s s' p w ws y ya) then show ?case + apply (clarsimp simp: do_store_action_def not_blocked_def fM_rel_def filter_empty_conv p_not_sys valid_W_invD + split: mem_store_action.splits if_splits) + apply (auto split: obj_at_splits) + done +next case (mark_loop_mo_mtest s s' p w ws y ya) then show ?case + apply (clarsimp simp: do_store_action_def not_blocked_def fM_rel_def filter_empty_conv p_not_sys valid_W_invD + split: mem_store_action.splits) + apply (auto split: obj_at_splits) + done +next case (mark_loop_mo_phase s s' p w ws y) then show ?case + apply (clarsimp simp: do_store_action_def not_blocked_def fM_rel_def filter_empty_conv p_not_sys valid_W_invD + split: mem_store_action.splits) + apply (auto split: obj_at_splits) + done +next case (mark_loop_mo_ptest s s' p w ws y) then show ?case + apply (clarsimp simp: do_store_action_def not_blocked_def fM_rel_def filter_empty_conv p_not_sys valid_W_invD + split: mem_store_action.splits) + apply (auto split: obj_at_splits) + done +next case (mark_loop_mo_co_lock s s' p w ws y) then show ?case + apply (clarsimp simp: do_store_action_def not_blocked_def fM_rel_def filter_empty_conv p_not_sys valid_W_invD + split: mem_store_action.splits) + apply (auto split: obj_at_splits) + done +next case (mark_loop_mo_co_cmark s s' p w ws y) then show ?case + apply (clarsimp simp: do_store_action_def not_blocked_def fM_rel_def filter_empty_conv p_not_sys valid_W_invD + split: mem_store_action.splits) + done +next case (mark_loop_mo_co_ctest s s' p w ws y) then show ?case + apply (clarsimp simp: do_store_action_def not_blocked_def fM_rel_def filter_empty_conv p_not_sys valid_W_invD + split: mem_store_action.splits) + done +next case (mark_loop_mo_co_mark s s' p w ws y) then show ?case + apply (clarsimp simp: do_store_action_def not_blocked_def fM_rel_def filter_empty_conv p_not_sys valid_W_invD + split: mem_store_action.splits) + done +next case (mark_loop_mo_co_unlock s s' p w ws y) then show ?case + apply (clarsimp simp: do_store_action_def not_blocked_def fM_rel_def filter_empty_conv p_not_sys + split: mem_store_action.splits) + apply auto + done +next case (mark_loop_mo_co_won s s' p w ws y) then show ?case + apply (clarsimp simp: do_store_action_def not_blocked_def fM_rel_def filter_empty_conv p_not_sys valid_W_invD + split: mem_store_action.splits) + apply (auto split: obj_at_splits) + done +next case (mark_loop_mo_co_W s s' p w ws y) then show ?case + apply (clarsimp simp: do_store_action_def not_blocked_def fM_rel_def filter_empty_conv p_not_sys valid_W_invD + split: mem_store_action.splits) + apply (auto split: obj_at_splits) + done +qed + +end + +lemma (in mut_m') mut_get_roots_mark_object_invL[intro]: + "\ mut_get_roots.mark_object_invL m \ mutator m'" +by vcg_chainsaw + +lemma (in mut_m') mut_store_ins_mark_object_invL[intro]: + "\ mut_store_ins.mark_object_invL m \ mutator m'" +by vcg_chainsaw + +lemma (in mut_m') mut_store_del_mark_object_invL[intro]: + "\ mut_store_del.mark_object_invL m \ mutator m'" +by vcg_chainsaw + +lemma (in gc) mut_get_roots_mark_object_invL[intro]: + "\ 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_chainsaw mut_m.handshake_invL_def mut_m.mut_get_roots_mark_object_invL_def2) 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 + "\ handshake_invL \<^bold>\ gc.handshake_invL \<^bold>\ gc.obj_fields_marked_invL + \<^bold>\ LSTP (tso_store_inv \<^bold>\ valid_refs_inv) \ + mutator m + \ gc.obj_fields_marked_invL \" +apply (vcg_chainsaw gc.obj_fields_marked_invL_def gc.handshake_invL_def) (* 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+ + apply (clarsimp simp: gc.obj_fields_marked_def fun_upd_apply) + 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 simp: fun_upd_apply split: obj_at_splits; fail) + apply (erule greyI) + apply (clarsimp simp: fun_upd_apply split: obj_at_splits; fail) + apply (clarsimp simp: fun_upd_apply split: obj_at_field_on_heap_splits; fail) +apply (clarsimp simp: fun_upd_apply) 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 +by (vcg_chainsaw gc.gc_mark_mark_object_invL_def2) -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,1008 +1,1080 @@ (*<*) (* * 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" + "HOL-Library.Sublist" begin +(* From 40e16c534243 by Makarius. Doesn't seem to have a huge impact on run time now (2021-01-07) *) declare subst_all [simp del] [[simproc del: defined_all]] (*>*) -section\The Schism garbage collector \label{sec:gc-model}\ +section\A model of a 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.} +@{bold \@{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" +type_synonym location = string 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 +datatype hs_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}. +\S\ref{sec:gc_handshakes}. \ -datatype handshake_phase - = hp_Idle \\ done 1 noop \ +datatype hs_phase + = hp_Idle \ \done 1 noop\ | hp_IdleInit | hp_InitMark - | hp_Mark \\ done 4 noops \ - | hp_IdleMarkSweep \\ done get roots \ + | hp_Mark \ \done 4 noops\ + | hp_IdleMarkSweep \ \done get roots\ -definition handshake_step :: "handshake_phase \ handshake_phase" where - "handshake_step ph \ case ph of +definition + hs_step :: "hs_phase \ hs_phase" +where + "hs_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" + | 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}. +An object consists of a garbage collection mark and two partial +maps. Firstly the types: - @{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. +\<^item> @{typ "'field"} is the abstract type of fields. +\<^item> @{typ "'ref"} is the abstract type of object references. +\<^item> @{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. +The maps: + +\<^item> \obj_fields\ maps @{typ "'field"}s to object + references (or @{const "None"} signifying \texttt{NULL} or type + error). +\<^item> \obj_payload\ maps a @{typ "'field"} to non-reference + data. For convenience we similarly allow that to be \texttt{NULL}. \ type_synonym gc_mark = bool -record ('field, 'ref) object = +record ('field, 'payload, 'ref) object = obj_mark :: "gc_mark" - obj_fields :: "'field \ 'ref option" + obj_fields :: "'field \ 'ref" + obj_payload :: "'field \ 'payload" text\ -The TSO store buffers track write actions, represented by \('field, 'ref) mem_write_action\. +The TSO store buffers track store actions, represented by \('field, 'ref) mem_store_action\. \ -datatype ('field, 'ref) mem_write_action +datatype ('field, 'payload, 'ref) mem_store_action = mw_Mark 'ref gc_mark | mw_Mutate 'ref 'field "'ref option" + | mw_Mutate_Payload 'ref 'field "'payload option" | mw_fA gc_mark | mw_fM gc_mark | mw_Phase gc_phase text\ +An action is a request by a mutator or the garbage collector to the +system. + +\ + +datatype ('field, 'ref) mem_load_action + = mr_Ref 'ref 'field + | mr_Payload 'ref 'field + | mr_Mark 'ref + | mr_Phase + | mr_fM + | mr_fA + +datatype ('field, 'mut, 'payload, 'ref) request_op + = ro_MFENCE + | ro_Load "('field, 'ref) mem_load_action" + | ro_Store "('field, 'payload, 'ref) mem_store_action" + | ro_Lock + | ro_Unlock + | ro_Alloc + | ro_Free 'ref + | ro_hs_gc_load_pending 'mut + | ro_hs_gc_store_type hs_type + | ro_hs_gc_store_pending 'mut + | ro_hs_gc_load_W + | ro_hs_mut_load_pending + | ro_hs_mut_load_type + | ro_hs_mut_done "'ref set" + +abbreviation "LoadfM \ ro_Load mr_fM" +abbreviation "LoadMark r \ ro_Load (mr_Mark r)" +abbreviation "LoadPayload r f \ ro_Load (mr_Payload r f)" +abbreviation "LoadPhase \ ro_Load mr_Phase" +abbreviation "LoadRef r f \ ro_Load (mr_Ref r f)" + +abbreviation "StorefA m \ ro_Store (mw_fA m)" +abbreviation "StorefM m \ ro_Store (mw_fM m)" +abbreviation "StoreMark r m \ ro_Store (mw_Mark r m)" +abbreviation "StorePayload r f pl \ ro_Store (mw_Mutate_Payload r f pl)" +abbreviation "StorePhase ph \ ro_Store (mw_Phase ph)" +abbreviation "StoreRef r f r' \ ro_Store (mw_Mutate r f r')" + +type_synonym ('field, 'mut, 'payload, 'ref) request + = "'mut process_name \ ('field, 'mut, 'payload, 'ref) request_op" + +datatype ('field, 'payload, 'ref) response + = mv_Bool "bool" + | mv_Mark "gc_mark option" + | mv_Payload "'payload option" \\ the requested reference might be invalid \ + | mv_Phase "gc_phase" + | mv_Ref "'ref option" + | mv_Refs "'ref set" + | mv_Void + | mv_hs_type hs_type + +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 = +record ('field, 'mut, 'payload, 'ref) local_state = \ \System-specific fields\ - heap :: "'ref \ ('field, 'ref) object option" + heap :: "'ref \ ('field, 'payload, 'ref) object" \ \TSO memory state\ - mem_write_buffers :: "'mut process_name \ ('field, 'ref) mem_write_action list" + mem_store_buffers :: "'mut process_name \ ('field, 'payload, 'ref) mem_store_action list" mem_lock :: "'mut process_name option" - \ \The state of the handshakes\ - handshake_type :: "handshake_type" - handshake_pending :: "'mut \ bool" + \ \Handshake state\ + hs_pending :: "'mut \ bool" \ \Ghost state\ - ghost_handshake_in_sync :: "'mut \ bool" - ghost_handshake_phase :: "handshake_phase" + ghost_hs_in_sync :: "'mut \ bool" + ghost_hs_phase :: "hs_phase" \ \Mutator-specific temporaries\ new_ref :: "'ref option" roots :: "'ref set" ghost_honorary_root :: "'ref set" + payload_value :: "'payload option" + mutator_data :: "'field \ 'payload" + mutator_hs_pending :: "bool" \ \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" + \ \Handshake state\ + hs_type :: "hs_type" \ \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, 'payload, 'ref) gc_com + = "(('field, 'payload, 'ref) response, location, ('field, 'mut, 'payload, 'ref) request, ('field, 'mut, 'payload, 'ref) local_state) com" +type_synonym ('field, 'mut, 'payload, 'ref) gc_loc_comp + = "(('field, 'payload, 'ref) response, location, ('field, 'mut, 'payload, 'ref) request, ('field, 'mut, 'payload, 'ref) local_state) loc_comp" +type_synonym ('field, 'mut, 'payload, 'ref) gc_pred + = "(('field, 'payload, 'ref) response, location, 'mut process_name, ('field, 'mut, 'payload, 'ref) request, ('field, 'mut, 'payload, 'ref) local_state) state_pred" +type_synonym ('field, 'mut, 'payload, 'ref) gc_system + = "(('field, 'payload, 'ref) response, location, 'mut process_name, ('field, 'mut, 'payload, 'ref) request, ('field, 'mut, 'payload, '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, 'payload, 'ref) gc_event + = "('field, 'mut, 'payload, 'ref) request \ ('field, 'payload, 'ref) response" +type_synonym ('field, 'mut, 'payload, 'ref) gc_history + = "('field, 'mut, 'payload, 'ref) gc_event list" -type_synonym ('field, 'mut, 'ref) lst_pred - = "('field, 'mut, 'ref) local_state \ bool" +type_synonym ('field, 'mut, 'payload, 'ref) lst_pred + = "('field, 'mut, 'payload, 'ref) local_state \ bool" -type_synonym ('field, 'mut, 'ref) lsts - = "'mut process_name \ ('field, 'mut, 'ref) local_state" +type_synonym ('field, 'mut, 'payload, 'ref) lsts + = "'mut process_name \ ('field, 'mut, 'payload, 'ref) local_state" -type_synonym ('field, 'mut, 'ref) lsts_pred - = "('field, 'mut, 'ref) lsts \ bool" +type_synonym ('field, 'mut, 'payload, 'ref) lsts_pred + = "('field, 'mut, 'payload, '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 +typically use their contents by prefixing identifiers with 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 lock_syn :: "location \ ('field, 'mut, 'payload, 'ref) gc_com" where + "lock_syn l \ \l\ Request (\s. (p, ro_Lock)) (\_ s. {s})" +notation lock_syn ("\_\ lock") -abbreviation unlock :: "location \ ('field, 'mut, 'ref) gc_com" where - "unlock l \ \l\ Request (\s. (p, ro_Unlock)) (\_ s. {s})" -notation unlock ("\_\ unlock") +abbreviation unlock_syn :: "location \ ('field, 'mut, 'payload, 'ref) gc_com" where + "unlock_syn l \ \l\ Request (\s. (p, ro_Unlock)) (\_ s. {s})" +notation unlock_syn ("\_\ unlock") abbreviation - read_mark :: "location \ (('field, 'mut, 'ref) local_state \ 'ref) + load_mark_syn :: "location \ (('field, 'mut, 'payload, '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" + \ ('field, 'mut, 'payload, 'ref) local_state + \ ('field, 'mut, 'payload, 'ref) local_state) \ ('field, 'mut, 'payload, '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") + "load_mark_syn l r upd \ \l\ Request (\s. (p, LoadMark (r s))) (\mv s. { upd \m\ s |m. mv = mv_Mark m })" +notation load_mark_syn ("\_\ load'_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 load_fM_syn :: "location \ ('field, 'mut, 'payload, 'ref) gc_com" where + "load_fM_syn l \ \l\ Request (\s. (p, ro_Load mr_fM)) (\mv s. { s\fM := m\ |m. mv = mv_Mark (Some m) })" +notation load_fM_syn ("\_\ load'_fM") abbreviation - read_phase :: "location \ ('field, 'mut, 'ref) gc_com" + load_phase :: "location \ ('field, 'mut, 'payload, '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") + "load_phase l \ \l\ Request (\s. (p, LoadPhase)) (\mv s. { s\phase := ph\ |ph. mv = mv_Phase ph })" +notation load_phase ("\_\ load'_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 + store_mark_syn :: "location \ (('field, 'mut, 'payload, 'ref) local_state \ 'ref) \ (('field, 'mut, 'payload, 'ref) local_state \ bool) \ ('field, 'mut, 'payload, 'ref) gc_com" +where + "store_mark_syn l r fl \ \l\ Request (\s. (p, StoreMark (r s) (fl s))) (\_ s. { s\ ghost_honorary_grey := {r s} \ })" +notation store_mark_syn ("\_\ store'_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") +abbreviation + add_to_W_syn :: "location \ (('field, 'mut, 'payload, 'ref) local_state \ 'ref) \ ('field, 'mut, 'payload, 'ref) gc_com" +where + "add_to_W_syn l r \ \l\ \\s. s\ W := W s \ {r s}, ghost_honorary_grey := {} \\" +notation add_to_W_syn ("\_\ 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 +TSO means we cannot avoid having the mark store 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 \ +definition + mark_object_fn :: "location \ ('field, 'mut, 'payload, '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_mark''\ load_mark (the \ ref) mark_update ;; + \l @ ''_mo_fM''\ load_fM ;; \l @ ''_mo_mtest''\ IF mark \<^bold>\ Some \ fM THEN - \l @ ''_mo_phase''\ read_phase ;; + \l @ ''_mo_phase''\ load_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_cmark''\ load_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 + \l @ ''_mo_co_mark''\ store_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}\ +subsection\Handshakes \label{sec:gc_handshakes}\ 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 +The garbage collector needs to synchronise with the mutators. +Here we do so by having the GC busy-wait: 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. +collector's local worklist during \get_roots\ and \get_work\ handshakes. +We carefully model the effect these handshakes have on the processes' TSO buffers. -The system and mutators track handshake phases using ghost state. +The system and mutators track handshake phases using ghost state; see +\S\ref{sec:phase-invariants}. + +The handshake type and handshake pending bit are not subject to TSO as we expect +a realistic implementation of handshakes would involve synchronisation. \ -abbreviation hp_step :: "handshake_type \ handshake_phase \ handshake_phase" where +abbreviation hp_step :: "hs_type \ hs_phase \ hs_phase" where "hp_step ht \ case ht of - ht_NOOP \ handshake_step - | ht_GetRoots \ handshake_step + ht_NOOP \ hs_step + | ht_GetRoots \ hs_step | ht_GetWork \ id" context sys begin -definition handshake :: "('field, 'mut, 'ref) gc_com" where - "handshake \ +definition + handshake :: "('field, 'mut, 'payload, '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) \, + (\req s. { (s\ hs_type := ht, + ghost_hs_in_sync := \False\, + ghost_hs_phase := hp_step ht (ghost_hs_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 + |ht. req = (gc, ro_hs_gc_store_type ht) }) + \ \''sys_hs_gc_mut_reqs''\ Response + (\req s. { (s\ hs_pending := (hs_pending s)(m := True) \, mv_Void) + |m. req = (gc, ro_hs_gc_store_pending m) }) + \ \''sys_hs_gc_done''\ Response + (\req s. { (s, mv_Bool (\hs_pending s m)) + |m. req = (gc, ro_hs_gc_load_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), + \ \''sys_hs_mut_pending''\ Response + (\req s. { (s, mv_Bool (hs_pending s m)) + |m. req = (mutator m, ro_hs_mut_load_pending) }) + \ \''sys_hs_mut''\ Response + (\req s. { (s, mv_hs_type (hs_type s)) + |m. req = (mutator m, ro_hs_mut_load_type) }) + \ \''sys_hs_mut_done''\ Response + (\req s. { (s\ hs_pending := (hs_pending s)(m := False), W := W s \ W', - ghost_handshake_in_sync := (ghost_handshake_in_sync s)(m := True) \, + ghost_hs_in_sync := (ghost_hs_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 +The mutators' 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"}. +Again we could make these subject to TSO, but that would be over specification. + \ context mut_m begin -abbreviation mark_object :: "location \ ('field, 'mut, 'ref) gc_com" ("\_\ mark'_object") where +abbreviation mark_object_syn :: "location \ ('field, 'mut, 'payload, 'ref) gc_com" ("\_\ mark'_object" [0] 71) where "\l\ mark_object \ mark_object_fn (mutator m) l" -abbreviation mfence :: "location \ ('field, 'mut, 'ref) gc_com" ("\_\ MFENCE") where +abbreviation mfence_syn :: "location \ ('field, 'mut, 'payload, 'ref) gc_com" ("\_\ MFENCE" [0] 71) 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_load_pending_syn :: "location \ ('field, 'mut, 'payload, 'ref) gc_com" ("\_\ hs'_load'_pending'_" [0] 71) where + "\l\ hs_load_pending_ \ \l\ Request (\s. (mutator m, ro_hs_mut_load_pending)) (\mv s. { s\ mutator_hs_pending := b \ |b. mv = mv_Bool b })" -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_load_type_syn :: "location \ ('field, 'mut, 'payload, 'ref) gc_com" ("\_\ hs'_load'_type" [0] 71) where + "\l\ hs_load_type \ \l\ Request (\s. (mutator m, ro_hs_mut_load_type)) (\mv s. { s\ hs_type := ht \ |ht. mv = mv_hs_type ht})" -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_noop_done_syn :: "location \ ('field, 'mut, 'payload, 'ref) gc_com" ("\_\ hs'_noop'_done'_") where + "\l\ hs_noop_done_ \ \l\ Request (\s. (mutator m, ro_hs_mut_done {})) + (\_ s. {s\ ghost_hs_phase := hs_step (ghost_hs_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 +abbreviation hs_get_roots_done_syn :: "location \ (('field, 'mut, 'payload, 'ref) local_state \ 'ref set) \ ('field, 'mut, 'payload, '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_hs_phase := hs_step (ghost_hs_phase s) \})" + +abbreviation hs_get_work_done_syn :: "location \ (('field, 'mut, 'payload, 'ref) local_state \ 'ref set) \ ('field, 'mut, 'payload, '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" +definition + handshake :: "('field, 'mut, 'payload, 'ref) gc_com" +where + "handshake = + \''hs_load_pending''\ hs_load_pending_ ;; + \''hs_pending''\ IF mutator_hs_pending + THEN + \''hs_mfence''\ MFENCE ;; + \''hs_load_ht''\ hs_load_type ;; + \''hs_noop''\ IF hs_type \<^bold>= \ht_NOOP\ + THEN + \''hs_noop_done''\ hs_noop_done_ + ELSE \''hs_get_roots''\ IF hs_type \<^bold>= \ht_GetRoots\ + THEN + \''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 + ELSE \''hs_get_work''\ IF hs_type \<^bold>= \ht_GetWork\ + THEN + \''hs_get_work_done''\ hs_get_work_done W + FI FI FI + FI" 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_type :: "location \ hs_type \ ('field, 'mut, 'payload, 'ref) gc_com" ("\_\ set'_hs'_type") where + "\l\ set_hs_type ht \ \l\ Request (\s. (gc, ro_hs_gc_store_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})" +abbreviation set_hs_pending :: "location \ (('field, 'mut, 'payload, 'ref) local_state \ 'mut) \ ('field, 'mut, 'payload, 'ref) gc_com" ("\_\ set'_hs'_pending") where + "\l\ set_hs_pending m \ \l\ Request (\s. (gc, ro_hs_gc_store_pending (m s))) (\_ s. {s})" + +abbreviation load_W :: "location \ ('field, 'mut, 'payload, '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, 'payload, 'ref) gc_com" ("\_\ MFENCE") where + "\l\ MFENCE \ \l\ Request (\s. (gc, ro_MFENCE)) (\_ s. {s})" definition - handshake_init :: "location \ handshake_type \ ('field, 'mut, 'ref) gc_com" ("\_\ handshake'_init") + handshake_init :: "location \ hs_type \ ('field, 'mut, 'payload, 'ref) gc_com" ("\_\ handshake'_init") where - "\l\ handshake_init req \ + "\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") + handshake_done :: "location \ ('field, 'mut, 'payload, 'ref) gc_com" ("\_\ handshake'_done") where - "\l\ handshake_done \ + "\l\ handshake_done = \l @ ''_done_muts''\ \muts := UNIV ;; - \l @ ''_done_loop''\ WHILE \<^bold>\(EMPTY muts) DO + \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))) + (\s. (gc, ro_hs_gc_load_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") + handshake_noop :: "location \ ('field, 'mut, 'payload, 'ref) gc_com" ("\_\ handshake'_noop") where - "\l\ handshake_noop \ + "\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") + handshake_get_roots :: "location \ ('field, 'mut, 'payload, 'ref) gc_com" ("\_\ handshake'_get'_roots") where - "\l\ handshake_get_roots \ + "\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") + handshake_get_work :: "location \ ('field, 'mut, 'payload, 'ref) gc_com" ("\_\ handshake'_get'_work") where - "\l\ handshake_get_work \ + "\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. +however, autonomously commit a store 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" + not_blocked :: "('field, 'mut, 'payload, 'ref) local_state \ 'mut process_name \ bool" where - "not_blocked s p \ case mem_lock s of None \ True | Some p' \ p = p'" + "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. +pending stores. \ -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 +definition + do_store_action :: "('field, 'payload, 'ref) mem_store_action \ ('field, 'mut, 'payload, 'ref) local_state \ ('field, 'mut, 'payload, 'ref) local_state" +where + "do_store_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_Mutate_Payload r f pl \ s\heap := (heap s)(r := map_option (\obj. obj\obj_payload := (obj_payload obj)(f := pl) \) (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\" + | 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" + fold_stores :: "('field, 'payload, 'ref) mem_store_action list \ ('field, 'mut, 'payload, 'ref) local_state \ ('field, 'mut, 'payload, 'ref) local_state" where - "fold_writes ws \ fold (\w. (\) (do_write_action w)) ws id" + "fold_stores ws = fold (\w. (\) (do_store_action w)) ws id" abbreviation - processors_view_of_memory :: "'mut process_name \ ('field, 'mut, 'ref) local_state \ ('field, 'mut, 'ref) local_state" + processors_view_of_memory :: "'mut process_name \ ('field, 'mut, 'payload, 'ref) local_state \ ('field, 'mut, 'payload, 'ref) local_state" where - "processors_view_of_memory p s \ fold_writes (mem_write_buffers s p) s" + "processors_view_of_memory p s \ fold_stores (mem_store_buffers s p) s" definition - do_read_action :: "('field, 'ref) mem_read_action - \ ('field, 'mut, 'ref) local_state - \ ('field, 'ref) response" + do_load_action :: "('field, 'ref) mem_load_action + \ ('field, 'mut, 'payload, 'ref) local_state + \ ('field, 'payload, '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)) + "do_load_action ract = + (\s. case ract of + mr_Ref r f \ mv_Ref (Option.bind (heap s r) (\obj. obj_fields obj f)) + | mr_Payload r f \ mv_Payload (Option.bind (heap s r) (\obj. obj_payload 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))" + | 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" + sys_load :: "'mut process_name + \ ('field, 'ref) mem_load_action + \ ('field, 'mut, 'payload, 'ref) local_state + \ ('field, 'payload, 'ref) response" where - "sys_read p ract \ do_read_action ract \ processors_view_of_memory p" + "sys_load p ract = do_load_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 +process with a non-empty store 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, +processes. The system can autonomously take the oldest store in the +store 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" + mem_TSO :: "('field, 'mut, 'payload, '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) + "mem_TSO = + \''tso_load''\ Response (\req s. { (s, sys_load p mr s) + |p mr. req = (p, ro_Load mr) \ not_blocked s p }) + \ \''tso_store''\ Response (\req s. { (s\ mem_store_buffers := (mem_store_buffers s)(p := mem_store_buffers s p @ [w]) \, mv_Void) + |p w. req = (p, ro_Store w) }) + \ \''tso_mfence''\ Response (\req s. { (s, mv_Void) + |p. req = (p, ro_MFENCE) \ mem_store_buffers s p = [] }) + \ \''tso_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 })" + \ \''tso_unlock''\ Response (\req s. { (s\ mem_lock := None \, mv_Void) + |p. req = (p, ro_Unlock) \ mem_lock s = Some p \ mem_store_buffers s p = [] }) + \ \''tso_dequeue_store_buffer''\ LocalOp (\s. { (do_store_action w s)\ mem_store_buffers := (mem_store_buffers s)(p := ws) \ + | p w ws. mem_store_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. +deallocates references. 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. +\texttt{Alloc}, avoiding deadlock. We instead signal the exhaustion of +the heap explicitly, i.e., the @{const "ro_Alloc"} action cannot fail. \ definition - alloc :: "('field, 'mut, 'ref) gc_com" + alloc :: "('field, 'mut, 'payload, '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 })" + "alloc = \''alloc''\ Response (\req s. + if dom (heap s) = UNIV + then {(s, mv_Ref None) |_::unit. snd req = ro_Alloc } + else { ( s\ heap := (heap s)(r := Some \ obj_mark = fA s, obj_fields = Map.empty, obj_payload = Map.empty \) \, 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" + free :: "('field, 'mut, 'payload, 'ref) gc_com" where - "free \ \''sys_free''\ Response (\req s. + "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" + com :: "('field, 'mut, 'payload, 'ref) gc_com" where - "com \ + "com = LOOP DO mem_TSO - \ alloc - \ free - \ handshake + \ 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 \ +abbreviation alloc :: "('field, 'mut, 'payload, 'ref) gc_com" where + "alloc \ \''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" + (\mv s. { s\ roots := roots s \ set_option opt_r \ |opt_r. mv = mv_Ref opt_r })" text\ The mutator can always discard any references it holds. \ -abbreviation discard :: "('field, 'mut, 'ref) gc_com" where +abbreviation discard :: "('field, 'mut, 'payload, '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 +abbreviation load :: "('field, 'mut, 'payload, '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 })" + \''mut_load_choose''\ LocalOp (\s. { s\ tmp_ref := r, field := f \ |r f. r \ roots s }) ;; + \''mut_load''\ Request (\s. (mutator m, LoadRef (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") + \ (('field, 'mut, 'payload, 'ref) local_state \ 'ref) + \ (('field, 'mut, 'payload, 'ref) local_state \ 'field) + \ (('ref option \ 'ref option) \ ('field, 'mut, 'payload, 'ref) local_state \ ('field, 'mut, 'payload, 'ref) local_state) \ ('field, 'mut, 'payload, 'ref) gc_com" ("\_\ deref") where - "\l\ deref r f upd \ \l\ Request (\s. (mutator m, ReadRef (r s) (f s))) + "\l\ deref r f upd \ \l\ Request (\s. (mutator m, LoadRef (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) + "_mut_fassign" :: "location \ idt \ 'ref \ 'field \ ('field, 'mut, 'payload, '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") + store_ref :: "location + \ (('field, 'mut, 'payload, 'ref) local_state \ 'ref) + \ (('field, 'mut, 'payload, 'ref) local_state \ 'field) + \ (('field, 'mut, 'payload, 'ref) local_state \ 'ref option) + \ ('field, 'mut, 'payload, 'ref) gc_com" ("\_\ store'_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 := {}\})" + "\l\ store_ref r f r' \ \l\ Request (\s. (mutator m, StoreRef (r s) (f s) (r' s))) (\_ s. {s\ghost_honorary_root := {}\})" definition - store :: "('field, 'mut, 'ref) gc_com" + store :: "('field, 'mut, 'payload, 'ref) gc_com" where - "store \ - \ \Choose vars for: \ref\field := new_ref\\ + "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" + \''store_ins''\ store_ref tmp_ref field new_ref" + +text\ + +Load and store payload data. + +\ + +abbreviation load_payload :: "('field, 'mut, 'payload, 'ref) gc_com" where + "load_payload \ + \''mut_load_payload_choose''\ LocalOp (\s. { s\ tmp_ref := r, field := f \ |r f. r \ roots s }) ;; + \''mut_load_payload''\ Request (\s. (mutator m, LoadPayload (tmp_ref s) (field s))) + (\mv s. { s\ mutator_data := (mutator_data s)(var := pl) \ + |var pl. mv = mv_Payload pl })" + +abbreviation store_payload :: "('field, 'mut, 'payload, 'ref) gc_com" where + "store_payload \ + \''mut_store_payload_choose''\ LocalOp (\s. { s\ tmp_ref := r, field := f, payload_value := pl s \ |r f pl. r \ roots s }) ;; + \''mut_store_payload''\ Request (\s. (mutator m, StorePayload (tmp_ref s) (field s) (payload_value s))) + (\mv s. { s\ mutator_data := (mutator_data s)(f := pl) \ + |f pl. mv = mv_Payload pl })" 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. \ +(* +FIXME add SKIP before alloc, mfence. handshake needs work too: want to +commit to checking for handshakes in a strongly fair way. A SKIP +at the top level of \handshake\ + a noop transition + appropriate fairness constraints might work. + +FIXME is mut local computation strong enough? only works on mutator data; not roots. +*) definition - com :: "('field, 'mut, 'ref) gc_com" + com :: "('field, 'mut, 'payload, 'ref) gc_com" where - "com \ + "com = LOOP DO - \''mut local computation''\ SKIP - \ alloc - \ discard - \ load - \ store - \ \''mut mfence''\ MFENCE - \ handshake + \''mut_local_computation''\ LocalOp (\s. {s\mutator_data := f (mutator_data s)\ |f. True}) + \ alloc + \ discard + \ load + \ store + \ load_payload + \ store_payload + \ \''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" + \ (('field, 'mut, 'payload, 'ref) local_state \ 'ref) + \ (('field, 'mut, 'payload, 'ref) local_state \ 'field) + \ (('ref option \ 'ref option) \ ('field, 'mut, 'payload, 'ref) local_state \ ('field, 'mut, 'payload, 'ref) local_state) \ ('field, 'mut, 'payload, 'ref) gc_com" where - "gc_deref l r f upd \ \l\ Request (\s. (gc, ReadRef (r s) (f s))) + "gc_deref l r f upd \ \l\ Request (\s. (gc, LoadRef (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" + gc_load_mark :: "location + \ (('field, 'mut, 'payload, 'ref) local_state \ 'ref) + \ ((gc_mark option \ gc_mark option) \ ('field, 'mut, 'payload, 'ref) local_state \ ('field, 'mut, 'payload, 'ref) local_state) + \ ('field, 'mut, 'payload, '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 })" + "gc_load_mark l r upd \ \l\ Request (\s. (gc, LoadMark (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) + "_gc_fassign" :: "location \ idt \ 'ref \ 'field \ ('field, 'mut, 'payload, 'ref) gc_com" ("\_\ \_ := \_ \ _" [0, 0, 0, 70] 71) + "_gc_massign" :: "location \ idt \ 'ref \ ('field, 'mut, 'payload, 'ref) gc_com" ("\_\ \_ := \_ \ flag" [0, 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)" + "\l\ \m := \r\flag" => "CONST gc_load_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 store_fA_syn :: "location \ (('field, 'mut, 'payload, 'ref) local_state \ gc_mark) \ ('field, 'mut, 'payload, 'ref) gc_com" ("\_\ store'_fA") where + "\l\ store_fA f \ \l\ Request (\s. (gc, StorefA (f s))) (\_ s. {s})" -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 load_fM_syn :: "location \ ('field, 'mut, 'payload, 'ref) gc_com" ("\_\ load'_fM") where + "\l\ load_fM \ \l\ Request (\s. (gc, LoadfM)) (\mv s. { s\fM := m\ |m. mv = mv_Mark (Some m) })" -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 store_fM_syn :: "location \ ('field, 'mut, 'payload, 'ref) gc_com" ("\_\ store'_fM") where + "\l\ store_fM \ \l\ Request (\s. (gc, StorefM (fM s))) (\_ s. {s})" -abbreviation mark_object :: "location \ ('field, 'mut, 'ref) gc_com" ("\_\ mark'_object") where +abbreviation store_phase_syn :: "location \ gc_phase \ ('field, 'mut, 'payload, 'ref) gc_com" ("\_\ store'_phase") where + "\l\ store_phase ph \ \l\ Request (\s. (gc, StorePhase ph)) (\_ s. {s\ phase := ph \})" + +abbreviation mark_object_syn :: "location \ ('field, 'mut, 'payload, '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 +abbreviation free_syn :: "location \ (('field, 'mut, 'payload, 'ref) local_state \ 'ref) \ ('field, 'mut, 'payload, '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" + com :: "('field, 'mut, 'payload, 'ref) gc_com" where - "com \ + "com = LOOP DO \''idle_noop''\ handshake_noop ;; \ \\hp_Idle\\ - \''idle_read_fM''\ read_fM ;; + \''idle_load_fM''\ load_fM ;; \''idle_invert_fM''\ \fM := (\ \fM) ;; - \''idle_write_fM''\ write_fM ;; + \''idle_store_fM''\ store_fM ;; \''idle_flip_noop''\ handshake_noop ;; \ \\hp_IdleInit\\ - \''idle_phase_init''\ write_phase ph_Init ;; + \''idle_phase_init''\ store_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 ;; + \''init_phase_mark''\ store_phase ph_Mark ;; + \''mark_load_fM''\ load_fM ;; + \''mark_store_fA''\ store_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''\ 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_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 ;; + \''mark_end''\ store_phase ph_Sweep ;; + \''sweep_load_fM''\ load_fM ;; \''sweep_refs''\ \refs := UNIV ;; - \''sweep_loop''\ WHILE \<^bold>\(EMPTY refs) DO + \''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_load_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 + \''sweep_idle''\ store_phase ph_Idle OD" end primrec - gc_pgms :: "'mut process_name \ ('field, 'mut, 'ref) gc_com" + gc_coms :: "'mut process_name \ ('field, 'mut, 'payload, 'ref) gc_com" where - "gc_pgms (mutator m) = mut_m.com m" -| "gc_pgms gc = gc.com" -| "gc_pgms sys = sys.com" + "gc_coms (mutator m) = mut_m.com m" +| "gc_coms gc = gc.com" +| "gc_coms sys = sys.com" (*<*) end (*>*) diff --git a/thys/ConcurrentGC/Noninterference.thy b/thys/ConcurrentGC/Noninterference.thy new file mode 100644 --- /dev/null +++ b/thys/ConcurrentGC/Noninterference.thy @@ -0,0 +1,572 @@ +(*<*) +(* + * 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 Noninterference +imports + Global_Invariants_Lemmas + Local_Invariants_Lemmas + Tactics +begin + +(*>*) +section\ Noninterference \ + +lemma mut_del_barrier1_subseteq_mut_mo_valid_ref_locs[locset_cache]: (* FIXME rename *) + "mut_m.del_barrier1_locs \ mut_m.mo_valid_ref_locs" +unfolding mut_m.del_barrier1_locs_def mut_m.mo_valid_ref_locs_def by (auto intro: append_prefixD) + +lemma mut_del_barrier2_subseteq_mut_mo_valid_ref[locset_cache]: (* FIXME rename *) + "mut_m.ins_barrier_locs \ mut_m.mo_valid_ref_locs" +unfolding mut_m.ins_barrier_locs_def mut_m.mo_valid_ref_locs_def by (auto intro: append_prefixD) + +context gc +begin + +lemma obj_fields_marked_locs_subseteq_hp_IdleMarkSweep_locs: + "obj_fields_marked_locs \ hp_IdleMarkSweep_locs" +unfolding gc.obj_fields_marked_locs_def gc.hp_IdleMarkSweep_locs_def gc.mark_loop_locs_def gc.mark_loop_mo_locs_def +apply (clarsimp simp: locset_cache loc_defs) +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" +unfolding obj_fields_marked_locs_def hs_in_sync_locs_def hs_done_locs_def mark_loop_mo_locs_def +by (auto simp: loc_defs dest: prefix_same_cases) + +lemma obj_fields_marked_good_ref_subseteq_hp_IdleMarkSweep_locs: + "obj_fields_marked_good_ref_locs \ hp_IdleMarkSweep_locs" +unfolding obj_fields_marked_good_ref_locs_def mark_loop_locs_def hp_IdleMarkSweep_locs_def mark_loop_mo_locs_def +apply (clarsimp simp: loc_defs) +apply (drule mp) +apply (auto intro: append_prefixD) +done + +lemma mark_loop_mo_mark_loop_field_done_subseteq_hs_in_sync_locs: + "obj_fields_marked_good_ref_locs \ hs_in_sync_locs" +unfolding obj_fields_marked_good_ref_locs_def hs_in_sync_locs_def mark_loop_mo_locs_def hs_done_locs_def +by (auto simp: loc_defs dest: prefix_same_cases) + +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 loc_defs + dest: prefix_same_cases) + +lemma get_roots_UN_get_work_locs_subseteq_gc_W_empty_locs: + "get_roots_UN_get_work_locs \ gc_W_empty_locs" +unfolding get_roots_UN_get_work_locs_def +by (auto simp: hs_get_roots_locs_def hs_get_work_locs_def gc_W_empty_locs_def) + +end + +declare + gc.obj_fields_marked_locs_subseteq_hp_IdleMarkSweep_locs[locset_cache] + gc.obj_fields_marked_locs_subseteq_hs_in_sync_locs[locset_cache] + gc.obj_fields_marked_good_ref_subseteq_hp_IdleMarkSweep_locs[locset_cache] + gc.mark_loop_mo_mark_loop_field_done_subseteq_hs_in_sync_locs[locset_cache] + gc.no_grey_refs_locs_subseteq_hs_in_sync_locs[locset_cache] + gc.get_roots_UN_get_work_locs_subseteq_gc_W_empty_locs[locset_cache] + +lemma handshake_obj_fields_markedD: + "\ atS gc gc.obj_fields_marked_locs s; gc.handshake_invL s \ \ sys_ghost_hs_phase s\ = hp_IdleMarkSweep \ All (ghost_hs_in_sync (s\ sys))" +unfolding gc.handshake_invL_def +by (metis (no_types, lifting) atS_mono gc.obj_fields_marked_locs_subseteq_hp_IdleMarkSweep_locs gc.obj_fields_marked_locs_subseteq_hs_in_sync_locs) + +lemma obj_fields_marked_good_ref_locs_hp_phaseD: + "\ atS gc gc.obj_fields_marked_good_ref_locs s; gc.handshake_invL s \ + \ sys_ghost_hs_phase s\ = hp_IdleMarkSweep \ All (ghost_hs_in_sync (s\ sys))" +unfolding gc.handshake_invL_def +by (metis (no_types, lifting) atS_mono gc.mark_loop_mo_mark_loop_field_done_subseteq_hs_in_sync_locs gc.obj_fields_marked_good_ref_subseteq_hp_IdleMarkSweep_locs) + +lemma gc_marking_reaches_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 := map_option (\obj. obj\obj_fields := (obj_fields obj)(f := opt_r')\) (sys_heap s r)), + mem_store_buffers := (mem_store_buffers (s sys))(p := ws)\))" + assumes sb: "sys_mem_store_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_store_refs m s \ (z reaches y) s \ valid_ref y s" + proof induct + case (refl x) then show ?case by auto + next + case (step x y z) with sb vri show ?case + apply (clarsimp simp: points_to_Mutate) + apply (elim disjE) + apply (metis (no_types, lifting) obj_at_cong reaches_def rtranclp.rtrancl_into_rtrancl) + apply (metis (no_types, lifting) obj_at_def option.case(2) reaches_def rtranclp.rtrancl_into_rtrancl valid_refs_invD(4)) + apply clarsimp + apply (elim disjE) + apply (rule exI[where x=z]) + apply (clarsimp simp: mut_m.tso_store_refs_def) + apply (rule valid_refs_invD(3)[where m=m and x=z], auto simp: mut_m.tso_store_refs_def; fail)[1] + apply (metis (no_types, lifting) obj_at_cong reaches_def rtranclp.rtrancl_into_rtrancl) + apply clarsimp + apply (elim disjE) + apply (rule exI[where x=z]) + apply (clarsimp simp: mut_m.tso_store_refs_def) + apply (rule valid_refs_invD(3)[where m=m and x=z], auto simp: mut_m.tso_store_refs_def)[1] + apply (metis (no_types, lifting) obj_at_def option.case(2) reaches_def rtranclp.rtrancl_into_rtrancl valid_refs_invD(4)) + done + qed + then show ?thesis by blast +qed + +lemma (in sys) gc_obj_fields_marked_invL[intro]: + notes filter_empty_conv[simp] + notes fun_upd_apply[simp] + 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_store_inv \<^bold>\ valid_refs_inv \<^bold>\ valid_W_inv) \ + sys + \ gc.obj_fields_marked_invL \" +proof(vcg_jackhammer (keep_locs) (no_thin_post_inv), vcg_name_cases) + case (tso_dequeue_store_buffer s s' p w ws) show ?case + proof(cases w) + case (mw_Mark ref mark) with tso_dequeue_store_buffer show ?thesis +apply - +apply (clarsimp simp: p_not_sys gc.obj_fields_marked_invL_def) +apply (intro conjI impI; clarsimp) + +apply (frule (1) handshake_obj_fields_markedD) +apply (clarsimp simp: gc.obj_fields_marked_def) +apply (frule (1) valid_W_invD) +apply (drule_tac x=x in spec) +apply clarsimp +apply (erule obj_at_field_on_heapE) + apply (force split: obj_at_splits) +apply (force split: obj_at_splits) + +apply (erule obj_at_field_on_heapE) + apply (clarsimp split: obj_at_splits; fail) +apply (clarsimp split: obj_at_splits) + apply (metis valid_W_invD(1)) +apply (metis valid_W_invD(1)) + +apply (force simp: valid_W_invD(1) split: obj_at_splits) +done + next case (mw_Mutate r f opt_r') with tso_dequeue_store_buffer show ?thesis +apply - +apply (clarsimp simp: p_not_sys gc.obj_fields_marked_invL_def) +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) +apply (drule_tac x=m in spec) +apply (intro conjI impI; clarsimp simp: obj_at_field_on_heap_imp_valid_ref gc_marking_reaches_Mutate split: option.splits) + +subgoal for m +apply (frule (1) handshake_obj_fields_markedD) +apply (elim disjE; auto simp: gc.obj_fields_marked_def split: option.splits) +done + +subgoal for m r' +apply (frule (1) obj_fields_marked_good_ref_locs_hp_phaseD) +apply (elim disjE; clarsimp simp: marked_insertionD) +done +done + next case (mw_Mutate_Payload r f pl) with tso_dequeue_store_buffer show ?thesis by - (erule gc_obj_fields_marked_invL_niE; clarsimp) + next case (mw_fA mark) with tso_dequeue_store_buffer show ?thesis by - (erule gc_obj_fields_marked_invL_niE; clarsimp) + next case (mw_fM mark) with tso_dequeue_store_buffer show ?thesis + apply - + apply (clarsimp simp: p_not_sys fM_rel_inv_def fM_rel_def gc.obj_fields_marked_invL_def) + apply (erule disjE; clarsimp) + apply (intro conjI impI; clarsimp) + apply (metis (no_types, lifting) handshake_obj_fields_markedD hs_phase.distinct(7)) + apply (metis (no_types, lifting) hs_phase.distinct(7) obj_fields_marked_good_ref_locs_hp_phaseD) + apply (metis (no_types, lifting) UnCI elem_set hs_phase.distinct(7) gc.obj_fields_marked_good_ref_locs_def obj_fields_marked_good_ref_locs_hp_phaseD option.simps(15) thin_locs_pre_keep_atSE) + done + next case (mw_Phase ph) with tso_dequeue_store_buffer show ?thesis + by - (erule gc_obj_fields_marked_invL_niE; clarsimp) + qed +qed + + +subsection\The infamous termination argument\ + +lemma (in mut_m) gc_W_empty_mut_inv_eq_imp: + "eq_imp (\m'. sys_W \<^bold>\ WL (mutator m') \<^bold>\ sys_ghost_hs_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_hs_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 gc_W_empty_mut_inv_load_W: + "\ \m. mut_m.gc_W_empty_mut_inv m s; \m. sys_ghost_hs_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 + +context gc +begin + +lemma gc_W_empty_mut_inv_hs_init[iff]: + "mut_m.gc_W_empty_mut_inv m (s(sys := s sys\hs_type := ht, ghost_hs_in_sync := \False\\))" + "mut_m.gc_W_empty_mut_inv m (s(sys := s sys\hs_type := ht, ghost_hs_in_sync := \False\, ghost_hs_phase := hp' \))" +by (simp_all add: mut_m.gc_W_empty_mut_inv_def) + +lemma gc_W_empty_invL[intro]: + notes fun_upd_apply[simp] + shows + "\ 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; (clarsimp elim: gc_W_empty_mut_inv_load_W simp: WL_def)?) +proof vcg_name_cases + case (mark_loop_get_work_done_loop s s') then show ?case + by (simp add: WL_def gc_W_empty_mut_inv_load_W valid_W_inv_sys_ghg_empty_iff) +next case (mark_loop_get_roots_done_loop s s') then show ?case + by (simp add: WL_def gc_W_empty_mut_inv_load_W valid_W_inv_sys_ghg_empty_iff) +qed + +end + +lemma (in sys) gc_gc_W_empty_invL[intro]: + notes fun_upd_apply[simp] + shows + "\ gc.gc_W_empty_invL \ sys" +by vcg_chainsaw + +lemma empty_WL_GC: + "\ atS gc gc.get_roots_UN_get_work_locs s; gc.obj_fields_marked_invL s \ \ gc_ghost_honorary_grey s\ = {}" +unfolding gc.obj_fields_marked_invL_def +using atS_mono[OF _ gc.get_roots_UN_get_work_locs_subseteq_ghost_honorary_grey_empty_locs] +apply metis +done + +lemma gc_hs_get_roots_get_workD: + "\ atS gc gc.get_roots_UN_get_work_locs s; gc.handshake_invL s \ + \ sys_ghost_hs_phase s\ = hp_IdleMarkSweep \ sys_hs_type s\ \ {ht_GetWork, ht_GetRoots}" +unfolding gc.handshake_invL_def +apply clarsimp +apply (metis (no_types, lifting) atS_mono atS_un gc.get_roots_UN_get_work_locs_def gc.hs_get_roots_locs_subseteq_hp_IdleMarkSweep_locs gc.hs_get_work_locs_subseteq_hp_IdleMarkSweep_locs) +done + + +context gc +begin + +lemma handshake_sweep_mark_endD: + "\ atS gc no_grey_refs_locs s; handshake_invL s; handshake_phase_inv s\ \ + \ mut_m.mut_ghost_hs_phase m s\ = hp_IdleMarkSweep \ All (ghost_hs_in_sync (s\ sys))" +apply (simp add: gc.handshake_invL_def) +apply (elim conjE) +apply (drule mp, erule atS_mono[OF _ gc.no_grey_refs_locs_subseteq_hs_in_sync_locs]) +apply (drule mut_m.handshake_phase_invD) +apply (simp only: gc.no_grey_refs_locs_def cong del: atS_state_weak_cong) +apply (clarsimp simp: atS_un) +apply (elim disjE) + apply (drule mp, erule atS_mono[where ls'="gc.hp_IdleMarkSweep_locs"]) + apply (clarsimp simp: gc.black_heap_locs_def locset_cache) + apply (clarsimp simp: hp_step_rel_def) + apply blast + apply (drule mp, erule atS_mono[where ls'="gc.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 locset_cache hp_step_rel_def) +apply blast +done + +lemma gc_W_empty_mut_mo_co_mark: + "\ \x. mut_m.gc_W_empty_mut_inv x s\; mutators_phase_inv s\; + mut_m.mut_ghost_honorary_grey m s\ = {}; + r \ mut_m.mut_roots m s\ \ mut_m.mut_ghost_honorary_root m s\; white r s\; + atS gc get_roots_UN_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_hs_get_roots_get_workD) +apply (frule_tac m=m in mut_m.handshake_phase_invD) +apply (clarsimp simp: hp_step_rel_def simp del: Un_iff) +apply (elim disjE, simp_all) +proof(goal_cases before_get_work past_get_work before_get_roots after_get_roots) + case before_get_work then show ?thesis + apply (clarsimp simp: mut_m.gc_W_empty_mut_inv_def) + apply blast + done +next case past_get_work then show ?thesis + apply (clarsimp simp: mut_m.gc_W_empty_mut_inv_def) + apply (frule spec[where x=m], clarsimp) + apply (frule (2) mut_m.reachable_snapshot_inv_white_root) + apply clarsimp + apply (drule grey_protects_whiteD) + apply (clarsimp simp: grey_def) + apply (rename_tac g p) + apply (case_tac p; clarsimp) + (* mutator *) + apply blast + (* Can't be the GC *) + apply (frule (1) empty_WL_GC) + apply (drule mp, erule atS_mono[OF _ get_roots_UN_get_work_locs_subseteq_gc_W_empty_locs]) + apply (clarsimp simp: WL_def; fail) + (* Can't be sys *) + apply (clarsimp simp: WL_def valid_W_inv_sys_ghg_empty_iff; fail) + done +next case before_get_roots then show ?case + apply (clarsimp simp: mut_m.gc_W_empty_mut_inv_def) + apply blast + done +next case after_get_roots then show ?case + apply (clarsimp simp: mut_m.gc_W_empty_mut_inv_def) + apply (frule spec[where x=m], clarsimp) + apply (frule (2) mut_m.reachable_snapshot_inv_white_root) + apply clarsimp + apply (drule grey_protects_whiteD) + apply (clarsimp simp: grey_def) + apply (rename_tac g p) + apply (case_tac p; clarsimp) + (* mutator *) + apply blast + (* Can't be the GC *) + apply (frule (1) empty_WL_GC) + apply (drule mp, erule atS_mono[OF _ get_roots_UN_get_work_locs_subseteq_gc_W_empty_locs]) + apply (clarsimp simp: WL_def; fail) + (* Can't be sys *) + apply (clarsimp simp: WL_def valid_W_inv_sys_ghg_empty_iff; fail) + done +qed + +lemma no_grey_refs_mo_co_mark: + "\ 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_m.mut_roots m s\ \ mut_m.mut_ghost_honorary_root m 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 locset_cache) + apply (frule mut_m.handshake_phase_invD) + apply (clarsimp simp: hp_step_rel_def) + apply (drule spec[where x=m]) + apply (clarsimp simp: conj_disj_distribR[symmetric]) + apply (simp add: handshake_in_syncD mut_m.no_grey_refs_not_rootD; fail) + apply (clarsimp simp: atS_simps gc.handshake_invL_def locset_cache) + apply (frule mut_m.handshake_phase_invD) + apply (clarsimp simp: hp_step_rel_def) + apply (drule spec[where x=m]) + apply (simp add: handshake_in_syncD mut_m.no_grey_refs_not_rootD; fail) + apply (clarsimp simp: atS_simps gc.handshake_invL_def locset_cache) + apply (frule mut_m.handshake_phase_invD) + apply (clarsimp simp: hp_step_rel_def) + apply (drule spec[where x=m]) + apply (simp add: handshake_in_syncD mut_m.no_grey_refs_not_rootD; fail) +apply (frule (2) handshake_sweep_mark_endD) +apply (drule spec[where x=m]) +apply clarsimp +apply (simp add: handshake_in_syncD mut_m.no_grey_refs_not_rootD; fail) +done + +end + +context mut_m +begin + +lemma gc_W_empty_invL[intro]: + notes gc.gc_W_empty_mut_mo_co_mark[simp] + notes gc.no_grey_refs_mo_co_mark[simp] + notes fun_upd_apply[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>\ 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 \" +proof(vcg_chainsaw gc.gc_W_empty_invL_def, vcg_name_cases) + case (hs_noop_done s s' x) then show ?case + unfolding gc.handshake_invL_def + by (metis atS_un gc.get_roots_UN_get_work_locs_def hs_type.distinct(1) hs_type.distinct(3)) +next case (hs_get_roots_done0 s s' x) then show ?case + apply (clarsimp simp: mut_m.gc_W_empty_mut_inv_def WL_def) + apply (metis (no_types, lifting)) + done +next case (hs_get_work_done0 s s' x) then show ?case + apply (clarsimp simp: mut_m.gc_W_empty_mut_inv_def WL_def) + apply (metis (no_types, lifting)) + done +qed (simp_all add: no_grey_refs_def) + +end + +context gc +begin + +lemma mut_store_old_mark_object_invL[intro]: + notes fun_upd_apply[simp] + 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_chainsaw mut_m.mark_object_invL_def mut_m.mut_store_del_mark_object_invL_def2) \ \\at gc sweep_loop_free s\\ + apply (metis (no_types, lifting) handshake_in_syncD mut_m.mutator_phase_inv_aux.simps(5) mut_m.no_grey_refs_not_rootD obj_at_cong white_def)+ +done + +lemma mut_store_ins_mark_object_invL[intro]: + "\ 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_chainsaw mut_m.mark_object_invL_def mut_m.mut_store_ins_mark_object_invL_def2) \ \\at gc sweep_loop_free s\\ + apply (metis (no_types, lifting) handshake_in_syncD mut_m.mutator_phase_inv_aux.simps(5) mut_m.no_grey_refs_not_rootD obj_at_cong white_def)+ +done + +lemma mut_mark_object_invL[intro]: + "\ 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 \" +proof(vcg_chainsaw mut_m.handshake_invL_def mut_m.mark_object_invL_def, vcg_name_cases "mutator m") \ \\at gc sweep_loop_free s\\ + case (ins_barrier_locs s s') then show ?case + apply - + apply (drule_tac x=m in spec) + apply (clarsimp simp: fun_upd_apply dest!: handshake_in_syncD split: obj_at_field_on_heap_splits) + apply (metis (no_types, lifting) mut_m.no_grey_refs_not_rootD obj_at_cong white_def) + apply (metis (no_types) marked_not_white mut_m.no_grey_refs_not_rootD whiteI) + done +next case (del_barrier1_locs s s') then show ?case + apply - + apply (drule_tac x=m in spec) + apply (clarsimp simp: fun_upd_apply dest!: handshake_in_syncD split: obj_at_field_on_heap_splits) + apply (metis (no_types, lifting) mut_m.no_grey_refs_not_rootD obj_at_cong white_def) + apply (metis (no_types, lifting) marked_not_white mut_m.no_grey_refs_not_rootD obj_at_cong white_def) + done +qed blast+ + +end + +lemma mut_m_get_roots_no_fM_write: + "\ mut_m.handshake_invL m s; handshake_phase_inv s\; fM_rel_inv s\; tso_store_inv s\ \ + \ atS (mutator m) mut_m.hs_get_roots_locs s \ p \ sys \ \sys_mem_store_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 (clarsimp simp: hp_step_rel_def fM_rel_def filter_empty_conv p_not_sys) +apply (metis (full_types) hs_phase.distinct(7) list.set_intros(1) tso_store_invD(4)) +done + +(* FIXME loads of cut-and-paste here *) +lemma (in sys) mut_mark_object_invL[intro]: + notes filter_empty_conv[simp] + notes fun_upd_apply[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_store_inv) \ + sys + \ mut_m.mark_object_invL m \" +proof(vcg_chainsaw mut_m.mark_object_invL_def, vcg_name_cases "mutator m") + case (hs_get_roots_loop_locs s s' p w ws x) then show ?case + apply - + apply (cases w; clarsimp split: obj_at_splits) + apply (meson valid_W_invD(1)) + apply (simp add: atS_mono mut_m.hs_get_roots_loop_locs_subseteq_hs_get_roots_locs mut_m_get_roots_no_fM_write) + done +next case (hs_get_roots_loop_done s s' p w ws y) then show ?case + apply - + apply (cases w; clarsimp simp: p_not_sys valid_W_invD split: obj_at_splits) + apply (rename_tac fl obj) + apply (drule_tac fl=fl and p=p and ws=ws in mut_m_get_roots_no_fM_write; clarsimp) + apply (drule mp, erule atS_simps, loc_mem) + apply blast + done +next case (hs_get_roots_done s s' p w ws x) then show ?case + apply - + apply (cases w; clarsimp simp: p_not_sys valid_W_invD split: obj_at_splits) + apply blast + apply (rename_tac fl) + apply (drule_tac fl=fl and p=p and ws=ws in mut_m_get_roots_no_fM_write; clarsimp) + apply (drule mp, erule atS_simps, loc_mem) + apply blast + done +next case (mo_ptest_locs s s' p ws ph') then show ?case by (clarsimp simp: p_not_sys; elim disjE; clarsimp simp: phase_rel_def handshake_in_syncD dest!: phase_rel_invD) +next case (store_ins s s' p w ws y) then show ?case + apply - + apply (cases w; clarsimp simp: p_not_sys valid_W_invD split: obj_at_splits) + apply (metis (no_types, lifting) hs_phase.distinct(3, 5) mut_m.mut_ghost_handshake_phase_idle mut_m_not_idle_no_fM_writeD store_ins(9)) + using valid_refs_invD(9) apply fastforce + apply (elim disjE; clarsimp simp: phase_rel_def handshake_in_syncD dest!: phase_rel_invD) + done +next case (del_barrier1_locs s s' p w ws) then show ?case + proof(cases w) + case (mw_Mutate r f opt_r') with del_barrier1_locs show ?thesis +apply (clarsimp simp: p_not_sys; elim disjE; clarsimp) +apply (intro conjI impI; clarsimp simp: obj_at_field_on_heap_imp_valid_ref split: option.splits) + apply (intro conjI impI; clarsimp) + apply (smt (z3) reachableI(1) valid_refs_invD(8)) + apply (metis (no_types, lifting) marked_insertionD mut_m.mutator_phase_inv_aux.simps(4) mut_m.mutator_phase_inv_aux.simps(5) obj_at_cong reachableI(1) valid_refs_invD(8)) +(* brutal *) +apply (rename_tac ma x2) +apply (frule_tac m=m in mut_m.handshake_phase_invD) +apply (frule_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; clarsimp simp: marked_insertionD mut_m.mut_ghost_handshake_phase_idle) +done + next case (mw_fM fl) with del_barrier1_locs mut_m_not_idle_no_fM_writeD show ?thesis by fastforce + next case (mw_Phase ph) with del_barrier1_locs show ?thesis by (clarsimp simp: p_not_sys; elim disjE; clarsimp simp: phase_rel_def handshake_in_syncD dest!: phase_rel_invD) + qed (fastforce simp: valid_W_invD split: obj_at_field_on_heap_splits obj_at_splits)+ +next case (ins_barrier_locs s s' p w ws) then show ?case + proof(cases w) + case (mw_Mutate r f opt_r') with ins_barrier_locs show ?thesis +apply (clarsimp simp: p_not_sys; elim disjE; clarsimp) +apply (intro conjI impI; clarsimp simp: obj_at_field_on_heap_imp_valid_ref split: option.splits) + apply (intro conjI impI; clarsimp) + apply (smt (z3) reachableI(1) valid_refs_invD(8)) + apply (metis (no_types, lifting) marked_insertionD mut_m.mutator_phase_inv_aux.simps(4) mut_m.mutator_phase_inv_aux.simps(5) obj_at_cong reachableI(1) valid_refs_invD(8)) +(* brutal *) +apply (rename_tac ma x2) +apply (frule_tac m=m in mut_m.handshake_phase_invD) +apply (frule_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; clarsimp simp: marked_insertionD mut_m.mut_ghost_handshake_phase_idle) +done + next case (mw_fM fl) with ins_barrier_locs mut_m_not_idle_no_fM_writeD show ?thesis by fastforce + next case (mw_Phase ph) with ins_barrier_locs show ?thesis by (clarsimp simp: p_not_sys; elim disjE; clarsimp simp: phase_rel_def handshake_in_syncD dest!: phase_rel_invD) + qed (fastforce simp: valid_W_invD split: obj_at_field_on_heap_splits obj_at_splits)+ +next case (lop_store_ins s s' p w ws y) then show ?case + apply - + apply (cases w; clarsimp simp: valid_W_invD(1) split: obj_at_splits) + apply (metis (no_types, hide_lams) hs_phase.distinct(5,7) mut_m_not_idle_no_fM_write) + apply (clarsimp simp: p_not_sys; elim disjE; clarsimp simp: phase_rel_def handshake_in_syncD dest!: phase_rel_invD; fail)+ + done +qed + + +(*<*) + +end +(*>*) diff --git a/thys/ConcurrentGC/Phases.thy b/thys/ConcurrentGC/Phases.thy new file mode 100644 --- /dev/null +++ b/thys/ConcurrentGC/Phases.thy @@ -0,0 +1,440 @@ +(*<*) +(* + * 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 Phases +imports + Global_Invariants_Lemmas + Local_Invariants_Lemmas + Tactics +begin + +(*>*) +section\Handshake phases\ + +text\ + +Reasoning about phases, handshakes. + +Tie the garbage collector's control location to the value of @{const +"gc_phase"}. + +\ + +lemma (in gc) phase_invL_eq_imp: + "eq_imp (\(_::unit) s. (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 (fastforce dest!: phase_rel_invD simp: phase_rel_def) + +lemma (in sys) gc_phase_invL[intro]: + notes fun_upd_apply[simp] + notes if_splits[split] + shows + "\ gc.phase_invL \ sys" +by (vcg_chainsaw gc.phase_invL_def) + +lemma (in mut_m) gc_phase_invL[intro]: + "\ gc.phase_invL \ mutator m" +by (vcg_chainsaw gc.phase_invL_def[inv]) + +lemma (in gc) phase_rel_inv[intro]: + "\ handshake_invL \<^bold>\ phase_invL \<^bold>\ LSTP phase_rel_inv \ gc \ LSTP phase_rel_inv \" +unfolding phase_rel_inv_def by (vcg_jackhammer (no_thin_post_inv); simp add: phase_rel_def; blast) + +lemma (in sys) phase_rel_inv[intro]: + notes gc.phase_invL_def[inv] + notes phase_rel_inv_def[inv] + notes fun_upd_apply[simp] + shows + "\ LSTP (phase_rel_inv \<^bold>\ tso_store_inv) \ sys \ LSTP phase_rel_inv \" +proof(vcg_jackhammer (no_thin_post_inv), vcg_name_cases) + case (tso_dequeue_store_buffer s s' p w ws) then show ?case + apply (simp add: phase_rel_def p_not_sys split: if_splits) + apply (elim disjE; auto split: if_splits) + done +qed + +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 \" +unfolding phase_rel_inv_def +proof (vcg_jackhammer (no_thin_post_inv), vcg_name_cases) + case (hs_noop_done s s') then show ?case + by (auto dest!: handshake_phase_invD + simp: handshake_phase_rel_def phase_rel_def hp_step_rel_def + split: hs_phase.splits) +next case (hs_get_roots_done s s') then show ?case + by (auto dest!: handshake_phase_invD + simp: handshake_phase_rel_def phase_rel_def hp_step_rel_def + split: hs_phase.splits) +next case (hs_get_work_done s s') then show ?case + by (auto dest!: handshake_phase_invD + simp: handshake_phase_rel_def phase_rel_def hp_step_rel_def + split: hs_phase.splits) +qed + +text\ + +Connect @{const "sys_ghost_hs_phase"} with locations in the GC. + +\ + +lemma gc_handshake_invL_eq_imp: + "eq_imp (\(_::unit) s. (AT s gc, s\ gc, sys_ghost_hs_phase s\, hs_pending (s\ sys), ghost_hs_in_sync (s\ sys), sys_hs_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 sys) gc_handshake_invL[intro]: + "\ gc.handshake_invL \ sys" +by (vcg_chainsaw gc.handshake_invL_def) + +lemma (in sys) handshake_phase_inv[intro]: + "\ LSTP handshake_phase_inv \ sys" +unfolding handshake_phase_inv_def by (vcg_jackhammer (no_thin_post_inv)) + +lemma (in gc) handshake_invL[intro]: + notes fun_upd_apply[simp] + shows + "\ handshake_invL \ gc" +by vcg_jackhammer fastforce+ + +lemma (in gc) handshake_phase_inv[intro]: + notes fun_upd_apply[simp] + shows + "\ handshake_invL \<^bold>\ LSTP handshake_phase_inv \ gc \ LSTP handshake_phase_inv \" +unfolding handshake_phase_inv_def by (vcg_jackhammer (no_thin_post_inv)) (auto simp: handshake_phase_inv_def hp_step_rel_def) + +text\ + +Local handshake phase invariant for the mutators. + +\ + +lemma (in mut_m) handshake_invL_eq_imp: + "eq_imp (\(_::unit) s. (AT s (mutator m), s\ (mutator m), sys_hs_type s\, sys_hs_pending m s\, mem_store_buffers (s\ sys) (mutator m))) + handshake_invL" +unfolding eq_imp_def handshake_invL_def by simp + +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_chainsaw + +lemma (in gc) mut_handshake_invL[intro]: + notes fun_upd_apply[simp] + shows + "\ handshake_invL \<^bold>\ mut_m.handshake_invL m \ gc \ mut_m.handshake_invL m \" +by (vcg_chainsaw mut_m.handshake_invL_def) + +lemma (in sys) mut_handshake_invL[intro]: + notes if_splits[split] + notes fun_upd_apply[simp] + shows + "\ mut_m.handshake_invL m \ sys" +by (vcg_chainsaw mut_m.handshake_invL_def) + +lemma (in mut_m) gc_handshake_invL[intro]: + notes fun_upd_apply[simp] + shows + "\ handshake_invL \<^bold>\ gc.handshake_invL \ mutator m \ gc.handshake_invL \" +by (vcg_chainsaw gc.handshake_invL_def) + +lemma (in mut_m) handshake_phase_inv[intro]: + notes fun_upd_apply[simp] + shows + "\ handshake_invL \<^bold>\ LSTP handshake_phase_inv \ mutator m \ LSTP handshake_phase_inv \" +unfolding handshake_phase_inv_def by (vcg_jackhammer (no_thin_post_inv)) (auto simp: hp_step_rel_def) + +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. + +\ + +lemma gc_fM_fA_invL_eq_imp: + "eq_imp (\(_::unit) s. (AT s gc, s\ gc, sys_fA s\, sys_fM s\, sys_mem_store_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] + +context gc +begin + +lemma fM_fA_invL[intro]: + "\ fM_fA_invL \ gc" +by vcg_jackhammer + +lemma fM_rel_inv[intro]: + notes fun_upd_apply[simp] + shows + "\ fM_fA_invL \<^bold>\ handshake_invL \<^bold>\ tso_lock_invL \<^bold>\ LSTP fM_rel_inv \ + gc + \ LSTP fM_rel_inv \" +by (vcg_jackhammer (no_thin_post_inv); simp add: fM_rel_inv_def fM_rel_def) + +lemma fA_rel_inv[intro]: + notes fun_upd_apply[simp] + shows + "\ fM_fA_invL \<^bold>\ handshake_invL \<^bold>\ LSTP fA_rel_inv \ + gc + \ LSTP fA_rel_inv \" +by (vcg_jackhammer (no_thin_post_inv); simp add: fA_rel_inv_def; auto simp: fA_rel_def) + +end + +context mut_m +begin + +lemma gc_fM_fA_invL[intro]: + "\ gc.fM_fA_invL \ mutator m" +by (vcg_chainsaw gc.fM_fA_invL_def) + +lemma fM_rel_inv[intro]: + notes fun_upd_apply[simp] + shows + "\ LSTP fM_rel_inv \ mutator m" +unfolding fM_rel_inv_def by (vcg_jackhammer (no_thin_post_inv); simp add: fM_rel_def; elim disjE; auto split: if_splits) + +lemma fA_rel_inv[intro]: + notes fun_upd_apply[simp] + shows + "\ LSTP fA_rel_inv \ mutator m" +unfolding fA_rel_inv_def by (vcg_jackhammer (no_thin_post_inv); simp add: fA_rel_def; elim disjE; auto split: if_splits) + +end + + +context gc +begin + +lemma fA_neq_locs_diff_fA_tso_empty_locs: + "fA_neq_locs - fA_tso_empty_locs = {}" +apply (simp add: fA_neq_locs_def fA_tso_empty_locs_def locset_cache) +apply (simp add: loc_defs) +done + +end + +context sys +begin + +lemma gc_fM_fA_invL[intro]: + "\ gc.fM_fA_invL \<^bold>\ LSTP (fA_rel_inv \<^bold>\ fM_rel_inv \<^bold>\ tso_store_inv) \ + sys + \ gc.fM_fA_invL \" +apply( vcg_chainsaw (no_thin) gc.fM_fA_invL_def + ; (simp add: p_not_sys)?; (erule disjE)?; clarsimp split: if_splits ) +proof(vcg_name_cases sys gc) + case (tso_dequeue_store_buffer_mark_noop_mfence s s' ws) then show ?case by (clarsimp simp: fA_rel_inv_def fA_rel_def) +next case (tso_dequeue_store_buffer_fA_neq_locs s s' ws) then show ?case + 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 gc.fA_neq_locs_diff_fA_tso_empty_locs) + done +next case (tso_dequeue_store_buffer_fA_eq_locs s s' ws) then show ?case by (clarsimp simp: fA_rel_inv_def fA_rel_def) +next case (tso_dequeue_store_buffer_idle_flip_noop_mfence s s' ws) then show ?case by (clarsimp simp: fM_rel_inv_def fM_rel_def) +next case (tso_dequeue_store_buffer_fM_eq_locs s s' ws) then show ?case by (clarsimp simp: fM_rel_inv_def fM_rel_def) +qed + +lemma fM_rel_inv[intro]: + notes fun_upd_apply[simp] + shows + "\ LSTP (fM_rel_inv \<^bold>\ tso_store_inv) \ sys \ LSTP fM_rel_inv \" +apply (vcg_jackhammer (no_thin_post_inv)) +apply (clarsimp simp: do_store_action_def fM_rel_inv_def fM_rel_def p_not_sys + split: mem_store_action.splits) +apply (intro allI conjI impI; clarsimp) +done + +lemma fA_rel_inv[intro]: + notes fun_upd_apply[simp] + shows + "\ LSTP (fA_rel_inv \<^bold>\ tso_store_inv) \ sys \ LSTP fA_rel_inv \" +apply (vcg_jackhammer (no_thin_post_inv)) +apply (clarsimp simp: do_store_action_def fA_rel_inv_def fA_rel_def p_not_sys + split: mem_store_action.splits) +apply (intro allI conjI impI; clarsimp) +done + +end + + +subsubsection\sys phase inv\ + +context mut_m +begin + +lemma sys_phase_inv[intro]: + notes if_split_asm[split del] + notes fun_upd_apply[simp] + shows + "\ 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 \" +proof( (vcg_jackhammer (no_thin_post_inv) + ; clarsimp simp: fA_rel_inv_def fM_rel_inv_def sys_phase_inv_aux_case heap_colours_colours + split: hs_phase.splits if_splits ) + , vcg_name_cases ) + case (alloc s s' rb) then show ?case + by (clarsimp simp: fA_rel_def fM_rel_def no_black_refs_def + dest!: handshake_phase_invD phase_rel_invD + split: hs_phase.splits) +next case (store_ins_mo_co_mark0 s s' y) then show ?case + by (fastforce simp: fA_rel_def fM_rel_def hp_step_rel_def + dest!: handshake_phase_invD phase_rel_invD) +next case (store_ins_mo_co_mark s s' y) then show ?case + apply - + 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; fail) + 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 +next case (store_del_mo_co_mark0 s s' y) then show ?case + apply (clarsimp simp: hp_step_rel_def dest!: handshake_phase_invD phase_rel_invD) + apply (metis (no_types, lifting) mut_m.no_grey_refs_not_rootD mutator_phase_inv_aux.simps(5)) + done +next case (store_del_mo_co_mark s s' y) then show ?case + apply - + 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; fail) + 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) + apply clarsimp + apply (elim disjE, simp_all add: no_grey_refs_not_rootD filter_empty_conv) + apply fastforce + done +next case (hs_get_roots_done s s') then show ?case + apply (clarsimp simp: hp_step_rel_def phase_rel_def filter_empty_conv + dest!: handshake_phase_invD phase_rel_invD) + apply auto + done +next case (hs_get_roots_loop_mo_co_mark s s' y) then show ?case + apply - + 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; fail) + 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) + apply clarsimp + apply (elim disjE, simp_all add: no_grey_refs_not_rootD filter_empty_conv)[1] + apply fastforce + done +next case (hs_get_work_done s s') then show ?case + apply (clarsimp simp: hp_step_rel_def phase_rel_def filter_empty_conv + dest!: handshake_phase_invD phase_rel_invD) + apply auto + done +qed (clarsimp simp: hp_step_rel_def dest!: handshake_phase_invD phase_rel_invD)+ + +end + +lemma (in gc) sys_phase_inv[intro]: + notes fun_upd_apply[simp] + shows + "\ 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_store_inv) \ + gc + \ LSTP sys_phase_inv \" +proof(vcg_jackhammer (no_thin_post_inv), vcg_name_cases) + case (mark_loop_get_work_load_W s s') then show ?case by (fastforce dest!: phase_rel_invD simp: phase_rel_def no_grey_refsD filter_empty_conv) +next case (mark_loop_blacken s s') then show ?case by (meson no_grey_refsD(1)) +next case (mark_loop_mo_co_W s s') then show ?case by (meson no_grey_refsD(1)) +next case (mark_loop_mo_co_mark s s') then show ?case by (meson no_grey_refsD(1)) +next case (mark_loop_get_roots_load_W s s') then show ?case by (fastforce dest!: phase_rel_invD simp: phase_rel_def no_grey_refsD filter_empty_conv) +next case (mark_loop_get_roots_init_type s s') then show ?case by (fastforce dest!: phase_rel_invD simp: phase_rel_def no_grey_refsD filter_empty_conv) +next case (idle_noop_init_type s s') then show ?case using black_heap_no_greys by blast +qed + +lemma no_grey_refs_no_marks[simp]: + "\ no_grey_refs s; valid_W_inv s \ \ \sys_mem_store_buffers p s = mw_Mark r fl # ws" +unfolding no_grey_refs_def by (metis greyI(1) list.set_intros(1) valid_W_invE(5)) +(* FIXME suggests redundancy in valid_W_inv rules: by (meson greyI(1) valid_W_invD(1)) *) + +context sys +begin + +lemma black_heap_dequeue_mark[iff]: + "\ sys_mem_store_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 := map_option (obj_mark_update (\_. fl)) (sys_heap s r)), mem_store_buffers := (mem_store_buffers (s sys))(p := ws)\))" +unfolding black_heap_def by (metis colours_distinct(4) valid_W_invD(1) white_valid_ref) + +lemma white_heap_dequeue_fM[iff]: + "black_heap s\ + \ white_heap (s\(sys := s\ sys\fM := \ sys_fM s\, mem_store_buffers := (mem_store_buffers (s\ sys))(gc := ws)\))" +unfolding black_heap_def white_heap_def black_def white_def by clarsimp (* FIXME rules? *) + +lemma black_heap_dequeue_fM[iff]: + "\ white_heap s\; no_grey_refs s\ \ + \ black_heap (s\(sys := s\ sys\fM := \ sys_fM s\, mem_store_buffers := (mem_store_buffers (s\ sys))(gc := ws)\))" +unfolding black_heap_def white_heap_def no_grey_refs_def by auto + +lemma sys_phase_inv[intro]: + notes if_split_asm[split del] + notes fun_upd_apply[simp] + shows + "\ 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_store_inv \<^bold>\ valid_W_inv) \ + sys + \ LSTP sys_phase_inv \" +proof(vcg_jackhammer (no_thin_post_inv) + , clarsimp simp: fA_rel_inv_def fM_rel_inv_def p_not_sys + , vcg_name_cases) + case (tso_dequeue_store_buffer s s' p w ws) then show ?case + apply (clarsimp simp: do_store_action_def sys_phase_inv_aux_case + split: mem_store_action.splits hs_phase.splits if_splits) + apply (clarsimp simp: fA_rel_def fM_rel_def; erule disjE; clarsimp simp: fA_rel_def fM_rel_def)+ + apply (metis (mono_tags, lifting) filter.simps(2) list.discI tso_store_invD(4)) + apply auto + done +qed + +end +(*<*) + +end +(*>*) diff --git a/thys/ConcurrentGC/Proofs.thy b/thys/ConcurrentGC/Proofs.thy --- a/thys/ConcurrentGC/Proofs.thy +++ b/thys/ConcurrentGC/Proofs.thy @@ -1,333 +1,164 @@ (*<*) (* * 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 + TSO + Phases + MarkObject StrongTricolour + Valid_Refs + Worklists + Global_Noninterference + Noninterference + Initial_Conditions 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 | 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 | fastforce )+ + done +qed + +(* FIXME split mutators_phase_inv from global invs to local invs. Move to StrongTricolour or similar. note dependence on I *) +lemma mutators_phase_inv[intro]: + "\ I \ mutator m \ LSTP (mut_m.mutator_phase_inv 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 (simp add: I_def gc.invsL_def invs_def Local_Invariants.invsL_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 +context gc_system 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) +theorem I: "gc_system \\<^bsub>pre\<^esub> I" +apply (rule 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\ +\label{sec:proofs-headline-safety} + 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) +corollary safety: "gc_system \\<^bsub>pre\<^esub> LSTP valid_refs" +using I unfolding I_def invs_def valid_refs_def prerun_valid_def apply clarsimp -apply (drule_tac x=xa in spec, fastforce simp: mut_m.reachable_def) +apply (drule_tac x=\ in spec) +apply (drule (1) mp) +apply (rule alwaysI) +apply (erule_tac i=i in alwaysE) +apply (clarsimp simp: valid_refs_invD(1)) done -(*>*) -text\\ - end text\ The GC is correct for the remaining fixed-but-arbitrary initial conditions. \ interpretation gc_system_interpretation: gc_system undefined . +(* unused_thms Main Sublist CIMP - *) -subsection\A concrete system state\ + +section\ A concrete system state \label{sec: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. +concrete initial state that satisfies the initial conditions. The heap +is shown in Figure~\ref{fig:concrete-heap}. We use Isabelle's notation +for types of a given size. + +\begin{figure} + \centering + \includegraphics{heap.pdf} + \caption{A concrete system state.} + \label{fig:concrete-heap} +\end{figure} \ (*<*) end (*>*) diff --git a/thys/ConcurrentGC/Proofs_Basis.thy b/thys/ConcurrentGC/Proofs_Basis.thy new file mode 100644 --- /dev/null +++ b/thys/ConcurrentGC/Proofs_Basis.thy @@ -0,0 +1,583 @@ +(*<*) +(* + * 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 + Model + "HOL-Library.Simps_Case_Conv" +begin + +(* From 40e16c534243 by Makarius. Doesn't seem to have a huge impact on run time now (2021-01-07) *) +declare subst_all [simp del] [[simproc del: defined_all]] + +(*>*) +section\ Proofs Basis \label{sec:proofs-basis}\ + +text\ + +Extra HOL. + +\ + +lemma Set_bind_insert[simp]: + "Set.bind (insert a A) B = B a \ (Set.bind A B)" +by (auto simp: Set.bind_def) + +lemma option_bind_invE[elim]: + "\ Option.bind f g = None; \a. \ f = Some a; g a = None \ \ Q; f = None \ Q \ \ Q" + "\ Option.bind f g = Some x; \a. \ f = Some a; g a = Some x \ \ Q \ \ Q" +by (case_tac [!] f) simp_all + +lemmas conj_explode = conj_imp_eq_imp_imp + +text\ + +Tweak the default simpset: +\<^item> "not in dom" as a premise negates the goal +\<^item> we always want to execute suffix +\<^item> we try to make simplification rules about @{const \fun_upd\} more stable + +\ + +declare dom_def[simp] +declare suffix_to_prefix[simp] +declare map_option.compositionality[simp] +declare o_def[simp] +declare Option.Option.option.set_map[simp] +declare bind_image[simp] + +declare fun_upd_apply[simp del] +declare fun_upd_same[simp] +declare fun_upd_other[simp] + +declare gc_phase.case_cong[cong] +declare mem_store_action.case_cong[cong] +declare process_name.case_cong[cong] +declare hs_phase.case_cong[cong] +declare hs_type.case_cong[cong] + +declare if_split_asm[split] + +text\ + +Collect the component definitions. Inline everything. This is what the proofs work on. +Observe we lean heavily on locales. + +\ + +context gc +begin + +lemmas all_com_defs = + (* gc.com_def *) handshake_done_def handshake_init_def handshake_noop_def handshake_get_roots_def handshake_get_work_def + mark_object_fn_def + +lemmas com_def2 = com_def[simplified all_com_defs append.simps if_True if_False] + +intern_com com_def2 + +end + +context mut_m +begin + +lemmas all_com_defs = + (* mut.com_def *) mut_m.handshake_def mut_m.store_def + mark_object_fn_def + +lemmas com_def2 = mut_m.com_def[simplified all_com_defs append.simps if_True if_False] + +intern_com com_def2 + +end + +context sys +begin + +lemmas all_com_defs = + (* sys.com_def *) sys.alloc_def sys.free_def sys.mem_TSO_def sys.handshake_def + +lemmas com_def2 = com_def[simplified all_com_defs append.simps if_True if_False] + +intern_com com_def2 + +end + +lemmas all_com_interned_defs = gc.com_interned mut_m.com_interned sys.com_interned + +named_theorems inv "Location-sensitive invariant definitions" +named_theorems nie "Non-interference elimination rules" + + +subsection\ Model-specific 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 prefixed :: "location \ location set" where + "prefixed p \ { l . prefix p l }" + +abbreviation suffixed :: "location \ location set" where + "suffixed p \ { l . suffix p l }" + +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_Mutate_Payload w \ \r f pl. w = mw_Mutate_Payload r f pl" +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, 'payload, '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, 'payload, 'ref) lsts_pred" (infix "in'_ghost'_honorary'_grey" 50) where + "r in_ghost_honorary_grey p \ \s. r \ ghost_honorary_grey (s p)" + +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, 'payload, 'ref) lsts_pred \ ('field, 'mut, 'payload, 'ref) gc_pred" where + "at_gc l P \ at gc l \<^bold>\ LSTP P" + +abbreviation atS_gc :: "location set \ ('field, 'mut, 'payload, 'ref) lsts_pred \ ('field, 'mut, 'payload, 'ref) gc_pred" where + "atS_gc ls P \ atS gc ls \<^bold>\ LSTP P" + +context mut_m +begin + +abbreviation at_mut :: "location \ ('field, 'mut, 'payload, 'ref) lsts_pred \ ('field, 'mut, 'payload, 'ref) gc_pred" where + "at_mut l P \ at (mutator m) l \<^bold>\ LSTP P" + +abbreviation atS_mut :: "location set \ ('field, 'mut, 'payload, 'ref) lsts_pred \ ('field, 'mut, 'payload, '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_hs_phase s \ ghost_hs_phase (s (mutator m))" +abbreviation "mut_ghost_honorary_root s \ ghost_honorary_root (s (mutator m))" +abbreviation "mut_hs_pending s \ mutator_hs_pending (s (mutator m))" +abbreviation "mut_hs_type s \ hs_type (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 + +abbreviation sys_heap :: "('field, 'mut, 'payload, 'ref) lsts \ 'ref \ ('field, 'payload, '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_hs_in_sync m s \ ghost_hs_in_sync (s sys) m" +abbreviation "sys_ghost_hs_phase s \ ghost_hs_phase (s sys)" +abbreviation "sys_hs_pending m s \ hs_pending (s sys) m" +abbreviation "sys_hs_type s \ hs_type (s sys)" +abbreviation "sys_mem_store_buffers p s \ mem_store_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, 'payload, 'ref) lsts_pred \ ('field, 'mut, 'payload, '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_store_buffers (s sys) p)" +abbreviation (input) "tso_pending_store p w s \ w \ set (mem_store_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_mw_mutate p \ tso_pending p is_mw_Mutate" +abbreviation (input) "tso_pending_mutate p \ tso_pending p (is_mw_Mutate \<^bold>\ is_mw_Mutate_Payload)" \\ TSO makes it (mostly) not worth distinguishing these. \ +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. In some sense this encodes a three-valued logic. + +\ + +definition obj_at :: "(('field, 'payload, 'ref) object \ bool) \ 'ref \ ('field, 'mut, 'payload, '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, 'payload, 'ref) lsts_pred" where + "valid_ref r \ obj_at \True\ r" + +definition valid_null_ref :: "'ref option \ ('field, 'mut, 'payload, '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, 'payload, '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. + +\ + +definition reaches :: "'ref \ 'ref \ ('field, 'mut, 'payload, '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\. + +\ + +definition obj_at_field_on_heap :: "('ref \ bool) \ 'ref \ 'field \ ('field, 'mut, 'payload, 'ref) lsts_pred" where + "obj_at_field_on_heap P r f \ \s. + case map_option obj_fields (sys_heap s r) of + None \ False + | Some fs \ (case fs f of None \ True + | Some r' \ P r')" + + +subsection\Object colours\ + +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 or @{const \ghost_honorary_grey\} +\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, 'payload, 'ref) lsts_pred" where + "marked r s \ obj_at (\obj. obj_mark obj = sys_fM s) r s" + +definition white :: "'ref \ ('field, 'mut, 'payload, '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, 'payload, 'ref) lsts \ 'ref set" where + "WL p = (\s. W (s p) \ ghost_honorary_grey (s p))" + +definition grey :: "'ref \ ('field, 'mut, 'payload, 'ref) lsts_pred" where + "grey r = (\<^bold>\p. \r\ \<^bold>\ WL p)" + +definition black :: "'ref \ ('field, 'mut, 'payload, 'ref) lsts_pred" where + "black r \ marked r \<^bold>\ \<^bold>\grey r" + +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 white_def obj_at_def split: option.splits) (* FIXME invisible *) + +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 white_def obj_at_def split: option.splits) (* FIXME invisible *) + +text\ + +In some phases the heap is monochrome. + +\ + +definition black_heap :: "('field, 'mut, 'payload, 'ref) lsts_pred" where + "black_heap = (\<^bold>\r. valid_ref r \<^bold>\ black r)" + +definition white_heap :: "('field, 'mut, 'payload, 'ref) lsts_pred" where + "white_heap = (\<^bold>\r. valid_ref r \<^bold>\ white r)" + +definition no_black_refs :: "('field, 'mut, 'payload, 'ref) lsts_pred" where + "no_black_refs = (\<^bold>\r. \<^bold>\black r)" + +definition no_grey_refs :: "('field, 'mut, 'payload, 'ref) lsts_pred" where + "no_grey_refs = (\<^bold>\r. \<^bold>\grey r)" + + +subsection\Reachability\ + +text\ + +We treat pending TSO heap mutations as extra mutator roots. + +\ + +abbreviation store_refs :: "('field, 'payload, 'ref) mem_store_action \ 'ref set" where + "store_refs w \ case w of mw_Mutate r f r' \ {r} \ Option.set_option r' | mw_Mutate_Payload r f pl \ {r} | _ \ {}" + +definition (in mut_m) tso_store_refs :: "('field, 'mut, 'payload, 'ref) lsts \ 'ref set" where + "tso_store_refs = (\s. \w \ set (sys_mem_store_buffers (mutator m) s). store_refs w)" + +abbreviation (in mut_m) root :: "'ref \ ('field, 'mut, 'payload, 'ref) lsts_pred" where + "root x \ \x\ \<^bold>\ mut_roots \<^bold>\ mut_ghost_honorary_root \<^bold>\ tso_store_refs" + +definition (in mut_m) reachable :: "'ref \ ('field, 'mut, 'payload, 'ref) lsts_pred" where + "reachable y = (\<^bold>\x. root x \<^bold>\ x reaches y)" + +definition grey_reachable :: "'ref \ ('field, 'mut, 'payload, 'ref) lsts_pred" where + "grey_reachable y = (\<^bold>\g. grey g \<^bold>\ g reaches y)" + + +subsection\ Sundry detritus \ + +lemmas eq_imp_simps = \\equations for deriving useful things from @{const \eq_imp\} facts\ + eq_imp_def + all_conj_distrib + split_paired_All split_def fst_conv snd_conv prod_eq_iff + conj_explode + simp_thms + +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 + +text\ obj at \ + +lemma obj_at_cong[cong]: + "\\obj. sys_heap s r = Some obj \ P obj = P' obj; r = r'; s = s'\ + \ obj_at P r s \ obj_at P' r' s'" +unfolding obj_at_def by (simp cong: option.case_cong) + +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 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] + +lemma obj_at_simps: + "obj_at (\obj. P obj \ Q obj) r s \ obj_at P r s \ obj_at Q r s" +(* "obj_at (\obj. R) r s \ valid_ref r s \ R" looks good but applies to valid_ref *) +by (simp_all split: obj_at_splits) + +text\ obj at field on heap \ + +lemma obj_at_field_on_heap_cong[cong]: + "\\r' obj. \sys_heap s r = Some obj; obj_fields obj f = Some r'\\ P r' = P' r'; r = r'; f = f'; s = s'\ + \ obj_at_field_on_heap P r f s \ obj_at_field_on_heap P' r' f' s'" +unfolding obj_at_field_on_heap_def by (simp cong: option.case_cong) + +lemma obj_at_field_on_heap_split: + "Q (obj_at_field_on_heap P r f s) \ ((sys_heap s r = None \ Q False) + \ (\obj. sys_heap s r = Some obj \ obj_fields obj f = None \ Q True) + \ (\r' obj. sys_heap s r = Some obj \ obj_fields obj f = Some r' \ Q (P r')))" +by (simp add: obj_at_field_on_heap_def split: option.splits) + +lemma obj_at_field_on_heap_split_asm: + "Q (obj_at_field_on_heap P r f s) \ (\ ((sys_heap s r = None \ \Q False) + \ (\obj. sys_heap s r = Some obj \ obj_fields obj f = None \ \ Q True) + \ (\r' obj. sys_heap s r = Some obj \ obj_fields obj f = Some r' \ \ Q (P r'))))" +by (simp add: obj_at_field_on_heap_def split: option.splits) + +lemmas obj_at_field_on_heap_splits = obj_at_field_on_heap_split obj_at_field_on_heap_split_asm + +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_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_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 valid_null_ref_eq_imp: + "eq_imp (\(_::unit) s. Option.bind r (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" +unfolding valid_null_ref_def by simp_all + +text\Derive simplification rules from \case\ expressions\ + +simps_of_case hs_step_simps[simp]: hs_step_def (splits: hs_phase.split) +simps_of_case do_load_action_simps[simp]: fun_cong[OF do_load_action_def[simplified atomize_eq]] (splits: mem_load_action.split) +simps_of_case do_store_action_simps[simp]: fun_cong[OF do_store_action_def[simplified atomize_eq]] (splits: mem_store_action.split) + +(* This gives some indication of how much we're cheating on the TSO front. *) +lemma do_store_action_prj_simps[simp]: + "fM (do_store_action w s) = fl \ (fM s = fl \ w \ mw_fM (\fM s)) \ w = mw_fM fl" + "fl = fM (do_store_action w s) \ (fl = fM s \ w \ mw_fM (\fM s)) \ w = mw_fM fl" + "fA (do_store_action w s) = fl \ (fA s = fl \ w \ mw_fA (\fA s)) \ w = mw_fA fl" + "fl = fA (do_store_action w s) \ (fl = fA s \ w \ mw_fA (\fA s)) \ w = mw_fA fl" + "ghost_hs_in_sync (do_store_action w s) = ghost_hs_in_sync s" + "ghost_hs_phase (do_store_action w s) = ghost_hs_phase s" + "ghost_honorary_grey (do_store_action w s) = ghost_honorary_grey s" + "hs_pending (do_store_action w s) = hs_pending s" + "hs_type (do_store_action w s) = hs_type s" + "heap (do_store_action w s) r = None \ heap s r = None" + "mem_lock (do_store_action w s) = mem_lock s" + "phase (do_store_action w s) = ph \ (phase s = ph \ (\ph'. w \ mw_Phase ph') \ w = mw_Phase ph)" + "ph = phase (do_store_action w s) \ (ph = phase s \ (\ph'. w \ mw_Phase ph') \ w = mw_Phase ph)" + "W (do_store_action w s) = W s" +by (auto simp: do_store_action_def fun_upd_apply split: mem_store_action.splits obj_at_splits) + + +text\ reaches \ + +lemma reaches_refl[iff]: + "(r reaches r) s" +unfolding reaches_def by blast + +lemma reaches_step[intro]: + "\(x reaches y) s; (y points_to z) s\ \ (x reaches z) s" + "\(y reaches z) s; (x points_to y) s\ \ (x reaches z) s" +unfolding reaches_def + apply (simp add: rtranclp.rtrancl_into_rtrancl) +apply (simp add: converse_rtranclp_into_rtranclp) +done + +lemma reaches_induct[consumes 1, case_names refl step, induct set: reaches]: + assumes "(x reaches y) s" + assumes "\x. P x x" + assumes "\x y z. \(x reaches y) s; P x y; (y points_to z) s\ \ P x z" + shows "P x y" +using assms unfolding reaches_def by (rule rtranclp.induct) + +lemma converse_reachesE[consumes 1, case_names base step]: + assumes "(x reaches z) s" + assumes "x = z \ P" + assumes "\y. \(x points_to y) s; (y reaches z) s\ \ P" + shows P +using assms unfolding reaches_def by (blast elim: converse_rtranclpE) + +lemma reaches_fields: \ \Complicated condition takes care of \alloc\: collapses no object and object with no fields\ + assumes "(x reaches y) s'" + assumes "\r'. \(ran ` obj_fields ` set_option (sys_heap s' r')) = \(ran ` obj_fields ` set_option (sys_heap s r'))" + shows "(x reaches y) s" +using assms +proof induct + case (step x y z) + then have "(y points_to z) s" + by (cases "sys_heap s y") + (auto 10 10 simp: ran_def obj_at_def split: option.splits dest!: spec[where x=y]) + with step show ?case by blast +qed simp + +lemma reaches_eq_imp: + "eq_imp (\r' s. \(ran ` obj_fields ` set_option (sys_heap s r'))) + (x reaches y)" +unfolding eq_imp_def by (metis reaches_fields) + +lemmas reaches_fun_upd[simp] = eq_imp_fun_upd[OF reaches_eq_imp, simplified eq_imp_simps, rule_format] + +text\ + +Location-specific facts. + +\ + +lemma obj_at_mark_dequeue[simp]: + "obj_at P r (s(sys := s sys\ heap := (sys_heap s)(r' := map_option (obj_mark_update (\_. fl)) (sys_heap s r')), mem_store_buffers := wb' \)) +\ obj_at (\obj. (P (if r = r' then obj\ obj_mark := fl \ else obj))) r s" +by (clarsimp simp: fun_upd_apply split: obj_at_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 := map_option (\obj :: ('field, 'payload, 'ref) object. obj\obj_fields := (obj_fields obj)(f := opt_r')\) (sys_heap s r)), + mem_store_buffers := (mem_store_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' := map_option (obj_mark_update (\_. fl)) (sys_heap s r')), mem_store_buffers := sb'\)) +\ obj_at_field_on_heap P r f s" +by (auto simp: obj_at_field_on_heap_def fun_upd_apply split: option.splits obj_at_splits) + +lemma obj_at_field_on_heap_no_pending_stores: + "\ sys_load (mutator m) (mr_Ref r f) (s sys) = mv_Ref opt_r'; \opt_r'. mw_Mutate r f opt_r' \ set (sys_mem_store_buffers (mutator m) s); valid_ref r s \ + \ obj_at_field_on_heap (\r. opt_r' = Some r) r f s" +unfolding sys_load_def fold_stores_def +apply clarsimp +apply (rule fold_invariant[where P="\fr. obj_at_field_on_heap (\r'. Option.bind (heap (fr (s sys)) r) (\obj. obj_fields obj f) = Some r') r f s" + and Q="\w. w \ set (sys_mem_store_buffers (mutator m) s)"]) + apply fastforce + apply (fastforce simp: obj_at_field_on_heap_def split: option.splits obj_at_splits) +apply (auto simp: do_store_action_def map_option_case fun_upd_apply + split: obj_at_field_on_heap_splits option.splits obj_at_splits mem_store_action.splits) +done +(*<*) + +end + +(*>*) diff --git a/thys/ConcurrentGC/Proofs_basis.thy b/thys/ConcurrentGC/Proofs_basis.thy deleted file mode 100644 --- a/thys/ConcurrentGC/Proofs_basis.thy +++ /dev/null @@ -1,351 +0,0 @@ -(*<*) -(* - * 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/ROOT b/thys/ConcurrentGC/ROOT --- a/thys/ConcurrentGC/ROOT +++ b/thys/ConcurrentGC/ROOT @@ -1,20 +1,34 @@ chapter AFP session ConcurrentGC (AFP slow) = ConcurrentIMP + options [timeout = 36000] directories "concrete" theories [show_question_marks = false, names_short] Model + Proofs_Basis + Global_Invariants + Local_Invariants Tactics - Proofs_basis + + Global_Invariants_Lemmas + Local_Invariants_Lemmas + + Initial_Conditions + + Noninterference + Global_Noninterference + MarkObject + Phases + StrongTricolour TSO - Handshakes - MarkObject - StrongTricolour + Valid_Refs + Worklists + Proofs "concrete/Concrete_heap" "concrete/Concrete" document_files "root.bib" "root.tex" + "heap.pdf" diff --git a/thys/ConcurrentGC/StrongTricolour.thy b/thys/ConcurrentGC/StrongTricolour.thy --- a/thys/ConcurrentGC/StrongTricolour.thy +++ b/thys/ConcurrentGC/StrongTricolour.thy @@ -1,3402 +1,532 @@ (*<*) (* * 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 + Global_Invariants_Lemmas + Local_Invariants_Lemmas + Tactics 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) +(* local lemma bucket *) -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 +context mut_m +begin -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)" +(* marked insertions *) -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 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_store_buffers := (mem_store_buffers (s sys))(mutator m := sys_mem_store_buffers (mutator m) s @ [mw_Mutate r f opt_r'])\))" +by (auto simp: marked_insertions_def + split: mem_store_action.splits option.splits) -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]: +lemma 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 (clarsimp simp: marked_insertions_def split: mem_store_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 +(* marked_deletions *) -lemmas marked_deletions_fun_upd[simp] = eq_imp_fun_upd[OF mut_m.marked_deletions_eq_imp, simplified eq_imp_simps, rule_format] +lemma 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_store_buffers := (mem_store_buffers (s sys))(mutator m := sys_mem_store_buffers (mutator m) s @ [mw_Mutate r f opt_r'])\))" +by (auto simp: marked_deletions_def + split: mem_store_action.splits option.splits) -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]: +lemma 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 (clarsimp simp: marked_deletions_def split: mem_store_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) +end - 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\)) + "eq_imp (\(_::unit) s. (AT s gc, s\ gc, sys_fM s\, 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) +proof(vcg_jackhammer, vcg_name_cases) + case sweep_loop_ref_done then show ?case by blast +next case sweep_loop_check then show ?case + apply (clarsimp split: obj_at_splits) + apply (metis (no_types, lifting) option.collapse option.inject) + done +next case sweep_loop_load_mark then show ?case by (clarsimp split: obj_at_splits) +qed -apply (rename_tac s s' x) -apply (case_tac "x \ gc_refs s\") - apply force -apply blast +context gc +begin -apply (clarsimp split: obj_at_splits) -done - -lemma sweep_loop_sweep_locs[iff]: +lemma sweep_loop_locs_subseteq_sweep_locs: "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]: +lemma sweep_locs_subseteq_fM_tso_empty_locs: "sweep_locs \ fM_tso_empty_locs" -by (auto simp: sweep_locs_def fM_tso_empty_locs_def) +by (auto simp: sweep_locs_def fM_tso_empty_locs_def loc_defs) 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) +by (auto simp: sweep_loop_locs_def fM_eq_locs_def sweep_locs_def loc_defs) 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]: +lemma black_heap_locs_subseteq_fM_tso_empty_locs: "black_heap_locs \ fM_tso_empty_locs" -by (auto simp: black_heap_locs_def fM_tso_empty_locs_def) +by (auto simp: black_heap_locs_def fM_tso_empty_locs_def loc_defs) lemma black_heap_locs_fM_eq_locs: "black_heap_locs \ fM_eq_locs" -by (simp add: black_heap_locs_def fM_eq_locs_def) +by (simp add: black_heap_locs_def fM_eq_locs_def loc_defs) 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) +by (simp add: black_heap_locs_def fA_eq_locs_def sweep_locs_def loc_defs) -lemma (in gc) fM_invL_tso_emptyD: +lemma fM_fA_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 (auto simp: locset_cache atS_simps dest: atS_mono) + apply (simp add: atS_mono gc.sweep_loop_locs_subseteq_sweep_locs; fail) 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) +apply (auto simp: sweep_locs_def sweep_loop_not_choose_ref_locs_def intro: append_prefixD) done +end + 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) \ + "\ gc.fM_fA_invL \<^bold>\ gc.gc_W_empty_invL \<^bold>\ gc.sweep_loop_invL + \<^bold>\ LSTP (tso_store_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 *) +proof(vcg_jackhammer (keep_locs) (no_thin_post_inv), vcg_name_cases) + case (tso_dequeue_store_buffer s s' p w ws) then show ?case + proof(cases w) + case (mw_Mark r fl) with tso_dequeue_store_buffer show ?thesis + apply - + apply (rule gc.gc_sweep_loop_invL_locsE) + apply (simp only: gc.gc_W_empty_invL_def gc.no_grey_refs_locs_def cong del: atS_state_weak_cong) + apply (clarsimp simp: atS_un) + apply (thin_tac "AT _ = _") (* FIXME speed the metis call up a bit *) + apply (thin_tac "at _ _ _ \ _")+ + apply (metis (mono_tags, lifting) filter.simps(2) loc_mem_tac_simps(4) no_grey_refs_no_pending_marks) + done + next case (mw_Mutate r f opt_r') with tso_dequeue_store_buffer show ?thesis by clarsimp (erule gc_sweep_loop_invL_niE; simp add: fun_eq_iff fun_upd_apply) + next case (mw_Mutate_Payload r f pl) with tso_dequeue_store_buffer show ?thesis by clarsimp (erule gc_sweep_loop_invL_niE; simp add: fun_eq_iff fun_upd_apply) + next case (mw_fA fl) with tso_dequeue_store_buffer show ?thesis by - (erule gc_sweep_loop_invL_niE; simp add: fun_eq_iff) + next case (mw_fM fl) with tso_dequeue_store_buffer show ?thesis + apply - + apply (rule gc.gc_sweep_loop_invL_locsE) + apply (case_tac p; clarsimp) + apply (drule (1) gc.fM_fA_invL_tso_emptyD) + apply simp_all + using gc.black_heap_locs_subseteq_fM_tso_empty_locs gc.sweep_locs_subseteq_fM_tso_empty_locs apply blast + done + next case (mw_Phase ph) with tso_dequeue_store_buffer show ?thesis by - (erule gc_sweep_loop_invL_niE; simp add: fun_eq_iff) + qed +qed 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 +proof( vcg_chainsaw (no_thin) gc.fM_fA_invL_def gc.sweep_loop_invL_def gc.handshake_invL_def, vcg_name_cases gc) + case (sweep_loop_locs s s' rb) then show ?case by (metis (no_types, lifting) atS_mono gc.sweep_loop_locs_fA_eq_locs gc.sweep_loop_locs_fM_eq_locs) +next case (black_heap_locs s s' rb) then show ?case by (metis (no_types, lifting) atS_mono gc.black_heap_locs_fA_eq_locs gc.black_heap_locs_fM_eq_locs) +qed - 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 +subsection\ Mutator proofs \ -lemma (in gc) no_black_refs_sweep_loop_free[simp]: +context mut_m +begin + +(* reachable snapshot inv *) + +lemma 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} \))" +unfolding in_snapshot_def reachable_snapshot_inv_def by (auto simp: fun_upd_apply) + +lemma reachable_snapshot_inv_hs_get_roots_done: + 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_hs_phase := ghp'\, + sys := s sys\hs_pending := hp', W := sys_W s \ mut_W s, ghost_hs_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) (* 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_store_refs_def) + apply (rename_tac w; case_tac w; fastforce) + done + next + case (reaches x y) + from reaches vri have "valid_ref x s" "valid_ref y s" + using reachable_points_to by fastforce+ + with reaches sti vri show ?case + apply (clarsimp simp: in_snapshot_def) + 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 reachable_snapshot_inv_hs_get_work_done: + "reachable_snapshot_inv s + \ reachable_snapshot_inv + (s(mutator m := s (mutator m)\W := {}\, + sys := s sys\hs_pending := pending', W := sys_W s \ mut_W s, + ghost_hs_in_sync := (ghost_hs_in_sync (s sys))(m := True)\))" +by (simp add: reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def) + +lemma reachable_snapshot_inv_deref_del: + "\ reachable_snapshot_inv s; sys_load (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'\))" +unfolding reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def by (clarsimp simp: fun_upd_apply) + +lemma mutator_phase_inv[intro]: + notes fun_upd_apply[simp] + notes reachable_snapshot_inv_deref_del[simp] + notes if_split_asm[split del] + shows + "\ 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 (handshake_phase_inv \<^bold>\ mutators_phase_inv \<^bold>\ phase_rel_inv \<^bold>\ sys_phase_inv \<^bold>\ fA_rel_inv \<^bold>\ fM_rel_inv \<^bold>\ valid_refs_inv \<^bold>\ strong_tricolour_inv \<^bold>\ valid_W_inv) \ + mutator m + \ LSTP mutator_phase_inv \" +proof( vcg_jackhammer (no_thin_post_inv) + , simp_all add: mutator_phase_inv_aux_case split: hs_phase.splits + , vcg_name_cases) + case alloc then show ?case + apply (drule_tac x=m in spec) + apply (drule handshake_phase_invD) + apply (clarsimp simp: fA_rel_inv_def fM_rel_inv_def fM_rel_def hp_step_rel_def split: if_split_asm) + apply (intro conjI impI; simp) + apply (elim disjE; force simp: fA_rel_def) + apply (rule reachable_snapshot_inv_alloc, simp_all) + apply (elim disjE; force simp: fA_rel_def) + done +next case (store_ins s s') then show ?case + apply (drule_tac x=m in spec) + apply (drule handshake_phase_invD) + apply (intro conjI impI; 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_store_buffers (mutator m) s\))"; clarsimp) + apply (force simp: marked_deletions_def) + 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; fail) + apply (rule marked_deletions_store_ins; clarsimp) (* FIXME as above *) + apply (erule disjE; clarsimp) + apply (drule phase_rel_invD) + apply (clarsimp simp: phase_rel_def) + apply (elim disjE; clarsimp) + apply (fastforce simp: hp_step_rel_def) + apply (clarsimp simp: hp_step_rel_def) + apply (case_tac "sys_ghost_hs_phase s\"; clarsimp) (* FIXME invert handshake_phase_rel *) + apply (clarsimp simp: obj_at_field_on_heap_def split: option.splits) + apply (rule conjI, fast, clarsimp) + apply (frule_tac r=x2a in blackD(1)[OF reachable_blackD], simp_all)[1] + apply (rule_tac x="mut_tmp_ref s\" in reachable_points_to; auto simp: ran_def split: obj_at_splits; fail) + apply (clarsimp simp: obj_at_field_on_heap_def split: option.splits) + apply (rule conjI, fast, clarsimp) + apply (frule_tac r=x2a in blackD(1)[OF reachable_blackD], simp_all)[1] + apply (rule_tac x="mut_tmp_ref s\" in reachable_points_to; auto simp: ran_def split: obj_at_splits; fail) + apply (force simp: marked_deletions_def) + done +next case (hs_noop_done s s') then show ?case + apply - + apply (drule_tac x=m in spec) + apply (drule handshake_phase_invD) + apply (simp add: fA_rel_def fM_rel_def hp_step_rel_def) + apply (cases "mut_ghost_hs_phase s\") (* FIXME invert handshake_step *) + apply auto + done +next case (hs_get_roots_done s s') then show ?case + apply - + apply (drule_tac x=m in spec) + apply (drule handshake_phase_invD) + apply (force simp: hp_step_rel_def reachable_snapshot_inv_hs_get_roots_done) + done +next case (hs_get_work_done s s') then show ?case + apply (drule_tac x=m in spec) + apply (drule handshake_phase_invD) + apply (force simp add: hp_step_rel_def reachable_snapshot_inv_hs_get_work_done) + done +qed + +end + +lemma (in mut_m') mutator_phase_inv[intro]: + notes mut_m.mark_object_invL_def[inv] + notes mut_m.handshake_invL_def[inv] + notes fun_upd_apply[simp] + 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 \" +proof( vcg_jackhammer (no_thin_post_inv) + , simp_all add: mutator_phase_inv_aux_case split: hs_phase.splits + , vcg_name_cases) + case (alloc s s' rb) then show ?case + apply - + apply (clarsimp simp: fA_rel_inv_def fM_rel_inv_def white_def) + apply (drule spec[where x=m]) + apply (intro conjI impI; clarsimp) + apply (clarsimp simp: hp_step_rel_def simp: fA_rel_def fM_rel_def dest!: handshake_phase_invD) + apply (elim disjE, auto; fail) + apply (rule reachable_snapshot_inv_alloc, simp_all) + apply (clarsimp simp: hp_step_rel_def simp: fA_rel_def fM_rel_def dest!: handshake_phase_invD) + apply (cases "sys_ghost_hs_phase s\"; clarsimp; blast) + done +next case (hs_get_roots_done s s') then show ?case + apply - + apply (drule spec[where x=m]) + apply (simp add: no_black_refs_def reachable_snapshot_inv_def in_snapshot_def) + done +next case (hs_get_work_done s s') then show ?case + apply - + 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 +qed + +(* FIXME Some of \mutator_phase_inv\, the rest in Global Noninterference *) + +lemma 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) +unfolding no_black_refs_def by simp -lemma (in gc) no_black_refs_load_W[simp]: +lemma 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) +unfolding no_black_refs_def by simp + +lemma marked_insertions_sweep_loop_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)\))" +unfolding mut_m.marked_insertions_def by (fastforce simp: fun_upd_apply split: mem_store_action.splits obj_at_splits option.splits) lemma marked_deletions_sweep_loop_free[simp]: + notes fun_upd_apply[simp] + shows "\ 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) +unfolding mut_m.marked_deletions_def +apply (clarsimp split: mem_store_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 mp, auto simp: mut_m.reachable_def mut_m.tso_store_refs_def split: mem_store_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 (clarsimp; fail) +apply (rule conjI; 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 mp, auto simp: mut_m.reachable_def mut_m.tso_store_refs_def split: mem_store_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 *) - +unfolding white_def apply (clarsimp split: obj_at_splits) 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 +context gc 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 obj_fields_marked_inv_blacken: + "\ gc_field_set s = {}; obj_fields_marked s; (gc_tmp_ref s points_to w) s; white w s \ \ False" +by (simp add: obj_fields_marked_def obj_at_field_on_heap_def ran_def white_def split: option.splits obj_at_splits) -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 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 s; valid_W_inv s \ \ w = gc_tmp_ref s" +by (metis (mono_tags, lifting) converse_rtranclpE gc.obj_fields_marked_inv_blacken has_white_path_to_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) +lemma mutator_phase_inv[intro]: + notes fun_upd_apply[simp] + shows + "\ 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>\ valid_refs_inv \<^bold>\ valid_W_inv) \ + gc + \ LSTP (mut_m.mutator_phase_inv m) \" +proof( vcg_jackhammer (no_thin_post_inv) + , simp_all add: mutator_phase_inv_aux_case white_def split: hs_phase.splits + , vcg_name_cases ) + case (sweep_loop_free s s') then show ?case + apply (intro allI conjI impI) + apply (drule mut_m.handshake_phase_invD[where m=m], clarsimp simp: hp_step_rel_def; fail) + apply (rule mut_m.reachable_snapshot_inv_sweep_loop_free, simp_all add: white_def) + done +next case (mark_loop_get_work_load_W s s') then show ?case + 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 +next case (mark_loop_blacken s s') then show ?case + apply - + apply (drule spec[where x=m]) + apply clarsimp + apply (intro allI conjI impI; 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 (metis (no_types, hide_lams) obj_fields_marked_inv_has_white_path_to_blacken) + done +next case (mark_loop_mo_co_mark s s' y) then show ?case by (clarsimp simp: handshake_in_syncD mut_m.reachable_snapshot_inv_mo_co_mark) +next case (mark_loop_get_roots_load_W s s') then show ?case + 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 +qed 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] +lemma (in gc) strong_tricolour_inv[intro]: + notes fun_upd_apply[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) \ + \<^bold>\ LSTP (strong_tricolour_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 +unfolding strong_tricolour_inv_def +proof(vcg_jackhammer (no_thin_post_inv), vcg_name_cases) + case (mark_loop_blacken s s' x xa) then show ?case by (fastforce elim!: obj_fields_marked_inv_blacken) +qed lemma (in mut_m) strong_tricolour[intro]: + notes fun_upd_apply[simp] + shows "\ 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) - +unfolding strong_tricolour_inv_def +proof(vcg_jackhammer (no_thin_post_inv), vcg_name_cases) + case (alloc s s' x xa rb) then show ?case +apply (clarsimp simp: fA_rel_inv_def fM_rel_inv_def) apply (drule handshake_phase_invD) +apply (drule spec[where x=m]) apply (clarsimp simp: sys_phase_inv_aux_case - split: handshake_phase.splits if_splits) + split: hs_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 *) + (* FIXME rule? *) + apply (metis (no_types, lifting) black_def no_black_refsD obj_at_cong option.simps(3)) + apply (metis (no_types, lifting) black_def no_black_refsD obj_at_cong option.distinct(1)) - 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,203 +1,90 @@ (*<*) (* * 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 + Global_Invariants_Lemmas + Local_Invariants_Lemmas + Tactics 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) +section\ Coarse TSO invariants \ -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 +context gc +begin -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 tso_lock_invL[intro]: + "\ tso_lock_invL \ gc" +by vcg_jackhammer -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 tso_store_inv[intro]: + "\ LSTP tso_store_inv \ gc" +unfolding tso_store_inv_def by vcg_jackhammer -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 mut_tso_lock_invL[intro]: + "\ mut_m.tso_lock_invL m \ gc" +by (vcg_chainsaw mut_m.tso_lock_invL_def) -lemma (in sys) tso_gc_writes_inv[intro]: - "\ LSTP tso_writes_inv \ sys" -apply (vcg_jackhammer simp: tso_writes_inv_def) +end + +context mut_m +begin + +lemma tso_store_inv[intro]: + notes fun_upd_apply[simp] + shows + "\ LSTP tso_store_inv \ mutator m" +unfolding tso_store_inv_def by vcg_jackhammer + +lemma gc_tso_lock_invL[intro]: + "\ gc.tso_lock_invL \ mutator m" +by (vcg_chainsaw gc.tso_lock_invL_def) + +lemma tso_lock_invL[intro]: + "\ tso_lock_invL \ mutator m" +by vcg_jackhammer + +end + +context mut_m' +begin + +lemma tso_lock_invL[intro]: + "\ tso_lock_invL \ mutator m'" +by (vcg_chainsaw tso_lock_invL) + +end + +context sys +begin + +lemma tso_gc_store_inv[intro]: + notes fun_upd_apply[simp] + shows + "\ LSTP tso_store_inv \ sys" +apply (vcg_chainsaw tso_store_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 gc_tso_lock_invL[intro]: + "\ gc.tso_lock_invL \ sys" +by (vcg_chainsaw gc.tso_lock_invL_def) -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)" +lemma mut_tso_lock_invL[intro]: + "\ mut_m.tso_lock_invL m \ sys" +by (vcg_chainsaw mut_m.tso_lock_invL_def) -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 (*<*) end (*>*) diff --git a/thys/ConcurrentGC/Tactics.thy b/thys/ConcurrentGC/Tactics.thy --- a/thys/ConcurrentGC/Tactics.thy +++ b/thys/ConcurrentGC/Tactics.thy @@ -1,309 +1,504 @@ (*<*) (* * 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" + Proofs_Basis begin -declare subst_all [simp del] [[simproc del: defined_all]] - (*>*) -section\Invariants and Proofs \label{sec:gc-invs}\ - -subsection\Constructors for sets of locations.\ +section\ CIMP specialisation \label{sec:cimp-specialisation} \ -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\ +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" ("\_\ _ \_\") + valid_proc :: "('field, 'mut, 'payload, 'ref) gc_pred \ 'mut process_name \ ('field, 'mut, 'payload, 'ref) gc_pred \ bool" ("\_\ _ \_\") where - "\P\ p \Q\ \ \(c, afts) \ vcg_fragments (gc_pgms p). (gc_pgms, p, afts \ \P\ c \Q\)" + "\P\ p \Q\ = (\(c, afts) \ vcg_fragments (gc_coms p). gc_coms, p, afts \ \P\ c \Q\)" abbreviation - valid_proc_inv_syn :: "('field, 'mut, 'ref) gc_pred \ 'mut process_name \ bool" ("\_\ _" [100,0] 100) + valid_proc_inv_syn :: "('field, 'mut, 'payload, '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] +subsection\ Tactics \ -(* 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) +subsubsection\ Model-specific \ 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; + "\ (\l\ Response action, afts) \ fragments (gc_coms p) {}; v \ action x s; \ p = sys; ?P \ \ Q \ \ Q" apply (cases p) -apply (simp_all add: gc_all_defs) +apply (simp_all add: all_com_interned_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 + "\ (\l\ Response action, afts) \ fragments (gc_coms p) {}; v \ action (pname, req) s; + \ p = sys; case_request_op ?P1 ?P2 ?P3 ?P4 ?P5 ?P6 ?P7 ?P8 ?P9 ?P10 ?P11 ?P12 ?P13 ?P14 req \ \ Q \ \ Q" +apply (erule(1) system_responds_actionE) +apply (cases req; simp only: request_op.simps prod.inject simp_thms fst_conv snd_conv if_cancel empty_def[symmetric] empty_iff) +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 + "\ (\l\ Response action, afts) \ fragments (gc_coms p) {}; v \ action x s; + \ p = sys; case_request_op ?P1 ?P2 ?P3 ?P4 ?P5 ?P6 ?P7 ?P8 ?P9 ?P10 ?P11 ?P12 ?P13 ?P14 (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 *) + +subsubsection \ Locations\ + +(* FIXME not automation friendly but used in some non-interference proofs *) +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) + +lemma schematic_prem: "\Q \ P; Q\ \ P" +by blast + +(* One way of instantiating a schematic prem. *) +lemma TrueE: "\True; P\ \ P" +by blast + +lemma thin_locs_pre_discardE: + "\at p l' s \ P; at p l s; l' \ l; Q\ \ Q" + "\atS p ls s \ P; at p l s; l \ ls; Q\ \ Q" +unfolding atS_def by blast+ + +lemma thin_locs_pre_keep_atE: + "\at p l s \ P; at p l s; P \ Q\ \ Q" +by blast + +lemma thin_locs_pre_keep_atSE: + "\atS p ls s \ P; at p l s; l \ ls; P \ Q\ \ Q" +unfolding atS_def by blast + +(* FIXME complete with symmetric rules on process names -- but probably not needed here *) +lemma thin_locs_post_discardE: + "\AT s' = (AT s)(p := lfn, q := lfn'); l' \ lfn; p \ q\ \ at p l' s' \ P" + "\AT s' = (AT s)(p := lfn); l' \ lfn\ \ at p l' s' \ P" + "\AT s' = (AT s)(p := lfn, q := lfn'); \l. l \ lfn \ l \ ls; p \ q\ \ atS p ls s' \ P" + "\AT s' = (AT s)(p := lfn); \l. l \ lfn \ l \ ls\ \ atS p ls s' \ P" +unfolding atS_def by (auto simp: fun_upd_apply) + +lemmas thin_locs_post_discard_conjE = + conjI[OF thin_locs_post_discardE(1)] + conjI[OF thin_locs_post_discardE(2)] + conjI[OF thin_locs_post_discardE(3)] + conjI[OF thin_locs_post_discardE(4)] + +lemma thin_locs_post_keep_locsE: + "\(L \ P) \ R; R \ Q\ \ (L \ P) \ Q" + "L \ P \ L \ P" +by blast+ + +lemma thin_locs_post_keepE: + "\P \ R; R \ Q\ \ (L \ P) \ Q" + "P \ L \ P" +by blast+ + +(* FIXME checking the fun_upds are irrelevant is not necessary, but ensures the keep rule applies. *) +(* FIXME consider atS (mutator m) mut_hs_get_roots_loop_locs s' -- same again but replace at proc l s' with atS proc ls s' *) +(* FIXME in general we'd need to consider intersections *) +lemma ni_thin_locs_discardE: + "\at proc l s \ P; AT s' = (AT s)(p := lfn, q := lfn'); at proc l' s'; l \ l'; proc \ p; proc \ q; Q\ \ Q" + "\at proc l s \ P; AT s' = (AT s)(p := lfn); at proc l' s'; l \ l'; proc \ p; Q\ \ Q" + "\atS proc ls s \ P; AT s' = (AT s)(p := lfn, q := lfn'); at proc l' s'; l' \ ls; proc \ p; proc \ q; Q\ \ Q" + "\atS proc ls s \ P; AT s' = (AT s)(p := lfn); at proc l' s'; l' \ ls; proc \ p; Q\ \ Q" + + "\at proc l s \ P; AT s' = (AT s)(p := lfn, q := lfn'); atS proc ls' s'; l \ ls'; proc \ p; proc \ q; Q\ \ Q" + "\at proc l s \ P; AT s' = (AT s)(p := lfn); atS proc ls' s'; l \ ls'; proc \ p; Q\ \ Q" +(* + "\atS proc ls s \ P; AT s' = (AT s)(p := lfn, q := lfn'); atS proc l s'; l \ ls; proc \ p; proc \ q; Q\ \ Q" + "\atS proc ls s \ P; AT s' = (AT s)(p := lfn); atS proc ls' s'; l \ ls; proc \ p; Q\ \ Q" +*) +unfolding atS_def by auto + +lemma ni_thin_locs_keep_atE: + "\at proc l s \ P; AT s' = (AT s)(p := lfn, q := lfn'); at proc l s'; proc \ p; proc \ q; P \ Q\ \ Q" + "\at proc l s \ P; AT s' = (AT s)(p := lfn); at proc l s'; proc \ p; P \ Q\ \ Q" +by (auto simp: fun_upd_apply) + +lemma ni_thin_locs_keep_atSE: + "\atS proc ls s \ P; AT s' = (AT s)(p := lfn, q := lfn'); at proc l' s'; l' \ ls; proc \ p; proc \ q; P \ Q\ \ Q" + "\atS proc ls s \ P; AT s' = (AT s)(p := lfn); at proc l' s'; l' \ ls; proc \ p; P \ Q\ \ Q" + "\atS proc ls s \ P; AT s' = (AT s)(p := lfn, q := lfn'); atS proc ls' s'; ls' \ ls; proc \ p; proc \ q; P \ Q\ \ Q" + "\atS proc ls s \ P; AT s' = (AT s)(p := lfn); atS proc ls' s'; ls' \ ls; proc \ p; P \ Q\ \ Q" +unfolding atS_def by (auto simp: fun_upd_apply) + +lemma loc_mem_tac_intros: + "\c \ A; c \ B\ \ c \ A \ B" + "c \ d \ c \ {d}" + "c \ A \ c \ - A" + "c \ A \ c \ - A" + "A \ A" +by blast+ + +(* FIXME these rules are dangerous if any other sets are lying around. Specialise the types? *) +lemmas loc_mem_tac_elims = + singletonE + UnE + +lemmas loc_mem_tac_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 + simp_thms + Eq_FalseI + not_Cons_self lemmas vcg_fragments'_simps = - valid_proc_def gc_pgms.simps vcg_fragments'.simps atC.simps + valid_proc_def gc_coms.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 +lemmas vcg_sem_simps = + lconst.simps simp_thms + True_implies_equals + prod.simps fst_conv snd_conv + gc_phase.simps process_name.simps hs_type.simps hs_phase.simps + mem_store_action.simps mem_load_action.simps request_op.simps response.simps -(* 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] +lemmas vcg_inv_simps = + simp_thms ML \ -(* Thanks BNF people. *) -fun ss_only thms ctxt = clear_simpset (put_simpset HOL_basic_ss ctxt) addsimps thms; +signature GC_VCG = +sig + (* Internals *) + val nuke_schematic_prems : Proof.context -> int -> tactic + val loc_mem_tac : Proof.context -> int -> tactic + val vcg_fragments_tac : Proof.context -> int -> tactic + val vcg_sem_tac : Proof.context -> int -> tactic + val thin_pre_inv_tac : Proof.context -> int -> tactic + val thin_post_inv_tac : bool -> Proof.context -> int -> tactic + val vcg_inv_tac : bool -> bool -> Proof.context -> int -> tactic + (* End-user tactics *) + val vcg_jackhammer_tac : bool -> bool -> Proof.context -> int -> tactic + val vcg_chainsaw_tac : bool -> thm list -> Proof.context -> int -> tactic + val vcg_name_cases_tac : term list -> thm list -> context_tactic +end -(* 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; +structure GC_VCG : GC_VCG = +struct -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) +(* Identify and remove schematic premises. FIXME reverses the prems *) +fun nuke_schematic_prems ctxt = + let + fun is_schematic_prem t = + case t of + Const ("HOL.Trueprop", _) $ t => is_schematic_prem t + | t $ _ => is_schematic_prem t + | Var _ => true + | _ => false + in + DETERM o filter_prems_tac ctxt (not o is_schematic_prem) + end; + +(* FIXME Want to unify only with a non-schematic prem. might get there with first order matching or some existing variant of assume. *) +fun assume_non_schematic_prem_tac ctxt = + (TRY o nuke_schematic_prems ctxt) THEN' assume_tac ctxt + +fun vcg_fragments_tac ctxt = + SELECT_GOAL (HEADGOAL (safe_simp_tac (ss_only (@{thms vcg_fragments'_simps} @ @{thms all_com_interned_defs}) ctxt) + THEN' SELECT_GOAL (safe_tac ctxt))); (* FIXME split the goal, simplify the sets away. FIXME try to nuke safe_tac *) + +fun vcg_sem_tac ctxt = + SELECT_GOAL (HEADGOAL (match_tac ctxt @{thms CIMP_vcg.vcg.intros} + THEN' (TRY o (ematch_tac ctxt @{thms system_responds_action_specE} THEN' assume_tac ctxt)) + THEN' Rule_Insts.thin_tac ctxt "HST s = h" [(@{binding s}, NONE, NoSyn), (@{binding h}, NONE, NoSyn)] (* discard history: we don't use it here *) + THEN' clarsimp_tac (ss_only @{thms vcg_sem_simps} ctxt) + THEN_ALL_NEW asm_simp_tac ctxt)); (* remove unused meta-bound vars FIXME subgoaler in HOL's usual simplifier setup, somehow lost by ss_only *) + +(* FIXME gingerly settle location set membership and (dis-)equalities *) +fun loc_mem_tac ctxt = + let + val loc_defs = Proof_Context.get_fact ctxt (Facts.named "loc_defs") + in + SELECT_GOAL (HEADGOAL ( (TRY o REPEAT_ALL_NEW (ematch_tac ctxt @{thms loc_mem_tac_elims})) + THEN_ALL_NEW (TRY o hyp_subst_tac ctxt) + THEN_ALL_NEW (TRY o REPEAT_ALL_NEW (match_tac ctxt @{thms loc_mem_tac_intros})) + THEN_ALL_NEW ( SOLVED' (match_tac ctxt (Named_Theorems.get ctxt \<^named_theorems>\locset_cache\) + ORELSE' safe_simp_tac (HOL_ss_only (@{thms loc_mem_tac_simps} @ loc_defs) ctxt) ) ))) + end; + +fun thin_pre_inv_tac ctxt = + SELECT_GOAL (HEADGOAL ( (* FIXME trying to scope the REPEAT_DETERM ala [1] *) + (REPEAT_DETERM o ematch_tac ctxt @{thms conjE}) + THEN' (REPEAT_DETERM o ( (ematch_tac ctxt @{thms thin_locs_pre_discardE} THEN' assume_tac ctxt THEN' loc_mem_tac ctxt) + ORELSE' (ematch_tac ctxt @{thms thin_locs_pre_keep_atE} THEN' assume_tac ctxt) + ORELSE' (ematch_tac ctxt @{thms thin_locs_pre_keep_atSE} THEN' assume_tac ctxt THEN' loc_mem_tac ctxt) )))); + +(* FIXME redo keep_postE: if at loc is provable, discard the at antecedent, otherwise keep it *) +(* if the post inv is an LSTP then the present fix is to say (no_thin_post_inv) -- would be good to automate that *) +fun thin_post_inv_tac keep_locs ctxt = + let + val keep_postE_thms = if keep_locs then @{thms thin_locs_post_keep_locsE} else @{thms thin_locs_post_keepE} + fun nail_discard_prems_tac ctxt = assume_non_schematic_prem_tac ctxt THEN' loc_mem_tac ctxt THEN' (TRY o match_tac ctxt @{thms process_name.simps}) + in + SELECT_GOAL (HEADGOAL ( (* FIXME trying to scope the REPEAT_DETERM ala [1] *) + resolve_tac ctxt @{thms schematic_prem} + THEN' REPEAT_DETERM o CHANGED o (* FIXME CHANGED? also check what happens for non-invL post invs -- aim to fail the ^^^ resolve_tac too *) + ( ( match_tac ctxt @{thms thin_locs_post_discard_conjE} THEN' nail_discard_prems_tac ctxt) + ORELSE' (eresolve_tac ctxt @{thms TrueE} THEN' match_tac ctxt @{thms thin_locs_post_discardE} THEN' nail_discard_prems_tac ctxt) + ORELSE' eresolve_tac ctxt keep_postE_thms ) + )) + end; + +fun vcg_inv_tac keep_locs no_thin_post_inv ctxt = + let + val invs = Named_Theorems.get ctxt \<^named_theorems>\inv\ + in + SELECT_GOAL (Local_Defs.unfold_tac ctxt invs) (* FIXME trying to say unfold in [1] only *) + THEN' thin_pre_inv_tac ctxt + THEN' ( if no_thin_post_inv + then SELECT_GOAL all_tac (* full_simp_tac (ss_only @{thms vcg_inv_simps} ctxt) (* FIXME maybe not? *) *) + else full_simp_tac (Splitter.add_split @{thm lcond_split_asm} (ss_only @{thms vcg_inv_simps} ctxt)) + THEN_ALL_NEW thin_post_inv_tac keep_locs ctxt ) + end; + +(* For showing local invariants. FIXME tack on (no_thin_post_inv) for universal/LSTP ones *) +fun vcg_jackhammer_tac keep_locs no_thin_post_inv ctxt = + SELECT_GOAL (HEADGOAL (vcg_fragments_tac ctxt) + THEN PARALLEL_ALLGOALS ( + vcg_sem_tac ctxt + THEN_ALL_NEW vcg_inv_tac keep_locs no_thin_post_inv ctxt + THEN_ALL_NEW (if keep_locs then SELECT_GOAL all_tac else Rule_Insts.thin_tac ctxt "AT _ = _" []) + THEN_ALL_NEW TRY o clarsimp_tac ctxt (* limply try to solve the remaining goals *) + )); + +(* For showing noninterference *) +fun vcg_chainsaw_tac no_thin unfold_foreign_inv_thms ctxt = + let + fun specialize_foreign_invs_tac ctxt = + (* FIXME split goal: makes sense because local procs control locs have changed? *) + REPEAT_ALL_NEW (match_tac ctxt @{thms conjI}) + THEN_ALL_NEW TRY o ( match_tac ctxt @{thms impI} (* FIXME could tweak rules vvvv *) + (* thin out the invariant we're showing non-interference for *) +(* FIXME look for reasons to retain the invariant, then do a big thin_tac at the end. +Intuitively we don't have enough info to settle atS v atS questions and it's too hard/not informative enough to try. +Let the user do it. +Maybe add an info thing that tells what was thinned. + +FIXME location-sensitive predicates are not amenable to +simplification: this is the cost of using projections on +pred_state. Instead use elimination rules \nie\. + + *) + THEN' ( REPEAT_DETERM o ( ( ematch_tac ctxt @{thms ni_thin_locs_discardE} THEN' assume_tac ctxt THEN' assume_tac ctxt THEN' loc_mem_tac ctxt THEN' match_tac ctxt @{thms process_name.simps} THEN' TRY o match_tac ctxt @{thms process_name.simps} ) + ORELSE' ( ematch_tac ctxt @{thms ni_thin_locs_keep_atE} THEN' assume_tac ctxt THEN' assume_tac ctxt THEN' match_tac ctxt @{thms process_name.simps} THEN' TRY o match_tac ctxt @{thms process_name.simps} ) + ORELSE' ( ematch_tac ctxt @{thms ni_thin_locs_keep_atSE} THEN' assume_tac ctxt THEN' assume_tac ctxt THEN' loc_mem_tac ctxt THEN' match_tac ctxt @{thms process_name.simps} THEN' TRY o match_tac ctxt @{thms process_name.simps} ) ) ) ) + in + SELECT_GOAL (HEADGOAL (vcg_fragments_tac ctxt) + THEN PARALLEL_ALLGOALS ( + vcg_sem_tac ctxt + (* nail cheaply with an nie fact + ambient clarsimp *) + THEN_ALL_NEW ( SOLVED' (ematch_tac ctxt (Named_Theorems.get ctxt @{named_theorems nie}) THEN_ALL_NEW clarsimp_tac ctxt) + ORELSE' ( (* do it the hard way: specialise any process-specific invariants. Similar to vcg_jackhammer but not the same *) + vcg_inv_tac false true ctxt + (* unfold the foreign inv *) + THEN' SELECT_GOAL (Local_Defs.unfold_tac ctxt unfold_foreign_inv_thms) (* FIXME trying to say [1] *) + THEN' (REPEAT_DETERM o ematch_tac ctxt @{thms conjE}) + THEN' specialize_foreign_invs_tac ctxt + (* limply try to solve the remaining goals; FIXME turn s' into s as much as easily possible *) + THEN_ALL_NEW (TRY o clarsimp_tac ctxt) + (* FIXME discard loc info. It is useful, this is a stopgap *) + THEN_ALL_NEW (if no_thin then SELECT_GOAL all_tac + else (Rule_Insts.thin_tac ctxt "AT _ = _" [] + THEN_ALL_NEW (REPEAT_DETERM o Rule_Insts.thin_tac ctxt "at _ _ _ \ _" []) + THEN_ALL_NEW (REPEAT_DETERM o Rule_Insts.thin_tac ctxt "atS _ _ _ \ _" []) )) + )))) + end; + +(* + +Scrutinise the goal state for an `AT` fact that tells us which label the process is at. + +It seems this is not kosher: + - for reasons unknown (Eisbach?) this tactic gets called with a bogus "TERM _" and then the real goal state. + - this tactic (sometimes) does not work if used with THEN_ALL_NEW ';' -- + chk_label does not manage to uniquify labels -- so be sure to + combine with ','. + - if two goals have the same `at` location then we disambiguate, but + perhaps not stably. Could imagine creating subcases, but + `Method.goal_cases_tac` is not yet capable of that. + - at communication steps we could get unlucky and choose the label from the other process. + +The user can supply a list of process names to somewhat address these issues. + +See Pure/Tools/rule_insts.ML for structurally similar tactics ("dynamic instantiation"). + +Limitations: + - only works with `Const` labels + - brittle: assumes things are very well-formed + +*) +fun vcg_name_cases_tac (proc_names: term list) _(*facts*) (ctxt, st) = + if Thm.nprems_of st = 0 + then Seq.empty (* no_tac *) + else + let + fun fst_ord ord ((x, _), (x', _)) = ord (x, x') + fun snd_ord ord ((_, y), (_, y')) = ord (y, y') + + (* FIXME this search is drecky *) + fun find_AT (t: term) : (term * string) option = + ( (* tracing ("scruting: " ^ Syntax.string_of_term ctxt t) ; *) + case t of Const ("HOL.Trueprop", _) $ (Const (@{const_name "Set.member"}, _) $ Const (l, _) $ (Const (@{const_name "CIMP_lang.AT"}, _) $ _ $ p)) => ((* tracing "HIT"; *) SOME (p, Long_Name.base_name l)) + | Const ("HOL.Trueprop", _) $ (Const (@{const_name "CIMP_lang.atS"}, _) $ p $ Const (ls, _) $ _) => ((* tracing "HIT"; *) SOME (p, Long_Name.base_name ls)) + | _ => NONE ) + + (* FIXME Isabelle makes it awkward to use polymorphic process names; paper over that crack here *) + val rec terms_eq_ignoring_types = + fn (Const (c0, _), Const (c1, _)) => fast_string_ord (c0, c1) = EQUAL + | (Free (f0, _), Free (f1, _)) => fast_string_ord (f0, f1) = EQUAL + | (Var (v0, _) , Var (v1, _)) => prod_ord fast_string_ord int_ord (v0, v1) = EQUAL + | (Bound i0, Bound i1) => i0 = i1 + | (Abs (b0, _, t0), Abs (b1, _, t1)) => fast_string_ord (b0, b1) = EQUAL andalso terms_eq_ignoring_types (t0, t1) + | (t0 $ u0, t1 $ u1) => terms_eq_ignoring_types (t0, t1) andalso terms_eq_ignoring_types (u0, u1) + | _ => false + + fun mk_label (default: string) (ats : (term * string) list) : string = + case (ats, proc_names) of + ((_, l)::_, []) => ((* tracing ("No proc_names, Using label: " ^ l); *) l) + | _ => + let + val ls = List.mapPartial (fn p => List.find (fn (p', _) => terms_eq_ignoring_types (p, p')) ats) proc_names + in + case ls of + [] => (warning ("vcg_name_cases: using the default name: " ^ default); default) + | _ => ls |> List.map snd |> String.concatWith "_" + end + + fun labels_for_cases (i: int) (acc: (int * string) list) : (int * string) list = + case i of + 0 => acc + | i => Thm.cprem_of st i |> Thm.term_of |> Logic.strip_assums_hyp + |> List.mapPartial find_AT |> mk_label (Int.toString i) + |> (fn l => labels_for_cases (i - 1) ((i, l) :: acc)) + + (* Make the labels unique if need be *) + fun uniquify (i: int) (ls: (int * string) list) : (int * string) list = + case ls of + [] => [] + | [l] => [l] + | l :: l' :: ls => (case fast_string_ord (snd l, snd l') of + EQUAL => (fst l, snd l ^ Int.toString i) :: uniquify (i + 1) (l' :: ls) + | _ => l :: uniquify 0 (l' :: ls)) + val labels = labels_for_cases (Thm.nprems_of st) [] + val labels = labels + |> sort (snd_ord fast_string_ord) |> uniquify 0 |> sort (fst_ord int_ord) + |> List.map (fn (_, l) => ((* tracing ("label: " ^ l); *) l)) + in + Method.goal_cases_tac labels (ctxt, st) + end; + +end 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") + Theory.setup (Method.setup @{binding loc_mem} + (Scan.succeed (SIMPLE_METHOD' o GC_VCG.loc_mem_tac)) + "solve location membership 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_fragments} + (Scan.succeed (SIMPLE_METHOD' o GC_VCG.vcg_fragments_tac)) + "unfold com defs, execute vcg_fragments' and split goals") 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") + (Scan.succeed (SIMPLE_METHOD' o GC_VCG.vcg_sem_tac)) + "reduce VCG goal to 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_inv} + (Scan.lift (Args.mode "keep_locs" -- Args.mode "no_thin_post_inv") >> (fn (keep_locs, no_thin_post_inv) => SIMPLE_METHOD' o GC_VCG.vcg_inv_tac keep_locs no_thin_post_inv)) + "specialise the invariants to the goal. (keep_locs) retains locs in post conds") 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") + (Scan.lift (Args.mode "keep_locs" -- Args.mode "no_thin_post_inv") >> (fn (keep_locs, no_thin_post_inv) => SIMPLE_METHOD' o GC_VCG.vcg_jackhammer_tac keep_locs no_thin_post_inv)) + "VCG tactic. (keep_locs) retains locs in post conds. (no_thin_post_inv) does not attempt to specalise the post condition.") 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 *) + Theory.setup (Method.setup @{binding vcg_chainsaw} + (Scan.lift (Args.mode "no_thin") -- Attrib.thms >> (fn (no_thin, thms) => SIMPLE_METHOD' o GC_VCG.vcg_chainsaw_tac no_thin thms)) + "VCG non-interference tactic. Tell it how to unfold the foreign local invs.") 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") + Theory.setup (Method.setup @{binding vcg_name_cases} + (Scan.repeat Args.term >> (fn proc_names => fn _ => CONTEXT_METHOD (GC_VCG.vcg_name_cases_tac proc_names))) + "divine canonical case names for outstanding VCG goals") -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/Valid_Refs.thy b/thys/ConcurrentGC/Valid_Refs.thy new file mode 100644 --- /dev/null +++ b/thys/ConcurrentGC/Valid_Refs.thy @@ -0,0 +1,130 @@ +(*<*) +(* + * 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 Valid_Refs +imports + Global_Invariants_Lemmas + Local_Invariants_Lemmas + Tactics +begin + +(*>*) +section\ Valid refs inv proofs \ + +lemma valid_refs_inv_sweep_loop_free: + 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 unfolding valid_refs_inv_def grey_reachable_def no_grey_refs_def +apply (clarsimp 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]: + notes fun_upd_apply[simp] + shows + "\ 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 \" +proof(vcg_jackhammer (no_thin_post_inv), vcg_name_cases) + case (sweep_loop_free s s') then show ?case +apply - +apply (drule (1) handshake_in_syncD) +apply (rule valid_refs_inv_sweep_loop_free, assumption, assumption) + apply (simp; fail) +apply (simp add: white_def) (* FIXME rule? *) +done +qed (auto simp: valid_refs_inv_def grey_reachable_def) + +context mut_m +begin + +lemma valid_refs_inv_discard_roots: + "\ valid_refs_inv s; roots' \ mut_roots s \ + \ valid_refs_inv (s(mutator m := s (mutator m)\roots := roots'\))" +unfolding valid_refs_inv_def mut_m.reachable_def by (auto simp: fun_upd_apply) + +lemma valid_refs_inv_load: + "\ valid_refs_inv s; sys_load (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'\))" +unfolding valid_refs_inv_def by (simp add: fun_upd_apply) + +lemma valid_refs_inv_alloc: + "\ 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, obj_payload = Map.empty\)\))" +unfolding valid_refs_inv_def mut_m.reachable_def +apply (clarsimp simp: fun_upd_apply) +apply (auto elim: converse_reachesE split: obj_at_splits) +done + +lemma valid_refs_inv_store_ins: + "\ 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_store_buffers := (mem_store_buffers (s sys))(mutator m := sys_mem_store_buffers (mutator m) s @ [mw_Mutate r f opt_r']) \))" +apply (subst valid_refs_inv_def) +apply (auto simp: grey_reachable_def mut_m.reachable_def fun_upd_apply) +(* FIXME what's gone wrong here? *) +apply (subst (asm) tso_store_refs_simps; force)+ +done + +lemma valid_refs_inv_deref_del: + "\ valid_refs_inv s; sys_load (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'\))" +unfolding valid_refs_inv_def by (simp add: fun_upd_apply) + +lemma valid_refs_inv_mo_co_mark: + "\ r \ mut_roots s \ mut_ghost_honorary_root s; mut_ghost_honorary_grey s = {}; valid_refs_inv s \ + \ valid_refs_inv (s(mutator m := s (mutator m)\ghost_honorary_grey := {r}\))" +unfolding valid_refs_inv_def +apply (clarsimp simp: grey_reachable_def fun_upd_apply) +apply (metis grey_reachable_def valid_refs_invD(1) valid_refs_invD(10) valid_refs_inv_def) +done + +lemma valid_refs_inv[intro]: + notes fun_upd_apply[simp] + notes valid_refs_inv_discard_roots[simp] + notes valid_refs_inv_load[simp] + notes valid_refs_inv_alloc[simp] + notes valid_refs_inv_store_ins[simp] + notes valid_refs_inv_deref_del[simp] + notes valid_refs_inv_mo_co_mark[simp] + shows + "\ 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 \" +proof(vcg_jackhammer (keep_locs) (no_thin_post_inv), vcg_name_cases) +next case (hs_get_roots_done s s') then show ?case by (clarsimp simp: valid_refs_inv_def grey_reachable_def) +next case (hs_get_work_done s s') then show ?case by (clarsimp simp: valid_refs_inv_def grey_reachable_def) +qed + +end + +lemma (in sys) valid_refs_inv[intro]: + "\ LSTP (valid_refs_inv \<^bold>\ tso_store_inv) \ sys \ LSTP valid_refs_inv \" +proof(vcg_jackhammer (no_thin_post_inv), vcg_name_cases) + case (tso_dequeue_store_buffer s s' p w ws) then show ?case + unfolding do_store_action_def + apply (auto simp: p_not_sys valid_refs_inv_dequeue_Mutate valid_refs_inv_dequeue_Mutate_Payload fun_upd_apply split: mem_store_action.splits) + done +qed + +(*<*) + +end +(*>*) diff --git a/thys/ConcurrentGC/Worklists.thy b/thys/ConcurrentGC/Worklists.thy new file mode 100644 --- /dev/null +++ b/thys/ConcurrentGC/Worklists.thy @@ -0,0 +1,230 @@ +(*<*) +(* + * 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 Worklists +imports + Global_Invariants_Lemmas + Local_Invariants_Lemmas + Tactics +begin + +(*>*) +section\Worklist invariants \label{sec:worklist-invariants} \ + +lemma valid_W_invD0: + "\ 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" +using marked_not_white unfolding valid_W_inv_def WL_def by (auto 0 5 split: obj_at_splits) + +lemma valid_W_distinct_simps: + "\r \ ghost_honorary_grey (s p); valid_W_inv s\ \ (r \ ghost_honorary_grey (s q)) \ (p = q)" + "\r \ W (s p); valid_W_inv s\ \ (r \ W (s q)) \ (p = q)" + "\r \ WL p s; valid_W_inv s\ \ (r \ WL q s) \ (p = q)" + using valid_W_invD0(4) apply fastforce + using valid_W_invD0(1) apply fastforce +apply (metis UnE WL_def valid_W_invD0(1) valid_W_invD0(4)) +done + +lemma valid_W_inv_sys_mem_store_buffersD: + "\ sys_mem_store_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_store_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_store_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_store_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]" +unfolding valid_W_inv_def white_def by (clarsimp dest!: spec[where x=p], blast)+ + +lemma valid_W_invE2: + "\ r \ W (s p); valid_W_inv s; \obj. obj_mark obj = sys_fM s \ P obj\ \ obj_at P r s" + "\ r \ ghost_honorary_grey (s p); sys_mem_lock s \ Some p; valid_W_inv s; \obj. obj_mark obj = sys_fM s \ P obj \ \ obj_at P r s" +unfolding valid_W_inv_def +apply (simp_all add: split: obj_at_splits) +apply blast+ +done + +lemma (in sys) valid_W_inv[intro]: + notes if_split_asm[split del] + notes fun_upd_apply[simp] + shows + "\ LSTP (fM_rel_inv \<^bold>\ sys_phase_inv \<^bold>\ tso_store_inv \<^bold>\ valid_refs_inv \<^bold>\ valid_W_inv) \ + sys + \ LSTP valid_W_inv \" +proof(vcg_jackhammer (no_thin_post_inv), vcg_name_cases) + case (tso_dequeue_store_buffer s s' p w ws) then show ?case + proof(cases w) + case (mw_Mark r fl) with tso_dequeue_store_buffer show ?thesis +apply (subst valid_W_inv_def) +apply clarsimp +apply (frule (1) valid_W_invD(1)) +apply (clarsimp simp: all_conj_distrib white_def valid_W_inv_sys_ghg_empty_iff filter_empty_conv obj_at_simps) +apply (intro allI conjI impI) +apply (auto elim: valid_W_invE2)[3] +apply (meson Int_emptyI valid_W_distinct_simps(3)) +apply (meson valid_W_invD0(2)) +apply (meson valid_W_invD0(2)) +using valid_W_invD(2) apply fastforce +apply auto[1] +using valid_W_invD(2) apply fastforce +done + next case (mw_fM fl) with tso_dequeue_store_buffer show ?thesis + apply (clarsimp simp: fM_rel_inv_def fM_rel_def p_not_sys) + apply (elim disjE; clarsimp) + apply (frule (1) no_grey_refs_no_pending_marks) + apply (subst valid_W_inv_def) + apply clarsimp + apply (meson Int_emptyI no_grey_refsD(1) no_grey_refsD(3) valid_W_distinct_simps(3) valid_W_invD(2) valid_W_inv_sys_ghg_empty_iff valid_W_inv_sys_mem_store_buffersD(3)) + done + qed simp_all +qed + +(* Lemmas for key mark_object transitions *) + +lemma valid_W_inv_ghg_disjoint: + "\ white y s; sys_mem_lock s = Some p; valid_W_inv s; p0 \ p1 \ + \ WL p0 (s(p := s p\ghost_honorary_grey := {y}\)) \ WL p1 (s(p := s p\ghost_honorary_grey := {y}\)) = {}" +unfolding valid_W_inv_def WL_def by (auto 5 5 simp: fun_upd_apply) + +lemma valid_W_inv_mo_co_mark: + "\ valid_W_inv s; white y s; sys_mem_lock s = Some p; filter is_mw_Mark (sys_mem_store_buffers p s) = []; p \ sys \ + \ valid_W_inv (s(p := s p\ghost_honorary_grey := {y}\, sys := s sys\mem_store_buffers := (mem_store_buffers (s sys))(p := sys_mem_store_buffers p s @ [mw_Mark y (sys_fM s)])\))" +apply (subst valid_W_inv_def) +apply (clarsimp simp: all_conj_distrib fun_upd_apply) +apply (intro allI conjI impI) +apply (auto simp: valid_W_invD valid_W_distinct_simps(3) valid_W_inv_sys_ghg_empty_iff valid_W_invD0 valid_W_inv_ghg_disjoint valid_W_inv_colours) +done + +lemma valid_W_inv_mo_co_lock: + "\ valid_W_inv s; sys_mem_lock s = None \ + \ valid_W_inv (s(sys := s sys\mem_lock := Some p\))" +by (auto simp: valid_W_inv_def fun_upd_apply) (* FIXME some eager rule expects valid_W_inv *) + +lemma valid_W_inv_mo_co_W: + "\ valid_W_inv s; marked y s; ghost_honorary_grey (s p) = {y}; p \ sys \ + \ valid_W_inv (s(p := s p\W := insert y (W (s p)), ghost_honorary_grey := {}\))" +apply (subst valid_W_inv_def) +apply (clarsimp simp: all_conj_distrib valid_W_invD0(2) fun_upd_apply) +apply (intro allI conjI impI) +apply (auto simp: valid_W_invD valid_W_invD0(2) valid_W_distinct_simps(3)) + using valid_W_distinct_simps(1) apply fastforce +apply (metis marked_not_white singletonD valid_W_invD(2)) +done + +lemma valid_W_inv_mo_co_unlock: + "\ sys_mem_lock s = Some p; sys_mem_store_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 simp: fun_upd_apply) (metis emptyE empty_set) + +lemma (in gc) valid_W_inv[intro]: + notes if_split_asm[split del] + notes fun_upd_apply[simp] + shows + "\ 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 valid_W_inv \ + gc + \ LSTP valid_W_inv \" +proof(vcg_jackhammer (no_thin_post_inv), vcg_name_cases) + case (sweep_loop_free s s') then show ?case + apply (subst valid_W_inv_def) + apply (clarsimp simp: all_conj_distrib white_def valid_W_inv_sys_ghg_empty_iff) + apply (meson disjoint_iff_not_equal no_grey_refsD(1) no_grey_refsD(2) no_grey_refsD(3) valid_W_invE(5)) + done +next case (mark_loop_get_work_load_W s s') then show ?case + apply (subst valid_W_inv_def) + apply (clarsimp simp: all_conj_distrib) + apply (intro allI conjI impI; auto dest: valid_W_invD0 valid_W_invD simp: valid_W_distinct_simps split: if_splits process_name.splits) + done +next case (mark_loop_blacken s s') then show ?case + apply (subst valid_W_inv_def) + apply (clarsimp simp: all_conj_distrib) + apply (intro allI conjI impI; auto dest: valid_W_invD0 valid_W_invD simp: valid_W_distinct_simps split: if_splits process_name.splits) + done +next case (mark_loop_mo_co_W s s' y) then show ?case by - (erule valid_W_inv_mo_co_W; blast) +next case (mark_loop_mo_co_unlock s s' y) then show ?case by - (erule valid_W_inv_mo_co_unlock; simp split: if_splits) +next case (mark_loop_mo_co_mark s s' y) then show ?case by - (erule valid_W_inv_mo_co_mark; blast) +next case (mark_loop_mo_co_lock s s' y) then show ?case by - (erule valid_W_inv_mo_co_lock; assumption+) +next case (mark_loop_get_roots_load_W s s') then show ?case +(* FIXME ran out of patience. Something makes auto diverge on some subgoals *) + apply (subst valid_W_inv_def) + apply (clarsimp simp: all_conj_distrib valid_W_inv_sys_ghg_empty_iff) + apply (intro allI conjI impI) +apply (auto simp: valid_W_invD valid_W_invD0 valid_W_inv_sys_ghg_empty_iff split: process_name.splits; fail)[1] +apply (auto simp: valid_W_invD valid_W_invD0 valid_W_inv_sys_ghg_empty_iff split: process_name.splits; fail)[1] +apply (auto simp: valid_W_invD valid_W_invD0 valid_W_inv_sys_ghg_empty_iff split: process_name.splits; fail)[1] + apply (clarsimp split: process_name.splits) + apply (meson Int_emptyI Un_iff process_name.distinct(4) valid_W_distinct_simps(3) valid_W_invD0(1)) +apply (auto simp: valid_W_invD valid_W_invD0 split: process_name.splits; fail)[1] +apply (auto simp: valid_W_invD valid_W_invD0 valid_W_inv_sys_ghg_empty_iff split: process_name.splits; fail)[1] + using valid_W_invD(2) valid_W_inv_sys_ghg_empty_iff apply fastforce +apply (auto simp: valid_W_invD valid_W_invD0 valid_W_inv_sys_ghg_empty_iff split: process_name.splits; fail)[1] +apply (auto simp: valid_W_invD valid_W_invD0 valid_W_inv_sys_ghg_empty_iff split: process_name.splits; fail)[1] +apply (auto simp: valid_W_invD valid_W_invD0 valid_W_inv_sys_ghg_empty_iff split: process_name.splits; fail)[1] +apply (auto simp: valid_W_invD valid_W_invD0 valid_W_inv_sys_ghg_empty_iff split: process_name.splits; fail)[1] +apply (auto simp: valid_W_invD valid_W_invD0 valid_W_inv_sys_ghg_empty_iff split: process_name.splits; fail)[1] +apply (auto simp: valid_W_invD valid_W_invD0 valid_W_inv_sys_ghg_empty_iff split: process_name.splits; fail)[1] +done +qed + +lemma (in mut_m) valid_W_inv[intro]: + notes if_split_asm[split del] + notes fun_upd_apply[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>\ valid_refs_inv \<^bold>\ valid_W_inv) \ + mutator m + \ LSTP valid_W_inv \" +proof(vcg_jackhammer (no_thin_post_inv), vcg_name_cases) + case (alloc s s' r) then show ?case + apply (subst valid_W_inv_def) + apply (clarsimp simp: all_conj_distrib split: if_split_asm) + apply (intro allI conjI impI) + apply (auto simp: valid_W_distinct_simps valid_W_invD0(3) valid_W_invD(2)) + done +next case (store_ins_mo_co_W s s' y) then show ?case by - (erule valid_W_inv_mo_co_W; blast+) +next case (store_ins_mo_co_unlock s s' y) then show ?case by - (erule valid_W_inv_mo_co_unlock; simp split: if_splits) +next case (store_ins_mo_co_mark s s' y) then show ?case by - (erule valid_W_inv_mo_co_mark; blast+) +next case (store_ins_mo_co_lock s s' y) then show ?case by - (erule valid_W_inv_mo_co_lock; assumption+) +next case (store_del_mo_co_W s s' y) then show ?case by - (erule valid_W_inv_mo_co_W; blast+) +next case (store_del_mo_co_unlock s s' y) then show ?case by - (erule valid_W_inv_mo_co_unlock; simp split: if_splits) +next case (store_del_mo_co_mark s s' y) then show ?case by - (erule valid_W_inv_mo_co_mark; blast+) +next case (store_del_mo_co_lock s s' y) then show ?case by - (erule valid_W_inv_mo_co_lock; assumption+) +next case (hs_get_roots_loop_mo_co_W s s' y) then show ?case by - (erule valid_W_inv_mo_co_W; blast+) +next case (hs_get_roots_loop_mo_co_unlock s s' y) then show ?case by - (erule valid_W_inv_mo_co_unlock; simp split: if_splits) +next case (hs_get_roots_loop_mo_co_mark s s' y) then show ?case by - (erule valid_W_inv_mo_co_mark; blast+) +next case (hs_get_roots_loop_mo_co_lock s s' y) then show ?case by - (erule valid_W_inv_mo_co_lock; assumption+) +next case (hs_get_roots_done s s') then show ?case + apply (subst valid_W_inv_def) + apply (simp add: all_conj_distrib) + apply (intro allI conjI impI; clarsimp simp: valid_W_inv_sys_ghg_empty_iff valid_W_invD(2) valid_W_invD0(2,3) filter_empty_conv dest!: valid_W_invE(5)) + apply (fastforce simp: valid_W_distinct_simps split: process_name.splits if_splits) + done +next case (hs_get_work_done s s') then show ?case + apply (subst valid_W_inv_def) + apply (simp add: all_conj_distrib) + apply (intro allI conjI impI; clarsimp simp: valid_W_inv_sys_ghg_empty_iff valid_W_invD(2) valid_W_invD0(2,3) filter_empty_conv dest!: valid_W_invE(5)) + apply (fastforce simp: valid_W_distinct_simps split: process_name.splits if_splits) + done +qed + +(*<*) + +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,79 +1,76 @@ (*<*) (* * 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 + , hs_pending := \False\ + , hs_type := ht_GetRoots , mem_lock := None - , mem_write_buffers := \[]\ + , mem_store_buffers := \[]\ , phase := ph_Idle , W := {} , ghost_honorary_grey := {} - , ghost_handshake_in_sync := \True\ - , ghost_handshake_phase := hp_IdleMarkSweep \" + , ghost_hs_in_sync := \True\ + , ghost_hs_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) +unfolding gc_system_init_def gc_initial_state_def mut_initial_state_def sys_initial_state_def +apply (clarsimp; intro conjI) + apply (auto simp: ran_def; fail) +unfolding valid_refs_def reaches_def +apply auto apply (erule rtranclp.cases; auto simp: ran_def split: if_splits obj_at_splits)+ done (*>*) text\\ end (*<*) end (*>*) diff --git a/thys/ConcurrentGC/concrete/Concrete_heap.thy b/thys/ConcurrentGC/concrete/Concrete_heap.thy --- a/thys/ConcurrentGC/concrete/Concrete_heap.thy +++ b/thys/ConcurrentGC/concrete/Concrete_heap.thy @@ -1,47 +1,62 @@ +(*<*) + theory Concrete_heap imports "HOL-Library.Saturated" - "../Proofs" + "../Global_Invariants" begin +(* From 40e16c534243 by Makarius. Doesn't seem to have a huge impact on run time now (2021-01-07) *) +declare subst_all [simp del] [[simproc del: defined_all]] + +(*>*) type_synonym field = "3" type_synonym mut = "2" +type_synonym payload = "unit" type_synonym ref = "5" -type_synonym concrete_local_state = "(field, mut, ref) local_state" -type_synonym clsts = "(field, mut, ref) lsts" +type_synonym concrete_local_state = "(field, mut, payload, ref) local_state" +type_synonym clsts = "(field, mut, payload, ref) lsts" abbreviation mut_common_init_state :: concrete_local_state where - "mut_common_init_state \ undefined\ ghost_handshake_phase := hp_IdleMarkSweep, ghost_honorary_grey := {}, ghost_honorary_root := {}, roots := {}, W := {} \" + "mut_common_init_state \ undefined\ ghost_hs_phase := hp_IdleMarkSweep, ghost_honorary_grey := {}, ghost_honorary_root := {}, roots := {}, W := {} \" context gc_system begin -abbreviation sys_init_heap :: "ref \ (field, ref) object option" where +abbreviation sys_init_heap :: "ref \ (field, payload, ref) object option" where "sys_init_heap \ [ 0 \ \ obj_mark = initial_mark, - obj_fields = [ 0 \ 5 ] \, + obj_fields = [ 0 \ 5 ], + obj_payload = Map.empty \, 1 \ \ obj_mark = initial_mark, - obj_fields = Map.empty \, + obj_fields = Map.empty, + obj_payload = Map.empty \, 2 \ \ obj_mark = initial_mark, - obj_fields = Map.empty \, + obj_fields = Map.empty, + obj_payload = Map.empty \, 3 \ \ obj_mark = initial_mark, - obj_fields = [ 0 \ 1 , 1 \ 2 ] \, + obj_fields = [ 0 \ 1 , 1 \ 2 ], + obj_payload = Map.empty \, 4 \ \ obj_mark = initial_mark, - obj_fields = [ 1 \ 0 ] \, + obj_fields = [ 1 \ 0 ], + obj_payload = Map.empty \, 5 \ \ obj_mark = initial_mark, - obj_fields = Map.empty \ + obj_fields = Map.empty, + obj_payload = Map.empty \ ]" abbreviation mut_init_state0 :: concrete_local_state where "mut_init_state0 \ mut_common_init_state \ roots := {1, 2, 3} \" abbreviation mut_init_state1 :: concrete_local_state where "mut_init_state1 \ mut_common_init_state \ roots := {3} \" abbreviation mut_init_state2 :: concrete_local_state where "mut_init_state2 \ mut_common_init_state \ roots := {2, 5} \" end +(*<*) end +(*>*) diff --git a/thys/ConcurrentGC/concrete/README b/thys/ConcurrentGC/concrete/README deleted file mode 100644 --- a/thys/ConcurrentGC/concrete/README +++ /dev/null @@ -1,52 +0,0 @@ -This directory contains the following three files. - -- Main.hs -- WriteDotFile.hs -- input_heap.txt - -Main.hs and WriteDotFile.hs generates an executable that takes -an input file and returns an output file whose extention is .dot. - -input_heap.txt is an example of input file. - -Input files are converted to pdf files by using Graphviz. -Suppose the executable has generated output.dot, and Graphviz is -installed properly, Graphviz generates a pdf file whose name is -output.pdf from the following command: - - dot -Tpdf input.dot -o output.pdf - -The input file for the executable should follow the specific format -that described in the following. - -The input file consists of two parts. -The first part describes the mutators and the lines starting from -mutators. - - 0 1 2 3 - 1 3 - 2 2 5 - -stands for the following: -There are three mutators called 0, 1, 2, respectively. The mutator 1 -has references to the object 1, 2, and 3. The mutator 1 has a -reference to the object 3. The mutator 2 has references to the -object 2 and 5. - -The first and second parts are split by "--". - -The second part describes the objects and the lines starting from -obejcts. - - 0 - 1 - 2 - 3 0 1 1 2 - 4 - 5 - -stands for the following: -There are six objects called 0, 1, 2, 3, 4, 5, respectively. The -objecct 3 has two fields called 0 and 1 respecitvely. In the field -0, a reference to the object 1 is stored, while the field 1 contains -a reference to the object 2. diff --git a/thys/ConcurrentGC/concrete/README.md b/thys/ConcurrentGC/concrete/README.md new file mode 100644 --- /dev/null +++ b/thys/ConcurrentGC/concrete/README.md @@ -0,0 +1,41 @@ +## Concrete heap translator + +This directory contains a Haskell program for translating high-level +descriptions of heaps into Isabelle `.thy` files and Graphviz `dot` +files. A sample input file is `input_heap.txt`. + +> ghci Main +> > main + +> dot -Tpdf input.dot -o output.pdf + +The input file consists of two parts split by `--`. + +The first part describes the mutators and their roots. For example: + + 0 1 2 3 + 1 3 + 2 2 5 + +stands for the following: + +There are three mutators called 0, 1, 2, respectively. The mutator 1 +has references to the object 1, 2, and 3. The mutator 1 has a +reference to the object 3. The mutator 2 has references to the +object 2 and 5. + +The second part describes the objects and their fields. + + 0 + 1 + 2 + 3 0 1 1 2 + 4 + 5 + +stands for the following: + +There are six objects called 0, 1, 2, 3, 4, 5, respectively. The +object 3 has two fields called 0 and 1 respectively. In the field 0, a +reference to the object 1 is stored, while the field 1 contains a +reference to the object 2. diff --git a/thys/ConcurrentGC/concrete/WriteDotFile.hs b/thys/ConcurrentGC/concrete/WriteDotFile.hs --- a/thys/ConcurrentGC/concrete/WriteDotFile.hs +++ b/thys/ConcurrentGC/concrete/WriteDotFile.hs @@ -1,139 +1,134 @@ {- - 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) -} module WriteDotFile ( write_dot_file ) where import WriteFileBasis -{- - - WriteDotFile provides functions that allow to write .dot files - - from its input file whose extension is .txt. - -} - -- the header of the output file. header :: String header = "digraph g {\n" ++ " graph [rankdir = \"LR\"];\n" ++ " node [fontsize = \"10.5\" shape = \"ellipse\"];\n" ++ " edge [];"; -- returns a string that describes one mutator node in the output file. write_mut :: String -> String write_mut m = " mut" ++ m ++ " [label = \" mut" ++ m ++ "\" shape = \"record\"];\n"; -- returns a string that describes all the mutator nodes in the output file. write_muts :: [[String]] -> String write_muts [] = ""; write_muts (m:ms) = (write_mut $ head m) ++ write_muts ms; -- returns a string that describes the subgraph for mutators. subgraph_mut :: String -> String subgraph_mut contents = "\n\nsubgraph clusterA{\n\n" ++ (write_muts $ map words $ lines $ get_mut_part contents) ++ "\n}\n"; -- auxiliary functions to draw fields in objects take_eventh :: [String] -> [String] take_eventh [] = []; take_eventh (x:[]) = []; take_eventh (x1:x2:xs) = x2:(take_eventh xs); -- returns a string that describes all the fields in one object. fields :: [String] -> String fields [] = [] fields (f:fs) = " | " ++ f ++ (fields fs); -- returns a string that describes one object in the .dot file. write_obj :: [String] -> String write_obj os = " obj" ++ (head os) ++ " [label = \" obj" ++ (head os) ++ (fields $ take_eventh os) ++ "\" shape = \"record\"];\n"; -- returns a string that describes all the objects in the .dot file. write_objs :: [[String]] -> String write_objs [] = ""; write_objs (o:os) = (write_obj o) ++ write_objs os; -- returns a string that describes all the objects in the .dot file. subgraph_obj :: String -> String subgraph_obj contents = "\n\nsubgraph clusterB{\n\n" ++ (write_objs $ map words $ lines $ get_obj_part contents) ++ "\n}\n"; -- returns a string that describe one line starting from a mutator. write_mut_1_line :: String -> String -> String write_mut_1_line from to = "\"mut" ++ from ++ "\":f -> \"obj" ++ to ++ "\":f;\n"; -- returns a string that describes all the lines starting from a mutator. write_mut_line :: [String] -> String write_mut_line [] = []; write_mut_line (l:[]) = []; write_mut_line (l1:l2:ls) = write_mut_1_line l1 l2 ++ write_mut_line (l1:ls); -- returns a string that describes all the lines starting from all the mutators. write_mut_lines :: [[String]] -> String write_mut_lines [] = ""; write_mut_lines (l:ls) = write_mut_line l ++ write_mut_lines ls; {- - returns a string that describes all the lines starting from - all the mutators. -} lines_from_mut :: String -> String lines_from_mut contents = "\n" ++ (write_mut_lines $ map words $ lines contents); -- returns a string that describes one line starting from a object. write_obj_1_line :: String -> String -> String -> String write_obj_1_line from_obj from_field to_obj = "\"obj" ++ from_obj ++ "\":f" ++ from_field ++ " -> \"obj" ++ to_obj ++ "\":f;\n"; -- returns a string that describes all the line starting from a object. write_obj_line :: [String] -> String write_obj_line [] = []; write_obj_line (l:[]) = []; write_obj_line (l1:l2:[]) = []; write_obj_line (l1:l2:l3:ls) = write_obj_1_line l1 l2 l3 ++ write_obj_line (l1:ls); -- returns a string that describes all the line starting from all the object. write_obj_lines :: [[String]] -> String write_obj_lines [] = ""; write_obj_lines (l:ls) = write_obj_line l ++ write_obj_lines ls; {- - returns a string that describess all the lines starting from - all the mutators. -} lines_from_obj :: String -> String lines_from_obj contents = "\n" ++ (write_obj_lines $ map words $ lines contents); -- .returns a string that describes all the lines specified in the input file. draw_lines :: String -> String draw_lines cs = (lines_from_mut $ get_mut_part cs) ++ (lines_from_obj $ get_obj_part cs); footer :: String footer = "}"; write_dot_file :: String -> String write_dot_file inp = header ++ subgraph_mut inp ++ subgraph_obj inp ++ draw_lines inp ++ footer; diff --git a/thys/ConcurrentGC/concrete/WriteThyFile.hs b/thys/ConcurrentGC/concrete/WriteThyFile.hs --- a/thys/ConcurrentGC/concrete/WriteThyFile.hs +++ b/thys/ConcurrentGC/concrete/WriteThyFile.hs @@ -1,117 +1,118 @@ {- - 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) -} module WriteThyFile ( write_thy_file ) where import WriteFileBasis intersperse :: a -> [a] -> [a] intersperse sep xxs = case xxs of [] -> [] x : xs -> x : prependToAll xs where prependToAll xxs = case xxs of [] -> [] x : xs -> sep : x : prependToAll xs -{- - - WriteThyFile provides functions that allow to write isabelle - - files from its input file whose extension is .txt. - -} - -- the header of the output file. -- FIXME in general the type synonyms are a function of the input. header :: String -> String header inp = unlines - [ "theory Concrete_heap" + [ "(*<*) + , "theory Concrete_heap" , "imports" - , " \"~~/src/HOL/Library/Saturated\"" - , " \"../Proofs\"" + , " \"HOL-Library.Saturated\"" + , " \"../Global_Invariants\"" , "begin" , "" + , "(*>*)" , "type_synonym field = \"3\"" , "type_synonym mut = \"2\"" + , "type_synonym payload = \"unit\"" , "type_synonym ref = \"5\"" , "" - , "type_synonym concrete_local_state = \"(field, mut, ref) local_state\"" - , "type_synonym clsts = \"(field, mut, ref) lsts\"" + , "type_synonym concrete_local_state = \"(field, mut, payload, ref) local_state\"" + , "type_synonym clsts = \"(field, mut, payload, ref) lsts\"" , "" , "abbreviation mut_common_init_state :: concrete_local_state where" - , " \"mut_common_init_state \\ undefined\\ ghost_handshake_phase := hp_IdleMarkSweep, ghost_honorary_grey := {}, ghost_honorary_root := {}, roots := {}, W := {} \\\"" + , " \"mut_common_init_state \\ undefined\\ ghost_hs_phase := hp_IdleMarkSweep, ghost_honorary_grey := {}, ghost_honorary_root := {}, roots := {}, W := {} \\\"" , "" , "context gc_system" , "begin" , "" ] write_mut :: [String] -> String write_mut [] = "" write_mut (m : ms) = unlines [ "abbreviation mut_init_state" ++ m ++ " :: concrete_local_state where" , " \"mut_init_state" ++ m ++ " \\ " ++ "mut_common_init_state \\ " ++ "roots := {" ++ concat (intersperse ", " ms) ++ "} \\\"" , "" ] write_muts :: [[String]] -> String write_muts [] = "" write_muts (m : ms) = write_mut m ++ write_muts ms write_mut_part :: String -> String write_mut_part = write_muts . map words . lines . get_mut_part write_refs' :: [String] -> [String] write_refs' [] = [] write_refs' (x1:x2:xs) = x1 : "\\" : x2 : write_refs' xs put_commas :: [String] -> [String] put_commas [] = []; put_commas (x:[]) = []; put_commas (x:y:[]) = []; put_commas (x:y:z:[]) = x : y : z : []; put_commas (x:y:z:ws) = x : y : z : "," : put_commas ws; write_refs :: [String] -> String write_refs [] = []; write_refs xs = concat $ intersperse " " $ put_commas $ write_refs' xs; write_obj :: [String] -> String write_obj [] = [] write_obj (o : os) = concat [ " " ++ o ++ " \\ \\ obj_mark = initial_mark,\n" - , " obj_fields = " - , if os == [] then "Map.empty" - else "[ " ++ write_refs os ++ " ]" + , " obj_fields = " + , if os == [] then "Map.empty,\n" + else "[ " ++ write_refs os ++ " ],\n" + , " obj_payload = Map.empty" , " \\" ] write_obj_part :: String -> String write_obj_part cs = unlines - [ "abbreviation sys_init_heap :: \"ref \\ (field, ref) object option\" where" + [ "abbreviation sys_init_heap :: \"ref \\ (field, payload, ref) object option\" where" , " \"sys_init_heap \\" , " [" ++ concat (intersperse ",\n " (map (write_obj . words) $ lines $ get_obj_part cs)) , " ]\"" , "" ] footer :: String footer = unlines [ "end" + , "(*<*)" , "" , "end" + , "(*>*)" ] write_thy_file :: String -> String write_thy_file inp = header inp ++ write_obj_part inp ++ write_mut_part inp ++ footer diff --git a/thys/ConcurrentGC/document/heap.pdf b/thys/ConcurrentGC/document/heap.pdf new file mode 100644 index e69de29bb2d1d6434b8b29ae775ad8c2e48c5391..45d1b874e7b6567c0576ac94adafb00c8ecad004 GIT binary patch literal 12776 zc$~F*WmsHG(=7=EclQhsTn8WAJ-E9~U~qT05ZqmYYmneha3=(UySoRMo8-uQp7(su zckiD&viH=g+Ev}v-E00VN<}dVMj#UhB4yoC#XTZBfCXS{Y=Owf2Vj;1+n70-12}#{ z%7_2}fLQ`!s(%Y4tQ4bvt50%9#N?kQh*ZITQ?vBNjP3$_*g+L5 zty~ay5ZdDHc2SjsLfp3`dSlnnv2xks`_#Tt%iH|i=!xeuOX7%obK)e*LL0UsS4(&@ zzCy}Kp50^2tr*^45>0v?OIf6=UC#fbLQ{Dt&7_1jnh+&xk-ts>O7a`Ja0u-VCV4}~ z8@_LX*TzacRtVA}_Imh$(Pe)udAo8{(#PRMWgTyrnb{VLg-;AG49iUqsH z-l{v^Hy}ICgP=Nb${~17mK#)y62iQ(I|`NTfM5thVkzs>SNZ-9%2GbiUIU61ily{Z z1BT_>2y(WsJj@>v?kSL0M?g?@NpucHv9<65q(YTzkro#DWySV{QgfvBi3D}zqgap| z%PKX^Dw7*}xvDmvIuv9vIObxt=1{_!5{4g`79L@Q#z2KHSExguJ(|%c)3m0^bLU*5 zi~KLe0c`?``Y1ir{&q5Vr%XrdnIWX;07%kA`FCyO9TfS9(t1o1_R24!JK4z--aVM_ z6GK|+h$bB1T)*)$%n3F$@Pj^cc$ipoo$EkmP8ZzA237fg=r73egt3}jGA|@n*CSF^ ziM?AH_~b!Lz-lS=v4VhuM?FTU)=KK*>=zIgW@VI|R0L6~{taEBpaSZYTuYxwhxrgw z)kpl|s!36LfZX?d8lp1Ag*Nn{8m^QV2}0T#8YR%3v6(gzns9DO!m*!WGwM1x!__j! z&B9Xh($lHUVlARk9azciqPRf~D0JTg4WhH(nINoreV`n2w|Qk2Hg+@>8yQc)0tSAt zp8!O>D%8#W4E0`ya5g<12w}nVBdEfYCc&YCW;uAV{bB}My*4zUERiZf{Ou`=kI>>! zig6u-g80#xLJ4;8|y|Bs11Z3GTIVZPyb0H zV;4U#I;!cyBN*RLSz85fXS+KB#N@r#s?ooFcz$t+eAa6`_tw%$=o6{ys!ySz`t?E2 zR|J$SVK%U|t!B7cdX*_@(`U~BuTav^wXG(Ihu^zp#j1;&ugths0)hg_9u@l=r2!mj z2DY&(YOXwmq5%4ScW|gIh$&e|`x=k>RY`m{>oO*7AiHSZg=Lyv-b zM6iwNpXBoM*DqS)_>-P~vJnu-!uc1>_BU(&<|%PECrK5jpBx4F1Nem*u>gK?ny4@U z2w*lb`g#4yTt6LuT$v?oZJhoQ{f?1^kp;l{7xVoL`uX{@80$YMQq|oK3}99?G6MrB z0l?px6pb9fKXU>9X3e)?Q;3m>ts6k+C&U6^XX6C0a&rOnemg~f*8Lgo2w?v^UedwV z+3p`>`J={f^e;%|C+c8i<7oG*mWlgcur$EQ!5RFQCHgZ;4D13i0V_+2{9~%hU`JbL z2NSR(fbDn9|ICc*-{jBzzmoqi{Et8&0La14@ehC%z|GFe^$&mz@GIb-e-}p{U{zwv zNk-g7GITFHvv+_-7&J_H;lN=QZe&*U45$|YI(DuNFs`mMfENM5fGo)^W!StTRcXijq=r$HWFBZQ-#1rE_U)PcvA2mtss``hxv##Y9R>i z>#D(NywIeNUMUd~`$=|0yS|8>rc*&YJX&3cK~FezTksqqf;{<=YtaPzxMNTVW{yyP zL{QrpF9q>DnBZ4o{coOe;2p7FrhvNSSO^LXuWX`{wsGyqIt`+sa8cA4)m>boI`|ER zg@ye-EOemzu}=+tTSV==L<^9tM4EzqXNy!9hyDf?|F6S zW^oarWC(P4x}LsLhhklF(a<6|zH+k*xP^hrSbS*+g@37;2Dom;0aXrjFri^cvDS?!0VXqTSp94>l9b#6r5J7*+qu?Ot~QV0)Fqu5%T<^_p=aJ=8zY43`am8KKixeVvW3~*IM|=n1wEytAQR-5o{VOA zxNCkGg5ZJuhKrfG6B^l#_)YP_>l!F1hq*7S(X>#WW{r5O_AWqQq}3)Us9jX&pvSAH zJ#-i-BbToohyB6+r>?#m(UmXY8|;FfbYQqv&m30|ZyEx3CXDvZU4vFP$f+ii*y?6P=0Xze~1$cYDM1QBi{i z=3?j>ipx;Gja?WdDUjbIOh+4Bi{Eff6LiJ$k0 zI>HO2vXyLE4};@Hn@#X!P{lLFTW2Y{&}U|r;Uw?YHz`XouNt@*ZLN~?zt-WNgKCF! zrH-Dh#AB-(8I%s*?#X9S_?|nwx_a1RJUqUgSG{w(QD*}*)=LC3;w^m*VK5KQV^^SL z4NvjuE`LkqU@OdfJd2*{?{>g2BH@YcMkFEY&of%GfwAspU05_){&Y#8lLJL*Ut({t zDogIs?nV}9Fxjrb!OrTaWM*@{vm>><^*JaifF|ScO&yjOw*QR2>+YzJUg0ept6quM zFnAxP1(#MR@x^S=L87T4>&cosi?}|c4Vjom5-earYnRVBE@L@h8wW!xQIXi;?i@k( z>6IFU^Lthed{DgB=HZ_6G3yxO-5cw<1%bHE+Oz4tbmh#fw$xf0N?q&J#QSTZuTC>! zL3>&ZvLj_-=Dx6<(qBU}lk64jxRB^-!!1wWhAI_OPU}TIsj;^YO`Bg_-WNN1oG`Mt ze%Fn*cIu3T$Lu29;i5I9vhb`3zOQQ>wFgFxppB2SHF=HC(SMyr`RFV}^^C#fW$v8# zmgD_;W8|mG){7%v)d`cwIL*{4y@dLG$JqpH>OYs#6NIL?=w40 z3YYE9ZDlftaxWp1EkY+={1YchkBSH6q{R#1@ODCCE#ZvQy&YH~SJ$eCuT%K9?jDfa zZS&Ffc%SIH#W_8CGnX|L3m6emzQ`ndi)7WQ!dd<3#u68%J`mB&Ba75kj=o+FS)}8k z#6takaU%8nk5{w^()vc0z75)D8q_t|7@oXvivEUk_eVS^&lhFMd}_JGQ58h0_0Gd6 zwKDK+F2+djPKeg)vQQv|2j>`d)v=#E6Nw8k;=G0(Z(_reRCsWAO~;2=>sx0OzJ!!N z-i9?zYWOBcKD;IPM8|37C6N*x3JSW>+Hh!xIq&3^d_wv#P#*ch6r9t&;Ac6|RGlO5 z-4K6s03P?~A%g%nkK+kYOA()_kiUP#(~$>u)LOyOXhA|LxD~z(M7XY?HciEU#^A?Y zjF51(5m{Eo9e-=(DcZh|mM=K4sP(?!zDAoO59bvlHUexb_QzQQ6|5A*Xd(DdJ6Ccd zp)FqLc7}A(7il~cZoX>{xHApw1vRy(<48_htmRgbTjoy6#V@tWY<_&pJ8>Qyd$pR+xIWKz`1@# zw1%3@@uS8gz7T7kAGMDs1jH!Jz3mJ4GJ^%$Tf=y=_V%x+kn( zLOu`DEexd%w$}4?x_OV>7GyJUtx@oA#KP>x8J+DfSJ#FVIzad8zwlsZ5pPPtrFU)BgKuC)7&+iKhe=JWq zHj7hLF`t6Tm@dRyvroEo|M=r7%0`5E7h0)gCqmUPY3ExoUVegOzEAfejkV5)luvI} zfc5CfpTJE@q)!<}1@#{uWDVO|!qx4~`0j31*arPAH1C$KkKJPFSvAB$yp%Kz9aYLU z$&o}r>Abe+4~D+l%s!nngsNV5 zy&sBBraTu+g3Lwjxo=+F^ES14-A9ME8wlUG_Fy{;hGM(jrrJ2-+N##av_8NS;JV_1 z0)Z7e#8Tf;Nw|r(n+F@~zvdl(_W&S~DTe5i@gb*exHSNLt` zwRmxR27im7k@^#6SC9^8%%~5bBzgROc&zbZc|o2k2Zsg!U{2H9SS?lPvyQXfONv`+)634cwsfthR4|^nnR+HE<91`rFN;(&Og-qkY^A4}1eL!CC zmL1#J(tyRT=>DjTD5nt7e)pjpV)4`N7ewEC zO1PB~KuFmO5u8*a3^035qLjm9-8dR0SXlOf-^Ph2qG&Ql(yVwbmOV>GD*@LxRPTPcDJ%@Gs46}=c^C4*9`#n-I&c!WwD-c0 zBc(FuZH^^Xghk ORi847kXG;M`pQ{J^#qzoDnHuIE=%`v@Pu-&e)&RFi)OU>cL`A%YY zn5AkP#`=kF=&-Rj4AXs$`ZM~2a4^~lIR%KC650V5v2VB>m2T2vVowM;2Ekjy&@DO| zO{A%;Vxrf4cfHKOgeg`(;m4?F^QOH?AnjVO>DW=(@ZASHKA}nWkV&v|qRl6gr%TtlgG){jyB5z{_*^Ez!UKHUT4E0AXgoyB zEK!|ybHz5$6ff1Xl2iBQ?r5Hcsj&R?58315?21Z^6M74m==hY?af{I`d01^jvFqBATLn9p2R+=%h3?aJt+|J$4r6Z>(`hRuiIcsc);1b1?5!%o!SJX&an;hHG zr|D=n+>^ng7?Z69xD0GDdqZ8i1iMiYe}iUPRubq&w>72JEZ0L$9v{R;#g;h9IJd<- zaM!u&EG^s{S;@lb6ze8}_Zgl@YLWAuzkZdUQN%*maWi$AIgiKrotbf?db z>7nki4`O1&?`??DWcRwN>NWQsdnq`Ms{pMlJ}mUJ!J6&18yVXz*-Omo^UI2DDuOi# zW@uGW*4(}Ghz_N&{;B46yeZ9%gQ0mNX4xueq9Q^9$sE7ZE>JUgZL0w$x<(;48x5%k zO`^pnEH*wu-Z*xDV1QlCgVf@-4k!Dyffjm)39_VE)P6!^lHou~>KTm5Oa(*15&&Kv zrv#2kK6^j;n!0HxU5Qxfq=sW(Z6s=1&nU#&&h0B!r82pb{R zQ{pKc@jt3%wVlKHvinfk7;0DL=Omp$Ih2m%r9Y;Ta~p8$p|G$2Wm7S9GD5W~FT05f zEQE8~w=-5nHwk8mNLyJpsDO)8PIUwJm!xNFHx4V$5ZrrG^tZ?^6MN$q=(!-rL0z=@ z#vA&a=fAFUi99S88ktP)@1d#@S)?Sixj7qNJbc;<-}>6?uH|Q|u6UfTVM+W2GU#^v zLak8yecsyAN^xEWZe|FUngLZbeS+m!`rr~s0-k}lgH+o|r#LowLRHhGZOnL(YI}k< zcY4)hdHvZrVv<_A@oS7AV9e1H)^|!P;Mh{yc>KwXc0+q$C^QYl#`I}R>8)IH2I9tp zp%rmfYH#h>dQj+uCr!FM(MrHw3vxtaZ-o{fwzs*J{OA%%;C-Z#kg?8&9h8su`@uQ5 zCAvc<>-c4bH=r^j4epclklu`x{yduqPo z{99ScN)L#Dl*x13O*D8&nml_ZX5imc%`x%}nT|Rb|#Y7t=Xxs@=Wn zO;ELVB_}zSv|#_{6oW|nD;|uhLqTu0g|18o-lgtFu2i)j=4NtAN8Y`8agd(AR+EOt z+unt#MTE0i`OF;8=qPtr|G2Wi;;Vt8=t`-=i01LEAgXY@^~hME4bsN9r5hRp;$2>& zeCKk`YL5jUQ!5e*opdK|>lzrgCF6RU7YrKL)>ad-&W6#JVc3jE3NNWC$Ztg_tJ>C_ zXLg(I)A(AG%aMsXCC)YG4=P7+EQb{6Xlk*)yolJwFo}|Cd9rnc6Hx!5X=vzGw(m1!kHwo2#oK0dClNdl!VhK(1rS5L;3v(w=wFvC9 z`%yT1I=E#v%cA~LS(xv*Y(ISYv7zZ@F8RbIZ97?f86$gUHfumapWoVIH;Lrn^!BdV z2rj#DW+_D@d#ks9ja3e-go(7LqYA(S#p49qqh;4P8P&Cei|-9i<~7}3s)(D-XX7Ph z-LPwCHwh8h?0CeTc=G$|9P}E7C1;)wRi~dO)3|}ksu7+WB<&R*8FAG6%7U$rjp?pu zoTtspZE+l@1IwIn1HCv1N#4UpEmjOS2`6~nH$IiP$*{oD(q3255fq?C1IX0lBP9O( zOBwq*-AOlk)A}ishc-`M)R2+ksb-`m@4mFd&(^qF1uYlMe=I$s%vYhjB1J2XauDXg z4V=w;P;q;P!XjrO{kXWrTM$ey>0gzNIW$54eHpLRE}8-;eN92~WyCyMO4T3C6Al%1 z>W_4y9)@O+no%H1@g@@S!`?>y0u%)LJ%c? zWiCDFt6aJ3ApkmNf)NN)C0kT%CC#C^j;U;DM7*!`!3^MJBAV| zhpJ6(i_)>8vn;Sx$_E~~&3PV$y6+B4`Vb>z3o~zIyn8Z>%&!1q>AqZxbm9Trds5FAIdJU8j1p+rru=%e4vzg6nSh`!!;qbOw%mFqqV3hfsQEuKL; zqZdcp9fs8n+JMn~&#v2(rVGP&d2;ARdaMmSUPWiLG4#@pJ;OA2+CTl6F3zbXNbzk6DsMJg7RQ@69LpbR&Q-=M!SlL!^% z$HC6M&zgj}6j2=#L`-P`S}A1cPD2^^6_Y4t;Z$FL%6XJS^gm-dD;)p)P_{yZf2?gl z<>P6h<#pqaI=jH#QNTmG@x-8R<#VA|80a8eI3{sc_wuG&@~bWd7w|wK(KhnPVZnzO z=XgLJq#;=5gurI*5)JaTJf+e|& zK~ifqcPYF%)_%lolc@4WbP6krT4`D(VHC|*(@;k2m=lsholvWqYA?KBMq1ohTg-d` zcEQOmI4ZUn+j^U)1xB#wqy}PoiK_VXxT{w{GdDbDGUj<95Q0*>xDtVj1F9I~?iVsk z?OKyfa1%LUv&$#*B=SdJqHgTzx7hEe(Sn*ZJrlg=`A{-5%1L5!!8gtlt+cRP>)$Ao z_vWfaVDUB)39#fx1ElXOY6W}RcN7$h16A$ObYprLdKPTj6v-%rZBV{;bM5I9XQ6zR zaindbb+Of1sjjHS4pX#wtD9$eQ@@$>o<*}k_i*I-jT^Sc$0BSFaKOV;Ku84|&oE*kiuG*J0SmyIHr= zPn%(qxu8Gtm?R|x4ydtl*#|GiQlD~=l=1u8*WwyTZtrNPeuhMd42!?6jY@QpR-5bY z3U}16t&et@&>E5w@*nfo52#Ff=e3G^czf28eLc8hhCCkGg+$6*2AI%T_&`$YmqZxVD@W-8GG>Y^X z9)uq)%P*LYOZh0W?S8L@tH->@ z&u`#>cSN+6k757T^dnlgYg9`c60a3&fa4~tZpg{Efi0k}l}=jAQ#Od$Fq2zm8t+hJ zn_M+a_y~N+Mw--I`-CgX4cl(hR)w|@iVK`oJ=wxCm^1)lZ(l^c>^NdbIMsdAEUq@r z0?gIPy;lpH&YJLZ_W$un3+(Q-jZx06xxNV6JTSA)ps4kaiH5qIz;+5R}a0ZgmooK7)^YfXe*DJu>@0Q z;xLZD6ly)+0{5hCTNXc4!*qJ2I#ExBiJ8}-81jBt1$}romSjWQ%z}8oq-JK?nbMr! za?r1t(6-C|1@v)l#Qo^xtRO5$?%US1&M6TR`%wS81EqPkE6GG08xb=~wcQ6U-~NU=%;S~RUdq?~5m7>1Yjlixv+1RVN!OcH2ey1r!dSAZ(D{e) zWm*ibx~sVwVMXaQ7V=OX$yGzg~~OFKlvmr zv{$GUH?EKw9^v<3Bz>sP9tomf{Hi-Ugh!>uOsq#cNg+b6oEp>}=;&Ee>3L%(ytsCb zDE572(Lhn)30-WZq_$4qLdx#`DiN9*#Q_c6O@OHVc=mA&1E_~l6oPr6*7`cIG);Od zj9gsL<8@W8yLWteK(U^_e3LY#GL@%8*4GafEPG8fN}E{tRPx#XiQI2AIKq->5~Vqz!ocngiL zFoJl?5JUYn^CVKaiyBb`<;#4e(1#+cTL0|Zhv}Cg$?weh?_@q_xL+)wHP>ho9WMcF zT$^huhCRoUlk55fM0wu3lu;nS7!BZr>eFDnI<7De$>~hW*CEh$!(>zEt+Uppzh`yu zt}1TYu{mFzH4a*w7JF@g5~Jp9Zir~kf`nXdcpB^DW>FGUM{2ipj1hPCW+{X2kq<6* z=4D1bII~8txTM!;T_#>AhA4M0di>}f#;Z63PtfG-CD$lP@~C9q;17-WG78DJ*utE+ z_8cX76eorfm(uX!mPT6pvcs&&e#pn~>6KNd$yHA%BLz0IDF?g>HOcrW4|(}qIL6>U z5CPQ*@=*l)#LJ+UAud`6_blPdyF1s8f@h>EFOWkW&OgN0l}Tb*$qwg{U%=?sIOkv)*=Dhb?4X?@v$@j2Le{(UVdZCjQ~AqG{9Z`Iw8 z0oyk^UX-5FMutE~OZLnS+e5CQw1rquHzu}x7VdC8HJOhWtIkBvcP^0BBpmCjWk_`J|&Ov?76vBa);WE6#caMEz@bwxMrn_G_uN zJ1e&heOT|t193_nVK^Dz==(gjYT)(Nd~AUy+uPj+29LTwVv51M?{hnjzYJNvgBku9 zG;eMBF`Ik#4z2g4redbV_i#HE5aCg{ufRdf15U z@^uYwX_|V6{e?Wf!ul2cmIqJQE9Vq?c(W zMvU@Bi6k3)fjZBjtew4pPTRL8Jm#=SnJ-YV9qtX_7DMFe-S1S0x2|j7PfzRoijz01 zEnpt7hS!JXH1_;S1dHG@dV1=~)X!sUZDwF($f{bXSK1iV;2yUGZBO$nnUM`NP_J!b zSkXWqvOx}cht4tP9Ym%@KVqC($_*HgUdMgJPuo&|SeRgR-{3f*P z(d1E}<^;KC*()`B@h$j90;q`_K8YcsdQGKdB*)nm-T=A1SJY)lsw+tZmm&6WTO|AJ zsCLBE+$u2)Qc5?<6wVPQBrX{*iMZhusFVh}-0qOktnagS7|!k?6D73hPJ#GZ_Ydj^ zill^ZC@0DbBa!H|zH=rqp2J!32<70O(RLVLPyYCEpK@oGwmQnd?1K1#u8Ih$O7%Ft z0{U_Q%a;)V68{*no{2aUw-%?At{MtydlG@L=!NmfWf7(wZXSRx3HQR|W%b(Zk~x`U*ogNEI#wn!^K zXUASW`O#h%Ak+AGf@)%|zkMkj%#s;Nt$z7NL9KTz1ryI9Tp9_#Gyiy7=y`}rOQs$} zN=@I5mgl>fzS$%Fj2^O}s|J@!7%)g`XSoMoItz_7X88M&cpSq~ukw-Ah5dAQlalq_ zq;>kiH`Q&++8``JUiE`)DMm4Gs+kr9;TIn7N;ps5IatXYkItNz3bB&%eJ%!1O$yDO zC%%C)StF~>P=(6h9Ym_8|7fXR%Brl)E!Re&ImzpNb0pFA{`e()a=L*dVPJRBQb!zu zefE^fjbYXbLxj}PLfc+NdXzFoZGES2>dSKWVU_eS-uzT8C9cu29jp~bUE?sB>gU#j zo@?w`ouSfjc?EZ=n%GwbRNdwBba}ZWae5bZ{I4k2$a&i>w%l3k7Ho^$!~;2k`Nrh3 zX!cPUH;X7N)WdJ~j*)^gstH%$>zwsrAecJyw6ca_&!-vYI*93HZa>*%VqbJ|wnGL{ zjpqh9YMImgCivd2{*V=mluO0pH!v7ro<%$SX6Hn~d5QsTEJny|7wyr)Nkv4B!){&9 zf5ef_eq2X(Q>Su$8EiTPtw(cJlv~`#&3fkjDE)@5AL$EKr9ZcVSiF|roT0|EIOV$h zH}kfH=BztgxOhoz8b!9t1{F(d$ za#k({Dz#ai)Ynax7LB-Hkor?2WFN!=xCLJ*eBTNVH4rjeuk3!14B-mM42f@od(*f0 zfI-y+Hg^WyRT=W@HY%;^&adaxNWW4jiPJI4M0#Okaqi7cFwFP(GN5-!`7+gfmC zC-LRFX@i(wpSS#PpAnw20e-VWktAy@Iwq2qa%hspSj5{m_mx&DDbGm!E}f$PR<3dX zN4W-KXXE+{2K`%d1!Vbu7ToH^cG&a)ko#P-x;C7-gM%zvK1zLe9{eh3F%P%enJAh3a0( za|E#e{fx*Tl0QZuCpUnRl^ghTUD<#D5G(ti%0Fo*BP)oBu#K4&7{Kxy5q32Br5yYj z^4t71j$ih#hyTzD8reyKA!g?P2~lwZTdM=OS^hSp{xW`z;OFiE1GxT69B}I%J!M~-E|F_9Q92}iQ&5axYY^;B1{*3;#fk3Xm2lSuL0)YR~M~d70 z+}98rGXS#KZ05%L%C+p7UFbuxGSQ#<*;y*XI^=kNco_rG(qa&iB@bN#;m)Bnr! z%l^IcKl^K|W&`wTAhFI$Y0u0$KRchg_!>8!!P-%^55!F=U-COeR#@xxt&XI?i`8R6jU}R?wF>z$Fbugo){Uhzq!{Y?8 zwGsOzp`{VyVPye=Sh#@PKrR+`AR9f9g_?zh_Mb)B{xxh-qo3Q`*6dG3?9VCuDTw{g zLv9Y>cZe(iAUh(<-!A|s2L~Gm;2q!(49E)PVh8-;vH25Y0kW~N{0n1c1^zeo6Xp0< z91z6)U)az6$@%Ym9PIy&1ODpme`9R_T>}VY2mL#a<6mR>Ib=4_zc4lq&i`Oe4n{vu zu*2`Jsz5xzzk2msz-w#!bJ5xUK0wko?|!Z=>tFqKbTV>q`rTb1D?1Aak&;qeK?3pr E0C;^KcmMzZ diff --git a/thys/ConcurrentGC/document/root.bib b/thys/ConcurrentGC/document/root.bib --- a/thys/ConcurrentGC/document/root.bib +++ b/thys/ConcurrentGC/document/root.bib @@ -1,567 +1,145 @@ @STRING{cacm="Communications of the ACM"} -@TechReport{LynchTuttle:1989, - author = "Lynch, N. A and Tuttle, M. R.", - title = "An introduction to - input/output automata", - institution = "CWI", - year = 1989, - month = sep, - annote = "CWI Quarterly 2 (3): 219–246", - url = "http://www.markrtuttle.com/data/papers/lt89-cwi.pdf" -} - -@book{Lynch:1996, - author = {N. A. Lynch}, - title = {Distributed Algorithms}, - year = {1996}, - publisher = {Morgan Kaufmann Publishers Inc.}, -} - -@InProceedings{PrenEsp00, - author = {Prensa Nieto, L. and J. Esparza}, - title = {Verifying Single and Multi-mutator Garbage Collectors - with {Owicki/Gries} in {Isabelle/HOL}}, - booktitle = {Mathematical Foundations of Computer Science (MFCS 2000)}, - editor = {M. Nielsen and B. Rovan}, - publisher = {Springer-Verlag}, - series = {LNCS}, - volume = 1893, - pages = {619--628}, - year = 2000 -} - -@PhdThesis{Prensa-PhD,author={Prensa Nieto, L.}, -title={Verification of Parallel Programs with the Owicki-Gries and -Rely-Guarantee Methods in Isabelle/HOL}, -school={Technische Universit{\"a}t M{\"u}nchen},year=2002} - -@inproceedings{Prensa-ESOP03,author={Prensa Nieto, L.}, -title={The {Rely-Guarantee} Method in {Isabelle/HOL}}, -booktitle={European Symposium on Programming (ESOP'03)},editor={P. Degano}, -publisher=Springer,series=LNCS,volume=2618,pages={348--362},year=2003} - -@book{Winskel:1993, - author = {Winskel, G.}, - title = {The Formal Semantics of Programming Languages}, - year = {1993}, - publisher = {MIT Press}, - address = "Cambridge, MA" -} - @article{DBLP:journals/cacm/DijkstraLMSS78, author = {E. W. Dijkstra and L. Lamport and A. J. Martin and C. S. Scholten and E. F. M. Steffens}, title = {On-the-Fly Garbage Collection: An Exercise in Cooperation}, journal = cacm, volume = {21}, number = {11}, year = {1978}, pages = {966-975}, ee = {http://doi.acm.org/10.1145/359642.359655}, bibsource = {DBLP, http://dblp.uni-trier.de} } @inproceedings{DBLP:conf/lpar/Schirmer04, author = {N. Schirmer}, title = {A Verification Environment for Sequential Imperative Programs in Isabelle/HOL}, booktitle = {LPAR}, year = {2004}, pages = {398-414}, ee = {https://doi.org/10.1007/978-3-540-32275-7_26}, crossref = {DBLP:conf/lpar/2004}, bibsource = {DBLP, http://dblp.uni-trier.de} } % isbn = {3-540-25236-3}, @proceedings{DBLP:conf/lpar/2004, editor = {F. Baader and A. Voronkov}, title = {Logic for Programming, Artificial Intelligence, and Reasoning, 11th International Conference, LPAR 2004, Montevideo, Uruguay, March 14-18, 2005, Proceedings}, booktitle = {LPAR}, publisher = {Springer}, series = lncs, volume = {3452}, year = {2005}, bibsource = {DBLP, http://dblp.uni-trier.de} } @article{DBLP:journals/cacm/SewellSONM10, author = {P. Sewell and S. Sarkar and S. Owens and F. Zappa Nardelli and M. O. Myreen}, title = {{x86-TSO}: a rigorous and usable programmer's model for x86 multiprocessors}, journal = cacm, volume = {53}, number = {7}, year = {2010}, pages = {89-97}, ee = {http://doi.acm.org/10.1145/1785414.1785443}, bibsource = {DBLP, http://dblp.uni-trier.de} } -@inproceedings{DBLP:conf/mfcs/BoerRH99, - author = {de Boer, F. S. and - de Roever, W. P. and - U. Hannemann}, - title = {The Semantic Foundations of a Compositional Proof Method - for Synchronously Communicating Processes}, - booktitle = {MFCS}, - year = {1999}, - pages = {343-353}, - ee = {https://doi.org/10.1007/3-540-48340-3_31}, - crossref = {DBLP:conf/mfcs/1999}, - bibsource = {DBLP, http://dblp.uni-trier.de} -} - -% isbn = {3-540-66408-4}, -@proceedings{DBLP:conf/mfcs/1999, - editor = {M. Kutylowski and - L. Pacholski and - T. Wierzbicki}, - title = {Mathematical Foundations of Computer Science 1999, 24th - International Symposium, MFCS'99, Szklarska Poreba, Poland, - September 6-10, 1999, Proceedings}, - booktitle = {MFCS}, - publisher = {Springer}, - series = lncs, - volume = {1672}, - year = {1999}, - bibsource = {DBLP, http://dblp.uni-trier.de} -} - -% isbn = {0-521-80608-9}, -@book{DBLP:books/cu/RoeverBH2001, - author = {de Roever, W. P. and - de Boer, F. S. and - U. Hannemann and - J. Hooman and - Y. Lakhnech and - M. Poel and - J. Zwiers}, - title = {Concurrency Verification: Introduction to Compositional - and Noncompositional Methods}, - publisher = {Cambridge University Press}, - series = {Cambridge Tracts in Theoretical Computer Science}, - volume = {54}, - year = {2001}, - bibsource = {DBLP, http://dblp.uni-trier.de} -} - @InProceedings{Pizlo+2010PLDI, author = {F. Pizlo and L. Ziarek and P. Maj and A. L. Hosking and E. Blanton and J. Vitek}, title = {Schism: Fragmentation-Tolerant Real-Time Garbage Collection}, booktitle = {Proceedings of the ACM SIGPLAN Conference on Programming Language Design and Implementation}, year = 2010, pages = "146--159", month = jun, address = {Toronto, Canada}, doi = {10.1145/1806596.1806615} } @PhdThesis{Pizlo201xPhd, author = {F. Pizlo}, title = {Fragmentation Tolerant Real Time Garbage Collection}, school = {Purdue University}, year = {201x} } -@article{DBLP:journals/toplas/LamportS84, - author = {L. Lamport and - F. B. Schneider}, - title = {The ``{Hoare Logic}'' of {CSP}, and All That}, - journal = toplas, - volume = {6}, - number = {2}, - year = {1984}, - pages = {281-296}, - ee = {http://doi.acm.org/10.1145/2993.357247}, - bibsource = {DBLP, http://dblp.uni-trier.de} -} - -@article{DBLP:journals/acta/Lamport80, - author = {L. Lamport}, - title = {The ``{Hoare Logic}'' of Concurrent Programs}, - journal = {Acta Informatica}, - volume = {14}, - year = {1980}, - pages = {21-37}, - ee = {https://doi.org/10.1007/BF00289062}, - bibsource = {DBLP, http://dblp.uni-trier.de} -} - -@article{DBLP:journals/jar/Berghofer12, - author = {S. Berghofer}, - title = {A Solution to the {PoplMark} Challenge Using {de Bruijn} Indices - in {Isabelle/HOL}}, - journal = {J. Autom. Reasoning}, - volume = {49}, - number = {3}, - year = {2012}, - pages = {303-326}, - ee = {https://doi.org/10.1007/s10817-011-9231-4}, - bibsource = {DBLP, http://dblp.uni-trier.de} -} - -% isbn = {3-540-44044-5} -@InCollection{PittsAM:opespe, - author = {A. M. Pitts}, - title = {Operational Semantics and Program Equivalence}, - booktitle = {Applied Semantics, Advanced Lectures}, - pages = {378--412}, - publisher = {Springer-Verlag}, - year = 2002, - editor = {G. Barthe and P. Dybjer and J. Saraiva}, - volume = 2395, - series = lncs, - note = {International Summer School, {APPSEM} 2000, Caminha, - Portugal, September 9--15, 2000}, -} - -% isbn = {978-0-387-94459-3}, -@book{DBLP:books/daglib/0080029, - author = {Z. Manna and - A. Pnueli}, - title = {Temporal verification of reactive systems - safety}, - publisher = {Springer}, - year = {1995}, - pages = {I-XV, 1-512}, - bibsource = {DBLP, http://dblp.uni-trier.de} -} - -@article{DBLP:journals/acta/LevinG81, - author = {G. Levin and - D. Gries}, - title = {A Proof Technique for Communicating Sequential Processes}, - journal = {Acta Inf.}, - volume = {15}, - year = {1981}, - pages = {281-302}, - ee = {https://doi.org/10.1007/BF00289266}, - bibsource = {DBLP, http://dblp.uni-trier.de} -} - -@inproceedings{DBLP:conf/icalp/CousotC80, - author = {P. Cousot and - R. Cousot}, - title = "Semantic Analysis of {Communicating Sequential Processes} - (Shortened Version)", - booktitle = {ICALP}, - year = {1980}, - pages = {119-133}, - ee = {https://doi.org/10.1007/3-540-10003-2_65}, - crossref = {DBLP:conf/icalp/1980}, - bibsource = {DBLP, http://dblp.uni-trier.de} -} - -% isbn = {3-540-10003-2}, -@proceedings{DBLP:conf/icalp/1980, - editor = {de Bakker, J. W. and van Leeuwen, J.}, - title = {Automata, Languages and Programming, 7th Colloquium, Noordweijkerhout, - The Netherland, July 14-18, 1980, Proceedings}, - booktitle = {ICALP}, - publisher = {Springer}, - series = lncs, - volume = {85}, - year = {1980}, - bibsource = {DBLP, http://dblp.uni-trier.de} -} - -@InCollection{MannaPnueli:1991, - author = "Z. Manna and A. Pnueli", - title = "Tools and rules for the practicing verifier", - booktitle = "CMU Computer Science: A 25th Anniversary Commemorative", - pages = "121–-156", - publisher = "ACM Press and Addison-Wesley", - year = 1991, - editor = "R. F. Rashid" -} - -% isbn = {0-3211-4306-X}, -@book{DBLP:books/aw/Lamport2002, - author = {L. Lamport}, - title = {Specifying Systems, The TLA+ Language and Tools for Hardware - and Software Engineers}, - publisher = {Addison-Wesley}, - year = {2002}, - bibsource = {DBLP, http://dblp.uni-trier.de} -} - -@article{TLA-AFP, - author = {G. Grov and S. Merz}, - title = {A Definitional Encoding of TLA* in Isabelle/HOL}, - journal = {Archive of Formal Proofs}, - month = nov, - year = 2011, - note = {\url{http://isa-afp.org/entries/TLA}, - Formal proof development}, - ISSN = {2150-914x}, -} - -@inproceedings{DBLP:conf/itp/CohenS10, - author = {E. Cohen and - B. Schirmer}, - title = {From Total Store Order to Sequential Consistency: A Practical - Reduction Theorem}, - booktitle = {ITP}, - year = {2010}, - pages = {403-418}, - ee = {https://doi.org/10.1007/978-3-642-14052-5_28}, - crossref = {DBLP:conf/itp/2010}, - bibsource = {DBLP, http://dblp.uni-trier.de} -} - -% isbn = {978-3-642-14051-8}, -@proceedings{DBLP:conf/itp/2010, - editor = {M. Kaufmann and - L. C. Paulson}, - title = {Interactive Theorem Proving, First International Conference, - ITP 2010, Edinburgh, UK, July 11-14, 2010. Proceedings}, - booktitle = {ITP}, - publisher = {Springer}, - series = lncs, - volume = {6172}, - year = {2010}, - ee = {https://doi.org/10.1007/978-3-642-14052-5}, - bibsource = {DBLP, http://dblp.uni-trier.de} -} - -@inproceedings{DBLP:conf/fase/NipkowN99, - author = {T. Nipkow and - Prensa Nieto, L.}, - title = "{Owicki/Gries} in {Isabelle/HOL}", - booktitle = {FASE}, - year = {1999}, - pages = {188-203}, - ee = {https://doi.org/10.1007/978-3-540-49020-3_13}, - crossref = {DBLP:conf/fase/1999}, - bibsource = {DBLP, http://dblp.uni-trier.de} -} - -% isbn = {3-540-65718-5}, -@proceedings{DBLP:conf/fase/1999, - editor = {J.-P. Finance}, - title = {Fundamental Approaches to Software Engineering, Second Internationsl - Conference, FASE'99, Held as Part of the European Joint - Conferences on the Theory and Practice of Software, ETAPS'99, - Amsterdam, The Netherlands, March 22-28, 1999, Proceedings}, - booktitle = {FASE}, - publisher = {Springer}, - series = lncs, - volume = {1577}, - year = {1999}, - bibsource = {DBLP, http://dblp.uni-trier.de} -} - -@article{CousotCousot89-IC, - author = {Cousot, P. and Cousot, R.}, - title = {A language independent proof of the soundness - and completeness of generalized {H}oare logic}, - journal = {Information and Computation}, - volume = 80, - number = 2, - pages = {165--191}, - month = feb, - year = 1989, -} - -@inproceedings{DBLP:conf/popl/DoligezG94, +@inproceedings{DoligezGonthier:1994, author = {D. Doligez and G. Gonthier}, title = {Portable, Unobtrusive Garbage Collection for Multiprocessor Systems}, - booktitle = {POPL}, + booktitle = {POPL'1994}, year = {1994}, pages = {70-83}, - ee = {http://doi.acm.org/10.1145/174675.174673}, - crossref = {DBLP:conf/popl/1994}, - bibsource = {DBLP, http://dblp.uni-trier.de} -} - -% isbn = {0-89791-636-0}, -@proceedings{DBLP:conf/popl/1994, - editor = {Boehm, H.-J. and - B. Lang and - D. M. Yellin}, - title = {Conference Record of POPL'94: 21st ACM SIGPLAN-SIGACT Symposium - on Principles of Programming Languages, Portland, Oregon, - USA, January 17-21, 1994}, - booktitle = popl, publisher = {ACM Press}, - year = {1994}, - ee = {http://dl.acm.org/citation.cfm?id=174675}, - bibsource = {DBLP, http://dblp.uni-trier.de} -} - -@article{DBLP:journals/jlp/Plotkin04, - author = {G. D. Plotkin}, - title = {The origins of structural operational semantics}, - journal = {Journal of Logic and Algebraic Programming}, - volume = {60-61}, - year = {2004}, - pages = {3-15}, - ee = {https://doi.org/10.1016/j.jlap.2004.03.009}, - bibsource = {DBLP, http://dblp.uni-trier.de} -} - -@Book{ConcreteSemantics:2014, - author = "T. Nipkow and G. Klein", - title = "Concrete Semantics: A Proof Assistant Approach", - publisher = Springer, - year = 2014, - url = "http://www.in.tum.de/~nipkow/Concrete-Semantics/" -} - -@Book{Hoare:1985, - author = {C.A.R. Hoare}, - title = "{Communicating Sequential Processes}", - publisher = {Prentice-Hall}, - year = 1985, - series = {International Series In Computer Science}, - url = "http://www.usingcsp.com/" -} - -@Book{Milner:1980, - author = "R. Milner", - title = "A Calculus of Communicating Systems", - publisher = Springer, - year = 1980 -} - -@book{Milner:1989, - author = {R. Milner}, - title = "Communication and Concurrency", - year = {1989}, - publisher = {Prentice-Hall}, - address = "Englewood Cliffs, NJ", -} - -@inproceedings{DBLP:conf/vstte/Ridge10, - author = {Tom Ridge}, - title = {A Rely-Guarantee Proof System for x86-TSO}, - booktitle = {VSTTE}, - year = {2010}, - pages = {55-70}, - ee = {https://doi.org/10.1007/978-3-642-15057-9_4}, - crossref = {DBLP:conf/vstte/2010}, - bibsource = {DBLP, http://dblp.uni-trier.de} -} - -% isbn = {978-3-642-15056-2}, -@proceedings{DBLP:conf/vstte/2010, - editor = {Gary T. Leavens and - Peter W. O'Hearn and - Sriram K. Rajamani}, - title = {Verified Software: Theories, Tools, Experiments, Third International - Conference, VSTTE 2010, Edinburgh, UK, August 16-19, 2010. - Proceedings}, - booktitle = {VSTTE}, - publisher = {Springer}, - series = {Lecture Notes in Computer Science}, - volume = {6217}, - year = {2010}, - ee = {https://doi.org/10.1007/978-3-642-15057-9}, - bibsource = {DBLP, http://dblp.uni-trier.de} -} - -@Article{FelleisenHieb:1992:TCS, - author = {M. Felleisen and R. Hieb}, - title = {The revised report on the syntactic theories of sequential - control and state}, - journal = {Theoretical Computer Science}, - year = 1992, - volume = 103, - number = 2, - pages = {235--271}, - month = sep, - doi = {10.1016/0304-3975(92)90014-7} + doi = {10.1145/174675.174673}, } @article{DBLP:journals/entcs/SchirmerW09, author = {N. Schirmer and M. Wenzel}, title = {State Spaces - The Locale Way}, journal = {Electr. Notes Theor. Comput. Sci.}, volume = {254}, year = {2009}, pages = {161-179}, ee = {https://doi.org/10.1016/j.entcs.2009.09.065}, bibsource = {DBLP, http://dblp.uni-trier.de} } -% url = {http://doi.acm.org/10.1145/1480881.1480934}, -@inproceedings{DBLP:conf/popl/Ridge09, - author = {T. Ridge}, - title = {Verifying distributed systems: the operational approach}, - booktitle = {Proceedings of the 36th {ACM} {SIGPLAN-SIGACT} Symposium on Principles - of Programming Languages, {POPL} 2009, Savannah, GA, USA, January - 21-23, 2009}, - pages = {429--440}, - year = {2009}, - crossref = {DBLP:conf/popl/2009}, - doi = {10.1145/1480881.1480934}, -} - -% isbn = {978-1-60558-379-2}, -@proceedings{DBLP:conf/popl/2009, - editor = {Z. Shao and - B. C. Pierce}, - title = {Proceedings of the 36th {ACM} {SIGPLAN-SIGACT} Symposium on Principles - of Programming Languages, {POPL} 2009, Savannah, GA, USA, January - 21-23, 2009}, - publisher = {{ACM}}, - year = {2009}, - url = {http://dl.acm.org/citation.cfm?id=1480881}, -} - % url = {https://doi.org/10.1007/978-3-642-03359-9_27}, @inproceedings{DBLP:conf/tphol/OwensSS09, author = {S. Owens and S. Sarkar and P. Sewell}, - title = {A Better x86 Memory Model: x86-TSO}, + title = {A Better x86 Memory Model: {x86-TSO}}, booktitle = {Theorem Proving in Higher Order Logics, 22nd International Conference, TPHOLs 2009, Munich, Germany, August 17-20, 2009. Proceedings}, pages = {391--407}, year = {2009}, crossref = {DBLP:conf/tphol/2009}, doi = {10.1007/978-3-642-03359-9_27}, } % isbn = {978-3-642-03358-2}, % url = {https://doi.org/10.1007/978-3-642-03359-9}, @proceedings{DBLP:conf/tphol/2009, editor = {S. Berghofer and T. Nipkow and C. Urban and M. Wenzel}, title = {Theorem Proving in Higher Order Logics, 22nd International Conference, TPHOLs 2009, Munich, Germany, August 17-20, 2009. Proceedings}, series = {Lecture Notes in Computer Science}, volume = {5674}, publisher = {Springer}, year = {2009}, doi = {10.1007/978-3-642-03359-9}, } - @article{ConcurrentIMP_AFP, - author = {Peter Gammie}, + author = {P. Gammie}, title = {Concurrent {IMP}}, journal = {Archive of Formal Proofs}, month = apr, year = 2015, note = {\url{http://isa-afp.org/entries/ConcurrentIMP.shtml}, Formal proof development}, ISSN = {2150-914x}, -} \ No newline at end of file +} diff --git a/thys/ConcurrentGC/document/root.tex b/thys/ConcurrentGC/document/root.tex --- a/thys/ConcurrentGC/document/root.tex +++ b/thys/ConcurrentGC/document/root.tex @@ -1,97 +1,86 @@ \documentclass[11pt,a4paper]{article} \usepackage[a4paper,margin=1cm,footskip=.5cm]{geometry} \usepackage{isabelle,isabellesym} - \usepackage{amssymb} \usepackage[english]{babel} -% FIXME lifted composition. -\newcommand{\isasymbigcirc}{\isamath{\circ}} - -% Bibliography +\usepackage{graphicx} \usepackage[authoryear,longnamesfirst,sort]{natbib} \bibpunct();A{}, % this should be the last package used \usepackage{pdfsetup} % urls in roman style, theory text in math-similar italics \urlstyle{rm} \isabellestyle{it} \renewcommand{\ttdefault}{cmtt} % CM rather than courier for \tt % for uniform font size \renewcommand{\isastyle}{\isastyleminor} % Abstract various things that might change. \newcommand{\ccode}[1]{\texttt{#1}} \newcommand{\isabelletype}[1]{\emph{#1}} \newcommand{\isabelleterm}[1]{\emph{#1}} % sane default for proof documents \parindent 0pt\parskip 0.5ex \begin{document} \title{Relaxing Safely: Verified On-the-Fly Garbage Collection for x86-TSO} \author{Peter Gammie, Tony Hosking and Kai Engelhardt} \maketitle \begin{abstract} - CIMP extends the small imperative language IMP with synchronous - message passing. - - We use CIMP to model Schism, a state-of-the-art real-time garbage + We model an instance of Schism, a state-of-the-art real-time garbage collection scheme for weak memory, and show that it is safe on x86-TSO. \end{abstract} \tableofcontents \section{Introduction} \label{sec:introduction} We verify the memory safety of one of the Schism garbage collectors as developed by \citet{Pizlo+2010PLDI,Pizlo201xPhd} with respect to the x86-TSO model (a total store order memory model for modern multicore Intel x86 architectures) developed and validated by \citet{DBLP:journals/cacm/SewellSONM10}. Our development is inspired by the original work on the verification of concurrent mark/sweep collectors by \citet{DBLP:journals/cacm/DijkstraLMSS78}, and the more realistic -models and proofs of \citet{DBLP:conf/popl/DoligezG94}. We leave a -thorough survey of formal garbage collection verification to future -work. +models and proofs of \citet{DoligezGonthier:1994}. We leave a thorough +survey of formal garbage collection verification to future work. -We present our model of the -garbage collector in \S\ref{sec:gc-model}, the detailed invariants in -\S\ref{sec:gc-invs}, and the high-level safety results in -\S\ref{sec:top-level-correctness}. This bottom-up presentation is how -we developed the proof; we have resisted the urge to hide the bodies -with a rational reconstruction, partly because we expect the current -structure to more readily support extensions and revisions. +We present our model of the garbage collector in \S\ref{sec:gc-model}, +the predicates we use in our assertions in \S\ref{sec:proofs-basis}, +the detailed invariants in \S\ref{sec:global-invariants} and +\S\ref{sec:local-invariants}, and the high-level safety results in +\S\ref{sec:top-level-correctness}. A concrete system state that +satisfies our invariants is exhibited in +\S\ref{sec:concrete-system-state}. The other sections contain the +often gnarly proofs and lemmas starring in supporting roles. -This document does not include the formal proofs that the model -satisfies these invariants; the curious reader is encouraged to peruse -the Isabelle sources. - -For details about the modelling language CIMP used in this development, -see the separate AFP entry ConcurrentIMP \citep{ConcurrentIMP_AFP}. +The modelling language CIMP used in this development is described in +the AFP entry ConcurrentIMP \citep{ConcurrentIMP_AFP}. % generated text of all theories \input{session} \bibliographystyle{plainnat} \bibliography{root} \addcontentsline{toc}{section}{References} \end{document} %%% Local Variables: %%% mode: latex %%% TeX-master: t %%% End: diff --git a/thys/ConcurrentIMP/CIMP.thy b/thys/ConcurrentIMP/CIMP.thy --- a/thys/ConcurrentIMP/CIMP.thy +++ b/thys/ConcurrentIMP/CIMP.thy @@ -1,134 +1,40 @@ (*<*) (* * 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 CIMP imports CIMP_pred CIMP_lang CIMP_vcg - "HOL-Library.Sublist" + CIMP_vcg_rules keywords - "inv_definition" "locset_definition" :: thy_defn + "locset_definition" :: thy_defn + and "intern_com" :: thy_decl begin text\ Infrastructure for reasoning about CIMP programs. See AFP entry \ConcurrentGC\ for examples of use. \ -named_theorems com "Command definitions" -named_theorems inv "Location-sensitive invariant definitions" -named_theorems loc "Location set membership cache" - -ML\ - -signature CIMP = -sig - val com_locs_fold : (term -> 'b -> 'b) -> 'b -> term -> 'b - val com_locs_map : (term -> 'b) -> term -> 'b list - val com_locs_fold_no_response : (term -> 'b -> 'b) -> 'b -> term -> 'b - val com_locs_map_no_response : (term -> 'b) -> term -> 'b list - val def_inv : thm -> local_theory -> local_theory - val def_locset : thm -> local_theory -> local_theory -end; - -structure Cimp : CIMP = -struct - -fun com_locs_fold f x (Const (@{const_name Request}, _) $ l $ _ $ _ ) = f l x - | com_locs_fold f x (Const (@{const_name Response}, _) $ l $ _) = f l x - | com_locs_fold f x (Const (@{const_name LocalOp}, _) $ l $ _) = f l x - | com_locs_fold f x (Const (@{const_name Cond1}, _) $ l $ _ $ c) = com_locs_fold f (f l x) c - | com_locs_fold f x (Const (@{const_name Cond2}, _) $ l $ _ $ c1 $ c2) = com_locs_fold f (com_locs_fold f (f l x) c1) c2 - | com_locs_fold f x (Const (@{const_name Loop}, _) $ c) = com_locs_fold f x c - | com_locs_fold f x (Const (@{const_name While}, _) $ l $ _ $ c) = com_locs_fold f (f l x) c - | com_locs_fold f x (Const (@{const_name Seq}, _) $ c1 $ c2) = com_locs_fold f (com_locs_fold f x c1) c2 - | com_locs_fold f x (Const (@{const_name Choose}, _) $ c1 $ c2) = com_locs_fold f (com_locs_fold f x c1) c2 - | com_locs_fold _ x _ = x; - -fun com_locs_map f com = com_locs_fold (fn l => fn acc => f l :: acc) [] com - -fun com_locs_fold_no_response f x (Const (@{const_name Request}, _) $ l $ _ $ _ ) = f l x - | com_locs_fold_no_response _ x (Const (@{const_name Response}, _) $ _ $ _) = x (* can often ignore \Response\ *) - | com_locs_fold_no_response f x (Const (@{const_name LocalOp}, _) $ l $ _) = f l x - | com_locs_fold_no_response f x (Const (@{const_name Cond1}, _) $ l $ _ $ c) = com_locs_fold_no_response f (f l x) c - | com_locs_fold_no_response f x (Const (@{const_name Cond2}, _) $ l $ _ $ c1 $ c2) = com_locs_fold_no_response f (com_locs_fold_no_response f (f l x) c1) c2 - | com_locs_fold_no_response f x (Const (@{const_name Loop}, _) $ c) = com_locs_fold_no_response f x c - | com_locs_fold_no_response f x (Const (@{const_name While}, _) $ l $ _ $ c) = com_locs_fold_no_response f (f l x) c - | com_locs_fold_no_response f x (Const (@{const_name Seq}, _) $ c1 $ c2) = com_locs_fold_no_response f (com_locs_fold_no_response f x c1) c2 - | com_locs_fold_no_response f x (Const (@{const_name Choose}, _) $ c1 $ c2) = com_locs_fold_no_response f (com_locs_fold_no_response f x c1) c2 - | com_locs_fold_no_response _ x _ = x; - -fun com_locs_map_no_response f com = com_locs_fold_no_response (fn l => fn acc => f l :: acc) [] com - -(* Cache location set membership facts. - -Decide membership in the given set for each label in the CIMP programs -in the Named_Theorems "com". +named_theorems locset_cache "Location set membership cache" -If the label set and com types differ, we probably get a nasty error. - -*) +lemmas cleanup_simps = + atomize_eq + simp_thms -fun def_locset thm ctxt = - let - val set_name = thm - |> Local_Defs.meta_rewrite_rule ctxt (* handle `=` or `\` *) - |> Thm.cprop_of |> Thm.dest_equals |> fst |> Thm.term_of - val set_typ = set_name |> type_of - val elt_typ = case set_typ of Type ("Set.set", [t]) => t | _ => raise Fail "thm should define a set" - val set_name_str = case set_name of Const (c, _) => c | Free (c, _) => c | _ => error ("Not an equation of the form x = set: " ^ Thm.string_of_thm ctxt thm) - val thm_name = Binding.qualify true set_name_str (Binding.name "memb") - fun mk_memb l = Thm.cterm_of ctxt (Const (@{const_name "Set.member"}, elt_typ --> set_typ --> @{typ "bool"}) $ l $ set_name) - val rewrite_tac = Simplifier.rewrite (ctxt addsimps ([thm] @ Named_Theorems.get ctxt @{named_theorems "loc"})) (* probably want the ambient simpset + some stuff *) - val coms = Named_Theorems.get ctxt @{named_theorems "com"} |> map (Thm.cprop_of #> Thm.dest_equals #> snd #> Thm.term_of) - val attrs = [(* Attrib.internal (K (Clasimp.iff_add)), *) Attrib.internal (K (Named_Theorems.add @{named_theorems "loc"}))] -(* Parallel *) - fun mk_thms coms = Par_List.map rewrite_tac (maps (com_locs_map_no_response mk_memb) coms) -(* Sequential *) -(* fun mk_thms coms = List.foldl (fn (c, thms) => com_locs_fold (fn l => fn thms => rewrite_tac (mk_memb l) :: thms) thms c) [] coms *) - in - ctxt - |> Local_Theory.note ((thm_name, attrs), mk_thms coms) - |> snd - end; - -(* FIXME later need to rewrite using interned labels (fold defs). *) -fun def_inv thm ctxt : local_theory = - let - val attrs = [Attrib.internal (K (Named_Theorems.add @{named_theorems "inv"}))] - in - ctxt - |> Local_Theory.note ((Binding.empty, attrs), [thm]) - |> snd - end; - -end; - -val _ = - Outer_Syntax.local_theory' \<^command_keyword>\inv_definition\ "constant definition for invariants" - (Scan.option Parse_Spec.constdecl -- (Parse_Spec.opt_thm_name ":" -- Parse.prop) -- - Parse_Spec.if_assumes -- Parse.for_fixes >> (fn (((decl, spec), prems), params) => fn b => fn lthy => - Specification.definition_cmd decl params prems spec b lthy - |> (fn ((_, (_, thm)), lthy) => (thm, lthy)) |> uncurry Cimp.def_inv)); - -val _ = - Outer_Syntax.local_theory' \<^command_keyword>\locset_definition\ "constant definition for sets of locations" - (Scan.option Parse_Spec.constdecl -- (Parse_Spec.opt_thm_name ":" -- Parse.prop) -- - Parse_Spec.if_assumes -- Parse.for_fixes >> (fn (((decl, spec), prems), params) => fn b => fn lthy => - Specification.definition_cmd decl params prems spec b lthy - |> (fn ((_, (_, thm)), lthy) => (thm, lthy)) |> uncurry Cimp.def_locset)); -\ +ML_file\cimp.ML\ +(*<*) end (*>*) diff --git a/thys/ConcurrentIMP/CIMP_lang.thy b/thys/ConcurrentIMP/CIMP_lang.thy --- a/thys/ConcurrentIMP/CIMP_lang.thy +++ b/thys/ConcurrentIMP/CIMP_lang.thy @@ -1,946 +1,510 @@ (*<*) (* * 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 CIMP_lang imports CIMP_pred + LTL begin (*>*) -section\CIMP syntax and semantics\ +section\CIMP syntax and semantics \label{sec:cimp-syntax-semantics}\ text\ -\label{sec:cimp-syntax-semantics} - We define a small sequential programming language with synchronous message passing primitives for describing the individual processes. This has the advantage over raw transition systems in that it is programmer-readable, includes sequential composition, supports a program logic and VCG (\S\ref{sec:cimp-vcg}), etc. These processes are composed in parallel at the top-level. CIMP is inspired by IMP, as presented by @{cite [cite_macro=citet] "Winskel:1993"} and @{cite [cite_macro=citet] "ConcreteSemantics:2014"}, and the classical process algebras CCS @{cite [cite_macro=citep] "Milner:1980" and "Milner:1989"} and CSP @{cite [cite_macro=citep] "Hoare:1985"}. Note that the algebraic properties of this language have not been developed. As we operate in a concurrent setting, we need to provide a small-step semantics (\S\ref{sec:cimp-semantics}), which we give in the style of \emph{structural operational semantics} (SOS) as popularised by @{cite [cite_macro=citet] "DBLP:journals/jlp/Plotkin04"}. The semantics of a complete system (\S\ref{sec:cimp-system-steps}) is presently taken simply to be the states reachable by interleaving the enabled steps of the individual processes, subject to message passing rendezvous. We leave a trace or branching semantics to future work. +This theory contains all the trusted definitions. The soundness of the other +theories supervenes upon this one. + \ + subsection\Syntax\ text\ Programs are represented using an explicit (deep embedding) of their syntax, as the semantics needs to track the progress of multiple threads of control. Each (atomic) \emph{basic command} (\S\ref{sec:cimp-decompose}) is annotated with a @{typ "'location"}, which we use in our assertions (\S\ref{sec:cimp-control-predicates}). These locations need not be unique, though in practice they likely will be. Processes maintain \emph{local states} of type @{typ "'state"}. These can be updated with arbitrary relations of @{typ "'state \ 'state set"} with \LocalOp\, and conditions of type @{typ "'s \ bool"} are similarly shallowly embedded. This arrangement allows the end-user to select their own level of atomicity. The sequential composition operator and control constructs are standard. We add the infinite looping construct \Loop\ so we can construct single-state reactive systems; this has implications for fairness assertions. \ type_synonym 's bexp = "'s \ bool" datatype ('answer, 'location, 'question, 'state) com = Request "'location" "'state \ 'question" "'answer \ 'state \ 'state set" ("\_\ Request _ _" [0, 70, 70] 71) | Response "'location" "'question \ 'state \ ('state \ 'answer) set" ("\_\ Response _" [0, 70] 71) | LocalOp "'location" "'state \ 'state set" ("\_\ LocalOp _" [0, 70] 71) - | Cond1 "'location" "'state bexp" "('answer, 'location, 'question, 'state) com" ("\_\ IF _ THEN _ FI" [0, 0] 71) + | Cond1 "'location" "'state bexp" "('answer, 'location, 'question, 'state) com" ("\_\ IF _ THEN _ FI" [0, 0, 0] 71) | Cond2 "'location" "'state bexp" "('answer, 'location, 'question, 'state) com" - "('answer, 'location, 'question, 'state) com" ("\_\ IF _/ THEN _/ ELSE _/ FI" [0, 0, 0] 71) + "('answer, 'location, 'question, 'state) com" ("\_\ IF _/ THEN _/ ELSE _/ FI" [0, 0, 0, 0] 71) | Loop "('answer, 'location, 'question, 'state) com" ("LOOP DO _/ OD" [0] 71) | While "'location" "'state bexp" "('answer, 'location, 'question, 'state) com" ("\_\ WHILE _/ DO _/ OD" [0, 0, 0] 71) | Seq "('answer, 'location, 'question, 'state) com" "('answer, 'location, 'question, 'state) com" (infixr ";;" 69) | Choose "('answer, 'location, 'question, 'state) com" - "('answer, 'location, 'question, 'state) com" (infixl "\" 68) + "('answer, 'location, 'question, 'state) com" (infixl "\" 68) text\ We provide a one-armed conditional as it is the common form and avoids the need to discover a label for an internal \SKIP\ and/or trickier proofs about the VCG. In contrast to classical process algebras, we have local state and -distinct send and receive actions. These provide an interface to +distinct request and response actions. These provide an interface to Isabelle/HOL's datatypes that avoids the need for binding (ala the $\pi$-calculus of @{cite [cite_macro=citet] "Milner:1989"}) or large non-deterministic sums (ala CCS @{cite [cite_macro=citep] \\S2.8\ -"Milner:1980"}). Intuitively the sender asks a @{typ "'question"} with +"Milner:1980"}). Intuitively the requester poses a @{typ "'question"} with a \Request\ command, which upon rendezvous with a -receiver's \Response\ command receives an @{typ +responder's \Response\ command receives an @{typ "'answer"}. The @{typ "'question"} is a deterministic function of the -sender's local state, whereas a receiver can respond -non-deterministically. Note that CIMP does not provide a notion of +requester's local state, whereas responses can be +non-deterministic. Note that CIMP does not provide a notion of channel; these can be modelled by a judicious choice of @{typ "'question"}. -We also provide a binary external choice operator. Internal choice can -be recovered in combination with local operations (see @{cite -[cite_macro=citet] \\S2.3\ "Milner:1980"}). +We also provide a binary external choice operator @{term\Choose\} (infix @{term\(\)\}). +Internal choice can be recovered in combination with local operations +(see @{cite [cite_macro=citet] \\S2.3\ "Milner:1980"}). We abbreviate some common commands: \SKIP\ is a local operation that does nothing, and the floor brackets simplify deterministic \LocalOp\s. We also adopt some syntax magic from -Makarius's Hoare and Multiquote theories in the Isabelle/HOL +Makarius's \Hoare\ and \Multiquote\ theories in the Isabelle/HOL distribution. \ -abbreviation SKIP_syn ("\_\/ SKIP" 70) where +abbreviation SKIP_syn ("\_\/ SKIP" [0] 71) where "\l\ SKIP \ \l\ LocalOp (\s. {s})" abbreviation (input) DetLocalOp :: "'location \ ('state \ 'state) - \ ('answer, 'location, 'question, 'state) com" ("\_\ \_\") where + \ ('answer, 'location, 'question, 'state) com" ("\_\ \_\" [0, 0] 71) where "\l\ \f\ \ \l\ LocalOp (\s. {f s})" syntax "_quote" :: "'b \ ('a \ 'b)" ("\_\" [0] 1000) "_antiquote" :: "('a \ 'b) \ 'b" ("\_" [1000] 1000) - "_Assign" :: "'location \ idt \ 'b \ ('answer, 'location, 'question, 'state) com" ("(\_\ \_ :=/ _)" [0, 0, 70] 71) (* FIXME precedence -- see massive ambiguities in the GC model *) + "_Assign" :: "'location \ idt \ 'b \ ('answer, 'location, 'question, 'state) com" ("(\_\ \_ :=/ _)" [0, 0, 70] 71) "_NonDetAssign" :: "'location \ idt \ 'b set \ ('answer, 'location, 'question, 'state) com" ("(\_\ \_ :\/ _)" [0, 0, 70] 71) abbreviation (input) NonDetAssign :: "'location \ (('val \ 'val) \ 'state \ 'state) \ ('state \ 'val set) \ ('answer, 'location, 'question, 'state) com" where "NonDetAssign l upd es \ \l\ LocalOp (\s. { upd \e\ s |e. e \ es s })" translations "\l\ \x := e" => "CONST DetLocalOp l \\(_update_name x (\_. e))\" "\l\ \x :\ es" => "CONST NonDetAssign l (_update_name x) \es\" parse_translation \ let fun antiquote_tr i (Const (@{syntax_const "_antiquote"}, _) $ (t as Const (@{syntax_const "_antiquote"}, _) $ _)) = skip_antiquote_tr i t | antiquote_tr i (Const (@{syntax_const "_antiquote"}, _) $ t) = antiquote_tr i t $ Bound i | antiquote_tr i (t $ u) = antiquote_tr i t $ antiquote_tr i u | antiquote_tr i (Abs (x, T, t)) = Abs (x, T, antiquote_tr (i + 1) t) | antiquote_tr _ a = a and skip_antiquote_tr i ((c as Const (@{syntax_const "_antiquote"}, _)) $ t) = c $ skip_antiquote_tr i t | skip_antiquote_tr i t = antiquote_tr i t; fun quote_tr [t] = Abs ("s", dummyT, antiquote_tr 0 (Term.incr_boundvars 1 t)) | quote_tr ts = raise TERM ("quote_tr", ts); in [(@{syntax_const "_quote"}, K quote_tr)] end \ -subsection\Process semantics\ +subsection\Process semantics \label{sec:cimp-semantics} \ text\ -\label{sec:cimp-semantics} - Here we define the semantics of a single process's program. We begin by defining the type of externally-visible behaviour: \ datatype ('answer, 'question) seq_label = sl_Internal ("\") | sl_Send 'question 'answer ("\_, _\") | sl_Receive 'question 'answer ("\_, _\") text\ We define a \emph{labelled transition system} (an LTS) using an execution-stack style of semantics that avoids special treatment of the \SKIP\s introduced by a traditional small step semantics (such as @{cite [cite_macro=citet] \Chapter~14\ "Winskel:1993"}) when a basic command is executed. This was suggested by Thomas Sewell; @{cite [cite_macro=citet] "PittsAM:opespe"} gave a semantics to an ML-like language using this approach. +We record the location of the command that was executed to support +fairness constraints. + \ type_synonym ('answer, 'location, 'question, 'state) local_state - = "('answer, 'location, 'question, 'state) com list \ 'state" + = "('answer, 'location, 'question, 'state) com list \ 'location option \ 'state" inductive small_step :: "('answer, 'location, 'question, 'state) local_state \ ('answer, 'question) seq_label \ ('answer, 'location, 'question, 'state) local_state \ bool" ("_ \\<^bsub>_\<^esub> _" [55, 0, 56] 55) where - Request: "\ \ = action s; s' \ val \ s \ \ (\l\ Request action val # cs, s) \\<^bsub>\\, \\\<^esub> (cs, s')" -| Response: "(s', \) \ action \ s \ (\l\ Response action # cs, s) \\<^bsub>\\, \\\<^esub> (cs, s')" - -| LocalOp: "s' \ R s \ (\l\ LocalOp R # cs, s) \\<^bsub>\\<^esub> (cs, s')" - -| Cond1True: "b s \ (\l\ IF b THEN c FI # cs, s) \\<^bsub>\\<^esub> (c # cs, s)" -| Cond1False: "\b s \ (\l\ IF b THEN c FI # cs, s) \\<^bsub>\\<^esub> (cs, s)" + "\ \ = action s; s' \ val \ s \ \ (\l\ Request action val # cs, _, s) \\<^bsub>\\, \\\<^esub> (cs, Some l, s')" +| "(s', \) \ action \ s \ (\l\ Response action # cs, _, s) \\<^bsub>\\, \\\<^esub> (cs, Some l, s')" -| Cond2True: "b s \ (\l\ IF b THEN c1 ELSE c2 FI # cs, s) \\<^bsub>\\<^esub> (c1 # cs, s)" -| Cond2False: "\b s \ (\l\ IF b THEN c1 ELSE c2 FI # cs, s) \\<^bsub>\\<^esub> (c2 # cs, s)" - -| Loop: "(c # LOOP DO c OD # cs, s) \\<^bsub>\\<^esub> (cs', s') \ (LOOP DO c OD # cs, s) \\<^bsub>\\<^esub> (cs', s')" +| "s' \ R s \ (\l\ LocalOp R # cs, _, s) \\<^bsub>\\<^esub> (cs, Some l, s')" -| While: "b s \ (\l\ WHILE b DO c OD # cs, s) \\<^bsub>\\<^esub> (c # \l\ WHILE b DO c OD # cs, s)" -| WhileFalse: "\ b s \ (\l\ WHILE b DO c OD # cs, s) \\<^bsub>\\<^esub> (cs, s)" +| "b s \ (\l\ IF b THEN c FI # cs, _, s) \\<^bsub>\\<^esub> (c # cs, Some l, s)" +| "\b s \ (\l\ IF b THEN c FI # cs, _, s) \\<^bsub>\\<^esub> (cs, Some l, s)" -| Seq: "(c1 # c2 # cs, s) \\<^bsub>\\<^esub> (cs', s') \ (c1;; c2 # cs, s) \\<^bsub>\\<^esub> (cs', s')" +| "b s \ (\l\ IF b THEN c1 ELSE c2 FI # cs, _, s) \\<^bsub>\\<^esub> (c1 # cs, Some l, s)" +| "\b s \ (\l\ IF b THEN c1 ELSE c2 FI # cs, _, s) \\<^bsub>\\<^esub> (c2 # cs, Some l, s)" -| Choose1: "(c1 # cs, s) \\<^bsub>\\<^esub> (cs', s') \ (c1 \ c2 # cs, s) \\<^bsub>\\<^esub> (cs', s')" -| Choose2: "(c2 # cs, s) \\<^bsub>\\<^esub> (cs', s') \ (c1 \ c2 # cs, s) \\<^bsub>\\<^esub> (cs', s')" +| "(c # LOOP DO c OD # cs, s) \\<^bsub>\\<^esub> (cs', s') \ (LOOP DO c OD # cs, s) \\<^bsub>\\<^esub> (cs', s')" + +| "b s \ (\l\ WHILE b DO c OD # cs, _, s) \\<^bsub>\\<^esub> (c # \l\ WHILE b DO c OD # cs, Some l, s)" +| "\ b s \ (\l\ WHILE b DO c OD # cs, _, s) \\<^bsub>\\<^esub> (cs, Some l, s)" + +| "(c1 # c2 # cs, s) \\<^bsub>\\<^esub> (cs', s') \ (c1;; c2 # cs, s) \\<^bsub>\\<^esub> (cs', s')" + +| Choose1: "(c1 # cs, s) \\<^bsub>\\<^esub> (cs', s') \ (c1 \ c2 # cs, s) \\<^bsub>\\<^esub> (cs', s')" +| Choose2: "(c2 # cs, s) \\<^bsub>\\<^esub> (cs', s') \ (c1 \ c2 # cs, s) \\<^bsub>\\<^esub> (cs', s')" text\ -The following projections operate on local states. These are internal -to CIMP and should not appear to the end-user. +The following projections operate on local states. These should not +appear to the end-user. \ abbreviation cPGM :: "('answer, 'location, 'question, 'state) local_state \ ('answer, 'location, 'question, 'state) com list" where "cPGM \ fst" -abbreviation cLST :: "('answer, 'location, 'question, 'state) local_state \ 'state" where - "cLST s \ snd s" -(*<*) - -declare small_step.intros[intro] - -inductive_cases small_step_inv[elim]: - "(\l\ Request action val # cs, ls) \\<^bsub>a\<^esub> s'" - "(\l\ Response action # cs, ls) \\<^bsub>a\<^esub> s'" - "(\l\ LocalOp R # cs, ls) \\<^bsub>a\<^esub> s'" - "(\l\ IF b THEN c FI # cs, ls) \\<^bsub>a\<^esub> s'" - "(\l\ IF b THEN c1 ELSE c2 FI # cs, ls) \\<^bsub>a\<^esub> s'" - "(\l\ WHILE b DO c OD # cs, ls) \\<^bsub>a\<^esub> s'" - "(LOOP DO c OD # cs, ls) \\<^bsub>a\<^esub> s'" - -lemma small_step_impossible[iff]: - "\(\l\ Request action val # cs, ls) \\<^bsub>\\<^esub> s'" - "\(\l\ Request action val # cs, ls) \\<^bsub>\\, \\\<^esub> s'" - "\(\l\ Response action' # cs, ls) \\<^bsub>\\<^esub> s'" - "\(\l\ Response action' # cs, ls) \\<^bsub>\\, \\\<^esub> s'" - "\(\l\ LocalOp R # cs, ls) \\<^bsub>\\, \\\<^esub> s'" - "\(\l\ LocalOp R # cs, ls) \\<^bsub>\\, \\\<^esub> s'" - "\(\l\ IF b THEN c FI # cs, ls) \\<^bsub>\\, \\\<^esub> s'" - "\(\l\ IF b THEN c FI # cs, ls) \\<^bsub>\\, \\\<^esub> s'" - "\(\l\ IF b THEN c1 ELSE c2 FI # cs, ls) \\<^bsub>\\, \\\<^esub> s'" - "\(\l\ IF b THEN c1 ELSE c2 FI # cs, ls) \\<^bsub>\\, \\\<^esub> s'" - "\(\l\ WHILE b DO c OD # cs, ls) \\<^bsub>\\, \\\<^esub> s'" - "\(\l\ WHILE b DO c OD # cs, ls) \\<^bsub>\\, \\\<^esub> s'" -by (auto elim: small_step.cases) - -lemma small_step_stuck[iff]: - "\ ([], l, s) \\<^bsub>\\<^esub> c'" -by (auto elim: small_step.cases) - -(*>*) -text\ - -\label{sec:cimp-decompose} - -To reason about system transitions we need to identify which basic -statement gets executed next. To that end we factor out the recursive -cases of the @{term "small_step"} semantics into \emph{contexts}, -which identify the \basic_com\ commands with immediate -externally-visible behaviour. Note that non-determinism means that -more than one \basic_com\ can be enabled at a time. - -The representation of evaluation contexts follows @{cite -[cite_macro=citet] "DBLP:journals/jar/Berghofer12"}. This style of -operational semantics was originated by @{cite [cite_macro=citet] -"DBLP:journals/tcs/FelleisenH92"}. - -\ - -type_synonym ('answer, 'location, 'question, 'state) ctxt - = "('answer, 'location, 'question, 'state) com \ ('answer, 'location, 'question, 'state) com" - -inductive_set - ctxt :: "( ('answer, 'location, 'question, 'state) ctxt - \ (('answer, 'location, 'question, 'state) com \ ('answer, 'location, 'question, 'state) com list) ) set" -where - C_Hole: "(id, \[]\) \ ctxt" -| C_Loop: "(E, fctxt) \ ctxt \ (\t. LOOP DO E t OD, \t. fctxt t @ [LOOP DO E t OD]) \ ctxt" -| C_Seq: "(E, fctxt) \ ctxt \ (\t. E t;; u, \t. fctxt t @ [u]) \ ctxt" -| C_Choose1: "(E, fctxt) \ ctxt \ (\t. E t \ u, fctxt) \ ctxt" -| C_Choose2: "(E, fctxt) \ ctxt \ (\t. u \ E t, fctxt) \ ctxt" - -inductive - basic_com :: "('answer, 'location, 'question, 'state) com \ bool" -where - "basic_com (\l\ Request action val)" -| "basic_com (\l\ Response action)" -| "basic_com (\l\ LocalOp R)" -| "basic_com (\l\ IF b THEN c FI)" -| "basic_com (\l\ IF b THEN c1 ELSE c2 FI)" -| "basic_com (\l\ WHILE b DO c OD)" -(*<*) - -declare basic_com.intros[intro!] basic_com.cases[elim] - -(*>*) -text\ +abbreviation cTKN :: "('answer, 'location, 'question, 'state) local_state \ 'location option" where + "cTKN s \ fst (snd s)" -We can decompose a small step into a context and a @{term -"basic_com"}. - -\ - -fun - decompose_com :: "('answer, 'location, 'question, 'state) com - \ ( ('answer, 'location, 'question, 'state) com - \ ('answer, 'location, 'question, 'state) ctxt - \ (('answer, 'location, 'question, 'state) com \ ('answer, 'location, 'question, 'state) com list) ) set" -where - "decompose_com (LOOP DO c1 OD) = { (c, \t. LOOP DO ictxt t OD, \t. fctxt t @ [LOOP DO ictxt t OD]) |c fctxt ictxt. (c, ictxt, fctxt) \ decompose_com c1 }" -| "decompose_com (c1;; c2) = { (c, \t. ictxt t;; c2, \t. fctxt t @ [c2]) |c fctxt ictxt. (c, ictxt, fctxt) \ decompose_com c1 }" -| "decompose_com (c1 \ c2) = { (c, \t. ictxt t \ c2, fctxt) |c fctxt ictxt. (c, ictxt, fctxt) \ decompose_com c1 } - \ { (c, \t. c1 \ ictxt t, fctxt) |c fctxt ictxt. (c, ictxt, fctxt) \ decompose_com c2 }" -| "decompose_com c = {(c, id, \[]\)}" - -definition - decomposeLS :: "('answer, 'location, 'question, 'state) local_state - \ ( ('answer, 'location, 'question, 'state) com - \ (('answer, 'location, 'question, 'state) com \ ('answer, 'location, 'question, 'state) com) - \ (('answer, 'location, 'question, 'state) com \ ('answer, 'location, 'question, 'state) com list) ) set" -where - "decomposeLS s \ case cPGM s of c # _ \ decompose_com c | _ \ {}" - -(*<*) -declare ctxt.intros[intro!] - -lemma ctxt_inj: - assumes "(E, fctxt) \ ctxt" - assumes "E x = E y" - shows "x = y" -using assms by (induct set: ctxt) auto - -lemma decompose_ctxt: - assumes "(c', ictxt, fctxt) \ decompose_com c" - shows "(ictxt, fctxt) \ ctxt" -using assms by (induct c arbitrary: c' ictxt fctxt) auto +abbreviation cLST :: "('answer, 'location, 'question, 'state) local_state \ 'state" where + "cLST s \ snd (snd s)" -lemma decompose_ictxt: - assumes "(c', ictxt, fctxt) \ decompose_com c" - shows "c = ictxt c'" -using assms by (induct c arbitrary: c' ictxt fctxt) auto -(*>*) -theorem context_decompose: - "s \\<^bsub>\\<^esub> s' \ (\(c, ictxt, fctxt) \ decomposeLS s. - cPGM s = ictxt c # tl (cPGM s) - \ basic_com c - \ (c # fctxt c @ tl (cPGM s), cLST s) \\<^bsub>\\<^esub> s')" -(*<*)(is "?lhs = ?rhs") -proof - assume ?lhs then show ?rhs - by (induct rule: small_step.induct) - (fastforce simp: decomposeLS_def)+ -next - have gen: "(c0 # cs, s) \\<^bsub>\\<^esub> s'" - if as: "(c # fctxt c @ cs, s) \\<^bsub>\\<^esub> s'" - and ds: "(c, ictxt, fctxt) \ decompose_com c0" - for cs c c0 ictxt fctxt l s s' - proof - - from ds have ic: "(ictxt, fctxt) \ ctxt" - unfolding decomposeLS_def by (auto intro: decompose_ctxt split: list.splits) - from ic as decompose_ictxt[OF ds] - show "(c0 # cs, s) \\<^bsub>\\<^esub> s'" - by (induct ictxt fctxt arbitrary: c0 cs set: ctxt) - (cases s', fastforce simp: fun_eq_iff dest: ctxt_inj)+ - qed - assume ?rhs then show ?lhs - by (cases s) - (auto simp: decomposeLS_def split: list.splits dest!: gen) -qed - -(*>*) -text\ - -While we only use this result left-to-right (to decompose a small step -into a basic one), this equivalence shows that we lose no information -in doing so. - -\ - -subsection\System steps\ +subsection\System steps \label{sec:cimp-system-steps}\ text\ -\label{sec:cimp-system-steps} - A global state maps process names to process' local states. One might hope to allow processes to have distinct types of local state, but there remains no good solution yet in a simply-typed setting; see @{cite [cite_macro=citet] "DBLP:journals/entcs/SchirmerW09"}. \ type_synonym ('answer, 'location, 'proc, 'question, 'state) global_state = "'proc \ ('answer, 'location, 'question, 'state) local_state" type_synonym ('proc, 'state) local_states = "'proc \ 'state" text\ An execution step of the overall system is either any enabled internal @{term "\"} step of any process, or a communication rendezvous between two processes. For the latter to occur, a @{const "Request"} action must be enabled in process @{term "p1"}, and a @{const "Response"} action in (distinct) process @{term "p2"}, where the request/response labels @{term "\"} and @{term "\"} (semantically) match. We also track global communication history here to support assertional -reasoning (see \S\ref{sec:cimp-assertions}). +reasoning (see \S\ref{sec:cimp-invariants}). \ type_synonym ('answer, 'question) event = "'question \ 'answer" type_synonym ('answer, 'question) history = "('answer, 'question) event list" -type_synonym ('answer, 'location, 'proc, 'question, 'state) system_state - = "('answer, 'location, 'proc, 'question, 'state) global_state - \ ('answer, 'question) history" - -inductive_set - system_step :: "( ('answer, 'ls, 'proc, 'question, 'state) system_state - \ ('answer, 'ls, 'proc, 'question, 'state) system_state ) set" -where - LocalStep: "\ s p \\<^bsub>\\<^esub> ls'; s' = s(p := ls'); h' = h \ \ ((s, h), (s', h')) \ system_step" -| CommunicationStep: "\ s p1 \\<^bsub>\\, \\\<^esub> ls1'; s p2 \\<^bsub>\\, \\\<^esub> ls2'; p1 \ p2; - s' = s(p1 := ls1', p2 := ls2'); h' = h @ [(\, \)] \ \ ((s, h), (s', h')) \ system_step" +record ('answer, 'location, 'proc, 'question, 'state) system_state = + GST :: "('answer, 'location, 'proc, 'question, 'state) global_state" + HST :: "('answer, 'question) history" -abbreviation - system_step_syn :: "('answer, 'ls, 'proc, 'question, 'state) system_state - \ ('answer, 'ls, 'proc, 'question, 'state) system_state \ bool" ("_ s\ _" [55, 56] 55) +inductive \\ This is a predicate of the current state, so the successor state comes first. \ + system_step :: "'proc set + \ ('answer, 'location, 'proc, 'question, 'state) system_state + \ ('answer, 'location, 'proc, 'question, 'state) system_state + \ bool" where - "sh s\ sh' \ (sh, sh') \ system_step" + LocalStep: "\ GST sh p \\<^bsub>\\<^esub> ls'; GST sh' = (GST sh)(p := ls'); HST sh' = HST sh \ \ system_step {p} sh' sh" +| CommunicationStep: "\ GST sh p \\<^bsub>\\, \\\<^esub> ls1'; GST sh q \\<^bsub>\\, \\\<^esub> ls2'; p \ q; + GST sh' = (GST sh)(p := ls1', q := ls2'); HST sh' = HST sh @ [(\, \)] \ \ system_step {p, q} sh' sh" -abbreviation - system_steps_syn :: "('answer, 'ls, 'proc, 'question, 'state) system_state - \ ('answer, 'ls, 'proc, 'question, 'state) system_state \ bool" ("_ s\\<^sup>* _" [55, 56] 55) -where - "sh s\\<^sup>* sh' \ (sh, sh') \ system_step\<^sup>*" -(*<*) - -(*>*) text\ In classical process algebras matching communication actions yield \\\ steps, which aids nested parallel composition and the restriction operation @{cite [cite_macro=citep] \\S2.2\ "Milner:1980"}. As CIMP does not provide either we do not need to hide communication labels. In CCS/CSP it is not clear how one reasons about the communication history, and it seems that assertional reasoning about these languages is not well developed. \ -subsection\Assertions\ +text\ + +We define predicates over communication histories and system states. These are uncurried to ease composition. + +\ + +type_synonym ('answer, 'location, 'proc, 'question, 'state) state_pred + = "('answer, 'location, 'proc, 'question, 'state) system_state \ bool" text\ -\label{sec:cimp-assertions} +The \LST\ operator (written as a postfix \\\) projects +the local states of the processes from a \<^typ>\('answer, 'location, 'proc, 'question, 'state) system_state\, i.e. it +discards control location information. -We now develop a technique for showing that a CIMP system satisfies a -single global invariant, following @{cite [cite_macro=citet] -"DBLP:journals/acta/Lamport80" and "DBLP:journals/toplas/LamportS84"} -(and the later @{cite [cite_macro=citet] "DBLP:books/aw/Lamport2002"}) -and closely related work by @{cite [cite_macro=citet] -"AptFrancezDeRoever:1980"}, @{cite [cite_macro=citet] -"CousotCousot:1980"} and @{cite [cite_macro=citet] -"DBLP:journals/acta/LevinG81"}, which suggest the incorporation of a -history variable. @{cite [cite_macro=citet] "CousotCousot:1980"} -apparently contains a completeness proof. Lamport mentions that this -technique was well-known in the mid-80s when he proposed the use of -prophecy variables (see his webpage bibliography). See also @{cite -[cite_macro=citet] "DBLP:books/cu/RoeverBH2001"} for an extended -discussion of some of this. +Conversely the \LSTP\ operator lifts predicates over +local states into predicates over \<^typ>\('answer, 'location, 'proc, 'question, 'state) system_state\. -Achieving the right level of abstraction is a bit fiddly; we want to -avoid revealing too much of the program text as it -executes. Intuitively we wish to expose the processes's present -control locations and local states only. @{cite [cite_macro=citet] -"DBLP:journals/acta/Lamport80"} avoids these issues by only providing -an axiomatic semantics for his language. +Predicates that do not depend on control locations were termed @{emph +\universal assertions\} by @{cite [cite_macro=citet] +\\S3.6\ "DBLP:journals/acta/LevinG81"}. \ -subsubsection\Control predicates\ +type_synonym ('proc, 'state) local_state_pred + = "('proc, 'state) local_states \ bool" + +definition LST :: "('answer, 'location, 'proc, 'question, 'state) system_state + \ ('proc, 'state) local_states" ("_\" [1000] 1000) where + "s\ = cLST \ GST s" + +abbreviation (input) LSTP :: "('proc, 'state) local_state_pred + \ ('answer, 'location, 'proc, 'question, 'state) state_pred" where + "LSTP P \ \s. P s\" + + +subsection\Control predicates \label{sec:cimp-control-predicates}\ text\ -\label{sec:cimp-control-predicates} - Following @{cite [cite_macro=citet] "DBLP:journals/acta/Lamport80"}\footnote{@{cite [cite_macro=citet] "MannaPnueli:1995"} also develop a theory of locations. I think Lamport attributes control predicates to Owicki in her PhD thesis (under Gries). I did not find a treatment of procedures. @{cite -[cite_macro=citet] "MannaPnueli:1991"} observe that a set notation for -spreading assertions over sets of locations reduces clutter +[cite_macro=citet] "MannaPnueli:1991"} observe that a notation for +making assertions over sets of locations reduces clutter significantly.}, we define the \at\ predicate, which holds of a process when control resides at that location. Due to non-determinism processes can be \at\ a set of locations; it is more like ``a statement with this location is enabled'', which incidentally handles non-unique locations. Lamport's language is deterministic, so he doesn't have this problem. This also allows him to develop a stronger theory about his control predicates. \ +type_synonym 'location label = "'location set" + primrec - atC :: "('answer, 'location, 'question, 'state) com \ 'location \ bool" + atC :: "('answer, 'location, 'question, 'state) com \ 'location label" where - "atC (\l'\ Request action val) = (\l. l = l')" -| "atC (\l'\ Response action) = (\l. l = l')" -| "atC (\l'\ LocalOp f) = (\l. l = l')" -| "atC (\l'\ IF _ THEN _ FI) = (\l. l = l')" -| "atC (\l'\ IF _ THEN _ ELSE _ FI) = (\l. l = l')" -| "atC (\l'\ WHILE _ DO _ OD) = (\l. l = l')" + "atC (\l\ Request action val) = {l}" +| "atC (\l\ Response action) = {l}" +| "atC (\l\ LocalOp f) = {l}" +| "atC (\l\ IF _ THEN _ FI) = {l}" +| "atC (\l\ IF _ THEN _ ELSE _ FI) = {l}" +| "atC (\l\ WHILE _ DO _ OD) = {l}" | "atC (LOOP DO c OD) = atC c" | "atC (c1;; c2) = atC c1" -| "atC (c1 \ c2) = (atC c1 \<^bold>\ atC c2)" - -primrec atL :: "('answer, 'location, 'question, 'state) com list \ 'location \ bool" where - "atL [] = \False\" -| "atL (c # _) = atC c" - -abbreviation atLS :: "('answer, 'location, 'question, 'state) local_state \ 'location \ bool" where - "atLS \ \s. atL (cPGM s)" -(*<*) - -lemma at_decompose: - "(c, ictxt, fctxt) \ decompose_com c0 \ (\l. atC c l \ atC c0 l)" -by (induct c0 arbitrary: c ictxt fctxt) fastforce+ - -lemma at_decomposeLS: - "(c, ictxt, fctxt) \ decomposeLS s \ (\l. atC c l \ atLS s l)" -by (auto simp: decomposeLS_def at_decompose split: list.splits) - -(*>*) -text\ +| "atC (c1 \ c2) = atC c1 \ atC c2" -We define predicates over communication histories and a projection of -global states. These are uncurried to ease composition. - -\ - -type_synonym ('location, 'proc, 'state) pred_local_state - = "'proc \ (('location \ bool) \ 'state)" - -record ('answer, 'location, 'proc, 'question, 'state) pred_state = - local_states :: "('location, 'proc, 'state) pred_local_state" - hist :: "('answer, 'question) history" +primrec atCs :: "('answer, 'location, 'question, 'state) com list \ 'location label" where + "atCs [] = {}" +| "atCs (c # _) = atC c" -type_synonym ('answer, 'location, 'proc, 'question, 'state) pred - = "('answer, 'location, 'proc, 'question, 'state) pred_state \ bool" - -definition mkP :: "('answer, 'location, 'proc, 'question, 'state) system_state \ ('answer, 'location, 'proc, 'question, 'state) pred_state" where - "mkP \ \(s, h). \ local_states = \p. case s p of (cs, ps) \ (atL cs, ps), hist = h \" -(*<*) - -lemma hist_mkP[iff]: - "hist (mkP (s, h)) = h" -by (simp add: mkP_def) - -(*>*) text\ We provide the following definitions to the end-user. \AT\ maps process names to a predicate that is true of -locations where control for that process resides. The abbreviation -\at\ shuffles its parameters; the former is -simplifier-friendly and eta-reduced, while the latter is convenient -for writing assertions. - -\ - -definition AT :: "('answer, 'location, 'proc, 'question, 'state) pred_state \ 'proc \ 'location \ bool" where - "AT \ \s p l. fst (local_states s p) l" - -abbreviation at :: "'proc \ 'location \ ('answer, 'location, 'proc, 'question, 'state) pred" where - "at p l s \ AT s p l" - -text\ - -Often we wish to talk about control residing at one of a set of -locations. This stands in for, and generalises, the \in\ -predicate of @{cite [cite_macro=citet] +locations where control for that process resides, and the abbreviation \at\ provides a conventional +way to use it. The constant \atS\ specifies that control for process \p\ resides at one of +the given locations. This stands in for, and generalises, the \in\ predicate of @{cite [cite_macro=citet] "DBLP:journals/acta/Lamport80"}. \ -definition atS :: "'proc \ 'location set \ ('answer, 'location, 'proc, 'question, 'state) pred" where - "atS \ \p ls s. \l\ls. at p l s" +definition AT :: "('answer, 'location, 'proc, 'question, 'state) system_state \ 'proc \ 'location label" where + "AT s p = atCs (cPGM (GST s p))" + +abbreviation at :: "'proc \ 'location \ ('answer, 'location, 'proc, 'question, 'state) state_pred" where + "at p l s \ l \ AT s p" + +definition atS :: "'proc \ 'location set \ ('answer, 'location, 'proc, 'question, 'state) state_pred" where + "atS p ls s = (\l\ls. at p l s)" + +(* FIXME rename, document, rules. Identifies a process's control state label precisely as one element of a set. *) +definition atLs :: "'proc \ 'location label set \ ('answer, 'location, 'proc, 'question, 'state) state_pred" where + "atLs p labels s = (AT s p \ labels)" + +(* FIXME rename, document. Identifies a process's control state label precisely. Relate atL to at/atS, ex/ *) +abbreviation (input) atL :: "'proc \ 'location label \ ('answer, 'location, 'proc, 'question, 'state) state_pred" where + "atL p label \ atLs p {label}" + +(* FIXME rename, document. Processes are at the given labels. *) +definition atPLs :: "('proc \ 'location label) set \ ('answer, 'location, 'proc, 'question, 'state) state_pred" where + "atPLs pls = (\<^bold>\p label. \(p, label) \ pls\ \<^bold>\ atL p label)" text\ +The constant \taken\ provides a way of identifying which transition was taken. It is somewhat like +Lamport's \after\, but not quite due to the presence of non-determinism here. This does not work well +for invariants or preconditions. + +\ + +definition taken :: "'proc \ 'location \ ('answer, 'location, 'proc, 'question, 'state) state_pred" where + "taken p l s \ cTKN (GST s p) = Some l" + +text\ + +\label{sec:cimp-termination} + A process is terminated if it not at any control location. \ -abbreviation terminated :: "'proc \ ('answer, 'location, 'proc, 'question, 'state) pred" where - "terminated p s \ \l. \at p l s" - -text\ - -The \LST\ operator (written as a postfix \\\) projects -the local states of the processes from a \pred_state\, i.e. it -discards control location information. - -Conversely the \LSTP\ operator lifts predicates over -local states into predicates over \pred_state\. @{cite -[cite_macro=citet] \\S3.6\ "DBLP:journals/acta/LevinG81"} -call such predicates \emph{universal assertions}. - -\ - -type_synonym ('proc, 'state) state_pred - = "('proc, 'state) local_states \ bool" - -definition LST :: "('answer, 'location, 'proc, 'question, 'state) pred_state - \ ('proc, 'state) local_states" ("_\" [1000] 1000) where - "s\ \ snd \ local_states s" - -abbreviation (input) LSTP :: "('proc, 'state) state_pred - \ ('answer, 'location, 'proc, 'question, 'state) pred" where - "LSTP P \ \s. P (LST s)" +abbreviation (input) terminated :: "'proc \ ('answer, 'location, 'proc, 'question, 'state) state_pred" where + "terminated p \ atL p {}" text\ -By default we ask the simplifier to rewrite @{const "atS"} using -ambient @{const "AT"} information. - -\ - -lemma atS_state_cong[cong]: - "\ AT s p = AT s' p \ \ atS p ls s \ atS p ls s'" -by (auto simp: atS_def) - -text\ - -We provide an incomplete set of basic rules for label sets. - -\ - -lemma atS_simps: - "\atS p {} s" - "atS p {l} s \ at p l s" - "\ at p l s; l \ ls \ \ atS p ls s \ True" - "(\l. at p l s \ l \ ls) \ atS p ls s \ False" -by (auto simp: atS_def) - -lemma atS_mono: - "\ atS p ls s; ls \ ls' \ \ atS p ls' s" -by (auto simp: atS_def) - -lemma atS_un: - "atS p (l \ l') s \ atS p l s \ atS p l' s" -by (auto simp: atS_def) - -subsubsection\Invariants\ - -text\ - -\label{sec:cimp-invariants} - A complete system consists of one program per process, and a (global) constraint on their initial local states. From these we can construct the set of initial global states and all those reachable by system steps (\S\ref{sec:cimp-system-steps}). \ type_synonym ('answer, 'location, 'proc, 'question, 'state) programs = "'proc \ ('answer, 'location, 'question, 'state) com" -type_synonym ('answer, 'location, 'proc, 'question, 'state) system - = "('answer, 'location, 'proc, 'question, 'state) programs - \ ('proc, 'state) state_pred" +record ('answer, 'location, 'proc, 'question, 'state) pre_system = + PGMs :: "('answer, 'location, 'proc, 'question, 'state) programs" + INIT :: "('proc, 'state) local_state_pred" definition - initial_states :: "('answer, 'location, 'proc, 'question, 'state) system - \ ('answer, 'location, 'proc, 'question, 'state) global_state set" + initial_state :: "('answer, 'location, 'proc, 'question, 'state, 'ext) pre_system_ext + \ ('answer, 'location, 'proc, 'question, 'state) global_state + \ bool" where - "initial_states sys \ - { f . (\p. cPGM (f p) = [fst sys p]) \ snd sys (cLST \ f) }" - -definition - reachable_states :: "('answer, 'location, 'proc, 'question, 'state) system - \ ('answer, 'location, 'proc, 'question, 'state) system_state set" -where - "reachable_states sys \ system_step\<^sup>* `` (initial_states sys \ {[]})" + "initial_state sys s = ((\p. cPGM (s p) = [PGMs sys p] \ cTKN (s p) = None) \ INIT sys (cLST \ s))" text\ -The following is a slightly more convenient induction rule for the set -of reachable states. - -\ - -lemma reachable_states_system_step_induct[consumes 1, - case_names init LocalStep CommunicationStep]: - assumes r: "(s, h) \ reachable_states sys" - assumes i: "\s. s \ initial_states sys \ P s []" - assumes l: "\s h ls' p. \ (s, h) \ reachable_states sys; P s h; s p \\<^bsub>\\<^esub> ls' \ - \ P (s(p := ls')) h" - assumes c: "\s h ls1' ls2' p1 p2 \ \. - \ (s, h) \ reachable_states sys; P s h; - s p1 \\<^bsub>\\, \\\<^esub> ls1'; s p2 \\<^bsub>\\, \\\<^esub> ls2'; p1 \ p2 \ - \ P (s(p1 := ls1', p2 := ls2')) (h @ [(\, \)])" - shows "P s h" -(*<*) -proof - - have "P s' h'" if "(s, []) s\\<^sup>* (s', h')" "s \ initial_states sys" for s s' h' - using that - by (induct rule: rtrancl_induct2) - (fastforce simp: reachable_states_def elim: system_step.cases intro: i l c)+ - with r show ?thesis by (clarsimp simp: reachable_states_def) -qed +We construct infinite runs of a system by allowing stuttering, i.e., arbitrary repetitions of states +following @{cite [cite_macro=citet] \Chapter~8\"Lamport:2002"}, by taking the reflexive +closure of the @{const system_step} relation. Therefore terminated +programs infinitely repeat their final state (but note our definition +of terminated processes in \S\ref{sec:cimp-termination}). -lemma initial_statesD: - "s \ initial_states sys \ AT (mkP (s, [])) = atC \ fst sys \ snd sys (mkP (s, []))\" -by (simp add: initial_states_def mkP_def split_def o_def LST_def AT_def) - -lemma initial_states_mkP[iff]: - "s \ initial_states sys \ at p l (mkP (s, [])) \ atC (fst sys p) l" -by (simp add: initial_states_def mkP_def split_def AT_def) - -(*>*) -subsubsection\Relating reachable states to the initial programs\ - -text\ - -\label{sec:cimp-decompose-small-step} - -To usefully reason about the control locations presumably embedded in -the single global invariant, we need to link the programs we have in -reachable state \s\ to the programs in the initial states. The -\fragments\ function decomposes the program into statements -that can be directly executed (\S\ref{sec:cimp-decompose}). We also -compute the locations we could be at after executing that statement as -a function of the process's local state. - -We could support Lamport's \after\ control predicate with more -syntactic analysis of this kind. +Some accounts define stuttering as the @{emph \finite\} repetition of states. With or without this constraint +\prerun\ contains @{emph \junk\} in the form of unfair runs, where particular processes do not progress. \ -fun - extract_cond :: "('answer, 'location, 'question, 'state) com \ 'state bexp" -where - "extract_cond (\l\ IF b THEN c FI) = b" -| "extract_cond (\l\ IF b THEN c1 ELSE c2 FI) = b" -| "extract_cond (\l\ WHILE b DO c OD) = b" -| "extract_cond c = \False\" - -type_synonym ('answer, 'location, 'question, 'state) loc_comp - = "('answer, 'location, 'question, 'state) com - \ 'state \ 'location \ bool" - -fun lconst :: "('location \ bool) \ ('answer, 'location, 'question, 'state) loc_comp" where - "lconst lp b s = lp" - -definition lcond :: "('location \ bool) \ ('location \ bool) - \ ('answer, 'location, 'question, 'state) loc_comp" where - "lcond lp lp' c s \ if extract_cond c s then lp else lp'" -(*<*) - -lemma lcond_split: - "Q (lcond lp lp' c s) \ (extract_cond c s \ Q lp) \ (\extract_cond c s \ Q lp')" -by (simp add: lcond_def split: if_splits) - -lemma lcond_split_asm: - "Q (lcond lp lp' c s) \ \ ((extract_cond c s \ \Q lp) \ (\extract_cond c s \ \ Q lp'))" -by (simp add: lcond_def split: if_splits) - -lemmas lcond_splits = lcond_split lcond_split_asm -(*>*) - -fun - fragments :: "('answer, 'location, 'question, 'state) com - \ ('location \ bool) - \ ( ('answer, 'location, 'question, 'state) com - \ ('answer, 'location, 'question, 'state) loc_comp ) set" +definition + system_step_reflclp :: "('answer, 'location, 'proc, 'question, 'state) system_state seq_pred" where - "fragments (\l\ IF b THEN c FI) ls - = { (\l\ IF b THEN c' FI, lcond (atC c) ls) |c'. True } - \ fragments c ls" -| "fragments (\l\ IF b THEN c1 ELSE c2 FI) ls - = { (\l\ IF b THEN c1' ELSE c2' FI, lcond (atC c1) (atC c2)) |c1' c2'. True } - \ fragments c1 ls \ fragments c2 ls" -| "fragments (LOOP DO c OD) ls = fragments c (atC c)" -| "fragments (\l'\ WHILE b DO c OD) ls - = { (\l'\ WHILE b DO c' OD, lcond (atC c) ls) |c'. True } \ fragments c (\l. l = l')" -| "fragments (c1;; c2) ls = fragments c1 (atC c2) \ fragments c2 ls" -| "fragments (c1 \ c2) ls = fragments c1 ls \ fragments c2 ls" -| "fragments c ls = { (c, lconst ls) }" + "system_step_reflclp \ \ (\sh sh'. \pls. system_step pls sh' sh)\<^sup>=\<^sup>= (\ 0) (\ 1)" -fun - fragmentsL :: "('answer, 'location, 'question, 'state) com list - \ ( ('answer, 'location, 'question, 'state) com - \ ('answer, 'location, 'question, 'state) loc_comp ) set" +definition + prerun :: "('answer, 'location, 'proc, 'question, 'state, 'ext) pre_system_ext + \ ('answer, 'location, 'proc, 'question, 'state) system_state seq_pred" where - "fragmentsL [] = {}" -| "fragmentsL [c] = fragments c \False\" -| "fragmentsL (c # c' # cs) = fragments c (atC c') \ fragmentsL (c' # cs)" - -abbreviation - fragmentsLS :: "('answer, 'location, 'question, 'state) local_state - \ ( ('answer, 'location, 'question, 'state) com - \ ('answer, 'location, 'question, 'state) loc_comp ) set" -where - "fragmentsLS s \ fragmentsL (cPGM s)" -(*<*) + "prerun sys = ((\\. initial_state sys (GST (\ 0)) \ HST (\ 0) = []) \<^bold>\ \system_step_reflclp)" -lemma fragmentsL_mono_Cons[iff]: - "fragmentsL cs \ fragmentsL (c # cs)" -by (induct cs) auto +definition \\ state-based invariants only \ + prerun_valid :: "('answer, 'location, 'proc, 'question, 'state, 'ext) pre_system_ext + \ ('answer, 'location, 'proc, 'question, 'state) state_pred \ bool" ("_ \\<^bsub>pre\<^esub> _" [11, 0] 11) (* FIXME priorities *) +where + "(sys \\<^bsub>pre\<^esub> \) \ (\\. prerun sys \ \ (\\\\) \)" -lemma small_step_fragmentsLS: - assumes "s \\<^bsub>\\<^esub> s'" - shows "fragmentsLS s' \ fragmentsLS s" -using assms by (induct rule: small_step.induct) (case_tac [!] cs, auto) - -lemmas small_step_fragmentsLS_mem = subsetD[OF small_step_fragmentsLS] - -(*>*) text\ -Eliding the bodies of \IF\ and \WHILE\ statements -yields smaller (but equivalent) proof obligations. - -We show that taking system steps preserves fragments. +A \run\ of a system is a @{const \prerun\} that satisfies the \FAIR\ requirement. +Typically this would include @{emph \weak fairness\} for every transition of every process. \ -lemma reachable_states_fragmentsLS: - assumes "(s, h) \ reachable_states sys" - shows "fragmentsLS (s p) \ fragments (fst sys p) \False\" -(*<*) -using assms -by (induct rule: reachable_states_system_step_induct) - (auto simp: initial_states_def dest: small_step_fragmentsLS_mem) -(*>*) - -text\ - -Decomposing a compound command preserves fragments too. - -\ - -fun - extract_inner_locations :: "('answer, 'location, 'question, 'state) com - \ ('answer, 'location, 'question, 'state) com list - \ ('answer, 'location, 'question, 'state) loc_comp" -where - "extract_inner_locations (\l\ IF b THEN c FI) cs = lcond (atC c) (atL cs)" -| "extract_inner_locations (\l\ IF b THEN c1 ELSE c2 FI) cs = lcond (atC c1) (atC c2)" -| "extract_inner_locations (LOOP DO c OD) cs = lconst (atC c)" -| "extract_inner_locations (\l\ WHILE b DO c OD) cs = lcond (atC c) (atL cs)" -| "extract_inner_locations c cs = lconst (atL cs)" - -(*<*) - -lemma decompose_fragments: - assumes "(c, ictxt, fctxt) \ decompose_com c0" - shows "(c, extract_inner_locations c (fctxt c @ cs)) \ fragments c0 (atL cs)" -using assms -proof(induct c0 arbitrary: c ictxt fctxt cs) - case (Loop c01 c ictxt fctxt cs) - from Loop.prems Loop.hyps(1)[where cs="ictxt c # cs"] show ?case by (auto simp: decompose_ictxt[symmetric]) -next - case (Seq c01 c02 c ictxt fctxt cs) - from Seq.prems Seq.hyps(1)[where cs="c02 # cs"] show ?case by auto -qed auto +record ('answer, 'location, 'proc, 'question, 'state) system = + "('answer, 'location, 'proc, 'question, 'state) pre_system" ++ FAIR :: "('answer, 'location, 'proc, 'question, 'state) system_state seq_pred" -lemma decomposeLS_fragmentsLS: - assumes "(c, ictxt, fctxt) \ decomposeLS s" - shows "(c, extract_inner_locations c (fctxt c @ tl (cPGM s))) \ fragmentsLS s" -using assms -proof(cases "cPGM s") - case (Cons d ds) - with assms decompose_fragments[where cs="ds"] - show ?thesis by (cases ds) (auto simp: decomposeLS_def) -qed (simp add: decomposeLS_def) -(*>*) - -lemma small_step_extract_inner_locations: - assumes "basic_com c" - assumes "(c # cs, ls) \\<^bsub>\\<^esub> ls'" - shows "extract_inner_locations c cs c ls = atLS ls'" -using assms by (fastforce split: lcond_splits) - -text\ - -The headline result allows us to constrain the initial and final states -of a given small step in terms of the original programs, provided the -initial state is reachable. - -\ +definition + run :: "('answer, 'location, 'proc, 'question, 'state) system + \ ('answer, 'location, 'proc, 'question, 'state) system_state seq_pred" +where + "run sys = (prerun sys \<^bold>\ FAIR sys)" -theorem decompose_small_step: - assumes "s p \\<^bsub>\\<^esub> ps'" - assumes "(s, h) \ reachable_states sys" - obtains c cs ls' - where "(c, ls') \ fragments (fst sys p) \False\" - and "basic_com c" - and "\l. atC c l \ atLS (s p) l" - and "ls' c (cLST (s p)) = atLS ps'" - and "(c # cs, cLST (s p)) \\<^bsub>\\<^esub> ps'" -(*<*) -using assms -apply - -apply (frule iffD1[OF context_decompose]) -apply clarsimp -apply (frule decomposeLS_fragmentsLS) -apply (frule at_decomposeLS) -apply (frule (1) subsetD[OF reachable_states_fragmentsLS, rotated]) -apply (frule (1) small_step_extract_inner_locations[rotated]) -apply auto -done - -(*>*) -text\ - -Reasoning with @{thm [source] "reachable_states_system_step_induct"} -and @{thm [source] "decompose_small_step"} is quite tedious. We -provide a very simple VCG that generates friendlier local proof -obligations. - -\ +definition + valid :: "('answer, 'location, 'proc, 'question, 'state) system + \ ('answer, 'location, 'proc, 'question, 'state) system_state seq_pred \ bool" ("_ \ _" [11, 0] 11) (* FIXME priorities *) +where + "(sys \ \) \ (\\. run sys \ \ \ \)" (*<*) end (*>*) diff --git a/thys/ConcurrentIMP/CIMP_one_place_buffer_ex.thy b/thys/ConcurrentIMP/CIMP_one_place_buffer_ex.thy deleted file mode 100644 --- a/thys/ConcurrentIMP/CIMP_one_place_buffer_ex.thy +++ /dev/null @@ -1,192 +0,0 @@ -(*<*) -(* - * 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 CIMP_one_place_buffer_ex -imports - CIMP -begin - -(*>*) -section\One-place buffer example\ - -text\ - -\label{sec:one_place_buffer} - -To demonstrate the CIMP reasoning infrastructure, we treat the trivial -one-place buffer example of @{cite [cite_macro=citet] -\\S3.3\ "DBLP:journals/toplas/LamportS84"}. Note that the -semantics for our language is different to @{cite -[cite_macro=citeauthor] "DBLP:journals/toplas/LamportS84"}'s, who -treated a historical variant of CSP (i.e., not the one in @{cite -"Hoare:1985"}). - -We introduce some syntax for fixed-topology (static channel-based) -scenarios. - -\ - -abbreviation - Receive :: "'location \ 'channel \ ('val \ 'state \ 'state) - \ (unit, 'location, 'channel \ 'val, 'state) com" ("\_\/ _\_") -where - "\l\ ch\f \ \l\ Response (\quest s. if fst quest = ch then {(f (snd quest) s, ())} else {})" - -abbreviation - Send :: "'location \ 'channel \ ('state \ 'val) - \ (unit, 'location, 'channel \ 'val, 'state) com" ("\_\/ _\_") -where - "\l\ ch\f \ \l\ Request (\s. (ch, f s)) (\ans s. {s})" - -text\ - -We further specialise these for our particular example. - -\ - -abbreviation - Receive' :: "'location \ 'channel \ (unit, 'location, 'channel \ 'state, 'state) com" ("\_\/ _\") -where - "\l\ ch\ \ \l\ ch\(\v _. v)" - -abbreviation - Send' :: "'location \ 'channel \ (unit, 'location, 'channel \ 'state, 'state) com" ("\_\/ _\") -where - "\l\ ch\ \ \l\ ch\id" - -text\ - -These definitions largely follow @{cite [cite_macro=citet] -"DBLP:journals/toplas/LamportS84"}. We have three processes -communicating over two channels. We enumerate program locations. - -\ - -datatype ex_chname = \12 | \23 -type_synonym ex_val = nat -type_synonym ex_ch = "ex_chname \ ex_val" -datatype ex_loc = r12 | r23 | s23 | s12 -datatype ex_proc = p1 | p2 | p3 - -type_synonym ex_pgm = "(unit, ex_loc, ex_ch, ex_val) com" -type_synonym ex_pred = "(unit, ex_loc, ex_proc, ex_ch, ex_val) pred" -type_synonym ex_state = "(unit, ex_loc, ex_proc, ex_ch, ex_val) global_state" -type_synonym ex_system = "(unit, ex_loc, ex_proc, ex_ch, ex_val) system" -type_synonym ex_history = "(ex_ch \ unit) list" - -primrec - ex_pgms :: "ex_proc \ ex_pgm" -where - "ex_pgms p1 = \s12\ \12\" -| "ex_pgms p2 = LOOP DO \r12\ \12\;; \s23\ \23\ OD" -| "ex_pgms p3 = \r23\ \23\" - -text\ - -Each process starts with an arbitrary initial local state. - -\ - -abbreviation ex_init :: "(ex_proc \ ex_val) \ bool" where - "ex_init \ \True\" - -abbreviation ex_system :: "ex_system" where - "ex_system \ (ex_pgms, ex_init)" - -text\ - -The following adapts Kai Engelhardt's, from his notes titled -\emph{Proving an Asynchronous Message Passing Program Correct}, -2011. The history variable tracks the causality of the system, which I -feel is missing in Lamport's treatment. We tack on Lamport's invariant -so we can establish \Etern_pred\. - -\ - -abbreviation - filter_on_channel :: "ex_chname \ ex_history \ ex_val list" -where - "filter_on_channel ch \ map (snd \ fst) \ filter ((=) ch \ fst \ fst)" - -definition Ip1_0 :: ex_pred where - "Ip1_0 \ at p1 s12 \<^bold>\ (\s. filter_on_channel \12 (hist s) = [])" -definition Ip1_1 :: ex_pred where - "Ip1_1 \ terminated p1 \<^bold>\ (\s. filter_on_channel \12 (hist s) = [LST s p1])" - -definition Ip2_0 :: ex_pred where - "Ip2_0 \ at p2 r12 \<^bold>\ (\s. filter_on_channel \12 (hist s) = filter_on_channel \23 (hist s))" -definition Ip2_1 :: ex_pred where - "Ip2_1 \ at p2 s23 \<^bold>\ (\s. filter_on_channel \12 (hist s) = filter_on_channel \23 (hist s) @ [LST s p2] - \ LST s p1 = LST s p2)" - -definition Ip3_0 :: ex_pred where - "Ip3_0 \ at p3 r23 \<^bold>\ (\s. filter_on_channel \23 (hist s) = [])" -definition Ip3_1 :: ex_pred where - "Ip3_1 \ terminated p3 \<^bold>\ (\s. filter_on_channel \23 (hist s) = [LST s p2] - \ LST s p1 = LST s p3)" - -definition I_pred :: ex_pred where - "I_pred \ pred_conjoin [ Ip1_0, Ip1_1, Ip2_0, Ip2_1, Ip3_0, Ip3_1 ]" - -lemmas I_defs = Ip1_0_def Ip1_1_def Ip2_0_def Ip2_1_def Ip3_0_def Ip3_1_def - -text\ - -If process three terminates, then it has process one's value. This is -stronger than @{cite [cite_macro=citeauthor] -"DBLP:journals/toplas/LamportS84"}'s as we don't ask that the first -process has also terminated. - -\ - -definition Etern_pred :: ex_pred where - "Etern_pred \ terminated p3 \<^bold>\ (\s. LST s p1 = LST s p3)" - -text\ - -Proofs from here down. - -\ - -lemma correct_system: - "I_pred sh \ Etern_pred sh" -apply (clarsimp simp: Etern_pred_def I_pred_def I_defs) -done - -lemma p1: "ex_pgms, p1, lconst \False\ \ \I_pred\ \s12\ \12\\s. s" -apply (rule vcg.intros) -apply (rename_tac p') -apply (case_tac p') - apply (auto simp: I_pred_def I_defs atS_def) -done - -lemma p2_1: "ex_pgms, p2, lconst (\l. l = r12) \ \I_pred\ \s23\ \23\\s. s" -apply (rule vcg.intros) -apply (rename_tac p') -apply (case_tac p') - apply (auto simp: I_pred_def I_defs atS_def) -done - -lemma "(s, h) \ reachable_states ex_system \ I_pred (mkP (s, h))" -apply (erule VCG) - apply (clarsimp simp: I_pred_def I_defs atS_def) -apply simp -apply (rename_tac p) -apply (case_tac p) - apply auto (* vcg_clarsimp_tac *) - apply (auto simp: p1 p2_1) -done - -text\\ -(*<*) - -end -(*>*) diff --git a/thys/ConcurrentIMP/CIMP_pred.thy b/thys/ConcurrentIMP/CIMP_pred.thy --- a/thys/ConcurrentIMP/CIMP_pred.thy +++ b/thys/ConcurrentIMP/CIMP_pred.thy @@ -1,138 +1,169 @@ (*<*) (* * 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 CIMP_pred imports Main begin +(* Extra HOL *) + +lemma triv: "P \ P" +by simp + +lemma always_eventually_pigeonhole: + "(\i. \n\i. \m\k. P m n) \ (\m\k::nat. \i::nat. \n\i. P m n)" +proof(induct k) + case (Suc k) then show ?case + apply (auto 8 0) + using le_SucI apply blast + apply (metis (full_types) le_Suc_eq nat_le_linear order_trans) + done +qed simp + (*>*) section\ Point-free notation \ text\ \label{sec:cimp-lifted-predicates} Typically we define predicates as functions of a state. The following provide a somewhat comfortable point-free imitation of Isabelle/HOL's operators. \ abbreviation (input) pred_K :: "'b \ 'a \ 'b" ("\_\") where "\f\ \ \s. f" abbreviation (input) - pred_not :: "('a \ bool) \ 'a \ bool" ("\<^bold>\") where + pred_not :: "('a \ bool) \ 'a \ bool" ("\<^bold>\ _" [40] 40) where "\<^bold>\a \ \s. \a s" abbreviation (input) pred_conj :: "('a \ bool) \ ('a \ bool) \ 'a \ bool" (infixr "\<^bold>\" 35) where "a \<^bold>\ b \ \s. a s \ b s" abbreviation (input) pred_disj :: "('a \ bool) \ ('a \ bool) \ 'a \ bool" (infixr "\<^bold>\" 30) where "a \<^bold>\ b \ \s. a s \ b s" abbreviation (input) pred_implies :: "('a \ bool) \ ('a \ bool) \ 'a \ bool" (infixr "\<^bold>\" 25) where "a \<^bold>\ b \ \s. a s \ b s" abbreviation (input) + pred_iff :: "('a \ bool) \ ('a \ bool) \ 'a \ bool" (infixr "\<^bold>\" 25) where + "a \<^bold>\ b \ \s. a s \ b s" + +abbreviation (input) pred_eq :: "('a \ 'b) \ ('a \ 'b) \ 'a \ bool" (infix "\<^bold>=" 40) where "a \<^bold>= b \ \s. a s = b s" abbreviation (input) pred_member :: "('a \ 'b) \ ('a \ 'b set) \ 'a \ bool" (infix "\<^bold>\" 40) where "a \<^bold>\ b \ \s. a s \ b s" abbreviation (input) pred_neq :: "('a \ 'b) \ ('a \ 'b) \ 'a \ bool" (infix "\<^bold>\" 40) where "a \<^bold>\ b \ \s. a s \ b s" abbreviation (input) pred_If :: "('a \ bool) \ ('a \ 'b) \ ('a \ 'b) \ 'a \ 'b" ("(If (_)/ Then (_)/ Else (_))" [0, 0, 10] 10) where "If P Then x Else y \ \s. if P s then x s else y s" abbreviation (input) pred_less :: "('a \ 'b::ord) \ ('a \ 'b) \ 'a \ bool" (infix "\<^bold><" 40) where "a \<^bold>< b \ \s. a s < b s" abbreviation (input) pred_le :: "('a \ 'b::ord) \ ('a \ 'b) \ 'a \ bool" (infix "\<^bold>\" 40) where "a \<^bold>\ b \ \s. a s \ b s" abbreviation (input) pred_plus :: "('a \ 'b::plus) \ ('a \ 'b) \ 'a \ 'b" (infixl "\<^bold>+" 65) where "a \<^bold>+ b \ \s. a s + b s" abbreviation (input) pred_minus :: "('a \ 'b::minus) \ ('a \ 'b) \ 'a \ 'b" (infixl "\<^bold>-" 65) where "a \<^bold>- b \ \s. a s - b s" abbreviation (input) fun_fanout :: "('a \ 'b) \ ('a \ 'c) \ 'a \ 'b \ 'c" (infix "\<^bold>\" 35) where "f \<^bold>\ g \ \x. (f x, g x)" abbreviation (input) pred_all :: "('b \ 'a \ bool) \ 'a \ bool" (binder "\<^bold>\" 10) where "\<^bold>\x. P x \ \s. \x. P x s" abbreviation (input) pred_ex :: "('b \ 'a \ bool) \ 'a \ bool" (binder "\<^bold>\" 10) where "\<^bold>\x. P x \ \s. \x. P x s" abbreviation (input) pred_app :: "('b \ 'a \ 'c) \ ('a \ 'b) \ 'a \ 'c" (infixl "\<^bold>$" 100) where "f \<^bold>$ g \ \s. f (g s) s" abbreviation (input) pred_subseteq :: "('a \ 'b set) \ ('a \ 'b set) \ 'a \ bool" (infix "\<^bold>\" 50) where "A \<^bold>\ B \ \s. A s \ B s" abbreviation (input) pred_union :: "('a \ 'b set) \ ('a \ 'b set) \ 'a \ 'b set" (infixl "\<^bold>\" 65) where "a \<^bold>\ b \ \s. a s \ b s" +abbreviation (input) + pred_inter :: "('a \ 'b set) \ ('a \ 'b set) \ 'a \ 'b set" (infixl "\<^bold>\" 65) where + "a \<^bold>\ b \ \s. a s \ b s" + text\ More application specific. \ abbreviation (input) pred_conjoin :: "('a \ bool) list \ 'a \ bool" where "pred_conjoin xs \ foldr (\<^bold>\) xs \True\" abbreviation (input) + pred_disjoin :: "('a \ bool) list \ 'a \ bool" where + "pred_disjoin xs \ foldr (\<^bold>\) xs \False\" + +abbreviation (input) pred_is_none :: "('a \ 'b option) \ 'a \ bool" ("NULL _" [40] 40) where "NULL a \ \s. a s = None" abbreviation (input) pred_empty :: "('a \ 'b set) \ 'a \ bool" ("EMPTY _" [40] 40) where "EMPTY a \ \s. a s = {}" abbreviation (input) pred_list_null :: "('a \ 'b list) \ 'a \ bool" ("LIST'_NULL _" [40] 40) where "LIST_NULL a \ \s. a s = []" abbreviation (input) + pred_list_append :: "('a \ 'b list) \ ('a \ 'b list) \ 'a \ 'b list" (infixr "\<^bold>@" 65) where + "xs \<^bold>@ ys \ \s. xs s @ ys s" + +abbreviation (input) pred_pair :: "('a \ 'b) \ ('a \ 'c) \ 'a \ 'b \ 'c" (infixr "\<^bold>\" 60) where "a \<^bold>\ b \ \s. (a s, b s)" abbreviation (input) pred_singleton :: "('a \ 'b) \ 'a \ 'b set" where "pred_singleton x \ \s. {x s}" (*<*) end (*>*) diff --git a/thys/ConcurrentIMP/CIMP_unbounded_buffer_ex.thy b/thys/ConcurrentIMP/CIMP_unbounded_buffer_ex.thy deleted file mode 100644 --- a/thys/ConcurrentIMP/CIMP_unbounded_buffer_ex.thy +++ /dev/null @@ -1,168 +0,0 @@ -(*<*) -(* - * 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 CIMP_unbounded_buffer_ex -imports - CIMP - "HOL-Library.Prefix_Order" -begin - -abbreviation - Receive :: "'location \ 'channel \ ('val \ 'state \ 'state) - \ (unit, 'location, 'channel \ 'val, 'state) com" ("\_\/ _\_") -where - "\l\ ch\f \ \l\ Response (\quest s. if fst quest = ch then {(f (snd quest) s, ())} else {})" - -abbreviation - Send :: "'location \ 'channel \ ('state \ 'val) - \ (unit, 'location, 'channel \ 'val, 'state) com" ("\_\/ _\_") -where - "\l\ ch\f \ \l\ Request (\s. (ch, f s)) (\ans s. {s})" - -lemma butlastE: - "\ butlast xs = ys; xs \ []; \z. xs = ys @ [z] \ P \ \ P" -by (induct xs rule: rev_induct) auto - -(*>*) -section\Unbounded buffer example\ - -text\ - -\label{sec:unbounded_place_buffer} - -This is more literally Kai's example from his notes titled -\emph{Proving an Asynchronous Message Passing Program Correct}, 2011. - -\ - -datatype ex_chname = \12 | \23 -type_synonym ex_val = nat -type_synonym ex_ls = "ex_val list" -type_synonym ex_ch = "ex_chname \ ex_val" -datatype ex_loc = \4 | \5 | c1 | r12 | r23 | s23 | s12 -datatype ex_proc = p1 | p2 | p3 - -type_synonym ex_pgm = "(unit, ex_loc, ex_ch, ex_ls) com" -type_synonym ex_pred = "(unit, ex_loc, ex_proc, ex_ch, ex_ls) pred" -type_synonym ex_state = "(unit, ex_loc, ex_proc, ex_ch, ex_ls) global_state" -type_synonym ex_system = "(unit, ex_loc, ex_proc, ex_ch, ex_ls) system" -type_synonym ex_history = "(ex_ch \ unit) list" - -text\ - -FIXME a bit fake: the local state for the producer process contains -all values produced. - -\ - -primrec ex_pgms :: "ex_proc \ ex_pgm" where - "ex_pgms p1 = LOOP DO \c1\ LocalOp (\xs. { xs @ [x] |x. True }) ;; \s12\ \12\last OD" -| "ex_pgms p2 = LOOP DO \r12\ \12\(\x xs. xs @ [x]) - \ \\4\ IF (\s. length s > 0) THEN \s23\ Request (\s. (\23, hd s)) (\ans s. {tl s}) FI - OD" -| "ex_pgms p3 = LOOP DO \r23\ \23\(\x xs. xs @ [x]) OD" - -abbreviation ex_init :: "(ex_proc \ ex_ls) \ bool" where - "ex_init f \ \p. f p = []" - -abbreviation ex_system :: "ex_system" where - "ex_system \ (ex_pgms, ex_init)" - -definition filter_on_channel :: "ex_chname \ ex_history \ ex_val list" where - "filter_on_channel ch \ map (snd \ fst) \ filter ((=) ch \ fst \ fst)" - -lemma filter_on_channel_simps [simp]: - "filter_on_channel ch [] = []" - "filter_on_channel ch (xs @ ys) = filter_on_channel ch xs @ filter_on_channel ch ys" - "filter_on_channel ch (((ch', v), resp) # vals) = (if ch' = ch then [v] else []) @ filter_on_channel ch vals" -by (simp_all add: filter_on_channel_def) - -definition Ip1_0 :: ex_pred where - "Ip1_0 \ \s. at p1 c1 s \ filter_on_channel \12 (hist s) = s\ p1" -definition Ip1_1 :: ex_pred where - "Ip1_1 \ \s. at p1 s12 s \ length (s\ p1) > 0 \ butlast (s\ p1) = filter_on_channel \12 (hist s)" -definition Ip1_2 :: ex_pred where - "Ip1_2 \ \s. filter_on_channel \12 (hist s) \ s\ p1" - -definition Ip2_0 :: ex_pred where - "Ip2_0 \ \s. filter_on_channel \12 (hist s) = filter_on_channel \23 (hist s) @ s\ p2" -(* We would get this for free from a proper VCG. *) -definition Ip2_1 :: ex_pred where - "Ip2_1 \ \s. at p2 s23 s \ length (s\ p2) > 0" - -definition Ip3_0 :: ex_pred where - "Ip3_0 \ \s. s\ p3 = filter_on_channel \23 (hist s)" - -definition I_pred :: ex_pred where - "I_pred \ pred_conjoin [ Ip1_0, Ip1_1, Ip1_2, Ip2_0, Ip2_1, Ip3_0 ]" - -lemmas I_defs = I_pred_def Ip1_0_def Ip1_1_def Ip1_2_def Ip2_0_def Ip2_1_def Ip3_0_def - -text\ - -The local state of @{const "p3"} is some prefix of the local state of -@{const "p1"}. - -\ - -definition Etern_pred :: ex_pred where - "Etern_pred \ \s. s\ p3 \ s\ p1" - -lemma correct_system: - "I_pred s \ Etern_pred s" -(*<*) -by (clarsimp simp: Etern_pred_def I_defs less_eq_list_def prefix_def) - -lemma p1_c1[simplified, intro]: - "ex_pgms, p1, lconst (\l. l = s12) \ \I_pred\ \c1\ LocalOp (\xs. { xs @ [x] |x. True })" -apply (rule vcg.intros) -apply (clarsimp simp: I_defs atS_def) -done - -lemma p1_s12[intro]: - "ex_pgms, p1, lconst (\l. l = c1) \ \I_pred\ \s12\ \12\last" -apply (rule vcg.intros) -apply (rename_tac p') -apply (case_tac p', simp_all) -apply (clarsimp simp: I_defs atS_def elim!: butlastE) -done - -lemma p2_s23[intro]: - "ex_pgms, p2, lconst (\l. l = r12 \ l = \4) \ \I_pred\ \s23\ Request (\s. (\23, hd s)) (\ans s. {tl s})" -apply (rule vcg.intros) -apply (rename_tac p') -apply (case_tac p', simp_all) -apply (clarsimp simp: I_defs atS_def) -done - -lemma p2_pi4[intro]: - "ex_pgms, p2, lcond (\l. l = s23) (\l. l = r12 \ l = \4) \ \I_pred\ \\4\ IF \s. s \ [] THEN c' FI" -apply (rule vcg.intros) -apply (clarsimp simp: I_defs atS_def split: lcond_splits) -done - -(*>*) -text\\ - -lemma "s \ reachable_states ex_system \ I_pred (mkP s)" -(*<*) -apply (erule VCG) - apply (clarsimp dest!: initial_statesD simp: I_defs atS_def) -apply simp -apply (rename_tac p) -apply (case_tac p) - apply auto -done -(*>*) -(*<*) - -end -(*>*) diff --git a/thys/ConcurrentIMP/CIMP_vcg.thy b/thys/ConcurrentIMP/CIMP_vcg.thy --- a/thys/ConcurrentIMP/CIMP_vcg.thy +++ b/thys/ConcurrentIMP/CIMP_vcg.thy @@ -1,436 +1,793 @@ (*<*) (* * 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 CIMP_vcg imports CIMP_lang begin (*>*) -subsection\Simple-minded Hoare Logic/VCG for CIMP\ +section\ State-based invariants \label{sec:cimp-invariants} \ + +text\ + +We provide a simple-minded verification condition generator (VCG) for this language, providing +support for establishing state-based invariants. It is just one way of reasoning about CIMP programs +and is proven sound wrt to the CIMP semantics. + +Our approach follows @{cite [cite_macro=citet] +"DBLP:journals/acta/Lamport80" and "DBLP:journals/toplas/LamportS84"} +(and the later @{cite [cite_macro=citet] "Lamport:2002"}) and closely +related work by @{cite [cite_macro=citet] "AptFrancezDeRoever:1980"}, +@{cite [cite_macro=citet] "CousotCousot:1980"} and @{cite +[cite_macro=citet] "DBLP:journals/acta/LevinG81"}, who suggest the +incorporation of a history variable. @{cite [cite_macro=citet] +"CousotCousot:1980"} apparently contains a completeness proof. +Lamport mentions that this technique was well-known in the mid-80s +when he proposed the use of prophecy variables\footnote{@{url +"https://lamport.azurewebsites.net/pubs/pubs.html"}}. See also @{cite +[cite_macro=citet] "deRoeverEtAl:2001"} for an extended discussion of +some of this. + +\ + +declare small_step.intros[intro] + +inductive_cases small_step_inv: + "(\l\ Request action val # cs, ls) \\<^bsub>a\<^esub> s'" + "(\l\ Response action # cs, ls) \\<^bsub>a\<^esub> s'" + "(\l\ LocalOp R # cs, ls) \\<^bsub>a\<^esub> s'" + "(\l\ IF b THEN c FI # cs, ls) \\<^bsub>a\<^esub> s'" + "(\l\ IF b THEN c1 ELSE c2 FI # cs, ls) \\<^bsub>a\<^esub> s'" + "(\l\ WHILE b DO c OD # cs, ls) \\<^bsub>a\<^esub> s'" + "(LOOP DO c OD # cs, ls) \\<^bsub>a\<^esub> s'" + +lemma small_step_stuck: + "\ ([], s) \\<^bsub>\\<^esub> c'" +by (auto elim: small_step.cases) + +declare system_step.intros[intro] + +text\ + +By default we ask the simplifier to rewrite @{const "atS"} using +ambient @{const "AT"} information. + +\ + +lemma atS_state_weak_cong[cong]: + "AT s p = AT s' p \ atS p ls s \ atS p ls s'" +by (auto simp: atS_def) + +text\ + +We provide an incomplete set of basic rules for label sets. + +\ + +lemma atS_simps: + "\atS p {} s" + "atS p {l} s \ at p l s" + "\at p l s; l \ ls\ \ atS p ls s" + "(\l. at p l s \ l \ ls) \ \atS p ls s" +by (auto simp: atS_def) + +lemma atS_mono: + "\atS p ls s; ls \ ls'\ \ atS p ls' s" +by (auto simp: atS_def) + +lemma atS_un: + "atS p (l \ l') s \ atS p l s \ atS p l' s" +by (auto simp: atS_def) + +lemma atLs_disj_union[simp]: + "(atLs p label0 \<^bold>\ atLs p label1) = atLs p (label0 \ label1)" +unfolding atLs_def by simp + +lemma atLs_insert_disj: + "atLs p (insert l label0) = (atL p l \<^bold>\ atLs p label0)" +by simp + +lemma small_step_terminated: + "s \\<^bsub>x\<^esub> s' \ atCs (fst s) = {} \ atCs (fst s') = {}" +by (induct pred: small_step) auto + +lemma atC_not_empty: + "atC c \ {}" +by (induct c) auto + +lemma atCs_empty: + "atCs cs = {} \ cs = []" +by (induct cs) (auto simp: atC_not_empty) + +lemma terminated_no_commands: + assumes "terminated p sh" + shows "\s. GST sh p = ([], s)" +using assms unfolding atLs_def AT_def by (metis atCs_empty prod.collapse singletonD) + +lemma terminated_GST_stable: + assumes "system_step q sh' sh" + assumes "terminated p sh" + shows "GST sh p = GST sh' p" +using assms by (auto dest!: terminated_no_commands simp: small_step_stuck elim!: system_step.cases) + +lemma terminated_stable: + assumes "system_step q sh' sh" + assumes "terminated p sh" + shows "terminated p sh'" +using assms unfolding atLs_def AT_def +by (fastforce split: if_splits prod.splits + dest: small_step_terminated + elim!: system_step.cases) + +lemma system_step_pls_nonempty: + assumes "system_step pls sh' sh" + shows "pls \ {}" +using assms by cases simp_all + +lemma system_step_no_change: + assumes "system_step ps sh' sh" + assumes "p \ ps" + shows "GST sh' p = GST sh p" +using assms by cases simp_all + +lemma initial_stateD: + assumes "initial_state sys s" + shows "AT (\GST = s, HST = []\) = atC \ PGMs sys \ INIT sys (\GST = s, HST = []\)\ \ (\p l. \taken p l \GST = s, HST = []\)" +using assms unfolding initial_state_def split_def o_def LST_def AT_def taken_def by simp + +lemma initial_states_initial[iff]: + assumes "initial_state sys s" + shows "at p l (\GST = s, HST = []\) \ l \ atC (PGMs sys p)" +using assms unfolding initial_state_def split_def AT_def by simp + +definition + reachable_state :: "('answer, 'location, 'proc, 'question, 'state, 'ext) pre_system_ext + \ ('answer, 'location, 'proc, 'question, 'state) state_pred" +where + "reachable_state sys s \ (\\ i. prerun sys \ \ \ i = s)" + +lemma reachable_stateE: + assumes "reachable_state sys sh" + assumes "\\ i. prerun sys \ \ P (\ i)" + shows "P sh" +using assms unfolding reachable_state_def by blast + +lemma prerun_reachable_state: + assumes "prerun sys \" + shows "reachable_state sys (\ i)" +using assms unfolding prerun_def LTL.defs system_step_reflclp_def reachable_state_def by auto + +lemma reachable_state_induct[consumes 1, case_names init LocalStep CommunicationStep, induct set: reachable_state]: + assumes r: "reachable_state sys sh" + assumes i: "\s. initial_state sys s \ P \GST = s, HST = []\" + assumes l: "\sh ls' p. \reachable_state sys sh; P sh; GST sh p \\<^bsub>\\<^esub> ls'\ \ P \GST = (GST sh)(p := ls'), HST = HST sh\" + assumes c: "\sh ls1' ls2' p1 p2 \ \. + \reachable_state sys sh; P sh; + GST sh p1 \\<^bsub>\\, \\\<^esub> ls1'; GST sh p2 \\<^bsub>\\, \\\<^esub> ls2'; p1 \ p2 \ + \ P \GST = (GST sh)(p1 := ls1', p2 := ls2'), HST = HST sh @ [(\, \)]\" + shows "P sh" +using r +proof(rule reachable_stateE) + fix \ i assume "prerun sys \" show "P (\ i)" + proof(induct i) + case 0 from \prerun sys \\ show ?case + unfolding prerun_def by (metis (full_types) i old.unit.exhaust system_state.surjective) + next + case (Suc i) with \prerun sys \\ show ?case +unfolding prerun_def LTL.defs system_step_reflclp_def reachable_state_def +apply clarsimp +apply (drule_tac x=i in spec) +apply (erule disjE; clarsimp) +apply (erule system_step.cases; clarsimp) + apply (metis (full_types) \prerun sys \\ l old.unit.exhaust prerun_reachable_state system_state.surjective) +apply (metis (full_types) \prerun sys \\ c old.unit.exhaust prerun_reachable_state system_state.surjective) +done + qed +qed + +lemma prerun_valid_TrueI: + shows "sys \\<^bsub>pre\<^esub> \True\" +unfolding prerun_valid_def by simp + +lemma prerun_valid_conjI: + assumes "sys \\<^bsub>pre\<^esub> P" + assumes "sys \\<^bsub>pre\<^esub> Q" + shows "sys \\<^bsub>pre\<^esub> P \<^bold>\ Q" +using assms unfolding prerun_valid_def always_def by simp + +lemma valid_prerun_lift: + assumes "sys \\<^bsub>pre\<^esub> I" + shows "sys \ \\I\" +using assms unfolding prerun_valid_def valid_def run_def by blast + +lemma prerun_valid_induct: + assumes "\\. prerun sys \ \ \I\ \" + assumes "\\. prerun sys \ \ (\I\ \<^bold>\ (\\I\)) \" + shows "sys \\<^bsub>pre\<^esub> I" +unfolding prerun_valid_def using assms by (simp add: always_induct) + +lemma prerun_validI: + assumes "\s. reachable_state sys s \ I s" + shows "sys \\<^bsub>pre\<^esub> I" +unfolding prerun_valid_def using assms by (simp add: alwaysI prerun_reachable_state) + +lemma prerun_validE: + assumes "reachable_state sys s" + assumes "sys \\<^bsub>pre\<^esub> I" + shows "I s" +using assms unfolding prerun_valid_def +by (metis alwaysE reachable_stateE suffix_state_prop) + + +subsubsection\Relating reachable states to the initial programs \label{sec:cimp-decompose-small-step}\ + +text\ + +To usefully reason about the control locations presumably embedded in +the single global invariant, we need to link the programs we have in +reachable state \s\ to the programs in the initial states. The +\fragments\ function decomposes the program into statements +that can be directly executed (\S\ref{sec:cimp-decompose}). We also +compute the locations we could be at after executing that statement as +a function of the process's local state. + +Eliding the bodies of \IF\ and \WHILE\ statements +yields smaller (but equivalent) proof obligations. + +\ + +type_synonym ('answer, 'location, 'question, 'state) loc_comp + = "'state \ 'location set" + +fun lconst :: "'location set \ ('answer, 'location, 'question, 'state) loc_comp" where + "lconst lp s = lp" + +definition lcond :: "'location set \ 'location set \ 'state bexp + \ ('answer, 'location, 'question, 'state) loc_comp" where + "lcond lp lp' b s = (if b s then lp else lp')" + +lemma lcond_split: + "Q (lcond lp lp' b s) \ (b s \ Q lp) \ (\b s \ Q lp')" +unfolding lcond_def by (simp split: if_splits) + +lemma lcond_split_asm: + "Q (lcond lp lp' b s) \ \ ((b s \ \Q lp) \ (\b s \ \ Q lp'))" +unfolding lcond_def by (simp split: if_splits) + +lemmas lcond_splits = lcond_split lcond_split_asm + +fun + fragments :: "('answer, 'location, 'question, 'state) com + \ 'location set + \ ( ('answer, 'location, 'question, 'state) com + \ ('answer, 'location, 'question, 'state) loc_comp ) set" +where + "fragments (\l\ IF b THEN c FI) aft + = { (\l\ IF b THEN c' FI, lcond (atC c) aft b) |c'. True } + \ fragments c aft" +| "fragments (\l\ IF b THEN c1 ELSE c2 FI) aft + = { (\l\ IF b THEN c1' ELSE c2' FI, lcond (atC c1) (atC c2) b) |c1' c2'. True } + \ fragments c1 aft \ fragments c2 aft" +| "fragments (LOOP DO c OD) aft = fragments c (atC c)" +| "fragments (\l\ WHILE b DO c OD) aft + = fragments c {l} \ { (\l\ WHILE b DO c' OD, lcond (atC c) aft b) |c'. True }" +| "fragments (c1;; c2) aft = fragments c1 (atC c2) \ fragments c2 aft" +| "fragments (c1 \ c2) aft = fragments c1 aft \ fragments c2 aft" +| "fragments c aft = { (c, lconst aft) }" + +fun + fragmentsL :: "('answer, 'location, 'question, 'state) com list + \ ( ('answer, 'location, 'question, 'state) com + \ ('answer, 'location, 'question, 'state) loc_comp ) set" +where + "fragmentsL [] = {}" +| "fragmentsL [c] = fragments c {}" +| "fragmentsL (c # c' # cs) = fragments c (atC c') \ fragmentsL (c' # cs)" + +abbreviation + fragmentsLS :: "('answer, 'location, 'question, 'state) local_state + \ ( ('answer, 'location, 'question, 'state) com + \ ('answer, 'location, 'question, 'state) loc_comp ) set" +where + "fragmentsLS s \ fragmentsL (cPGM s)" + +text\ + +We show that taking system steps preserves fragments. + +\ + +lemma small_step_fragmentsLS: + assumes "s \\<^bsub>\\<^esub> s'" + shows "fragmentsLS s' \ fragmentsLS s" +using assms by induct (case_tac [!] cs, auto) + +lemma reachable_state_fragmentsLS: + assumes "reachable_state sys sh" + shows "fragmentsLS (GST sh p) \ fragments (PGMs sys p) {}" +using assms +by (induct rule: reachable_state_induct) + (auto simp: initial_state_def dest: subsetD[OF small_step_fragmentsLS]) + +inductive + basic_com :: "('answer, 'location, 'question, 'state) com \ bool" +where + "basic_com (\l\ Request action val)" +| "basic_com (\l\ Response action)" +| "basic_com (\l\ LocalOp R)" +| "basic_com (\l\ IF b THEN c FI)" +| "basic_com (\l\ IF b THEN c1 ELSE c2 FI)" +| "basic_com (\l\ WHILE b DO c OD)" + +lemma fragments_basic_com: + assumes "(c', aft') \ fragments c aft" + shows "basic_com c'" +using assms by (induct c arbitrary: aft) (auto intro: basic_com.intros) + +lemma fragmentsL_basic_com: + assumes "(c', aft') \ fragmentsL cs" + shows "basic_com c'" +using assms +apply (induct cs) + apply simp +apply (case_tac cs) + apply (auto simp: fragments_basic_com) +done + +text\ + +To reason about system transitions we need to identify which basic +statement gets executed next. To that end we factor out the recursive +cases of the @{term "small_step"} semantics into \emph{contexts}, +which isolate the \basic_com\ commands with immediate +externally-visible behaviour. Note that non-determinism means that +more than one \basic_com\ can be enabled at a time. + +The representation of evaluation contexts follows @{cite +[cite_macro=citet] "DBLP:journals/jar/Berghofer12"}. This style of +operational semantics was originated by @{cite [cite_macro=citet] +"FelleisenHieb:1992"}. + +\ + +type_synonym ('answer, 'location, 'question, 'state) ctxt + = "(('answer, 'location, 'question, 'state) com \ ('answer, 'location, 'question, 'state) com) + \ (('answer, 'location, 'question, 'state) com \ ('answer, 'location, 'question, 'state) com list)" + +inductive_set + ctxt :: "('answer, 'location, 'question, 'state) ctxt set" +where + C_Hole: "(id, \[]\) \ ctxt" +| C_Loop: "(E, fctxt) \ ctxt \ (\c1. LOOP DO E c1 OD, \c1. fctxt c1 @ [LOOP DO E c1 OD]) \ ctxt" +| C_Seq: "(E, fctxt) \ ctxt \ (\c1. E c1;; c2, \c1. fctxt c1 @ [c2]) \ ctxt" +| C_Choose1: "(E, fctxt) \ ctxt \ (\c1. E c1 \ c2, fctxt) \ ctxt" +| C_Choose2: "(E, fctxt) \ ctxt \ (\c2. c1 \ E c2, fctxt) \ ctxt" + +text\ + +We can decompose a small step into a context and a @{const "basic_com"}. + +\ + +fun + decompose_com :: "('answer, 'location, 'question, 'state) com + \ ( ('answer, 'location, 'question, 'state) com + \ ('answer, 'location, 'question, 'state) ctxt ) set" +where + "decompose_com (LOOP DO c1 OD) = { (c, \t. LOOP DO ictxt t OD, \t. fctxt t @ [LOOP DO ictxt t OD]) |c fctxt ictxt. (c, ictxt, fctxt) \ decompose_com c1 }" +| "decompose_com (c1;; c2) = { (c, \t. ictxt t;; c2, \t. fctxt t @ [c2]) |c fctxt ictxt. (c, ictxt, fctxt) \ decompose_com c1 }" +| "decompose_com (c1 \ c2) = { (c, \t. ictxt t \ c2, fctxt) |c fctxt ictxt. (c, ictxt, fctxt) \ decompose_com c1 } + \ { (c, \t. c1 \ ictxt t, fctxt) |c fctxt ictxt. (c, ictxt, fctxt) \ decompose_com c2 }" +| "decompose_com c = {(c, id, \[]\)}" + +definition + decomposeLS :: "('answer, 'location, 'question, 'state) local_state + \ ( ('answer, 'location, 'question, 'state) com + \ (('answer, 'location, 'question, 'state) com \ ('answer, 'location, 'question, 'state) com) + \ (('answer, 'location, 'question, 'state) com \ ('answer, 'location, 'question, 'state) com list) ) set" +where + "decomposeLS s = (case cPGM s of c # _ \ decompose_com c | _ \ {})" + +lemma ctxt_inj: + assumes "(E, fctxt) \ ctxt" + assumes "E x = E y" + shows "x = y" +using assms by (induct set: ctxt) auto + +lemma decompose_com_non_empty: "decompose_com c \ {}" +by (induct c) auto + +lemma decompose_com_basic_com: + assumes "(c', ctxts) \ decompose_com c" + shows "basic_com c'" +using assms by (induct c arbitrary: c' ctxts) (auto intro: basic_com.intros) + +lemma decomposeLS_basic_com: + assumes "(c', ctxts) \ decomposeLS s" + shows "basic_com c'" +using assms unfolding decomposeLS_def by (simp add: decompose_com_basic_com split: list.splits) + +lemma decompose_com_ctxt: + assumes "(c', ctxts) \ decompose_com c" + shows "ctxts \ ctxt" +using assms by (induct c arbitrary: c' ctxts) (auto intro: ctxt.intros) + +lemma decompose_com_ictxt: + assumes "(c', ictxt, fctxt) \ decompose_com c" + shows "ictxt c' = c" +using assms by (induct c arbitrary: c' ictxt fctxt) auto + +lemma decompose_com_small_step: + assumes as: "(c' # fctxt c' @ cs, s) \\<^bsub>\\<^esub> s'" + assumes ds: "(c', ictxt, fctxt) \ decompose_com c" + shows "(c # cs, s) \\<^bsub>\\<^esub> s'" +using decompose_com_ctxt[OF ds] as decompose_com_ictxt[OF ds] +by (induct ictxt fctxt arbitrary: c cs) + (cases s', fastforce simp: fun_eq_iff dest: ctxt_inj)+ + +theorem context_decompose: + "s \\<^bsub>\\<^esub> s' \ (\(c, ictxt, fctxt) \ decomposeLS s. + cPGM s = ictxt c # tl (cPGM s) + \ (c # fctxt c @ tl (cPGM s), cTKN s, cLST s) \\<^bsub>\\<^esub> s' + \ (\l\atC c. cTKN s' = Some l))" (is "?lhs = ?rhs") +proof(rule iffI) + assume ?lhs then show ?rhs + unfolding decomposeLS_def + proof(induct rule: small_step.induct) + case (Choose1 c1 cs s \ cs' s' c2) then show ?case + apply clarsimp + apply (rename_tac c ictxt fctxt) + apply (rule_tac x="(c, \t. ictxt t \ c2, fctxt)" in bexI) + apply auto + done + next + case (Choose2 c2 cs s \ cs' s' c1) then show ?case + apply clarsimp + apply (rename_tac c ictxt fctxt) + apply (rule_tac x="(c, \t. c1 \ ictxt t, fctxt)" in bexI) + apply auto + done + qed fastforce+ +next + assume ?rhs then show ?lhs + unfolding decomposeLS_def + by (cases s) (auto split: list.splits dest: decompose_com_small_step) +qed + +text\ + +While we only use this result left-to-right (to decompose a small step +into a basic one), this equivalence shows that we lose no information +in doing so. + +Decomposing a compound command preserves @{const \fragments\} too. + +\ + +fun + loc_compC :: "('answer, 'location, 'question, 'state) com + \ ('answer, 'location, 'question, 'state) com list + \ ('answer, 'location, 'question, 'state) loc_comp" +where + "loc_compC (\l\ IF b THEN c FI) cs = lcond (atC c) (atCs cs) b" +| "loc_compC (\l\ IF b THEN c1 ELSE c2 FI) cs = lcond (atC c1) (atC c2) b" +| "loc_compC (LOOP DO c OD) cs = lconst (atC c)" +| "loc_compC (\l\ WHILE b DO c OD) cs = lcond (atC c) (atCs cs) b" +| "loc_compC c cs = lconst (atCs cs)" + +lemma decompose_fragments: + assumes "(c, ictxt, fctxt) \ decompose_com c0" + shows "(c, loc_compC c (fctxt c @ cs)) \ fragments c0 (atCs cs)" +using assms +proof(induct c0 arbitrary: c ictxt fctxt cs) + case (Loop c01 c ictxt fctxt cs) + from Loop.prems Loop.hyps(1)[where cs="ictxt c # cs"] show ?case by (auto simp: decompose_com_ictxt) +next + case (Seq c01 c02 c ictxt fctxt cs) + from Seq.prems Seq.hyps(1)[where cs="c02 # cs"] show ?case by auto +qed auto + +lemma at_decompose: + assumes "(c, ictxt, fctxt) \ decompose_com c0" + shows "atC c \ atC c0" +using assms by (induct c0 arbitrary: c ictxt fctxt; fastforce) + +lemma at_decomposeLS: + assumes "(c, ictxt, fctxt) \ decomposeLS s" + shows "atC c \ atCs (cPGM s)" +using assms unfolding decomposeLS_def by (auto simp: at_decompose split: list.splits) + +lemma decomposeLS_fragmentsLS: + assumes "(c, ictxt, fctxt) \ decomposeLS s" + shows "(c, loc_compC c (fctxt c @ tl (cPGM s))) \ fragmentsLS s" +using assms +proof(cases "cPGM s") + case (Cons d ds) + with assms decompose_fragments[where cs="ds"] show ?thesis + by (cases ds) (auto simp: decomposeLS_def) +qed (simp add: decomposeLS_def) + +lemma small_step_loc_compC: + assumes "basic_com c" + assumes "(c # cs, ls) \\<^bsub>\\<^esub> ls'" + shows "loc_compC c cs (snd ls) = atCs (cPGM ls')" +using assms by (fastforce elim: basic_com.cases elim!: small_step_inv split: lcond_splits) + +text\ + +The headline result allows us to constrain the initial and final states +of a given small step in terms of the original programs, provided the +initial state is reachable. + +\ + +theorem decompose_small_step: + assumes "GST sh p \\<^bsub>\\<^esub> ps'" + assumes "reachable_state sys sh" + obtains c cs aft + where "(c, aft) \ fragments (PGMs sys p) {}" + and "atC c \ atCs (cPGM (GST sh p))" + and "aft (cLST (GST sh p)) = atCs (cPGM ps')" + and "(c # cs, cTKN (GST sh p), cLST (GST sh p)) \\<^bsub>\\<^esub> ps'" + and "\l\atC c. cTKN ps' = Some l" +using assms +apply - +apply (frule iffD1[OF context_decompose]) +apply clarsimp +apply (frule decomposeLS_fragmentsLS) +apply (frule at_decomposeLS) +apply (frule (1) subsetD[OF reachable_state_fragmentsLS]) +apply (frule decomposeLS_basic_com) +apply (frule (1) small_step_loc_compC) +apply simp +done + +text\ + +Reasoning by induction over the reachable states +with @{thm [source] "decompose_small_step"} is quite tedious. We +provide a very simple VCG that generates friendlier local proof +obligations in \S\ref{sec:vcg}. + +\ + + +subsection\Simple-minded Hoare Logic/VCG for CIMP \label{sec:vcg}\ text\ \label{sec:cimp-vcg} We do not develop a proper Hoare logic or full VCG for CIMP: this machinery merely packages up the subgoals that arise from induction over the reachable states (\S\ref{sec:cimp-invariants}). This is somewhat in the spirit of @{cite [cite_macro=citet] "Ridge:2009"}. Note that this approach is not compositional: it consults the original system to find matching communicating pairs, and \aft\ tracks the labels of possible successor statements. More serious Hoare logics are provided by @{cite [cite_macro=citet] "DBLP:journals/acta/Lamport80" and "DBLP:journals/toplas/LamportS84" and "CousotCousot89-IC"}. Intuitively we need to discharge a proof obligation for either @{const "Request"}s or @{const "Response"}s but not both. Here we choose to focus on @{const "Request"}s as we expect to have more local information available about these. \ inductive vcg :: "('answer, 'location, 'proc, 'question, 'state) programs \ 'proc \ ('answer, 'location, 'question, 'state) loc_comp - \ ('answer, 'location, 'proc, 'question, 'state) pred + \ ('answer, 'location, 'proc, 'question, 'state) state_pred \ ('answer, 'location, 'question, 'state) com - \ ('answer, 'location, 'proc, 'question, 'state) pred - \ bool" ("_, _, _ \/ \_\/ _/ \_\") + \ ('answer, 'location, 'proc, 'question, 'state) state_pred + \ bool" ("_, _, _ \/ \_\/ _/ \_\" [11,0,0,0,0,0] 11) where - Request: "\ \aft' action' s ps' p's' l' \ s' p'. - \ pre s; (\l'\ Response action', aft') \ fragments (pgms p') \False\; p \ p'; - ps' \ val \ (LST s p); (p's', \) \ action' (action (LST s p)) (LST s p'); - at p l s; at p' l' s; - AT s' = (AT s)(p := aft (\l\ Request action val) (LST s p), - p' := aft' (\l'\ Response action') (LST s p')); - LST s' = (LST s)(p := ps', p' := p's'); - hist s' = hist s @ [(action (LST s p), \)] - \ \ post s' - \ \ pgms, p, aft \ \pre\ \l\ Request action val \post\" -| LocalOp: "\ \s ps' s'. \ pre s; ps' \ f (LST s p); - at p l s; - AT s' = (AT s)(p := aft (\l\ LocalOp f) (LST s p)); - LST s' = (LST s)(p := ps'); - hist s' = hist s - \ \ post s' - \ \ pgms, p, aft \ \pre\ \l\ LocalOp f \post\" -| Cond1: "\ \s s'. \ pre s; - at p l s; - AT s' = (AT s)(p := aft (\l\ IF b THEN t FI) (LST s p)); - LST s' = LST s; - hist s' = hist s - \ \ post s' - \ \ pgms, p, aft \ \pre\ \l\ IF b THEN t FI \post\" -| Cond2: "\ \s s'. \ pre s; - at p l s; - AT s' = (AT s)(p := aft (\l\ IF b THEN t ELSE e FI) (LST s p)); - LST s' = LST s; - hist s' = hist s - \ \ post s' - \ \ pgms, p, aft \ \pre\ \l\ IF b THEN t ELSE e FI \post\" -| While: "\ \s s'. \ pre s; - at p l s; - AT s' = (AT s)(p := aft (\l\ WHILE b DO c OD) (LST s p)); - LST s' = LST s; - hist s' = hist s - \ \ post s' - \ \ pgms, p, aft \ \pre\ \l\ WHILE b DO c OD \post\" -\ \There are no proof obligations for the following commands.\ -| Response: "pgms, p, aft \ \pre\ \l\ Response action \post\" -| Seq: "pgms, p, aft \ \pre\ c1 ;; c2 \post\" -| Loop: "pgms, p, aft \ \pre\ LOOP DO c OD \post\" -| Choose: "pgms, p, aft \ \pre\ c1 \ c2 \post\" + "\ \aft' action' s ps' p's' l' \ s' p'. + \ pre s; (\l'\ Response action', aft') \ fragments (coms p') {}; p \ p'; + ps' \ val \ (s\ p); (p's', \) \ action' (action (s\ p)) (s\ p'); + at p l s; at p' l' s; + AT s' = (AT s)(p := aft (s\ p), p' := aft' (s\ p')); + s'\ = s\(p := ps', p' := p's'); + taken p l s'; + HST s' = HST s @ [(action (s\ p), \)]; + \p''\-{p,p'}. GST s' p'' = GST s p'' + \ \ post s' + \ \ coms, p, aft \ \pre\ \l\ Request action val \post\" +| "\ \s ps' s'. + \ pre s; ps' \ f (s\ p); + at p l s; + AT s' = (AT s)(p := aft (s\ p)); + s'\ = s\(p := ps'); + taken p l s'; + HST s' = HST s; + \p''\-{p}. GST s' p'' = GST s p'' + \ \ post s' + \ \ coms, p, aft \ \pre\ \l\ LocalOp f \post\" +| "\ \s s'. + \ pre s; + at p l s; + AT s' = (AT s)(p := aft (s\ p)); + s'\ = s\; + taken p l s'; + HST s' = HST s; + \p''\-{p}. GST s' p'' = GST s p'' + \ \ post s' + \ \ coms, p, aft \ \pre\ \l\ IF b THEN t FI \post\" +| "\ \s s'. + \ pre s; + at p l s; + AT s' = (AT s)(p := aft (s\ p)); + s'\ = s\; + taken p l s'; + HST s' = HST s; + \p''\-{p}. GST s' p'' = GST s p'' + \ \ post s' + \ \ coms, p, aft \ \pre\ \l\ IF b THEN t ELSE e FI \post\" +| "\ \s s'. + \ pre s; + at p l s; + AT s' = (AT s)(p := aft (s\ p)); + s'\ = s\; + taken p l s'; + HST s' = HST s; + \p''\-{p}. GST s' p'' = GST s p'' + \ \ post s' + \ \ coms, p, aft \ \pre\ \l\ WHILE b DO c OD \post\" +\ \There are no proof obligations for the following commands, but including them makes some basic rules hold (\S\ref{sec:cimp:vcg_rules}):\ +| "coms, p, aft \ \pre\ \l\ Response action \post\" +| "coms, p, aft \ \pre\ c1 ;; c2 \post\" +| "coms, p, aft \ \pre\ LOOP DO c OD \post\" +| "coms, p, aft \ \pre\ c1 \ c2 \post\" text\ We abbreviate invariance with one-sided validity syntax. \ -abbreviation valid_inv ("_, _, _ \/ \_\/ _") where - "pgms, p, aft \ \I\ c \ pgms, p, aft \ \I\ c \I\" -(*<*) +abbreviation valid_inv ("_, _, _ \/ \_\/ _" [11,0,0,0,0] 11) where + "coms, p, aft \ \I\ c \ coms, p, aft \ \I\ c \I\" inductive_cases vcg_inv: - "pgms, p, aft \ \pre\ \l\ Request action val \post\" - "pgms, p, aft \ \pre\ \l\ LocalOp f \post\" - "pgms, p, aft \ \pre\ \l\ IF b THEN t FI \post\" - "pgms, p, aft \ \pre\ \l\ IF b THEN t ELSE e FI \post\" - "pgms, p, aft \ \pre\ LOOP DO c OD \post\" - "pgms, p, aft \ \pre\ \l\ WHILE b DO c OD \post\" - "pgms, p, aft \ \pre\ \l\ Response action \post\" - "pgms, p, aft \ \pre\ c1 ;; c2 \post\" - "pgms, p, aft \ \pre\ c1 \ c2 \post\" + "coms, p, aft \ \pre\ \l\ Request action val \post\" + "coms, p, aft \ \pre\ \l\ LocalOp f \post\" + "coms, p, aft \ \pre\ \l\ IF b THEN t FI \post\" + "coms, p, aft \ \pre\ \l\ IF b THEN t ELSE e FI \post\" + "coms, p, aft \ \pre\ \l\ WHILE b DO c OD \post\" + "coms, p, aft \ \pre\ LOOP DO c OD \post\" + "coms, p, aft \ \pre\ \l\ Response action \post\" + "coms, p, aft \ \pre\ c1 ;; c2 \post\" + "coms, p, aft \ \pre\ Choose c1 c2 \post\" -lemma vcg_precond_cong: - "P = P' \ (pgms, p, aft \ \P\ c \Q\) \ (pgms, p, aft \ \P'\ c \Q\)" -by simp - -lemma vcg_postcond_cong: - "Q = Q' \ (pgms, p, aft \ \P\ c \Q\) \ (pgms, p, aft \ \P\ c \Q'\)" -by simp - -(*>*) text\ We tweak @{const "fragments"} by omitting @{const "Response"}s, -yielding fewer obligations. +yielding fewer obligations \ fun vcg_fragments' :: "('answer, 'location, 'question, 'state) com - \ ('location \ bool) + \ 'location set \ ( ('answer, 'location, 'question, 'state) com \ ('answer, 'location, 'question, 'state) loc_comp ) set" where - "vcg_fragments' (\l\ Response action) ls = {}" -| "vcg_fragments' (\l\ IF b THEN c FI) ls - = vcg_fragments' c ls - \ { (\l\ IF b THEN c' FI, lcond (atC c) ls) |c'. True }" -| "vcg_fragments' (\l\ IF b THEN c1 ELSE c2 FI) ls - = vcg_fragments' c2 ls \ vcg_fragments' c1 ls - \ { (\l\ IF b THEN c1' ELSE c2' FI, lcond (atC c1) (atC c2)) |c1' c2'. True }" -| "vcg_fragments' (LOOP DO c OD) ls = vcg_fragments' c (atC c)" -| "vcg_fragments' (\l'\ WHILE b DO c OD) ls - = vcg_fragments' c (\l. l = l') \ { (\l'\ WHILE b DO c' OD, lcond (atC c) ls) |c'. True }" -| "vcg_fragments' (c1 ;; c2) ls = vcg_fragments' c2 ls \ vcg_fragments' c1 (atC c2)" -| "vcg_fragments' (c1 \ c2) ls = vcg_fragments' c1 ls \ vcg_fragments' c2 ls" -| "vcg_fragments' c ls = {(c, lconst ls)}" + "vcg_fragments' (\l\ Response action) aft = {}" +| "vcg_fragments' (\l\ IF b THEN c FI) aft + = vcg_fragments' c aft + \ { (\l\ IF b THEN c' FI, lcond (atC c) aft b) |c'. True }" +| "vcg_fragments' (\l\ IF b THEN c1 ELSE c2 FI) aft + = vcg_fragments' c2 aft \ vcg_fragments' c1 aft + \ { (\l\ IF b THEN c1' ELSE c2' FI, lcond (atC c1) (atC c2) b) |c1' c2'. True }" +| "vcg_fragments' (LOOP DO c OD) aft = vcg_fragments' c (atC c)" +| "vcg_fragments' (\l\ WHILE b DO c OD) aft + = vcg_fragments' c {l} \ { (\l\ WHILE b DO c' OD, lcond (atC c) aft b) |c'. True }" +| "vcg_fragments' (c1 ;; c2) aft = vcg_fragments' c2 aft \ vcg_fragments' c1 (atC c2)" +| "vcg_fragments' (c1 \ c2) aft = vcg_fragments' c1 aft \ vcg_fragments' c2 aft" +| "vcg_fragments' c aft = {(c, lconst aft)}" abbreviation vcg_fragments :: "('answer, 'location, 'question, 'state) com \ ( ('answer, 'location, 'question, 'state) com \ ('answer, 'location, 'question, 'state) loc_comp ) set" where - "vcg_fragments c \ vcg_fragments' c \False\" -(*<*) + "vcg_fragments c \ vcg_fragments' c {}" -lemma vcg_fragments'_inc: - "\ (c', lp) \ fragments c ls; \l action. c' \ \l\ Response action \ \ (c', lp) \ vcg_fragments' c ls" -by (induct c arbitrary: ls) (auto split: if_splits) +fun isResponse :: "('answer, 'location, 'question, 'state) com \ bool" where + "isResponse (\l\ Response action) \ True" +| "isResponse _ \ False" + +lemma fragments_vcg_fragments': + "\ (c, aft) \ fragments c' aft'; \isResponse c \ \ (c, aft) \ vcg_fragments' c' aft'" +by (induct c' arbitrary: aft') auto + +lemma vcg_fragments'_fragments: + "vcg_fragments' c' aft' \ fragments c' aft'" +by (induct c' arbitrary: aft') (auto 10 0) lemma VCG_step: - assumes V: "\p. \(c, afts) \ vcg_fragments (fst sys p). ((fst sys), p, afts \ \pre\ c \post\)" - assumes S: "(s, h) s\ (s', h')" - assumes R: "(s, h) \ reachable_states sys" - assumes P: "pre (mkP (s, h))" - shows "post (mkP (s', h'))" + assumes V: "\p. \(c, aft) \ vcg_fragments (PGMs sys p). PGMs sys, p, aft \ \pre\ c \post\" + assumes S: "system_step p sh' sh" + assumes R: "reachable_state sys sh" + assumes P: "pre sh" + shows "post sh'" using S proof cases - case (LocalStep p ps') with P show ?thesis + case LocalStep with P show ?thesis apply - apply (erule decompose_small_step[OF _ R]) - apply (elim basic_com.cases, simp_all) - apply (fastforce dest!: V[rule_format] vcg_fragments'_inc - elim!: small_step_inv vcg_inv - simp: fun_eq_iff LST_def AT_def mkP_def split_def)+ + apply (frule fragments_basic_com) + apply (erule basic_com.cases) + apply (fastforce dest!: fragments_vcg_fragments' V[rule_format] + elim: vcg_inv elim!: small_step_inv + simp: LST_def AT_def taken_def fun_eq_iff)+ done next - case (CommunicationStep p1 \ \ ls1' p2 ls2) with P show ?thesis + case CommunicationStep with P show ?thesis apply - apply (erule decompose_small_step[OF _ R]) apply (erule decompose_small_step[OF _ R]) - apply (elim basic_com.cases, simp_all) - apply (drule vcg_fragments'_inc, simp) - apply (drule V[rule_format]) - apply clarsimp - apply (erule vcg_inv) - apply (elim small_step_inv) - apply (fastforce simp add: fun_eq_iff LST_def AT_def mkP_def split_def) + subgoal for c cs aft c' cs' aft' + apply (frule fragments_basic_com[where c'=c]) + apply (frule fragments_basic_com[where c'=c']) + apply (elim basic_com.cases; clarsimp elim!: small_step_inv) + apply (drule fragments_vcg_fragments') + apply (fastforce dest!: V[rule_format] + elim: vcg_inv elim!: small_step_inv + simp: LST_def AT_def taken_def fun_eq_iff)+ + done done qed -(*>*) text\ -The user sees the conclusion of \V\ for each element of \vcg_fragments\. - -\ - -lemma VCG: - assumes R: "s \ reachable_states sys" - assumes I: "\s \ initial_states sys. I (mkP (s, []))" - assumes V: "\p. \(c, afts) \ vcg_fragments (fst sys p). ((fst sys), p, afts \ \I\ c)" - shows "I (mkP s)" -(*<*) -proof - - have "I (mkP (s', h'))" if B: "s0 \ initial_states sys" and S: "(s0, []) s\\<^sup>* (s', h')" for s0 s' h' - using S - proof(induct rule: rtrancl_induct2) - case refl with B show ?case by (rule I[rule_format]) - next - case (step s' h' s'' h'') with B V show ?case - by - (erule (1) VCG_step; auto simp: reachable_states_def) - qed - with I R show ?thesis by (cases s) (clarsimp simp: reachable_states_def) -qed -(*>*) - -subsubsection\VCG rules\ - -text\ - -We can develop some (but not all) of the familiar Hoare rules; see -@{cite [cite_macro=citet] "DBLP:journals/acta/Lamport80"} and the -seL4/l4.verified lemma buckets for inspiration. We avoid many of the -issues Lamport mentions as we only treat basic (atomic) commands. +The user sees the conclusion of \V\ for each element of @{const \vcg_fragments\}. \ -context - fixes pgms :: "('answer, 'location, 'proc, 'question, 'state) programs" - fixes p :: "'proc" - fixes afts :: "('answer, 'location, 'question, 'state) loc_comp" -begin - -abbreviation - valid_syn :: "('answer, 'location, 'proc, 'question, 'state) pred - \ ('answer, 'location, 'question, 'state) com - \ ('answer, 'location, 'proc, 'question, 'state) pred \ bool" where - "valid_syn P c Q \ pgms, p, afts \ \P\ c \Q\" -notation valid_syn ("\_\/ _/ \_\") - -abbreviation - valid_inv_syn :: "('answer, 'location, 'proc, 'question, 'state) pred - \ ('answer, 'location, 'question, 'state) com \ bool" where - "valid_inv_syn P c \ \P\ c \P\" -notation valid_inv_syn ("\_\/ _") - -lemma vcg_True: - "\P\ c \\True\\" -by (cases c) (fastforce elim!: vcg_inv intro: vcg.intros)+ - -lemma vcg_conj: - "\ \I\ c \Q\; \I\ c \R\ \ \ \I\ c \Q \<^bold>\ R\" -by (cases c) (fastforce elim!: vcg_inv intro: vcg.intros)+ - -lemma vcg_pre_imp: - "\ \s. P s \ Q s; \Q\ c \R\ \ \ \P\ c \R\" -by (cases c) (fastforce elim!: vcg_inv intro: vcg.intros)+ - -lemmas vcg_pre = vcg_pre_imp[rotated] - -lemma vcg_post_imp: - "\ \s. Q s \ R s; \P\ c \Q\ \ \ \P\ c \R\" -by (cases c) (fastforce elim!: vcg_inv intro: vcg.intros)+ - -lemma vcg_prop[intro]: - "\\P\\ c" -by (cases c) (fastforce intro: vcg.intros)+ - -lemma vcg_drop_imp: - assumes "\P\ c \Q\" - shows "\P\ c \R \<^bold>\ Q\" -using assms -by (cases c) (fastforce elim!: vcg_inv intro: vcg.intros)+ - -lemma vcg_conj_lift: - assumes x: "\P\ c \Q\" - assumes y: "\P'\ c \Q'\" - shows "\P \<^bold>\ P'\ c \Q \<^bold>\ Q'\" -apply (rule vcg_conj) - apply (rule vcg_pre[OF x], simp) -apply (rule vcg_pre[OF y], simp) +lemma VCG_step_inv_stable: + assumes V: "\p. \(c, aft) \ vcg_fragments (PGMs sys p). PGMs sys, p, aft \ \I\ c" + assumes "prerun sys \" + shows "(\I\ \<^bold>\ \\I\) \" +apply (rule alwaysI) +apply clarsimp +apply (rule nextI) +apply clarsimp +using assms(2) unfolding prerun_def +apply clarsimp +apply (erule_tac i=i in alwaysE) +unfolding system_step_reflclp_def +apply clarsimp +apply (erule disjE; clarsimp) +using VCG_step[where pre=I and post=I] V assms(2) prerun_reachable_state +apply blast done -lemma vcg_disj_lift: - assumes x: "\P\ c \Q\" - assumes y: "\P'\ c \Q'\" - shows "\P \<^bold>\ P'\ c \Q \<^bold>\ Q'\" -using assms -by (cases c) (fastforce elim!: vcg_inv intro: vcg.intros)+ - -lemma vcg_imp_lift: - assumes "\P'\ c \\<^bold>\ P\" - assumes "\Q'\ c \Q\" - shows "\P' \<^bold>\ Q'\ c \P \<^bold>\ Q\" -by (simp only: imp_conv_disj vcg_disj_lift[OF assms]) - -lemma vcg_ex_lift: - assumes "\x. \P x\ c \Q x\" - shows "\\s. \x. P x s\ c \\s. \x. Q x s\" -using assms -by (cases c) (fastforce elim!: vcg_inv intro: vcg.intros)+ - -lemma vcg_all_lift: - assumes "\x. \P x\ c \Q x\" - shows "\\s. \x. P x s\ c \\s. \x. Q x s\" -using assms -by (cases c) (fastforce elim!: vcg_inv intro: vcg.intros)+ - -lemma vcg_name_pre_state: - assumes "\s. P s \ \(=) s\ c \Q\" - shows "\P\ c \Q\" -using assms -by (cases c) (fastforce elim!: vcg_inv intro: vcg.intros)+ - -lemma vcg_lift_comp: - assumes f: "\P. \\s. P (f s :: 'a :: type)\ c" - assumes P: "\x. \Q x\ c \P x\" - shows "\\s. Q (f s) s\ c \\s. P (f s) s\" - apply (rule vcg_name_pre_state) - apply (rename_tac s) - apply (rule vcg_pre) - apply (rule vcg_post_imp[rotated]) - apply (rule vcg_conj_lift) - apply (rule_tac x="f s" in P) - apply (rule_tac P="\fs. fs = f s" in f) - apply simp - apply simp - done - -(* **************************************** *) - -subsubsection\Cheap non-interference rules\ - -text\ - -These rules magically construct VCG lifting rules from the easier to -prove \eq_imp\ facts. We don't actually use these in the GC, -but we do derive @{const "fun_upd"} equations using the same -mechanism. Thanks to Thomas Sewell for the requisite syntax magic. - -As these \eq_imp\ facts do not usefully compose, we make the -definition asymmetric (i.e., \g\ does not get a bundle of -parameters). - -\ - -definition eq_imp :: "('a \ 'b \ 'c) \ ('b \ 'e) \ bool" where - "eq_imp f g \ (\s s'. (\x. f x s = f x s') \ (g s = g s'))" - -lemma eq_impD: - "\ eq_imp f g; \x. f x s = f x s' \ \ g s = g s'" -by (simp add: eq_imp_def) - -lemma eq_imp_vcg: - assumes g: "eq_imp f g" - assumes f: "\x P. \P \ (f x)\ c" - shows "\P \ g\ c" - apply (rule vcg_name_pre_state) - apply (rename_tac s) - apply (rule vcg_pre) - apply (rule vcg_post_imp[rotated]) - apply (rule vcg_all_lift[where 'a='a]) - apply (rule_tac x=x and P="\fs. fs = f x s" in f[rule_format]) - apply simp - apply (frule eq_impD[where f=f, OF g]) - apply simp - apply simp - done - -lemma eq_imp_vcg_LST: - assumes g: "eq_imp f g" - assumes f: "\x P. \P \ (f x) \ LST\ c" - shows "\P \ g \ LST\ c" - apply (rule vcg_name_pre_state) - apply (rename_tac s) - apply (rule vcg_pre) - apply (rule vcg_post_imp[rotated]) - apply (rule vcg_all_lift[where 'a='a]) - apply (rule_tac x=x and P="\fs. fs = f x s\" in f[rule_format]) - apply simp - apply (frule eq_impD[where f=f, OF g]) - apply simp - apply simp - done - -lemma eq_imp_fun_upd: - assumes g: "eq_imp f g" - assumes f: "\x. f x (s(fld := val)) = f x s" - shows "g (s(fld := val)) = g s" -apply (rule eq_impD[OF g]) -apply (rule f) +lemma VCG: + assumes I: "\s. initial_state sys s \ I (\GST = s, HST = []\)" + assumes V: "\p. \(c, aft) \ vcg_fragments (PGMs sys p). PGMs sys, p, aft \ \I\ c" + shows "sys \\<^bsub>pre\<^esub> I" +apply (rule prerun_valid_induct) + apply (clarsimp simp: prerun_def state_prop_def) + apply (metis (full_types) I old.unit.exhaust system_state.surjective) +using VCG_step_inv_stable[OF V] apply blast done -lemma curry_forall_eq: - "(\f. P f) = (\f. P (case_prod f))" - apply safe - apply simp_all - apply (rename_tac f) - apply (drule_tac x="\x y. f (x, y)" in spec) - apply simp - done - -lemma pres_tuple_vcg: - "(\P. \P \ (\s. (f s, g s))\ c) - \ ((\P. \P \ f\ c) \ (\P. \P \ g\ c))" - apply (simp add: curry_forall_eq o_def) - apply safe - apply fast - apply fast - apply (rename_tac P) - apply (rule_tac f="f" and P="\fs s. P fs (g s)" in vcg_lift_comp, simp, simp) - done - -lemma pres_tuple_vcg_LST: - "(\P. \P \ (\s. (f s, g s)) \ LST\ c) - \ ((\P. \P \ f \ LST\ c) \ (\P. \P \ g \ LST\ c))" - apply (simp add: curry_forall_eq o_def) - apply safe - apply fast - apply fast - apply (rename_tac P) - apply (rule_tac f="\s. f s\" and P="\fs s. P fs (g s)" for g in vcg_lift_comp, simp, simp) - done - -lemmas conj_explode = conj_imp_eq_imp_imp - -end +lemmas VCG_valid = valid_prerun_lift[OF VCG, of sys I] for sys I (*<*) end (*>*) diff --git a/thys/ConcurrentIMP/CIMP_vcg_rules.thy b/thys/ConcurrentIMP/CIMP_vcg_rules.thy new file mode 100644 --- /dev/null +++ b/thys/ConcurrentIMP/CIMP_vcg_rules.thy @@ -0,0 +1,229 @@ +(*<*) +(* + * 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 CIMP_vcg_rules +imports + CIMP_vcg +begin + +(*>*) +subsubsection\ VCG rules \label{sec:cimp:vcg_rules} \ + +text\ + +We can develop some (but not all) of the familiar Hoare rules; see +@{cite [cite_macro=citet] "DBLP:journals/acta/Lamport80"} and the +seL4/l4.verified lemma buckets for inspiration. We avoid many of the +issues Lamport mentions as we only treat basic (atomic) commands. + +\ + +context + fixes coms :: "('answer, 'location, 'proc, 'question, 'state) programs" + fixes p :: "'proc" + fixes aft :: "('answer, 'location, 'question, 'state) loc_comp" +begin + +abbreviation + valid_syn :: "('answer, 'location, 'proc, 'question, 'state) state_pred + \ ('answer, 'location, 'question, 'state) com + \ ('answer, 'location, 'proc, 'question, 'state) state_pred \ bool" where + "valid_syn P c Q \ coms, p, aft \ \P\ c \Q\" +notation valid_syn ("\_\/ _/ \_\") + +abbreviation + valid_inv_syn :: "('answer, 'location, 'proc, 'question, 'state) state_pred + \ ('answer, 'location, 'question, 'state) com \ bool" where + "valid_inv_syn P c \ \P\ c \P\" +notation valid_inv_syn ("\_\/ _") + +lemma vcg_True: + "\P\ c \\True\\" +by (cases c) (fastforce elim!: vcg_inv intro: vcg.intros)+ + +lemma vcg_conj: + "\ \I\ c \Q\; \I\ c \R\ \ \ \I\ c \Q \<^bold>\ R\" +by (cases c) (fastforce elim!: vcg_inv intro: vcg.intros)+ + +lemma vcg_pre_imp: + "\ \s. P s \ Q s; \Q\ c \R\ \ \ \P\ c \R\" +by (cases c) (fastforce elim!: vcg_inv intro: vcg.intros)+ + +lemmas vcg_pre = vcg_pre_imp[rotated] + +lemma vcg_post_imp: + "\ \s. Q s \ R s; \P\ c \Q\ \ \ \P\ c \R\" +by (cases c) (fastforce elim!: vcg_inv intro: vcg.intros)+ + +lemma vcg_prop[intro]: + "\\P\\ c" +by (cases c) (fastforce intro: vcg.intros)+ + +lemma vcg_drop_imp: + assumes "\P\ c \Q\" + shows "\P\ c \R \<^bold>\ Q\" +using assms +by (cases c) (fastforce elim!: vcg_inv intro: vcg.intros)+ + +lemma vcg_conj_lift: + assumes x: "\P\ c \Q\" + assumes y: "\P'\ c \Q'\" + shows "\P \<^bold>\ P'\ c \Q \<^bold>\ Q'\" +apply (rule vcg_conj) + apply (rule vcg_pre[OF x], simp) +apply (rule vcg_pre[OF y], simp) +done + +lemma vcg_disj_lift: + assumes x: "\P\ c \Q\" + assumes y: "\P'\ c \Q'\" + shows "\P \<^bold>\ P'\ c \Q \<^bold>\ Q'\" +using assms +by (cases c) (fastforce elim!: vcg_inv intro: vcg.intros)+ + +lemma vcg_imp_lift: + assumes "\P'\ c \\<^bold>\ P\" + assumes "\Q'\ c \Q\" + shows "\P' \<^bold>\ Q'\ c \P \<^bold>\ Q\" +by (simp only: imp_conv_disj vcg_disj_lift[OF assms]) + +lemma vcg_ex_lift: + assumes "\x. \P x\ c \Q x\" + shows "\\s. \x. P x s\ c \\s. \x. Q x s\" +using assms +by (cases c) (fastforce elim!: vcg_inv intro: vcg.intros)+ + +lemma vcg_all_lift: + assumes "\x. \P x\ c \Q x\" + shows "\\s. \x. P x s\ c \\s. \x. Q x s\" +using assms +by (cases c) (fastforce elim!: vcg_inv intro: vcg.intros)+ + +lemma vcg_name_pre_state: + assumes "\s. P s \ \(=) s\ c \Q\" + shows "\P\ c \Q\" +using assms +by (cases c) (fastforce elim!: vcg_inv intro: vcg.intros)+ + +lemma vcg_lift_comp: + assumes f: "\P. \\s. P (f s :: 'a :: type)\ c" + assumes P: "\x. \Q x\ c \P x\" + shows "\\s. Q (f s) s\ c \\s. P (f s) s\" +apply (rule vcg_name_pre_state) +apply (rename_tac s) +apply (rule vcg_pre) + apply (rule vcg_post_imp[rotated]) + apply (rule vcg_conj_lift) + apply (rule_tac x="f s" in P) + apply (rule_tac P="\fs. fs = f s" in f) + apply simp +apply simp +done + + +subsubsection\Cheap non-interference rules\ + +text\ + +These rules magically construct VCG lifting rules from the easier to +prove \eq_imp\ facts. We don't actually use these in the GC, +but we do derive @{const "fun_upd"} equations using the same +mechanism. Thanks to Thomas Sewell for the requisite syntax magic. + +As these \eq_imp\ facts do not usefully compose, we make the +definition asymmetric (i.e., \g\ does not get a bundle of +parameters). + +Note that these are effectively parametricity rules. + +\ + +definition eq_imp :: "('a \ 'b \ 'c) \ ('b \ 'e) \ bool" where + "eq_imp f g \ (\s s'. (\x. f x s = f x s') \ (g s = g s'))" + +lemma eq_impD: + "\ eq_imp f g; \x. f x s = f x s' \ \ g s = g s'" +by (simp add: eq_imp_def) + +lemma eq_imp_vcg: + assumes g: "eq_imp f g" + assumes f: "\x P. \P \ (f x)\ c" + shows "\P \ g\ c" +apply (rule vcg_name_pre_state) +apply (rename_tac s) +apply (rule vcg_pre) + apply (rule vcg_post_imp[rotated]) + apply (rule vcg_all_lift[where 'a='a]) + apply (rule_tac x=x and P="\fs. fs = f x s" in f[rule_format]) + apply simp + apply (frule eq_impD[where f=f, OF g]) + apply simp +apply simp +done + +lemma eq_imp_vcg_LST: + assumes g: "eq_imp f g" + assumes f: "\x P. \P \ (f x) \ LST\ c" + shows "\P \ g \ LST\ c" +apply (rule vcg_name_pre_state) +apply (rename_tac s) +apply (rule vcg_pre) + apply (rule vcg_post_imp[rotated]) + apply (rule vcg_all_lift[where 'a='a]) + apply (rule_tac x=x and P="\fs. fs = f x s\" in f[rule_format]) + apply simp + apply (frule eq_impD[where f=f, OF g]) + apply simp +apply simp +done + +lemma eq_imp_fun_upd: + assumes g: "eq_imp f g" + assumes f: "\x. f x (s(fld := val)) = f x s" + shows "g (s(fld := val)) = g s" +apply (rule eq_impD[OF g]) +apply (rule f) +done + +lemma curry_forall_eq: + "(\f. P f) = (\f. P (case_prod f))" +by (metis case_prod_curry) + +lemma pres_tuple_vcg: + "(\P. \P \ (\s. (f s, g s))\ c) + \ ((\P. \P \ f\ c) \ (\P. \P \ g\ c))" +apply (simp add: curry_forall_eq o_def) +apply safe + apply fast + apply fast +apply (rename_tac P) +apply (rule_tac f="f" and P="\fs s. P fs (g s)" in vcg_lift_comp; simp) +done + +lemma pres_tuple_vcg_LST: + "(\P. \P \ (\s. (f s, g s)) \ LST\ c) + \ ((\P. \P \ f \ LST\ c) \ (\P. \P \ g \ LST\ c))" +apply (simp add: curry_forall_eq o_def) +apply safe + apply fast + apply fast +apply (rename_tac P) +apply (rule_tac f="\s. f s\" and P="\fs s. P fs (g s)" for g in vcg_lift_comp; simp) +done + +no_notation valid_syn ("\_\/ _/ \_\") + +end + +(*<*) + +end +(*>*) diff --git a/thys/ConcurrentIMP/Infinite_Sequences.thy b/thys/ConcurrentIMP/Infinite_Sequences.thy new file mode 100644 --- /dev/null +++ b/thys/ConcurrentIMP/Infinite_Sequences.thy @@ -0,0 +1,333 @@ +(*<*) +theory Infinite_Sequences +imports + CIMP_pred +begin + +(*>*) +section\ Infinite Sequences \label{sec:infinite_sequences}\ + +text\ + +Infinite sequences and some operations on them. + +We use the customary function-based representation. + +\ + +type_synonym 'a seq = "nat \ 'a" +type_synonym 'a seq_pred = "'a seq \ bool" + +definition suffix :: "'a seq \ nat \ 'a seq" (infixl "|\<^sub>s" 60) where + "\ |\<^sub>s i \ \j. \ (j + i)" + +primrec stake :: "nat \ 'a seq \ 'a list" where + "stake 0 \ = []" +| "stake (Suc n) \ = \ 0 # stake n (\ |\<^sub>s 1)" + +primrec shift :: "'a list \ 'a seq \ 'a seq" (infixr \@-\ 65) where + "shift [] \ = \" +| "shift (x # xs) \ = (\i. case i of 0 \ x | Suc i \ shift xs \ i)" + +(* FIXME misleading: this is \(i, i+j). Use a bundle for this notation. FIXME move *) +abbreviation interval_syn ("_'(_ \ _')" [69, 0, 0] 70) where (* FIXME priorities *) + "\(i \ j) \ stake j (\ |\<^sub>s i)" + +lemma suffix_eval: "(\ |\<^sub>s i) j = \ (j + i)" +unfolding suffix_def by simp + +lemma suffix_plus: "\ |\<^sub>s n |\<^sub>s m = \ |\<^sub>s (m + n)" +unfolding suffix_def by (simp add: add.assoc) + +lemma suffix_commute: "((\ |\<^sub>s n) |\<^sub>s m) = ((\ |\<^sub>s m) |\<^sub>s n)" +by (simp add: suffix_plus add.commute) + +lemma suffix_plus_com: "\ |\<^sub>s m |\<^sub>s n = \ |\<^sub>s (m + n)" +proof - + have "\ |\<^sub>s n |\<^sub>s m = \ |\<^sub>s (m + n)" by (rule suffix_plus) + then show "\ |\<^sub>s m |\<^sub>s n = \ |\<^sub>s (m + n)" by (simp add: suffix_commute) +qed + +lemma suffix_zero: "\ |\<^sub>s 0 = \" +unfolding suffix_def by simp + +lemma comp_suffix: "f \ \ |\<^sub>s i = (f \ \) |\<^sub>s i" +unfolding suffix_def comp_def by simp + +lemmas suffix_simps[simp] = + comp_suffix + suffix_eval + suffix_plus_com + suffix_zero + +lemma length_stake[simp]: "length (stake n s) = n" +by (induct n arbitrary: s) auto + +lemma shift_simps[simp]: + "(xs @- \) 0 = (if xs = [] then \ 0 else hd xs)" + "(xs @- \) |\<^sub>s Suc 0 = (if xs = [] then \ |\<^sub>s Suc 0 else tl xs @- \)" +by (induct xs) auto + +lemma stake_nil[simp]: + "stake i \ = [] \ i = 0" +by (cases i; clarsimp) + +lemma stake_shift: + "stake i (w @- \) = take i w @ stake (i - length w) \" +by (induct i arbitrary: w) (auto simp: neq_Nil_conv) + +lemma shift_snth_less[simp]: + assumes "i < length xs" + shows "(xs @- \) i = xs ! i" +using assms +proof(induct i arbitrary: xs) + case (Suc i xs) then show ?case by (cases xs) simp_all +qed (simp add: hd_conv_nth nth_tl) + +lemma shift_snth_ge[simp]: + assumes "i \ length xs" + shows "(xs @- \) i = \ (i - length xs)" +using assms +proof(induct i arbitrary: xs) + case (Suc i xs) then show ?case by (cases xs) simp_all +qed simp + +lemma shift_snth: + "(xs @- \) i = (if i < length xs then xs ! i else \ (i - length xs))" +by simp + +lemma suffix_shift: + "(xs @- \) |\<^sub>s i = drop i xs @- (\ |\<^sub>s i - length xs)" +proof(induct i arbitrary: xs) + case (Suc i xs) then show ?case by (cases xs) simp_all +qed simp + +lemma stake_nth[simp]: + assumes "i < j" + shows "stake j s ! i = s i" +using assms by (induct j arbitrary: s i) (simp_all add: nth_Cons') + +lemma stake_suffix_id: + "stake i \ @- (\ |\<^sub>s i) = \" +by (induct i) (simp_all add: fun_eq_iff shift_snth split: nat.splits) + +lemma id_stake_snth_suffix: + "\ = (stake i \ @ [\ i]) @- (\ |\<^sub>s Suc i)" +using stake_suffix_id +apply (metis Suc_diff_le append_Nil2 diff_is_0_eq length_stake lessI nat.simps(3) nat_le_linear shift_snth stake_nil stake_shift take_Suc_conv_app_nth) +done + +lemma stake_add[simp]: + "stake i \ @ stake j (\ |\<^sub>s i) = stake (i + j) \" +apply (induct i arbitrary: \) + apply simp +apply auto +apply (metis One_nat_def plus_1_eq_Suc suffix_plus_com) +done + +lemma stake_append: "stake n (u @- s) = take (min (length u) n) u @ stake (n - length u) s" +proof (induct n arbitrary: u) + case (Suc n) then show ?case + apply clarsimp + apply (cases u) + apply auto + done +qed auto + +lemma stake_shift_stake_shift: + "stake i \ @- stake j (\ |\<^sub>s i) @- \ = stake (i + j) \ @- \" +apply (induct i arbitrary: \) + apply simp +apply auto +apply (metis One_nat_def plus_1_eq_Suc suffix_plus_com) +done + +lemma stake_suffix_drop: + "stake i (\ |\<^sub>s j) = drop j (stake (i + j) \)" +by (metis append_eq_conv_conj length_stake semiring_normalization_rules(24) stake_add) + +lemma stake_suffix: + assumes "i \ j" + shows "stake j \ @- u |\<^sub>s i = \(i \ j - i) @- u" +by (simp add: assms stake_suffix_drop suffix_shift) + + +subsection\Decomposing safety and liveness \label{sec:infinite_sequences-safety-liveness}\ + +text\ + +Famously properties on infinite sequences can be decomposed into +@{emph \safety\} and @{emph \liveness\} +properties @{cite "AlpernSchneider:1985" and "Schneider:1987"}. See +@{cite [cite_macro=citet] "Kindler:1994"} for an overview. + +\ + +definition safety :: "'a seq_pred \ bool" where + "safety P \ (\\. \P \ \ (\i. \\. \P (stake i \ @- \)))" + +lemma safety_def2: \ \Contraposition gives the customary prefix-closure definition\ + "safety P \ (\\. (\i. \\. P (stake i \ @- \)) \ P \)" +unfolding safety_def by blast + +definition liveness :: "'a seq_pred \ bool" where + "liveness P \ (\\. \\. P (\ @- \))" + +lemmas safetyI = iffD2[OF safety_def, rule_format] +lemmas safetyI2 = iffD2[OF safety_def2, rule_format] +lemmas livenessI = iffD2[OF liveness_def, rule_format] + +lemma safety_False: + shows "safety (\\. False)" +by (rule safetyI) simp + +lemma safety_True: + shows "safety (\\. True)" +by (rule safetyI) simp + +lemma safety_state_prop: + shows "safety (\\. P (\ 0))" +by (rule safetyI) auto + +lemma safety_invariant: + shows "safety (\\. \i. P (\ i))" +apply (rule safetyI) +apply clarsimp +apply (metis length_stake lessI shift_snth_less stake_nth) +done + +lemma safety_transition_relation: + shows "safety (\\. \i. (\ i, \ (i + 1)) \ R)" +apply (rule safetyI) +apply clarsimp +apply (metis (no_types, hide_lams) Suc_eq_plus1 add.left_neutral add_Suc_right add_diff_cancel_left' le_add1 list.sel(1) list.simps(3) shift_simps(1) stake.simps(2) stake_suffix suffix_def) +done + +lemma safety_conj: + assumes "safety P" + assumes "safety Q" + shows "safety (P \<^bold>\ Q)" +using assms unfolding safety_def by blast + +lemma safety_always_eventually[simplified]: + assumes "safety P" + assumes "\i. \j\i. \\. P (\(0 \ j) @- \)" + shows "P \" +using assms unfolding safety_def2 +apply - +apply (drule_tac x=\ in spec) +apply clarsimp +apply (drule_tac x=i in spec) +apply clarsimp +apply (rule_tac x="(stake j \ @- \) |\<^sub>s i" in exI) +apply (simp add: stake_shift_stake_shift stake_suffix) +done + +lemma safety_disj: + assumes "safety P" + assumes "safety Q" + shows "safety (P \<^bold>\ Q)" +unfolding safety_def2 using assms +by (metis safety_always_eventually add_diff_cancel_right' diff_le_self le_add_same_cancel2) + +text\ + +The decomposition is given by a form of closure. + +\ + +definition M\<^sub>p :: "'a seq_pred \ 'a seq_pred" where + "M\<^sub>p P = (\\. \i. \\. P (stake i \ @- \))" + +definition Safe :: "'a seq_pred \ 'a seq_pred" where + "Safe P = (P \<^bold>\ M\<^sub>p P)" + +definition Live :: "'a seq_pred \ 'a seq_pred" where + "Live P = (P \<^bold>\ \<^bold>\M\<^sub>p P)" + +lemma decomp: + "P = (Safe P \<^bold>\ Live P)" +unfolding Safe_def Live_def by blast + +lemma safe: + "safety (Safe P)" +unfolding Safe_def safety_def M\<^sub>p_def +apply clarsimp +apply (simp add: stake_shift) +apply (rule_tac x=i in exI) +apply clarsimp +apply (rule_tac x=i in exI) +apply clarsimp +done + +lemma live: + "liveness (Live P)" +proof(rule livenessI) + fix \ + have "(\\. P (\ @- \)) \ \(\\. P (\ @- \))" by blast + also have "?this \ (\\. P (\ @- \) \ (\\. \P (\ @- \)))" by blast + also have "\ \ (\\. P (\ @- \) \ (\i. i = length \ \ (\\. \P (stake i (\ @- \) @- \))))" by (simp add: stake_shift) + also have "\ \ (\\. P (\ @- \) \ (\i. (\\. \P (stake i (\ @- \) @- \))))" by blast + finally have "\\. P (\ @- \) \ (\i. \\. \ P (stake i (\ @- \) @- \))" . + then show "\\. Live P (\ @- \)" unfolding Live_def M\<^sub>p_def by simp +qed + +text\ + +@{cite "Sistla:1994"} proceeds to give a topological analysis of fairness. An \<^emph>\absolute\ +liveness property is a liveness property whose complement is stable. + +\ + +definition absolute_liveness :: "'a seq_pred \ bool" where \ \ closed under prepending any finite sequence \ + "absolute_liveness P \ (\\. P \) \ (\\ \. P \ \ P (\ @- \))" + +definition stable :: "'a seq_pred \ bool" where \ \ closed under suffixes \ + "stable P \ (\\. P \) \ (\\ i. P \ \ P (\ |\<^sub>s i))" + +lemma absolute_liveness_liveness: + assumes "absolute_liveness P" + shows "liveness P" +using assms unfolding absolute_liveness_def liveness_def by blast + +lemma stable_absolute_liveness: + assumes "P \" + assumes "\P \'" \\ extra hypothesis \ + shows "stable P \ absolute_liveness (\<^bold>\ P)" +using assms unfolding stable_def absolute_liveness_def +apply auto + apply (metis cancel_comm_monoid_add_class.diff_cancel drop_eq_Nil order_refl shift.simps(1) suffix_shift suffix_zero) +apply (metis stake_suffix_id) +done + +(* +text\ + +Fairness ala Sistla. Unmotivated. + +FIXME safety properties are insensitive to fairness. +FIXME typically we prove \sys \ safety\. The result below doesn't appear strong enough. +FIXME observe fairness is a special liveness property. + +\ +*) + +definition fairness :: "'a seq_pred \ bool" where + "fairness P \ stable P \ absolute_liveness P" + +lemma fairness_safety: + assumes "safety P" + assumes "fairness F" + shows "(\\. F \ \ P \) \ (\\. P \)" +apply rule +using assms +apply clarsimp +unfolding safety_def fairness_def stable_def absolute_liveness_def +apply clarsimp +apply blast+ +done + +(*<*) + +end +(*>*) diff --git a/thys/ConcurrentIMP/LTL.thy b/thys/ConcurrentIMP/LTL.thy new file mode 100644 --- /dev/null +++ b/thys/ConcurrentIMP/LTL.thy @@ -0,0 +1,828 @@ +(*<*) +theory LTL +imports + CIMP_pred + Infinite_Sequences +begin + +(*>*) +section\ Linear Temporal Logic \label{sec:ltl}\ + +text\ + +To talk about liveness we need to consider infinitary behaviour on +sequences. Traditionally future-time linear temporal logic (LTL) is used to do +this @{cite "MannaPnueli:1991" and "OwickiLamport:1982"}. + +The following is a straightforward shallow embedding of the +now-traditional anchored semantics of LTL @{cite "MannaPnueli:1988"}. Some of it is adapted +from the sophisticated TLA development in the AFP due to @{cite [cite_macro=citet] "TLA-AFP"}. + +Unlike @{cite [cite_macro=citet] "Lamport:2002"}, include the +next operator, which is convenient for stating rules. Sometimes it allows us to +ignore the system, i.e. to state rules as temporally valid +(LTL-valid) rather than just temporally program valid (LTL-cimp-), in Jackson's terminology. + +\ + +definition state_prop :: "('a \ bool) \ 'a seq_pred" ("\_\") where + "\P\ = (\\. P (\ 0))" + +definition "next" :: "'a seq_pred \ 'a seq_pred" ("\_" [80] 80) where + "(\P) = (\\. P (\ |\<^sub>s 1))" + +definition always :: "'a seq_pred \ 'a seq_pred" ("\_" [80] 80) where + "(\P) = (\\. \i. P (\ |\<^sub>s i))" + +definition until :: "'a seq_pred \ 'a seq_pred \ 'a seq_pred" (infixr "\" 30) where (* FIXME priority, binds tighter than \<^bold>\ *) + "(P \ Q) = (\\. \i. Q (\ |\<^sub>s i) \ (\k |\<^sub>s k)))" + +definition eventually :: "'a seq_pred \ 'a seq_pred" ("\_" [80] 80) where (* FIXME priority, consider making an abbreviation *) + "(\P) = (\True\ \ P)" + +definition release :: "'a seq_pred \ 'a seq_pred \ 'a seq_pred" (infixr "\" 30) where (* FIXME priority, dual of Until *) + "(P \ Q) = (\<^bold>\(\<^bold>\P \ \<^bold>\Q))" + +definition unless :: "'a seq_pred \ 'a seq_pred \ 'a seq_pred" (infixr "\" 30) where (* FIXME priority, aka weak until *) + "(P \ Q) = ((P \ Q) \<^bold>\ \P)" + +abbreviation (input) + pred_always_imp_syn :: "'a seq_pred \ 'a seq_pred \ 'a seq_pred" (infixr "\<^bold>\" 25) where + "P \<^bold>\ Q \ \(P \<^bold>\ Q)" + +lemmas defs = + state_prop_def + always_def + eventually_def + next_def + release_def + unless_def + until_def + +lemma suffix_state_prop[simp]: + shows "\P\ (\ |\<^sub>s i) = P (\ i)" +unfolding defs by simp + +lemma alwaysI[intro]: + assumes "\i. P (\ |\<^sub>s i)" + shows "(\P) \" +unfolding defs using assms by blast + +lemma alwaysD: + assumes "(\P) \" + shows "P (\ |\<^sub>s i)" +using assms unfolding defs by blast + +lemma alwaysE: "\(\P) \; P (\ |\<^sub>s i) \ Q\ \ Q" +unfolding defs by blast + +lemma always_induct: + assumes "P \" + assumes "(\(P \<^bold>\ \P)) \" + shows "(\P) \" +proof(rule alwaysI) + fix i from assms show "P (\ |\<^sub>s i)" + unfolding defs by (induct i) simp_all +qed + +lemma seq_comp: + fixes \ :: "'a seq" + fixes P :: "'b seq_pred" + fixes f :: "'a \ 'b" + shows + "(\P) (f \ \) \ (\(\\. P (f \ \))) \" + "(\P) (f \ \) \ (\(\\. P (f \ \))) \" + "(P \ Q) (f \ \) \ ((\\. P (f \ \)) \ (\\. Q (f \ \))) \" + "(P \ Q) (f \ \) \ ((\\. P (f \ \)) \ (\\. Q (f \ \))) \" +unfolding defs by simp_all + +lemma nextI[intro]: + assumes "P (\ |\<^sub>s Suc 0)" + shows "(\P) \" +using assms unfolding defs by simp + +lemma untilI[intro]: + assumes "Q (\ |\<^sub>s i)" + assumes "\k |\<^sub>s k)" + shows "(P \ Q) \" +unfolding defs using assms by blast + +lemma untilE: + assumes "(P \ Q) \" + obtains i where "Q (\ |\<^sub>s i)" and "\k |\<^sub>s k)" +using assms unfolding until_def by blast + +lemma eventuallyI[intro]: + assumes "P (\ |\<^sub>s i)" + shows "(\P) \" +unfolding eventually_def using assms by blast + +lemma eventuallyE[elim]: + assumes "(\P) \" + obtains i where "P (\ |\<^sub>s i)" +using assms unfolding defs by (blast elim: untilE) + +lemma unless_alwaysI: + assumes "(\ P) \" + shows "(P \ Q) \" +using assms unfolding defs by blast + +lemma unless_untilI: + assumes "Q (\ |\<^sub>s j)" + assumes "\i. i < j \ P (\ |\<^sub>s i)" + shows "(P \ Q) \" +unfolding defs using assms by blast + +lemma always_imp_refl[iff]: + shows "(P \<^bold>\ P) \" +unfolding defs by blast + +lemma always_imp_trans: + assumes "(P \<^bold>\ Q) \" + assumes "(Q \<^bold>\ R) \" + shows "(P \<^bold>\ R) \" +using assms unfolding defs by blast + +lemma always_imp_mp: + assumes "(P \<^bold>\ Q) \" + assumes "P \" + shows "Q \" +using assms unfolding defs by (metis suffix_zero) + +lemma always_imp_mp_suffix: + assumes "(P \<^bold>\ Q) \" + assumes "P (\ |\<^sub>s i)" + shows "Q (\ |\<^sub>s i)" +using assms unfolding defs by metis + +text\ + +Some basic facts and equivalences, mostly sanity. + +\ + +lemma necessitation: + "(\s. P s) \ (\P) \" + "(\s. P s) \ (\P) \" + "(\s. P s) \ (P \ Q) \" + "(\s. Q s) \ (P \ Q) \" +unfolding defs by auto + +lemma cong: + "(\s. P s = P' s) \ \P\ = \P'\" + "(\\. P \ = P' \) \ (\P) = (\P')" + "(\\. P \ = P' \) \ (\P) = (\P')" + "(\\. P \ = P' \) \ (\P) = (\P')" + "\\\. P \ = P' \; \\. Q \ = Q' \\ \ (P \ Q) = (P' \ Q')" + "\\\. P \ = P' \; \\. Q \ = Q' \\ \ (P \ Q) = (P' \ Q')" +unfolding defs by auto + +lemma norm[simp]: + "\\False\\ = \False\" + "\\True\\ = \True\" + "(\<^bold>\\p\) = \\<^bold>\p\" + "(\p\ \<^bold>\ \q\) = \p \<^bold>\ q\" + "(\p\ \<^bold>\ \q\) = \p \<^bold>\ q\" + "(\p\ \<^bold>\ \q\) = \p \<^bold>\ q\" + "(\p\ \ \ \q\ \) = \p \<^bold>\ q\ \" + "(\p\ \ \ \q\ \) = \p \<^bold>\ q\ \" + "(\p\ \ \ \q\ \) = \p \<^bold>\ q\ \" + + "(\\False\) = \False\" + "(\\True\) = \True\" + + "(\\False\) = \False\" + "(\\True\) = \True\" + "(\<^bold>\\ P) \ = (\ (\<^bold>\ P)) \" + "(\\ P) = (\ P)" + + "(\\False\) = \False\" + "(\\True\) = \True\" + "(\<^bold>\\ P) = (\ (\<^bold>\ P))" + "(\\ P) = (\ P)" + + "(P \ \False\) = (\ P)" + + "(\<^bold>\(P \ Q)) \ = (\<^bold>\P \ \<^bold>\Q) \" + "(\False\ \ P) = P" + "(P \ \False\) = \False\" + "(P \ \True\) = \True\" + "(\True\ \ P) = (\ P)" + "(P \ (P \ Q)) = (P \ Q)" + + "(\<^bold>\(P \ Q)) \ = (\<^bold>\P \ \<^bold>\Q) \" + "(\False\ \ P) = (\P)" + "(P \ \False\) = \False\" + "(\True\ \ P) = P" + "(P \ \True\) = \True\" +(* + "(P \ (P \ Q)) \ = (P \ Q) \" FIXME for Release +*) +unfolding defs +apply (auto simp: fun_eq_iff) +apply (metis suffix_plus suffix_zero) +apply (metis suffix_plus suffix_zero) + apply fastforce + apply force +apply (metis add.commute add_diff_inverse_nat less_diff_conv2 not_le) +apply (metis add.right_neutral not_less0) + apply force + apply fastforce +done + +lemma always_conj_distrib: "(\(P \<^bold>\ Q)) = (\P \<^bold>\ \Q)" +unfolding defs by auto + +lemma eventually_disj_distrib: "(\(P \<^bold>\ Q)) = (\P \<^bold>\ \Q)" +unfolding defs by auto + +lemma always_eventually[elim!]: + assumes "(\P) \" + shows "(\P) \" +using assms unfolding defs by auto + +lemma eventually_imp_conv_disj: "(\(P \<^bold>\ Q)) = (\(\<^bold>\P) \<^bold>\ \Q)" +unfolding defs by auto + +lemma eventually_imp_distrib: + "(\(P \<^bold>\ Q)) = (\P \<^bold>\ \Q)" +unfolding defs by auto + +lemma unfold: + "(\ P) \ = (P \<^bold>\ \\P) \" + "(\ P) \ = (P \<^bold>\ \\P) \" + "(P \ Q) \ = (Q \<^bold>\ (P \<^bold>\ \(P \ Q))) \" + "(P \ Q) \ = (Q \<^bold>\ (P \<^bold>\ \(P \ Q))) \" + "(P \ Q) \ = (Q \<^bold>\ (P \<^bold>\ \(P \ Q))) \" +unfolding defs +apply - +apply (metis (full_types) add.commute add_diff_inverse_nat less_one suffix_plus suffix_zero) +apply (metis (full_types) One_nat_def add.right_neutral add_Suc_right lessI less_Suc_eq_0_disj suffix_plus suffix_zero) + +apply auto + apply fastforce + apply (metis gr0_conv_Suc nat_neq_iff not_less_eq suffix_zero) + apply (metis suffix_zero) + apply force + using less_Suc_eq_0_disj apply fastforce + apply (metis gr0_conv_Suc nat_neq_iff not_less0 suffix_zero) + + apply fastforce + apply (case_tac i; auto) + apply force + using less_Suc_eq_0_disj apply force + + apply force + using less_Suc_eq_0_disj apply fastforce + apply fastforce + apply (case_tac i; auto) +done + +lemma mono: + "\(\P) \; \\. P \ \ P' \\ \ (\P') \" + "\(\P) \; \\. P \ \ P' \\ \ (\P') \" + "\(P \ Q) \; \\. P \ \ P' \; \\. Q \ \ Q' \\ \ (P' \ Q') \" + "\(P \ Q) \; \\. P \ \ P' \; \\. Q \ \ Q' \\ \ (P' \ Q') \" +unfolding defs by force+ + +lemma always_imp_mono: + "\(\P) \; (P \<^bold>\ P') \\ \ (\P') \" + "\(\P) \; (P \<^bold>\ P') \\ \ (\P') \" + "\(P \ Q) \; (P \<^bold>\ P') \; (Q \<^bold>\ Q') \\ \ (P' \ Q') \" + "\(P \ Q) \; (P \<^bold>\ P') \; (Q \<^bold>\ Q') \\ \ (P' \ Q') \" +unfolding defs by force+ + +lemma next_conj_distrib: + "(\(P \<^bold>\ Q)) = (\P \<^bold>\ \Q)" +unfolding defs by auto + +lemma next_disj_distrib: + "(\(P \<^bold>\ Q)) = (\P \<^bold>\ \Q)" +unfolding defs by auto + +lemma until_next_distrib: + "(\(P \ Q)) = (\P \ \Q)" +unfolding defs by (auto simp: fun_eq_iff) + +lemma until_imp_eventually: + "((P \ Q) \<^bold>\ \Q) \" +unfolding defs by auto + +lemma until_until_disj: + assumes "(P \ Q \ R) \" + shows "((P \<^bold>\ Q) \ R) \" +using assms unfolding defs +apply clarsimp +apply (metis (full_types) add_diff_inverse_nat nat_add_left_cancel_less) +done + +lemma unless_unless_disj: + assumes "(P \ Q \ R) \" + shows "((P \<^bold>\ Q) \ R) \" +using assms unfolding defs +apply auto + apply (metis add.commute add_diff_inverse_nat leI less_diff_conv2) +apply (metis add_diff_inverse_nat) +done + +lemma until_conj_distrib: + "((P \<^bold>\ Q) \ R) = ((P \ R) \<^bold>\ (Q \ R))" +unfolding defs +apply (auto simp: fun_eq_iff) +apply (metis dual_order.strict_trans nat_neq_iff) +done + +lemma until_disj_distrib: + "(P \ (Q \<^bold>\ R)) = ((P \ Q) \<^bold>\ (P \ R))" +unfolding defs by (auto simp: fun_eq_iff) + +lemma eventually_until: + "(\P) = (\<^bold>\P \ P)" +unfolding defs +apply (auto simp: fun_eq_iff) +apply (case_tac "P (x |\<^sub>s 0)") + apply blast +apply (drule (1) ex_least_nat_less) +apply (metis le_simps(2)) +done + +lemma eventually_until_eventually: + "(\(P \ Q)) = (\Q)" +unfolding defs by force + +lemma eventually_unless_until: + "((P \ Q) \<^bold>\ \Q) = (P \ Q)" +unfolding defs by force + +lemma eventually_always_imp_always_eventually: + assumes "(\\P) \" + shows "(\\P) \" +using assms unfolding defs by (metis suffix_commute) + +lemma eventually_always_next_stable: + assumes "(\P) \" + assumes "(P \<^bold>\ \P) \" + shows "(\\P) \" +using assms by (metis (no_types) eventuallyI alwaysD always_induct eventuallyE norm(15)) + +lemma next_stable_imp_eventually_always: + assumes "(P \<^bold>\ \P) \" + shows "(\P \<^bold>\ \\P) \" +using assms eventually_always_next_stable by blast + +lemma always_eventually_always: + "\\\P = \\P" +unfolding defs by (clarsimp simp: fun_eq_iff) (metis add.left_commute semiring_normalization_rules(25)) + +(* FIXME define "stable", develop more rules for it *) +lemma stable_unless: + assumes "(P \<^bold>\ \(P \<^bold>\ Q)) \" + shows "(P \<^bold>\ (P \ Q)) \" +using assms unfolding defs +apply - +apply (rule ccontr) +apply clarsimp +apply (drule (1) ex_least_nat_less[where P="\j. \P (\ |\<^sub>s i + j)" for i, simplified]) +apply clarsimp +apply (metis add_Suc_right le_less less_Suc_eq) +done + +lemma unless_induct: \\ Rule \texttt{WAIT} from @{cite [cite_macro=citet] \Fig~3.3\ "MannaPnueli:1995"}\ + assumes I: "(I \<^bold>\ \(I \<^bold>\ R)) \" + assumes P: "(P \<^bold>\ I \<^bold>\ R) \" + assumes Q: "(I \<^bold>\ Q) \" + shows "(P \<^bold>\ Q \ R) \" +apply (intro alwaysI impI) +apply (erule impE[OF alwaysD[OF P]]) +apply (erule disjE) + apply (rule always_imp_mono(4)[where P=I and Q=R]) + apply (erule mp[OF alwaysD[OF stable_unless[OF I]]]) + apply (simp add: Q alwaysD) + apply blast +apply (simp add: unfold) +done + + +subsection\ Leads-to and leads-to-via \label{sec:leads-to} \ + +text\ + +Most of our assertions will be of the form @{term "A \<^bold>\ \C"} (pronounced ``\A\ leads to \C\'') +or @{term "A \<^bold>\ B \ C"} (``\A\ leads to \C\ via \B\''). + +Most of these rules are due to @{cite [cite_macro=citet] +"Jackson:1998"} who used leads-to-via in a sequential setting. Others +are due to @{cite [cite_macro=citet] "MannaPnueli:1991"}. + +The leads-to-via connective is similar to the ``ensures'' modality of @{cite [cite_macro=citet] \\S3.4.4\ "ChandyMisra:1989"}. + +\ + +abbreviation (input) + leads_to :: "'a seq_pred \ 'a seq_pred \ 'a seq_pred" (infixr "\<^bold>\" 25) where (* FIXME priority *) + "P \<^bold>\ Q \ P \<^bold>\ \Q" + +lemma leads_to_refl: + shows "(P \<^bold>\ P) \" +by (metis (no_types, lifting) necessitation(1) unfold(2)) + +lemma leads_to_trans: + assumes "(P \<^bold>\ Q) \" + assumes "(Q \<^bold>\ R) \" + shows "(P \<^bold>\ R) \" +using assms unfolding defs by clarsimp (metis semiring_normalization_rules(25)) + +lemma leads_to_eventuallyE: + assumes "(P \<^bold>\ Q) \" + assumes "(\P) \" + shows "(\Q) \" +using assms unfolding defs by auto + +lemma leads_to_mono: + assumes "(P' \<^bold>\ P) \" + assumes "(Q \<^bold>\ Q') \" + assumes "(P \<^bold>\ Q) \" + shows "(P' \<^bold>\ Q') \" +using assms unfolding defs by clarsimp blast + +lemma leads_to_eventually: + shows "(P \<^bold>\ Q \<^bold>\ \P \<^bold>\ \Q) \" +by (metis (no_types, lifting) alwaysI unfold(2)) + +lemma leads_to_disj: + assumes "(P \<^bold>\ R) \" + assumes "(Q \<^bold>\ R) \" + shows "((P \<^bold>\ Q) \<^bold>\ R) \" +using assms unfolding defs by simp + +lemma leads_to_leads_to_viaE: + shows "((P \<^bold>\ P \ Q) \<^bold>\ P \<^bold>\ Q) \" +unfolding defs by clarsimp blast + +lemma leads_to_via_concl_weaken: + assumes "(R \<^bold>\ R') \" + assumes "(P \<^bold>\ Q \ R) \" + shows "(P \<^bold>\ Q \ R') \" +using assms unfolding LTL.defs by force + +lemma leads_to_via_trans: + assumes "(A \<^bold>\ B \ C) \" + assumes "(C \<^bold>\ D \ E) \" + shows "(A \<^bold>\ (B \<^bold>\ D) \ E) \" +proof(rule alwaysI, rule impI) + fix i assume "A (\ |\<^sub>s i)" with assms show "((B \<^bold>\ D) \ E) (\ |\<^sub>s i)" + apply - + apply (erule alwaysE[where i=i]) + apply clarsimp + apply (erule untilE) + apply clarsimp (* suffix *) + apply (drule (1) always_imp_mp_suffix) + apply (erule untilE) + apply clarsimp (* suffix *) + apply (rule_tac i="ia + iaa" in untilI; simp add: ac_simps) + apply (metis (full_types) add.assoc leI le_Suc_ex nat_add_left_cancel_less) (* arithmetic *) + done +qed + +lemma leads_to_via_disj: \ \ useful for case distinctions \ + assumes "(P \<^bold>\ Q \ R) \" + assumes "(P' \<^bold>\ Q' \ R) \" + shows "(P \<^bold>\ P' \<^bold>\ (Q \<^bold>\ Q') \ R) \" +using assms unfolding defs by (auto 10 0) + +lemma leads_to_via_disj': \ \ more like a chaining rule \ + assumes "(A \<^bold>\ B \ C) \" + assumes "(C \<^bold>\ D \ E) \" + shows "(A \<^bold>\ C \<^bold>\ (B \<^bold>\ D) \ E) \" +proof(rule alwaysI, rule impI, erule disjE) + fix i assume "A (\ |\<^sub>s i)" with assms show "((B \<^bold>\ D) \ E) (\ |\<^sub>s i)" + apply - + apply (erule alwaysE[where i=i]) + apply clarsimp + apply (erule untilE) + apply clarsimp (* suffix *) + apply (drule (1) always_imp_mp_suffix) + apply (erule untilE) + apply clarsimp (* suffix *) + apply (rule_tac i="ia + iaa" in untilI; simp add: ac_simps) + apply (metis (full_types) add.assoc leI le_Suc_ex nat_add_left_cancel_less) (* arithmetic *) + done +next + fix i assume "C (\ |\<^sub>s i)" with assms(2) show "((B \<^bold>\ D) \ E) (\ |\<^sub>s i)" + apply - + apply (erule alwaysE[where i=i]) + apply (simp add: mono) + done +qed + +lemma leads_to_via_stable_augmentation: + assumes stable: "(P \<^bold>\ Q \<^bold>\ \Q) \" + assumes U: "(A \<^bold>\ P \ C) \" + shows "((A \<^bold>\ Q) \<^bold>\ P \ (C \<^bold>\ Q)) \" +proof(intro alwaysI impI, elim conjE) + fix i assume AP: "A (\ |\<^sub>s i)" "Q (\ |\<^sub>s i)" + have "Q (\ |\<^sub>s (j + i))" if "Q (\ |\<^sub>s i)" and "\k |\<^sub>s (k + i))" for j + using that stable by (induct j; force simp: defs) + with U AP show "(P \ (\\. C \ \ Q \)) (\ |\<^sub>s i)" + unfolding defs by clarsimp (metis (full_types) add.commute) +qed + +lemma leads_to_via_wf: + assumes "wf R" + assumes indhyp: "\t. (A \<^bold>\ \\ \<^bold>= \t\\ \<^bold>\ B \ (A \<^bold>\ \\ \<^bold>\ \t\ \<^bold>\ \R\\ \<^bold>\ C)) \" + shows "(A \<^bold>\ B \ C) \" +proof(intro alwaysI impI) + fix i assume "A (\ |\<^sub>s i)" with \wf R\ show "(B \ C) (\ |\<^sub>s i)" + proof(induct "\ (\ i)" arbitrary: i) + case (less i) with indhyp[where t="\ (\ i)"] show ?case + apply - + apply (drule alwaysD[where i=i]) + apply clarsimp + apply (erule untilE) + apply (rename_tac j) + apply (erule disjE; clarsimp) + apply (drule_tac x="i + j" in meta_spec; clarsimp) + apply (erule untilE; clarsimp) + apply (rename_tac j k) + apply (rule_tac i="j + k" in untilI) + apply (simp add: add.assoc) + apply clarsimp + apply (metis add.assoc add.commute add_diff_inverse_nat less_diff_conv2 not_le) + apply auto + done + qed +qed + +text\ + +The well-founded response rule due to @{cite [cite_macro=citet] \Fig~1.23: \texttt{WELL} (well-founded response)\"MannaPnueli:2010"}, +generalised to an arbitrary set of assertions and sequence predicates. +\<^item> \W1\ generalised to be contingent. +\<^item> \W2\ is a well-founded set of assertions that by \W1\ includes \P\ + +\ + +(* FIXME: Does \Is\ need to be consistent? *) +lemma leads_to_wf: + fixes Is :: "('a seq_pred \ ('a \ 'b)) set" + assumes "wf (R :: 'b rel)" + assumes W1: "(\(\<^bold>\\. \\\\fst ` Is\\ \<^bold>\ (P \<^bold>\ \))) \" + assumes W2: "\(\, \)\Is. \(\', \')\insert (Q, \0) Is. \t. (\ \<^bold>\ \\ \<^bold>= \t\\ \<^bold>\ \' \<^bold>\ \\' \<^bold>\ \t\ \<^bold>\ \R\\) \" + shows "(P \<^bold>\ Q) \" +proof - + have "(\ \<^bold>\ \\ \<^bold>= \t\\ \<^bold>\ Q) \" if "(\, \) \ Is" for \ \ t + using \wf R\ that W2 + apply (induct t arbitrary: \ \) + unfolding LTL.defs split_def + apply clarsimp + apply (metis (no_types, hide_lams) ab_semigroup_add_class.add_ac(1) fst_eqD snd_conv surjective_pairing) + done + with W1 show ?thesis + apply - + apply (rule alwaysI) + apply clarsimp + apply (erule_tac i=i in alwaysE) + apply clarsimp + using alwaysD suffix_state_prop apply blast + done +qed + + +subsection\Fairness\ + +text\ + +A few renderings of weak fairness. @{cite [cite_macro=citet] "vanGlabbeekHofner:2019"} call this +"response to insistence" as a generalisation of weak fairness. + +\ + +definition weakly_fair :: "'a seq_pred \ 'a seq_pred \ 'a seq_pred" where + "weakly_fair enabled taken = (\enabled \<^bold>\ \taken)" + +lemma weakly_fair_def2: + shows "weakly_fair enabled taken = \(\<^bold>\\(enabled \<^bold>\ \<^bold>\taken))" +unfolding weakly_fair_def by (metis (full_types) always_conj_distrib norm(18)) + +lemma weakly_fair_def3: + shows "weakly_fair enabled taken = (\\enabled \<^bold>\ \\taken)" +unfolding weakly_fair_def2 +apply (clarsimp simp: fun_eq_iff) + +unfolding defs (* True, but can we get there deductively? *) +apply auto +apply (metis (full_types) add.left_commute semiring_normalization_rules(25)) +done + +lemma weakly_fair_def4: + shows "weakly_fair enabled taken = \\(enabled \<^bold>\ taken)" +using weakly_fair_def2 by force + +lemma mp_weakly_fair: + assumes "weakly_fair enabled taken \" + assumes "(\enabled) \" + shows "(\taken) \" +using assms unfolding weakly_fair_def using always_imp_mp by blast + +lemma always_weakly_fair: + shows "\(weakly_fair enabled taken) = weakly_fair enabled taken" +unfolding weakly_fair_def by simp + +lemma eventually_weakly_fair: + shows "\(weakly_fair enabled taken) = weakly_fair enabled taken" +unfolding weakly_fair_def2 by (simp add: always_eventually_always) + +lemma weakly_fair_weaken: + assumes "(enabled' \<^bold>\ enabled) \" + assumes "(taken \<^bold>\ taken') \" + shows "(weakly_fair enabled taken \<^bold>\ weakly_fair enabled' taken') \" +using assms unfolding weakly_fair_def defs by simp blast + +lemma weakly_fair_unless_until: + shows "(weakly_fair enabled taken \<^bold>\ (enabled \<^bold>\ enabled \ taken)) = (enabled \<^bold>\ enabled \ taken)" +unfolding defs weakly_fair_def +apply (auto simp: fun_eq_iff) +apply (metis add.right_neutral) +done + +lemma stable_leads_to_eventually: + assumes "(enabled \<^bold>\ \(enabled \<^bold>\ taken)) \" + shows "(enabled \<^bold>\ (\enabled \<^bold>\ \taken)) \" +using assms unfolding defs +apply - +apply (rule ccontr) +apply clarsimp +apply (drule (1) ex_least_nat_less[where P="\j. \ enabled (\ |\<^sub>s i + j)" for i, simplified]) +apply clarsimp +apply (metis add_Suc_right leI less_irrefl_nat) +done + +lemma weakly_fair_stable_leads_to: + assumes "(weakly_fair enabled taken) \" + assumes "(enabled \<^bold>\ \(enabled \<^bold>\ taken)) \" + shows "(enabled \<^bold>\ taken) \" +using stable_leads_to_eventually[OF assms(2)] assms(1) unfolding defs weakly_fair_def +by (auto simp: fun_eq_iff) + +lemma weakly_fair_stable_leads_to_via: + assumes "(weakly_fair enabled taken) \" + assumes "(enabled \<^bold>\ \(enabled \<^bold>\ taken)) \" + shows "(enabled \<^bold>\ enabled \ taken) \" +using stable_unless[OF assms(2)] assms(1) by (metis (mono_tags) weakly_fair_unless_until) + +text\ + +Similarly for strong fairness. @{cite [cite_macro=citet] "vanGlabbeekHofner:2019"} call this +"response to persistence" as a generalisation of strong fairness. + +\ + +definition strongly_fair :: "'a seq_pred \ 'a seq_pred \ 'a seq_pred" where + "strongly_fair enabled taken = (\\enabled \<^bold>\ \taken)" + +lemma strongly_fair_def2: + "strongly_fair enabled taken = \(\<^bold>\\(\enabled \<^bold>\ \<^bold>\taken))" +unfolding strongly_fair_def by (metis weakly_fair_def weakly_fair_def2) + +lemma strongly_fair_def3: + "strongly_fair enabled taken = (\\enabled \<^bold>\ \\taken)" +unfolding strongly_fair_def2 by (metis (full_types) always_eventually_always weakly_fair_def2 weakly_fair_def3) + +lemma always_strongly_fair: + "\(strongly_fair enabled taken) = strongly_fair enabled taken" +unfolding strongly_fair_def by simp + +lemma eventually_strongly_fair: + "\(strongly_fair enabled taken) = strongly_fair enabled taken" +unfolding strongly_fair_def2 by (simp add: always_eventually_always) + +lemma strongly_fair_disj_distrib: \ \not true for \weakly_fair\\ + "strongly_fair (enabled1 \<^bold>\ enabled2) taken = (strongly_fair enabled1 taken \<^bold>\ strongly_fair enabled2 taken)" +unfolding strongly_fair_def defs +apply (auto simp: fun_eq_iff) + apply blast + apply blast + apply (metis (full_types) semiring_normalization_rules(25)) +done + +lemma strongly_fair_imp_weakly_fair: + assumes "strongly_fair enabled taken \" + shows "weakly_fair enabled taken \" +using assms unfolding strongly_fair_def3 weakly_fair_def3 by (simp add: eventually_always_imp_always_eventually) + +lemma always_enabled_weakly_fair_strongly_fair: + assumes "(\enabled) \" + shows "weakly_fair enabled taken \ = strongly_fair enabled taken \" +using assms by (metis strongly_fair_def3 strongly_fair_imp_weakly_fair unfold(2) weakly_fair_def3) + + +subsection\Safety and liveness \label{sec:ltl-safety-liveness}\ + +text\ + +@{cite [cite_macro=citet] "Sistla:1994"} shows some characterisations +of LTL formulas in terms of safety and liveness. Note his @{term +"(\)"} is actually @{term "(\)"}. + +See also @{cite [cite_macro=citet] "ChangMannaPnueli:1992"}. + +\ + +lemma safety_state_prop: + shows "safety \P\" +unfolding defs by (rule safety_state_prop) + +lemma safety_Next: + assumes "safety P" + shows "safety (\P)" +using assms unfolding defs safety_def +apply clarsimp +apply (metis (mono_tags) One_nat_def list.sel(3) nat.simps(3) stake.simps(2)) +done + +lemma safety_unless: + assumes "safety P" + assumes "safety Q" + shows "safety (P \ Q)" +proof(rule safetyI2) + fix \ assume X: "\\. (P \ Q) (stake i \ @- \)" for i + then show "(P \ Q) \" + proof(cases "\i j. \\. P (\(i \ j) @- \)") + case True + with \safety P\ have "\i. P (\ |\<^sub>s i)" unfolding safety_def2 by blast + then show ?thesis by (blast intro: unless_alwaysI) + next + case False + then obtain k k' where "\\. \ P (\(k \ k') @- \)" by clarsimp + then have "\i u. k + k' \ i \ \P ((stake i \ @- u) |\<^sub>s k)" + by (metis add.commute diff_add stake_shift_stake_shift stake_suffix_drop suffix_shift) + then have "\i u. k + k' \ i \ (P \ Q) (stake i \ @- u) \ (\m\k. Q ((stake i \ @- u) |\<^sub>s m) \ (\p @- u) |\<^sub>s p)))" + unfolding defs using leI by blast + then have "\i u. k + k' \ i \ (P \ Q) (stake i \ @- u) \ (\m\k. Q (\(m \ i - m) @- u) \ (\p(p \ i - p) @- u)))" + by (metis stake_suffix add_leE nat_less_le order_trans) + then have "\i. \n\i. \m\k. \u. Q (\(m \ n - m) @- u) \ (\p(p \ n - p) @- u))" + using X by (metis add.commute le_add1) + then have "\m\k. \i. \n\i. \u. Q (\(m \ n - m) @- u) \ (\p(p \ n - p) @- u))" + by (simp add: always_eventually_pigeonhole) + with \safety P\ \safety Q\ show "(P \ Q) \" + unfolding defs by (metis Nat.le_diff_conv2 add_leE safety_always_eventually) + qed +qed + +lemma safety_always: + assumes "safety P" + shows "safety (\P)" +using assms by (metis norm(20) safety_def safety_unless) + +lemma absolute_liveness_eventually: + shows "absolute_liveness P \ (\\. P \) \ P = \P" +unfolding absolute_liveness_def defs +by (metis cancel_comm_monoid_add_class.diff_cancel drop_eq_Nil order_refl shift.simps(1) stake_suffix_id suffix_shift suffix_zero) + +lemma stable_always: + shows "stable P \ (\\. P \) \ P = \P" +unfolding stable_def defs by (metis suffix_zero) + +(* FIXME Sistla's examples of stable properties are boring and follow directly from this lemma. + FIXME the fairness "type of formulas" follow from the above and the fairness def. *) + +text\ + +To show that @{const \weakly_fair\} is a @{const \fairness\} property requires some constraints on \enabled\ and \taken\: +\<^item> it is reasonable to assume they are state formulas +\<^item> \taken\ must be satisfiable + +\ + +lemma fairness_weakly_fair: + assumes "\s. taken s" + shows "fairness (weakly_fair \enabled\ \taken\)" +unfolding fairness_def stable_def absolute_liveness_def weakly_fair_def +using assms +apply auto + apply (rule_tac x="\_ .s" in exI) + apply fastforce + apply (simp add: alwaysD) + apply (rule_tac x="\_ .s" in exI) + apply fastforce +apply (metis (full_types) absolute_liveness_def absolute_liveness_eventually eventually_weakly_fair weakly_fair_def) +done + +lemma fairness_strongly_fair: + assumes "\s. taken s" + shows "fairness (strongly_fair \enabled\ \taken\)" +unfolding fairness_def stable_def absolute_liveness_def strongly_fair_def +using assms +apply auto + apply (rule_tac x="\_ .s" in exI) + apply fastforce + apply (simp add: alwaysD) + apply (rule_tac x="\_ .s" in exI) + apply fastforce +apply (metis (full_types) absolute_liveness_def absolute_liveness_eventually eventually_weakly_fair weakly_fair_def) +done + +(*<*) + +end +(*>*) diff --git a/thys/ConcurrentIMP/ROOT b/thys/ConcurrentIMP/ROOT --- a/thys/ConcurrentIMP/ROOT +++ b/thys/ConcurrentIMP/ROOT @@ -1,11 +1,14 @@ chapter AFP session ConcurrentIMP (AFP) = "HOL-Library" + options [timeout = 300] + directories + "ex" theories [show_question_marks = false, names_short] CIMP - CIMP_one_place_buffer_ex - CIMP_unbounded_buffer_ex + "ex/CIMP_locales" + "ex/CIMP_one_place_buffer" + "ex/CIMP_unbounded_buffer" document_files "root.bib" "root.tex" diff --git a/thys/ConcurrentIMP/cimp.ML b/thys/ConcurrentIMP/cimp.ML new file mode 100644 --- /dev/null +++ b/thys/ConcurrentIMP/cimp.ML @@ -0,0 +1,140 @@ +(* Pollutes the global namespace, but we use them everywhere *) +fun ss_only thms ctxt = clear_simpset (put_simpset HOL_basic_ss ctxt) addsimps thms; +fun HOL_ss_only thms ctxt = clear_simpset (put_simpset HOL_ss ctxt) addsimps thms; + +signature CIMP = +sig + val com_locs_fold : (term * 'b -> 'b) -> 'b -> term -> 'b + val com_locs_map : (term -> 'b) -> term -> 'b list + val com_locs_fold_no_response : (term * 'b -> 'b) -> 'b -> term -> 'b + val com_locs_map_no_response : (term -> 'b) -> term -> 'b list + val intern_com : Facts.ref -> local_theory -> local_theory + val def_locset : thm -> local_theory -> local_theory +end; + +structure Cimp : CIMP = +struct + +fun com_locs_fold f x (Const (@{const_name Request}, _) $ l $ _ $ _ ) = f (l, x) + | com_locs_fold f x (Const (@{const_name Response}, _) $ l $ _) = f (l, x) + | com_locs_fold f x (Const (@{const_name LocalOp}, _) $ l $ _) = f (l, x) + | com_locs_fold f x (Const (@{const_name Cond1}, _) $ l $ _ $ c) = com_locs_fold f (f (l, x)) c + | com_locs_fold f x (Const (@{const_name Cond2}, _) $ l $ _ $ c1 $ c2) = com_locs_fold f (com_locs_fold f (f (l, x)) c1) c2 + | com_locs_fold f x (Const (@{const_name Loop}, _) $ c) = com_locs_fold f x c + | com_locs_fold f x (Const (@{const_name While}, _) $ l $ _ $ c) = com_locs_fold f (f (l, x)) c + | com_locs_fold f x (Const (@{const_name Seq}, _) $ c1 $ c2) = com_locs_fold f (com_locs_fold f x c1) c2 + | com_locs_fold f x (Const (@{const_name Choose}, _) $ c1 $ c2) = com_locs_fold f (com_locs_fold f x c1) c2 + | com_locs_fold _ x _ = x; + +fun com_locs_map f com = com_locs_fold (fn (l, acc) => f l :: acc) [] com + +fun com_locs_fold_no_response f x (Const (@{const_name Request}, _) $ l $ _ $ _ ) = f (l, x) + | com_locs_fold_no_response _ x (Const (@{const_name Response}, _) $ _ $ _) = x (* can often ignore \Response\ *) + | com_locs_fold_no_response f x (Const (@{const_name LocalOp}, _) $ l $ _) = f (l, x) + | com_locs_fold_no_response f x (Const (@{const_name Cond1}, _) $ l $ _ $ c) = com_locs_fold_no_response f (f (l, x)) c + | com_locs_fold_no_response f x (Const (@{const_name Cond2}, _) $ l $ _ $ c1 $ c2) = com_locs_fold_no_response f (com_locs_fold_no_response f (f (l, x)) c1) c2 + | com_locs_fold_no_response f x (Const (@{const_name Loop}, _) $ c) = com_locs_fold_no_response f x c + | com_locs_fold_no_response f x (Const (@{const_name While}, _) $ l $ _ $ c) = com_locs_fold_no_response f (f (l, x)) c + | com_locs_fold_no_response f x (Const (@{const_name Seq}, _) $ c1 $ c2) = com_locs_fold_no_response f (com_locs_fold_no_response f x c1) c2 + | com_locs_fold_no_response f x (Const (@{const_name Choose}, _) $ c1 $ c2) = com_locs_fold_no_response f (com_locs_fold_no_response f x c1) c2 + | com_locs_fold_no_response _ x _ = x; + +fun com_locs_map_no_response f com = com_locs_fold_no_response (fn (l, acc) => f l :: acc) [] com + +fun cprop_of_equality ctxt : thm -> cterm = + Local_Defs.meta_rewrite_rule ctxt (* handle `=` or `\` *) + #> Thm.cprop_of + +fun equality_lhs ctxt : thm -> term = + cprop_of_equality ctxt #> Thm.dest_equals_lhs #> Thm.term_of + +fun equality_rhs ctxt : thm -> term = + cprop_of_equality ctxt #> Thm.dest_equals_rhs #> Thm.term_of + +(* Intern all labels mentioned in CIMP programs \facts\ + +FIXME can only use \com_intern\ once per locale +FIXME forces all labels to be unique and distinct from other constants in the locale. +FIXME assumes the labels are character strings +*) +fun intern_com facts ctxt : local_theory = + let + val thms = Proof_Context.get_fact ctxt facts + (* Define constants with defs of the form loc.XXX_def: "XXX \ ''XXX''. *) + val attrs = [] + fun add_literal_def (literal, (loc_defs, ctxt)) : thm list * local_theory = + let + val literal_name = HOLogic.dest_string literal (* FIXME might not be a nice name, but the error is readable so shrug. FIXME generalise to other label types *) + val literal_def_binding = Binding.empty (* Binding.qualify true "loc" (Binding.name (Thm.def_name literal_name)) No need to name individual defs *) + val ((_, (_, loc_def)), ctxt) = Local_Theory.define ((Binding.name literal_name, Mixfix.NoSyn), ((literal_def_binding, attrs), literal)) ctxt + in + (loc_def :: loc_defs, ctxt) + end; + val (loc_defs, ctxt) = List.foldl (fn (com, acc) => com_locs_fold add_literal_def acc (equality_rhs ctxt com)) ([], ctxt) thms + + val coms_interned = List.map (Local_Defs.fold ctxt loc_defs) thms + val attrs = [] + val (_, ctxt) = Local_Theory.note ((@{binding "loc_defs"}, attrs), loc_defs) ctxt + val (_, ctxt) = Local_Theory.note ((@{binding "com_interned"}, attrs), coms_interned) ctxt + in + ctxt + end; + +(* Cache location set membership facts. + +Decide membership in the given set for each label in the CIMP programs +in the Named_Theorems "com". + +If the label set and com types differ, we probably get a nasty error. + +*) + +fun def_locset thm ctxt = + let + val set_name = equality_lhs ctxt thm + val set_name_str = case set_name of Const (c, _) => c | Free (c, _) => c | _ => error ("Not an equation of the form x = set: " ^ Thm.string_of_thm ctxt thm) + val memb_thm_name = Binding.qualify true set_name_str (Binding.name "memb") + fun mk_memb l = Thm.cterm_of ctxt (HOLogic.mk_mem (l, set_name)) +(* +1. solve atomic membership yielding \''label'' \ set\ or \''label'' \ set\. +2. fold \loc_defs\ +3. cleanup with the existing \locset_cache\. +FIXME trim simpsets: 1: sets 2: not much 3: not much +*) + val loc_defs = Proof_Context.get_fact ctxt (Facts.named "loc_defs") + val membership_ctxt = ctxt addsimps ([thm] @ loc_defs) + val cleanup_ctxt = HOL_ss_only (@{thms cleanup_simps} @ Named_Theorems.get ctxt \<^named_theorems>\locset_cache\) ctxt + val rewrite_tac = + Simplifier.rewrite membership_ctxt + #> Local_Defs.fold ctxt loc_defs + #> Simplifier.simplify cleanup_ctxt + val coms = Proof_Context.get_fact ctxt (Facts.named "com_interned") +(* Parallel *) + fun mk_thms coms : thm list = Par_List.map rewrite_tac (maps (equality_rhs ctxt #> com_locs_map_no_response mk_memb) coms) +(* Sequential *) +(* fun mk_thms coms = List.foldl (fn (c, thms) => com_locs_fold (fn l => fn thms => rewrite_tac (mk_memb l) :: thms) thms c) [] coms *) + val attrs = [] + val (_, ctxt) = ctxt |> Local_Theory.note ((memb_thm_name, attrs), mk_thms coms) +(* Add \memb_thms\ to the global (and inherited by locales) \locset_cache\ *) + val memb_thm_full_name = Local_Theory.full_name ctxt memb_thm_name + val (finish, ctxt) = Target_Context.switch_named_cmd (SOME ("-", Position.none)) (Context.Proof ctxt) (* switch to the "root" local theory *) + val memb_thms = Proof_Context.get_fact ctxt (Facts.named memb_thm_full_name) + val attrs = [Attrib.internal (K (Named_Theorems.add \<^named_theorems>\locset_cache\))] + val (_, ctxt) = ctxt |> Local_Theory.note ((Binding.empty, attrs), memb_thms) + val ctxt = case finish ctxt of Context.Proof ctxt => ctxt | _ => error "Context.generic failure" (* Return to the locale we were in *) + in + ctxt + end; + +end; + +val _ = + Outer_Syntax.local_theory \<^command_keyword>\intern_com\ "intern labels in CIMP commands" + (Parse.thms1 >> (fn facts => fn ctxt => List.foldl (fn ((f, _), ctxt) => Cimp.intern_com f ctxt) ctxt facts)); + +val _ = + Outer_Syntax.local_theory' \<^command_keyword>\locset_definition\ "constant definition for sets of locations" + (Scan.option Parse_Spec.constdecl -- (Parse_Spec.opt_thm_name ":" -- Parse.prop) -- + Parse_Spec.if_assumes -- Parse.for_fixes >> (fn (((decl, spec), prems), params) => fn b => fn lthy => + Specification.definition_cmd decl params prems spec b lthy + |> (fn ((_, (_, thm)), lthy) => (thm, lthy)) |> uncurry Cimp.def_locset)); diff --git a/thys/ConcurrentIMP/document/root.bib b/thys/ConcurrentIMP/document/root.bib --- a/thys/ConcurrentIMP/document/root.bib +++ b/thys/ConcurrentIMP/document/root.bib @@ -1,533 +1,656 @@ @STRING{cacm="Communications of the ACM"} +@STRING{csur="{ACM} Computing Surveys"} +@STRING{fac="Formal Aspects of Computing"} +@STRING{ipl="Information Processing Letters"} +@STRING{lncs="LNCS"} +@STRING{tcs="Theoretical Computer Science"} @article{AptFrancezDeRoever:1980, author = {K. R. Apt and Francez, N. and de Roever, W. P.}, title = {A Proof System for Communicating Sequential Processes}, journal = toplas, year = 1980, volume = 2, number = 3, pages = {359--385}, doi = {10.1145/357103.357110} } @TechReport{LynchTuttle:1989, author = "Lynch, N. A and Tuttle, M. R.", - title = "An introduction to - input/output automata", + title = "An introduction to input/output automata", institution = "CWI", year = 1989, month = sep, annote = "CWI Quarterly 2 (3): 219–246", url = "http://www.markrtuttle.com/data/papers/lt89-cwi.pdf" } @book{Lynch:1996, author = {N. A. Lynch}, title = {Distributed Algorithms}, - year = {1996}, + year = 1996, publisher = {Morgan Kaufmann Publishers Inc.}, } @InProceedings{PrenEsp00, author = {Prensa Nieto, L. and J. Esparza}, title = {Verifying Single and Multi-mutator Garbage Collectors with {Owicki/Gries} in {Isabelle/HOL}}, booktitle = {Mathematical Foundations of Computer Science (MFCS 2000)}, editor = {M. Nielsen and B. Rovan}, - publisher = {Springer-Verlag}, - series = {LNCS}, + publisher = {Springer}, + series = lncs, volume = 1893, pages = {619--628}, year = 2000 } -@PhdThesis{Prensa-PhD,author={Prensa Nieto, L.}, -title={Verification of Parallel Programs with the Owicki-Gries and -Rely-Guarantee Methods in Isabelle/HOL}, -school={Technische Universit{\"a}t M{\"u}nchen},year=2002} +@PhdThesis{Prensa-PhD, + author={Prensa Nieto, L.}, + title={Verification of Parallel Programs with the Owicki-Gries and Rely-Guarantee Methods in Isabelle/HOL}, + school={Technische Universit{\"a}t M{\"u}nchen}, + year=2002 +} -@inproceedings{Prensa-ESOP03,author={Prensa Nieto, L.}, -title={The {Rely-Guarantee} Method in {Isabelle/HOL}}, -booktitle={European Symposium on Programming (ESOP'03)},editor={P. Degano}, -publisher=Springer,series=LNCS,volume=2618,pages={348--362},year=2003} +@inproceedings{Prensa-ESOP03, + author={Prensa Nieto, L.}, + title={The {Rely-Guarantee} Method in {Isabelle/HOL}}, + booktitle={ESOP'2003}, + editor={P. Degano}, + publisher="Springer", + series=lncs,volume=2618, + pages={348--362}, + year=2003 +} @book{Winskel:1993, author = {Winskel, G.}, title = {The Formal Semantics of Programming Languages}, - year = {1993}, + year = 1993, publisher = {MIT Press}, address = "Cambridge, MA" } @article{DBLP:journals/cacm/DijkstraLMSS78, author = {E. W. Dijkstra and L. Lamport and A. J. Martin and C. S. Scholten and E. F. M. Steffens}, title = {On-the-Fly Garbage Collection: An Exercise in Cooperation}, journal = cacm, - volume = {21}, - number = {11}, - year = {1978}, + volume = 21, + number = 11, + year = 1978, pages = {966-975}, ee = {http://doi.acm.org/10.1145/359642.359655}, bibsource = {DBLP, http://dblp.uni-trier.de} } @inproceedings{DBLP:conf/lpar/Schirmer04, author = {N. Schirmer}, title = {A Verification Environment for Sequential Imperative Programs in Isabelle/HOL}, booktitle = {LPAR}, - year = {2004}, + year = 2004, pages = {398-414}, ee = {https://doi.org/10.1007/978-3-540-32275-7_26}, crossref = {DBLP:conf/lpar/2004}, bibsource = {DBLP, http://dblp.uni-trier.de} } % isbn = {3-540-25236-3}, @proceedings{DBLP:conf/lpar/2004, editor = {F. Baader and A. Voronkov}, title = {Logic for Programming, Artificial Intelligence, and Reasoning, 11th International Conference, LPAR 2004, Montevideo, Uruguay, March 14-18, 2005, Proceedings}, booktitle = {LPAR}, publisher = {Springer}, series = lncs, volume = {3452}, year = {2005}, bibsource = {DBLP, http://dblp.uni-trier.de} } @article{DBLP:journals/cacm/SewellSONM10, author = {P. Sewell and S. Sarkar and S. Owens and F. Zappa Nardelli and M. O. Myreen}, title = {{x86-TSO}: a rigorous and usable programmer's model for x86 multiprocessors}, journal = cacm, volume = {53}, number = {7}, year = {2010}, pages = {89-97}, ee = {http://doi.acm.org/10.1145/1785414.1785443}, bibsource = {DBLP, http://dblp.uni-trier.de} } @inproceedings{DBLP:conf/mfcs/BoerRH99, author = {de Boer, F. S. and de Roever, W. P. and U. Hannemann}, title = {The Semantic Foundations of a Compositional Proof Method for Synchronously Communicating Processes}, booktitle = {MFCS}, year = {1999}, pages = {343-353}, ee = {https://doi.org/10.1007/3-540-48340-3_31}, crossref = {DBLP:conf/mfcs/1999}, bibsource = {DBLP, http://dblp.uni-trier.de} } % isbn = {3-540-66408-4}, @proceedings{DBLP:conf/mfcs/1999, editor = {M. Kutylowski and L. Pacholski and T. Wierzbicki}, title = {Mathematical Foundations of Computer Science 1999, 24th International Symposium, MFCS'99, Szklarska Poreba, Poland, September 6-10, 1999, Proceedings}, booktitle = {MFCS}, publisher = {Springer}, series = lncs, volume = {1672}, year = {1999}, bibsource = {DBLP, http://dblp.uni-trier.de} } % isbn = {0-521-80608-9}, -@book{DBLP:books/cu/RoeverBH2001, +@book{deRoeverEtAl:2001, author = {de Roever, W. P. and de Boer, F. S. and U. Hannemann and J. Hooman and Y. Lakhnech and M. Poel and J. Zwiers}, title = {Concurrency Verification: Introduction to Compositional and Noncompositional Methods}, publisher = {Cambridge University Press}, series = {Cambridge Tracts in Theoretical Computer Science}, - volume = {54}, - year = {2001}, - bibsource = {DBLP, http://dblp.uni-trier.de} + volume = 54, + year = 2001, } @InProceedings{Pizlo+2010PLDI, author = {F. Pizlo and L. Ziarek and P. Maj and A. L. Hosking and E. Blanton and J. Vitek}, title = {Schism: Fragmentation-Tolerant Real-Time Garbage Collection}, booktitle = {Proceedings of the ACM SIGPLAN Conference on Programming Language Design and Implementation}, year = 2010, pages = "146--159", month = jun, address = {Toronto, Canada}, doi = {10.1145/1806596.1806615} } @PhdThesis{Pizlo201xPhd, author = {F. Pizlo}, title = {Fragmentation Tolerant Real Time Garbage Collection}, school = {Purdue University}, year = {201x} } @article{DBLP:journals/toplas/LamportS84, - author = {L. Lamport and - F. B. Schneider}, + author = {L. Lamport and F. B. Schneider}, title = {The ``{Hoare Logic}'' of {CSP}, and All That}, journal = toplas, volume = {6}, number = {2}, year = {1984}, pages = {281-296}, ee = {http://doi.acm.org/10.1145/2993.357247}, bibsource = {DBLP, http://dblp.uni-trier.de} } @article{DBLP:journals/acta/Lamport80, author = {L. Lamport}, title = {The ``{Hoare Logic}'' of Concurrent Programs}, journal = {Acta Informatica}, volume = {14}, year = {1980}, pages = {21-37}, ee = {https://doi.org/10.1007/BF00289062}, bibsource = {DBLP, http://dblp.uni-trier.de} } @article{DBLP:journals/jar/Berghofer12, author = {S. Berghofer}, title = {A Solution to the {PoplMark} Challenge Using {de Bruijn} Indices in {Isabelle/HOL}}, journal = {J. Autom. Reasoning}, volume = {49}, number = {3}, year = {2012}, pages = {303-326}, ee = {https://doi.org/10.1007/s10817-011-9231-4}, bibsource = {DBLP, http://dblp.uni-trier.de} } % isbn = {3-540-44044-5} @InCollection{PittsAM:opespe, author = {A. M. Pitts}, title = {Operational Semantics and Program Equivalence}, booktitle = {Applied Semantics, Advanced Lectures}, pages = {378--412}, - publisher = {Springer-Verlag}, + publisher = {Springer}, year = 2002, editor = {G. Barthe and P. Dybjer and J. Saraiva}, volume = 2395, series = lncs, note = {International Summer School, {APPSEM} 2000, Caminha, Portugal, September 9--15, 2000}, } +@InCollection{MannaPnueli:1991, + author = "Z. Manna and A. Pnueli", + title = "Tools and rules for the practicing verifier", + booktitle = "CMU Computer Science: A 25th Anniversary Commemorative", + pages = "121–-156", + publisher = "ACM Press and Addison-Wesley", + year = 1991, + editor = "R. F. Rashid", + note = "Also Technical Report STAN-CS-90-1321" +} + % isbn = {978-0-387-94459-3}, @book{MannaPnueli:1995, author = {Z. Manna and A. Pnueli}, - title = {Temporal verification of reactive systems - safety}, + title = {Temporal verification of reactive systems - {Safety}}, publisher = {Springer}, year = 1995, pages = {I-XV, 1-512}, } +@inproceedings{MannaPnueli:2010, + author = {Z. Manna and A. Pnueli}, + title = {Temporal Verification of Reactive Systems: Response}, + booktitle = {Time for Verification, Essays in Memory of Amir Pnueli}, + series = lncs, + volume = 6200, + publisher = {Springer}, + pages = {279--361}, + year = 2010, + doi = {10.1007/978-3-642-13754-9\_13}, +} + @article{DBLP:journals/acta/LevinG81, author = {G. Levin and D. Gries}, title = {A Proof Technique for Communicating Sequential Processes}, journal = {Acta Inf.}, volume = {15}, year = {1981}, pages = {281-302}, ee = {https://doi.org/10.1007/BF00289266}, bibsource = {DBLP, http://dblp.uni-trier.de} } @inproceedings{CousotCousot:1980, author = {P. Cousot and R. Cousot}, title = "Semantic Analysis of {Communicating Sequential Processes} (Shortened Version)", booktitle = {ICALP}, series = lncs, volume = {85}, publisher = {Springer}, year = 1980, pages = {119-133}, ee = {https://doi.org/10.1007/3-540-10003-2_65}, } -@InCollection{MannaPnueli:1991, - author = "Z. Manna and A. Pnueli", - title = "Tools and rules for the practicing verifier", - booktitle = "CMU Computer Science: A 25th Anniversary Commemorative", - pages = "121–-156", - publisher = "ACM Press and Addison-Wesley", - year = 1991, - editor = "R. F. Rashid" -} - % isbn = {0-3211-4306-X}, -@book{DBLP:books/aw/Lamport2002, +@book{Lamport:2002, author = {L. Lamport}, title = {Specifying Systems, The TLA+ Language and Tools for Hardware and Software Engineers}, publisher = {Addison-Wesley}, - year = {2002}, - bibsource = {DBLP, http://dblp.uni-trier.de} + year = 2002 } @article{TLA-AFP, author = {G. Grov and S. Merz}, title = {A Definitional Encoding of TLA* in Isabelle/HOL}, journal = {Archive of Formal Proofs}, month = nov, year = 2011, note = {\url{http://isa-afp.org/entries/TLA}, Formal proof development}, ISSN = {2150-914x}, } @inproceedings{DBLP:conf/itp/CohenS10, author = {E. Cohen and B. Schirmer}, title = {From Total Store Order to Sequential Consistency: A Practical Reduction Theorem}, booktitle = {ITP}, year = {2010}, pages = {403-418}, ee = {https://doi.org/10.1007/978-3-642-14052-5_28}, crossref = {DBLP:conf/itp/2010}, bibsource = {DBLP, http://dblp.uni-trier.de} } % isbn = {978-3-642-14051-8}, @proceedings{DBLP:conf/itp/2010, editor = {M. Kaufmann and L. C. Paulson}, title = {Interactive Theorem Proving, First International Conference, ITP 2010, Edinburgh, UK, July 11-14, 2010. Proceedings}, booktitle = {ITP}, publisher = {Springer}, series = lncs, volume = {6172}, year = {2010}, ee = {https://doi.org/10.1007/978-3-642-14052-5}, bibsource = {DBLP, http://dblp.uni-trier.de} } @inproceedings{DBLP:conf/fase/NipkowN99, author = {T. Nipkow and Prensa Nieto, L.}, title = "{Owicki/Gries} in {Isabelle/HOL}", booktitle = {FASE}, year = {1999}, pages = {188-203}, ee = {https://doi.org/10.1007/978-3-540-49020-3_13}, crossref = {DBLP:conf/fase/1999}, bibsource = {DBLP, http://dblp.uni-trier.de} } % isbn = {3-540-65718-5}, @proceedings{DBLP:conf/fase/1999, editor = {J.-P. Finance}, title = {Fundamental Approaches to Software Engineering, Second Internationsl Conference, FASE'99, Held as Part of the European Joint Conferences on the Theory and Practice of Software, ETAPS'99, Amsterdam, The Netherlands, March 22-28, 1999, Proceedings}, booktitle = {FASE}, publisher = {Springer}, series = lncs, volume = {1577}, year = {1999}, bibsource = {DBLP, http://dblp.uni-trier.de} } @article{CousotCousot89-IC, author = {Cousot, P. and Cousot, R.}, title = {A language independent proof of the soundness and completeness of generalized {H}oare logic}, journal = {Information and Computation}, volume = 80, number = 2, pages = {165--191}, month = feb, year = 1989, } -@inproceedings{DBLP:conf/popl/DoligezG94, - author = {D. Doligez and - G. Gonthier}, - title = {Portable, Unobtrusive Garbage Collection for Multiprocessor - Systems}, - booktitle = {POPL}, - year = {1994}, - pages = {70-83}, - ee = {http://doi.acm.org/10.1145/174675.174673}, - crossref = {DBLP:conf/popl/1994}, - bibsource = {DBLP, http://dblp.uni-trier.de} -} - -% isbn = {0-89791-636-0}, -@proceedings{DBLP:conf/popl/1994, - editor = {Boehm, H.-J. and - B. Lang and - D. M. Yellin}, - title = {Conference Record of POPL'94: 21st ACM SIGPLAN-SIGACT Symposium - on Principles of Programming Languages, Portland, Oregon, - USA, January 17-21, 1994}, - booktitle = popl, - publisher = {ACM Press}, - year = {1994}, - ee = {http://dl.acm.org/citation.cfm?id=174675}, - bibsource = {DBLP, http://dblp.uni-trier.de} -} - @article{DBLP:journals/jlp/Plotkin04, author = {G. D. Plotkin}, title = {The origins of structural operational semantics}, journal = {Journal of Logic and Algebraic Programming}, volume = {60-61}, year = {2004}, pages = {3-15}, ee = {https://doi.org/10.1016/j.jlap.2004.03.009}, bibsource = {DBLP, http://dblp.uni-trier.de} } @Book{ConcreteSemantics:2014, - author = "T. Nipkow and G. Klein", + author = "T. Nipkow and G. Klein", title = "Concrete Semantics: A Proof Assistant Approach", - publisher = Springer, + publisher = "Springer", year = 2014, url = "http://www.in.tum.de/~nipkow/Concrete-Semantics/" } @Book{Hoare:1985, author = {C.A.R. Hoare}, title = "{Communicating Sequential Processes}", publisher = {Prentice-Hall}, - year = 1985, + year = 1985, series = {International Series In Computer Science}, url = "http://www.usingcsp.com/" } @Book{Milner:1980, - author = "R. Milner", + author = "R. Milner", title = "A Calculus of Communicating Systems", - publisher = Springer, + publisher = "Springer", year = 1980 } @book{Milner:1989, author = {R. Milner}, title = "Communication and Concurrency", year = {1989}, publisher = {Prentice-Hall}, address = "Englewood Cliffs, NJ", } -@article{DBLP:journals/tcs/FelleisenH92, +@article{FelleisenHieb:1992, author = {M. Felleisen and R. Hieb}, title = {The Revised Report on the Syntactic Theories of Sequential Control and State}, - journal = {Theor. Comput. Sci.}, - volume = {103}, - number = {2}, - year = {1992}, + journal = tcs, + volume = 103, + number = 2, + year = 1992, pages = {235-271}, - ee = {https://doi.org/10.1016/0304-3975(92)90014-7}, - bibsource = {DBLP, http://dblp.uni-trier.de} + doi = {10.1016/0304-3975(92)90014-7}, } -@inproceedings{DBLP:conf/tphol/Jackson98, - author = {Paul B. Jackson}, +@inproceedings{Jackson:1998, + author = {P. B. Jackson}, title = {Verifying a Garbage Collection Algorithm}, booktitle = {TPHOLs}, - year = {1998}, - pages = {225-244}, - ee = {https://doi.org/10.1007/BFb0055139}, - crossref = {DBLP:conf/tphol/1998}, - bibsource = {DBLP, http://dblp.uni-trier.de} -} - -@proceedings{DBLP:conf/tphol/1998, - editor = {Jim Grundy and - Malcolm C. Newey}, - title = {Theorem Proving in Higher Order Logics, 11th International - Conference, TPHOLs'98, Canberra, Australia, September 27 - - October 1, 1998, Proceedings}, - booktitle = {TPHOLs}, publisher = {Springer}, - series = {Lecture Notes in Computer Science}, + series = lncs, volume = {1479}, - year = {1998}, - isbn = {3-540-64987-5}, - bibsource = {DBLP, http://dblp.uni-trier.de} + year = 1998, + pages = {225-244}, + doi = {10.1007/BFb0055139}, } -@article{DBLP:journals/logcom/Coupet-GrimalN03, - author = {Coupet-Grimal, S. and - Nouvet, C.}, - title = {Formal Verification of an Incremental Garbage Collector}, - journal = {J. Log. Comput.}, - volume = {13}, - number = {6}, - year = {2003}, - pages = {815-833}, - ee = {https://doi.org/10.1093/logcom/13.6.815}, - bibsource = {DBLP, http://dblp.uni-trier.de} -} - -@article{DBLP:journals/logcom/Coupet-Grimal03, +@article{CoupetGrimal:2003, author = {Coupet-Grimal, S.}, - title = {An Axiomatization of Linear Temporal Logic in the Calculus - of Inductive Constructions}, - journal = {J. Log. Comput.}, - volume = {13}, - number = {6}, - year = {2003}, + title = {An Axiomatization of {Linear Temporal Logic} in the {Calculus + of Inductive Constructions}}, + journal = {Journal of Logic and Computation}, + volume = 13, + number = 6, + year = 2003, pages = {801-813}, - ee = {https://doi.org/10.1093/logcom/13.6.801}, - bibsource = {DBLP, http://dblp.uni-trier.de} + doi = {10.1093/logcom/13.6.801}, } @article{DBLP:journals/entcs/SchirmerW09, author = {N. Schirmer and M. Wenzel}, title = {State Spaces - The Locale Way}, journal = {Electr. Notes Theor. Comput. Sci.}, volume = {254}, year = {2009}, pages = {161-179}, ee = {https://doi.org/10.1016/j.entcs.2009.09.065}, bibsource = {DBLP, http://dblp.uni-trier.de} } @inproceedings{Ridge:2009, author = {T. Ridge}, title = {Verifying distributed systems: the operational approach}, booktitle = {POPL'2009}, pages = {429--440}, publisher = {{ACM}}, year = 2009, doi = {10.1145/1480881.1480934}, } + +@book{ChandyMisra:1989, + author = {Chandy, K. M. and Misra, J.}, + title = {Parallel program design - a foundation}, + publisher = {Addison-Wesley}, + year = 1989, + isbn = {978-0-201-05866-6}, +} + +% 2019-06 FIXME not yet published? +@Article{vanGlabbeekHofner:2019, + author = {van Glabbeek, R. J. and H\"{o}fner, P.}, + title = {Progress, Justness and Fairness}, + journal = csur, + year = 2019, + url = {http://arxiv.org/abs/1810.07414v1} +} + +@article{AptDelporteGallet:1986, + author = {Krzysztof R. Apt and + Carole Delporte{-}Gallet}, + title = {Syntax Directed Analysis of Liveness Properties}, + journal = {Information and Control}, + volume = 68, + number = {1-3}, + pages = {223--253}, + year = 1986, + doi = {10.1016/S0019-9958(86)80037-7}, +} + +@inproceedings{MisraChandySmith:1982, + author = {J. Misra and + K. M. Chandy and + T. Smith}, + title = {Proving Safety and Liveness of Communicating Processes with Examples}, + booktitle = {PODC'1982}, + publisher = {{ACM}}, + pages = {201--208}, + year = 1982, + doi = {10.1145/800220.806698}, +} + +@article{AlpernSchneider:1985, + author = {B. Alpern and + F. B. Schneider}, + title = {Defining Liveness}, + journal = ipl, + volume = 21, + number = 4, + pages = {181--185}, + year = 1985, + doi = {10.1016/0020-0190(85)90056-0}, +} + +@TechReport{Schneider:1987, + author = {F. B. Schneider}, + title = {Decomposing Properties into Safety and Liveness using Predicate Logic}, + institution = {Department of Computer Science, Cornell University}, + year = 1987, + number = {87-874}, + month = oct +} + +% FIXME guesswork on some of the biblio data +@article{Kindler:1994, + author = {E. Kindler}, + title = {Safety and liveness properties: A survey}, + journal = {Bulletin of the European Association for Theoretical Computer Science}, + year = 1994, + volume = 53, + number = 30, + pages = {268--272}, + month = 6, + url = "http://citeseerx.ist.psu.edu/viewdoc/download?doi=10.1.1.43.8206&rep=rep1&type=pdf", +} + +@article{Sistla:1994, + author = {Sistla, A. P.}, + title = {Safety, Liveness and Fairness in Temporal Logic}, + journal = fac, + volume = 6, + number = 5, + pages = {495--512}, + year = 1994, + doi = {10.1007/BF01211865}, +} + +@article{OwickiLamport:1982, + author = {S. S. Owicki and + L. Lamport}, + title = {Proving Liveness Properties of Concurrent Programs}, + journal = toplas, + volume = 4, + number = 3, + pages = {455--495}, + year = 1982, + doi = {10.1145/357172.357178}, +} + +@inproceedings{ChangMannaPnueli:1992, + author = {E. Y. Chang and + Z. Manna and + A. Pnueli}, + title = {Characterization of Temporal Property Classes}, + booktitle = {ICALP'1992}, + pages = {474--486}, + year = 1992, + series = lncs, + volume = 623, + publisher = {Springer}, + doi = {10.1007/3-540-55719-9\_97}, +} + +@inproceedings{MannaPnueli:1988, + author = {Z. Manna and + A. Pnueli}, + title = {The anchored version of the temporal framework}, + booktitle = {Linear Time, Branching Time and Partial Order in Logics and Models + for Concurrency, School/Workshop, Noordwijkerhout, The Netherlands, + May 30 - June 3, 1988, Proceedings}, + editor = {de Bakker, J. W. and + de Roever, W. P. and + Rozenberg, G.}, + pages = {201--284}, + year = 1988, + doi = {10.1007/BFb0013024}, + volume = 354, + publisher = {Springer}, + series = lncs, +} + +@inproceedings{PnueliArons:2003, + author="Pnueli, A. and Arons, T.", + title="{TLPVS}: A {PVS}-Based {LTL} Verification System", + bookTitle="Verification: Theory and Practice: Essays Dedicated to Zohar Manna on the Occasion of His 64th Birthday", + year=2003, + publisher="Springer", + series=lncs, + volume=2772, + address="Berlin, Heidelberg", + pages="598--625", + abstract="In this paper we present our pvs implementation of a linear temporal logic verification system. The system includes a set of theories defining a temporal logic, a number of proof rules for proving soundness and response properties, and strategies which aid in conducting the proofs. In addition to implementing a framework for existing rules, we have also derived new methods which are particularly useful in a deductive ltl system. A distributed rank rule for the verification of response properties in parameterized systems is presented, and a methodology is detailed for reducing compassion requirements to justice requirements (strong fairness to weak fairness). Special attention has been paid to the verification of unbounded systems --- systems in which the number of processes is unbounded --- and our verification rules are appropriate to such systems.", + isbn="978-3-540-39910-0", + doi="10.1007/978-3-540-39910-0_26", +} + +@article{AbadiPlotkin:1993, + author = {Mart{\'{\i}}n Abadi and + Gordon D. Plotkin}, + title = {A Logical View of Composition}, + journal = tcs, + volume = 114, + number = 1, + pages = {3--30}, + year = 1993, + doi = {10.1016/0304-3975(93)90151-I}, + note = {Journal version of FIXME Abadi, Plotkin (POPL, 1991)} +} diff --git a/thys/ConcurrentIMP/document/root.tex b/thys/ConcurrentIMP/document/root.tex --- a/thys/ConcurrentIMP/document/root.tex +++ b/thys/ConcurrentIMP/document/root.tex @@ -1,87 +1,88 @@ \documentclass[11pt,a4paper]{article} \usepackage[a4paper,margin=1cm,footskip=.5cm]{geometry} \usepackage{isabelle,isabellesym} \usepackage{amssymb} +\usepackage{wasysym} \usepackage[english]{babel} % lifted composition. \newcommand{\isasymbigcirc}{\isamath{\circ}} % Bibliography \usepackage[authoryear,longnamesfirst,sort]{natbib} \bibpunct();A{}, % this should be the last package used \usepackage{pdfsetup} % urls in roman style, theory text in math-similar italics \urlstyle{rm} \isabellestyle{it} \renewcommand{\ttdefault}{cmtt} % CM rather than courier for \tt % for uniform font size \renewcommand{\isastyle}{\isastyleminor} % Abstract various things that might change. \newcommand{\ccode}[1]{\texttt{#1}} \newcommand{\isabelletype}[1]{\emph{#1}} \newcommand{\isabelleterm}[1]{\emph{#1}} \begin{document} \title{CIMP} \author{Peter Gammie} \maketitle \begin{abstract} CIMP extends the small imperative language IMP with control non-determinism and constructs for synchronous message passing. \end{abstract} \tableofcontents % sane default for proof documents \parindent 0pt\parskip 0.5ex % generated text of all theories \input{session} \section{Concluding remarks} Previously \citet{DBLP:conf/fase/NipkowN99,Prensa-PhD,Prensa-ESOP03}\footnote{The theories are in \texttt{\$ISABELLE/src/HOL/Hoare\_Parallel}.} have developed the classical Owicki/Gries and Rely-Guarantee paradigms for the verification of shared-variable concurrent programs in Isabelle/HOL. These have been used to show the correctness of a garbage collector \citep{PrenEsp00}. We instead use synchronous message passing, which is significantly less explored. \citet{DBLP:conf/mfcs/BoerRH99,DBLP:books/cu/RoeverBH2001} provide compositional systems for \emph{terminating} systems. We have instead adopted Lamport's paradigm of a single global invariant and local proof obligations as the systems we have in mind are tightly coupled and it is not obvious that the proofs would be easier on a decomposed system; see \citet[\S1.6.6]{DBLP:books/cu/RoeverBH2001} for a concurring opinion. Unlike the generic sequential program verification framework Simpl \citep{DBLP:conf/lpar/Schirmer04}, we do not support function calls, or a sophisticated account of state spaces. Moreover we do no meta-theory beyond showing the simple VCG is sound (\S\ref{sec:cimp-vcg}). \bibliographystyle{plainnat} \bibliography{root} \addcontentsline{toc}{section}{References} \end{document} %%% Local Variables: %%% mode: latex %%% TeX-master: t %%% End: diff --git a/thys/ConcurrentIMP/ex/CIMP_locales.thy b/thys/ConcurrentIMP/ex/CIMP_locales.thy new file mode 100644 --- /dev/null +++ b/thys/ConcurrentIMP/ex/CIMP_locales.thy @@ -0,0 +1,60 @@ +(*<*) +theory CIMP_locales +imports + "../CIMP" +begin + +(*>*) +section\ One locale per process \ + +text\ + +A sketch of what we're doing in \ConcurrentGC\, for quicker testing. + +FIXME write some lemmas that further exercise the generated thms. + +\ + +locale P1 +begin + +definition com :: "(unit, string, unit, nat) com" where + "com = \''A''\ WHILE ((<) 0) DO \''B''\ \\s. s - 1\ OD" + +intern_com com_def +print_theorems (* P1.com_interned, P1.loc_defs *) + +locset_definition "loop = {B}" +print_theorems (* P1.loop.memb, P1.loop_def *) +thm locset_cache (* the two facts in P1.loop.memb *) + +definition "assertion = atS False loop" + +end + +thm locset_cache (* the two facts in P1.loop.memb *) + +locale P2 +begin + +thm locset_cache (* the two facts in P1.loop.memb *) + +definition com :: "(unit, string, unit, nat) com" where + "com = \''C''\ WHILE ((<) 0) DO \''A''\ \Suc\ OD" + +intern_com com_def +locset_definition "loop = {A}" +print_theorems + +end + +thm locset_cache (* four facts: P1.loop.memb, P2.loop.memb *) + +primrec coms :: "bool \ (unit, string, unit, nat) com" where + "coms False = P1.com" +| "coms True = P2.com" + +(*<*) + +end +(*>*) diff --git a/thys/ConcurrentIMP/ex/CIMP_one_place_buffer.thy b/thys/ConcurrentIMP/ex/CIMP_one_place_buffer.thy new file mode 100644 --- /dev/null +++ b/thys/ConcurrentIMP/ex/CIMP_one_place_buffer.thy @@ -0,0 +1,165 @@ +(*<*) +(* + * 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 CIMP_one_place_buffer +imports + "../CIMP" +begin + +(*>*) +section\Example: a one-place buffer \label{sec:one_place_buffer}\ + +text\ + +To demonstrate the CIMP reasoning infrastructure, we treat the trivial +one-place buffer example of @{cite [cite_macro=citet] +\\S3.3\ "DBLP:journals/toplas/LamportS84"}. Note that the +semantics for our language is different to @{cite +[cite_macro=citeauthor] "DBLP:journals/toplas/LamportS84"}'s, who +treated a historical variant of CSP (i.e., not the one in @{cite +"Hoare:1985"}). + +We introduce some syntax for fixed-topology (static channel-based) +scenarios. + +\ + +abbreviation + rcv_syn :: "'location \ 'channel \ ('val \ 'state \ 'state) + \ (unit, 'location, 'channel \ 'val, 'state) com" ("\_\/ _\_" [0,0,81] 81) +where + "\l\ ch\f \ \l\ Response (\q s. if fst q = ch then {(f (snd q) s, ())} else {})" + +abbreviation + snd_syn :: "'location \ 'channel \ ('state \ 'val) + \ (unit, 'location, 'channel \ 'val, 'state) com" ("\_\/ _\_" [0,0,81] 81) +where + "\l\ ch\f \ \l\ Request (\s. (ch, f s)) (\ans s. {s})" + +text\ + +These definitions largely follow @{cite [cite_macro=citet] +"DBLP:journals/toplas/LamportS84"}. We have three processes +communicating over two channels. We enumerate program locations. + +\ + +datatype ex_chname = \12 | \23 +type_synonym ex_val = nat +type_synonym ex_ch = "ex_chname \ ex_val" +datatype ex_loc = r12 | r23 | s23 | s12 +datatype ex_proc = p1 | p2 | p3 + +type_synonym ex_pgm = "(unit, ex_loc, ex_ch, ex_val) com" +type_synonym ex_pred = "(unit, ex_loc, ex_proc, ex_ch, ex_val) state_pred" +type_synonym ex_state = "(unit, ex_loc, ex_proc, ex_ch, ex_val) system_state" +type_synonym ex_sys = "(unit, ex_loc, ex_proc, ex_ch, ex_val) system" +type_synonym ex_history = "(ex_ch \ unit) list" + +text\ + +We further specialise these for our particular example. + +\ + +primrec + ex_coms :: "ex_proc \ ex_pgm" +where + "ex_coms p1 = \s12\ \12\id" +| "ex_coms p2 = LOOP DO \r12\ \12\(\v _. v) ;; \s23\ \23\id OD" +| "ex_coms p3 = \r23\ \23\(\v _. v)" + +text\ + +Each process starts with an arbitrary initial local state. + +\ + +abbreviation ex_init :: "(ex_proc \ ex_val) \ bool" where + "ex_init \ \True\" + +abbreviation sys :: ex_sys where + "sys \ \PGMs = ex_coms, INIT = ex_init, FAIR = \True\\" (* FIXME add fairness hypotheses *) + +text\ + +The following adapts Kai Engelhardt's, from his notes titled +\emph{Proving an Asynchronous Message Passing Program Correct}, +2011. The history variable tracks the causality of the system, which I +feel is missing in Lamport's treatment. We tack on Lamport's invariant +so we can establish \Etern_pred\. + +\ + +abbreviation + filter_on_channel :: "ex_chname \ ex_state \ ex_val list" ("\_" [100] 101) +where + "\ch \ map (snd \ fst) \ filter ((=) ch \ fst \ fst) \ HST" + +definition IL :: ex_pred where + "IL = pred_conjoin [ + at p1 s12 \<^bold>\ LIST_NULL \\12 + , terminated p1 \<^bold>\ \\12 \<^bold>= (\s. [s\ p1]) + , at p2 r12 \<^bold>\ \\12 \<^bold>= \\23 + , at p2 s23 \<^bold>\ \\12 \<^bold>= \\23 \<^bold>@ (\s. [s\ p2]) \<^bold>\ (\s. s\ p1 = s\ p2) + , at p3 r23 \<^bold>\ LIST_NULL \\23 + , terminated p3 \<^bold>\ \\23 \<^bold>= (\s. [s\ p2]) \<^bold>\ (\s. s\ p1 = s\ p3) + ]" + +text\ + +If @{const p3} terminates, then it has @{const p1}'s value. This is +stronger than @{cite [cite_macro=citeauthor] +"DBLP:journals/toplas/LamportS84"}'s as we don't ask that the first +process has also terminated. + +\ + +definition Etern_pred :: ex_pred where + "Etern_pred = (terminated p3 \<^bold>\ (\s. s\ p1 = s\ p3))" + +text\ + +Proofs from here down. + +\ + +lemma correct_system: + assumes "IL sh" + shows "Etern_pred sh" +using assms unfolding Etern_pred_def IL_def by simp + +lemma IL_p1: "ex_coms, p1, lconst {} \ \IL\ \s12\ \12\(\s. s)" +apply (rule vcg.intros) +apply (rename_tac p') +apply (case_tac p'; clarsimp simp: IL_def atLs_def) +done + +lemma IL_p2: "ex_coms, p2, lconst {r12} \ \IL\ \s23\ \23\(\s. s)" +apply (rule vcg.intros) +apply (rename_tac p') +apply (case_tac p'; clarsimp simp: IL_def) +done + +lemma IL: "sys \\<^bsub>pre\<^esub> IL" +apply (rule VCG) + apply (clarsimp simp: IL_def atLs_def dest!: initial_stateD) +apply (rename_tac p) +apply (case_tac p; clarsimp simp: IL_p1 IL_p2) +done + +lemma IL_valid: "sys \ \\IL\" +by (rule valid_prerun_lift[OF IL]) + +(*<*) + +end +(*>*) diff --git a/thys/ConcurrentIMP/ex/CIMP_unbounded_buffer.thy b/thys/ConcurrentIMP/ex/CIMP_unbounded_buffer.thy new file mode 100644 --- /dev/null +++ b/thys/ConcurrentIMP/ex/CIMP_unbounded_buffer.thy @@ -0,0 +1,145 @@ +(*<*) +(* + * 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 CIMP_unbounded_buffer +imports + "../CIMP" + "HOL-Library.Prefix_Order" +begin + +abbreviation (input) + Receive :: "'location \ 'channel \ ('val \ 'state \ 'state) + \ (unit, 'location, 'channel \ 'val, 'state) com" ("\_\/ _\_" [0,0,81] 81) +where + "\l\ ch\f \ \l\ Response (\q s. if fst q = ch then {(f (snd q) s, ())} else {})" + +abbreviation (input) + Send :: "'location \ 'channel \ ('state \ 'val) \ ('state \ 'state) + \ (unit, 'location, 'channel \ 'val, 'state) com" ("\_\/ _\_" [0,0,81] 81) +where + "\l\ ch\f \ \l\ Request (\s. (ch, fst f s)) (\ans s. {snd f s})" + +(*>*) +section\Example: an unbounded buffer \label{sec:unbounded_buffer}\ + +text\ + +This is more literally Kai Engelhardt's example from his notes titled +\emph{Proving an Asynchronous Message Passing Program Correct}, 2011. + +\ + +datatype ex_chname = \12 | \23 +type_synonym ex_val = nat +type_synonym ex_ls = "ex_val list" +type_synonym ex_ch = "ex_chname \ ex_val" +datatype ex_loc = c1 | r12 | r23 | s23 | s12 +datatype ex_proc = p1 | p2 | p3 + +type_synonym ex_pgm = "(unit, ex_loc, ex_ch, ex_ls) com" +type_synonym ex_pred = "(unit, ex_loc, ex_proc, ex_ch, ex_ls) state_pred" +type_synonym ex_state = "(unit, ex_loc, ex_proc, ex_ch, ex_ls) system_state" +type_synonym ex_sys = "(unit, ex_loc, ex_proc, ex_ch, ex_ls) system" +type_synonym ex_history = "(ex_ch \ unit) list" + +text\ + +The local state for the producer process contains all values produced; consider that ghost state. + +\ + +abbreviation (input) snoc :: "'a \ 'a list \ 'a list" where "snoc x xs \ xs @ [x]" + +primrec ex_coms :: "ex_proc \ ex_pgm" where + "ex_coms p1 = LOOP DO \c1\ LocalOp (\xs. {snoc x xs |x. True}) ;; \s12\ \12\(last, id) OD" +| "ex_coms p2 = LOOP DO \r12\ \12\snoc + \ \c1\ IF (\s. length s > 0) THEN \s23\ \12\(hd, tl) FI + OD" +| "ex_coms p3 = LOOP DO \r23\ \23\snoc OD" + +abbreviation ex_init :: "(ex_proc \ ex_ls) \ bool" where + "ex_init s \ \p. s p = []" + +abbreviation sys :: ex_sys where + "sys \ \PGMs = ex_coms, INIT = ex_init, FAIR = \True\\" (* FIXME add fairness hypotheses *) + +abbreviation + filter_on_channel :: "ex_chname \ ex_state \ ex_val list" ("\_" [100] 101) +where + "\ch \ map (snd \ fst) \ filter ((=) ch \ fst \ fst) \ HST" + +definition I_pred :: ex_pred where + "I_pred = pred_conjoin [ + at p1 c1 \<^bold>\ \\12 \<^bold>= (\s. s\ p1) + , at p1 s12 \<^bold>\ (\s. length (s\ p1) > 0 \ butlast (s\ p1) = (\\12) s) + , \\12 \<^bold>\ (\s. s\ p1) + , \\12 \<^bold>= \\23 \<^bold>@ (\s. s\ p2) + , at p2 s23 \<^bold>\ (\s. length (s\ p2) > 0) + , (\s. s\ p3) \<^bold>= \\23 + ]" + +text\ + +The local state of @{const "p3"} is some prefix of the local state of +@{const "p1"}. + +\ + +definition Etern_pred :: ex_pred where + "Etern_pred \ \s. s\ p3 \ s\ p1" + +lemma correct_system: + assumes "I_pred s" + shows "Etern_pred s" +using assms unfolding Etern_pred_def I_pred_def less_eq_list_def prefix_def by clarsimp + +lemma p1_c1[simplified, intro]: + "ex_coms, p1, lconst {s12} \ \I_pred\ \c1\ LocalOp (\xs. { snoc x xs |x. True })" +apply (rule vcg.intros) +apply (clarsimp simp: I_pred_def atS_def) +done + +lemma p1_s12[simplified, intro]: + "ex_coms, p1, lconst {c1} \ \I_pred\ \s12\ \12\(last, id)" +apply (rule vcg.intros) +apply (rename_tac p') +apply (case_tac p'; clarsimp) +apply (clarsimp simp: I_pred_def atS_def) +apply (metis Prefix_Order.prefix_snoc append.assoc append_butlast_last_id) +done + +lemma p2_s23[simplified, intro]: + "ex_coms, p2, lconst {c1, r12} \ \I_pred\ \s23\ \12\(hd, tl)" +apply (rule vcg.intros) +apply (rename_tac p') +apply (case_tac p'; clarsimp) +done + +lemma p2_pi4[intro]: + "ex_coms, p2, lcond {s23} {c1, r12} (\s. s \ []) \ \I_pred\ \c1\ IF (\s. s \ []) THEN c' FI" +apply (rule vcg.intros) +apply (clarsimp simp: I_pred_def atS_def split: lcond_splits) +done + +lemma I: "sys \\<^bsub>pre\<^esub> I_pred" +apply (rule VCG) + apply (clarsimp dest!: initial_stateD simp: I_pred_def atS_def) +apply (rename_tac p) +apply (case_tac p; auto) +done + +lemma I_valid: "sys \ \\I_pred\" +by (rule valid_prerun_lift[OF I]) + +(*<*) + +end +(*>*) diff --git a/thys/Transition_Systems_and_Automata/Automata/Nondeterministic.thy b/thys/Transition_Systems_and_Automata/Automata/Nondeterministic.thy --- a/thys/Transition_Systems_and_Automata/Automata/Nondeterministic.thy +++ b/thys/Transition_Systems_and_Automata/Automata/Nondeterministic.thy @@ -1,761 +1,761 @@ section \Nondeterministic Automata\ theory Nondeterministic imports "../Transition_Systems/Transition_System" "../Transition_Systems/Transition_System_Extra" "../Transition_Systems/Transition_System_Construction" "../Basic/Degeneralization" begin locale automaton = fixes automaton :: "'label set \ 'state set \ ('label \ 'state \ 'state set) \ 'condition \ 'automaton" fixes alphabet initial transition condition assumes automaton[simp]: "automaton (alphabet A) (initial A) (transition A) (condition A) = A" assumes alphabet[simp]: "alphabet (automaton a i t c) = a" assumes initial[simp]: "initial (automaton a i t c) = i" assumes transition[simp]: "transition (automaton a i t c) = t" assumes condition[simp]: "condition (automaton a i t c) = c" begin sublocale transition_system_initial "\ a p. snd a" "\ a p. fst a \ alphabet A \ snd a \ transition A (fst a) p" "\ p. p \ initial A" for A defines path' = path and run' = run and reachable' = reachable and nodes' = nodes by this lemma states_alt_def: "states r p = map snd r" by (induct r arbitrary: p) (auto) lemma trace_alt_def: "trace r p = smap snd r" by (coinduction arbitrary: r p) (auto) lemma successors_alt_def: "successors A p = (\ a \ alphabet A. transition A a p)" by auto lemma reachable_transition[intro]: assumes "a \ alphabet A" "q \ reachable A p" "r \ transition A a q" shows "r \ reachable A p" using reachable.execute assms by force lemma nodes_transition[intro]: assumes "a \ alphabet A" "p \ nodes A" "q \ transition A a p" shows "q \ nodes A" using nodes.execute assms by force lemma path_alphabet: assumes "length r = length w" "path A (w || r) p" shows "w \ lists (alphabet A)" using assms by (induct arbitrary: p rule: list_induct2) (auto) lemma run_alphabet: assumes "run A (w ||| r) p" shows "w \ streams (alphabet A)" using assms by (coinduction arbitrary: w r p) (metis run.cases stream.map szip_smap szip_smap_fst) definition restrict :: "'automaton \ 'automaton" where "restrict A \ automaton (alphabet A) (initial A) (\ a p. if a \ alphabet A then transition A a p else {}) (condition A)" lemma restrict_simps[simp]: "alphabet (restrict A) = alphabet A" "initial (restrict A) = initial A" "transition (restrict A) a p = (if a \ alphabet A then transition A a p else {})" "condition (restrict A) = condition A" unfolding restrict_def by auto lemma restrict_path[simp]: "path (restrict A) = path A" proof (intro ext iffI) show "path A wr p" if "path (restrict A) wr p" for wr p using that by induct auto show "path (restrict A) wr p" if "path A wr p" for wr p using that by induct auto qed lemma restrict_run[simp]: "run (restrict A) = run A" proof (intro ext iffI) show "run A wr p" if "run (restrict A) wr p" for wr p using that by coinduct auto show "run (restrict A) wr p" if "run A wr p" for wr p using that by coinduct auto qed end locale automaton_path = automaton automaton alphabet initial transition condition for automaton :: "'label set \ 'state set \ ('label \ 'state \ 'state set) \ 'condition \ 'automaton" and alphabet initial transition condition + fixes test :: "'condition \ 'label list \ 'state list \ 'state \ bool" begin definition language :: "'automaton \ 'label list set" where "language A \ {w |w r p. length r = length w \ p \ initial A \ path A (w || r) p \ test (condition A) w r p}" lemma language[intro]: assumes "length r = length w" "p \ initial A" "path A (w || r) p" "test (condition A) w r p" shows "w \ language A" using assms unfolding language_def by auto lemma language_elim[elim]: assumes "w \ language A" obtains r p where "length r = length w" "p \ initial A" "path A (w || r) p" "test (condition A) w r p" using assms unfolding language_def by auto lemma language_alphabet: "language A \ lists (alphabet A)" by (auto dest: path_alphabet) lemma restrict_language[simp]: "language (restrict A) = language A" by force end locale automaton_run = automaton automaton alphabet initial transition condition for automaton :: "'label set \ 'state set \ ('label \ 'state \ 'state set) \ 'condition \ 'automaton" and alphabet initial transition condition + fixes test :: "'condition \ 'label stream \ 'state stream \ 'state \ bool" begin definition language :: "'automaton \ 'label stream set" where "language A \ {w |w r p. p \ initial A \ run A (w ||| r) p \ test (condition A) w r p}" lemma language[intro]: assumes "p \ initial A" "run A (w ||| r) p" "test (condition A) w r p" shows "w \ language A" using assms unfolding language_def by auto lemma language_elim[elim]: assumes "w \ language A" obtains r p where "p \ initial A" "run A (w ||| r) p" "test (condition A) w r p" using assms unfolding language_def by auto lemma language_alphabet: "language A \ streams (alphabet A)" by (auto dest: run_alphabet) lemma restrict_language[simp]: "language (restrict A) = language A" by force end locale automaton_degeneralization = a: automaton automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 + b: automaton automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 for automaton\<^sub>1 :: "'label set \ 'state set \ ('label \ 'state \ 'state set) \ 'item pred gen \ 'automaton\<^sub>1" and alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 and automaton\<^sub>2 :: "'label set \ 'state degen set \ ('label \ 'state degen \ 'state degen set) \ 'item_degen pred \ 'automaton\<^sub>2" and alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 + fixes item :: "'state \ 'label \ 'state \ 'item" fixes translate :: "'item_degen \ 'item degen" begin definition degeneralize :: "'automaton\<^sub>1 \ 'automaton\<^sub>2" where "degeneralize A \ automaton\<^sub>2 (alphabet\<^sub>1 A) (initial\<^sub>1 A \ {0}) (\ a (p, k). {(q, count (condition\<^sub>1 A) (item (p, a, q)) k) |q. q \ transition\<^sub>1 A a p}) (degen (condition\<^sub>1 A) \ translate)" lemma degeneralize_simps[simp]: "alphabet\<^sub>2 (degeneralize A) = alphabet\<^sub>1 A" "initial\<^sub>2 (degeneralize A) = initial\<^sub>1 A \ {0}" "transition\<^sub>2 (degeneralize A) a (p, k) = {(q, count (condition\<^sub>1 A) (item (p, a, q)) k) |q. q \ transition\<^sub>1 A a p}" "condition\<^sub>2 (degeneralize A) = degen (condition\<^sub>1 A) \ translate" unfolding degeneralize_def by auto lemma run_degeneralize: assumes "a.run A (w ||| r) p" shows "b.run (degeneralize A) (w ||| r ||| sscan (count (condition\<^sub>1 A) \ item) (p ## r ||| w ||| r) k) (p, k)" using assms by (coinduction arbitrary: w r p k) (force elim: a.run.cases) lemma degeneralize_run: assumes "b.run (degeneralize A) (w ||| rs) pk" obtains r s p k where "rs = r ||| s" "pk = (p, k)" "a.run A (w ||| r) p" "s = sscan (count (condition\<^sub>1 A) \ item) (p ## r ||| w ||| r) k" proof - obtain r s p k where 1: "rs = r ||| s" "pk = (p, k)" using szip_smap surjective_pairing by metis show ?thesis proof show "rs = r ||| s" "pk = (p, k)" using 1 by this show "a.run A (w ||| r) p" using assms unfolding 1 by (coinduction arbitrary: w r s p k) (force elim: b.run.cases) show "s = sscan (count (condition\<^sub>1 A) \ item) (p ## r ||| w ||| r) k" using assms unfolding 1 by (coinduction arbitrary: w r s p k) (erule b.run.cases, force) qed qed lemma degeneralize_nodes: "b.nodes (degeneralize A) \ a.nodes A \ insert 0 {0 ..< length (condition\<^sub>1 A)}" proof fix pk assume "pk \ b.nodes (degeneralize A)" then show "pk \ a.nodes A \ insert 0 {0 ..< length (condition\<^sub>1 A)}" by (induct) (force, cases "condition\<^sub>1 A = []", auto) qed lemma nodes_degeneralize: "a.nodes A \ fst ` b.nodes (degeneralize A)" proof fix p assume "p \ a.nodes A" then show "p \ fst ` b.nodes (degeneralize A)" proof induct case (initial p) have "(p, 0) \ b.nodes (degeneralize A)" using initial by auto then show ?case using image_iff fst_conv by force next case (execute p aq) obtain k where "(p, k) \ b.nodes (degeneralize A)" using execute(2) by auto then have "(snd aq, count (condition\<^sub>1 A) (item (p, aq)) k) \ b.nodes (degeneralize A)" using execute(3) by auto then show ?case using image_iff snd_conv by force qed qed lemma degeneralize_nodes_finite[iff]: "finite (b.nodes (degeneralize A)) \ finite (a.nodes A)" proof show "finite (a.nodes A)" if "finite (b.nodes (degeneralize A))" using finite_subset nodes_degeneralize that by blast show "finite (b.nodes (degeneralize A))" if "finite (a.nodes A)" using finite_subset degeneralize_nodes that by blast qed end locale automaton_degeneralization_run = automaton_degeneralization automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 item translate + a: automaton_run automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_run automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 for automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 and automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 and item translate + assumes test[iff]: "test\<^sub>2 (degen cs \ translate) w (r ||| sscan (count cs \ item) (p ## r ||| w ||| r) k) (p, k) \ test\<^sub>1 cs w r p" begin lemma degeneralize_language[simp]: "b.language (degeneralize A) = a.language A" unfolding a.language_def b.language_def by (auto dest: run_degeneralize elim!: degeneralize_run) end locale automaton_product = a: automaton automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 + b: automaton automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 + c: automaton automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 for automaton\<^sub>1 :: "'label set \ 'state\<^sub>1 set \ ('label \ 'state\<^sub>1 \ 'state\<^sub>1 set) \ 'condition\<^sub>1 \ 'automaton\<^sub>1" and alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 and automaton\<^sub>2 :: "'label set \ 'state\<^sub>2 set \ ('label \ 'state\<^sub>2 \ 'state\<^sub>2 set) \ 'condition\<^sub>2 \ 'automaton\<^sub>2" and alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 and automaton\<^sub>3 :: "'label set \ ('state\<^sub>1 \ 'state\<^sub>2) set \ ('label \ 'state\<^sub>1 \ 'state\<^sub>2 \ ('state\<^sub>1 \ 'state\<^sub>2) set) \ 'condition\<^sub>3 \ 'automaton\<^sub>3" and alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 + fixes condition :: "'condition\<^sub>1 \ 'condition\<^sub>2 \ 'condition\<^sub>3" begin definition product :: "'automaton\<^sub>1 \ 'automaton\<^sub>2 \ 'automaton\<^sub>3" where "product A B \ automaton\<^sub>3 (alphabet\<^sub>1 A \ alphabet\<^sub>2 B) (initial\<^sub>1 A \ initial\<^sub>2 B) (\ a (p, q). transition\<^sub>1 A a p \ transition\<^sub>2 B a q) (condition (condition\<^sub>1 A) (condition\<^sub>2 B))" lemma product_simps[simp]: "alphabet\<^sub>3 (product A B) = alphabet\<^sub>1 A \ alphabet\<^sub>2 B" "initial\<^sub>3 (product A B) = initial\<^sub>1 A \ initial\<^sub>2 B" "transition\<^sub>3 (product A B) a (p, q) = transition\<^sub>1 A a p \ transition\<^sub>2 B a q" "condition\<^sub>3 (product A B) = condition (condition\<^sub>1 A) (condition\<^sub>2 B)" unfolding product_def by auto lemma product_target[simp]: assumes "length w = length r" "length r = length s" shows "c.target (w || r || s) (p, q) = (a.target (w || r) p, b.target (w || s) q)" using assms by (induct arbitrary: p q rule: list_induct3) (auto) lemma product_path[iff]: assumes "length w = length r" "length r = length s" shows "c.path (product A B) (w || r || s) (p, q) \ a.path A (w || r) p \ b.path B (w || s) q" using assms by (induct arbitrary: p q rule: list_induct3) (auto) lemma product_run[iff]: "c.run (product A B) (w ||| r ||| s) (p, q) \ a.run A (w ||| r) p \ b.run B (w ||| s) q" proof safe show "a.run A (w ||| r) p" if "c.run (product A B) (w ||| r ||| s) (p, q)" using that by (coinduction arbitrary: w r s p q) (force elim: c.run.cases) show "b.run B (w ||| s) q" if "c.run (product A B) (w ||| r ||| s) (p, q)" using that by (coinduction arbitrary: w r s p q) (force elim: c.run.cases) show "c.run (product A B) (w ||| r ||| s) (p, q)" if "a.run A (w ||| r) p" "b.run B (w ||| s) q" using that by (coinduction arbitrary: w r s p q) (auto elim: a.run.cases b.run.cases) qed lemma product_nodes: "c.nodes (product A B) \ a.nodes A \ b.nodes B" proof fix pq assume "pq \ c.nodes (product A B)" then show "pq \ a.nodes A \ b.nodes B" by induct auto qed lemma product_nodes_finite[intro]: assumes "finite (a.nodes A)" "finite (b.nodes B)" shows "finite (c.nodes (product A B))" using finite_subset product_nodes assms by blast end locale automaton_intersection_path = automaton_product automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 condition + a: automaton_path automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_path automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 + c: automaton_path automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 test\<^sub>3 for automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 and automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 and automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 test\<^sub>3 and condition + assumes test[iff]: "length r = length w \ length s = length w \ test\<^sub>3 (condition c\<^sub>1 c\<^sub>2) w (r || s) (p, q) \ test\<^sub>1 c\<^sub>1 w r p \ test\<^sub>2 c\<^sub>2 w s q" begin lemma product_language[simp]: "c.language (product A B) = a.language A \ b.language B" unfolding a.language_def b.language_def c.language_def by (force iff: split_zip) end locale automaton_intersection_run = automaton_product automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 condition + a: automaton_run automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_run automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 + c: automaton_run automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 test\<^sub>3 for automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 and automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 and automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 test\<^sub>3 and condition + assumes test[iff]: "test\<^sub>3 (condition c\<^sub>1 c\<^sub>2) w (r ||| s) (p, q) \ test\<^sub>1 c\<^sub>1 w r p \ test\<^sub>2 c\<^sub>2 w s q" begin lemma product_language[simp]: "c.language (product A B) = a.language A \ b.language B" unfolding a.language_def b.language_def c.language_def by (fastforce iff: split_szip) end locale automaton_sum = a: automaton automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 + b: automaton automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 + c: automaton automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 for automaton\<^sub>1 :: "'label set \ 'state\<^sub>1 set \ ('label \ 'state\<^sub>1 \ 'state\<^sub>1 set) \ 'condition\<^sub>1 \ 'automaton\<^sub>1" and alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 and automaton\<^sub>2 :: "'label set \ 'state\<^sub>2 set \ ('label \ 'state\<^sub>2 \ 'state\<^sub>2 set) \ 'condition\<^sub>2 \ 'automaton\<^sub>2" and alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 and automaton\<^sub>3 :: "'label set \ ('state\<^sub>1 + 'state\<^sub>2) set \ ('label \ 'state\<^sub>1 + 'state\<^sub>2 \ ('state\<^sub>1 + 'state\<^sub>2) set) \ 'condition\<^sub>3 \ 'automaton\<^sub>3" and alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 + fixes condition :: "'condition\<^sub>1 \ 'condition\<^sub>2 \ 'condition\<^sub>3" begin definition sum :: "'automaton\<^sub>1 \ 'automaton\<^sub>2 \ 'automaton\<^sub>3" where "sum A B \ automaton\<^sub>3 (alphabet\<^sub>1 A \ alphabet\<^sub>2 B) (initial\<^sub>1 A <+> initial\<^sub>2 B) (\ a. \ Inl p \ Inl ` transition\<^sub>1 A a p | Inr q \ Inr ` transition\<^sub>2 B a q) (condition (condition\<^sub>1 A) (condition\<^sub>2 B))" lemma sum_simps[simp]: "alphabet\<^sub>3 (sum A B) = alphabet\<^sub>1 A \ alphabet\<^sub>2 B" "initial\<^sub>3 (sum A B) = initial\<^sub>1 A <+> initial\<^sub>2 B" "transition\<^sub>3 (sum A B) a (Inl p) = Inl ` transition\<^sub>1 A a p" "transition\<^sub>3 (sum A B) a (Inr q) = Inr ` transition\<^sub>2 B a q" "condition\<^sub>3 (sum A B) = condition (condition\<^sub>1 A) (condition\<^sub>2 B)" unfolding sum_def by auto lemma path_sum_a: assumes "length r = length w" "a.path A (w || r) p" shows "c.path (sum A B) (w || map Inl r) (Inl p)" using assms by (induct arbitrary: p rule: list_induct2) (auto) lemma path_sum_b: assumes "length s = length w" "b.path B (w || s) q" shows "c.path (sum A B) (w || map Inr s) (Inr q)" using assms by (induct arbitrary: q rule: list_induct2) (auto) lemma sum_path: assumes "alphabet\<^sub>1 A = alphabet\<^sub>2 B" assumes "length rs = length w" "c.path (sum A B) (w || rs) pq" obtains (a) r p where "rs = map Inl r" "pq = Inl p" "a.path A (w || r) p" | (b) s q where "rs = map Inr s" "pq = Inr q" "b.path B (w || s) q" proof (cases pq) case (Inl p) have 1: "rs = map Inl (map projl rs)" using assms(2, 3) unfolding Inl by (induct arbitrary: p rule: list_induct2) (auto) have 2: "a.path A (w || map projl rs) p" using assms(2, 1, 3) unfolding Inl by (induct arbitrary: p rule: list_induct2) (auto) show ?thesis using a 1 Inl 2 by this next case (Inr q) have 1: "rs = map Inr (map projr rs)" using assms(2, 3) unfolding Inr by (induct arbitrary: q rule: list_induct2) (auto) have 2: "b.path B (w || map projr rs) q" using assms(2, 1, 3) unfolding Inr by (induct arbitrary: q rule: list_induct2) (auto) show ?thesis using b 1 Inr 2 by this qed lemma run_sum_a: assumes "a.run A (w ||| r) p" shows "c.run (sum A B) (w ||| smap Inl r) (Inl p)" using assms by (coinduction arbitrary: w r p) (force elim: a.run.cases) lemma run_sum_b: assumes "b.run B (w ||| s) q" shows "c.run (sum A B) (w ||| smap Inr s) (Inr q)" using assms by (coinduction arbitrary: w s q) (force elim: b.run.cases) lemma sum_run: assumes "alphabet\<^sub>1 A = alphabet\<^sub>2 B" assumes "c.run (sum A B) (w ||| rs) pq" obtains (a) r p where "rs = smap Inl r" "pq = Inl p" "a.run A (w ||| r) p" | (b) s q where "rs = smap Inr s" "pq = Inr q" "b.run B (w ||| s) q" proof (cases pq) case (Inl p) have 1: "rs = smap Inl (smap projl rs)" using assms(2) unfolding Inl by (coinduction arbitrary: w rs p) (force elim: c.run.cases) have 2: "a.run A (w ||| smap projl rs) p" using assms unfolding Inl by (coinduction arbitrary: w rs p) (force elim: c.run.cases) show ?thesis using a 1 Inl 2 by this next case (Inr q) have 1: "rs = smap Inr (smap projr rs)" using assms(2) unfolding Inr by (coinduction arbitrary: w rs q) (force elim: c.run.cases) have 2: "b.run B (w ||| smap projr rs) q" using assms unfolding Inr by (coinduction arbitrary: w rs q) (force elim: c.run.cases) show ?thesis using b 1 Inr 2 by this qed lemma sum_nodes: assumes "alphabet\<^sub>1 A = alphabet\<^sub>2 B" shows "c.nodes (sum A B) \ a.nodes A <+> b.nodes B" proof fix pq assume "pq \ c.nodes (sum A B)" then show "pq \ a.nodes A <+> b.nodes B" using assms by (induct) (auto 0 3) qed lemma sum_nodes_finite[intro]: assumes "alphabet\<^sub>1 A = alphabet\<^sub>2 B" assumes "finite (a.nodes A)" "finite (b.nodes B)" shows "finite (c.nodes (sum A B))" using finite_subset sum_nodes assms by (auto intro: finite_Plus) end locale automaton_union_path = automaton_sum automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 condition + a: automaton_path automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_path automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 + c: automaton_path automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 test\<^sub>3 for automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 and automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 and automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 test\<^sub>3 and condition + assumes test\<^sub>1[iff]: "length r = length w \ test\<^sub>3 (condition c\<^sub>1 c\<^sub>2) w (map Inl r) (Inl p) \ test\<^sub>1 c\<^sub>1 w r p" assumes test\<^sub>2[iff]: "length s = length w \ test\<^sub>3 (condition c\<^sub>1 c\<^sub>2) w (map Inr s) (Inr q) \ test\<^sub>2 c\<^sub>2 w s q" begin lemma sum_language[simp]: assumes "alphabet\<^sub>1 A = alphabet\<^sub>2 B" shows "c.language (sum A B) = a.language A \ b.language B" using assms unfolding a.language_def b.language_def c.language_def by (force intro: path_sum_a path_sum_b elim!: sum_path) end locale automaton_union_run = automaton_sum automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 condition + a: automaton_run automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_run automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 + c: automaton_run automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 test\<^sub>3 for automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 and automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 and automaton\<^sub>3 alphabet\<^sub>3 initial\<^sub>3 transition\<^sub>3 condition\<^sub>3 test\<^sub>3 and condition + assumes test\<^sub>1[iff]: "test\<^sub>3 (condition c\<^sub>1 c\<^sub>2) w (smap Inl r) (Inl p) \ test\<^sub>1 c\<^sub>1 w r p" assumes test\<^sub>2[iff]: "test\<^sub>3 (condition c\<^sub>1 c\<^sub>2) w (smap Inr s) (Inr q) \ test\<^sub>2 c\<^sub>2 w s q" begin lemma sum_language[simp]: assumes "alphabet\<^sub>1 A = alphabet\<^sub>2 B" shows "c.language (sum A B) = a.language A \ b.language B" using assms unfolding a.language_def b.language_def c.language_def by (auto intro: run_sum_a run_sum_b elim!: sum_run) end locale automaton_product_list = a: automaton automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 + b: automaton automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 for automaton\<^sub>1 :: "'label set \ 'state set \ ('label \ 'state \ 'state set) \ 'condition\<^sub>1 \ 'automaton\<^sub>1" and alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 and automaton\<^sub>2 :: "'label set \ 'state list set \ ('label \ 'state list \ 'state list set) \ 'condition\<^sub>2 \ 'automaton\<^sub>2" and alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 + fixes condition :: "'condition\<^sub>1 list \ 'condition\<^sub>2" begin definition product :: "'automaton\<^sub>1 list \ 'automaton\<^sub>2" where "product AA \ automaton\<^sub>2 (\ (alphabet\<^sub>1 ` set AA)) (listset (map initial\<^sub>1 AA)) (\ a ps. listset (map2 (\ A p. transition\<^sub>1 A a p) AA ps)) (condition (map condition\<^sub>1 AA))" lemma product_simps[simp]: "alphabet\<^sub>2 (product AA) = \ (alphabet\<^sub>1 ` set AA)" "initial\<^sub>2 (product AA) = listset (map initial\<^sub>1 AA)" "transition\<^sub>2 (product AA) a ps = listset (map2 (\ A p. transition\<^sub>1 A a p) AA ps)" "condition\<^sub>2 (product AA) = condition (map condition\<^sub>1 AA)" unfolding product_def by auto lemma product_run_length: assumes "length ps = length AA" assumes "b.run (product AA) (w ||| r) ps" assumes "qs \ sset r" shows "length qs = length AA" proof - have "pred_stream (\ qs. length qs = length AA) r" using assms(1, 2) by (coinduction arbitrary: w r ps) (force elim: b.run.cases simp: listset_member list_all2_conv_all_nth) then show ?thesis using assms(3) unfolding stream.pred_set by auto qed lemma product_run_stranspose: assumes "length ps = length AA" assumes "b.run (product AA) (w ||| r) ps" obtains rs where "r = stranspose rs" "length rs = length AA" proof define rs where "rs \ map (\ k. smap (\ ps. ps ! k) r) [0 ..< length AA]" have "length qs = length AA" if "qs \ sset r" for qs using product_run_length assms that by this then show "r = stranspose rs" - unfolding rs_def by (coinduction arbitrary: r) (auto intro: nth_equalityI simp: comp_def) + unfolding rs_def by (coinduction arbitrary: r) (force intro: nth_equalityI simp: comp_def) show "length rs = length AA" unfolding rs_def by auto qed lemma run_product: assumes "length rs = length AA" "length ps = length AA" assumes "\ k. k < length AA \ a.run (AA ! k) (w ||| rs ! k) (ps ! k)" shows "b.run (product AA) (w ||| stranspose rs) ps" using assms proof (coinduction arbitrary: w rs ps) case (run ap r) then show ?case proof (intro conjI exI) show "fst ap \ alphabet\<^sub>2 (product AA)" using run by (force elim: a.run.cases simp: set_conv_nth) show "snd ap \ transition\<^sub>2 (product AA) (fst ap) ps" using run by (force elim: a.run.cases simp: listset_member list_all2_conv_all_nth) show "\ k < length AA. a.run' (AA ! k) (stl w ||| map stl rs ! k) (map shd rs ! k)" using run by (force elim: a.run.cases) qed auto qed lemma product_run: assumes "length rs = length AA" "length ps = length AA" assumes "b.run (product AA) (w ||| stranspose rs) ps" shows "k < length AA \ a.run (AA ! k) (w ||| rs ! k) (ps ! k)" using assms proof (coinduction arbitrary: w rs ps) case (run ap wr) then show ?case proof (intro exI conjI) show "fst ap \ alphabet\<^sub>1 (AA ! k)" using run by (force elim: b.run.cases) show "snd ap \ transition\<^sub>1 (AA ! k) (fst ap) (ps ! k)" using run by (force elim: b.run.cases simp: listset_member list_all2_conv_all_nth) show "b.run' (product AA) (stl w ||| stranspose (map stl rs)) (shd (stranspose rs))" using run by (force elim: b.run.cases) qed auto qed lemma product_nodes: "b.nodes (product AA) \ listset (map a.nodes AA)" proof show "ps \ listset (map a.nodes AA)" if "ps \ b.nodes (product AA)" for ps using that by (induct) (auto 0 3 simp: listset_member list_all2_conv_all_nth) qed lemma product_nodes_finite[intro]: assumes "list_all (finite \ a.nodes) AA" shows "finite (b.nodes (product AA))" using list.pred_map product_nodes assms by (blast dest: finite_subset) lemma product_nodes_card: assumes "list_all (finite \ a.nodes) AA" shows "card (b.nodes (product AA)) \ prod_list (map (card \ a.nodes) AA)" proof - have "card (b.nodes (product AA)) \ card (listset (map a.nodes AA))" using list.pred_map product_nodes assms by (blast intro: card_mono) also have "\ = prod_list (map (card \ a.nodes) AA)" by simp finally show ?thesis by this qed end locale automaton_intersection_list_run = automaton_product_list automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 condition + a: automaton_run automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_run automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 for automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 and automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 and condition + assumes test[iff]: "length rs = length cs \ length ps = length cs \ test\<^sub>2 (condition cs) w (stranspose rs) ps \ list_all (\ (c, r, p). test\<^sub>1 c w r p) (cs || rs || ps)" begin lemma product_language[simp]: "b.language (product AA) = \ (a.language ` set AA)" proof safe fix A w assume 1: "w \ b.language (product AA)" "A \ set AA" obtain r ps where 2: "ps \ initial\<^sub>2 (product AA)" "b.run (product AA) (w ||| r) ps" "test\<^sub>2 (condition\<^sub>2 (product AA)) w r ps" using 1(1) by auto have 3: "length ps = length AA" using 2(1) by (simp add: listset_member list_all2_conv_all_nth) obtain rs where 4: "r = stranspose rs" "length rs = length AA" using product_run_stranspose 3 2(2) by this obtain k where 5: "k < length AA" "A = AA ! k" using 1(2) unfolding set_conv_nth by auto show "w \ a.language A" proof show "ps ! k \ initial\<^sub>1 A" using 2(1) 5 by (auto simp: listset_member list_all2_conv_all_nth) show "a.run A (w ||| rs ! k) (ps ! k)" using 2(2) 3 4 5 by (auto intro: product_run) show "test\<^sub>1 (condition\<^sub>1 A) w (rs ! k) (ps ! k)" using 2(3) 3 4 5 by (simp add: list_all_length) qed next fix w assume 1: "w \ \ (a.language ` set AA)" have 2: "\ A \ set AA. \ r p. p \ initial\<^sub>1 A \ a.run A (w ||| r) p \ test\<^sub>1 (condition\<^sub>1 A) w r p" using 1 by blast obtain rs ps where 3: "length rs = length AA" "length ps = length AA" "\ k. k < length AA \ ps ! k \ initial\<^sub>1 (AA ! k)" "\ k. k < length AA \ a.run (AA ! k) (w ||| rs ! k) (ps ! k)" "\ k. k < length AA \ test\<^sub>1 (condition\<^sub>1 (AA ! k)) w (rs ! k) (ps ! k)" using 2 unfolding Ball_set list_choice_zip list_choice_pair unfolding list.pred_set set_conv_nth by force show "w \ b.language (product AA)" proof show "ps \ initial\<^sub>2 (product AA)" using 3 by (auto simp: listset_member list_all2_conv_all_nth) show "b.run (product AA) (w ||| stranspose rs) ps" using 3 by (auto intro: run_product) show "test\<^sub>2 (condition\<^sub>2 (product AA)) w (stranspose rs) ps" using 3 by (auto simp: list_all_length) qed qed end locale automaton_sum_list = a: automaton automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 + b: automaton automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 for automaton\<^sub>1 :: "'label set \ 'state set \ ('label \ 'state \ 'state set) \ 'condition\<^sub>1 \ 'automaton\<^sub>1" and alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 and automaton\<^sub>2 :: "'label set \ (nat \ 'state) set \ ('label \ nat \ 'state \ (nat \ 'state) set) \ 'condition\<^sub>2 \ 'automaton\<^sub>2" and alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 + fixes condition :: "'condition\<^sub>1 list \ 'condition\<^sub>2" begin definition sum :: "'automaton\<^sub>1 list \ 'automaton\<^sub>2" where "sum AA \ automaton\<^sub>2 (\ (alphabet\<^sub>1 ` set AA)) (\ k < length AA. {k} \ initial\<^sub>1 (AA ! k)) (\ a (k, p). {k} \ transition\<^sub>1 (AA ! k) a p) (condition (map condition\<^sub>1 AA))" lemma sum_simps[simp]: "alphabet\<^sub>2 (sum AA) = \ (alphabet\<^sub>1 ` set AA)" "initial\<^sub>2 (sum AA) = (\ k < length AA. {k} \ initial\<^sub>1 (AA ! k))" "transition\<^sub>2 (sum AA) a (k, p) = {k} \ transition\<^sub>1 (AA ! k) a p" "condition\<^sub>2 (sum AA) = condition (map condition\<^sub>1 AA)" unfolding sum_def by auto lemma run_sum: assumes "\ (alphabet\<^sub>1 ` set AA) = \ (alphabet\<^sub>1 ` set AA)" assumes "A \ set AA" assumes "a.run A (w ||| s) p" obtains k where "k < length AA" "A = AA ! k" "b.run (sum AA) (w ||| sconst k ||| s) (k, p)" proof - obtain k where 1: "k < length AA" "A = AA ! k" using assms(2) unfolding set_conv_nth by auto show ?thesis proof show "k < length AA" "A = AA ! k" using 1 by this show "b.run (sum AA) (w ||| sconst k ||| s) (k, p)" using assms 1(2) by (coinduction arbitrary: w s p) (force elim: a.run.cases) qed qed lemma sum_run: assumes "\ (alphabet\<^sub>1 ` set AA) = \ (alphabet\<^sub>1 ` set AA)" assumes "k < length AA" assumes "b.run (sum AA) (w ||| r) (k, p)" obtains s where "r = sconst k ||| s" "a.run (AA ! k) (w ||| s) p" proof show "r = sconst k ||| smap snd r" using assms by (coinduction arbitrary: w r p) (force elim: b.run.cases) show "a.run (AA ! k) (w ||| smap snd r) p" using assms by (coinduction arbitrary: w r p) (force elim: b.run.cases) qed lemma sum_nodes: assumes "\ (alphabet\<^sub>1 ` set AA) = \ (alphabet\<^sub>1 ` set AA)" shows "b.nodes (sum AA) \ (\ k < length AA. {k} \ a.nodes (AA ! k))" proof show "kp \ (\ k < length AA. {k} \ a.nodes (AA ! k))" if "kp \ b.nodes (sum AA)" for kp using that assms by (induct) (auto 0 4) qed lemma sum_nodes_finite[intro]: assumes "\ (alphabet\<^sub>1 ` set AA) = \ (alphabet\<^sub>1 ` set AA)" assumes "list_all (finite \ a.nodes) AA" shows "finite (b.nodes (sum AA))" proof (rule finite_subset) show "b.nodes (sum AA) \ (\ k < length AA. {k} \ a.nodes (AA ! k))" using sum_nodes assms(1) by this show "finite (\ k < length AA. {k} \ a.nodes' (AA ! k))" using assms(2) unfolding list_all_length by auto qed end locale automaton_union_list_run = automaton_sum_list automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 condition + a: automaton_run automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 + b: automaton_run automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 for automaton\<^sub>1 alphabet\<^sub>1 initial\<^sub>1 transition\<^sub>1 condition\<^sub>1 test\<^sub>1 and automaton\<^sub>2 alphabet\<^sub>2 initial\<^sub>2 transition\<^sub>2 condition\<^sub>2 test\<^sub>2 and condition + assumes test[iff]: "k < length cs \ test\<^sub>2 (condition cs) w (sconst k ||| r) (k, p) \ test\<^sub>1 (cs ! k) w r p" begin lemma sum_language[simp]: assumes "\ (alphabet\<^sub>1 ` set AA) = \ (alphabet\<^sub>1 ` set AA)" shows "b.language (sum AA) = \ (a.language ` set AA)" proof show "b.language (sum AA) \ \ (a.language ` set AA)" using assms unfolding a.language_def b.language_def by (force elim: sum_run) show "\ (a.language ` set AA) \ b.language (sum AA)" using assms unfolding a.language_def b.language_def by (force elim!: run_sum) qed end end \ No newline at end of file