Skip to content
Snippets Groups Projects
Commit fd89aa52 authored by Robbert Krebbers's avatar Robbert Krebbers
Browse files

Invariants over states in WP and get rid of heap_ctx.

The WP construction now takes an invariant on states as a parameter
(part of the irisG class) and no longer builds in the authoritative
ownership of the entire state. When instantiating WP with a concrete
language on can choose its state invariant. For example, for heap_lang
we directly use `auth (gmap loc (frac * dec_agree val))`, and avoid
the indirection through invariants entirely.

As a result, we no longer have to carry `heap_ctx` around.
parent 617a69b4
No related branches found
No related tags found
No related merge requests found
......@@ -26,11 +26,11 @@ Definition rev : val :=
end.
Lemma rev_acc_wp hd acc xs ys (Φ : val iProp Σ) :
heap_ctx is_list hd xs is_list acc ys
is_list hd xs is_list acc ys
( w, is_list w (reverse xs ++ ys) -∗ Φ w)
WP rev hd acc {{ Φ }}.
Proof.
iIntros "(#Hh & Hxs & Hys & HΦ)".
iIntros "(Hxs & Hys & HΦ)".
iLöb as "IH" forall (hd acc xs ys Φ). wp_rec. wp_let.
destruct xs as [|x xs]; iSimplifyEq.
- wp_match. by iApply "HΦ".
......@@ -42,11 +42,11 @@ Proof.
Qed.
Lemma rev_wp hd xs (Φ : val iProp Σ) :
heap_ctx is_list hd xs ( w, is_list w (reverse xs) -∗ Φ w)
is_list hd xs ( w, is_list w (reverse xs) -∗ Φ w)
WP rev hd (InjL #()) {{ Φ }}.
Proof.
iIntros "(#Hh & Hxs & )".
iApply (rev_acc_wp hd NONEV xs [] with "[- $Hh $Hxs]").
iIntros "[Hxs HΦ]".
iApply (rev_acc_wp hd NONEV xs [] with "[- $Hxs]").
iSplit; first done. iIntros (w). rewrite right_id_L. iApply "HΦ".
Qed.
End list_reverse.
......@@ -29,18 +29,18 @@ Instance subG_one_shotΣ {Σ} : subG one_shotΣ Σ → one_shotG Σ.
Proof. intros [?%subG_inG _]%subG_inv. split; apply _. Qed.
Section proof.
Context `{!heapG Σ, !one_shotG Σ} (N : namespace) (HN : heapN N).
Context `{!heapG Σ, !one_shotG Σ} (N : namespace).
Definition one_shot_inv (γ : gname) (l : loc) : iProp Σ :=
(l NONEV own γ Pending n : Z, l SOMEV #n own γ (Shot n))%I.
Lemma wp_one_shot (Φ : val iProp Σ) :
heap_ctx ( f1 f2 : val,
( f1 f2 : val,
( n : Z, WP f1 #n {{ w, w = #true w = #false }})
WP f2 #() {{ g, WP g #() {{ _, True }} }} -∗ Φ (f1,f2)%V)
WP one_shot_example #() {{ Φ }}.
Proof.
iIntros "[#? Hf] /=".
iIntros "Hf /=".
rewrite -wp_fupd /one_shot_example /=. wp_seq. wp_alloc l as "Hl". wp_let.
iMod (own_alloc Pending) as (γ) "Hγ"; first done.
iMod (inv_alloc N _ (one_shot_inv γ l) with "[Hl Hγ]") as "#HN".
......@@ -83,14 +83,13 @@ Proof.
Qed.
Lemma ht_one_shot (Φ : val iProp Σ) :
heap_ctx {{ True }} one_shot_example #()
{{ True }} one_shot_example #()
{{ ff,
( n : Z, {{ True }} Fst ff #n {{ w, w = #true w = #false }})
{{ True }} Snd ff #() {{ g, {{ True }} g #() {{ _, True }} }}
}}.
Proof.
iIntros "#? !# _". iApply wp_one_shot. iSplit; first done.
iIntros (f1 f2) "[#Hf1 #Hf2]"; iSplit.
iIntros "!# _". iApply wp_one_shot. iIntros (f1 f2) "[#Hf1 #Hf2]"; iSplit.
- iIntros (n) "!# _". wp_proj. iApply "Hf1".
- iIntros "!# _". wp_proj.
iApply (wp_wand with "Hf2"). by iIntros (v) "#? !# _".
......
......@@ -34,11 +34,10 @@ Definition sum' : val := λ: "t",
!"l".
Lemma sum_loop_wp `{!heapG Σ} v t l (n : Z) (Φ : val iProp Σ) :
heap_ctx l #n is_tree v t
(l #(sum t + n) -∗ is_tree v t -∗ Φ #())
l #n is_tree v t (l #(sum t + n) -∗ is_tree v t -∗ Φ #())
WP sum_loop v #l {{ Φ }}.
Proof.
iIntros "(#Hh & Hl & Ht & HΦ)".
iIntros "(Hl & Ht & HΦ)".
iLöb as "IH" forall (v t l n Φ). wp_rec. wp_let.
destruct t as [n'|tl tr]; simpl in *.
- iDestruct "Ht" as "%"; subst.
......@@ -55,12 +54,11 @@ Proof.
Qed.
Lemma sum_wp `{!heapG Σ} v t Φ :
heap_ctx is_tree v t (is_tree v t -∗ Φ #(sum t))
WP sum' v {{ Φ }}.
is_tree v t (is_tree v t -∗ Φ #(sum t)) WP sum' v {{ Φ }}.
Proof.
iIntros "(#Hh & Ht & )". rewrite /sum' /=.
iIntros "[Ht HΦ]". rewrite /sum' /=.
wp_let. wp_alloc l as "Hl". wp_let.
wp_apply (sum_loop_wp with "[- $Hh $Ht $Hl]").
wp_apply (sum_loop_wp with "[- $Ht $Hl]").
rewrite Z.add_0_r.
iIntros "Hl Ht". wp_seq. wp_load. by iApply "HΦ".
Qed.
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