diff --git a/theories/coneris/examples/two_die.v b/theories/coneris/examples/two_die.v new file mode 100644 index 000000000..bebec419e --- /dev/null +++ b/theories/coneris/examples/two_die.v @@ -0,0 +1,151 @@ +From iris.algebra Require Import excl_auth. +From iris.base_logic.lib Require Import invariants. +From clutch.coneris Require Import coneris par spawn. + +Local Open Scope Z. + +Set Default Proof Using "Type*". + +Section lemmas. + Context `{!inG Σ (excl_authR (option natO))}. + + (* Helpful lemmas *) + Lemma ghost_var_alloc b : + ⊢ |==> ∃ γ, own γ (●E b) ∗ own γ (◯E b). + Proof. + iMod (own_alloc (●E b ⋅ ◯E b)) as (γ) "[??]". + - by apply excl_auth_valid. + - by eauto with iFrame. + Qed. + + Lemma ghost_var_agree γ b c : + own γ (●E b) -∗ own γ (◯E c) -∗ ⌜ b = c ⌝. + Proof. + iIntros "Hγ● Hγ◯". + by iCombine "Hγ● Hγ◯" gives %->%excl_auth_agree_L. + Qed. + + Lemma ghost_var_update γ b' b c : + own γ (●E b) -∗ own γ (◯E c) ==∗ own γ (●E b') ∗ own γ (◯E b'). + Proof. + iIntros "Hγ● Hγ◯". + iMod (own_update_2 _ _ _ (●E b' ⋅ ◯E b') with "Hγ● Hγ◯") as "[$$]". + { by apply excl_auth_update. } + done. + Qed. +End lemmas. + +Definition two_die_prog : expr := + let, ("v1", "v2") := ((rand #5) ||| rand #5) in + "v1" + "v2". + +Section simple. + Context `{!conerisGS Σ, !spawnG Σ}. + Lemma simple_parallel_add_spec: + {{{ ↯ (1/3) }}} + two_die_prog + {{{ (n:nat), RET #n; ⌜(0&%)(%&->&%)]". + iNext. + wp_pures. + iModIntro. + rewrite -Nat2Z.inj_add. iApply "HΦ". + iPureIntro. lia. + Qed. +End simple. + +Section complex. + Context `{!conerisGS Σ, !spawnG Σ, !inG Σ (excl_authR (option natO))}. + + Definition parallel_add_inv (γ1 γ2 : gname) : iProp Σ := + ∃ (n1 n2 : option nat) , + own γ1 (●E n1) ∗ own γ2 (●E n2) ∗ + if bool_decide ((∃ x, n1 = Some x /\ (0(%&%&Hauth1&Hauth2&Herr)" "Hclose". + admit. + + admit. + + iIntros (??) "[(%n&->&Hfrag1) (%n'&->&Hfrag2)]". + iNext. + wp_pures. + iInv "I" as ">(%&%&Hauth1&Hauth2&Herr)" "Hclose". + iDestruct (ghost_var_agree with "[$Hauth1][$]") as "->". + iDestruct (ghost_var_agree with "[$Hauth2][$]") as "->". + iAssert (⌜0)". + rewrite !bool_decide_eq_true_2; last done. + simpl. + replace (_+_-_)%R with 0%R; last lra. + rewrite Rpower_O; last lra. + iDestruct (ec_contradict with "[$]") as "[]". + simpl. lra. + - intros [(?&?&?)|(?&?&?)]; simplify_eq; lia. + } + all: iPureIntro; lia. + } + iMod ("Hclose" with "[$Herr $Hauth1 $Hauth2]"). + rewrite -Nat2Z.inj_add. iApply "HΦ". + iPureIntro. lia. + Admitted. +End complex. + + + + +