aneris.examples.ping_pong_done.heap_lang_threads
From iris.program_logic Require Import hoare weakestpre.
From iris.base_logic Require Export upred.
From iris.base_logic.lib Require Export invariants.
From iris.algebra Require Import cmra agree frac .
From iris.heap_lang Require Import proofmode lang par assert notation.
(* Strings are not a part of heap_lang - we use integers instead *)
Notation MSG_PING := #0%V.
Notation MSG_PONG := #1%V.
Notation MSG_DONE := #2%V.
Section code.
Definition listen_wait : val :=
rec: "loop" "ℓ" :=
let: "m" := !"ℓ" in
match: "m" with
SOME "msg" =>
if: CAS "ℓ" "m" NONE then "msg" else "loop" "ℓ"
| NONE => "loop" "ℓ"
end.
Definition send : val :=
rec: "loop" "ℓ" "v" :=
if: CAS "ℓ" NONE (SOME "v") then #() else "loop" "ℓ" "v".
Definition pong : val := λ: "ℓ_recv" "ℓ_send",
let: "msg" := listen_wait "ℓ_recv" in
(if: "msg" = MSG_PING then send "ℓ_send" MSG_PONG else assert #false);;
let: "ack" := listen_wait "ℓ_recv" in
(if: "ack" = MSG_DONE then #() else assert #false).
Definition ping : val := λ: "ℓ_recv" "ℓ_send",
send "ℓ_send" MSG_PING;;
let: "msg" := listen_wait "ℓ_recv" in
if: "msg" = MSG_PONG then send "ℓ_send" MSG_DONE else assert #false.
Definition ping_pong_runner : expr :=
let: "ℓ_pongrecv" := ref NONE in
let: "ℓ_pingrecv" := ref NONE in
(pong "ℓ_pongrecv" "ℓ_pingrecv") ||| (ping "ℓ_pingrecv" "ℓ_pongrecv").
End code.
Inductive msg_received := MSG_NONE | PING | PONG | DONE.
Definition ppCmra : cmraT := (prodR fracR (agreeR (leibnizO msg_received))).
Class ppdG Σ := PpdG {
pongG_inG :> inG Σ ppCmra;
pingG_inG :> inG Σ ppCmra
}.
Definition pingpongΣ : gFunctors := #[GFunctor ppCmra].
Instance subG_ppdΣ {Σ} : subG pingpongΣ Σ → ppdG Σ.
Proof. solve_inG. Qed.
Class pingpongG Σ := MppdG {
pong_inG :> inG Σ ppCmra;
ping_inG :> inG Σ ppCmra;
pong_name : gname;
ping_name : gname
}.
Section ghost_state.
Context `{!pingpongG Σ}.
(* Corresponds to a view of what the last message pong received is *)
Definition ownPong (q : Qp) (s : msg_received) :=
own (A := ppCmra) pong_name (q, to_agree s).
(* Corresponds to a view of what the last message ping received is *)
Definition ownPing (q : Qp) (s : msg_received) :=
own (A := ppCmra) ping_name (q, to_agree s).
Lemma ownPong_eq p q s1 s2:
ownPong p s1 -∗ ownPong q s2 -∗ ⌜s1 = s2⌝.
Proof.
iIntros "H1 H2". rewrite /ownPong.
iCombine "H1" "H2" as "H".
iDestruct (own_valid with "H") as %HValid.
destruct HValid as [_ H2]; simpl in H2;
iIntros "!%". by apply (agree_op_invL' (A:=leibnizO msg_received)).
Qed.
Lemma ownPing_eq p q s1 s2:
ownPing p s1 -∗ ownPing q s2 -∗ ⌜s1 = s2⌝.
Proof.
iIntros "H1 H2". rewrite /ownPing.
iCombine "H1" "H2" as "H".
iDestruct (own_valid with "H") as %HValid.
destruct HValid as [_ H2]; simpl in H2;
iIntros "!%". by apply (agree_op_invL' (A:=leibnizO msg_received)).
Qed.
Lemma ownPong_update s1 s2:
ownPong (1/2) s1 -∗ ownPong (1/2) s1 ==∗ ownPong (1/2) s2 ∗ ownPong (1/2) s2.
Proof.
iIntros "H1 H2". rewrite /ownPong. iCombine "H1 H2" as "H".
rewrite -own_op -pair_op frac_op' Qp_half_half agree_idemp.
iApply (own_update with "H"). by apply cmra_update_exclusive.
Qed.
Lemma ownPing_update s1 s2:
ownPing (1/2) s1 -∗ ownPing (1/2) s1 ==∗ ownPing (1/2) s2 ∗ ownPing (1/2) s2.
Proof.
iIntros "H1 H2". rewrite /ownPing. iCombine "H1 H2" as "H".
rewrite -own_op -pair_op frac_op' Qp_half_half agree_idemp.
iApply (own_update with "H"). by apply cmra_update_exclusive.
Qed.
End ghost_state.
Section proof.
Context `{!heapG Σ, !spawnG Σ, !pingpongG Σ}.
Context (N : namespace).
Definition iProp := iProp Σ.
Definition pong_inv (ℓ : loc) : iProp :=
(ℓ ↦ NONEV ∨
ℓ ↦ SOMEV MSG_PING ∗ ownPong (1/2) MSG_NONE ∨
ℓ ↦ SOMEV MSG_DONE ∗ ownPong (1/2) PING ∗ ownPing (1/2) PONG)%I.
Definition ping_inv (ℓ : loc) : iProp :=
(ℓ ↦ NONEV ∨
ℓ ↦ SOMEV MSG_PONG ∗ ownPong (1/2) PING ∗ ownPing (1/2) MSG_NONE)%I.
From iris.base_logic Require Export upred.
From iris.base_logic.lib Require Export invariants.
From iris.algebra Require Import cmra agree frac .
From iris.heap_lang Require Import proofmode lang par assert notation.
(* Strings are not a part of heap_lang - we use integers instead *)
Notation MSG_PING := #0%V.
Notation MSG_PONG := #1%V.
Notation MSG_DONE := #2%V.
Section code.
Definition listen_wait : val :=
rec: "loop" "ℓ" :=
let: "m" := !"ℓ" in
match: "m" with
SOME "msg" =>
if: CAS "ℓ" "m" NONE then "msg" else "loop" "ℓ"
| NONE => "loop" "ℓ"
end.
Definition send : val :=
rec: "loop" "ℓ" "v" :=
if: CAS "ℓ" NONE (SOME "v") then #() else "loop" "ℓ" "v".
Definition pong : val := λ: "ℓ_recv" "ℓ_send",
let: "msg" := listen_wait "ℓ_recv" in
(if: "msg" = MSG_PING then send "ℓ_send" MSG_PONG else assert #false);;
let: "ack" := listen_wait "ℓ_recv" in
(if: "ack" = MSG_DONE then #() else assert #false).
Definition ping : val := λ: "ℓ_recv" "ℓ_send",
send "ℓ_send" MSG_PING;;
let: "msg" := listen_wait "ℓ_recv" in
if: "msg" = MSG_PONG then send "ℓ_send" MSG_DONE else assert #false.
Definition ping_pong_runner : expr :=
let: "ℓ_pongrecv" := ref NONE in
let: "ℓ_pingrecv" := ref NONE in
(pong "ℓ_pongrecv" "ℓ_pingrecv") ||| (ping "ℓ_pingrecv" "ℓ_pongrecv").
End code.
Inductive msg_received := MSG_NONE | PING | PONG | DONE.
Definition ppCmra : cmraT := (prodR fracR (agreeR (leibnizO msg_received))).
Class ppdG Σ := PpdG {
pongG_inG :> inG Σ ppCmra;
pingG_inG :> inG Σ ppCmra
}.
Definition pingpongΣ : gFunctors := #[GFunctor ppCmra].
Instance subG_ppdΣ {Σ} : subG pingpongΣ Σ → ppdG Σ.
Proof. solve_inG. Qed.
Class pingpongG Σ := MppdG {
pong_inG :> inG Σ ppCmra;
ping_inG :> inG Σ ppCmra;
pong_name : gname;
ping_name : gname
}.
Section ghost_state.
Context `{!pingpongG Σ}.
(* Corresponds to a view of what the last message pong received is *)
Definition ownPong (q : Qp) (s : msg_received) :=
own (A := ppCmra) pong_name (q, to_agree s).
(* Corresponds to a view of what the last message ping received is *)
Definition ownPing (q : Qp) (s : msg_received) :=
own (A := ppCmra) ping_name (q, to_agree s).
Lemma ownPong_eq p q s1 s2:
ownPong p s1 -∗ ownPong q s2 -∗ ⌜s1 = s2⌝.
Proof.
iIntros "H1 H2". rewrite /ownPong.
iCombine "H1" "H2" as "H".
iDestruct (own_valid with "H") as %HValid.
destruct HValid as [_ H2]; simpl in H2;
iIntros "!%". by apply (agree_op_invL' (A:=leibnizO msg_received)).
Qed.
Lemma ownPing_eq p q s1 s2:
ownPing p s1 -∗ ownPing q s2 -∗ ⌜s1 = s2⌝.
Proof.
iIntros "H1 H2". rewrite /ownPing.
iCombine "H1" "H2" as "H".
iDestruct (own_valid with "H") as %HValid.
destruct HValid as [_ H2]; simpl in H2;
iIntros "!%". by apply (agree_op_invL' (A:=leibnizO msg_received)).
Qed.
Lemma ownPong_update s1 s2:
ownPong (1/2) s1 -∗ ownPong (1/2) s1 ==∗ ownPong (1/2) s2 ∗ ownPong (1/2) s2.
Proof.
iIntros "H1 H2". rewrite /ownPong. iCombine "H1 H2" as "H".
rewrite -own_op -pair_op frac_op' Qp_half_half agree_idemp.
iApply (own_update with "H"). by apply cmra_update_exclusive.
Qed.
Lemma ownPing_update s1 s2:
ownPing (1/2) s1 -∗ ownPing (1/2) s1 ==∗ ownPing (1/2) s2 ∗ ownPing (1/2) s2.
Proof.
iIntros "H1 H2". rewrite /ownPing. iCombine "H1 H2" as "H".
rewrite -own_op -pair_op frac_op' Qp_half_half agree_idemp.
iApply (own_update with "H"). by apply cmra_update_exclusive.
Qed.
End ghost_state.
Section proof.
Context `{!heapG Σ, !spawnG Σ, !pingpongG Σ}.
Context (N : namespace).
Definition iProp := iProp Σ.
Definition pong_inv (ℓ : loc) : iProp :=
(ℓ ↦ NONEV ∨
ℓ ↦ SOMEV MSG_PING ∗ ownPong (1/2) MSG_NONE ∨
ℓ ↦ SOMEV MSG_DONE ∗ ownPong (1/2) PING ∗ ownPing (1/2) PONG)%I.
Definition ping_inv (ℓ : loc) : iProp :=
(ℓ ↦ NONEV ∨
ℓ ↦ SOMEV MSG_PONG ∗ ownPong (1/2) PING ∗ ownPing (1/2) MSG_NONE)%I.
Lemma listen_wait_spec1 (ℓ_recv : loc) :
{{{ inv N (pong_inv ℓ_recv) ∗ ownPong (1/2) MSG_NONE }}}
listen_wait #ℓ_recv
{{{ v, RET v; ⌜v = MSG_PING⌝ ∗ ownPong (1/2) PING ∗ ownPong (1/2) PING }}}.
Proof.
iIntros (Φ) "(#Hinv & Hpong) HΦ". iLöb as "IH". rewrite /listen_wait.
wp_rec. wp_bind (!_)%E.
iInv N as ">[Hℓ | [(Hℓ & Hpong') | (Hℓ & Hpong' & Hping')]]" "Hclose".
- wp_load. iMod ("Hclose" with "[Hℓ]").
{ by iNext; iLeft. }
iModIntro. wp_let. wp_match. by iApply ("IH" with "Hpong").
- wp_load. iMod ("Hclose" with "[Hℓ Hpong']") as "_".
{ iNext. iRight; iLeft. by iFrame. }
iModIntro. wp_pures. wp_bind (CmpXchg _ _ _)%E.
iInv N as ">[Hℓ | [(Hℓ & Hpong') | (Hℓ & Hpong' & Hping')]]" "Hclose".
+ wp_cmpxchg_fail. iMod ("Hclose" with "[Hℓ]") as "_".
{ iNext. iLeft. by iFrame. }
iModIntro. wp_proj. wp_if. by iApply ("IH" with "Hpong").
+ wp_cmpxchg_suc.
iDestruct (ownPong_update _ PING with "Hpong' Hpong") as ">[Hpong' Hpong]".
iMod ("Hclose" with "[Hℓ]") as "_".
{ iNext. iLeft. by iFrame. }
iModIntro. wp_proj. wp_if. iApply "HΦ". by iFrame.
+ iDestruct (ownPong_eq with "Hpong Hpong'") as %Hcontr; inversion Hcontr.
- wp_load. iDestruct (ownPong_eq with "Hpong Hpong'") as %Hcontr; inversion Hcontr.
Qed.
Lemma listen_wait_spec2 (ℓ_recv : loc) :
{{{ inv N (pong_inv ℓ_recv) ∗ ownPong (1/2) PING }}}
listen_wait #ℓ_recv
{{{ v, RET v; ⌜v = MSG_DONE⌝
∗ ownPong (1/2) PING ∗ ownPong (1/2) PING ∗ ownPing (1/2) PONG }}}.
Proof.
iIntros (Φ) "(#Hinv & Hpong) HΦ". iLöb as "IH". rewrite /listen_wait.
wp_rec. wp_bind (!_)%E.
iInv N as ">[Hℓ | [(Hℓ & Hpong') | (Hℓ & Hpong' & Hping')]]" "Hclose".
- wp_load. iMod ("Hclose" with "[Hℓ]") as "_".
{ by iNext; iLeft. }
iModIntro. wp_let. wp_match. by iApply ("IH" with "Hpong").
- iDestruct (ownPong_eq with "Hpong Hpong'") as %Hcontr; inversion Hcontr.
- wp_load. iMod ("Hclose" with "[Hℓ Hpong' Hping']") as "_".
{ iNext. iRight; iRight. by iFrame. }
iModIntro. wp_pures. wp_bind (CmpXchg _ _ _)%E.
iInv N as ">[Hℓ | [(Hℓ & Hpong') | (Hℓ & Hpong' & Hping')]]" "Hclose".
+ wp_cmpxchg_fail. iMod ("Hclose" with "[Hℓ]") as "_".
{ by iNext; iLeft. }
iModIntro. wp_proj. wp_if. by iApply ("IH" with "Hpong").
+ wp_cmpxchg_fail. iMod ("Hclose" with "[Hℓ Hpong']") as "_".
{ iNext; iRight; iLeft. by iFrame. }
iModIntro. wp_proj. wp_if. by iApply ("IH" with "Hpong").
+ wp_cmpxchg_suc. iMod ("Hclose" with "[Hℓ]") as "_".
{ by iNext; iLeft. }
iModIntro. wp_proj. wp_if. iApply "HΦ". by iFrame.
Qed.
Lemma send_spec1 (ℓ_send : loc) :
{{{ inv N (ping_inv ℓ_send) ∗ ownPong (1/2) PING ∗ ownPing (1/2) MSG_NONE }}}
send #ℓ_send MSG_PONG
{{{ v, RET v; True }}}.
Proof.
iIntros (Φ) "(#Hinv & Hpong & Hping) HΦ". iLöb as "IH".
wp_lam. wp_pures. wp_bind (CmpXchg _ _ _).
iInv N as ">[Hℓ | (Hℓ & Hpong' & Hping')]" "Hclose".
- wp_cmpxchg_suc. iMod ("Hclose" with "[Hℓ Hpong Hping]") as "_".
{ iNext. iRight. by iFrame. }
iModIntro. wp_pures. by iApply "HΦ".
- wp_cmpxchg_fail. iMod ("Hclose" with "[Hℓ Hpong Hping]") as "_".
{ iNext. iRight. by iFrame. }
iModIntro. wp_pures. by iApply ("IH" with "[$Hpong'] [$Hping']").
Qed.
Lemma pong_spec (ℓ_recv ℓ_send : loc) :
{{{ inv N (pong_inv ℓ_recv)
∗ inv N (ping_inv ℓ_send)
∗ ownPong (1/2) MSG_NONE
∗ ownPing (1/2) MSG_NONE }}}
pong #ℓ_recv #ℓ_send
{{{ v, RET #v; ownPong 1 DONE }}}.
Proof.
iIntros (Φ) "(#Hrecv & #Hsend & Hpong & Hping) HΦ".
rewrite /pong. wp_pures.
wp_bind (listen_wait _). iApply (listen_wait_spec1 with "[$]").
iNext. iIntros (v) "(-> & Hpong & Hpong')". wp_let. wp_op. wp_if. wp_bind (send _ _).
iApply (send_spec1 with "[$]").
iNext. iIntros (v) "_". wp_seq. wp_bind (listen_wait _).
iApply (listen_wait_spec2 with "[$]").
iNext. iIntros (v') "(-> & Hpong & Hpong' & Hping)".
iDestruct (ownPong_update _ DONE with "Hpong' Hpong") as ">Hpong".
rewrite -own_op -pair_op frac_op' Qp_half_half agree_idemp.
wp_let. wp_op. wp_if. by iApply "HΦ".
Qed.
{{{ inv N (pong_inv ℓ_recv) ∗ ownPong (1/2) MSG_NONE }}}
listen_wait #ℓ_recv
{{{ v, RET v; ⌜v = MSG_PING⌝ ∗ ownPong (1/2) PING ∗ ownPong (1/2) PING }}}.
Proof.
iIntros (Φ) "(#Hinv & Hpong) HΦ". iLöb as "IH". rewrite /listen_wait.
wp_rec. wp_bind (!_)%E.
iInv N as ">[Hℓ | [(Hℓ & Hpong') | (Hℓ & Hpong' & Hping')]]" "Hclose".
- wp_load. iMod ("Hclose" with "[Hℓ]").
{ by iNext; iLeft. }
iModIntro. wp_let. wp_match. by iApply ("IH" with "Hpong").
- wp_load. iMod ("Hclose" with "[Hℓ Hpong']") as "_".
{ iNext. iRight; iLeft. by iFrame. }
iModIntro. wp_pures. wp_bind (CmpXchg _ _ _)%E.
iInv N as ">[Hℓ | [(Hℓ & Hpong') | (Hℓ & Hpong' & Hping')]]" "Hclose".
+ wp_cmpxchg_fail. iMod ("Hclose" with "[Hℓ]") as "_".
{ iNext. iLeft. by iFrame. }
iModIntro. wp_proj. wp_if. by iApply ("IH" with "Hpong").
+ wp_cmpxchg_suc.
iDestruct (ownPong_update _ PING with "Hpong' Hpong") as ">[Hpong' Hpong]".
iMod ("Hclose" with "[Hℓ]") as "_".
{ iNext. iLeft. by iFrame. }
iModIntro. wp_proj. wp_if. iApply "HΦ". by iFrame.
+ iDestruct (ownPong_eq with "Hpong Hpong'") as %Hcontr; inversion Hcontr.
- wp_load. iDestruct (ownPong_eq with "Hpong Hpong'") as %Hcontr; inversion Hcontr.
Qed.
Lemma listen_wait_spec2 (ℓ_recv : loc) :
{{{ inv N (pong_inv ℓ_recv) ∗ ownPong (1/2) PING }}}
listen_wait #ℓ_recv
{{{ v, RET v; ⌜v = MSG_DONE⌝
∗ ownPong (1/2) PING ∗ ownPong (1/2) PING ∗ ownPing (1/2) PONG }}}.
Proof.
iIntros (Φ) "(#Hinv & Hpong) HΦ". iLöb as "IH". rewrite /listen_wait.
wp_rec. wp_bind (!_)%E.
iInv N as ">[Hℓ | [(Hℓ & Hpong') | (Hℓ & Hpong' & Hping')]]" "Hclose".
- wp_load. iMod ("Hclose" with "[Hℓ]") as "_".
{ by iNext; iLeft. }
iModIntro. wp_let. wp_match. by iApply ("IH" with "Hpong").
- iDestruct (ownPong_eq with "Hpong Hpong'") as %Hcontr; inversion Hcontr.
- wp_load. iMod ("Hclose" with "[Hℓ Hpong' Hping']") as "_".
{ iNext. iRight; iRight. by iFrame. }
iModIntro. wp_pures. wp_bind (CmpXchg _ _ _)%E.
iInv N as ">[Hℓ | [(Hℓ & Hpong') | (Hℓ & Hpong' & Hping')]]" "Hclose".
+ wp_cmpxchg_fail. iMod ("Hclose" with "[Hℓ]") as "_".
{ by iNext; iLeft. }
iModIntro. wp_proj. wp_if. by iApply ("IH" with "Hpong").
+ wp_cmpxchg_fail. iMod ("Hclose" with "[Hℓ Hpong']") as "_".
{ iNext; iRight; iLeft. by iFrame. }
iModIntro. wp_proj. wp_if. by iApply ("IH" with "Hpong").
+ wp_cmpxchg_suc. iMod ("Hclose" with "[Hℓ]") as "_".
{ by iNext; iLeft. }
iModIntro. wp_proj. wp_if. iApply "HΦ". by iFrame.
Qed.
Lemma send_spec1 (ℓ_send : loc) :
{{{ inv N (ping_inv ℓ_send) ∗ ownPong (1/2) PING ∗ ownPing (1/2) MSG_NONE }}}
send #ℓ_send MSG_PONG
{{{ v, RET v; True }}}.
Proof.
iIntros (Φ) "(#Hinv & Hpong & Hping) HΦ". iLöb as "IH".
wp_lam. wp_pures. wp_bind (CmpXchg _ _ _).
iInv N as ">[Hℓ | (Hℓ & Hpong' & Hping')]" "Hclose".
- wp_cmpxchg_suc. iMod ("Hclose" with "[Hℓ Hpong Hping]") as "_".
{ iNext. iRight. by iFrame. }
iModIntro. wp_pures. by iApply "HΦ".
- wp_cmpxchg_fail. iMod ("Hclose" with "[Hℓ Hpong Hping]") as "_".
{ iNext. iRight. by iFrame. }
iModIntro. wp_pures. by iApply ("IH" with "[$Hpong'] [$Hping']").
Qed.
Lemma pong_spec (ℓ_recv ℓ_send : loc) :
{{{ inv N (pong_inv ℓ_recv)
∗ inv N (ping_inv ℓ_send)
∗ ownPong (1/2) MSG_NONE
∗ ownPing (1/2) MSG_NONE }}}
pong #ℓ_recv #ℓ_send
{{{ v, RET #v; ownPong 1 DONE }}}.
Proof.
iIntros (Φ) "(#Hrecv & #Hsend & Hpong & Hping) HΦ".
rewrite /pong. wp_pures.
wp_bind (listen_wait _). iApply (listen_wait_spec1 with "[$]").
iNext. iIntros (v) "(-> & Hpong & Hpong')". wp_let. wp_op. wp_if. wp_bind (send _ _).
iApply (send_spec1 with "[$]").
iNext. iIntros (v) "_". wp_seq. wp_bind (listen_wait _).
iApply (listen_wait_spec2 with "[$]").
iNext. iIntros (v') "(-> & Hpong & Hpong' & Hping)".
iDestruct (ownPong_update _ DONE with "Hpong' Hpong") as ">Hpong".
rewrite -own_op -pair_op frac_op' Qp_half_half agree_idemp.
wp_let. wp_op. wp_if. by iApply "HΦ".
Qed.
Lemma send_spec2 (ℓ_send : loc) :
{{{ inv N (pong_inv ℓ_send) ∗ ownPong (1/2) MSG_NONE }}}
send #ℓ_send MSG_PING
{{{ v, RET v; True }}}.
Proof.
iIntros (Φ) "[#Hinv Hpong] HΦ". iLöb as "IH".
rewrite /send. wp_pures. wp_bind (CmpXchg _ _ _).
iInv N as ">[Hℓ | [(Hℓ & Hpong') | (Hℓ & Hpong' & Hping')]]" "Hclose".
- wp_cmpxchg_suc. iMod ("Hclose" with "[Hℓ Hpong]") as "_".
{ iNext. iRight; iLeft. by iFrame. }
iModIntro. wp_proj. wp_if. by iApply "HΦ".
- wp_cmpxchg_fail. iMod ("Hclose" with "[Hℓ Hpong']") as "_".
{ iNext. iRight; iLeft. by iFrame. }
iModIntro. wp_proj. wp_if. by iApply ("IH" with "Hpong").
- wp_cmpxchg_fail. iMod ("Hclose" with "[Hℓ Hpong' Hping']") as "_".
{ iNext. iRight; iRight. by iFrame. }
iModIntro. wp_proj. wp_if. by iApply ("IH" with "Hpong").
Qed.
Lemma send_spec3 (ℓ_send : loc) :
{{{ inv N (pong_inv ℓ_send) ∗ ownPong (1/2) PING ∗ ownPing (1/2) PONG }}}
send #ℓ_send MSG_DONE
{{{ v, RET v; True }}}.
Proof.
iIntros (Φ) "(#Hinv & Hpong & Hping) HΦ". iLöb as "IH".
rewrite /send. wp_pures. wp_bind (CmpXchg _ _ _ ).
iInv N as ">[Hℓ | [(Hℓ & Hpong') | (Hℓ & Hpong' & Hping')]]" "Hclose".
- wp_cmpxchg_suc. iMod ("Hclose" with "[Hℓ Hpong Hping]") as "_".
{ iNext. iRight; iRight. by iFrame. }
iModIntro. wp_pures. by iApply "HΦ".
- wp_cmpxchg_fail. iMod ("Hclose" with "[Hℓ Hpong']") as "_".
{ iNext. iRight; iLeft. by iFrame. }
iModIntro. wp_proj. wp_if. by iApply ("IH" with "[$] [$]").
- wp_cmpxchg_fail. iMod ("Hclose" with "[Hℓ Hpong' Hping']") as "_".
{ iNext. iRight; iRight. by iFrame. }
iModIntro. wp_proj. wp_if. by iApply ("IH" with "[$] [$]").
Qed.
Lemma listen_wait_spec3 (ℓ_recv : loc) :
{{{ inv N (ping_inv ℓ_recv) }}}
listen_wait #ℓ_recv
{{{ v, RET v; ⌜v = MSG_PONG⌝ ∗ ownPong (1/2) PING ∗ ownPing (1/2) MSG_NONE }}}.
Proof.
iIntros (Φ) "#Hinv HΦ". iLöb as "IH". rewrite /listen_wait.
wp_rec. wp_bind (! _)%E.
iInv N as ">[Hℓ | (Hℓ & Hpong' & Hping')]" "Hclose".
- wp_load. iMod ("Hclose" with "[Hℓ]") as "_".
{ iNext. iLeft. done. }
iModIntro. wp_let. wp_match. by iApply ("IH" with "[$]").
- wp_load. iMod ("Hclose" with "[Hℓ Hpong' Hping']") as "_".
{ iNext. iRight. by iFrame. }
iModIntro. wp_pures. wp_bind (CmpXchg _ _ _).
iInv N as ">[Hℓ | (Hℓ & Hpong' & Hping')]" "Hclose".
+ wp_cmpxchg_fail. iMod ("Hclose" with "[Hℓ]") as "_".
{ iNext. iLeft. done. }
iModIntro. wp_proj. wp_if. by iApply ("IH" with "[$]").
+ wp_cmpxchg_suc. iMod ("Hclose" with "[Hℓ]") as "_".
{ iNext. iLeft. done. }
iModIntro. wp_proj. wp_if. iApply "HΦ". by iFrame.
Qed.
Lemma ping_spec (ℓ_recv ℓ_send : loc) :
{{{ inv N (pong_inv ℓ_send)
∗ inv N (ping_inv ℓ_recv)
∗ ownPong (1/2) MSG_NONE
∗ ownPing (1/2) MSG_NONE }}}
ping #ℓ_recv #ℓ_send
{{{ v, RET v; True }}}.
Proof.
iIntros (Φ) "(#Hrecv & #Hsend & Hpong & Hping) HΦ".
rewrite /ping. wp_pures. wp_bind (send _ _).
iApply (send_spec2 with "[$]").
iNext. iIntros (v) "_". wp_seq. wp_bind (listen_wait _).
iApply (listen_wait_spec3 with "[$]").
iNext. iIntros (v') "(-> & Hpong & Hping')".
iDestruct (ownPing_update _ PONG with "Hping Hping'") as ">[Hping Hping']".
wp_let. wp_op. wp_if. iApply (send_spec3 with "[$]"). iAssumption.
Qed.
Theorem ping_pong_done_runner_spec :
{{{ ownPong 1 MSG_NONE ∗ ownPing 1 MSG_NONE }}}
ping_pong_runner
{{{ v, RET v; True }}}.
Proof.
iIntros (Φ) "[Hpong Hping] HΦ". rewrite /ping_pong_runner.
wp_alloc ℓpong as "Hℓpong". wp_let.
wp_alloc ℓping as "Hℓping". wp_let.
iMod (inv_alloc N _ (pong_inv ℓpong) with "[Hℓpong]") as "#Hinv1".
{ iNext. by iLeft. }
iMod (inv_alloc N _ (ping_inv ℓping) with "[Hℓping]") as "#Hinv2".
{ iNext. by iLeft. }
rewrite /ownPong /ownPing -Qp_half_half -frac_op'
-(agree_idemp (@to_agree (leibnizO _) _))
(pair_op _ _ (@to_agree (leibnizO _ ) _)) !own_op.
iDestruct "Hpong" as "[Hpong1 Hpong2]". iDestruct "Hping" as "[Hping1 Hping2]".
wp_apply ((wp_par (λ _, True)%I (λ _, True))%I with "[Hpong1 Hping1] [Hpong2 Hping2]").
- iApply (pong_spec with "[$]"). iNext. iIntros (_) "_". done.
- iApply (ping_spec with "[$]"). iNext. iIntros (_) "_". done.
- iIntros (? ?) "_". by iApply "HΦ".
Qed.
End proof.
{{{ inv N (pong_inv ℓ_send) ∗ ownPong (1/2) MSG_NONE }}}
send #ℓ_send MSG_PING
{{{ v, RET v; True }}}.
Proof.
iIntros (Φ) "[#Hinv Hpong] HΦ". iLöb as "IH".
rewrite /send. wp_pures. wp_bind (CmpXchg _ _ _).
iInv N as ">[Hℓ | [(Hℓ & Hpong') | (Hℓ & Hpong' & Hping')]]" "Hclose".
- wp_cmpxchg_suc. iMod ("Hclose" with "[Hℓ Hpong]") as "_".
{ iNext. iRight; iLeft. by iFrame. }
iModIntro. wp_proj. wp_if. by iApply "HΦ".
- wp_cmpxchg_fail. iMod ("Hclose" with "[Hℓ Hpong']") as "_".
{ iNext. iRight; iLeft. by iFrame. }
iModIntro. wp_proj. wp_if. by iApply ("IH" with "Hpong").
- wp_cmpxchg_fail. iMod ("Hclose" with "[Hℓ Hpong' Hping']") as "_".
{ iNext. iRight; iRight. by iFrame. }
iModIntro. wp_proj. wp_if. by iApply ("IH" with "Hpong").
Qed.
Lemma send_spec3 (ℓ_send : loc) :
{{{ inv N (pong_inv ℓ_send) ∗ ownPong (1/2) PING ∗ ownPing (1/2) PONG }}}
send #ℓ_send MSG_DONE
{{{ v, RET v; True }}}.
Proof.
iIntros (Φ) "(#Hinv & Hpong & Hping) HΦ". iLöb as "IH".
rewrite /send. wp_pures. wp_bind (CmpXchg _ _ _ ).
iInv N as ">[Hℓ | [(Hℓ & Hpong') | (Hℓ & Hpong' & Hping')]]" "Hclose".
- wp_cmpxchg_suc. iMod ("Hclose" with "[Hℓ Hpong Hping]") as "_".
{ iNext. iRight; iRight. by iFrame. }
iModIntro. wp_pures. by iApply "HΦ".
- wp_cmpxchg_fail. iMod ("Hclose" with "[Hℓ Hpong']") as "_".
{ iNext. iRight; iLeft. by iFrame. }
iModIntro. wp_proj. wp_if. by iApply ("IH" with "[$] [$]").
- wp_cmpxchg_fail. iMod ("Hclose" with "[Hℓ Hpong' Hping']") as "_".
{ iNext. iRight; iRight. by iFrame. }
iModIntro. wp_proj. wp_if. by iApply ("IH" with "[$] [$]").
Qed.
Lemma listen_wait_spec3 (ℓ_recv : loc) :
{{{ inv N (ping_inv ℓ_recv) }}}
listen_wait #ℓ_recv
{{{ v, RET v; ⌜v = MSG_PONG⌝ ∗ ownPong (1/2) PING ∗ ownPing (1/2) MSG_NONE }}}.
Proof.
iIntros (Φ) "#Hinv HΦ". iLöb as "IH". rewrite /listen_wait.
wp_rec. wp_bind (! _)%E.
iInv N as ">[Hℓ | (Hℓ & Hpong' & Hping')]" "Hclose".
- wp_load. iMod ("Hclose" with "[Hℓ]") as "_".
{ iNext. iLeft. done. }
iModIntro. wp_let. wp_match. by iApply ("IH" with "[$]").
- wp_load. iMod ("Hclose" with "[Hℓ Hpong' Hping']") as "_".
{ iNext. iRight. by iFrame. }
iModIntro. wp_pures. wp_bind (CmpXchg _ _ _).
iInv N as ">[Hℓ | (Hℓ & Hpong' & Hping')]" "Hclose".
+ wp_cmpxchg_fail. iMod ("Hclose" with "[Hℓ]") as "_".
{ iNext. iLeft. done. }
iModIntro. wp_proj. wp_if. by iApply ("IH" with "[$]").
+ wp_cmpxchg_suc. iMod ("Hclose" with "[Hℓ]") as "_".
{ iNext. iLeft. done. }
iModIntro. wp_proj. wp_if. iApply "HΦ". by iFrame.
Qed.
Lemma ping_spec (ℓ_recv ℓ_send : loc) :
{{{ inv N (pong_inv ℓ_send)
∗ inv N (ping_inv ℓ_recv)
∗ ownPong (1/2) MSG_NONE
∗ ownPing (1/2) MSG_NONE }}}
ping #ℓ_recv #ℓ_send
{{{ v, RET v; True }}}.
Proof.
iIntros (Φ) "(#Hrecv & #Hsend & Hpong & Hping) HΦ".
rewrite /ping. wp_pures. wp_bind (send _ _).
iApply (send_spec2 with "[$]").
iNext. iIntros (v) "_". wp_seq. wp_bind (listen_wait _).
iApply (listen_wait_spec3 with "[$]").
iNext. iIntros (v') "(-> & Hpong & Hping')".
iDestruct (ownPing_update _ PONG with "Hping Hping'") as ">[Hping Hping']".
wp_let. wp_op. wp_if. iApply (send_spec3 with "[$]"). iAssumption.
Qed.
Theorem ping_pong_done_runner_spec :
{{{ ownPong 1 MSG_NONE ∗ ownPing 1 MSG_NONE }}}
ping_pong_runner
{{{ v, RET v; True }}}.
Proof.
iIntros (Φ) "[Hpong Hping] HΦ". rewrite /ping_pong_runner.
wp_alloc ℓpong as "Hℓpong". wp_let.
wp_alloc ℓping as "Hℓping". wp_let.
iMod (inv_alloc N _ (pong_inv ℓpong) with "[Hℓpong]") as "#Hinv1".
{ iNext. by iLeft. }
iMod (inv_alloc N _ (ping_inv ℓping) with "[Hℓping]") as "#Hinv2".
{ iNext. by iLeft. }
rewrite /ownPong /ownPing -Qp_half_half -frac_op'
-(agree_idemp (@to_agree (leibnizO _) _))
(pair_op _ _ (@to_agree (leibnizO _ ) _)) !own_op.
iDestruct "Hpong" as "[Hpong1 Hpong2]". iDestruct "Hping" as "[Hping1 Hping2]".
wp_apply ((wp_par (λ _, True)%I (λ _, True))%I with "[Hpong1 Hping1] [Hpong2 Hping2]").
- iApply (pong_spec with "[$]"). iNext. iIntros (_) "_". done.
- iApply (ping_spec with "[$]"). iNext. iIntros (_) "_". done.
- iIntros (? ?) "_". by iApply "HΦ".
Qed.
End proof.