diff --git a/theories/channel.v b/theories/channel.v index 33f0f48d8ed1e74cb096040136411d85b0d22040..f30b49d90628ebe298bf92455e929f82b9e2c30a 100644 --- a/theories/channel.v +++ b/theories/channel.v @@ -20,6 +20,7 @@ Definition new_chan : val := Coercion side_to_bool (s : side) : bool := match s with Left => true | Right => false end. +Arguments side_to_bool : simpl never. Definition side_elim {A} (s : side) (l r : A) : A := match s with Left => l | Right => r end. diff --git a/theories/examples.v b/theories/examples.v new file mode 100644 index 0000000000000000000000000000000000000000..c6519796f1937079c163dd76623479432c87ee35 --- /dev/null +++ b/theories/examples.v @@ -0,0 +1,106 @@ +From iris.proofmode Require Import tactics. +From iris.program_logic Require Export weakestpre. +From iris.heap_lang Require Import proofmode notation. +From osiris Require Import typing channel logrel. +From iris.algebra Require Import list auth excl. +From iris.base_logic Require Import invariants. + +Section Examples. + Context `{!heapG Σ} (N : namespace). + Context `{!logrelG Σ}. + + Definition seq_example : expr := + (let: "c" := new_chan #() in + send "c" #Left #5;; + recv "c" #Right)%E. + + Notation "⟦ c @ s : sÏ„ ⟧{ γ }" := (interp_st N γ sÏ„ c s) + (at level 10, s at next level, sÏ„ at next level, γ at next level, + format "⟦ c @ s : sÏ„ ⟧{ γ }"). + + Lemma seq_proof : + {{{ True }}} seq_example {{{ v, RET v; ⌜v = #5⌠}}}. + Proof. + iIntros (Φ H) "HΦ". + rewrite /seq_example. + wp_bind (new_chan _). + wp_apply (new_chan_st_spec N (TSR Send (λ v, ⌜v = #5âŒ%I) (λ v, TEnd)))=> //. + iIntros (c γ) "[Hstl Hstr]"=> /=. + wp_pures. + wp_bind (send _ _ _). + wp_apply (send_st_spec N with "[Hstl]"). by iFrame. + iIntros "Hstl". + wp_pures. + wp_apply (recv_st_spec _ with "[Hstr]"). by iFrame. + iIntros (v) "[Hstr HP]". + by iApply "HΦ". + Qed. + + Definition par_example : expr := + (let: "c" := new_chan #() in + Fork(send "c" #Left #5);; + recv "c" #Right)%E. + + Lemma par_proof : + {{{ True }}} par_example {{{ v, RET v; ⌜v = #5⌠}}}. + Proof. + iIntros (Φ H) "HΦ". + rewrite /par_example. + wp_bind (new_chan _). + wp_apply (new_chan_st_spec N (TSR Send (λ v, ⌜v = #5âŒ%I) (λ v, TEnd)))=> //. + iIntros (c γ) "[Hstl Hstr]"=> /=. + wp_pures. + wp_bind (Fork _). + wp_apply (wp_fork with "[Hstl]"). + - iNext. wp_apply (send_st_spec N with "[Hstl]"). by iFrame. eauto. + - wp_apply (recv_st_spec _ with "[Hstr]"). by iFrame. + iIntros (v) "[Hstr HP]". by iApply "HΦ". + Qed. + + Definition par_2_example : expr := + (let: "c" := new_chan #() in + Fork(let: "v" := recv "c" #Right in send "c" #Right ("v"+#2));; + send "c" #Left #5;;recv "c" #Left)%E. + + Lemma par_2_proof : + {{{ True }}} + par_2_example + {{{ v, RET v; ∃ w, ⌜v = LitV $ LitInt $ w ∧ w >= 7⌠}}}. + Proof. + iIntros (Φ H) "HΦ". + rewrite /par_2_example. + wp_bind (new_chan _). + wp_apply (new_chan_st_spec N + (TSR Send + (λ v, ⌜v = #5âŒ%I) + (λ v, TSR Receive + (λ v', + (∃ w, ⌜v = LitV $ LitInt $ w⌠∧ + ∃ w', ⌜v' = LitV $ LitInt $ w' ∧ w' >= w+2âŒ)%I) + (λ v', TEnd))))=> //. + iIntros (c γ) "[Hstl Hstr]"=> /=. + wp_pures. + wp_bind (Fork _). + wp_apply (wp_fork with "[Hstr]"). + - iNext. + wp_apply (recv_st_spec _ with "[Hstr]"). by iFrame. + iIntros (v) "[Hstr HP]". + wp_pures. + iDestruct "HP" as %->. + wp_apply (send_st_spec N with "[Hstr]"). iFrame. + iExists _. iSplit=> //. + iExists _. iSplit=> //. + eauto. + - wp_apply (send_st_spec _ with "[Hstl]"). by iFrame. + iIntros "Hstl". + wp_apply (recv_st_spec _ with "[Hstl]"). by iFrame. + iIntros (v) "[Hstl HP]". + iApply "HΦ". + iDestruct "HP" as %[w [HP [w' [HQ HQ']]]]. + iExists _. + iSplit=> //. iPureIntro. + inversion HP. rewrite -H1 in HQ'. + eauto. + Qed. + +End Examples. \ No newline at end of file diff --git a/theories/logrel.v b/theories/logrel.v index c2b55f5d580f3d5b042fe667e23c18777c3e045a..5017f5b407de34c401470f01374fe8899f351cf5 100644 --- a/theories/logrel.v +++ b/theories/logrel.v @@ -18,7 +18,7 @@ Instance subG_chanΣ {Σ} : subG logrelΣ Σ → logrelG Σ. Proof. intros [??%subG_auth_exclG]%subG_inv. constructor; apply _. Qed. Section logrel. - Context `{!heapG Σ, !lockG Σ} (N : namespace). + Context `{!heapG Σ} (N : namespace). Context `{!logrelG Σ}. Notation stype_iprop := (@stypeC (laterC (iProp Σ))).