Theory Soundness
subsection ‹Soundness of the Rules›
text ‹In this file, we prove that each rule of the logic is sound. We do this by assuming that the Hoare
triples in the premise of the rule hold semantically (as defined in Safety.thy), and then proving
that the Hoare triple in the conclusion also holds semantically. We prove soundness of the logic
(with some corollaries) at the end of the file.›
text ‹For each rule, we first prove an important lemma about the safety of the statement (i.e., under
which conditions is executing this statement safe, and what conditions will hold about the set of states
that can be reached by executing this statement). We then use this lemma to prove the rule of the logic,
by constructing the set of states that will be reached, proving that safety holds, and proving that the
final set of states satisfies the postcondition.›
theory Soundness
imports Safety AbstractCommutativity
begin
subsubsection Skip
lemma safe_skip:
fixes Δ :: "('i, 'a, nat) cont"
assumes "(s, h) ∈ S"
shows "safe n Δ Cskip (s, h) S"
using assms
proof (induct n)
case (Suc n)
then show ?case
proof (cases Δ)
case None
then show ?thesis
by (simp add: Suc.prems)
next
case (Some a)
then show ?thesis
by (simp add: assms)
qed
qed (simp)
theorem rule_skip:
"hoare_triple_valid Γ P Cskip P"
proof (rule hoare_triple_validI)
let ?Σ = "λσ. {σ}"
show "⋀s h n. (s, h), (s, h) ⊨ P ⟹ safe n Γ Cskip (s, h) (?Σ (s, h))"
by (simp add: safe_skip)
show "⋀s h s' h'. (s, h), (s', h') ⊨ P ⟹ pair_sat {(s, h)} {(s', h')} P"
by (metis pair_sat_smallerI singleton_iff)
qed
subsubsection ‹Assign›
inductive_cases red_assign_cases: "red (Cassign x E) σ C' σ'"
inductive_cases aborts_assign_cases: "aborts (Cassign x E) σ"
lemma safe_assign:
fixes Δ :: "('i, 'a, nat) cont"
assumes "⋀Γ. Δ = Some Γ ⟹ x ∉ fvA (invariant Γ)"
shows "safe m Δ (Cassign x E) (s, h) { (s(x := edenot E s), h) }"
proof (induct m)
case (Suc n)
show "safe (Suc n) Δ (Cassign x E) (s, h) {(s(x := edenot E s), h)}"
proof (rule safeI)
show "no_abort Δ (Cassign x E) s h"
using aborts_assign_cases no_abortI by blast
show "⋀H hf C' s' h'.
Δ = None ⟹
Some H = Some h ⊕ Some hf ∧ full_ownership (get_fh H) ∧ no_guard H ∧ red (Cassign x E) (s, FractionalHeap.normalize (get_fh H)) C' (s', h') ⟹
∃h'' H'.
full_ownership (get_fh H') ∧
no_guard H' ∧ h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hf ∧ safe n None C' (s', h'') {(s(x := edenot E s), h)}"
by (metis Pair_inject insertI1 red_assign_cases safe_skip)
show "accesses (Cassign x E) s ⊆ dom (fst h) ∧ writes (Cassign x E) s ⊆ fpdom (fst h)"
by simp
fix H hf C' s' h' hj v0 Γ
assume asm0: "Δ = Some Γ" "Some H = Some h ⊕ Some hj ⊕ Some hf ∧
full_ownership (get_fh H) ∧ semi_consistent Γ v0 H ∧ sat_inv s hj Γ ∧ red (Cassign x E) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
then have "sat_inv (s(x := edenot E s)) hj Γ"
by (meson agrees_update assms sat_inv_agrees)
then show "∃h'' H' hj'. full_ownership (get_fh H') ∧ semi_consistent Γ v0 H' ∧ sat_inv s' hj' Γ ∧
h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hj' ⊕ Some hf ∧ safe n (Some Γ) C' (s', h'') {(s(x := edenot E s), h)}"
by (metis (no_types, lifting) asm0(2) insertI1 old.prod.inject red_assign_cases safe_skip)
qed (simp)
qed (simp)
theorem assign_rule:
fixes Δ :: "('i, 'a, nat) cont"
assumes "⋀Γ. Δ = Some Γ ⟹ x ∉ fvA (invariant Γ)"
and "collect_existentials P ∩ fvE E = {}"
shows "hoare_triple_valid Δ (subA x E P) (Cassign x E) P"
proof -
define Σ :: "store × ('i, 'a) heap ⇒ (store × ('i, 'a) heap) set " where "Σ = (λσ. { ((fst σ)(x := edenot E (fst σ)), snd σ) })"
show ?thesis
proof (rule hoare_triple_validI)
show "⋀s h n. (s, h), (s, h) ⊨ subA x E P ⟹ safe n Δ (Cassign x E) (s, h) (Σ (s, h))"
using assms safe_assign by (metis Σ_def fst_eqD snd_eqD)
show "⋀s h s' h'. (s, h), (s', h') ⊨ subA x E P ⟹ pair_sat (Σ (s, h)) (Σ (s', h')) P"
by (metis assms(2) Σ_def fst_conv pair_sat_smallerI singleton_iff snd_conv subA_assign)
qed
qed
subsubsection ‹Alloc›
inductive_cases red_alloc_cases: "red (Calloc x E) σ C' σ'"
inductive_cases aborts_alloc_cases: "aborts (Calloc x E) σ"
lemma safe_new_None:
"safe n (None :: ('i, 'a, nat) cont) (Calloc x E) (s, (Map.empty, gs, gu)) { (s(x := a), (Map.empty(a ↦ (pwrite, edenot E s)), gs, gu)) |a. True }"
proof (induct n)
case (Suc n)
show ?case
proof (rule safeNoneI)
show "Calloc x E = Cskip ⟹ (s, Map.empty, gs, gu) ∈ {(s(x := a), [a ↦ (pwrite, edenot E s)], gs, gu) |a. True}" by simp
show "no_abort None (Calloc x E) s (Map.empty, gs, gu)"
using aborts_alloc_cases no_abort.simps(1) by blast
fix H hf C' s' h'
assume asm0: "Some H = Some (Map.empty, gs, gu) ⊕ Some hf ∧
full_ownership (get_fh H) ∧ no_guard H ∧ red (Calloc x E) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
show "∃h'' H'.
full_ownership (get_fh H') ∧
no_guard H' ∧
h' = FractionalHeap.normalize (get_fh H') ∧
Some H' = Some h'' ⊕ Some hf ∧ safe n (None :: ('i, 'a, nat) cont) C' (s', h'') {(s(x := a), [a ↦ (pwrite, edenot E s)], gs, gu) |a. True}"
proof (rule red_alloc_cases)
show "red (Calloc x E) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
using asm0 by blast
fix sa h v
assume asm1: "(s, FractionalHeap.normalize (get_fh H)) = (sa, h)" "C' = Cskip" "(s', h') = (sa(x := v), h(v ↦ edenot E sa))"
"v ∉ dom h"
then have "v ∉ dom (get_fh H)"
by (simp add: dom_normalize)
then have "v ∉ dom (get_fh hf)"
by (metis asm0 fst_conv get_fh.simps no_guard_and_no_heap no_guard_then_smaller_same no_guards_remove plus_comm)
moreover have "(Map.empty(v ↦ (pwrite, edenot E sa)), gs, gu) ## hf"
proof (rule compatibleI)
show "compatible_fract_heaps (get_fh ([v ↦ (pwrite, edenot E sa)], gs, gu)) (get_fh hf)"
proof (rule compatible_fract_heapsI)
fix l p p'
assume asm0: "get_fh ([v ↦ (pwrite, edenot E sa)], gs, gu) l = Some p ∧ get_fh hf l = Some p'"
then show "pgte pwrite (padd (fst p) (fst p'))"
by (metis calculation domIff fst_conv fun_upd_other get_fh.elims option.distinct(1))
show "snd p = snd p'"
by (metis asm0 calculation domIff fst_conv fun_upd_other get_fh.elims option.distinct(1))
qed
show "⋀k. get_gu ([v ↦ (pwrite, edenot E sa)], gs, gu) k = None ∨ get_gu hf k = None"
by (metis asm0 compatible_def compatible_eq get_gu.simps option.discI snd_conv)
show "⋀p p'. get_gs ([v ↦ (pwrite, edenot E sa)], gs, gu) = Some p ∧ get_gs hf = Some p' ⟹ pgte pwrite (padd (fst p) (fst p'))"
by (metis asm0 no_guard_def no_guard_then_smaller_same option.simps(3) plus_comm)
qed
then obtain H' where "Some H' = Some (Map.empty(v ↦ (pwrite, edenot E sa)), gs, gu) ⊕ Some hf"
by auto
moreover have "(s', (Map.empty(v ↦ (pwrite, edenot E sa)), gs, gu)) ∈ {(s(x := a), [a ↦ (pwrite, edenot E s)], gs, gu) |a. True}"
using asm1(1) asm1(3) by blast
then have "safe n (None :: ('i, 'a, nat) cont) C' (s', (Map.empty(v ↦ (pwrite, edenot E sa)), gs, gu)) {(s(x := a), [a ↦ (pwrite, edenot E s)], gs, gu) |a. True}"
by (simp add: asm1(2) safe_skip)
moreover have "full_ownership (get_fh H') ∧ no_guard H' ∧ h' = FractionalHeap.normalize (get_fh H')"
proof -
have "full_ownership (get_fh H')"
proof (rule full_ownershipI)
fix l p
assume "get_fh H' l = Some p"
show "fst p = pwrite"
proof (cases "l = v")
case True
then have "get_fh hf l = None"
using calculation(1) by blast
then have "get_fh H' l = (Map.empty(v ↦ (pwrite, edenot E sa))) l"
by (metis calculation(2) fst_conv get_fh.simps sum_second_none_get_fh)
then show ?thesis
using True ‹get_fh H' l = Some p› by fastforce
next
case False
then have "get_fh ([v ↦ (pwrite, edenot E sa)], gs, gu) l = None"
by simp
then show "fst p = pwrite"
by (metis (mono_tags, lifting) ‹get_fh H' l = Some p› asm0 calculation(2) fst_conv full_ownership_def get_fh.elims sum_first_none_get_fh)
qed
qed
moreover have "no_guard H'"
proof -
have "no_guard hf"
by (metis asm0 no_guard_then_smaller_same plus_comm)
moreover have "no_guard (Map.empty, gs, gu)"
using asm0 no_guard_then_smaller_same by blast
ultimately show ?thesis
by (metis ‹Some H' = Some ([v ↦ (pwrite, edenot E sa)], gs, gu) ⊕ Some hf› decompose_heap_triple no_guard_remove(1) no_guard_remove(2) no_guards_remove remove_guards_def snd_conv)
qed
moreover have "h' = FractionalHeap.normalize (get_fh H')"
proof (rule ext)
fix l show "h' l = FractionalHeap.normalize (get_fh H') l"
proof (cases "l = v")
case True
then have "get_fh (Map.empty(v ↦ (pwrite, edenot E sa)), gs, gu) l = Some (pwrite, edenot E sa)"
by auto
then have "get_fh hf l = None"
using True ‹v ∉ dom (get_fh hf)› by force
then show "h' l = FractionalHeap.normalize (get_fh H') l"
apply (cases "h' l")
using True asm1(3) apply auto[1]
by (metis (no_types, lifting) FractionalHeap.normalize_def True ‹Some H' = Some ([v ↦ (pwrite, edenot E sa)], gs, gu) ⊕ Some hf› ‹get_fh ([v ↦ (pwrite, edenot E sa)], gs, gu) l = Some (pwrite, edenot E sa)› apply_opt.simps(2) asm1(3) fun_upd_same snd_conv sum_second_none_get_fh)
next
case False
then have "get_fh (Map.empty(v ↦ (pwrite, edenot E sa)), gs, gu) l = None"
by simp
then have "get_fh H' l = get_fh hf l"
using ‹Some H' = Some ([v ↦ (pwrite, edenot E sa)], gs, gu) ⊕ Some hf› sum_first_none_get_fh by blast
moreover have "get_fh H l = get_fh hf l"
by (metis asm0 fst_conv get_fh.elims plus_comm sum_second_none_get_fh)
ultimately show ?thesis
proof (cases "get_fh hf l")
case None
then show ?thesis
by (metis False FractionalHeap.normalize_eq(1) ‹get_fh H l = get_fh hf l› ‹get_fh H' l = get_fh hf l› asm1(1) asm1(3) fun_upd_apply old.prod.inject)
next
case (Some f)
then show ?thesis
by (metis (no_types, lifting) False FractionalHeap.normalize_eq(1) FractionalHeap.normalize_eq(2) ‹get_fh H l = get_fh hf l› ‹get_fh H' l = get_fh hf l› asm1(1) asm1(3) domD not_in_dom fun_upd_apply old.prod.inject)
qed
qed
qed
ultimately show ?thesis
by auto
qed
ultimately show "∃h'' H'. full_ownership (get_fh H') ∧ no_guard H' ∧
h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hf ∧ safe n (None :: ('i, 'a, nat) cont) C' (s', h'') {(s(x := a), [a ↦ (pwrite, edenot E s)], gs, gu) |a. True}"
by blast
qed
qed (simp)
qed (simp)
lemma safe_new_Some:
assumes "x ∉ fvA (invariant Γ)"
and "view_function_of_inv Γ"
shows "safe n (Some Γ) (Calloc x E) (s, (Map.empty, gs, gu)) { (s(x := a), (Map.empty(a ↦ (pwrite, edenot E s)), gs, gu)) |a. True }"
proof (induct n)
case (Suc n)
show ?case
proof (rule safeSomeI)
show "Calloc x E = Cskip ⟹ (s, Map.empty, gs, gu) ∈ {(s(x := a), [a ↦ (pwrite, edenot E s)], gs, gu) |a. True}" by simp
show "no_abort (Some Γ) (Calloc x E) s (Map.empty, gs, gu)"
using aborts_alloc_cases no_abort.simps(2) by blast
fix H hf C' s' h' hj v0
assume asm0: "Some H = Some (Map.empty, gs, gu) ⊕ Some hj ⊕ Some hf ∧
full_ownership (get_fh H) ∧ semi_consistent Γ v0 H ∧ sat_inv s hj Γ ∧ red (Calloc x E) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
then obtain hjf where "Some hjf = Some hj ⊕ Some hf"
by (metis plus.simps(2) plus.simps(3) plus_asso)
then have "Some H = Some (Map.empty, gs, gu) ⊕ Some hjf"
by (metis asm0 plus_asso)
show "∃h'' H' hj'.
full_ownership (get_fh H') ∧
semi_consistent Γ v0 H' ∧
sat_inv s' hj' Γ ∧
h' = FractionalHeap.normalize (get_fh H') ∧
Some H' = Some h'' ⊕ Some hj' ⊕ Some hf ∧ safe n (Some Γ) C' (s', h'') {(s(x := a), [a ↦ (pwrite, edenot E s)], gs, gu) |a. True}"
proof (rule red_alloc_cases)
show "red (Calloc x E) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
using asm0 by blast
fix sa h v
assume asm1: "(s, FractionalHeap.normalize (get_fh H)) = (sa, h)" "C' = Cskip" "(s', h') = (sa(x := v), h(v ↦ edenot E sa))"
"v ∉ dom h"
then have "v ∉ dom (get_fh H)"
by (simp add: dom_normalize)
then have "v ∉ dom (get_fh hjf)"
by (metis (no_types, lifting) ‹Some H = Some (Map.empty, gs, gu) ⊕ Some hjf› addition_smaller_domain in_mono plus_comm)
moreover have "(Map.empty(v ↦ (pwrite, edenot E sa)), gs, gu) ## hjf"
proof (rule compatibleI)
show "compatible_fract_heaps (get_fh ([v ↦ (pwrite, edenot E sa)], gs, gu)) (get_fh hjf)"
proof (rule compatible_fract_heapsI)
fix l p p'
assume asm2: "get_fh ([v ↦ (pwrite, edenot E sa)], gs, gu) l = Some p ∧ get_fh hjf l = Some p'"
then show "pgte pwrite (padd (fst p) (fst p'))"
by (metis calculation domIff fst_conv fun_upd_other get_fh.elims option.distinct(1))
show "snd p = snd p'"
using asm2 calculation domIff fst_conv fun_upd_other get_fh.elims option.distinct(1) by metis
qed
show "⋀k. get_gu ([v ↦ (pwrite, edenot E sa)], gs, gu) k = None ∨ get_gu hjf k = None"
by (metis ‹Some H = Some (Map.empty, gs, gu) ⊕ Some hjf› compatible_def compatible_eq get_gu.simps option.discI snd_conv)
show "⋀p p'. get_gs ([v ↦ (pwrite, edenot E sa)], gs, gu) = Some p ∧ get_gs hjf = Some p' ⟹ pgte pwrite (padd (fst p) (fst p'))"
by (metis ‹Some H = Some (Map.empty, gs, gu) ⊕ Some hjf› compatible_def compatible_eq get_gs.simps option.simps(3) snd_eqD)
qed
then obtain H' where "Some H' = Some (Map.empty(v ↦ (pwrite, edenot E sa)), gs, gu) ⊕ Some hjf"
by auto
moreover have "(s', (Map.empty(v ↦ (pwrite, edenot E sa)), gs, gu)) ∈ {(s(x := a), [a ↦ (pwrite, edenot E s)], gs, gu) |a. True}"
using asm1(1) asm1(3) by blast
then have "safe n (Some Γ) C' (s', (Map.empty(v ↦ (pwrite, edenot E sa)), gs, gu)) {(s(x := a), [a ↦ (pwrite, edenot E s)], gs, gu) |a. True}"
by (simp add: asm1(2) safe_skip)
moreover have "full_ownership (get_fh H') ∧ semi_consistent Γ v0 H' ∧ h' = FractionalHeap.normalize (get_fh H')"
proof -
have "full_ownership (get_fh H')"
proof (rule full_ownershipI)
fix l p
assume "get_fh H' l = Some p"
show "fst p = pwrite"
proof (cases "l = v")
case True
then have "get_fh hjf l = None"
using calculation(1) by blast
then have "get_fh H' l = (Map.empty(v ↦ (pwrite, edenot E sa))) l"
by (metis calculation(2) fst_conv get_fh.simps sum_second_none_get_fh)
then show ?thesis
using True ‹get_fh H' l = Some p› by fastforce
next
case False
then have "get_fh H' l = get_fh hjf l" using sum_first_none_get_fh[of H' _ hjf l]
using calculation(2) by force
then show ?thesis
by (metis (no_types, lifting) ‹Some H = Some (Map.empty, gs, gu) ⊕ Some hjf› ‹get_fh H' l = Some p› asm0 fst_conv full_ownership_def get_fh.elims plus_comm sum_second_none_get_fh)
qed
qed
moreover have "h' = FractionalHeap.normalize (get_fh H')"
proof (rule ext)
fix l show "h' l = FractionalHeap.normalize (get_fh H') l"
proof (cases "l = v")
case True
then have "get_fh (Map.empty(v ↦ (pwrite, edenot E sa)), gs, gu) l = Some (pwrite, edenot E sa)"
by auto
then have "get_fh hjf l = None"
using True ‹v ∉ dom (get_fh hjf)› by force
then show ?thesis
apply (cases "h' l")
using True asm1(3) apply auto[1]
by (metis (no_types, lifting) FractionalHeap.normalize_def True ‹Some H' = Some ([v ↦ (pwrite, edenot E sa)], gs, gu) ⊕ Some hjf› ‹get_fh ([v ↦ (pwrite, edenot E sa)], gs, gu) l = Some (pwrite, edenot E sa)› apply_opt.simps(2) asm1(3) fun_upd_same snd_conv sum_second_none_get_fh)
next
case False
then have "get_fh (Map.empty(v ↦ (pwrite, edenot E sa)), gs, gu) l = None"
by simp
then have "get_fh H' l = get_fh hjf l"
using ‹Some H' = Some ([v ↦ (pwrite, edenot E sa)], gs, gu) ⊕ Some hjf› sum_first_none_get_fh by blast
moreover have "get_fh H l = get_fh hjf l"
by (metis ‹Some H = Some (Map.empty, gs, gu) ⊕ Some hjf› fst_eqD get_fh.simps sum_first_none_get_fh)
ultimately show ?thesis
proof (cases "get_fh hjf l")
case None
then show ?thesis
by (metis False FractionalHeap.normalize_eq(1) ‹get_fh H l = get_fh hjf l› ‹get_fh H' l = get_fh hjf l› asm1(1) asm1(3) fun_upd_apply old.prod.inject)
next
case (Some f)
then show ?thesis
by (metis (no_types, lifting) False FractionalHeap.normalize_eq(1) FractionalHeap.normalize_eq(2) ‹get_fh H l = get_fh hjf l› ‹get_fh H' l = get_fh hjf l› asm1(1) asm1(3) domD not_in_dom fun_upd_apply old.prod.inject)
qed
qed
qed
moreover have "semi_consistent Γ v0 H'"
proof (rule semi_consistentI)
have "get_gs H' = get_gs H"
by (metis ‹Some H = Some (Map.empty, gs, gu) ⊕ Some hjf› ‹Some H' = Some ([v ↦ (pwrite, edenot E sa)], gs, gu) ⊕ Some hjf› fst_conv get_gs.simps option.discI option.sel plus.simps(3) snd_conv)
moreover have "get_gu H' = get_gu H"
by (metis ‹Some H = Some (Map.empty, gs, gu) ⊕ Some hjf› ‹Some H' = Some ([v ↦ (pwrite, edenot E sa)], gs, gu) ⊕ Some hjf› get_gu.simps option.discI option.sel plus.simps(3) snd_conv)
ultimately show "all_guards H'"
by (metis all_guards_def asm0 semi_consistent_def)
show "reachable Γ v0 H'"
proof (rule reachableI)
fix sargs uargs
assume "get_gs H' = Some (pwrite, sargs) ∧ (∀k. get_gu H' k = Some (uargs k))"
then have "reachable_value (saction Γ) (uaction Γ) v0 sargs uargs (view Γ (FractionalHeap.normalize (get_fh H)))"
by (metis ‹get_gs H' = get_gs H› ‹get_gu H' = get_gu H› asm0 reachableE semi_consistent_def)
moreover have "view Γ (FractionalHeap.normalize (get_fh H)) = view Γ (FractionalHeap.normalize (get_fh H'))"
proof -
have "view Γ (FractionalHeap.normalize (get_fh H)) = view Γ (FractionalHeap.normalize (get_fh hj))"
using view_function_of_invE[of Γ s hj H] by (simp add: asm0 assms(2) larger3)
moreover have "view Γ (FractionalHeap.normalize (get_fh H')) = view Γ (FractionalHeap.normalize (get_fh hj))"
using view_function_of_invE[of Γ s hj H']
by (metis ‹Some H' = Some ([v ↦ (pwrite, edenot E sa)], gs, gu) ⊕ Some hjf› ‹Some hjf = Some hj ⊕ Some hf› asm0 assms(2) larger3 plus_comm)
ultimately show ?thesis by simp
qed
ultimately show "reachable_value (saction Γ) (uaction Γ) v0 sargs uargs (view Γ (FractionalHeap.normalize (get_fh H')))"
by simp
qed
qed
ultimately show ?thesis
by auto
qed
moreover have "sat_inv s' hj Γ"
proof (rule sat_invI)
show "no_guard hj"
using asm0 sat_inv_def by blast
have "agrees (fvA (invariant Γ)) s s'"
using asm1(1) asm1(3) assms
by (simp add: agrees_update)
then show "(s', hj), (s', hj) ⊨ invariant Γ"
using asm0 sat_inv_agrees sat_inv_def by blast
qed
ultimately show "∃h'' H' hj'. full_ownership (get_fh H') ∧ semi_consistent Γ v0 H' ∧ sat_inv s' hj' Γ ∧ h' = FractionalHeap.normalize (get_fh H') ∧
Some H' = Some h'' ⊕ Some hj' ⊕ Some hf ∧ safe n (Some Γ) C' (s', h'') {(s(x := a), [a ↦ (pwrite, edenot E s)], gs, gu) |a. True}"
by (metis (no_types, lifting) ‹Some hjf = Some hj ⊕ Some hf› plus_asso)
qed
qed (simp)
qed (simp)
lemma safe_new:
fixes Δ :: "('i, 'a, nat) cont"
assumes "⋀Γ. Δ = Some Γ ⟹ x ∉ fvA (invariant Γ) ∧ view_function_of_inv Γ"
shows "safe n Δ (Calloc x E) (s, (Map.empty, gs, gu)) { (s(x := a), (Map.empty(a ↦ (pwrite, edenot E s)), gs, gu)) |a. True }"
apply (cases Δ)
using safe_new_None safe_new_Some assms by blast+
theorem new_rule:
fixes Δ :: "('i, 'a, nat) cont"
assumes "x ∉ fvE E"
and "⋀Γ. Δ = Some Γ ⟹ x ∉ fvA (invariant Γ) ∧ view_function_of_inv Γ"
shows "hoare_triple_valid Δ Emp (Calloc x E) (PointsTo (Evar x) pwrite E)"
proof (rule hoare_triple_validI)
define Σ :: "store × ('i, 'a) heap ⇒ (store × ('i, 'a) heap) set" where "Σ = (λ(s, h). { (s(x := a), (Map.empty(a ↦ (pwrite, edenot E s)), get_gs h, get_gu h)) |a. True })"
show "⋀s h n. (s, h), (s, h) ⊨ Emp ⟹ safe n Δ (Calloc x E) (s, h) (Σ (s, h))"
proof -
fix s h n assume "(s, h :: ('i, 'a) heap), (s, h) ⊨ Emp" then have "get_fh h = Map.empty"
by simp
then have "h = (Map.empty, get_gs h, get_gu h)" using decompose_heap_triple
by metis
moreover have "safe n Δ (Calloc x E) (s, Map.empty, get_gs h, get_gu h) {(s(x := a), [a ↦ (pwrite, edenot E s)], get_gs h, get_gu h) |a. True}"
using safe_new assms(2) by blast
moreover have "Σ (s, h) = { (s(x := a), (Map.empty(a ↦ (pwrite, edenot E s)), get_gs h, get_gu h)) |a. True }"
using Σ_def by force
ultimately show "safe n Δ (Calloc x E) (s, h) (Σ (s, h))"
by presburger
qed
fix s1 h1 s2 h2
assume "(s1, h1 :: ('i, 'a) heap), (s2, h2) ⊨ Emp"
show "pair_sat (case (s1, h1) of (s, h) ⇒ {(s(x := a), [a ↦ (pwrite, edenot E s)], get_gs h, get_gu h) |a. True})
(case (s2, h2) of (s, h) ⇒ {(s(x := a), [a ↦ (pwrite, edenot E s)], get_gs h, get_gu h) |a. True}) (PointsTo (Evar x) pwrite E)"
proof (rule pair_satI)
fix s1' h1' s2' h2'
assume asm0: "(s1', h1') ∈ (case (s1, h1) of (s, h) ⇒ {(s(x := a), [a ↦ (pwrite, edenot E s)], get_gs h, get_gu h) |a. True}) ∧
(s2', h2') ∈ (case (s2, h2) of (s, h) ⇒ {(s(x := a), [a ↦ (pwrite, edenot E s)], get_gs h, get_gu h) |a. True})"
then obtain a1 a2 where "s1' = s1(x := a1)" "s2' = s2(x := a2)" "h1' = ([a1 ↦ (pwrite, edenot E s1)], get_gs h1, get_gu h1)"
"h2' = ([a2 ↦ (pwrite, edenot E s2)], get_gs h2, get_gu h2)"
by blast
then show "(s1', h1'), (s2', h2') ⊨ PointsTo (Evar x) pwrite E"
by (simp add: assms(1))
qed
qed
subsubsection ‹Write›
inductive_cases red_write_cases: "red (Cwrite x E) σ C' σ'"
inductive_cases aborts_write_cases: "aborts (Cwrite x E) σ"
lemma safe_write_None:
assumes "fh (edenot loc s) = Some (pwrite, v)"
shows "safe n (None :: ('i, 'a, nat) cont) (Cwrite loc E) (s, (fh, gs, gu)) { (s, (fh(edenot loc s ↦ (pwrite, edenot E s)), gs, gu)) }"
using assms
proof (induct n)
case (Suc n)
show ?case
proof (rule safeNoneI)
show "Cwrite loc E = Cskip ⟹ (s, fh, gs, gu) ∈ {(s, fh(edenot loc s ↦ (pwrite, edenot E s)), gs, gu)}"
by simp
show "no_abort None (Cwrite loc E) s (fh, gs, gu)"
proof (rule no_abortNoneI)
fix hf H assume asm0: "Some H = Some (fh, gs, gu) ⊕ Some hf ∧ full_ownership (get_fh H) ∧ no_guard H"
then have "edenot loc s ∈ dom (normalize (get_fh H))"
by (metis (mono_tags, lifting) Suc.prems addition_smaller_domain dom_def dom_normalize fst_conv get_fh.simps mem_Collect_eq option.discI subsetD)
then show "¬ aborts (Cwrite loc E) (s, normalize (get_fh H))"
by (metis aborts_write_cases fst_eqD snd_eqD)
qed
show "accesses (Cwrite loc E) s ⊆ dom (fst (fh, gs, gu)) ∧ writes (Cwrite loc E) s ⊆ fpdom (fst (fh, gs, gu))"
using assms fpdom_def by fastforce
fix H hf C' s' h'
assume asm0: "Some H = Some (fh, gs, gu) ⊕ Some hf ∧ full_ownership (get_fh H) ∧ no_guard H
∧ red (Cwrite loc E) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
then have "get_fh hf (edenot loc s) = None"
proof -
have "compatible_fract_heaps fh (get_fh hf)"
by (metis asm0 compatible_def compatible_eq fst_conv get_fh.elims option.discI)
then show ?thesis using compatible_then_dom_disjoint(2)[of fh "get_fh hf"]
assms disjoint_iff_not_equal[of "dom (get_fh hf)" "fpdom fh"] not_in_dom fpdom_def mem_Collect_eq
by fastforce
qed
show "∃h'' H'. full_ownership (get_fh H') ∧ no_guard H' ∧ h' = FractionalHeap.normalize (get_fh H')
∧ Some H' = Some h'' ⊕ Some hf ∧ safe n None C' (s', h'') {(s, fh(edenot loc s ↦ (pwrite, edenot E s)), gs, gu)}"
proof (rule red_write_cases)
show "red (Cwrite loc E) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
using asm0 by blast
fix sa h
assume asm1: "(s, FractionalHeap.normalize (get_fh H)) = (sa, h)" "C' = Cskip"
"(s', h') = (sa, h(edenot loc sa ↦ edenot E sa))"
then obtain "s = sa" "h' = h(edenot loc s ↦ edenot E s)" by blast
let ?h = "(fh(edenot loc s ↦ (pwrite, edenot E s)), gs, gu)"
have "?h ## hf"
proof (rule compatibleI)
show "compatible_fract_heaps (get_fh (fh(edenot loc s ↦ (pwrite, edenot E s)), gs, gu)) (get_fh hf)"
proof (rule compatible_fract_heapsI)
fix l p p' assume asm2: "get_fh (fh(edenot loc s ↦ (pwrite, edenot E s)), gs, gu) l = Some p ∧ get_fh hf l = Some p'"
then show "pgte pwrite (padd (fst p) (fst p'))"
apply (cases "l = edenot loc s")
apply (metis Suc.prems asm0 fst_conv fun_upd_same get_fh.elims option.sel plus_extract_point_fh)
by (metis asm0 fst_conv fun_upd_other get_fh.elims plus_extract_point_fh)
show "snd p = snd p'"
apply (cases "l = edenot loc s")
using ‹pgte pwrite (padd (fst p) (fst p'))› asm2 not_pgte_charact sum_larger apply fastforce
by (metis (mono_tags, opaque_lifting) asm0 asm2 fst_eqD get_fh.simps map_upd_Some_unfold plus_extract_point_fh)
qed
show "⋀k. get_gu (fh(edenot loc s ↦ (pwrite, edenot E s)), gs, gu) k = None ∨ get_gu hf k = None"
by (metis asm0 compatible_def compatible_eq get_gu.simps option.discI snd_conv)
show "⋀p p'. get_gs (fh(edenot loc s ↦ (pwrite, edenot E s)), gs, gu) = Some p ∧ get_gs hf = Some p' ⟹ pgte pwrite (padd (fst p) (fst p'))"
by (metis asm0 no_guard_def no_guard_then_smaller_same option.simps(3) plus_comm)
qed
then obtain H' where "Some H' = Some ?h ⊕ Some hf" by auto
moreover have "H' = ((get_fh H)(edenot loc s ↦ (pwrite, edenot E s)), get_gs H, get_gu H)"
proof (rule heap_ext)
show "get_fh H' = get_fh ((get_fh H)(edenot loc s ↦ (pwrite, edenot E s)), get_gs H, get_gu H)"
using calculation asm0 by (metis ‹get_fh hf (edenot loc s) = None› add_fh_update add_get_fh fst_conv get_fh.simps)
show "get_gs H' = get_gs ((get_fh H)(edenot loc s ↦ (pwrite, edenot E s)), get_gs H, get_gu H)"
using calculation asm0
by (metis fst_conv get_gs.simps plus_extract(2) snd_conv)
show "get_gu H' = get_gu ((get_fh H)(edenot loc s ↦ (pwrite, edenot E s)), get_gs H, get_gu H)"
using add_fh_update[of "get_fh hf" "edenot E s" fh "(pwrite, edenot E s)"] asm0 calculation
by (metis get_gu.elims plus_extract(3) snd_conv)
qed
moreover have "safe n (None :: ('i, 'a, nat) cont) C' (s', ?h) {(s, fh(edenot loc s ↦ (pwrite, edenot E s)), gs, gu)}"
using ‹s = sa› asm1(2) asm1(3) safe_skip by fastforce
moreover have "full_ownership (get_fh H') ∧ no_guard H' ∧ h' = FractionalHeap.normalize (get_fh H')"
proof -
have "full_ownership (get_fh H')"
proof (rule full_ownershipI)
fix l p
assume asm: "get_fh H' l = Some p"
then show "fst p = pwrite"
proof (cases "l = edenot loc s")
case True
then show ?thesis
using asm calculation(2) by fastforce
next
case False
then show ?thesis
by (metis (mono_tags, lifting) asm asm0 calculation(2) fst_eqD full_ownership_def get_fh.simps map_upd_Some_unfold)
qed
qed
moreover have "no_guard H'" using asm0
by (simp add: ‹H' = ((get_fh H)(edenot loc s ↦ (pwrite, edenot E s)), get_gs H, get_gu H)› no_guard_def)
moreover have "h' = FractionalHeap.normalize (get_fh H')"
proof (rule ext)
fix l show "h' l = FractionalHeap.normalize (get_fh H') l"
proof(cases "l = edenot loc s")
case True
then show ?thesis
by (metis (no_types, lifting) FractionalHeap.normalize_eq(2) ‹H' = ((get_fh H)(edenot loc s ↦ (pwrite, edenot E s)), get_gs H, get_gu H)› ‹h' = h(edenot loc s ↦ edenot E s)› fst_conv fun_upd_same get_fh.elims)
next
case False
then have "FractionalHeap.normalize (get_fh H') l = FractionalHeap.normalize (get_fh H) l"
using FractionalHeap.normalize_eq(2)[of "get_fh H'" l]
FractionalHeap.normalize_eq(2)[of "get_fh H" l] ‹H' = ((get_fh H)(edenot loc s ↦ (pwrite, edenot E s)), get_gs H, get_gu H)›
fst_conv fun_upd_other[of l "edenot loc s" "get_fh H"] get_fh.simps option.exhaust
by metis
then show ?thesis
using False ‹h' = h(edenot loc s ↦ edenot E s)› asm1(1) by force
qed
qed
ultimately show ?thesis
by auto
qed
ultimately show "∃h'' H'. full_ownership (get_fh H') ∧ no_guard H' ∧ h' = FractionalHeap.normalize (get_fh H') ∧
Some H' = Some h'' ⊕ Some hf ∧ safe n None C' (s', h'') {(s, fh(edenot loc s ↦ (pwrite, edenot E s)), gs, gu)}"
by (metis ‹Some H' = Some (fh(edenot loc s ↦ (pwrite, edenot E s)), gs, gu) ⊕ Some hf› ‹s = sa› asm1(2) asm1(3) fst_conv insertI1 safe_skip)
qed
qed
qed (simp)
lemma safe_write_Some:
assumes "fh (edenot loc s) = Some (pwrite, v)"
and "view_function_of_inv Γ"
shows "safe n (Some Γ) (Cwrite loc E) (s, (fh, gs, gu)) { (s, (fh(edenot loc s ↦ (pwrite, edenot E s)), gs, gu)) }"
using assms
proof (induct n)
case (Suc n)
show ?case
proof (rule safeSomeI)
show "Cwrite loc E = Cskip ⟹ (s, fh, gs, gu) ∈ {(s, fh(edenot loc s ↦ (pwrite, edenot E s)), gs, gu)}"
by simp
show "no_abort (Some Γ) (Cwrite loc E) s (fh, gs, gu)"
proof (rule no_abortSomeI)
fix H hf hj v0
assume asm0: "Some H = Some (fh, gs, gu) ⊕ Some hj ⊕ Some hf ∧ full_ownership (get_fh H) ∧ semi_consistent Γ v0 H ∧ sat_inv s hj Γ"
then have "edenot loc s ∈ dom (get_fh H)"
by (metis Un_iff assms(1) domI dom_three_sum fst_conv get_fh.simps)
then have "edenot loc s ∈ dom (normalize (get_fh H))"
by (simp add: dom_normalize)
then show "¬ aborts (Cwrite loc E) (s, FractionalHeap.normalize (get_fh H))"
by (metis aborts_write_cases fst_eqD snd_eqD)
qed
show "accesses (Cwrite loc E) s ⊆ dom (fst (fh, gs, gu)) ∧ writes (Cwrite loc E) s ⊆ fpdom (fst (fh, gs, gu))"
using Suc.prems(1) fpdom_def by fastforce
fix H hf C' s' h' hj v0
assume asm0: "Some H = Some (fh, gs, gu) ⊕ Some hj ⊕ Some hf ∧
full_ownership (get_fh H) ∧ semi_consistent Γ v0 H ∧ sat_inv s hj Γ ∧ red (Cwrite loc E) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
then obtain hjf where hjf_def: "Some hjf = Some hj ⊕ Some hf"
by (metis (no_types, opaque_lifting) option.exhaust_sel plus.simps(1) plus_asso plus_comm)
then have asm00: "Some H = Some (fh, gs, gu) ⊕ Some hjf"
by (metis asm0 plus_asso)
then have "get_fh hjf (edenot loc s) = None"
proof -
have "compatible_fract_heaps fh (get_fh hjf)"
by (metis asm00 compatible_def compatible_eq fst_conv get_fh.elims option.discI)
then show ?thesis using compatible_then_dom_disjoint(2)[of fh "get_fh hjf"]
assms disjoint_iff_not_equal[of "dom (get_fh hjf)" "fpdom fh"] not_in_dom fpdom_def mem_Collect_eq
by fastforce
qed
show "∃h'' H' hj'. full_ownership (get_fh H') ∧ semi_consistent Γ v0 H' ∧ sat_inv s' hj' Γ ∧ h' = FractionalHeap.normalize (get_fh H') ∧
Some H' = Some h'' ⊕ Some hj' ⊕ Some hf ∧ safe n (Some Γ) C' (s', h'') {(s, fh(edenot loc s ↦ (pwrite, edenot E s)), gs, gu)}"
proof (rule red_write_cases)
show "red (Cwrite loc E) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
using asm0 by blast
fix sa h
assume asm1: "(s, FractionalHeap.normalize (get_fh H)) = (sa, h)" "C' = Cskip"
"(s', h') = (sa, h(edenot loc sa ↦ edenot E sa))"
then obtain "s = sa" "h' = h(edenot loc s ↦ edenot E s)" by blast
let ?h = "(fh(edenot loc s ↦ (pwrite, edenot E s)), gs, gu)"
have "?h ## hjf"
proof (rule compatibleI)
show "compatible_fract_heaps (get_fh (fh(edenot loc s ↦ (pwrite, edenot E s)), gs, gu)) (get_fh hjf)"
proof (rule compatible_fract_heapsI)
fix l p p' assume asm2: "get_fh (fh(edenot loc s ↦ (pwrite, edenot E s)), gs, gu) l = Some p ∧ get_fh hjf l = Some p'"
then show "pgte pwrite (padd (fst p) (fst p'))"
apply (cases "l = edenot loc s")
apply (metis Suc.prems(1) asm00 fst_conv fun_upd_same get_fh.elims option.sel plus_extract_point_fh)
by (metis asm00 fst_conv fun_upd_other get_fh.elims plus_extract_point_fh)
show "snd p = snd p'"
apply (cases "l = edenot loc s")
using ‹pgte pwrite (padd (fst p) (fst p'))› asm2 not_pgte_charact sum_larger apply fastforce
by (metis (mono_tags, opaque_lifting) asm00 asm2 fst_eqD get_fh.simps map_upd_Some_unfold plus_extract_point_fh)
qed
show "⋀k. get_gu (fh(edenot loc s ↦ (pwrite, edenot E s)), gs, gu) k = None ∨ get_gu hjf k = None"
by (metis asm00 compatible_def compatible_eq get_gu.simps option.discI snd_conv)
show "⋀p p'. get_gs (fh(edenot loc s ↦ (pwrite, edenot E s)), gs, gu) = Some p ∧ get_gs hjf = Some p' ⟹ pgte pwrite (padd (fst p) (fst p'))"
by (metis asm00 compatible_def compatible_eq get_gs.simps option.discI snd_conv)
qed
then obtain H' where "Some H' = Some ?h ⊕ Some hjf" by auto
moreover have "H' = ((get_fh H)(edenot loc s ↦ (pwrite, edenot E s)), get_gs H, get_gu H)"
proof (rule heap_ext)
show "get_gs H' = get_gs ((get_fh H)(edenot loc s ↦ (pwrite, edenot E s)), get_gs H, get_gu H)"
using asm00 calculation
by (metis fst_conv get_gs.simps plus_extract(2) snd_conv)
show "get_gu H' = get_gu ((get_fh H)(edenot loc s ↦ (pwrite, edenot E s)), get_gs H, get_gu H)"
using asm00 calculation
by (metis get_gu.simps plus_extract(3) snd_conv)
show "get_fh H' = get_fh ((get_fh H)(edenot loc s ↦ (pwrite, edenot E s)), get_gs H, get_gu H)"
proof (rule ext)
fix l show "get_fh H' l = get_fh ((get_fh H)(edenot loc s ↦ (pwrite, edenot E s)), get_gs H, get_gu H) l"
using add_fh_update[of "get_fh hjf" "edenot E s" fh "(pwrite, edenot E s)"]
by (metis ‹get_fh hjf (edenot loc s) = None› add_fh_update add_get_fh asm00 calculation fst_conv get_fh.elims)
qed
qed
moreover have "safe n (Some Γ) C' (s', ?h) {(s, fh(edenot loc s ↦ (pwrite, edenot E s)), gs, gu)}"
using ‹s = sa› asm1(2) asm1(3) safe_skip by fastforce
moreover have "full_ownership (get_fh H') ∧ h' = FractionalHeap.normalize (get_fh H')"
proof -
have "full_ownership (get_fh H')"
proof (rule full_ownershipI)
fix l p
assume asm: "get_fh H' l = Some p"
then show "fst p = pwrite"
proof (cases "l = edenot loc s")
case True
then show ?thesis
using asm calculation(2) by fastforce
next
case False
then show ?thesis
by (metis (mono_tags, lifting) asm asm0 calculation(2) fst_eqD full_ownership_def get_fh.simps map_upd_Some_unfold)
qed
qed
moreover have "h' = FractionalHeap.normalize (get_fh H')"
proof (rule ext)
fix l show "h' l = FractionalHeap.normalize (get_fh H') l"
proof(cases "l = edenot loc s")
case True
then show ?thesis
by (metis (no_types, lifting) FractionalHeap.normalize_eq(2) ‹H' = ((get_fh H)(edenot loc s ↦ (pwrite, edenot E s)), get_gs H, get_gu H)› ‹h' = h(edenot loc s ↦ edenot E s)› fst_conv fun_upd_same get_fh.elims)
next
case False
then have "FractionalHeap.normalize (get_fh H') l = FractionalHeap.normalize (get_fh H) l"
using FractionalHeap.normalize_eq(2)[of "get_fh H'" l]
FractionalHeap.normalize_eq(2)[of "get_fh H" l] ‹H' = ((get_fh H)(edenot loc s ↦ (pwrite, edenot E s)), get_gs H, get_gu H)›
fst_conv fun_upd_other[of l "edenot loc s" "get_fh H"] get_fh.simps option.exhaust
by metis
then show ?thesis
using False ‹h' = h(edenot loc s ↦ edenot E s)› asm1(1) by force
qed
qed
ultimately show ?thesis
by auto
qed
moreover have "Some H' = Some ?h ⊕ Some hj ⊕ Some hf"
by (metis calculation(1) hjf_def simpler_asso)
moreover have "semi_consistent Γ v0 H'"
proof (rule semi_consistentI)
show "all_guards H'"
by (metis all_guards_def asm0 calculation(2) fst_conv get_gs.simps get_gu.simps semi_consistent_def snd_conv)
have "view Γ (normalize (get_fh H')) = view Γ (normalize (get_fh H))"
proof -
have "view Γ (normalize (get_fh H')) = view Γ (normalize (get_fh hj))"
by (metis asm0 assms(2) calculation(5) larger3 view_function_of_invE)
then show ?thesis using assms(2) larger3 view_function_of_invE
by (metis asm0)
qed
then show "reachable Γ v0 H'"
by (metis asm0 calculation(2) fst_eqD get_gs.simps get_gu.simps reachableE reachableI semi_consistent_def snd_eqD)
qed
ultimately show "∃h'' H' hj'.
full_ownership (get_fh H') ∧
semi_consistent Γ v0 H' ∧
sat_inv s' hj' Γ ∧
h' = FractionalHeap.normalize (get_fh H') ∧
Some H' = Some h'' ⊕ Some hj' ⊕ Some hf ∧ safe n (Some Γ) C' (s', h'') {(s, fh(edenot loc s ↦ (pwrite, edenot E s)), gs, gu)}"
using ‹s = sa› asm0 asm1(2) asm1(3) by blast
qed
qed
qed (simp)
lemma safe_write:
fixes Δ :: "('i, 'a, nat) cont"
assumes "fh (edenot loc s) = Some (pwrite, v)"
and "⋀Γ. Δ = Some Γ ⟹ view_function_of_inv Γ"
shows "safe n Δ (Cwrite loc E) (s, (fh, gs, gu)) { (s, (fh(edenot loc s ↦ (pwrite, edenot E s)), gs, gu)) }"
apply (cases Δ)
using safe_write_None safe_write_Some assms by blast+
theorem write_rule:
fixes Δ :: "('i, 'a, nat) cont"
assumes "⋀Γ. Δ = Some Γ ⟹ view_function_of_inv Γ"
and "v ∉ fvE loc"
shows "hoare_triple_valid Δ (Exists v (PointsTo loc pwrite (Evar v))) (Cwrite loc E) (PointsTo loc pwrite E)"
proof (rule hoare_triple_validI)
define Σ :: "store × ('i, 'a) heap ⇒ (store × ('i, 'a) heap) set" where
"Σ = (λ(s, h). { (s, ((get_fh h)(edenot loc s ↦ (pwrite, edenot E s)), get_gs h, get_gu h)) })"
show "⋀s h n. (s, h), (s, h) ⊨ Exists v (PointsTo loc pwrite (Evar v)) ⟹ safe n Δ (Cwrite loc E) (s, h) (Σ (s, h))"
proof -
fix s h n assume "(s, h :: ('i, 'a) heap), (s, h) ⊨ Exists v (PointsTo loc pwrite (Evar v))"
then obtain vv where "(s(v := vv), h), (s(v := vv), h) ⊨ PointsTo loc pwrite (Evar v)"
by (meson hyper_sat.simps(6) hyper_sat.simps(7))
then have "get_fh h (edenot loc (s(v := vv))) = Some (pwrite, vv)"
by simp
then have "get_fh h (edenot loc s) = Some (pwrite, vv)"
using assms(2) by auto
then show "safe n Δ (Cwrite loc E) (s, h) (Σ (s, h))"
by (metis (mono_tags, lifting) Σ_def assms(1) decompose_heap_triple old.prod.case safe_write)
qed
fix s1 h1 s2 h2
assume "(s1, h1 :: ('i, 'a) heap), (s2, h2) ⊨ Exists v (PointsTo loc pwrite (Evar v))"
then obtain v1 v2 where "get_fh h1 (edenot loc s1) = Some (pwrite, v1)" "get_fh h2 (edenot loc s2) = Some (pwrite, v2)"
using assms(2) by auto
show "pair_sat (case (s1, h1) of (s, h) ⇒ {(s, (get_fh h)(edenot loc s ↦ (pwrite, edenot E s)), get_gs h, get_gu h)})
(case (s2, h2) of (s, h) ⇒ {(s, (get_fh h)(edenot loc s ↦ (pwrite, edenot E s)), get_gs h, get_gu h)}) (PointsTo loc pwrite E)"
proof (rule pair_satI)
fix s1' h1' s2' h2'
assume asm0: "(s1', h1') ∈ (case (s1, h1) of (s, h) ⇒ {(s, (get_fh h)(edenot loc s ↦ (pwrite, edenot E s)), get_gs h, get_gu h)}) ∧
(s2', h2') ∈ (case (s2, h2) of (s, h) ⇒ {(s, (get_fh h)(edenot loc s ↦ (pwrite, edenot E s)), get_gs h, get_gu h)})"
then show "(s1', h1'), (s2', h2') ⊨ PointsTo loc pwrite E"
using ‹(s1, h1), (s2, h2) ⊨ Exists v (PointsTo loc pwrite (Evar v))› assms(2) by auto
qed
qed
subsubsection ‹Read›
inductive_cases red_read_cases: "red (Cread x E) σ C' σ'"
inductive_cases aborts_read_cases: "aborts (Cread x E) σ"
lemma safe_read_None:
"safe n (None :: ('i, 'a, nat) cont) (Cread x E) (s, ([edenot E s ↦ (π, v)], gs, gu)) { (s(x := v), ([edenot E s ↦ (π, v)], gs, gu)) }"
proof (induct n)
case (Suc n)
show ?case
proof (rule safeNoneI)
show "no_abort (None :: ('i, 'a, nat) cont) (Cread x E) s ([edenot E s ↦ (π, v)], gs, gu)"
proof (rule no_abortNoneI)
fix hf H
assume asm0: "Some H = Some ([edenot E s ↦ (π, v)], gs, gu) ⊕ Some hf ∧ full_ownership (get_fh H) ∧ no_guard H"
then have "edenot E s ∈ dom (get_fh H)"
by (metis Un_iff dom_eq_singleton_conv dom_sum_two fst_eqD get_fh.elims insert_iff)
then have "edenot E s ∈ dom (FractionalHeap.normalize (get_fh H))"
by (simp add: dom_normalize)
then show "¬ aborts (Cread x E) (s, FractionalHeap.normalize (get_fh H))"
by (metis aborts_read_cases fst_eqD snd_eqD)
qed
show "accesses (Cread x E) s ⊆ dom (fst ([edenot E s ↦ (π, v)], gs, gu)) ∧ writes (Cread x E) s ⊆ fpdom (fst ([edenot E s ↦ (π, v)], gs, gu))"
by simp
fix H hf C' s' h'
assume asm0: "Some H = Some ([edenot E s ↦ (π, v)], gs, gu) ⊕ Some hf ∧
full_ownership (get_fh H) ∧ no_guard H ∧ red (Cread x E) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
let ?S = "{ (s(x := v), ([edenot E s ↦ (π, v)], gs, gu)) }"
show "∃h'' H'.
full_ownership (get_fh H') ∧
no_guard H' ∧
h' = FractionalHeap.normalize (get_fh H') ∧
Some H' = Some h'' ⊕ Some hf ∧ safe n (None :: ('i, 'a, nat) cont) C' (s', h'') ?S"
proof (rule red_read_cases)
show "red (Cread x E) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
using asm0 by blast
fix sa h va
assume "(s, FractionalHeap.normalize (get_fh H)) = (sa, h)" "C' = Cskip" "(s', h') = (sa(x := va), h)"
"h (edenot E sa) = Some va"
then have "s = sa"
by force
then have "va = v"
proof -
have "∃π'. get_fh H (edenot E s) = Some (π', v)"
proof (rule one_value_sum_same)
show "Some H = Some ([edenot E s ↦ (π, v)], gs, gu) ⊕ Some hf"
using asm0 by fastforce
qed (simp)
then show ?thesis
by (metis FractionalHeap.normalize_eq(2) Pair_inject ‹(s, FractionalHeap.normalize (get_fh H)) = (sa, h)› ‹h (edenot E sa) = Some va› option.sel)
qed
then have "safe n (None :: ('i, 'a, nat) cont) C' (s', ([edenot E s ↦ (π, v)], gs, gu)) ?S"
using ‹(s', h') = (sa(x := va), h)› ‹C' = Cskip› ‹s = sa› safe_skip by fastforce
then show "∃h'' H'.
full_ownership (get_fh H') ∧ no_guard H' ∧
h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hf ∧ safe n (None :: ('i, 'a, nat) cont) C' (s', h'') {(s(x := v), [edenot E s ↦ (π, v)], gs, gu)}"
using ‹(s', h') = (sa(x := va), h)› ‹(s, FractionalHeap.normalize (get_fh H)) = (sa, h)› asm0 by blast
qed
qed (simp)
qed (simp)
lemma safe_read_Some:
assumes "view_function_of_inv Γ"
and "x ∉ fvA (invariant Γ)"
shows "safe n (Some Γ) (Cread x E) (s, ([edenot E s ↦ (π, v)], gs, gu)) { (s(x := v), ([edenot E s ↦ (π, v)], gs, gu)) }"
proof (induct n)
case (Suc n)
show ?case
proof (rule safeSomeI)
show "no_abort (Some Γ) (Cread x E) s ([edenot E s ↦ (π, v)], gs, gu)"
proof (rule no_abortSomeI)
fix hf H hj v0
assume asm0: "Some H = Some ([edenot E s ↦ (π, v)], gs, gu) ⊕ Some hj ⊕ Some hf ∧ full_ownership (get_fh H) ∧ semi_consistent Γ v0 H ∧ sat_inv s hj Γ"
then obtain hjf where "Some H = Some ([edenot E s ↦ (π, v)], gs, gu) ⊕ Some hjf"
by (metis (no_types, lifting) plus.simps(2) plus.simps(3) plus_asso)
then have "edenot E s ∈ dom (get_fh H)"
by (metis Un_iff dom_eq_singleton_conv dom_sum_two fst_eqD get_fh.elims insert_iff)
then have "edenot E s ∈ dom (FractionalHeap.normalize (get_fh H))"
by (simp add: dom_normalize)
then show "¬ aborts (Cread x E) (s, FractionalHeap.normalize (get_fh H))"
by (metis aborts_read_cases fst_eqD snd_eqD)
qed
show "accesses (Cread x E) s ⊆ dom (fst ([edenot E s ↦ (π, v)], gs, gu)) ∧ writes (Cread x E) s ⊆ fpdom (fst ([edenot E s ↦ (π, v)], gs, gu))"
by simp
fix H hf C' s' h' hj v0
assume asm0: "Some H = Some ([edenot E s ↦ (π, v)], gs, gu) ⊕ Some hj ⊕ Some hf ∧
full_ownership (get_fh H) ∧ semi_consistent Γ v0 H ∧ sat_inv s hj Γ ∧ red (Cread x E) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
then obtain hjf where "Some hjf = Some hj ⊕ Some hf"
using compatible_last_two by (metis plus.simps(3) plus_asso)
then have "Some H = Some ([edenot E s ↦ (π, v)], gs, gu) ⊕ Some hjf"
by (metis asm0 plus_asso)
let ?S = "{ (s(x := v), ([edenot E s ↦ (π, v)], gs, gu)) }"
show "∃h'' H' hj'. full_ownership (get_fh H') ∧ semi_consistent Γ v0 H' ∧ sat_inv s' hj' Γ ∧ h' = FractionalHeap.normalize (get_fh H') ∧
Some H' = Some h'' ⊕ Some hj' ⊕ Some hf ∧ safe n (Some Γ) C' (s', h'') {(s(x := v), [edenot E s ↦ (π, v)], gs, gu)}"
proof (rule red_read_cases)
show "red (Cread x E) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
using asm0 by blast
fix sa h va
assume "(s, FractionalHeap.normalize (get_fh H)) = (sa, h)" "C' = Cskip" "(s', h') = (sa(x := va), h)"
"h (edenot E sa) = Some va"
then have "s = sa"
by force
then have "va = v"
proof -
have "∃π'. get_fh H (edenot E s) = Some (π', v)"
proof (rule one_value_sum_same)
show "Some H = Some ([edenot E s ↦ (π, v)], gs, gu) ⊕ Some hjf"
by (simp add: ‹Some H = Some ([edenot E s ↦ (π, v)], gs, gu) ⊕ Some hjf›)
qed (simp)
then show ?thesis
by (metis FractionalHeap.normalize_eq(2) Pair_inject ‹(s, FractionalHeap.normalize (get_fh H)) = (sa, h)› ‹h (edenot E sa) = Some va› option.sel)
qed
then have "safe n (Some Γ) C' (s', ([edenot E s ↦ (π, v)], gs, gu)) ?S"
using ‹(s', h') = (sa(x := va), h)› ‹C' = Cskip› ‹s = sa› safe_skip by fastforce
moreover have "sat_inv s' hj Γ"
by (metis ‹(s', h') = (sa(x := va), h)› ‹s = sa› agrees_update asm0 assms(2) prod.inject sat_inv_agrees)
ultimately show "∃h'' H' hj'.
full_ownership (get_fh H') ∧ semi_consistent Γ v0 H' ∧ sat_inv s' hj' Γ ∧ h' = FractionalHeap.normalize (get_fh H') ∧
Some H' = Some h'' ⊕ Some hj' ⊕ Some hf ∧ safe n (Some Γ) C' (s', h'') {(s(x := v), [edenot E s ↦ (π, v)], gs, gu)}"
using ‹(s', h') = (sa(x := va), h)› ‹(s, FractionalHeap.normalize (get_fh H)) = (sa, h)› asm0 by blast
qed
qed (simp)
qed (simp)
lemma safe_read:
fixes Δ :: "('i, 'a, nat) cont"
assumes "⋀Γ. Δ = Some Γ ⟹ x ∉ fvA (invariant Γ) ∧ view_function_of_inv Γ"
shows "safe n Δ (Cread x E) (s, ([edenot E s ↦ (π, v)], gs, gu)) { (s(x := v), ([edenot E s ↦ (π, v)], gs, gu)) }"
apply (cases Δ)
using safe_read_None safe_read_Some assms by blast+
theorem read_rule:
fixes Δ :: "('i, 'a, nat) cont"
assumes "⋀Γ. Δ = Some Γ ⟹ x ∉ fvA (invariant Γ) ∧ view_function_of_inv Γ"
and "x ∉ fvE E ∪ fvE e"
shows "hoare_triple_valid Δ (PointsTo E π e) (Cread x E) (And (PointsTo E π e) (Bool (Beq (Evar x) e)))"
proof (rule hoare_triple_validI)
define Σ :: "store × ('i, 'a) heap ⇒ (store × ('i, 'a) heap) set" where
"Σ = (λ(s, h). { (s(x := edenot e s), ([edenot E s ↦ (π, edenot e s)], get_gs h, get_gu h)) })"
show "⋀s h n. (s, h), (s, h) ⊨ PointsTo E π e ⟹ safe n Δ (Cread x E) (s, h) (Σ (s, h))"
proof -
fix s h n
assume "(s, h :: ('i, 'a) heap), (s, h) ⊨ PointsTo E π e"
then have "get_fh h = [edenot E s ↦ (π, edenot e s)]"
using sat_points_to by blast
then have "h = ([edenot E s ↦ (π, edenot e s)], get_gs h, get_gu h)"
by (metis decompose_heap_triple)
then have "safe n Δ (Cread x E) (s, ([edenot E s ↦ (π, edenot e s)], get_gs h, get_gu h))
{ (s(x := edenot e s), ([edenot E s ↦ (π, edenot e s)], get_gs h, get_gu h)) }"
using assms safe_read by blast
then show "safe n Δ (Cread x E) (s, h) (Σ (s, h))"
using Σ_def ‹h = ([edenot E s ↦ (π, edenot e s)], get_gs h, get_gu h)› by auto
qed
fix s1 h1 s2 h2
assume "(s1, h1 :: ('i, 'a) heap), (s2, h2) ⊨ PointsTo E π e"
show "pair_sat (case (s1, h1) of (s, h) ⇒ {(s(x := edenot e s), [edenot E s ↦ (π, edenot e s)], get_gs h, get_gu h)})
(case (s2, h2) of (s, h) ⇒ {(s(x := edenot e s), [edenot E s ↦ (π, edenot e s)], get_gs h, get_gu h)}) (And (PointsTo E π e) (Bool (Beq (Evar x) e)))"
proof (rule pair_satI)
fix s1' h1' s2' h2'
assume asm0: "(s1', h1') ∈ (case (s1, h1) of (s, h) ⇒ {(s(x := edenot e s), [edenot E s ↦ (π, edenot e s)], get_gs h, get_gu h)}) ∧
(s2', h2') ∈ (case (s2, h2) of (s, h) ⇒ {(s(x := edenot e s), [edenot E s ↦ (π, edenot e s)], get_gs h, get_gu h)})"
then obtain "s1' = s1(x := edenot e s1)" "h1' = ([edenot E s1 ↦ (π, edenot e s1)], get_gs h1, get_gu h1)"
"s2' = s2(x := edenot e s2)" "h2' = ([edenot E s2 ↦ (π, edenot e s2)], get_gs h2, get_gu h2)"
by force
then show "(s1', h1'), (s2', h2') ⊨ And (PointsTo E π e) (Bool (Beq (Evar x) e))"
using assms(2) by auto
qed
qed
subsubsection ‹Share›
lemma share_no_abort:
assumes "no_abort (Some Γ) C s (h :: ('i, 'a) heap)"
and "Some (h' :: ('i, 'a) heap) = Some h ⊕ Some hj"
and "sat_inv s hj Γ"
and "get_gs h = Some (pwrite, sargs)"
and "⋀k. get_gu h k = Some (uargs k)"
and "reachable_value (saction Γ) (uaction Γ) v0 sargs uargs (view Γ (normalize (get_fh hj)))"
and "view_function_of_inv Γ"
shows "no_abort None C s (remove_guards h')"
proof (rule no_abortI)
show "⋀H hf hj v0 Γ.
None = Some Γ ∧
Some H = Some (remove_guards h') ⊕ Some hj ⊕ Some hf ∧ full_ownership (get_fh H) ∧ semi_consistent Γ v0 H ∧ sat_inv s hj Γ ⟹
¬ aborts C (s, FractionalHeap.normalize (get_fh H))" by blast
fix hf H :: "('i, 'a) heap"
assume asm0: "Some H = Some (remove_guards h') ⊕ Some hf ∧ None = None ∧ full_ownership (get_fh H) ∧ no_guard H"
have "compatible h' hf"
proof (rule compatibleI)
show "compatible_fract_heaps (get_fh h') (get_fh hf)"
by (metis asm0 compatible_def compatible_eq fst_eqD get_fh.simps option.distinct(1) remove_guards_def)
show "⋀k. get_gu h' k = None ∨ get_gu hf k = None"
by (metis asm0 no_guard_def no_guard_then_smaller_same plus_comm)
fix p p' assume "get_gs h' = Some p ∧ get_gs hf = Some p'"
then show "pgte pwrite (padd (fst p) (fst p'))"
by (metis asm0 no_guard_def no_guard_then_smaller_same option.distinct(1) plus_comm)
qed
then obtain H' where "Some H' = Some h' ⊕ Some hf"
by simp
then have "get_fh H' = get_fh H"
by (metis asm0 fst_eqD get_fh.elims option.discI remove_guards_def option.sel plus.simps(3))
have "¬ aborts C (s, FractionalHeap.normalize (get_fh H'))"
proof (rule no_abortE(2))
show "no_abort (Some Γ) C s h"
using assms by blast
show "Some Γ = Some Γ" by blast
show "full_ownership (get_fh H')"
using ‹get_fh H' = get_fh H› asm0 by presburger
show "semi_consistent Γ v0 H'"
proof (rule semi_consistentI)
show "all_guards H'"
by (metis ‹Some H' = Some h' ⊕ Some hf› all_guards_def all_guards_same assms(2) assms(4) assms(5) option.discI)
have "view Γ (normalize (get_fh hj)) = view Γ (normalize (get_fh H'))"
using assms(7)
proof (rule view_function_of_invE)
show "H' ≽ hj"
using larger_trans
by (simp add: ‹Some H' = Some h' ⊕ Some hf› assms(2) larger3)
show "sat_inv s hj Γ"
by (simp add: assms(3))
qed
show "reachable Γ v0 H'"
proof (rule reachableI)
fix sargs' uargs'
assume asm1: "get_gs H' = Some (pwrite, sargs') ∧ (∀k. get_gu H' k = Some (uargs' k))"
then have "sargs = sargs'"
by (metis ‹Some H' = Some h' ⊕ Some hf› assms(2) assms(4) full_sguard_sum_same option.inject snd_conv)
moreover have "uargs = uargs'"
proof (rule ext)
fix k
show "uargs k = uargs' k"
using full_uguard_sum_same[of h' k _ H' hf]
by (metis ‹Some H' = Some h' ⊕ Some hf› asm1 assms(2) assms(5) full_uguard_sum_same option.inject)
qed
ultimately show "reachable_value (saction Γ) (uaction Γ) v0 sargs' uargs' (view Γ (FractionalHeap.normalize (get_fh H')))"
using ‹view Γ (FractionalHeap.normalize (get_fh hj)) = view Γ (FractionalHeap.normalize (get_fh H'))› assms(6) by presburger
qed
qed
show "Some H' = Some h ⊕ Some hj ⊕ Some hf"
using ‹Some H' = Some h' ⊕ Some hf› assms(2) by presburger
show "sat_inv s hj Γ"
by (simp add: assms(3))
qed
then show "¬ aborts C (s, FractionalHeap.normalize (get_fh H))"
using ‹get_fh H' = get_fh H› by auto
qed
definition S_after_share where
"S_after_share S Γ v0 = { (s, remove_guards h') |h hj h' s. semi_consistent Γ v0 h' ∧ Some h' = Some h ⊕ Some hj ∧ (s, h) ∈ S ∧ sat_inv s hj Γ }"
lemma share_lemma:
assumes "safe n (Some Γ) C (s, h :: ('i, 'a) heap) S"
and "Some (h' :: ('i, 'a) heap) = Some h ⊕ Some hj"
and "sat_inv s hj Γ"
and "semi_consistent Γ v0 h'"
and "view_function_of_inv Γ"
and "bounded h'"
shows "safe n (None :: ('i, 'a, nat) cont) C (s, remove_guards h') (S_after_share S Γ v0)"
using assms
proof (induct n arbitrary: C s h h' hj)
case (Suc n)
let ?S' = "S_after_share S Γ v0"
have is_in_s': "⋀h hj h'. Some h' = Some h ⊕ Some hj ∧ (s, h) ∈ S ∧ sat_inv s hj Γ ∧ semi_consistent Γ v0 h' ⟹ (s, remove_guards h') ∈ ?S'"
proof -
fix h hj h' assume "Some h' = Some h ⊕ Some hj ∧ (s, h) ∈ S ∧ sat_inv s hj Γ ∧ semi_consistent Γ v0 h'"
then show "(s, remove_guards h') ∈ ?S'"
using S_after_share_def[of S Γ v0] mem_Collect_eq by blast
qed
show ?case
proof (rule safeNoneI)
show "C = Cskip ⟹ (s, remove_guards h') ∈ ?S'"
proof -
assume "C = Cskip"
show "(s, remove_guards h') ∈ ?S'"
proof (rule is_in_s')
show "Some h' = Some h ⊕ Some hj ∧ (s, h) ∈ S ∧ sat_inv s hj Γ ∧ semi_consistent Γ v0 h'"
using Suc.prems ‹C = Cskip› safeSomeE(1) sat_inv_def by blast
qed
qed
obtain sargs uargs where " get_gs h' = Some (pwrite, sargs) ∧
(∀k. get_gu h' k = Some (uargs k)) ∧ reachable_value (saction Γ) (uaction Γ) v0 sargs uargs (view Γ (FractionalHeap.normalize (get_fh h')))"
by (meson Suc.prems(4) semi_consistentE)
show "no_abort None C s (remove_guards h')"
proof (rule share_no_abort)
show "no_abort (Some Γ) C s h"
using Suc.prems(1) safeSomeE(2) by blast
show "Some h' = Some h ⊕ Some hj"
using Suc.prems(2) by blast
show "sat_inv s hj Γ"
using Suc.prems(3) by auto
show "get_gs h = Some (pwrite, sargs)"
by (metis Suc.prems(2) ‹get_gs h' = Some (pwrite, sargs) ∧ (∀k. get_gu h' k = Some (uargs k)) ∧ reachable_value (saction Γ) (uaction Γ) v0 sargs uargs (view Γ (FractionalHeap.normalize (get_fh h')))› ‹sat_inv s hj Γ› no_guard_remove(1) sat_inv_def)
show "⋀k. get_gu h k = Some (uargs k)"
by (metis Suc.prems(2) ‹get_gs h' = Some (pwrite, sargs) ∧ (∀k. get_gu h' k = Some (uargs k)) ∧ reachable_value (saction Γ) (uaction Γ) v0 sargs uargs (view Γ (FractionalHeap.normalize (get_fh h')))› ‹sat_inv s hj Γ› no_guard_remove(2) sat_inv_def)
show "reachable_value (saction Γ) (uaction Γ) v0 sargs uargs (view Γ (FractionalHeap.normalize (get_fh hj)))"
by (metis Suc.prems(2) Suc.prems(3) ‹get_gs h' = Some (pwrite, sargs) ∧ (∀k. get_gu h' k = Some (uargs k)) ∧ reachable_value (saction Γ) (uaction Γ) v0 sargs uargs (view Γ (FractionalHeap.normalize (get_fh h')))› assms(5) larger_def plus_comm view_function_of_invE)
show "view_function_of_inv Γ"
by (simp add: assms(5))
qed
have "accesses C s ⊆ dom (fst h) ∧ writes C s ⊆ fpdom (fst h)"
using Suc.prems(1) by force
moreover have "dom (fst h) ⊆ dom (fst h')"
by (metis Suc.prems(2) addition_smaller_domain get_fh.simps)
moreover have "fpdom (fst h) ⊆ fpdom (fst h')"
by (simp add: Suc.prems(2) Suc(7) fpdom_inclusion)
ultimately show "accesses C s ⊆ dom (fst (remove_guards h')) ∧ writes C s ⊆ fpdom (fst (remove_guards h'))"
unfolding remove_guards_def by force
fix H hf C' s' h'a
assume asm0: "Some H = Some (remove_guards h') ⊕ Some hf ∧
full_ownership (get_fh H) ∧ no_guard H ∧ red C (s, FractionalHeap.normalize (get_fh H)) C' (s', h'a)"
have "compatible h' hf"
proof (rule compatibleI)
show "compatible_fract_heaps (get_fh h') (get_fh hf)"
by (metis asm0 compatible_def compatible_eq fst_eqD get_fh.simps option.distinct(1) remove_guards_def)
show "⋀k. get_gu h' k = None ∨ get_gu hf k = None"
by (metis asm0 no_guard_def no_guard_then_smaller_same plus_comm)
fix p p' assume "get_gs h' = Some p ∧ get_gs hf = Some p'"
then show "pgte pwrite (padd (fst p) (fst p'))"
by (metis asm0 no_guard_def no_guard_then_smaller_same option.distinct(1) plus_comm)
qed
then obtain Hg where "Some Hg = Some h' ⊕ Some hf"
by simp
then have "get_fh Hg = get_fh H"
by (metis asm0 fst_eqD get_fh.elims option.discI remove_guards_def option.sel plus.simps(3))
have "∃h'' H' hj'.
full_ownership (get_fh H') ∧
semi_consistent Γ v0 H' ∧
sat_inv s' hj' Γ ∧ h'a = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hj' ⊕ Some hf ∧ safe n (Some Γ) C' (s', h'') S"
using Suc(2)
proof (rule safeSomeE(3)[of n Γ C s h S Hg hj hf v0 C' s' h'a])
show "Some Hg = Some h ⊕ Some hj ⊕ Some hf"
by (simp add: Suc.prems(2) ‹Some Hg = Some h' ⊕ Some hf›)
show "full_ownership (get_fh Hg)"
using ‹get_fh Hg = get_fh H› asm0 by presburger
show "sat_inv s hj Γ"
by (simp add: Suc.prems(3))
show "red C (s, FractionalHeap.normalize (get_fh Hg)) C' (s', h'a)"
using ‹get_fh Hg = get_fh H› asm0 by presburger
show "semi_consistent Γ v0 Hg"
proof (rule semi_consistentI)
show "all_guards Hg"
by (meson Suc.prems(4) ‹Some Hg = Some h' ⊕ Some hf› all_guards_same semi_consistent_def)
have "view Γ (normalize (get_fh hj)) = view Γ (normalize (get_fh Hg))"
using assms(5)
proof (rule view_function_of_invE)
show "Hg ≽ hj"
using larger_trans
using ‹Some Hg = Some h ⊕ Some hj ⊕ Some hf› larger3 by blast
show "sat_inv s hj Γ"
by (simp add: ‹sat_inv s hj Γ›)
qed
show "reachable Γ v0 Hg"
proof (rule reachableI)
fix sargs' uargs'
assume asm1: "get_gs Hg = Some (pwrite, sargs') ∧ (∀k. get_gu Hg k = Some (uargs' k))"
then have "sargs = sargs'"
by (metis Pair_inject ‹Some Hg = Some h' ⊕ Some hf› ‹get_gs h' = Some (pwrite, sargs) ∧ (∀k. get_gu h' k = Some (uargs k)) ∧ reachable_value (saction Γ) (uaction Γ) v0 sargs uargs (view Γ (FractionalHeap.normalize (get_fh h')))› full_sguard_sum_same option.inject)
moreover have "uargs = uargs'"
proof (rule ext)
fix k
show "uargs k = uargs' k"
by (metis ‹Some Hg = Some h' ⊕ Some hf› ‹get_gs h' = Some (pwrite, sargs) ∧ (∀k. get_gu h' k = Some (uargs k)) ∧ reachable_value (saction Γ) (uaction Γ) v0 sargs uargs (view Γ (FractionalHeap.normalize (get_fh h')))› asm1 full_uguard_sum_same option.inject)
qed
ultimately show "reachable_value (saction Γ) (uaction Γ) v0 sargs' uargs' (view Γ (FractionalHeap.normalize (get_fh Hg)))"
by (metis Suc.prems(2) ‹get_gs h' = Some (pwrite, sargs) ∧ (∀k. get_gu h' k = Some (uargs k)) ∧ reachable_value (saction Γ) (uaction Γ) v0 sargs uargs (view Γ (FractionalHeap.normalize (get_fh h')))› ‹sat_inv s hj Γ› ‹view Γ (FractionalHeap.normalize (get_fh hj)) = view Γ (FractionalHeap.normalize (get_fh Hg))› assms(5) larger_def plus_comm view_function_of_invE)
qed
qed
qed
then obtain h'' H' hj' where asm1: "full_ownership (get_fh H') ∧ semi_consistent Γ v0 H' ∧
sat_inv s' hj' Γ ∧ h'a = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hj' ⊕ Some hf ∧ safe n (Some Γ) C' (s', h'') S"
by blast
obtain hj'' where "Some hj'' = Some h'' ⊕ Some hj'"
by (metis asm1 not_Some_eq plus.simps(1))
moreover obtain sargs' uargs' where new_guards_def:
"get_gs H' = Some (pwrite, sargs') ∧ (∀k. get_gu H' k = Some (uargs' k)) ∧ reachable_value (saction Γ) (uaction Γ) v0 sargs' uargs' (view Γ (FractionalHeap.normalize (get_fh H')))"
by (meson asm1 semi_consistentE)
have "safe n (None :: ('i, 'a, nat) cont) C' (s', remove_guards hj'') ?S'"
proof (rule Suc(1)[of C' s' h'' hj'' hj'])
show "safe n (Some Γ) C' (s', h'') S"
using asm1 by blast
show "Some hj'' = Some h'' ⊕ Some hj'"
using ‹Some hj'' = Some h'' ⊕ Some hj'› by blast
show "sat_inv s' hj' Γ"
using asm1 by fastforce
have "no_guard hf"
by (metis asm0 no_guard_then_smaller_same plus_comm)
moreover have "no_guard hj'"
using ‹sat_inv s' hj' Γ› sat_inv_def by blast
have "view Γ (normalize (get_fh hj')) = view Γ (normalize (get_fh H'))"
using assms(5)
proof (rule view_function_of_invE)
show "H' ≽ hj'"
using larger_trans
using asm1 larger3 by blast
show "sat_inv s' hj' Γ"
by (simp add: asm1)
qed
obtain uargs' sargs' where args': "get_gs H' = Some (pwrite, sargs') ∧ (∀k. get_gu H' k = Some (uargs' k)) ∧ reachable_value (saction Γ) (uaction Γ) v0 sargs' uargs'
(view Γ (FractionalHeap.normalize (get_fh H')))"
using semi_consistentE[of Γ v0 H'] asm1
by blast
then have "get_gs hj'' = Some (pwrite, sargs') ∧ (∀k. get_gu hj'' k = Some (uargs' k))"
by (metis ‹Some hj'' = Some h'' ⊕ Some hj'› asm1 calculation no_guard_remove(1) no_guard_remove(2))
show "semi_consistent Γ v0 hj''"
proof (rule semi_consistentI)
show "all_guards hj''"
by (metis ‹get_gs hj'' = Some (pwrite, sargs') ∧ (∀k. get_gu hj'' k = Some (uargs' k))› all_guards_def option.discI)
have "view Γ (FractionalHeap.normalize (get_fh H')) = view Γ (FractionalHeap.normalize (get_fh hj''))"
by (metis ‹Some hj'' = Some h'' ⊕ Some hj'› ‹view Γ (FractionalHeap.normalize (get_fh hj')) = view Γ (FractionalHeap.normalize (get_fh H'))› asm1 assms(5) larger_def plus_comm view_function_of_invE)
then show "reachable Γ v0 hj''"
by (metis ‹get_gs hj'' = Some (pwrite, sargs') ∧ (∀k. get_gu hj'' k = Some (uargs' k))› args' asm1 ext get_fh.simps new_guards_def option.sel reachable_def snd_conv)
qed
show "view_function_of_inv Γ"
by (simp add: assms(5))
show "bounded hj''"
proof (rule bounded_smaller)
show "bounded H'"
by (metis asm1 full_ownership_then_bounded get_fh.simps)
show "H' ≽ hj''"
by (metis ‹Some hj'' = Some h'' ⊕ Some hj'› asm1 larger_def)
qed
qed
let ?h'' = "remove_guards hj''"
have "hj'' ## hf"
by (metis asm1 calculation option.simps(3) plus.simps(3))
then obtain H'' where "Some H'' = Some ?h'' ⊕ Some hf"
by (simp add: remove_guards_smaller smaller_more_compatible)
then have "get_fh H'' = get_fh H'"
by (metis asm1 calculation equiv_sum_get_fh get_fh_remove_guards)
moreover have "no_guard H''"
by (metis ‹Some H'' = Some (remove_guards hj'') ⊕ Some hf› asm0 no_guard_remove_guards no_guard_then_smaller_same plus_comm sum_of_no_guards)
ultimately show "∃h'' H'.
full_ownership (get_fh H') ∧
no_guard H' ∧ h'a = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hf ∧ safe n (None :: ('i, 'a, nat) cont) C' (s', h'') ?S'"
by (metis ‹Some H'' = Some (remove_guards hj'') ⊕ Some hf› ‹safe n None C' (s', remove_guards hj'') ?S'› asm1)
qed
qed (simp)
definition no_need_guards where
"no_need_guards A ⟷ (∀s1 h1 s2 h2. (s1, h1), (s2, h2) ⊨ A ⟶ (s1, remove_guards h1), (s2, remove_guards h2) ⊨ A)"
lemma has_guard_then_safe_none:
assumes "¬ no_guard h"
and "C = Cskip ⟹ (s, h) ∈ S"
and "accesses C s ⊆ dom (fst h) ∧ writes C s ⊆ fpdom (fst h)"
shows "safe n (None :: ('i, 'a, nat) cont) C (s, h) S"
proof (induct n)
case (Suc n)
show ?case
proof (rule safeNoneI)
show "C = Cskip ⟹ (s, h) ∈ S"
by (simp add: assms(2))
show "no_abort None C s h"
using assms(1) no_abortNoneI no_guard_then_smaller_same by blast
show "accesses C s ⊆ dom (fst h) ∧ writes C s ⊆ fpdom (fst h)"
using assms(3) by blast
show "⋀H hf C' s' h'.
Some H = Some h ⊕ Some hf ∧ full_ownership (get_fh H) ∧ no_guard H ∧ red C (s, FractionalHeap.normalize (get_fh H)) C' (s', h') ⟹
∃h'' H'.
full_ownership (get_fh H') ∧ no_guard H' ∧ h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hf ∧ safe n None C' (s', h'') S"
using assms(1) no_guard_then_smaller_same by blast
qed
qed (simp)
theorem share_rule:
fixes Γ :: "('i, 'a, nat) single_context"
assumes "Γ = ⦇ view = f, abstract_view = α, saction = sact, uaction = uact, invariant = J ⦈"
and "all_axioms α sact spre uact upre"
and "hoare_triple_valid (Some Γ) (Star P EmptyFullGuards) C (Star Q (And (PreSharedGuards (Abs_precondition spre)) (PreUniqueGuards (Abs_indexed_precondition upre))))"
and "view_function_of_inv Γ"
and "unary J ∧ precise J"
and "wf_indexed_precondition upre ∧ wf_precondition spre"
and "x ∉ fvA J"
and "no_guard_assertion (Star P (LowView (α ∘ f) J x))"
shows "hoare_triple_valid (None :: ('i, 'a, nat) cont) (Star P (LowView (α ∘ f) J x)) C (Star Q (LowView (α ∘ f) J x))"
proof -
let ?P = "Star P EmptyFullGuards"
let ?Q = "Star Q (And (PreSharedGuards (Abs_precondition spre)) (PreUniqueGuards (Abs_indexed_precondition upre)))"
obtain Σ where asm0: "⋀σ n. σ, σ ⊨ Star P EmptyFullGuards ⟹ bounded (snd σ) ⟹ safe n (Some Γ) C σ (Σ σ)"
"⋀σ σ'. σ, σ' ⊨ Star P EmptyFullGuards ⟹ pair_sat (Σ σ) (Σ σ') (Star Q (And (PreSharedGuards (Abs_precondition spre)) (PreUniqueGuards (Abs_indexed_precondition upre))))"
using hoare_triple_validE[of "Some Γ" ?P C ?Q] assms(3) by blast
text ‹Steps:
1) Remove the hj and add empty-guards
2) Apply sigma
3) Remove the guards and add hj, using S-after-share›
define input_Σ where "input_Σ = (λσ. { (fst σ, add_empty_guards hp) |hp hj. Some (snd σ) = Some hp ⊕ Some hj ∧
(fst σ, hp), (fst σ, hp) ⊨ P ∧ sat_inv (fst σ) hj Γ})"
define Σ' where "Σ' = (λσ. ⋃p ∈ input_Σ σ. S_after_share (Σ p) Γ (f (normalize (get_fh (snd σ)))))"
show ?thesis
proof (rule hoare_triple_validI_bounded)
show "⋀s h n. (s, h), (s, h) ⊨ Star P (LowView (α ∘ f) J x) ⟹ bounded h ⟹ safe n (None :: ('i, 'a, nat) cont) C (s, h) (Σ' (s, h))"
proof -
fix s h n assume asm1: "(s, h), (s, h) ⊨ Star P (LowView (α ∘ f) J x)" "bounded h"
then obtain hp hj where "no_guard h" "Some h = Some hp ⊕ Some hj" "(s, hp), (s, hp) ⊨ P"
"(s, hj), (s, hj) ⊨ LowView (α ∘ f) J x"
by (meson always_sat_refl assms(8) hyper_sat.simps(4) no_guard_assertion_def)
then have "sat_inv s hj Γ"
by (metis LowViewE assms(1) assms(7) no_guard_then_smaller_same plus_comm sat_inv_def select_convs(5))
then have "(s, add_empty_guards hp) ∈ input_Σ (s, h)"
using ‹(s, hp), (s, hp) ⊨ P› ‹Some h = Some hp ⊕ Some hj› input_Σ_def by force
let ?v0 = "f (normalize (get_fh h))"
let ?p = "(s, add_empty_guards hp)"
have "safe n (None :: ('i, 'a, nat) cont) C (s, remove_guards (add_empty_guards h)) (S_after_share (Σ ?p) Γ ?v0)"
proof (rule share_lemma)
show "safe n (Some Γ) C ?p (Σ ?p)"
proof (rule asm0(1))
show "(s, add_empty_guards hp), (s, add_empty_guards hp) ⊨ Star P EmptyFullGuards"
using ‹(s, hp), (s, hp) ⊨ P› ‹Some h = Some hp ⊕ Some hj› ‹no_guard h› no_guard_and_sat_p_empty_guards no_guard_then_smaller_same by blast
show "bounded (snd (s, add_empty_guards hp))"
unfolding bounded_def add_empty_guards_def apply simp
by (metis ‹Some h = Some hp ⊕ Some hj› asm1(2) boundedE bounded_smaller_sum fst_eqD)
qed
show "Some (add_empty_guards h) = Some (add_empty_guards hp) ⊕ Some hj"
using ‹Some h = Some hp ⊕ Some hj› ‹no_guard h› no_guard_add_empty_guards_sum by blast
show "sat_inv s hj Γ"
using ‹sat_inv s hj Γ› by auto
show "view_function_of_inv Γ"
by (simp add: assms(4))
show "semi_consistent Γ (f (FractionalHeap.normalize (get_fh h))) (add_empty_guards h)"
by (metis ‹no_guard h› assms(1) select_convs(1) semi_consistent_empty_no_guard_initial_value)
show "bounded (add_empty_guards h)"
unfolding bounded_def add_empty_guards_def apply simp
by (metis asm1(2) boundedE fst_eqD)
qed
moreover have "(S_after_share (Σ ?p) Γ ?v0) ⊆ Σ' (s, h)"
using Σ'_def ‹(s, add_empty_guards hp) ∈ input_Σ (s, h)› by auto
ultimately show "safe n (None :: ('i, 'a, nat) cont) C (s, h) (Σ' (s, h))"
by (metis ‹no_guard h› no_guards_remove_same safe_larger_set)
qed
fix s1 h1 s2 h2
assume "(s1, h1), (s2, h2) ⊨ Star P (LowView (α ∘ f) J x)"
then obtain hp1 hj1 hp2 hj2 where asm1: "Some h1 = Some hp1 ⊕ Some hj1"
"Some h2 = Some hp2 ⊕ Some hj2" "(s1, hp1), (s2, hp2) ⊨ P" "no_guard h1" "no_guard h2"
"(s1, hj1), (s2, hj2) ⊨ LowView (α ∘ f) J x"
using assms(8) hyper_sat.simps(4) no_guard_assertion_def by blast
then obtain "(s1, hj1), (s2, hj2) ⊨ J" "α (f (normalize (get_fh hj1))) = α (f (normalize (get_fh hj2)))"
by (metis LowViewE assms(7) comp_apply)
show "pair_sat (Σ' (s1, h1)) (Σ' (s2, h2)) (Star Q (LowView (α ∘ f) J x))"
proof (rule pair_satI)
fix s1' h1' s2' h2'
assume asm2: "(s1', h1') ∈ Σ' (s1, h1) ∧ (s2', h2') ∈ Σ' (s2, h2)"
then obtain p1 p2 where p_assms: "p1 ∈ input_Σ (s1, h1)" "p2 ∈ input_Σ (s2, h2)"
"(s1', h1') ∈ S_after_share (Σ p1) Γ (f (normalize (get_fh h1)))"
"(s2', h2') ∈ S_after_share (Σ p2) Γ (f (normalize (get_fh h2)))"
using Σ'_def by force
moreover have "pair_sat (Σ p1) (Σ p2) (Star Q (And (PreSharedGuards (Abs_precondition spre)) (PreUniqueGuards (Abs_indexed_precondition upre))))"
proof (rule asm0(2))
obtain hj1' hj2' hp1' hp2' where "snd p1 = add_empty_guards hp1'" "snd p2 = add_empty_guards hp2'"
"Some h1 = Some hp1' ⊕ Some hj1'" "Some h2 = Some hp2' ⊕ Some hj2'" "sat_inv s1 hj1' Γ" "sat_inv s2 hj2' Γ"
"fst p1 = s1" "fst p2 = s2"
using p_assms(1) p_assms(2) input_Σ_def by auto
moreover have "hj1 = hj1' ∧ hj2 = hj2'"
proof (rule preciseE)
show "precise J"
by (simp add: assms(5))
show "h1 ≽ hj1' ∧ h1 ≽ hj1 ∧ h2 ≽ hj2' ∧ h2 ≽ hj2"
by (metis asm1(1) asm1(2) calculation(3) calculation(4) larger_def plus_comm)
show "(s1, hj1'), (s2, hj2') ⊨ J ∧ (s1, hj1), (s2, hj2) ⊨ J"
by (metis ‹(s1, hj1), (s2, hj2) ⊨ J› assms(1) assms(5) calculation(5) calculation(6) sat_inv_def select_convs(5) unaryE)
qed
then have "hp1 = hp1' ∧ hp2 = hp2'"
using addition_cancellative asm1(1) asm1(2) calculation(3) calculation(4) by blast
then show "p1, p2 ⊨ Star P EmptyFullGuards"
using no_guard_and_sat_p_empty_guards[of "fst p1" "snd p1" "fst p2" "snd p2" P]
by (metis asm1(3) asm1(4) asm1(5) calculation(1) calculation(2) calculation(3) calculation(4) calculation(7) calculation(8) no_guard_and_sat_p_empty_guards no_guard_then_smaller_same prod.exhaust_sel)
qed
let ?v1 = "f (normalize (get_fh h1))"
let ?v2 = "f (normalize (get_fh h2))"
obtain hj1' hg1 H1 hj2' hg2 H2 where asm3: "h1' = remove_guards H1" "semi_consistent Γ ?v1 H1"
"Some H1 = Some hg1 ⊕ Some hj1'" "(s1', hg1) ∈ Σ p1" "sat_inv s1' hj1' Γ"
"h2' = remove_guards H2" "semi_consistent Γ ?v2 H2"
"Some H2 = Some hg2 ⊕ Some hj2'" "(s2', hg2) ∈ Σ p2" "sat_inv s2' hj2' Γ"
using p_assms(3) S_after_share_def[of "Σ p1" Γ ?v1] p_assms(4) S_after_share_def[of "Σ p2" Γ ?v2] by blast
then have "(s1', hg1), (s2', hg2) ⊨ Star Q (And (PreSharedGuards (Abs_precondition spre)) (PreUniqueGuards (Abs_indexed_precondition upre)))"
using ‹pair_sat (Σ p1) (Σ p2) (Star Q (And (PreSharedGuards (Abs_precondition spre)) (PreUniqueGuards (Abs_indexed_precondition upre))))› pair_satE by blast
then obtain q1 g1 q2 g2 where "Some hg1 = Some q1 ⊕ Some g1" "Some hg2 = Some q2 ⊕ Some g2"
"(s1', q1), (s2', q2) ⊨ Q" "(s1', g1), (s2', g2) ⊨ PreSharedGuards (Abs_precondition spre)" "(s1', g1), (s2', g2) ⊨ PreUniqueGuards (Abs_indexed_precondition upre)"
by (meson hyper_sat.simps(3) hyper_sat.simps(4))
moreover have "Rep_precondition (Abs_precondition spre) = spre ∧ Rep_indexed_precondition (Abs_indexed_precondition upre) = upre"
by (simp add: assms(6) wf_indexed_precondition_rep_prec wf_precondition_rep_prec)
ultimately obtain sargs1 sargs2 where
"get_gs g1 = Some (pwrite, sargs1)" "get_gs g2 = Some (pwrite, sargs2)" "PRE_shared_simpler spre sargs1 sargs2"
"get_fh g1 = Map.empty" "get_fh g2 = Map.empty"
by auto
moreover obtain uargs1 uargs2 where
unique_facts: "⋀k. get_gu g1 k = Some (uargs1 k) ∧ get_gu g2 k = Some (uargs2 k) ∧ PRE_unique (upre k) (uargs1 k) (uargs2 k)"
using sat_PreUniqueE[OF ‹(s1', g1), (s2', g2) ⊨ PreUniqueGuards (Abs_indexed_precondition upre)›]
by (metis ‹Rep_precondition (Abs_precondition spre) = spre ∧ Rep_indexed_precondition (Abs_indexed_precondition upre) = upre›)
moreover obtain "get_gs H1 = Some (pwrite, sargs1)" "⋀k. get_gu H1 k = Some (uargs1 k)"
by (metis (no_types, opaque_lifting) ‹Some hg1 = Some q1 ⊕ Some g1› asm3(3) calculation(1) calculation(6) full_sguard_sum_same full_uguard_sum_same plus_comm)
then have reach1: "reachable_value sact uact ?v1 sargs1 uargs1 (f (normalize (get_fh H1)))"
by (metis asm3(2) assms(1) reachableE select_convs(1) select_convs(3) select_convs(4) semi_consistent_def)
moreover obtain "get_gs H2 = Some (pwrite, sargs2)" "⋀k. get_gu H2 k = Some (uargs2 k)"
by (metis (no_types, lifting) ‹Some hg2 = Some q2 ⊕ Some g2› asm3(8) calculation(2) calculation(6) full_sguard_sum_same full_uguard_sum_same plus_comm)
then have reach2: "reachable_value sact uact ?v2 sargs2 uargs2 (f (normalize (get_fh H2)))"
by (metis asm3(7) assms(1) reachableE semi_consistent_def simps(1) simps(3) simps(4))
moreover have "α (f (normalize (get_fh h1))) = α (f (normalize (get_fh hj1)))"
using view_function_of_invE[of Γ s1 hj1 h1]
by (metis ‹(s1, hj1), (s2, hj2) ⊨ J› always_sat_refl asm1(1) asm1(4) assms(1) assms(4) larger_def no_guard_then_smaller_same plus_comm sat_inv_def select_convs(1) select_convs(5))
moreover have "α (f (normalize (get_fh h2))) = α (f (normalize (get_fh hj2)))"
using view_function_of_invE[of Γ s2 hj2 h2]
by (metis ‹(s1, hj1), (s2, hj2) ⊨ J› always_sat_refl asm1(2) asm1(5) assms(1) assms(4) larger_def no_guard_then_smaller_same plus_comm sat_comm sat_inv_def select_convs(1) select_convs(5))
ultimately have low_abstract_view: "α (f (FractionalHeap.normalize (get_fh H1))) = α (f (FractionalHeap.normalize (get_fh H2)))"
using reach1 reach2 main_result[of sact uact ?v1 sargs1 uargs1 "f (normalize (get_fh H1))" ?v2 sargs2 uargs2 "f (normalize (get_fh H2))" spre upre α]
using ‹α (f (FractionalHeap.normalize (get_fh hj1))) = α (f (FractionalHeap.normalize (get_fh hj2)))› assms(2) by presburger
moreover have "α (f (normalize (get_fh H1))) = α (f (normalize (get_fh hj1')))"
using view_function_of_invE[of Γ s1' hj1' H1]
by (metis asm3(3) asm3(5) assms(1) assms(4) larger_def plus_comm select_convs(1))
moreover have "α (f (normalize (get_fh H2))) = α (f (normalize (get_fh hj2')))"
using view_function_of_invE[of Γ s2' hj2' H2]
by (metis asm3(10) asm3(8) assms(1) assms(4) larger_def plus_comm select_convs(1))
moreover have "(s1', hj1'), (s2', hj2') ⊨ J"
by (metis asm3(10) asm3(5) assms(1) assms(5) sat_inv_def select_convs(5) unaryE)
ultimately have "(s1', hj1'), (s2', hj2') ⊨ LowView (α ∘ f) J x"
by (simp add: LowViewI assms(7))
moreover have "Some h1' = Some q1 ⊕ Some hj1'"
proof -
have "Some h1' = Some (remove_guards hg1) ⊕ Some (remove_guards hj1')"
using asm3(1) asm3(3) remove_guards_sum by blast
moreover have "remove_guards hg1 = remove_guards q1"
by (metis ‹Some hg1 = Some q1 ⊕ Some g1› ‹get_fh g1 = Map.empty› get_fh_remove_guards no_guard_and_no_heap no_guard_remove_guards no_guards_remove remove_guards_sum)
moreover have "remove_guards hj1' = hj1'"
by (metis asm3(5) no_guards_remove sat_inv_def)
ultimately show ?thesis
by (metis ‹Some hg1 = Some q1 ⊕ Some g1› ‹get_gs g1 = Some (pwrite, sargs1)› unique_facts all_guards_def full_guard_comp_then_no no_guards_remove option.distinct(1) plus.simps(3) plus_comm)
qed
moreover have "Some h2' = Some q2 ⊕ Some hj2'"
proof -
have "Some h2' = Some (remove_guards hg2) ⊕ Some (remove_guards hj2')"
using asm3(6) asm3(8) remove_guards_sum by blast
moreover have "remove_guards hg2 = remove_guards q2"
by (metis ‹Some hg2 = Some q2 ⊕ Some g2› ‹get_fh g2 = Map.empty› get_fh_remove_guards no_guard_and_no_heap no_guard_remove_guards no_guards_remove remove_guards_sum)
moreover have "remove_guards hj2' = hj2'"
by (metis asm3(10) no_guards_remove sat_inv_def)
ultimately show ?thesis
by (metis ‹Some hg2 = Some q2 ⊕ Some g2› ‹get_gs g2 = Some (pwrite, sargs2)› unique_facts all_guards_def full_guard_comp_then_no no_guards_remove option.distinct(1) plus.simps(3) plus_comm)
qed
ultimately show "(s1', h1'), (s2', h2') ⊨ Star Q (LowView (α ∘ f) J x)"
by (meson LowViewI ‹(s1', q1), (s2', q2) ⊨ Q› assms(7) hyper_sat.simps(9) hyper_sat.simps(4))
qed
qed
qed
subsubsection ‹Atomic›
lemma red_rtrans_induct:
assumes "red_rtrans C σ C' σ'"
and "⋀C σ. P C σ C σ"
and "⋀C σ C' σ' C'' σ''. red C σ C' σ' ⟹ red_rtrans C' σ' C'' σ'' ⟹ P C' σ' C'' σ'' ⟹ P C σ C'' σ''"
shows "P C σ C' σ'"
using red_red_rtrans.inducts[of _ _ _ _ "λ_ _ _ _. True" P] assms by auto
lemma safe_atomic:
assumes "red_rtrans C1 σ1 C2 σ2"
and "σ1 = (s1, H1)"
and "σ2 = (s2, H2)"
and "⋀n. safe n (None :: ('i, 'a, nat) cont) C1 (s1, h) S"
and "H = denormalize H1"
and "Some H = Some h ⊕ Some hf"
and "full_ownership (get_fh H) ∧ no_guard H"
shows "¬ aborts C2 σ2 ∧ (C2 = Cskip ⟶
(∃h1 H'. Some H' = Some h1 ⊕ Some hf ∧ H2 = normalize (get_fh (H')) ∧ no_guard H' ∧ full_ownership (get_fh H') ∧ (s2, h1) ∈ S))"
using assms
proof (induction arbitrary: s1 H1 H h rule: red_rtrans_induct[of C1 σ1 C2 σ2])
case 1
then show ?case by (simp add: assms(1))
next
case (2 C σ)
then have "¬ aborts C (s1, FractionalHeap.normalize (get_fh H))"
using no_abortE(1) safe.simps(2) by blast
then have "¬ aborts C σ"
by (metis "2.prems"(2) "2.prems"(5) denormalize_properties(3))
moreover have "safe (Suc 1) (None :: ('i, 'a, nat) cont) C (s1, h) S"
using "2.prems"(4) by blast
then have "C = Cskip ⟹ (s2, h) ∈ S"
by (metis "2.prems"(2) "2.prems"(3) Pair_inject safeNoneE(1))
then have "C = Cskip ⟹ Some H = Some h ⊕ Some hf ∧ H2 = FractionalHeap.normalize (get_fh H) ∧ no_guard H ∧ full_ownership (get_fh H) ∧ (s2, h) ∈ S"
by (metis "2.prems"(3) "2.prems"(2) "2.prems"(6) "2.prems"(5) "2.prems"(7) denormalize_properties(3) old.prod.inject)
ultimately show ?case
by blast
next
case (3 C σ C' σ' C'' σ'')
obtain s0 H0 where "σ' = (s0, H0)" using prod.exhaust_sel by blast
have "safe (Suc 0) (None :: ('i, 'a, nat) cont) C (s1, h) S"
using "3.prems"(4) by force
then have "∃h'' H'. full_ownership (get_fh H') ∧ no_guard H' ∧ H0 = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hf ∧ safe 0 (None :: ('i, 'a, nat) cont) C' (s0, h'') S"
proof (rule safeNoneE(3)[of 0 C s1 h S H hf C' s0 H0])
show "Some H = Some h ⊕ Some hf" using "3.prems"(6) by blast
show "full_ownership (get_fh H)" using "3.prems"(7) by blast
show "no_guard H" using "3.prems"(7) by auto
show "red C (s1, FractionalHeap.normalize (get_fh H)) C' (s0, H0)"
by (metis "3.hyps"(1) "3.prems"(2) "3.prems"(5) ‹σ' = (s0, H0)› denormalize_properties(3))
qed
then obtain h0 H0' where
r1: "full_ownership (get_fh H0') ∧ no_guard H0' ∧ H0 = FractionalHeap.normalize (get_fh H0') ∧ Some H0' = Some h0 ⊕ Some hf ∧ safe 0 (None :: ('i, 'a, nat) cont) C' (s0, h0) S"
by blast
then have "Some (denormalize H0) = Some h0 ⊕ Some hf"
by (metis denormalize_properties(4))
have ih:
"¬ aborts C'' σ'' ∧ (C'' = Cskip ⟶
(∃h1 H'. Some H' = Some h1 ⊕ Some hf ∧ H2 = FractionalHeap.normalize (get_fh H') ∧ no_guard H' ∧ full_ownership (get_fh H') ∧ (s2, h1) ∈ S))"
proof (rule 3(3)[of s0 H0 h0 H0'])
show "σ' = (s0, H0)" by (simp add: ‹σ' = (s0, H0)›)
show "σ'' = (s2, H2)"
by (simp add: "3.prems"(3))
show "H0' = denormalize H0" by (metis denormalize_properties(4) r1)
show "Some H0' = Some h0 ⊕ Some hf" using r1 by blast
show "full_ownership (get_fh H0') ∧ no_guard H0'" using r1 by blast
show "red_rtrans C' σ' C'' σ''"
by (simp add: "3.hyps"(2))
fix n
have "safe (Suc n) (None :: ('i, 'a, nat) cont) C (s1, h) S"
using "3.prems"(4) by force
then have "∃h'' H'. full_ownership (get_fh H') ∧ no_guard H' ∧ H0 = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hf ∧ safe n (None :: ('i, 'a, nat) cont) C' (s0, h'') S"
proof (rule safeNoneE(3)[of n C s1 h S H hf C' s0 H0])
show "Some H = Some h ⊕ Some hf" using "3.prems"(6) by blast
show "full_ownership (get_fh H)" using "3.prems"(7) by blast
show "no_guard H" using "3.prems"(7) by auto
show "red C (s1, FractionalHeap.normalize (get_fh H)) C' (s0, H0)"
by (metis "3.hyps"(1) "3.prems"(2) "3.prems"(5) ‹σ' = (s0, H0)› denormalize_properties(3))
qed
then obtain h3 H3' where
r2: "full_ownership (get_fh H3') ∧ no_guard H3' ∧ H0 = FractionalHeap.normalize (get_fh H3') ∧ Some H3' = Some h3 ⊕ Some hf ∧ safe n (None :: ('i, 'a, nat) cont) C' (s0, h3) S"
by blast
then have "h3 = h0"
by (metis ‹Some (denormalize H0) = Some h0 ⊕ Some hf› addition_cancellative denormalize_properties(4))
moreover have "H3' = H0'"
by (metis ‹Some H0' = Some h0 ⊕ Some hf› calculation option.inject r2)
ultimately show "safe n (None :: ('i, 'a, nat) cont) C' (s0, h0) S" using r2 by blast
qed
then show ?case by blast
qed
theorem atomic_rule_unique:
fixes Γ :: "('i, 'a, nat) single_context"
fixes map_to_list :: "nat ⇒ 'a list"
fixes map_to_arg :: "nat ⇒ 'a"
assumes "Γ = ⦇ view = f, abstract_view = α, saction = sact, uaction = uact, invariant = J ⦈"
and "hoare_triple_valid (None :: ('i, 'a, nat) cont) (Star P (View f J (λs. s x)))
C (Star Q (View f J (λs. uact index (s x) (map_to_arg (s uarg)))))"
and "precise J ∧ unary J"
and "view_function_of_inv Γ"
and "x ∉ fvC C ∪ fvA P ∪ fvA Q ∪ fvA J"
and "uarg ∉ fvC C"
and "l ∉ fvC C"
and "x ∉ fvS (λs. map_to_list (s l))"
and "x ∉ fvS (λs. map_to_arg (s uarg) # map_to_list (s l))"
and "no_guard_assertion P"
and "no_guard_assertion Q"
shows "hoare_triple_valid (Some Γ) (Star P (UniqueGuard index (λs. map_to_list (s l)))) (Catomic C)
(Star Q (UniqueGuard index (λs. map_to_arg (s uarg) # map_to_list (s l))))"
proof -
let ?J = "View f J (λs. s x)"
let ?J' = "View f J (λs. uact index (s x) (map_to_arg (s uarg)))"
let ?pre_l = "(λs. map_to_list (s l))"
let ?G = "UniqueGuard index ?pre_l"
let ?l = "λs. map_to_arg (s uarg) # map_to_list (s l)"
let ?G' = "UniqueGuard index ?l"
have unaries: "unary ?J ∧ unary ?J'"
by (simp add: assms(3) unary_inv_then_view)
moreover have precises: "precise ?J ∧ precise ?J'"
by (simp add: assms(3) precise_inv_then_view)
obtain Σ where asm0: "⋀n σ. σ, σ ⊨ Star P ?J ⟹ bounded (snd σ) ⟹ safe n (None :: ('i, 'a, nat) cont) C σ (Σ σ)"
"⋀σ σ'. σ, σ' ⊨ Star P ?J ⟹ pair_sat (Σ σ) (Σ σ') (Star Q ?J')"
using assms(2) hoare_triple_valid_def by blast
define start where "start = (λσ. { (s, h) |s h hj. agrees (- {x}) (fst σ) s ∧ Some h = Some (remove_guards (snd σ)) ⊕ Some hj ∧ (s, hj), (s, hj) ⊨ ?J})"
define end_qj where "end_qj = (λσ. ⋃σ' ∈ start σ. Σ σ')"
define Σ' where "Σ' = (λσ. { (s, add_uguard_to_no_guard index hq (?l s)) |s hq h hj. (s, h) ∈ end_qj σ ∧ Some h = Some hq ⊕ Some hj ∧ (s, hj), (s, hj) ⊨ ?J' })"
let ?Σ' = "λσ. close_var (Σ' σ) x"
show "hoare_triple_valid (Some Γ) (Star P ?G) (Catomic C) (Star Q ?G')"
proof (rule hoare_triple_validI_bounded)
show "⋀s h s' h'. (s, h), (s', h') ⊨ Star P ?G ⟹ pair_sat (?Σ' (s, h)) (?Σ' (s', h')) (Star Q ?G')"
proof -
fix s1 h1 s2 h2
assume asm1: "(s1, h1), (s2, h2) ⊨ Star P ?G"
then obtain p1 p2 g1 g2 where r0: "Some h1 = Some p1 ⊕ Some g1"
"Some h2 = Some p2 ⊕ Some g2"
"(s1, p1), (s2, p2) ⊨ P" "(s1, g1), (s2, g2) ⊨ ?G"
using hyper_sat.simps(4) by auto
then obtain "remove_guards h1 = p1" "remove_guards h2 = p2"
by (meson assms(10) hyper_sat.simps(13) no_guard_and_no_heap no_guard_assertion_def)
have "pair_sat (Σ' (s1, h1)) (Σ' (s2, h2)) (Star Q ?G')"
proof (rule pair_satI)
fix s1' hqg1 s2' hqg2 σ2'
assume asm2: "(s1', hqg1) ∈ Σ' (s1, h1) ∧ (s2', hqg2) ∈ Σ' (s2, h2)"
then obtain h1' hj1' h2' hj2' hq1 hq2 where r: "(s1', h1') ∈ end_qj (s1, h1)" "Some h1' = Some hq1 ⊕ Some hj1'"
"(s1', hj1'), (s1', hj1') ⊨ ?J'" "(s2', h2') ∈ end_qj (s2, h2)" "Some h2' = Some hq2 ⊕ Some hj2'" "(s2', hj2'), (s2', hj2') ⊨ ?J'"
"hqg1 = add_uguard_to_no_guard index hq1 (?l s1')" "hqg2 = add_uguard_to_no_guard index hq2 (?l s2')"
using Σ'_def by blast
then obtain σ1' σ2' where "σ1' ∈ start (s1, h1)" "σ2' ∈ start (s2, h2)" "(s1', h1') ∈ Σ σ1'" "(s2', h2') ∈ Σ σ2'"
using end_qj_def by blast
then obtain hj1 hj2 where "agrees (- {x}) s1 (fst σ1')" "Some (snd σ1') = Some p1 ⊕ Some hj1" "(fst σ1', hj1), (fst σ1', hj1) ⊨ ?J"
"agrees (- {x}) s2 (fst σ2')" "Some (snd σ2') = Some p2 ⊕ Some hj2" "(fst σ2', hj2), (fst σ2', hj2) ⊨ ?J"
using start_def ‹remove_guards h1 = p1› ‹remove_guards h2 = p2› by force
moreover have "(fst σ1', hj1), (fst σ2', hj2) ⊨ ?J"
using calculation(3) calculation(6) unaries unaryE by blast
moreover have "(fst σ1', p1), (fst σ2', p2) ⊨ P"
proof -
have "fvA P ⊆ - {x}"
using assms(5) by force
then have "agrees (fvA P) (fst σ1') s1 ∧ agrees (fvA P) (fst σ2') s2"
using calculation(1) calculation(4)
by (metis agrees_comm agrees_union subset_Un_eq)
then show ?thesis using r0(3)
by (meson agrees_same sat_comm)
qed
ultimately have "σ1', σ2' ⊨ Star P ?J" using hyper_sat.simps(4)[of "fst σ1'" "snd σ1'" "fst σ2'" "snd σ2'"] prod.collapse
by metis
then have "pair_sat (Σ σ1') (Σ σ2') (Star Q ?J')"
using asm0(2)[of σ1' σ2'] by blast
then have "(s1', h1'), (s2', h2') ⊨ Star Q ?J'"
using ‹(s1', h1') ∈ Σ σ1'› ‹(s2', h2') ∈ Σ σ2'› pair_sat_def by blast
moreover have "(s1', hj1'), (s2', hj2') ⊨ ?J'"
using r(3) r(6) unaries unaryE by blast
moreover have "(s1', hq1), (s2', hq2) ⊨ Q" using magic_lemma
using calculation(1) calculation(2) precises r(2) r(5) by blast
have "(s1', add_uguard_to_no_guard index hq1 (?l s1')), (s2', add_uguard_to_no_guard index hq2 (?l s2')) ⊨ Star Q ?G'"
proof (rule no_guard_then_sat_star_uguard)
show "no_guard hq1 ∧ no_guard hq2"
using ‹(s1', hq1), (s2', hq2) ⊨ Q› assms(11) no_guard_assertion_def by blast
show "(s1', hq1), (s2', hq2) ⊨ Q"
using ‹(s1', hq1), (s2', hq2) ⊨ Q› by auto
qed
then show "(s1', hqg1), (s2', hqg2) ⊨ Star Q ?G'"
using r(7) r(8) by force
qed
then show "pair_sat (?Σ' (s1, h1)) (?Σ' (s2, h2)) (Star Q ?G')"
proof (rule pair_sat_close_var_double)
show "x ∉ fvA (Star Q (UniqueGuard index (λs. map_to_arg (s uarg) # map_to_list (s l))))"
using assms(5) assms(9) by auto
qed
qed
fix pre_s h k
assume "(pre_s, h), (pre_s, h) ⊨ Star P ?G"
then obtain pp gg where "Some h = Some pp ⊕ Some gg" "(pre_s, pp), (pre_s, pp) ⊨ P" "(pre_s, gg), (pre_s, gg) ⊨ ?G"
using always_sat_refl hyper_sat.simps(4) by blast
then have "remove_guards h = pp"
using assms(10) hyper_sat.simps(13) no_guard_and_no_heap no_guard_assertion_def by metis
then have "(pre_s, remove_guards h), (pre_s, remove_guards h) ⊨ P"
using ‹(pre_s, pp), (pre_s, pp) ⊨ P› hyper_sat.simps(9) by blast
then have "(pre_s, remove_guards h), (pre_s, remove_guards h) ⊨ P"
by (simp add: no_guard_remove_guards)
show "safe k (Some Γ) (Catomic C) (pre_s, h) (?Σ' (pre_s, h))"
proof (cases k)
case (Suc n)
moreover have "safe (Suc n) (Some Γ) (Catomic C) (pre_s, h) (?Σ' (pre_s, h))"
proof (rule safeSomeAltI)
show "Catomic C = Cskip ⟹ (pre_s, h) ∈ ?Σ' (pre_s, h)" by simp
fix H hf hj v0
assume asm2: "Some H = Some h ⊕ Some hj ⊕ Some hf ∧ full_ownership (get_fh H) ∧ semi_consistent Γ v0 H ∧ sat_inv pre_s hj Γ"
define v where "v = f (normalize (get_fh H))"
define s where "s = pre_s(x := v)"
then have "v = s x" by simp
moreover have agreements: "agrees (fvC C ∪ fvA P ∪ fvA Q ∪ fvA J ∪ fvA (UniqueGuard k ?pre_l)) s pre_s"
by (metis UnE agrees_comm agrees_update assms(5) assms(8) fvA.simps(9) s_def)
have asm1: "(s, h), (s, h) ⊨ Star P ?G"
using Un_iff[of x] ‹(pre_s, h), (pre_s, h) ⊨ Star P (UniqueGuard index (λs. map_to_list (s l)))›
agrees_same agrees_update[of x] always_sat_refl assms(5) assms(8) fvA.simps(3)[of P "UniqueGuard index (λs. map_to_list (s l))"]
fvA.simps(9)[of index " (λs. map_to_list (s l))"] s_def
by metis
moreover have asm2_bis: "sat_inv s hj Γ"
proof (rule sat_inv_agrees)
show "sat_inv pre_s hj Γ" using asm2 by simp
show "agrees (fvA (invariant Γ)) pre_s s"
using assms(1) assms(5) s_def
by (simp add: agrees_update)
qed
moreover have "(s, remove_guards h), (s, remove_guards h) ⊨ P"
by (meson ‹(pre_s, remove_guards h), (pre_s, remove_guards h) ⊨ P› agreements agrees_same agrees_union always_sat_refl)
moreover have "agrees (- {x}) pre_s s"
proof (rule agreesI)
fix y assume "y ∈ - {x}"
then have "y ≠ x"
by force
then show "pre_s y = s y"
by (simp add: s_def)
qed
moreover obtain "(s, pp), (s, pp) ⊨ P" "(s, gg), (s, gg) ⊨ ?G"
by (metis ‹(pre_s, gg), (pre_s, gg) ⊨ UniqueGuard index (λs. map_to_list (s l))› ‹remove_guards h = pp› agrees_same_aux agrees_update always_sat_refl_aux assms(8) calculation(4) fvA.simps(9) s_def)
let ?hf = "remove_guards hf"
let ?H = "remove_guards H"
let ?h = "remove_guards h"
obtain hhj where "Some hhj = Some h ⊕ Some hj"
by (metis asm2 plus.simps(2) plus.simps(3) plus_comm)
then have "Some H = Some hhj ⊕ Some hf"
using asm2 by presburger
then have "Some (remove_guards hhj) = Some ?h ⊕ Some hj"
by (metis ‹Some hhj = Some h ⊕ Some hj› asm2 no_guards_remove remove_guards_sum sat_inv_def)
moreover have "f (normalize (get_fh hj)) = v"
proof -
have "view Γ (normalize (get_fh hj)) = view Γ (normalize (get_fh H))"
using assms(4) view_function_of_invE
by (metis (no_types, opaque_lifting) ‹Some hhj = Some h ⊕ Some hj› asm2 larger_def larger_trans plus_comm)
then show ?thesis using assms(1) v_def by fastforce
qed
then have "(s, hj), (s, hj) ⊨ ?J"
by (metis ‹v = s x› asm2_bis assms(1) hyper_sat.simps(11) sat_inv_def select_convs(5))
ultimately have "(s, remove_guards hhj), (s, remove_guards hhj) ⊨ Star P ?J"
using ‹(s, remove_guards h), (s, remove_guards h) ⊨ P› hyper_sat.simps(4) by blast
moreover have "bounded hhj"
apply (rule bounded_smaller_sum[of H])
apply (metis asm2 full_ownership_then_bounded get_fh.simps)
using ‹Some H = Some hhj ⊕ Some hf› by blast
ultimately have all_safes: "⋀n. safe n (None :: ('i, 'a, nat) cont) C (s, remove_guards hhj) (Σ (s, remove_guards hhj))"
using asm0(1) unfolding bounded_def remove_guards_def by simp
then have "⋀σ1 H1 σ2 H2 s2 C2. red_rtrans C σ1 C2 σ2 ⟹ σ1 = (s, H1) ⟹ σ2 = (s2, H2) ⟹
?H = denormalize H1 ⟹
¬ aborts C2 σ2 ∧ (C2 = Cskip ⟶ (∃h1 H'. Some H' = Some h1 ⊕ Some ?hf ∧ H2 = FractionalHeap.normalize (get_fh H')
∧ no_guard H' ∧ full_ownership (get_fh H') ∧ (s2, h1) ∈ Σ (s, remove_guards hhj)))"
proof -
fix σ1 H1 σ2 H2 s2 C2
assume "?H = denormalize H1"
assume "red_rtrans C σ1 C2 σ2" "σ1 = (s, H1)" "σ2 = (s2, H2)"
then show "¬ aborts C2 σ2 ∧
(C2 = Cskip ⟶
(∃h1 H'.
Some H' = Some h1 ⊕ Some (remove_guards hf) ∧
H2 = FractionalHeap.normalize (get_fh H') ∧ no_guard H' ∧ full_ownership (get_fh H') ∧ (s2, h1) ∈ Σ (s, remove_guards hhj)))"
using all_safes
proof (rule safe_atomic)
show "?H = denormalize H1" using ‹?H = denormalize H1› by simp
show "Some ?H = Some (remove_guards hhj) ⊕ Some ?hf"
using ‹Some H = Some hhj ⊕ Some hf› remove_guards_sum by blast
show "full_ownership (get_fh (remove_guards H)) ∧ no_guard (remove_guards H)"
by (metis asm2 get_fh_remove_guards no_guard_remove_guards)
qed
qed
moreover have "?H = denormalize (normalize (get_fh H))"
by (metis asm2 denormalize_properties(5))
ultimately have safe_atomic_simplified: "⋀σ2 H2 s2 C2. red_rtrans C (s, normalize (get_fh H)) C2 σ2
⟹ σ2 = (s2, H2) ⟹ ¬ aborts C2 σ2 ∧ (C2 = Cskip ⟶ (∃h1 H'. Some H' = Some h1 ⊕ Some ?hf ∧ H2 = FractionalHeap.normalize (get_fh H')
∧ no_guard H' ∧ full_ownership (get_fh H') ∧ (s2, h1) ∈ Σ (s, remove_guards hhj)))"
by presburger
have "¬ aborts (Catomic C) (s, normalize (get_fh H))"
proof (rule ccontr)
assume "¬ ¬ aborts (Catomic C) (s, normalize (get_fh H))"
then obtain C' σ' where asm3: "red_rtrans C (s, FractionalHeap.normalize (get_fh H)) C' σ'"
"aborts C' σ'"
using abort_atomic_cases by blast
then have "¬ aborts C' σ'" using safe_atomic_simplified[of C' σ' "fst σ'" "snd σ'"] by simp
then show False using asm3(2) by simp
qed
then show "¬ aborts (Catomic C) (pre_s, normalize (get_fh H))"
by (metis agreements aborts_agrees agrees_comm agrees_union fst_eqD fvC.simps(11) snd_conv)
fix C' pre_s' h'
assume "red (Catomic C) (pre_s, FractionalHeap.normalize (get_fh H)) C' (pre_s', h')"
then obtain s' where "red (Catomic C) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
"agrees (- {x}) s' pre_s'"
by (metis (no_types, lifting) UnI1 ‹agrees (- {x}) pre_s s› agrees_comm assms(5) fst_eqD fvC.simps(11) red_agrees snd_conv subset_Compl_singleton)
then obtain h1 H' where asm3: "Some H' = Some h1 ⊕ Some (remove_guards hf)" "C' = Cskip"
"h' = FractionalHeap.normalize (get_fh H')" "no_guard H' ∧ full_ownership (get_fh H')" "(s', h1) ∈ Σ (s, remove_guards hhj)"
using safe_atomic_simplified[of C' "(s', h')" s' h'] by (metis red_atomic_cases)
moreover have "s x = s' x ∧ s' uarg = s uarg ∧ s l = s' l" using red_not_in_fv_not_touched
using ‹red (Catomic C) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')›
by (metis Un_iff assms(5) assms(6) assms(7) fst_conv fvC.simps(11))
have "∃hq' hj'. Some h1 = Some hq' ⊕ Some hj' ∧ (s', add_uguard_to_no_guard index hq' (?l s')) ∈ Σ' (pre_s, h) ∧ sat_inv s' hj' Γ
∧ f (normalize (get_fh hj')) = uact index (s' x) (map_to_arg (s' uarg))"
proof -
have "pair_sat (Σ (s, remove_guards hhj)) (Σ (s, remove_guards hhj)) (Star Q ?J')"
using asm0(2)[of "(s, remove_guards hhj)" "(s, remove_guards hhj)"]
using ‹(s, remove_guards hhj), (s, remove_guards hhj) ⊨ Star P ?J› by blast
then have "(s', h1), (s', h1) ⊨ Star Q ?J'"
using asm3(5) pair_sat_def by blast
then obtain hq' hj' where "Some h1 = Some hq' ⊕ Some hj'" "(s', hq'), (s', hq') ⊨ Q" "(s', hj'), (s', hj') ⊨ ?J'"
using always_sat_refl hyper_sat.simps(4) by blast
then have "no_guard hj'"
by (metis (no_types, opaque_lifting) calculation(1) calculation(4) no_guard_then_smaller_same plus_comm)
moreover have "f (normalize (get_fh hj')) = uact index (s' x) (map_to_arg (s' uarg))"
using ‹(s', hj'), (s', hj') ⊨ View f J (λs. uact index (s x) (map_to_arg (s uarg)))› by auto
moreover have "(s, remove_guards hhj) ∈ start (pre_s, h)"
proof -
have "Some (remove_guards hhj) = Some ?h ⊕ Some hj"
using ‹Some (remove_guards hhj) = Some (remove_guards h) ⊕ Some hj› by blast
moreover have "(s, hj), (s, hj) ⊨ ?J"
using ‹(s, hj), (s, hj) ⊨ ?J› by fastforce
ultimately show ?thesis using start_def
using ‹agrees (- {x}) pre_s s› by fastforce
qed
then have "(s', h1) ∈ end_qj (pre_s, h)"
using ‹end_qj ≡ λσ. ⋃ (Σ ` start σ)› asm3(5) by blast
then have "(s', add_uguard_to_no_guard index hq' (?l s')) ∈ Σ' (pre_s, h)"
using Σ'_def ‹(s', hj'), (s', hj') ⊨ ?J'› ‹Some h1 = Some hq' ⊕ Some hj'› by blast
ultimately show "∃hq' hj'.
Some h1 = Some hq' ⊕ Some hj' ∧
(s', add_uguard_to_no_guard index hq' (map_to_arg (s' uarg) # map_to_list (s' l))) ∈ Σ' (pre_s, h) ∧
sat_inv s' hj' Γ ∧ f (FractionalHeap.normalize (get_fh hj')) = uact index (s' x) (map_to_arg (s' uarg))"
using ‹(s', hj'), (s', hj') ⊨ ?J'› ‹Some h1 = Some hq' ⊕ Some hj'›
assms(1) hyper_sat.simps(11) sat_inv_def select_convs(5)
by fastforce
qed
then obtain hq' hj' where "Some h1 = Some hq' ⊕ Some hj'" "(s', add_uguard_to_no_guard index hq' (?l s')) ∈ Σ' (pre_s, h)" "sat_inv s' hj' Γ"
"f (normalize (get_fh hj')) = uact index (s' x) (map_to_arg (s' uarg))"
by blast
then have "safe n (Some Γ) C' (s', add_uguard_to_no_guard index hq' (?l s')) (Σ' (pre_s, h))"
using asm3(2) safe_skip by blast
moreover have "∃H''. semi_consistent Γ v0 H'' ∧ Some H'' = Some (add_uguard_to_no_guard index hq' (?l s')) ⊕ Some hj' ⊕ Some hf"
proof -
have "Some (add_uguard_to_no_guard index hq' (?l s')) = Some hq' ⊕ Some (Map.empty, None, [index ↦ ?l s'])"
by (metis ‹Some h1 = Some hq' ⊕ Some hj'› add_uguard_as_sum calculation(1) calculation(4) no_guard_then_smaller_same)
obtain hhf where "Some hhf = Some h ⊕ Some hf"
by (metis (no_types, opaque_lifting) ‹Some H = Some hhj ⊕ Some hf› ‹Some hhj = Some h ⊕ Some hj› option.exhaust_sel plus.simps(1) plus_asso plus_comm)
then have "all_guards hhf"
by (metis (no_types, lifting) all_guards_no_guard_propagates asm2 plus_asso plus_comm sat_inv_def semi_consistent_def)
moreover have "get_gs h = None ∧ get_gu h index = Some (?pre_l s)"
proof -
have "no_guard pp"
using ‹remove_guards h = pp› no_guard_remove_guards by blast
then show ?thesis
by (metis (no_types, lifting) ‹Some h = Some pp ⊕ Some gg› ‹⋀thesis. (⟦(s, pp), (s, pp) ⊨ P; (s, gg), (s, gg) ⊨ UniqueGuard index (λs. map_to_list (s l))⟧ ⟹ thesis) ⟹ thesis› full_uguard_sum_same hyper_sat.simps(13) no_guard_remove(1) plus_comm)
qed
moreover have "⋀i'. i' ≠ index ⟹ get_gu h i' = None"
by (metis ‹Some h = Some pp ⊕ Some gg› ‹⋀thesis. (⟦(s, pp), (s, pp) ⊨ P; (s, gg), (s, gg) ⊨ UniqueGuard index (λs. map_to_list (s l))⟧ ⟹ thesis) ⟹ thesis› ‹remove_guards h = pp› hyper_sat.simps(13) no_guard_remove(2) no_guard_remove_guards plus_comm)
then obtain sargs where "get_gu hf index = None ∧ get_gs hf = Some (pwrite, sargs)"
by (metis (no_types, opaque_lifting) ‹Some hhf = Some h ⊕ Some hf› add_gs.simps(1) all_guards_def calculation(1) calculation(2) compatible_def compatible_eq option.distinct(1) plus_extract(2))
moreover obtain uargs where "⋀i'. i' ≠ index ⟹ get_gu hf i' = Some (uargs i')"
by (metis (no_types, opaque_lifting) ‹Some hhf = Some h ⊕ Some hf› ‹⋀i'. i' ≠ index ⟹ get_gu h i' = None› add_gu_def add_gu_single.simps(1) all_guards_exists_uargs calculation(1) plus_extract(3))
then obtain ghf where ghf_def: "Some hf = Some (remove_guards hf) ⊕ Some ghf"
"get_fh ghf = Map.empty" "get_gu ghf index = None"
"get_gs ghf = Some (pwrite, sargs)" "⋀i'. i' ≠ index ⟹ get_gu ghf i' = Some (uargs i')"
using decompose_guard_remove_easy[of hf]
using calculation(3) by auto
have "(Map.empty, None, [index ↦ ?l s']) ## ghf"
proof (rule compatibleI)
show "compatible_fract_heaps (get_fh (Map.empty, None, [index ↦ map_to_arg (s' uarg) # map_to_list (s' l)])) (get_fh ghf)"
using compatible_fract_heapsI by fastforce
show "⋀k. get_gu (Map.empty, None, [index ↦ map_to_arg (s' uarg) # map_to_list (s' l)]) k = None ∨ get_gu ghf k = None"
using ghf_def(3) by auto
qed (simp)
then obtain g where g_def: "Some g = Some (Map.empty, None, [index ↦ ?l s']) ⊕ Some ghf"
by simp
moreover have "H' ## g"
proof (rule compatibleI)
have "get_fh g = add_fh Map.empty Map.empty"
using add_get_fh[of g "(Map.empty, None, [index ↦ ?l s'])" ghf]
g_def ‹get_fh ghf = Map.empty›
by fastforce
then have "get_fh g = Map.empty"
using add_fh_map_empty by auto
then show "compatible_fract_heaps (get_fh H') (get_fh g)"
using compatible_fract_heapsI by force
show "⋀k. get_gu H' k = None ∨ get_gu g k = None"
by (meson asm3(4) no_guard_def)
show "⋀p p'. get_gs H' = Some p ∧ get_gs g = Some p' ⟹ pgte pwrite (padd (fst p) (fst p'))"
by (metis asm3(4) no_guard_def option.simps(3))
qed
then obtain H'' where "Some H'' = Some H' ⊕ Some g"
by simp
then have "Some H'' = Some (add_uguard_to_no_guard index hq' (?l s')) ⊕ Some hj' ⊕ Some hf"
proof -
have "Some H'' = Some h1 ⊕ Some g ⊕ Some (remove_guards hf)"
by (metis ‹Some H'' = Some H' ⊕ Some g› asm3(1) plus_comm simpler_asso)
moreover have "Some (add_uguard_to_no_guard index hq' (?l s')) = Some hq' ⊕ Some (Map.empty, None, [index ↦ ?l s'])"
using ‹Some (add_uguard_to_no_guard index hq' (map_to_arg (s' uarg) # map_to_list (s' l))) = Some hq' ⊕ Some (Map.empty, None, [index ↦ map_to_arg (s' uarg) # map_to_list (s' l)])› by blast
ultimately show ?thesis
by (metis (no_types, lifting) ‹Some h1 = Some hq' ⊕ Some hj'› g_def ghf_def(1) plus_comm simpler_asso)
qed
moreover have "semi_consistent Γ v0 H''"
proof (rule semi_consistentI)
have "get_gs g = Some (pwrite, sargs)"
by (metis full_sguard_sum_same g_def ghf_def(4) plus_comm)
moreover have "get_gu g index = Some (?l s')"
proof (rule full_uguard_sum_same)
show "get_gu (Map.empty, None, [index ↦ ?l s']) index = Some (?l s')"
using get_gu.simps by auto
show "Some g = Some (Map.empty, None, [index ↦ ?l s']) ⊕ Some ghf"
using g_def by auto
qed
moreover have "⋀i'. i' ≠ index ⟹ get_gu g i' = Some (uargs i')"
by (metis full_uguard_sum_same g_def ghf_def(5) plus_comm)
ultimately have "all_guards g"
by (metis all_guardsI option.discI)
then show "all_guards H''"
by (metis ‹Some H'' = Some H' ⊕ Some g› all_guards_same plus_comm)
show "reachable Γ v0 H''"
proof (rule reachableI)
fix sargs' uargs'
assume "get_gs H'' = Some (pwrite, sargs') ∧ (∀k. get_gu H'' k = Some (uargs' k))"
then have "sargs = sargs'"
by (metis (no_types, opaque_lifting) Pair_inject ‹Some H'' = Some H' ⊕ Some g› ‹get_gs g = Some (pwrite, sargs)› full_sguard_sum_same option.inject plus_comm)
moreover have "uargs' index = ?l s'"
by (metis ‹Some H'' = Some H' ⊕ Some g› ‹get_gs H'' = Some (pwrite, sargs') ∧ (∀k. get_gu H'' k = Some (uargs' k))› ‹get_gu g index = Some (map_to_arg (s' uarg) # map_to_list (s' l))› asm3(4) no_guard_remove(2) option.inject plus_comm)
moreover have "⋀i'. i' ≠ index ⟹ uargs' i' = uargs i'"
by (metis ‹Some H'' = Some H' ⊕ Some g› ‹⋀i'. i' ≠ index ⟹ get_gu g i' = Some (uargs i')› ‹get_gs H'' = Some (pwrite, sargs') ∧ (∀k. get_gu H'' k = Some (uargs' k))› asm3(4) no_guard_remove(2) option.sel plus_comm)
moreover have "view Γ (FractionalHeap.normalize (get_fh hj')) = view Γ (FractionalHeap.normalize (get_fh H''))"
using assms(4) ‹sat_inv s' hj' Γ›
proof (rule view_function_of_invE)
show "H'' ≽ hj'"
by (metis (no_types, opaque_lifting) ‹Some H'' = Some H' ⊕ Some g› ‹Some h1 = Some hq' ⊕ Some hj'› asm3(1) larger_def larger_trans plus_comm)
qed
moreover have "reachable_value (saction Γ) (uaction Γ) v0 sargs (uargs(index := ?l s')) (uact index (s' x) (map_to_arg (s' uarg)))"
proof -
have "reachable_value (saction Γ) (uaction Γ) v0 sargs (uargs(index := ?pre_l s')) (view Γ (FractionalHeap.normalize (get_fh H)))"
proof -
have "reachable Γ v0 H"
by (meson asm2 semi_consistent_def)
moreover have "get_gs H = Some (pwrite, sargs)"
by (metis ‹Some H = Some hhj ⊕ Some hf› ‹get_gu hf index = None ∧ get_gs hf = Some (pwrite, sargs)› full_sguard_sum_same plus_comm)
moreover have "get_gu H index = Some (?pre_l s')"
by (metis ‹Some H = Some hhj ⊕ Some hf› ‹Some hhj = Some h ⊕ Some hj› ‹get_gs h = None ∧ get_gu h index = Some (map_to_list (s l))› ‹s x = s' x ∧ s' uarg = s uarg ∧ s l = s' l› full_uguard_sum_same)
moreover have "⋀i. i ≠ index ⟹ get_gu H i = Some (uargs i)"
by (metis ‹Some H = Some hhj ⊕ Some hf› ‹⋀i'. i' ≠ index ⟹ get_gu hf i' = Some (uargs i')› full_uguard_sum_same plus_comm)
ultimately show ?thesis
by (simp add: reachable_def)
qed
moreover have "view Γ (FractionalHeap.normalize (get_fh hj)) = view Γ (FractionalHeap.normalize (get_fh H))"
using assms(4)
proof (rule view_function_of_invE)
show "sat_inv s hj Γ"
by (simp add: asm2_bis)
show "H ≽ hj"
by (metis (no_types, opaque_lifting) ‹Some H = Some hhj ⊕ Some hf› ‹Some hhj = Some h ⊕ Some hj› larger_def larger_trans plus_comm)
qed
moreover have "s' x = v"
using ‹s x = s' x ∧ s' uarg = s uarg ∧ s l = s' l› ‹v = s x› by presburger
ultimately have "reachable_value (saction Γ) (uaction Γ) v0 sargs (uargs(index := ?pre_l s')) v"
using ‹f (FractionalHeap.normalize (get_fh hj)) = v› assms(1) by auto
then show ?thesis
by (metis UniqueStep ‹s' x = v› assms(1) fun_upd_same fun_upd_upd select_convs(4))
qed
moreover have "uargs' = (uargs(index := map_to_arg (s' uarg) # map_to_list (s' l)))"
proof (rule ext)
fix i show "uargs' i = (uargs(index := map_to_arg (s' uarg) # map_to_list (s' l))) i"
apply (cases "i = index")
using calculation(2) apply auto[1]
using calculation(3) by force
qed
ultimately show "reachable_value (saction Γ) (uaction Γ) v0 sargs' uargs' (view Γ (FractionalHeap.normalize (get_fh H'')))"
using ‹f (FractionalHeap.normalize (get_fh hj')) = uact index (s' x) (map_to_arg (s' uarg))› assms(1) by force
qed
qed
ultimately show "∃H''. semi_consistent Γ v0 H'' ∧ Some H'' = Some (add_uguard_to_no_guard index hq' (map_to_arg (s' uarg) # map_to_list (s' l))) ⊕ Some hj' ⊕ Some hf"
by blast
qed
ultimately obtain H'' where "semi_consistent Γ v0 H'' ∧
Some H'' = Some (add_uguard_to_no_guard index hq' (map_to_arg (s' uarg) # map_to_list (s' l))) ⊕ Some hj' ⊕ Some hf" by blast
moreover have "full_ownership (get_fh H'') ∧ h' = FractionalHeap.normalize (get_fh H'')"
proof -
obtain x where "Some x = Some (add_uguard_to_no_guard index hq' (?l s')) ⊕ Some hj'"
by (metis calculation not_Some_eq plus.simps(1))
then have "get_fh H'' = add_fh (add_fh (get_fh (add_uguard_to_no_guard index hq' (?l s'))) (get_fh hj')) (get_fh hf)"
by (metis add_get_fh calculation)
moreover have "get_fh (add_uguard_to_no_guard index hq' (?l s')) = get_fh hq' ∧ get_fh hf = get_fh (remove_guards hf)"
by (metis get_fh_add_uguard get_fh_remove_guards)
ultimately show ?thesis
by (metis ‹Some h1 = Some hq' ⊕ Some hj'› add_get_fh asm3(1) asm3(3) asm3(4))
qed
moreover have "sat_inv pre_s' hj' Γ"
proof (rule sat_inv_agrees)
show "sat_inv s' hj' Γ"
by (simp add: ‹sat_inv s' hj' Γ›)
show "agrees (fvA (invariant Γ)) s' pre_s'"
using UnCI ‹agrees (- {x}) s' pre_s'› assms(1) assms(5) select_convs(5) subset_Compl_singleton
by (metis agrees_union sup.orderE)
qed
moreover have "safe n (Some Γ) C' (pre_s', add_uguard_to_no_guard index hq' (?l s')) (?Σ' (pre_s, h))"
proof (rule safe_free_vars_Some)
show "safe n (Some Γ) C' (s', add_uguard_to_no_guard index hq' (?l s')) (?Σ' (pre_s, h))"
by (meson ‹safe n (Some Γ) C' (s', add_uguard_to_no_guard index hq' (map_to_arg (s' uarg) # map_to_list (s' l))) (Σ' (pre_s, h))› close_var_subset safe_larger_set)
show "agrees (fvC C' ∪ (- {x}) ∪ fvA (invariant Γ)) s' pre_s'"
by (metis UnI2 Un_absorb1 ‹agrees (- {x}) s' pre_s'› asm3(2) assms(1) assms(5) empty_iff fvC.simps(1) inf_sup_aci(5) select_convs(5) subset_Compl_singleton)
show "upper_fvs (close_var (Σ' (pre_s, h)) x) (- {x})"
by (simp add: upper_fvs_close_vars)
qed
ultimately show "∃h'' H' hj'.
full_ownership (get_fh H') ∧
semi_consistent Γ v0 H' ∧
sat_inv pre_s' hj' Γ ∧ h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hj' ⊕ Some hf
∧ safe n (Some Γ) C' (pre_s', h'') (?Σ' (pre_s, h))" using ‹sat_inv s' hj' Γ› by blast
qed (simp)
ultimately show "safe k (Some Γ) (Catomic C) (pre_s, h) (?Σ' (pre_s, h))" by blast
qed (simp)
qed
qed
theorem atomic_rule_shared:
fixes Γ :: "('i, 'a, nat) single_context"
fixes map_to_multiset :: "nat ⇒ 'a multiset"
fixes map_to_arg :: "nat ⇒ 'a"
assumes "Γ = ⦇ view = f, abstract_view = α, saction = sact, uaction = uact, invariant = J ⦈"
and "hoare_triple_valid (None :: ('i, 'a, nat) cont) (Star P (View f J (λs. s x))) C
(Star Q (View f J (λs. sact (s x) (map_to_arg (s sarg)))))"
and "precise J ∧ unary J"
and "view_function_of_inv Γ"
and "x ∉ fvC C ∪ fvA P ∪ fvA Q ∪ fvA J"
and "sarg ∉ fvC C"
and "ms ∉ fvC C"
and "x ∉ fvS (λs. map_to_multiset (s ms))"
and "x ∉ fvS (λs. {# map_to_arg (s sarg) #} + map_to_multiset (s ms))"
and "no_guard_assertion P"
and "no_guard_assertion Q"
shows "hoare_triple_valid (Some Γ) (Star P (SharedGuard π (λs. map_to_multiset (s ms)))) (Catomic C)
(Star Q (SharedGuard π (λs. {# map_to_arg (s sarg) #} + map_to_multiset (s ms))))"
proof -
let ?J = "View f J (λs. s x)"
let ?J' = "View f J (λs. sact (s x) (map_to_arg (s sarg)))"
let ?pre_ms = "λs. map_to_multiset (s ms)"
let ?G = "SharedGuard π ?pre_ms"
let ?ms = "λs. {# map_to_arg (s sarg) #} + map_to_multiset (s ms)"
let ?G' = "SharedGuard π ?ms"
have unaries: "unary ?J ∧ unary ?J'"
by (simp add: assms(3) unary_inv_then_view)
moreover have precises: "precise ?J ∧ precise ?J'"
by (simp add: assms(3) precise_inv_then_view)
obtain Σ where asm0: "⋀n σ. σ, σ ⊨ Star P ?J ⟹ bounded (snd σ) ⟹ safe n (None :: ('i, 'a, nat) cont) C σ (Σ σ)"
"⋀σ σ'. σ, σ' ⊨ Star P ?J ⟹ pair_sat (Σ σ) (Σ σ') (Star Q ?J')"
using assms(2) hoare_triple_valid_def by blast
define start where "start = (λσ. { (s, h) |s h hj. agrees (- {x}) (fst σ) s ∧ Some h = Some (remove_guards (snd σ)) ⊕ Some hj ∧ (s, hj), (s, hj) ⊨ ?J})"
define end_qj where "end_qj = (λσ. ⋃σ' ∈ start σ. Σ σ')"
define Σ' where "Σ' = (λσ. { (s, add_sguard_to_no_guard hq π (?ms s)) |s hq h hj. (s, h) ∈ end_qj σ ∧ Some h = Some hq ⊕ Some hj ∧ (s, hj), (s, hj) ⊨ ?J' })"
let ?Σ' = "λσ. close_var (Σ' σ) x"
show "hoare_triple_valid (Some Γ) (Star P ?G) (Catomic C) (Star Q ?G')"
proof (rule hoare_triple_validI)
show "⋀s h s' h'. (s, h), (s', h') ⊨ Star P ?G ⟹ pair_sat (?Σ' (s, h)) (?Σ' (s', h')) (Star Q ?G')"
proof -
fix s1 h1 s2 h2
assume asm1: "(s1, h1), (s2, h2) ⊨ Star P ?G"
then obtain p1 p2 g1 g2 where r0: "Some h1 = Some p1 ⊕ Some g1"
"Some h2 = Some p2 ⊕ Some g2"
"(s1, p1), (s2, p2) ⊨ P" "(s1, g1), (s2, g2) ⊨ ?G"
using hyper_sat.simps(4) by auto
then obtain "remove_guards h1 = p1" "remove_guards h2 = p2"
using assms(10) hyper_sat.simps(12) no_guard_and_no_heap no_guard_assertion_def
by metis
have "pair_sat (Σ' (s1, h1)) (Σ' (s2, h2)) (Star Q ?G')"
proof (rule pair_satI)
fix s1' hqg1 s2' hqg2 σ2'
assume asm2: "(s1', hqg1) ∈ Σ' (s1, h1) ∧ (s2', hqg2) ∈ Σ' (s2, h2)"
then obtain h1' hj1' h2' hj2' hq1 hq2 where r: "(s1', h1') ∈ end_qj (s1, h1)" "Some h1' = Some hq1 ⊕ Some hj1'"
"(s1', hj1'), (s1', hj1') ⊨ ?J'" "(s2', h2') ∈ end_qj (s2, h2)" "Some h2' = Some hq2 ⊕ Some hj2'" "(s2', hj2'), (s2', hj2') ⊨ ?J'"
"hqg1 = add_sguard_to_no_guard hq1 π (?ms s1')" "hqg2 = add_sguard_to_no_guard hq2 π (?ms s2')"
using Σ'_def by blast
then obtain σ1' σ2' where "σ1' ∈ start (s1, h1)" "σ2' ∈ start (s2, h2)" "(s1', h1') ∈ Σ σ1'" "(s2', h2') ∈ Σ σ2'"
using end_qj_def by blast
then obtain hj1 hj2 where "agrees (- {x}) s1 (fst σ1')" "Some (snd σ1') = Some p1 ⊕ Some hj1" "(fst σ1', hj1), (fst σ1', hj1) ⊨ ?J"
"agrees (- {x}) s2 (fst σ2')" "Some (snd σ2') = Some p2 ⊕ Some hj2" "(fst σ2', hj2), (fst σ2', hj2) ⊨ ?J"
using start_def ‹remove_guards h1 = p1› ‹remove_guards h2 = p2› by force
moreover have "(fst σ1', hj1), (fst σ2', hj2) ⊨ ?J"
using calculation(3) calculation(6) unaries unaryE by blast
moreover have "(fst σ1', p1), (fst σ2', p2) ⊨ P"
proof -
have "fvA P ⊆ - {x}"
using assms(5) by force
then have "agrees (fvA P) (fst σ1') s1 ∧ agrees (fvA P) (fst σ2') s2"
using calculation(1) calculation(4)
by (metis agrees_comm agrees_union subset_Un_eq)
then show ?thesis using r0(3)
by (meson agrees_same sat_comm)
qed
ultimately have "σ1', σ2' ⊨ Star P ?J" using hyper_sat.simps(4)[of "fst σ1'" "snd σ1'" "fst σ2'" "snd σ2'"] prod.collapse
by metis
then have "pair_sat (Σ σ1') (Σ σ2') (Star Q ?J')"
using asm0(2)[of σ1' σ2'] by blast
then have "(s1', h1'), (s2', h2') ⊨ Star Q ?J'"
using ‹(s1', h1') ∈ Σ σ1'› ‹(s2', h2') ∈ Σ σ2'› pair_sat_def by blast
moreover have "(s1', hj1'), (s2', hj2') ⊨ ?J'"
using r(3) r(6) unaries unaryE by blast
moreover have "(s1', hq1), (s2', hq2) ⊨ Q" using magic_lemma
using calculation(1) calculation(2) precises r(2) r(5) by blast
moreover have "no_guard hq1 ∧ no_guard hq2"
using assms(11) calculation(3) no_guard_assertion_def by blast
ultimately show "(s1', hqg1), (s2', hqg2) ⊨ Star Q ?G'"
using no_guard_then_sat_star r(7) r(8)
by (metis (mono_tags, lifting))
qed
then show "pair_sat (?Σ' (s1, h1)) (?Σ' (s2, h2)) (Star Q ?G')"
proof (rule pair_sat_close_var_double)
show "x ∉ fvA (Star Q (SharedGuard π (λs. {#map_to_arg (s sarg)#} + map_to_multiset (s ms))))"
using assms(5) assms(9) by auto
qed
qed
fix pre_s h k
assume "(pre_s, h), (pre_s, h) ⊨ Star P ?G"
then obtain pp gg where "Some h = Some pp ⊕ Some gg" "(pre_s, pp), (pre_s, pp) ⊨ P" "(pre_s, gg), (pre_s, gg) ⊨ ?G"
using always_sat_refl hyper_sat.simps(4) by blast
then have "remove_guards h = pp"
by (meson assms(10) hyper_sat.simps(12) no_guard_and_no_heap no_guard_assertion_def)
then have "(pre_s, remove_guards h), (pre_s, remove_guards h) ⊨ P"
using ‹(pre_s, pp), (pre_s, pp) ⊨ P› hyper_sat.simps(9) by blast
then have "(pre_s, remove_guards h), (pre_s, remove_guards h) ⊨ P"
by (simp add: no_guard_remove_guards)
show "safe k (Some Γ) (Catomic C) (pre_s, h) (?Σ' (pre_s, h))"
proof (cases k)
case (Suc n)
moreover have "safe (Suc n) (Some Γ) (Catomic C) (pre_s, h) (?Σ' (pre_s, h))"
proof (rule safeSomeAltI)
show "Catomic C = Cskip ⟹ (pre_s, h) ∈ ?Σ' (pre_s, h)" by simp
fix H hf hj v0
assume asm2: "Some H = Some h ⊕ Some hj ⊕ Some hf ∧ full_ownership (get_fh H) ∧ semi_consistent Γ v0 H ∧ sat_inv pre_s hj Γ"
define v where "v = f (normalize (get_fh H))"
define s where "s = pre_s(x := v)"
then have "v = s x" by simp
moreover have agreements: "agrees (fvC C ∪ fvA P ∪ fvA Q ∪ fvA J ∪ fvA (SharedGuard π (λs. map_to_multiset (s ms)))) s pre_s"
by (metis (mono_tags, lifting) Un_iff agrees_def assms(5) assms(8) fun_upd_other fvA.simps(8) s_def)
then have asm1: "(s, h), (s, h) ⊨ Star P ?G"
by (metis (mono_tags, lifting) ‹(pre_s, h), (pre_s, h) ⊨ Star P (SharedGuard π (λs. map_to_multiset (s ms)))› agrees_same agrees_union fvA.simps(3) fvA.simps(8) sat_comm)
moreover have asm2_bis: "sat_inv s hj Γ"
proof (rule sat_inv_agrees)
show "sat_inv pre_s hj Γ" using asm2 by simp
show "agrees (fvA (invariant Γ)) pre_s s"
using assms(1) assms(5) s_def
by (simp add: agrees_update)
qed
moreover have "(s, remove_guards h), (s, remove_guards h) ⊨ P"
by (meson ‹(pre_s, remove_guards h), (pre_s, remove_guards h) ⊨ P› agreements agrees_same agrees_union always_sat_refl)
then have "(s, remove_guards h), (s, remove_guards h) ⊨ P"
by (simp add: no_guard_remove_guards)
moreover have "agrees (- {x}) pre_s s"
proof (rule agreesI)
fix y assume "y ∈ - {x}"
then have "y ≠ x"
by force
then show "pre_s y = s y"
by (simp add: s_def)
qed
moreover obtain "(s, pp), (s, pp) ⊨ P" "(s, gg), (s, gg) ⊨ ?G"
using ‹(pre_s, gg), (pre_s, gg) ⊨ SharedGuard π (λs. map_to_multiset (s ms))› ‹remove_guards h = pp› agreements agrees_same agrees_union always_sat_refl_aux calculation(4) by blast
let ?hf = "remove_guards hf"
let ?H = "remove_guards H"
let ?h = "remove_guards h"
obtain hhj where "Some hhj = Some h ⊕ Some hj"
by (metis asm2 plus.simps(2) plus.simps(3) plus_comm)
then have "Some H = Some hhj ⊕ Some hf"
using asm2 by presburger
then have "Some (remove_guards hhj) = Some ?h ⊕ Some hj"
by (metis ‹Some hhj = Some h ⊕ Some hj› asm2 no_guards_remove remove_guards_sum sat_inv_def)
moreover have "f (normalize (get_fh hj)) = v"
proof -
have "view Γ (normalize (get_fh hj)) = view Γ (normalize (get_fh H))"
using assms(4) view_function_of_invE
by (metis (no_types, opaque_lifting) ‹Some hhj = Some h ⊕ Some hj› asm2 larger_def larger_trans plus_comm)
then show ?thesis using assms(1) v_def by fastforce
qed
then have "(s, hj), (s, hj) ⊨ ?J"
by (metis ‹v = s x› asm2_bis assms(1) hyper_sat.simps(11) sat_inv_def select_convs(5))
ultimately have "(s, remove_guards hhj), (s, remove_guards hhj) ⊨ Star P ?J"
using ‹(s, remove_guards h), (s, remove_guards h) ⊨ P› hyper_sat.simps(4) by blast
moreover have "bounded hhj"
apply (rule bounded_smaller_sum[OF _ ‹Some H = Some hhj ⊕ Some hf›])
by (metis asm2 full_ownership_then_bounded get_fh.simps)
ultimately have all_safes: "⋀n. safe n (None :: ('i, 'a, nat) cont) C (s, remove_guards hhj) (Σ (s, remove_guards hhj))"
using asm0(1) unfolding bounded_def remove_guards_def by simp
then have "⋀σ1 H1 σ2 H2 s2 C2. red_rtrans C σ1 C2 σ2 ⟹ σ1 = (s, H1) ⟹ σ2 = (s2, H2) ⟹
?H = denormalize H1 ⟹
¬ aborts C2 σ2 ∧ (C2 = Cskip ⟶ (∃h1 H'. Some H' = Some h1 ⊕ Some ?hf ∧ H2 = FractionalHeap.normalize (get_fh H')
∧ no_guard H' ∧ full_ownership (get_fh H') ∧ (s2, h1) ∈ Σ (s, remove_guards hhj)))"
proof -
fix σ1 H1 σ2 H2 s2 C2
assume "?H = denormalize H1"
assume "red_rtrans C σ1 C2 σ2" "σ1 = (s, H1)" "σ2 = (s2, H2)"
then show "¬ aborts C2 σ2 ∧
(C2 = Cskip ⟶
(∃h1 H'.
Some H' = Some h1 ⊕ Some (remove_guards hf) ∧
H2 = FractionalHeap.normalize (get_fh H') ∧ no_guard H' ∧ full_ownership (get_fh H') ∧ (s2, h1) ∈ Σ (s, remove_guards hhj)))"
using all_safes
proof (rule safe_atomic)
show "?H = denormalize H1" using ‹?H = denormalize H1› by simp
show "Some ?H = Some (remove_guards hhj) ⊕ Some ?hf"
using ‹Some H = Some hhj ⊕ Some hf› remove_guards_sum by blast
show "full_ownership (get_fh (remove_guards H)) ∧ no_guard (remove_guards H)"
by (metis asm2 get_fh_remove_guards no_guard_remove_guards)
qed
qed
moreover have "?H = denormalize (normalize (get_fh H))"
by (metis asm2 denormalize_properties(5))
ultimately have safe_atomic_simplified: "⋀σ2 H2 s2 C2. red_rtrans C (s, normalize (get_fh H)) C2 σ2
⟹ σ2 = (s2, H2) ⟹ ¬ aborts C2 σ2 ∧ (C2 = Cskip ⟶ (∃h1 H'. Some H' = Some h1 ⊕ Some ?hf ∧ H2 = FractionalHeap.normalize (get_fh H')
∧ no_guard H' ∧ full_ownership (get_fh H') ∧ (s2, h1) ∈ Σ (s, remove_guards hhj)))"
by presburger
have "¬ aborts (Catomic C) (s, normalize (get_fh H))"
proof (rule ccontr)
assume "¬ ¬ aborts (Catomic C) (s, normalize (get_fh H))"
then obtain C' σ' where asm3: "red_rtrans C (s, FractionalHeap.normalize (get_fh H)) C' σ'"
"aborts C' σ'"
using abort_atomic_cases by blast
then have "¬ aborts C' σ'" using safe_atomic_simplified[of C' σ' "fst σ'" "snd σ'"] by simp
then show False using asm3(2) by simp
qed
then show "¬ aborts (Catomic C) (pre_s, normalize (get_fh H))"
by (metis agreements aborts_agrees agrees_comm agrees_union fst_eqD fvC.simps(11) snd_conv)
fix C' pre_s' h'
assume "red (Catomic C) (pre_s, FractionalHeap.normalize (get_fh H)) C' (pre_s', h')"
then obtain s' where "red (Catomic C) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
"agrees (- {x}) s' pre_s'"
by (metis (no_types, lifting) UnI1 ‹agrees (- {x}) pre_s s› agrees_comm assms(5) fst_eqD fvC.simps(11) red_agrees snd_conv subset_Compl_singleton)
then obtain h1 H' where asm3: "Some H' = Some h1 ⊕ Some (remove_guards hf)" "C' = Cskip"
"h' = FractionalHeap.normalize (get_fh H')" "no_guard H' ∧ full_ownership (get_fh H')" "(s', h1) ∈ Σ (s, remove_guards hhj)"
using safe_atomic_simplified[of C' "(s', h')" s' h'] by (metis red_atomic_cases)
moreover have "s x = s' x ∧ s sarg = s' sarg ∧ s ms = s' ms" using red_not_in_fv_not_touched
using ‹red (Catomic C) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')›
by (metis UnI1 assms(5) assms(6) assms(7) fst_eqD fvC.simps(11))
have "∃hq' hj'. Some h1 = Some hq' ⊕ Some hj' ∧ (s', add_sguard_to_no_guard hq' π (?ms s')) ∈ Σ' (pre_s, h)
∧ sat_inv s' hj' Γ ∧ f (normalize (get_fh hj')) = sact v (map_to_arg (s' sarg))"
proof -
have "pair_sat (Σ (s, remove_guards hhj)) (Σ (s, remove_guards hhj)) (Star Q ?J')"
using asm0(2)[of "(s, remove_guards hhj)" "(s, remove_guards hhj)"]
using ‹(s, remove_guards hhj), (s, remove_guards hhj) ⊨ Star P ?J› by blast
then have "(s', h1), (s', h1) ⊨ Star Q ?J'"
using asm3(5) pair_sat_def by blast
then obtain hq' hj' where "Some h1 = Some hq' ⊕ Some hj'" "(s', hq'), (s', hq') ⊨ Q" "(s', hj'), (s', hj') ⊨ ?J'"
using always_sat_refl hyper_sat.simps(4) by blast
then have "no_guard hj'"
by (metis (no_types, opaque_lifting) calculation(1) calculation(4) no_guard_then_smaller_same plus_comm)
moreover have "f (normalize (get_fh hj')) = sact v (map_to_arg (s' sarg))"
using ‹(s', hj'), (s', hj') ⊨ View f J (λs. sact (s x) (map_to_arg (s sarg)))› ‹s x = s' x ∧ s sarg = s' sarg ∧ s ms = s' ms› ‹v = s x› by fastforce
moreover have "(s, remove_guards hhj) ∈ start (pre_s, h)"
proof -
have "Some (remove_guards hhj) = Some ?h ⊕ Some hj"
using ‹Some (remove_guards hhj) = Some (remove_guards h) ⊕ Some hj› by blast
moreover have "(s, hj), (s, hj) ⊨ ?J"
using ‹(s, hj), (s, hj) ⊨ ?J› by fastforce
ultimately show ?thesis using start_def
using ‹agrees (- {x}) pre_s s› by fastforce
qed
then have "(s', h1) ∈ end_qj (pre_s, h)"
using ‹end_qj ≡ λσ. ⋃ (Σ ` start σ)› asm3(5) by blast
then have "(s', add_sguard_to_no_guard hq' π (?ms s')) ∈ Σ' (pre_s, h)"
using Σ'_def ‹(s', hj'), (s', hj') ⊨ ?J'› ‹Some h1 = Some hq' ⊕ Some hj'› by blast
ultimately show "∃hq' hj'.
Some h1 = Some hq' ⊕ Some hj' ∧
(s', add_sguard_to_no_guard hq' π ({#map_to_arg (s' sarg)#} + map_to_multiset (s' ms))) ∈ Σ' (pre_s, h) ∧
sat_inv s' hj' Γ ∧ f (FractionalHeap.normalize (get_fh hj')) = sact v (map_to_arg (s' sarg))"
using ‹(s', hj'), (s', hj') ⊨ View f J (λs. sact (s x) (map_to_arg (s sarg)))› ‹Some h1 = Some hq' ⊕ Some hj'› assms(1) sat_inv_def by fastforce
qed
then obtain hq' hj' where "Some h1 = Some hq' ⊕ Some hj'" "(s', add_sguard_to_no_guard hq' π (?ms s')) ∈ Σ' (pre_s, h)" "sat_inv s' hj' Γ"
"f (FractionalHeap.normalize (get_fh hj')) = sact v (map_to_arg (s' sarg))"
by blast
then have "safe n (Some Γ) C' (s', add_sguard_to_no_guard hq' π (?ms s')) (Σ' (pre_s, h))"
using asm3(2) safe_skip by blast
moreover have "∃H''. semi_consistent Γ v0 H'' ∧ Some H'' = Some (add_sguard_to_no_guard hq' π (?ms s')) ⊕ Some hj' ⊕ Some hf"
proof -
have "Some (add_sguard_to_no_guard hq' π (?ms s')) = Some hq' ⊕ Some (Map.empty, Some (π, ?ms s'), (λ_. None))"
using ‹Some h1 = Some hq' ⊕ Some hj'› add_sguard_as_sum asm3(1) asm3(4) no_guard_then_smaller_same by blast
obtain hhf where "Some hhf = Some h ⊕ Some hf"
by (metis (no_types, opaque_lifting) ‹Some H = Some hhj ⊕ Some hf› ‹Some hhj = Some h ⊕ Some hj› option.exhaust_sel plus.simps(1) plus_asso plus_comm)
then have "all_guards hhf"
by (metis (no_types, lifting) all_guards_no_guard_propagates asm2 plus_asso plus_comm sat_inv_def semi_consistent_def)
moreover have "get_gu h = (λ_. None) ∧ get_gs h = Some (π, ?pre_ms s)"
proof -
have "no_guard pp"
using ‹(pre_s, pp), (pre_s, pp) ⊨ P› assms(10) no_guard_assertion_def by blast
then show ?thesis
by (metis ‹Some h = Some pp ⊕ Some gg› ‹⋀thesis. (⟦(s, pp), (s, pp) ⊨ P; (s, gg), (s, gg) ⊨ SharedGuard π (λs. map_to_multiset (s ms))⟧ ⟹ thesis) ⟹ thesis› ‹remove_guards h = pp› decompose_heap_triple fst_conv hyper_sat.simps(12) no_guard_remove(2) plus_comm remove_guards_def snd_conv sum_gs_one_none)
qed
then have " ∃π' msf uargs. (∀k. get_gu hf k = Some (uargs k)) ∧
(π = pwrite ∧ get_gs hf = None ∧ msf = {#} ∨ pwrite = padd π π' ∧ get_gs hf = Some (π', msf))"
using all_guards_sum_known_one[of hhf h hf π]
using ‹Some hhf = Some h ⊕ Some hf› calculation by fastforce
then obtain π' uargs msf where "(∀k. get_gu hf k = Some (uargs k)) ∧ ((π = pwrite ∧ get_gs hf = None ∧ msf = {#}) ∨ (pwrite = padd π π' ∧ get_gs hf = Some (π', msf)))"
by blast
then obtain ghf where ghf_def: "Some hf = Some (remove_guards hf) ⊕ Some ghf"
"get_fh ghf = Map.empty" "(π = pwrite ∧ get_gs ghf = None ∧ msf = {#}) ∨ (padd π π' = pwrite ∧ get_gs ghf = Some (π', msf))"
"⋀i. get_gu ghf i = Some (uargs i)"
using decompose_guard_remove_easy[of hf]
by (metis fst_conv get_fh.elims get_gs.elims get_gu.simps snd_conv)
have "(Map.empty, Some (π, ?ms s'), (λ_. None)) ## ghf"
proof (rule compatibleI)
show "compatible_fract_heaps (get_fh (Map.empty, Some (π, ?ms s'), (λ_. None))) (get_fh ghf)"
using compatible_fract_heapsI by fastforce
show "⋀k. get_gu (Map.empty, Some (π, {#map_to_arg (s' sarg)#} + map_to_multiset (s' ms)), Map.empty) k = None ∨ get_gu ghf k = None"
by simp
fix p p'
assume "get_gs (Map.empty, Some (π, {#map_to_arg (s' sarg)#} + map_to_multiset (s' ms)), Map.empty) = Some p ∧ get_gs ghf = Some p'"
then have "p = (π, ?ms s') ∧ p' = (π', msf) ∧ padd π π' = pwrite"
using ghf_def by auto
then show "pgte pwrite (padd (fst p) (fst p'))"
using not_pgte_charact pgt_implies_pgte by auto
qed
then obtain g where g_def: "Some g = Some (Map.empty, Some (π, ?ms s'), (λ_. None)) ⊕ Some ghf"
by simp
moreover have "H' ## g"
proof (rule compatibleI)
have "get_fh g = add_fh Map.empty Map.empty" using add_get_fh[of g "(Map.empty, Some (π, ?ms s'), (λ_. None))" ghf]
g_def ‹get_fh ghf = Map.empty›
by fastforce
then have "get_fh g = Map.empty"
using add_fh_map_empty by auto
then show "compatible_fract_heaps (get_fh H') (get_fh g)"
using compatible_fract_heapsI by force
show "⋀k. get_gu H' k = None ∨ get_gu g k = None"
by (meson asm3(4) no_guard_def)
show "⋀p p'. get_gs H' = Some p ∧ get_gs g = Some p' ⟹ pgte pwrite (padd (fst p) (fst p'))"
by (metis asm3(4) no_guard_def option.simps(3))
qed
then obtain H'' where "Some H'' = Some H' ⊕ Some g"
by simp
then have "Some H'' = Some (add_sguard_to_no_guard hq' π (?ms s')) ⊕ Some hj' ⊕ Some hf"
proof -
have "Some H'' = Some h1 ⊕ Some g ⊕ Some (remove_guards hf)"
by (metis ‹Some H'' = Some H' ⊕ Some g› asm3(1) plus_comm simpler_asso)
moreover have "Some (add_sguard_to_no_guard hq' π (?ms s')) = Some hq' ⊕ Some (Map.empty, Some (π, ?ms s'), (λ_. None))"
using ‹Some (add_sguard_to_no_guard hq' π ({#map_to_arg (s' sarg)#} + map_to_multiset (s' ms))) = Some hq' ⊕ Some (Map.empty, Some (π, {#map_to_arg (s' sarg)#} + map_to_multiset (s' ms)), (λ_. None))› by blast
ultimately show ?thesis
by (metis (no_types, lifting) ‹Some h1 = Some hq' ⊕ Some hj'› g_def ghf_def(1) plus_comm simpler_asso)
qed
moreover have "semi_consistent Γ v0 H''"
proof (rule semi_consistentI)
have "get_gs g = Some (pwrite, ?ms s' + msf)"
proof (cases "π = pwrite")
case True
then have "π = pwrite ∧ get_gs ghf = None ∧ msf = {#}" using ghf_def(3)
by (metis not_pgte_charact pgt_implies_pgte sum_larger)
then show ?thesis
by (metis add.right_neutral fst_conv g_def get_gs.simps snd_conv sum_gs_one_none)
next
case False
then have "padd π π' = pwrite ∧ get_gs ghf = Some (π', msf)"
using ghf_def(3) by blast
then show ?thesis
by (metis calculation(2) fst_conv get_gs.elims snd_conv sum_gs_one_some)
qed
moreover have "⋀i. get_gu g i = Some (uargs i)"
by (metis full_uguard_sum_same ghf_def(4) g_def plus_comm)
ultimately have "all_guards g"
using all_guards_def by blast
then show "all_guards H''"
by (metis ‹Some H'' = Some H' ⊕ Some g› all_guards_same plus_comm)
show "reachable Γ v0 H''"
proof (rule reachableI)
fix sargs uargs'
assume "get_gs H'' = Some (pwrite, sargs) ∧ (∀k. get_gu H'' k = Some (uargs' k))"
then have "sargs = ?ms s' + msf"
by (metis (no_types, opaque_lifting) ‹Some H'' = Some H' ⊕ Some g› ‹get_gs g = Some (pwrite, {#map_to_arg (s' sarg)#} + map_to_multiset (s' ms) + msf)› asm3(4) no_guard_remove(1) option.inject plus_comm snd_conv)
moreover have "uargs = uargs'"
apply (rule ext)
by (metis ‹Some H'' = Some H' ⊕ Some g› ‹⋀i. get_gu g i = Some (uargs i)› ‹get_gs H'' = Some (pwrite, sargs) ∧ (∀k. get_gu H'' k = Some (uargs' k))› asm3(4) no_guard_remove(2) option.sel plus_comm)
moreover have "view Γ (FractionalHeap.normalize (get_fh hj')) = view Γ (FractionalHeap.normalize (get_fh H''))"
using assms(4) ‹sat_inv s' hj' Γ›
proof (rule view_function_of_invE)
show "H'' ≽ hj'"
by (metis (no_types, opaque_lifting) ‹Some H'' = Some H' ⊕ Some g› ‹Some h1 = Some hq' ⊕ Some hj'› asm3(1) larger_def larger_trans plus_comm)
qed
moreover have "reachable_value (saction Γ) (uaction Γ) v0 (?ms s' + msf) uargs (sact v (map_to_arg (s' sarg)))"
proof -
have "reachable_value (saction Γ) (uaction Γ) v0 (?pre_ms s + msf) uargs (view Γ (FractionalHeap.normalize (get_fh H)))"
proof -
have "reachable Γ v0 H"
by (meson asm2 semi_consistent_def)
moreover have "get_gs H = Some (pwrite, ?pre_ms s + msf) ∧ (∀k. get_gu H k = Some (uargs k))"
proof (rule conjI)
show "∀k. get_gu H k = Some (uargs k)"
by (metis ‹Some H = Some hhj ⊕ Some hf› full_uguard_sum_same ghf_def(1) ghf_def(4) plus_comm)
moreover have "get_gs hhj = Some (π, ?pre_ms s)"
proof -
have "get_gs hj = None"
using asm2 no_guard_def sat_inv_def by blast
moreover have "get_gs h = Some (π, ?pre_ms s)"
using ‹get_gu h = Map.empty ∧ get_gs h = Some (π, map_to_multiset (s ms))› by blast
ultimately show ?thesis
by (metis ‹Some hhj = Some h ⊕ Some hj› sum_gs_one_none)
qed
ultimately show "get_gs H = Some (pwrite, ?pre_ms s + msf)"
proof (cases "π = pwrite")
case True
then have "π = pwrite ∧ get_gs ghf = None ∧ msf = {#}" using ghf_def(3)
by (metis not_pgte_charact pgt_implies_pgte sum_larger)
then show ?thesis
by (metis ‹Some H = Some hhj ⊕ Some hf› ‹get_gs hhj = Some (π, map_to_multiset (s ms))› add.right_neutral full_sguard_sum_same)
next
case False
then have "padd π π' = pwrite ∧ get_gs ghf = Some (π', msf)"
using ghf_def(3) by blast
then show ?thesis using ‹Some H = Some hhj ⊕ Some hf› sum_gs_one_some ghf_def(1)
‹get_gs hhj = Some (π, ?pre_ms s)› asm3(1) asm3(4) no_guard_remove(1)[of hf ghf "remove_guards hf"] no_guard_then_smaller_same plus_comm
by metis
qed
qed
ultimately show ?thesis
by (meson reachableE)
qed
moreover have "view Γ (FractionalHeap.normalize (get_fh hj)) = view Γ (FractionalHeap.normalize (get_fh H))"
using assms(4)
proof (rule view_function_of_invE)
show "sat_inv s hj Γ"
by (simp add: asm2_bis)
show "H ≽ hj"
by (metis (no_types, opaque_lifting) ‹Some H = Some hhj ⊕ Some hf› ‹Some hhj = Some h ⊕ Some hj› larger_def larger_trans plus_comm)
qed
ultimately have "reachable_value (saction Γ) (uaction Γ) v0 (?pre_ms s + msf) uargs v"
using ‹f (FractionalHeap.normalize (get_fh hj)) = v› assms(1) by auto
then show ?thesis
using SharedStep assms(1)
using ‹s x = s' x ∧ s sarg = s' sarg ∧ s ms = s' ms› by fastforce
qed
ultimately show "reachable_value (saction Γ) (uaction Γ) v0 sargs uargs' (view Γ (FractionalHeap.normalize (get_fh H'')))"
using ‹f (FractionalHeap.normalize (get_fh hj')) = sact v (map_to_arg (s' sarg))› assms(1) by force
qed
qed
ultimately show "∃H''. semi_consistent Γ v0 H'' ∧ Some H'' = Some (add_sguard_to_no_guard hq' π ({#map_to_arg (s' sarg)#} + map_to_multiset (s' ms))) ⊕ Some hj' ⊕ Some hf"
by blast
qed
ultimately obtain H'' where "semi_consistent Γ v0 H'' ∧ Some H'' = Some (add_sguard_to_no_guard hq' π (?ms s')) ⊕ Some hj' ⊕ Some hf
∧ safe n (Some Γ) C' (s', add_sguard_to_no_guard hq' π (?ms s')) (Σ' (pre_s, h))" by blast
moreover have "full_ownership (get_fh H'') ∧ h' = FractionalHeap.normalize (get_fh H'')"
proof -
obtain x where "Some x = Some (add_sguard_to_no_guard hq' π (?ms s')) ⊕ Some hj'"
by (metis calculation not_Some_eq plus.simps(1))
then have "get_fh H'' = add_fh (add_fh (get_fh (add_sguard_to_no_guard hq' π (?ms s'))) (get_fh hj')) (get_fh hf)"
by (metis add_get_fh calculation)
moreover have "get_fh (add_sguard_to_no_guard hq' π (?ms s')) = get_fh hq' ∧ get_fh hf = get_fh (remove_guards hf)"
by (metis get_fh_add_sguard get_fh_remove_guards)
ultimately show ?thesis
by (metis ‹Some h1 = Some hq' ⊕ Some hj'› add_get_fh asm3(1) asm3(3) asm3(4))
qed
moreover have "sat_inv pre_s' hj' Γ"
proof (rule sat_inv_agrees)
show "sat_inv s' hj' Γ"
by (simp add: ‹sat_inv s' hj' Γ›)
show "agrees (fvA (invariant Γ)) s' pre_s'"
using UnCI ‹agrees (- {x}) s' pre_s'› assms(1) assms(5) select_convs(5) subset_Compl_singleton
by (metis (mono_tags, lifting) agrees_def in_mono)
qed
moreover have "safe n (Some Γ) C' (pre_s', add_sguard_to_no_guard hq' π (?ms s')) (?Σ' (pre_s, h))"
proof (rule safe_free_vars_Some)
show "safe n (Some Γ) C' (s', add_sguard_to_no_guard hq' π (?ms s')) (?Σ' (pre_s, h))"
by (meson ‹safe n (Some Γ) C' (s', add_sguard_to_no_guard hq' π ({#map_to_arg (s' sarg)#} + map_to_multiset (s' ms))) (Σ' (pre_s, h))› close_var_subset safe_larger_set)
show "agrees (fvC C' ∪ (- {x}) ∪ fvA (invariant Γ)) s' pre_s'"
by (metis UnI2 Un_absorb1 ‹agrees (- {x}) s' pre_s'› asm3(2) assms(1) assms(5) empty_iff fvC.simps(1) inf_sup_aci(5) select_convs(5) subset_Compl_singleton)
show "upper_fvs (close_var (Σ' (pre_s, h)) x) (- {x})"
by (simp add: upper_fvs_close_vars)
qed
ultimately show "∃h'' H' hj'.
full_ownership (get_fh H') ∧
semi_consistent Γ v0 H' ∧
sat_inv pre_s' hj' Γ ∧ h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hj' ⊕ Some hf
∧ safe n (Some Γ) C' (pre_s', h'') (?Σ' (pre_s, h))" using ‹sat_inv s' hj' Γ› by blast
qed (simp)
ultimately show "safe k (Some Γ) (Catomic C) (pre_s, h) (?Σ' (pre_s, h))" by blast
qed (simp)
qed
qed
subsubsection ‹Parallel›
lemma par_cases:
assumes "red (Cpar C1 C2) σ C' σ'"
and "⋀C1'. C' = Cpar C1' C2 ∧ red C1 σ C1' σ' ⟹ P"
and "⋀C2'. C' = Cpar C1 C2' ∧ red C2 σ C2' σ' ⟹ P"
and "C1 = Cskip ∧ C2 = Cskip ∧ C' = Cskip ∧ σ = σ' ⟹ P"
shows P
using assms(1)
apply (rule red.cases)
apply blast+
apply (simp add: assms(2))
apply (simp add: assms(3))
apply (simp add: assms(4))
apply blast+
done
lemma no_abort_par:
assumes "no_abort Γ C1 s h"
and "no_abort Γ C2 s h"
and "safe (Suc n) Δ C1 (s, h1) S1"
and "safe (Suc n) Δ C2 (s, h2) S2"
and "Some h = Some h1 ⊕ Some h2"
and "bounded h"
shows "no_abort Γ (Cpar C1 C2) s h ∧ accesses (Cpar C1 C2) s ⊆ dom (fst h) ∧ writes (Cpar C1 C2) s ⊆ fpdom (fst h)"
proof (rule conjI)
have r: "accesses C1 s ⊆ dom (fst h1) ∧ accesses C2 s ⊆ dom (fst h2) ∧ writes C1 s ⊆ fpdom (fst h1) ∧ writes C2 s ⊆ fpdom (fst h2)"
using assms(3-4)
by (metis fst_eqD safeAccessesE snd_eqD)
then have "accesses (Cpar C1 C2) s ⊆ dom (fst h)"
by (metis Un_mono accesses.simps(8) assms(5) dom_sum_two get_fh.elims)
moreover have "fpdom (fst h1) ∪ fpdom (fst h2) ⊆ fpdom (fst h)"
using assms(5) fpdom_dom_union assms(6) by blast
then have "writes (Cpar C1 C2) s ⊆ fpdom (fst h)"
using ‹accesses C1 s ⊆ dom (fst h1) ∧ accesses C2 s ⊆ dom (fst h2) ∧ writes C1 s ⊆ fpdom (fst h1) ∧ writes C2 s ⊆ fpdom (fst h2)› dual_order.trans by auto
ultimately show "accesses (Cpar C1 C2) s ⊆ dom (fst h) ∧ writes (Cpar C1 C2) s ⊆ fpdom (fst h)" by simp
have r: "accesses C1 s ∩ writes C2 s = {} ∧ accesses C2 s ∩ writes C1 s = {}"
apply (rule conjI)
using fpdom_dom_disjoint[OF assms(5)] r
apply blast
using assms(5) fpdom_dom_disjoint[of h h2 h1] r plus_comm[of "Some h1" "Some h2"]
by force
show "no_abort Γ (Cpar C1 C2) s h"
proof (rule no_abortI)
show "⋀hf H.
Some H = Some h ⊕ Some hf ∧ Γ = None ∧ full_ownership (get_fh H) ∧ no_guard H ⟹
¬ aborts (Cpar C1 C2) (s, FractionalHeap.normalize (get_fh H))"
proof -
fix hf H assume asm0: "Some H = Some h ⊕ Some hf ∧ Γ = None ∧ full_ownership (get_fh H) ∧ no_guard H"
let ?H = "FractionalHeap.normalize (get_fh H)"
show "¬ aborts (Cpar C1 C2) (s, FractionalHeap.normalize (get_fh H))"
proof (rule ccontr)
assume "¬ ¬ aborts (Cpar C1 C2) (s, FractionalHeap.normalize (get_fh H))"
then have "aborts (Cpar C1 C2) (s, FractionalHeap.normalize (get_fh H))" by simp
then have "aborts C1 (s, ?H) ∨ aborts C2 (s, ?H)"
apply (rule aborts.cases)
apply simp_all
using r by force+
then show False
using asm0 assms(1) assms(2) no_abortE(1) by blast
qed
qed
fix H hf hj v0 Γ'
assume asm0: "Γ = Some Γ' ∧ Some H = Some h ⊕ Some hj ⊕ Some hf ∧ full_ownership (get_fh H) ∧ semi_consistent Γ' v0 H ∧ sat_inv s hj Γ'"
let ?H = "FractionalHeap.normalize (get_fh H)"
show "¬ aborts (Cpar C1 C2) (s, FractionalHeap.normalize (get_fh H))"
proof (rule ccontr)
assume "¬ ¬ aborts (Cpar C1 C2) (s, FractionalHeap.normalize (get_fh H))"
then have "aborts (Cpar C1 C2) (s, FractionalHeap.normalize (get_fh H))" by simp
then have "aborts C1 (s, ?H) ∨ aborts C2 (s, ?H)"
apply (rule aborts.cases)
apply simp_all
using r by force+
then show False
using asm0 assms(1) assms(2) no_abortE(2) by blast
qed
qed
qed
lemma parallel_comp_none:
assumes "safe n (None :: ('i, 'a, nat) cont) C1 (s, h1) S1"
and "safe n (None :: ('i, 'a, nat) cont) C2 (s, h2) S2"
and "Some h = Some h1 ⊕ Some h2"
and "disjoint (fvC C1 ∪ vars1) (wrC C2)"
and "disjoint (fvC C2 ∪ vars2) (wrC C1)"
and "upper_fvs S1 vars1"
and "upper_fvs S2 vars2"
and "bounded h"
shows "safe n (None :: ('i, 'a, nat) cont) (Cpar C1 C2) (s, h) (add_states S1 S2)"
using assms
proof (induct n arbitrary: C1 h1 C2 h2 s h S1 S2)
case (Suc n)
show ?case
proof (rule safeNoneI)
show "Cpar C1 C2 = Cskip ⟹ (s, h) ∈ add_states S1 S2"
by simp
have r: "no_abort (None :: ('i, 'a, nat) cont) (Cpar C1 C2) s h ∧ accesses (Cpar C1 C2) s ⊆ dom (fst h) ∧ writes (Cpar C1 C2) s ⊆ fpdom (fst h)"
proof (rule no_abort_par)
show "no_abort (None :: ('i, 'a, nat) cont) C1 s h"
using Suc.prems(1) Suc.prems(3) larger_def no_abort_larger safe.simps(2) by blast
have "h ≽ h2"
by (metis Suc.prems(3) larger_def plus_comm)
then show "no_abort (None :: ('i, 'a, nat) cont) C2 s h"
using Suc.prems(2) no_abort_larger safeNoneE_bis(2) by blast
show "safe (Suc n) None C1 (s, h1) S1" using Suc(2) by simp
show "safe (Suc n) None C2 (s, h2) S2" using Suc(3) by simp
qed (simp_all add: Suc)
then show "no_abort (None :: ('i, 'a, nat) cont) (Cpar C1 C2) s h" by simp
show "accesses (Cpar C1 C2) s ⊆ dom (fst h) ∧ writes (Cpar C1 C2) s ⊆ fpdom (fst h)" using r by simp
fix H hf C' s' h'
assume asm0: "Some H = Some h ⊕ Some hf ∧
full_ownership (get_fh H) ∧ no_guard H ∧ red (Cpar C1 C2) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
obtain hf1 where "Some hf1 = Some h1 ⊕ Some hf"
by (metis (no_types, opaque_lifting) Suc.prems(3) asm0 plus.simps(1) plus.simps(3) plus_asso plus_comm)
then have "Some H = Some h2 ⊕ Some hf1"
by (metis (no_types, lifting) Suc.prems(3) asm0 plus_asso plus_comm)
obtain hf2 where "Some hf2 = Some h2 ⊕ Some hf"
by (metis (no_types, opaque_lifting) ‹Some H = Some h2 ⊕ Some hf1› ‹Some hf1 = Some h1 ⊕ Some hf› option.exhaust_sel plus.simps(1) plus_asso plus_comm)
then have "Some H = Some h1 ⊕ Some hf2"
by (metis Suc.prems(3) asm0 plus_asso)
let ?H = "normalize (get_fh H)"
show "∃h'' H'.
full_ownership (get_fh H') ∧
no_guard H' ∧ h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hf ∧ safe n (None :: ('i, 'a, nat) cont) C' (s', h'') (add_states S1 S2)"
proof (rule par_cases)
show "red (Cpar C1 C2) (s, ?H) C' (s', h')"
using asm0 by blast
show "C1 = Cskip ∧ C2 = Cskip ∧ C' = Cskip ∧ (s, FractionalHeap.normalize (get_fh H)) = (s', h') ⟹
∃h'' H'. full_ownership (get_fh H') ∧
no_guard H' ∧ h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hf ∧ safe n (None :: ('i, 'a, nat) cont) C' (s', h'') (add_states S1 S2)"
proof -
assume asm1: "C1 = Cskip ∧ C2 = Cskip ∧ C' = Cskip ∧ (s, FractionalHeap.normalize (get_fh H)) = (s', h')"
then have "(s, h1) ∈ S1 ∧ (s, h2) ∈ S2"
using Suc.prems(1) Suc.prems(2) safe.simps(2) by blast
moreover have "(s, h) ∈ add_states S1 S2"
by (metis (mono_tags, lifting) Suc.prems(3) add_states_def calculation mem_Collect_eq)
ultimately show "∃h'' H'.
full_ownership (get_fh H') ∧
no_guard H' ∧ h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hf ∧ safe n (None :: ('i, 'a, nat) cont) C' (s', h'') (add_states S1 S2)"
by (metis asm0 asm1 old.prod.inject safe_skip)
qed
show "⋀C1'. C' = Cpar C1' C2 ∧ red C1 (s, FractionalHeap.normalize (get_fh H)) C1' (s', h') ⟹
∃h'' H'.
full_ownership (get_fh H') ∧
no_guard H' ∧ h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hf ∧ safe n (None :: ('i, 'a, nat) cont) C' (s', h'') (add_states S1 S2)"
proof -
fix C1'
assume asm1: "C' = Cpar C1' C2 ∧ red C1 (s, FractionalHeap.normalize (get_fh H)) C1' (s', h')"
then obtain h1' H' where asm2: "full_ownership (get_fh H')" "no_guard H'" "h' = FractionalHeap.normalize (get_fh H')"
"Some H' = Some h1' ⊕ Some hf2" "safe n (None :: ('i, 'a, nat) cont) C1' (s', h1') S1"
using Suc.prems(1) asm0 safeNoneE(3)[of n C1 s h1 S1 H hf2 C1' s' h'] ‹Some H = Some h1 ⊕ Some hf2› by blast
moreover have "safe n (None :: ('i, 'a, nat) cont) C2 (s, h2) S2"
by (meson Suc.prems(2) Suc_leD le_Suc_eq safe_smaller)
then have "safe n (None :: ('i, 'a, nat) cont) C2 (s', h2) S2"
proof (rule safe_free_vars_None)
show "agrees (fvC C2 ∪ vars2) s s'"
using Suc.prems(5) agrees_minusD[of ] agrees_comm asm1 fst_eqD red_properties(1) disjoint_def inf_commute
by metis
show "upper_fvs S2 vars2"
by (simp add: Suc.prems(7))
qed
moreover obtain h'' where "Some h'' = Some h1' ⊕ Some h2"
by (metis ‹Some hf2 = Some h2 ⊕ Some hf› calculation(4) not_Some_eq plus.simps(1) plus_asso)
have "safe n (None :: ('i, 'a, nat) cont) (Cpar C1' C2) (s', h'') (add_states S1 S2)"
proof (rule Suc.hyps)
show "safe n (None :: ('i, 'a, nat) cont) C1' (s', h1') S1"
using calculation(5) by blast
show "safe n (None :: ('i, 'a, nat) cont) C2 (s', h2) S2"
using calculation(6) by auto
show "Some h'' = Some h1' ⊕ Some h2"
using ‹Some h'' = Some h1' ⊕ Some h2› by blast
show "disjoint (fvC C1' ∪ vars1) (wrC C2)"
using Suc.prems(4) asm1 red_properties(1) Un_iff disjoint_def[of "fvC C1 ∪ vars1" "wrC C2"]
disjoint_def[of "fvC C1' ∪ vars1" "wrC C2"]
inf_shunt subset_iff by blast
show "disjoint (fvC C2 ∪ vars2) (wrC C1')"
by (metis (no_types, lifting) Suc.prems(5) asm1 disjoint_def inf_commute inf_shunt red_properties(1) subset_Un_eq sup_assoc)
show "upper_fvs S1 vars1"
by (simp add: Suc.prems(6))
show "upper_fvs S2 vars2"
by (simp add: Suc.prems(7))
show "bounded h''"
apply (rule bounded_smaller[of H'])
using calculation(1) full_ownership_then_bounded apply fastforce
by (metis ‹Some h'' = Some h1' ⊕ Some h2› ‹Some hf2 = Some h2 ⊕ Some hf› calculation(4) larger_def plus_asso)
qed
ultimately show "∃h'' H'.
full_ownership (get_fh H') ∧
no_guard H' ∧
h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hf ∧ safe n (None :: ('i, 'a, nat) cont) C' (s', h'') (add_states S1 S2)"
by (metis ‹Some h'' = Some h1' ⊕ Some h2› ‹Some hf2 = Some h2 ⊕ Some hf› asm1 plus_asso)
qed
show "⋀C2'. C' = Cpar C1 C2' ∧ red C2 (s, FractionalHeap.normalize (get_fh H)) C2' (s', h') ⟹
∃h'' H'.
full_ownership (get_fh H') ∧
no_guard H' ∧ h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hf ∧ safe n (None :: ('i, 'a, nat) cont) C' (s', h'') (add_states S1 S2)"
proof -
fix C2'
assume asm1: "C' = Cpar C1 C2' ∧ red C2 (s, FractionalHeap.normalize (get_fh H)) C2' (s', h')"
then obtain h2' H' where asm2: "full_ownership (get_fh H')" "no_guard H'" "h' = FractionalHeap.normalize (get_fh H')"
"Some H' = Some h2' ⊕ Some hf1" "safe n (None :: ('i, 'a, nat) cont) C2' (s', h2') S2"
using Suc.prems(1) asm0 safeNoneE(3) Suc.prems(2) ‹Some H = Some h2 ⊕ Some hf1› by blast
moreover have "safe n (None :: ('i, 'a, nat) cont) C1 (s, h1) S1"
by (meson Suc.prems(1) Suc_leD le_Suc_eq safe_smaller)
then have "safe n (None :: ('i, 'a, nat) cont) C1 (s', h1) S1"
proof (rule safe_free_vars_None)
show "agrees (fvC C1 ∪ vars1) s s'"
using Suc.prems(4) agrees_comm asm1 fst_eqD red_properties(1) disjoint_def[of "fvC C1 ∪ vars1" "wrC C2"]
agrees_minusD by (metis inf_commute)
show "upper_fvs S1 vars1"
by (simp add: Suc.prems(6))
qed
moreover obtain h'' where "Some h'' = Some h2' ⊕ Some h1"
by (metis ‹Some hf1 = Some h1 ⊕ Some hf› calculation(4) not_Some_eq plus.simps(1) plus_asso)
have "safe n (None :: ('i, 'a, nat) cont) (Cpar C1 C2') (s', h'') (add_states S1 S2)"
proof (rule Suc.hyps)
show "safe n (None :: ('i, 'a, nat) cont) C1 (s', h1) S1"
using calculation(6) by blast
show "safe n (None :: ('i, 'a, nat) cont) C2' (s', h2') S2"
using calculation(5) by auto
show "Some h'' = Some h1 ⊕ Some h2'"
by (simp add: ‹Some h'' = Some h2' ⊕ Some h1› plus_comm)
show "disjoint (fvC C2' ∪ vars2) (wrC C1)"
using Suc.prems(5) asm1 disjoint_def[of "fvC C2 ∪ vars2" "wrC C1"] disjoint_def[of "fvC C2' ∪ vars2" "wrC C1"]
inf_shunt inf_sup_aci(5) red_properties(1) subset_Un_eq sup.idem sup_assoc
by fast
show "disjoint (fvC C1 ∪ vars1) (wrC C2')"
by (metis (no_types, lifting) Suc.prems(4) asm1 disjoint_def inf_commute inf_shunt red_properties(1) subset_Un_eq sup_assoc)
show "upper_fvs S1 vars1"
by (simp add: Suc.prems(6))
show "upper_fvs S2 vars2"
by (simp add: Suc.prems(7))
show "bounded h''"
apply (rule bounded_smaller[of H'])
using calculation(1) full_ownership_then_bounded apply fastforce
by (metis ‹Some h'' = Some h2' ⊕ Some h1› ‹Some hf1 = Some h1 ⊕ Some hf› calculation(4) larger_def plus_asso)
qed
ultimately show "∃h'' H'.
full_ownership (get_fh H') ∧
no_guard H' ∧
h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hf ∧ safe n (None :: ('i, 'a, nat) cont) C' (s', h'') (add_states S1 S2)"
by (metis ‹Some h'' = Some h2' ⊕ Some h1› ‹Some hf1 = Some h1 ⊕ Some hf› asm1 plus_asso)
qed
qed
qed
qed (simp)
lemma parallel_comp_some:
assumes "safe n (Some Γ) C1 (s, h1) S1"
and "safe n (Some Γ) C2 (s, h2) S2"
and "Some h = Some h1 ⊕ Some h2"
and "disjoint (fvC C1 ∪ vars1) (wrC C2)"
and "disjoint (fvC C2 ∪ vars2) (wrC C1)"
and "upper_fvs S1 vars1"
and "upper_fvs S2 vars2"
and "disjoint (fvA (invariant Γ)) (wrC C2)"
and "disjoint (fvA (invariant Γ)) (wrC C1)"
and "bounded h"
shows "safe n (Some Γ) (Cpar C1 C2) (s, h) (add_states S1 S2)"
using assms
proof (induct n arbitrary: C1 h1 C2 h2 s h S1 S2)
case (Suc n)
show ?case
proof (rule safeSomeI)
show "Cpar C1 C2 = Cskip ⟹ (s, h) ∈ add_states S1 S2"
by simp
have r: "no_abort (Some Γ) (Cpar C1 C2) s h ∧ accesses (Cpar C1 C2) s ⊆ dom (fst h) ∧ writes (Cpar C1 C2) s ⊆ fpdom (fst h)"
proof (rule no_abort_par)
show "no_abort (Some Γ) C1 s h"
using Suc.prems(1) Suc.prems(3) larger_def no_abort_larger safe.simps(3) by blast
have "h ≽ h2"
by (metis Suc.prems(3) larger_def plus_comm)
then show "no_abort (Some Γ) C2 s h"
using Suc.prems(2) no_abort_larger safeSomeE(2) by blast
show "safe (Suc n) (Some Γ) C1 (s, h1) S1" using Suc(2) by simp
show "safe (Suc n) (Some Γ) C2 (s, h2) S2" using Suc(3) by simp
qed (simp_all add: Suc)
then show "accesses (Cpar C1 C2) s ⊆ dom (fst h) ∧ writes (Cpar C1 C2) s ⊆ fpdom (fst h)"
by simp
show "no_abort (Some Γ) (Cpar C1 C2) s h" using r by simp
fix H hf C' s' h' hj v0
assume asm0: "Some H = Some h ⊕ Some hj ⊕ Some hf ∧ full_ownership (get_fh H) ∧
semi_consistent Γ v0 H ∧ sat_inv s hj Γ ∧ red (Cpar C1 C2) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
obtain hf1 where "Some hf1 = Some h1 ⊕ Some hf"
by (metis (no_types, opaque_lifting) Suc.prems(3) asm0 plus.simps(1) plus.simps(3) plus_asso plus_comm)
then have "Some H = Some h2 ⊕ Some hf1 ⊕ Some hj"
by (metis (no_types, lifting) Suc.prems(3) asm0 plus_asso plus_comm)
then have "Some H = Some h2 ⊕ Some hj ⊕ Some hf1"
by (metis plus_asso plus_comm)
obtain hf2 where "Some hf2 = Some h2 ⊕ Some hf"
by (metis (no_types, opaque_lifting) ‹Some H = Some h2 ⊕ Some hf1 ⊕ Some hj› ‹Some hf1 = Some h1 ⊕ Some hf› not_Some_eq plus.simps(1) plus_asso plus_comm)
then have "Some H = Some h1 ⊕ Some hf2 ⊕ Some hj"
by (metis (no_types, opaque_lifting) ‹Some H = Some h2 ⊕ Some hf1 ⊕ Some hj› ‹Some hf1 = Some h1 ⊕ Some hf› plus_asso plus_comm)
then have "Some H = Some h1 ⊕ Some hj ⊕ Some hf2"
by (metis plus_asso plus_comm)
let ?H = "normalize (get_fh H)"
show "∃h'' H' hj'.
full_ownership (get_fh H') ∧
semi_consistent Γ v0 H' ∧
sat_inv s' hj' Γ ∧
h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hj' ⊕ Some hf ∧ safe n (Some Γ) C' (s', h'') (add_states S1 S2)"
proof (rule par_cases)
show "red (Cpar C1 C2) (s, ?H) C' (s', h')"
using asm0 by blast
show "C1 = Cskip ∧ C2 = Cskip ∧ C' = Cskip ∧ (s, FractionalHeap.normalize (get_fh H)) = (s', h') ⟹
∃h'' H' hj'. full_ownership (get_fh H') ∧
semi_consistent Γ v0 H' ∧
sat_inv s' hj' Γ ∧ h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hj' ⊕ Some hf ∧ safe n (Some Γ) C' (s', h'') (add_states S1 S2)"
proof -
assume asm1: "C1 = Cskip ∧ C2 = Cskip ∧ C' = Cskip ∧ (s, FractionalHeap.normalize (get_fh H)) = (s', h')"
then have "(s, h1) ∈ S1 ∧ (s, h2) ∈ S2"
using Suc.prems(1) Suc.prems(2) safe.simps(3) by blast
moreover have "(s, h) ∈ add_states S1 S2"
by (metis (mono_tags, lifting) Suc.prems(3) add_states_def calculation mem_Collect_eq)
ultimately show "∃h'' H' hj'. full_ownership (get_fh H') ∧
semi_consistent Γ v0 H' ∧
sat_inv s' hj' Γ ∧ h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hj' ⊕ Some hf ∧ safe n (Some Γ) C' (s', h'') (add_states S1 S2)"
by (metis asm0 asm1 old.prod.inject safe_skip)
qed
show "⋀C1'. C' = Cpar C1' C2 ∧ red C1 (s, FractionalHeap.normalize (get_fh H)) C1' (s', h') ⟹
∃h'' H' hj'. full_ownership (get_fh H') ∧ semi_consistent Γ v0 H' ∧
sat_inv s' hj' Γ ∧ h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hj' ⊕ Some hf ∧ safe n (Some Γ) C' (s', h'') (add_states S1 S2)"
proof -
fix C1'
assume asm1: "C' = Cpar C1' C2 ∧ red C1 (s, FractionalHeap.normalize (get_fh H)) C1' (s', h')"
then obtain h1' H' hj' where asm2: "full_ownership (get_fh H')" "h' = FractionalHeap.normalize (get_fh H')"
"semi_consistent Γ v0 H'" "sat_inv s' hj' Γ" "Some H' = Some h1' ⊕ Some hj' ⊕ Some hf2" "safe n (Some Γ) C1' (s', h1') S1"
using safeSomeE(3)[of n Γ C1 s h1 S1 H hj hf2 v0 C1' s' h'] Suc.prems(1) asm0
using ‹Some H = Some h1 ⊕ Some hj ⊕ Some hf2› by blast
moreover have "safe n (Some Γ) C2 (s, h2) S2"
by (meson Suc.prems(2) Suc_leD le_Suc_eq safe_smaller)
then have "safe n (Some Γ) C2 (s', h2) S2"
proof (rule safe_free_vars_Some)
show "agrees (fvC C2 ∪ vars2 ∪ fvA (invariant Γ)) s s'"
using Suc.prems(5) Suc.prems(9) agrees_minusD agrees_comm asm1 disjoint_def fst_eqD red_properties(1)
by (metis agrees_union inf_commute)
show "upper_fvs S2 vars2"
by (simp add: Suc.prems(7))
qed
moreover have "h1' ## h2"
by (metis (no_types, opaque_lifting) ‹Some hf2 = Some h2 ⊕ Some hf› calculation(5) compatible_eq option.discI plus.simps(1) plus_asso plus_comm)
then obtain h'' where "Some h'' = Some h1' ⊕ Some h2" by simp
have "safe n (Some Γ) (Cpar C1' C2) (s', h'') (add_states S1 S2)"
proof (rule Suc.hyps)
show "safe n (Some Γ) C1' (s', h1') S1"
using calculation(6) by blast
show "safe n (Some Γ) C2 (s', h2) S2"
using calculation(7) by auto
show "Some h'' = Some h1' ⊕ Some h2"
using ‹Some h'' = Some h1' ⊕ Some h2› by blast
show "disjoint (fvC C1' ∪ vars1) (wrC C2)"
by (metis (no_types, opaque_lifting) Suc.prems(4) asm1 disjnt_Un1 disjnt_def disjoint_def red_properties(1) sup.orderE)
show "disjoint (fvC C2 ∪ vars2) (wrC C1')"
by (metis (no_types, lifting) Suc.prems(5) asm1 disjoint_def inf_commute inf_shunt red_properties(1) subset_Un_eq sup_assoc)
show "upper_fvs S1 vars1"
by (simp add: Suc.prems(6))
show "upper_fvs S2 vars2"
by (simp add: Suc.prems(7))
show "disjoint (fvA (invariant Γ)) (wrC C2)"
by (simp add: Suc.prems(8))
show "disjoint (fvA (invariant Γ)) (wrC C1')"
by (metis (no_types, lifting) Suc.prems(9) asm1 disjoint_def inf_commute inf_shunt red_properties(1) subset_Un_eq sup_assoc)
show "bounded h''"
apply (rule bounded_smaller[of H'])
using calculation(1) full_ownership_then_bounded apply fastforce
using ‹Some h'' = Some h1' ⊕ Some h2› ‹Some hf2 = Some h2 ⊕ Some hf› calculation(5) larger3[of H' _ h'' hf] plus_asso plus_comm[of "Some hj'"]
by metis
qed
moreover have "Some H' = Some h'' ⊕ Some hj' ⊕ Some hf"
by (metis (no_types, opaque_lifting) ‹Some h'' = Some h1' ⊕ Some h2› ‹Some hf2 = Some h2 ⊕ Some hf› calculation(5) plus_asso plus_comm)
ultimately show "∃h'' H' hj'.
full_ownership (get_fh H') ∧
semi_consistent Γ v0 H' ∧
sat_inv s' hj' Γ ∧ h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hj' ⊕ Some hf ∧ safe n (Some Γ) C' (s', h'') (add_states S1 S2)"
using asm1 by blast
qed
show "⋀C2'. C' = Cpar C1 C2' ∧ red C2 (s, FractionalHeap.normalize (get_fh H)) C2' (s', h') ⟹
∃h'' H' hj'.
full_ownership (get_fh H') ∧
semi_consistent Γ v0 H' ∧
sat_inv s' hj' Γ ∧ h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hj' ⊕ Some hf ∧ safe n (Some Γ) C' (s', h'') (add_states S1 S2)"
proof -
fix C2'
assume asm1: "C' = Cpar C1 C2' ∧ red C2 (s, FractionalHeap.normalize (get_fh H)) C2' (s', h')"
then obtain h2' H' hj' where asm2: "full_ownership (get_fh H')" "h' = FractionalHeap.normalize (get_fh H')"
"semi_consistent Γ v0 H'" "sat_inv s' hj' Γ" "Some H' = Some h2' ⊕ Some hj' ⊕ Some hf1" "safe n (Some Γ) C2' (s', h2') S2"
using safeSomeE(3)[of n Γ C2 s h2 S2 H hj hf1 v0 C2' s' h'] Suc.prems(2) Suc.prems(3)
using ‹Some H = Some h2 ⊕ Some hj ⊕ Some hf1› asm0 by blast
moreover have "safe n (Some Γ) C1 (s, h1) S1"
by (meson Suc.prems(1) Suc_leD le_Suc_eq safe_smaller)
then have "safe n (Some Γ) C1 (s', h1) S1"
proof (rule safe_free_vars_Some)
show "agrees (fvC C1 ∪ vars1 ∪ fvA (invariant Γ)) s s'"
using Suc.prems(4) Suc.prems(8) agrees_minusD agrees_comm asm1 fst_eqD red_properties(1)
by (metis agrees_union disjoint_def inf_commute)
show "upper_fvs S1 vars1"
by (simp add: Suc.prems(6))
qed
moreover have "h1 ## h2'"
by (metis (no_types, opaque_lifting) ‹Some hf1 = Some h1 ⊕ Some hf› calculation(5) compatible_eq option.distinct(1) plus.simps(1) plus_asso plus_comm)
then obtain h'' where "Some h'' = Some h1 ⊕ Some h2'" by simp
have "safe n (Some Γ) (Cpar C1 C2') (s', h'') (add_states S1 S2)"
proof (rule Suc.hyps)
show "safe n (Some Γ) C1 (s', h1) S1"
using calculation(7) by blast
show "safe n (Some Γ) C2' (s', h2') S2"
using calculation(6) by auto
show "Some h'' = Some h1 ⊕ Some h2'"
using ‹Some h'' = Some h1 ⊕ Some h2'› by blast
show "disjoint (fvC C1 ∪ vars1) (wrC C2')"
by (metis (no_types, lifting) Suc.prems(4) asm1 disjoint_def inf_commute inf_shunt red_properties(1) subset_Un_eq sup_assoc)
show "disjoint (fvC C2' ∪ vars2) (wrC C1)"
using Suc.prems(5) asm1 red_properties(1)
by (metis (no_types, lifting) Un_subset_iff disjoint_def inf_shunt subset_Un_eq)
show "disjoint (fvA (invariant Γ)) (wrC C2')"
using Suc.prems(8) asm1 red_properties(1)
by (metis (no_types, lifting) Un_subset_iff disjoint_def inf_commute inf_shunt subset_Un_eq)
show "disjoint (fvA (invariant Γ)) (wrC C1)"
by (simp add: Suc.prems(9))
show "upper_fvs S1 vars1"
by (simp add: Suc.prems(6))
show "upper_fvs S2 vars2"
by (simp add: Suc.prems(7))
show "bounded h''"
apply (rule bounded_smaller[of H'])
using calculation(1) full_ownership_then_bounded apply fastforce
proof -
have "Some H' = Some h1 ⊕ Some h2' ⊕ Some hj' ⊕ Some hf"
by (metis (no_types, lifting) ‹Some hf1 = Some h1 ⊕ Some hf› calculation(5) plus_asso plus_comm)
then show "H' ≽ h''"
by (simp add: ‹Some h'' = Some h1 ⊕ Some h2'› larger3 plus_comm)
qed
qed
moreover have "Some H' = Some h'' ⊕ Some hj' ⊕ Some hf"
by (metis ‹Some h'' = Some h1 ⊕ Some h2'› ‹Some hf1 = Some h1 ⊕ Some hf› calculation(5) plus_comm simpler_asso)
ultimately show "∃h'' H' hj'.
full_ownership (get_fh H') ∧
semi_consistent Γ v0 H' ∧
sat_inv s' hj' Γ ∧
h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hj' ⊕ Some hf ∧ safe n (Some Γ) C' (s', h'') (add_states S1 S2)"
using asm1 by blast
qed
qed
qed
qed (simp)
lemma parallel_comp:
fixes Δ :: "('i, 'a, nat) cont"
assumes "safe n Δ C1 (s, h1) S1"
and "safe n Δ C2 (s, h2) S2"
and "Some h = Some h1 ⊕ Some h2"
and "disjoint (fvC C1 ∪ vars1) (wrC C2)"
and "disjoint (fvC C2 ∪ vars2) (wrC C1)"
and "upper_fvs S1 vars1"
and "upper_fvs S2 vars2"
and "⋀Γ. Δ = Some Γ ⟹ disjoint (fvA (invariant Γ)) (wrC C2)"
and "⋀Γ. Δ = Some Γ ⟹ disjoint (fvA (invariant Γ)) (wrC C1)"
and "bounded h"
shows "safe n Δ (Cpar C1 C2) (s, h) (add_states S1 S2)"
proof (cases Δ)
case None
then show ?thesis
using assms parallel_comp_none by blast
next
case (Some Γ)
then show ?thesis
using assms parallel_comp_some by blast
qed
theorem rule_par:
fixes Δ :: "('i, 'a, nat) cont"
assumes "hoare_triple_valid Δ P1 C1 Q1"
and "hoare_triple_valid Δ P2 C2 Q2"
and "disjoint (fvA P1 ∪ fvC C1 ∪ fvA Q1) (wrC C2)"
and "disjoint (fvA P2 ∪ fvC C2 ∪ fvA Q2) (wrC C1)"
and "⋀Γ. Δ = Some Γ ⟹ disjoint (fvA (invariant Γ)) (wrC C2)"
and "⋀Γ. Δ = Some Γ ⟹ disjoint (fvA (invariant Γ)) (wrC C1)"
and "precise P1 ∨ precise P2"
shows "hoare_triple_valid Δ (Star P1 P2) (Cpar C1 C2) (Star Q1 Q2)"
proof -
obtain Σ1 where r1: "⋀σ n. σ, σ ⊨ P1 ⟹ bounded (snd σ) ⟹ safe n Δ C1 σ (Σ1 σ)" "⋀σ σ'. σ, σ' ⊨ P1 ⟹ pair_sat (Σ1 σ) (Σ1 σ') Q1"
using assms(1) hoare_triple_validE by blast
obtain Σ2 where r2: "⋀σ n. σ, σ ⊨ P2 ⟹ bounded (snd σ) ⟹ safe n Δ C2 σ (Σ2 σ)" "⋀σ σ'. σ, σ' ⊨ P2 ⟹ pair_sat (Σ2 σ) (Σ2 σ') Q2"
using assms(2) hoare_triple_validE by blast
define pairs where "pairs = (λ(s, h). { ((s, h1), (s, h2)) |h1 h2. Some h = Some h1 ⊕ Some h2 ∧ (s, h1), (s, h1) ⊨ P1 ∧ (s, h2), (s, h2) ⊨ P2 })"
define Σ where "Σ = (λσ. ⋃(σ1, σ2) ∈ pairs σ. add_states (upperize (Σ1 σ1) (fvA Q1)) (upperize (Σ2 σ2) (fvA Q2)))"
show ?thesis
proof (rule hoare_triple_validI_bounded)
show "⋀s h n. (s, h), (s, h) ⊨ Star P1 P2 ⟹ bounded h ⟹ safe n Δ (Cpar C1 C2) (s, h) (Σ (s, h))"
proof -
fix s h n assume "(s, h), (s, h) ⊨ Star P1 P2" "bounded h"
then obtain h1 h2 where asm0: "Some h = Some h1 ⊕ Some h2" "(s, h1), (s, h1) ⊨ P1"
"(s, h2), (s, h2) ⊨ P2"
using always_sat_refl hyper_sat.simps(4) by blast
then have "((s, h1), (s, h2)) ∈ pairs (s, h)"
using pairs_def by blast
then have "add_states (upperize (Σ1 (s, h1)) (fvA Q1)) (upperize (Σ2 (s, h2)) (fvA Q2)) ⊆ Σ (s, h)"
using Σ_def by blast
moreover have "safe n Δ (Cpar C1 C2) (s, h) (add_states (upperize (Σ1 (s, h1)) (fvA Q1)) (upperize (Σ2 (s, h2)) (fvA Q2)))"
proof (rule parallel_comp)
show "safe n Δ C1 (s, h1) (upperize (Σ1 (s, h1)) (fvA Q1))"
by (metis ‹bounded h› asm0(1) asm0(2) bounded_smaller_sum r1(1) safe_larger_set_aux snd_conv upperize_larger)
show "safe n Δ C2 (s, h2) (upperize (Σ2 (s, h2)) (fvA Q2))"
by (metis ‹bounded h› asm0(1) asm0(3) bounded_smaller_sum plus_comm r2(1) safe_larger_set_aux snd_conv upperize_larger)
show "Some h = Some h1 ⊕ Some h2" using asm0 by simp
show "disjoint (fvC C1 ∪ fvA Q1) (wrC C2)"
by (metis Un_subset_iff assms(3) disjoint_def inf_shunt)
show "disjoint (fvC C2 ∪ fvA Q2) (wrC C1)"
by (metis Un_subset_iff assms(4) disjoint_def inf_shunt)
show "upper_fvs (upperize (Σ1 (s, h1)) (fvA Q1)) (fvA Q1)"
by (simp add: upper_fvs_upperize)
show "upper_fvs (upperize (Σ2 (s, h2)) (fvA Q2)) (fvA Q2)"
using upper_fvs_upperize by auto
show "⋀Γ. Δ = Some Γ ⟹ disjoint (fvA (invariant Γ)) (wrC C2)"
using assms(5) by auto
show "⋀Γ. Δ = Some Γ ⟹ disjoint (fvA (invariant Γ)) (wrC C1)"
using assms(6) by blast
show "bounded h"
by (simp add: ‹bounded h›)
qed
ultimately show "safe n Δ (Cpar C1 C2) (s, h) (Σ (s, h))"
using safe_larger_set by blast
qed
fix s h s' h'
assume "(s, h), (s', h') ⊨ Star P1 P2"
then obtain h1 h2 h1' h2' where asm0: "Some h = Some h1 ⊕ Some h2" "Some h' = Some h1' ⊕ Some h2'"
"(s, h1), (s', h1') ⊨ P1" "(s, h2), (s', h2') ⊨ P2"
by auto
show "pair_sat (Σ (s, h)) (Σ (s', h')) (Star Q1 Q2)"
proof (rule pair_satI)
fix ss hh ss' hh' assume asm1: "(ss, hh) ∈ Σ (s, h) ∧ (ss', hh') ∈ Σ (s', h')"
then obtain σ1 σ2 σ1' σ2' where "(σ1, σ2) ∈ pairs (s, h)" "(σ1', σ2') ∈ pairs (s', h')"
"(ss, hh) ∈ add_states (upperize (Σ1 σ1) (fvA Q1)) (upperize (Σ2 σ2) (fvA Q2))"
"(ss', hh') ∈ add_states (upperize (Σ1 σ1') (fvA Q1)) (upperize (Σ2 σ2') (fvA Q2))"
using Σ_def by blast
then obtain "fst σ1 = s" "fst σ2 = s" "fst σ1' = s'" "fst σ2' = s'" "Some h = Some (snd σ1) ⊕ Some (snd σ2)"
"Some h' = Some (snd σ1') ⊕ Some (snd σ2')"
"(s, snd σ1), (s, snd σ1) ⊨ P1 ∧ (s, snd σ2), (s, snd σ2) ⊨ P2"
"(s', snd σ1'), (s', snd σ1') ⊨ P1 ∧ (s', snd σ2'), (s', snd σ2') ⊨ P2"
using case_prod_conv pairs_def by auto
moreover have "snd σ1 = h1 ∧ snd σ2 = h2 ∧ snd σ1' = h1' ∧ snd σ2' = h2'"
proof (cases "precise P1")
case True
then have "snd σ1 = h1 ∧ snd σ1' = h1'"
proof (rule preciseE)
show "h ≽ h1 ∧ h ≽ snd σ1 ∧ h' ≽ h1' ∧ h' ≽ snd σ1'"
using asm0(1) asm0(2) calculation(5) calculation(6) larger_def by blast
show "(s, h1), (s', h1') ⊨ P1 ∧ (s, snd σ1), (s', snd σ1') ⊨ P1"
by (metis True ‹h ≽ h1 ∧ h ≽ snd σ1 ∧ h' ≽ h1' ∧ h' ≽ snd σ1'› always_sat_refl asm0(3) calculation(7) calculation(8) preciseE sat_comm)
qed
then show ?thesis
by (metis addition_cancellative asm0(1) asm0(2) calculation(5) calculation(6) plus_comm)
next
case False
then have "precise P2"
using assms(7) by blast
then have "snd σ2 = h2 ∧ snd σ2' = h2'"
proof (rule preciseE)
show "h ≽ h2 ∧ h ≽ snd σ2 ∧ h' ≽ h2' ∧ h' ≽ snd σ2'"
by (metis asm0(1) asm0(2) calculation(5) calculation(6) larger_def plus_comm)
show "(s, h2), (s', h2') ⊨ P2 ∧ (s, snd σ2), (s', snd σ2') ⊨ P2"
by (metis ‹h ≽ h2 ∧ h ≽ snd σ2 ∧ h' ≽ h2' ∧ h' ≽ snd σ2'› ‹precise P2› always_sat_refl asm0(4) calculation(7) calculation(8) preciseE sat_comm)
qed
then show ?thesis
using addition_cancellative asm0(1) asm0(2) calculation(5) calculation(6) by blast
qed
ultimately have "pair_sat (Σ1 σ1) (Σ1 σ1') Q1 ∧ pair_sat (Σ2 σ2) (Σ2 σ2') Q2"
by (metis asm0(3) asm0(4) prod.exhaust_sel r1(2) r2(2))
then show "(ss, hh), (ss', hh') ⊨ Star Q1 Q2"
by (metis (no_types, opaque_lifting) ‹(ss', hh') ∈ add_states (upperize (Σ1 σ1') (fvA Q1)) (upperize (Σ2 σ2') (fvA Q2))› ‹(ss, hh) ∈ add_states (upperize (Σ1 σ1) (fvA Q1)) (upperize (Σ2 σ2) (fvA Q2))› add_states_sat_star pair_sat_comm pair_sat_def pair_sat_upperize)
qed
qed
qed
subsubsection ‹If›
lemma if_cases:
assumes "red (Cif b C1 C2) (s, h) C' (s', h')"
and "C' = C1 ⟹ s = s' ∧ h = h' ⟹ bdenot b s ⟹ P"
and "C' = C2 ⟹ s = s' ∧ h = h' ⟹ ¬ bdenot b s ⟹ P"
shows P
using assms(1)
apply (rule red.cases)
apply blast+
using assms(2) apply fastforce
using assms(3) apply fastforce
apply blast+
done
lemma if_safe_None:
fixes Δ :: "('i, 'a, nat) cont"
assumes "bdenot b s ⟹ safe n Δ C1 (s, h) S"
and "¬ bdenot b s ⟹ safe n Δ C2 (s, h) S"
and "Δ = None"
shows "safe (Suc n) (None :: ('i, 'a, nat) cont) (Cif b C1 C2) (s, h) S"
proof (rule safeNoneI)
show "Cif b C1 C2 = Cskip ⟹ (s, h) ∈ S" by simp
show "no_abort (None :: ('i, 'a, nat) cont) (Cif b C1 C2) s h"
proof (rule no_abortNoneI)
fix hf H assume "Some H = Some h ⊕ Some hf ∧ full_ownership (get_fh H) ∧ no_guard H"
show "¬ aborts (Cif b C1 C2) (s, FractionalHeap.normalize (get_fh H))"
proof (rule ccontr)
assume "¬ ¬ aborts (Cif b C1 C2) (s, FractionalHeap.normalize (get_fh H))"
then have "aborts (Cif b C1 C2) (s, FractionalHeap.normalize (get_fh H))" by simp
then show False
by (rule aborts.cases) auto
qed
qed
fix H hf C' s' h'
assume asm0: "Some H = Some h ⊕ Some hf ∧ full_ownership (get_fh H) ∧ no_guard H
∧ red (Cif b C1 C2) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
show "∃h'' H'.
full_ownership (get_fh H') ∧
no_guard H' ∧ h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hf ∧ safe n (None :: ('i, 'a, nat) cont) C' (s', h'') S"
by (metis asm0 assms(1) assms(2) assms(3) if_cases)
qed (simp)
lemma if_safe_Some:
assumes "bdenot b s ⟹ safe n (Some Γ) C1 (s, h) S"
and "¬ bdenot b s ⟹ safe n (Some Γ) C2 (s, h) S"
shows "safe (Suc n) (Some Γ) (Cif b C1 C2) (s, h) S"
proof (rule safeSomeI)
show "Cif b C1 C2 = Cskip ⟹ (s, h) ∈ S" by simp
show "no_abort (Some Γ) (Cif b C1 C2) s h"
proof (rule no_abortSomeI)
fix H hf hj v0
assume asm0: "Some H = Some h ⊕ Some hj ⊕ Some hf ∧ full_ownership (get_fh H) ∧ semi_consistent Γ v0 H ∧ sat_inv s hj Γ"
show "¬ aborts (Cif b C1 C2) (s, FractionalHeap.normalize (get_fh H))"
proof (rule ccontr)
assume "¬ ¬ aborts (Cif b C1 C2) (s, FractionalHeap.normalize (get_fh H))"
then have "aborts (Cif b C1 C2) (s, FractionalHeap.normalize (get_fh H))" by simp
then show False
by (rule aborts.cases) auto
qed
qed
fix H hf C' s' h' hj v0
assume asm0: "Some H = Some h ⊕ Some hj ⊕ Some hf ∧ full_ownership (get_fh H) ∧ semi_consistent Γ v0 H
∧ sat_inv s hj Γ ∧ red (Cif b C1 C2) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
show "∃h'' H' hj'.
full_ownership (get_fh H') ∧
semi_consistent Γ v0 H' ∧
sat_inv s' hj' Γ ∧ h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hj' ⊕ Some hf ∧ safe n (Some Γ) C' (s', h'') S"
by (metis asm0 assms(1) assms(2) if_cases)
qed (simp)
lemma if_safe:
fixes Δ :: "('i, 'a, nat) cont"
assumes "bdenot b s ⟹ safe n Δ C1 (s, h) S"
and "¬ bdenot b s ⟹ safe n Δ C2 (s, h) S"
shows "safe (Suc n) Δ (Cif b C1 C2) (s, h) S"
apply (cases Δ)
using assms(1) assms(2) if_safe_None apply blast
using assms(1) assms(2) if_safe_Some by blast
theorem if1_rule:
fixes Δ :: "('i, 'a, nat) cont"
assumes "hoare_triple_valid Δ (And P (Bool b)) C1 Q"
and "hoare_triple_valid Δ (And P (Bool (Bnot b))) C2 Q"
shows "hoare_triple_valid Δ (And P (Low b)) (Cif b C1 C2) Q"
proof -
obtain Σt where safe_t: "⋀σ n. σ, σ ⊨ And P (Bool b) ⟹ bounded (snd σ) ⟹ safe n Δ C1 σ (Σt σ)"
"⋀σ σ'. σ, σ' ⊨ And P (Bool b) ⟹ pair_sat (Σt σ) (Σt σ') Q"
using assms(1) hoare_triple_validE by blast
obtain Σf where safe_f: "⋀σ n. σ, σ ⊨ And P (Bool (Bnot b)) ⟹ bounded (snd σ) ⟹ safe n Δ C2 σ (Σf σ)"
"⋀σ σ'. σ, σ' ⊨ And P (Bool (Bnot b)) ⟹ pair_sat (Σf σ) (Σf σ') Q"
using assms(2) hoare_triple_validE by blast
define Σ where "Σ = (λσ. if bdenot b (fst σ) then Σt σ else Σf σ)"
show ?thesis
proof (rule hoare_triple_valid_smallerI_bounded)
show "⋀σ n. σ, σ ⊨ And P (Low b) ⟹ bounded (snd σ) ⟹ safe n Δ (Cif b C1 C2) σ (Σ σ)"
proof -
fix σ n
assume asm0: "σ, σ ⊨ And P (Low b)" "bounded (snd σ)"
show "safe n Δ (Cif b C1 C2) σ (Σ σ)"
proof (cases "bdenot b (fst σ)")
case True
then have "safe n Δ C1 σ (Σ σ)"
by (metis Σ_def asm0 hyper_sat.simps(1) hyper_sat.simps(3) prod.exhaust_sel safe_t(1))
then show ?thesis
by (metis (no_types, lifting) Suc_n_not_le_n True if_safe nat_le_linear prod.exhaust_sel safe_smaller)
next
case False
then have "safe n Δ C2 σ (Σ σ)"
by (metis Σ_def asm0 bdenot.simps(3) hyper_sat.simps(1) hyper_sat.simps(3) prod.exhaust_sel safe_f(1))
then show ?thesis
by (metis (mono_tags) False Suc_n_not_le_n if_safe nat_le_linear prod.exhaust_sel safe_smaller)
qed
qed
fix σ σ' assume asm0: "σ, σ' ⊨ And P (Low b)"
show "pair_sat (Σ σ) (Σ σ') Q"
proof (cases "bdenot b (fst σ)")
case True
then show ?thesis
by (metis (no_types, lifting) Σ_def asm0 hyper_sat.simps(1) hyper_sat.simps(3) hyper_sat.simps(5) prod.exhaust_sel safe_t(2))
next
case False
then show ?thesis
by (metis (no_types, lifting) Σ_def asm0 bdenot.simps(3) hyper_sat.simps(1) hyper_sat.simps(3) hyper_sat.simps(5) prod.exhaust_sel safe_f(2))
qed
qed
qed
theorem if2_rule:
fixes Δ :: "('i, 'a, nat) cont"
assumes "hoare_triple_valid Δ (And P (Bool b)) C1 Q"
and "hoare_triple_valid Δ (And P (Bool (Bnot b))) C2 Q"
and "unary Q"
shows "hoare_triple_valid Δ P (Cif b C1 C2) Q"
proof -
obtain Σt where safe_t: "⋀σ n. σ, σ ⊨ And P (Bool b) ⟹ bounded (snd σ) ⟹ safe n Δ C1 σ (Σt σ)"
"⋀σ σ'. σ, σ' ⊨ And P (Bool b) ⟹ pair_sat (Σt σ) (Σt σ') Q"
using assms(1) hoare_triple_validE by blast
obtain Σf where safe_f: "⋀σ n. σ, σ ⊨ And P (Bool (Bnot b)) ⟹ bounded (snd σ) ⟹ safe n Δ C2 σ (Σf σ)"
"⋀σ σ'. σ, σ' ⊨ And P (Bool (Bnot b)) ⟹ pair_sat (Σf σ) (Σf σ') Q"
using assms(2) hoare_triple_validE by blast
define Σ where "Σ = (λσ. if bdenot b (fst σ) then Σt σ else Σf σ)"
show ?thesis
proof (rule hoare_triple_valid_smallerI_bounded)
show "⋀σ n. σ, σ ⊨ P ⟹ bounded (snd σ) ⟹ safe n Δ (Cif b C1 C2) σ (Σ σ)"
proof -
fix σ n
assume asm0: "σ, σ ⊨ P" "bounded (snd σ)"
show "safe n Δ (Cif b C1 C2) σ (Σ σ)"
proof (cases "bdenot b (fst σ)")
case True
then have "safe n Δ C1 σ (Σ σ)"
by (metis Σ_def asm0 hyper_sat.simps(1) hyper_sat.simps(3) prod.exhaust_sel safe_t(1))
then show ?thesis
by (metis (no_types, lifting) Suc_n_not_le_n True if_safe nat_le_linear prod.exhaust_sel safe_smaller)
next
case False
then have "safe n Δ C2 σ (Σ σ)"
by (metis Σ_def asm0 bdenot.simps(3) hyper_sat.simps(1) hyper_sat.simps(3) prod.exhaust_sel safe_f(1))
then show ?thesis
by (metis (mono_tags) False Suc_n_not_le_n if_safe nat_le_linear prod.exhaust_sel safe_smaller)
qed
qed
fix σ1 σ2 assume asm0: "σ1, σ2 ⊨ P"
then have asm0_bis: "σ2, σ1 ⊨ P"
by (simp add: sat_comm)
show "pair_sat (Σ σ1) (Σ σ2) Q"
proof (rule pair_sat_smallerI)
fix σ1' σ2'
assume asm1: "σ1' ∈ Σ σ1 ∧ σ2' ∈ Σ σ2"
then have "σ1', σ1' ⊨ Q"
apply (cases "bdenot b (fst σ1)")
apply (metis (no_types, lifting) Σ_def always_sat_refl asm0 hyper_sat.simps(1) hyper_sat.simps(3) pair_sat_def safe_t(2) surjective_pairing)
by (metis (no_types, lifting) Σ_def always_sat_refl asm0 bdenot.simps(3) hyper_sat.simps(1) hyper_sat.simps(3) pair_satE prod.collapse safe_f(2))
moreover have "σ2', σ2' ⊨ Q"
apply (cases "bdenot b (fst σ2)")
apply (metis (mono_tags) Σ_def always_sat_refl asm0_bis asm1 entailsI entails_def fst_conv hyper_sat.simps(1) hyper_sat.simps(3) old.prod.exhaust pair_sat_def safe_t(2))
using Σ_def always_sat_refl asm0_bis bdenot.simps(3) hyper_sat.simps(1) hyper_sat.simps(3) pair_satE prod.collapse safe_f(2)
by (metis (no_types, lifting) asm1)
ultimately show "σ1', σ2' ⊨ Q"
by (metis assms(3) eq_fst_iff unaryE)
qed
qed
qed
subsubsection ‹Sequential composition›
inductive_cases red_seq_cases: "red (Cseq C1 C2) σ C' σ'"
lemma aborts_seq_aborts_C1:
assumes "aborts (Cseq C1 C2) σ"
shows "aborts C1 σ"
using aborts.simps assms cmd.inject(6) by blast
lemma safe_seq_None:
assumes "safe n (None :: ('i, 'a, nat) cont) C1 (s, h) S1"
and "⋀m s' h'. m ≤ n ∧ (s', h') ∈ S1 ⟹ safe m (None :: ('i, 'a, nat) cont) C2 (s', h') S2"
shows "safe n (None :: ('i, 'a, nat) cont) (Cseq C1 C2) (s, h) S2"
using assms
proof (induct n arbitrary: C1 s h)
case (Suc n)
show ?case
proof (rule safeNoneI)
show "no_abort (None :: ('i, 'a, nat) cont) (Cseq C1 C2) s h"
by (meson Suc.prems(1) aborts_seq_aborts_C1 no_abort.simps(1) safeNoneE_bis(2))
show "accesses (Cseq C1 C2) s ⊆ dom (fst h) ∧ writes (Cseq C1 C2) s ⊆ fpdom (fst h)"
by (metis Suc.prems(1) accesses.simps(7) fst_conv safeAccessesE snd_conv writes.simps(7))
fix H hf C' s' h'
assume asm0: "Some H = Some h ⊕ Some hf ∧
full_ownership (get_fh H) ∧ no_guard H ∧ red (Cseq C1 C2) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
show "∃h'' H'.
full_ownership (get_fh H') ∧
no_guard H' ∧ h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hf ∧ safe n (None :: ('i, 'a, nat) cont) C' (s', h'') S2"
proof (rule red_seq_cases)
show "red (Cseq C1 C2) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
using asm0 by blast
show "C1 = Cskip ⟹
C' = C2 ⟹
(s', h') = (s, FractionalHeap.normalize (get_fh H)) ⟹
∃h'' H'. full_ownership (get_fh H') ∧
no_guard H' ∧ h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hf ∧ safe n (None :: ('i, 'a, nat) cont) C' (s', h'') S2"
using Suc.prems(1) Suc.prems(2) asm0 order_refl prod.inject safeNoneE_bis(1)
by (metis le_SucI)
fix C1' assume "C' = Cseq C1' C2" "red C1 (s, FractionalHeap.normalize (get_fh H)) C1' (s', h')"
obtain H' h'' where asm1: "full_ownership (get_fh H')" "no_guard H'" "h' = FractionalHeap.normalize (get_fh H')"
"Some H' = Some h'' ⊕ Some hf" "safe n (None :: ('i, 'a, nat) cont) C1' (s', h'') S1"
using Suc(2) safeNoneE(3)[of n C1 s h S1 H hf C1' s' h']
using ‹red C1 (s, FractionalHeap.normalize (get_fh H)) C1' (s', h')› asm0 by blast
moreover have "safe n (None :: ('i, 'a, nat) cont) (Cseq C1' C2) (s', h'') S2"
using Suc.hyps Suc.prems(2) calculation(5)
using le_Suc_eq by presburger
ultimately show ?thesis
using ‹C' = Cseq C1' C2› by blast
qed
qed (simp)
qed (simp)
lemma safe_seq_Some:
assumes "safe n (Some Γ) C1 (s, h) S1"
and "⋀m s' h'. m ≤ n ∧ (s', h') ∈ S1 ⟹ safe m (Some Γ) C2 (s', h') S2"
shows "safe n (Some Γ) (Cseq C1 C2) (s, h) S2"
using assms
proof (induct n arbitrary: C1 s h)
case (Suc n)
show ?case
proof (rule safeSomeI)
show "no_abort (Some Γ) (Cseq C1 C2) s h"
by (meson Suc.prems(1) aborts_seq_aborts_C1 no_abort.simps(2) safeSomeE(2))
show "accesses (Cseq C1 C2) s ⊆ dom (fst h) ∧ writes (Cseq C1 C2) s ⊆ fpdom (fst h)"
by (metis Suc.prems(1) accesses.simps(7) fst_conv safeAccessesE snd_conv writes.simps(7))
fix H hf C' s' h' hj v0
assume asm0: "Some H = Some h ⊕ Some hj ⊕ Some hf ∧
full_ownership (get_fh H) ∧ semi_consistent Γ v0 H ∧ sat_inv s hj Γ ∧ red (Cseq C1 C2) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
show "∃h'' H' hj'. full_ownership (get_fh H') ∧
semi_consistent Γ v0 H' ∧ sat_inv s' hj' Γ ∧ h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hj' ⊕ Some hf ∧ safe n (Some Γ) C' (s', h'') S2"
proof (rule red_seq_cases)
show "red (Cseq C1 C2) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
using asm0 by blast
show "C1 = Cskip ⟹
C' = C2 ⟹
(s', h') = (s, FractionalHeap.normalize (get_fh H)) ⟹ ∃h'' H' hj'. full_ownership (get_fh H') ∧
semi_consistent Γ v0 H' ∧ sat_inv s' hj' Γ ∧ h' = FractionalHeap.normalize (get_fh H')
∧ Some H' = Some h'' ⊕ Some hj' ⊕ Some hf ∧ safe n (Some Γ) C' (s', h'') S2"
using Pair_inject Suc.prems(1) Suc_n_not_le_n asm0 assms(2) not_less_eq_eq safeSomeE(1)
by (metis (no_types, lifting) Suc.prems(2) nat_le_linear)
fix C1' assume "C' = Cseq C1' C2" "red C1 (s, FractionalHeap.normalize (get_fh H)) C1' (s', h')"
obtain H' h'' hj' where asm1: "full_ownership (get_fh H') ∧
semi_consistent Γ v0 H' ∧ sat_inv s' hj' Γ ∧ h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hj' ⊕ Some hf ∧ safe n (Some Γ) C1' (s', h'') S1"
using Suc(2) safeSomeE(3)[of n Γ C1 s h S1 H hj hf v0 C1' s' h']
using ‹red C1 (s, FractionalHeap.normalize (get_fh H)) C1' (s', h')› asm0 by blast
moreover have "safe n (Some Γ) (Cseq C1' C2) (s', h'') S2"
by (simp add: Suc.hyps Suc.prems(2) calculation)
ultimately show ?thesis
using ‹C' = Cseq C1' C2› by blast
qed
qed (simp)
qed (simp)
lemma seq_safe:
fixes Δ :: "('i, 'a, nat) cont"
assumes "safe n Δ C1 (s, h) S1"
and "⋀m s' h'. m ≤ n ∧ (s', h') ∈ S1 ⟹ safe m Δ C2 (s', h') S2"
shows "safe n Δ (Cseq C1 C2) (s, h) S2"
apply (cases Δ)
using assms(1) assms(2) safe_seq_None apply blast
using assms(1) assms(2) safe_seq_Some by blast
theorem seq_rule:
fixes Δ :: "('i, 'a, nat) cont"
assumes "hoare_triple_valid Δ P C1 R"
and "hoare_triple_valid Δ R C2 Q"
shows "hoare_triple_valid Δ P (Cseq C1 C2) Q"
proof -
obtain Σ1 where safe_1: "⋀σ n. σ, σ ⊨ P ⟹ bounded (snd σ) ⟹ safe n Δ C1 σ (Σ1 σ)"
"⋀σ σ'. σ, σ' ⊨ P ⟹ pair_sat (Σ1 σ) (Σ1 σ') R"
using assms(1) hoare_triple_validE by blast
obtain Σ2 where safe_2: "⋀σ n. σ, σ ⊨ R ⟹ bounded (snd σ) ⟹ safe n Δ C2 σ (Σ2 σ)"
"⋀σ σ'. σ, σ' ⊨ R ⟹ pair_sat (Σ2 σ) (Σ2 σ') Q"
using assms(2) hoare_triple_validE by blast
define Σ where "Σ = (λσ. (⋃σ' ∈ Σ1 σ. Σ2 σ'))"
show ?thesis
proof (rule hoare_triple_valid_smallerI_bounded)
show "⋀σ n. σ, σ ⊨ P ⟹ bounded (snd σ) ⟹ safe n Δ (Cseq C1 C2) σ (Σ σ)"
proof -
fix σ n assume asm0: "σ, σ ⊨ P" "bounded (snd σ)"
then have "pair_sat (Σ1 σ) (Σ1 σ) R"
using safe_1(2) by blast
have "safe n Δ (Cseq C1 C2) (fst σ, snd σ) (Σ σ)"
proof (rule seq_safe)
thm restrict_safe_to_bounded
show "safe n Δ C1 (fst σ, snd σ) (Set.filter (bounded ∘ snd) (Σ1 σ))"
apply (rule restrict_safe_to_bounded)
using asm0 safe_1(1) by simp_all
fix m s' h'
assume "m ≤ n ∧ (s', h') ∈ Set.filter (bounded ∘ snd) (Σ1 σ)"
then have "safe m Δ C2 (s', h') (Σ2 (s', h'))"
using safe_2(1)[of "(s', h')" m] ‹pair_sat (Σ1 σ) (Σ1 σ) R› unfolding pair_sat_def
by simp
then show "safe m Δ C2 (s', h') (Σ σ)"
using Sup_upper ‹Σ ≡ λσ. ⋃ (Σ2 ` Σ1 σ)› ‹pair_sat (Σ1 σ) (Σ1 σ) R› image_iff safe_larger_set asm0(2)
by (metis (no_types, lifting) ‹m ≤ n ∧ (s', h') ∈ Set.filter (bounded ∘ snd) (Σ1 σ)› member_filter)
qed
then show "safe n Δ (Cseq C1 C2) σ (Σ σ)" by auto
qed
fix σ1 σ2
assume asm0: "σ1, σ2 ⊨ P"
show "pair_sat (Σ σ1) (Σ σ2) Q"
proof (rule pair_sat_smallerI)
fix σ1'' σ2''
assume asm1: "σ1'' ∈ Σ σ1 ∧ σ2'' ∈ Σ σ2"
then obtain σ1' σ2' where "σ1'' ∈ Σ2 σ1'" "σ1' ∈ Σ1 σ1" "σ2'' ∈ Σ2 σ2'" "σ2' ∈ Σ1 σ2"
using ‹Σ ≡ λσ. ⋃ (Σ2 ` Σ1 σ)› by blast
then show "σ1'', σ2'' ⊨ Q"
by (meson asm0 pair_sat_def safe_1(2) safe_2(2))
qed
qed
qed
subsubsection ‹Frame rule›
lemma safe_frame_None:
assumes "safe n (None :: ('i, 'a, nat) cont) C (s, h) S"
and "Some H = Some h ⊕ Some hf0"
and "bounded H"
shows "safe n (None :: ('i, 'a, nat) cont) C (s, H) (add_states S {(s'', hf0) |s''. agrees (- wrC C) s s''})"
using assms
proof (induct n arbitrary: s h H C)
case (Suc n)
show "safe (Suc n) (None :: ('i, 'a, nat) cont) C (s, H) (add_states S {(s'', hf0) |s''. agrees (- wrC C) s s''} )"
proof (rule safeNoneI)
show "C = Cskip ⟹ (s, H) ∈ add_states S {(s', hf0) |s'. agrees (- wrC C) s s'}"
using CollectI Suc.prems(1) Suc.prems(2) add_states_def agrees_def[of "- wrC C" s] safeNoneE(1)[of n C s h S]
by fast
show "no_abort (None :: ('i, 'a, nat) cont) C s H"
using Suc.prems(1) Suc.prems(2) larger_def no_abort_larger safeNoneE(2) by blast
have "accesses C s ⊆ dom (fst h) ∧ writes C s ⊆ fpdom (fst h)"
using Suc.prems(1) by auto
moreover have "dom (fst h) ⊆ dom (fst H)"
by (metis Suc.prems(2) addition_smaller_domain get_fh.simps)
moreover have "fpdom (fst h) ⊆ fpdom (fst H)"
using Suc.prems(2) fpdom_inclusion
using Suc.prems(3) by blast
ultimately show "accesses C s ⊆ dom (fst H) ∧ writes C s ⊆ fpdom (fst H)"
by blast
fix H1 hf1 C' s' h'
assume asm0: "Some H1 = Some H ⊕ Some hf1 ∧ full_ownership (get_fh H1) ∧ no_guard H1 ∧ red C (s, FractionalHeap.normalize (get_fh H1)) C' (s', h')"
then obtain hf where "Some hf = Some hf0 ⊕ Some hf1"
by (metis (no_types, opaque_lifting) Suc.prems(2) option.collapse plus.simps(1) plus_asso plus_comm)
then have "Some H1 = Some h ⊕ Some hf"
by (metis Suc.prems(2) asm0 plus_asso)
then obtain h'' H' where r: "full_ownership (get_fh H')"
"no_guard H'" "h' = FractionalHeap.normalize (get_fh H')" "Some H' = Some h'' ⊕ Some hf" "safe n (None :: ('i, 'a, nat) cont) C' (s', h'') S"
using safeNoneE(3)[of n C s h S H1 hf C' s'] Suc.prems(1) asm0 by blast
then obtain h''' where "Some h''' = Some h'' ⊕ Some hf0"
by (metis ‹Some hf = Some hf0 ⊕ Some hf1› not_Some_eq plus.simps(1) plus_asso)
then have "Some H' = Some h''' ⊕ Some hf1"
by (metis ‹Some hf = Some hf0 ⊕ Some hf1› plus_asso r(4))
moreover have "safe n (None :: ('i, 'a, nat) cont) C' (s', h''') (add_states S {(s'', hf0) |s''. agrees (- wrC C') s' s''})"
proof (rule Suc.hyps)
show "safe n (None :: ('i, 'a, nat) cont) C' (s', h'') S"
using r by simp
show "Some h''' = Some h'' ⊕ Some hf0"
by (simp add: ‹Some h''' = Some h'' ⊕ Some hf0›)
show "bounded h'''"
by (metis bounded_smaller_sum calculation full_ownership_then_bounded get_fh.simps r(1))
qed
moreover have "add_states S {(s'', hf0) |s''. agrees (- wrC C') s' s''} ⊆ add_states S {(s'', hf0) |s''. agrees (- wrC C) s s''}"
proof -
have "wrC C' ⊆ wrC C"
using asm0 red_properties(1) by blast
have "{(s'', hf0) |s''. agrees (- wrC C') s' s''} ⊆ {(s'', hf0) |s''. agrees (- wrC C) s s''}"
proof
fix x assume "x ∈ {(s'', hf0) |s''. agrees (- wrC C') s' s''}"
then have "agrees (- wrC C') s' (fst x) ∧ snd x = hf0" by force
moreover have "fvC C' ⊆ fvC C ∧ wrC C' ⊆ wrC C ∧ agrees (- wrC C) s' s"
using asm0 red_properties(1) by force
moreover have "agrees (- wrC C) s (fst x)"
proof (rule agreesI)
fix y assume "y ∈ - wrC C"
show "s y = fst x y"
by (metis (no_types, lifting) Compl_subset_Compl_iff ‹y ∈ - wrC C› agrees_def calculation(1) calculation(2) in_mono)
qed
then show "x ∈ {(s'', hf0) |s''. agrees (- wrC C) s s''}"
using ‹agrees (- wrC C') s' (fst x) ∧ snd x = hf0› by force
qed
then show ?thesis
by (metis (no_types, lifting) add_states_comm add_states_subset)
qed
ultimately have "safe n (None :: ('i, 'a, nat) cont) C' (s', h''') (add_states S {(s'', hf0) |s''. agrees (- wrC C) s s''})"
using safe_larger_set by blast
then show "∃h'' H'.
full_ownership (get_fh H') ∧
no_guard H' ∧
h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hf1 ∧ safe n (None :: ('i, 'a, nat) cont) C' (s', h'') (add_states S {(s'', hf0) |s''. agrees (- wrC C) s s''})"
using ‹Some H' = Some h''' ⊕ Some hf1› r(1) r(2) r(3) by blast
qed
qed (simp)
lemma safe_frame_Some:
assumes "safe n (Some Γ) C (s, h) S"
and "Some H = Some h ⊕ Some hf0"
and "bounded H"
shows "safe n (Some Γ) C (s, H) (add_states S {(s'', hf0) |s''. agrees (- wrC C) s s''})"
using assms
proof (induct n arbitrary: s h H C)
case (Suc n)
let ?R = "{(s'', hf0) |s''. agrees (- wrC C) s s''}"
show "safe (Suc n) (Some Γ) C (s, H) (add_states S ?R)"
proof (rule safeSomeI)
show "C = Cskip ⟹ (s, H) ∈ add_states S ?R"
using CollectI Suc.prems(1) Suc.prems(2) add_states_def[of S ?R] agrees_def[of "- wrC C" s]
safeSomeE(1)[of n Γ C s h S] by fast
show "no_abort (Some Γ) C s H"
using Suc.prems(1) Suc.prems(2) larger_def no_abort_larger safeSomeE(2) by blast
have "accesses C s ⊆ dom (fst h) ∧ writes C s ⊆ fpdom (fst h)"
using Suc.prems(1) by auto
moreover have "dom (fst h) ⊆ dom (fst H)"
by (metis Suc.prems(2) addition_smaller_domain get_fh.simps)
moreover have "fpdom (fst h) ⊆ fpdom (fst H)"
using Suc.prems(2) Suc.prems(3) fpdom_inclusion by blast
ultimately show "accesses C s ⊆ dom (fst H) ∧ writes C s ⊆ fpdom (fst H)"
by blast
fix H1 hf1 C' s' h' hj v0
assume asm0: "Some H1 = Some H ⊕ Some hj ⊕ Some hf1 ∧
full_ownership (get_fh H1) ∧ semi_consistent Γ v0 H1 ∧ sat_inv s hj Γ ∧ red C (s, FractionalHeap.normalize (get_fh H1)) C' (s', h')"
then obtain hf where "Some hf = Some hf0 ⊕ Some hf1"
by (metis (no_types, opaque_lifting) Suc.prems(2) option.collapse plus.simps(1) plus_asso plus_comm)
then have "Some H1 = Some h ⊕ Some hj ⊕ Some hf"
by (metis (no_types, opaque_lifting) Suc.prems(2) asm0 plus_asso plus_comm)
then obtain h'' H' hj' where r: "full_ownership (get_fh H') ∧
semi_consistent Γ v0 H' ∧ sat_inv s' hj' Γ ∧ h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hj' ⊕ Some hf ∧ safe n (Some Γ) C' (s', h'') S"
using safeSomeE(3)[of n Γ C s h S H1 hj hf v0 C' s' h'] Suc.prems(1) asm0 by blast
then obtain h''' where "Some h''' = Some h'' ⊕ Some hf0"
by (metis (no_types, lifting) ‹Some hf = Some hf0 ⊕ Some hf1› plus.simps(2) plus.simps(3) plus_asso plus_comm)
then have "Some H' = Some h''' ⊕ Some hj' ⊕ Some hf1"
by (metis (no_types, lifting) ‹Some hf = Some hf0 ⊕ Some hf1› plus_asso plus_comm r)
moreover have "safe n (Some Γ) C' (s', h''') (add_states S {(s'', hf0) |s''. agrees (- wrC C') s' s''})"
proof (rule Suc.hyps)
show "safe n (Some Γ) C' (s', h'') S"
using r by simp
show "Some h''' = Some h'' ⊕ Some hf0"
by (simp add: ‹Some h''' = Some h'' ⊕ Some hf0›)
show "bounded h'''"
apply (rule bounded_smaller[of H'])
apply (metis full_ownership_then_bounded get_fh.simps r)
by (simp add: calculation larger3 plus_comm)
qed
moreover have "add_states S {(s'', hf0) |s''. agrees (- wrC C') s' s''} ⊆ add_states S {(s'', hf0) |s''. agrees (- wrC C) s s''}"
proof -
have "wrC C' ⊆ wrC C"
using asm0 red_properties(1) by blast
have "{(s'', hf0) |s''. agrees (- wrC C') s' s''} ⊆ {(s'', hf0) |s''. agrees (- wrC C) s s''}"
proof
fix x assume "x ∈ {(s'', hf0) |s''. agrees (- wrC C') s' s''}"
then have "agrees (- wrC C') s' (fst x) ∧ snd x = hf0" by force
moreover have "fvC C' ⊆ fvC C ∧ wrC C' ⊆ wrC C ∧ agrees (- wrC C) s' s"
using asm0 red_properties(1) by force
moreover have "agrees (- wrC C) s (fst x)"
proof (rule agreesI)
fix y assume "y ∈ - wrC C"
then show "s y = fst x y"
by (metis (mono_tags, opaque_lifting) Compl_iff agrees_def calculation(1) calculation(2) in_mono)
qed
then show "x ∈ {(s'', hf0) |s''. agrees (- wrC C) s s''}"
using ‹agrees (- wrC C') s' (fst x) ∧ snd x = hf0› by force
qed
then show ?thesis
by (metis (no_types, lifting) add_states_comm add_states_subset)
qed
ultimately have "safe n (Some Γ) C' (s', h''') (add_states S ?R)"
using safe_larger_set by blast
then show "∃h'' H' hj'. full_ownership (get_fh H') ∧
semi_consistent Γ v0 H' ∧
sat_inv s' hj' Γ ∧ h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hj' ⊕ Some hf1 ∧ safe n (Some Γ) C' (s', h'') (add_states S ?R)"
using ‹Some H' = Some h''' ⊕ Some hj' ⊕ Some hf1› r by blast
qed
qed (simp)
lemma safe_frame:
fixes Δ :: "('i, 'a, nat) cont"
assumes "safe n Δ C (s, h) S"
and "Some H = Some h ⊕ Some hf0"
and "bounded H"
shows "safe n Δ C (s, H) (add_states S {(s'', hf0) |s''. agrees (- wrC C) s s''})"
apply (cases Δ)
using assms safe_frame_None apply blast
using assms safe_frame_Some by blast
theorem frame_rule:
fixes Δ :: "('i, 'a, nat) cont"
assumes "hoare_triple_valid Δ P C Q"
and "disjoint (fvA R) (wrC C)"
and "precise P ∨ precise R"
shows "hoare_triple_valid Δ (Star P R) C (Star Q R)"
proof -
obtain Σ where asm0: "⋀σ n. σ, σ ⊨ P ⟹ bounded (snd σ) ⟹ safe n Δ C σ (Σ σ)" "⋀σ σ'. σ, σ' ⊨ P ⟹ pair_sat (Σ σ) (Σ σ') Q"
using assms(1) hoare_triple_validE by blast
define pairs where "pairs = (λσ. { (p, r) |p r. Some (snd σ) = Some p ⊕ Some r ∧ (fst σ, p), (fst σ, p) ⊨ P
∧ (fst σ, r), (fst σ, r) ⊨ R })"
define Σ' where "Σ' = (λσ. (⋃(p, r) ∈ pairs σ. add_states (Σ (fst σ, p)) {(s'', r) |s''. agrees (- wrC C) (fst σ) s''}))"
show ?thesis
proof (rule hoare_triple_validI_bounded)
show "⋀s h n. (s, h), (s, h) ⊨ Star P R ⟹ bounded h ⟹ safe n Δ C (s, h) (Σ' (s, h))"
proof -
fix s h n assume asm1: "(s, h), (s, h) ⊨ Star P R" "bounded h"
then obtain p r where "Some h = Some p ⊕ Some r" "(s, p), (s, p) ⊨ P" "(s, r), (s, r) ⊨ R"
using always_sat_refl hyper_sat.simps(4) by blast
then have "safe n Δ C (s, p) (Σ (s, p))"
using asm0 asm1
by (metis bounded_smaller_sum snd_conv)
then have "safe n Δ C (s, h) (add_states (Σ (s, p)) {(s'', r) |s''. agrees (- wrC C) s s''})"
using safe_frame[of n Δ C s p "Σ (s, p)" h r] ‹Some h = Some p ⊕ Some r›
using asm1(2) by blast
moreover have "(add_states (Σ (s, p)) {(s'', r) |s''. agrees (- wrC C) s s''}) ⊆ Σ' (s, h)"
proof -
have "(p, r) ∈ pairs (s, h)"
using ‹(s, p), (s, p) ⊨ P› ‹(s, r), (s, r) ⊨ R› ‹Some h = Some p ⊕ Some r› pairs_def by force
then show ?thesis
using Σ'_def by auto
qed
ultimately show "safe n Δ C (s, h) (Σ' (s, h))"
using safe_larger_set by blast
qed
fix s1 h1 s2 h2
assume asm1: "(s1, h1), (s2, h2) ⊨ Star P R"
then obtain p1 p2 r1 r2 where "Some h1 = Some p1 ⊕ Some r1" "Some h2 = Some p2 ⊕ Some r2"
"(s1, p1), (s2, p2) ⊨ P" "(s1, r1), (s2, r2) ⊨ R"
by auto
then have "(s1, p1), (s1, p1) ⊨ P ∧ (s1, r1), (s1, r1) ⊨ R ∧ (s2, p2), (s2, p2) ⊨ P ∧ (s2, r2), (s2, r2) ⊨ R"
using always_sat_refl sat_comm by blast
show "pair_sat (Σ' (s1, h1)) (Σ' (s2, h2)) (Star Q R)"
proof (rule pair_satI)
fix s1' h1' s2' h2'
assume asm2: "(s1', h1') ∈ Σ' (s1, h1) ∧ (s2', h2') ∈ Σ' (s2, h2)"
then obtain p1' r1' p2' r2' where "(p1', r1') ∈ pairs (s1, h1)" "(p2', r2') ∈ pairs (s2, h2)"
"(s1', h1') ∈ add_states (Σ (s1, p1')) {(s'', r1') |s''. agrees (- wrC C) s1 s''}"
"(s2', h2') ∈ add_states (Σ (s2, p2')) {(s'', r2') |s''. agrees (- wrC C) s2 s''}"
using Σ'_def by force
moreover obtain "(s1, p1'), (s1, p1') ⊨ P" "(s1, r1'), (s1, r1') ⊨ R" "(s2, p2'), (s2, p2') ⊨ P" "(s2, r2'), (s2, r2') ⊨ R"
"Some h1 = Some p1' ⊕ Some r1'" "Some h2 = Some p2' ⊕ Some r2'"
using calculation(1) calculation(2) pairs_def by auto
ultimately have "p1 = p1' ∧ p2 = p2' ∧ r1 = r1' ∧ r2 = r2'"
proof (cases "precise P")
case True
then have "p1 = p1' ∧ p2 = p2'" using preciseE
by (metis ‹(s1, p1), (s1, p1) ⊨ P ∧ (s1, r1), (s1, r1) ⊨ R ∧ (s2, p2), (s2, p2) ⊨ P ∧ (s2, r2), (s2, r2) ⊨ R› ‹Some h1 = Some p1 ⊕ Some r1› ‹Some h2 = Some p2 ⊕ Some r2› ‹⋀thesis. (⟦(s1, p1'), (s1, p1') ⊨ P; (s1, r1'), (s1, r1') ⊨ R; (s2, p2'), (s2, p2') ⊨ P; (s2, r2'), (s2, r2') ⊨ R; Some h1 = Some p1' ⊕ Some r1'; Some h2 = Some p2' ⊕ Some r2'⟧ ⟹ thesis) ⟹ thesis› larger_def)
then show ?thesis
by (metis ‹Some h1 = Some p1 ⊕ Some r1› ‹Some h1 = Some p1' ⊕ Some r1'› ‹Some h2 = Some p2 ⊕ Some r2› ‹Some h2 = Some p2' ⊕ Some r2'› addition_cancellative plus_comm)
next
case False
then have "precise R"
using assms(3) by auto
then show ?thesis
by (metis (no_types, opaque_lifting) ‹(s1, p1), (s1, p1) ⊨ P ∧ (s1, r1), (s1, r1) ⊨ R ∧ (s2, p2), (s2, p2) ⊨ P ∧ (s2, r2), (s2, r2) ⊨ R› ‹Some h1 = Some p1 ⊕ Some r1› ‹Some h2 = Some p2 ⊕ Some r2› ‹⋀thesis. (⟦(s1, p1'), (s1, p1') ⊨ P; (s1, r1'), (s1, r1') ⊨ R; (s2, p2'), (s2, p2') ⊨ P; (s2, r2'), (s2, r2') ⊨ R; Some h1 = Some p1' ⊕ Some r1'; Some h2 = Some p2' ⊕ Some r2'⟧ ⟹ thesis) ⟹ thesis› addition_cancellative larger_def plus_comm preciseE)
qed
then have "pair_sat (Σ (s1, p1')) (Σ (s2, p2')) Q"
using ‹(s1, p1), (s2, p2) ⊨ P› asm0(2) by blast
moreover have "pair_sat {(s'', r1') |s''. agrees (- wrC C) s1 s''} {(s'', r2') |s''. agrees (- wrC C) s2 s''} R"
(is "pair_sat ?R1 ?R2 R")
proof (rule pair_satI)
fix s1'' r1'' s2'' r2'' assume "(s1'', r1'') ∈ {(s'', r1') |s''. agrees (- wrC C) s1 s''} ∧ (s2'', r2'') ∈ {(s'', r2') |s''. agrees (- wrC C) s2 s''}"
then obtain "r1'' = r1'" "r2'' = r2'" "agrees (- wrC C) s1 s1''" "agrees (- wrC C) s2 s2''"
by fastforce
then show "(s1'', r1''), (s2'', r2'') ⊨ R"
using ‹(s1, r1), (s2, r2) ⊨ R› ‹p1 = p1' ∧ p2 = p2' ∧ r1 = r1' ∧ r2 = r2'› agrees_minusD agrees_same
assms(2) sat_comm
by (metis (no_types, opaque_lifting) disjoint_def inf_commute)
qed
ultimately have "pair_sat (add_states (Σ (s1, p1')) ?R1) (add_states (Σ (s2, p2')) ?R2) (Star Q R)"
using add_states_sat_star by blast
then show "(s1', h1'), (s2', h2') ⊨ Star Q R"
using ‹(s1', h1') ∈ add_states (Σ (s1, p1')) {(s'', r1') |s''. agrees (- wrC C) s1 s''}› ‹(s2', h2') ∈ add_states (Σ (s2, p2')) {(s'', r2') |s''. agrees (- wrC C) s2 s''}› pair_sat_def by blast
qed
qed
qed
subsubsection ‹Consequence›
theorem consequence_rule:
fixes Δ :: "('i, 'a, nat) cont"
assumes "hoare_triple_valid Δ P' C Q'"
and "entails P P'"
and "entails Q' Q"
shows "hoare_triple_valid Δ P C Q"
proof -
obtain Σ where asm0: "⋀σ n. σ, σ ⊨ P' ⟹ bounded (snd σ) ⟹ safe n Δ C σ (Σ σ)" "⋀σ σ'. σ, σ' ⊨ P' ⟹ pair_sat (Σ σ) (Σ σ') Q'"
using assms(1) hoare_triple_validE by blast
show ?thesis
proof (rule hoare_triple_validI_bounded)
show "⋀s h n. (s, h), (s, h) ⊨ P ⟹ bounded h ⟹ safe n Δ C (s, h) (Σ (s, h))"
using asm0(1) assms(2) entails_def
by fastforce
show "⋀s h s' h'. (s, h), (s', h') ⊨ P ⟹ pair_sat (Σ (s, h)) (Σ (s', h')) Q"
by (meson asm0(2) assms(2) assms(3) entails_def pair_sat_def)
qed
qed
subsubsection ‹Existential›
theorem existential_rule:
fixes Δ :: "('i, 'a, nat) cont"
assumes "hoare_triple_valid Δ P C Q"
and "x ∉ fvC C"
and "⋀Γ. Δ = Some Γ ⟹ x ∉ fvA (invariant Γ)"
and "unambiguous P x"
shows "hoare_triple_valid Δ (Exists x P) C (Exists x Q)"
proof -
obtain Σ where asm0: "⋀σ n. σ, σ ⊨ P ⟹ bounded (snd σ) ⟹ safe n Δ C σ (Σ σ)" "⋀σ σ'. σ, σ' ⊨ P ⟹ pair_sat (Σ σ) (Σ σ') Q"
using assms(1) hoare_triple_validE by blast
define Σ' where "Σ' = (λσ. ⋃v ∈ { v |v. ((fst σ)(x := v), snd σ), ((fst σ)(x := v), snd σ) ⊨ P }. upperize (Σ ((fst σ)(x := v), snd σ)) (fvA Q - {x}))"
show ?thesis
proof (rule hoare_triple_validI_bounded)
show "⋀s h n. (s, h), (s, h) ⊨ Exists x P ⟹ bounded h ⟹ safe n Δ C (s, h) (Σ' (s, h))"
proof -
fix s h n assume "(s, h), (s, h) ⊨ Exists x P" "bounded h"
then obtain v where "(s(x := v), h), (s(x := v), h) ⊨ P"
using always_sat_refl hyper_sat.simps(7) by blast
then have "Σ (s(x := v), h) ⊆ Σ' (s, h)"
using upperize_larger SUP_upper2 Σ'_def by fastforce
moreover have "safe n Δ C (s(x := v), h) (Σ (s(x := v), h))"
by (simp add: ‹(s(x := v), h), (s(x := v), h) ⊨ P› ‹bounded h› asm0(1))
ultimately have "safe n Δ C (s(x := v), h) (Σ' (s, h))"
using safe_larger_set by blast
then have "safe n Δ C (s, h) (Σ' (s, h))"
proof (rule safe_free_vars)
show "⋀Γ. Δ = Some Γ ⟹ agrees (fvA (invariant Γ)) (s(x := v)) s"
by (meson agrees_comm agrees_update assms(3))
show "agrees (fvC C ∪ (fvA Q - {x})) (s(x := v)) s"
by (simp add: agrees_def assms(2))
show "upper_fvs (Σ' (s, h)) (fvA Q - {x})"
proof (rule upper_fvsI)
fix sa s' ha
assume asm0: "(sa, ha) ∈ Σ' (s, h) ∧ agrees (fvA Q - {x}) sa s'"
then obtain v where "(s(x := v), h), (s(x := v), h) ⊨ P" "(sa, ha) ∈ upperize (Σ (s(x := v), h)) (fvA Q - {x})"
using Σ'_def by force
then have "(s', ha) ∈ upperize (Σ (s(x := v), h)) (fvA Q - {x})"
using asm0 upper_fvs_def upper_fvs_upperize by blast
then show "(s', ha) ∈ Σ' (s, h)"
using ‹(s(x := v), h), (s(x := v), h) ⊨ P› Σ'_def by force
qed
qed
then show "safe n Δ C (s, h) (Σ' (s, h))"
by auto
qed
fix s1 h1 s2 h2
assume asm1: "(s1, h1), (s2, h2) ⊨ Exists x P"
then obtain v1' v2' where "(s1(x := v1'), h1), (s2(x := v2'), h2) ⊨ P" by auto
show "pair_sat (Σ' (s1, h1)) (Σ' (s2, h2)) (Exists x Q)"
proof (rule pair_satI)
fix s1' h1' s2' h2'
assume asm2: "(s1', h1') ∈ Σ' (s1, h1) ∧ (s2', h2') ∈ Σ' (s2, h2)"
then obtain v1 v2 where
r: "(s1(x := v1), h1), (s1(x := v1), h1) ⊨ P" "(s1', h1') ∈ upperize (Σ (s1(x := v1), h1)) (fvA Q - {x})"
"(s2(x := v2), h2), (s2(x := v2), h2) ⊨ P" "(s2', h2') ∈ upperize (Σ (s2(x := v2), h2)) (fvA Q - {x})"
using Σ'_def by auto
then obtain s1'' s2'' where "agrees (fvA Q - {x}) s1'' s1'" "(s1'', h1') ∈ Σ (s1(x := v1), h1)"
"agrees (fvA Q - {x}) s2'' s2'" "(s2'', h2') ∈ Σ (s2(x := v2), h2)"
using in_upperize by (metis (no_types, lifting))
moreover have "(s1(x := v1), h1), (s2(x := v2), h2) ⊨ P"
proof -
have "v1 = v1'"
using ‹(s1(x := v1'), h1), (s2(x := v2'), h2) ⊨ P› always_sat_refl assms(4) r(1) unambiguous_def by blast
moreover have "v2 = v2'"
using ‹(s1(x := v1'), h1), (s2(x := v2'), h2) ⊨ P› always_sat_refl assms(4) r(3) sat_comm_aux unambiguous_def by blast
ultimately show ?thesis
by (simp add: ‹(s1(x := v1'), h1), (s2(x := v2'), h2) ⊨ P›)
qed
then have "pair_sat (Σ (s1(x := v1), h1)) (Σ (s2(x := v2), h2)) Q"
using asm0 by simp
then have "(s1'', h1'), (s2'', h2') ⊨ Q"
using calculation(2) calculation(4) pair_sat_def by blast
moreover have "agrees (fvA Q) s1'' (s1'(x := s1'' x))"
proof (rule agreesI)
fix y assume "y ∈ fvA Q"
then show "s1'' y = (s1'(x := s1'' x)) y"
apply (cases "x = y")
apply auto[1]
by (metis (mono_tags, lifting) DiffI agrees_def calculation(1) fun_upd_other singleton_iff)
qed
moreover have "agrees (fvA Q) s2'' (s2'(x := s2'' x))"
proof (rule agreesI)
fix y assume "y ∈ fvA Q"
then show "s2'' y = (s2'(x := s2'' x)) y"
apply (cases "x = y")
apply auto[1]
by (metis (mono_tags, lifting) DiffI agrees_def calculation(3) fun_upd_other singleton_iff)
qed
ultimately have "(s1'(x := s1'' x), h1'), (s2'(x := s2'' x), h2') ⊨ Q"
by (meson agrees_same sat_comm)
then show "(s1', h1'), (s2', h2') ⊨ Exists x Q"
using hyper_sat.simps(7) by blast
qed
qed
qed
subsubsection ‹While loops›
inductive leads_to_loop where
"leads_to_loop b I Σ σ σ"
| "⟦ leads_to_loop b I Σ σ σ' ; bdenot b (fst σ') ; σ'' ∈ Σ σ' ⟧ ⟹ leads_to_loop b I Σ σ σ''"
definition leads_to_loop_set where
"leads_to_loop_set b I Σ σ = { σ' |σ'. leads_to_loop b I Σ σ σ'}"
definition trans_Σ where
"trans_Σ b I Σ σ = Set.filter (λσ. ¬ bdenot b (fst σ)) (leads_to_loop_set b I Σ σ)"
inductive_cases red_while_cases: "red (Cwhile b s) σ C' σ'"
inductive_cases abort_while_cases: "aborts (Cwhile b s) σ"
lemma safe_while_None:
assumes "⋀σ m. σ, σ ⊨ And I (Bool b) ⟹ bounded (snd σ) ⟹ safe n (None :: ('i, 'a, nat) cont) C σ (Σ σ)"
and "⋀σ σ'. σ, σ' ⊨ And I (Bool b) ⟹ pair_sat (Σ σ) (Σ σ') I"
and "(s, h), (s, h) ⊨ I"
and "leads_to_loop b I Σ σ (s, h)"
and "bounded h"
shows "safe n (None :: ('i, 'a, nat) cont) (Cwhile b C) (s, h) (trans_Σ b I Σ σ)"
using assms
proof (induct n arbitrary: s h)
let ?S = "trans_Σ b I Σ σ"
case (Suc n)
show ?case
proof (rule safeNoneI)
show "no_abort (None :: ('i, 'a, nat) cont) (Cwhile b C) s h"
using abort_while_cases no_abortNoneI by blast
fix H hf C' s' h'
assume asm0: "Some H = Some h ⊕ Some hf ∧ full_ownership (get_fh H) ∧ no_guard H ∧ red (Cwhile b C) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
show "∃h'' H'. full_ownership (get_fh H') ∧ no_guard H' ∧ h' = FractionalHeap.normalize (get_fh H')
∧ Some H' = Some h'' ⊕ Some hf ∧ safe n (None :: ('i, 'a, nat) cont) C' (s', h'') (trans_Σ b I Σ σ)"
proof (rule red_while_cases)
show "red (Cwhile b C) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
using asm0 by linarith
assume asm1: "C' = Cif b (Cseq C (Cwhile b C)) Cskip" "(s', h') = (s, FractionalHeap.normalize (get_fh H))"
have "safe n (None :: ('i, 'a, nat) cont) C' (s, h) ?S"
proof (cases n)
case (Suc k)
have "safe (Suc k) (None :: ('i, 'a, nat) cont) (Cif b (Cseq C (Cwhile b C)) Cskip) (s, h) ?S"
proof (rule if_safe)
have "¬ bdenot b s ⟹ (s, h) ∈ ?S"
by (metis CollectI Suc.prems(4) asm1(2) fst_eqD leads_to_loop_set_def member_filter trans_Σ_def)
then show "¬ bdenot b s ⟹ safe k (None :: ('i, 'a, nat) cont) Cskip (s, h) (trans_Σ b I Σ σ)"
by (metis Pair_inject asm1(2) safe_skip)
assume asm2: "bdenot b s"
then have "(s, h), (s, h) ⊨ And I (Bool b)"
by (simp add: Suc.prems(3))
then have r: "safe (Suc n) (None :: ('i, 'a, nat) cont) C (s, h) (Σ (s, h))"
using Suc.prems(1)
by (metis Suc.prems(5) snd_eqD)
show "safe k (None :: ('i, 'a, nat) cont) (Cseq C (Cwhile b C)) (s, h) (trans_Σ b I Σ σ)"
proof (rule seq_safe)
show "safe k (None :: ('i, 'a, nat) cont) C (s, h) (Set.filter (bounded ∘ snd) (Σ (s, h)))"
apply (rule restrict_safe_to_bounded)
using Suc Suc_n_not_le_n nat_le_linear r safe_smaller apply metis
by (simp add: Suc.prems(5))
fix m s' h' assume asm3: "m ≤ k ∧ (s', h') ∈ Set.filter (bounded ∘ snd) (Σ (s, h))"
have "safe n (None :: ('i, 'a, nat) cont) (Cwhile b C) (s', h') (trans_Σ b I Σ σ)"
proof (rule Suc.hyps)
show "leads_to_loop b I Σ σ (s', h')"
by (metis Suc.prems(4) asm2 asm3 fst_conv leads_to_loop.simps member_filter)
show "(s', h'), (s', h') ⊨ I"
using ‹(s, h), (s, h) ⊨ And I (Bool b)› asm3 assms(2) pair_satE
by (metis member_filter)
show "⋀σ. σ, σ ⊨ And I (Bool b) ⟹ bounded (snd σ) ⟹ safe n (None :: ('i, 'a, nat) cont) C σ (Σ σ)"
by (meson Suc.prems(1) Suc_n_not_le_n nat_le_linear safe_smaller)
show "bounded h'"
using asm3 by auto
qed (auto simp add: assms)
then show "safe m (None :: ('i, 'a, nat) cont) (Cwhile b C) (s', h') (trans_Σ b I Σ σ)"
using Suc asm3 le_SucI safe_smaller by blast
qed
qed
then show ?thesis
using Suc asm1(1) by blast
qed (simp)
then show "∃h'' H'. full_ownership (get_fh H') ∧
no_guard H' ∧ h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hf ∧ safe n None C' (s', h'') (trans_Σ b I Σ σ)"
using asm0 asm1(2) by blast
qed
qed (simp_all)
qed (simp)
lemma safe_while_Some:
assumes "⋀σ m. σ, σ ⊨ And I (Bool b) ⟹ bounded (snd σ) ⟹ safe n (Some Γ) C σ (Σ σ)"
and "⋀σ σ'. σ, σ' ⊨ And I (Bool b) ⟹ pair_sat (Σ σ) (Σ σ') I"
and "(s, h), (s, h) ⊨ I"
and "leads_to_loop b I Σ σ (s, h)"
shows "safe n (Some Γ) (Cwhile b C) (s, h) (trans_Σ b I Σ σ)"
using assms
proof (induct n arbitrary: s h)
let ?S = "trans_Σ b I Σ σ"
case (Suc n)
show ?case
proof (rule safeSomeI)
show "no_abort (Some Γ) (Cwhile b C) s h"
using abort_while_cases no_abortSomeI by blast
fix H hf C' s' h' hj v0
assume asm0: "Some H = Some h ⊕ Some hj ⊕ Some hf ∧
full_ownership (get_fh H) ∧ semi_consistent Γ v0 H ∧ sat_inv s hj Γ ∧ red (Cwhile b C) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
show "∃h'' H' hj'. full_ownership (get_fh H') ∧ semi_consistent Γ v0 H' ∧ sat_inv s' hj' Γ ∧
h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hj' ⊕ Some hf ∧ safe n (Some Γ) C' (s', h'') (trans_Σ b I Σ σ)"
proof (rule red_while_cases)
show "red (Cwhile b C) (s, FractionalHeap.normalize (get_fh H)) C' (s', h')"
using asm0 by linarith
assume asm1: "C' = Cif b (Cseq C (Cwhile b C)) Cskip" "(s', h') = (s, FractionalHeap.normalize (get_fh H))"
have "safe n (Some Γ) C' (s, h) ?S"
proof (cases n)
case (Suc k)
have "safe (Suc k) (Some Γ) (Cif b (Cseq C (Cwhile b C)) Cskip) (s, h) ?S"
proof (rule if_safe)
have "¬ bdenot b s ⟹ (s, h) ∈ ?S"
by (metis CollectI Suc.prems(4) asm1(2) fst_eqD leads_to_loop_set_def member_filter trans_Σ_def)
then show "¬ bdenot b s ⟹ safe k (Some Γ) Cskip (s, h) (trans_Σ b I Σ σ)"
by (metis Pair_inject asm1(2) safe_skip)
assume asm2: "bdenot b s"
then have "(s, h), (s, h) ⊨ And I (Bool b)"
by (simp add: Suc.prems(3))
moreover have "bounded h"
apply (rule bounded_smaller[of H])
using asm0 full_ownership_then_bounded apply fastforce
using asm0 plus_asso[of "Some h" "Some hj" "Some hf"] unfolding larger_def by auto
ultimately have r: "safe (Suc n) (Some Γ) C (s, h) (Σ (s, h))"
using Suc.prems(1)[of "(s, h)"] by fastforce
show "safe k (Some Γ) (Cseq C (Cwhile b C)) (s, h) (trans_Σ b I Σ σ)"
proof (rule seq_safe)
show "safe k (Some Γ) C (s, h) (Σ (s, h))"
by (metis Suc Suc_n_not_le_n nat_le_linear r safe_smaller)
fix m s' h' assume asm3: "m ≤ k ∧ (s', h') ∈ Σ (s, h)"
have "safe n (Some Γ) (Cwhile b C) (s', h') (trans_Σ b I Σ σ)"
proof (rule Suc.hyps)
show "leads_to_loop b I Σ σ (s', h')"
by (metis Suc.prems(4) asm2 asm3 fst_conv leads_to_loop.intros(2))
show "(s', h'), (s', h') ⊨ I"
using ‹(s, h), (s, h) ⊨ And I (Bool b)› asm3 assms(2) pair_satE by blast
show "⋀σ. σ, σ ⊨ And I (Bool b) ⟹ bounded (snd σ) ⟹ safe n (Some Γ) C σ (Σ σ)"
by (meson Suc.prems(1) Suc_n_not_le_n nat_le_linear safe_smaller)
qed (auto simp add: assms)
then show "safe m (Some Γ) (Cwhile b C) (s', h') (trans_Σ b I Σ σ)"
using Suc asm3 le_SucI safe_smaller by blast
qed
qed
then show ?thesis
using Suc asm1(1) by blast
qed (simp)
then show "∃h'' H' hj'.
full_ownership (get_fh H') ∧
semi_consistent Γ v0 H' ∧
sat_inv s' hj' Γ ∧ h' = FractionalHeap.normalize (get_fh H') ∧ Some H' = Some h'' ⊕ Some hj' ⊕ Some hf ∧ safe n (Some Γ) C' (s', h'') (trans_Σ b I Σ σ)"
using asm0 asm1(2) by blast
qed
qed (simp_all)
qed (simp)
lemma safe_while:
fixes Δ :: "('i, 'a, nat) cont"
assumes "⋀σ m. σ, σ ⊨ And I (Bool b) ⟹ bounded (snd σ) ⟹ safe n Δ C σ (Σ σ)"
and "⋀σ σ'. σ, σ' ⊨ And I (Bool b) ⟹ pair_sat (Σ σ) (Σ σ') I"
and "(s, h), (s, h) ⊨ I"
and "leads_to_loop b I Σ σ (s, h)"
and "bounded h"
shows "safe n Δ (Cwhile b C) (s, h) (trans_Σ b I Σ σ)"
apply (cases Δ)
using assms safe_while_None apply blast
using assms safe_while_Some by blast
lemma leads_to_sat_inv_unary:
assumes "leads_to_loop b I Σ σ σ'"
and "⋀σ σ'. σ, σ' ⊨ (And I (Bool b)) ⟹ pair_sat (Σ σ) (Σ σ') I"
and "σ, σ ⊨ I"
shows "σ', σ' ⊨ I"
using assms
proof (induct arbitrary: rule: leads_to_loop.induct)
case (2 b I Σ σ0 σ1 σ2)
then have "pair_sat (Σ σ1) (Σ σ1) I"
by (metis hyper_sat.simps(1) hyper_sat.simps(3) prod.collapse)
then show ?case
using "2.hyps"(4) pair_sat_def by blast
qed (simp)
theorem while_rule2:
fixes Δ :: "('i, 'a, nat) cont"
assumes "unary I"
and "hoare_triple_valid Δ (And I (Bool b)) C I"
shows "hoare_triple_valid Δ I (Cwhile b C) (And I (Bool (Bnot b)))"
proof -
obtain Σ where asm0: "⋀σ n. σ, σ ⊨ (And I (Bool b)) ⟹ bounded (snd σ) ⟹ safe n Δ C σ (Σ σ)"
and "⋀σ σ'. σ, σ' ⊨ (And I (Bool b)) ⟹ pair_sat (Σ σ) (Σ σ') I"
using assms(2) hoare_triple_validE by blast
let ?Σ = "trans_Σ b I Σ"
show ?thesis
proof (rule hoare_triple_validI_bounded)
show "⋀s h s' h'. (s, h), (s', h') ⊨ I ⟹ pair_sat (?Σ (s, h)) (?Σ (s', h')) (And I (Bool (Bnot b)))"
proof -
fix s1 h1 s2 h2 assume asm0: "(s1, h1), (s2, h2) ⊨ I"
show "pair_sat (trans_Σ b I Σ (s1, h1)) (trans_Σ b I Σ (s2, h2)) (And I (Bool (Bnot b)))"
proof (rule pair_satI)
fix s1' h1' s2' h2'
assume asm1: "(s1', h1') ∈ trans_Σ b I Σ (s1, h1) ∧ (s2', h2') ∈ trans_Σ b I Σ (s2, h2)"
then obtain "leads_to_loop b I Σ (s1, h1) (s1', h1')" "¬ bdenot b s1'"
"leads_to_loop b I Σ (s2, h2) (s2', h2')" "¬ bdenot b s2'"
using trans_Σ_def leads_to_loop_set_def
by (metis fst_conv mem_Collect_eq member_filter)
then have "(s1', h1'), (s1', h1') ⊨ I ∧ (s2', h2'), (s2', h2') ⊨ I"
by (meson ‹⋀σ' σ. σ, σ' ⊨ And I (Bool b) ⟹ pair_sat (Σ σ) (Σ σ') I› always_sat_refl asm0 leads_to_sat_inv_unary sat_comm_aux)
then show "(s1', h1'), (s2', h2') ⊨ And I (Bool (Bnot b))"
by (metis ‹¬ bdenot b s1'› ‹¬ bdenot b s2'› assms(1) bdenot.simps(3) hyper_sat.simps(1) hyper_sat.simps(3) unaryE)
qed
qed
fix s h n
assume asm1: "(s, h), (s, h) ⊨ I" "bounded h"
show "safe n Δ (Cwhile b C) (s, h) (trans_Σ b I Σ (s, h))"
proof (rule safe_while)
show "⋀σ σ'. σ, σ' ⊨ And I (Bool b) ⟹ pair_sat (Σ σ) (Σ σ') I"
by (simp add: ‹⋀σ' σ. σ, σ' ⊨ And I (Bool b) ⟹ pair_sat (Σ σ) (Σ σ') I›)
show "(s, h), (s, h) ⊨ I"
using asm1 by auto
show "leads_to_loop b I Σ (s, h) (s, h)"
by (simp add: leads_to_loop.intros(1))
show "⋀σ m. σ, σ ⊨ And I (Bool b) ⟹ bounded (snd σ) ⟹ safe n Δ C σ (Σ σ)"
by (simp add: asm0 asm1(2))
show "bounded h"
using asm1(2) by blast
qed
qed
qed
fun iterate_sigma :: "nat ⇒ bexp ⇒ ('i, 'a, nat) assertion ⇒ ((store × ('i, 'a) heap) ⇒ (store × ('i, 'a) heap) set) ⇒ (store × ('i, 'a) heap) ⇒ (store × ('i, 'a) heap) set"
where
"iterate_sigma 0 b I Σ σ = {σ}"
| "iterate_sigma (Suc n) b I Σ σ = (⋃σ' ∈ Set.filter (λσ. bdenot b (fst σ)) (iterate_sigma n b I Σ σ). Σ σ')"
lemma union_of_iterate_sigma_is_leads_to_loop_set:
assumes "leads_to_loop b I Σ σ σ'"
shows "σ' ∈ (⋃n. iterate_sigma n b I Σ σ)"
using assms
proof (induct rule: leads_to_loop.induct)
case (1 b I Σ σ)
have "σ ∈ iterate_sigma 0 b I Σ σ"
by simp
then show ?case
by blast
next
case (2 b I Σ σ σ' σ'')
then obtain n where "σ' ∈ iterate_sigma n b I Σ σ" by blast
then have "σ'' ∈ iterate_sigma (Suc n) b I Σ σ" using 2 by auto
then show ?case by blast
qed
lemma trans_included:
"trans_Σ b I Σ σ ⊆ Set.filter (λσ. ¬ bdenot b (fst σ)) (⋃n. iterate_sigma n b I Σ σ)"
proof
fix x assume "x ∈ trans_Σ b I Σ σ"
then have "¬ bdenot b (fst x) ∧ leads_to_loop b I Σ σ x"
by (simp add: leads_to_loop_set_def trans_Σ_def)
then show "x ∈ Set.filter (λσ. ¬ bdenot b (fst σ)) (⋃n. iterate_sigma n b I Σ σ)"
by (metis member_filter union_of_iterate_sigma_is_leads_to_loop_set)
qed
lemma iterate_sigma_low_all_sat_I_and_low:
assumes "⋀σ σ'. σ, σ' ⊨ (And I (Bool b)) ⟹ pair_sat (Σ σ) (Σ σ') (And I (Low b))"
and "σ1, σ2 ⊨ I"
and "bdenot b (fst σ1) = bdenot b (fst σ2)"
shows "pair_sat (iterate_sigma n b I Σ σ1) (iterate_sigma n b I Σ σ2) (And I (Low b))"
using assms
proof (induct n)
case 0
then show ?case
by (metis (mono_tags, lifting) hyper_sat.simps(3) hyper_sat.simps(5) iterate_sigma.simps(1) pair_satI prod.exhaust_sel singletonD)
next
case (Suc n)
show ?case
proof (rule pair_satI)
fix s1 h1 s2 h2
assume asm0: "(s1, h1) ∈ iterate_sigma (Suc n) b I Σ σ1 ∧ (s2, h2) ∈ iterate_sigma (Suc n) b I Σ σ2"
then obtain σ1' σ2' where "bdenot b (fst σ1')" "bdenot b (fst σ2')"
"σ1' ∈ iterate_sigma n b I Σ σ1" "σ2' ∈ iterate_sigma n b I Σ σ2"
"(s1, h1) ∈ Σ σ1'" "(s2, h2) ∈ Σ σ2'"
by auto
then have "pair_sat (iterate_sigma n b I Σ σ1) (iterate_sigma n b I Σ σ2) (And I (Low b))"
using Suc.hyps
using Suc.prems(3) assms(1) assms(2) by blast
moreover have "pair_sat (Σ σ1') (Σ σ2') (And I (Low b))"
proof (rule Suc.prems)
show "σ1', σ2' ⊨ And I (Bool b)"
by (metis ‹σ1' ∈ iterate_sigma n b I Σ σ1› ‹σ2' ∈ iterate_sigma n b I Σ σ2› ‹bdenot b (fst σ1')› ‹bdenot b (fst σ2')› calculation hyper_sat.simps(1) hyper_sat.simps(3) pair_sat_def prod.exhaust_sel)
qed
ultimately show "(s1, h1), (s2, h2) ⊨ And I (Low b)"
using ‹(s1, h1) ∈ Σ σ1'› ‹(s2, h2) ∈ Σ σ2'› pair_sat_def by blast
qed
qed
lemma iterate_empty_later_empty:
assumes "iterate_sigma n b I Σ σ = {}"
and "m ≥ n"
shows "iterate_sigma m b I Σ σ = {}"
using assms
proof (induct "m - n" arbitrary: n m)
case (Suc k)
then obtain mm where "m = Suc mm"
by (metis iterate_sigma.elims zero_diff)
then have "iterate_sigma mm b I Σ σ = {}"
by (metis Suc.hyps(1) Suc.hyps(2) Suc.prems(1) Suc.prems(2) Suc_le_mono diff_Suc_Suc diff_diff_cancel diff_le_self)
then show ?case
using ‹m = Suc mm› by force
qed (simp)
lemma all_same:
assumes "⋀σ σ'. σ, σ' ⊨ (And I (Bool b)) ⟹ pair_sat (Σ σ) (Σ σ') (And I (Low b))"
and "σ1, σ2 ⊨ I"
and "bdenot b (fst σ1) = bdenot b (fst σ2)"
and "x1 ∈ iterate_sigma n b I Σ σ1"
and "x2 ∈ iterate_sigma n b I Σ σ2"
shows "bdenot b (fst x1) = bdenot b (fst x2)"
proof -
have "x1, x2 ⊨ (And I (Low b))"
using assms(1) assms(2) assms(3) assms(4) assms(5) iterate_sigma_low_all_sat_I_and_low pair_sat_def by blast
then show ?thesis
by (metis (no_types, lifting) hyper_sat.simps(3) hyper_sat.simps(5) surjective_pairing)
qed
lemma non_empty_at_most_once:
assumes "⋀σ σ'. σ, σ' ⊨ (And I (Bool b)) ⟹ pair_sat (Σ σ) (Σ σ') (And I (Low b))"
and "σ, σ ⊨ I"
and "Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma n1 b I Σ σ) ≠ {}"
and "Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma n2 b I Σ σ) ≠ {}"
shows "n1 = n2"
proof -
let ?n = "min n1 n2"
obtain σ' where "σ' ∈ Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma ?n b I Σ σ)"
by (metis assms(3) assms(4) equals0I min.orderE min_def)
then have "¬ bdenot b (fst σ')"
by fastforce
moreover have "pair_sat (iterate_sigma ?n b I Σ σ) (iterate_sigma ?n b I Σ σ) (And I (Low b))"
using assms(1) assms(2) assms(3) iterate_sigma_low_all_sat_I_and_low by blast
then have r: "⋀x. x ∈ iterate_sigma ?n b I Σ σ ⟹ ¬ bdenot b (fst x)"
by (metis ‹σ' ∈ Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma (min n1 n2) b I Σ σ)› all_same assms(1) assms(2) member_filter)
then have "iterate_sigma (Suc ?n) b I Σ σ = {}" by auto
then have "¬ (n1 > ?n) ∧ ¬ (n2 > ?n)" using iterate_empty_later_empty[of "Suc ?n" b I Σ σ]
assms by (metis (no_types, lifting) Set.filter_def empty_Collect_eq empty_def le_simps(3) mem_Collect_eq)
then show ?thesis by linarith
qed
lemma one_non_empty_union:
assumes "⋀σ σ'. σ, σ' ⊨ (And I (Bool b)) ⟹ pair_sat (Σ σ) (Σ σ') (And I (Low b))"
and "σ, σ ⊨ I"
and "Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma k b I Σ σ) ≠ {}"
shows "Set.filter (λσ. ¬ bdenot b (fst σ)) (⋃n. iterate_sigma n b I Σ σ) = Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma k b I Σ σ)"
proof
show "Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma k b I Σ σ) ⊆ Set.filter (λσ. ¬ bdenot b (fst σ)) (⋃n. iterate_sigma n b I Σ σ)"
by auto
show "Set.filter (λσ. ¬ bdenot b (fst σ)) (⋃n. iterate_sigma n b I Σ σ) ⊆ Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma k b I Σ σ)"
proof
fix x assume "x ∈ Set.filter (λσ. ¬ bdenot b (fst σ)) (⋃n. iterate_sigma n b I Σ σ)"
then obtain k' where "x ∈ iterate_sigma k' b I Σ σ" "¬ bdenot b (fst x)"
by auto
then have "x ∈ Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma k' b I Σ σ)"
by fastforce
then have "k = k'"
using non_empty_at_most_once assms(1) assms(2) assms(3) by blast
then show "x ∈ Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma k b I Σ σ)"
using ‹x ∈ Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma k' b I Σ σ)› by blast
qed
qed
definition not_set where
"not_set b S = Set.filter (λσ. ¬ bdenot b (fst σ)) S"
lemma union_exists_at_some_point_exactly:
assumes "⋀σ σ'. σ, σ' ⊨ (And I (Bool b)) ⟹ pair_sat (Σ σ) (Σ σ') (And I (Low b))"
and "σ1, σ2 ⊨ I"
and "bdenot b (fst σ1) = bdenot b (fst σ2)"
and "Set.filter (λσ. ¬ bdenot b (fst σ)) (⋃n. iterate_sigma n b I Σ σ1) ≠ {}"
and "Set.filter (λσ. ¬ bdenot b (fst σ)) (⋃n. iterate_sigma n b I Σ σ2) ≠ {}"
shows "∃k. not_set b (⋃n. iterate_sigma n b I Σ σ1) = not_set b (iterate_sigma k b I Σ σ1) ∧ not_set b (⋃n. iterate_sigma n b I Σ σ2) = not_set b (iterate_sigma k b I Σ σ2)"
proof -
obtain k1 where "Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma k1 b I Σ σ1) ≠ {}"
using assms(4) by fastforce
moreover obtain k2 where "Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma k2 b I Σ σ2) ≠ {}"
using assms(5) by fastforce
show ?thesis
proof (cases "k1 ≤ k2")
case True
then have "iterate_sigma k1 b I Σ σ2 ≠ {}"
by (metis (no_types, lifting) Collect_cong Set.filter_def ‹Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma k2 b I Σ σ2) ≠ {}› empty_def iterate_empty_later_empty mem_Collect_eq)
then obtain σ1' σ2' where "σ1' ∈ Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma k1 b I Σ σ1) ∧ σ2' ∈ iterate_sigma k1 b I Σ σ2"
using calculation by blast
then have "¬ bdenot b (fst σ1')"
by fastforce
moreover have "pair_sat (iterate_sigma k1 b I Σ σ1) (iterate_sigma k1 b I Σ σ2) (And I (Low b))"
using assms(1) assms(2) assms(3) iterate_sigma_low_all_sat_I_and_low by blast
then have r: "⋀x1 x2. x1 ∈ iterate_sigma k1 b I Σ σ1 ∧ x2 ∈ iterate_sigma k1 b I Σ σ2 ⟹ bdenot b (fst x1) ⟷ bdenot b (fst x2)"
by (metis (no_types, opaque_lifting) eq_fst_iff hyper_sat.simps(3) hyper_sat.simps(5) pair_sat_def)
then have "¬ bdenot b (fst σ2')"
by (metis ‹σ1' ∈ Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma k1 b I Σ σ1) ∧ σ2' ∈ iterate_sigma k1 b I Σ σ2› member_filter)
then have "⋀x1. x1 ∈ iterate_sigma k1 b I Σ σ1 ⟹ ¬ bdenot b (fst x1)"
using ‹σ1' ∈ Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma k1 b I Σ σ1) ∧ σ2' ∈ iterate_sigma k1 b I Σ σ2› r by blast
then have "iterate_sigma (Suc k1) b I Σ σ1 = {}" by auto
moreover have "⋀x2. x2 ∈ iterate_sigma k1 b I Σ σ2 ⟹ ¬ bdenot b (fst x2)"
by (metis ‹σ1' ∈ Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma k1 b I Σ σ1) ∧ σ2' ∈ iterate_sigma k1 b I Σ σ2› member_filter r)
then have "iterate_sigma (Suc k1) b I Σ σ2 = {}" by auto
then have "k1 = k2"
using True ‹Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma k2 b I Σ σ2) ≠ {}› dual_order.antisym[of k1 k2]
ex_in_conv iterate_empty_later_empty[of _ b I Σ σ2] member_filter not_less_eq_eq
by metis
moreover have "Set.filter (λσ. ¬ bdenot b (fst σ)) (⋃n. iterate_sigma n b I Σ σ1) = Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma k1 b I Σ σ1)"
using one_non_empty_union[of I b Σ σ1]
using ‹Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma k1 b I Σ σ1) ≠ {}› always_sat_refl assms(1) assms(2) by blast
moreover have "Set.filter (λσ. ¬ bdenot b (fst σ)) (⋃n. iterate_sigma n b I Σ σ2) = Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma k1 b I Σ σ2)"
using one_non_empty_union[of I b Σ σ2]
using ‹Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma k2 b I Σ σ2) ≠ {}› always_sat_refl assms(1) assms(2) calculation(3) sat_comm by blast
ultimately show ?thesis
by (metis not_set_def)
next
case False
then have "iterate_sigma k2 b I Σ σ1 ≠ {}"
by (metis (no_types, lifting) Collect_cong Set.filter_def calculation empty_def iterate_empty_later_empty linorder_le_cases mem_Collect_eq)
then obtain σ1' σ2' where "σ1' ∈ iterate_sigma k2 b I Σ σ1 ∧ σ2' ∈ not_set b (iterate_sigma k2 b I Σ σ2)"
by (metis ‹Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma k2 b I Σ σ2) ≠ {}› ex_in_conv not_set_def)
then have "¬ bdenot b (fst σ2')"
using not_set_def by fastforce
then have "¬ bdenot b (fst σ1')"
by (metis ‹σ1' ∈ iterate_sigma k2 b I Σ σ1 ∧ σ2' ∈ not_set b (iterate_sigma k2 b I Σ σ2)› all_same assms(1) assms(2) assms(3) member_filter not_set_def)
then have "⋀x1. x1 ∈ iterate_sigma k2 b I Σ σ1 ⟹ ¬ bdenot b (fst x1)"
using ‹σ1' ∈ iterate_sigma k2 b I Σ σ1 ∧ σ2' ∈ not_set b (iterate_sigma k2 b I Σ σ2)› all_same always_sat_refl assms(1) assms(2) by blast
then have "iterate_sigma (Suc k2) b I Σ σ1 = {}" by auto
moreover have "⋀x2. x2 ∈ iterate_sigma k2 b I Σ σ2 ⟹ ¬ bdenot b (fst x2)"
using ‹¬ bdenot b (fst σ1')› ‹σ1' ∈ iterate_sigma k2 b I Σ σ1 ∧ σ2' ∈ not_set b (iterate_sigma k2 b I Σ σ2)› all_same assms(1) assms(2) assms(3) by blast
then have "iterate_sigma (Suc k2) b I Σ σ2 = {}" by auto
then show ?thesis
by (metis (no_types, lifting) Collect_empty_eq False Set.filter_def ‹Set.filter (λσ. ¬ bdenot b (fst σ)) (iterate_sigma k1 b I Σ σ1) ≠ {}› calculation empty_iff iterate_empty_later_empty not_less_eq_eq)
qed
qed
theorem while_rule1:
fixes Δ :: "('i, 'a, nat) cont"
assumes "hoare_triple_valid Δ (And I (Bool b)) C (And I (Low b))"
shows "hoare_triple_valid Δ (And I (Low b)) (Cwhile b C) (And I (Bool (Bnot b)))"
proof -
obtain Σ where asm0: "⋀σ n. σ, σ ⊨ (And I (Bool b)) ⟹ bounded (snd σ) ⟹ safe n Δ C σ (Σ σ)"
and "⋀σ σ'. σ, σ' ⊨ (And I (Bool b)) ⟹ pair_sat (Σ σ) (Σ σ') (And I (Low b))"
using assms(1) hoare_triple_validE by blast
let ?Σ = "λσ. not_set b (⋃n. iterate_sigma n b I Σ σ)"
show ?thesis
proof (rule hoare_triple_validI_bounded)
show "⋀s h s' h'. (s, h), (s', h') ⊨ And I (Low b) ⟹ pair_sat (?Σ (s, h)) (?Σ (s', h')) (And I (Bool (Bnot b)))"
proof -
fix s1 h1 s2 h2 assume asm0: "(s1, h1), (s2, h2) ⊨ And I (Low b)"
then have asm0_bis: "(s1, h1), (s2, h2) ⊨ I ∧ bdenot b (fst (s1, h1)) = bdenot b (fst (s2, h2))" by auto
show "pair_sat (not_set b (⋃n. iterate_sigma n b I Σ (s1, h1))) (not_set b (⋃n. iterate_sigma n b I Σ (s2, h2))) (And I (Bool (Bnot b)))"
proof (rule pair_satI)
fix s1' h1' s2' h2'
assume asm1: "(s1', h1') ∈ not_set b (⋃n. iterate_sigma n b I Σ (s1, h1)) ∧ (s2', h2') ∈ not_set b (⋃n. iterate_sigma n b I Σ (s2, h2))"
then obtain k where "not_set b (⋃n. iterate_sigma n b I Σ (s1, h1)) = not_set b (iterate_sigma k b I Σ (s1, h1))"
"not_set b (⋃n. iterate_sigma n b I Σ (s2, h2)) = not_set b (iterate_sigma k b I Σ (s2, h2))"
using union_exists_at_some_point_exactly[of I b Σ "(s1, h1)" "(s2, h2)"] asm0_bis not_set_def
using ‹⋀σ' σ. σ, σ' ⊨ And I (Bool b) ⟹ pair_sat (Σ σ) (Σ σ') (And I (Low b))› by blast
moreover have "pair_sat (iterate_sigma k b I Σ (s1, h1)) (iterate_sigma k b I Σ (s2, h2)) (And I (Low b))"
using ‹⋀σ' σ. σ, σ' ⊨ And I (Bool b) ⟹ pair_sat (Σ σ) (Σ σ') (And I (Low b))› asm0_bis iterate_sigma_low_all_sat_I_and_low by blast
ultimately show "(s1', h1'), (s2', h2') ⊨ And I (Bool (Bnot b))"
by (metis (no_types, lifting) asm1 bdenot.simps(3) fst_conv hyper_sat.simps(1) hyper_sat.simps(3) member_filter not_set_def pair_satE)
qed
qed
fix s h n
assume asm1: "(s, h), (s, h) ⊨ And I (Low b)" "bounded h"
have "safe n Δ (Cwhile b C) (s, h) (trans_Σ b I Σ (s, h))"
proof (rule safe_while)
show "⋀σ σ'. σ, σ' ⊨ And I (Bool b) ⟹ pair_sat (Σ σ) (Σ σ') I"
by (meson ‹⋀σ' σ. σ, σ' ⊨ And I (Bool b) ⟹ pair_sat (Σ σ) (Σ σ') (And I (Low b))› hyper_sat.simps(3) pair_sat_def)
show "(s, h), (s, h) ⊨ I"
using asm1 by auto
show "leads_to_loop b I Σ (s, h) (s, h)"
by (simp add: leads_to_loop.intros(1))
show "⋀σ m. σ, σ ⊨ And I (Bool b) ⟹ bounded (snd σ) ⟹ safe n Δ C σ (Σ σ)"
by (simp add: asm0 asm1)
show "bounded h"
by (simp add: asm1(2))
qed
then show "safe n Δ (Cwhile b C) (s, h) (not_set b (⋃n. iterate_sigma n b I Σ (s, h)))"
by (simp add: not_set_def safe_larger_set trans_included)
qed
qed
lemma entails_smallerI:
assumes "⋀s1 h1 s2 h2. (s1, h1), (s2, h2) ⊨ A ⟹ (s1, h1), (s2, h2) ⊨ B"
shows "entails A B"
by (simp add: assms entails_def)
corollary while_rule:
fixes Δ :: "('i, 'a, nat) cont"
assumes "entails P (Star P' R)"
and "unary P'"
and "fvA R ∩ wrC C = {}"
and "hoare_triple_valid Δ (And P' (Bool e)) C P'"
and "hoare_triple_valid Δ (And P (Bool (Band e e'))) C (And P (Low (Band e e')))"
and "precise P' ∨ precise R"
shows "hoare_triple_valid Δ (And P (Low (Band e e'))) (Cseq (Cwhile (Band e e') C) (Cwhile e C)) (And (Star P' R) (Bool (Bnot e)))"
proof (rule seq_rule)
show "hoare_triple_valid Δ (And P (Low (Band e e'))) (Cwhile (Band e e') C) (And P (Bool (Bnot (Band e e'))))"
proof (rule while_rule1)
show "hoare_triple_valid Δ (And P (Bool (Band e e'))) C (And P (Low (Band e e')))"
by (simp add: assms(5))
qed
show "hoare_triple_valid Δ (And P (Bool (Bnot (Band e e')))) (Cwhile e C) (And (Star P' R) (Bool (Bnot e)))"
proof (rule consequence_rule)
show "hoare_triple_valid Δ (Star P' R) (Cwhile e C) (Star (And P' (Bool (Bnot e))) R)"
proof (rule frame_rule)
show "precise P' ∨ precise R"
by (simp add: assms(6))
show "disjoint (fvA R) (wrC (Cwhile e C))"
by (simp add: assms(3) disjoint_def)
show "hoare_triple_valid Δ P' (Cwhile e C) (And P' (Bool (Bnot e)))"
proof (rule while_rule2)
show "hoare_triple_valid Δ (And P' (Bool e)) C P'"
by (simp add: assms(4))
show "unary P'" using assms(2) by auto
qed
qed
show "entails (And P (Bool (Bnot (Band e e')))) (Star P' R)"
using assms(1) entails_def hyper_sat.simps(3) by blast
show "entails (Star (And P' (Bool (Bnot e))) R) (And (Star P' R) (Bool (Bnot e)))"
proof (rule entails_smallerI)
fix s1 h1 s2 h2
assume asm0: "(s1, h1), (s2, h2) ⊨ Star (And P' (Bool (Bnot e))) R"
then obtain hp1 hr1 hp2 hr2 where "Some h1 = Some hp1 ⊕ Some hr1" "Some h2 = Some hp2 ⊕ Some hr2"
"(s1, hp1), (s2, hp2) ⊨ And P' (Bool (Bnot e))" "(s1, hr1), (s2, hr2) ⊨ R"
using hyper_sat.simps(4) by blast
then show "(s1, h1), (s2, h2) ⊨ And (Star P' R) (Bool (Bnot e))"
by fastforce
qed
qed
qed
subsubsection ‹CommCSL is sound›
theorem soundness:
assumes "Δ ⊢ {P} C {Q}"
shows "Δ ⊨ {P} C {Q}"
using assms
proof (induct rule: CommCSL.induct)
case (RuleAtomicShared Γ f α sact uact J P Q x C map_to_arg sarg ms map_to_multiset π)
then show ?case using atomic_rule_shared by blast
qed (simp_all add: rule_skip assign_rule new_rule write_rule read_rule share_rule atomic_rule_unique
rule_par if1_rule if2_rule seq_rule frame_rule consequence_rule existential_rule while_rule1 while_rule2)
subsection ‹Corollaries›
text ‹The two following corollaries express what proving a Hoare triple in CommCSL with no invariant (initially)
guarantees, i.e., that if C is executed in two states that together satisfy the precondition P,
then no execution will abort, and any pair of final states will satisfy together the postcondition Q.›
text ‹This first corollary considers that the heap h1 is part of a larger execution with heap H1.›
theorem safety:
assumes "hoare_triple_valid (None :: ('i, 'a, nat) cont) P C Q"
and "(s1, h1), (s2, h2) ⊨ P"
and "Some H1 = Some h1 ⊕ Some hf1 ∧ full_ownership (get_fh H1) ∧ no_guard H1"
and "Some H2 = Some h2 ⊕ Some hf2 ∧ full_ownership (get_fh H2) ∧ no_guard H2"
shows "⋀σ' C'. red_rtrans C (s1, normalize (get_fh H1)) C' σ' ⟹ ¬ aborts C' σ'"
and "⋀σ' C'. red_rtrans C (s2, normalize (get_fh H2)) C' σ' ⟹ ¬ aborts C' σ'"
and "⋀σ1' σ2'. red_rtrans C (s1, normalize (get_fh H1)) Cskip σ1'
⟹ red_rtrans C (s2, normalize (get_fh H2)) Cskip σ2'
⟹ (∃h1' h2' H1' H2'. no_guard H1' ∧ full_ownership (get_fh H1') ∧ snd σ1' = normalize (get_fh H1') ∧ Some H1' = Some h1' ⊕ Some hf1
∧ no_guard H2' ∧ full_ownership (get_fh H2') ∧ snd σ2' = normalize (get_fh H2') ∧ Some H2' = Some h2' ⊕ Some hf2
∧ (fst σ1', h1'), (fst σ2', h2') ⊨ Q)"
proof -
obtain Σ where asm0: "⋀σ n. σ, σ ⊨ P ⟹ bounded (snd σ) ⟹ safe n (None :: ('i, 'a, nat) cont) C σ (Σ σ)"
"⋀σ σ'. σ, σ' ⊨ P ⟹ pair_sat (Σ σ) (Σ σ') Q"
using assms(1) hoare_triple_validE by blast
then have "pair_sat (Σ (s1, h1)) (Σ (s2, h2)) Q"
using assms(2) by blast
moreover have "bounded h1"
by (metis assms(3) bounded_smaller_sum full_ownership_then_bounded get_fh.simps)
then have "⋀n. safe n (None :: ('i, 'a, nat) cont) C (s1, h1) (Σ (s1, h1))"
using always_sat_refl asm0(1) assms(2)
by (metis snd_conv)
then show "⋀σ' C'. red_rtrans C (s1, FractionalHeap.normalize (get_fh H1)) C' σ' ⟹ ¬ aborts C' σ'"
proof -
fix σ' C'
assume "red_rtrans C (s1, FractionalHeap.normalize (get_fh H1)) C' σ'"
then show "¬ aborts C' σ'"
using safe_atomic[of C "(s1, FractionalHeap.normalize (get_fh H1))" C' σ' s1 "FractionalHeap.normalize (get_fh H1)" "fst σ'" "snd σ'"]
by (metis ‹⋀n. safe n None C (s1, h1) (Σ (s1, h1))› assms(3) denormalize_properties(4) prod.exhaust_sel)
qed
moreover have "⋀n. safe n (None :: ('i, 'a, nat) cont) C (s2, h2) (Σ (s2, h2))"
using always_sat_refl asm0(1) assms(2) sat_comm_aux
by (metis assms(4) bounded_smaller_sum full_ownership_then_bounded get_fh.elims snd_conv)
then show "⋀σ' C'. red_rtrans C (s2, FractionalHeap.normalize (get_fh H2)) C' σ' ⟹ ¬ aborts C' σ'"
proof -
fix σ' C'
assume "red_rtrans C (s2, FractionalHeap.normalize (get_fh H2)) C' σ'"
then show "¬ aborts C' σ'"
using safe_atomic[of C "(s2, FractionalHeap.normalize (get_fh H2))" C' σ' s2 "FractionalHeap.normalize (get_fh H2)" "fst σ'" "snd σ'"]
by (metis ‹⋀n. safe n None C (s2, h2) (Σ (s2, h2))› assms(4) denormalize_properties(4) prod.exhaust_sel)
qed
fix σ1'
assume "red_rtrans C (s1, FractionalHeap.normalize (get_fh H1)) Cskip σ1'"
then obtain h1' H1' where r1: "Some H1' = Some h1' ⊕ Some hf1" "snd σ1' = FractionalHeap.normalize (get_fh H1')"
"no_guard H1' ∧ full_ownership (get_fh H1')" "(fst σ1', h1') ∈ Σ (s1, h1)"
using safe_atomic[of C "(s1, FractionalHeap.normalize (get_fh H1))" Cskip σ1' s1 _ "fst σ1'" "snd σ1'" h1 "Σ (s1, h1)" H1 hf1]
by (metis ‹⋀n. safe n None C (s1, h1) (Σ (s1, h1))› assms(3) denormalize_properties(4) surjective_pairing)
fix σ2'
assume "red_rtrans C (s2, FractionalHeap.normalize (get_fh H2)) Cskip σ2'"
then obtain h2' H2' where r2: "Some H2' = Some h2' ⊕ Some hf2" "snd σ2' = FractionalHeap.normalize (get_fh H2')"
"no_guard H2' ∧ full_ownership (get_fh H2')" "(fst σ2', h2') ∈ Σ (s2, h2)"
using safe_atomic[of C "(s2, FractionalHeap.normalize (get_fh H2))" Cskip σ2' s2 _ "fst σ2'" "snd σ2'" h2 "Σ (s2, h2)" H2 hf2]
by (metis ‹⋀n. safe n None C (s2, h2) (Σ (s2, h2))› assms(4) denormalize_properties(4) surjective_pairing)
then have "(fst σ1', h1'), (fst σ2', h2') ⊨ Q"
using calculation(1) pair_satE r1(4) by blast
then show "∃h1' h2' H1' H2'.
no_guard H1' ∧
full_ownership (get_fh H1') ∧
snd σ1' = FractionalHeap.normalize (get_fh H1') ∧
Some H1' = Some h1' ⊕ Some hf1 ∧
no_guard H2' ∧
full_ownership (get_fh H2') ∧ snd σ2' = FractionalHeap.normalize (get_fh H2') ∧ Some H2' = Some h2' ⊕ Some hf2 ∧ (fst σ1', h1'), (fst σ2', h2') ⊨ Q"
using r1 r2 by blast
qed
lemma neutral_add:
"Some h = Some h ⊕ Some (Map.empty, None, (λ_. None))"
proof -
have "h ## (Map.empty, None, (λ_. None))"
by (metis compatibleI compatible_fract_heapsI empty_heap_def fst_conv get_fh.elims get_gs.simps get_gu.simps option.distinct(1) snd_conv)
then obtain x where "Some x = Some h ⊕ Some (Map.empty, None, (λ_. None))"
by simp
moreover have "x = h"
by (metis (no_types, lifting) addition_cancellative calculation decompose_guard_remove_easy fst_eqD get_gs.simps get_gu.simps no_guard_def no_guards_remove prod.sel(2) simpler_asso)
ultimately show ?thesis by blast
qed
text ‹This second corollary considers that the heap h1 is the only execution that matters, and
thus it ignores any frame. It corresponds to Corollary 4.5 in the paper.›
corollary safety_no_frame:
assumes "hoare_triple_valid (None :: ('i, 'a, nat) cont) P C Q"
and "(s1, H1), (s2, H2) ⊨ P"
and "full_ownership (get_fh H1) ∧ no_guard H1"
and "full_ownership (get_fh H2) ∧ no_guard H2"
shows "⋀σ' C'. red_rtrans C (s1, normalize (get_fh H1)) C' σ' ⟹ ¬ aborts C' σ'"
and "⋀σ' C'. red_rtrans C (s2, normalize (get_fh H2)) C' σ' ⟹ ¬ aborts C' σ'"
and "⋀σ1' σ2'. red_rtrans C (s1, normalize (get_fh H1)) Cskip σ1'
⟹ red_rtrans C (s2, normalize (get_fh H2)) Cskip σ2'
⟹ (∃H1' H2'. no_guard H1' ∧ full_ownership (get_fh H1') ∧ snd σ1' = normalize (get_fh H1')
∧ no_guard H2' ∧ full_ownership (get_fh H2') ∧ snd σ2' = normalize (get_fh H2')
∧ (fst σ1', H1'), (fst σ2', H2') ⊨ Q)"
proof -
have "Some H1 = Some H1 ⊕ Some (Map.empty, None, (λ_. None))"
using neutral_add by blast
moreover have "Some H2 = Some H2 ⊕ Some (Map.empty, None, (λ_. None))"
using neutral_add by blast
show "⋀σ' C'. red_rtrans C (s1, FractionalHeap.normalize (get_fh H1)) C' σ' ⟹ ¬ aborts C' σ'"
using always_sat_refl_aux assms(1) assms(2) assms(3) calculation safety(2) by blast
show "⋀σ' C'. red_rtrans C (s2, FractionalHeap.normalize (get_fh H2)) C' σ' ⟹ ¬ aborts C' σ'"
using ‹Some H2 = Some H2 ⊕ Some (Map.empty, None, (λ_. None))› assms(1) assms(2) assms(3) assms(4) calculation safety(2) by blast
fix σ1' σ2'
assume "red_rtrans C (s1, FractionalHeap.normalize (get_fh H1)) Cskip σ1'"
"red_rtrans C (s2, FractionalHeap.normalize (get_fh H2)) Cskip σ2'"
then obtain h1' h2' H1' H2' where asm0: "no_guard H1' ∧ full_ownership (get_fh H1') ∧ snd σ1' = normalize (get_fh H1') ∧ Some H1' = Some h1' ⊕ Some (Map.empty, None, (λ_. None))
∧ no_guard H2' ∧ full_ownership (get_fh H2') ∧ snd σ2' = normalize (get_fh H2') ∧ Some H2' = Some h2' ⊕ Some (Map.empty, None, (λ_. None))
∧ (fst σ1', h1'), (fst σ2', h2') ⊨ Q"
using safety[of P C Q s1 H1 s2 H2 H1 "(Map.empty, None, (λ_. None))" H2 "(Map.empty, None, (λ_. None))"] assms
by (metis (no_types, lifting) ‹Some H2 = Some H2 ⊕ Some (Map.empty, None, (λ_. None))› calculation)
then have "H1' = h1'"
using addition_cancellative decompose_guard_remove_easy denormalize_properties(4) denormalize_properties(5)
by (metis denormalize_def get_gs.simps get_gu.simps prod.exhaust_sel snd_conv)
moreover have "H2' = h2'"
by (metis asm0 denormalize_properties(4) denormalize_properties(5) fst_eqD get_fh.elims no_guard_and_no_heap no_guard_then_smaller_same)
ultimately show "∃H1' H2'.
no_guard H1' ∧
full_ownership (get_fh H1') ∧
snd σ1' = FractionalHeap.normalize (get_fh H1') ∧
no_guard H2' ∧ full_ownership (get_fh H2') ∧ snd σ2' = FractionalHeap.normalize (get_fh H2') ∧ (fst σ1', H1'), (fst σ2', H2') ⊨ Q"
using asm0 by blast
qed
end