Skip to content
Snippets Groups Projects
Commit 9646293e authored by Jacques-Henri Jourdan's avatar Jacques-Henri Jourdan
Browse files

A specific constructor for injecting values in expressions

We add a specific constructor to the type of expressions for injecting
values in expressions.

The advantage are :
- Values can be assumed to be always closed when performing
  substitutions (even though they could contain free variables, but it
  turns out it does not cause any problem in the proofs in
  practice). This means that we no longer need the `Closed` typeclass
  and everything that comes with it (all the reflection-based machinery
  contained in tactics.v is no longer necessary). I have not measured
  anything, but I guess this would have a significant performance
  impact.

- There is only one constructor for values. As a result, the AsVal and
  IntoVal typeclasses are no longer necessary: an expression which is
  a value will always unify with `Val _`, and therefore lemmas can be
  stated using this constructor.

Of course, this means that there are two ways of writing such a thing
as "The pair of integers 1 and 2": Either by using the value
constructor applied to the pair represented as a value, or by using
the expression pair constructor. So we add reduction rules that
transform reduced pair, injection and closure expressions into values.
At first, this seems weird, because of the redundancy. But in fact,
this has some meaning, since the machine migth actually be doing
something to e.g., allocate the pair or the closure.

These additional steps of computation show up in the proofs, and some
additional wp_* tactics need to be called.
parent 1f796221
No related branches found
No related tags found
No related merge requests found
Showing
with 384 additions and 414 deletions
...@@ -81,6 +81,7 @@ theories/program_logic/total_lifting.v ...@@ -81,6 +81,7 @@ theories/program_logic/total_lifting.v
theories/program_logic/total_ectx_lifting.v theories/program_logic/total_ectx_lifting.v
theories/program_logic/atomic.v theories/program_logic/atomic.v
theories/heap_lang/lang.v theories/heap_lang/lang.v
theories/heap_lang/metatheory.v
theories/heap_lang/tactics.v theories/heap_lang/tactics.v
theories/heap_lang/lifting.v theories/heap_lang/lifting.v
theories/heap_lang/notation.v theories/heap_lang/notation.v
......
...@@ -50,6 +50,18 @@ ...@@ -50,6 +50,18 @@
--------------------------------------∗ --------------------------------------∗
True True
The command has indeed failed with message:
Ltac call to "wp_pure (open_constr)" failed.
Tactic failure: wp_pure: cannot find ?y in (Var "x") or
?y is not a redex.
1 subgoal
Σ : gFunctors
H : heapG Σ
============================
--------------------------------------∗
WP "x" {{ _, True }}
1 subgoal 1 subgoal
Σ : gFunctors Σ : gFunctors
...@@ -104,4 +116,4 @@ ...@@ -104,4 +116,4 @@
: string : string
The command has indeed failed with message: The command has indeed failed with message:
Ltac call to "wp_cas_suc" failed. Ltac call to "wp_cas_suc" failed.
Tactic failure: wp_cas_suc: cannot find 'CAS' in (#())%E. Tactic failure: wp_cas_suc: cannot find 'CAS' in (Val #()).
...@@ -131,6 +131,10 @@ Section tests. ...@@ -131,6 +131,10 @@ Section tests.
WP Alloc #0 {{ _, True }}%I. WP Alloc #0 {{ _, True }}%I.
Proof. wp_alloc l as "_". Show. done. Qed. Proof. wp_alloc l as "_". Show. done. Qed.
Lemma wp_nonclosed_value :
WP let: "x" := #() in (λ: "y", "x")%V #() {{ _, True }}%I.
Proof. wp_let. wp_lam. Fail wp_pure _. Show. Abort.
End tests. End tests.
Section printing_tests. Section printing_tests.
......
...@@ -82,14 +82,14 @@ Section list_reverse. ...@@ -82,14 +82,14 @@ Section list_reverse.
destruct xs as [|x xs]; iSimplifyEq. destruct xs as [|x xs]; iSimplifyEq.
- (* nil *) by wp_match. - (* nil *) by wp_match.
- (* cons *) iDestruct "Hxs" as (l hd') "(% & Hx & Hxs)"; iSimplifyEq. - (* cons *) iDestruct "Hxs" as (l hd') "(% & Hx & Hxs)"; iSimplifyEq.
wp_match. wp_load. wp_proj. wp_let. wp_load. wp_proj. wp_let. wp_store. wp_match. wp_load. wp_proj. wp_let. wp_load. wp_proj. wp_let. wp_pair. wp_store.
rewrite reverse_cons -assoc. rewrite reverse_cons -assoc.
iApply ("IH" $! hd' (InjRV #l) xs (x :: ys) with "Hxs [Hx Hys]"). iApply ("IH" $! hd' (InjRV #l) xs (x :: ys) with "Hxs [Hx Hys]").
iExists l, acc; by iFrame. iExists l, acc; by iFrame.
Qed. Qed.
Lemma rev_ht hd xs : Lemma rev_ht hd xs :
{{ is_list hd xs }} rev hd NONE {{ w, is_list w (reverse xs) }}. {{ is_list hd xs }} rev hd NONEV {{ w, is_list w (reverse xs) }}.
Proof. Proof.
iIntros "!# Hxs". rewrite -(right_id_L [] (++) (reverse xs)). iIntros "!# Hxs". rewrite -(right_id_L [] (++) (reverse xs)).
iApply (rev_acc_ht hd NONEV with "[Hxs]"); simpl; by iFrame. iApply (rev_acc_ht hd NONEV with "[Hxs]"); simpl; by iFrame.
...@@ -204,7 +204,7 @@ Section counter_proof. ...@@ -204,7 +204,7 @@ Section counter_proof.
Lemma newcounter_spec : Lemma newcounter_spec :
{{ True }} newcounter #() {{ v, l, v = #l C l 0 }}. {{ True }} newcounter #() {{ v, l, v = #l C l 0 }}.
Proof. Proof.
iIntros "!# _ /=". rewrite -wp_fupd /newcounter /=. wp_seq. wp_alloc l as "Hl". iIntros "!# _ /=". rewrite -wp_fupd /newcounter /=. wp_lam. wp_alloc l as "Hl".
iMod (own_alloc (Auth 0)) as (γ) "Hγ"; first done. iMod (own_alloc (Auth 0)) as (γ) "Hγ"; first done.
rewrite (auth_frag_op 0 0) //; iDestruct "Hγ" as "[Hγ Hγf]". rewrite (auth_frag_op 0 0) //; iDestruct "Hγ" as "[Hγ Hγf]".
set (N:= nroot .@ "counter"). set (N:= nroot .@ "counter").
...@@ -242,7 +242,7 @@ Section counter_proof. ...@@ -242,7 +242,7 @@ Section counter_proof.
{{ C l n }} read #l {{ v, m : nat, v = #m n m C l m }}. {{ C l n }} read #l {{ v, m : nat, v = #m n m C l m }}.
Proof. Proof.
iIntros "!# Hl /=". iDestruct "Hl" as (N γ) "[#Hinv Hγf]". iIntros "!# Hl /=". iDestruct "Hl" as (N γ) "[#Hinv Hγf]".
rewrite /read /=. wp_let. Show. iApply wp_inv_open; last iFrame "Hinv"; auto. rewrite /read /=. wp_lam. Show. iApply wp_inv_open; last iFrame "Hinv"; auto.
iDestruct 1 as (c) "[Hl Hγ]". wp_load. Show. iDestruct 1 as (c) "[Hl Hγ]". wp_load. Show.
iDestruct (own_valid γ (Frag n Auth c) with "[-]") as % ?%auth_frag_valid. iDestruct (own_valid γ (Frag n Auth c) with "[-]") as % ?%auth_frag_valid.
{ iApply own_op. by iFrame. } { iApply own_op. by iFrame. }
......
...@@ -23,11 +23,11 @@ ...@@ -23,11 +23,11 @@
"Hys" : is_list acc ys "Hys" : is_list acc ys
"HΦ" : ∀ w : val, is_list w ys -∗ Φ w "HΦ" : ∀ w : val, is_list w ys -∗ Φ w
--------------------------------------∗ --------------------------------------∗
WP match: InjL #() with WP match: InjLV #() with
InjL <> => acc InjL <> => acc
| InjR "l" => | InjR "l" =>
let: "tmp1" := Fst ! "l" in let: "tmp1" := Fst ! "l" in
let: "tmp2" := Snd ! "l" in let: "tmp2" := Snd ! "l" in
"l" <- ("tmp1", acc);; (rev "tmp2") (InjL #()) "l" <- ("tmp1", acc);; (rev "tmp2") (InjLV #())
end [{ v, Φ v }] end [{ v, Φ v }]
...@@ -36,14 +36,14 @@ Proof. ...@@ -36,14 +36,14 @@ Proof.
iSimplifyEq; wp_rec; wp_let. iSimplifyEq; wp_rec; wp_let.
- Show. wp_match. by iApply "HΦ". - Show. wp_match. by iApply "HΦ".
- iDestruct "Hxs" as (l hd' ->) "[Hx Hxs]". - iDestruct "Hxs" as (l hd' ->) "[Hx Hxs]".
wp_match. wp_load. wp_proj. wp_let. wp_load. wp_proj. wp_let. wp_store. wp_match. wp_load. wp_proj. wp_let. wp_load. wp_proj. wp_let. wp_pair. wp_store.
iApply ("IH" $! hd' (SOMEV #l) (x :: ys) with "Hxs [Hx Hys]"); simpl. iApply ("IH" $! hd' (SOMEV #l) (x :: ys) with "Hxs [Hx Hys]"); simpl.
{ iExists l, acc; by iFrame. } { iExists l, acc; by iFrame. }
iIntros (w). rewrite cons_middle assoc -reverse_cons. iApply "HΦ". iIntros (w). rewrite cons_middle assoc -reverse_cons. iApply "HΦ".
Qed. Qed.
Lemma rev_wp hd xs : Lemma rev_wp hd xs :
[[{ is_list hd xs }]] rev hd NONE [[{ w, RET w; is_list w (reverse xs) }]]. [[{ is_list hd xs }]] rev hd NONEV [[{ w, RET w; is_list w (reverse xs) }]].
Proof. Proof.
iIntros (Φ) "Hxs HΦ". iIntros (Φ) "Hxs HΦ".
iApply (rev_acc_wp hd NONEV xs [] with "[$Hxs //]"). iApply (rev_acc_wp hd NONEV xs [] with "[$Hxs //]").
......
...@@ -35,7 +35,7 @@ ...@@ -35,7 +35,7 @@
"Hγ" : own γ (Shot m') "Hγ" : own γ (Shot m')
--------------------------------------∗ --------------------------------------∗
|={⊤ ∖ ↑N}=> ▷ one_shot_inv γ l |={⊤ ∖ ↑N}=> ▷ one_shot_inv γ l
∗ WP match: InjR #m' with ∗ WP match: InjRV #m' with
InjL <> => assert: #false InjL <> => assert: #false
| InjR "m" => assert: #m = "m" | InjR "m" => assert: #m = "m"
end {{ _, True }} end {{ _, True }}
......
...@@ -43,12 +43,12 @@ Lemma wp_one_shot (Φ : val → iProp Σ) : ...@@ -43,12 +43,12 @@ Lemma wp_one_shot (Φ : val → iProp Σ) :
WP one_shot_example #() {{ Φ }}. WP one_shot_example #() {{ Φ }}.
Proof. Proof.
iIntros "Hf /=". pose proof (nroot .@ "N") as N. iIntros "Hf /=". pose proof (nroot .@ "N") as N.
rewrite -wp_fupd /one_shot_example /=. wp_seq. wp_alloc l as "Hl". wp_let. rewrite -wp_fupd /one_shot_example /=. wp_lam. wp_inj. wp_alloc l as "Hl". wp_let.
iMod (own_alloc Pending) as (γ) "Hγ"; first done. iMod (own_alloc Pending) as (γ) "Hγ"; first done.
iMod (inv_alloc N _ (one_shot_inv γ l) with "[Hl Hγ]") as "#HN". iMod (inv_alloc N _ (one_shot_inv γ l) with "[Hl Hγ]") as "#HN".
{ iNext. iLeft. by iSplitL "Hl". } { iNext. iLeft. by iSplitL "Hl". }
iModIntro. iApply "Hf"; iSplit. wp_closure. wp_closure. wp_pair. iModIntro. iApply "Hf"; iSplit.
- iIntros (n) "!#". wp_let. - iIntros (n) "!#". wp_lam. wp_inj. wp_inj.
iInv N as ">[[Hl Hγ]|H]"; last iDestruct "H" as (m) "[Hl Hγ]". iInv N as ">[[Hl Hγ]|H]"; last iDestruct "H" as (m) "[Hl Hγ]".
+ iMod (own_update with "Hγ") as "Hγ". + iMod (own_update with "Hγ") as "Hγ".
{ by apply cmra_update_exclusive with (y:=Shot n). } { by apply cmra_update_exclusive with (y:=Shot n). }
...@@ -56,7 +56,7 @@ Proof. ...@@ -56,7 +56,7 @@ Proof.
iModIntro. iNext; iRight; iExists n; by iFrame. iModIntro. iNext; iRight; iExists n; by iFrame.
+ wp_cas_fail. iSplitL; last eauto. + wp_cas_fail. iSplitL; last eauto.
rewrite /one_shot_inv; eauto 10. rewrite /one_shot_inv; eauto 10.
- iIntros "!# /=". wp_seq. wp_bind (! _)%E. - iIntros "!# /=". wp_lam. wp_bind (! _)%E.
iInv N as ">Hγ". iInv N as ">Hγ".
iAssert ( v, l v ((v = NONEV own γ Pending) iAssert ( v, l v ((v = NONEV own γ Pending)
n : Z, v = SOMEV #n own γ (Shot n)))%I with "[Hγ]" as "Hv". n : Z, v = SOMEV #n own γ (Shot n)))%I with "[Hγ]" as "Hv".
...@@ -70,7 +70,7 @@ Proof. ...@@ -70,7 +70,7 @@ Proof.
+ Show. iSplit. iLeft; by iSplitL "Hl". eauto. + Show. iSplit. iLeft; by iSplitL "Hl". eauto.
+ iSplit. iRight; iExists m; by iSplitL "Hl". eauto. } + iSplit. iRight; iExists m; by iSplitL "Hl". eauto. }
iSplitL "Hinv"; first by eauto. iSplitL "Hinv"; first by eauto.
iModIntro. wp_let. iIntros "!#". wp_seq. iModIntro. wp_let. wp_closure. iIntros "!#". wp_lam.
iDestruct "Hv" as "[%|Hv]"; last iDestruct "Hv" as (m) "[% Hγ']"; subst. iDestruct "Hv" as "[%|Hv]"; last iDestruct "Hv" as (m) "[% Hγ']"; subst.
{ by wp_match. } { by wp_match. }
wp_match. wp_bind (! _)%E. wp_match. wp_bind (! _)%E.
......
...@@ -58,7 +58,7 @@ Lemma sum_wp `{!heapG Σ} v t : ...@@ -58,7 +58,7 @@ Lemma sum_wp `{!heapG Σ} v t :
[[{ is_tree v t }]] sum' v [[{ RET #(sum t); is_tree v t }]]. [[{ is_tree v t }]] sum' v [[{ RET #(sum t); is_tree v t }]].
Proof. Proof.
iIntros (Φ) "Ht HΦ". rewrite /sum' /=. iIntros (Φ) "Ht HΦ". rewrite /sum' /=.
wp_let. wp_alloc l as "Hl". wp_let. wp_lam. wp_alloc l as "Hl". wp_let.
wp_apply (sum_loop_wp with "[$Hl $Ht]"). wp_apply (sum_loop_wp with "[$Hl $Ht]").
rewrite Z.add_0_r. rewrite Z.add_0_r.
iIntros "[Hl Ht]". wp_seq. wp_load. by iApply "HΦ". iIntros "[Hl Ht]". wp_seq. wp_load. by iApply "HΦ".
......
This diff is collapsed.
...@@ -9,16 +9,16 @@ Definition assert : val := ...@@ -9,16 +9,16 @@ Definition assert : val :=
(* just below ;; *) (* just below ;; *)
Notation "'assert:' e" := (assert (λ: <>, e))%E (at level 99) : expr_scope. Notation "'assert:' e" := (assert (λ: <>, e))%E (at level 99) : expr_scope.
Lemma twp_assert `{heapG Σ} E (Φ : val iProp Σ) e `{!Closed [] e} : Lemma twp_assert `{heapG Σ} E (Φ : val iProp Σ) e :
WP e @ E [{ v, v = #true Φ #() }] -∗ WP assert: e @ E [{ Φ }]. WP e @ E [{ v, v = #true Φ #() }] -∗ WP assert: e @ E [{ Φ }].
Proof. Proof.
iIntros "HΦ". rewrite /assert. wp_let. wp_seq. iIntros "HΦ". rewrite /assert. wp_closure. wp_lam. wp_lam.
wp_apply (twp_wand with "HΦ"). iIntros (v) "[% ?]"; subst. by wp_if. wp_apply (twp_wand with "HΦ"). iIntros (v) "[% ?]"; subst. by wp_if.
Qed. Qed.
Lemma wp_assert `{heapG Σ} E (Φ : val iProp Σ) e `{!Closed [] e} : Lemma wp_assert `{heapG Σ} E (Φ : val iProp Σ) e :
WP e @ E {{ v, v = #true Φ #() }} -∗ WP assert: e @ E {{ Φ }}. WP e @ E {{ v, v = #true Φ #() }} -∗ WP assert: e @ E {{ Φ }}.
Proof. Proof.
iIntros "HΦ". rewrite /assert. wp_let. wp_seq. iIntros "HΦ". rewrite /assert. wp_closure. wp_lam. wp_lam.
wp_apply (wp_wand with "HΦ"). iIntros (v) "[% ?]"; subst. by wp_if. wp_apply (wp_wand with "HΦ"). iIntros (v) "[% ?]"; subst. by wp_if.
Qed. Qed.
...@@ -21,21 +21,20 @@ Class atomic_heap {Σ} `{!heapG Σ} := AtomicHeap { ...@@ -21,21 +21,20 @@ Class atomic_heap {Σ} `{!heapG Σ} := AtomicHeap {
AsFractional (mapsto l q v) (λ q, mapsto l q v) q; AsFractional (mapsto l q v) (λ q, mapsto l q v) q;
mapsto_agree l q1 q2 v1 v2 :> mapsto l q1 v1 -∗ mapsto l q2 v2 -∗ v1 = v2; mapsto_agree l q1 q2 v1 v2 :> mapsto l q1 v1 -∗ mapsto l q2 v2 -∗ v1 = v2;
(* -- operation specs -- *) (* -- operation specs -- *)
alloc_spec e v : alloc_spec (v : val) :
IntoVal e v {{{ True }}} alloc e {{{ l, RET #l; mapsto l 1 v }}}; {{{ True }}} alloc v {{{ l, RET #l; mapsto l 1 v }}};
load_spec (l : loc) : load_spec (l : loc) :
<<< (v : val) q, mapsto l q v >>> load #l @ <<< mapsto l q v, RET v >>>; <<< (v : val) q, mapsto l q v >>> load #l @ <<< mapsto l q v, RET v >>>;
store_spec (l : loc) (e : expr) (w : val) : store_spec (l : loc) (w : val) :
IntoVal e w <<< v, mapsto l 1 v >>> store #l w @
<<< v, mapsto l 1 v >>> store #l e @
<<< mapsto l 1 w, RET #() >>>; <<< mapsto l 1 w, RET #() >>>;
(* This spec is slightly weaker than it could be: It is sufficient for [w1] (* This spec is slightly weaker than it could be: It is sufficient for [w1]
*or* [v] to be unboxed. However, by writing it this way the [val_is_unboxed] *or* [v] to be unboxed. However, by writing it this way the [val_is_unboxed]
is outside the atomic triple, which makes it much easier to use -- and the is outside the atomic triple, which makes it much easier to use -- and the
spec is still good enough for all our applications. *) spec is still good enough for all our applications. *)
cas_spec (l : loc) (e1 e2 : expr) (w1 w2 : val) : cas_spec (l : loc) (w1 w2 : val) :
IntoVal e1 w1 IntoVal e2 w2 val_is_unboxed w1 val_is_unboxed w1
<<< v, mapsto l 1 v >>> cas #l e1 e2 @ <<< v, mapsto l 1 v >>> cas #l w1 w2 @
<<< if decide (v = w1) then mapsto l 1 w2 else mapsto l 1 v, <<< if decide (v = w1) then mapsto l 1 w2 else mapsto l 1 v,
RET #(if decide (v = w1) then true else false) >>>; RET #(if decide (v = w1) then true else false) >>>;
}. }.
...@@ -72,39 +71,38 @@ Definition primitive_cas : val := ...@@ -72,39 +71,38 @@ Definition primitive_cas : val :=
Section proof. Section proof.
Context `{!heapG Σ}. Context `{!heapG Σ}.
Lemma primitive_alloc_spec e v : Lemma primitive_alloc_spec (v : val) :
IntoVal e v {{{ True }}} primitive_alloc e {{{ l, RET #l; l v }}}. {{{ True }}} primitive_alloc v {{{ l, RET #l; l v }}}.
Proof. Proof.
iIntros (<- Φ) "_ HΦ". wp_let. wp_alloc l. iApply "HΦ". done. iIntros (Φ) "_ HΦ". wp_lam. wp_alloc l. iApply "HΦ". done.
Qed. Qed.
Lemma primitive_load_spec (l : loc) : Lemma primitive_load_spec (l : loc) :
<<< (v : val) q, l {q} v >>> primitive_load #l @ <<< (v : val) q, l {q} v >>> primitive_load #l @
<<< l {q} v, RET v >>>. <<< l {q} v, RET v >>>.
Proof. Proof.
iIntros (Q Φ) "? AU". wp_let. iIntros (Q Φ) "? AU". wp_lam.
iMod "AU" as (v q) "[H↦ [_ Hclose]]". iMod "AU" as (v q) "[H↦ [_ Hclose]]".
wp_load. iMod ("Hclose" with "H↦") as "HΦ". by iApply "HΦ". wp_load. iMod ("Hclose" with "H↦") as "HΦ". by iApply "HΦ".
Qed. Qed.
Lemma primitive_store_spec (l : loc) (e : expr) (w : val) : Lemma primitive_store_spec (l : loc) (w : val) :
IntoVal e w <<< v, l v >>> primitive_store #l w @
<<< v, l v >>> primitive_store #l e @
<<< l w, RET #() >>>. <<< l w, RET #() >>>.
Proof. Proof.
iIntros (<- Q Φ) "? AU". wp_lam. wp_let. iIntros (Q Φ) "? AU". wp_lam. wp_let.
iMod "AU" as (v) "[H↦ [_ Hclose]]". iMod "AU" as (v) "[H↦ [_ Hclose]]".
wp_store. iMod ("Hclose" with "H↦") as "HΦ". by iApply "HΦ". wp_store. iMod ("Hclose" with "H↦") as "HΦ". by iApply "HΦ".
Qed. Qed.
Lemma primitive_cas_spec (l : loc) e1 e2 (w1 w2 : val) : Lemma primitive_cas_spec (l : loc) (w1 w2 : val) :
IntoVal e1 w1 IntoVal e2 w2 val_is_unboxed w1 val_is_unboxed w1
<<< (v : val), l v >>> <<< (v : val), l v >>>
primitive_cas #l e1 e2 @ primitive_cas #l w1 w2 @
<<< if decide (v = w1) then l w2 else l v, <<< if decide (v = w1) then l w2 else l v,
RET #(if decide (v = w1) then true else false) >>>. RET #(if decide (v = w1) then true else false) >>>.
Proof. Proof.
iIntros (<- <- ? Q Φ) "? AU". wp_lam. wp_let. wp_let. iIntros (? Q Φ) "? AU". wp_lam. wp_let. wp_let.
iMod "AU" as (v) "[H↦ [_ Hclose]]". iMod "AU" as (v) "[H↦ [_ Hclose]]".
destruct (decide (v = w1)) as [<-|Hv]; [wp_cas_suc|wp_cas_fail]; destruct (decide (v = w1)) as [<-|Hv]; [wp_cas_suc|wp_cas_fail];
iMod ("Hclose" with "H↦") as "HΦ"; by iApply "HΦ". iMod ("Hclose" with "H↦") as "HΦ"; by iApply "HΦ".
......
...@@ -36,12 +36,11 @@ Section coinflip. ...@@ -36,12 +36,11 @@ Section coinflip.
Lemma rand_spec : Lemma rand_spec :
{{{ True }}} rand #() {{{ (b : bool), RET #b; True }}}. {{{ True }}} rand #() {{{ (b : bool), RET #b; True }}}.
Proof. Proof.
iIntros (Φ) "_ HP". iIntros (Φ) "_ HP". wp_lam. wp_alloc l as "Hl". wp_let.
wp_lam. wp_alloc l as "Hl". wp_lam.
iMod (inv_alloc N _ ( (b: bool), l #b)%I with "[Hl]") as "#Hinv"; first by eauto. iMod (inv_alloc N _ ( (b: bool), l #b)%I with "[Hl]") as "#Hinv"; first by eauto.
wp_apply wp_fork. wp_apply wp_fork.
- iInv N as (b) ">Hl". wp_store. iModIntro. iSplitL; eauto. - iInv N as (b) ">Hl". wp_store. iModIntro. iSplitL; eauto.
- wp_lam. iInv N as (b) ">Hl". wp_load. iModIntro. iSplitL "Hl"; first by eauto. - wp_seq. iInv N as (b) ">Hl". wp_load. iModIntro. iSplitL "Hl"; first by eauto.
iApply "HP". done. iApply "HP". done.
Qed. Qed.
......
...@@ -35,7 +35,7 @@ Section mono_proof. ...@@ -35,7 +35,7 @@ Section mono_proof.
Lemma newcounter_mono_spec : Lemma newcounter_mono_spec :
{{{ True }}} newcounter #() {{{ l, RET #l; mcounter l 0 }}}. {{{ True }}} newcounter #() {{{ l, RET #l; mcounter l 0 }}}.
Proof. Proof.
iIntros (Φ) "_ HΦ". rewrite -wp_fupd /newcounter /=. wp_seq. wp_alloc l as "Hl". iIntros (Φ) "_ HΦ". rewrite -wp_fupd /newcounter /=. wp_lam. wp_alloc l as "Hl".
iMod (own_alloc ( (O:mnat) (O:mnat))) as (γ) "[Hγ Hγ']"; first done. iMod (own_alloc ( (O:mnat) (O:mnat))) as (γ) "[Hγ Hγ']"; first done.
iMod (inv_alloc N _ (mcounter_inv γ l) with "[Hl Hγ]"). iMod (inv_alloc N _ (mcounter_inv γ l) with "[Hl Hγ]").
{ iNext. iExists 0%nat. by iFrame. } { iNext. iExists 0%nat. by iFrame. }
...@@ -71,7 +71,7 @@ Section mono_proof. ...@@ -71,7 +71,7 @@ Section mono_proof.
{{{ mcounter l j }}} read #l {{{ i, RET #i; j i⌝%nat mcounter l i }}}. {{{ mcounter l j }}} read #l {{{ i, RET #i; j i⌝%nat mcounter l i }}}.
Proof. Proof.
iIntros (ϕ) "Hc HΦ". iDestruct "Hc" as (γ) "[#Hinv Hγf]". iIntros (ϕ) "Hc HΦ". iDestruct "Hc" as (γ) "[#Hinv Hγf]".
rewrite /read /=. wp_let. iInv N as (c) ">[Hγ Hl]". rewrite /read /=. wp_lam. iInv N as (c) ">[Hγ Hl]".
wp_load. wp_load.
iDestruct (own_valid_2 with "Hγ Hγf") iDestruct (own_valid_2 with "Hγ Hγf")
as %[?%mnat_included _]%auth_valid_discrete_2. as %[?%mnat_included _]%auth_valid_discrete_2.
...@@ -112,7 +112,7 @@ Section contrib_spec. ...@@ -112,7 +112,7 @@ Section contrib_spec.
{{{ True }}} newcounter #() {{{ True }}} newcounter #()
{{{ γ l, RET #l; ccounter_ctx γ l ccounter γ 1 0 }}}. {{{ γ l, RET #l; ccounter_ctx γ l ccounter γ 1 0 }}}.
Proof. Proof.
iIntros (Φ) "_ HΦ". rewrite -wp_fupd /newcounter /=. wp_seq. wp_alloc l as "Hl". iIntros (Φ) "_ HΦ". rewrite -wp_fupd /newcounter /=. wp_lam. wp_alloc l as "Hl".
iMod (own_alloc (●! O%nat ◯! 0%nat)) as (γ) "[Hγ Hγ']"; first done. iMod (own_alloc (●! O%nat ◯! 0%nat)) as (γ) "[Hγ Hγ']"; first done.
iMod (inv_alloc N _ (ccounter_inv γ l) with "[Hl Hγ]"). iMod (inv_alloc N _ (ccounter_inv γ l) with "[Hl Hγ]").
{ iNext. iExists 0%nat. by iFrame. } { iNext. iExists 0%nat. by iFrame. }
...@@ -144,7 +144,7 @@ Section contrib_spec. ...@@ -144,7 +144,7 @@ Section contrib_spec.
{{{ c, RET #c; n c⌝%nat ccounter γ q n }}}. {{{ c, RET #c; n c⌝%nat ccounter γ q n }}}.
Proof. Proof.
iIntros (Φ) "[#? Hγf] HΦ". iIntros (Φ) "[#? Hγf] HΦ".
rewrite /read /=. wp_let. iInv N as (c) ">[Hγ Hl]". wp_load. rewrite /read /=. wp_lam. iInv N as (c) ">[Hγ Hl]". wp_load.
iDestruct (own_valid_2 with "Hγ Hγf") as % ?%frac_auth_included_total%nat_included. iDestruct (own_valid_2 with "Hγ Hγf") as % ?%frac_auth_included_total%nat_included.
iModIntro. iSplitL "Hl Hγ"; [iNext; iExists c; by iFrame|]. iModIntro. iSplitL "Hl Hγ"; [iNext; iExists c; by iFrame|].
iApply ("HΦ" with "[-]"); rewrite /ccounter; eauto 10. iApply ("HΦ" with "[-]"); rewrite /ccounter; eauto 10.
...@@ -155,7 +155,7 @@ Section contrib_spec. ...@@ -155,7 +155,7 @@ Section contrib_spec.
{{{ n, RET #n; ccounter γ 1 n }}}. {{{ n, RET #n; ccounter γ 1 n }}}.
Proof. Proof.
iIntros (Φ) "[#? Hγf] HΦ". iIntros (Φ) "[#? Hγf] HΦ".
rewrite /read /=. wp_let. iInv N as (c) ">[Hγ Hl]". wp_load. rewrite /read /=. wp_lam. iInv N as (c) ">[Hγ Hl]". wp_load.
iDestruct (own_valid_2 with "Hγ Hγf") as % <-%frac_auth_agreeL. iDestruct (own_valid_2 with "Hγ Hγf") as % <-%frac_auth_agreeL.
iModIntro. iSplitL "Hl Hγ"; [iNext; iExists c; by iFrame|]. iModIntro. iSplitL "Hl Hγ"; [iNext; iExists c; by iFrame|].
by iApply "HΦ". by iApply "HΦ".
......
...@@ -23,15 +23,14 @@ Section increment. ...@@ -23,15 +23,14 @@ Section increment.
Lemma incr_spec (l: loc) : Lemma incr_spec (l: loc) :
<<< (v : Z), l #v >>> incr #l @ <<< l #(v + 1), RET #v >>>. <<< (v : Z), l #v >>> incr #l @ <<< l #(v + 1), RET #v >>>.
Proof. Proof.
iApply wp_atomic_intro. iIntros (Φ) "AU". iLöb as "IH". wp_let. iApply wp_atomic_intro. iIntros (Φ) "AU". iLöb as "IH". wp_lam.
wp_apply load_spec; first by iAccu. wp_apply load_spec; first by iAccu.
(* Prove the atomic update for load *) (* Prove the atomic update for load *)
iAuIntro. iApply (aacc_aupd_abort with "AU"); first done. iAuIntro. iApply (aacc_aupd_abort with "AU"); first done.
iIntros (x) "H↦". iAaccIntro with "H↦"; first by eauto with iFrame. iIntros (x) "H↦". iAaccIntro with "H↦"; first by eauto with iFrame.
iIntros "$ !> AU !> _". iIntros "$ !> AU !> _".
(* Now go on *) (* Now go on *)
wp_let. wp_op. wp_bind (CAS _ _ _)%I. wp_let. wp_op. wp_apply cas_spec; [done|iAccu|].
wp_apply cas_spec; [done|iAccu|].
(* Prove the atomic update for CAS *) (* Prove the atomic update for CAS *)
iAuIntro. iApply (aacc_aupd with "AU"); first done. iAuIntro. iApply (aacc_aupd with "AU"); first done.
iIntros (x') "H↦". iAaccIntro with "H↦"; first by eauto with iFrame. iIntros (x') "H↦". iAaccIntro with "H↦"; first by eauto with iFrame.
...@@ -57,7 +56,7 @@ Section increment. ...@@ -57,7 +56,7 @@ Section increment.
weak_incr #l @ weak_incr #l @
<<< v = v' l #(v + 1), RET #v >>>. <<< v = v' l #(v + 1), RET #v >>>.
Proof. Proof.
iIntros "Hl". iApply wp_atomic_intro. iIntros (Φ) "AU". wp_let. iIntros "Hl". iApply wp_atomic_intro. iIntros (Φ) "AU". wp_lam.
wp_apply (atomic_wp_seq $! (load_spec _) with "Hl"). wp_apply (atomic_wp_seq $! (load_spec _) with "Hl").
iIntros "Hl". wp_let. wp_op. iIntros "Hl". wp_let. wp_op.
wp_apply store_spec; first by iAccu. wp_apply store_spec; first by iAccu.
...@@ -86,7 +85,7 @@ Section increment_client. ...@@ -86,7 +85,7 @@ Section increment_client.
Lemma incr_client_safe (x: Z): Lemma incr_client_safe (x: Z):
WP incr_client #x {{ _, True }}%I. WP incr_client #x {{ _, True }}%I.
Proof using Type*. Proof using Type*.
wp_let. wp_alloc l as "Hl". wp_let. wp_lam. wp_alloc l as "Hl". wp_let.
iMod (inv_alloc nroot _ (x':Z, l #x')%I with "[Hl]") as "#Hinv"; first eauto. iMod (inv_alloc nroot _ (x':Z, l #x')%I with "[Hl]") as "#Hinv"; first eauto.
(* FIXME: I am only using persistent stuff, so I should be allowed (* FIXME: I am only using persistent stuff, so I should be allowed
to move this to the persisten context even without the additional □. *) to move this to the persisten context even without the additional □. *)
......
...@@ -6,12 +6,12 @@ Import uPred. ...@@ -6,12 +6,12 @@ Import uPred.
Definition parN : namespace := nroot .@ "par". Definition parN : namespace := nroot .@ "par".
Definition par : val := Definition par : val :=
λ: "fs", λ: "e1" "e2",
let: "handle" := spawn (Fst "fs") in let: "handle" := spawn "e1" in
let: "v2" := Snd "fs" #() in let: "v2" := "e2" #() in
let: "v1" := join "handle" in let: "v1" := join "handle" in
("v1", "v2"). ("v1", "v2").
Notation "e1 ||| e2" := (par (Pair (λ: <>, e1) (λ: <>, e2)))%E : expr_scope. Notation "e1 ||| e2" := (par (λ: <>, e1)%E (λ: <>, e2)%E) : expr_scope.
Section proof. Section proof.
Local Set Default Proof Using "Type*". Local Set Default Proof Using "Type*".
...@@ -21,28 +21,26 @@ Context `{!heapG Σ, !spawnG Σ}. ...@@ -21,28 +21,26 @@ Context `{!heapG Σ, !spawnG Σ}.
brought together. That is strictly stronger than first stripping a later brought together. That is strictly stronger than first stripping a later
and then merging them, as demonstrated by [tests/joining_existentials.v]. and then merging them, as demonstrated by [tests/joining_existentials.v].
This is why these are not Texan triples. *) This is why these are not Texan triples. *)
Lemma par_spec (Ψ1 Ψ2 : val iProp Σ) e (f1 f2 : val) (Φ : val iProp Σ) : Lemma par_spec (Ψ1 Ψ2 : val iProp Σ) (f1 f2 : val) (Φ : val iProp Σ) :
IntoVal e (f1,f2)
WP f1 #() {{ Ψ1 }} -∗ WP f2 #() {{ Ψ2 }} -∗ WP f1 #() {{ Ψ1 }} -∗ WP f2 #() {{ Ψ2 }} -∗
( v1 v2, Ψ1 v1 Ψ2 v2 -∗ Φ (v1,v2)%V) -∗ ( v1 v2, Ψ1 v1 Ψ2 v2 -∗ Φ (v1,v2)%V) -∗
WP par e {{ Φ }}. WP par f1 f2 {{ Φ }}.
Proof. Proof.
iIntros (<-) "Hf1 Hf2 HΦ". iIntros "Hf1 Hf2 HΦ".
rewrite /par /=. wp_let. wp_proj. rewrite /par /=. wp_lam. wp_let.
wp_apply (spawn_spec parN with "Hf1"). wp_apply (spawn_spec parN with "Hf1").
iIntros (l) "Hl". wp_let. wp_proj. wp_bind (f2 _). iIntros (l) "Hl". wp_let. wp_bind (f2 _).
iApply (wp_wand with "Hf2"); iIntros (v) "H2". wp_let. wp_apply (wp_wand with "Hf2"); iIntros (v) "H2". wp_let.
wp_apply (join_spec with "[$Hl]"). iIntros (w) "H1". wp_apply (join_spec with "[$Hl]"). iIntros (w) "H1".
iSpecialize ("HΦ" with "[-]"); first by iSplitL "H1". by wp_let. iSpecialize ("HΦ" with "[-]"); first by iSplitL "H1". wp_let. by wp_pair.
Qed. Qed.
Lemma wp_par (Ψ1 Ψ2 : val iProp Σ) Lemma wp_par (Ψ1 Ψ2 : val iProp Σ) (e1 e2 : expr) (Φ : val iProp Σ) :
(e1 e2 : expr) `{!Closed [] e1, Closed [] e2} (Φ : val iProp Σ) :
WP e1 {{ Ψ1 }} -∗ WP e2 {{ Ψ2 }} -∗ WP e1 {{ Ψ1 }} -∗ WP e2 {{ Ψ2 }} -∗
( v1 v2, Ψ1 v1 Ψ2 v2 -∗ Φ (v1,v2)%V) -∗ ( v1 v2, Ψ1 v1 Ψ2 v2 -∗ Φ (v1,v2)%V) -∗
WP e1 ||| e2 {{ Φ }}. WP e1 ||| e2 {{ Φ }}.
Proof. Proof.
iIntros "H1 H2 H". iApply (par_spec Ψ1 Ψ2 with "[H1] [H2] [H]"). iIntros "H1 H2 H". do 2 wp_closure.
by wp_let. by wp_let. auto. iApply (par_spec Ψ1 Ψ2 with "[H1] [H2] [H]"); [by wp_lam..|auto].
Qed. Qed.
End proof. End proof.
...@@ -49,13 +49,13 @@ Lemma spawn_spec (Ψ : val → iProp Σ) e (f : val) : ...@@ -49,13 +49,13 @@ Lemma spawn_spec (Ψ : val → iProp Σ) e (f : val) :
{{{ WP f #() {{ Ψ }} }}} spawn e {{{ l, RET #l; join_handle l Ψ }}}. {{{ WP f #() {{ Ψ }} }}} spawn e {{{ l, RET #l; join_handle l Ψ }}}.
Proof. Proof.
iIntros (<- Φ) "Hf HΦ". rewrite /spawn /=. iIntros (<- Φ) "Hf HΦ". rewrite /spawn /=.
wp_let. wp_alloc l as "Hl". wp_let. wp_lam. wp_inj. wp_alloc l as "Hl". wp_let.
iMod (own_alloc (Excl ())) as (γ) "Hγ"; first done. iMod (own_alloc (Excl ())) as (γ) "Hγ"; first done.
iMod (inv_alloc N _ (spawn_inv γ l Ψ) with "[Hl]") as "#?". iMod (inv_alloc N _ (spawn_inv γ l Ψ) with "[Hl]") as "#?".
{ iNext. iExists NONEV. iFrame; eauto. } { iNext. iExists NONEV. iFrame; eauto. }
wp_apply (wp_fork with "[Hf]"). wp_apply (wp_fork with "[Hf]").
- iNext. wp_bind (f _). iApply (wp_wand with "Hf"); iIntros (v) "Hv". - iNext. wp_bind (f _). iApply (wp_wand with "Hf"); iIntros (v) "Hv".
iInv N as (v') "[Hl _]". wp_inj. iInv N as (v') "[Hl _]".
wp_store. iSplitL; last done. iIntros "!> !>". iExists (SOMEV v). iFrame. eauto. wp_store. iSplitL; last done. iIntros "!> !>". iExists (SOMEV v). iFrame. eauto.
- wp_seq. iApply "HΦ". rewrite /join_handle. eauto. - wp_seq. iApply "HΦ". rewrite /join_handle. eauto.
Qed. Qed.
......
...@@ -83,7 +83,7 @@ Section proof. ...@@ -83,7 +83,7 @@ Section proof.
Proof. Proof.
iIntros (Φ) "(Hlock & Hlocked & HR) HΦ". iIntros (Φ) "(Hlock & Hlocked & HR) HΦ".
iDestruct "Hlock" as (l ->) "#Hinv". iDestruct "Hlock" as (l ->) "#Hinv".
rewrite /release /=. wp_let. iInv N as (b) "[Hl _]". rewrite /release /=. wp_lam. iInv N as (b) "[Hl _]".
wp_store. iSplitR "HΦ"; last by iApply "HΦ". wp_store. iSplitR "HΦ"; last by iApply "HΦ".
iModIntro. iNext. iExists false. by iFrame. iModIntro. iNext. iExists false. by iFrame.
Qed. Qed.
......
...@@ -73,14 +73,14 @@ Section proof. ...@@ -73,14 +73,14 @@ Section proof.
Lemma newlock_spec (R : iProp Σ) : Lemma newlock_spec (R : iProp Σ) :
{{{ R }}} newlock #() {{{ lk γ, RET lk; is_lock γ lk R }}}. {{{ R }}} newlock #() {{{ lk γ, RET lk; is_lock γ lk R }}}.
Proof. Proof.
iIntros (Φ) "HR HΦ". rewrite -wp_fupd /newlock /=. iIntros (Φ) "HR HΦ". rewrite -wp_fupd /newlock /=. repeat wp_proj.
wp_seq. wp_alloc ln as "Hln". wp_alloc lo as "Hlo". wp_lam. wp_alloc ln as "Hln". wp_alloc lo as "Hlo".
iMod (own_alloc ( (Excl' 0%nat, GSet ) (Excl' 0%nat, GSet ))) as (γ) "[Hγ Hγ']". iMod (own_alloc ( (Excl' 0%nat, GSet ) (Excl' 0%nat, GSet ))) as (γ) "[Hγ Hγ']".
{ by rewrite -auth_both_op. } { by rewrite -auth_both_op. }
iMod (inv_alloc _ _ (lock_inv γ lo ln R) with "[-HΦ]"). iMod (inv_alloc _ _ (lock_inv γ lo ln R) with "[-HΦ]").
{ iNext. rewrite /lock_inv. { iNext. rewrite /lock_inv.
iExists 0%nat, 0%nat. iFrame. iLeft. by iFrame. } iExists 0%nat, 0%nat. iFrame. iLeft. by iFrame. }
iModIntro. iApply ("HΦ" $! (#lo, #ln)%V γ). iExists lo, ln. eauto. wp_pair. iModIntro. iApply ("HΦ" $! (#lo, #ln)%V γ). iExists lo, ln. eauto.
Qed. Qed.
Lemma wait_loop_spec γ lk x R : Lemma wait_loop_spec γ lk x R :
...@@ -137,7 +137,7 @@ Section proof. ...@@ -137,7 +137,7 @@ Section proof.
Proof. Proof.
iIntros (Φ) "(Hl & Hγ & HR) HΦ". iDestruct "Hl" as (lo ln ->) "#Hinv". iIntros (Φ) "(Hl & Hγ & HR) HΦ". iDestruct "Hl" as (lo ln ->) "#Hinv".
iDestruct "Hγ" as (o) "Hγo". iDestruct "Hγ" as (o) "Hγo".
wp_let. wp_proj. wp_bind (! _)%E. wp_lam. wp_proj. wp_bind (! _)%E.
iInv N as (o' n) "(>Hlo & >Hln & >Hauth & Haown)". iInv N as (o' n) "(>Hlo & >Hln & >Hauth & Haown)".
wp_load. wp_load.
iDestruct (own_valid_2 with "Hauth Hγo") as iDestruct (own_valid_2 with "Hauth Hγo") as
......
...@@ -44,7 +44,6 @@ Ltac inv_head_step := ...@@ -44,7 +44,6 @@ Ltac inv_head_step :=
inversion H; subst; clear H inversion H; subst; clear H
end. end.
Local Hint Extern 0 (atomic _ _) => solve_atomic.
Local Hint Extern 0 (head_reducible _ _) => eexists _, _, _, _; simpl. Local Hint Extern 0 (head_reducible _ _) => eexists _, _, _, _; simpl.
Local Hint Extern 0 (head_reducible_no_obs _ _) => eexists _, _, _; simpl. Local Hint Extern 0 (head_reducible_no_obs _ _) => eexists _, _, _; simpl.
...@@ -59,51 +58,61 @@ Local Hint Resolve to_of_val. ...@@ -59,51 +58,61 @@ Local Hint Resolve to_of_val.
Local Ltac solve_exec_safe := intros; subst; do 3 eexists; econstructor; eauto. Local Ltac solve_exec_safe := intros; subst; do 3 eexists; econstructor; eauto.
Local Ltac solve_exec_puredet := simpl; intros; by inv_head_step. Local Ltac solve_exec_puredet := simpl; intros; by inv_head_step.
Local Ltac solve_pure_exec := Local Ltac solve_pure_exec :=
unfold IntoVal in *; subst; intros ?; apply nsteps_once, pure_head_step_pure_step;
repeat match goal with H : AsVal _ |- _ => destruct H as [??] end; subst;
intros ?; apply nsteps_once, pure_head_step_pure_step;
constructor; [solve_exec_safe | solve_exec_puredet]. constructor; [solve_exec_safe | solve_exec_puredet].
Class AsRec (e : expr) (f x : binder) (erec : expr) := Class AsRecV (v : val) (f x : binder) (erec : expr) :=
as_rec : e = Rec f x erec. as_recv : v = RecV f x erec.
Instance AsRec_rec f x e : AsRec (Rec f x e) f x e := eq_refl. Instance AsRecV_recv f x e : AsRecV (RecV f x e) f x e := eq_refl.
Instance AsRec_rec_locked_val v f x e : Instance AsRecV_recv_locked v f x e :
AsRec (of_val v) f x e AsRec (of_val (locked v)) f x e. AsRecV v f x e AsRecV (locked v) f x e.
Proof. by unlock. Qed. Proof. by unlock. Qed.
Instance pure_rec f x (erec e1 e2 : expr) Instance pure_recc f x (erec : expr) :
`{!AsVal e2, AsRec e1 f x erec, Closed (f :b: x :b: []) erec} : PureExec True 1 (Rec f x erec) (Val $ RecV f x erec).
PureExec True 1 (App e1 e2) (subst' x e2 (subst' f e1 erec)). Proof. solve_pure_exec. Qed.
Proof. unfold AsRec in *; solve_pure_exec. Qed. Instance pure_pairc (v1 v2 : val) :
PureExec True 1 (Pair (Val v1) (Val v2)) (Val $ PairV v1 v2).
Proof. solve_pure_exec. Qed.
Instance pure_injlc (v : val) :
PureExec True 1 (InjL $ Val v) (Val $ InjLV v).
Proof. solve_pure_exec. Qed.
Instance pure_injrc (v : val) :
PureExec True 1 (InjR $ Val v) (Val $ InjRV v).
Proof. solve_pure_exec. Qed.
Instance pure_unop op e v v' `{!IntoVal e v} : Instance pure_beta f x (erec : expr) (v1 v2 : val) `{AsRecV v1 f x erec} :
PureExec (un_op_eval op v = Some v') 1 (UnOp op e) (of_val v'). PureExec True 1 (App (Val v1) (Val v2)) (subst' x v2 (subst' f v1 erec)).
Proof. unfold AsRecV in *. solve_pure_exec. Qed.
Instance pure_unop op v v' :
PureExec (un_op_eval op v = Some v') 1 (UnOp op (Val v)) (Val v').
Proof. solve_pure_exec. Qed. Proof. solve_pure_exec. Qed.
Instance pure_binop op e1 e2 v1 v2 v' `{!IntoVal e1 v1, !IntoVal e2 v2} : Instance pure_binop op v1 v2 v' :
PureExec (bin_op_eval op v1 v2 = Some v') 1 (BinOp op e1 e2) (of_val v'). PureExec (bin_op_eval op v1 v2 = Some v') 1 (BinOp op (Val v1) (Val v2)) (Val v').
Proof. solve_pure_exec. Qed. Proof. solve_pure_exec. Qed.
Instance pure_if_true e1 e2 : PureExec True 1 (If (Lit (LitBool true)) e1 e2) e1. Instance pure_if_true e1 e2 : PureExec True 1 (If (Val $ LitV $ LitBool true) e1 e2) e1.
Proof. solve_pure_exec. Qed. Proof. solve_pure_exec. Qed.
Instance pure_if_false e1 e2 : PureExec True 1 (If (Lit (LitBool false)) e1 e2) e2. Instance pure_if_false e1 e2 : PureExec True 1 (If (Val $ LitV $ LitBool false) e1 e2) e2.
Proof. solve_pure_exec. Qed. Proof. solve_pure_exec. Qed.
Instance pure_fst e1 e2 v1 `{!IntoVal e1 v1, !AsVal e2} : Instance pure_fst v1 v2 :
PureExec True 1 (Fst (Pair e1 e2)) e1. PureExec True 1 (Fst (Val $ PairV v1 v2)) (Val v1).
Proof. solve_pure_exec. Qed. Proof. solve_pure_exec. Qed.
Instance pure_snd e1 e2 v2 `{!AsVal e1, !IntoVal e2 v2} : Instance pure_snd v1 v2 :
PureExec True 1 (Snd (Pair e1 e2)) e2. PureExec True 1 (Snd (Val $ PairV v1 v2)) (Val v2).
Proof. solve_pure_exec. Qed. Proof. solve_pure_exec. Qed.
Instance pure_case_inl e0 v e1 e2 `{!IntoVal e0 v} : Instance pure_case_inl v e1 e2 :
PureExec True 1 (Case (InjL e0) e1 e2) (App e1 e0). PureExec True 1 (Case (Val $ InjLV v) e1 e2) (App e1 (Val v)).
Proof. solve_pure_exec. Qed. Proof. solve_pure_exec. Qed.
Instance pure_case_inr e0 v e1 e2 `{!IntoVal e0 v} : Instance pure_case_inr v e1 e2 :
PureExec True 1 (Case (InjR e0) e1 e2) (App e2 e0). PureExec True 1 (Case (Val $ InjRV v) e1 e2) (App e2 (Val v)).
Proof. solve_pure_exec. Qed. Proof. solve_pure_exec. Qed.
Section lifting. Section lifting.
...@@ -131,21 +140,19 @@ Proof. ...@@ -131,21 +140,19 @@ Proof.
Qed. Qed.
(** Heap *) (** Heap *)
Lemma wp_alloc s E e v : Lemma wp_alloc s E v :
IntoVal e v {{{ True }}} Alloc (Val v) @ s; E {{{ l, RET LitV (LitLoc l); l v }}}.
{{{ True }}} Alloc e @ s; E {{{ l, RET LitV (LitLoc l); l v }}}.
Proof. Proof.
iIntros (<- Φ) "_ HΦ". iApply wp_lift_atomic_head_step_no_fork; auto. iIntros (Φ) "_ HΦ". iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs) "[Hσ Hκs] !>"; iSplit; first by eauto. iIntros (σ1 κ κs) "[Hσ Hκs] !>"; iSplit; first by auto.
iNext; iIntros (v2 σ2 efs Hstep); inv_head_step. iNext; iIntros (v2 σ2 efs Hstep); inv_head_step.
iMod (@gen_heap_alloc with "Hσ") as "[Hσ Hl]"; first done. iMod (@gen_heap_alloc with "Hσ") as "[Hσ Hl]"; first done.
iModIntro; iSplit=> //. iFrame. by iApply "HΦ". iModIntro; iSplit=> //. iFrame. by iApply "HΦ".
Qed. Qed.
Lemma twp_alloc s E e v : Lemma twp_alloc s E v :
IntoVal e v [[{ True }]] Alloc (Val v) @ s; E [[{ l, RET LitV (LitLoc l); l v }]].
[[{ True }]] Alloc e @ s; E [[{ l, RET LitV (LitLoc l); l v }]].
Proof. Proof.
iIntros (<- Φ) "_ HΦ". iApply twp_lift_atomic_head_step_no_fork; auto. iIntros (Φ) "_ HΦ". iApply twp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κs) "[Hσ Hκs] !>"; iSplit; first by eauto. iIntros (σ1 κs) "[Hσ Hκs] !>"; iSplit; first by eauto.
iIntros (κ v2 σ2 efs Hstep); inv_head_step. iIntros (κ v2 σ2 efs Hstep); inv_head_step.
iMod (@gen_heap_alloc with "Hσ") as "[Hσ Hl]"; first done. iMod (@gen_heap_alloc with "Hσ") as "[Hσ Hl]"; first done.
...@@ -153,7 +160,7 @@ Proof. ...@@ -153,7 +160,7 @@ Proof.
Qed. Qed.
Lemma wp_load s E l q v : Lemma wp_load s E l q v :
{{{ l {q} v }}} Load (Lit (LitLoc l)) @ s; E {{{ RET v; l {q} v }}}. {{{ l {q} v }}} Load (Val $ LitV $ LitLoc l) @ s; E {{{ RET v; l {q} v }}}.
Proof. Proof.
iIntros (Φ) ">Hl HΦ". iApply wp_lift_atomic_head_step_no_fork; auto. iIntros (Φ) ">Hl HΦ". iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs) "[Hσ Hκs] !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?. iIntros (σ1 κ κs) "[Hσ Hκs] !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?.
...@@ -161,7 +168,7 @@ Proof. ...@@ -161,7 +168,7 @@ Proof.
iModIntro; iSplit=> //. iFrame. by iApply "HΦ". iModIntro; iSplit=> //. iFrame. by iApply "HΦ".
Qed. Qed.
Lemma twp_load s E l q v : Lemma twp_load s E l q v :
[[{ l {q} v }]] Load (Lit (LitLoc l)) @ s; E [[{ RET v; l {q} v }]]. [[{ l {q} v }]] Load (Val $ LitV $ LitLoc l) @ s; E [[{ RET v; l {q} v }]].
Proof. Proof.
iIntros (Φ) "Hl HΦ". iApply twp_lift_atomic_head_step_no_fork; auto. iIntros (Φ) "Hl HΦ". iApply twp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κs) "[Hσ Hκs] !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?. iIntros (σ1 κs) "[Hσ Hκs] !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?.
...@@ -169,22 +176,22 @@ Proof. ...@@ -169,22 +176,22 @@ Proof.
iModIntro; iSplit=> //. iSplit; first done. iFrame. by iApply "HΦ". iModIntro; iSplit=> //. iSplit; first done. iFrame. by iApply "HΦ".
Qed. Qed.
Lemma wp_store s E l v' e v : Lemma wp_store s E l v' v :
IntoVal e v {{{ l v' }}} Store (Val $ LitV (LitLoc l)) (Val v) @ s; E
{{{ l v' }}} Store (Lit (LitLoc l)) e @ s; E {{{ RET LitV LitUnit; l v }}}. {{{ RET LitV LitUnit; l v }}}.
Proof. Proof.
iIntros (<- Φ) ">Hl HΦ". iIntros (Φ) ">Hl HΦ".
iApply wp_lift_atomic_head_step_no_fork; auto. iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs) "[Hσ Hκs] !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?. iIntros (σ1 κ κs) "[Hσ Hκs] !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?.
iSplit; first by eauto. iNext; iIntros (v2 σ2 efs Hstep); inv_head_step. iSplit; first by eauto. iNext; iIntros (v2 σ2 efs Hstep); inv_head_step.
iMod (@gen_heap_update with "Hσ Hl") as "[$ Hl]". iMod (@gen_heap_update with "Hσ Hl") as "[$ Hl]".
iModIntro. iSplit=>//. iFrame. by iApply "HΦ". iModIntro. iSplit=>//. iFrame. by iApply "HΦ".
Qed. Qed.
Lemma twp_store s E l v' e v : Lemma twp_store s E l v' v :
IntoVal e v [[{ l v' }]] Store (Val $ LitV $ LitLoc l) (Val v) @ s; E
[[{ l v' }]] Store (Lit (LitLoc l)) e @ s; E [[{ RET LitV LitUnit; l v }]]. [[{ RET LitV LitUnit; l v }]].
Proof. Proof.
iIntros (<- Φ) "Hl HΦ". iIntros (Φ) "Hl HΦ".
iApply twp_lift_atomic_head_step_no_fork; auto. iApply twp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κs) "[Hσ Hκs] !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?. iIntros (σ1 κs) "[Hσ Hκs] !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?.
iSplit; first by eauto. iIntros (κ v2 σ2 efs Hstep); inv_head_step. iSplit; first by eauto. iIntros (κ v2 σ2 efs Hstep); inv_head_step.
...@@ -192,73 +199,65 @@ Proof. ...@@ -192,73 +199,65 @@ Proof.
iModIntro. iSplit=>//. iSplit; first done. iFrame. by iApply "HΦ". iModIntro. iSplit=>//. iSplit; first done. iFrame. by iApply "HΦ".
Qed. Qed.
Lemma wp_cas_fail s E l q v' e1 v1 e2 : Lemma wp_cas_fail s E l q v' v1 v2 :
IntoVal e1 v1 AsVal e2 v' v1 vals_cas_compare_safe v' v1 v' v1 vals_cas_compare_safe v' v1
{{{ l {q} v' }}} CAS (Lit (LitLoc l)) e1 e2 @ s; E {{{ l {q} v' }}} CAS (Val $ LitV $ LitLoc l) (Val v1) (Val v2) @ s; E
{{{ RET LitV (LitBool false); l {q} v' }}}. {{{ RET LitV (LitBool false); l {q} v' }}}.
Proof. Proof.
iIntros (<- [v2 <-] ?? Φ) ">Hl HΦ". iIntros (?? Φ) ">Hl HΦ". iApply wp_lift_atomic_head_step_no_fork; auto.
iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs) "[Hσ Hκs] !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?. iIntros (σ1 κ κs) "[Hσ Hκs] !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?.
iSplit; first by eauto. iNext; iIntros (v2' σ2 efs Hstep); inv_head_step. iSplit; first by eauto. iNext; iIntros (v2' σ2 efs Hstep); inv_head_step.
iModIntro; iSplit=> //. iFrame. by iApply "HΦ". iModIntro; iSplit=> //. iFrame. by iApply "HΦ".
Qed. Qed.
Lemma twp_cas_fail s E l q v' e1 v1 e2 : Lemma twp_cas_fail s E l q v' v1 v2 :
IntoVal e1 v1 AsVal e2 v' v1 vals_cas_compare_safe v' v1 v' v1 vals_cas_compare_safe v' v1
[[{ l {q} v' }]] CAS (Lit (LitLoc l)) e1 e2 @ s; E [[{ l {q} v' }]] CAS (Val $ LitV $ LitLoc l) (Val v1) (Val v2) @ s; E
[[{ RET LitV (LitBool false); l {q} v' }]]. [[{ RET LitV (LitBool false); l {q} v' }]].
Proof. Proof.
iIntros (<- [v2 <-] ?? Φ) "Hl HΦ". iIntros (?? Φ) "Hl HΦ". iApply twp_lift_atomic_head_step_no_fork; auto.
iApply twp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κs) "[Hσ Hκs] !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?. iIntros (σ1 κs) "[Hσ Hκs] !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?.
iSplit; first by eauto. iIntros (κ v2' σ2 efs Hstep); inv_head_step. iSplit; first by eauto. iIntros (κ v2' σ2 efs Hstep); inv_head_step.
iModIntro; iSplit=> //. iSplit; first done. iFrame. by iApply "HΦ". iModIntro; iSplit=> //. iSplit; first done. iFrame. by iApply "HΦ".
Qed. Qed.
Lemma wp_cas_suc s E l e1 v1 e2 v2 : Lemma wp_cas_suc s E l v1 v2 :
IntoVal e1 v1 IntoVal e2 v2 vals_cas_compare_safe v1 v1 vals_cas_compare_safe v1 v1
{{{ l v1 }}} CAS (Lit (LitLoc l)) e1 e2 @ s; E {{{ l v1 }}} CAS (Val $ LitV $ LitLoc l) (Val v1) (Val v2) @ s; E
{{{ RET LitV (LitBool true); l v2 }}}. {{{ RET LitV (LitBool true); l v2 }}}.
Proof. Proof.
iIntros (<- <- ? Φ) ">Hl HΦ". iIntros (? Φ) ">Hl HΦ". iApply wp_lift_atomic_head_step_no_fork; auto.
iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs) "[Hσ Hκs] !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?. iIntros (σ1 κ κs) "[Hσ Hκs] !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?.
iSplit; first by eauto. iNext; iIntros (v2' σ2 efs Hstep); inv_head_step. iSplit; first by eauto. iNext; iIntros (v2' σ2 efs Hstep); inv_head_step.
iMod (@gen_heap_update with "Hσ Hl") as "[$ Hl]". iMod (@gen_heap_update with "Hσ Hl") as "[$ Hl]".
iModIntro. iSplit=>//. iFrame. by iApply "HΦ". iModIntro. iSplit=>//. iFrame. by iApply "HΦ".
Qed. Qed.
Lemma twp_cas_suc s E l e1 v1 e2 v2 : Lemma twp_cas_suc s E l v1 v2 :
IntoVal e1 v1 IntoVal e2 v2 vals_cas_compare_safe v1 v1 vals_cas_compare_safe v1 v1
[[{ l v1 }]] CAS (Lit (LitLoc l)) e1 e2 @ s; E [[{ l v1 }]] CAS (Val $ LitV $ LitLoc l) (Val v1) (Val v2) @ s; E
[[{ RET LitV (LitBool true); l v2 }]]. [[{ RET LitV (LitBool true); l v2 }]].
Proof. Proof.
iIntros (<- <- ? Φ) "Hl HΦ". iIntros (? Φ) "Hl HΦ". iApply twp_lift_atomic_head_step_no_fork; auto.
iApply twp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κs) "[Hσ Hκs] !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?. iIntros (σ1 κs) "[Hσ Hκs] !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?.
iSplit; first by eauto. iIntros (κ v2' σ2 efs Hstep); inv_head_step. iSplit; first by eauto. iIntros (κ v2' σ2 efs Hstep); inv_head_step.
iMod (@gen_heap_update with "Hσ Hl") as "[$ Hl]". iMod (@gen_heap_update with "Hσ Hl") as "[$ Hl]".
iModIntro. iSplit=>//. iSplit; first done. iFrame. by iApply "HΦ". iModIntro. iSplit=>//. iSplit; first done. iFrame. by iApply "HΦ".
Qed. Qed.
Lemma wp_faa s E l i1 e2 i2 : Lemma wp_faa s E l i1 i2 :
IntoVal e2 (LitV (LitInt i2)) {{{ l LitV (LitInt i1) }}} FAA (Val $ LitV $ LitLoc l) (Val $ LitV $ LitInt i2) @ s; E
{{{ l LitV (LitInt i1) }}} FAA (Lit (LitLoc l)) e2 @ s; E
{{{ RET LitV (LitInt i1); l LitV (LitInt (i1 + i2)) }}}. {{{ RET LitV (LitInt i1); l LitV (LitInt (i1 + i2)) }}}.
Proof. Proof.
iIntros (<- Φ) ">Hl HΦ". iIntros (Φ) ">Hl HΦ". iApply wp_lift_atomic_head_step_no_fork; auto.
iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs) "[Hσ Hκs] !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?. iIntros (σ1 κ κs) "[Hσ Hκs] !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?.
iSplit; first by eauto. iNext; iIntros (v2' σ2 efs Hstep); inv_head_step. iSplit; first by eauto. iNext; iIntros (v2' σ2 efs Hstep); inv_head_step.
iMod (@gen_heap_update with "Hσ Hl") as "[$ Hl]". iMod (@gen_heap_update with "Hσ Hl") as "[$ Hl]".
iModIntro. iSplit=>//. iFrame. by iApply "HΦ". iModIntro. iSplit=>//. iFrame. by iApply "HΦ".
Qed. Qed.
Lemma twp_faa s E l i1 e2 i2 : Lemma twp_faa s E l i1 i2 :
IntoVal e2 (LitV (LitInt i2)) [[{ l LitV (LitInt i1) }]] FAA (Val $ LitV $ LitLoc l) (Val $ LitV $ LitInt i2) @ s; E
[[{ l LitV (LitInt i1) }]] FAA (Lit (LitLoc l)) e2 @ s; E
[[{ RET LitV (LitInt i1); l LitV (LitInt (i1 + i2)) }]]. [[{ RET LitV (LitInt i1); l LitV (LitInt (i1 + i2)) }]].
Proof. Proof.
iIntros (<- Φ) "Hl HΦ". iIntros (Φ) "Hl HΦ". iApply twp_lift_atomic_head_step_no_fork; auto.
iApply twp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κs) "[Hσ Hκs] !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?. iIntros (σ1 κs) "[Hσ Hκs] !>". iDestruct (@gen_heap_valid with "Hσ Hl") as %?.
iSplit; first by eauto. iIntros (κ e2 σ2 efs Hstep); inv_head_step. iSplit; first by eauto. iIntros (κ e2 σ2 efs Hstep); inv_head_step.
iMod (@gen_heap_update with "Hσ Hl") as "[$ Hl]". iMod (@gen_heap_update with "Hσ Hl") as "[$ Hl]".
......
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment