From c3a22c76aebf9c440b990d631f6b47abb629013f Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Sat, 20 Jan 2024 13:21:37 +0000 Subject: [PATCH 01/81] Synchronous ghost theory for binary session types --- theories/channel/proto.v | 1361 ++++++++++++++++++++------------------ 1 file changed, 706 insertions(+), 655 deletions(-) diff --git a/theories/channel/proto.v b/theories/channel/proto.v index 92c4c00..c12d547 100644 --- a/theories/channel/proto.v +++ b/theories/channel/proto.v @@ -245,38 +245,53 @@ Definition iProto_dual_if {Σ V} (d : bool) (p : iProto Σ V) : iProto Σ V := Arguments iProto_dual_if {_ _} _ _%proto. Global Instance: Params (@iProto_dual_if) 3 := {}. -(** * Protocol entailment *) -Definition iProto_le_pre {Σ V} - (rec : iProto Σ V → iProto Σ V → iProp Σ) (p1 p2 : iProto Σ V) : iProp Σ := - (p1 ≡ END ∗ p2 ≡ END) ∨ - ∃ a1 a2 m1 m2, - (p1 ≡ <a1> m1) ∗ (p2 ≡ <a2> m2) ∗ - match a1, a2 with - | Recv, Recv => ∀ v p1', - iMsg_car m1 v (Next p1') -∗ ∃ p2', â–· rec p1' p2' ∗ iMsg_car m2 v (Next p2') - | Send, Send => ∀ v p2', - iMsg_car m2 v (Next p2') -∗ ∃ p1', â–· rec p1' p2' ∗ iMsg_car m1 v (Next p1') - | Recv, Send => ∀ v1 v2 p1' p2', - iMsg_car m1 v1 (Next p1') -∗ iMsg_car m2 v2 (Next p2') -∗ ∃ pt, - â–· rec p1' (<!> MSG v2; pt) ∗ â–· rec (<?> MSG v1; pt) p2' - | Send, Recv => False - end. -Global Instance iProto_le_pre_ne {Σ V} (rec : iProto Σ V → iProto Σ V → iProp Σ) : - NonExpansive2 (iProto_le_pre rec). +Definition iProto_consistent_pre {Σ V} (rec : iProto Σ V → iProto Σ V → iProp Σ) + (pl pr : iProto Σ V) : iProp Σ := + (∀ a1 a2 m1 m2, + (pl ≡ <a1> m1) -∗ (pr ≡ <a2> m2) -∗ + match a1,a2 with + | Send,Recv => ∀ v p1, iMsg_car m1 v (Next p1) -∗ + ∃ p2, iMsg_car m2 v (Next p2) ∗ â–· (rec p1 p2) + | Recv,Send => ∀ v p2, iMsg_car m2 v (Next p2) -∗ + ∃ p1, iMsg_car m1 v (Next p1) ∗ â–· (rec p1 p2) + | _, _ => True + end). + +Global Instance iProto_consistent_pre_ne {Σ V} + (rec : iProto Σ V -n> iProto Σ V -n> iPropO Σ) : + NonExpansive2 (iProto_consistent_pre (λ p1 p2, rec p1 p2)). Proof. solve_proper. Qed. -Program Definition iProto_le_pre' {Σ V} - (rec : iProto Σ V -n> iProto Σ V -n> iPropO Σ) : - iProto Σ V -n> iProto Σ V -n> iPropO Σ := λne p1 p2, - iProto_le_pre (λ p1' p2', rec p1' p2') p1 p2. +Program Definition iProto_consistent_pre' {Σ V} + (rec : iProto Σ V -n> iProto Σ V -n> iPropO Σ) : + iProto Σ V -n> iProto Σ V -n> iPropO Σ := + λne p1 p2, iProto_consistent_pre (λ p1 p2, rec p1 p2) p1 p2. Solve Obligations with solve_proper. -Local Instance iProto_le_pre_contractive {Σ V} : Contractive (@iProto_le_pre' Σ V). + +Local Instance iProto_consistent_pre_contractive {Σ V} : Contractive (@iProto_consistent_pre' Σ V). Proof. - intros n rec1 rec2 Hrec p1 p2. rewrite /iProto_le_pre' /iProto_le_pre /=. - by repeat (f_contractive || f_equiv). + intros n rec1 rec2 Hrec pl pr. rewrite /iProto_consistent_pre' /iProto_consistent_pre /=. + repeat (f_contractive || f_equiv). done. done. +Qed. + +Definition iProto_consistent {Σ V} (pl pr : iProto Σ V) : iProp Σ := + fixpoint iProto_consistent_pre' pl pr. + +Arguments iProto_consistent {_ _} _%proto _%proto. +Global Instance: Params (@iProto_consistent) 2 := {}. + +Global Instance iProto_consistent_ne {Σ V} : NonExpansive2 (@iProto_consistent Σ V). +Proof. solve_proper. Qed. +Global Instance iProto_consistent_proper {Σ V} : Proper ((≡) ==> (≡) ==> (⊣⊢)) (@iProto_consistent Σ V). +Proof. solve_proper. Qed. + +Lemma iProto_consistent_unfold {Σ V} (pl pr : iProto Σ V) : + iProto_consistent pl pr ≡ (iProto_consistent_pre iProto_consistent) pl pr. +Proof. + apply: (fixpoint_unfold iProto_consistent_pre'). Qed. Definition iProto_le {Σ V} (p1 p2 : iProto Σ V) : iProp Σ := - fixpoint iProto_le_pre' p1 p2. + ∀ p3, iProto_consistent p1 p3 -∗ iProto_consistent p2 p3. Arguments iProto_le {_ _} _%proto _%proto. Global Instance: Params (@iProto_le) 2 := {}. Notation "p ⊑ q" := (iProto_le p q) : bi_scope. @@ -291,8 +306,6 @@ Fixpoint iProto_app_recvs {Σ V} (vs : list V) (p : iProto Σ V) : iProto Σ V : | [] => p | v :: vs => <?> MSG v; iProto_app_recvs vs p end. -Definition iProto_interp {Σ V} (vsl vsr : list V) (pl pr : iProto Σ V) : iProp Σ := - ∃ p, iProto_app_recvs vsr p ⊑ pl ∗ iProto_app_recvs vsl (iProto_dual p) ⊑ pr. Record proto_name := ProtName { proto_l_name : gname; proto_r_name : gname }. Global Instance proto_name_inhabited : Inhabited proto_name := @@ -325,11 +338,11 @@ Definition iProto_own_auth `{!protoG Σ V} (γ : proto_name) (s : side) own (side_elim s proto_l_name proto_r_name γ) (â—E (Next p)). Definition iProto_ctx `{protoG Σ V} - (γ : proto_name) (vsl vsr : list V) : iProp Σ := + (γ : proto_name) : iProp Σ := ∃ pl pr, iProto_own_auth γ Left pl ∗ iProto_own_auth γ Right pr ∗ - â–· iProto_interp vsl vsr pl pr. + â–· iProto_consistent pl pr. (** * The connective for ownership of channel ends *) Definition iProto_own `{!protoG Σ V} @@ -604,413 +617,531 @@ Section proto. iRewrite "Hp12". iIntros "!>". iRewrite "Hp1'". by iRewrite ("IH" $! p1d p2). Qed. - (** ** Protocol entailment **) - Lemma iProto_le_unfold p1 p2 : iProto_le p1 p2 ≡ iProto_le_pre iProto_le p1 p2. - Proof. apply: (fixpoint_unfold iProto_le_pre'). Qed. - - Lemma iProto_le_end : ⊢ END ⊑ (END : iProto Σ V). - Proof. rewrite iProto_le_unfold. iLeft. auto 10. Qed. - - Lemma iProto_le_send m1 m2 : - (∀ v p2', iMsg_car m2 v (Next p2') -∗ ∃ p1', - â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1')) -∗ - (<!> m1) ⊑ (<!> m2). - Proof. rewrite iProto_le_unfold. iIntros "H". iRight. auto 10. Qed. - Lemma iProto_le_recv m1 m2 : - (∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', - â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2')) -∗ - (<?> m1) ⊑ (<?> m2). - Proof. rewrite iProto_le_unfold. iIntros "H". iRight. auto 10. Qed. - Lemma iProto_le_swap m1 m2 : - (∀ v1 v2 p1' p2', - iMsg_car m1 v1 (Next p1') -∗ iMsg_car m2 v2 (Next p2') -∗ ∃ pt, - â–· (p1' ⊑ <!> MSG v2; pt) ∗ â–· ((<?> MSG v1; pt) ⊑ p2')) -∗ - (<?> m1) ⊑ (<!> m2). - Proof. rewrite iProto_le_unfold. iIntros "H". iRight. auto 10. Qed. - - Lemma iProto_le_end_inv_l p : p ⊑ END -∗ (p ≡ END). - Proof. - rewrite iProto_le_unfold. iIntros "[[Hp _]|H] //". - iDestruct "H" as (a1 a2 m1 m2) "(_ & Heq & _)". - by iDestruct (iProto_end_message_equivI with "Heq") as %[]. - Qed. - - Lemma iProto_le_end_inv_r p : END ⊑ p -∗ (p ≡ END). - Proof. - rewrite iProto_le_unfold. iIntros "[[_ Hp]|H] //". - iDestruct "H" as (a1 a2 m1 m2) "(Heq & _ & _)". - iDestruct (iProto_end_message_equivI with "Heq") as %[]. - Qed. - - Lemma iProto_le_send_inv p1 m2 : - p1 ⊑ (<!> m2) -∗ ∃ a1 m1, - (p1 ≡ <a1> m1) ∗ - match a1 with - | Send => ∀ v p2', - iMsg_car m2 v (Next p2') -∗ ∃ p1', - â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1') - | Recv => ∀ v1 v2 p1' p2', - iMsg_car m1 v1 (Next p1') -∗ iMsg_car m2 v2 (Next p2') -∗ ∃ pt, - â–· (p1' ⊑ <!> MSG v2; pt) ∗ â–· ((<?> MSG v1; pt) ⊑ p2') - end. - Proof. - rewrite iProto_le_unfold. iIntros "[[_ Heq]|H]". - { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } - iDestruct "H" as (a1 a2 m1 m2') "(Hp1 & Hp2 & H)". - iExists _, _; iSplit; [done|]. destruct a1, a2. + Lemma iProto_consistent_flip p1 p2 : + iProto_consistent p2 p1 -∗ + iProto_consistent p1 p2. + Proof. + iLöb as "IH" forall (p1 p2). + iIntros "Hprot". + rewrite iProto_consistent_unfold /iProto_consistent_pre. + rewrite iProto_consistent_unfold /iProto_consistent_pre. + iIntros (a1 a2 m1 m2) "Hm1 Hm2". + iDestruct ("Hprot" with "Hm2 Hm1") as "Hprot". + destruct a1, a2; [done| | |done]. + - iIntros (v p1') "Hm1". + iDestruct ("Hprot" with "Hm1") as (p2') "[Hprot Hrec]". + iExists p2'. iFrame. by iApply "IH". - iIntros (v p2') "Hm2". - iDestruct (iProto_message_equivI with "Hp2") as (_) "{Hp2} #Hm". - iApply "H". by iRewrite -("Hm" $! v (Next p2')). - - done. - - iIntros (v1 v2 p1' p2') "Hm1 Hm2". - iDestruct (iProto_message_equivI with "Hp2") as (_) "{Hp2} #Hm". - iApply ("H" with "Hm1"). by iRewrite -("Hm" $! v2 (Next p2')). - - iDestruct (iProto_message_equivI with "Hp2") as ([=]) "_". - Qed. - Lemma iProto_le_send_send_inv m1 m2 v p2' : - (<!> m1) ⊑ (<!> m2) -∗ - iMsg_car m2 v (Next p2') -∗ ∃ p1', â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1'). - Proof. - iIntros "H Hm2". iDestruct (iProto_le_send_inv with "H") as (a m1') "[Hm1 H]". - iDestruct (iProto_message_equivI with "Hm1") as (<-) "Hm1". - iDestruct ("H" with "Hm2") as (p1') "[Hle Hm]". - iRewrite -("Hm1" $! v (Next p1')) in "Hm". auto with iFrame. - Qed. - Lemma iProto_le_recv_send_inv m1 m2 v1 v2 p1' p2' : - (<?> m1) ⊑ (<!> m2) -∗ - iMsg_car m1 v1 (Next p1') -∗ iMsg_car m2 v2 (Next p2') -∗ ∃ pt, - â–· (p1' ⊑ <!> MSG v2; pt) ∗ â–· ((<?> MSG v1; pt) ⊑ p2'). - Proof. - iIntros "H Hm1 Hm2". iDestruct (iProto_le_send_inv with "H") as (a m1') "[Hm H]". - iDestruct (iProto_message_equivI with "Hm") as (<-) "{Hm} #Hm". - iApply ("H" with "[Hm1] Hm2"). by iRewrite -("Hm" $! v1 (Next p1')). - Qed. - - Lemma iProto_le_recv_inv p1 m2 : - p1 ⊑ (<?> m2) -∗ ∃ m1, - (p1 ≡ <?> m1) ∗ - ∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', - â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). - Proof. - rewrite iProto_le_unfold. iIntros "[[_ Heq]|H]". - { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } - iDestruct "H" as (a1 a2 m1 m2') "(Hp1 & Hp2 & H)". - iExists m1. - iDestruct (iProto_message_equivI with "Hp2") as (<-) "{Hp2} #Hm2". - destruct a1; [done|]. iSplit; [done|]. - iIntros (v p1') "Hm". iDestruct ("H" with "Hm") as (p2') "[Hle Hm]". - iExists p2'. iIntros "{$Hle}". by iRewrite ("Hm2" $! v (Next p2')). - Qed. - Lemma iProto_le_recv_recv_inv m1 m2 v p1' : - (<?> m1) ⊑ (<?> m2) -∗ - iMsg_car m1 v (Next p1') -∗ ∃ p2', â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). - Proof. - iIntros "H Hm2". iDestruct (iProto_le_recv_inv with "H") as (m1') "[Hm1 H]". - iApply "H". iDestruct (iProto_message_equivI with "Hm1") as (_) "Hm1". - by iRewrite -("Hm1" $! v (Next p1')). + iDestruct ("Hprot" with "Hm2") as (p1') "[Hprot Hrec]". + iExists p1'. iFrame. by iApply "IH". + Qed. + + Lemma iProto_consistent_step m1 m2 v p1 : + iProto_consistent (<!> m1) (<?> m2) -∗ + iMsg_car m1 v (Next p1) -∗ + ∃ p2, iMsg_car m2 v (Next p2) ∗ â–· iProto_consistent p1 p2. + Proof. + iIntros "Hprot Hm1". + rewrite iProto_consistent_unfold /iProto_consistent_pre. + iDestruct ("Hprot" with "[//] [//] Hm1") as (p2) "[Hm2 Hprot]". + iExists p2. iFrame. + Qed. + + Lemma iProto_consistent_dual p : + ⊢ iProto_consistent p (iProto_dual p). + Proof. + iLöb as "IH" forall (p). + rewrite iProto_consistent_unfold. + iIntros (a1 a2 m1 m2) "#Hmeq1 #Hmeq2". + destruct a1, a2; [done| | |done]. + - iIntros (v p') "Hm1". + iRewrite "Hmeq1" in "Hmeq2". + rewrite iProto_dual_message. simpl. + rewrite iProto_message_equivI. + iDestruct "Hmeq2" as (_) "Hmeq2". + iSpecialize ("Hmeq2" $! v (Next (iProto_dual p'))). + iExists (iProto_dual p'). + iRewrite -"Hmeq2". + iSplitL; [|iNext; by iApply "IH"]. + simpl. iExists p'. iFrame. done. + - iIntros (v p') "Hm2". + iRewrite "Hmeq1" in "Hmeq2". + rewrite iProto_dual_message. simpl. + rewrite iProto_message_equivI. + iDestruct "Hmeq2" as (_) "Hmeq2". + iSpecialize ("Hmeq2" $! v (Next p')). + iRewrite -"Hmeq2" in "Hm2". + simpl. + iDestruct "Hm2" as (p'') "[Hm2 Hmeq']". + iExists p''. iFrame. + iNext. iRewrite "Hmeq'". + iApply "IH". + Qed. + + Lemma iProto_consistent_le_l pl pl' pr : + iProto_consistent pl pr -∗ pl ⊑ pl' -∗ iProto_consistent pl' pr. + Proof. iIntros "Hprot Hle". by iApply "Hle". Qed. + Lemma iProto_consistent_le_r pl pr pr' : + iProto_consistent pl pr -∗ pr ⊑ pr' -∗ iProto_consistent pl pr'. + Proof. + iIntros "H Hle". iApply iProto_consistent_flip. + iApply "Hle". by iApply iProto_consistent_flip. Qed. Lemma iProto_le_refl p : ⊢ p ⊑ p. - Proof. - iLöb as "IH" forall (p). destruct (iProto_case p) as [->|([]&m&->)]. - - iApply iProto_le_end. - - iApply iProto_le_send. auto 10 with iFrame. - - iApply iProto_le_recv. auto 10 with iFrame. - Qed. - - Lemma iProto_le_trans p1 p2 p3 : p1 ⊑ p2 -∗ p2 ⊑ p3 -∗ p1 ⊑ p3. - Proof. - iIntros "H1 H2". iLöb as "IH" forall (p1 p2 p3). - destruct (iProto_case p3) as [->|([]&m3&->)]. - - iDestruct (iProto_le_end_inv_l with "H2") as "H2". by iRewrite "H2" in "H1". - - iDestruct (iProto_le_send_inv with "H2") as (a2 m2) "[Hp2 H2]". - iRewrite "Hp2" in "H1"; clear p2. destruct a2. - + iDestruct (iProto_le_send_inv with "H1") as (a1 m1) "[Hp1 H1]". - iRewrite "Hp1"; clear p1. destruct a1. - * iApply iProto_le_send. iIntros (v p3') "Hm3". - iDestruct ("H2" with "Hm3") as (p2') "[Hle Hm2]". - iDestruct ("H1" with "Hm2") as (p1') "[Hle' Hm1]". - iExists p1'. iIntros "{$Hm1} !>". by iApply ("IH" with "Hle'"). - * iApply iProto_le_swap. iIntros (v1 v3 p1' p3') "Hm1 Hm3". - iDestruct ("H2" with "Hm3") as (p2') "[Hle Hm2]". - iDestruct ("H1" with "Hm1 Hm2") as (pt) "[Hp1' Hp2']". - iExists pt. iIntros "{$Hp1'} !>". by iApply ("IH" with "Hp2'"). - + iDestruct (iProto_le_recv_inv with "H1") as (m1) "[Hp1 H1]". - iRewrite "Hp1"; clear p1. iApply iProto_le_swap. - iIntros (v1 v3 p1' p3') "Hm1 Hm3". - iDestruct ("H1" with "Hm1") as (p2') "[Hle Hm2]". - iDestruct ("H2" with "Hm2 Hm3") as (pt) "[Hp2' Hp3']". - iExists pt. iIntros "{$Hp3'} !>". by iApply ("IH" with "Hle"). - - iDestruct (iProto_le_recv_inv with "H2") as (m2) "[Hp2 H3]". - iRewrite "Hp2" in "H1". - iDestruct (iProto_le_recv_inv with "H1") as (m1) "[Hp1 H2]". - iRewrite "Hp1". iApply iProto_le_recv. iIntros (v p1') "Hm1". - iDestruct ("H2" with "Hm1") as (p2') "[Hle Hm2]". - iDestruct ("H3" with "Hm2") as (p3') "[Hle' Hm3]". - iExists p3'. iIntros "{$Hm3} !>". by iApply ("IH" with "Hle"). - Qed. - - Lemma iProto_le_payload_elim_l a m v P p : - (P -∗ (<?> MSG v; p) ⊑ (<a> m)) ⊢ - (<?> MSG v {{ P }}; p) ⊑ (<a> m). - Proof. - rewrite iMsg_base_eq. iIntros "H". destruct a. - - iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= (#?&#?&HP) Hm2 /=". - iApply (iProto_le_recv_send_inv with "(H HP)"); simpl; auto. - - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&HP)". - iApply (iProto_le_recv_recv_inv with "(H HP)"); simpl; auto. - Qed. - Lemma iProto_le_payload_elim_r a m v P p : - (P -∗ (<a> m) ⊑ (<!> MSG v; p)) ⊢ - (<a> m) ⊑ (<!> MSG v {{ P }}; p). - Proof. - rewrite iMsg_base_eq. iIntros "H". destruct a. - - iApply iProto_le_send. iIntros (v' p') "(->&Hp&HP)". - iApply (iProto_le_send_send_inv with "(H HP)"); simpl; auto. - - iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1 (->&#?&HP) /=". - iApply (iProto_le_recv_send_inv with "(H HP) Hm1"); simpl; auto. - Qed. - Lemma iProto_le_payload_intro_l v P p : - P -∗ (<!> MSG v {{ P }}; p) ⊑ (<!> MSG v; p). - Proof. - rewrite iMsg_base_eq. - iIntros "HP". iApply iProto_le_send. iIntros (v' p') "(->&Hp&_) /=". - iExists p'. iSplitR; [iApply iProto_le_refl|]. auto. - Qed. - Lemma iProto_le_payload_intro_r v P p : - P -∗ (<?> MSG v; p) ⊑ (<?> MSG v {{ P }}; p). - Proof. - rewrite iMsg_base_eq. - iIntros "HP". iApply iProto_le_recv. iIntros (v' p') "(->&Hp&_) /=". - iExists p'. iSplitR; [iApply iProto_le_refl|]. auto. - Qed. - - Lemma iProto_le_exist_elim_l {A} (m1 : A → iMsg Σ V) a m2 : - (∀ x, (<?> m1 x) ⊑ (<a> m2)) ⊢ - (<? x> m1 x) ⊑ (<a> m2). - Proof. - rewrite iMsg_exist_eq. iIntros "H". destruct a. - - iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1 Hm2 /=". - iDestruct "Hm1" as (x) "Hm1". - iApply (iProto_le_recv_send_inv with "H Hm1 Hm2"). - - iApply iProto_le_recv. iIntros (v p1') "/=". iDestruct 1 as (x) "Hm". - by iApply (iProto_le_recv_recv_inv with "H"). - Qed. - - Lemma iProto_le_exist_elim_l_inhabited `{!Inhabited A} (m : A → iMsg Σ V) p : - (∀ x, (<?> m x) ⊑ p) ⊢ - (<? x> m x) ⊑ p. - Proof. - rewrite iMsg_exist_eq. iIntros "H". - destruct (iProto_case p) as [Heq | [a [m' Heq]]]. - - unshelve iSpecialize ("H" $!inhabitant); first by apply _. - rewrite Heq. - iDestruct (iProto_le_end_inv_l with "H") as "H". - rewrite iProto_end_eq iProto_message_eq. - iDestruct (proto_message_end_equivI with "H") as "[]". - - iEval (rewrite Heq). destruct a. - + iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1 Hm2 /=". - iDestruct "Hm1" as (x) "Hm1". - iSpecialize ("H" $! x). rewrite Heq. - iApply (iProto_le_recv_send_inv with "H Hm1 Hm2"). - + iApply iProto_le_recv. iIntros (v p1') "/=". iDestruct 1 as (x) "Hm". - iSpecialize ("H" $! x). rewrite Heq. - by iApply (iProto_le_recv_recv_inv with "H"). - Qed. - - Lemma iProto_le_exist_elim_r {A} a m1 (m2 : A → iMsg Σ V) : - (∀ x, (<a> m1) ⊑ (<!> m2 x)) ⊢ - (<a> m1) ⊑ (<! x> m2 x). - Proof. - rewrite iMsg_exist_eq. iIntros "H". destruct a. - - iApply iProto_le_send. iIntros (v p2'). iDestruct 1 as (x) "Hm". - by iApply (iProto_le_send_send_inv with "H"). - - iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1". - iDestruct 1 as (x) "Hm2". - iApply (iProto_le_recv_send_inv with "H Hm1 Hm2"). - Qed. - Lemma iProto_le_exist_elim_r_inhabited `{Hinh : Inhabited A} p (m : A → iMsg Σ V) : - (∀ x, p ⊑ (<!> m x)) ⊢ - p ⊑ (<! x> m x). - Proof. - rewrite iMsg_exist_eq. iIntros "H". - destruct (iProto_case p) as [Heq | [a [m' Heq]]]. - - unshelve iSpecialize ("H" $!inhabitant); first by apply _. - rewrite Heq. - iDestruct (iProto_le_end_inv_r with "H") as "H". - rewrite iProto_end_eq iProto_message_eq. - iDestruct (proto_message_end_equivI with "H") as "[]". - - iEval (rewrite Heq). destruct a. - + iApply iProto_le_send. iIntros (v p2'). iDestruct 1 as (x) "Hm". - iSpecialize ("H" $! x). rewrite Heq. - by iApply (iProto_le_send_send_inv with "H"). - + iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1". - iDestruct 1 as (x) "Hm2". - iSpecialize ("H" $! x). rewrite Heq. - iApply (iProto_le_recv_send_inv with "H Hm1 Hm2"). - Qed. - Lemma iProto_le_exist_intro_l {A} (m : A → iMsg Σ V) a : - ⊢ (<! x> m x) ⊑ (<!> m a). - Proof. - rewrite iMsg_exist_eq. iApply iProto_le_send. iIntros (v p') "Hm /=". - iExists p'. iSplitR; last by auto. iApply iProto_le_refl. - Qed. - Lemma iProto_le_exist_intro_r {A} (m : A → iMsg Σ V) a : - ⊢ (<?> m a) ⊑ (<? x> m x). - Proof. - rewrite iMsg_exist_eq. iApply iProto_le_recv. iIntros (v p') "Hm /=". - iExists p'. iSplitR; last by auto. iApply iProto_le_refl. - Qed. - - Lemma iProto_le_texist_elim_l {TT : tele} (m1 : TT → iMsg Σ V) a m2 : - (∀ x, (<?> m1 x) ⊑ (<a> m2)) ⊢ - (<?.. x> m1 x) ⊑ (<a> m2). - Proof. - iIntros "H". iInduction TT as [|T TT] "IH"; simpl; [done|]. - iApply iProto_le_exist_elim_l; iIntros (x). - iApply "IH". iIntros (xs). iApply "H". - Qed. - Lemma iProto_le_texist_elim_r {TT : tele} a m1 (m2 : TT → iMsg Σ V) : - (∀ x, (<a> m1) ⊑ (<!> m2 x)) -∗ - (<a> m1) ⊑ (<!.. x> m2 x). - Proof. - iIntros "H". iInduction TT as [|T TT] "IH"; simpl; [done|]. - iApply iProto_le_exist_elim_r; iIntros (x). - iApply "IH". iIntros (xs). iApply "H". - Qed. - - Lemma iProto_le_texist_intro_l {TT : tele} (m : TT → iMsg Σ V) x : - ⊢ (<!.. x> m x) ⊑ (<!> m x). - Proof. - induction x as [|T TT x xs IH] using tele_arg_ind; simpl. - { iApply iProto_le_refl. } - iApply iProto_le_trans; [by iApply iProto_le_exist_intro_l|]. iApply IH. - Qed. - Lemma iProto_le_texist_intro_r {TT : tele} (m : TT → iMsg Σ V) x : - ⊢ (<?> m x) ⊑ (<?.. x> m x). - Proof. - induction x as [|T TT x xs IH] using tele_arg_ind; simpl. - { iApply iProto_le_refl. } - iApply iProto_le_trans; [|by iApply iProto_le_exist_intro_r]. iApply IH. - Qed. + Proof. iIntros (p') "$". Qed. - Lemma iProto_le_base a v P p1 p2 : - â–· (p1 ⊑ p2) ⊢ - (<a> MSG v {{ P }}; p1) ⊑ (<a> MSG v {{ P }}; p2). - Proof. - rewrite iMsg_base_eq. iIntros "H". destruct a. - - iApply iProto_le_send. iIntros (v' p') "(->&Hp&$)". - iExists p1. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". - - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&$)". - iExists p2. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". - Qed. - - Lemma iProto_le_base_swap v1 v2 P1 P2 p : - ⊢ (<?> MSG v1 {{ P1 }}; <!> MSG v2 {{ P2 }}; p) - ⊑ (<!> MSG v2 {{ P2 }}; <?> MSG v1 {{ P1 }}; p). - Proof. - rewrite {1 3}iMsg_base_eq. iApply iProto_le_swap. - iIntros (v1' v2' p1' p2') "/= (->&#Hp1&HP1) (->&#Hp2&HP2)". iExists p. - iSplitL "HP2". - - iIntros "!>". iRewrite -"Hp1". by iApply iProto_le_payload_intro_l. - - iIntros "!>". iRewrite -"Hp2". by iApply iProto_le_payload_intro_r. - Qed. - - Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. - Proof. - iIntros "H". iLöb as "IH" forall (p1 p2). - destruct (iProto_case p1) as [->|([]&m1&->)]. - - iDestruct (iProto_le_end_inv_l with "H") as "H". - iRewrite "H". iApply iProto_le_refl. - - iDestruct (iProto_le_send_inv with "H") as (a2 m2) "[Hp2 H]". - iRewrite "Hp2"; clear p2. iEval (rewrite !iProto_dual_message). - destruct a2; simpl. - + iApply iProto_le_recv. iIntros (v p1d). - iDestruct 1 as (p1') "[Hm1 #Hp1d]". - iDestruct ("H" with "Hm1") as (p2') "[H Hm2]". - iDestruct ("IH" with "H") as "H". iExists (iProto_dual p2'). - iSplitL "H"; [iIntros "!>"; by iRewrite "Hp1d"|]. simpl; auto. - + iApply iProto_le_swap. iIntros (v1 v2 p1d p2d). - iDestruct 1 as (p1') "[Hm1 #Hp1d]". iDestruct 1 as (p2') "[Hm2 #Hp2d]". - iDestruct ("H" with "Hm2 Hm1") as (pt) "[H1 H2]". - iDestruct ("IH" with "H1") as "H1". iDestruct ("IH" with "H2") as "H2 {IH}". - rewrite !iProto_dual_message /=. iExists (iProto_dual pt). iSplitL "H2". - * iIntros "!>". iRewrite "Hp1d". by rewrite -iMsg_dual_base. - * iIntros "!>". iRewrite "Hp2d". by rewrite -iMsg_dual_base. - - iDestruct (iProto_le_recv_inv with "H") as (m2) "[Hp2 H]". - iRewrite "Hp2"; clear p2. iEval (rewrite !iProto_dual_message /=). - iApply iProto_le_send. iIntros (v p2d). - iDestruct 1 as (p2') "[Hm2 #Hp2d]". - iDestruct ("H" with "Hm2") as (p1') "[H Hm1]". - iDestruct ("IH" with "H") as "H". iExists (iProto_dual p1'). - iSplitL "H"; [iIntros "!>"; by iRewrite "Hp2d"|]. simpl; auto. - Qed. - - Lemma iProto_le_amber_internal (p1 p2 : iProto Σ V → iProto Σ V) - `{Contractive p1, Contractive p2}: - â–¡ (∀ rec1 rec2, â–· (rec1 ⊑ rec2) → p1 rec1 ⊑ p2 rec2) ⊢ - fixpoint p1 ⊑ fixpoint p2. - Proof. - iIntros "#H". iLöb as "IH". - iEval (rewrite (fixpoint_unfold p1)). - iEval (rewrite (fixpoint_unfold p2)). - iApply "H". iApply "IH". - Qed. - Lemma iProto_le_amber_external (p1 p2 : iProto Σ V → iProto Σ V) - `{Contractive p1, Contractive p2}: - (∀ rec1 rec2, (⊢ rec1 ⊑ rec2) → ⊢ p1 rec1 ⊑ p2 rec2) → - ⊢ fixpoint p1 ⊑ fixpoint p2. - Proof. - intros IH. apply fixpoint_ind. - - by intros p1' p2' -> ?. - - exists (fixpoint p2). iApply iProto_le_refl. - - intros p' ?. rewrite (fixpoint_unfold p2). by apply IH. - - apply bi.limit_preserving_entails; [done|solve_proper]. - Qed. - - Lemma iProto_le_dual_l p1 p2 : iProto_dual p2 ⊑ p1 ⊢ iProto_dual p1 ⊑ p2. - Proof. - iIntros "H". iEval (rewrite -(involutive iProto_dual p2)). - by iApply iProto_le_dual. - Qed. - Lemma iProto_le_dual_r p1 p2 : p2 ⊑ iProto_dual p1 ⊢ p1 ⊑ iProto_dual p2. - Proof. - iIntros "H". iEval (rewrite -(involutive iProto_dual p1)). - by iApply iProto_le_dual. - Qed. + Lemma iProto_le_end : ⊢ END ⊑ (END : iProto Σ V). + Proof. iApply iProto_le_refl. Qed. + + (* Lemma iProto_le_send m1 m2 : *) + (* (∀ v p2', iMsg_car m2 v (Next p2') -∗ ∃ p1', *) + (* â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1')) -∗ *) + (* (<!> m1) ⊑ (<!> m2). *) + (* Proof. *) + (* iLöb as "IH" forall (m1 m2). *) + (* iIntros "Hle". *) + (* repeat (unfold iProto_le at 3). *) + (* iIntros (p) "Hprot". *) + (* repeat rewrite iProto_consistent_unfold; unfold iProto_consistent_pre. *) + (* iSplit; [iDestruct "Hprot" as "[Hprot _]"|iDestruct "Hprot" as "[_ Hprot]"]. *) + (* { iIntros (a m) "H". *) + (* iDestruct (iProto_message_equivI with "H") as (Heq) "{H} #Hmeq"; subst. *) + (* iIntros (v p') "Hm". *) + (* iSpecialize ("Hprot" with "[//]"). *) + (* iSpecialize ("Hmeq" $! v (Next p')). iRewrite -"Hmeq" in "Hm". *) + (* iSpecialize ("Hle" $! _ _ with "Hm"). iDestruct "Hle" as (p'') "[Hle Hm]". *) + (* iApply "Hle". iApply "Hprot". iApply "Hm". } *) + (* { iIntros (a m) "H". *) + (* iSpecialize ("Hprot" $! _ _ with "H"). *) + (* destruct a. *) + (* - iIntros (v p') "Hm". *) + (* iApply ("IH" with "[Hle]"). *) + (* { iNext. iIntros (v' p'') "Hm". *) + (* iSpecialize ("Hle" with "Hm"). *) + (* iDestruct "Hle" as (p''') "[Hle Hm]". *) + (* iExists p'''. iSplitR "Hm"; [iApply "Hle" | iApply "Hm"]. } *) + (* iApply "Hprot". iApply "Hm". *) + (* - iIntros (v vs Heq); subst. *) + (* iSpecialize ("Hprot" $! _ _ eq_refl). *) + (* iDestruct "Hprot" as (p') "[Hm Hprot]". *) + (* iExists (p'). iFrame "Hm". *) + (* iApply ("IH" with "[Hle] [Hprot]"); [|iApply "Hprot"]. *) + (* iNext. iIntros (v' p'') "Hm". *) + (* iSpecialize ("Hle" with "Hm"). iDestruct "Hle" as (p''') "[Hle Hm]". *) + (* iExists p'''; iSplitR "Hm"; [iApply "Hle" |iApply "Hm"]. } *) + (* Qed. *) + + (* Lemma iProto_le_recv m1 m2 : *) + (* (∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', *) + (* â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2')) -∗ *) + (* (<?> m1) ⊑ (<?> m2). *) + (* Proof. *) + (* iIntros "H" (vsl vsr p) "Hprot". *) + (* iLöb as "IH" forall (vsl vsr m2 p). *) + (* iEval (rewrite iProto_consistent_unfold). *) + (* rewrite /iProto_consistent_pre. *) + (* rewrite iProto_consistent_unfold /iProto_consistent_pre. *) + (* iSplit; last first. *) + (* { *) + (* iDestruct "Hprot" as "[_ Hprot]". *) + (* iIntros ([] m) "Heq". *) + (* { iIntros (v vs) "Hm". *) + (* iDestruct ("Hprot" with "Heq Hm") as "Hprot". *) + (* iIntros "!>". *) + (* by iApply ("IH" with "H"). } *) + (* { iIntros (v vs) "Hm". *) + (* iDestruct ("Hprot" with "Heq Hm") as (p') "[Hm Hprot]". *) + (* iExists p'. iFrame. *) + (* iIntros "!>". *) + (* by iApply ("IH" with "H"). } *) + (* } *) + (* destruct vsr. *) + (* { iIntros (a m) "Heq". rewrite iProto_message_equivI. *) + (* iDestruct "Heq" as (<-) "Heq". *) + (* iIntros (v vs Heq). done. } *) + (* iIntros (a m) "Heq". rewrite iProto_message_equivI. *) + (* iDestruct "Heq" as (<-) "Heq". *) + (* iIntros (w vs Heq). *) + (* assert (v = w) as <- by set_solver. *) + (* assert (vsr = vs) as <- by set_solver. *) + (* iDestruct "Hprot" as "[Hprot _]". *) + (* iDestruct ("Hprot" with "[//] [//]") as (p') "[Hm Hprot]". *) + (* iDestruct ("H" with "Hm") as (p'') "[Hle H]". *) + (* iExists p''. *) + (* iSpecialize ("Heq" $! v (Next p'')). iRewrite -"Heq". *) + (* iFrame. iIntros "!>". *) + (* iApply (iProto_consistent_le_l with "Hprot Hle"). *) + (* Qed. *) + + (* Lemma iProto_le_base a v P p1 p2 : *) + (* â–· (p1 ⊑ p2) -∗ *) + (* (<a> MSG v {{ P }}; p1) ⊑ (<a> MSG v {{ P }}; p2). *) + (* Proof. *) + (* rewrite iMsg_base_eq. iIntros "H". destruct a. *) + (* - iApply iProto_le_send. iIntros (v' p') "(->&Hp&$)". *) + (* iExists p1. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". *) + (* - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&$)". *) + (* iExists p2. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". *) + (* Qed. *) + + (* Lemma iProto_le_end_inv_l p : p ⊑ END -∗ (p ≡ END). *) + (* Proof. *) + (* rewrite iProto_le_unfold. iIntros "[[Hp _]|H] //". *) + (* iDestruct "H" as (a1 a2 m1 m2) "(_ & Heq & _)". *) + (* by iDestruct (iProto_end_message_equivI with "Heq") as %[]. *) + (* Qed. *) + + (* Lemma iProto_le_end_inv_r p : END ⊑ p -∗ (p ≡ END). *) + (* Proof. *) + (* rewrite iProto_le_unfold. iIntros "[[_ Hp]|H] //". *) + (* iDestruct "H" as (a1 a2 m1 m2) "(Heq & _ & _)". *) + (* iDestruct (iProto_end_message_equivI with "Heq") as %[]. *) + (* Qed. *) + + (* Lemma iProto_le_send_inv p1 m2 : *) + (* p1 ⊑ (<!> m2) -∗ ∃ a1 m1, *) + (* (p1 ≡ <a1> m1) ∗ *) + (* match a1 with *) + (* | Send => ∀ v p2', *) + (* iMsg_car m2 v (Next p2') -∗ ∃ p1', *) + (* â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1') *) + (* | Recv => ∀ v1 v2 p1' p2', *) + (* iMsg_car m1 v1 (Next p1') -∗ iMsg_car m2 v2 (Next p2') -∗ ∃ pt, *) + (* â–· (p1' ⊑ <!> MSG v2; pt) ∗ â–· ((<?> MSG v1; pt) ⊑ p2') *) + (* end. *) + (* Proof. *) + (* rewrite iProto_le_unfold. iIntros "[[_ Heq]|H]". *) + (* { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } *) + (* iDestruct "H" as (a1 a2 m1 m2') "(Hp1 & Hp2 & H)". *) + (* iExists _, _; iSplit; [done|]. destruct a1, a2. *) + (* - iIntros (v p2') "Hm2". *) + (* iDestruct (iProto_message_equivI with "Hp2") as (_) "{Hp2} #Hm". *) + (* iApply "H". by iRewrite -("Hm" $! v (Next p2')). *) + (* - done. *) + (* - iIntros (v1 v2 p1' p2') "Hm1 Hm2". *) + (* iDestruct (iProto_message_equivI with "Hp2") as (_) "{Hp2} #Hm". *) + (* iApply ("H" with "Hm1"). by iRewrite -("Hm" $! v2 (Next p2')). *) + (* - iDestruct (iProto_message_equivI with "Hp2") as ([=]) "_". *) + (* Qed. *) + (* Lemma iProto_le_send_send_inv m1 m2 v p2' : *) + (* (<!> m1) ⊑ (<!> m2) -∗ *) + (* iMsg_car m2 v (Next p2') -∗ ∃ p1', â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1'). *) + (* Proof. *) + (* iIntros "H Hm2". iDestruct (iProto_le_send_inv with "H") as (a m1') "[Hm1 H]". *) + (* iDestruct (iProto_message_equivI with "Hm1") as (<-) "Hm1". *) + (* iDestruct ("H" with "Hm2") as (p1') "[Hle Hm]". *) + (* iRewrite -("Hm1" $! v (Next p1')) in "Hm". auto with iFrame. *) + (* Qed. *) + (* Lemma iProto_le_recv_send_inv m1 m2 v1 v2 p1' p2' : *) + (* (<?> m1) ⊑ (<!> m2) -∗ *) + (* iMsg_car m1 v1 (Next p1') -∗ iMsg_car m2 v2 (Next p2') -∗ ∃ pt, *) + (* â–· (p1' ⊑ <!> MSG v2; pt) ∗ â–· ((<?> MSG v1; pt) ⊑ p2'). *) + (* Proof. *) + (* iIntros "H Hm1 Hm2". iDestruct (iProto_le_send_inv with "H") as (a m1') "[Hm H]". *) + (* iDestruct (iProto_message_equivI with "Hm") as (<-) "{Hm} #Hm". *) + (* iApply ("H" with "[Hm1] Hm2"). by iRewrite -("Hm" $! v1 (Next p1')). *) + (* Qed. *) + + (* Lemma iProto_le_recv_inv p1 m2 : *) + (* p1 ⊑ (<?> m2) -∗ ∃ m1, *) + (* (p1 ≡ <?> m1) ∗ *) + (* ∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', *) + (* â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). *) + (* Proof. *) + (* rewrite iProto_le_unfold. iIntros "[[_ Heq]|H]". *) + (* { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } *) + (* iDestruct "H" as (a1 a2 m1 m2') "(Hp1 & Hp2 & H)". *) + (* iExists m1. *) + (* iDestruct (iProto_message_equivI with "Hp2") as (<-) "{Hp2} #Hm2". *) + (* destruct a1; [done|]. iSplit; [done|]. *) + (* iIntros (v p1') "Hm". iDestruct ("H" with "Hm") as (p2') "[Hle Hm]". *) + (* iExists p2'. iIntros "{$Hle}". by iRewrite ("Hm2" $! v (Next p2')). *) + (* Qed. *) + (* Lemma iProto_le_recv_recv_inv m1 m2 v p1' : *) + (* (<?> m1) ⊑ (<?> m2) -∗ *) + (* iMsg_car m1 v (Next p1') -∗ ∃ p2', â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). *) + (* Proof. *) + (* iIntros "H Hm2". iDestruct (iProto_le_recv_inv with "H") as (m1') "[Hm1 H]". *) + (* iApply "H". iDestruct (iProto_message_equivI with "Hm1") as (_) "Hm1". *) + (* by iRewrite -("Hm1" $! v (Next p1')). *) + (* Qed. *) + + (* Lemma iProto_le_refl p : ⊢ p ⊑ p. *) + (* Proof. *) + (* iLöb as "IH" forall (p). destruct (iProto_case p) as [->|([]&m&->)]. *) + (* - iApply iProto_le_end. *) + (* - iApply iProto_le_send. auto 10 with iFrame. *) + (* - iApply iProto_le_recv. auto 10 with iFrame. *) + (* Qed. *) - Lemma iProto_le_app p1 p2 p3 p4 : - p1 ⊑ p2 -∗ p3 ⊑ p4 -∗ p1 <++> p3 ⊑ p2 <++> p4. + Lemma iProto_le_trans p1 p2 p3 : p1 ⊑ p2 -∗ p2 ⊑ p3 -∗ p1 ⊑ p3. Proof. - iIntros "H1 H2". iLöb as "IH" forall (p1 p2 p3 p4). - destruct (iProto_case p2) as [->|([]&m2&->)]. - - iDestruct (iProto_le_end_inv_l with "H1") as "H1". - iRewrite "H1". by rewrite !left_id. - - iDestruct (iProto_le_send_inv with "H1") as (a1 m1) "[Hp1 H1]". - iRewrite "Hp1"; clear p1. rewrite !iProto_app_message. destruct a1; simpl. - + iApply iProto_le_send. iIntros (v p24). - iDestruct 1 as (p2') "[Hm2 #Hp24]". - iDestruct ("H1" with "Hm2") as (p1') "[H1 Hm1]". - iExists (p1' <++> p3). iSplitR "Hm1"; [|by simpl; eauto]. - iIntros "!>". iRewrite "Hp24". by iApply ("IH" with "H1"). - + iApply iProto_le_swap. iIntros (v1 v2 p13 p24). - iDestruct 1 as (p1') "[Hm1 #Hp13]". iDestruct 1 as (p2') "[Hm2 #Hp24]". - iSpecialize ("H1" with "Hm1 Hm2"). - iDestruct "H1" as (pt) "[H1 H1']". - iExists (pt <++> p3). iSplitL "H1". - * iIntros "!>". iRewrite "Hp13". - rewrite /= -iMsg_app_base -iProto_app_message. - iApply ("IH" with "H1"). iApply iProto_le_refl. - * iIntros "!>". iRewrite "Hp24". - rewrite /= -iMsg_app_base -iProto_app_message. - iApply ("IH" with "H1' H2"). - - iDestruct (iProto_le_recv_inv with "H1") as (m1) "[Hp1 H1]". - iRewrite "Hp1"; clear p1. rewrite !iProto_app_message. iApply iProto_le_recv. - iIntros (v p13). iDestruct 1 as (p1') "[Hm1 #Hp13]". - iDestruct ("H1" with "Hm1") as (p2'') "[H1 Hm2]". - iExists (p2'' <++> p4). iSplitR "Hm2"; [|by simpl; eauto]. - iIntros "!>". iRewrite "Hp13". by iApply ("IH" with "H1"). - Qed. + iIntros "H1 H2" (p) "Hprot". + iApply "H2". iApply "H1". done. + Qed. + + (* Lemma iProto_le_payload_elim_l a m v P p : *) + (* (P -∗ (<?> MSG v; p) ⊑ (<a> m)) -∗ *) + (* (<?> MSG v {{ P }}; p) ⊑ (<a> m). *) + (* Proof. *) + (* rewrite iMsg_base_eq. iIntros "H". destruct a. *) + (* - iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= (#?&#?&HP) Hm2 /=". *) + (* iApply (iProto_le_recv_send_inv with "(H HP)"); simpl; auto. *) + (* - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&HP)". *) + (* iApply (iProto_le_recv_recv_inv with "(H HP)"); simpl; auto. *) + (* Qed. *) + (* Lemma iProto_le_payload_elim_r a m v P p : *) + (* (P -∗ (<a> m) ⊑ (<!> MSG v; p)) -∗ *) + (* (<a> m) ⊑ (<!> MSG v {{ P }}; p). *) + (* Proof. *) + (* rewrite iMsg_base_eq. iIntros "H". destruct a. *) + (* - iApply iProto_le_send. iIntros (v' p') "(->&Hp&HP)". *) + (* iApply (iProto_le_send_send_inv with "(H HP)"); simpl; auto. *) + (* - iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1 (->&#?&HP) /=". *) + (* iApply (iProto_le_recv_send_inv with "(H HP) Hm1"); simpl; auto. *) + (* Qed. *) + (* Lemma iProto_le_payload_intro_l v P p : *) + (* P -∗ (<!> MSG v {{ P }}; p) ⊑ (<!> MSG v; p). *) + (* Proof. *) + (* rewrite iMsg_base_eq. *) + (* iIntros "HP". iApply iProto_le_send. iIntros (v' p') "(->&Hp&_) /=". *) + (* iExists p'. iSplitR; [iApply iProto_le_refl|]. auto. *) + (* Qed. *) + (* Lemma iProto_le_payload_intro_r v P p : *) + (* P -∗ (<?> MSG v; p) ⊑ (<?> MSG v {{ P }}; p). *) + (* Proof. *) + (* rewrite iMsg_base_eq. *) + (* iIntros "HP". iApply iProto_le_recv. iIntros (v' p') "(->&Hp&_) /=". *) + (* iExists p'. iSplitR; [iApply iProto_le_refl|]. auto. *) + (* Qed. *) + + (* Lemma iProto_le_exist_elim_l {A} (m1 : A → iMsg Σ V) a m2 : *) + (* (∀ x, (<?> m1 x) ⊑ (<a> m2)) -∗ *) + (* (<? x> m1 x) ⊑ (<a> m2). *) + (* Proof. *) + (* rewrite iMsg_exist_eq. iIntros "H". destruct a. *) + (* - iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1 Hm2 /=". *) + (* iDestruct "Hm1" as (x) "Hm1". *) + (* iApply (iProto_le_recv_send_inv with "H Hm1 Hm2"). *) + (* - iApply iProto_le_recv. iIntros (v p1') "/=". iDestruct 1 as (x) "Hm". *) + (* by iApply (iProto_le_recv_recv_inv with "H"). *) + (* Qed. *) + + (* Lemma iProto_le_exist_elim_l_inhabited `{!Inhabited A} (m : A → iMsg Σ V) p : *) + (* (∀ x, (<?> m x) ⊑ p) -∗ *) + (* (<? x> m x) ⊑ p. *) + (* Proof. *) + (* rewrite iMsg_exist_eq. iIntros "H". *) + (* destruct (iProto_case p) as [Heq | [a [m' Heq]]]. *) + (* - unshelve iSpecialize ("H" $!inhabitant); first by apply _. *) + (* rewrite Heq. *) + (* iDestruct (iProto_le_end_inv_l with "H") as "H". *) + (* rewrite iProto_end_eq iProto_message_eq. *) + (* iDestruct (proto_message_end_equivI with "H") as "[]". *) + (* - iEval (rewrite Heq). destruct a. *) + (* + iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1 Hm2 /=". *) + (* iDestruct "Hm1" as (x) "Hm1". *) + (* iSpecialize ("H" $! x). rewrite Heq. *) + (* iApply (iProto_le_recv_send_inv with "H Hm1 Hm2"). *) + (* + iApply iProto_le_recv. iIntros (v p1') "/=". iDestruct 1 as (x) "Hm". *) + (* iSpecialize ("H" $! x). rewrite Heq. *) + (* by iApply (iProto_le_recv_recv_inv with "H"). *) + (* Qed. *) + + (* Lemma iProto_le_exist_elim_r {A} a m1 (m2 : A → iMsg Σ V) : *) + (* (∀ x, (<a> m1) ⊑ (<!> m2 x)) -∗ *) + (* (<a> m1) ⊑ (<! x> m2 x). *) + (* Proof. *) + (* rewrite iMsg_exist_eq. iIntros "H". destruct a. *) + (* - iApply iProto_le_send. iIntros (v p2'). iDestruct 1 as (x) "Hm". *) + (* by iApply (iProto_le_send_send_inv with "H"). *) + (* - iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1". *) + (* iDestruct 1 as (x) "Hm2". *) + (* iApply (iProto_le_recv_send_inv with "H Hm1 Hm2"). *) + (* Qed. *) + (* Lemma iProto_le_exist_elim_r_inhabited `{Hinh : Inhabited A} p (m : A → iMsg Σ V) : *) + (* (∀ x, p ⊑ (<!> m x)) -∗ *) + (* p ⊑ (<! x> m x). *) + (* Proof. *) + (* rewrite iMsg_exist_eq. iIntros "H". *) + (* destruct (iProto_case p) as [Heq | [a [m' Heq]]]. *) + (* - unshelve iSpecialize ("H" $!inhabitant); first by apply _. *) + (* rewrite Heq. *) + (* iDestruct (iProto_le_end_inv_r with "H") as "H". *) + (* rewrite iProto_end_eq iProto_message_eq. *) + (* iDestruct (proto_message_end_equivI with "H") as "[]". *) + (* - iEval (rewrite Heq). destruct a. *) + (* + iApply iProto_le_send. iIntros (v p2'). iDestruct 1 as (x) "Hm". *) + (* iSpecialize ("H" $! x). rewrite Heq. *) + (* by iApply (iProto_le_send_send_inv with "H"). *) + (* + iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1". *) + (* iDestruct 1 as (x) "Hm2". *) + (* iSpecialize ("H" $! x). rewrite Heq. *) + (* iApply (iProto_le_recv_send_inv with "H Hm1 Hm2"). *) + (* Qed. *) + (* Lemma iProto_le_exist_intro_l {A} (m : A → iMsg Σ V) a : *) + (* ⊢ (<! x> m x) ⊑ (<!> m a). *) + (* Proof. *) + (* rewrite iMsg_exist_eq. iApply iProto_le_send. iIntros (v p') "Hm /=". *) + (* iExists p'. iSplitR; last by auto. iApply iProto_le_refl. *) + (* Qed. *) + (* Lemma iProto_le_exist_intro_r {A} (m : A → iMsg Σ V) a : *) + (* ⊢ (<?> m a) ⊑ (<? x> m x). *) + (* Proof. *) + (* rewrite iMsg_exist_eq. iApply iProto_le_recv. iIntros (v p') "Hm /=". *) + (* iExists p'. iSplitR; last by auto. iApply iProto_le_refl. *) + (* Qed. *) + + (* Lemma iProto_le_texist_elim_l {TT : tele} (m1 : TT → iMsg Σ V) a m2 : *) + (* (∀ x, (<?> m1 x) ⊑ (<a> m2)) -∗ *) + (* (<?.. x> m1 x) ⊑ (<a> m2). *) + (* Proof. *) + (* iIntros "H". iInduction TT as [|T TT] "IH"; simpl; [done|]. *) + (* iApply iProto_le_exist_elim_l; iIntros (x). *) + (* iApply "IH". iIntros (xs). iApply "H". *) + (* Qed. *) + (* Lemma iProto_le_texist_elim_r {TT : tele} a m1 (m2 : TT → iMsg Σ V) : *) + (* (∀ x, (<a> m1) ⊑ (<!> m2 x)) -∗ *) + (* (<a> m1) ⊑ (<!.. x> m2 x). *) + (* Proof. *) + (* iIntros "H". iInduction TT as [|T TT] "IH"; simpl; [done|]. *) + (* iApply iProto_le_exist_elim_r; iIntros (x). *) + (* iApply "IH". iIntros (xs). iApply "H". *) + (* Qed. *) + + (* Lemma iProto_le_texist_intro_l {TT : tele} (m : TT → iMsg Σ V) x : *) + (* ⊢ (<!.. x> m x) ⊑ (<!> m x). *) + (* Proof. *) + (* induction x as [|T TT x xs IH] using tele_arg_ind; simpl. *) + (* { iApply iProto_le_refl. } *) + (* iApply iProto_le_trans; [by iApply iProto_le_exist_intro_l|]. iApply IH. *) + (* Qed. *) + (* Lemma iProto_le_texist_intro_r {TT : tele} (m : TT → iMsg Σ V) x : *) + (* ⊢ (<?> m x) ⊑ (<?.. x> m x). *) + (* Proof. *) + (* induction x as [|T TT x xs IH] using tele_arg_ind; simpl. *) + (* { iApply iProto_le_refl. } *) + (* iApply iProto_le_trans; [|by iApply iProto_le_exist_intro_r]. iApply IH. *) + (* Qed. *) + + (* Lemma iProto_le_base a v P p1 p2 : *) + (* â–· (p1 ⊑ p2) -∗ *) + (* (<a> MSG v {{ P }}; p1) ⊑ (<a> MSG v {{ P }}; p2). *) + (* Proof. *) + (* rewrite iMsg_base_eq. iIntros "H". destruct a. *) + (* - iApply iProto_le_send. iIntros (v' p') "(->&Hp&$)". *) + (* iExists p1. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". *) + (* - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&$)". *) + (* iExists p2. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". *) + (* Qed. *) + + (* Lemma iProto_le_base_swap v1 v2 P1 P2 p : *) + (* ⊢ (<?> MSG v1 {{ P1 }}; <!> MSG v2 {{ P2 }}; p) *) + (* ⊑ (<!> MSG v2 {{ P2 }}; <?> MSG v1 {{ P1 }}; p). *) + (* Proof. *) + (* rewrite {1 3}iMsg_base_eq. iApply iProto_le_swap. *) + (* iIntros (v1' v2' p1' p2') "/= (->&#Hp1&HP1) (->&#Hp2&HP2)". iExists p. *) + (* iSplitL "HP2". *) + (* - iIntros "!>". iRewrite -"Hp1". by iApply iProto_le_payload_intro_l. *) + (* - iIntros "!>". iRewrite -"Hp2". by iApply iProto_le_payload_intro_r. *) + (* Qed. *) + + (* Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. *) + (* Proof. *) + (* iIntros "H". iLöb as "IH" forall (p1 p2). *) + (* destruct (iProto_case p1) as [->|([]&m1&->)]. *) + (* - iDestruct (iProto_le_end_inv_l with "H") as "H". *) + (* iRewrite "H". iApply iProto_le_refl. *) + (* - iDestruct (iProto_le_send_inv with "H") as (a2 m2) "[Hp2 H]". *) + (* iRewrite "Hp2"; clear p2. iEval (rewrite !iProto_dual_message). *) + (* destruct a2; simpl. *) + (* + iApply iProto_le_recv. iIntros (v p1d). *) + (* iDestruct 1 as (p1') "[Hm1 #Hp1d]". *) + (* iDestruct ("H" with "Hm1") as (p2') "[H Hm2]". *) + (* iDestruct ("IH" with "H") as "H". iExists (iProto_dual p2'). *) + (* iSplitL "H"; [iIntros "!>"; by iRewrite "Hp1d"|]. simpl; auto. *) + (* + iApply iProto_le_swap. iIntros (v1 v2 p1d p2d). *) + (* iDestruct 1 as (p1') "[Hm1 #Hp1d]". iDestruct 1 as (p2') "[Hm2 #Hp2d]". *) + (* iDestruct ("H" with "Hm2 Hm1") as (pt) "[H1 H2]". *) + (* iDestruct ("IH" with "H1") as "H1". iDestruct ("IH" with "H2") as "H2 {IH}". *) + (* rewrite !iProto_dual_message /=. iExists (iProto_dual pt). iSplitL "H2". *) + (* * iIntros "!>". iRewrite "Hp1d". by rewrite -iMsg_dual_base. *) + (* * iIntros "!>". iRewrite "Hp2d". by rewrite -iMsg_dual_base. *) + (* - iDestruct (iProto_le_recv_inv with "H") as (m2) "[Hp2 H]". *) + (* iRewrite "Hp2"; clear p2. iEval (rewrite !iProto_dual_message /=). *) + (* iApply iProto_le_send. iIntros (v p2d). *) + (* iDestruct 1 as (p2') "[Hm2 #Hp2d]". *) + (* iDestruct ("H" with "Hm2") as (p1') "[H Hm1]". *) + (* iDestruct ("IH" with "H") as "H". iExists (iProto_dual p1'). *) + (* iSplitL "H"; [iIntros "!>"; by iRewrite "Hp2d"|]. simpl; auto. *) + (* Qed. *) + + (* Lemma iProto_le_amber_internal (p1 p2 : iProto Σ V → iProto Σ V) *) + (* `{Contractive p1, Contractive p2}: *) + (* â–¡ (∀ rec1 rec2, â–· (rec1 ⊑ rec2) → p1 rec1 ⊑ p2 rec2) -∗ *) + (* fixpoint p1 ⊑ fixpoint p2. *) + (* Proof. *) + (* iIntros "#H". iLöb as "IH". *) + (* iEval (rewrite (fixpoint_unfold p1)). *) + (* iEval (rewrite (fixpoint_unfold p2)). *) + (* iApply "H". iApply "IH". *) + (* Qed. *) + (* Lemma iProto_le_amber_external (p1 p2 : iProto Σ V → iProto Σ V) *) + (* `{Contractive p1, Contractive p2}: *) + (* (∀ rec1 rec2, (⊢ rec1 ⊑ rec2) → ⊢ p1 rec1 ⊑ p2 rec2) → *) + (* ⊢ fixpoint p1 ⊑ fixpoint p2. *) + (* Proof. *) + (* intros IH. apply fixpoint_ind. *) + (* - by intros p1' p2' -> ?. *) + (* - exists (fixpoint p2). iApply iProto_le_refl. *) + (* - intros p' ?. rewrite (fixpoint_unfold p2). by apply IH. *) + (* - apply bi.limit_preserving_entails; [done|solve_proper]. *) + (* Qed. *) + + (* Lemma iProto_le_dual_l p1 p2 : iProto_dual p2 ⊑ p1 -∗ iProto_dual p1 ⊑ p2. *) + (* Proof. *) + (* iIntros "H". iEval (rewrite -(involutive iProto_dual p2)). *) + (* by iApply iProto_le_dual. *) + (* Qed. *) + (* Lemma iProto_le_dual_r p1 p2 : p2 ⊑ iProto_dual p1 -∗ p1 ⊑ iProto_dual p2. *) + (* Proof. *) + (* iIntros "H". iEval (rewrite -(involutive iProto_dual p1)). *) + (* by iApply iProto_le_dual. *) + (* Qed. *) + + (* Lemma iProto_le_app p1 p2 p3 p4 : *) + (* p1 ⊑ p2 -∗ p3 ⊑ p4 -∗ p1 <++> p3 ⊑ p2 <++> p4. *) + (* Proof. *) + (* iIntros "H1 H2". iLöb as "IH" forall (p1 p2 p3 p4). *) + (* destruct (iProto_case p2) as [->|([]&m2&->)]. *) + (* - iDestruct (iProto_le_end_inv_l with "H1") as "H1". *) + (* iRewrite "H1". by rewrite !left_id. *) + (* - iDestruct (iProto_le_send_inv with "H1") as (a1 m1) "[Hp1 H1]". *) + (* iRewrite "Hp1"; clear p1. rewrite !iProto_app_message. destruct a1; simpl. *) + (* + iApply iProto_le_send. iIntros (v p24). *) + (* iDestruct 1 as (p2') "[Hm2 #Hp24]". *) + (* iDestruct ("H1" with "Hm2") as (p1') "[H1 Hm1]". *) + (* iExists (p1' <++> p3). iSplitR "Hm1"; [|by simpl; eauto]. *) + (* iIntros "!>". iRewrite "Hp24". by iApply ("IH" with "H1"). *) + (* + iApply iProto_le_swap. iIntros (v1 v2 p13 p24). *) + (* iDestruct 1 as (p1') "[Hm1 #Hp13]". iDestruct 1 as (p2') "[Hm2 #Hp24]". *) + (* iSpecialize ("H1" with "Hm1 Hm2"). *) + (* iDestruct "H1" as (pt) "[H1 H1']". *) + (* iExists (pt <++> p3). iSplitL "H1". *) + (* * iIntros "!>". iRewrite "Hp13". *) + (* rewrite /= -iMsg_app_base -iProto_app_message. *) + (* iApply ("IH" with "H1"). iApply iProto_le_refl. *) + (* * iIntros "!>". iRewrite "Hp24". *) + (* rewrite /= -iMsg_app_base -iProto_app_message. *) + (* iApply ("IH" with "H1' H2"). *) + (* - iDestruct (iProto_le_recv_inv with "H1") as (m1) "[Hp1 H1]". *) + (* iRewrite "Hp1"; clear p1. rewrite !iProto_app_message. iApply iProto_le_recv. *) + (* iIntros (v p13). iDestruct 1 as (p1') "[Hm1 #Hp13]". *) + (* iDestruct ("H1" with "Hm1") as (p2'') "[H1 Hm2]". *) + (* iExists (p2'' <++> p4). iSplitR "Hm2"; [|by simpl; eauto]. *) + (* iIntros "!>". iRewrite "Hp13". by iApply ("IH" with "H1"). *) + (* Qed. *) (** ** Lemmas about the auxiliary definitions and invariants *) Global Instance iProto_app_recvs_ne vs : @@ -1019,12 +1150,6 @@ Section proto. Global Instance iProto_app_recvs_proper vs : Proper ((≡) ==> (≡)) (iProto_app_recvs (Σ:=Σ) (V:=V) vs). Proof. induction vs; solve_proper. Qed. - Global Instance iProto_interp_ne vsl vsr : - NonExpansive2 (iProto_interp (Σ:=Σ) (V:=V) vsl vsr). - Proof. solve_proper. Qed. - Global Instance iProto_interp_proper vsl vsr : - Proper ((≡) ==> (≡) ==> (≡)) (iProto_interp (Σ:=Σ) (V:=V) vsl vsr). - Proof. apply (ne_proper_2 _). Qed. Global Instance iProto_own_frag_ne γ s : NonExpansive (iProto_own_frag γ s). Proof. solve_proper. Qed. @@ -1032,8 +1157,8 @@ Section proto. Lemma iProto_own_auth_agree γ s p p' : iProto_own_auth γ s p -∗ iProto_own_frag γ s p' -∗ â–· (p ≡ p'). Proof. - iIntros "Hâ— Hâ—¯". iCombine "Hâ— Hâ—¯" gives "H✓". - iDestruct (excl_auth_agreeI with "H✓") as "{H✓} H✓". + iIntros "Hâ— Hâ—¯". iDestruct (own_valid_2 with "Hâ— Hâ—¯") as "H✓". + iDestruct (excl_auth_agreeI with "H✓") as "H✓". iApply (later_equivI_1 with "H✓"). Qed. @@ -1046,67 +1171,6 @@ Section proto. by rewrite own_op. Qed. - Lemma iProto_interp_nil p : ⊢ iProto_interp [] [] p (iProto_dual p). - Proof. iExists p; simpl. iSplitL; iApply iProto_le_refl. Qed. - - Lemma iProto_interp_flip vsl vsr pl pr : - iProto_interp vsl vsr pl pr -∗ iProto_interp vsr vsl pr pl. - Proof. - iDestruct 1 as (p) "[Hp Hdp]". iExists (iProto_dual p). - rewrite (involutive iProto_dual). iFrame. - Qed. - - Lemma iProto_interp_le_l vsl vsr pl pl' pr : - iProto_interp vsl vsr pl pr -∗ pl ⊑ pl' -∗ iProto_interp vsl vsr pl' pr. - Proof. - iDestruct 1 as (p) "[Hp Hdp]". iIntros "Hle". iExists p. - iFrame "Hdp". by iApply (iProto_le_trans with "Hp"). - Qed. - Lemma iProto_interp_le_r vsl vsr pl pr pr' : - iProto_interp vsl vsr pl pr -∗ pr ⊑ pr' -∗ iProto_interp vsl vsr pl pr'. - Proof. - iIntros "H Hle". iApply iProto_interp_flip. - iApply (iProto_interp_le_l with "[H] Hle"). by iApply iProto_interp_flip. - Qed. - - Lemma iProto_interp_send vl ml vsl vsr pr pl' : - iProto_interp vsl vsr (<!> ml) pr -∗ - iMsg_car ml vl (Next pl') -∗ - â–·^(length vsr) iProto_interp (vsl ++ [vl]) vsr pl' pr. - Proof. - iDestruct 1 as (p) "[Hp Hdp] /="; iIntros "Hml". - iDestruct (iProto_le_trans _ _ (<!> MSG vl; pl') with "Hp [Hml]") as "Hp". - { iApply iProto_le_send. rewrite iMsg_base_eq. iIntros (v' p') "(->&Hp&_) /=". - iExists p'. iSplitR; [iApply iProto_le_refl|]. by iRewrite -"Hp". } - iInduction vsr as [|vr vsr] "IH" forall (pl'); simpl. - { iExists pl'; simpl. iSplitR; [iApply iProto_le_refl|]. - iApply (iProto_le_trans with "[Hp] Hdp"). - iInduction vsl as [|vl' vsl] "IH"; simpl. - { iApply iProto_le_dual_r. rewrite iProto_dual_message iMsg_dual_base /=. - by rewrite involutive. } - iApply iProto_le_base; iIntros "!>". by iApply "IH". } - iDestruct (iProto_le_recv_send_inv _ _ vr vl - (iProto_app_recvs vsr p) pl' with "Hp [] []") as (p') "[H1 H2]"; - [rewrite iMsg_base_eq; by auto..|]. - iIntros "!>". iSpecialize ("IH" with "Hdp H1"). iIntros "!>". - iDestruct "IH" as (p'') "[Hp'' Hdp'']". iExists p''. iFrame "Hdp''". - iApply (iProto_le_trans with "[Hp''] H2"); simpl. by iApply iProto_le_base. - Qed. - - Lemma iProto_interp_recv vl vsl vsr pl mr : - iProto_interp (vl :: vsl) vsr pl (<?> mr) -∗ - ∃ pr, iMsg_car mr vl (Next pr) ∗ â–· iProto_interp vsl vsr pl pr. - Proof. - iDestruct 1 as (p) "[Hp Hdp] /=". - iDestruct (iProto_le_recv_inv with "Hdp") as (m) "[#Hm Hpr]". - iDestruct (iProto_message_equivI with "Hm") as (_) "{Hm} #Hm". - iDestruct ("Hpr" $! vl (iProto_app_recvs vsl (iProto_dual p)) with "[]") - as (pr'') "[Hler Hpr]". - { iRewrite -("Hm" $! vl (Next (iProto_app_recvs vsl (iProto_dual p)))). - rewrite iMsg_base_eq; auto. } - iExists pr''. iIntros "{$Hpr} !>". iExists p. iFrame. - Qed. - Global Instance iProto_own_ne γ s : NonExpansive (iProto_own γ s). Proof. solve_proper. Qed. Global Instance iProto_own_proper γ s : Proper ((≡) ==> (≡)) (iProto_own γ s). @@ -1121,8 +1185,7 @@ Section proto. Lemma iProto_init p : ⊢ |==> ∃ γ, - iProto_ctx γ [] [] ∗ - iProto_own γ Left p ∗ iProto_own γ Right (iProto_dual p). + iProto_ctx γ ∗ iProto_own γ Left p ∗ iProto_own γ Right (iProto_dual p). Proof. iMod (own_alloc (â—E (Next p) â‹… â—¯E (Next p))) as (lγ) "[Hâ—l Hâ—¯l]". { by apply excl_auth_valid. } @@ -1130,166 +1193,154 @@ Section proto. â—¯E (Next (iProto_dual p)))) as (rγ) "[Hâ—r Hâ—¯r]". { by apply excl_auth_valid. } pose (ProtName lγ rγ) as γ. iModIntro. iExists γ. iSplitL "Hâ—l Hâ—r". - { iExists p, (iProto_dual p). iFrame. iApply iProto_interp_nil. } + { iExists p, (iProto_dual p). iFrame. iApply iProto_consistent_dual. } iSplitL "Hâ—¯l"; iExists _; iFrame; iApply iProto_le_refl. Qed. - Lemma iProto_send_l γ m vsr vsl vl p : - iProto_ctx γ vsl vsr -∗ - iProto_own γ Left (<!> m) -∗ - iMsg_car m vl (Next p) ==∗ - â–·^(length vsr) iProto_ctx γ (vsl ++ [vl]) vsr ∗ - iProto_own γ Left p. - Proof. - iDestruct 1 as (pl pr) "(Hâ—l & Hâ—r & Hinterp)". - iDestruct 1 as (pl') "[Hle Hâ—¯]". iIntros "Hm". - iDestruct (iProto_own_auth_agree with "Hâ—l Hâ—¯") as "#Hp". - iAssert (â–· (pl ⊑ <!> m))%I - with "[Hle]" as "{Hp} Hle"; first (iNext; by iRewrite "Hp"). - iDestruct (iProto_interp_le_l with "Hinterp Hle") as "Hinterp". - iDestruct (iProto_interp_send with "Hinterp [Hm //]") as "Hinterp". - iMod (iProto_own_auth_update _ _ _ _ p with "Hâ—l Hâ—¯") as "[Hâ—l Hâ—¯]". - iIntros "!>". iSplitR "Hâ—¯". - - iIntros "!>". iExists p, pr. iFrame. - - iExists p. iFrame. iApply iProto_le_refl. - Qed. - - Lemma iProto_send_r γ m vsr vsl vr p : - iProto_ctx γ vsl vsr -∗ - iProto_own γ Right (<!> m) -∗ - iMsg_car m vr (Next p) ==∗ - â–·^(length vsl) iProto_ctx γ vsl (vsr ++ [vr]) ∗ - iProto_own γ Right p. - Proof. - iDestruct 1 as (pl pr) "(Hâ—l & Hâ—r & Hinterp)". - iDestruct 1 as (pr') "[Hle Hâ—¯]". iIntros "Hm". - iDestruct (iProto_own_auth_agree with "Hâ—r Hâ—¯") as "#Hp". - iAssert (â–· (pr ⊑ <!> m))%I - with "[Hle]" as "{Hp} Hle"; first (iNext; by iRewrite "Hp"). - iDestruct (iProto_interp_flip with "Hinterp") as "Hinterp". - iDestruct (iProto_interp_le_l with "Hinterp Hle") as "Hinterp". - iDestruct (iProto_interp_send with "Hinterp [Hm //]") as "Hinterp". - iMod (iProto_own_auth_update _ _ _ _ p with "Hâ—r Hâ—¯") as "[Hâ—r Hâ—¯]". - iIntros "!>". iSplitR "Hâ—¯". - - iIntros "!>". iExists pl, p. iFrame. by iApply iProto_interp_flip. - - iExists p. iFrame. iApply iProto_le_refl. - Qed. - - Lemma iProto_recv_l γ m vr vsr vsl : - iProto_ctx γ vsl (vr :: vsr) -∗ - iProto_own γ Left (<?> m) ==∗ - â–· ∃ p, - iProto_ctx γ vsl vsr ∗ - iProto_own γ Left p ∗ - iMsg_car m vr (Next p). - Proof. - iDestruct 1 as (pl pr) "(Hâ—l & Hâ—r & Hinterp)". - iDestruct 1 as (p) "[Hle Hâ—¯]". - iDestruct (iProto_own_auth_agree with "Hâ—l Hâ—¯") as "#Hp". - iDestruct (iProto_interp_le_l with "Hinterp [Hle]") as "Hinterp". - { iIntros "!>". by iRewrite "Hp". } - iDestruct (iProto_interp_flip with "Hinterp") as "Hinterp". - iDestruct (iProto_interp_recv with "Hinterp") as (q) "[Hm Hinterp]". - iMod (iProto_own_auth_update _ _ _ _ q with "Hâ—l Hâ—¯") as "[Hâ—l Hâ—¯]". - iIntros "!> !> /=". iExists q. iFrame "Hm". iSplitR "Hâ—¯". - - iExists q, pr. iFrame. by iApply iProto_interp_flip. - - iExists q. iIntros "{$Hâ—¯} !>". iApply iProto_le_refl. - Qed. - - Lemma iProto_recv_r γ m vl vsr vsl : - iProto_ctx γ (vl :: vsl) vsr -∗ - iProto_own γ Right (<?> m) ==∗ - â–· ∃ p, - iProto_ctx γ vsl vsr ∗ - iProto_own γ Right p ∗ - iMsg_car m vl (Next p). - Proof. - iDestruct 1 as (pl pr) "(Hâ—l & Hâ—r & Hinterp)". - iDestruct 1 as (p) "[Hle Hâ—¯]". - iDestruct (iProto_own_auth_agree with "Hâ—r Hâ—¯") as "#Hp". - iDestruct (iProto_interp_le_r with "Hinterp [Hle]") as "Hinterp". - { iIntros "!>". by iRewrite "Hp". } - iDestruct (iProto_interp_recv with "Hinterp") as (q) "[Hm Hinterp]". - iMod (iProto_own_auth_update _ _ _ _ q with "Hâ—r Hâ—¯") as "[Hâ—r Hâ—¯]". - iIntros "!> !> /=". iExists q. iFrame "Hm". iSplitR "Hâ—¯". - - iExists pl, q. iFrame. - - iExists q. iIntros "{$Hâ—¯} !>". iApply iProto_le_refl. - Qed. - - (** The instances below make it possible to use the tactics [iIntros], - [iExist], [iSplitL]/[iSplitR], [iFrame] and [iModIntro] on [iProto_le] goals. *) - Global Instance iProto_le_from_forall_l {A} a (m1 : A → iMsg Σ V) m2 name : - AsIdentName m1 name → - FromForall (iProto_message Recv (iMsg_exist m1) ⊑ (<a> m2)) - (λ x, (<?> m1 x) ⊑ (<a> m2))%I name | 10. - Proof. intros _. apply iProto_le_exist_elim_l. Qed. - Global Instance iProto_le_from_forall_r {A} a m1 (m2 : A → iMsg Σ V) name : - AsIdentName m2 name → - FromForall ((<a> m1) ⊑ iProto_message Send (iMsg_exist m2)) - (λ x, (<a> m1) ⊑ (<!> m2 x))%I name | 11. - Proof. intros _. apply iProto_le_exist_elim_r. Qed. - - Global Instance iProto_le_from_wand_l a m v P p : - TCIf (TCEq P True%I) False TCTrue → - FromWand ((<?> MSG v {{ P }}; p) ⊑ (<a> m)) P ((<?> MSG v; p) ⊑ (<a> m)) | 10. - Proof. intros _. apply iProto_le_payload_elim_l. Qed. - Global Instance iProto_le_from_wand_r a m v P p : - FromWand ((<a> m) ⊑ (<!> MSG v {{ P }}; p)) P ((<a> m) ⊑ (<!> MSG v; p)) | 11. - Proof. apply iProto_le_payload_elim_r. Qed. - - Global Instance iProto_le_from_exist_l {A} (m : A → iMsg Σ V) p : - FromExist ((<! x> m x) ⊑ p) (λ a, (<!> m a) ⊑ p)%I | 10. - Proof. - rewrite /FromExist. iDestruct 1 as (x) "H". - iApply (iProto_le_trans with "[] H"). iApply iProto_le_exist_intro_l. - Qed. - Global Instance iProto_le_from_exist_r {A} (m : A → iMsg Σ V) p : - FromExist (p ⊑ <? x> m x) (λ a, p ⊑ (<?> m a))%I | 11. - Proof. - rewrite /FromExist. iDestruct 1 as (x) "H". - iApply (iProto_le_trans with "H"). iApply iProto_le_exist_intro_r. - Qed. - - Global Instance iProto_le_from_sep_l m v P p : - FromSep ((<!> MSG v {{ P }}; p) ⊑ (<!> m)) P ((<!> MSG v; p) ⊑ (<!> m)) | 10. - Proof. - rewrite /FromSep. iIntros "[HP H]". - iApply (iProto_le_trans with "[HP] H"). by iApply iProto_le_payload_intro_l. - Qed. - Global Instance iProto_le_from_sep_r m v P p : - FromSep ((<?> m) ⊑ (<?> MSG v {{ P }}; p)) P ((<?> m) ⊑ (<?> MSG v; p)) | 11. - Proof. - rewrite /FromSep. iIntros "[HP H]". - iApply (iProto_le_trans with "H"). by iApply iProto_le_payload_intro_r. - Qed. - - Global Instance iProto_le_frame_l q m v R P Q p : - Frame q R P Q → - Frame q R ((<!> MSG v {{ P }}; p) ⊑ (<!> m)) - ((<!> MSG v {{ Q }}; p) ⊑ (<!> m)) | 10. - Proof. - rewrite /Frame /=. iIntros (HP) "[HR H]". - iApply (iProto_le_trans with "[HR] H"). iApply iProto_le_payload_elim_r. - iIntros "HQ". iApply iProto_le_payload_intro_l. iApply HP; iFrame. - Qed. - Global Instance iProto_le_frame_r q m v R P Q p : - Frame q R P Q → - Frame q R ((<?> m) ⊑ (<?> MSG v {{ P }}; p)) - ((<?> m) ⊑ (<?> MSG v {{ Q }}; p)) | 11. - Proof. - rewrite /Frame /=. iIntros (HP) "[HR H]". - iApply (iProto_le_trans with "H"). iApply iProto_le_payload_elim_l. - iIntros "HQ". iApply iProto_le_payload_intro_r. iApply HP; iFrame. - Qed. - - Global Instance iProto_le_from_modal a v p1 p2 : - FromModal True (modality_instances.modality_laterN 1) (p1 ⊑ p2) - ((<a> MSG v; p1) ⊑ (<a> MSG v; p2)) (p1 ⊑ p2). - Proof. intros _. apply iProto_le_base. Qed. + Definition side_dual s := + match s with + | Left => Right + | Right => Left + end. + + Lemma iProto_step_l γ m1 m2 p1 v : + iProto_ctx γ -∗ iProto_own γ Left (<!> m1) -∗ iProto_own γ Right (<?> m2) -∗ + iMsg_car m1 v (Next p1) ==∗ + â–· ∃ p2, iMsg_car m2 v (Next p2) ∗ â–· iProto_ctx γ ∗ + iProto_own γ Left p1 ∗ iProto_own γ Right p2. + Proof. + iDestruct 1 as (pl pr) "(Hâ—l & Hâ—r & Hconsistent)". + iDestruct 1 as (pl') "[Hlel Hâ—¯l]". + iDestruct 1 as (pr') "[Hler Hâ—¯r]". + iIntros "Hm". + iDestruct (iProto_own_auth_agree with "Hâ—l Hâ—¯l") as "#Hpl". + iDestruct (iProto_own_auth_agree with "Hâ—r Hâ—¯r") as "#Hpr". + iAssert (â–· (pl ⊑ <!> m1))%I + with "[Hlel]" as "{Hpl} Hlel"; first (iNext; by iRewrite "Hpl"). + iAssert (â–· (pr ⊑ <?> m2))%I + with "[Hler]" as "{Hpr} Hler"; first (iNext; by iRewrite "Hpr"). + iDestruct (iProto_consistent_le_l with "Hconsistent Hlel") as "Hconsistent". + iDestruct (iProto_consistent_le_r with "Hconsistent Hler") as "Hconsistent". + iDestruct (iProto_consistent_step with "Hconsistent [Hm //]") as + (p2) "[Hm2 Hconsistent]". + iMod (iProto_own_auth_update _ _ _ _ p1 with "Hâ—l Hâ—¯l") as "[Hâ—l Hâ—¯l]". + iMod (iProto_own_auth_update _ _ _ _ p2 with "Hâ—r Hâ—¯r") as "[Hâ—r Hâ—¯r]". + iIntros "!>!>". + iExists p2. iFrame. + iSplitL "Hconsistent Hâ—l Hâ—r". + - iExists _, _. iFrame. + - iSplitL "Hâ—¯l". + + iExists _. iFrame. iApply iProto_le_refl. + + iExists _. iFrame. iApply iProto_le_refl. + Qed. + + Lemma iProto_step_r γ m1 m2 p2 v : + iProto_ctx γ -∗ iProto_own γ Left (<?> m1) -∗ iProto_own γ Right (<!> m2) -∗ + iMsg_car m2 v (Next p2) ==∗ + â–· ∃ p1, iMsg_car m1 v (Next p1) ∗ â–· iProto_ctx γ ∗ + iProto_own γ Left p1 ∗ iProto_own γ Right p2. + Proof. + iDestruct 1 as (pl pr) "(Hâ—l & Hâ—r & Hconsistent)". + iDestruct 1 as (pl') "[Hlel Hâ—¯l]". + iDestruct 1 as (pr') "[Hler Hâ—¯r]". + iIntros "Hm". + iDestruct (iProto_own_auth_agree with "Hâ—l Hâ—¯l") as "#Hpl". + iDestruct (iProto_own_auth_agree with "Hâ—r Hâ—¯r") as "#Hpr". + iAssert (â–· (pl ⊑ <?> m1))%I + with "[Hlel]" as "{Hpl} Hlel"; first (iNext; by iRewrite "Hpl"). + iAssert (â–· (pr ⊑ <!> m2))%I + with "[Hler]" as "{Hpr} Hler"; first (iNext; by iRewrite "Hpr"). + iDestruct (iProto_consistent_le_l with "Hconsistent Hlel") as "Hconsistent". + iDestruct (iProto_consistent_le_r with "Hconsistent Hler") as "Hconsistent". + iDestruct (iProto_consistent_flip with "Hconsistent") as + "Hconsistent". + iDestruct (iProto_consistent_step with "Hconsistent [Hm //]") as + (p1) "[Hm1 Hconsistent]". + iMod (iProto_own_auth_update _ _ _ _ p1 with "Hâ—l Hâ—¯l") as "[Hâ—l Hâ—¯l]". + iMod (iProto_own_auth_update _ _ _ _ p2 with "Hâ—r Hâ—¯r") as "[Hâ—r Hâ—¯r]". + iIntros "!>!>". + iExists p1. iFrame. + iSplitL "Hconsistent Hâ—l Hâ—r". + - iExists _, _. iFrame. iApply iProto_consistent_flip. iFrame. + - iSplitL "Hâ—¯l". + + iExists _. iFrame. iApply iProto_le_refl. + + iExists _. iFrame. iApply iProto_le_refl. + Qed. + + (* (** The instances below make it possible to use the tactics [iIntros], *) + (* [iExist], [iSplitL]/[iSplitR], [iFrame] and [iModIntro] on [iProto_le] goals. *) *) + (* Global Instance iProto_le_from_forall_l {A} a (m1 : A → iMsg Σ V) m2 name : *) + (* AsIdentName m1 name → *) + (* FromForall (iProto_message Recv (iMsg_exist m1) ⊑ (<a> m2)) *) + (* (λ x, (<?> m1 x) ⊑ (<a> m2))%I name | 10. *) + (* Proof. intros _. apply iProto_le_exist_elim_l. Qed. *) + (* Global Instance iProto_le_from_forall_r {A} a m1 (m2 : A → iMsg Σ V) name : *) + (* AsIdentName m2 name → *) + (* FromForall ((<a> m1) ⊑ iProto_message Send (iMsg_exist m2)) *) + (* (λ x, (<a> m1) ⊑ (<!> m2 x))%I name | 11. *) + (* Proof. intros _. apply iProto_le_exist_elim_r. Qed. *) + + (* Global Instance iProto_le_from_wand_l a m v P p : *) + (* TCIf (TCEq P True%I) False TCTrue → *) + (* FromWand ((<?> MSG v {{ P }}; p) ⊑ (<a> m)) P ((<?> MSG v; p) ⊑ (<a> m)) | 10. *) + (* Proof. intros _. apply iProto_le_payload_elim_l. Qed. *) + (* Global Instance iProto_le_from_wand_r a m v P p : *) + (* FromWand ((<a> m) ⊑ (<!> MSG v {{ P }}; p)) P ((<a> m) ⊑ (<!> MSG v; p)) | 11. *) + (* Proof. apply iProto_le_payload_elim_r. Qed. *) + + (* Global Instance iProto_le_from_exist_l {A} (m : A → iMsg Σ V) p : *) + (* FromExist ((<! x> m x) ⊑ p) (λ a, (<!> m a) ⊑ p)%I | 10. *) + (* Proof. *) + (* rewrite /FromExist. iDestruct 1 as (x) "H". *) + (* iApply (iProto_le_trans with "[] H"). iApply iProto_le_exist_intro_l. *) + (* Qed. *) + (* Global Instance iProto_le_from_exist_r {A} (m : A → iMsg Σ V) p : *) + (* FromExist (p ⊑ <? x> m x) (λ a, p ⊑ (<?> m a))%I | 11. *) + (* Proof. *) + (* rewrite /FromExist. iDestruct 1 as (x) "H". *) + (* iApply (iProto_le_trans with "H"). iApply iProto_le_exist_intro_r. *) + (* Qed. *) + + (* Global Instance iProto_le_from_sep_l m v P p : *) + (* FromSep ((<!> MSG v {{ P }}; p) ⊑ (<!> m)) P ((<!> MSG v; p) ⊑ (<!> m)) | 10. *) + (* Proof. *) + (* rewrite /FromSep. iIntros "[HP H]". *) + (* iApply (iProto_le_trans with "[HP] H"). by iApply iProto_le_payload_intro_l. *) + (* Qed. *) + (* Global Instance iProto_le_from_sep_r m v P p : *) + (* FromSep ((<?> m) ⊑ (<?> MSG v {{ P }}; p)) P ((<?> m) ⊑ (<?> MSG v; p)) | 11. *) + (* Proof. *) + (* rewrite /FromSep. iIntros "[HP H]". *) + (* iApply (iProto_le_trans with "H"). by iApply iProto_le_payload_intro_r. *) + (* Qed. *) + + (* Global Instance iProto_le_frame_l q m v R P Q p : *) + (* Frame q R P Q → *) + (* Frame q R ((<!> MSG v {{ P }}; p) ⊑ (<!> m)) *) + (* ((<!> MSG v {{ Q }}; p) ⊑ (<!> m)) | 10. *) + (* Proof. *) + (* rewrite /Frame /=. iIntros (HP) "[HR H]". *) + (* iApply (iProto_le_trans with "[HR] H"). iApply iProto_le_payload_elim_r. *) + (* iIntros "HQ". iApply iProto_le_payload_intro_l. iApply HP; iFrame. *) + (* Qed. *) + (* Global Instance iProto_le_frame_r q m v R P Q p : *) + (* Frame q R P Q → *) + (* Frame q R ((<?> m) ⊑ (<?> MSG v {{ P }}; p)) *) + (* ((<?> m) ⊑ (<?> MSG v {{ Q }}; p)) | 11. *) + (* Proof. *) + (* rewrite /Frame /=. iIntros (HP) "[HR H]". *) + (* iApply (iProto_le_trans with "H"). iApply iProto_le_payload_elim_l. *) + (* iIntros "HQ". iApply iProto_le_payload_intro_r. iApply HP; iFrame. *) + (* Qed. *) + + (* Global Instance iProto_le_from_modal a v p1 p2 : *) + (* FromModal True (modality_instances.modality_laterN 1) (p1 ⊑ p2) *) + (* ((<a> MSG v; p1) ⊑ (<a> MSG v; p2)) (p1 ⊑ p2). *) + (* Proof. intros _. apply iProto_le_base. Qed. *) End proto. -Global Typeclasses Opaque iProto_ctx iProto_own. +Typeclasses Opaque iProto_ctx iProto_own. Global Hint Extern 0 (environments.envs_entails _ (?x ⊑ ?y)) => first [is_evar x; fail 1 | is_evar y; fail 1|iApply iProto_le_refl] : core. -- GitLab From eda6b39599b8471315cb0011cee48001bbb21997 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Sun, 21 Jan 2024 01:27:03 +0100 Subject: [PATCH 02/81] WIP: Multiparty definition --- _CoqProject | 2 + theories/channel/multi_proto.v | 1724 ++++++++++++++++++++++++++ theories/channel/multi_proto_model.v | 307 +++++ 3 files changed, 2033 insertions(+) create mode 100644 theories/channel/multi_proto.v create mode 100644 theories/channel/multi_proto_model.v diff --git a/_CoqProject b/_CoqProject index bdcd5f1..330aec8 100644 --- a/_CoqProject +++ b/_CoqProject @@ -16,6 +16,8 @@ theories/utils/cofe_solver_2.v theories/utils/switch.v theories/channel/proto_model.v theories/channel/proto.v +theories/channel/multi_proto_model.v +theories/channel/multi_proto.v theories/channel/channel.v theories/channel/proofmode.v theories/examples/basics.v diff --git a/theories/channel/multi_proto.v b/theories/channel/multi_proto.v new file mode 100644 index 0000000..c2fbacb --- /dev/null +++ b/theories/channel/multi_proto.v @@ -0,0 +1,1724 @@ +(** This file defines the core of the Actris logic: It defines dependent +separation protocols and the various operations on it like dual, append, and +branching. + +Dependent separation protocols [iProto] are defined by instantiating the +parameterized version in [proto_model] with the type of propositions [iProp] of Iris. +We define ways of constructing instances of the instantiated type via two +constructors: +- [iProto_end], which is identical to [proto_end]. +- [iProto_message], which takes an [action] and an [iMsg]. The type [iMsg] is a + sequence of binders [iMsg_exist], terminated by the payload constructed with + [iMsg_base] based on arguments [v], [P] and [prot], which are the value, the + carried proposition and the [iProto] tail, respectively. + +For convenience sake, we provide the following notations: +- [END], which is simply [iProto_end]. +- [∃ x, m], which is [iMsg_exist] with argument [m]. +- [MSG v {{ P }}; prot], which is [iMsg_Base] with arguments [v], [P] and [prot]. +- [<a> m], which is [iProto_message] with arguments [a] and [m]. + +We also include custom notation to more easily construct complete constructions: +- [<a x1 .. xn> m], which is [<a> ∃ x1, .. ∃ xn, m]. +- [<a x1 .. xn> MSG v; {{ P }}; prot], which constructs a full protocol. + +Futhermore, we define the following operations: +- [iProto_dual], which turns all [Send] of a protocol into [Recv] and vice-versa. +- [iProto_app], which appends two protocols. + +In addition we define the subprotocol relation [iProto_le] [⊑], which generalises +the following inductive definition for asynchronous subtyping on session types: + + p1 <: p2 p1 <: p2 p1 <: !B.p3 ?A.p3 <: p2 +---------- ---------------- ---------------- ---------------------------- +end <: end !A.p1 <: !A.p2 ?A.p1 <: ?A.p2 ?A.p1 <: !B.p2 + +Example: + +!R <: !R ?Q <: ?Q ?Q <: ?Q +------------------- -------------- +?Q.!R <: !R.?Q ?P.?Q <: ?P.?Q +------------------------------------ + ?P.?Q.!R <: !R.?P.?Q + +Lastly, relevant type classes instances are defined for each of the above +notions, such as contractiveness and non-expansiveness, after which the +specifications of the message-passing primitives are defined in terms of the +protocol connectives. *) +From iris.algebra Require Import gmap excl_auth. +From iris.proofmode Require Import proofmode. +From iris.base_logic Require Export lib.iprop. +From iris.base_logic Require Import lib.own. +From iris.program_logic Require Import language. +From actris.channel Require Import multi_proto_model. +Set Default Proof Using "Type". +Export action. + +(** * Setup of Iris's cameras *) +Class protoG Σ V := + protoG_authG :: + inG Σ (excl_authR (laterO (proto (leibnizO V) (iPropO Σ) (iPropO Σ)))). + +Definition protoΣ V := #[ + GFunctor (authRF (optionURF (exclRF (laterOF (protoOF (leibnizO V) idOF idOF))))) +]. +Global Instance subG_chanΣ {Σ V} : subG (protoΣ V) Σ → protoG Σ V. +Proof. solve_inG. Qed. + +(** * Types *) +Definition iProto Σ V := proto V (iPropO Σ) (iPropO Σ). +Declare Scope proto_scope. +Delimit Scope proto_scope with proto. +Bind Scope proto_scope with iProto. +Local Open Scope proto. + +(** * Messages *) +Section iMsg. + Set Primitive Projections. + Record iMsg Σ V := IMsg { iMsg_car : V → laterO (iProto Σ V) -n> iPropO Σ }. +End iMsg. +Arguments IMsg {_ _} _. +Arguments iMsg_car {_ _} _. + +Declare Scope msg_scope. +Delimit Scope msg_scope with msg. +Bind Scope msg_scope with iMsg. +Global Instance iMsg_inhabited {Σ V} : Inhabited (iMsg Σ V) := populate (IMsg inhabitant). + +Section imsg_ofe. + Context {Σ : gFunctors} {V : Type}. + + Instance iMsg_equiv : Equiv (iMsg Σ V) := λ m1 m2, + ∀ w p, iMsg_car m1 w p ≡ iMsg_car m2 w p. + Instance iMsg_dist : Dist (iMsg Σ V) := λ n m1 m2, + ∀ w p, iMsg_car m1 w p ≡{n}≡ iMsg_car m2 w p. + + Lemma iMsg_ofe_mixin : OfeMixin (iMsg Σ V). + Proof. by apply (iso_ofe_mixin (iMsg_car : _ → V -d> _ -n> _)). Qed. + Canonical Structure iMsgO := Ofe (iMsg Σ V) iMsg_ofe_mixin. + + Global Instance iMsg_cofe : Cofe iMsgO. + Proof. by apply (iso_cofe (IMsg : (V -d> _ -n> _) → _) iMsg_car). Qed. +End imsg_ofe. + +Program Definition iMsg_base_def {Σ V} + (v : V) (P : iProp Σ) (p : iProto Σ V) : iMsg Σ V := + IMsg (λ v', λne p', ⌜ v = v' ⌠∗ Next p ≡ p' ∗ P)%I. +Next Obligation. solve_proper. Qed. +Definition iMsg_base_aux : seal (@iMsg_base_def). by eexists. Qed. +Definition iMsg_base := iMsg_base_aux.(unseal). +Definition iMsg_base_eq : @iMsg_base = @iMsg_base_def := iMsg_base_aux.(seal_eq). +Arguments iMsg_base {_ _} _%V _%I _%proto. +Global Instance: Params (@iMsg_base) 3 := {}. + +Program Definition iMsg_exist_def {Σ V A} (m : A → iMsg Σ V) : iMsg Σ V := + IMsg (λ v', λne p', ∃ x, iMsg_car (m x) v' p')%I. +Next Obligation. solve_proper. Qed. +Definition iMsg_exist_aux : seal (@iMsg_exist_def). by eexists. Qed. +Definition iMsg_exist := iMsg_exist_aux.(unseal). +Definition iMsg_exist_eq : @iMsg_exist = @iMsg_exist_def := iMsg_exist_aux.(seal_eq). +Arguments iMsg_exist {_ _ _} _%msg. +Global Instance: Params (@iMsg_exist) 3 := {}. + +Definition iMsg_texist {Σ V} {TT : tele} (m : TT → iMsg Σ V) : iMsg Σ V := + tele_fold (@iMsg_exist Σ V) (λ x, x) (tele_bind m). +Arguments iMsg_texist {_ _ !_} _%msg /. + +Notation "'MSG' v {{ P } } ; p" := (iMsg_base v P p) + (at level 200, v at level 20, right associativity, + format "MSG v {{ P } } ; p") : msg_scope. +Notation "'MSG' v ; p" := (iMsg_base v True p) + (at level 200, v at level 20, right associativity, + format "MSG v ; p") : msg_scope. +Notation "∃ x .. y , m" := + (iMsg_exist (λ x, .. (iMsg_exist (λ y, m)) ..)%msg) : msg_scope. +Notation "'∃..' x .. y , m" := + (iMsg_texist (λ x, .. (iMsg_texist (λ y, m)) .. )%msg) + (at level 200, x binder, y binder, right associativity, + format "∃.. x .. y , m") : msg_scope. + +Lemma iMsg_texist_exist {Σ V} {TT : tele} w lp (m : TT → iMsg Σ V) : + iMsg_car (∃.. x, m x)%msg w lp ⊣⊢ (∃.. x, iMsg_car (m x) w lp). +Proof. + rewrite /iMsg_texist iMsg_exist_eq. + induction TT as [|T TT IH]; simpl; [done|]. f_equiv=> x. apply IH. +Qed. + +(** * Operators *) +Definition iProto_end_def {Σ V} : iProto Σ V := proto_end. +Definition iProto_end_aux : seal (@iProto_end_def). by eexists. Qed. +Definition iProto_end := iProto_end_aux.(unseal). +Definition iProto_end_eq : @iProto_end = @iProto_end_def := iProto_end_aux.(seal_eq). +Arguments iProto_end {_ _}. + +Definition iProto_message_def {Σ V} (a : action) (m : iMsg Σ V) : iProto Σ V := + proto_message a (iMsg_car m). +Definition iProto_message_aux : seal (@iProto_message_def). by eexists. Qed. +Definition iProto_message := iProto_message_aux.(unseal). +Definition iProto_message_eq : + @iProto_message = @iProto_message_def := iProto_message_aux.(seal_eq). +Arguments iProto_message {_ _} _ _%msg. +Global Instance: Params (@iProto_message) 3 := {}. + +Notation "'END'" := iProto_end : proto_scope. + +Notation "< a > m" := (iProto_message a m) + (at level 200, a at level 10, m at level 200, + format "< a > m") : proto_scope. +Notation "< a @ x1 .. xn > m" := (iProto_message a (∃ x1, .. (∃ xn, m) ..)) + (at level 200, a at level 10, x1 closed binder, xn closed binder, m at level 200, + format "< a @ x1 .. xn > m") : proto_scope. +Notation "< a @.. x1 .. xn > m" := (iProto_message a (∃.. x1, .. (∃.. xn, m) ..)) + (at level 200, a at level 10, x1 closed binder, xn closed binder, m at level 200, + format "< a @.. x1 .. xn > m") : proto_scope. + +Notation "<!> m" := (<Send> m) (at level 200, m at level 200) : proto_scope. +Notation "<! x1 .. xn > m" := (<!> ∃ x1, .. (∃ xn, m) ..) + (at level 200, x1 closed binder, xn closed binder, m at level 200, + format "<! x1 .. xn > m") : proto_scope. +Notation "<!.. x1 .. xn > m" := (<!> ∃.. x1, .. (∃.. xn, m) ..) + (at level 200, x1 closed binder, xn closed binder, m at level 200, + format "<!.. x1 .. xn > m") : proto_scope. + +Notation "<?> m" := (<Recv> m) (at level 200, m at level 200) : proto_scope. +Notation "<? x1 .. xn > m" := (<?> ∃ x1, .. (∃ xn, m) ..) + (at level 200, x1 closed binder, xn closed binder, m at level 200, + format "<? x1 .. xn > m") : proto_scope. +Notation "<?.. x1 .. xn > m" := (<?> ∃.. x1, .. (∃.. xn, m) ..) + (at level 200, x1 closed binder, xn closed binder, m at level 200, + format "<?.. x1 .. xn > m") : proto_scope. + +Class MsgTele {Σ V} {TT : tele} (m : iMsg Σ V) + (tv : TT -t> V) (tP : TT -t> iProp Σ) (tp : TT -t> iProto Σ V) := + msg_tele : m ≡ (∃.. x, MSG tele_app tv x {{ tele_app tP x }}; tele_app tp x)%msg. +Global Hint Mode MsgTele ! ! - ! - - - : typeclass_instances. + +(** * Operations *) +Program Definition iMsg_map {Σ V} + (rec : iProto Σ V → iProto Σ V) (m : iMsg Σ V) : iMsg Σ V := + IMsg (λ v, λne p1', ∃ p1, iMsg_car m v (Next p1) ∗ p1' ≡ Next (rec p1))%I. +Next Obligation. solve_proper. Qed. + +Program Definition iProto_map_app_aux {Σ V} + (f : action → action) (p2 : iProto Σ V) + (rec : iProto Σ V -n> iProto Σ V) : iProto Σ V -n> iProto Σ V := λne p, + proto_elim p2 (λ a m, + proto_message (f a) (iMsg_car (iMsg_map rec (IMsg m)))) p. +Next Obligation. + intros Σ V f p2 rec n p1 p1' Hp. apply proto_elim_ne=> // a m1 m2 Hm. + apply proto_message_ne=> v p' /=. by repeat f_equiv. +Qed. + +Global Instance iProto_map_app_aux_contractive {Σ V} f (p2 : iProto Σ V) : + Contractive (iProto_map_app_aux f p2). +Proof. + intros n rec1 rec2 Hrec p1; simpl. apply proto_elim_ne=> // a m1 m2 Hm. + apply proto_message_ne=> v p' /=. by repeat (f_contractive || f_equiv). +Qed. + +Definition iProto_map_app {Σ V} (f : action → action) + (p2 : iProto Σ V) : iProto Σ V -n> iProto Σ V := + fixpoint (iProto_map_app_aux f p2). + +Definition iProto_app_def {Σ V} (p1 p2 : iProto Σ V) : iProto Σ V := + iProto_map_app id p2 p1. +Definition iProto_app_aux : seal (@iProto_app_def). Proof. by eexists. Qed. +Definition iProto_app := iProto_app_aux.(unseal). +Definition iProto_app_eq : @iProto_app = @iProto_app_def := iProto_app_aux.(seal_eq). +Arguments iProto_app {_ _} _%proto _%proto. +Global Instance: Params (@iProto_app) 2 := {}. +Infix "<++>" := iProto_app (at level 60) : proto_scope. +Notation "m <++> p" := (iMsg_map (flip iProto_app p) m) : msg_scope. + +Definition iProto_dual_def {Σ V} (p : iProto Σ V) : iProto Σ V := + iProto_map_app action_dual proto_end p. +Definition iProto_dual_aux : seal (@iProto_dual_def). Proof. by eexists. Qed. +Definition iProto_dual := iProto_dual_aux.(unseal). +Definition iProto_dual_eq : + @iProto_dual = @iProto_dual_def := iProto_dual_aux.(seal_eq). +Arguments iProto_dual {_ _} _%proto. +Global Instance: Params (@iProto_dual) 2 := {}. +Notation iMsg_dual := (iMsg_map iProto_dual). + +Definition iProto_dual_if {Σ V} (d : bool) (p : iProto Σ V) : iProto Σ V := + if d then iProto_dual p else p. +Arguments iProto_dual_if {_ _} _ _%proto. +Global Instance: Params (@iProto_dual_if) 3 := {}. + +(** * Proofs *) +Section proto. + Context `{!protoG Σ V}. + Implicit Types v : V. + Implicit Types p pl pr : iProto Σ V. + Implicit Types m : iMsg Σ V. + + (** ** Equality *) + Lemma iProto_case p : p ≡ END ∨ ∃ a m, p ≡ <a> m. + Proof. + rewrite iProto_message_eq iProto_end_eq. + destruct (proto_case p) as [|(a&m&?)]; [by left|right]. + by exists a, (IMsg m). + Qed. + Lemma iProto_message_equivI `{!BiInternalEq SPROP} a1 a2 m1 m2 : + (<a1> m1) ≡ (<a2> m2) ⊣⊢@{SPROP} ⌜ a1 = a2 ⌠∧ + (∀ v lp, iMsg_car m1 v lp ≡ iMsg_car m2 v lp). + Proof. rewrite iProto_message_eq. apply proto_message_equivI. Qed. + + Lemma iProto_message_end_equivI `{!BiInternalEq SPROP} a m : + (<a> m) ≡ END ⊢@{SPROP} False. + Proof. rewrite iProto_message_eq iProto_end_eq. apply proto_message_end_equivI. Qed. + Lemma iProto_end_message_equivI `{!BiInternalEq SPROP} a m : + END ≡ (<a> m) ⊢@{SPROP} False. + Proof. by rewrite internal_eq_sym iProto_message_end_equivI. Qed. + + (** ** Non-expansiveness of operators *) + Global Instance iMsg_car_proper : + Proper (iMsg_equiv ==> (=) ==> (≡) ==> (≡)) (iMsg_car (Σ:=Σ) (V:=V)). + Proof. + intros m1 m2 meq v1 v2 veq p1 p2 peq. rewrite meq. + f_equiv; [ by f_equiv | done ]. + Qed. + Global Instance iMsg_car_ne n : + Proper (iMsg_dist n ==> (=) ==> (dist n) ==> (dist n)) (iMsg_car (Σ:=Σ) (V:=V)). + Proof. + intros m1 m2 meq v1 v2 veq p1 p2 peq. rewrite meq. + f_equiv; [ by f_equiv | done ]. + Qed. + + Global Instance iMsg_contractive v n : + Proper (dist n ==> dist_later n ==> dist n) (iMsg_base (Σ:=Σ) (V:=V) v). + Proof. rewrite iMsg_base_eq=> P1 P2 HP p1 p2 Hp w q /=. solve_contractive. Qed. + Global Instance iMsg_ne v : NonExpansive2 (iMsg_base (Σ:=Σ) (V:=V) v). + Proof. rewrite iMsg_base_eq=> P1 P2 HP p1 p2 Hp w q /=. solve_proper. Qed. + Global Instance iMsg_proper v : + Proper ((≡) ==> (≡) ==> (≡)) (iMsg_base (Σ:=Σ) (V:=V) v). + Proof. apply (ne_proper_2 _). Qed. + + Global Instance iMsg_exist_ne A n : + Proper (pointwise_relation _ (dist n) ==> (dist n)) (@iMsg_exist Σ V A). + Proof. rewrite iMsg_exist_eq=> m1 m2 Hm v p /=. f_equiv=> x. apply Hm. Qed. + Global Instance iMsg_exist_proper A : + Proper (pointwise_relation _ (≡) ==> (≡)) (@iMsg_exist Σ V A). + Proof. rewrite iMsg_exist_eq=> m1 m2 Hm v p /=. f_equiv=> x. apply Hm. Qed. + + Global Instance msg_tele_base (v:V) (P : iProp Σ) (p : iProto Σ V) : + MsgTele (TT:=TeleO) (MSG v {{ P }}; p) v P p. + Proof. done. Qed. + Global Instance msg_tele_exist {A} {TT : A → tele} (m : A → iMsg Σ V) tv tP tp : + (∀ x, MsgTele (TT:=TT x) (m x) (tv x) (tP x) (tp x)) → + MsgTele (TT:=TeleS TT) (∃ x, m x) tv tP tp. + Proof. intros Hm. rewrite /MsgTele /=. f_equiv=> x. apply Hm. Qed. + + Global Instance iProto_message_ne a : + NonExpansive (iProto_message (Σ:=Σ) (V:=V) a). + Proof. rewrite iProto_message_eq. solve_proper. Qed. + Global Instance iProto_message_proper a : + Proper ((≡) ==> (≡)) (iProto_message (Σ:=Σ) (V:=V) a). + Proof. apply (ne_proper _). Qed. + + Lemma iProto_message_equiv {TT1 TT2 : tele} a1 a2 + (m1 m2 : iMsg Σ V) + (v1 : TT1 -t> V) (v2 : TT2 -t> V) + (P1 : TT1 -t> iProp Σ) (P2 : TT2 -t> iProp Σ) + (prot1 : TT1 -t> iProto Σ V) (prot2 : TT2 -t> iProto Σ V) : + MsgTele m1 v1 P1 prot1 → + MsgTele m2 v2 P2 prot2 → + ⌜ a1 = a2 ⌠-∗ + (■∀.. (xs1 : TT1), tele_app P1 xs1 -∗ + ∃.. (xs2 : TT2), ⌜tele_app v1 xs1 = tele_app v2 xs2⌠∗ + â–· (tele_app prot1 xs1 ≡ tele_app prot2 xs2) ∗ + tele_app P2 xs2) -∗ + (■∀.. (xs2 : TT2), tele_app P2 xs2 -∗ + ∃.. (xs1 : TT1), ⌜tele_app v1 xs1 = tele_app v2 xs2⌠∗ + â–· (tele_app prot1 xs1 ≡ tele_app prot2 xs2) ∗ + tele_app P1 xs1) -∗ + (<a1> m1) ≡ (<a2> m2). + Proof. + iIntros (Hm1 Hm2 Heq) "#Heq1 #Heq2". + unfold MsgTele in Hm1. rewrite Hm1. clear Hm1. + unfold MsgTele in Hm2. rewrite Hm2. clear Hm2. + rewrite iProto_message_eq proto_message_equivI. + iSplit; [ done | ]. + iIntros (v p'). + do 2 rewrite iMsg_texist_exist. + rewrite iMsg_base_eq /=. + iApply prop_ext. + iIntros "!>". iSplit. + - iDestruct 1 as (xs1 Hveq1) "[Hrec1 HP1]". + iDestruct ("Heq1" with "HP1") as (xs2 Hveq2) "[Hrec2 HP2]". + iExists xs2. rewrite -Hveq1 Hveq2. + iSplitR; [ done | ]. iSplitR "HP2"; [ | done ]. + iRewrite -"Hrec1". iApply later_equivI. iIntros "!>". by iRewrite "Hrec2". + - iDestruct 1 as (xs2 Hveq2) "[Hrec2 HP2]". + iDestruct ("Heq2" with "HP2") as (xs1 Hveq1) "[Hrec1 HP1]". + iExists xs1. rewrite -Hveq2 Hveq1. + iSplitR; [ done | ]. iSplitR "HP1"; [ | done ]. + iRewrite -"Hrec2". iApply later_equivI. iIntros "!>". by iRewrite "Hrec1". + Qed. + + (** Helpers *) + Lemma iMsg_map_base f v P p : + NonExpansive f → + iMsg_map f (MSG v {{ P }}; p) ≡ (MSG v {{ P }}; f p)%msg. + Proof. + rewrite iMsg_base_eq. intros ? v' p'; simpl. iSplit. + - iDestruct 1 as (p'') "[(->&Hp&$) Hp']". iSplit; [done|]. + iRewrite "Hp'". iIntros "!>". by iRewrite "Hp". + - iIntros "(->&Hp'&$)". iExists p. iRewrite -"Hp'". auto. + Qed. + Lemma iMsg_map_exist {A} f (m : A → iMsg Σ V) : + iMsg_map f (∃ x, m x) ≡ (∃ x, iMsg_map f (m x))%msg. + Proof. + rewrite iMsg_exist_eq. intros v' p'; simpl. iSplit. + - iDestruct 1 as (p'') "[H Hp']". iDestruct "H" as (x) "H"; auto. + - iDestruct 1 as (x p'') "[Hm Hp']". auto. + Qed. + + (** ** Dual *) + Global Instance iProto_dual_ne : NonExpansive (@iProto_dual Σ V). + Proof. rewrite iProto_dual_eq. solve_proper. Qed. + Global Instance iProto_dual_proper : Proper ((≡) ==> (≡)) (@iProto_dual Σ V). + Proof. apply (ne_proper _). Qed. + Global Instance iProto_dual_if_ne d : NonExpansive (@iProto_dual_if Σ V d). + Proof. solve_proper. Qed. + Global Instance iProto_dual_if_proper d : + Proper ((≡) ==> (≡)) (@iProto_dual_if Σ V d). + Proof. apply (ne_proper _). Qed. + + Lemma iProto_dual_end : iProto_dual (Σ:=Σ) (V:=V) END ≡ END. + Proof. + rewrite iProto_end_eq iProto_dual_eq /iProto_dual_def /iProto_map_app. + etrans; [apply (fixpoint_unfold (iProto_map_app_aux _ _))|]; simpl. + by rewrite proto_elim_end. + Qed. + Lemma iProto_dual_message a m : + iProto_dual (<a> m) ≡ <action_dual a> iMsg_dual m. + Proof. + rewrite iProto_message_eq iProto_dual_eq /iProto_dual_def /iProto_map_app. + etrans; [apply (fixpoint_unfold (iProto_map_app_aux _ _))|]; simpl. + rewrite /iProto_message_def. rewrite ->proto_elim_message; [done|]. + intros a' m1 m2 Hm; f_equiv; solve_proper. + Qed. + Lemma iMsg_dual_base v P p : + iMsg_dual (MSG v {{ P }}; p) ≡ (MSG v {{ P }}; iProto_dual p)%msg. + Proof. apply iMsg_map_base, _. Qed. + Lemma iMsg_dual_exist {A} (m : A → iMsg Σ V) : + iMsg_dual (∃ x, m x) ≡ (∃ x, iMsg_dual (m x))%msg. + Proof. apply iMsg_map_exist. Qed. + + Global Instance iProto_dual_involutive : Involutive (≡) (@iProto_dual Σ V). + Proof. + intros p. apply (uPred.internal_eq_soundness (M:=iResUR Σ)). + iLöb as "IH" forall (p). destruct (iProto_case p) as [->|(a&m&->)]. + { by rewrite !iProto_dual_end. } + rewrite !iProto_dual_message involutive. + iApply iProto_message_equivI; iSplit; [done|]; iIntros (v p') "/=". + iApply prop_ext; iIntros "!>"; iSplit. + - iDestruct 1 as (pd) "[H Hp']". iRewrite "Hp'". + iDestruct "H" as (pdd) "[H #Hpd]". + iApply (internal_eq_rewrite); [|done]; iIntros "!>". + iRewrite "Hpd". by iRewrite ("IH" $! pdd). + - iIntros "H". destruct (Next_uninj p') as [p'' Hp']. iExists _. + rewrite Hp'. iSplitL; [by auto|]. iIntros "!>". by iRewrite ("IH" $! p''). + Qed. + + (** ** Append *) + Global Instance iProto_app_end_l : LeftId (≡) END (@iProto_app Σ V). + Proof. + intros p. rewrite iProto_end_eq iProto_app_eq /iProto_app_def /iProto_map_app. + etrans; [apply (fixpoint_unfold (iProto_map_app_aux _ _))|]; simpl. + by rewrite proto_elim_end. + Qed. + Lemma iProto_app_message a m p2 : (<a> m) <++> p2 ≡ <a> m <++> p2. + Proof. + rewrite iProto_message_eq iProto_app_eq /iProto_app_def /iProto_map_app. + etrans; [apply (fixpoint_unfold (iProto_map_app_aux _ _))|]; simpl. + rewrite /iProto_message_def. rewrite ->proto_elim_message; [done|]. + intros a' m1 m2 Hm. f_equiv; solve_proper. + Qed. + + Global Instance iProto_app_ne : NonExpansive2 (@iProto_app Σ V). + Proof. + assert (∀ n, Proper (dist n ==> (=) ==> dist n) (@iProto_app Σ V)). + { intros n p1 p1' Hp1 p2 p2' <-. by rewrite iProto_app_eq /iProto_app_def Hp1. } + assert (Proper ((≡) ==> (=) ==> (≡)) (@iProto_app Σ V)). + { intros p1 p1' Hp1 p2 p2' <-. by rewrite iProto_app_eq /iProto_app_def Hp1. } + intros n p1 p1' Hp1 p2 p2' Hp2. rewrite Hp1. clear p1 Hp1. + revert p1'. induction (lt_wf n) as [n _ IH]; intros p1. + destruct (iProto_case p1) as [->|(a&m&->)]. + { by rewrite !left_id. } + rewrite !iProto_app_message. f_equiv=> v p' /=. do 4 f_equiv. + f_contractive. apply IH; eauto using dist_le with lia. + Qed. + Global Instance iProto_app_proper : Proper ((≡) ==> (≡) ==> (≡)) (@iProto_app Σ V). + Proof. apply (ne_proper_2 _). Qed. + + Lemma iMsg_app_base v P p1 p2 : + ((MSG v {{ P }}; p1) <++> p2)%msg ≡ (MSG v {{ P }}; p1 <++> p2)%msg. + Proof. apply: iMsg_map_base. Qed. + Lemma iMsg_app_exist {A} (m : A → iMsg Σ V) p2 : + ((∃ x, m x) <++> p2)%msg ≡ (∃ x, m x <++> p2)%msg. + Proof. apply iMsg_map_exist. Qed. + + Global Instance iProto_app_end_r : RightId (≡) END (@iProto_app Σ V). + Proof. + intros p. apply (uPred.internal_eq_soundness (M:=iResUR Σ)). + iLöb as "IH" forall (p). destruct (iProto_case p) as [->|(a&m&->)]. + { by rewrite left_id. } + rewrite iProto_app_message. + iApply iProto_message_equivI; iSplit; [done|]; iIntros (v p') "/=". + iApply prop_ext; iIntros "!>". iSplit. + - iDestruct 1 as (p1') "[H Hp']". iRewrite "Hp'". + iApply (internal_eq_rewrite); [|done]; iIntros "!>". + by iRewrite ("IH" $! p1'). + - iIntros "H". destruct (Next_uninj p') as [p'' Hp']. iExists p''. + rewrite Hp'. iSplitL; [by auto|]. iIntros "!>". by iRewrite ("IH" $! p''). + Qed. + Global Instance iProto_app_assoc : Assoc (≡) (@iProto_app Σ V). + Proof. + intros p1 p2 p3. apply (uPred.internal_eq_soundness (M:=iResUR Σ)). + iLöb as "IH" forall (p1). destruct (iProto_case p1) as [->|(a&m&->)]. + { by rewrite !left_id. } + rewrite !iProto_app_message. + iApply iProto_message_equivI; iSplit; [done|]; iIntros (v p123) "/=". + iApply prop_ext; iIntros "!>". iSplit. + - iDestruct 1 as (p1') "[H #Hp']". + iExists (p1' <++> p2). iSplitL; [by auto|]. + iRewrite "Hp'". iIntros "!>". iApply "IH". + - iDestruct 1 as (p12) "[H #Hp123]". iDestruct "H" as (p1') "[H #Hp12]". + iExists p1'. iFrame "H". iRewrite "Hp123". + iIntros "!>". iRewrite "Hp12". by iRewrite ("IH" $! p1'). + Qed. + + Lemma iProto_dual_app p1 p2 : + iProto_dual (p1 <++> p2) ≡ iProto_dual p1 <++> iProto_dual p2. + Proof. + apply (uPred.internal_eq_soundness (M:=iResUR Σ)). + iLöb as "IH" forall (p1 p2). destruct (iProto_case p1) as [->|(a&m&->)]. + { by rewrite iProto_dual_end !left_id. } + rewrite iProto_dual_message !iProto_app_message iProto_dual_message /=. + iApply iProto_message_equivI; iSplit; [done|]; iIntros (v p12) "/=". + iApply prop_ext; iIntros "!>". iSplit. + - iDestruct 1 as (p12d) "[H #Hp12]". iDestruct "H" as (p1') "[H #Hp12d]". + iExists (iProto_dual p1'). iSplitL; [by auto|]. + iRewrite "Hp12". iIntros "!>". iRewrite "Hp12d". iApply "IH". + - iDestruct 1 as (p1') "[H #Hp12]". iDestruct "H" as (p1d) "[H #Hp1']". + iExists (p1d <++> p2). iSplitL; [by auto|]. + iRewrite "Hp12". iIntros "!>". iRewrite "Hp1'". by iRewrite ("IH" $! p1d p2). + Qed. + +End proto. + +Definition can_step {Σ V} (rec : gmap nat (iProto Σ V) → iProp Σ) + (ps : gmap nat (iProto Σ V)) (i j : nat) : iProp Σ := + (∀ a1 a2 m1 m2, + (ps !!! i ≡ <a1> m1) -∗ (ps !!! j ≡ <a2> m2) -∗ + match a1,a2 with + | Send j, Recv i => ∀ v p1, iMsg_car m1 v (Next p1) -∗ + ∃ p2, iMsg_car m2 v (Next p2) ∗ + â–· (rec (<[i:=p1]>(<[j:=p2]>ps))) + | Recv j, Send i => ∀ v p2, iMsg_car m2 v (Next p2) -∗ + ∃ p1, iMsg_car m1 v (Next p1) ∗ + â–· (rec (<[i:=p1]>(<[j:=p2]>ps))) + | _, _ => True + end). + +Definition iProto_consistent_pre {Σ V} (rec : gmap nat (iProto Σ V) → iProp Σ) + (ps : gmap nat (iProto Σ V)) : iProp Σ := + ∀ i j, ⌜i ∈ dom ps⌠-∗ ⌜j ∈ dom ps⌠-∗ can_step rec ps i j. + +Global Instance iProto_consistent_pre_ne {Σ V} + (rec : gmapO natO (iProto Σ V) -n> iPropO Σ) : + NonExpansive (iProto_consistent_pre (λ ps, rec ps)). +Proof. Admitted. + +Program Definition iProto_consistent_pre' {Σ V} + (rec : gmapO natO (iProto Σ V) -n> iPropO Σ) : + gmapO natO (iProto Σ V) -n> iPropO Σ := + λne ps, iProto_consistent_pre (λ ps, rec ps) ps. + +Local Instance iProto_consistent_pre_contractive {Σ V} : Contractive (@iProto_consistent_pre' Σ V). +Proof. Admitted. + +Definition iProto_consistent {Σ V} (ps : gmap nat (iProto Σ V)) : iProp Σ := + fixpoint iProto_consistent_pre' ps. + +Arguments iProto_consistent {_ _} _%proto. +Global Instance: Params (@iProto_consistent) 1 := {}. + +Global Instance iProto_consistent_ne {Σ V} : NonExpansive (@iProto_consistent Σ V). +Proof. solve_proper. Qed. +Global Instance iProto_consistent_proper {Σ V} : Proper ((≡) ==> (⊣⊢)) (@iProto_consistent Σ V). +Proof. solve_proper. Qed. + +Lemma iProto_consistent_unfold {Σ V} (ps : gmap nat (iProto Σ V)) : + iProto_consistent ps ≡ (iProto_consistent_pre iProto_consistent) ps. +Proof. + apply: (fixpoint_unfold iProto_consistent_pre'). +Qed. +(* Definition iProto_le {Σ V} (p1 p2 : iProto Σ V) : iProp Σ := *) +(* ∀ p3, iProto_consistent p1 p3 -∗ iProto_consistent p2 p3. *) +(* Arguments iProto_le {_ _} _%proto _%proto. *) +(* Global Instance: Params (@iProto_le) 2 := {}. *) +(* Notation "p ⊑ q" := (iProto_le p q) : bi_scope. *) + +(* Global Instance iProto_le_ne {Σ V} : NonExpansive2 (@iProto_le Σ V). *) +(* Proof. solve_proper. Qed. *) +(* Global Instance iProto_le_proper {Σ V} : Proper ((≡) ==> (≡) ==> (⊣⊢)) (@iProto_le Σ V). *) +(* Proof. solve_proper. Qed. *) + +Definition iProto_example1 {Σ V} : gmap nat (iProto Σ V) := + ∅. + +Lemma iProto_example1_consistent {Σ V} : + ⊢ iProto_consistent (@iProto_example1 Σ V). +Proof. + rewrite iProto_consistent_unfold. + iIntros (i j Hi Hj). set_solver. +Qed. + +Definition iProto_example2 `{!invGS Σ} (P : iProp Σ) : gmap nat (iProto Σ Z) := + <[0 := (<(Send 1) @ (x:Z)> MSG x {{ P }} ; END)%proto ]> + (<[1 := (<(Recv 0) @ (x:Z)> MSG x {{ P }} ; END)%proto ]> + ∅). + +Lemma iProto_example2_consistent `{!invGS Σ} (P : iProp Σ) : + ⊢ iProto_consistent (@iProto_example2 _ Σ invGS0 P). +Proof. + rewrite iProto_consistent_unfold. + rewrite /iProto_example2. + iIntros (i j Hi Hj). + iIntros (a1 a2 m1 m2) "Hm1 Hm2". + destruct i, j. + - rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + iDestruct "Hm1" as (<-) "#Hm1". + iDestruct "Hm2" as (<-) "#Hm2". + done. + - destruct j; [|set_solver]. + rewrite !lookup_total_insert. + rewrite lookup_total_insert_ne; [|done]. + rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + iDestruct "Hm1" as (<-) "#Hm1". + iDestruct "Hm2" as (<-) "#Hm2". + iIntros (v p1) "Hm1'". + iSpecialize ("Hm1" $!v (Next p1)). + rewrite iMsg_base_eq. simpl. + rewrite iMsg_exist_eq. simpl. + iRewrite -"Hm1" in "Hm1'". + iExists END. + iSpecialize ("Hm2" $!v (Next END)). + iRewrite -"Hm2". + simpl. + iDestruct "Hm1'" as (x Heq) "[#Heq HP]". + iSplitL. + { iExists x. iSplit; [done|]. iFrame. done. } + iNext. iRewrite -"Heq". + rewrite insert_commute; [|done]. + rewrite insert_insert. + rewrite insert_commute; [|done]. + rewrite insert_insert. + rewrite iProto_consistent_unfold. + iIntros (i' j' Hi' Hj'). + iIntros (a1 a2 m1' m2') "Hm1' Hm2'". + destruct i'. + { rewrite lookup_total_insert. + iDestruct (iProto_end_message_equivI with "Hm1'") as "H". + done. } + destruct i'. + { rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert. + iDestruct (iProto_end_message_equivI with "Hm1'") as "H". + done. } + set_solver. + - destruct i; [|set_solver]. + rewrite !lookup_total_insert. + rewrite lookup_total_insert_ne; [|done]. + rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + iDestruct "Hm1" as (<-) "#Hm1". + iDestruct "Hm2" as (<-) "#Hm2". + iIntros (v p1) "Hm1'". + iSpecialize ("Hm2" $!v (Next p1)). + rewrite iMsg_base_eq. simpl. + rewrite iMsg_exist_eq. simpl. + iRewrite -"Hm2" in "Hm1'". + iExists END. + iSpecialize ("Hm1" $!v (Next END)). + iRewrite -"Hm1". + simpl. + iDestruct "Hm1'" as (x Heq) "[#Heq HP]". + iSplitL. + { iExists x. iSplit; [done|]. iFrame. done. } + iNext. iRewrite -"Heq". + rewrite insert_insert. + rewrite insert_commute; [|done]. + rewrite insert_insert. + rewrite iProto_consistent_unfold. + iIntros (i' j' Hi' Hj'). + iIntros (a1 a2 m1' m2') "Hm1' Hm2'". + destruct i'. + { rewrite lookup_total_insert. + iDestruct (iProto_end_message_equivI with "Hm1'") as "H". + done. } + destruct i'. + { rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert. + iDestruct (iProto_end_message_equivI with "Hm1'") as "H". + done. } + set_solver. + - destruct i; [|set_solver]. + destruct j; [|set_solver]. + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert. + rewrite !iProto_message_equivI. + iDestruct "Hm1" as (<-) "Hm1". + iDestruct "Hm2" as (<-) "Hm2". + done. +Qed. + +(* Lemma iProto_consistent_step {Σ V} (ps : gmap nat (iProto Σ V)) : *) +(* (∀ i j m1 m2, *) +(* ⌜ps !! i = Some (<(Send j)> m1)%proto⌠-∗ *) +(* ⌜ps !! j = Some (<(Recv i)> m2)⌠-∗ *) +(* iProto_consistent (ps))%I -∗ *) +(* iProto_consistent ps. *) + +Record proto_name := ProtName { proto_l_name : gname; proto_r_name : gname }. +Global Instance proto_name_inhabited : Inhabited proto_name := + populate (ProtName inhabitant inhabitant). +Global Instance proto_name_eq_dec : EqDecision proto_name. +Proof. solve_decision. Qed. +Global Instance proto_name_countable : Countable proto_name. +Proof. + refine (inj_countable (λ '(ProtName γl γr), (γl,γr)) + (λ '(γl, γr), Some (ProtName γl γr)) _); by intros []. +Qed. + +Inductive side := Left | Right. +Global Instance side_inhabited : Inhabited side := populate Left. +Global Instance side_eq_dec : EqDecision side. +Proof. solve_decision. Qed. +Global Instance side_countable : Countable side. +Proof. + refine (inj_countable (λ s, if s is Left then true else false) + (λ b, Some (if b then Left else Right)) _); by intros []. +Qed. +Definition side_elim {A} (s : side) (l r : A) : A := + match s with Left => l | Right => r end. + +Definition iProto_own_frag `{!protoG Σ V} (γ : proto_name) (s : side) + (p : iProto Σ V) : iProp Σ := + own (side_elim s proto_l_name proto_r_name γ) (â—¯E (Next p)). +Definition iProto_own_auth `{!protoG Σ V} (γ : proto_name) (s : side) + (p : iProto Σ V) : iProp Σ := + own (side_elim s proto_l_name proto_r_name γ) (â—E (Next p)). + +(* Definition iProto_ctx `{protoG Σ V} *) +(* (γ : proto_name) : iProp Σ := *) +(* ∃ ps, *) +(* iProto_own_auth γ Left pl ∗ *) +(* iProto_own_auth γ Right pr ∗ *) +(* â–· iProto_consistent ps. *) + +(* (** * The connective for ownership of channel ends *) *) +(* Definition iProto_own `{!protoG Σ V} *) +(* (γ : proto_name) (s : side) (p : iProto Σ V) : iProp Σ := *) +(* ∃ p', â–· (p' ⊑ p) ∗ iProto_own_frag γ s p'. *) +(* Arguments iProto_own {_ _ _} _ _%proto. *) +(* Global Instance: Params (@iProto_own) 3 := {}. *) + +(* Global Instance iProto_own_contractive `{protoG Σ V} γ s : *) +(* Contractive (iProto_own γ s). *) +(* Proof. solve_contractive. Qed. *) + +(** * Proofs *) +Section proto. + Context `{!protoG Σ V}. + Implicit Types v : V. + Implicit Types p pl pr : iProto Σ V. + Implicit Types m : iMsg Σ V. + + (** ** Equality *) + Lemma iProto_case p : p ≡ END ∨ ∃ a m, p ≡ <a> m. + Proof. + rewrite iProto_message_eq iProto_end_eq. + destruct (proto_case p) as [|(a&m&?)]; [by left|right]. + by exists a, (IMsg m). + Qed. + Lemma iProto_message_equivI `{!BiInternalEq SPROP} a1 a2 m1 m2 : + (<a1> m1) ≡ (<a2> m2) ⊣⊢@{SPROP} ⌜ a1 = a2 ⌠∧ + (∀ v lp, iMsg_car m1 v lp ≡ iMsg_car m2 v lp). + Proof. rewrite iProto_message_eq. apply proto_message_equivI. Qed. + + Lemma iProto_message_end_equivI `{!BiInternalEq SPROP} a m : + (<a> m) ≡ END ⊢@{SPROP} False. + Proof. rewrite iProto_message_eq iProto_end_eq. apply proto_message_end_equivI. Qed. + Lemma iProto_end_message_equivI `{!BiInternalEq SPROP} a m : + END ≡ (<a> m) ⊢@{SPROP} False. + Proof. by rewrite internal_eq_sym iProto_message_end_equivI. Qed. + + (** ** Non-expansiveness of operators *) + Global Instance iMsg_car_proper : + Proper (iMsg_equiv ==> (=) ==> (≡) ==> (≡)) (iMsg_car (Σ:=Σ) (V:=V)). + Proof. + intros m1 m2 meq v1 v2 veq p1 p2 peq. rewrite meq. + f_equiv; [ by f_equiv | done ]. + Qed. + Global Instance iMsg_car_ne n : + Proper (iMsg_dist n ==> (=) ==> (dist n) ==> (dist n)) (iMsg_car (Σ:=Σ) (V:=V)). + Proof. + intros m1 m2 meq v1 v2 veq p1 p2 peq. rewrite meq. + f_equiv; [ by f_equiv | done ]. + Qed. + + Global Instance iMsg_contractive v n : + Proper (dist n ==> dist_later n ==> dist n) (iMsg_base (Σ:=Σ) (V:=V) v). + Proof. rewrite iMsg_base_eq=> P1 P2 HP p1 p2 Hp w q /=. solve_contractive. Qed. + Global Instance iMsg_ne v : NonExpansive2 (iMsg_base (Σ:=Σ) (V:=V) v). + Proof. rewrite iMsg_base_eq=> P1 P2 HP p1 p2 Hp w q /=. solve_proper. Qed. + Global Instance iMsg_proper v : + Proper ((≡) ==> (≡) ==> (≡)) (iMsg_base (Σ:=Σ) (V:=V) v). + Proof. apply (ne_proper_2 _). Qed. + + Global Instance iMsg_exist_ne A n : + Proper (pointwise_relation _ (dist n) ==> (dist n)) (@iMsg_exist Σ V A). + Proof. rewrite iMsg_exist_eq=> m1 m2 Hm v p /=. f_equiv=> x. apply Hm. Qed. + Global Instance iMsg_exist_proper A : + Proper (pointwise_relation _ (≡) ==> (≡)) (@iMsg_exist Σ V A). + Proof. rewrite iMsg_exist_eq=> m1 m2 Hm v p /=. f_equiv=> x. apply Hm. Qed. + + Global Instance msg_tele_base (v:V) (P : iProp Σ) (p : iProto Σ V) : + MsgTele (TT:=TeleO) (MSG v {{ P }}; p) v P p. + Proof. done. Qed. + Global Instance msg_tele_exist {A} {TT : A → tele} (m : A → iMsg Σ V) tv tP tp : + (∀ x, MsgTele (TT:=TT x) (m x) (tv x) (tP x) (tp x)) → + MsgTele (TT:=TeleS TT) (∃ x, m x) tv tP tp. + Proof. intros Hm. rewrite /MsgTele /=. f_equiv=> x. apply Hm. Qed. + + Global Instance iProto_message_ne a : + NonExpansive (iProto_message (Σ:=Σ) (V:=V) a). + Proof. rewrite iProto_message_eq. solve_proper. Qed. + Global Instance iProto_message_proper a : + Proper ((≡) ==> (≡)) (iProto_message (Σ:=Σ) (V:=V) a). + Proof. apply (ne_proper _). Qed. + + Lemma iProto_message_equiv {TT1 TT2 : tele} a1 a2 + (m1 m2 : iMsg Σ V) + (v1 : TT1 -t> V) (v2 : TT2 -t> V) + (P1 : TT1 -t> iProp Σ) (P2 : TT2 -t> iProp Σ) + (prot1 : TT1 -t> iProto Σ V) (prot2 : TT2 -t> iProto Σ V) : + MsgTele m1 v1 P1 prot1 → + MsgTele m2 v2 P2 prot2 → + ⌜ a1 = a2 ⌠-∗ + (■∀.. (xs1 : TT1), tele_app P1 xs1 -∗ + ∃.. (xs2 : TT2), ⌜tele_app v1 xs1 = tele_app v2 xs2⌠∗ + â–· (tele_app prot1 xs1 ≡ tele_app prot2 xs2) ∗ + tele_app P2 xs2) -∗ + (■∀.. (xs2 : TT2), tele_app P2 xs2 -∗ + ∃.. (xs1 : TT1), ⌜tele_app v1 xs1 = tele_app v2 xs2⌠∗ + â–· (tele_app prot1 xs1 ≡ tele_app prot2 xs2) ∗ + tele_app P1 xs1) -∗ + (<a1> m1) ≡ (<a2> m2). + Proof. + iIntros (Hm1 Hm2 Heq) "#Heq1 #Heq2". + unfold MsgTele in Hm1. rewrite Hm1. clear Hm1. + unfold MsgTele in Hm2. rewrite Hm2. clear Hm2. + rewrite iProto_message_eq proto_message_equivI. + iSplit; [ done | ]. + iIntros (v p'). + do 2 rewrite iMsg_texist_exist. + rewrite iMsg_base_eq /=. + iApply prop_ext. + iIntros "!>". iSplit. + - iDestruct 1 as (xs1 Hveq1) "[Hrec1 HP1]". + iDestruct ("Heq1" with "HP1") as (xs2 Hveq2) "[Hrec2 HP2]". + iExists xs2. rewrite -Hveq1 Hveq2. + iSplitR; [ done | ]. iSplitR "HP2"; [ | done ]. + iRewrite -"Hrec1". iApply later_equivI. iIntros "!>". by iRewrite "Hrec2". + - iDestruct 1 as (xs2 Hveq2) "[Hrec2 HP2]". + iDestruct ("Heq2" with "HP2") as (xs1 Hveq1) "[Hrec1 HP1]". + iExists xs1. rewrite -Hveq2 Hveq1. + iSplitR; [ done | ]. iSplitR "HP1"; [ | done ]. + iRewrite -"Hrec2". iApply later_equivI. iIntros "!>". by iRewrite "Hrec1". + Qed. + + (** Helpers *) + Lemma iMsg_map_base f v P p : + NonExpansive f → + iMsg_map f (MSG v {{ P }}; p) ≡ (MSG v {{ P }}; f p)%msg. + Proof. + rewrite iMsg_base_eq. intros ? v' p'; simpl. iSplit. + - iDestruct 1 as (p'') "[(->&Hp&$) Hp']". iSplit; [done|]. + iRewrite "Hp'". iIntros "!>". by iRewrite "Hp". + - iIntros "(->&Hp'&$)". iExists p. iRewrite -"Hp'". auto. + Qed. + Lemma iMsg_map_exist {A} f (m : A → iMsg Σ V) : + iMsg_map f (∃ x, m x) ≡ (∃ x, iMsg_map f (m x))%msg. + Proof. + rewrite iMsg_exist_eq. intros v' p'; simpl. iSplit. + - iDestruct 1 as (p'') "[H Hp']". iDestruct "H" as (x) "H"; auto. + - iDestruct 1 as (x p'') "[Hm Hp']". auto. + Qed. + + (** ** Dual *) + Global Instance iProto_dual_ne : NonExpansive (@iProto_dual Σ V). + Proof. rewrite iProto_dual_eq. solve_proper. Qed. + Global Instance iProto_dual_proper : Proper ((≡) ==> (≡)) (@iProto_dual Σ V). + Proof. apply (ne_proper _). Qed. + Global Instance iProto_dual_if_ne d : NonExpansive (@iProto_dual_if Σ V d). + Proof. solve_proper. Qed. + Global Instance iProto_dual_if_proper d : + Proper ((≡) ==> (≡)) (@iProto_dual_if Σ V d). + Proof. apply (ne_proper _). Qed. + + Lemma iProto_dual_end : iProto_dual (Σ:=Σ) (V:=V) END ≡ END. + Proof. + rewrite iProto_end_eq iProto_dual_eq /iProto_dual_def /iProto_map_app. + etrans; [apply (fixpoint_unfold (iProto_map_app_aux _ _))|]; simpl. + by rewrite proto_elim_end. + Qed. + Lemma iProto_dual_message a m : + iProto_dual (<a> m) ≡ <action_dual a> iMsg_dual m. + Proof. + rewrite iProto_message_eq iProto_dual_eq /iProto_dual_def /iProto_map_app. + etrans; [apply (fixpoint_unfold (iProto_map_app_aux _ _))|]; simpl. + rewrite /iProto_message_def. rewrite ->proto_elim_message; [done|]. + intros a' m1 m2 Hm; f_equiv; solve_proper. + Qed. + Lemma iMsg_dual_base v P p : + iMsg_dual (MSG v {{ P }}; p) ≡ (MSG v {{ P }}; iProto_dual p)%msg. + Proof. apply iMsg_map_base, _. Qed. + Lemma iMsg_dual_exist {A} (m : A → iMsg Σ V) : + iMsg_dual (∃ x, m x) ≡ (∃ x, iMsg_dual (m x))%msg. + Proof. apply iMsg_map_exist. Qed. + + Global Instance iProto_dual_involutive : Involutive (≡) (@iProto_dual Σ V). + Proof. + intros p. apply (uPred.internal_eq_soundness (M:=iResUR Σ)). + iLöb as "IH" forall (p). destruct (iProto_case p) as [->|(a&m&->)]. + { by rewrite !iProto_dual_end. } + rewrite !iProto_dual_message involutive. + iApply iProto_message_equivI; iSplit; [done|]; iIntros (v p') "/=". + iApply prop_ext; iIntros "!>"; iSplit. + - iDestruct 1 as (pd) "[H Hp']". iRewrite "Hp'". + iDestruct "H" as (pdd) "[H #Hpd]". + iApply (internal_eq_rewrite); [|done]; iIntros "!>". + iRewrite "Hpd". by iRewrite ("IH" $! pdd). + - iIntros "H". destruct (Next_uninj p') as [p'' Hp']. iExists _. + rewrite Hp'. iSplitL; [by auto|]. iIntros "!>". by iRewrite ("IH" $! p''). + Qed. + + (** ** Append *) + Global Instance iProto_app_end_l : LeftId (≡) END (@iProto_app Σ V). + Proof. + intros p. rewrite iProto_end_eq iProto_app_eq /iProto_app_def /iProto_map_app. + etrans; [apply (fixpoint_unfold (iProto_map_app_aux _ _))|]; simpl. + by rewrite proto_elim_end. + Qed. + Lemma iProto_app_message a m p2 : (<a> m) <++> p2 ≡ <a> m <++> p2. + Proof. + rewrite iProto_message_eq iProto_app_eq /iProto_app_def /iProto_map_app. + etrans; [apply (fixpoint_unfold (iProto_map_app_aux _ _))|]; simpl. + rewrite /iProto_message_def. rewrite ->proto_elim_message; [done|]. + intros a' m1 m2 Hm. f_equiv; solve_proper. + Qed. + + Global Instance iProto_app_ne : NonExpansive2 (@iProto_app Σ V). + Proof. + assert (∀ n, Proper (dist n ==> (=) ==> dist n) (@iProto_app Σ V)). + { intros n p1 p1' Hp1 p2 p2' <-. by rewrite iProto_app_eq /iProto_app_def Hp1. } + assert (Proper ((≡) ==> (=) ==> (≡)) (@iProto_app Σ V)). + { intros p1 p1' Hp1 p2 p2' <-. by rewrite iProto_app_eq /iProto_app_def Hp1. } + intros n p1 p1' Hp1 p2 p2' Hp2. rewrite Hp1. clear p1 Hp1. + revert p1'. induction (lt_wf n) as [n _ IH]; intros p1. + destruct (iProto_case p1) as [->|(a&m&->)]. + { by rewrite !left_id. } + rewrite !iProto_app_message. f_equiv=> v p' /=. do 4 f_equiv. + f_contractive. apply IH; eauto using dist_le with lia. + Qed. + Global Instance iProto_app_proper : Proper ((≡) ==> (≡) ==> (≡)) (@iProto_app Σ V). + Proof. apply (ne_proper_2 _). Qed. + + Lemma iMsg_app_base v P p1 p2 : + ((MSG v {{ P }}; p1) <++> p2)%msg ≡ (MSG v {{ P }}; p1 <++> p2)%msg. + Proof. apply: iMsg_map_base. Qed. + Lemma iMsg_app_exist {A} (m : A → iMsg Σ V) p2 : + ((∃ x, m x) <++> p2)%msg ≡ (∃ x, m x <++> p2)%msg. + Proof. apply iMsg_map_exist. Qed. + + Global Instance iProto_app_end_r : RightId (≡) END (@iProto_app Σ V). + Proof. + intros p. apply (uPred.internal_eq_soundness (M:=iResUR Σ)). + iLöb as "IH" forall (p). destruct (iProto_case p) as [->|(a&m&->)]. + { by rewrite left_id. } + rewrite iProto_app_message. + iApply iProto_message_equivI; iSplit; [done|]; iIntros (v p') "/=". + iApply prop_ext; iIntros "!>". iSplit. + - iDestruct 1 as (p1') "[H Hp']". iRewrite "Hp'". + iApply (internal_eq_rewrite); [|done]; iIntros "!>". + by iRewrite ("IH" $! p1'). + - iIntros "H". destruct (Next_uninj p') as [p'' Hp']. iExists p''. + rewrite Hp'. iSplitL; [by auto|]. iIntros "!>". by iRewrite ("IH" $! p''). + Qed. + Global Instance iProto_app_assoc : Assoc (≡) (@iProto_app Σ V). + Proof. + intros p1 p2 p3. apply (uPred.internal_eq_soundness (M:=iResUR Σ)). + iLöb as "IH" forall (p1). destruct (iProto_case p1) as [->|(a&m&->)]. + { by rewrite !left_id. } + rewrite !iProto_app_message. + iApply iProto_message_equivI; iSplit; [done|]; iIntros (v p123) "/=". + iApply prop_ext; iIntros "!>". iSplit. + - iDestruct 1 as (p1') "[H #Hp']". + iExists (p1' <++> p2). iSplitL; [by auto|]. + iRewrite "Hp'". iIntros "!>". iApply "IH". + - iDestruct 1 as (p12) "[H #Hp123]". iDestruct "H" as (p1') "[H #Hp12]". + iExists p1'. iFrame "H". iRewrite "Hp123". + iIntros "!>". iRewrite "Hp12". by iRewrite ("IH" $! p1'). + Qed. + + Lemma iProto_dual_app p1 p2 : + iProto_dual (p1 <++> p2) ≡ iProto_dual p1 <++> iProto_dual p2. + Proof. + apply (uPred.internal_eq_soundness (M:=iResUR Σ)). + iLöb as "IH" forall (p1 p2). destruct (iProto_case p1) as [->|(a&m&->)]. + { by rewrite iProto_dual_end !left_id. } + rewrite iProto_dual_message !iProto_app_message iProto_dual_message /=. + iApply iProto_message_equivI; iSplit; [done|]; iIntros (v p12) "/=". + iApply prop_ext; iIntros "!>". iSplit. + - iDestruct 1 as (p12d) "[H #Hp12]". iDestruct "H" as (p1') "[H #Hp12d]". + iExists (iProto_dual p1'). iSplitL; [by auto|]. + iRewrite "Hp12". iIntros "!>". iRewrite "Hp12d". iApply "IH". + - iDestruct 1 as (p1') "[H #Hp12]". iDestruct "H" as (p1d) "[H #Hp1']". + iExists (p1d <++> p2). iSplitL; [by auto|]. + iRewrite "Hp12". iIntros "!>". iRewrite "Hp1'". by iRewrite ("IH" $! p1d p2). + Qed. + + Lemma iProto_consistent_flip p1 p2 : + iProto_consistent p2 p1 -∗ + iProto_consistent p1 p2. + Proof. + iLöb as "IH" forall (p1 p2). + iIntros "Hprot". + rewrite iProto_consistent_unfold /iProto_consistent_pre. + rewrite iProto_consistent_unfold /iProto_consistent_pre. + iIntros (a1 a2 m1 m2) "Hm1 Hm2". + iDestruct ("Hprot" with "Hm2 Hm1") as "Hprot". + destruct a1, a2; [done| | |done]. + - iIntros (v p1') "Hm1". + iDestruct ("Hprot" with "Hm1") as (p2') "[Hprot Hrec]". + iExists p2'. iFrame. by iApply "IH". + - iIntros (v p2') "Hm2". + iDestruct ("Hprot" with "Hm2") as (p1') "[Hprot Hrec]". + iExists p1'. iFrame. by iApply "IH". + Qed. + + Lemma iProto_consistent_step m1 m2 v p1 : + iProto_consistent (<!> m1) (<?> m2) -∗ + iMsg_car m1 v (Next p1) -∗ + ∃ p2, iMsg_car m2 v (Next p2) ∗ â–· iProto_consistent p1 p2. + Proof. + iIntros "Hprot Hm1". + rewrite iProto_consistent_unfold /iProto_consistent_pre. + iDestruct ("Hprot" with "[//] [//] Hm1") as (p2) "[Hm2 Hprot]". + iExists p2. iFrame. + Qed. + + Lemma iProto_consistent_dual p : + ⊢ iProto_consistent p (iProto_dual p). + Proof. + iLöb as "IH" forall (p). + rewrite iProto_consistent_unfold. + iIntros (a1 a2 m1 m2) "#Hmeq1 #Hmeq2". + destruct a1, a2; [done| | |done]. + - iIntros (v p') "Hm1". + iRewrite "Hmeq1" in "Hmeq2". + rewrite iProto_dual_message. simpl. + rewrite iProto_message_equivI. + iDestruct "Hmeq2" as (_) "Hmeq2". + iSpecialize ("Hmeq2" $! v (Next (iProto_dual p'))). + iExists (iProto_dual p'). + iRewrite -"Hmeq2". + iSplitL; [|iNext; by iApply "IH"]. + simpl. iExists p'. iFrame. done. + - iIntros (v p') "Hm2". + iRewrite "Hmeq1" in "Hmeq2". + rewrite iProto_dual_message. simpl. + rewrite iProto_message_equivI. + iDestruct "Hmeq2" as (_) "Hmeq2". + iSpecialize ("Hmeq2" $! v (Next p')). + iRewrite -"Hmeq2" in "Hm2". + simpl. + iDestruct "Hm2" as (p'') "[Hm2 Hmeq']". + iExists p''. iFrame. + iNext. iRewrite "Hmeq'". + iApply "IH". + Qed. + + Lemma iProto_consistent_le_l pl pl' pr : + iProto_consistent pl pr -∗ pl ⊑ pl' -∗ iProto_consistent pl' pr. + Proof. iIntros "Hprot Hle". by iApply "Hle". Qed. + Lemma iProto_consistent_le_r pl pr pr' : + iProto_consistent pl pr -∗ pr ⊑ pr' -∗ iProto_consistent pl pr'. + Proof. + iIntros "H Hle". iApply iProto_consistent_flip. + iApply "Hle". by iApply iProto_consistent_flip. + Qed. + + Lemma iProto_le_refl p : ⊢ p ⊑ p. + Proof. iIntros (p') "$". Qed. + + Lemma iProto_le_end : ⊢ END ⊑ (END : iProto Σ V). + Proof. iApply iProto_le_refl. Qed. + + (* Lemma iProto_le_send m1 m2 : *) + (* (∀ v p2', iMsg_car m2 v (Next p2') -∗ ∃ p1', *) + (* â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1')) -∗ *) + (* (<!> m1) ⊑ (<!> m2). *) + (* Proof. *) + (* iLöb as "IH" forall (m1 m2). *) + (* iIntros "Hle". *) + (* repeat (unfold iProto_le at 3). *) + (* iIntros (p) "Hprot". *) + (* repeat rewrite iProto_consistent_unfold; unfold iProto_consistent_pre. *) + (* iSplit; [iDestruct "Hprot" as "[Hprot _]"|iDestruct "Hprot" as "[_ Hprot]"]. *) + (* { iIntros (a m) "H". *) + (* iDestruct (iProto_message_equivI with "H") as (Heq) "{H} #Hmeq"; subst. *) + (* iIntros (v p') "Hm". *) + (* iSpecialize ("Hprot" with "[//]"). *) + (* iSpecialize ("Hmeq" $! v (Next p')). iRewrite -"Hmeq" in "Hm". *) + (* iSpecialize ("Hle" $! _ _ with "Hm"). iDestruct "Hle" as (p'') "[Hle Hm]". *) + (* iApply "Hle". iApply "Hprot". iApply "Hm". } *) + (* { iIntros (a m) "H". *) + (* iSpecialize ("Hprot" $! _ _ with "H"). *) + (* destruct a. *) + (* - iIntros (v p') "Hm". *) + (* iApply ("IH" with "[Hle]"). *) + (* { iNext. iIntros (v' p'') "Hm". *) + (* iSpecialize ("Hle" with "Hm"). *) + (* iDestruct "Hle" as (p''') "[Hle Hm]". *) + (* iExists p'''. iSplitR "Hm"; [iApply "Hle" | iApply "Hm"]. } *) + (* iApply "Hprot". iApply "Hm". *) + (* - iIntros (v vs Heq); subst. *) + (* iSpecialize ("Hprot" $! _ _ eq_refl). *) + (* iDestruct "Hprot" as (p') "[Hm Hprot]". *) + (* iExists (p'). iFrame "Hm". *) + (* iApply ("IH" with "[Hle] [Hprot]"); [|iApply "Hprot"]. *) + (* iNext. iIntros (v' p'') "Hm". *) + (* iSpecialize ("Hle" with "Hm"). iDestruct "Hle" as (p''') "[Hle Hm]". *) + (* iExists p'''; iSplitR "Hm"; [iApply "Hle" |iApply "Hm"]. } *) + (* Qed. *) + + (* Lemma iProto_le_recv m1 m2 : *) + (* (∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', *) + (* â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2')) -∗ *) + (* (<?> m1) ⊑ (<?> m2). *) + (* Proof. *) + (* iIntros "H" (vsl vsr p) "Hprot". *) + (* iLöb as "IH" forall (vsl vsr m2 p). *) + (* iEval (rewrite iProto_consistent_unfold). *) + (* rewrite /iProto_consistent_pre. *) + (* rewrite iProto_consistent_unfold /iProto_consistent_pre. *) + (* iSplit; last first. *) + (* { *) + (* iDestruct "Hprot" as "[_ Hprot]". *) + (* iIntros ([] m) "Heq". *) + (* { iIntros (v vs) "Hm". *) + (* iDestruct ("Hprot" with "Heq Hm") as "Hprot". *) + (* iIntros "!>". *) + (* by iApply ("IH" with "H"). } *) + (* { iIntros (v vs) "Hm". *) + (* iDestruct ("Hprot" with "Heq Hm") as (p') "[Hm Hprot]". *) + (* iExists p'. iFrame. *) + (* iIntros "!>". *) + (* by iApply ("IH" with "H"). } *) + (* } *) + (* destruct vsr. *) + (* { iIntros (a m) "Heq". rewrite iProto_message_equivI. *) + (* iDestruct "Heq" as (<-) "Heq". *) + (* iIntros (v vs Heq). done. } *) + (* iIntros (a m) "Heq". rewrite iProto_message_equivI. *) + (* iDestruct "Heq" as (<-) "Heq". *) + (* iIntros (w vs Heq). *) + (* assert (v = w) as <- by set_solver. *) + (* assert (vsr = vs) as <- by set_solver. *) + (* iDestruct "Hprot" as "[Hprot _]". *) + (* iDestruct ("Hprot" with "[//] [//]") as (p') "[Hm Hprot]". *) + (* iDestruct ("H" with "Hm") as (p'') "[Hle H]". *) + (* iExists p''. *) + (* iSpecialize ("Heq" $! v (Next p'')). iRewrite -"Heq". *) + (* iFrame. iIntros "!>". *) + (* iApply (iProto_consistent_le_l with "Hprot Hle"). *) + (* Qed. *) + + (* Lemma iProto_le_base a v P p1 p2 : *) + (* â–· (p1 ⊑ p2) -∗ *) + (* (<a> MSG v {{ P }}; p1) ⊑ (<a> MSG v {{ P }}; p2). *) + (* Proof. *) + (* rewrite iMsg_base_eq. iIntros "H". destruct a. *) + (* - iApply iProto_le_send. iIntros (v' p') "(->&Hp&$)". *) + (* iExists p1. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". *) + (* - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&$)". *) + (* iExists p2. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". *) + (* Qed. *) + + (* Lemma iProto_le_end_inv_l p : p ⊑ END -∗ (p ≡ END). *) + (* Proof. *) + (* rewrite iProto_le_unfold. iIntros "[[Hp _]|H] //". *) + (* iDestruct "H" as (a1 a2 m1 m2) "(_ & Heq & _)". *) + (* by iDestruct (iProto_end_message_equivI with "Heq") as %[]. *) + (* Qed. *) + + (* Lemma iProto_le_end_inv_r p : END ⊑ p -∗ (p ≡ END). *) + (* Proof. *) + (* rewrite iProto_le_unfold. iIntros "[[_ Hp]|H] //". *) + (* iDestruct "H" as (a1 a2 m1 m2) "(Heq & _ & _)". *) + (* iDestruct (iProto_end_message_equivI with "Heq") as %[]. *) + (* Qed. *) + + (* Lemma iProto_le_send_inv p1 m2 : *) + (* p1 ⊑ (<!> m2) -∗ ∃ a1 m1, *) + (* (p1 ≡ <a1> m1) ∗ *) + (* match a1 with *) + (* | Send => ∀ v p2', *) + (* iMsg_car m2 v (Next p2') -∗ ∃ p1', *) + (* â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1') *) + (* | Recv => ∀ v1 v2 p1' p2', *) + (* iMsg_car m1 v1 (Next p1') -∗ iMsg_car m2 v2 (Next p2') -∗ ∃ pt, *) + (* â–· (p1' ⊑ <!> MSG v2; pt) ∗ â–· ((<?> MSG v1; pt) ⊑ p2') *) + (* end. *) + (* Proof. *) + (* rewrite iProto_le_unfold. iIntros "[[_ Heq]|H]". *) + (* { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } *) + (* iDestruct "H" as (a1 a2 m1 m2') "(Hp1 & Hp2 & H)". *) + (* iExists _, _; iSplit; [done|]. destruct a1, a2. *) + (* - iIntros (v p2') "Hm2". *) + (* iDestruct (iProto_message_equivI with "Hp2") as (_) "{Hp2} #Hm". *) + (* iApply "H". by iRewrite -("Hm" $! v (Next p2')). *) + (* - done. *) + (* - iIntros (v1 v2 p1' p2') "Hm1 Hm2". *) + (* iDestruct (iProto_message_equivI with "Hp2") as (_) "{Hp2} #Hm". *) + (* iApply ("H" with "Hm1"). by iRewrite -("Hm" $! v2 (Next p2')). *) + (* - iDestruct (iProto_message_equivI with "Hp2") as ([=]) "_". *) + (* Qed. *) + (* Lemma iProto_le_send_send_inv m1 m2 v p2' : *) + (* (<!> m1) ⊑ (<!> m2) -∗ *) + (* iMsg_car m2 v (Next p2') -∗ ∃ p1', â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1'). *) + (* Proof. *) + (* iIntros "H Hm2". iDestruct (iProto_le_send_inv with "H") as (a m1') "[Hm1 H]". *) + (* iDestruct (iProto_message_equivI with "Hm1") as (<-) "Hm1". *) + (* iDestruct ("H" with "Hm2") as (p1') "[Hle Hm]". *) + (* iRewrite -("Hm1" $! v (Next p1')) in "Hm". auto with iFrame. *) + (* Qed. *) + (* Lemma iProto_le_recv_send_inv m1 m2 v1 v2 p1' p2' : *) + (* (<?> m1) ⊑ (<!> m2) -∗ *) + (* iMsg_car m1 v1 (Next p1') -∗ iMsg_car m2 v2 (Next p2') -∗ ∃ pt, *) + (* â–· (p1' ⊑ <!> MSG v2; pt) ∗ â–· ((<?> MSG v1; pt) ⊑ p2'). *) + (* Proof. *) + (* iIntros "H Hm1 Hm2". iDestruct (iProto_le_send_inv with "H") as (a m1') "[Hm H]". *) + (* iDestruct (iProto_message_equivI with "Hm") as (<-) "{Hm} #Hm". *) + (* iApply ("H" with "[Hm1] Hm2"). by iRewrite -("Hm" $! v1 (Next p1')). *) + (* Qed. *) + + (* Lemma iProto_le_recv_inv p1 m2 : *) + (* p1 ⊑ (<?> m2) -∗ ∃ m1, *) + (* (p1 ≡ <?> m1) ∗ *) + (* ∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', *) + (* â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). *) + (* Proof. *) + (* rewrite iProto_le_unfold. iIntros "[[_ Heq]|H]". *) + (* { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } *) + (* iDestruct "H" as (a1 a2 m1 m2') "(Hp1 & Hp2 & H)". *) + (* iExists m1. *) + (* iDestruct (iProto_message_equivI with "Hp2") as (<-) "{Hp2} #Hm2". *) + (* destruct a1; [done|]. iSplit; [done|]. *) + (* iIntros (v p1') "Hm". iDestruct ("H" with "Hm") as (p2') "[Hle Hm]". *) + (* iExists p2'. iIntros "{$Hle}". by iRewrite ("Hm2" $! v (Next p2')). *) + (* Qed. *) + (* Lemma iProto_le_recv_recv_inv m1 m2 v p1' : *) + (* (<?> m1) ⊑ (<?> m2) -∗ *) + (* iMsg_car m1 v (Next p1') -∗ ∃ p2', â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). *) + (* Proof. *) + (* iIntros "H Hm2". iDestruct (iProto_le_recv_inv with "H") as (m1') "[Hm1 H]". *) + (* iApply "H". iDestruct (iProto_message_equivI with "Hm1") as (_) "Hm1". *) + (* by iRewrite -("Hm1" $! v (Next p1')). *) + (* Qed. *) + + (* Lemma iProto_le_refl p : ⊢ p ⊑ p. *) + (* Proof. *) + (* iLöb as "IH" forall (p). destruct (iProto_case p) as [->|([]&m&->)]. *) + (* - iApply iProto_le_end. *) + (* - iApply iProto_le_send. auto 10 with iFrame. *) + (* - iApply iProto_le_recv. auto 10 with iFrame. *) + (* Qed. *) + + Lemma iProto_le_trans p1 p2 p3 : p1 ⊑ p2 -∗ p2 ⊑ p3 -∗ p1 ⊑ p3. + Proof. + iIntros "H1 H2" (p) "Hprot". + iApply "H2". iApply "H1". done. + Qed. + + (* Lemma iProto_le_payload_elim_l a m v P p : *) + (* (P -∗ (<?> MSG v; p) ⊑ (<a> m)) -∗ *) + (* (<?> MSG v {{ P }}; p) ⊑ (<a> m). *) + (* Proof. *) + (* rewrite iMsg_base_eq. iIntros "H". destruct a. *) + (* - iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= (#?&#?&HP) Hm2 /=". *) + (* iApply (iProto_le_recv_send_inv with "(H HP)"); simpl; auto. *) + (* - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&HP)". *) + (* iApply (iProto_le_recv_recv_inv with "(H HP)"); simpl; auto. *) + (* Qed. *) + (* Lemma iProto_le_payload_elim_r a m v P p : *) + (* (P -∗ (<a> m) ⊑ (<!> MSG v; p)) -∗ *) + (* (<a> m) ⊑ (<!> MSG v {{ P }}; p). *) + (* Proof. *) + (* rewrite iMsg_base_eq. iIntros "H". destruct a. *) + (* - iApply iProto_le_send. iIntros (v' p') "(->&Hp&HP)". *) + (* iApply (iProto_le_send_send_inv with "(H HP)"); simpl; auto. *) + (* - iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1 (->&#?&HP) /=". *) + (* iApply (iProto_le_recv_send_inv with "(H HP) Hm1"); simpl; auto. *) + (* Qed. *) + (* Lemma iProto_le_payload_intro_l v P p : *) + (* P -∗ (<!> MSG v {{ P }}; p) ⊑ (<!> MSG v; p). *) + (* Proof. *) + (* rewrite iMsg_base_eq. *) + (* iIntros "HP". iApply iProto_le_send. iIntros (v' p') "(->&Hp&_) /=". *) + (* iExists p'. iSplitR; [iApply iProto_le_refl|]. auto. *) + (* Qed. *) + (* Lemma iProto_le_payload_intro_r v P p : *) + (* P -∗ (<?> MSG v; p) ⊑ (<?> MSG v {{ P }}; p). *) + (* Proof. *) + (* rewrite iMsg_base_eq. *) + (* iIntros "HP". iApply iProto_le_recv. iIntros (v' p') "(->&Hp&_) /=". *) + (* iExists p'. iSplitR; [iApply iProto_le_refl|]. auto. *) + (* Qed. *) + + (* Lemma iProto_le_exist_elim_l {A} (m1 : A → iMsg Σ V) a m2 : *) + (* (∀ x, (<?> m1 x) ⊑ (<a> m2)) -∗ *) + (* (<? x> m1 x) ⊑ (<a> m2). *) + (* Proof. *) + (* rewrite iMsg_exist_eq. iIntros "H". destruct a. *) + (* - iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1 Hm2 /=". *) + (* iDestruct "Hm1" as (x) "Hm1". *) + (* iApply (iProto_le_recv_send_inv with "H Hm1 Hm2"). *) + (* - iApply iProto_le_recv. iIntros (v p1') "/=". iDestruct 1 as (x) "Hm". *) + (* by iApply (iProto_le_recv_recv_inv with "H"). *) + (* Qed. *) + + (* Lemma iProto_le_exist_elim_l_inhabited `{!Inhabited A} (m : A → iMsg Σ V) p : *) + (* (∀ x, (<?> m x) ⊑ p) -∗ *) + (* (<? x> m x) ⊑ p. *) + (* Proof. *) + (* rewrite iMsg_exist_eq. iIntros "H". *) + (* destruct (iProto_case p) as [Heq | [a [m' Heq]]]. *) + (* - unshelve iSpecialize ("H" $!inhabitant); first by apply _. *) + (* rewrite Heq. *) + (* iDestruct (iProto_le_end_inv_l with "H") as "H". *) + (* rewrite iProto_end_eq iProto_message_eq. *) + (* iDestruct (proto_message_end_equivI with "H") as "[]". *) + (* - iEval (rewrite Heq). destruct a. *) + (* + iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1 Hm2 /=". *) + (* iDestruct "Hm1" as (x) "Hm1". *) + (* iSpecialize ("H" $! x). rewrite Heq. *) + (* iApply (iProto_le_recv_send_inv with "H Hm1 Hm2"). *) + (* + iApply iProto_le_recv. iIntros (v p1') "/=". iDestruct 1 as (x) "Hm". *) + (* iSpecialize ("H" $! x). rewrite Heq. *) + (* by iApply (iProto_le_recv_recv_inv with "H"). *) + (* Qed. *) + + (* Lemma iProto_le_exist_elim_r {A} a m1 (m2 : A → iMsg Σ V) : *) + (* (∀ x, (<a> m1) ⊑ (<!> m2 x)) -∗ *) + (* (<a> m1) ⊑ (<! x> m2 x). *) + (* Proof. *) + (* rewrite iMsg_exist_eq. iIntros "H". destruct a. *) + (* - iApply iProto_le_send. iIntros (v p2'). iDestruct 1 as (x) "Hm". *) + (* by iApply (iProto_le_send_send_inv with "H"). *) + (* - iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1". *) + (* iDestruct 1 as (x) "Hm2". *) + (* iApply (iProto_le_recv_send_inv with "H Hm1 Hm2"). *) + (* Qed. *) + (* Lemma iProto_le_exist_elim_r_inhabited `{Hinh : Inhabited A} p (m : A → iMsg Σ V) : *) + (* (∀ x, p ⊑ (<!> m x)) -∗ *) + (* p ⊑ (<! x> m x). *) + (* Proof. *) + (* rewrite iMsg_exist_eq. iIntros "H". *) + (* destruct (iProto_case p) as [Heq | [a [m' Heq]]]. *) + (* - unshelve iSpecialize ("H" $!inhabitant); first by apply _. *) + (* rewrite Heq. *) + (* iDestruct (iProto_le_end_inv_r with "H") as "H". *) + (* rewrite iProto_end_eq iProto_message_eq. *) + (* iDestruct (proto_message_end_equivI with "H") as "[]". *) + (* - iEval (rewrite Heq). destruct a. *) + (* + iApply iProto_le_send. iIntros (v p2'). iDestruct 1 as (x) "Hm". *) + (* iSpecialize ("H" $! x). rewrite Heq. *) + (* by iApply (iProto_le_send_send_inv with "H"). *) + (* + iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1". *) + (* iDestruct 1 as (x) "Hm2". *) + (* iSpecialize ("H" $! x). rewrite Heq. *) + (* iApply (iProto_le_recv_send_inv with "H Hm1 Hm2"). *) + (* Qed. *) + (* Lemma iProto_le_exist_intro_l {A} (m : A → iMsg Σ V) a : *) + (* ⊢ (<! x> m x) ⊑ (<!> m a). *) + (* Proof. *) + (* rewrite iMsg_exist_eq. iApply iProto_le_send. iIntros (v p') "Hm /=". *) + (* iExists p'. iSplitR; last by auto. iApply iProto_le_refl. *) + (* Qed. *) + (* Lemma iProto_le_exist_intro_r {A} (m : A → iMsg Σ V) a : *) + (* ⊢ (<?> m a) ⊑ (<? x> m x). *) + (* Proof. *) + (* rewrite iMsg_exist_eq. iApply iProto_le_recv. iIntros (v p') "Hm /=". *) + (* iExists p'. iSplitR; last by auto. iApply iProto_le_refl. *) + (* Qed. *) + + (* Lemma iProto_le_texist_elim_l {TT : tele} (m1 : TT → iMsg Σ V) a m2 : *) + (* (∀ x, (<?> m1 x) ⊑ (<a> m2)) -∗ *) + (* (<?.. x> m1 x) ⊑ (<a> m2). *) + (* Proof. *) + (* iIntros "H". iInduction TT as [|T TT] "IH"; simpl; [done|]. *) + (* iApply iProto_le_exist_elim_l; iIntros (x). *) + (* iApply "IH". iIntros (xs). iApply "H". *) + (* Qed. *) + (* Lemma iProto_le_texist_elim_r {TT : tele} a m1 (m2 : TT → iMsg Σ V) : *) + (* (∀ x, (<a> m1) ⊑ (<!> m2 x)) -∗ *) + (* (<a> m1) ⊑ (<!.. x> m2 x). *) + (* Proof. *) + (* iIntros "H". iInduction TT as [|T TT] "IH"; simpl; [done|]. *) + (* iApply iProto_le_exist_elim_r; iIntros (x). *) + (* iApply "IH". iIntros (xs). iApply "H". *) + (* Qed. *) + + (* Lemma iProto_le_texist_intro_l {TT : tele} (m : TT → iMsg Σ V) x : *) + (* ⊢ (<!.. x> m x) ⊑ (<!> m x). *) + (* Proof. *) + (* induction x as [|T TT x xs IH] using tele_arg_ind; simpl. *) + (* { iApply iProto_le_refl. } *) + (* iApply iProto_le_trans; [by iApply iProto_le_exist_intro_l|]. iApply IH. *) + (* Qed. *) + (* Lemma iProto_le_texist_intro_r {TT : tele} (m : TT → iMsg Σ V) x : *) + (* ⊢ (<?> m x) ⊑ (<?.. x> m x). *) + (* Proof. *) + (* induction x as [|T TT x xs IH] using tele_arg_ind; simpl. *) + (* { iApply iProto_le_refl. } *) + (* iApply iProto_le_trans; [|by iApply iProto_le_exist_intro_r]. iApply IH. *) + (* Qed. *) + + (* Lemma iProto_le_base a v P p1 p2 : *) + (* â–· (p1 ⊑ p2) -∗ *) + (* (<a> MSG v {{ P }}; p1) ⊑ (<a> MSG v {{ P }}; p2). *) + (* Proof. *) + (* rewrite iMsg_base_eq. iIntros "H". destruct a. *) + (* - iApply iProto_le_send. iIntros (v' p') "(->&Hp&$)". *) + (* iExists p1. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". *) + (* - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&$)". *) + (* iExists p2. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". *) + (* Qed. *) + + (* Lemma iProto_le_base_swap v1 v2 P1 P2 p : *) + (* ⊢ (<?> MSG v1 {{ P1 }}; <!> MSG v2 {{ P2 }}; p) *) + (* ⊑ (<!> MSG v2 {{ P2 }}; <?> MSG v1 {{ P1 }}; p). *) + (* Proof. *) + (* rewrite {1 3}iMsg_base_eq. iApply iProto_le_swap. *) + (* iIntros (v1' v2' p1' p2') "/= (->&#Hp1&HP1) (->&#Hp2&HP2)". iExists p. *) + (* iSplitL "HP2". *) + (* - iIntros "!>". iRewrite -"Hp1". by iApply iProto_le_payload_intro_l. *) + (* - iIntros "!>". iRewrite -"Hp2". by iApply iProto_le_payload_intro_r. *) + (* Qed. *) + + (* Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. *) + (* Proof. *) + (* iIntros "H". iLöb as "IH" forall (p1 p2). *) + (* destruct (iProto_case p1) as [->|([]&m1&->)]. *) + (* - iDestruct (iProto_le_end_inv_l with "H") as "H". *) + (* iRewrite "H". iApply iProto_le_refl. *) + (* - iDestruct (iProto_le_send_inv with "H") as (a2 m2) "[Hp2 H]". *) + (* iRewrite "Hp2"; clear p2. iEval (rewrite !iProto_dual_message). *) + (* destruct a2; simpl. *) + (* + iApply iProto_le_recv. iIntros (v p1d). *) + (* iDestruct 1 as (p1') "[Hm1 #Hp1d]". *) + (* iDestruct ("H" with "Hm1") as (p2') "[H Hm2]". *) + (* iDestruct ("IH" with "H") as "H". iExists (iProto_dual p2'). *) + (* iSplitL "H"; [iIntros "!>"; by iRewrite "Hp1d"|]. simpl; auto. *) + (* + iApply iProto_le_swap. iIntros (v1 v2 p1d p2d). *) + (* iDestruct 1 as (p1') "[Hm1 #Hp1d]". iDestruct 1 as (p2') "[Hm2 #Hp2d]". *) + (* iDestruct ("H" with "Hm2 Hm1") as (pt) "[H1 H2]". *) + (* iDestruct ("IH" with "H1") as "H1". iDestruct ("IH" with "H2") as "H2 {IH}". *) + (* rewrite !iProto_dual_message /=. iExists (iProto_dual pt). iSplitL "H2". *) + (* * iIntros "!>". iRewrite "Hp1d". by rewrite -iMsg_dual_base. *) + (* * iIntros "!>". iRewrite "Hp2d". by rewrite -iMsg_dual_base. *) + (* - iDestruct (iProto_le_recv_inv with "H") as (m2) "[Hp2 H]". *) + (* iRewrite "Hp2"; clear p2. iEval (rewrite !iProto_dual_message /=). *) + (* iApply iProto_le_send. iIntros (v p2d). *) + (* iDestruct 1 as (p2') "[Hm2 #Hp2d]". *) + (* iDestruct ("H" with "Hm2") as (p1') "[H Hm1]". *) + (* iDestruct ("IH" with "H") as "H". iExists (iProto_dual p1'). *) + (* iSplitL "H"; [iIntros "!>"; by iRewrite "Hp2d"|]. simpl; auto. *) + (* Qed. *) + + (* Lemma iProto_le_amber_internal (p1 p2 : iProto Σ V → iProto Σ V) *) + (* `{Contractive p1, Contractive p2}: *) + (* â–¡ (∀ rec1 rec2, â–· (rec1 ⊑ rec2) → p1 rec1 ⊑ p2 rec2) -∗ *) + (* fixpoint p1 ⊑ fixpoint p2. *) + (* Proof. *) + (* iIntros "#H". iLöb as "IH". *) + (* iEval (rewrite (fixpoint_unfold p1)). *) + (* iEval (rewrite (fixpoint_unfold p2)). *) + (* iApply "H". iApply "IH". *) + (* Qed. *) + (* Lemma iProto_le_amber_external (p1 p2 : iProto Σ V → iProto Σ V) *) + (* `{Contractive p1, Contractive p2}: *) + (* (∀ rec1 rec2, (⊢ rec1 ⊑ rec2) → ⊢ p1 rec1 ⊑ p2 rec2) → *) + (* ⊢ fixpoint p1 ⊑ fixpoint p2. *) + (* Proof. *) + (* intros IH. apply fixpoint_ind. *) + (* - by intros p1' p2' -> ?. *) + (* - exists (fixpoint p2). iApply iProto_le_refl. *) + (* - intros p' ?. rewrite (fixpoint_unfold p2). by apply IH. *) + (* - apply bi.limit_preserving_entails; [done|solve_proper]. *) + (* Qed. *) + + (* Lemma iProto_le_dual_l p1 p2 : iProto_dual p2 ⊑ p1 -∗ iProto_dual p1 ⊑ p2. *) + (* Proof. *) + (* iIntros "H". iEval (rewrite -(involutive iProto_dual p2)). *) + (* by iApply iProto_le_dual. *) + (* Qed. *) + (* Lemma iProto_le_dual_r p1 p2 : p2 ⊑ iProto_dual p1 -∗ p1 ⊑ iProto_dual p2. *) + (* Proof. *) + (* iIntros "H". iEval (rewrite -(involutive iProto_dual p1)). *) + (* by iApply iProto_le_dual. *) + (* Qed. *) + + (* Lemma iProto_le_app p1 p2 p3 p4 : *) + (* p1 ⊑ p2 -∗ p3 ⊑ p4 -∗ p1 <++> p3 ⊑ p2 <++> p4. *) + (* Proof. *) + (* iIntros "H1 H2". iLöb as "IH" forall (p1 p2 p3 p4). *) + (* destruct (iProto_case p2) as [->|([]&m2&->)]. *) + (* - iDestruct (iProto_le_end_inv_l with "H1") as "H1". *) + (* iRewrite "H1". by rewrite !left_id. *) + (* - iDestruct (iProto_le_send_inv with "H1") as (a1 m1) "[Hp1 H1]". *) + (* iRewrite "Hp1"; clear p1. rewrite !iProto_app_message. destruct a1; simpl. *) + (* + iApply iProto_le_send. iIntros (v p24). *) + (* iDestruct 1 as (p2') "[Hm2 #Hp24]". *) + (* iDestruct ("H1" with "Hm2") as (p1') "[H1 Hm1]". *) + (* iExists (p1' <++> p3). iSplitR "Hm1"; [|by simpl; eauto]. *) + (* iIntros "!>". iRewrite "Hp24". by iApply ("IH" with "H1"). *) + (* + iApply iProto_le_swap. iIntros (v1 v2 p13 p24). *) + (* iDestruct 1 as (p1') "[Hm1 #Hp13]". iDestruct 1 as (p2') "[Hm2 #Hp24]". *) + (* iSpecialize ("H1" with "Hm1 Hm2"). *) + (* iDestruct "H1" as (pt) "[H1 H1']". *) + (* iExists (pt <++> p3). iSplitL "H1". *) + (* * iIntros "!>". iRewrite "Hp13". *) + (* rewrite /= -iMsg_app_base -iProto_app_message. *) + (* iApply ("IH" with "H1"). iApply iProto_le_refl. *) + (* * iIntros "!>". iRewrite "Hp24". *) + (* rewrite /= -iMsg_app_base -iProto_app_message. *) + (* iApply ("IH" with "H1' H2"). *) + (* - iDestruct (iProto_le_recv_inv with "H1") as (m1) "[Hp1 H1]". *) + (* iRewrite "Hp1"; clear p1. rewrite !iProto_app_message. iApply iProto_le_recv. *) + (* iIntros (v p13). iDestruct 1 as (p1') "[Hm1 #Hp13]". *) + (* iDestruct ("H1" with "Hm1") as (p2'') "[H1 Hm2]". *) + (* iExists (p2'' <++> p4). iSplitR "Hm2"; [|by simpl; eauto]. *) + (* iIntros "!>". iRewrite "Hp13". by iApply ("IH" with "H1"). *) + (* Qed. *) + + (** ** Lemmas about the auxiliary definitions and invariants *) + Global Instance iProto_app_recvs_ne vs : + NonExpansive (iProto_app_recvs (Σ:=Σ) (V:=V) vs). + Proof. induction vs; solve_proper. Qed. + Global Instance iProto_app_recvs_proper vs : + Proper ((≡) ==> (≡)) (iProto_app_recvs (Σ:=Σ) (V:=V) vs). + Proof. induction vs; solve_proper. Qed. + + Global Instance iProto_own_frag_ne γ s : NonExpansive (iProto_own_frag γ s). + Proof. solve_proper. Qed. + + Lemma iProto_own_auth_agree γ s p p' : + iProto_own_auth γ s p -∗ iProto_own_frag γ s p' -∗ â–· (p ≡ p'). + Proof. + iIntros "Hâ— Hâ—¯". iDestruct (own_valid_2 with "Hâ— Hâ—¯") as "H✓". + iDestruct (excl_auth_agreeI with "H✓") as "H✓". + iApply (later_equivI_1 with "H✓"). + Qed. + + Lemma iProto_own_auth_update γ s p p' p'' : + iProto_own_auth γ s p -∗ iProto_own_frag γ s p' ==∗ + iProto_own_auth γ s p'' ∗ iProto_own_frag γ s p''. + Proof. + iIntros "Hâ— Hâ—¯". iDestruct (own_update_2 with "Hâ— Hâ—¯") as "H". + { eapply (excl_auth_update _ _ (Next p'')). } + by rewrite own_op. + Qed. + + Global Instance iProto_own_ne γ s : NonExpansive (iProto_own γ s). + Proof. solve_proper. Qed. + Global Instance iProto_own_proper γ s : Proper ((≡) ==> (≡)) (iProto_own γ s). + Proof. apply (ne_proper _). Qed. + + Lemma iProto_own_le γ s p1 p2 : + iProto_own γ s p1 -∗ â–· (p1 ⊑ p2) -∗ iProto_own γ s p2. + Proof. + iDestruct 1 as (p1') "[Hle H]". iIntros "Hle'". + iExists p1'. iFrame "H". by iApply (iProto_le_trans with "Hle"). + Qed. + + Lemma iProto_init p : + ⊢ |==> ∃ γ, + iProto_ctx γ ∗ iProto_own γ Left p ∗ iProto_own γ Right (iProto_dual p). + Proof. + iMod (own_alloc (â—E (Next p) â‹… â—¯E (Next p))) as (lγ) "[Hâ—l Hâ—¯l]". + { by apply excl_auth_valid. } + iMod (own_alloc (â—E (Next (iProto_dual p)) â‹… + â—¯E (Next (iProto_dual p)))) as (rγ) "[Hâ—r Hâ—¯r]". + { by apply excl_auth_valid. } + pose (ProtName lγ rγ) as γ. iModIntro. iExists γ. iSplitL "Hâ—l Hâ—r". + { iExists p, (iProto_dual p). iFrame. iApply iProto_consistent_dual. } + iSplitL "Hâ—¯l"; iExists _; iFrame; iApply iProto_le_refl. + Qed. + + Definition side_dual s := + match s with + | Left => Right + | Right => Left + end. + + Lemma iProto_step_l γ m1 m2 p1 v : + iProto_ctx γ -∗ iProto_own γ Left (<!> m1) -∗ iProto_own γ Right (<?> m2) -∗ + iMsg_car m1 v (Next p1) ==∗ + â–· ∃ p2, iMsg_car m2 v (Next p2) ∗ â–· iProto_ctx γ ∗ + iProto_own γ Left p1 ∗ iProto_own γ Right p2. + Proof. + iDestruct 1 as (pl pr) "(Hâ—l & Hâ—r & Hconsistent)". + iDestruct 1 as (pl') "[Hlel Hâ—¯l]". + iDestruct 1 as (pr') "[Hler Hâ—¯r]". + iIntros "Hm". + iDestruct (iProto_own_auth_agree with "Hâ—l Hâ—¯l") as "#Hpl". + iDestruct (iProto_own_auth_agree with "Hâ—r Hâ—¯r") as "#Hpr". + iAssert (â–· (pl ⊑ <!> m1))%I + with "[Hlel]" as "{Hpl} Hlel"; first (iNext; by iRewrite "Hpl"). + iAssert (â–· (pr ⊑ <?> m2))%I + with "[Hler]" as "{Hpr} Hler"; first (iNext; by iRewrite "Hpr"). + iDestruct (iProto_consistent_le_l with "Hconsistent Hlel") as "Hconsistent". + iDestruct (iProto_consistent_le_r with "Hconsistent Hler") as "Hconsistent". + iDestruct (iProto_consistent_step with "Hconsistent [Hm //]") as + (p2) "[Hm2 Hconsistent]". + iMod (iProto_own_auth_update _ _ _ _ p1 with "Hâ—l Hâ—¯l") as "[Hâ—l Hâ—¯l]". + iMod (iProto_own_auth_update _ _ _ _ p2 with "Hâ—r Hâ—¯r") as "[Hâ—r Hâ—¯r]". + iIntros "!>!>". + iExists p2. iFrame. + iSplitL "Hconsistent Hâ—l Hâ—r". + - iExists _, _. iFrame. + - iSplitL "Hâ—¯l". + + iExists _. iFrame. iApply iProto_le_refl. + + iExists _. iFrame. iApply iProto_le_refl. + Qed. + + Lemma iProto_step_r γ m1 m2 p2 v : + iProto_ctx γ -∗ iProto_own γ Left (<?> m1) -∗ iProto_own γ Right (<!> m2) -∗ + iMsg_car m2 v (Next p2) ==∗ + â–· ∃ p1, iMsg_car m1 v (Next p1) ∗ â–· iProto_ctx γ ∗ + iProto_own γ Left p1 ∗ iProto_own γ Right p2. + Proof. + iDestruct 1 as (pl pr) "(Hâ—l & Hâ—r & Hconsistent)". + iDestruct 1 as (pl') "[Hlel Hâ—¯l]". + iDestruct 1 as (pr') "[Hler Hâ—¯r]". + iIntros "Hm". + iDestruct (iProto_own_auth_agree with "Hâ—l Hâ—¯l") as "#Hpl". + iDestruct (iProto_own_auth_agree with "Hâ—r Hâ—¯r") as "#Hpr". + iAssert (â–· (pl ⊑ <?> m1))%I + with "[Hlel]" as "{Hpl} Hlel"; first (iNext; by iRewrite "Hpl"). + iAssert (â–· (pr ⊑ <!> m2))%I + with "[Hler]" as "{Hpr} Hler"; first (iNext; by iRewrite "Hpr"). + iDestruct (iProto_consistent_le_l with "Hconsistent Hlel") as "Hconsistent". + iDestruct (iProto_consistent_le_r with "Hconsistent Hler") as "Hconsistent". + iDestruct (iProto_consistent_flip with "Hconsistent") as + "Hconsistent". + iDestruct (iProto_consistent_step with "Hconsistent [Hm //]") as + (p1) "[Hm1 Hconsistent]". + iMod (iProto_own_auth_update _ _ _ _ p1 with "Hâ—l Hâ—¯l") as "[Hâ—l Hâ—¯l]". + iMod (iProto_own_auth_update _ _ _ _ p2 with "Hâ—r Hâ—¯r") as "[Hâ—r Hâ—¯r]". + iIntros "!>!>". + iExists p1. iFrame. + iSplitL "Hconsistent Hâ—l Hâ—r". + - iExists _, _. iFrame. iApply iProto_consistent_flip. iFrame. + - iSplitL "Hâ—¯l". + + iExists _. iFrame. iApply iProto_le_refl. + + iExists _. iFrame. iApply iProto_le_refl. + Qed. + + (* (** The instances below make it possible to use the tactics [iIntros], *) + (* [iExist], [iSplitL]/[iSplitR], [iFrame] and [iModIntro] on [iProto_le] goals. *) *) + (* Global Instance iProto_le_from_forall_l {A} a (m1 : A → iMsg Σ V) m2 name : *) + (* AsIdentName m1 name → *) + (* FromForall (iProto_message Recv (iMsg_exist m1) ⊑ (<a> m2)) *) + (* (λ x, (<?> m1 x) ⊑ (<a> m2))%I name | 10. *) + (* Proof. intros _. apply iProto_le_exist_elim_l. Qed. *) + (* Global Instance iProto_le_from_forall_r {A} a m1 (m2 : A → iMsg Σ V) name : *) + (* AsIdentName m2 name → *) + (* FromForall ((<a> m1) ⊑ iProto_message Send (iMsg_exist m2)) *) + (* (λ x, (<a> m1) ⊑ (<!> m2 x))%I name | 11. *) + (* Proof. intros _. apply iProto_le_exist_elim_r. Qed. *) + + (* Global Instance iProto_le_from_wand_l a m v P p : *) + (* TCIf (TCEq P True%I) False TCTrue → *) + (* FromWand ((<?> MSG v {{ P }}; p) ⊑ (<a> m)) P ((<?> MSG v; p) ⊑ (<a> m)) | 10. *) + (* Proof. intros _. apply iProto_le_payload_elim_l. Qed. *) + (* Global Instance iProto_le_from_wand_r a m v P p : *) + (* FromWand ((<a> m) ⊑ (<!> MSG v {{ P }}; p)) P ((<a> m) ⊑ (<!> MSG v; p)) | 11. *) + (* Proof. apply iProto_le_payload_elim_r. Qed. *) + + (* Global Instance iProto_le_from_exist_l {A} (m : A → iMsg Σ V) p : *) + (* FromExist ((<! x> m x) ⊑ p) (λ a, (<!> m a) ⊑ p)%I | 10. *) + (* Proof. *) + (* rewrite /FromExist. iDestruct 1 as (x) "H". *) + (* iApply (iProto_le_trans with "[] H"). iApply iProto_le_exist_intro_l. *) + (* Qed. *) + (* Global Instance iProto_le_from_exist_r {A} (m : A → iMsg Σ V) p : *) + (* FromExist (p ⊑ <? x> m x) (λ a, p ⊑ (<?> m a))%I | 11. *) + (* Proof. *) + (* rewrite /FromExist. iDestruct 1 as (x) "H". *) + (* iApply (iProto_le_trans with "H"). iApply iProto_le_exist_intro_r. *) + (* Qed. *) + + (* Global Instance iProto_le_from_sep_l m v P p : *) + (* FromSep ((<!> MSG v {{ P }}; p) ⊑ (<!> m)) P ((<!> MSG v; p) ⊑ (<!> m)) | 10. *) + (* Proof. *) + (* rewrite /FromSep. iIntros "[HP H]". *) + (* iApply (iProto_le_trans with "[HP] H"). by iApply iProto_le_payload_intro_l. *) + (* Qed. *) + (* Global Instance iProto_le_from_sep_r m v P p : *) + (* FromSep ((<?> m) ⊑ (<?> MSG v {{ P }}; p)) P ((<?> m) ⊑ (<?> MSG v; p)) | 11. *) + (* Proof. *) + (* rewrite /FromSep. iIntros "[HP H]". *) + (* iApply (iProto_le_trans with "H"). by iApply iProto_le_payload_intro_r. *) + (* Qed. *) + + (* Global Instance iProto_le_frame_l q m v R P Q p : *) + (* Frame q R P Q → *) + (* Frame q R ((<!> MSG v {{ P }}; p) ⊑ (<!> m)) *) + (* ((<!> MSG v {{ Q }}; p) ⊑ (<!> m)) | 10. *) + (* Proof. *) + (* rewrite /Frame /=. iIntros (HP) "[HR H]". *) + (* iApply (iProto_le_trans with "[HR] H"). iApply iProto_le_payload_elim_r. *) + (* iIntros "HQ". iApply iProto_le_payload_intro_l. iApply HP; iFrame. *) + (* Qed. *) + (* Global Instance iProto_le_frame_r q m v R P Q p : *) + (* Frame q R P Q → *) + (* Frame q R ((<?> m) ⊑ (<?> MSG v {{ P }}; p)) *) + (* ((<?> m) ⊑ (<?> MSG v {{ Q }}; p)) | 11. *) + (* Proof. *) + (* rewrite /Frame /=. iIntros (HP) "[HR H]". *) + (* iApply (iProto_le_trans with "H"). iApply iProto_le_payload_elim_l. *) + (* iIntros "HQ". iApply iProto_le_payload_intro_r. iApply HP; iFrame. *) + (* Qed. *) + + (* Global Instance iProto_le_from_modal a v p1 p2 : *) + (* FromModal True (modality_instances.modality_laterN 1) (p1 ⊑ p2) *) + (* ((<a> MSG v; p1) ⊑ (<a> MSG v; p2)) (p1 ⊑ p2). *) + (* Proof. intros _. apply iProto_le_base. Qed. *) + +End proto. + +Typeclasses Opaque iProto_ctx iProto_own. + +Global Hint Extern 0 (environments.envs_entails _ (?x ⊑ ?y)) => + first [is_evar x; fail 1 | is_evar y; fail 1|iApply iProto_le_refl] : core. diff --git a/theories/channel/multi_proto_model.v b/theories/channel/multi_proto_model.v new file mode 100644 index 0000000..9c6c8be --- /dev/null +++ b/theories/channel/multi_proto_model.v @@ -0,0 +1,307 @@ +(** This file defines the model of dependent separation protocols as the +solution of a recursive domain equation, along with various primitive +operations, such as append and map. + +Important: This file should not be used directly, but rather the wrappers in +[proto] should be used. + +Dependent Separation Protocols are modeled as the solution of the following +recursive domain equation: + +[proto = 1 + (action * (V → â–¶ proto → PROP))] + +Here, the left-hand side of the sum is used for the terminated process, while +the right-hand side is used for the communication constructors. The type +[action] is an inductively defined datatype with two constructors [Send] and +[Recv]. Compared to having an additional sum in [proto], this makes it +possible to factorize the code in a better way. + +The remainder [V → â–¶ proto → PROP] is a predicate that ranges over the +communicated value [V] and the tail protocol [proto]. Note that to solve this +recursive domain equation using Iris's COFE solver, the recursive occurrence +of [proto] appear under the later [â–¶]. + +On top of the type [proto], we define the constructors: + +- [proto_end], which constructs the left-side of the sum. +- [proto_msg], which takes an action and a predicate and constructs the + right-hand side of the sum accordingly. + +The defined functions on the type [proto] are: + +- [proto_map], which can be used to map the actions and the propositions of + a given protocol. +- [proto_app], which appends two protocols [p1] and [p2], by substituting + all terminations [END] in [p1] with [p2]. *) +From iris.base_logic Require Import base_logic. +From iris.proofmode Require Import proofmode. +From actris.utils Require Import cofe_solver_2. +Set Default Proof Using "Type". + +Module Export action. + Inductive action := Send (n:nat) | Recv (n:nat). + Global Instance action_inhabited : Inhabited action := populate (Send 0). + Definition action_dual (a : action) : action := + match a with Send n => Recv n | Recv n => Send n end. + Global Instance action_dual_involutive : Involutive (=) action_dual. + Proof. by intros []. Qed. + Canonical Structure actionO := leibnizO action. +End action. + +Definition proto_auxO (V : Type) (PROP : ofe) (A : ofe) : ofe := + optionO (prodO actionO (V -d> laterO A -n> PROP)). +Definition proto_auxOF (V : Type) (PROP : ofe) : oFunctor := + optionOF (actionO * (V -d> â–¶ ∙ -n> PROP)). + +Definition proto_result (V : Type) := result_2 (proto_auxOF V). +Definition proto (V : Type) (PROPn PROP : ofe) `{!Cofe PROPn, !Cofe PROP} : ofe := + solution_2_car (proto_result V) PROPn _ PROP _. +Global Instance proto_cofe {V} `{!Cofe PROPn, !Cofe PROP} : Cofe (proto V PROPn PROP). +Proof. apply _. Qed. +Lemma proto_iso {V} `{!Cofe PROPn, !Cofe PROP} : + ofe_iso (proto_auxO V PROP (proto V PROP PROPn)) (proto V PROPn PROP). +Proof. apply proto_result. Qed. + +Definition proto_fold {V} `{!Cofe PROPn, !Cofe PROP} : + proto_auxO V PROP (proto V PROP PROPn) -n> proto V PROPn PROP := ofe_iso_1 proto_iso. +Definition proto_unfold {V} `{!Cofe PROPn, !Cofe PROP} : + proto V PROPn PROP -n> proto_auxO V PROP (proto V PROP PROPn) := ofe_iso_2 proto_iso. +Lemma proto_fold_unfold {V} `{!Cofe PROPn, !Cofe PROP} (p : proto V PROPn PROP) : + proto_fold (proto_unfold p) ≡ p. +Proof. apply (ofe_iso_12 proto_iso). Qed. +Lemma proto_unfold_fold {V} `{!Cofe PROPn, !Cofe PROP} + (p : proto_auxO V PROP (proto V PROP PROPn)) : + proto_unfold (proto_fold p) ≡ p. +Proof. apply (ofe_iso_21 proto_iso). Qed. + +Definition proto_end {V} `{!Cofe PROPn, !Cofe PROP} : proto V PROPn PROP := + proto_fold None. +Definition proto_message {V} `{!Cofe PROPn, !Cofe PROP} (a : action) + (m : V → laterO (proto V PROP PROPn) -n> PROP) : proto V PROPn PROP := + proto_fold (Some (a, m)). + +Global Instance proto_message_ne {V} `{!Cofe PROPn, !Cofe PROP} a n : + Proper (pointwise_relation V (dist n) ==> dist n) + (proto_message (PROPn:=PROPn) (PROP:=PROP) a). +Proof. intros c1 c2 Hc. rewrite /proto_message. f_equiv. by repeat constructor. Qed. +Global Instance proto_message_proper {V} `{!Cofe PROPn, !Cofe PROP} a : + Proper (pointwise_relation V (≡) ==> (≡)) + (proto_message (PROPn:=PROPn) (PROP:=PROP) a). +Proof. intros c1 c2 Hc. rewrite /proto_message. f_equiv. by repeat constructor. Qed. + +Lemma proto_case {V} `{!Cofe PROPn, !Cofe PROP} (p : proto V PROPn PROP) : + p ≡ proto_end ∨ ∃ a m, p ≡ proto_message a m. +Proof. + destruct (proto_unfold p) as [[a m]|] eqn:E; simpl in *; last first. + - left. by rewrite -(proto_fold_unfold p) E. + - right. exists a, m. by rewrite /proto_message -E proto_fold_unfold. +Qed. +Global Instance proto_inhabited {V} `{!Cofe PROPn, !Cofe PROP} : + Inhabited (proto V PROPn PROP) := populate proto_end. + +Lemma proto_message_equivI `{!BiInternalEq SPROP} {V} `{!Cofe PROPn, !Cofe PROP} a1 a2 m1 m2 : + proto_message (V:=V) (PROPn:=PROPn) (PROP:=PROP) a1 m1 ≡ proto_message a2 m2 + ⊣⊢@{SPROP} ⌜ a1 = a2 ⌠∧ (∀ v p', m1 v p' ≡ m2 v p'). +Proof. + trans (proto_unfold (proto_message a1 m1) ≡ proto_unfold (proto_message a2 m2) : SPROP)%I. + { iSplit. + - iIntros "Heq". by iRewrite "Heq". + - iIntros "Heq". rewrite -{2}(proto_fold_unfold (proto_message _ _)). + iRewrite "Heq". by rewrite proto_fold_unfold. } + rewrite /proto_message !proto_unfold_fold option_equivI prod_equivI /=. + rewrite discrete_eq discrete_fun_equivI. + by setoid_rewrite ofe_morO_equivI. +Qed. +Lemma proto_message_end_equivI `{!BiInternalEq SPROP} {V} `{!Cofe PROPn, !Cofe PROP} a m : + proto_message (V:=V) (PROPn:=PROPn) (PROP:=PROP) a m ≡ proto_end ⊢@{SPROP} False. +Proof. + trans (proto_unfold (proto_message a m) ≡ proto_unfold proto_end : SPROP)%I. + { iIntros "Heq". by iRewrite "Heq". } + by rewrite /proto_message !proto_unfold_fold option_equivI. +Qed. +Lemma proto_end_message_equivI `{!BiInternalEq SPROP} {V} `{!Cofe PROPn, !Cofe PROP} a m : + proto_end ≡ proto_message (V:=V) (PROPn:=PROPn) (PROP:=PROP) a m ⊢@{SPROP} False. +Proof. by rewrite internal_eq_sym proto_message_end_equivI. Qed. + +(** The eliminator [proto_elim x f p] is only well-behaved if the function [f] +is contractive *) +Definition proto_elim {V} `{!Cofe PROPn, !Cofe PROP} {A} + (x : A) (f : action → (V → laterO (proto V PROP PROPn) -n> PROP) → A) + (p : proto V PROPn PROP) : A := + match proto_unfold p with None => x | Some (a, m) => f a m end. + +Lemma proto_elim_ne {V} `{!Cofe PROPn, !Cofe PROP} {A : ofe} + (x : A) (f1 f2 : action → (V → laterO (proto V PROP PROPn) -n> PROP) → A) p1 p2 n : + (∀ a m1 m2, (∀ v, m1 v ≡{n}≡ m2 v) → f1 a m1 ≡{n}≡ f2 a m2) → + p1 ≡{n}≡ p2 → + proto_elim x f1 p1 ≡{n}≡ proto_elim x f2 p2. +Proof. + intros Hf Hp. rewrite /proto_elim. + apply (_ : NonExpansive proto_unfold) in Hp + as [[a1 m1] [a2 m2] [-> ?]|]; simplify_eq/=; [|done]. + apply Hf. destruct n; by simpl. +Qed. + +Lemma proto_elim_end {V} `{!Cofe PROPn, !Cofe PROP} {A : ofe} + (x : A) (f : action → (V → laterO (proto V PROP PROPn) -n> PROP) → A) : + proto_elim x f proto_end ≡ x. +Proof. + rewrite /proto_elim /proto_end. + pose proof (proto_unfold_fold (V:=V) (PROPn:=PROPn) (PROP:=PROP) None) as Hfold. + by destruct (proto_unfold (proto_fold None)) as [[??]|] eqn:E; inversion Hfold. +Qed. +Lemma proto_elim_message {V} `{!Cofe PROPn, !Cofe PROP} {A : ofe} + (x : A) (f : action → (V → laterO (proto V PROP PROPn) -n> PROP) → A) + `{Hf : ∀ a, Proper (pointwise_relation _ (≡) ==> (≡)) (f a)} a m : + proto_elim x f (proto_message a m) ≡ f a m. +Proof. + rewrite /proto_elim /proto_message /=. + pose proof (proto_unfold_fold (V:=V) (PROPn:=PROPn) + (PROP:=PROP) (Some (a, m))) as Hfold. + destruct (proto_unfold (proto_fold (Some (a, m)))) + as [[??]|] eqn:E; inversion Hfold as [?? [Ha Hc]|]; simplify_eq/=. + by f_equiv=> v. +Qed. + +(** Functor *) +Program Definition proto_map_aux {V} `{!Cofe PROPn, !Cofe PROPn', !Cofe PROP, !Cofe PROP'} + (g : PROP -n> PROP') (rec : proto V PROP' PROPn' -n> proto V PROP PROPn) : + proto V PROPn PROP -n> proto V PROPn' PROP' := λne p, + proto_elim proto_end (λ a m, proto_message a (λ v, g â—Ž m v â—Ž laterO_map rec)) p. +Next Obligation. + intros V PROPn ? PROPn' ? PROP ? PROP' ? g rec n p1 p2 Hp. + apply proto_elim_ne=> // a m1 m2 Hm. by repeat f_equiv. +Qed. + +Global Instance proto_map_aux_contractive {V} + `{!Cofe PROPn, !Cofe PROPn', !Cofe PROP, !Cofe PROP'} (g : PROP -n> PROP') : + Contractive (proto_map_aux (V:=V) (PROPn:=PROPn) (PROPn':=PROPn') g). +Proof. + intros n rec1 rec2 Hrec p. simpl. apply proto_elim_ne=> // a m1 m2 Hm. + f_equiv=> v p' /=. do 2 f_equiv; [done|]. + apply Next_contractive; by dist_later_intro as n' Hn'. +Qed. + +Definition proto_map_aux_2 {V} + `{!Cofe PROPn, !Cofe PROPn', !Cofe PROP, !Cofe PROP'} + (gn : PROPn' -n> PROPn) (g : PROP -n> PROP') + (rec : proto V PROPn PROP -n> proto V PROPn' PROP') : + proto V PROPn PROP -n> proto V PROPn' PROP' := + proto_map_aux g (proto_map_aux gn rec). +Global Instance proto_map_aux_2_contractive {V} + `{!Cofe PROPn, !Cofe PROPn', !Cofe PROP, !Cofe PROP'} + (gn : PROPn' -n> PROPn) (g : PROP -n> PROP') : + Contractive (proto_map_aux_2 (V:=V) gn g). +Proof. + intros n rec1 rec2 Hrec. rewrite /proto_map_aux_2. + f_equiv. by apply proto_map_aux_contractive. +Qed. +Definition proto_map {V} + `{!Cofe PROPn, !Cofe PROPn', !Cofe PROP, !Cofe PROP'} + (gn : PROPn' -n> PROPn) (g : PROP -n> PROP') : + proto V PROPn PROP -n> proto V PROPn' PROP' := + fixpoint (proto_map_aux_2 gn g). + +Lemma proto_map_unfold {V} + `{Hcn:!Cofe PROPn, Hcn':!Cofe PROPn', Hc:!Cofe PROP, Hc':!Cofe PROP'} + (gn : PROPn' -n> PROPn) (g : PROP -n> PROP') p : + proto_map (V:=V) gn g p ≡ proto_map_aux g (proto_map g gn) p. +Proof. + apply equiv_dist=> n. revert PROPn Hcn PROPn' Hcn' PROP Hc PROP' Hc' gn g p. + induction (lt_wf n) as [n _ IH]=> + PROPn Hcn PROPn' Hcn' PROP Hc PROP' Hc' gn g p. + etrans; [apply equiv_dist, (fixpoint_unfold (proto_map_aux_2 gn g))|]. + apply proto_map_aux_contractive; constructor=> n' ?. symmetry. apply: IH. lia. +Qed. +Lemma proto_map_end {V} `{!Cofe PROPn, !Cofe PROPn', !Cofe PROP, !Cofe PROP'} + (gn : PROPn' -n> PROPn) (g : PROP -n> PROP') : + proto_map (V:=V) gn g proto_end ≡ proto_end. +Proof. by rewrite proto_map_unfold /proto_map_aux /= proto_elim_end. Qed. +Lemma proto_map_message {V} `{!Cofe PROPn, !Cofe PROPn', !Cofe PROP, !Cofe PROP'} + (gn : PROPn' -n> PROPn) (g : PROP -n> PROP') a m : + proto_map (V:=V) gn g (proto_message a m) + ≡ proto_message a (λ v, g â—Ž m v â—Ž laterO_map (proto_map g gn)). +Proof. + rewrite proto_map_unfold /proto_map_aux /=. + rewrite ->proto_elim_message; [done|]. + intros a' m1 m2 Hm. f_equiv; solve_proper. +Qed. + +Lemma proto_map_ne {V} + `{Hcn:!Cofe PROPn, Hcn':!Cofe PROPn', Hc:!Cofe PROP, Hc':!Cofe PROP'} + (gn1 gn2 : PROPn' -n> PROPn) (g1 g2 : PROP -n> PROP') p n : + gn1 ≡{n}≡ gn2 → g1 ≡{n}≡ g2 → + proto_map (V:=V) gn1 g1 p ≡{n}≡ proto_map (V:=V) gn2 g2 p. +Proof. + revert PROPn Hcn PROPn' Hcn' PROP Hc PROP' Hc' gn1 gn2 g1 g2 p. + induction (lt_wf n) as [n _ IH]=> + PROPn ? PROPn' ? PROP ? PROP' ? gn1 gn2 g1 g2 p Hgn Hg /=. + destruct (proto_case p) as [->|(a & m & ->)]; [by rewrite !proto_map_end|]. + rewrite !proto_map_message /=. + apply proto_message_ne=> // v p' /=. f_equiv; [done|]. f_equiv. + apply Next_contractive; dist_later_intro as n' Hn'; eauto using dist_le with lia. +Qed. +Lemma proto_map_ext {V} `{!Cofe PROPn, !Cofe PROPn', !Cofe PROP, !Cofe PROP'} + (gn1 gn2 : PROPn' -n> PROPn) (g1 g2 : PROP -n> PROP') p : + gn1 ≡ gn2 → g1 ≡ g2 → proto_map (V:=V) gn1 g1 p ≡ proto_map (V:=V) gn2 g2 p. +Proof. + intros Hgn Hg. apply equiv_dist=> n. + apply proto_map_ne=> // ?; by apply equiv_dist. +Qed. +Lemma proto_map_id {V} `{Hcn:!Cofe PROPn, Hc:!Cofe PROP} (p : proto V PROPn PROP) : + proto_map cid cid p ≡ p. +Proof. + apply equiv_dist=> n. revert PROPn Hcn PROP Hc p. + induction (lt_wf n) as [n _ IH]=> PROPn ? PROP ? p /=. + destruct (proto_case p) as [->|(a & m & ->)]; [by rewrite !proto_map_end|]. + rewrite !proto_map_message /=. apply proto_message_ne=> // v p' /=. f_equiv. + apply Next_contractive; dist_later_intro as n' Hn'; auto. +Qed. +Lemma proto_map_compose {V} + `{Hcn:!Cofe PROPn, Hcn':!Cofe PROPn', Hcn'':!Cofe PROPn'', + Hc:!Cofe PROP, Hc':!Cofe PROP', Hc'':!Cofe PROP''} + (gn1 : PROPn'' -n> PROPn') (gn2 : PROPn' -n> PROPn) + (g1 : PROP -n> PROP') (g2 : PROP' -n> PROP'') (p : proto V PROPn PROP) : + proto_map (gn2 â—Ž gn1) (g2 â—Ž g1) p ≡ proto_map gn1 g2 (proto_map gn2 g1 p). +Proof. + apply equiv_dist=> n. revert PROPn Hcn PROPn' Hcn' PROPn'' Hcn'' + PROP Hc PROP' Hc' PROP'' Hc'' gn1 gn2 g1 g2 p. + induction (lt_wf n) as [n _ IH]=> PROPn ? PROPn' ? PROPn'' ? + PROP ? PROP' ? PROP'' ? gn1 gn2 g1 g2 p /=. + destruct (proto_case p) as [->|(a & c & ->)]; [by rewrite !proto_map_end|]. + rewrite !proto_map_message /=. apply proto_message_ne=> // v p' /=. + do 3 f_equiv. apply Next_contractive; dist_later_intro as n' Hn'; simpl; auto. +Qed. + +Program Definition protoOF (V : Type) (Fn F : oFunctor) + `{!∀ A B `{!Cofe A, !Cofe B}, Cofe (oFunctor_car Fn A B)} + `{!∀ A B `{!Cofe A, !Cofe B}, Cofe (oFunctor_car F A B)} : oFunctor := {| + oFunctor_car A _ B _ := proto V (oFunctor_car Fn B A) (oFunctor_car F A B); + oFunctor_map A1 _ A2 _ B1 _ B2 _ fg := + proto_map (oFunctor_map Fn (fg.2, fg.1)) (oFunctor_map F fg) +|}. +Next Obligation. + intros V Fn F ?? A1 ? A2 ? B1 ? B2 ? n f g [??] p; simpl in *. + apply proto_map_ne=> // y; by apply oFunctor_map_ne. +Qed. +Next Obligation. + intros V Fn F ?? A ? B ? p; simpl in *. rewrite /= -{2}(proto_map_id p). + apply proto_map_ext=> //= y; by rewrite oFunctor_map_id. +Qed. +Next Obligation. + intros V Fn F ?? A1 ? A2 ? A3 ? B1 ? B2 ? B3 ? f g f' g' p; simpl in *. + rewrite -proto_map_compose. + apply proto_map_ext=> //= y; by rewrite ofe.oFunctor_map_compose. +Qed. + +Global Instance protoOF_contractive (V : Type) (Fn F : oFunctor) + `{!∀ A B `{!Cofe A, !Cofe B}, Cofe (oFunctor_car Fn A B)} + `{!∀ A B `{!Cofe A, !Cofe B}, Cofe (oFunctor_car F A B)} : + oFunctorContractive Fn → oFunctorContractive F → + oFunctorContractive (protoOF V Fn F). +Proof. + intros HFn HF A1 ? A2 ? B1 ? B2 ? n f g Hfg p; simpl in *. + apply proto_map_ne=> y //=. + + apply HFn. dist_later_intro as n' Hn'. f_equiv; apply Hfg. + + apply HF. by dist_later_intro as n' Hn'. +Qed. -- GitLab From 0baa9a22c10a75aca4ef6f0ead52977ff9740f96 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Sun, 21 Jan 2024 16:07:51 +0100 Subject: [PATCH 03/81] WIP: Proof of iProto_consistent_step --- theories/channel/multi_proto.v | 987 ++++----------------------------- 1 file changed, 110 insertions(+), 877 deletions(-) diff --git a/theories/channel/multi_proto.v b/theories/channel/multi_proto.v index c2fbacb..878792e 100644 --- a/theories/channel/multi_proto.v +++ b/theories/channel/multi_proto.v @@ -509,6 +509,8 @@ Section proto. End proto. +Instance iProto_inhabited {Σ V} : Inhabited (iProto Σ V) := populate (END). + Definition can_step {Σ V} (rec : gmap nat (iProto Σ V) → iProp Σ) (ps : gmap nat (iProto Σ V)) (i j : nat) : iProp Σ := (∀ a1 a2 m1 m2, @@ -740,795 +742,26 @@ Section proto. Implicit Types p pl pr : iProto Σ V. Implicit Types m : iMsg Σ V. - (** ** Equality *) - Lemma iProto_case p : p ≡ END ∨ ∃ a m, p ≡ <a> m. - Proof. - rewrite iProto_message_eq iProto_end_eq. - destruct (proto_case p) as [|(a&m&?)]; [by left|right]. - by exists a, (IMsg m). - Qed. - Lemma iProto_message_equivI `{!BiInternalEq SPROP} a1 a2 m1 m2 : - (<a1> m1) ≡ (<a2> m2) ⊣⊢@{SPROP} ⌜ a1 = a2 ⌠∧ - (∀ v lp, iMsg_car m1 v lp ≡ iMsg_car m2 v lp). - Proof. rewrite iProto_message_eq. apply proto_message_equivI. Qed. - - Lemma iProto_message_end_equivI `{!BiInternalEq SPROP} a m : - (<a> m) ≡ END ⊢@{SPROP} False. - Proof. rewrite iProto_message_eq iProto_end_eq. apply proto_message_end_equivI. Qed. - Lemma iProto_end_message_equivI `{!BiInternalEq SPROP} a m : - END ≡ (<a> m) ⊢@{SPROP} False. - Proof. by rewrite internal_eq_sym iProto_message_end_equivI. Qed. - - (** ** Non-expansiveness of operators *) - Global Instance iMsg_car_proper : - Proper (iMsg_equiv ==> (=) ==> (≡) ==> (≡)) (iMsg_car (Σ:=Σ) (V:=V)). - Proof. - intros m1 m2 meq v1 v2 veq p1 p2 peq. rewrite meq. - f_equiv; [ by f_equiv | done ]. - Qed. - Global Instance iMsg_car_ne n : - Proper (iMsg_dist n ==> (=) ==> (dist n) ==> (dist n)) (iMsg_car (Σ:=Σ) (V:=V)). - Proof. - intros m1 m2 meq v1 v2 veq p1 p2 peq. rewrite meq. - f_equiv; [ by f_equiv | done ]. - Qed. - - Global Instance iMsg_contractive v n : - Proper (dist n ==> dist_later n ==> dist n) (iMsg_base (Σ:=Σ) (V:=V) v). - Proof. rewrite iMsg_base_eq=> P1 P2 HP p1 p2 Hp w q /=. solve_contractive. Qed. - Global Instance iMsg_ne v : NonExpansive2 (iMsg_base (Σ:=Σ) (V:=V) v). - Proof. rewrite iMsg_base_eq=> P1 P2 HP p1 p2 Hp w q /=. solve_proper. Qed. - Global Instance iMsg_proper v : - Proper ((≡) ==> (≡) ==> (≡)) (iMsg_base (Σ:=Σ) (V:=V) v). - Proof. apply (ne_proper_2 _). Qed. - - Global Instance iMsg_exist_ne A n : - Proper (pointwise_relation _ (dist n) ==> (dist n)) (@iMsg_exist Σ V A). - Proof. rewrite iMsg_exist_eq=> m1 m2 Hm v p /=. f_equiv=> x. apply Hm. Qed. - Global Instance iMsg_exist_proper A : - Proper (pointwise_relation _ (≡) ==> (≡)) (@iMsg_exist Σ V A). - Proof. rewrite iMsg_exist_eq=> m1 m2 Hm v p /=. f_equiv=> x. apply Hm. Qed. - - Global Instance msg_tele_base (v:V) (P : iProp Σ) (p : iProto Σ V) : - MsgTele (TT:=TeleO) (MSG v {{ P }}; p) v P p. - Proof. done. Qed. - Global Instance msg_tele_exist {A} {TT : A → tele} (m : A → iMsg Σ V) tv tP tp : - (∀ x, MsgTele (TT:=TT x) (m x) (tv x) (tP x) (tp x)) → - MsgTele (TT:=TeleS TT) (∃ x, m x) tv tP tp. - Proof. intros Hm. rewrite /MsgTele /=. f_equiv=> x. apply Hm. Qed. - - Global Instance iProto_message_ne a : - NonExpansive (iProto_message (Σ:=Σ) (V:=V) a). - Proof. rewrite iProto_message_eq. solve_proper. Qed. - Global Instance iProto_message_proper a : - Proper ((≡) ==> (≡)) (iProto_message (Σ:=Σ) (V:=V) a). - Proof. apply (ne_proper _). Qed. - - Lemma iProto_message_equiv {TT1 TT2 : tele} a1 a2 - (m1 m2 : iMsg Σ V) - (v1 : TT1 -t> V) (v2 : TT2 -t> V) - (P1 : TT1 -t> iProp Σ) (P2 : TT2 -t> iProp Σ) - (prot1 : TT1 -t> iProto Σ V) (prot2 : TT2 -t> iProto Σ V) : - MsgTele m1 v1 P1 prot1 → - MsgTele m2 v2 P2 prot2 → - ⌜ a1 = a2 ⌠-∗ - (■∀.. (xs1 : TT1), tele_app P1 xs1 -∗ - ∃.. (xs2 : TT2), ⌜tele_app v1 xs1 = tele_app v2 xs2⌠∗ - â–· (tele_app prot1 xs1 ≡ tele_app prot2 xs2) ∗ - tele_app P2 xs2) -∗ - (■∀.. (xs2 : TT2), tele_app P2 xs2 -∗ - ∃.. (xs1 : TT1), ⌜tele_app v1 xs1 = tele_app v2 xs2⌠∗ - â–· (tele_app prot1 xs1 ≡ tele_app prot2 xs2) ∗ - tele_app P1 xs1) -∗ - (<a1> m1) ≡ (<a2> m2). - Proof. - iIntros (Hm1 Hm2 Heq) "#Heq1 #Heq2". - unfold MsgTele in Hm1. rewrite Hm1. clear Hm1. - unfold MsgTele in Hm2. rewrite Hm2. clear Hm2. - rewrite iProto_message_eq proto_message_equivI. - iSplit; [ done | ]. - iIntros (v p'). - do 2 rewrite iMsg_texist_exist. - rewrite iMsg_base_eq /=. - iApply prop_ext. - iIntros "!>". iSplit. - - iDestruct 1 as (xs1 Hveq1) "[Hrec1 HP1]". - iDestruct ("Heq1" with "HP1") as (xs2 Hveq2) "[Hrec2 HP2]". - iExists xs2. rewrite -Hveq1 Hveq2. - iSplitR; [ done | ]. iSplitR "HP2"; [ | done ]. - iRewrite -"Hrec1". iApply later_equivI. iIntros "!>". by iRewrite "Hrec2". - - iDestruct 1 as (xs2 Hveq2) "[Hrec2 HP2]". - iDestruct ("Heq2" with "HP2") as (xs1 Hveq1) "[Hrec1 HP1]". - iExists xs1. rewrite -Hveq2 Hveq1. - iSplitR; [ done | ]. iSplitR "HP1"; [ | done ]. - iRewrite -"Hrec2". iApply later_equivI. iIntros "!>". by iRewrite "Hrec1". - Qed. - - (** Helpers *) - Lemma iMsg_map_base f v P p : - NonExpansive f → - iMsg_map f (MSG v {{ P }}; p) ≡ (MSG v {{ P }}; f p)%msg. - Proof. - rewrite iMsg_base_eq. intros ? v' p'; simpl. iSplit. - - iDestruct 1 as (p'') "[(->&Hp&$) Hp']". iSplit; [done|]. - iRewrite "Hp'". iIntros "!>". by iRewrite "Hp". - - iIntros "(->&Hp'&$)". iExists p. iRewrite -"Hp'". auto. - Qed. - Lemma iMsg_map_exist {A} f (m : A → iMsg Σ V) : - iMsg_map f (∃ x, m x) ≡ (∃ x, iMsg_map f (m x))%msg. - Proof. - rewrite iMsg_exist_eq. intros v' p'; simpl. iSplit. - - iDestruct 1 as (p'') "[H Hp']". iDestruct "H" as (x) "H"; auto. - - iDestruct 1 as (x p'') "[Hm Hp']". auto. - Qed. - - (** ** Dual *) - Global Instance iProto_dual_ne : NonExpansive (@iProto_dual Σ V). - Proof. rewrite iProto_dual_eq. solve_proper. Qed. - Global Instance iProto_dual_proper : Proper ((≡) ==> (≡)) (@iProto_dual Σ V). - Proof. apply (ne_proper _). Qed. - Global Instance iProto_dual_if_ne d : NonExpansive (@iProto_dual_if Σ V d). - Proof. solve_proper. Qed. - Global Instance iProto_dual_if_proper d : - Proper ((≡) ==> (≡)) (@iProto_dual_if Σ V d). - Proof. apply (ne_proper _). Qed. - - Lemma iProto_dual_end : iProto_dual (Σ:=Σ) (V:=V) END ≡ END. - Proof. - rewrite iProto_end_eq iProto_dual_eq /iProto_dual_def /iProto_map_app. - etrans; [apply (fixpoint_unfold (iProto_map_app_aux _ _))|]; simpl. - by rewrite proto_elim_end. - Qed. - Lemma iProto_dual_message a m : - iProto_dual (<a> m) ≡ <action_dual a> iMsg_dual m. - Proof. - rewrite iProto_message_eq iProto_dual_eq /iProto_dual_def /iProto_map_app. - etrans; [apply (fixpoint_unfold (iProto_map_app_aux _ _))|]; simpl. - rewrite /iProto_message_def. rewrite ->proto_elim_message; [done|]. - intros a' m1 m2 Hm; f_equiv; solve_proper. - Qed. - Lemma iMsg_dual_base v P p : - iMsg_dual (MSG v {{ P }}; p) ≡ (MSG v {{ P }}; iProto_dual p)%msg. - Proof. apply iMsg_map_base, _. Qed. - Lemma iMsg_dual_exist {A} (m : A → iMsg Σ V) : - iMsg_dual (∃ x, m x) ≡ (∃ x, iMsg_dual (m x))%msg. - Proof. apply iMsg_map_exist. Qed. - - Global Instance iProto_dual_involutive : Involutive (≡) (@iProto_dual Σ V). - Proof. - intros p. apply (uPred.internal_eq_soundness (M:=iResUR Σ)). - iLöb as "IH" forall (p). destruct (iProto_case p) as [->|(a&m&->)]. - { by rewrite !iProto_dual_end. } - rewrite !iProto_dual_message involutive. - iApply iProto_message_equivI; iSplit; [done|]; iIntros (v p') "/=". - iApply prop_ext; iIntros "!>"; iSplit. - - iDestruct 1 as (pd) "[H Hp']". iRewrite "Hp'". - iDestruct "H" as (pdd) "[H #Hpd]". - iApply (internal_eq_rewrite); [|done]; iIntros "!>". - iRewrite "Hpd". by iRewrite ("IH" $! pdd). - - iIntros "H". destruct (Next_uninj p') as [p'' Hp']. iExists _. - rewrite Hp'. iSplitL; [by auto|]. iIntros "!>". by iRewrite ("IH" $! p''). - Qed. - - (** ** Append *) - Global Instance iProto_app_end_l : LeftId (≡) END (@iProto_app Σ V). - Proof. - intros p. rewrite iProto_end_eq iProto_app_eq /iProto_app_def /iProto_map_app. - etrans; [apply (fixpoint_unfold (iProto_map_app_aux _ _))|]; simpl. - by rewrite proto_elim_end. - Qed. - Lemma iProto_app_message a m p2 : (<a> m) <++> p2 ≡ <a> m <++> p2. - Proof. - rewrite iProto_message_eq iProto_app_eq /iProto_app_def /iProto_map_app. - etrans; [apply (fixpoint_unfold (iProto_map_app_aux _ _))|]; simpl. - rewrite /iProto_message_def. rewrite ->proto_elim_message; [done|]. - intros a' m1 m2 Hm. f_equiv; solve_proper. - Qed. - - Global Instance iProto_app_ne : NonExpansive2 (@iProto_app Σ V). - Proof. - assert (∀ n, Proper (dist n ==> (=) ==> dist n) (@iProto_app Σ V)). - { intros n p1 p1' Hp1 p2 p2' <-. by rewrite iProto_app_eq /iProto_app_def Hp1. } - assert (Proper ((≡) ==> (=) ==> (≡)) (@iProto_app Σ V)). - { intros p1 p1' Hp1 p2 p2' <-. by rewrite iProto_app_eq /iProto_app_def Hp1. } - intros n p1 p1' Hp1 p2 p2' Hp2. rewrite Hp1. clear p1 Hp1. - revert p1'. induction (lt_wf n) as [n _ IH]; intros p1. - destruct (iProto_case p1) as [->|(a&m&->)]. - { by rewrite !left_id. } - rewrite !iProto_app_message. f_equiv=> v p' /=. do 4 f_equiv. - f_contractive. apply IH; eauto using dist_le with lia. - Qed. - Global Instance iProto_app_proper : Proper ((≡) ==> (≡) ==> (≡)) (@iProto_app Σ V). - Proof. apply (ne_proper_2 _). Qed. - - Lemma iMsg_app_base v P p1 p2 : - ((MSG v {{ P }}; p1) <++> p2)%msg ≡ (MSG v {{ P }}; p1 <++> p2)%msg. - Proof. apply: iMsg_map_base. Qed. - Lemma iMsg_app_exist {A} (m : A → iMsg Σ V) p2 : - ((∃ x, m x) <++> p2)%msg ≡ (∃ x, m x <++> p2)%msg. - Proof. apply iMsg_map_exist. Qed. - - Global Instance iProto_app_end_r : RightId (≡) END (@iProto_app Σ V). - Proof. - intros p. apply (uPred.internal_eq_soundness (M:=iResUR Σ)). - iLöb as "IH" forall (p). destruct (iProto_case p) as [->|(a&m&->)]. - { by rewrite left_id. } - rewrite iProto_app_message. - iApply iProto_message_equivI; iSplit; [done|]; iIntros (v p') "/=". - iApply prop_ext; iIntros "!>". iSplit. - - iDestruct 1 as (p1') "[H Hp']". iRewrite "Hp'". - iApply (internal_eq_rewrite); [|done]; iIntros "!>". - by iRewrite ("IH" $! p1'). - - iIntros "H". destruct (Next_uninj p') as [p'' Hp']. iExists p''. - rewrite Hp'. iSplitL; [by auto|]. iIntros "!>". by iRewrite ("IH" $! p''). - Qed. - Global Instance iProto_app_assoc : Assoc (≡) (@iProto_app Σ V). - Proof. - intros p1 p2 p3. apply (uPred.internal_eq_soundness (M:=iResUR Σ)). - iLöb as "IH" forall (p1). destruct (iProto_case p1) as [->|(a&m&->)]. - { by rewrite !left_id. } - rewrite !iProto_app_message. - iApply iProto_message_equivI; iSplit; [done|]; iIntros (v p123) "/=". - iApply prop_ext; iIntros "!>". iSplit. - - iDestruct 1 as (p1') "[H #Hp']". - iExists (p1' <++> p2). iSplitL; [by auto|]. - iRewrite "Hp'". iIntros "!>". iApply "IH". - - iDestruct 1 as (p12) "[H #Hp123]". iDestruct "H" as (p1') "[H #Hp12]". - iExists p1'. iFrame "H". iRewrite "Hp123". - iIntros "!>". iRewrite "Hp12". by iRewrite ("IH" $! p1'). - Qed. - - Lemma iProto_dual_app p1 p2 : - iProto_dual (p1 <++> p2) ≡ iProto_dual p1 <++> iProto_dual p2. - Proof. - apply (uPred.internal_eq_soundness (M:=iResUR Σ)). - iLöb as "IH" forall (p1 p2). destruct (iProto_case p1) as [->|(a&m&->)]. - { by rewrite iProto_dual_end !left_id. } - rewrite iProto_dual_message !iProto_app_message iProto_dual_message /=. - iApply iProto_message_equivI; iSplit; [done|]; iIntros (v p12) "/=". - iApply prop_ext; iIntros "!>". iSplit. - - iDestruct 1 as (p12d) "[H #Hp12]". iDestruct "H" as (p1') "[H #Hp12d]". - iExists (iProto_dual p1'). iSplitL; [by auto|]. - iRewrite "Hp12". iIntros "!>". iRewrite "Hp12d". iApply "IH". - - iDestruct 1 as (p1') "[H #Hp12]". iDestruct "H" as (p1d) "[H #Hp1']". - iExists (p1d <++> p2). iSplitL; [by auto|]. - iRewrite "Hp12". iIntros "!>". iRewrite "Hp1'". by iRewrite ("IH" $! p1d p2). - Qed. - - Lemma iProto_consistent_flip p1 p2 : - iProto_consistent p2 p1 -∗ - iProto_consistent p1 p2. - Proof. - iLöb as "IH" forall (p1 p2). - iIntros "Hprot". - rewrite iProto_consistent_unfold /iProto_consistent_pre. - rewrite iProto_consistent_unfold /iProto_consistent_pre. - iIntros (a1 a2 m1 m2) "Hm1 Hm2". - iDestruct ("Hprot" with "Hm2 Hm1") as "Hprot". - destruct a1, a2; [done| | |done]. - - iIntros (v p1') "Hm1". - iDestruct ("Hprot" with "Hm1") as (p2') "[Hprot Hrec]". - iExists p2'. iFrame. by iApply "IH". - - iIntros (v p2') "Hm2". - iDestruct ("Hprot" with "Hm2") as (p1') "[Hprot Hrec]". - iExists p1'. iFrame. by iApply "IH". - Qed. - - Lemma iProto_consistent_step m1 m2 v p1 : - iProto_consistent (<!> m1) (<?> m2) -∗ + Lemma iProto_consistent_step ps m1 m2 i j v p1 : + ps !!! i ≡ (<(Send j)> m1) -∗ + ps !!! j ≡ (<(Recv i)> m2) -∗ + iProto_consistent ps -∗ iMsg_car m1 v (Next p1) -∗ - ∃ p2, iMsg_car m2 v (Next p2) ∗ â–· iProto_consistent p1 p2. + ∃ p2, iMsg_car m2 v (Next p2) ∗ + â–· iProto_consistent (<[i := p1]>(<[j := p2]>ps)). Proof. - iIntros "Hprot Hm1". + iIntros "#Hi #Hj Hprot Hm1". rewrite iProto_consistent_unfold /iProto_consistent_pre. - iDestruct ("Hprot" with "[//] [//] Hm1") as (p2) "[Hm2 Hprot]". + iDestruct ("Hprot" $!i j with "[] [] Hi Hj Hm1") as (p2) "[Hm2 Hprot]". + { rewrite !lookup_total_alt elem_of_dom. + destruct (ps !! i) eqn: Hi; [done|]=> /=. + by rewrite iProto_end_message_equivI. } + { rewrite !lookup_total_alt elem_of_dom. + destruct (ps !! j) eqn: Hj; [done|]=> /=. + by rewrite iProto_end_message_equivI. } iExists p2. iFrame. Qed. - Lemma iProto_consistent_dual p : - ⊢ iProto_consistent p (iProto_dual p). - Proof. - iLöb as "IH" forall (p). - rewrite iProto_consistent_unfold. - iIntros (a1 a2 m1 m2) "#Hmeq1 #Hmeq2". - destruct a1, a2; [done| | |done]. - - iIntros (v p') "Hm1". - iRewrite "Hmeq1" in "Hmeq2". - rewrite iProto_dual_message. simpl. - rewrite iProto_message_equivI. - iDestruct "Hmeq2" as (_) "Hmeq2". - iSpecialize ("Hmeq2" $! v (Next (iProto_dual p'))). - iExists (iProto_dual p'). - iRewrite -"Hmeq2". - iSplitL; [|iNext; by iApply "IH"]. - simpl. iExists p'. iFrame. done. - - iIntros (v p') "Hm2". - iRewrite "Hmeq1" in "Hmeq2". - rewrite iProto_dual_message. simpl. - rewrite iProto_message_equivI. - iDestruct "Hmeq2" as (_) "Hmeq2". - iSpecialize ("Hmeq2" $! v (Next p')). - iRewrite -"Hmeq2" in "Hm2". - simpl. - iDestruct "Hm2" as (p'') "[Hm2 Hmeq']". - iExists p''. iFrame. - iNext. iRewrite "Hmeq'". - iApply "IH". - Qed. - - Lemma iProto_consistent_le_l pl pl' pr : - iProto_consistent pl pr -∗ pl ⊑ pl' -∗ iProto_consistent pl' pr. - Proof. iIntros "Hprot Hle". by iApply "Hle". Qed. - Lemma iProto_consistent_le_r pl pr pr' : - iProto_consistent pl pr -∗ pr ⊑ pr' -∗ iProto_consistent pl pr'. - Proof. - iIntros "H Hle". iApply iProto_consistent_flip. - iApply "Hle". by iApply iProto_consistent_flip. - Qed. - - Lemma iProto_le_refl p : ⊢ p ⊑ p. - Proof. iIntros (p') "$". Qed. - - Lemma iProto_le_end : ⊢ END ⊑ (END : iProto Σ V). - Proof. iApply iProto_le_refl. Qed. - - (* Lemma iProto_le_send m1 m2 : *) - (* (∀ v p2', iMsg_car m2 v (Next p2') -∗ ∃ p1', *) - (* â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1')) -∗ *) - (* (<!> m1) ⊑ (<!> m2). *) - (* Proof. *) - (* iLöb as "IH" forall (m1 m2). *) - (* iIntros "Hle". *) - (* repeat (unfold iProto_le at 3). *) - (* iIntros (p) "Hprot". *) - (* repeat rewrite iProto_consistent_unfold; unfold iProto_consistent_pre. *) - (* iSplit; [iDestruct "Hprot" as "[Hprot _]"|iDestruct "Hprot" as "[_ Hprot]"]. *) - (* { iIntros (a m) "H". *) - (* iDestruct (iProto_message_equivI with "H") as (Heq) "{H} #Hmeq"; subst. *) - (* iIntros (v p') "Hm". *) - (* iSpecialize ("Hprot" with "[//]"). *) - (* iSpecialize ("Hmeq" $! v (Next p')). iRewrite -"Hmeq" in "Hm". *) - (* iSpecialize ("Hle" $! _ _ with "Hm"). iDestruct "Hle" as (p'') "[Hle Hm]". *) - (* iApply "Hle". iApply "Hprot". iApply "Hm". } *) - (* { iIntros (a m) "H". *) - (* iSpecialize ("Hprot" $! _ _ with "H"). *) - (* destruct a. *) - (* - iIntros (v p') "Hm". *) - (* iApply ("IH" with "[Hle]"). *) - (* { iNext. iIntros (v' p'') "Hm". *) - (* iSpecialize ("Hle" with "Hm"). *) - (* iDestruct "Hle" as (p''') "[Hle Hm]". *) - (* iExists p'''. iSplitR "Hm"; [iApply "Hle" | iApply "Hm"]. } *) - (* iApply "Hprot". iApply "Hm". *) - (* - iIntros (v vs Heq); subst. *) - (* iSpecialize ("Hprot" $! _ _ eq_refl). *) - (* iDestruct "Hprot" as (p') "[Hm Hprot]". *) - (* iExists (p'). iFrame "Hm". *) - (* iApply ("IH" with "[Hle] [Hprot]"); [|iApply "Hprot"]. *) - (* iNext. iIntros (v' p'') "Hm". *) - (* iSpecialize ("Hle" with "Hm"). iDestruct "Hle" as (p''') "[Hle Hm]". *) - (* iExists p'''; iSplitR "Hm"; [iApply "Hle" |iApply "Hm"]. } *) - (* Qed. *) - - (* Lemma iProto_le_recv m1 m2 : *) - (* (∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', *) - (* â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2')) -∗ *) - (* (<?> m1) ⊑ (<?> m2). *) - (* Proof. *) - (* iIntros "H" (vsl vsr p) "Hprot". *) - (* iLöb as "IH" forall (vsl vsr m2 p). *) - (* iEval (rewrite iProto_consistent_unfold). *) - (* rewrite /iProto_consistent_pre. *) - (* rewrite iProto_consistent_unfold /iProto_consistent_pre. *) - (* iSplit; last first. *) - (* { *) - (* iDestruct "Hprot" as "[_ Hprot]". *) - (* iIntros ([] m) "Heq". *) - (* { iIntros (v vs) "Hm". *) - (* iDestruct ("Hprot" with "Heq Hm") as "Hprot". *) - (* iIntros "!>". *) - (* by iApply ("IH" with "H"). } *) - (* { iIntros (v vs) "Hm". *) - (* iDestruct ("Hprot" with "Heq Hm") as (p') "[Hm Hprot]". *) - (* iExists p'. iFrame. *) - (* iIntros "!>". *) - (* by iApply ("IH" with "H"). } *) - (* } *) - (* destruct vsr. *) - (* { iIntros (a m) "Heq". rewrite iProto_message_equivI. *) - (* iDestruct "Heq" as (<-) "Heq". *) - (* iIntros (v vs Heq). done. } *) - (* iIntros (a m) "Heq". rewrite iProto_message_equivI. *) - (* iDestruct "Heq" as (<-) "Heq". *) - (* iIntros (w vs Heq). *) - (* assert (v = w) as <- by set_solver. *) - (* assert (vsr = vs) as <- by set_solver. *) - (* iDestruct "Hprot" as "[Hprot _]". *) - (* iDestruct ("Hprot" with "[//] [//]") as (p') "[Hm Hprot]". *) - (* iDestruct ("H" with "Hm") as (p'') "[Hle H]". *) - (* iExists p''. *) - (* iSpecialize ("Heq" $! v (Next p'')). iRewrite -"Heq". *) - (* iFrame. iIntros "!>". *) - (* iApply (iProto_consistent_le_l with "Hprot Hle"). *) - (* Qed. *) - - (* Lemma iProto_le_base a v P p1 p2 : *) - (* â–· (p1 ⊑ p2) -∗ *) - (* (<a> MSG v {{ P }}; p1) ⊑ (<a> MSG v {{ P }}; p2). *) - (* Proof. *) - (* rewrite iMsg_base_eq. iIntros "H". destruct a. *) - (* - iApply iProto_le_send. iIntros (v' p') "(->&Hp&$)". *) - (* iExists p1. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". *) - (* - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&$)". *) - (* iExists p2. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". *) - (* Qed. *) - - (* Lemma iProto_le_end_inv_l p : p ⊑ END -∗ (p ≡ END). *) - (* Proof. *) - (* rewrite iProto_le_unfold. iIntros "[[Hp _]|H] //". *) - (* iDestruct "H" as (a1 a2 m1 m2) "(_ & Heq & _)". *) - (* by iDestruct (iProto_end_message_equivI with "Heq") as %[]. *) - (* Qed. *) - - (* Lemma iProto_le_end_inv_r p : END ⊑ p -∗ (p ≡ END). *) - (* Proof. *) - (* rewrite iProto_le_unfold. iIntros "[[_ Hp]|H] //". *) - (* iDestruct "H" as (a1 a2 m1 m2) "(Heq & _ & _)". *) - (* iDestruct (iProto_end_message_equivI with "Heq") as %[]. *) - (* Qed. *) - - (* Lemma iProto_le_send_inv p1 m2 : *) - (* p1 ⊑ (<!> m2) -∗ ∃ a1 m1, *) - (* (p1 ≡ <a1> m1) ∗ *) - (* match a1 with *) - (* | Send => ∀ v p2', *) - (* iMsg_car m2 v (Next p2') -∗ ∃ p1', *) - (* â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1') *) - (* | Recv => ∀ v1 v2 p1' p2', *) - (* iMsg_car m1 v1 (Next p1') -∗ iMsg_car m2 v2 (Next p2') -∗ ∃ pt, *) - (* â–· (p1' ⊑ <!> MSG v2; pt) ∗ â–· ((<?> MSG v1; pt) ⊑ p2') *) - (* end. *) - (* Proof. *) - (* rewrite iProto_le_unfold. iIntros "[[_ Heq]|H]". *) - (* { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } *) - (* iDestruct "H" as (a1 a2 m1 m2') "(Hp1 & Hp2 & H)". *) - (* iExists _, _; iSplit; [done|]. destruct a1, a2. *) - (* - iIntros (v p2') "Hm2". *) - (* iDestruct (iProto_message_equivI with "Hp2") as (_) "{Hp2} #Hm". *) - (* iApply "H". by iRewrite -("Hm" $! v (Next p2')). *) - (* - done. *) - (* - iIntros (v1 v2 p1' p2') "Hm1 Hm2". *) - (* iDestruct (iProto_message_equivI with "Hp2") as (_) "{Hp2} #Hm". *) - (* iApply ("H" with "Hm1"). by iRewrite -("Hm" $! v2 (Next p2')). *) - (* - iDestruct (iProto_message_equivI with "Hp2") as ([=]) "_". *) - (* Qed. *) - (* Lemma iProto_le_send_send_inv m1 m2 v p2' : *) - (* (<!> m1) ⊑ (<!> m2) -∗ *) - (* iMsg_car m2 v (Next p2') -∗ ∃ p1', â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1'). *) - (* Proof. *) - (* iIntros "H Hm2". iDestruct (iProto_le_send_inv with "H") as (a m1') "[Hm1 H]". *) - (* iDestruct (iProto_message_equivI with "Hm1") as (<-) "Hm1". *) - (* iDestruct ("H" with "Hm2") as (p1') "[Hle Hm]". *) - (* iRewrite -("Hm1" $! v (Next p1')) in "Hm". auto with iFrame. *) - (* Qed. *) - (* Lemma iProto_le_recv_send_inv m1 m2 v1 v2 p1' p2' : *) - (* (<?> m1) ⊑ (<!> m2) -∗ *) - (* iMsg_car m1 v1 (Next p1') -∗ iMsg_car m2 v2 (Next p2') -∗ ∃ pt, *) - (* â–· (p1' ⊑ <!> MSG v2; pt) ∗ â–· ((<?> MSG v1; pt) ⊑ p2'). *) - (* Proof. *) - (* iIntros "H Hm1 Hm2". iDestruct (iProto_le_send_inv with "H") as (a m1') "[Hm H]". *) - (* iDestruct (iProto_message_equivI with "Hm") as (<-) "{Hm} #Hm". *) - (* iApply ("H" with "[Hm1] Hm2"). by iRewrite -("Hm" $! v1 (Next p1')). *) - (* Qed. *) - - (* Lemma iProto_le_recv_inv p1 m2 : *) - (* p1 ⊑ (<?> m2) -∗ ∃ m1, *) - (* (p1 ≡ <?> m1) ∗ *) - (* ∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', *) - (* â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). *) - (* Proof. *) - (* rewrite iProto_le_unfold. iIntros "[[_ Heq]|H]". *) - (* { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } *) - (* iDestruct "H" as (a1 a2 m1 m2') "(Hp1 & Hp2 & H)". *) - (* iExists m1. *) - (* iDestruct (iProto_message_equivI with "Hp2") as (<-) "{Hp2} #Hm2". *) - (* destruct a1; [done|]. iSplit; [done|]. *) - (* iIntros (v p1') "Hm". iDestruct ("H" with "Hm") as (p2') "[Hle Hm]". *) - (* iExists p2'. iIntros "{$Hle}". by iRewrite ("Hm2" $! v (Next p2')). *) - (* Qed. *) - (* Lemma iProto_le_recv_recv_inv m1 m2 v p1' : *) - (* (<?> m1) ⊑ (<?> m2) -∗ *) - (* iMsg_car m1 v (Next p1') -∗ ∃ p2', â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). *) - (* Proof. *) - (* iIntros "H Hm2". iDestruct (iProto_le_recv_inv with "H") as (m1') "[Hm1 H]". *) - (* iApply "H". iDestruct (iProto_message_equivI with "Hm1") as (_) "Hm1". *) - (* by iRewrite -("Hm1" $! v (Next p1')). *) - (* Qed. *) - - (* Lemma iProto_le_refl p : ⊢ p ⊑ p. *) - (* Proof. *) - (* iLöb as "IH" forall (p). destruct (iProto_case p) as [->|([]&m&->)]. *) - (* - iApply iProto_le_end. *) - (* - iApply iProto_le_send. auto 10 with iFrame. *) - (* - iApply iProto_le_recv. auto 10 with iFrame. *) - (* Qed. *) - - Lemma iProto_le_trans p1 p2 p3 : p1 ⊑ p2 -∗ p2 ⊑ p3 -∗ p1 ⊑ p3. - Proof. - iIntros "H1 H2" (p) "Hprot". - iApply "H2". iApply "H1". done. - Qed. - - (* Lemma iProto_le_payload_elim_l a m v P p : *) - (* (P -∗ (<?> MSG v; p) ⊑ (<a> m)) -∗ *) - (* (<?> MSG v {{ P }}; p) ⊑ (<a> m). *) - (* Proof. *) - (* rewrite iMsg_base_eq. iIntros "H". destruct a. *) - (* - iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= (#?&#?&HP) Hm2 /=". *) - (* iApply (iProto_le_recv_send_inv with "(H HP)"); simpl; auto. *) - (* - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&HP)". *) - (* iApply (iProto_le_recv_recv_inv with "(H HP)"); simpl; auto. *) - (* Qed. *) - (* Lemma iProto_le_payload_elim_r a m v P p : *) - (* (P -∗ (<a> m) ⊑ (<!> MSG v; p)) -∗ *) - (* (<a> m) ⊑ (<!> MSG v {{ P }}; p). *) - (* Proof. *) - (* rewrite iMsg_base_eq. iIntros "H". destruct a. *) - (* - iApply iProto_le_send. iIntros (v' p') "(->&Hp&HP)". *) - (* iApply (iProto_le_send_send_inv with "(H HP)"); simpl; auto. *) - (* - iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1 (->&#?&HP) /=". *) - (* iApply (iProto_le_recv_send_inv with "(H HP) Hm1"); simpl; auto. *) - (* Qed. *) - (* Lemma iProto_le_payload_intro_l v P p : *) - (* P -∗ (<!> MSG v {{ P }}; p) ⊑ (<!> MSG v; p). *) - (* Proof. *) - (* rewrite iMsg_base_eq. *) - (* iIntros "HP". iApply iProto_le_send. iIntros (v' p') "(->&Hp&_) /=". *) - (* iExists p'. iSplitR; [iApply iProto_le_refl|]. auto. *) - (* Qed. *) - (* Lemma iProto_le_payload_intro_r v P p : *) - (* P -∗ (<?> MSG v; p) ⊑ (<?> MSG v {{ P }}; p). *) - (* Proof. *) - (* rewrite iMsg_base_eq. *) - (* iIntros "HP". iApply iProto_le_recv. iIntros (v' p') "(->&Hp&_) /=". *) - (* iExists p'. iSplitR; [iApply iProto_le_refl|]. auto. *) - (* Qed. *) - - (* Lemma iProto_le_exist_elim_l {A} (m1 : A → iMsg Σ V) a m2 : *) - (* (∀ x, (<?> m1 x) ⊑ (<a> m2)) -∗ *) - (* (<? x> m1 x) ⊑ (<a> m2). *) - (* Proof. *) - (* rewrite iMsg_exist_eq. iIntros "H". destruct a. *) - (* - iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1 Hm2 /=". *) - (* iDestruct "Hm1" as (x) "Hm1". *) - (* iApply (iProto_le_recv_send_inv with "H Hm1 Hm2"). *) - (* - iApply iProto_le_recv. iIntros (v p1') "/=". iDestruct 1 as (x) "Hm". *) - (* by iApply (iProto_le_recv_recv_inv with "H"). *) - (* Qed. *) - - (* Lemma iProto_le_exist_elim_l_inhabited `{!Inhabited A} (m : A → iMsg Σ V) p : *) - (* (∀ x, (<?> m x) ⊑ p) -∗ *) - (* (<? x> m x) ⊑ p. *) - (* Proof. *) - (* rewrite iMsg_exist_eq. iIntros "H". *) - (* destruct (iProto_case p) as [Heq | [a [m' Heq]]]. *) - (* - unshelve iSpecialize ("H" $!inhabitant); first by apply _. *) - (* rewrite Heq. *) - (* iDestruct (iProto_le_end_inv_l with "H") as "H". *) - (* rewrite iProto_end_eq iProto_message_eq. *) - (* iDestruct (proto_message_end_equivI with "H") as "[]". *) - (* - iEval (rewrite Heq). destruct a. *) - (* + iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1 Hm2 /=". *) - (* iDestruct "Hm1" as (x) "Hm1". *) - (* iSpecialize ("H" $! x). rewrite Heq. *) - (* iApply (iProto_le_recv_send_inv with "H Hm1 Hm2"). *) - (* + iApply iProto_le_recv. iIntros (v p1') "/=". iDestruct 1 as (x) "Hm". *) - (* iSpecialize ("H" $! x). rewrite Heq. *) - (* by iApply (iProto_le_recv_recv_inv with "H"). *) - (* Qed. *) - - (* Lemma iProto_le_exist_elim_r {A} a m1 (m2 : A → iMsg Σ V) : *) - (* (∀ x, (<a> m1) ⊑ (<!> m2 x)) -∗ *) - (* (<a> m1) ⊑ (<! x> m2 x). *) - (* Proof. *) - (* rewrite iMsg_exist_eq. iIntros "H". destruct a. *) - (* - iApply iProto_le_send. iIntros (v p2'). iDestruct 1 as (x) "Hm". *) - (* by iApply (iProto_le_send_send_inv with "H"). *) - (* - iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1". *) - (* iDestruct 1 as (x) "Hm2". *) - (* iApply (iProto_le_recv_send_inv with "H Hm1 Hm2"). *) - (* Qed. *) - (* Lemma iProto_le_exist_elim_r_inhabited `{Hinh : Inhabited A} p (m : A → iMsg Σ V) : *) - (* (∀ x, p ⊑ (<!> m x)) -∗ *) - (* p ⊑ (<! x> m x). *) - (* Proof. *) - (* rewrite iMsg_exist_eq. iIntros "H". *) - (* destruct (iProto_case p) as [Heq | [a [m' Heq]]]. *) - (* - unshelve iSpecialize ("H" $!inhabitant); first by apply _. *) - (* rewrite Heq. *) - (* iDestruct (iProto_le_end_inv_r with "H") as "H". *) - (* rewrite iProto_end_eq iProto_message_eq. *) - (* iDestruct (proto_message_end_equivI with "H") as "[]". *) - (* - iEval (rewrite Heq). destruct a. *) - (* + iApply iProto_le_send. iIntros (v p2'). iDestruct 1 as (x) "Hm". *) - (* iSpecialize ("H" $! x). rewrite Heq. *) - (* by iApply (iProto_le_send_send_inv with "H"). *) - (* + iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1". *) - (* iDestruct 1 as (x) "Hm2". *) - (* iSpecialize ("H" $! x). rewrite Heq. *) - (* iApply (iProto_le_recv_send_inv with "H Hm1 Hm2"). *) - (* Qed. *) - (* Lemma iProto_le_exist_intro_l {A} (m : A → iMsg Σ V) a : *) - (* ⊢ (<! x> m x) ⊑ (<!> m a). *) - (* Proof. *) - (* rewrite iMsg_exist_eq. iApply iProto_le_send. iIntros (v p') "Hm /=". *) - (* iExists p'. iSplitR; last by auto. iApply iProto_le_refl. *) - (* Qed. *) - (* Lemma iProto_le_exist_intro_r {A} (m : A → iMsg Σ V) a : *) - (* ⊢ (<?> m a) ⊑ (<? x> m x). *) - (* Proof. *) - (* rewrite iMsg_exist_eq. iApply iProto_le_recv. iIntros (v p') "Hm /=". *) - (* iExists p'. iSplitR; last by auto. iApply iProto_le_refl. *) - (* Qed. *) - - (* Lemma iProto_le_texist_elim_l {TT : tele} (m1 : TT → iMsg Σ V) a m2 : *) - (* (∀ x, (<?> m1 x) ⊑ (<a> m2)) -∗ *) - (* (<?.. x> m1 x) ⊑ (<a> m2). *) - (* Proof. *) - (* iIntros "H". iInduction TT as [|T TT] "IH"; simpl; [done|]. *) - (* iApply iProto_le_exist_elim_l; iIntros (x). *) - (* iApply "IH". iIntros (xs). iApply "H". *) - (* Qed. *) - (* Lemma iProto_le_texist_elim_r {TT : tele} a m1 (m2 : TT → iMsg Σ V) : *) - (* (∀ x, (<a> m1) ⊑ (<!> m2 x)) -∗ *) - (* (<a> m1) ⊑ (<!.. x> m2 x). *) - (* Proof. *) - (* iIntros "H". iInduction TT as [|T TT] "IH"; simpl; [done|]. *) - (* iApply iProto_le_exist_elim_r; iIntros (x). *) - (* iApply "IH". iIntros (xs). iApply "H". *) - (* Qed. *) - - (* Lemma iProto_le_texist_intro_l {TT : tele} (m : TT → iMsg Σ V) x : *) - (* ⊢ (<!.. x> m x) ⊑ (<!> m x). *) - (* Proof. *) - (* induction x as [|T TT x xs IH] using tele_arg_ind; simpl. *) - (* { iApply iProto_le_refl. } *) - (* iApply iProto_le_trans; [by iApply iProto_le_exist_intro_l|]. iApply IH. *) - (* Qed. *) - (* Lemma iProto_le_texist_intro_r {TT : tele} (m : TT → iMsg Σ V) x : *) - (* ⊢ (<?> m x) ⊑ (<?.. x> m x). *) - (* Proof. *) - (* induction x as [|T TT x xs IH] using tele_arg_ind; simpl. *) - (* { iApply iProto_le_refl. } *) - (* iApply iProto_le_trans; [|by iApply iProto_le_exist_intro_r]. iApply IH. *) - (* Qed. *) - - (* Lemma iProto_le_base a v P p1 p2 : *) - (* â–· (p1 ⊑ p2) -∗ *) - (* (<a> MSG v {{ P }}; p1) ⊑ (<a> MSG v {{ P }}; p2). *) - (* Proof. *) - (* rewrite iMsg_base_eq. iIntros "H". destruct a. *) - (* - iApply iProto_le_send. iIntros (v' p') "(->&Hp&$)". *) - (* iExists p1. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". *) - (* - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&$)". *) - (* iExists p2. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". *) - (* Qed. *) - - (* Lemma iProto_le_base_swap v1 v2 P1 P2 p : *) - (* ⊢ (<?> MSG v1 {{ P1 }}; <!> MSG v2 {{ P2 }}; p) *) - (* ⊑ (<!> MSG v2 {{ P2 }}; <?> MSG v1 {{ P1 }}; p). *) - (* Proof. *) - (* rewrite {1 3}iMsg_base_eq. iApply iProto_le_swap. *) - (* iIntros (v1' v2' p1' p2') "/= (->&#Hp1&HP1) (->&#Hp2&HP2)". iExists p. *) - (* iSplitL "HP2". *) - (* - iIntros "!>". iRewrite -"Hp1". by iApply iProto_le_payload_intro_l. *) - (* - iIntros "!>". iRewrite -"Hp2". by iApply iProto_le_payload_intro_r. *) - (* Qed. *) - - (* Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. *) - (* Proof. *) - (* iIntros "H". iLöb as "IH" forall (p1 p2). *) - (* destruct (iProto_case p1) as [->|([]&m1&->)]. *) - (* - iDestruct (iProto_le_end_inv_l with "H") as "H". *) - (* iRewrite "H". iApply iProto_le_refl. *) - (* - iDestruct (iProto_le_send_inv with "H") as (a2 m2) "[Hp2 H]". *) - (* iRewrite "Hp2"; clear p2. iEval (rewrite !iProto_dual_message). *) - (* destruct a2; simpl. *) - (* + iApply iProto_le_recv. iIntros (v p1d). *) - (* iDestruct 1 as (p1') "[Hm1 #Hp1d]". *) - (* iDestruct ("H" with "Hm1") as (p2') "[H Hm2]". *) - (* iDestruct ("IH" with "H") as "H". iExists (iProto_dual p2'). *) - (* iSplitL "H"; [iIntros "!>"; by iRewrite "Hp1d"|]. simpl; auto. *) - (* + iApply iProto_le_swap. iIntros (v1 v2 p1d p2d). *) - (* iDestruct 1 as (p1') "[Hm1 #Hp1d]". iDestruct 1 as (p2') "[Hm2 #Hp2d]". *) - (* iDestruct ("H" with "Hm2 Hm1") as (pt) "[H1 H2]". *) - (* iDestruct ("IH" with "H1") as "H1". iDestruct ("IH" with "H2") as "H2 {IH}". *) - (* rewrite !iProto_dual_message /=. iExists (iProto_dual pt). iSplitL "H2". *) - (* * iIntros "!>". iRewrite "Hp1d". by rewrite -iMsg_dual_base. *) - (* * iIntros "!>". iRewrite "Hp2d". by rewrite -iMsg_dual_base. *) - (* - iDestruct (iProto_le_recv_inv with "H") as (m2) "[Hp2 H]". *) - (* iRewrite "Hp2"; clear p2. iEval (rewrite !iProto_dual_message /=). *) - (* iApply iProto_le_send. iIntros (v p2d). *) - (* iDestruct 1 as (p2') "[Hm2 #Hp2d]". *) - (* iDestruct ("H" with "Hm2") as (p1') "[H Hm1]". *) - (* iDestruct ("IH" with "H") as "H". iExists (iProto_dual p1'). *) - (* iSplitL "H"; [iIntros "!>"; by iRewrite "Hp2d"|]. simpl; auto. *) - (* Qed. *) - - (* Lemma iProto_le_amber_internal (p1 p2 : iProto Σ V → iProto Σ V) *) - (* `{Contractive p1, Contractive p2}: *) - (* â–¡ (∀ rec1 rec2, â–· (rec1 ⊑ rec2) → p1 rec1 ⊑ p2 rec2) -∗ *) - (* fixpoint p1 ⊑ fixpoint p2. *) - (* Proof. *) - (* iIntros "#H". iLöb as "IH". *) - (* iEval (rewrite (fixpoint_unfold p1)). *) - (* iEval (rewrite (fixpoint_unfold p2)). *) - (* iApply "H". iApply "IH". *) - (* Qed. *) - (* Lemma iProto_le_amber_external (p1 p2 : iProto Σ V → iProto Σ V) *) - (* `{Contractive p1, Contractive p2}: *) - (* (∀ rec1 rec2, (⊢ rec1 ⊑ rec2) → ⊢ p1 rec1 ⊑ p2 rec2) → *) - (* ⊢ fixpoint p1 ⊑ fixpoint p2. *) - (* Proof. *) - (* intros IH. apply fixpoint_ind. *) - (* - by intros p1' p2' -> ?. *) - (* - exists (fixpoint p2). iApply iProto_le_refl. *) - (* - intros p' ?. rewrite (fixpoint_unfold p2). by apply IH. *) - (* - apply bi.limit_preserving_entails; [done|solve_proper]. *) - (* Qed. *) - - (* Lemma iProto_le_dual_l p1 p2 : iProto_dual p2 ⊑ p1 -∗ iProto_dual p1 ⊑ p2. *) - (* Proof. *) - (* iIntros "H". iEval (rewrite -(involutive iProto_dual p2)). *) - (* by iApply iProto_le_dual. *) - (* Qed. *) - (* Lemma iProto_le_dual_r p1 p2 : p2 ⊑ iProto_dual p1 -∗ p1 ⊑ iProto_dual p2. *) - (* Proof. *) - (* iIntros "H". iEval (rewrite -(involutive iProto_dual p1)). *) - (* by iApply iProto_le_dual. *) - (* Qed. *) - - (* Lemma iProto_le_app p1 p2 p3 p4 : *) - (* p1 ⊑ p2 -∗ p3 ⊑ p4 -∗ p1 <++> p3 ⊑ p2 <++> p4. *) - (* Proof. *) - (* iIntros "H1 H2". iLöb as "IH" forall (p1 p2 p3 p4). *) - (* destruct (iProto_case p2) as [->|([]&m2&->)]. *) - (* - iDestruct (iProto_le_end_inv_l with "H1") as "H1". *) - (* iRewrite "H1". by rewrite !left_id. *) - (* - iDestruct (iProto_le_send_inv with "H1") as (a1 m1) "[Hp1 H1]". *) - (* iRewrite "Hp1"; clear p1. rewrite !iProto_app_message. destruct a1; simpl. *) - (* + iApply iProto_le_send. iIntros (v p24). *) - (* iDestruct 1 as (p2') "[Hm2 #Hp24]". *) - (* iDestruct ("H1" with "Hm2") as (p1') "[H1 Hm1]". *) - (* iExists (p1' <++> p3). iSplitR "Hm1"; [|by simpl; eauto]. *) - (* iIntros "!>". iRewrite "Hp24". by iApply ("IH" with "H1"). *) - (* + iApply iProto_le_swap. iIntros (v1 v2 p13 p24). *) - (* iDestruct 1 as (p1') "[Hm1 #Hp13]". iDestruct 1 as (p2') "[Hm2 #Hp24]". *) - (* iSpecialize ("H1" with "Hm1 Hm2"). *) - (* iDestruct "H1" as (pt) "[H1 H1']". *) - (* iExists (pt <++> p3). iSplitL "H1". *) - (* * iIntros "!>". iRewrite "Hp13". *) - (* rewrite /= -iMsg_app_base -iProto_app_message. *) - (* iApply ("IH" with "H1"). iApply iProto_le_refl. *) - (* * iIntros "!>". iRewrite "Hp24". *) - (* rewrite /= -iMsg_app_base -iProto_app_message. *) - (* iApply ("IH" with "H1' H2"). *) - (* - iDestruct (iProto_le_recv_inv with "H1") as (m1) "[Hp1 H1]". *) - (* iRewrite "Hp1"; clear p1. rewrite !iProto_app_message. iApply iProto_le_recv. *) - (* iIntros (v p13). iDestruct 1 as (p1') "[Hm1 #Hp13]". *) - (* iDestruct ("H1" with "Hm1") as (p2'') "[H1 Hm2]". *) - (* iExists (p2'' <++> p4). iSplitR "Hm2"; [|by simpl; eauto]. *) - (* iIntros "!>". iRewrite "Hp13". by iApply ("IH" with "H1"). *) - (* Qed. *) - - (** ** Lemmas about the auxiliary definitions and invariants *) - Global Instance iProto_app_recvs_ne vs : - NonExpansive (iProto_app_recvs (Σ:=Σ) (V:=V) vs). - Proof. induction vs; solve_proper. Qed. - Global Instance iProto_app_recvs_proper vs : - Proper ((≡) ==> (≡)) (iProto_app_recvs (Σ:=Σ) (V:=V) vs). - Proof. induction vs; solve_proper. Qed. - Global Instance iProto_own_frag_ne γ s : NonExpansive (iProto_own_frag γ s). Proof. solve_proper. Qed. @@ -1549,101 +782,101 @@ Section proto. by rewrite own_op. Qed. - Global Instance iProto_own_ne γ s : NonExpansive (iProto_own γ s). - Proof. solve_proper. Qed. - Global Instance iProto_own_proper γ s : Proper ((≡) ==> (≡)) (iProto_own γ s). - Proof. apply (ne_proper _). Qed. + (* Global Instance iProto_own_ne γ s : NonExpansive (iProto_own γ s). *) + (* Proof. solve_proper. Qed. *) + (* Global Instance iProto_own_proper γ s : Proper ((≡) ==> (≡)) (iProto_own γ s). *) + (* Proof. apply (ne_proper _). Qed. *) - Lemma iProto_own_le γ s p1 p2 : - iProto_own γ s p1 -∗ â–· (p1 ⊑ p2) -∗ iProto_own γ s p2. - Proof. - iDestruct 1 as (p1') "[Hle H]". iIntros "Hle'". - iExists p1'. iFrame "H". by iApply (iProto_le_trans with "Hle"). - Qed. + (* Lemma iProto_own_le γ s p1 p2 : *) + (* iProto_own γ s p1 -∗ â–· (p1 ⊑ p2) -∗ iProto_own γ s p2. *) + (* Proof. *) + (* iDestruct 1 as (p1') "[Hle H]". iIntros "Hle'". *) + (* iExists p1'. iFrame "H". by iApply (iProto_le_trans with "Hle"). *) + (* Qed. *) - Lemma iProto_init p : - ⊢ |==> ∃ γ, - iProto_ctx γ ∗ iProto_own γ Left p ∗ iProto_own γ Right (iProto_dual p). - Proof. - iMod (own_alloc (â—E (Next p) â‹… â—¯E (Next p))) as (lγ) "[Hâ—l Hâ—¯l]". - { by apply excl_auth_valid. } - iMod (own_alloc (â—E (Next (iProto_dual p)) â‹… - â—¯E (Next (iProto_dual p)))) as (rγ) "[Hâ—r Hâ—¯r]". - { by apply excl_auth_valid. } - pose (ProtName lγ rγ) as γ. iModIntro. iExists γ. iSplitL "Hâ—l Hâ—r". - { iExists p, (iProto_dual p). iFrame. iApply iProto_consistent_dual. } - iSplitL "Hâ—¯l"; iExists _; iFrame; iApply iProto_le_refl. - Qed. + (* Lemma iProto_init p : *) + (* ⊢ |==> ∃ γ, *) + (* iProto_ctx γ ∗ iProto_own γ Left p ∗ iProto_own γ Right (iProto_dual p). *) + (* Proof. *) + (* iMod (own_alloc (â—E (Next p) â‹… â—¯E (Next p))) as (lγ) "[Hâ—l Hâ—¯l]". *) + (* { by apply excl_auth_valid. } *) + (* iMod (own_alloc (â—E (Next (iProto_dual p)) â‹… *) + (* â—¯E (Next (iProto_dual p)))) as (rγ) "[Hâ—r Hâ—¯r]". *) + (* { by apply excl_auth_valid. } *) + (* pose (ProtName lγ rγ) as γ. iModIntro. iExists γ. iSplitL "Hâ—l Hâ—r". *) + (* { iExists p, (iProto_dual p). iFrame. iApply iProto_consistent_dual. } *) + (* iSplitL "Hâ—¯l"; iExists _; iFrame; iApply iProto_le_refl. *) + (* Qed. *) - Definition side_dual s := - match s with - | Left => Right - | Right => Left - end. + (* Definition side_dual s := *) + (* match s with *) + (* | Left => Right *) + (* | Right => Left *) + (* end. *) - Lemma iProto_step_l γ m1 m2 p1 v : - iProto_ctx γ -∗ iProto_own γ Left (<!> m1) -∗ iProto_own γ Right (<?> m2) -∗ - iMsg_car m1 v (Next p1) ==∗ - â–· ∃ p2, iMsg_car m2 v (Next p2) ∗ â–· iProto_ctx γ ∗ - iProto_own γ Left p1 ∗ iProto_own γ Right p2. - Proof. - iDestruct 1 as (pl pr) "(Hâ—l & Hâ—r & Hconsistent)". - iDestruct 1 as (pl') "[Hlel Hâ—¯l]". - iDestruct 1 as (pr') "[Hler Hâ—¯r]". - iIntros "Hm". - iDestruct (iProto_own_auth_agree with "Hâ—l Hâ—¯l") as "#Hpl". - iDestruct (iProto_own_auth_agree with "Hâ—r Hâ—¯r") as "#Hpr". - iAssert (â–· (pl ⊑ <!> m1))%I - with "[Hlel]" as "{Hpl} Hlel"; first (iNext; by iRewrite "Hpl"). - iAssert (â–· (pr ⊑ <?> m2))%I - with "[Hler]" as "{Hpr} Hler"; first (iNext; by iRewrite "Hpr"). - iDestruct (iProto_consistent_le_l with "Hconsistent Hlel") as "Hconsistent". - iDestruct (iProto_consistent_le_r with "Hconsistent Hler") as "Hconsistent". - iDestruct (iProto_consistent_step with "Hconsistent [Hm //]") as - (p2) "[Hm2 Hconsistent]". - iMod (iProto_own_auth_update _ _ _ _ p1 with "Hâ—l Hâ—¯l") as "[Hâ—l Hâ—¯l]". - iMod (iProto_own_auth_update _ _ _ _ p2 with "Hâ—r Hâ—¯r") as "[Hâ—r Hâ—¯r]". - iIntros "!>!>". - iExists p2. iFrame. - iSplitL "Hconsistent Hâ—l Hâ—r". - - iExists _, _. iFrame. - - iSplitL "Hâ—¯l". - + iExists _. iFrame. iApply iProto_le_refl. - + iExists _. iFrame. iApply iProto_le_refl. - Qed. - - Lemma iProto_step_r γ m1 m2 p2 v : - iProto_ctx γ -∗ iProto_own γ Left (<?> m1) -∗ iProto_own γ Right (<!> m2) -∗ - iMsg_car m2 v (Next p2) ==∗ - â–· ∃ p1, iMsg_car m1 v (Next p1) ∗ â–· iProto_ctx γ ∗ - iProto_own γ Left p1 ∗ iProto_own γ Right p2. - Proof. - iDestruct 1 as (pl pr) "(Hâ—l & Hâ—r & Hconsistent)". - iDestruct 1 as (pl') "[Hlel Hâ—¯l]". - iDestruct 1 as (pr') "[Hler Hâ—¯r]". - iIntros "Hm". - iDestruct (iProto_own_auth_agree with "Hâ—l Hâ—¯l") as "#Hpl". - iDestruct (iProto_own_auth_agree with "Hâ—r Hâ—¯r") as "#Hpr". - iAssert (â–· (pl ⊑ <?> m1))%I - with "[Hlel]" as "{Hpl} Hlel"; first (iNext; by iRewrite "Hpl"). - iAssert (â–· (pr ⊑ <!> m2))%I - with "[Hler]" as "{Hpr} Hler"; first (iNext; by iRewrite "Hpr"). - iDestruct (iProto_consistent_le_l with "Hconsistent Hlel") as "Hconsistent". - iDestruct (iProto_consistent_le_r with "Hconsistent Hler") as "Hconsistent". - iDestruct (iProto_consistent_flip with "Hconsistent") as - "Hconsistent". - iDestruct (iProto_consistent_step with "Hconsistent [Hm //]") as - (p1) "[Hm1 Hconsistent]". - iMod (iProto_own_auth_update _ _ _ _ p1 with "Hâ—l Hâ—¯l") as "[Hâ—l Hâ—¯l]". - iMod (iProto_own_auth_update _ _ _ _ p2 with "Hâ—r Hâ—¯r") as "[Hâ—r Hâ—¯r]". - iIntros "!>!>". - iExists p1. iFrame. - iSplitL "Hconsistent Hâ—l Hâ—r". - - iExists _, _. iFrame. iApply iProto_consistent_flip. iFrame. - - iSplitL "Hâ—¯l". - + iExists _. iFrame. iApply iProto_le_refl. - + iExists _. iFrame. iApply iProto_le_refl. - Qed. + (* Lemma iProto_step_l γ m1 m2 p1 v : *) + (* iProto_ctx γ -∗ iProto_own γ Left (<!> m1) -∗ iProto_own γ Right (<?> m2) -∗ *) + (* iMsg_car m1 v (Next p1) ==∗ *) + (* â–· ∃ p2, iMsg_car m2 v (Next p2) ∗ â–· iProto_ctx γ ∗ *) + (* iProto_own γ Left p1 ∗ iProto_own γ Right p2. *) + (* Proof. *) + (* iDestruct 1 as (pl pr) "(Hâ—l & Hâ—r & Hconsistent)". *) + (* iDestruct 1 as (pl') "[Hlel Hâ—¯l]". *) + (* iDestruct 1 as (pr') "[Hler Hâ—¯r]". *) + (* iIntros "Hm". *) + (* iDestruct (iProto_own_auth_agree with "Hâ—l Hâ—¯l") as "#Hpl". *) + (* iDestruct (iProto_own_auth_agree with "Hâ—r Hâ—¯r") as "#Hpr". *) + (* iAssert (â–· (pl ⊑ <!> m1))%I *) + (* with "[Hlel]" as "{Hpl} Hlel"; first (iNext; by iRewrite "Hpl"). *) + (* iAssert (â–· (pr ⊑ <?> m2))%I *) + (* with "[Hler]" as "{Hpr} Hler"; first (iNext; by iRewrite "Hpr"). *) + (* iDestruct (iProto_consistent_le_l with "Hconsistent Hlel") as "Hconsistent". *) + (* iDestruct (iProto_consistent_le_r with "Hconsistent Hler") as "Hconsistent". *) + (* iDestruct (iProto_consistent_step with "Hconsistent [Hm //]") as *) + (* (p2) "[Hm2 Hconsistent]". *) + (* iMod (iProto_own_auth_update _ _ _ _ p1 with "Hâ—l Hâ—¯l") as "[Hâ—l Hâ—¯l]". *) + (* iMod (iProto_own_auth_update _ _ _ _ p2 with "Hâ—r Hâ—¯r") as "[Hâ—r Hâ—¯r]". *) + (* iIntros "!>!>". *) + (* iExists p2. iFrame. *) + (* iSplitL "Hconsistent Hâ—l Hâ—r". *) + (* - iExists _, _. iFrame. *) + (* - iSplitL "Hâ—¯l". *) + (* + iExists _. iFrame. iApply iProto_le_refl. *) + (* + iExists _. iFrame. iApply iProto_le_refl. *) + (* Qed. *) + + (* Lemma iProto_step_r γ m1 m2 p2 v : *) + (* iProto_ctx γ -∗ iProto_own γ Left (<?> m1) -∗ iProto_own γ Right (<!> m2) -∗ *) + (* iMsg_car m2 v (Next p2) ==∗ *) + (* â–· ∃ p1, iMsg_car m1 v (Next p1) ∗ â–· iProto_ctx γ ∗ *) + (* iProto_own γ Left p1 ∗ iProto_own γ Right p2. *) + (* Proof. *) + (* iDestruct 1 as (pl pr) "(Hâ—l & Hâ—r & Hconsistent)". *) + (* iDestruct 1 as (pl') "[Hlel Hâ—¯l]". *) + (* iDestruct 1 as (pr') "[Hler Hâ—¯r]". *) + (* iIntros "Hm". *) + (* iDestruct (iProto_own_auth_agree with "Hâ—l Hâ—¯l") as "#Hpl". *) + (* iDestruct (iProto_own_auth_agree with "Hâ—r Hâ—¯r") as "#Hpr". *) + (* iAssert (â–· (pl ⊑ <?> m1))%I *) + (* with "[Hlel]" as "{Hpl} Hlel"; first (iNext; by iRewrite "Hpl"). *) + (* iAssert (â–· (pr ⊑ <!> m2))%I *) + (* with "[Hler]" as "{Hpr} Hler"; first (iNext; by iRewrite "Hpr"). *) + (* iDestruct (iProto_consistent_le_l with "Hconsistent Hlel") as "Hconsistent". *) + (* iDestruct (iProto_consistent_le_r with "Hconsistent Hler") as "Hconsistent". *) + (* iDestruct (iProto_consistent_flip with "Hconsistent") as *) + (* "Hconsistent". *) + (* iDestruct (iProto_consistent_step with "Hconsistent [Hm //]") as *) + (* (p1) "[Hm1 Hconsistent]". *) + (* iMod (iProto_own_auth_update _ _ _ _ p1 with "Hâ—l Hâ—¯l") as "[Hâ—l Hâ—¯l]". *) + (* iMod (iProto_own_auth_update _ _ _ _ p2 with "Hâ—r Hâ—¯r") as "[Hâ—r Hâ—¯r]". *) + (* iIntros "!>!>". *) + (* iExists p1. iFrame. *) + (* iSplitL "Hconsistent Hâ—l Hâ—r". *) + (* - iExists _, _. iFrame. iApply iProto_consistent_flip. iFrame. *) + (* - iSplitL "Hâ—¯l". *) + (* + iExists _. iFrame. iApply iProto_le_refl. *) + (* + iExists _. iFrame. iApply iProto_le_refl. *) + (* Qed. *) (* (** The instances below make it possible to use the tactics [iIntros], *) (* [iExist], [iSplitL]/[iSplitR], [iFrame] and [iModIntro] on [iProto_le] goals. *) *) @@ -1718,7 +951,7 @@ Section proto. End proto. -Typeclasses Opaque iProto_ctx iProto_own. +(* Typeclasses Opaque iProto_ctx iProto_own. *) -Global Hint Extern 0 (environments.envs_entails _ (?x ⊑ ?y)) => - first [is_evar x; fail 1 | is_evar y; fail 1|iApply iProto_le_refl] : core. +(* Global Hint Extern 0 (environments.envs_entails _ (?x ⊑ ?y)) => *) +(* first [is_evar x; fail 1 | is_evar y; fail 1|iApply iProto_le_refl] : core. *) -- GitLab From 4d5fbe48c807923e466e991c9571532b46f55d49 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Sun, 21 Jan 2024 16:50:06 +0100 Subject: [PATCH 04/81] Wrapped up API for multiparty ghost theory --- theories/channel/multi_proto.v | 208 +++++++++++++++------------------ 1 file changed, 91 insertions(+), 117 deletions(-) diff --git a/theories/channel/multi_proto.v b/theories/channel/multi_proto.v index 878792e..d13ba79 100644 --- a/theories/channel/multi_proto.v +++ b/theories/channel/multi_proto.v @@ -57,14 +57,39 @@ Export action. (** * Setup of Iris's cameras *) Class protoG Σ V := protoG_authG :: - inG Σ (excl_authR (laterO (proto (leibnizO V) (iPropO Σ) (iPropO Σ)))). + inG Σ (authUR (gmapR natO + (optionUR (exclR (laterO (proto (leibnizO V) (iPropO Σ) (iPropO Σ))))))). Definition protoΣ V := #[ - GFunctor (authRF (optionURF (exclRF (laterOF (protoOF (leibnizO V) idOF idOF))))) + GFunctor (authRF (gmapURF natO (optionRF (exclRF (laterOF (protoOF (leibnizO V) idOF idOF)))))) ]. Global Instance subG_chanΣ {Σ V} : subG (protoΣ V) Σ → protoG Σ V. Proof. solve_inG. Qed. +(* (** * Setup of Iris's cameras *) *) +(* Class protoG Σ V := *) +(* protoG_authG :: *) +(* inG Σ (gmapR natO *) +(* (excl_authR (laterO (proto (leibnizO V) (iPropO Σ) (iPropO Σ))))). *) + +(* Definition protoΣ V := #[ *) +(* GFunctor (gmapRF natO (authRF (optionURF (exclRF (laterOF (protoOF (leibnizO V) idOF idOF)))))) *) +(* ]. *) +(* Global Instance subG_chanΣ {Σ V} : subG (protoΣ V) Σ → protoG Σ V. *) +(* Proof. solve_inG. Qed. *) + + +(* (** * Setup of Iris's cameras *) *) +(* Class protoG Σ V := *) +(* protoG_authG :: *) +(* inG Σ (excl_authR (laterO (proto (leibnizO V) (iPropO Σ) (iPropO Σ)))). *) + +(* Definition protoΣ V := #[ *) +(* GFunctor (authRF (optionURF (exclRF (laterOF (protoOF (leibnizO V) idOF idOF))))) *) +(* ]. *) +(* Global Instance subG_chanΣ {Σ V} : subG (protoΣ V) Σ → protoG Σ V. *) +(* Proof. solve_inG. Qed. *) + (** * Types *) Definition iProto Σ V := proto V (iPropO Σ) (iPropO Σ). Declare Scope proto_scope. @@ -687,53 +712,45 @@ Qed. (* iProto_consistent (ps))%I -∗ *) (* iProto_consistent ps. *) -Record proto_name := ProtName { proto_l_name : gname; proto_r_name : gname }. +Record proto_name := ProtName { proto_names : gmap nat gname }. Global Instance proto_name_inhabited : Inhabited proto_name := - populate (ProtName inhabitant inhabitant). + populate (ProtName inhabitant). Global Instance proto_name_eq_dec : EqDecision proto_name. Proof. solve_decision. Qed. Global Instance proto_name_countable : Countable proto_name. Proof. - refine (inj_countable (λ '(ProtName γl γr), (γl,γr)) - (λ '(γl, γr), Some (ProtName γl γr)) _); by intros []. + refine (inj_countable (λ '(ProtName γs), (γs)) + (λ '(γs), Some (ProtName γs)) _); by intros []. Qed. -Inductive side := Left | Right. -Global Instance side_inhabited : Inhabited side := populate Left. -Global Instance side_eq_dec : EqDecision side. -Proof. solve_decision. Qed. -Global Instance side_countable : Countable side. -Proof. - refine (inj_countable (λ s, if s is Left then true else false) - (λ b, Some (if b then Left else Right)) _); by intros []. -Qed. -Definition side_elim {A} (s : side) (l r : A) : A := - match s with Left => l | Right => r end. - -Definition iProto_own_frag `{!protoG Σ V} (γ : proto_name) (s : side) - (p : iProto Σ V) : iProp Σ := - own (side_elim s proto_l_name proto_r_name γ) (â—¯E (Next p)). -Definition iProto_own_auth `{!protoG Σ V} (γ : proto_name) (s : side) - (p : iProto Σ V) : iProp Σ := - own (side_elim s proto_l_name proto_r_name γ) (â—E (Next p)). - -(* Definition iProto_ctx `{protoG Σ V} *) -(* (γ : proto_name) : iProp Σ := *) -(* ∃ ps, *) -(* iProto_own_auth γ Left pl ∗ *) -(* iProto_own_auth γ Right pr ∗ *) -(* â–· iProto_consistent ps. *) - -(* (** * The connective for ownership of channel ends *) *) -(* Definition iProto_own `{!protoG Σ V} *) -(* (γ : proto_name) (s : side) (p : iProto Σ V) : iProp Σ := *) -(* ∃ p', â–· (p' ⊑ p) ∗ iProto_own_frag γ s p'. *) -(* Arguments iProto_own {_ _ _} _ _%proto. *) -(* Global Instance: Params (@iProto_own) 3 := {}. *) - -(* Global Instance iProto_own_contractive `{protoG Σ V} γ s : *) -(* Contractive (iProto_own γ s). *) -(* Proof. solve_contractive. Qed. *) +Definition iProto_own_frag `{!protoG Σ V} (γ : gname) + (i : nat) (p : iProto Σ V) : iProp Σ := + own γ (â—¯ ({[i := Excl' (Next p)]})). + +(* TODO: Fix this def. *) +Definition iProto_own_auth `{!protoG Σ V} (γ : gname) + (ps : gmap nat (iProto Σ V)) : iProp Σ := + own γ (◠∅). + (* own γ (â— ((λ (p : iProto Σ V), Excl' (Next p)) <$> ps)). *) + +Definition iProto_ctx `{protoG Σ V} + (γ : gname) : iProp Σ := + ∃ ps, iProto_own_auth γ ps ∗ â–· iProto_consistent ps. + +(** * The connective for ownership of channel ends *) +Definition iProto_own `{!protoG Σ V} + (γ : gname) (i : nat) (p : iProto Σ V) : iProp Σ := + iProto_own_frag γ i p. +Arguments iProto_own {_ _ _} _ _ _%proto. +Global Instance: Params (@iProto_own) 3 := {}. + +Global Instance iProto_own_frag_contractive `{protoG Σ V} γ i : + Contractive (iProto_own_frag γ i). +Proof. solve_contractive. Qed. + +Global Instance iProto_own_contractive `{protoG Σ V} γ i : + Contractive (iProto_own γ i). +Proof. solve_contractive. Qed. (** * Proofs *) Section proto. @@ -765,39 +782,33 @@ Section proto. Global Instance iProto_own_frag_ne γ s : NonExpansive (iProto_own_frag γ s). Proof. solve_proper. Qed. - Lemma iProto_own_auth_agree γ s p p' : - iProto_own_auth γ s p -∗ iProto_own_frag γ s p' -∗ â–· (p ≡ p'). - Proof. - iIntros "Hâ— Hâ—¯". iDestruct (own_valid_2 with "Hâ— Hâ—¯") as "H✓". - iDestruct (excl_auth_agreeI with "H✓") as "H✓". - iApply (later_equivI_1 with "H✓"). - Qed. - - Lemma iProto_own_auth_update γ s p p' p'' : - iProto_own_auth γ s p -∗ iProto_own_frag γ s p' ==∗ - iProto_own_auth γ s p'' ∗ iProto_own_frag γ s p''. - Proof. - iIntros "Hâ— Hâ—¯". iDestruct (own_update_2 with "Hâ— Hâ—¯") as "H". - { eapply (excl_auth_update _ _ (Next p'')). } - by rewrite own_op. - Qed. - - (* Global Instance iProto_own_ne γ s : NonExpansive (iProto_own γ s). *) - (* Proof. solve_proper. Qed. *) - (* Global Instance iProto_own_proper γ s : Proper ((≡) ==> (≡)) (iProto_own γ s). *) - (* Proof. apply (ne_proper _). Qed. *) + (* TODO: Relies on above def *) + Lemma iProto_own_auth_agree γ ps i p : + iProto_own_auth γ ps -∗ iProto_own_frag γ i p -∗ â–· (ps !!! i ≡ p). + Proof. Admitted. + (* iIntros "Hâ— Hâ—¯". iDestruct (own_valid_2 with "Hâ— Hâ—¯") as "H✓". *) + (* iDestruct (excl_auth_agreeI with "H✓") as "H✓". *) + (* iApply (later_equivI_1 with "H✓"). *) + (* Qed. *) - (* Lemma iProto_own_le γ s p1 p2 : *) - (* iProto_own γ s p1 -∗ â–· (p1 ⊑ p2) -∗ iProto_own γ s p2. *) - (* Proof. *) - (* iDestruct 1 as (p1') "[Hle H]". iIntros "Hle'". *) - (* iExists p1'. iFrame "H". by iApply (iProto_le_trans with "Hle"). *) + Lemma iProto_own_auth_update γ ps i p p' : + iProto_own_auth γ ps -∗ iProto_own_frag γ i p ==∗ + iProto_own_auth γ (<[i := p']>ps) ∗ iProto_own_frag γ i p'. + Proof. Admitted. + (* iIntros "Hâ— Hâ—¯". iDestruct (own_update_2 with "Hâ— Hâ—¯") as "H". *) + (* { eapply (excl_auth_update _ _ (Next p'')). } *) + (* by rewrite own_op. *) (* Qed. *) - (* Lemma iProto_init p : *) - (* ⊢ |==> ∃ γ, *) - (* iProto_ctx γ ∗ iProto_own γ Left p ∗ iProto_own γ Right (iProto_dual p). *) - (* Proof. *) + Global Instance iProto_own_ne γ s : NonExpansive (iProto_own γ s). + Proof. solve_proper. Qed. + Global Instance iProto_own_proper γ s : Proper ((≡) ==> (≡)) (iProto_own γ s). + Proof. apply (ne_proper _). Qed. + + Lemma iProto_init ps : + iProto_consistent ps -∗ + |==> ∃ γ, iProto_ctx γ ∗ [∗ map] i ↦p ∈ ps, iProto_own γ i p. + Proof. Admitted. (* iMod (own_alloc (â—E (Next p) â‹… â—¯E (Next p))) as (lγ) "[Hâ—l Hâ—¯l]". *) (* { by apply excl_auth_valid. } *) (* iMod (own_alloc (â—E (Next (iProto_dual p)) â‹… *) @@ -808,18 +819,14 @@ Section proto. (* iSplitL "Hâ—¯l"; iExists _; iFrame; iApply iProto_le_refl. *) (* Qed. *) - (* Definition side_dual s := *) - (* match s with *) - (* | Left => Right *) - (* | Right => Left *) - (* end. *) - - (* Lemma iProto_step_l γ m1 m2 p1 v : *) - (* iProto_ctx γ -∗ iProto_own γ Left (<!> m1) -∗ iProto_own γ Right (<?> m2) -∗ *) - (* iMsg_car m1 v (Next p1) ==∗ *) - (* â–· ∃ p2, iMsg_car m2 v (Next p2) ∗ â–· iProto_ctx γ ∗ *) - (* iProto_own γ Left p1 ∗ iProto_own γ Right p2. *) - (* Proof. *) + Lemma iProto_step γ i j m1 m2 p1 v : + iProto_ctx γ -∗ + iProto_own γ i (<(Send j)> m1) -∗ + iProto_own γ j (<(Recv i)> m2) -∗ + iMsg_car m1 v (Next p1) ==∗ + â–· ∃ p2, iMsg_car m2 v (Next p2) ∗ â–· iProto_ctx γ ∗ + iProto_own γ i p1 ∗ iProto_own γ j p2. + Proof. Admitted. (* iDestruct 1 as (pl pr) "(Hâ—l & Hâ—r & Hconsistent)". *) (* iDestruct 1 as (pl') "[Hlel Hâ—¯l]". *) (* iDestruct 1 as (pr') "[Hler Hâ—¯r]". *) @@ -845,39 +852,6 @@ Section proto. (* + iExists _. iFrame. iApply iProto_le_refl. *) (* Qed. *) - (* Lemma iProto_step_r γ m1 m2 p2 v : *) - (* iProto_ctx γ -∗ iProto_own γ Left (<?> m1) -∗ iProto_own γ Right (<!> m2) -∗ *) - (* iMsg_car m2 v (Next p2) ==∗ *) - (* â–· ∃ p1, iMsg_car m1 v (Next p1) ∗ â–· iProto_ctx γ ∗ *) - (* iProto_own γ Left p1 ∗ iProto_own γ Right p2. *) - (* Proof. *) - (* iDestruct 1 as (pl pr) "(Hâ—l & Hâ—r & Hconsistent)". *) - (* iDestruct 1 as (pl') "[Hlel Hâ—¯l]". *) - (* iDestruct 1 as (pr') "[Hler Hâ—¯r]". *) - (* iIntros "Hm". *) - (* iDestruct (iProto_own_auth_agree with "Hâ—l Hâ—¯l") as "#Hpl". *) - (* iDestruct (iProto_own_auth_agree with "Hâ—r Hâ—¯r") as "#Hpr". *) - (* iAssert (â–· (pl ⊑ <?> m1))%I *) - (* with "[Hlel]" as "{Hpl} Hlel"; first (iNext; by iRewrite "Hpl"). *) - (* iAssert (â–· (pr ⊑ <!> m2))%I *) - (* with "[Hler]" as "{Hpr} Hler"; first (iNext; by iRewrite "Hpr"). *) - (* iDestruct (iProto_consistent_le_l with "Hconsistent Hlel") as "Hconsistent". *) - (* iDestruct (iProto_consistent_le_r with "Hconsistent Hler") as "Hconsistent". *) - (* iDestruct (iProto_consistent_flip with "Hconsistent") as *) - (* "Hconsistent". *) - (* iDestruct (iProto_consistent_step with "Hconsistent [Hm //]") as *) - (* (p1) "[Hm1 Hconsistent]". *) - (* iMod (iProto_own_auth_update _ _ _ _ p1 with "Hâ—l Hâ—¯l") as "[Hâ—l Hâ—¯l]". *) - (* iMod (iProto_own_auth_update _ _ _ _ p2 with "Hâ—r Hâ—¯r") as "[Hâ—r Hâ—¯r]". *) - (* iIntros "!>!>". *) - (* iExists p1. iFrame. *) - (* iSplitL "Hconsistent Hâ—l Hâ—r". *) - (* - iExists _, _. iFrame. iApply iProto_consistent_flip. iFrame. *) - (* - iSplitL "Hâ—¯l". *) - (* + iExists _. iFrame. iApply iProto_le_refl. *) - (* + iExists _. iFrame. iApply iProto_le_refl. *) - (* Qed. *) - (* (** The instances below make it possible to use the tactics [iIntros], *) (* [iExist], [iSplitL]/[iSplitR], [iFrame] and [iModIntro] on [iProto_le] goals. *) *) (* Global Instance iProto_le_from_forall_l {A} a (m1 : A → iMsg Σ V) m2 name : *) -- GitLab From 6d189579dea6ce9b01f8b729a9d102dd7038421d Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Sun, 21 Jan 2024 16:57:50 +0100 Subject: [PATCH 05/81] Proved ghost theory stepping relation --- theories/channel/multi_proto.v | 42 ++++++++++++---------------------- 1 file changed, 14 insertions(+), 28 deletions(-) diff --git a/theories/channel/multi_proto.v b/theories/channel/multi_proto.v index d13ba79..6f35f29 100644 --- a/theories/channel/multi_proto.v +++ b/theories/channel/multi_proto.v @@ -760,14 +760,14 @@ Section proto. Implicit Types m : iMsg Σ V. Lemma iProto_consistent_step ps m1 m2 i j v p1 : + iProto_consistent ps -∗ ps !!! i ≡ (<(Send j)> m1) -∗ ps !!! j ≡ (<(Recv i)> m2) -∗ - iProto_consistent ps -∗ iMsg_car m1 v (Next p1) -∗ ∃ p2, iMsg_car m2 v (Next p2) ∗ â–· iProto_consistent (<[i := p1]>(<[j := p2]>ps)). Proof. - iIntros "#Hi #Hj Hprot Hm1". + iIntros "Hprot #Hi #Hj Hm1". rewrite iProto_consistent_unfold /iProto_consistent_pre. iDestruct ("Hprot" $!i j with "[] [] Hi Hj Hm1") as (p2) "[Hm2 Hprot]". { rewrite !lookup_total_alt elem_of_dom. @@ -824,33 +824,19 @@ Section proto. iProto_own γ i (<(Send j)> m1) -∗ iProto_own γ j (<(Recv i)> m2) -∗ iMsg_car m1 v (Next p1) ==∗ - â–· ∃ p2, iMsg_car m2 v (Next p2) ∗ â–· iProto_ctx γ ∗ + â–· ∃ p2, iMsg_car m2 v (Next p2) ∗ â–· iProto_ctx γ ∗ iProto_own γ i p1 ∗ iProto_own γ j p2. - Proof. Admitted. - (* iDestruct 1 as (pl pr) "(Hâ—l & Hâ—r & Hconsistent)". *) - (* iDestruct 1 as (pl') "[Hlel Hâ—¯l]". *) - (* iDestruct 1 as (pr') "[Hler Hâ—¯r]". *) - (* iIntros "Hm". *) - (* iDestruct (iProto_own_auth_agree with "Hâ—l Hâ—¯l") as "#Hpl". *) - (* iDestruct (iProto_own_auth_agree with "Hâ—r Hâ—¯r") as "#Hpr". *) - (* iAssert (â–· (pl ⊑ <!> m1))%I *) - (* with "[Hlel]" as "{Hpl} Hlel"; first (iNext; by iRewrite "Hpl"). *) - (* iAssert (â–· (pr ⊑ <?> m2))%I *) - (* with "[Hler]" as "{Hpr} Hler"; first (iNext; by iRewrite "Hpr"). *) - (* iDestruct (iProto_consistent_le_l with "Hconsistent Hlel") as "Hconsistent". *) - (* iDestruct (iProto_consistent_le_r with "Hconsistent Hler") as "Hconsistent". *) - (* iDestruct (iProto_consistent_step with "Hconsistent [Hm //]") as *) - (* (p2) "[Hm2 Hconsistent]". *) - (* iMod (iProto_own_auth_update _ _ _ _ p1 with "Hâ—l Hâ—¯l") as "[Hâ—l Hâ—¯l]". *) - (* iMod (iProto_own_auth_update _ _ _ _ p2 with "Hâ—r Hâ—¯r") as "[Hâ—r Hâ—¯r]". *) - (* iIntros "!>!>". *) - (* iExists p2. iFrame. *) - (* iSplitL "Hconsistent Hâ—l Hâ—r". *) - (* - iExists _, _. iFrame. *) - (* - iSplitL "Hâ—¯l". *) - (* + iExists _. iFrame. iApply iProto_le_refl. *) - (* + iExists _. iFrame. iApply iProto_le_refl. *) - (* Qed. *) + Proof. + iIntros "Hctx Hi Hj Hm". + iDestruct "Hctx" as (ps) "[Hauth Hconsistent]". + iDestruct (iProto_own_auth_agree with "Hauth Hi") as "#Hpi". + iDestruct (iProto_own_auth_agree with "Hauth Hj") as "#Hpj". + iDestruct (iProto_consistent_step with "Hconsistent [//] [//] [Hm //]") as + (p2) "[Hm2 Hconsistent]". + iMod (iProto_own_auth_update _ _ _ _ p1 with "Hauth Hi") as "[Hauth Hi]". + iMod (iProto_own_auth_update _ _ _ _ p2 with "Hauth Hj") as "[Hauth Hj]". + iIntros "!>!>". iExists p2. iFrame. iIntros "!>". iExists _. iFrame. + Qed. (* (** The instances below make it possible to use the tactics [iIntros], *) (* [iExist], [iSplitL]/[iSplitR], [iFrame] and [iModIntro] on [iProto_le] goals. *) *) -- GitLab From 207ec3c0c3699a8c0fe3ef35819f891d741c8e65 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Mon, 22 Jan 2024 05:43:52 +0100 Subject: [PATCH 06/81] Added proof of consistency for roundtrip example --- theories/channel/multi_proto.v | 319 ++++++++++++++++++++++++--------- 1 file changed, 237 insertions(+), 82 deletions(-) diff --git a/theories/channel/multi_proto.v b/theories/channel/multi_proto.v index 6f35f29..89da164 100644 --- a/theories/channel/multi_proto.v +++ b/theories/channel/multi_proto.v @@ -538,21 +538,15 @@ Instance iProto_inhabited {Σ V} : Inhabited (iProto Σ V) := populate (END). Definition can_step {Σ V} (rec : gmap nat (iProto Σ V) → iProp Σ) (ps : gmap nat (iProto Σ V)) (i j : nat) : iProp Σ := - (∀ a1 a2 m1 m2, - (ps !!! i ≡ <a1> m1) -∗ (ps !!! j ≡ <a2> m2) -∗ - match a1,a2 with - | Send j, Recv i => ∀ v p1, iMsg_car m1 v (Next p1) -∗ - ∃ p2, iMsg_car m2 v (Next p2) ∗ - â–· (rec (<[i:=p1]>(<[j:=p2]>ps))) - | Recv j, Send i => ∀ v p2, iMsg_car m2 v (Next p2) -∗ - ∃ p1, iMsg_car m1 v (Next p1) ∗ - â–· (rec (<[i:=p1]>(<[j:=p2]>ps))) - | _, _ => True - end). + ∀ m1 m2, + (ps !!! i ≡ <(Send j)> m1) -∗ (ps !!! j ≡ <Recv i> m2) -∗ + ∀ v p1, iMsg_car m1 v (Next p1) -∗ + ∃ p2, iMsg_car m2 v (Next p2) ∗ + â–· (rec (<[i:=p1]>(<[j:=p2]>ps))). Definition iProto_consistent_pre {Σ V} (rec : gmap nat (iProto Σ V) → iProp Σ) (ps : gmap nat (iProto Σ V)) : iProp Σ := - ∀ i j, ⌜i ∈ dom ps⌠-∗ ⌜j ∈ dom ps⌠-∗ can_step rec ps i j. + ∀ i j, can_step rec ps i j. Global Instance iProto_consistent_pre_ne {Σ V} (rec : gmapO natO (iProto Σ V) -n> iPropO Σ) : @@ -601,7 +595,9 @@ Lemma iProto_example1_consistent {Σ V} : ⊢ iProto_consistent (@iProto_example1 Σ V). Proof. rewrite iProto_consistent_unfold. - iIntros (i j Hi Hj). set_solver. + iIntros (i j m1 m2) "Hi Hj". + rewrite lookup_total_empty. + by rewrite iProto_end_message_equivI. Qed. Definition iProto_example2 `{!invGS Σ} (P : iProp Σ) : gmap nat (iProto Σ Z) := @@ -614,103 +610,268 @@ Lemma iProto_example2_consistent `{!invGS Σ} (P : iProp Σ) : Proof. rewrite iProto_consistent_unfold. rewrite /iProto_example2. - iIntros (i j Hi Hj). - iIntros (a1 a2 m1 m2) "Hm1 Hm2". + iIntros (i j m1 m2) "Hm1 Hm2". destruct i, j. - rewrite !lookup_total_insert. rewrite !iProto_message_equivI. - iDestruct "Hm1" as (<-) "#Hm1". - iDestruct "Hm2" as (<-) "#Hm2". + iDestruct "Hm1" as (Hieq) "#Hm1". + iDestruct "Hm2" as (Hjeq) "#Hm2". done. - - destruct j; [|set_solver]. + - destruct j; last first. + { rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + iDestruct "Hm1" as (Hieq) "#Hm1". done. } rewrite !lookup_total_insert. rewrite lookup_total_insert_ne; [|done]. rewrite !lookup_total_insert. rewrite !iProto_message_equivI. - iDestruct "Hm1" as (<-) "#Hm1". - iDestruct "Hm2" as (<-) "#Hm2". + iDestruct "Hm1" as (Hieq) "#Hm1". + iDestruct "Hm2" as (Hjeq) "#Hm2". iIntros (v p1) "Hm1'". iSpecialize ("Hm1" $!v (Next p1)). rewrite iMsg_base_eq. simpl. rewrite iMsg_exist_eq. simpl. iRewrite -"Hm1" in "Hm1'". - iExists END. + iExists END. iSpecialize ("Hm2" $!v (Next END)). - iRewrite -"Hm2". + iRewrite -"Hm2". simpl. iDestruct "Hm1'" as (x Heq) "[#Heq HP]". iSplitL. { iExists x. iSplit; [done|]. iFrame. done. } - iNext. iRewrite -"Heq". + iNext. iRewrite -"Heq". rewrite insert_commute; [|done]. rewrite insert_insert. rewrite insert_commute; [|done]. rewrite insert_insert. rewrite iProto_consistent_unfold. - iIntros (i' j' Hi' Hj'). - iIntros (a1 a2 m1' m2') "Hm1' Hm2'". + iIntros (i' j' m1' m2') "Hm1' Hm2'". destruct i'. - { rewrite lookup_total_insert. + { rewrite lookup_total_insert. iDestruct (iProto_end_message_equivI with "Hm1'") as "H". done. } destruct i'. - { rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert. + { rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert. iDestruct (iProto_end_message_equivI with "Hm1'") as "H". done. } - set_solver. - - destruct i; [|set_solver]. + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_alt. simpl. + iDestruct (iProto_end_message_equivI with "Hm1'") as "H". + done. + - destruct i; last first. + { rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_alt. simpl. + iDestruct (iProto_end_message_equivI with "Hm1") as "H". + done. } rewrite !lookup_total_insert. rewrite lookup_total_insert_ne; [|done]. rewrite !lookup_total_insert. rewrite !iProto_message_equivI. - iDestruct "Hm1" as (<-) "#Hm1". - iDestruct "Hm2" as (<-) "#Hm2". - iIntros (v p1) "Hm1'". - iSpecialize ("Hm2" $!v (Next p1)). - rewrite iMsg_base_eq. simpl. - rewrite iMsg_exist_eq. simpl. - iRewrite -"Hm2" in "Hm1'". - iExists END. - iSpecialize ("Hm1" $!v (Next END)). - iRewrite -"Hm1". - simpl. - iDestruct "Hm1'" as (x Heq) "[#Heq HP]". - iSplitL. - { iExists x. iSplit; [done|]. iFrame. done. } - iNext. iRewrite -"Heq". - rewrite insert_insert. - rewrite insert_commute; [|done]. - rewrite insert_insert. - rewrite iProto_consistent_unfold. - iIntros (i' j' Hi' Hj'). - iIntros (a1 a2 m1' m2') "Hm1' Hm2'". - destruct i'. - { rewrite lookup_total_insert. - iDestruct (iProto_end_message_equivI with "Hm1'") as "H". - done. } - destruct i'. - { rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert. - iDestruct (iProto_end_message_equivI with "Hm1'") as "H". - done. } - set_solver. - - destruct i; [|set_solver]. - destruct j; [|set_solver]. + iDestruct "Hm1" as (Hieq) "#Hm1". done. + - destruct i. + { rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert. + rewrite !iProto_message_equivI. + iDestruct "Hm1" as (Hieq) "#Hm1". done. } rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert. - rewrite !iProto_message_equivI. - iDestruct "Hm1" as (<-) "Hm1". - iDestruct "Hm2" as (<-) "Hm2". + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_alt. simpl. + iDestruct (iProto_end_message_equivI with "Hm1") as "H". done. Qed. -(* Lemma iProto_consistent_step {Σ V} (ps : gmap nat (iProto Σ V)) : *) -(* (∀ i j m1 m2, *) -(* ⌜ps !! i = Some (<(Send j)> m1)%proto⌠-∗ *) -(* ⌜ps !! j = Some (<(Recv i)> m2)⌠-∗ *) -(* iProto_consistent (ps))%I -∗ *) -(* iProto_consistent ps. *) +Definition iProto_example3 `{!invGS Σ} : gmap nat (iProto Σ Z) := + <[0 := <(Send 1) @ (x:Z)> MSG x ; <(Recv 2)> MSG x; END ]> + (<[1 := <(Recv 0) @ (x:Z)> MSG x ; <(Send 2)> MSG x; END ]> + (<[2 := <(Recv 1) @ (x:Z)> MSG x ; <(Send 0)> MSG x; END ]> + ∅)). + +Lemma iProto_example3_consistent `{!invGS Σ} : + ⊢ iProto_consistent (@iProto_example3 _ Σ invGS0). +Proof. + rewrite iProto_consistent_unfold. + rewrite /iProto_example3. + iIntros (i j m1 m2) "Hm1 Hm2". + destruct i; last first. + { destruct i. + { rewrite lookup_total_insert_ne; [|done]. + rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm1" as (Hieq) "#Hm1". } + destruct i. + { rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + rewrite !iProto_message_equivI. + by iDestruct "Hm1" as (Hieq) "#Hm1". } + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_empty=> /=. + by rewrite iProto_end_message_equivI. } + destruct j. + { rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm2" as (Hjeq) "#Hm2". } + destruct j; last first. + { rewrite !lookup_total_insert. + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + destruct j. + { rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm2" as (Hjeq) "#Hm2". } + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_empty=> /=. + by rewrite iProto_end_message_equivI. } + rewrite !lookup_total_insert. + rewrite lookup_total_insert_ne; [|done]. + rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + iDestruct "Hm1" as (Hieq) "#Hm1". + iDestruct "Hm2" as (Hjeq) "#Hm2". + iIntros (v p1) "Hm1'". + iSpecialize ("Hm1" $!v (Next p1)). + rewrite !iMsg_base_eq. + rewrite !iMsg_exist_eq. + iRewrite -"Hm1" in "Hm1'". + iDestruct "Hm1'" as (x Heq) "[#Hm1' _]". + iSpecialize ("Hm2" $!v (Next (<Send 2> iMsg_base_def x True END))). + iExists (<Send 2> iMsg_base_def x True END). + iRewrite -"Hm2". + simpl. + iSplitL. + { iExists x. iSplit; [done|]. iFrame. done. } + iNext. + rewrite iProto_consistent_unfold. + rewrite insert_commute; [|done]. + rewrite insert_insert. + rewrite insert_commute; [|done]. + rewrite insert_insert. + iRewrite -"Hm1'". + iClear (m1 m2 Hieq Hjeq p1 v Heq) "Hm1 Hm2 Hm1'". + iIntros (i j m1 m2) "Hm1 Hm2". + destruct i. + { rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm1" as (Hieq) "#Hm1". } + rewrite lookup_total_insert_ne; [|done]. + destruct i; last first. + { destruct i. + { rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm1" as (Hieq) "#Hm1". } + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_empty=> /=. + by rewrite iProto_end_message_equivI. } + rewrite lookup_total_insert. + destruct j. + { rewrite lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm2" as (Hjeq) "#Hm2". } + rewrite lookup_total_insert_ne; [|done]. + destruct j. + { rewrite lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm2" as (Hjeq) "#Hm2". } + rewrite lookup_total_insert_ne; [|done]. + destruct j; last first. + { rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_empty. + by rewrite iProto_end_message_equivI. } + rewrite lookup_total_insert. + rewrite !iProto_message_equivI. + iDestruct "Hm1" as (Hieq) "#Hm1". + iDestruct "Hm2" as (Hjeq) "#Hm2". + iIntros (v p1) "Hm1'". + iSpecialize ("Hm1" $!v (Next p1)). + iRewrite -"Hm1" in "Hm1'". + iDestruct "Hm1'" as (Heq) "[#Hm1' _]". + iSpecialize ("Hm2" $!v (Next (<Send 0> iMsg_base_def x True END))). + iExists (<Send 0> iMsg_base_def x True END). + iRewrite -"Hm2". + simpl. + iSplitL. + { iExists x. iSplit; [done|]. iFrame. done. } + iNext. + rewrite iProto_consistent_unfold. + rewrite (insert_commute _ 1 2); [|done]. + rewrite (insert_commute _ 1 0); [|done]. + rewrite insert_insert. + rewrite (insert_commute _ 2 0); [|done]. + rewrite (insert_commute _ 2 1); [|done]. + rewrite insert_insert. + iRewrite -"Hm1'". + iClear (m1 m2 Hieq Hjeq p1 v Heq) "Hm1 Hm2 Hm1'". + iIntros (i j m1 m2) "Hm1 Hm2". + destruct i. + { rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm1" as (Hieq) "#Hm1". } + rewrite lookup_total_insert_ne; [|done]. + destruct i. + { rewrite lookup_total_insert. + by rewrite iProto_end_message_equivI. } + rewrite lookup_total_insert_ne; [|done]. + destruct i; last first. + { rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_empty. + by rewrite iProto_end_message_equivI. } + rewrite lookup_total_insert. + destruct j; last first. + { rewrite lookup_total_insert_ne; [|done]. + destruct j. + { rewrite lookup_total_insert. + by rewrite iProto_end_message_equivI. } + rewrite lookup_total_insert_ne; [|done]. + destruct j. + { rewrite lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm2" as (Hjeq) "#Hm2". } + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_empty. + by rewrite iProto_end_message_equivI. } + rewrite lookup_total_insert. + rewrite !iProto_message_equivI. + iDestruct "Hm1" as (Hieq) "#Hm1". + iDestruct "Hm2" as (Hjeq) "#Hm2". + iIntros (v p1) "Hm1'". + iSpecialize ("Hm1" $!v (Next p1)). + iRewrite -"Hm1" in "Hm1'". + iDestruct "Hm1'" as (Heq) "[#Hm1' _]". + iSpecialize ("Hm2" $!v (Next END)). + iExists END. + iRewrite -"Hm2". + simpl. + iSplitL; [done|]. + rewrite insert_insert. + rewrite insert_commute; [|done]. + rewrite (insert_commute _ 2 1); [|done]. + rewrite insert_insert. + iNext. + iRewrite -"Hm1'". + rewrite iProto_consistent_unfold. + iClear (m1 m2 Hieq Hjeq p1 v Heq) "Hm1 Hm2 Hm1'". + iIntros (i j m1 m2) "Hm1 Hm2". + destruct i. + { rewrite lookup_total_insert. + by rewrite !iProto_end_message_equivI. } + rewrite lookup_total_insert_ne; [|done]. + destruct i. + { rewrite lookup_total_insert. + by rewrite !iProto_end_message_equivI. } + rewrite lookup_total_insert_ne; [|done]. + destruct i. + { rewrite lookup_total_insert. + by rewrite !iProto_end_message_equivI. } + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_empty. + by rewrite iProto_end_message_equivI. +Qed. Record proto_name := ProtName { proto_names : gmap nat gname }. Global Instance proto_name_inhabited : Inhabited proto_name := @@ -769,13 +930,7 @@ Section proto. Proof. iIntros "Hprot #Hi #Hj Hm1". rewrite iProto_consistent_unfold /iProto_consistent_pre. - iDestruct ("Hprot" $!i j with "[] [] Hi Hj Hm1") as (p2) "[Hm2 Hprot]". - { rewrite !lookup_total_alt elem_of_dom. - destruct (ps !! i) eqn: Hi; [done|]=> /=. - by rewrite iProto_end_message_equivI. } - { rewrite !lookup_total_alt elem_of_dom. - destruct (ps !! j) eqn: Hj; [done|]=> /=. - by rewrite iProto_end_message_equivI. } + iDestruct ("Hprot" with "Hi Hj Hm1") as (p2) "[Hm2 Hprot]". iExists p2. iFrame. Qed. @@ -809,8 +964,8 @@ Section proto. iProto_consistent ps -∗ |==> ∃ γ, iProto_ctx γ ∗ [∗ map] i ↦p ∈ ps, iProto_own γ i p. Proof. Admitted. - (* iMod (own_alloc (â—E (Next p) â‹… â—¯E (Next p))) as (lγ) "[Hâ—l Hâ—¯l]". *) - (* { by apply excl_auth_valid. } *) + (* iMod (own_alloc (â—E (Next p) â‹… â—¯E (Next p))) as (lγ) "[Hâ—l Hâ—¯l]". *) + (* { by apply excl_auth_valid. } *) (* iMod (own_alloc (â—E (Next (iProto_dual p)) â‹… *) (* â—¯E (Next (iProto_dual p)))) as (rγ) "[Hâ—r Hâ—¯r]". *) (* { by apply excl_auth_valid. } *) -- GitLab From 1fab9b4d1f44df9e45dcc65c776cf51554949ee3 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Mon, 22 Jan 2024 06:51:55 +0100 Subject: [PATCH 07/81] Solve remaining admits --- theories/channel/multi_proto.v | 122 +++++++++++++++++---------------- 1 file changed, 62 insertions(+), 60 deletions(-) diff --git a/theories/channel/multi_proto.v b/theories/channel/multi_proto.v index 89da164..e5507e8 100644 --- a/theories/channel/multi_proto.v +++ b/theories/channel/multi_proto.v @@ -45,7 +45,7 @@ Lastly, relevant type classes instances are defined for each of the above notions, such as contractiveness and non-expansiveness, after which the specifications of the message-passing primitives are defined in terms of the protocol connectives. *) -From iris.algebra Require Import gmap excl_auth. +From iris.algebra Require Import gmap excl_auth gmap_view. From iris.proofmode Require Import proofmode. From iris.base_logic Require Export lib.iprop. From iris.base_logic Require Import lib.own. @@ -57,39 +57,15 @@ Export action. (** * Setup of Iris's cameras *) Class protoG Σ V := protoG_authG :: - inG Σ (authUR (gmapR natO - (optionUR (exclR (laterO (proto (leibnizO V) (iPropO Σ) (iPropO Σ))))))). + inG Σ (gmap_viewR natO + (optionUR (exclR (laterO (proto (leibnizO V) (iPropO Σ) (iPropO Σ)))))). Definition protoΣ V := #[ - GFunctor (authRF (gmapURF natO (optionRF (exclRF (laterOF (protoOF (leibnizO V) idOF idOF)))))) + GFunctor ((gmap_viewRF natO (optionRF (exclRF (laterOF (protoOF (leibnizO V) idOF idOF)))))) ]. Global Instance subG_chanΣ {Σ V} : subG (protoΣ V) Σ → protoG Σ V. Proof. solve_inG. Qed. -(* (** * Setup of Iris's cameras *) *) -(* Class protoG Σ V := *) -(* protoG_authG :: *) -(* inG Σ (gmapR natO *) -(* (excl_authR (laterO (proto (leibnizO V) (iPropO Σ) (iPropO Σ))))). *) - -(* Definition protoΣ V := #[ *) -(* GFunctor (gmapRF natO (authRF (optionURF (exclRF (laterOF (protoOF (leibnizO V) idOF idOF)))))) *) -(* ]. *) -(* Global Instance subG_chanΣ {Σ V} : subG (protoΣ V) Σ → protoG Σ V. *) -(* Proof. solve_inG. Qed. *) - - -(* (** * Setup of Iris's cameras *) *) -(* Class protoG Σ V := *) -(* protoG_authG :: *) -(* inG Σ (excl_authR (laterO (proto (leibnizO V) (iPropO Σ) (iPropO Σ)))). *) - -(* Definition protoΣ V := #[ *) -(* GFunctor (authRF (optionURF (exclRF (laterOF (protoOF (leibnizO V) idOF idOF))))) *) -(* ]. *) -(* Global Instance subG_chanΣ {Σ V} : subG (protoΣ V) Σ → protoG Σ V. *) -(* Proof. solve_inG. Qed. *) - (** * Types *) Definition iProto Σ V := proto V (iPropO Σ) (iPropO Σ). Declare Scope proto_scope. @@ -550,8 +526,8 @@ Definition iProto_consistent_pre {Σ V} (rec : gmap nat (iProto Σ V) → iProp Global Instance iProto_consistent_pre_ne {Σ V} (rec : gmapO natO (iProto Σ V) -n> iPropO Σ) : - NonExpansive (iProto_consistent_pre (λ ps, rec ps)). -Proof. Admitted. + NonExpansive (iProto_consistent_pre rec). +Proof. rewrite /iProto_consistent_pre /can_step. solve_proper. Qed. Program Definition iProto_consistent_pre' {Σ V} (rec : gmapO natO (iProto Σ V) -n> iPropO Σ) : @@ -559,7 +535,10 @@ Program Definition iProto_consistent_pre' {Σ V} λne ps, iProto_consistent_pre (λ ps, rec ps) ps. Local Instance iProto_consistent_pre_contractive {Σ V} : Contractive (@iProto_consistent_pre' Σ V). -Proof. Admitted. +Proof. + rewrite /iProto_consistent_pre' /iProto_consistent_pre /can_step. + solve_contractive. +Qed. Definition iProto_consistent {Σ V} (ps : gmap nat (iProto Σ V)) : iProp Σ := fixpoint iProto_consistent_pre' ps. @@ -872,7 +851,7 @@ Proof. rewrite lookup_total_empty. by rewrite iProto_end_message_equivI. Qed. - + Record proto_name := ProtName { proto_names : gmap nat gname }. Global Instance proto_name_inhabited : Inhabited proto_name := populate (ProtName inhabitant). @@ -886,13 +865,12 @@ Qed. Definition iProto_own_frag `{!protoG Σ V} (γ : gname) (i : nat) (p : iProto Σ V) : iProp Σ := - own γ (â—¯ ({[i := Excl' (Next p)]})). + own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p))). -(* TODO: Fix this def. *) Definition iProto_own_auth `{!protoG Σ V} (γ : gname) (ps : gmap nat (iProto Σ V)) : iProp Σ := - own γ (◠∅). - (* own γ (â— ((λ (p : iProto Σ V), Excl' (Next p)) <$> ps)). *) + (* own γ (◠∅). *) + own γ (gmap_view_auth (DfracOwn 1) (((λ p, Excl' (Next p)) <$> ps) : gmap _ _)). Definition iProto_ctx `{protoG Σ V} (γ : gname) : iProp Σ := @@ -937,42 +915,66 @@ Section proto. Global Instance iProto_own_frag_ne γ s : NonExpansive (iProto_own_frag γ s). Proof. solve_proper. Qed. - (* TODO: Relies on above def *) Lemma iProto_own_auth_agree γ ps i p : iProto_own_auth γ ps -∗ iProto_own_frag γ i p -∗ â–· (ps !!! i ≡ p). - Proof. Admitted. - (* iIntros "Hâ— Hâ—¯". iDestruct (own_valid_2 with "Hâ— Hâ—¯") as "H✓". *) - (* iDestruct (excl_auth_agreeI with "H✓") as "H✓". *) - (* iApply (later_equivI_1 with "H✓"). *) - (* Qed. *) - + Proof. + iIntros "Hâ— Hâ—¯". + iDestruct (own_valid_2 with "Hâ— Hâ—¯") as "H✓". + rewrite gmap_view_both_validI. + iDestruct "H✓" as "[_ [H1 H2]]". + rewrite lookup_total_alt. + rewrite lookup_fmap. + destruct (ps !! i); last first. + { simpl. rewrite !option_equivI. done. } + simpl. + rewrite !option_equivI excl_equivI. + by iNext. + Qed. + Lemma iProto_own_auth_update γ ps i p p' : iProto_own_auth γ ps -∗ iProto_own_frag γ i p ==∗ iProto_own_auth γ (<[i := p']>ps) ∗ iProto_own_frag γ i p'. - Proof. Admitted. - (* iIntros "Hâ— Hâ—¯". iDestruct (own_update_2 with "Hâ— Hâ—¯") as "H". *) - (* { eapply (excl_auth_update _ _ (Next p'')). } *) - (* by rewrite own_op. *) - (* Qed. *) + Proof. + iIntros "Hâ— Hâ—¯". + iMod (own_update_2 with "Hâ— Hâ—¯") as "[H1 H2]"; [|iModIntro]. + { eapply (gmap_view_replace _ _ _ (Excl' (Next p'))). done. } + iFrame. rewrite -fmap_insert. done. + Qed. Global Instance iProto_own_ne γ s : NonExpansive (iProto_own γ s). Proof. solve_proper. Qed. Global Instance iProto_own_proper γ s : Proper ((≡) ==> (≡)) (iProto_own γ s). Proof. apply (ne_proper _). Qed. + Lemma iProto_own_auth_alloc ps : + ⊢ |==> ∃ γ, iProto_own_auth γ ps ∗ [∗ map] i ↦p ∈ ps, iProto_own γ i p. + Proof. + iMod (own_alloc (gmap_view_auth (DfracOwn 1) ∅)) as (γ) "Hauth". + { apply gmap_view_auth_valid. } + iExists γ. + iInduction ps as [|i p ps Hin] "IH" using map_ind. + { iModIntro. iFrame. + by iApply big_sepM_empty. } + iMod ("IH" with "Hauth") as "[Hauth Hfrags]". + rewrite big_sepM_insert; [|done]. iFrame "Hfrags". + iMod (own_update with "Hauth") as "[Hauth Hfrag]". + { apply (gmap_view_alloc _ i (DfracOwn 1) (Excl' (Next p))). + - rewrite lookup_fmap. rewrite Hin. done. + - done. + - done. } + iFrame. + iModIntro. + rewrite -fmap_insert. iFrame. + Qed. + Lemma iProto_init ps : - iProto_consistent ps -∗ + â–· iProto_consistent ps -∗ |==> ∃ γ, iProto_ctx γ ∗ [∗ map] i ↦p ∈ ps, iProto_own γ i p. - Proof. Admitted. - (* iMod (own_alloc (â—E (Next p) â‹… â—¯E (Next p))) as (lγ) "[Hâ—l Hâ—¯l]". *) - (* { by apply excl_auth_valid. } *) - (* iMod (own_alloc (â—E (Next (iProto_dual p)) â‹… *) - (* â—¯E (Next (iProto_dual p)))) as (rγ) "[Hâ—r Hâ—¯r]". *) - (* { by apply excl_auth_valid. } *) - (* pose (ProtName lγ rγ) as γ. iModIntro. iExists γ. iSplitL "Hâ—l Hâ—r". *) - (* { iExists p, (iProto_dual p). iFrame. iApply iProto_consistent_dual. } *) - (* iSplitL "Hâ—¯l"; iExists _; iFrame; iApply iProto_le_refl. *) - (* Qed. *) + Proof. + iIntros "Hconsistnet". + iMod iProto_own_auth_alloc as (γ) "[Hauth Hfrags]". + iExists γ. iFrame. iExists _. iFrame. done. + Qed. Lemma iProto_step γ i j m1 m2 p1 v : iProto_ctx γ -∗ @@ -988,8 +990,8 @@ Section proto. iDestruct (iProto_own_auth_agree with "Hauth Hj") as "#Hpj". iDestruct (iProto_consistent_step with "Hconsistent [//] [//] [Hm //]") as (p2) "[Hm2 Hconsistent]". - iMod (iProto_own_auth_update _ _ _ _ p1 with "Hauth Hi") as "[Hauth Hi]". iMod (iProto_own_auth_update _ _ _ _ p2 with "Hauth Hj") as "[Hauth Hj]". + iMod (iProto_own_auth_update _ _ _ _ p1 with "Hauth Hi") as "[Hauth Hi]". iIntros "!>!>". iExists p2. iFrame. iIntros "!>". iExists _. iFrame. Qed. -- GitLab From 4af1c72f8abb90b036463e032dd1e2fda6cd35fc Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Mon, 22 Jan 2024 07:16:44 +0100 Subject: [PATCH 08/81] Moved examples to other file. Added naive def of iProto_le --- theories/channel/multi_proto.v | 322 ++---------------- .../multi_proto_consistency_examples.v | 293 ++++++++++++++++ 2 files changed, 317 insertions(+), 298 deletions(-) create mode 100644 theories/channel/multi_proto_consistency_examples.v diff --git a/theories/channel/multi_proto.v b/theories/channel/multi_proto.v index e5507e8..984b84c 100644 --- a/theories/channel/multi_proto.v +++ b/theories/channel/multi_proto.v @@ -556,302 +556,17 @@ Lemma iProto_consistent_unfold {Σ V} (ps : gmap nat (iProto Σ V)) : Proof. apply: (fixpoint_unfold iProto_consistent_pre'). Qed. -(* Definition iProto_le {Σ V} (p1 p2 : iProto Σ V) : iProp Σ := *) -(* ∀ p3, iProto_consistent p1 p3 -∗ iProto_consistent p2 p3. *) -(* Arguments iProto_le {_ _} _%proto _%proto. *) -(* Global Instance: Params (@iProto_le) 2 := {}. *) -(* Notation "p ⊑ q" := (iProto_le p q) : bi_scope. *) - -(* Global Instance iProto_le_ne {Σ V} : NonExpansive2 (@iProto_le Σ V). *) -(* Proof. solve_proper. Qed. *) -(* Global Instance iProto_le_proper {Σ V} : Proper ((≡) ==> (≡) ==> (⊣⊢)) (@iProto_le Σ V). *) -(* Proof. solve_proper. Qed. *) - -Definition iProto_example1 {Σ V} : gmap nat (iProto Σ V) := - ∅. - -Lemma iProto_example1_consistent {Σ V} : - ⊢ iProto_consistent (@iProto_example1 Σ V). -Proof. - rewrite iProto_consistent_unfold. - iIntros (i j m1 m2) "Hi Hj". - rewrite lookup_total_empty. - by rewrite iProto_end_message_equivI. -Qed. - -Definition iProto_example2 `{!invGS Σ} (P : iProp Σ) : gmap nat (iProto Σ Z) := - <[0 := (<(Send 1) @ (x:Z)> MSG x {{ P }} ; END)%proto ]> - (<[1 := (<(Recv 0) @ (x:Z)> MSG x {{ P }} ; END)%proto ]> - ∅). - -Lemma iProto_example2_consistent `{!invGS Σ} (P : iProp Σ) : - ⊢ iProto_consistent (@iProto_example2 _ Σ invGS0 P). -Proof. - rewrite iProto_consistent_unfold. - rewrite /iProto_example2. - iIntros (i j m1 m2) "Hm1 Hm2". - destruct i, j. - - rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - iDestruct "Hm1" as (Hieq) "#Hm1". - iDestruct "Hm2" as (Hjeq) "#Hm2". - done. - - destruct j; last first. - { rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - iDestruct "Hm1" as (Hieq) "#Hm1". done. } - rewrite !lookup_total_insert. - rewrite lookup_total_insert_ne; [|done]. - rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - iDestruct "Hm1" as (Hieq) "#Hm1". - iDestruct "Hm2" as (Hjeq) "#Hm2". - iIntros (v p1) "Hm1'". - iSpecialize ("Hm1" $!v (Next p1)). - rewrite iMsg_base_eq. simpl. - rewrite iMsg_exist_eq. simpl. - iRewrite -"Hm1" in "Hm1'". - iExists END. - iSpecialize ("Hm2" $!v (Next END)). - iRewrite -"Hm2". - simpl. - iDestruct "Hm1'" as (x Heq) "[#Heq HP]". - iSplitL. - { iExists x. iSplit; [done|]. iFrame. done. } - iNext. iRewrite -"Heq". - rewrite insert_commute; [|done]. - rewrite insert_insert. - rewrite insert_commute; [|done]. - rewrite insert_insert. - rewrite iProto_consistent_unfold. - iIntros (i' j' m1' m2') "Hm1' Hm2'". - destruct i'. - { rewrite lookup_total_insert. - iDestruct (iProto_end_message_equivI with "Hm1'") as "H". - done. } - destruct i'. - { rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert. - iDestruct (iProto_end_message_equivI with "Hm1'") as "H". - done. } - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_alt. simpl. - iDestruct (iProto_end_message_equivI with "Hm1'") as "H". - done. - - destruct i; last first. - { rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_alt. simpl. - iDestruct (iProto_end_message_equivI with "Hm1") as "H". - done. } - rewrite !lookup_total_insert. - rewrite lookup_total_insert_ne; [|done]. - rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - iDestruct "Hm1" as (Hieq) "#Hm1". done. - - destruct i. - { rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert. - rewrite !iProto_message_equivI. - iDestruct "Hm1" as (Hieq) "#Hm1". done. } - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_alt. simpl. - iDestruct (iProto_end_message_equivI with "Hm1") as "H". - done. -Qed. +Definition iProto_le {Σ V} (i:nat) (p1 p2 : iProto Σ V) : iProp Σ := + ∀ ps, iProto_consistent (<[i:=p1]>ps) -∗ iProto_consistent (<[i:=p2]>ps). +Arguments iProto_le {_ _} _ _%proto _%proto. +Global Instance: Params (@iProto_le) 3 := {}. +(* Notation "p ⊑ q" := (iProto_le i p q) : bi_scope. *) -Definition iProto_example3 `{!invGS Σ} : gmap nat (iProto Σ Z) := - <[0 := <(Send 1) @ (x:Z)> MSG x ; <(Recv 2)> MSG x; END ]> - (<[1 := <(Recv 0) @ (x:Z)> MSG x ; <(Send 2)> MSG x; END ]> - (<[2 := <(Recv 1) @ (x:Z)> MSG x ; <(Send 0)> MSG x; END ]> - ∅)). +Global Instance iProto_le_ne {Σ V} n : Proper ((=) ==> (dist n) ==> (dist n) ==> (dist n)) (@iProto_le Σ V). +Proof. solve_proper. Qed. +Global Instance iProto_le_proper {Σ V} : Proper ((=) ==> (≡) ==> (≡) ==> (⊣⊢)) (@iProto_le Σ V). +Proof. solve_proper. Qed. -Lemma iProto_example3_consistent `{!invGS Σ} : - ⊢ iProto_consistent (@iProto_example3 _ Σ invGS0). -Proof. - rewrite iProto_consistent_unfold. - rewrite /iProto_example3. - iIntros (i j m1 m2) "Hm1 Hm2". - destruct i; last first. - { destruct i. - { rewrite lookup_total_insert_ne; [|done]. - rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm1" as (Hieq) "#Hm1". } - destruct i. - { rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - rewrite !iProto_message_equivI. - by iDestruct "Hm1" as (Hieq) "#Hm1". } - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_empty=> /=. - by rewrite iProto_end_message_equivI. } - destruct j. - { rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm2" as (Hjeq) "#Hm2". } - destruct j; last first. - { rewrite !lookup_total_insert. - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - destruct j. - { rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm2" as (Hjeq) "#Hm2". } - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_empty=> /=. - by rewrite iProto_end_message_equivI. } - rewrite !lookup_total_insert. - rewrite lookup_total_insert_ne; [|done]. - rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - iDestruct "Hm1" as (Hieq) "#Hm1". - iDestruct "Hm2" as (Hjeq) "#Hm2". - iIntros (v p1) "Hm1'". - iSpecialize ("Hm1" $!v (Next p1)). - rewrite !iMsg_base_eq. - rewrite !iMsg_exist_eq. - iRewrite -"Hm1" in "Hm1'". - iDestruct "Hm1'" as (x Heq) "[#Hm1' _]". - iSpecialize ("Hm2" $!v (Next (<Send 2> iMsg_base_def x True END))). - iExists (<Send 2> iMsg_base_def x True END). - iRewrite -"Hm2". - simpl. - iSplitL. - { iExists x. iSplit; [done|]. iFrame. done. } - iNext. - rewrite iProto_consistent_unfold. - rewrite insert_commute; [|done]. - rewrite insert_insert. - rewrite insert_commute; [|done]. - rewrite insert_insert. - iRewrite -"Hm1'". - iClear (m1 m2 Hieq Hjeq p1 v Heq) "Hm1 Hm2 Hm1'". - iIntros (i j m1 m2) "Hm1 Hm2". - destruct i. - { rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm1" as (Hieq) "#Hm1". } - rewrite lookup_total_insert_ne; [|done]. - destruct i; last first. - { destruct i. - { rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm1" as (Hieq) "#Hm1". } - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_empty=> /=. - by rewrite iProto_end_message_equivI. } - rewrite lookup_total_insert. - destruct j. - { rewrite lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm2" as (Hjeq) "#Hm2". } - rewrite lookup_total_insert_ne; [|done]. - destruct j. - { rewrite lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm2" as (Hjeq) "#Hm2". } - rewrite lookup_total_insert_ne; [|done]. - destruct j; last first. - { rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_empty. - by rewrite iProto_end_message_equivI. } - rewrite lookup_total_insert. - rewrite !iProto_message_equivI. - iDestruct "Hm1" as (Hieq) "#Hm1". - iDestruct "Hm2" as (Hjeq) "#Hm2". - iIntros (v p1) "Hm1'". - iSpecialize ("Hm1" $!v (Next p1)). - iRewrite -"Hm1" in "Hm1'". - iDestruct "Hm1'" as (Heq) "[#Hm1' _]". - iSpecialize ("Hm2" $!v (Next (<Send 0> iMsg_base_def x True END))). - iExists (<Send 0> iMsg_base_def x True END). - iRewrite -"Hm2". - simpl. - iSplitL. - { iExists x. iSplit; [done|]. iFrame. done. } - iNext. - rewrite iProto_consistent_unfold. - rewrite (insert_commute _ 1 2); [|done]. - rewrite (insert_commute _ 1 0); [|done]. - rewrite insert_insert. - rewrite (insert_commute _ 2 0); [|done]. - rewrite (insert_commute _ 2 1); [|done]. - rewrite insert_insert. - iRewrite -"Hm1'". - iClear (m1 m2 Hieq Hjeq p1 v Heq) "Hm1 Hm2 Hm1'". - iIntros (i j m1 m2) "Hm1 Hm2". - destruct i. - { rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm1" as (Hieq) "#Hm1". } - rewrite lookup_total_insert_ne; [|done]. - destruct i. - { rewrite lookup_total_insert. - by rewrite iProto_end_message_equivI. } - rewrite lookup_total_insert_ne; [|done]. - destruct i; last first. - { rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_empty. - by rewrite iProto_end_message_equivI. } - rewrite lookup_total_insert. - destruct j; last first. - { rewrite lookup_total_insert_ne; [|done]. - destruct j. - { rewrite lookup_total_insert. - by rewrite iProto_end_message_equivI. } - rewrite lookup_total_insert_ne; [|done]. - destruct j. - { rewrite lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm2" as (Hjeq) "#Hm2". } - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_empty. - by rewrite iProto_end_message_equivI. } - rewrite lookup_total_insert. - rewrite !iProto_message_equivI. - iDestruct "Hm1" as (Hieq) "#Hm1". - iDestruct "Hm2" as (Hjeq) "#Hm2". - iIntros (v p1) "Hm1'". - iSpecialize ("Hm1" $!v (Next p1)). - iRewrite -"Hm1" in "Hm1'". - iDestruct "Hm1'" as (Heq) "[#Hm1' _]". - iSpecialize ("Hm2" $!v (Next END)). - iExists END. - iRewrite -"Hm2". - simpl. - iSplitL; [done|]. - rewrite insert_insert. - rewrite insert_commute; [|done]. - rewrite (insert_commute _ 2 1); [|done]. - rewrite insert_insert. - iNext. - iRewrite -"Hm1'". - rewrite iProto_consistent_unfold. - iClear (m1 m2 Hieq Hjeq p1 v Heq) "Hm1 Hm2 Hm1'". - iIntros (i j m1 m2) "Hm1 Hm2". - destruct i. - { rewrite lookup_total_insert. - by rewrite !iProto_end_message_equivI. } - rewrite lookup_total_insert_ne; [|done]. - destruct i. - { rewrite lookup_total_insert. - by rewrite !iProto_end_message_equivI. } - rewrite lookup_total_insert_ne; [|done]. - destruct i. - { rewrite lookup_total_insert. - by rewrite !iProto_end_message_equivI. } - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_empty. - by rewrite iProto_end_message_equivI. -Qed. - Record proto_name := ProtName { proto_names : gmap nat gname }. Global Instance proto_name_inhabited : Inhabited proto_name := populate (ProtName inhabitant). @@ -898,6 +613,17 @@ Section proto. Implicit Types p pl pr : iProto Σ V. Implicit Types m : iMsg Σ V. + Lemma iProto_consistent_le ps i p1 p2 : + ps !! i = Some p1 → + iProto_le i p1 p2 -∗ + iProto_consistent ps -∗ + iProto_consistent (<[i := p2]>ps). + Proof. + iIntros (HSome) "Hle". + rewrite -{1}(insert_id ps i p1); [|done]. + iApply "Hle". + Qed. + Lemma iProto_consistent_step ps m1 m2 i j v p1 : iProto_consistent ps -∗ ps !!! i ≡ (<(Send j)> m1) -∗ @@ -918,7 +644,7 @@ Section proto. Lemma iProto_own_auth_agree γ ps i p : iProto_own_auth γ ps -∗ iProto_own_frag γ i p -∗ â–· (ps !!! i ≡ p). Proof. - iIntros "Hâ— Hâ—¯". + iIntros "Hâ— Hâ—¯". iDestruct (own_valid_2 with "Hâ— Hâ—¯") as "H✓". rewrite gmap_view_both_validI. iDestruct "H✓" as "[_ [H1 H2]]". @@ -930,7 +656,7 @@ Section proto. rewrite !option_equivI excl_equivI. by iNext. Qed. - + Lemma iProto_own_auth_update γ ps i p p' : iProto_own_auth γ ps -∗ iProto_own_frag γ i p ==∗ iProto_own_auth γ (<[i := p']>ps) ∗ iProto_own_frag γ i p'. @@ -965,7 +691,7 @@ Section proto. iFrame. iModIntro. rewrite -fmap_insert. iFrame. - Qed. + Qed. Lemma iProto_init ps : â–· iProto_consistent ps -∗ @@ -974,7 +700,7 @@ Section proto. iIntros "Hconsistnet". iMod iProto_own_auth_alloc as (γ) "[Hauth Hfrags]". iExists γ. iFrame. iExists _. iFrame. done. - Qed. + Qed. Lemma iProto_step γ i j m1 m2 p1 v : iProto_ctx γ -∗ diff --git a/theories/channel/multi_proto_consistency_examples.v b/theories/channel/multi_proto_consistency_examples.v new file mode 100644 index 0000000..4863f4a --- /dev/null +++ b/theories/channel/multi_proto_consistency_examples.v @@ -0,0 +1,293 @@ +From iris.algebra Require Import gmap excl_auth gmap_view. +From iris.proofmode Require Import proofmode. +From iris.base_logic Require Export lib.iprop. +From iris.base_logic Require Import lib.own. +From iris.program_logic Require Import language. +From actris.channel Require Import multi_proto_model multi_proto. +Set Default Proof Using "Type". +Export action. + +Definition iProto_example1 {Σ V} : gmap nat (iProto Σ V) := + ∅. + +Lemma iProto_example1_consistent {Σ V} : + ⊢ iProto_consistent (@iProto_example1 Σ V). +Proof. + rewrite iProto_consistent_unfold. + iIntros (i j m1 m2) "Hi Hj". + rewrite lookup_total_empty. + by rewrite iProto_end_message_equivI. +Qed. + +Definition iProto_example2 `{!invGS Σ} (P : iProp Σ) : gmap nat (iProto Σ Z) := + <[0 := (<(Send 1) @ (x:Z)> MSG x {{ P }} ; END)%proto ]> + (<[1 := (<(Recv 0) @ (x:Z)> MSG x {{ P }} ; END)%proto ]> + ∅). + +Lemma iProto_example2_consistent `{!invGS Σ} (P : iProp Σ) : + ⊢ iProto_consistent (@iProto_example2 _ Σ invGS0 P). +Proof. + rewrite iProto_consistent_unfold. + rewrite /iProto_example2. + iIntros (i j m1 m2) "Hm1 Hm2". + destruct i, j. + - rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + iDestruct "Hm1" as (Hieq) "#Hm1". + iDestruct "Hm2" as (Hjeq) "#Hm2". + done. + - destruct j; last first. + { rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + iDestruct "Hm1" as (Hieq) "#Hm1". done. } + rewrite !lookup_total_insert. + rewrite lookup_total_insert_ne; [|done]. + rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + iDestruct "Hm1" as (Hieq) "#Hm1". + iDestruct "Hm2" as (Hjeq) "#Hm2". + iIntros (v p1) "Hm1'". + iSpecialize ("Hm1" $!v (Next p1)). + rewrite iMsg_base_eq. simpl. + rewrite iMsg_exist_eq. simpl. + iRewrite -"Hm1" in "Hm1'". + iExists END%proto. + iSpecialize ("Hm2" $!v (Next END%proto)). + iRewrite -"Hm2". + simpl. + iDestruct "Hm1'" as (x Heq) "[#Heq HP]". + iSplitL. + { iExists x. iSplit; [done|]. iFrame. done. } + iNext. iRewrite -"Heq". + rewrite insert_commute; [|done]. + rewrite insert_insert. + rewrite insert_commute; [|done]. + rewrite insert_insert. + rewrite iProto_consistent_unfold. + iIntros (i' j' m1' m2') "Hm1' Hm2'". + destruct i'. + { rewrite lookup_total_insert. + iDestruct (iProto_end_message_equivI with "Hm1'") as "H". + done. } + destruct i'. + { rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert. + iDestruct (iProto_end_message_equivI with "Hm1'") as "H". + done. } + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_alt. simpl. + iDestruct (iProto_end_message_equivI with "Hm1'") as "H". + done. + - destruct i; last first. + { rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_alt. simpl. + iDestruct (iProto_end_message_equivI with "Hm1") as "H". + done. } + rewrite !lookup_total_insert. + rewrite lookup_total_insert_ne; [|done]. + rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + iDestruct "Hm1" as (Hieq) "#Hm1". done. + - destruct i. + { rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert. + rewrite !iProto_message_equivI. + iDestruct "Hm1" as (Hieq) "#Hm1". done. } + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_alt. simpl. + iDestruct (iProto_end_message_equivI with "Hm1") as "H". + done. +Qed. + +Definition iProto_example3 `{!invGS Σ} : gmap nat (iProto Σ Z) := + <[0 := (<(Send 1) @ (x:Z)> MSG x ; <(Recv 2)> MSG x; END)%proto ]> + (<[1 := (<(Recv 0) @ (x:Z)> MSG x ; <(Send 2)> MSG x; END)%proto ]> + (<[2 := (<(Recv 1) @ (x:Z)> MSG x ; <(Send 0)> MSG x; END)%proto ]> + ∅)). + +Lemma iProto_example3_consistent `{!invGS Σ} : + ⊢ iProto_consistent (@iProto_example3 _ Σ invGS0). +Proof. + rewrite iProto_consistent_unfold. + rewrite /iProto_example3. + iIntros (i j m1 m2) "Hm1 Hm2". + destruct i; last first. + { destruct i. + { rewrite lookup_total_insert_ne; [|done]. + rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm1" as (Hieq) "#Hm1". } + destruct i. + { rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + rewrite !iProto_message_equivI. + by iDestruct "Hm1" as (Hieq) "#Hm1". } + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_empty=> /=. + by rewrite iProto_end_message_equivI. } + destruct j. + { rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm2" as (Hjeq) "#Hm2". } + destruct j; last first. + { rewrite !lookup_total_insert. + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + destruct j. + { rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm2" as (Hjeq) "#Hm2". } + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_empty=> /=. + by rewrite iProto_end_message_equivI. } + rewrite !lookup_total_insert. + rewrite lookup_total_insert_ne; [|done]. + rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + iDestruct "Hm1" as (Hieq) "#Hm1". + iDestruct "Hm2" as (Hjeq) "#Hm2". + iIntros (v p1) "Hm1'". + iSpecialize ("Hm1" $!v (Next p1)). + rewrite !iMsg_base_eq. + rewrite !iMsg_exist_eq. + iRewrite -"Hm1" in "Hm1'". + iDestruct "Hm1'" as (x Heq) "[#Hm1' _]". + iSpecialize ("Hm2" $!v (Next (<Send 2> iMsg_base_def x True END)))%proto. + iExists (<Send 2> iMsg_base_def x True END)%proto. + iRewrite -"Hm2". + simpl. + iSplitL. + { iExists x. iSplit; [done|]. iFrame. done. } + iNext. + rewrite iProto_consistent_unfold. + rewrite insert_commute; [|done]. + rewrite insert_insert. + rewrite insert_commute; [|done]. + rewrite insert_insert. + iRewrite -"Hm1'". + iClear (m1 m2 Hieq Hjeq p1 v Heq) "Hm1 Hm2 Hm1'". + iIntros (i j m1 m2) "Hm1 Hm2". + destruct i. + { rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm1" as (Hieq) "#Hm1". } + rewrite lookup_total_insert_ne; [|done]. + destruct i; last first. + { destruct i. + { rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm1" as (Hieq) "#Hm1". } + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_empty=> /=. + by rewrite iProto_end_message_equivI. } + rewrite lookup_total_insert. + destruct j. + { rewrite lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm2" as (Hjeq) "#Hm2". } + rewrite lookup_total_insert_ne; [|done]. + destruct j. + { rewrite lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm2" as (Hjeq) "#Hm2". } + rewrite lookup_total_insert_ne; [|done]. + destruct j; last first. + { rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_empty. + by rewrite iProto_end_message_equivI. } + rewrite lookup_total_insert. + rewrite !iProto_message_equivI. + iDestruct "Hm1" as (Hieq) "#Hm1". + iDestruct "Hm2" as (Hjeq) "#Hm2". + iIntros (v p1) "Hm1'". + iSpecialize ("Hm1" $!v (Next p1)). + iRewrite -"Hm1" in "Hm1'". + iDestruct "Hm1'" as (Heq) "[#Hm1' _]". + iSpecialize ("Hm2" $!v (Next (<Send 0> iMsg_base_def x True END)))%proto. + iExists (<Send 0> iMsg_base_def x True END)%proto. + iRewrite -"Hm2". + simpl. + iSplitL. + { iExists x. iSplit; [done|]. iFrame. done. } + iNext. + rewrite iProto_consistent_unfold. + rewrite (insert_commute _ 1 2); [|done]. + rewrite (insert_commute _ 1 0); [|done]. + rewrite insert_insert. + rewrite (insert_commute _ 2 0); [|done]. + rewrite (insert_commute _ 2 1); [|done]. + rewrite insert_insert. + iRewrite -"Hm1'". + iClear (m1 m2 Hieq Hjeq p1 v Heq) "Hm1 Hm2 Hm1'". + iIntros (i j m1 m2) "Hm1 Hm2". + destruct i. + { rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm1" as (Hieq) "#Hm1". } + rewrite lookup_total_insert_ne; [|done]. + destruct i. + { rewrite lookup_total_insert. + by rewrite iProto_end_message_equivI. } + rewrite lookup_total_insert_ne; [|done]. + destruct i; last first. + { rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_empty. + by rewrite iProto_end_message_equivI. } + rewrite lookup_total_insert. + destruct j; last first. + { rewrite lookup_total_insert_ne; [|done]. + destruct j. + { rewrite lookup_total_insert. + by rewrite iProto_end_message_equivI. } + rewrite lookup_total_insert_ne; [|done]. + destruct j. + { rewrite lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm2" as (Hjeq) "#Hm2". } + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_empty. + by rewrite iProto_end_message_equivI. } + rewrite lookup_total_insert. + rewrite !iProto_message_equivI. + iDestruct "Hm1" as (Hieq) "#Hm1". + iDestruct "Hm2" as (Hjeq) "#Hm2". + iIntros (v p1) "Hm1'". + iSpecialize ("Hm1" $!v (Next p1)). + iRewrite -"Hm1" in "Hm1'". + iDestruct "Hm1'" as (Heq) "[#Hm1' _]". + iSpecialize ("Hm2" $!v (Next END))%proto. + iExists END%proto. + iRewrite -"Hm2". + simpl. + iSplitL; [done|]. + rewrite insert_insert. + rewrite insert_commute; [|done]. + rewrite (insert_commute _ 2 1); [|done]. + rewrite insert_insert. + iNext. + iRewrite -"Hm1'". + rewrite iProto_consistent_unfold. + iClear (m1 m2 Hieq Hjeq p1 v Heq) "Hm1 Hm2 Hm1'". + iIntros (i j m1 m2) "Hm1 Hm2". + destruct i. + { rewrite lookup_total_insert. + by rewrite !iProto_end_message_equivI. } + rewrite lookup_total_insert_ne; [|done]. + destruct i. + { rewrite lookup_total_insert. + by rewrite !iProto_end_message_equivI. } + rewrite lookup_total_insert_ne; [|done]. + destruct i. + { rewrite lookup_total_insert. + by rewrite !iProto_end_message_equivI. } + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_empty. + by rewrite iProto_end_message_equivI. +Qed. -- GitLab From 6a4c7570c444aa329c720069c26ae67118f81e98 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Mon, 22 Jan 2024 14:29:41 +0100 Subject: [PATCH 09/81] Proof of concept for synchronous resource transfer --- theories/channel/channel.v | 487 +++++++++++++++++-------------------- theories/channel/proto.v | 8 +- 2 files changed, 224 insertions(+), 271 deletions(-) diff --git a/theories/channel/channel.v b/theories/channel/channel.v index e8d038e..1dbb8d8 100644 --- a/theories/channel/channel.v +++ b/theories/channel/channel.v @@ -21,88 +21,85 @@ In this file we define the three message-passing connectives: It is additionaly shown that the channel ownership [c ↣ prot] is closed under the subprotocol relation [⊑] *) +From iris.algebra Require Import gmap excl_auth gmap_view. From iris.heap_lang Require Export primitive_laws notation. From iris.heap_lang Require Export proofmode. From iris.heap_lang.lib Require Import spin_lock. +From actris.utils Require Export llist. +From actris.channel Require Import proto_model. From actris.channel Require Export proto. -From actris.utils Require Import llist skip. Set Default Proof Using "Type". -Local Existing Instance spin_lock. - (** * The definition of the message-passing connectives *) Definition new_chan : val := - λ: <>, - let: "l" := lnil #() in - let: "r" := lnil #() in - let: "lk" := newlock #() in - ((("l","r"),"lk"), (("r","l"),"lk")). + λ: <>, let: "l1" := ref NONEV in + let: "l2" := ref NONEV in + (("l1","l2"), ("l2","l1")). Definition fork_chan : val := λ: "f", let: "cc" := new_chan #() in Fork ("f" (Snd "cc"));; Fst "cc". +Definition wait : val := + rec: "go" "l" := + match: !"l" with + NONE => #() + | SOME <> => "go" "l" + end. + Definition send : val := λ: "c" "v", - let: "lk" := Snd "c" in - let: "l" := Fst (Fst "c") in - let: "r" := Snd (Fst "c") in - acquire "lk";; - lsnoc "l" "v";; - release "lk". + let: "l" := Fst "c" in + "l" <- SOME "v";; wait "l". -Definition try_recv : val := - λ: "c", - let: "lk" := Snd "c" in - acquire "lk";; - let: "l" := Snd (Fst "c") in - let: "ret" := if: lisnil "l" then NONE else SOME (lpop "l") in - release "lk";; "ret". +(* Definition recv : val := *) +(* rec: "go" "c" "i" := *) +(* let: "l" := Snd (llookup "c" "i") in *) +(* match: !"l" with *) +(* NONE => "go" "c" *) +(* | SOME "v" => "c" <- NONE;; "v" *) +(* end. *) Definition recv : val := rec: "go" "c" := - match: try_recv "c" with - SOME "p" => "p" - | NONE => "go" "c" + let: "l" := Snd "c" in + let: "v" := Xchg "l" NONEV in + match: "v" with + NONE => "go" "c" + | SOME "v" => "v" end. -(** * Setup of Iris's cameras *) Class chanG Σ := { - chanG_lockG :: lockG Σ; + chanG_tokG :: inG Σ (exclR unitO); chanG_protoG :: protoG Σ val; }. -Definition chanΣ : gFunctors := #[ spin_lockΣ; protoΣ val ]. +Definition chanΣ : gFunctors := #[ protoΣ val; GFunctor (exclR unitO) ]. Global Instance subG_chanΣ {Σ} : subG chanΣ Σ → chanG Σ. Proof. solve_inG. Qed. -Record chan_name := ChanName { - chan_lock_name : gname; - chan_proto_name : proto_name; -}. -Global Instance chan_name_inhabited : Inhabited chan_name := - populate (ChanName inhabitant inhabitant). -Global Instance chan_name_eq_dec : EqDecision chan_name. -Proof. solve_decision. Qed. -Global Instance chan_name_countable : Countable chan_name. -Proof. - refine (inj_countable (λ '(ChanName γl γr), (γl,γr)) - (λ '(γl, γr), Some (ChanName γl γr)) _); by intros []. -Qed. - (** * Definition of the pointsto connective *) Notation iProto Σ := (iProto Σ val). Notation iMsg Σ := (iMsg Σ val). +Definition tok `{!chanG Σ} (γ : gname) : iProp Σ := own γ (Excl ()). + +Definition chan_inv `{!heapGS Σ, !chanG Σ} γ γE γt s (l:loc) : iProp Σ := + (l ↦ NONEV ∗ tok γt) ∨ + (∃ v m, l ↦ SOMEV v ∗ + iProto_own γ s (<!> m)%proto ∗ + (∃ p, iMsg_car m v (Next p) ∗ own γE (â—E (Next p)))) ∨ + (∃ p, l ↦ NONEV ∗ + iProto_own γ s p ∗ own γE (â—E (Next p))). + Definition iProto_pointsto_def `{!heapGS Σ, !chanG Σ} (c : val) (p : iProto Σ) : iProp Σ := - ∃ γ s (l r : loc) (lk : val), - ⌜ c = ((#(side_elim s l r), #(side_elim s r l)), lk)%V ⌠∗ - is_lock (chan_lock_name γ) lk (∃ vsl vsr, - llist internal_eq l vsl ∗ - llist internal_eq r vsr ∗ - steps_lb (length vsl) ∗ steps_lb (length vsr) ∗ - iProto_ctx (chan_proto_name γ) vsl vsr) ∗ - iProto_own (chan_proto_name γ) s p. + ∃ γ γE1 γE2 γt1 γt2 s (l1 l2:loc), + ⌜ c = PairV #l1 #l2 ⌠∗ + inv (nroot.@"ctx") (iProto_ctx γ) ∗ + inv (nroot.@"p") (chan_inv γ γE1 γt1 s l1) ∗ + inv (nroot.@"p") (chan_inv γ γE2 γt2 (side_dual s) l2) ∗ + own γE1 (â—E (Next p)) ∗ own γE1 (â—¯E (Next p)) ∗ + iProto_own γ s p. Definition iProto_pointsto_aux : seal (@iProto_pointsto_def). by eexists. Qed. Definition iProto_pointsto := iProto_pointsto_aux.(unseal). Definition iProto_pointsto_eq : @@ -112,25 +109,6 @@ Global Instance: Params (@iProto_pointsto) 4 := {}. Notation "c ↣ p" := (iProto_pointsto c p) (at level 20, format "c ↣ p"). -Global Instance iProto_pointsto_contractive `{!heapGS Σ, !chanG Σ} c : - Contractive (iProto_pointsto c). -Proof. rewrite iProto_pointsto_eq. solve_contractive. Qed. - -Definition iProto_choice {Σ} (a : action) (P1 P2 : iProp Σ) - (p1 p2 : iProto Σ) : iProto Σ := - (<a @ (b : bool)> MSG #b {{ if b then P1 else P2 }}; if b then p1 else p2)%proto. -Global Typeclasses Opaque iProto_choice. -Arguments iProto_choice {_} _ _%I _%I _%proto _%proto. -Global Instance: Params (@iProto_choice) 2 := {}. -Infix "<{ P1 }+{ P2 }>" := (iProto_choice Send P1 P2) (at level 85) : proto_scope. -Infix "<{ P1 }&{ P2 }>" := (iProto_choice Recv P1 P2) (at level 85) : proto_scope. -Infix "<+{ P2 }>" := (iProto_choice Send True P2) (at level 85) : proto_scope. -Infix "<&{ P2 }>" := (iProto_choice Recv True P2) (at level 85) : proto_scope. -Infix "<{ P1 }+>" := (iProto_choice Send P1 True) (at level 85) : proto_scope. -Infix "<{ P1 }&>" := (iProto_choice Recv P1 True) (at level 85) : proto_scope. -Infix "<+>" := (iProto_choice Send True True) (at level 85) : proto_scope. -Infix "<&>" := (iProto_choice Recv True True) (at level 85) : proto_scope. - Section channel. Context `{!heapGS Σ, !chanG Σ}. Implicit Types p : iProto Σ. @@ -141,64 +119,12 @@ Section channel. Global Instance iProto_pointsto_proper c : Proper ((≡) ==> (≡)) (iProto_pointsto c). Proof. apply (ne_proper _). Qed. - Lemma iProto_pointsto_le c p1 p2 : c ↣ p1 ⊢ â–· (p1 ⊑ p2) -∗ c ↣ p2. - Proof. - rewrite iProto_pointsto_eq. iDestruct 1 as (γ s l r lk ->) "[Hlk H]". - iIntros "Hle'". iExists γ, s, l, r, lk. iSplit; [done|]. iFrame "Hlk". - by iApply (iProto_own_le with "H"). - Qed. - - Global Instance iProto_choice_contractive n a : - Proper (dist n ==> dist n ==> - dist_later n ==> dist_later n ==> dist n) (@iProto_choice Σ a). - Proof. solve_contractive. Qed. - Global Instance iProto_choice_ne n a : - Proper (dist n ==> dist n ==> dist n ==> dist n ==> dist n) (@iProto_choice Σ a). - Proof. solve_proper. Qed. - Global Instance iProto_choice_proper a : - Proper ((≡) ==> (≡) ==> (≡) ==> (≡) ==> (≡)) (@iProto_choice Σ a). - Proof. solve_proper. Qed. - - Lemma iProto_choice_equiv a1 a2 (P11 P12 P21 P22 : iProp Σ) - (p11 p12 p21 p22 : iProto Σ) : - ⌜a1 = a2⌠-∗ ((P11 ≡ P12):iProp Σ) -∗ (P21 ≡ P22) -∗ - â–· (p11 ≡ p12) -∗ â–· (p21 ≡ p22) -∗ - iProto_choice a1 P11 P21 p11 p21 ≡ iProto_choice a2 P12 P22 p12 p22. - Proof. - iIntros (->) "#HP1 #HP2 #Hp1 #Hp2". - rewrite /iProto_choice. iApply iProto_message_equiv; [ eauto | | ]. - - iIntros "!>" (b) "H". iExists b. iSplit; [ done | ]. - destruct b; - [ iRewrite -"HP1"; iFrame "H Hp1" | iRewrite -"HP2"; iFrame "H Hp2" ]. - - iIntros "!>" (b) "H". iExists b. iSplit; [ done | ]. - destruct b; - [ iRewrite "HP1"; iFrame "H Hp1" | iRewrite "HP2"; iFrame "H Hp2" ]. - Qed. - - Lemma iProto_dual_choice a P1 P2 p1 p2 : - iProto_dual (iProto_choice a P1 P2 p1 p2) - ≡ iProto_choice (action_dual a) P1 P2 (iProto_dual p1) (iProto_dual p2). - Proof. - rewrite /iProto_choice iProto_dual_message /= iMsg_dual_exist. - f_equiv; f_equiv=> -[]; by rewrite iMsg_dual_base. - Qed. - - Lemma iProto_app_choice a P1 P2 p1 p2 q : - (iProto_choice a P1 P2 p1 p2 <++> q)%proto - ≡ (iProto_choice a P1 P2 (p1 <++> q) (p2 <++> q))%proto. - Proof. - rewrite /iProto_choice iProto_app_message /= iMsg_app_exist. - f_equiv; f_equiv=> -[]; by rewrite iMsg_app_base. - Qed. - - Lemma iProto_le_choice a P1 P2 p1 p2 p1' p2' : - (P1 -∗ P1 ∗ â–· (p1 ⊑ p1')) ∧ (P2 -∗ P2 ∗ â–· (p2 ⊑ p2')) -∗ - iProto_choice a P1 P2 p1 p2 ⊑ iProto_choice a P1 P2 p1' p2'. - Proof. - iIntros "H". rewrite /iProto_choice. destruct a; - iIntros (b) "HP"; iExists b; destruct b; - iDestruct ("H" with "HP") as "[$ ?]"; by iModIntro. - Qed. + (* Lemma iProto_pointsto_le c p1 p2 : c ↣ p1 ⊢ â–· (p1 ⊑ p2) -∗ c ↣ p2. *) + (* Proof. *) + (* rewrite iProto_pointsto_eq. iDestruct 1 as (γ s l r lk ->) "[Hlk H]". *) + (* iIntros "Hle'". iExists γ, s, l, r, lk. iSplit; [done|]. iFrame "Hlk". *) + (* by iApply (iProto_own_le with "H"). *) + (* Qed. *) (** ** Specifications of [send] and [recv] *) Lemma new_chan_spec p : @@ -206,19 +132,36 @@ Section channel. new_chan #() {{{ c1 c2, RET (c1,c2); c1 ↣ p ∗ c2 ↣ iProto_dual p }}}. Proof. - iIntros (Φ _) "HΦ". wp_lam. iMod (steps_lb_0) as "#Hlb". - wp_smart_apply (lnil_spec internal_eq with "[//]"); iIntros (l) "Hl". - wp_smart_apply (lnil_spec internal_eq with "[//]"); iIntros (r) "Hr". + iIntros (Φ _) "HΦ". wp_lam. + wp_alloc l1 as "Hl1". + wp_alloc l2 as "Hl2". iMod (iProto_init p) as (γp) "(Hctx & Hcl & Hcr)". - wp_smart_apply (newlock_spec (∃ vsl vsr, - llist internal_eq l vsl ∗ llist internal_eq r vsr ∗ - steps_lb (length vsl) ∗ steps_lb (length vsr) ∗ - iProto_ctx γp vsl vsr) with "[Hl Hr Hctx]"). - { iExists [], []. iFrame "#∗". } - iIntros (lk γlk) "#Hlk". wp_pures. iApply "HΦ". - set (γ := ChanName γlk γp). iSplitL "Hcl". - - rewrite iProto_pointsto_eq. iExists γ, Left, l, r, lk. by iFrame "Hcl #". - - rewrite iProto_pointsto_eq. iExists γ, Right, l, r, lk. by iFrame "Hcr #". + iMod (own_alloc (â—E (Next p) â‹… â—¯E (Next p))) as (γl) "[Hâ—l Hâ—¯l]". + { by apply excl_auth_valid. } + iMod (own_alloc (â—E (Next (iProto_dual p)) â‹… â—¯E (Next (iProto_dual p)))) as (γr) "[Hâ—r Hâ—¯r]". + { by apply excl_auth_valid. } + iMod (own_alloc (Excl ())) as (γtl) "Htokl". + { done. } + iMod (own_alloc (Excl ())) as (γtr) "Htokr". + { done. } + wp_pures. + iMod (inv_alloc _ ⊤ (iProto_ctx γp) with "[Hctx]") + as "#IH". + { done. } + iMod (inv_alloc _ ⊤ (chan_inv γp γl γtl Left l1) with "[Hl1 Htokl]") + as "#IHl". + { iLeft. iFrame. } + iMod (inv_alloc _ ⊤ (chan_inv γp γr γtr Right l2) with "[Hl2 Htokr]") + as "#IHr". + { iLeft. iFrame. } + iModIntro. + iApply "HΦ". + iSplitL "Hcl Hâ—l Hâ—¯l". + - rewrite iProto_pointsto_eq. + iExists _, _, _, _, _, _, _, _. eauto with iFrame. + - rewrite iProto_pointsto_eq. + iExists _, _, _, _, _, _, _, _. + iSplit; [done|]. iFrame "#∗". Qed. Lemma fork_chan_spec p Φ (f : val) : @@ -233,149 +176,159 @@ Section channel. wp_pures. iApply ("HΦ" with "Hc1"). Qed. + Lemma own_prot_excl γ (p1 p2 : iProto Σ) : + own γ (â—¯E (Next p1)) -∗ own γ (â—¯E (Next p2)) -∗ False. + Proof. Admitted. + Lemma send_spec c v p : {{{ c ↣ <!> MSG v; p }}} send c v {{{ RET #(); c ↣ p }}}. Proof. rewrite iProto_pointsto_eq. iIntros (Φ) "Hc HΦ". wp_lam; wp_pures. - iDestruct "Hc" as (γ s l r lk ->) "[#Hlk H]"; wp_pures. - wp_smart_apply (acquire_spec with "Hlk"); iIntros "[Hlkd Hinv]". - iDestruct "Hinv" as (vsl vsr) "(Hl & Hr & #Hlbl & #Hlbr & Hctx)". - destruct s; simpl. - - wp_pures. wp_bind (lsnoc _ _). - iApply (wp_step_fupdN_lb with "Hlbr [Hctx H]"); [done| |]. - { iApply fupd_mask_intro; [set_solver|]. simpl. - iIntros "Hclose !>!>". - iMod (iProto_send_l with "Hctx H []") as "[Hctx H]". - { rewrite iMsg_base_eq /=; auto. } - iModIntro. - iApply step_fupdN_intro; [done|]. - iIntros "!>". iMod "Hclose". - iCombine ("Hctx H") as "H". - iExact "H". } - iApply (wp_lb_update with "Hlbl"). - wp_smart_apply (lsnoc_spec with "[$Hl //]"); iIntros "Hl". - iIntros "#Hlbl' [Hctx H] !>". - wp_smart_apply (release_spec with "[Hl Hr Hctx $Hlk $Hlkd]"). - { iExists (vsl ++ [v]), vsr. - rewrite app_length /=. - replace (length vsl + 1) with (S (length vsl)) by lia. - iFrame "#∗". } - iIntros "_". iApply "HΦ". iExists γ, Left, l, r, lk. eauto 10 with iFrame. - - wp_pures. wp_bind (lsnoc _ _). - iApply (wp_step_fupdN_lb with "Hlbl [Hctx H]"); [done| |]. - { iApply fupd_mask_intro; [set_solver|]. simpl. - iIntros "Hclose !>!>". - iMod (iProto_send_r with "Hctx H []") as "[Hctx H]". - { rewrite iMsg_base_eq /=; auto. } - iModIntro. - iApply step_fupdN_intro; [done|]. - iIntros "!>". iMod "Hclose". - iCombine ("Hctx H") as "H". - iExact "H". } - iApply (wp_lb_update with "Hlbr"). - wp_smart_apply (lsnoc_spec with "[$Hr //]"); iIntros "Hr". - iIntros "#Hlbr' [Hctx H] !>". - wp_smart_apply (release_spec with "[Hl Hr Hctx $Hlk $Hlkd]"). - { iExists vsl, (vsr ++ [v]). - rewrite app_length /=. - replace (length vsr + 1) with (S (length vsr)) by lia. - iFrame "#∗". } - iIntros "_". iApply "HΦ". iExists γ, Right, l, r, lk. eauto 10 with iFrame. - Qed. - - Lemma send_spec_tele {TT} c (tt : TT) - (v : TT → val) (P : TT → iProp Σ) (p : TT → iProto Σ) : - {{{ c ↣ (<!.. x > MSG v x {{ P x }}; p x) ∗ P tt }}} - send c (v tt) - {{{ RET #(); c ↣ (p tt) }}}. - Proof. - iIntros (Φ) "[Hc HP] HΦ". - iDestruct (iProto_pointsto_le _ _ (<!> MSG v tt; p tt)%proto with "Hc [HP]") - as "Hc". - { iIntros "!>". - iApply iProto_le_trans. - iApply iProto_le_texist_intro_l. - by iFrame "HP". } - by iApply (send_spec with "Hc"). + iDestruct "Hc" as (γ γE1 γE2 γt1 γt2 s l1 l2 ->) + "(#IH & #IHl & #IHr & Hâ— & Hâ—¯ & Hown)". + wp_pures. + wp_bind (Store _ _). + iInv "IHl" as "HIp". + iDestruct "HIp" as "[HIp|HIp]"; last first. + { iDestruct "HIp" as "[HIp|HIp]". + - iDestruct "HIp" as (? m) "(>Hl & Hown' & HIp)". + wp_store. + rewrite /iProto_own. + iDestruct "Hown" as (p') "[_ Hown]". + iDestruct "Hown'" as (p'') "[_ Hown']". + iDestruct (own_prot_excl with "Hown Hown'") as "H". done. + - iDestruct "HIp" as (p') "(>Hl & Hown' & HIp)". + wp_store. + rewrite /iProto_own. + iDestruct "Hown" as (p'') "[_ Hown]". + iDestruct "Hown'" as (p''') "[_ Hown']". + iDestruct (own_prot_excl with "Hown Hown'") as "H". done. } + iDestruct "HIp" as "[>Hl Htok]". + wp_store. + iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]". + { apply excl_auth_update. } + iModIntro. + iSplitL "Hl Hâ— Hown". + { iRight. iLeft. iIntros "!>". iExists _, _. iFrame. + iExists _. iFrame. rewrite iMsg_base_eq. simpl. done. } + wp_pures. + iLöb as "HL". + wp_lam. + wp_bind (Load _). + iInv "IHl" as "HIp". + iDestruct "HIp" as "[HIp|HIp]". + { iDestruct "HIp" as ">[Hl Htok']". + iDestruct (own_valid_2 with "Htok Htok'") as %H. done. } + iDestruct "HIp" as "[HIp|HIp]". + - iDestruct "HIp" as (? m) "(>Hl & Hown & HIp)". + wp_load. iModIntro. + iSplitL "Hl Hown HIp". + { iRight. iLeft. iExists _, _. iFrame. } + wp_pures. iApply ("HL" with "HΦ Htok Hâ—¯"). + - iDestruct "HIp" as (p') "(>Hl & Hown & Hâ—)". + wp_load. + iModIntro. + iSplitL "Hl Htok". + { iLeft. iFrame. } + iDestruct (own_valid_2 with "Hâ— Hâ—¯") as "#Hagree". + iDestruct (excl_auth_agreeI with "Hagree") as "Hagree'". + wp_pures. + iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]". + { apply excl_auth_update. } + iModIntro. + iApply "HΦ". + iExists _, _, _, _, _, _, _, _. + iSplit; [done|]. iFrame "#∗". + iRewrite -"Hagree'". done. Qed. - Lemma try_recv_spec {TT} c (v : TT → val) (P : TT → iProp Σ) (p : TT → iProto Σ) : - {{{ c ↣ <?.. x> MSG v x {{ P x }}; p x }}} - try_recv c - {{{ w, RET w; (⌜w = NONEV⌠∗ c ↣ <?.. x> MSG v x {{ P x }}; p x) ∨ - (∃.. x, ⌜w = SOMEV (v x)⌠∗ c ↣ p x ∗ P x) }}}. - Proof. - rewrite iProto_pointsto_eq. iIntros (Φ) "Hc HΦ". wp_lam; wp_pures. - iDestruct "Hc" as (γ s l r lk ->) "[#Hlk H]"; wp_pures. - wp_smart_apply (acquire_spec with "Hlk"); iIntros "[Hlkd Hinv]". - iDestruct "Hinv" as (vsl vsr) "(Hl & Hr & #Hlbl & #Hlbr & Hctx)". destruct s; simpl. - - wp_smart_apply (lisnil_spec with "Hr"); iIntros "Hr". - destruct vsr as [|vr vsr]; wp_pures. - { wp_smart_apply (release_spec with "[Hl Hr Hctx $Hlk $Hlkd]"); [by eauto with iFrame|]. - iIntros "_". wp_pures. iModIntro. iApply "HΦ". iLeft. iSplit; [done|]. - iExists γ, Left, l, r, lk. eauto 10 with iFrame. } - wp_smart_apply (lpop_spec with "Hr"); iIntros (v') "[% Hr]"; simplify_eq/=. - iMod (iProto_recv_l with "Hctx H") as (q) "(Hctx & H & Hm)". wp_pures. - rewrite iMsg_base_eq. - iDestruct (iMsg_texist_exist with "Hm") as (x <-) "[Hp HP]". - iDestruct (steps_lb_le _ (length vsr) with "Hlbr") as "#Hlbr'"; [lia|]. - wp_smart_apply (release_spec with "[Hl Hr Hctx $Hlk $Hlkd]"); [by eauto with iFrame|]. - iIntros "_". wp_pures. iModIntro. iApply "HΦ". iRight. iExists x. iSplit; [done|]. - iFrame "HP". iExists γ, Left, l, r, lk. iSplit; [done|]. iFrame "Hlk". - by iRewrite "Hp". - - wp_smart_apply (lisnil_spec with "Hl"); iIntros "Hl". - destruct vsl as [|vl vsl]; wp_pures. - { wp_smart_apply (release_spec with "[Hl Hr Hctx $Hlk $Hlkd]"); [by eauto with iFrame|]. - iIntros "_". wp_pures. iModIntro. iApply "HΦ". iLeft. iSplit; [done|]. - iExists γ, Right, l, r, lk. eauto 10 with iFrame. } - wp_smart_apply (lpop_spec with "Hl"); iIntros (v') "[% Hl]"; simplify_eq/=. - iMod (iProto_recv_r with "Hctx H") as (q) "(Hctx & H & Hm)". wp_pures. - rewrite iMsg_base_eq. - iDestruct (iMsg_texist_exist with "Hm") as (x <-) "[Hp HP]". - iDestruct (steps_lb_le _ (length vsl) with "Hlbl") as "#Hlbl'"; [lia|]. - wp_smart_apply (release_spec with "[Hl Hr Hctx $Hlk $Hlkd]"); [by eauto with iFrame|]. - iIntros "_". wp_pures. iModIntro. iApply "HΦ". iRight. iExists x. iSplit; [done|]. - iFrame "HP". iExists γ, Right, l, r, lk. iSplit; [done|]. iFrame "Hlk". - by iRewrite "Hp". - Qed. + (* Lemma send_spec_tele {TT} c (tt : TT) *) + (* (v : TT → val) (P : TT → iProp Σ) (p : TT → iProto Σ) : *) + (* {{{ c ↣ (<!.. x > MSG v x {{ P x }}; p x) ∗ P tt }}} *) + (* send c (v tt) *) + (* {{{ RET #(); c ↣ (p tt) }}}. *) + (* Proof. *) + (* iIntros (Φ) "[Hc HP] HΦ". *) + (* iDestruct (iProto_pointsto_le _ _ (<!> MSG v tt; p tt)%proto with "Hc [HP]") *) + (* as "Hc". *) + (* { iIntros "!>". *) + (* iApply iProto_le_trans. *) + (* iApply iProto_le_texist_intro_l. *) + (* by iFrame "HP". } *) + (* by iApply (send_spec with "Hc"). *) + (* Qed. *) Lemma recv_spec {TT} c (v : TT → val) (P : TT → iProp Σ) (p : TT → iProto Σ) : {{{ c ↣ <?.. x> MSG v x {{ â–· P x }}; p x }}} recv c {{{ x, RET v x; c ↣ p x ∗ P x }}}. Proof. - iIntros (Φ) "Hc HΦ". iLöb as "IH". wp_lam. - wp_smart_apply (try_recv_spec with "Hc"); iIntros (w) "[[-> H]|H]". - { wp_pures. by iApply ("IH" with "[$]"). } - iDestruct "H" as (x ->) "[Hc HP]". wp_pures. iApply "HΦ". by iFrame. - Qed. - - (** ** Specifications for choice *) - Lemma select_spec c (b : bool) P1 P2 p1 p2 : - {{{ c ↣ (p1 <{P1}+{P2}> p2) ∗ if b then P1 else P2 }}} - send c #b - {{{ RET #(); c ↣ (if b then p1 else p2) }}}. - Proof. - rewrite /iProto_choice. iIntros (Φ) "[Hc HP] HΦ". - iApply (send_spec with "[Hc HP] HΦ"). - iApply (iProto_pointsto_le with "Hc"). - iIntros "!>". iExists b. by iFrame "HP". + iIntros (Φ) "Hc HΦ". iLöb as "HL". wp_lam. + rewrite iProto_pointsto_eq. + iDestruct "Hc" as (γ E1 γE2 γt1 γt2 s l1 l2 ->) + "(#IH & #IHl & #IHr & Hâ— & Hâ—¯ & Hown)". + wp_pures. + wp_bind (Xchg _ _). + iInv "IHr" as "HIp". + iDestruct "HIp" as "[HIp|HIp]". + { iDestruct "HIp" as ">[Hl Htok]". + wp_xchg. + iModIntro. + iSplitL "Hl Htok". + { iLeft. iFrame. } + wp_pures. iApply ("HL" with "[Hâ— Hâ—¯ Hown] HΦ"). + iExists _, _, _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". } + iDestruct "HIp" as "[HIp|HIp]"; last first. + { iDestruct "HIp" as (p') "[>Hl [Hown' Hâ—¯']]". + wp_xchg. + iModIntro. + iSplitL "Hl Hown' Hâ—¯'". + { iRight. iRight. iExists _. iFrame. } + wp_pures. iApply ("HL" with "[Hâ— Hâ—¯ Hown] HΦ"). + iExists _, _, _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". } + iDestruct "HIp" as (w m) "(>Hl & Hown' & HIp)". + iDestruct "HIp" as (p') "[Hm Hp']". + iInv "IH" as "Hctx". + wp_xchg. + destruct s. + - simpl. + iMod (iProto_step_r with "Hctx Hown Hown' Hm") as + (p'') "(Hm & Hctx & Hown & Hown')". + iModIntro. + iSplitL "Hctx"; [done|]. + iModIntro. + iSplitL "Hl Hown' Hp'". + { iRight. iRight. iExists _. iFrame. } + wp_pure _. + rewrite iMsg_base_eq. + iDestruct (iMsg_texist_exist with "Hm") as (x <-) "[Hp HP]". + wp_pures. + iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]". + { apply excl_auth_update. } + iModIntro. iApply "HΦ". + iFrame. + iExists _, _, _, _, _, _, _, _. iSplit; [done|]. + iRewrite "Hp". iFrame "#∗". done. + - simpl. + iMod (iProto_step_l with "Hctx Hown' Hown Hm") as + (p'') "(Hm & Hctx & Hown & Hown')". + iModIntro. + iSplitL "Hctx"; [done|]. + iModIntro. + iSplitL "Hl Hown Hp'". + { iRight. iRight. iExists _. iFrame. } + wp_pure _. + rewrite iMsg_base_eq. + iDestruct (iMsg_texist_exist with "Hm") as (x <-) "[Hp HP]". + wp_pures. + iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]". + { apply excl_auth_update. } + iModIntro. iApply "HΦ". + iFrame. + iExists _, _, _, _, _, _, _, _. iSplit; [done|]. + iRewrite "Hp". iFrame "#∗". done. Qed. - Lemma branch_spec c P1 P2 p1 p2 : - {{{ c ↣ (p1 <{P1}&{P2}> p2) }}} - recv c - {{{ b, RET #b; c ↣ (if b : bool then p1 else p2) ∗ if b then P1 else P2 }}}. - Proof. - rewrite /iProto_choice. iIntros (Φ) "Hc HΦ". - iApply (recv_spec _ (tele_app _) - (tele_app (TT:=[tele _ : bool]) (λ b, if b then P1 else P2))%I - (tele_app _) with "[Hc]"). - { iApply (iProto_pointsto_le with "Hc"). - iIntros "!> /=" (b) "HP". iExists b. by iSplitL. } - rewrite -bi_tforall_forall. - iIntros "!>" (x) "[Hc H]". iApply "HΦ". iFrame. - Qed. End channel. diff --git a/theories/channel/proto.v b/theories/channel/proto.v index c12d547..f7b54fe 100644 --- a/theories/channel/proto.v +++ b/theories/channel/proto.v @@ -1206,8 +1206,8 @@ Section proto. Lemma iProto_step_l γ m1 m2 p1 v : iProto_ctx γ -∗ iProto_own γ Left (<!> m1) -∗ iProto_own γ Right (<?> m2) -∗ iMsg_car m1 v (Next p1) ==∗ - â–· ∃ p2, iMsg_car m2 v (Next p2) ∗ â–· iProto_ctx γ ∗ - iProto_own γ Left p1 ∗ iProto_own γ Right p2. + â–· ∃ p2, iMsg_car m2 v (Next p2) ∗ iProto_ctx γ ∗ + iProto_own γ Left p1 ∗ iProto_own γ Right p2. Proof. iDestruct 1 as (pl pr) "(Hâ—l & Hâ—r & Hconsistent)". iDestruct 1 as (pl') "[Hlel Hâ—¯l]". @@ -1237,8 +1237,8 @@ Section proto. Lemma iProto_step_r γ m1 m2 p2 v : iProto_ctx γ -∗ iProto_own γ Left (<?> m1) -∗ iProto_own γ Right (<!> m2) -∗ iMsg_car m2 v (Next p2) ==∗ - â–· ∃ p1, iMsg_car m1 v (Next p1) ∗ â–· iProto_ctx γ ∗ - iProto_own γ Left p1 ∗ iProto_own γ Right p2. + â–· ∃ p1, iMsg_car m1 v (Next p1) ∗ iProto_ctx γ ∗ + iProto_own γ Left p1 ∗ iProto_own γ Right p2. Proof. iDestruct 1 as (pl pr) "(Hâ—l & Hâ—r & Hconsistent)". iDestruct 1 as (pl') "[Hlel Hâ—¯l]". -- GitLab From cb563064db46beee2102087acc420202e954950b Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Mon, 22 Jan 2024 18:35:57 +0100 Subject: [PATCH 10/81] WIP multi_channel --- _CoqProject | 3 ++- theories/channel/multi_proto.v | 4 ++-- 2 files changed, 4 insertions(+), 3 deletions(-) diff --git a/_CoqProject b/_CoqProject index 330aec8..2fcbc7e 100644 --- a/_CoqProject +++ b/_CoqProject @@ -16,9 +16,10 @@ theories/utils/cofe_solver_2.v theories/utils/switch.v theories/channel/proto_model.v theories/channel/proto.v +theories/channel/channel.v theories/channel/multi_proto_model.v theories/channel/multi_proto.v -theories/channel/channel.v +theories/channel/multi_channel.v theories/channel/proofmode.v theories/examples/basics.v theories/examples/equivalence.v diff --git a/theories/channel/multi_proto.v b/theories/channel/multi_proto.v index 984b84c..05b31f7 100644 --- a/theories/channel/multi_proto.v +++ b/theories/channel/multi_proto.v @@ -707,7 +707,7 @@ Section proto. iProto_own γ i (<(Send j)> m1) -∗ iProto_own γ j (<(Recv i)> m2) -∗ iMsg_car m1 v (Next p1) ==∗ - â–· ∃ p2, iMsg_car m2 v (Next p2) ∗ â–· iProto_ctx γ ∗ + â–· ∃ p2, iMsg_car m2 v (Next p2) ∗ iProto_ctx γ ∗ iProto_own γ i p1 ∗ iProto_own γ j p2. Proof. iIntros "Hctx Hi Hj Hm". @@ -718,7 +718,7 @@ Section proto. (p2) "[Hm2 Hconsistent]". iMod (iProto_own_auth_update _ _ _ _ p2 with "Hauth Hj") as "[Hauth Hj]". iMod (iProto_own_auth_update _ _ _ _ p1 with "Hauth Hi") as "[Hauth Hi]". - iIntros "!>!>". iExists p2. iFrame. iIntros "!>". iExists _. iFrame. + iIntros "!>!>". iExists p2. iFrame. iExists _. iFrame. Qed. (* (** The instances below make it possible to use the tactics [iIntros], *) -- GitLab From e7a4245076ba9cf7896e2fa8fc7d4fe402ed94fa Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Mon, 22 Jan 2024 19:30:03 +0100 Subject: [PATCH 11/81] Added multi channels and more --- theories/channel/multi_channel.v | 327 ++++++++++++++++++ .../multi_proto_consistency_examples.v | 97 +++++- 2 files changed, 407 insertions(+), 17 deletions(-) create mode 100644 theories/channel/multi_channel.v diff --git a/theories/channel/multi_channel.v b/theories/channel/multi_channel.v new file mode 100644 index 0000000..878772e --- /dev/null +++ b/theories/channel/multi_channel.v @@ -0,0 +1,327 @@ +(** This file contains the definition of the channels, encoded as a pair of +lock-protected buffers, and their primitive proof rules. Moreover: + +- It defines the connective [c ↣ prot] for ownership of channel endpoints, + which describes that channel endpoint [c] adheres to protocol [prot]. +- It proves Actris's specifications of [send] and [recv] w.r.t. dependent + separation protocols. + +An encoding of the usual (binary) choice connectives [prot1 <{Q1}+{Q2}> prot2] +and [prot1 <{Q1}&{Q2}> prot2], inspired by session types, is also included in +this file. + +In this file we define the three message-passing connectives: + +- [new_chan] creates references to two empty buffers and a lock, and returns a + pair of endpoints, where the order of the two references determines the + polarity of the endpoints. +- [send] takes an endpoint and adds an element to the first buffer. +- [recv] performs a busy loop until there is something in the second buffer, + which it pops and returns, locking during each peek. + +It is additionaly shown that the channel ownership [c ↣ prot] is closed under +the subprotocol relation [⊑] *) +From iris.algebra Require Import gmap excl_auth gmap_view. +From iris.heap_lang Require Export primitive_laws notation. +From iris.heap_lang Require Export proofmode. +From iris.heap_lang.lib Require Import spin_lock. +From actris.utils Require Export llist. +From actris.channel Require Import multi_proto_model. +From actris.channel Require Export multi_proto. +Set Default Proof Using "Type". + +(* TODO: Update new_chan definition to use pointers with offsets *) +(** * The definition of the message-passing connectives *) +Definition new_chan : val := + λ: "n", + let: "l" := AllocN ("n"*"n") NONEV in + let: "xxs" := lnil #() in + (rec: "go1" "i" := if: "i" = "n" then #() else + let: "xs" := lnil #() in + (rec: "go2" "j" := if: "j" = "n" then #() else + lcons ("l" +â‚— ("i"*"n"+"j"), "l" +â‚— ("j"*"n"+"i")) "xs";; + "go2" ("j"+#1)) #0;; + lcons "xs" "xxs";; + "go1" ("i"+#1)) #0;; "xxs". + +Definition wait : val := + rec: "go" "c" := + match: !"c" with + NONE => #() + | SOME <> => "go" "c" + end. + +Definition send : val := + λ: "c" "i" "v", + let: "len" := Fst "c" in + if: "i" < "len" then + let: "l" := Fst (! ((Snd "c") +â‚— "i")) in + "l" <- SOME "v";; wait "l" + (* OBS: Hacky *) + else (rec: "go" <> := "go" #())%V #(). + +Definition recv : val := + rec: "go" "c" "i" := + let: "len" := Fst "c" in + if: "i" < "len" then + let: "l" := Snd (! ((Snd "c") +â‚— "i")) in + let: "v" := Xchg "l" NONEV in + match: "v" with + NONE => "go" "c" "i" + | SOME "v" => "v" + end + (* OBS: Hacky *) + else (rec: "go" <> := "go" #())%V #(). + +(** * Setup of Iris's cameras *) +Class proto_exclG Σ V := + protoG_exclG :: + inG Σ (excl_authR (laterO (proto (leibnizO V) (iPropO Σ) (iPropO Σ)))). + +Definition proto_exclΣ V := #[ + GFunctor (authRF (optionURF (exclRF (laterOF (protoOF (leibnizO V) idOF idOF))))) +]. + +Class chanG Σ := { + chanG_proto_exclG :: proto_exclG Σ val; + chanG_tokG :: inG Σ (exclR unitO); + chanG_protoG :: protoG Σ val; +}. +Definition chanΣ : gFunctors := #[ proto_exclΣ val; protoΣ val; GFunctor (exclR unitO) ]. +Global Instance subG_chanΣ {Σ} : subG chanΣ Σ → chanG Σ. +Proof. solve_inG. Qed. + +(** * Definition of the pointsto connective *) +Notation iProto Σ := (iProto Σ val). +Notation iMsg Σ := (iMsg Σ val). + +Definition tok `{!chanG Σ} (γ : gname) : iProp Σ := own γ (Excl ()). + +Definition chan_inv `{!heapGS Σ, !chanG Σ} γ γE γt i j (l:loc) : iProp Σ := + (l ↦ NONEV ∗ tok γt) ∨ + (∃ v m, l ↦ SOMEV v ∗ + iProto_own γ i (<Send j> m)%proto ∗ + (∃ p, iMsg_car m v (Next p) ∗ own γE (â—E (Next p)))) ∨ + (∃ p, l ↦ NONEV ∗ + iProto_own γ i p ∗ own γE (â—E (Next p))). + +Definition iProto_pointsto_def `{!heapGS Σ, !chanG Σ} + (c : val) (i:nat) (p : iProto Σ) : iProp Σ := + ∃ γ γE1 γE2 γt1 γt2 (l:loc) ls, + ⌜ c = PairV #(length ls) #l ⌠∗ + inv (nroot.@"ctx") (iProto_ctx γ) ∗ + l ↦∗ ls ∗ + ([∗list] j ↦ v ∈ ls, + ∃ (l1 l2 : loc), + ⌜v = PairV #l1 #l2⌠∗ + inv (nroot.@"p") (chan_inv γ γE1 γt1 i j l1) ∗ + inv (nroot.@"p") (chan_inv γ γE2 γt2 j i l2)) ∗ + (* llist (λ l v, *) + (* ∃ (j:nat) (l1 l2 : loc), *) + (* ⌜l = (j,(l1,l2))⌠∗ ⌜v = PairV #l1 #l2⌠∗ *) + (* inv (nroot.@"p") (chan_inv γ γE1 γt1 i j l1) ∗ *) + (* inv (nroot.@"p") (chan_inv γ γE2 γt2 j i l2)) l *) + (* (zip (seq 0 (length ls)) ls) ∗ *) + own γE1 (â—E (Next p)) ∗ own γE1 (â—¯E (Next p)) ∗ + iProto_own γ i p. +Definition iProto_pointsto_aux : seal (@iProto_pointsto_def). by eexists. Qed. +Definition iProto_pointsto := iProto_pointsto_aux.(unseal). +Definition iProto_pointsto_eq : + @iProto_pointsto = @iProto_pointsto_def := iProto_pointsto_aux.(seal_eq). +Arguments iProto_pointsto {_ _ _} _ _ _%proto. +Global Instance: Params (@iProto_pointsto) 4 := {}. +Notation "c ↣{ i } p" := (iProto_pointsto c i p) + (at level 20, format "c ↣{ i } p"). + +Section channel. + Context `{!heapGS Σ, !chanG Σ}. + Implicit Types p : iProto Σ. + Implicit Types TT : tele. + + Global Instance iProto_pointsto_ne c i : NonExpansive (iProto_pointsto c i). + Proof. rewrite iProto_pointsto_eq. solve_proper. Qed. + Global Instance iProto_pointsto_proper c i : Proper ((≡) ==> (≡)) (iProto_pointsto c i). + Proof. apply (ne_proper _). Qed. + + (* Lemma iProto_pointsto_le c p1 p2 : c ↣ p1 ⊢ â–· (p1 ⊑ p2) -∗ c ↣ p2. *) + (* Proof. *) + (* rewrite iProto_pointsto_eq. iDestruct 1 as (γ s l r lk ->) "[Hlk H]". *) + (* iIntros "Hle'". iExists γ, s, l, r, lk. iSplit; [done|]. iFrame "Hlk". *) + (* by iApply (iProto_own_le with "H"). *) + (* Qed. *) + + (** ** Specifications of [send] and [recv] *) + Lemma new_chan_spec (n:nat) (ps:gmap nat (iProto Σ)) : + (∀ i, i < n → is_Some (ps !! i)) → + {{{ iProto_consistent ps }}} + new_chan #n + {{{ cs ls, RET #cs; + ⌜length ls = n⌠∗ cs ↦∗ ls ∗ + [∗list] i ↦ l ∈ ls, l ↣{i} (ps !!! i) }}}. + Proof. Admitted. + + Lemma own_prot_excl γ i (p1 p2 : iProto Σ) : + own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p1))) -∗ + own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p2))) -∗ + False. + Proof. Admitted. + + Lemma send_spec c i j v p : + {{{ c ↣{i} <Send j> MSG v; p }}} + send c #j v + {{{ RET #(); c ↣{i} p }}}. + Proof. + rewrite iProto_pointsto_eq. iIntros (Φ) "Hc HΦ". wp_lam; wp_pures. + iDestruct "Hc" as + (γ γE1 γE2 γt1 γt2 l ls ->) "(#IH & Hl & Hls & Hâ— & Hâ—¯ & Hown)". + wp_pures. + case_bool_decide; last first. + { wp_pures. iClear "IH Hâ— Hâ—¯ Hown HΦ Hls Hl". + iLöb as "IH". wp_lam. iApply "IH". } + (* assert (is_Some ((zip (seq 0 (length ls)) ls) !! j)) as [l' HSome]. *) + (* { apply lookup_lt_is_Some_2. lia. } *) + (* wp_smart_apply (llookup_spec with "Hls"); [done|]. *) + assert (is_Some (ls !! j)) as [l' HSome]. + { apply lookup_lt_is_Some_2. lia. } + wp_pures. + wp_smart_apply (wp_load_offset with "Hl"). + { done. } + iIntros "Hl". wp_pures. + iDestruct (big_sepL_lookup_acc with "Hls") as "[Hj Hls]"; [done|]. + iDestruct "Hj" as (l1 l2 ->) "#[IHl1 IHl2]". + iDestruct ("Hls" with "[]") as "Hls". + { iExists _, _. iFrame "#". done. } + wp_pures. + wp_bind (Store _ _). + iInv "IHl1" as "HIp". + iDestruct "HIp" as "[HIp|HIp]"; last first. + { iDestruct "HIp" as "[HIp|HIp]". + - iDestruct "HIp" as (? m) "(>Hl' & Hown' & HIp)". + wp_store. + rewrite /iProto_own. + iDestruct (own_prot_excl with "Hown Hown'") as "H". done. + - iDestruct "HIp" as (p') "(>Hl' & Hown' & HIp)". + wp_store. + rewrite /iProto_own. + iDestruct (own_prot_excl with "Hown Hown'") as "H". done. } + iDestruct "HIp" as "[>Hl' Htok]". + wp_store. + iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]". + { apply excl_auth_update. } + iModIntro. + iSplitL "Hl' Hâ— Hown". + { iRight. iLeft. iIntros "!>". iExists _, _. iFrame. + iExists _. iFrame. rewrite iMsg_base_eq. simpl. done. } + wp_pures. + iLöb as "HL". + wp_lam. + wp_bind (Load _). + iInv "IHl1" as "HIp". + iDestruct "HIp" as "[HIp|HIp]". + { iDestruct "HIp" as ">[Hl' Htok']". + iDestruct (own_valid_2 with "Htok Htok'") as %[]. } + iDestruct "HIp" as "[HIp|HIp]". + - iDestruct "HIp" as (? m) "(>Hl' & Hown & HIp)". + wp_load. iModIntro. + iSplitL "Hl' Hown HIp". + { iRight. iLeft. iExists _, _. iFrame. } + wp_pures. iApply ("HL" with "HΦ Hl Hls Htok Hâ—¯"). + - iDestruct "HIp" as (p') "(>Hl' & Hown & Hâ—)". + wp_load. + iModIntro. + iSplitL "Hl' Htok". + { iLeft. iFrame. } + iDestruct (own_valid_2 with "Hâ— Hâ—¯") as "#Hagree". + iDestruct (excl_auth_agreeI with "Hagree") as "Hagree'". + wp_pures. + iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]". + { apply excl_auth_update. } + iModIntro. + iApply "HΦ". + iExists _, _, _, _, _, _, _. + iSplit; [done|]. iFrame "#∗". + iRewrite -"Hagree'". done. + Qed. + + (* Lemma send_spec_tele {TT} c (tt : TT) *) + (* (v : TT → val) (P : TT → iProp Σ) (p : TT → iProto Σ) : *) + (* {{{ c ↣ (<!.. x > MSG v x {{ P x }}; p x) ∗ P tt }}} *) + (* send c (v tt) *) + (* {{{ RET #(); c ↣ (p tt) }}}. *) + (* Proof. *) + (* iIntros (Φ) "[Hc HP] HΦ". *) + (* iDestruct (iProto_pointsto_le _ _ (<!> MSG v tt; p tt)%proto with "Hc [HP]") *) + (* as "Hc". *) + (* { iIntros "!>". *) + (* iApply iProto_le_trans. *) + (* iApply iProto_le_texist_intro_l. *) + (* by iFrame "HP". } *) + (* by iApply (send_spec with "Hc"). *) + (* Qed. *) + + Lemma recv_spec {TT} c i j (v : TT → val) (P : TT → iProp Σ) (p : TT → iProto Σ) : + {{{ c ↣{i} <(Recv j) @.. x> MSG v x {{ â–· P x }}; p x }}} + recv c #j + {{{ x, RET v x; c ↣{i} p x ∗ P x }}}. + Proof. + iIntros (Φ) "Hc HΦ". iLöb as "HL". wp_lam. + rewrite iProto_pointsto_eq. + iDestruct "Hc" as + (γ γE1 γE2 γt1 γt2 l ls ->) "(#IH & Hl & Hls & Hâ— & Hâ—¯ & Hown)". + wp_pures. + case_bool_decide; last first. + { wp_pures. iClear "IH Hâ— Hâ—¯ Hown HΦ Hls Hl". + iLöb as "IH". wp_lam. iApply "IH". } + wp_pures. + assert (is_Some (ls !! j)) as [l' HSome]. + { apply lookup_lt_is_Some_2. lia. } + wp_smart_apply (wp_load_offset with "Hl"). + { done. } + iIntros "Hl". wp_pures. + iDestruct (big_sepL_lookup_acc with "Hls") as "[Hj Hls]"; [done|]. + iDestruct "Hj" as (l1 l2 ->) "#[IHl1 IHl2]". + iDestruct ("Hls" with "[]") as "Hls". + { iExists _, _. iFrame "#". done. } + wp_pures. + wp_bind (Xchg _ _). + iInv "IHl2" as "HIp". + iDestruct "HIp" as "[HIp|HIp]". + { iDestruct "HIp" as ">[Hl' Htok]". + wp_xchg. + iModIntro. + iSplitL "Hl' Htok". + { iLeft. iFrame. } + wp_pures. iApply ("HL" with "[Hâ— Hâ—¯ Hown Hls Hl] HΦ"). + iExists _, _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". } + iDestruct "HIp" as "[HIp|HIp]"; last first. + { iDestruct "HIp" as (p') "[>Hl' [Hown' Hâ—¯']]". + wp_xchg. + iModIntro. + iSplitL "Hl' Hown' Hâ—¯'". + { iRight. iRight. iExists _. iFrame. } + wp_pures. iApply ("HL" with "[Hâ— Hâ—¯ Hown Hls Hl] HΦ"). + iExists _, _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". } + iDestruct "HIp" as (w m) "(>Hl' & Hown' & HIp)". + iDestruct "HIp" as (p') "[Hm Hp']". + iInv "IH" as "Hctx". + wp_xchg. + iMod (iProto_step with "Hctx Hown' Hown Hm") as + (p'') "(Hm & Hctx & Hown & Hown')". + iModIntro. + iSplitL "Hctx"; [done|]. + iModIntro. + iSplitL "Hl' Hown Hp'". + { iRight. iRight. iExists _. iFrame. } + wp_pure _. + rewrite iMsg_base_eq. + iDestruct (iMsg_texist_exist with "Hm") as (x <-) "[Hp HP]". + wp_pures. + iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]". + { apply excl_auth_update. } + iModIntro. iApply "HΦ". + iFrame. + iExists _, _, _, _, _, _, _. iSplit; [done|]. + iRewrite "Hp". iFrame "#∗". + Qed. + +End channel. diff --git a/theories/channel/multi_proto_consistency_examples.v b/theories/channel/multi_proto_consistency_examples.v index 4863f4a..345c5c8 100644 --- a/theories/channel/multi_proto_consistency_examples.v +++ b/theories/channel/multi_proto_consistency_examples.v @@ -3,15 +3,15 @@ From iris.proofmode Require Import proofmode. From iris.base_logic Require Export lib.iprop. From iris.base_logic Require Import lib.own. From iris.program_logic Require Import language. -From actris.channel Require Import multi_proto_model multi_proto. +From actris.channel Require Import multi_proto_model multi_proto multi_channel. Set Default Proof Using "Type". Export action. -Definition iProto_example1 {Σ V} : gmap nat (iProto Σ V) := +Definition iProto_example1 {Σ} : gmap nat (iProto Σ) := ∅. -Lemma iProto_example1_consistent {Σ V} : - ⊢ iProto_consistent (@iProto_example1 Σ V). +Lemma iProto_example1_consistent {Σ} : + ⊢ iProto_consistent (@iProto_example1 Σ). Proof. rewrite iProto_consistent_unfold. iIntros (i j m1 m2) "Hi Hj". @@ -19,13 +19,13 @@ Proof. by rewrite iProto_end_message_equivI. Qed. -Definition iProto_example2 `{!invGS Σ} (P : iProp Σ) : gmap nat (iProto Σ Z) := - <[0 := (<(Send 1) @ (x:Z)> MSG x {{ P }} ; END)%proto ]> - (<[1 := (<(Recv 0) @ (x:Z)> MSG x {{ P }} ; END)%proto ]> +Definition iProto_example2 `{!invGS Σ} (P : iProp Σ) : gmap nat (iProto Σ) := + <[0 := (<(Send 1) @ (x:Z)> MSG #x {{ P }} ; END)%proto ]> + (<[1 := (<(Recv 0) @ (x:Z)> MSG #x {{ P }} ; END)%proto ]> ∅). Lemma iProto_example2_consistent `{!invGS Σ} (P : iProp Σ) : - ⊢ iProto_consistent (@iProto_example2 _ Σ invGS0 P). + ⊢ iProto_consistent (@iProto_example2 Σ invGS0 P). Proof. rewrite iProto_consistent_unfold. rewrite /iProto_example2. @@ -102,14 +102,14 @@ Proof. done. Qed. -Definition iProto_example3 `{!invGS Σ} : gmap nat (iProto Σ Z) := - <[0 := (<(Send 1) @ (x:Z)> MSG x ; <(Recv 2)> MSG x; END)%proto ]> - (<[1 := (<(Recv 0) @ (x:Z)> MSG x ; <(Send 2)> MSG x; END)%proto ]> - (<[2 := (<(Recv 1) @ (x:Z)> MSG x ; <(Send 0)> MSG x; END)%proto ]> +Definition iProto_example3 `{!invGS Σ} : gmap nat (iProto Σ) := + <[0 := (<(Send 1) @ (x:Z)> MSG #x ; <(Recv 2)> MSG #x; END)%proto ]> + (<[1 := (<(Recv 0) @ (x:Z)> MSG #x ; <(Send 2)> MSG #x; END)%proto ]> + (<[2 := (<(Recv 1) @ (x:Z)> MSG #x ; <(Send 0)> MSG #x; END)%proto ]> ∅)). Lemma iProto_example3_consistent `{!invGS Σ} : - ⊢ iProto_consistent (@iProto_example3 _ Σ invGS0). + ⊢ iProto_consistent (@iProto_example3 Σ invGS0). Proof. rewrite iProto_consistent_unfold. rewrite /iProto_example3. @@ -157,8 +157,8 @@ Proof. rewrite !iMsg_exist_eq. iRewrite -"Hm1" in "Hm1'". iDestruct "Hm1'" as (x Heq) "[#Hm1' _]". - iSpecialize ("Hm2" $!v (Next (<Send 2> iMsg_base_def x True END)))%proto. - iExists (<Send 2> iMsg_base_def x True END)%proto. + iSpecialize ("Hm2" $!v (Next (<Send 2> iMsg_base_def #x True END)))%proto. + iExists (<Send 2> iMsg_base_def #x True END)%proto. iRewrite -"Hm2". simpl. iSplitL. @@ -210,8 +210,8 @@ Proof. iSpecialize ("Hm1" $!v (Next p1)). iRewrite -"Hm1" in "Hm1'". iDestruct "Hm1'" as (Heq) "[#Hm1' _]". - iSpecialize ("Hm2" $!v (Next (<Send 0> iMsg_base_def x True END)))%proto. - iExists (<Send 0> iMsg_base_def x True END)%proto. + iSpecialize ("Hm2" $!v (Next (<Send 0> iMsg_base_def #x True END)))%proto. + iExists (<Send 0> iMsg_base_def #x True END)%proto. iRewrite -"Hm2". simpl. iSplitL. @@ -291,3 +291,66 @@ Proof. rewrite lookup_total_empty. by rewrite iProto_end_message_equivI. Qed. + +Definition roundtrip_prog : val := + λ: <>, + let: "cs" := new_chan #3 in + let: "c0" := ! ("cs" +â‚— #0) in + let: "c1" := ! ("cs" +â‚— #1) in + let: "c2" := ! ("cs" +â‚— #2) in + Fork (let: "x" := recv "c1" #0 in send "c1" #2 "x");; + Fork (let: "x" := recv "c2" #1 in send "c2" #0 "x");; + send "c0" #1 #42;; + recv "c0" #2. + +Section channel. + Context `{!heapGS Σ, !chanG Σ}. + Implicit Types p : iProto Σ. + Implicit Types TT : tele. + + Lemma roundtrip_prog_spec : + {{{ True }}} roundtrip_prog #() {{{ RET #42 ; True }}}. + Proof. + iIntros (Φ) "_ HΦ". wp_lam. + wp_pures. + wp_apply (new_chan_spec 3 iProto_example3 with "[]"). + { intros i Hle. destruct i as [|[|[]]]; try set_solver. lia. } + { iApply iProto_example3_consistent. } + iIntros (cs ls) "[%Hlen [Hcs Hls]]". + assert (is_Some (ls !! 0)) as [c0 HSome0]. + { apply lookup_lt_is_Some_2. lia. } + assert (is_Some (ls !! 1)) as [c1 HSome1]. + { apply lookup_lt_is_Some_2. lia. } + assert (is_Some (ls !! 2)) as [c2 HSome2]. + { apply lookup_lt_is_Some_2. lia. } + wp_smart_apply (wp_load_offset _ _ _ _ 0 with "Hcs"); [done|]. + iIntros "Hcs". + wp_smart_apply (wp_load_offset _ _ _ _ 1 with "Hcs"); [done|]. + iIntros "Hcs". + wp_smart_apply (wp_load_offset _ _ _ _ 2 with "Hcs"); [done|]. + iIntros "Hcs". + iDestruct (big_sepL_delete' _ _ 0 with "Hls") as "[Hc0 Hls]"; [set_solver|]. + iDestruct (big_sepL_delete' _ _ 1 with "Hls") as "[Hc1 Hls]"; [set_solver|]. + iDestruct (big_sepL_delete' _ _ 2 with "Hls") as "[Hc2 _]"; [set_solver|]. + iDestruct ("Hc1" with "[//]") as "Hc1". + iDestruct ("Hc2" with "[//] [//]") as "Hc2". + rewrite /iProto_example3. + rewrite !lookup_total_insert. + rewrite lookup_total_insert_ne; [|done]. + rewrite !lookup_total_insert. + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + rewrite !lookup_total_insert. + wp_smart_apply (wp_fork with "[Hc1]"). + { iIntros "!>". wp_pures. + wp_apply (recv_spec c1 _ 0 with "[Hc1]"); [admit|]. + admit. } + wp_smart_apply (wp_fork with "[Hc2]"). + { iIntros "!>". wp_pures. + wp_apply (recv_spec c2 _ 1 with "[Hc2]"); [admit|]. + admit. } + wp_pures. wp_apply (send_spec c0 0 1 #42 with "[Hc0]"); [admit|]. + iIntros "Hc0". + Admitted. + +End channel. -- GitLab From 4229311f43e326535ea4e583020de66248fe9b60 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Mon, 22 Jan 2024 21:32:34 +0100 Subject: [PATCH 12/81] First closed proof! --- theories/channel/multi_channel.v | 99 +++++++++++++++---- .../multi_proto_consistency_examples.v | 42 ++++++-- 2 files changed, 113 insertions(+), 28 deletions(-) diff --git a/theories/channel/multi_channel.v b/theories/channel/multi_channel.v index 878772e..763567f 100644 --- a/theories/channel/multi_channel.v +++ b/theories/channel/multi_channel.v @@ -178,9 +178,6 @@ Section channel. case_bool_decide; last first. { wp_pures. iClear "IH Hâ— Hâ—¯ Hown HΦ Hls Hl". iLöb as "IH". wp_lam. iApply "IH". } - (* assert (is_Some ((zip (seq 0 (length ls)) ls) !! j)) as [l' HSome]. *) - (* { apply lookup_lt_is_Some_2. lia. } *) - (* wp_smart_apply (llookup_spec with "Hls"); [done|]. *) assert (is_Some (ls !! j)) as [l' HSome]. { apply lookup_lt_is_Some_2. lia. } wp_pures. @@ -243,24 +240,88 @@ Section channel. iRewrite -"Hagree'". done. Qed. - (* Lemma send_spec_tele {TT} c (tt : TT) *) - (* (v : TT → val) (P : TT → iProp Σ) (p : TT → iProto Σ) : *) - (* {{{ c ↣ (<!.. x > MSG v x {{ P x }}; p x) ∗ P tt }}} *) - (* send c (v tt) *) - (* {{{ RET #(); c ↣ (p tt) }}}. *) - (* Proof. *) - (* iIntros (Φ) "[Hc HP] HΦ". *) - (* iDestruct (iProto_pointsto_le _ _ (<!> MSG v tt; p tt)%proto with "Hc [HP]") *) - (* as "Hc". *) - (* { iIntros "!>". *) - (* iApply iProto_le_trans. *) - (* iApply iProto_le_texist_intro_l. *) - (* by iFrame "HP". } *) - (* by iApply (send_spec with "Hc"). *) - (* Qed. *) + Lemma send_spec_tele {TT} c j i (tt : TT) + (v : TT → val) (P : TT → iProp Σ) (p : TT → iProto Σ) : + {{{ c ↣{i} (<(Send j) @.. x > MSG v x {{ P x }}; p x) ∗ P tt }}} + send c #j (v tt) + {{{ RET #(); c ↣{i} (p tt) }}}. + Proof. + rewrite iProto_pointsto_eq. iIntros (Φ) "[Hc HP] HΦ". wp_lam; wp_pures. + iDestruct "Hc" as + (γ γE1 γE2 γt1 γt2 l ls ->) "(#IH & Hl & Hls & Hâ— & Hâ—¯ & Hown)". + wp_pures. + case_bool_decide; last first. + { wp_pures. iClear "IH Hâ— Hâ—¯ Hown HΦ Hls Hl". + iLöb as "IH". wp_lam. iApply "IH". iFrame. } + assert (is_Some (ls !! j)) as [l' HSome]. + { apply lookup_lt_is_Some_2. lia. } + wp_pures. + wp_smart_apply (wp_load_offset with "Hl"). + { done. } + iIntros "Hl". wp_pures. + iDestruct (big_sepL_lookup_acc with "Hls") as "[Hj Hls]"; [done|]. + iDestruct "Hj" as (l1 l2 ->) "#[IHl1 IHl2]". + iDestruct ("Hls" with "[]") as "Hls". + { iExists _, _. iFrame "#". done. } + wp_pures. + wp_bind (Store _ _). + iInv "IHl1" as "HIp". + iDestruct "HIp" as "[HIp|HIp]"; last first. + { iDestruct "HIp" as "[HIp|HIp]". + - iDestruct "HIp" as (? m) "(>Hl' & Hown' & HIp)". + wp_store. + rewrite /iProto_own. + iDestruct (own_prot_excl with "Hown Hown'") as "H". done. + - iDestruct "HIp" as (p') "(>Hl' & Hown' & HIp)". + wp_store. + rewrite /iProto_own. + iDestruct (own_prot_excl with "Hown Hown'") as "H". done. } + iDestruct "HIp" as "[>Hl' Htok]". + wp_store. + iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]". + { apply excl_auth_update. } + iModIntro. + iSplitL "Hl' Hâ— Hown HP". + { iRight. iLeft. iIntros "!>". iExists _, _. iFrame. + iExists _. iFrame. rewrite iMsg_base_eq. simpl. + iApply iMsg_texist_exist. + simpl. iExists tt. + iSplit; [done|]. + iSplit; [done|]. + done. } + wp_pures. + iLöb as "HL". + wp_lam. + wp_bind (Load _). + iInv "IHl1" as "HIp". + iDestruct "HIp" as "[HIp|HIp]". + { iDestruct "HIp" as ">[Hl' Htok']". + iDestruct (own_valid_2 with "Htok Htok'") as %[]. } + iDestruct "HIp" as "[HIp|HIp]". + - iDestruct "HIp" as (? m) "(>Hl' & Hown & HIp)". + wp_load. iModIntro. + iSplitL "Hl' Hown HIp". + { iRight. iLeft. iExists _, _. iFrame. } + wp_pures. iApply ("HL" with "HΦ Hl Hls Htok Hâ—¯"). + - iDestruct "HIp" as (p') "(>Hl' & Hown & Hâ—)". + wp_load. + iModIntro. + iSplitL "Hl' Htok". + { iLeft. iFrame. } + iDestruct (own_valid_2 with "Hâ— Hâ—¯") as "#Hagree". + iDestruct (excl_auth_agreeI with "Hagree") as "Hagree'". + wp_pures. + iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]". + { apply excl_auth_update. } + iModIntro. + iApply "HΦ". + iExists _, _, _, _, _, _, _. + iSplit; [done|]. iFrame "#∗". + iRewrite -"Hagree'". done. + Qed. Lemma recv_spec {TT} c i j (v : TT → val) (P : TT → iProp Σ) (p : TT → iProto Σ) : - {{{ c ↣{i} <(Recv j) @.. x> MSG v x {{ â–· P x }}; p x }}} + {{{ c ↣{i} <(Recv j) @.. x> MSG v x {{ P x }}; p x }}} recv c #j {{{ x, RET v x; c ↣{i} p x ∗ P x }}}. Proof. diff --git a/theories/channel/multi_proto_consistency_examples.v b/theories/channel/multi_proto_consistency_examples.v index 345c5c8..e934156 100644 --- a/theories/channel/multi_proto_consistency_examples.v +++ b/theories/channel/multi_proto_consistency_examples.v @@ -310,7 +310,7 @@ Section channel. Lemma roundtrip_prog_spec : {{{ True }}} roundtrip_prog #() {{{ RET #42 ; True }}}. - Proof. + Proof using chanG0 heapGS0 Σ. iIntros (Φ) "_ HΦ". wp_lam. wp_pures. wp_apply (new_chan_spec 3 iProto_example3 with "[]"). @@ -342,15 +342,39 @@ Section channel. rewrite lookup_total_insert_ne; [|done]. rewrite !lookup_total_insert. wp_smart_apply (wp_fork with "[Hc1]"). - { iIntros "!>". wp_pures. - wp_apply (recv_spec c1 _ 0 with "[Hc1]"); [admit|]. - admit. } + { iIntros "!>". + wp_smart_apply + (recv_spec (TT:=[tele Z]) c1 1 0 + (tele_app (λ (x:Z), #x)) (λ _, True)%I (tele_app (λ (x:Z), _)) + with "Hc1"). + iIntros (x') "[Hc1 _]". + epose proof (tele_arg_S_inv x') as [x [[] ->]]. simpl. + wp_smart_apply (send_spec c1 1 2 with "Hc1"). + by iIntros "_". } wp_smart_apply (wp_fork with "[Hc2]"). - { iIntros "!>". wp_pures. - wp_apply (recv_spec c2 _ 1 with "[Hc2]"); [admit|]. - admit. } - wp_pures. wp_apply (send_spec c0 0 1 #42 with "[Hc0]"); [admit|]. + { iIntros "!>". + wp_smart_apply + (recv_spec (TT:=[tele Z]) c2 2 1 + (tele_app (λ (x:Z), #x)) (λ _, True)%I (tele_app (λ (x:Z), _)) + with "Hc2"). + iIntros (x') "[Hc1 _]". + epose proof (tele_arg_S_inv x') as [x [[] ->]]. simpl. + wp_smart_apply (send_spec c2 2 0 with "Hc1"). + by iIntros "_". } + wp_smart_apply + (send_spec_tele (TT:=[tele Z]) c0 1 0 ([tele_arg 42%Z]) + (tele_app (λ (x:Z), #x)) (λ _, True)%I + (tele_app (λ (x:Z), _)) + with "[Hc0]"). + { iSplitL; [|done]. simpl. iFrame "Hc0". } iIntros "Hc0". - Admitted. + wp_smart_apply (recv_spec (TT:=[tele]) c0 0 2 + (λ _, #42) + (λ _, True)%I + (λ _, _) + with "Hc0"). + iIntros (_) "Hc0". + by iApply "HΦ". + Qed. End channel. -- GitLab From d6df07bf72487fdbb4cdf6eff281a3ff7aebf369 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Mon, 22 Jan 2024 22:07:17 +0100 Subject: [PATCH 13/81] Hid the channel index on the channel ownership --- theories/channel/multi_channel.v | 52 +++++++++---------- .../multi_proto_consistency_examples.v | 24 ++++----- 2 files changed, 35 insertions(+), 41 deletions(-) diff --git a/theories/channel/multi_channel.v b/theories/channel/multi_channel.v index 763567f..04b3286 100644 --- a/theories/channel/multi_channel.v +++ b/theories/channel/multi_channel.v @@ -106,8 +106,8 @@ Definition chan_inv `{!heapGS Σ, !chanG Σ} γ γE γt i j (l:loc) : iProp Σ : iProto_own γ i p ∗ own γE (â—E (Next p))). Definition iProto_pointsto_def `{!heapGS Σ, !chanG Σ} - (c : val) (i:nat) (p : iProto Σ) : iProp Σ := - ∃ γ γE1 γE2 γt1 γt2 (l:loc) ls, + (c : val) (p : iProto Σ) : iProp Σ := + ∃ γ γE1 γE2 γt1 γt2 i (l:loc) ls, ⌜ c = PairV #(length ls) #l ⌠∗ inv (nroot.@"ctx") (iProto_ctx γ) ∗ l ↦∗ ls ∗ @@ -128,19 +128,19 @@ Definition iProto_pointsto_aux : seal (@iProto_pointsto_def). by eexists. Qed. Definition iProto_pointsto := iProto_pointsto_aux.(unseal). Definition iProto_pointsto_eq : @iProto_pointsto = @iProto_pointsto_def := iProto_pointsto_aux.(seal_eq). -Arguments iProto_pointsto {_ _ _} _ _ _%proto. -Global Instance: Params (@iProto_pointsto) 4 := {}. -Notation "c ↣{ i } p" := (iProto_pointsto c i p) - (at level 20, format "c ↣{ i } p"). +Arguments iProto_pointsto {_ _ _} _ _%proto. +Global Instance: Params (@iProto_pointsto) 5 := {}. +Notation "c ↣ p" := (iProto_pointsto c p) + (at level 20, format "c ↣ p"). Section channel. Context `{!heapGS Σ, !chanG Σ}. Implicit Types p : iProto Σ. Implicit Types TT : tele. - Global Instance iProto_pointsto_ne c i : NonExpansive (iProto_pointsto c i). + Global Instance iProto_pointsto_ne c : NonExpansive (iProto_pointsto c). Proof. rewrite iProto_pointsto_eq. solve_proper. Qed. - Global Instance iProto_pointsto_proper c i : Proper ((≡) ==> (≡)) (iProto_pointsto c i). + Global Instance iProto_pointsto_proper c : Proper ((≡) ==> (≡)) (iProto_pointsto c). Proof. apply (ne_proper _). Qed. (* Lemma iProto_pointsto_le c p1 p2 : c ↣ p1 ⊢ â–· (p1 ⊑ p2) -∗ c ↣ p2. *) @@ -157,7 +157,7 @@ Section channel. new_chan #n {{{ cs ls, RET #cs; ⌜length ls = n⌠∗ cs ↦∗ ls ∗ - [∗list] i ↦ l ∈ ls, l ↣{i} (ps !!! i) }}}. + [∗list] i ↦ l ∈ ls, l ↣ (ps !!! i) }}}. Proof. Admitted. Lemma own_prot_excl γ i (p1 p2 : iProto Σ) : @@ -166,14 +166,14 @@ Section channel. False. Proof. Admitted. - Lemma send_spec c i j v p : - {{{ c ↣{i} <Send j> MSG v; p }}} + Lemma send_spec c j v p : + {{{ c ↣ <Send j> MSG v; p }}} send c #j v - {{{ RET #(); c ↣{i} p }}}. + {{{ RET #(); c ↣ p }}}. Proof. rewrite iProto_pointsto_eq. iIntros (Φ) "Hc HΦ". wp_lam; wp_pures. iDestruct "Hc" as - (γ γE1 γE2 γt1 γt2 l ls ->) "(#IH & Hl & Hls & Hâ— & Hâ—¯ & Hown)". + (γ γE1 γE2 γt1 γt2 i l ls ->) "(#IH & Hl & Hls & Hâ— & Hâ—¯ & Hown)". wp_pures. case_bool_decide; last first. { wp_pures. iClear "IH Hâ— Hâ—¯ Hown HΦ Hls Hl". @@ -235,20 +235,20 @@ Section channel. { apply excl_auth_update. } iModIntro. iApply "HΦ". - iExists _, _, _, _, _, _, _. + iExists _, _, _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". iRewrite -"Hagree'". done. Qed. - Lemma send_spec_tele {TT} c j i (tt : TT) + Lemma send_spec_tele {TT} c j (tt : TT) (v : TT → val) (P : TT → iProp Σ) (p : TT → iProto Σ) : - {{{ c ↣{i} (<(Send j) @.. x > MSG v x {{ P x }}; p x) ∗ P tt }}} + {{{ c ↣ (<(Send j) @.. x > MSG v x {{ P x }}; p x) ∗ P tt }}} send c #j (v tt) - {{{ RET #(); c ↣{i} (p tt) }}}. + {{{ RET #(); c ↣ (p tt) }}}. Proof. rewrite iProto_pointsto_eq. iIntros (Φ) "[Hc HP] HΦ". wp_lam; wp_pures. iDestruct "Hc" as - (γ γE1 γE2 γt1 γt2 l ls ->) "(#IH & Hl & Hls & Hâ— & Hâ—¯ & Hown)". + (γ γE1 γE2 γt1 γt2 i l ls ->) "(#IH & Hl & Hls & Hâ— & Hâ—¯ & Hown)". wp_pures. case_bool_decide; last first. { wp_pures. iClear "IH Hâ— Hâ—¯ Hown HΦ Hls Hl". @@ -315,20 +315,20 @@ Section channel. { apply excl_auth_update. } iModIntro. iApply "HΦ". - iExists _, _, _, _, _, _, _. + iExists _, _, _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". iRewrite -"Hagree'". done. Qed. - Lemma recv_spec {TT} c i j (v : TT → val) (P : TT → iProp Σ) (p : TT → iProto Σ) : - {{{ c ↣{i} <(Recv j) @.. x> MSG v x {{ P x }}; p x }}} + Lemma recv_spec {TT} c j (v : TT → val) (P : TT → iProp Σ) (p : TT → iProto Σ) : + {{{ c ↣ <(Recv j) @.. x> MSG v x {{ P x }}; p x }}} recv c #j - {{{ x, RET v x; c ↣{i} p x ∗ P x }}}. + {{{ x, RET v x; c ↣ p x ∗ P x }}}. Proof. iIntros (Φ) "Hc HΦ". iLöb as "HL". wp_lam. rewrite iProto_pointsto_eq. iDestruct "Hc" as - (γ γE1 γE2 γt1 γt2 l ls ->) "(#IH & Hl & Hls & Hâ— & Hâ—¯ & Hown)". + (γ γE1 γE2 γt1 γt2 i l ls ->) "(#IH & Hl & Hls & Hâ— & Hâ—¯ & Hown)". wp_pures. case_bool_decide; last first. { wp_pures. iClear "IH Hâ— Hâ—¯ Hown HΦ Hls Hl". @@ -353,7 +353,7 @@ Section channel. iSplitL "Hl' Htok". { iLeft. iFrame. } wp_pures. iApply ("HL" with "[Hâ— Hâ—¯ Hown Hls Hl] HΦ"). - iExists _, _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". } + iExists _, _, _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". } iDestruct "HIp" as "[HIp|HIp]"; last first. { iDestruct "HIp" as (p') "[>Hl' [Hown' Hâ—¯']]". wp_xchg. @@ -361,7 +361,7 @@ Section channel. iSplitL "Hl' Hown' Hâ—¯'". { iRight. iRight. iExists _. iFrame. } wp_pures. iApply ("HL" with "[Hâ— Hâ—¯ Hown Hls Hl] HΦ"). - iExists _, _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". } + iExists _, _, _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". } iDestruct "HIp" as (w m) "(>Hl' & Hown' & HIp)". iDestruct "HIp" as (p') "[Hm Hp']". iInv "IH" as "Hctx". @@ -381,7 +381,7 @@ Section channel. { apply excl_auth_update. } iModIntro. iApply "HΦ". iFrame. - iExists _, _, _, _, _, _, _. iSplit; [done|]. + iExists _, _, _, _, _, _, _, _. iSplit; [done|]. iRewrite "Hp". iFrame "#∗". Qed. diff --git a/theories/channel/multi_proto_consistency_examples.v b/theories/channel/multi_proto_consistency_examples.v index e934156..2813020 100644 --- a/theories/channel/multi_proto_consistency_examples.v +++ b/theories/channel/multi_proto_consistency_examples.v @@ -300,8 +300,7 @@ Definition roundtrip_prog : val := let: "c2" := ! ("cs" +â‚— #2) in Fork (let: "x" := recv "c1" #0 in send "c1" #2 "x");; Fork (let: "x" := recv "c2" #1 in send "c2" #0 "x");; - send "c0" #1 #42;; - recv "c0" #2. + send "c0" #1 #42;; recv "c0" #2. Section channel. Context `{!heapGS Σ, !chanG Σ}. @@ -312,8 +311,7 @@ Section channel. {{{ True }}} roundtrip_prog #() {{{ RET #42 ; True }}}. Proof using chanG0 heapGS0 Σ. iIntros (Φ) "_ HΦ". wp_lam. - wp_pures. - wp_apply (new_chan_spec 3 iProto_example3 with "[]"). + wp_smart_apply (new_chan_spec 3 iProto_example3 with "[]"). { intros i Hle. destruct i as [|[|[]]]; try set_solver. lia. } { iApply iProto_example3_consistent. } iIntros (cs ls) "[%Hlen [Hcs Hls]]". @@ -344,37 +342,33 @@ Section channel. wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". wp_smart_apply - (recv_spec (TT:=[tele Z]) c1 1 0 + (recv_spec (TT:=[tele Z]) c1 0 (tele_app (λ (x:Z), #x)) (λ _, True)%I (tele_app (λ (x:Z), _)) with "Hc1"). iIntros (x') "[Hc1 _]". epose proof (tele_arg_S_inv x') as [x [[] ->]]. simpl. - wp_smart_apply (send_spec c1 1 2 with "Hc1"). + wp_smart_apply (send_spec c1 2 with "Hc1"). by iIntros "_". } wp_smart_apply (wp_fork with "[Hc2]"). { iIntros "!>". wp_smart_apply - (recv_spec (TT:=[tele Z]) c2 2 1 + (recv_spec (TT:=[tele Z]) c2 1 (tele_app (λ (x:Z), #x)) (λ _, True)%I (tele_app (λ (x:Z), _)) with "Hc2"). iIntros (x') "[Hc1 _]". epose proof (tele_arg_S_inv x') as [x [[] ->]]. simpl. - wp_smart_apply (send_spec c2 2 0 with "Hc1"). + wp_smart_apply (send_spec c2 0 with "Hc1"). by iIntros "_". } wp_smart_apply - (send_spec_tele (TT:=[tele Z]) c0 1 0 ([tele_arg 42%Z]) + (send_spec_tele (TT:=[tele Z]) c0 1 ([tele_arg 42%Z]) (tele_app (λ (x:Z), #x)) (λ _, True)%I (tele_app (λ (x:Z), _)) with "[Hc0]"). { iSplitL; [|done]. simpl. iFrame "Hc0". } iIntros "Hc0". - wp_smart_apply (recv_spec (TT:=[tele]) c0 0 2 - (λ _, #42) - (λ _, True)%I - (λ _, _) + wp_smart_apply (recv_spec (TT:=[tele]) c0 2 (λ _, #42) (λ _, True)%I (λ _, _) with "Hc0"). - iIntros (_) "Hc0". - by iApply "HΦ". + iIntros (_) "Hc0". by iApply "HΦ". Qed. End channel. -- GitLab From b93dcfb8ee1c13775442d0e26f2dc8a0ee22d87c Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Mon, 22 Jan 2024 22:13:43 +0100 Subject: [PATCH 14/81] Updated definition of new_chan --- theories/channel/multi_channel.v | 17 +++++++---------- 1 file changed, 7 insertions(+), 10 deletions(-) diff --git a/theories/channel/multi_channel.v b/theories/channel/multi_channel.v index 04b3286..c78f0e7 100644 --- a/theories/channel/multi_channel.v +++ b/theories/channel/multi_channel.v @@ -22,12 +22,9 @@ In this file we define the three message-passing connectives: It is additionaly shown that the channel ownership [c ↣ prot] is closed under the subprotocol relation [⊑] *) From iris.algebra Require Import gmap excl_auth gmap_view. -From iris.heap_lang Require Export primitive_laws notation. -From iris.heap_lang Require Export proofmode. -From iris.heap_lang.lib Require Import spin_lock. -From actris.utils Require Export llist. -From actris.channel Require Import multi_proto_model. -From actris.channel Require Export multi_proto. +From iris.base_logic.lib Require Import invariants. +From iris.heap_lang Require Export primitive_laws notation proofmode. +From actris.channel Require Import multi_proto_model multi_proto. Set Default Proof Using "Type". (* TODO: Update new_chan definition to use pointers with offsets *) @@ -35,13 +32,13 @@ Set Default Proof Using "Type". Definition new_chan : val := λ: "n", let: "l" := AllocN ("n"*"n") NONEV in - let: "xxs" := lnil #() in + let: "xxs" := AllocN "n" NONEV in (rec: "go1" "i" := if: "i" = "n" then #() else - let: "xs" := lnil #() in + let: "xs" := AllocN "n" NONEV in (rec: "go2" "j" := if: "j" = "n" then #() else - lcons ("l" +â‚— ("i"*"n"+"j"), "l" +â‚— ("j"*"n"+"i")) "xs";; + ("xs" +â‚— "j") <- ("l" +â‚— ("i"*"n"+"j"), "l" +â‚— ("j"*"n"+"i"));; "go2" ("j"+#1)) #0;; - lcons "xs" "xxs";; + ("xxs" +â‚— "i") <- "xs";; "go1" ("i"+#1)) #0;; "xxs". Definition wait : val := -- GitLab From 8d3cbcfbd5f50a2984a81fcfaa4ea59646fffc03 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Tue, 23 Jan 2024 14:55:13 +0100 Subject: [PATCH 15/81] Added example with resource transfer --- .../multi_proto_consistency_examples.v | 301 ++++++++++++++++++ 1 file changed, 301 insertions(+) diff --git a/theories/channel/multi_proto_consistency_examples.v b/theories/channel/multi_proto_consistency_examples.v index 2813020..c3c3743 100644 --- a/theories/channel/multi_proto_consistency_examples.v +++ b/theories/channel/multi_proto_consistency_examples.v @@ -372,3 +372,304 @@ Section channel. Qed. End channel. + +Section example4. + Context `{!heapGS Σ}. + + Definition iProto_example4 : gmap nat (iProto Σ) := + <[0 := (<(Send 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; + <(Recv 2)> MSG #() {{ l ↦ #(x+2) }} ; END)%proto]> + (<[1 := (<(Recv 0) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; + <(Send 2)> MSG #l {{ l ↦ #(x+1) }}; END)%proto]> + (<[2 := (<(Recv 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; + <(Send 0)> MSG #() {{ l ↦ #(x+1) }}; END)%proto]> + ∅)). + + Lemma iProto_example4_consistent : + ⊢ iProto_consistent (iProto_example4). + Proof. + rewrite iProto_consistent_unfold. + rewrite /iProto_example4. + iIntros (i j m1 m2) "Hm1 Hm2". + destruct i; last first. + { destruct i. + { rewrite lookup_total_insert_ne; [|done]. + rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm1" as (Hieq) "#Hm1". } + destruct i. + { rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + rewrite !iProto_message_equivI. + by iDestruct "Hm1" as (Hieq) "#Hm1". } + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_empty=> /=. + by rewrite iProto_end_message_equivI. } + destruct j. + { rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm2" as (Hjeq) "#Hm2". } + destruct j; last first. + { rewrite !lookup_total_insert. + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + destruct j. + { rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm2" as (Hjeq) "#Hm2". } + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_empty=> /=. + by rewrite iProto_end_message_equivI. } + rewrite !lookup_total_insert. + rewrite lookup_total_insert_ne; [|done]. + rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + iDestruct "Hm1" as (Hieq) "#Hm1". + iDestruct "Hm2" as (Hjeq) "#Hm2". + iIntros (v p1) "Hm1'". + iSpecialize ("Hm1" $!v (Next p1)). + rewrite !iMsg_base_eq. + rewrite !iMsg_exist_eq. + iRewrite -"Hm1" in "Hm1'". + iDestruct "Hm1'" as (l x <-) "[#Hm1' Hl]". simpl. + iSpecialize ("Hm2" $!#l (Next (<Send 2> iMsg_base_def #l + (l ↦ #(x+1)) END)))%proto. + Unshelve. 2-4: apply _. (* Why is this needed? *) + iExists (<Send 2> iMsg_base_def #l (l ↦ #(x+1)) END)%proto. + simpl. + iSplitL. + { iRewrite -"Hm2". iExists l, x. iSplit; [done|]. iFrame. done. } + iNext. + rewrite iProto_consistent_unfold. + rewrite insert_commute; [|done]. + rewrite insert_insert. + rewrite insert_commute; [|done]. + rewrite insert_insert. + iRewrite -"Hm1'". + iClear (m1 m2 Hieq Hjeq p1) "Hm1 Hm2 Hm1'". + iIntros (i j m1 m2) "Hm1 Hm2". + destruct i. + { rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm1" as (Hieq) "#Hm1". } + rewrite lookup_total_insert_ne; [|done]. + destruct i; last first. + { destruct i. + { rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm1" as (Hieq) "#Hm1". } + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_empty=> /=. + by rewrite iProto_end_message_equivI. } + rewrite lookup_total_insert. + destruct j. + { rewrite lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm2" as (Hjeq) "#Hm2". } + rewrite lookup_total_insert_ne; [|done]. + destruct j. + { rewrite lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm2" as (Hjeq) "#Hm2". } + rewrite lookup_total_insert_ne; [|done]. + destruct j; last first. + { rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_empty. + by rewrite iProto_end_message_equivI. } + rewrite lookup_total_insert. + rewrite !iProto_message_equivI. + iDestruct "Hm1" as (Hieq) "#Hm1". + iDestruct "Hm2" as (Hjeq) "#Hm2". + iIntros (v p1) "Hm1'". + iSpecialize ("Hm1" $!v (Next p1)). + iRewrite -"Hm1" in "Hm1'". + iDestruct "Hm1'" as (<-) "[#Hm1' Hl]". + simpl. + iSpecialize ("Hm2" $!#l (Next (<Send 0> iMsg_base_def #() (l ↦ #(x+1+1)) END)))%proto. + iExists (<Send 0> iMsg_base_def #() (l ↦ #(x+1+1)) END)%proto. + Unshelve. 2-4: apply _. + iRewrite -"Hm2". + simpl. + iSplitL. + { iExists l, (x+1)%Z. iSplit; [done|]. iFrame. done. } + iNext. + rewrite iProto_consistent_unfold. + rewrite (insert_commute _ 1 2); [|done]. + rewrite (insert_commute _ 1 0); [|done]. + rewrite insert_insert. + rewrite (insert_commute _ 2 0); [|done]. + rewrite (insert_commute _ 2 1); [|done]. + rewrite insert_insert. + iRewrite -"Hm1'". + iClear (m1 m2 Hieq Hjeq p1) "Hm1 Hm2 Hm1'". + iIntros (i j m1 m2) "Hm1 Hm2". + destruct i. + { rewrite !lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm1" as (Hieq) "#Hm1". } + rewrite lookup_total_insert_ne; [|done]. + destruct i. + { rewrite lookup_total_insert. + by rewrite iProto_end_message_equivI. } + rewrite lookup_total_insert_ne; [|done]. + destruct i; last first. + { rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_empty. + by rewrite iProto_end_message_equivI. } + rewrite lookup_total_insert. + destruct j; last first. + { rewrite lookup_total_insert_ne; [|done]. + destruct j. + { rewrite lookup_total_insert. + by rewrite iProto_end_message_equivI. } + rewrite lookup_total_insert_ne; [|done]. + destruct j. + { rewrite lookup_total_insert. + rewrite !iProto_message_equivI. + by iDestruct "Hm2" as (Hjeq) "#Hm2". } + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_empty. + by rewrite iProto_end_message_equivI. } + rewrite lookup_total_insert. + rewrite !iProto_message_equivI. + iDestruct "Hm1" as (Hieq) "#Hm1". + iDestruct "Hm2" as (Hjeq) "#Hm2". + iIntros (v p1) "Hm1'". + iSpecialize ("Hm1" $!v (Next p1)). + iRewrite -"Hm1" in "Hm1'". + iDestruct "Hm1'" as (<-) "[#Hm1' Hl]". + iSpecialize ("Hm2" $!#() (Next END))%proto. + iExists END%proto. + iRewrite -"Hm2". + simpl. + replace (x + 1 + 1)%Z with (x+2)%Z by lia. + iSplitL; [iFrame;done|]. + rewrite insert_insert. + rewrite insert_commute; [|done]. + rewrite (insert_commute _ 2 1); [|done]. + rewrite insert_insert. + iNext. + iRewrite -"Hm1'". + rewrite iProto_consistent_unfold. + iClear (m1 m2 Hieq Hjeq p1) "Hm1 Hm2 Hm1'". + iIntros (i j m1 m2) "Hm1 Hm2". + destruct i. + { rewrite lookup_total_insert. + by rewrite !iProto_end_message_equivI. } + rewrite lookup_total_insert_ne; [|done]. + destruct i. + { rewrite lookup_total_insert. + by rewrite !iProto_end_message_equivI. } + rewrite lookup_total_insert_ne; [|done]. + destruct i. + { rewrite lookup_total_insert. + by rewrite !iProto_end_message_equivI. } + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_empty. + by rewrite iProto_end_message_equivI. + Qed. + +End example4. + +Definition roundtrip_ref_prog : val := + λ: <>, + let: "cs" := new_chan #3 in + let: "c0" := ! ("cs" +â‚— #0) in + let: "c1" := ! ("cs" +â‚— #1) in + let: "c2" := ! ("cs" +â‚— #2) in + Fork (let: "l" := recv "c1" #0 in "l" <- !"l" + #1;; send "c1" #2 "l");; + Fork (let: "l" := recv "c2" #1 in "l" <- !"l" + #1;; send "c2" #0 #());; + let: "l" := ref #40 in send "c0" #1 "l";; recv "c0" #2;; !"l". + +Section proof. + Context `{!heapGS Σ, !chanG Σ}. + Implicit Types p : iProto Σ. + Implicit Types TT : tele. + + Lemma roundtrip_ref_prog_spec : + {{{ True }}} roundtrip_ref_prog #() {{{ RET #42 ; True }}}. + Proof using chanG0. + iIntros (Φ) "_ HΦ". wp_lam. + wp_smart_apply (new_chan_spec 3 iProto_example4 with "[]"). + { intros i Hle. destruct i as [|[|[]]]; try set_solver. lia. } + { iApply iProto_example4_consistent. } + iIntros (cs ls) "[%Hlen [Hcs Hls]]". + assert (is_Some (ls !! 0)) as [c0 HSome0]. + { apply lookup_lt_is_Some_2. lia. } + assert (is_Some (ls !! 1)) as [c1 HSome1]. + { apply lookup_lt_is_Some_2. lia. } + assert (is_Some (ls !! 2)) as [c2 HSome2]. + { apply lookup_lt_is_Some_2. lia. } + wp_smart_apply (wp_load_offset _ _ _ _ 0 with "Hcs"); [done|]. + iIntros "Hcs". + wp_smart_apply (wp_load_offset _ _ _ _ 1 with "Hcs"); [done|]. + iIntros "Hcs". + wp_smart_apply (wp_load_offset _ _ _ _ 2 with "Hcs"); [done|]. + iIntros "Hcs". + iDestruct (big_sepL_delete' _ _ 0 with "Hls") as "[Hc0 Hls]"; [set_solver|]. + iDestruct (big_sepL_delete' _ _ 1 with "Hls") as "[Hc1 Hls]"; [set_solver|]. + iDestruct (big_sepL_delete' _ _ 2 with "Hls") as "[Hc2 _]"; [set_solver|]. + iDestruct ("Hc1" with "[//]") as "Hc1". + iDestruct ("Hc2" with "[//] [//]") as "Hc2". + rewrite /iProto_example4. + rewrite !lookup_total_insert. + rewrite lookup_total_insert_ne; [|done]. + rewrite !lookup_total_insert. + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + rewrite !lookup_total_insert. + wp_smart_apply (wp_fork with "[Hc1]"). + { iIntros "!>". + wp_smart_apply + (recv_spec (TT:=[tele loc Z]) c1 0 + (tele_app (λ (l : loc) (x:Z), #l)) + (tele_app (λ (l : loc) (x:Z), l ↦ #x)%I) + (tele_app (λ (l : loc) (x:Z), _)) + with "Hc1"). + iIntros (x') "[Hc1 Hl]". + epose proof (tele_arg_S_inv x') as [l [y' ->]]. simpl. + epose proof (tele_arg_S_inv y') as [x [[] ->]]. simpl. + wp_load. wp_store. + wp_smart_apply (send_spec_tele (TT:=[tele]) c1 2 + ([tele_arg]) + (λ _, #l) + (λ _, l ↦ #(x+1))%I + (λ _, _) with "[$Hc1 $Hl]"). + by iIntros "_". } + wp_smart_apply (wp_fork with "[Hc2]"). + { iIntros "!>". + wp_smart_apply + (recv_spec (TT:=[tele loc Z]) c2 1 + (tele_app (λ (l : loc) (x:Z), #l)) + (tele_app (λ (l : loc) (x:Z), l ↦ #x)%I) + (tele_app (λ (l : loc) (x:Z), _)) + with "Hc2"). + iIntros (x') "[Hc2 Hl]". + epose proof (tele_arg_S_inv x') as [l [y' ->]]. simpl. + epose proof (tele_arg_S_inv y') as [x [[] ->]]. simpl. + wp_load. wp_store. + wp_smart_apply (send_spec_tele (TT:=[tele]) c2 0 + ([tele_arg]) + (λ _, #()) + (λ _, l ↦ #(x+1))%I + (λ _, _) with "[$Hc2 $Hl]"). + by iIntros "_". } + wp_alloc l as "Hl". + wp_smart_apply + (send_spec_tele (TT:=[tele l Z]) c0 1 ([tele_arg l ; 40%Z]) + (tele_app (λ (l:loc) (x:Z), #l)) + (tele_app (λ (l:loc) (x:Z), l ↦ #x)%I) + (tele_app (λ (l:loc) (x:Z), _)) + with "[$Hc0 $Hl]"). + iIntros "Hc0". + wp_smart_apply (recv_spec (TT:=[tele]) c0 2 + (λ _, #()) (λ _, l ↦ #(40 + 2))%I (λ _, _) + with "Hc0"). + iIntros (_) "[Hc0 Hl]". wp_load. by iApply "HΦ". + Qed. + +End proof. -- GitLab From ad01352aa1385f0915ba5be614b93df0c29dfb73 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Tue, 23 Jan 2024 16:01:39 +0100 Subject: [PATCH 16/81] Proved subprotocol lemmas --- theories/channel/multi_channel.v | 6 -- theories/channel/multi_proto.v | 117 ++++++++++++++++++++++++++++++- 2 files changed, 114 insertions(+), 9 deletions(-) diff --git a/theories/channel/multi_channel.v b/theories/channel/multi_channel.v index c78f0e7..157bf8e 100644 --- a/theories/channel/multi_channel.v +++ b/theories/channel/multi_channel.v @@ -113,12 +113,6 @@ Definition iProto_pointsto_def `{!heapGS Σ, !chanG Σ} ⌜v = PairV #l1 #l2⌠∗ inv (nroot.@"p") (chan_inv γ γE1 γt1 i j l1) ∗ inv (nroot.@"p") (chan_inv γ γE2 γt2 j i l2)) ∗ - (* llist (λ l v, *) - (* ∃ (j:nat) (l1 l2 : loc), *) - (* ⌜l = (j,(l1,l2))⌠∗ ⌜v = PairV #l1 #l2⌠∗ *) - (* inv (nroot.@"p") (chan_inv γ γE1 γt1 i j l1) ∗ *) - (* inv (nroot.@"p") (chan_inv γ γE2 γt2 j i l2)) l *) - (* (zip (seq 0 (length ls)) ls) ∗ *) own γE1 (â—E (Next p)) ∗ own γE1 (â—¯E (Next p)) ∗ iProto_own γ i p. Definition iProto_pointsto_aux : seal (@iProto_pointsto_def). by eexists. Qed. diff --git a/theories/channel/multi_proto.v b/theories/channel/multi_proto.v index 05b31f7..d0e29de 100644 --- a/theories/channel/multi_proto.v +++ b/theories/channel/multi_proto.v @@ -559,8 +559,8 @@ Qed. Definition iProto_le {Σ V} (i:nat) (p1 p2 : iProto Σ V) : iProp Σ := ∀ ps, iProto_consistent (<[i:=p1]>ps) -∗ iProto_consistent (<[i:=p2]>ps). Arguments iProto_le {_ _} _ _%proto _%proto. -Global Instance: Params (@iProto_le) 3 := {}. -(* Notation "p ⊑ q" := (iProto_le i p q) : bi_scope. *) +Global Instance: Params (@iProto_le) 2 := {}. +Notation "p ⊑{ i } q" := (iProto_le i p q) (at level 20) : bi_scope. Global Instance iProto_le_ne {Σ V} n : Proper ((=) ==> (dist n) ==> (dist n) ==> (dist n)) (@iProto_le Σ V). Proof. solve_proper. Qed. @@ -615,7 +615,7 @@ Section proto. Lemma iProto_consistent_le ps i p1 p2 : ps !! i = Some p1 → - iProto_le i p1 p2 -∗ + p1 ⊑{i} p2 -∗ iProto_consistent ps -∗ iProto_consistent (<[i := p2]>ps). Proof. @@ -624,6 +624,117 @@ Section proto. iApply "Hle". Qed. + Lemma iProto_le_refl i p : ⊢ iProto_le i p p. + Proof. iIntros (ps) "$". Qed. + + Lemma iProto_le_send i j m1 m2 : + (∀ v p2', iMsg_car m2 v (Next p2') -∗ ∃ p1', + â–· (p1' ⊑{i} p2') ∗ iMsg_car m1 v (Next p1')) -∗ + (<Send j> m1) ⊑{i} (<Send j> m2). + Proof. + iIntros "Hle". + iIntros (ps) "H". + iLöb as "IH" forall (ps m1 m2). + rewrite !iProto_consistent_unfold. + iIntros (i' j' m1' m2') "#Hm1' #Hm2'". + iAssert (⌜i ≠j'âŒ)%I as %Hneq'. + { destruct (decide (i = j')) as [<-|Hneq]; [|done]. + rewrite lookup_total_insert. rewrite iProto_message_equivI. + iDestruct "Hm2'" as (Heq) "_". inversion Heq. } + destruct (decide (i = i')) as [<-|Hne]; last first. + { rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + iDestruct ("H" $! i' j' with "[Hm1'] [Hm2']") as "H". + { rewrite lookup_total_insert_ne; [|done]. done. } + { rewrite lookup_total_insert_ne; [|done]. done. } + iIntros (v p1) "Hm1". + iDestruct ("H" with "Hm1") as (p2) "[Hm2 H]". + iExists p2. iFrame. iNext. + rewrite !(insert_commute _ j' i); [|done..]. + rewrite !(insert_commute _ i' i); [|done..]. + iApply ("IH" with "Hle H"). } + rewrite lookup_total_insert. + rewrite lookup_total_insert_ne; [|done]. + iIntros (v p1) "Hm1". + rewrite iProto_message_equivI. + iDestruct "Hm1'" as "[%Heq #Hm1']". + inversion Heq. simplify_eq. + iSpecialize ("Hm1'" $! v (Next p1)). + iRewrite -"Hm1'" in "Hm1". + iDestruct ("Hle" with "Hm1") as (p2) "[Hle Hm2]". + iDestruct ("H" $!i with "[] [] Hm2") as (p2') "[Hm2 H]". + { rewrite lookup_total_insert. done. } + { rewrite lookup_total_insert_ne; [|done]. done. } + iExists p2'. + iFrame. + iNext. iApply "Hle". + rewrite !(insert_commute _ j' i); [|done..]. + rewrite !insert_insert. done. + Qed. + + Lemma iProto_le_recv i j m1 m2 : + (∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', + â–· (p1' ⊑{i} p2') ∗ iMsg_car m2 v (Next p2')) -∗ + (<Recv j> m1) ⊑{i} (<Recv j> m2). + Proof. + iIntros "Hle". + iIntros (ps) "H". + iLöb as "IH" forall (ps m1 m2). + rewrite !iProto_consistent_unfold. + iIntros (i' j' m1' m2') "#Hm1' #Hm2'". + iAssert (⌜i ≠i'âŒ)%I as %Hneq'. + { destruct (decide (i = i')) as [<-|Hneq]; [|done]. + rewrite lookup_total_insert. rewrite iProto_message_equivI. + iDestruct "Hm1'" as (Heq) "_". inversion Heq. } + destruct (decide (i = j')) as [<-|Hne]; last first. + { rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + iDestruct ("H" $! i' j' with "[Hm1'] [Hm2']") as "H". + { rewrite lookup_total_insert_ne; [|done]. done. } + { rewrite lookup_total_insert_ne; [|done]. done. } + iIntros (v p1) "Hm1". + iDestruct ("H" with "Hm1") as (p2) "[Hm2 H]". + iExists p2. iFrame. iNext. + rewrite !(insert_commute _ j' i); [|done..]. + rewrite !(insert_commute _ i' i); [|done..]. + iApply ("IH" with "Hle H"). } + rewrite lookup_total_insert. + rewrite lookup_total_insert_ne; [|done]. + iIntros (v p1) "Hm1". + rewrite iProto_message_equivI. + iDestruct "Hm2'" as "[%Heq #Hm2']". + inversion Heq. simplify_eq. + iDestruct ("H" $!i' with "[] [] Hm1") as (p2') "[Hm2 H]". + { rewrite lookup_total_insert_ne; [|done]. done. } + { rewrite lookup_total_insert. done. } + iDestruct ("Hle" with "Hm2") as (p2) "[Hle Hm2]". + iSpecialize ("Hm2'" $! v (Next p2)). + iRewrite "Hm2'" in "Hm2". + iExists p2. + iFrame. + iNext. + rewrite !insert_insert. + rewrite !(insert_commute _ i' i); [|done..]. + iApply "Hle". done. + Qed. + + Lemma iProto_le_base i a v P p1 p2 : + â–· (p1 ⊑{i} p2) -∗ + (<a> MSG v {{ P }}; p1) ⊑{i} (<a> MSG v {{ P }}; p2). + Proof. + rewrite iMsg_base_eq. iIntros "H". destruct a. + - iApply iProto_le_send. iIntros (v' p') "(->&Hp&$)". + iExists p1. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". + - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&$)". + iExists p2. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". + Qed. + + Lemma iProto_le_trans i p1 p2 p3 : p1 ⊑{i} p2 -∗ p2 ⊑{i} p3 -∗ p1 ⊑{i} p3. + Proof. + iIntros "H1 H2" (p) "Hprot". + iApply "H2". iApply "H1". done. + Qed. + Lemma iProto_consistent_step ps m1 m2 i j v p1 : iProto_consistent ps -∗ ps !!! i ≡ (<(Send j)> m1) -∗ -- GitLab From 5770cc5f767e2f85ecb74cad8a14699e83d7a880 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Wed, 24 Jan 2024 13:53:38 +0100 Subject: [PATCH 17/81] Updated subprotocol definition and proved lemmas --- _CoqProject | 1 + theories/channel/multi_channel.v | 8 +- theories/channel/multi_proto.v | 491 +++++++++++++----- .../multi_proto_consistency_examples.v | 38 +- theories/channel/multi_proto_model.v | 9 +- 5 files changed, 404 insertions(+), 143 deletions(-) diff --git a/_CoqProject b/_CoqProject index 2fcbc7e..1b9d3a3 100644 --- a/_CoqProject +++ b/_CoqProject @@ -20,6 +20,7 @@ theories/channel/channel.v theories/channel/multi_proto_model.v theories/channel/multi_proto.v theories/channel/multi_channel.v +theories/channel/multi_proto_consistency_examples.v theories/channel/proofmode.v theories/examples/basics.v theories/examples/equivalence.v diff --git a/theories/channel/multi_channel.v b/theories/channel/multi_channel.v index 157bf8e..40f61e1 100644 --- a/theories/channel/multi_channel.v +++ b/theories/channel/multi_channel.v @@ -97,7 +97,7 @@ Definition tok `{!chanG Σ} (γ : gname) : iProp Σ := own γ (Excl ()). Definition chan_inv `{!heapGS Σ, !chanG Σ} γ γE γt i j (l:loc) : iProp Σ := (l ↦ NONEV ∗ tok γt) ∨ (∃ v m, l ↦ SOMEV v ∗ - iProto_own γ i (<Send j> m)%proto ∗ + iProto_own γ i (<(Send, j)> m)%proto ∗ (∃ p, iMsg_car m v (Next p) ∗ own γE (â—E (Next p)))) ∨ (∃ p, l ↦ NONEV ∗ iProto_own γ i p ∗ own γE (â—E (Next p))). @@ -158,7 +158,7 @@ Section channel. Proof. Admitted. Lemma send_spec c j v p : - {{{ c ↣ <Send j> MSG v; p }}} + {{{ c ↣ <(Send, j)> MSG v; p }}} send c #j v {{{ RET #(); c ↣ p }}}. Proof. @@ -233,7 +233,7 @@ Section channel. Lemma send_spec_tele {TT} c j (tt : TT) (v : TT → val) (P : TT → iProp Σ) (p : TT → iProto Σ) : - {{{ c ↣ (<(Send j) @.. x > MSG v x {{ P x }}; p x) ∗ P tt }}} + {{{ c ↣ (<(Send, j) @.. x > MSG v x {{ P x }}; p x) ∗ P tt }}} send c #j (v tt) {{{ RET #(); c ↣ (p tt) }}}. Proof. @@ -312,7 +312,7 @@ Section channel. Qed. Lemma recv_spec {TT} c j (v : TT → val) (P : TT → iProp Σ) (p : TT → iProto Σ) : - {{{ c ↣ <(Recv j) @.. x> MSG v x {{ P x }}; p x }}} + {{{ c ↣ <(Recv, j) @.. x> MSG v x {{ P x }}; p x }}} recv c #j {{{ x, RET v x; c ↣ p x ∗ P x }}}. Proof. diff --git a/theories/channel/multi_proto.v b/theories/channel/multi_proto.v index d0e29de..a24b94e 100644 --- a/theories/channel/multi_proto.v +++ b/theories/channel/multi_proto.v @@ -254,11 +254,11 @@ Section proto. Implicit Types m : iMsg Σ V. (** ** Equality *) - Lemma iProto_case p : p ≡ END ∨ ∃ a m, p ≡ <a> m. + Lemma iProto_case p : p ≡ END ∨ ∃ t n m, p ≡ <(t,n)> m. Proof. rewrite iProto_message_eq iProto_end_eq. - destruct (proto_case p) as [|(a&m&?)]; [by left|right]. - by exists a, (IMsg m). + destruct (proto_case p) as [|([a n]&m&?)]; [by left|right]. + by exists a, n, (IMsg m). Qed. Lemma iProto_message_equivI `{!BiInternalEq SPROP} a1 a2 m1 m2 : (<a1> m1) ≡ (<a2> m2) ⊣⊢@{SPROP} ⌜ a1 = a2 ⌠∧ @@ -410,7 +410,7 @@ Section proto. Global Instance iProto_dual_involutive : Involutive (≡) (@iProto_dual Σ V). Proof. intros p. apply (uPred.internal_eq_soundness (M:=iResUR Σ)). - iLöb as "IH" forall (p). destruct (iProto_case p) as [->|(a&m&->)]. + iLöb as "IH" forall (p). destruct (iProto_case p) as [->|(a&n&m&->)]. { by rewrite !iProto_dual_end. } rewrite !iProto_dual_message involutive. iApply iProto_message_equivI; iSplit; [done|]; iIntros (v p') "/=". @@ -446,7 +446,7 @@ Section proto. { intros p1 p1' Hp1 p2 p2' <-. by rewrite iProto_app_eq /iProto_app_def Hp1. } intros n p1 p1' Hp1 p2 p2' Hp2. rewrite Hp1. clear p1 Hp1. revert p1'. induction (lt_wf n) as [n _ IH]; intros p1. - destruct (iProto_case p1) as [->|(a&m&->)]. + destruct (iProto_case p1) as [->|(a&i&m&->)]. { by rewrite !left_id. } rewrite !iProto_app_message. f_equiv=> v p' /=. do 4 f_equiv. f_contractive. apply IH; eauto using dist_le with lia. @@ -464,7 +464,7 @@ Section proto. Global Instance iProto_app_end_r : RightId (≡) END (@iProto_app Σ V). Proof. intros p. apply (uPred.internal_eq_soundness (M:=iResUR Σ)). - iLöb as "IH" forall (p). destruct (iProto_case p) as [->|(a&m&->)]. + iLöb as "IH" forall (p). destruct (iProto_case p) as [->|(a&i&m&->)]. { by rewrite left_id. } rewrite iProto_app_message. iApply iProto_message_equivI; iSplit; [done|]; iIntros (v p') "/=". @@ -478,7 +478,7 @@ Section proto. Global Instance iProto_app_assoc : Assoc (≡) (@iProto_app Σ V). Proof. intros p1 p2 p3. apply (uPred.internal_eq_soundness (M:=iResUR Σ)). - iLöb as "IH" forall (p1). destruct (iProto_case p1) as [->|(a&m&->)]. + iLöb as "IH" forall (p1). destruct (iProto_case p1) as [->|(a&i&m&->)]. { by rewrite !left_id. } rewrite !iProto_app_message. iApply iProto_message_equivI; iSplit; [done|]; iIntros (v p123) "/=". @@ -495,7 +495,7 @@ Section proto. iProto_dual (p1 <++> p2) ≡ iProto_dual p1 <++> iProto_dual p2. Proof. apply (uPred.internal_eq_soundness (M:=iResUR Σ)). - iLöb as "IH" forall (p1 p2). destruct (iProto_case p1) as [->|(a&m&->)]. + iLöb as "IH" forall (p1 p2). destruct (iProto_case p1) as [->|(a&i&m&->)]. { by rewrite iProto_dual_end !left_id. } rewrite iProto_dual_message !iProto_app_message iProto_dual_message /=. iApply iProto_message_equivI; iSplit; [done|]; iIntros (v p12) "/=". @@ -515,7 +515,7 @@ Instance iProto_inhabited {Σ V} : Inhabited (iProto Σ V) := populate (END). Definition can_step {Σ V} (rec : gmap nat (iProto Σ V) → iProp Σ) (ps : gmap nat (iProto Σ V)) (i j : nat) : iProp Σ := ∀ m1 m2, - (ps !!! i ≡ <(Send j)> m1) -∗ (ps !!! j ≡ <Recv i> m2) -∗ + (ps !!! i ≡ <(Send, j)> m1) -∗ (ps !!! j ≡ <(Recv, i)> m2) -∗ ∀ v p1, iMsg_car m1 v (Next p1) -∗ ∃ p2, iMsg_car m2 v (Next p2) ∗ â–· (rec (<[i:=p1]>(<[j:=p2]>ps))). @@ -556,17 +556,56 @@ Lemma iProto_consistent_unfold {Σ V} (ps : gmap nat (iProto Σ V)) : Proof. apply: (fixpoint_unfold iProto_consistent_pre'). Qed. -Definition iProto_le {Σ V} (i:nat) (p1 p2 : iProto Σ V) : iProp Σ := - ∀ ps, iProto_consistent (<[i:=p1]>ps) -∗ iProto_consistent (<[i:=p2]>ps). -Arguments iProto_le {_ _} _ _%proto _%proto. + +(** * Protocol entailment *) +Definition iProto_le_pre {Σ V} + (rec : iProto Σ V → iProto Σ V → iProp Σ) (p1 p2 : iProto Σ V) : iProp Σ := + (p1 ≡ END ∗ p2 ≡ END) ∨ + ∃ a1 a2 m1 m2, + (p1 ≡ <a1> m1) ∗ (p2 ≡ <a2> m2) ∗ + match a1, a2 with + | (Recv,i), (Recv,j) => ⌜i = j⌠∗ ∀ v p1', + iMsg_car m1 v (Next p1') -∗ ∃ p2', â–· rec p1' p2' ∗ iMsg_car m2 v (Next p2') + | (Send,i), (Send,j) => ⌜i = j⌠∗ ∀ v p2', + iMsg_car m2 v (Next p2') -∗ ∃ p1', â–· rec p1' p2' ∗ iMsg_car m1 v (Next p1') + | _, _ => False + end. +Global Instance iProto_le_pre_ne {Σ V} (rec : iProto Σ V → iProto Σ V → iProp Σ) : + NonExpansive2 (iProto_le_pre rec). +Proof. solve_proper. Qed. + +Program Definition iProto_le_pre' {Σ V} + (rec : iProto Σ V -n> iProto Σ V -n> iPropO Σ) : + iProto Σ V -n> iProto Σ V -n> iPropO Σ := λne p1 p2, + iProto_le_pre (λ p1' p2', rec p1' p2') p1 p2. +Solve Obligations with solve_proper. +Local Instance iProto_le_pre_contractive {Σ V} : Contractive (@iProto_le_pre' Σ V). +Proof. + intros n rec1 rec2 Hrec p1 p2. rewrite /iProto_le_pre' /iProto_le_pre /=. + by repeat (f_contractive || f_equiv). +Qed. +Definition iProto_le {Σ V} (p1 p2 : iProto Σ V) : iProp Σ := + fixpoint iProto_le_pre' p1 p2. +Arguments iProto_le {_ _} _%proto _%proto. Global Instance: Params (@iProto_le) 2 := {}. -Notation "p ⊑{ i } q" := (iProto_le i p q) (at level 20) : bi_scope. +Notation "p ⊑ q" := (iProto_le p q) : bi_scope. -Global Instance iProto_le_ne {Σ V} n : Proper ((=) ==> (dist n) ==> (dist n) ==> (dist n)) (@iProto_le Σ V). +Global Instance iProto_le_ne {Σ V} : NonExpansive2 (@iProto_le Σ V). Proof. solve_proper. Qed. -Global Instance iProto_le_proper {Σ V} : Proper ((=) ==> (≡) ==> (≡) ==> (⊣⊢)) (@iProto_le Σ V). +Global Instance iProto_le_proper {Σ V} : Proper ((≡) ==> (≡) ==> (⊣⊢)) (@iProto_le Σ V). Proof. solve_proper. Qed. +(* Definition iProto_le {Σ V} (i:nat) (p1 p2 : iProto Σ V) : iProp Σ := *) +(* ∀ ps, iProto_consistent (<[i:=p1]>ps) -∗ iProto_consistent (<[i:=p2]>ps). *) +(* Arguments iProto_le {_ _} _ _%proto _%proto. *) +(* Global Instance: Params (@iProto_le) 2 := {}. *) +(* Notation "p ⊑{ i } q" := (iProto_le i p q) (at level 20) : bi_scope. *) + +(* Global Instance iProto_le_ne {Σ V} n : Proper ((=) ==> (dist n) ==> (dist n) ==> (dist n)) (@iProto_le Σ V). *) +(* Proof. solve_proper. Qed. *) +(* Global Instance iProto_le_proper {Σ V} : Proper ((=) ==> (≡) ==> (≡) ==> (⊣⊢)) (@iProto_le Σ V). *) +(* Proof. solve_proper. Qed. *) + Record proto_name := ProtName { proto_names : gmap nat gname }. Global Instance proto_name_inhabited : Inhabited proto_name := populate (ProtName inhabitant). @@ -613,132 +652,352 @@ Section proto. Implicit Types p pl pr : iProto Σ V. Implicit Types m : iMsg Σ V. - Lemma iProto_consistent_le ps i p1 p2 : - ps !! i = Some p1 → - p1 ⊑{i} p2 -∗ - iProto_consistent ps -∗ - iProto_consistent (<[i := p2]>ps). + (** ** Protocol entailment **) + Lemma iProto_le_unfold p1 p2 : iProto_le p1 p2 ≡ iProto_le_pre iProto_le p1 p2. + Proof. apply: (fixpoint_unfold iProto_le_pre'). Qed. + + Lemma iProto_le_end : ⊢ END ⊑ (END : iProto Σ V). + Proof. rewrite iProto_le_unfold. iLeft. auto 10. Qed. + + Lemma iProto_le_end_inv_l p : p ⊑ END -∗ (p ≡ END). Proof. - iIntros (HSome) "Hle". - rewrite -{1}(insert_id ps i p1); [|done]. - iApply "Hle". + rewrite iProto_le_unfold. iIntros "[[Hp _]|H] //". + iDestruct "H" as (a1 a2 m1 m2) "(_ & Heq & _)". + by iDestruct (iProto_end_message_equivI with "Heq") as %[]. Qed. - Lemma iProto_le_refl i p : ⊢ iProto_le i p p. - Proof. iIntros (ps) "$". Qed. + Lemma iProto_le_end_inv_r p : END ⊑ p -∗ (p ≡ END). + Proof. + rewrite iProto_le_unfold. iIntros "[[_ Hp]|H] //". + iDestruct "H" as (a1 a2 m1 m2) "(Heq & _ & _)". + iDestruct (iProto_end_message_equivI with "Heq") as %[]. + Qed. - Lemma iProto_le_send i j m1 m2 : - (∀ v p2', iMsg_car m2 v (Next p2') -∗ ∃ p1', - â–· (p1' ⊑{i} p2') ∗ iMsg_car m1 v (Next p1')) -∗ - (<Send j> m1) ⊑{i} (<Send j> m2). + Lemma iProto_le_send_inv i p1 m2 : + p1 ⊑ (<(Send,i)> m2) -∗ ∃ m1, + (p1 ≡ <(Send,i)> m1) ∗ + ∀ v p2', iMsg_car m2 v (Next p2') -∗ ∃ p1', + â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1'). Proof. - iIntros "Hle". - iIntros (ps) "H". - iLöb as "IH" forall (ps m1 m2). - rewrite !iProto_consistent_unfold. - iIntros (i' j' m1' m2') "#Hm1' #Hm2'". - iAssert (⌜i ≠j'âŒ)%I as %Hneq'. - { destruct (decide (i = j')) as [<-|Hneq]; [|done]. - rewrite lookup_total_insert. rewrite iProto_message_equivI. - iDestruct "Hm2'" as (Heq) "_". inversion Heq. } - destruct (decide (i = i')) as [<-|Hne]; last first. - { rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - iDestruct ("H" $! i' j' with "[Hm1'] [Hm2']") as "H". - { rewrite lookup_total_insert_ne; [|done]. done. } - { rewrite lookup_total_insert_ne; [|done]. done. } - iIntros (v p1) "Hm1". - iDestruct ("H" with "Hm1") as (p2) "[Hm2 H]". - iExists p2. iFrame. iNext. - rewrite !(insert_commute _ j' i); [|done..]. - rewrite !(insert_commute _ i' i); [|done..]. - iApply ("IH" with "Hle H"). } - rewrite lookup_total_insert. - rewrite lookup_total_insert_ne; [|done]. - iIntros (v p1) "Hm1". + rewrite iProto_le_unfold. + iIntros "[[_ Heq]|H]". + { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } + iDestruct "H" as (a1 a2 m1 m2') "(Hp1 & Hp2 & H)". + rewrite iProto_message_equivI. iDestruct "Hp2" as "[%Heq Hm2]". + simplify_eq. + destruct a1 as [[]]; [|done]. + iDestruct "H" as (->) "H". + iExists m1. iFrame. + iIntros (v p2). + iSpecialize ("Hm2" $! v (Next p2)). + iRewrite "Hm2". iApply "H". + Qed. + + Lemma iProto_le_send_send_inv i m1 m2 v p2' : + (<(Send,i)> m1) ⊑ (<(Send,i)> m2) -∗ + iMsg_car m2 v (Next p2') -∗ ∃ p1', â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1'). + Proof. + iIntros "H Hm2". iDestruct (iProto_le_send_inv with "H") as (m1') "[Hm1 H]". + iDestruct (iProto_message_equivI with "Hm1") as (Heq) "Hm1". + iDestruct ("H" with "Hm2") as (p1') "[Hle Hm]". + iRewrite -("Hm1" $! v (Next p1')) in "Hm". auto with iFrame. + Qed. + + + Lemma iProto_le_recv_inv_l i m1 p2 : + (<(Recv,i)> m1) ⊑ p2 -∗ ∃ m2, + (p2 ≡ <(Recv,i)> m2) ∗ + ∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', + â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). + Proof. + rewrite iProto_le_unfold. + iIntros "[[Heq _]|H]". + { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } + iDestruct "H" as (a1 a2 m1' m2) "(Hp1 & Hp2 & H)". + rewrite iProto_message_equivI. iDestruct "Hp1" as "[%Heq Hm1]". + simplify_eq. + destruct a2 as [[]]; [done|]. + iDestruct "H" as (->) "H". + iExists m2. iFrame. + iIntros (v p1). + iSpecialize ("Hm1" $! v (Next p1)). + iRewrite "Hm1". iApply "H". + Qed. + + Lemma iProto_le_recv_inv_r i p1 m2 : + (p1 ⊑ <(Recv,i)> m2) -∗ ∃ m1, + (p1 ≡ <(Recv,i)> m1) ∗ + ∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', + â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). + Proof. + rewrite iProto_le_unfold. + iIntros "[[_ Heq]|H]". + { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } + iDestruct "H" as (a1 a2 m1 m2') "(Hp1 & Hp2 & H)". rewrite iProto_message_equivI. - iDestruct "Hm1'" as "[%Heq #Hm1']". - inversion Heq. simplify_eq. - iSpecialize ("Hm1'" $! v (Next p1)). - iRewrite -"Hm1'" in "Hm1". - iDestruct ("Hle" with "Hm1") as (p2) "[Hle Hm2]". - iDestruct ("H" $!i with "[] [] Hm2") as (p2') "[Hm2 H]". - { rewrite lookup_total_insert. done. } - { rewrite lookup_total_insert_ne; [|done]. done. } - iExists p2'. - iFrame. - iNext. iApply "Hle". - rewrite !(insert_commute _ j' i); [|done..]. - rewrite !insert_insert. done. + iDestruct "Hp2" as "[%Heq Hm2]". + simplify_eq. + destruct a1 as [[]]; [done|]. + iDestruct "H" as (->) "H". + iExists m1. iFrame. + iIntros (v p2). + iIntros "Hm1". iDestruct ("H" with "Hm1") as (p2') "[Hle H]". + iSpecialize ("Hm2" $! v (Next p2')). + iExists p2'. iFrame. + iRewrite "Hm2". iApply "H". Qed. - Lemma iProto_le_recv i j m1 m2 : - (∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', - â–· (p1' ⊑{i} p2') ∗ iMsg_car m2 v (Next p2')) -∗ - (<Recv j> m1) ⊑{i} (<Recv j> m2). + Lemma iProto_le_recv_recv_inv i m1 m2 v p1' : + (<(Recv, i)> m1) ⊑ (<(Recv, i)> m2) -∗ + iMsg_car m1 v (Next p1') -∗ ∃ p2', â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). + Proof. + iIntros "H Hm2". iDestruct (iProto_le_recv_inv_r with "H") as (m1') "[Hm1 H]". + iApply "H". iDestruct (iProto_message_equivI with "Hm1") as (_) "Hm1". + by iRewrite -("Hm1" $! v (Next p1')). + Qed. + + Lemma iProto_consistent_le ps i p1 p2 : + ps !! i = Some p1 → + p1 ⊑ p2 -∗ + iProto_consistent ps -∗ + iProto_consistent (<[i := p2]>ps). Proof. - iIntros "Hle". - iIntros (ps) "H". - iLöb as "IH" forall (ps m1 m2). + iIntros (HSome) "Hle Hprot". + iLöb as "IH" forall (p1 p2 ps HSome). rewrite !iProto_consistent_unfold. - iIntros (i' j' m1' m2') "#Hm1' #Hm2'". - iAssert (⌜i ≠i'âŒ)%I as %Hneq'. - { destruct (decide (i = i')) as [<-|Hneq]; [|done]. - rewrite lookup_total_insert. rewrite iProto_message_equivI. - iDestruct "Hm1'" as (Heq) "_". inversion Heq. } - destruct (decide (i = j')) as [<-|Hne]; last first. - { rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - iDestruct ("H" $! i' j' with "[Hm1'] [Hm2']") as "H". - { rewrite lookup_total_insert_ne; [|done]. done. } + iIntros (i' j' m1 m2) "#Hm1 #Hm2". + destruct (decide (i = i')) as [<-|Hneq]. + { rewrite lookup_total_insert. + pose proof (iProto_case p2) as [Hend|Hmsg]. + { setoid_rewrite Hend. rewrite iProto_end_message_equivI. done. } + destruct Hmsg as (a&?&m&Hmsg). + setoid_rewrite Hmsg. + destruct a; last first. + { rewrite iProto_message_equivI. + iDestruct "Hm1" as "[%Htag Hm1]". done. } + rewrite iProto_message_equivI. + iDestruct "Hm1" as "[%Htag Hm1]". + inversion Htag. simplify_eq. + iIntros (v p) "Hm1'". + iSpecialize ("Hm1" $! v (Next p)). + iDestruct (iProto_le_send_inv with "Hle") as "Hle". + iRewrite -"Hm1" in "Hm1'". + iDestruct "Hle" as (m') "[#Heq H]". + iDestruct ("H" with "Hm1'") as (p') "[Hle H]". + destruct (decide (i = j')) as [<-|Hneq]. + { rewrite lookup_total_insert. rewrite iProto_message_equivI. + iDestruct "Hm2" as "[%Heq _]". done. } + iDestruct ("Hprot" $!i j' with "[] [] H") as "Hprot". + { iRewrite -"Heq". rewrite !lookup_total_alt. rewrite HSome. done. } { rewrite lookup_total_insert_ne; [|done]. done. } - iIntros (v p1) "Hm1". - iDestruct ("H" with "Hm1") as (p2) "[Hm2 H]". - iExists p2. iFrame. iNext. - rewrite !(insert_commute _ j' i); [|done..]. - rewrite !(insert_commute _ i' i); [|done..]. - iApply ("IH" with "Hle H"). } - rewrite lookup_total_insert. + iDestruct "Hprot" as (p'') "[Hm Hprot]". + iExists p''. iFrame. + iNext. + iDestruct ("IH" with "[] Hle Hprot") as "HI". + { iPureIntro. by rewrite lookup_insert. } + iClear "IH Hm1 Hm2 Heq". + rewrite insert_insert. + rewrite (insert_commute _ j' i); [|done]. + rewrite insert_insert. done. } rewrite lookup_total_insert_ne; [|done]. - iIntros (v p1) "Hm1". - rewrite iProto_message_equivI. - iDestruct "Hm2'" as "[%Heq #Hm2']". - inversion Heq. simplify_eq. - iDestruct ("H" $!i' with "[] [] Hm1") as (p2') "[Hm2 H]". - { rewrite lookup_total_insert_ne; [|done]. done. } - { rewrite lookup_total_insert. done. } - iDestruct ("Hle" with "Hm2") as (p2) "[Hle Hm2]". - iSpecialize ("Hm2'" $! v (Next p2)). - iRewrite "Hm2'" in "Hm2". - iExists p2. - iFrame. + destruct (decide (i = j')) as [<-|Hneq']. + { rewrite lookup_total_insert. + pose proof (iProto_case p2) as [Hend|Hmsg]. + { setoid_rewrite Hend. rewrite iProto_end_message_equivI. done. } + destruct Hmsg as (a&?&m&Hmsg). + setoid_rewrite Hmsg. + destruct a. + { rewrite iProto_message_equivI. + iDestruct "Hm2" as "[%Htag Hm2]". done. } + rewrite iProto_message_equivI. + iDestruct "Hm2" as "[%Htag Hm2]". + inversion Htag. simplify_eq. + iIntros (v p) "Hm1'". + iDestruct (iProto_le_recv_inv_r with "Hle") as "Hle". + (* iRewrite -"Hm2" in "Hm2'". *) + iDestruct "Hle" as (m') "[#Heq Hle]". + + iDestruct ("Hprot" $!i' with "[] [] Hm1'") as "Hprot". + { done. } + { rewrite !lookup_total_alt. rewrite HSome. done. } + iDestruct ("Hprot") as (p') "[Hm1' Hprot]". + iDestruct ("Hle" with "Hm1'") as (p2') "[Hle Hm']". + iSpecialize ("Hm2" $! v (Next p2')). + iExists p2'. + iRewrite -"Hm2". iFrame. + iDestruct ("IH" with "[] Hle Hprot") as "HI". + { iPureIntro. rewrite lookup_insert_ne; [|done]. rewrite lookup_insert. done. } + rewrite insert_commute; [|done]. + rewrite !insert_insert. done. } + rewrite lookup_total_insert_ne; [|done]. + iIntros (v p) "Hm1'". + iDestruct ("Hprot" $!i' j' with "[//] [//] Hm1'") as "Hprot". + iDestruct "Hprot" as (p') "[Hm2' Hprot]". + iExists p'. iFrame. iNext. - rewrite !insert_insert. - rewrite !(insert_commute _ i' i); [|done..]. - iApply "Hle". done. + rewrite (insert_commute _ j' i); [|done]. + rewrite (insert_commute _ i' i); [|done]. + iApply ("IH" with "[] Hle Hprot"). + iPureIntro. + rewrite lookup_insert_ne; [|done]. + rewrite lookup_insert_ne; [|done]. + done. + Qed. + + Lemma iProto_le_send i m1 m2 : + (∀ v p2', iMsg_car m2 v (Next p2') -∗ ∃ p1', + â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1')) -∗ + (<(Send,i)> m1) ⊑ (<(Send,i)> m2). + Proof. + iIntros "Hle". rewrite iProto_le_unfold. + iRight. iExists (Send, i), (Send, i), m1, m2. by eauto. + Qed. + + Lemma iProto_le_recv i m1 m2 : + (∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', + â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2')) -∗ + (<(Recv,i)> m1) ⊑ (<(Recv,i)> m2). + Proof. + iIntros "Hle". rewrite iProto_le_unfold. + iRight. iExists (Recv, i), (Recv, i), m1, m2. by eauto. Qed. - Lemma iProto_le_base i a v P p1 p2 : - â–· (p1 ⊑{i} p2) -∗ - (<a> MSG v {{ P }}; p1) ⊑{i} (<a> MSG v {{ P }}; p2). + Lemma iProto_le_base a v P p1 p2 : + â–· (p1 ⊑ p2) -∗ + (<a> MSG v {{ P }}; p1) ⊑ (<a> MSG v {{ P }}; p2). Proof. - rewrite iMsg_base_eq. iIntros "H". destruct a. + rewrite iMsg_base_eq. iIntros "H". destruct a as [[]]. - iApply iProto_le_send. iIntros (v' p') "(->&Hp&$)". iExists p1. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&$)". iExists p2. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". Qed. - Lemma iProto_le_trans i p1 p2 p3 : p1 ⊑{i} p2 -∗ p2 ⊑{i} p3 -∗ p1 ⊑{i} p3. + Lemma iProto_le_trans p1 p2 p3 : p1 ⊑ p2 -∗ p2 ⊑ p3 -∗ p1 ⊑ p3. + Proof. + iIntros "H1 H2". iLöb as "IH" forall (p1 p2 p3). + destruct (iProto_case p3) as [->|([]&i&m3&->)]. + - iDestruct (iProto_le_end_inv_l with "H2") as "H2". by iRewrite "H2" in "H1". + - iDestruct (iProto_le_send_inv with "H2") as (m2) "[Hp2 H2]". + iRewrite "Hp2" in "H1"; clear p2. + iDestruct (iProto_le_send_inv with "H1") as (m1) "[Hp1 H1]". + iRewrite "Hp1"; clear p1. + iApply iProto_le_send. iIntros (v p3') "Hm3". + iDestruct ("H2" with "Hm3") as (p2') "[Hle Hm2]". + iDestruct ("H1" with "Hm2") as (p1') "[Hle' Hm1]". + iExists p1'. iIntros "{$Hm1} !>". by iApply ("IH" with "Hle'"). + - iDestruct (iProto_le_recv_inv_r with "H2") as (m2) "[Hp2 H3]". + iRewrite "Hp2" in "H1". + iDestruct (iProto_le_recv_inv_r with "H1") as (m1) "[Hp1 H2]". + iRewrite "Hp1". iApply iProto_le_recv. iIntros (v p1') "Hm1". + iDestruct ("H2" with "Hm1") as (p2') "[Hle Hm2]". + iDestruct ("H3" with "Hm2") as (p3') "[Hle' Hm3]". + iExists p3'. iIntros "{$Hm3} !>". by iApply ("IH" with "Hle"). + Qed. + + Lemma iProto_le_refl p : ⊢ p ⊑ p. + Proof. + iLöb as "IH" forall (p). destruct (iProto_case p) as [->|([]&i&m&->)]. + - iApply iProto_le_end. + - iApply iProto_le_send. auto 10 with iFrame. + - iApply iProto_le_recv. auto 10 with iFrame. + Qed. + + Lemma iProto_le_payload_elim_l i m v P p : + (P -∗ (<(Recv,i)> MSG v; p) ⊑ (<(Recv,i)> m)) ⊢ + (<(Recv,i)> MSG v {{ P }}; p) ⊑ <(Recv,i)> m. + Proof. + rewrite iMsg_base_eq. iIntros "H". + iApply iProto_le_recv. iIntros (v' p') "(->&Hp&HP)". + iApply (iProto_le_recv_recv_inv with "(H HP)"); simpl; auto. + Qed. + Lemma iProto_le_payload_elim_r i m v P p : + (P -∗ (<(Send, i)> m) ⊑ (<(Send, i)> MSG v; p)) ⊢ + (<(Send,i)> m) ⊑ (<(Send,i)> MSG v {{ P }}; p). + Proof. + rewrite iMsg_base_eq. iIntros "H". + iApply iProto_le_send. iIntros (v' p') "(->&Hp&HP)". + iApply (iProto_le_send_send_inv with "(H HP)"); simpl; auto. + Qed. + Lemma iProto_le_payload_intro_l i v P p : + P -∗ (<(Send,i)> MSG v {{ P }}; p) ⊑ (<(Send,i)> MSG v; p). + Proof. + rewrite iMsg_base_eq. + iIntros "HP". iApply iProto_le_send. iIntros (v' p') "(->&Hp&_) /=". + iExists p'. iSplitR; [iApply iProto_le_refl|]. auto. + Qed. + Lemma iProto_le_payload_intro_r i v P p : + P -∗ (<(Recv,i)> MSG v; p) ⊑ (<(Recv,i)> MSG v {{ P }}; p). + Proof. + rewrite iMsg_base_eq. + iIntros "HP". iApply iProto_le_recv. iIntros (v' p') "(->&Hp&_) /=". + iExists p'. iSplitR; [iApply iProto_le_refl|]. auto. + Qed. + Lemma iProto_le_exist_elim_l {A} i (m1 : A → iMsg Σ V) m2 : + (∀ x, (<(Recv,i)> m1 x) ⊑ (<(Recv,i)> m2)) ⊢ + (<(Recv,i) @ x> m1 x) ⊑ (<(Recv,i)> m2). + Proof. + rewrite iMsg_exist_eq. iIntros "H". + iApply iProto_le_recv. iIntros (v p1') "/=". iDestruct 1 as (x) "Hm". + by iApply (iProto_le_recv_recv_inv with "H"). + Qed. + Lemma iProto_le_exist_elim_r {A} i m1 (m2 : A → iMsg Σ V) : + (∀ x, (<(Send,i)> m1) ⊑ (<(Send,i)> m2 x)) ⊢ + (<(Send,i)> m1) ⊑ (<(Send,i) @ x> m2 x). + Proof. + rewrite iMsg_exist_eq. iIntros "H". + iApply iProto_le_send. iIntros (v p2'). iDestruct 1 as (x) "Hm". + by iApply (iProto_le_send_send_inv with "H"). + Qed. + Lemma iProto_le_exist_intro_l {A} i (m : A → iMsg Σ V) a : + ⊢ (<(Send,i) @ x> m x) ⊑ (<(Send,i)> m a). + Proof. + rewrite iMsg_exist_eq. iApply iProto_le_send. iIntros (v p') "Hm /=". + iExists p'. iSplitR; last by auto. iApply iProto_le_refl. + Qed. + Lemma iProto_le_exist_intro_r {A} i (m : A → iMsg Σ V) a : + ⊢ (<(Recv,i)> m a) ⊑ (<(Recv,i) @ x> m x). + Proof. + rewrite iMsg_exist_eq. iApply iProto_le_recv. iIntros (v p') "Hm /=". + iExists p'. iSplitR; last by auto. iApply iProto_le_refl. + Qed. + + Lemma iProto_le_texist_elim_l {TT : tele} i (m1 : TT → iMsg Σ V) m2 : + (∀ x, (<(Recv,i)> m1 x) ⊑ (<(Recv,i)> m2)) ⊢ + (<(Recv,i) @.. x> m1 x) ⊑ (<(Recv,i)> m2). + Proof. + iIntros "H". iInduction TT as [|T TT] "IH"; simpl; [done|]. + iApply iProto_le_exist_elim_l; iIntros (x). + iApply "IH". iIntros (xs). iApply "H". + Qed. + Lemma iProto_le_texist_elim_r {TT : tele} i m1 (m2 : TT → iMsg Σ V) : + (∀ x, (<(Send,i)> m1) ⊑ (<(Send,i)> m2 x)) -∗ + (<(Send,i)> m1) ⊑ (<(Send,i) @.. x> m2 x). + Proof. + iIntros "H". iInduction TT as [|T TT] "IH"; simpl; [done|]. + iApply iProto_le_exist_elim_r; iIntros (x). + iApply "IH". iIntros (xs). iApply "H". + Qed. + + Lemma iProto_le_texist_intro_l {TT : tele} i (m : TT → iMsg Σ V) x : + ⊢ (<(Send,i) @.. x> m x) ⊑ (<(Send,i)> m x). + Proof. + induction x as [|T TT x xs IH] using tele_arg_ind; simpl. + { iApply iProto_le_refl. } + iApply iProto_le_trans; [by iApply iProto_le_exist_intro_l|]. iApply IH. + Qed. + Lemma iProto_le_texist_intro_r {TT : tele} i (m : TT → iMsg Σ V) x : + ⊢ (<(Recv,i)> m x) ⊑ (<(Recv,i) @.. x> m x). Proof. - iIntros "H1 H2" (p) "Hprot". - iApply "H2". iApply "H1". done. + induction x as [|T TT x xs IH] using tele_arg_ind; simpl. + { iApply iProto_le_refl. } + iApply iProto_le_trans; [|by iApply iProto_le_exist_intro_r]. iApply IH. Qed. Lemma iProto_consistent_step ps m1 m2 i j v p1 : iProto_consistent ps -∗ - ps !!! i ≡ (<(Send j)> m1) -∗ - ps !!! j ≡ (<(Recv i)> m2) -∗ + ps !!! i ≡ (<(Send, j)> m1) -∗ + ps !!! j ≡ (<(Recv, i)> m2) -∗ iMsg_car m1 v (Next p1) -∗ ∃ p2, iMsg_car m2 v (Next p2) ∗ â–· iProto_consistent (<[i := p1]>(<[j := p2]>ps)). @@ -815,8 +1074,8 @@ Section proto. Lemma iProto_step γ i j m1 m2 p1 v : iProto_ctx γ -∗ - iProto_own γ i (<(Send j)> m1) -∗ - iProto_own γ j (<(Recv i)> m2) -∗ + iProto_own γ i (<(Send, j)> m1) -∗ + iProto_own γ j (<(Recv, i)> m2) -∗ iMsg_car m1 v (Next p1) ==∗ â–· ∃ p2, iMsg_car m2 v (Next p2) ∗ iProto_ctx γ ∗ iProto_own γ i p1 ∗ iProto_own γ j p2. diff --git a/theories/channel/multi_proto_consistency_examples.v b/theories/channel/multi_proto_consistency_examples.v index c3c3743..f9ec28f 100644 --- a/theories/channel/multi_proto_consistency_examples.v +++ b/theories/channel/multi_proto_consistency_examples.v @@ -20,8 +20,8 @@ Proof. Qed. Definition iProto_example2 `{!invGS Σ} (P : iProp Σ) : gmap nat (iProto Σ) := - <[0 := (<(Send 1) @ (x:Z)> MSG #x {{ P }} ; END)%proto ]> - (<[1 := (<(Recv 0) @ (x:Z)> MSG #x {{ P }} ; END)%proto ]> + <[0 := (<(Send, 1) @ (x:Z)> MSG #x {{ P }} ; END)%proto ]> + (<[1 := (<(Recv, 0) @ (x:Z)> MSG #x {{ P }} ; END)%proto ]> ∅). Lemma iProto_example2_consistent `{!invGS Σ} (P : iProp Σ) : @@ -103,9 +103,9 @@ Proof. Qed. Definition iProto_example3 `{!invGS Σ} : gmap nat (iProto Σ) := - <[0 := (<(Send 1) @ (x:Z)> MSG #x ; <(Recv 2)> MSG #x; END)%proto ]> - (<[1 := (<(Recv 0) @ (x:Z)> MSG #x ; <(Send 2)> MSG #x; END)%proto ]> - (<[2 := (<(Recv 1) @ (x:Z)> MSG #x ; <(Send 0)> MSG #x; END)%proto ]> + <[0 := (<(Send, 1) @ (x:Z)> MSG #x ; <(Recv, 2)> MSG #x; END)%proto ]> + (<[1 := (<(Recv, 0) @ (x:Z)> MSG #x ; <(Send, 2)> MSG #x; END)%proto ]> + (<[2 := (<(Recv, 1) @ (x:Z)> MSG #x ; <(Send, 0)> MSG #x; END)%proto ]> ∅)). Lemma iProto_example3_consistent `{!invGS Σ} : @@ -157,8 +157,8 @@ Proof. rewrite !iMsg_exist_eq. iRewrite -"Hm1" in "Hm1'". iDestruct "Hm1'" as (x Heq) "[#Hm1' _]". - iSpecialize ("Hm2" $!v (Next (<Send 2> iMsg_base_def #x True END)))%proto. - iExists (<Send 2> iMsg_base_def #x True END)%proto. + iSpecialize ("Hm2" $!v (Next (<(Send, 2)> iMsg_base_def #x True END)))%proto. + iExists (<(Send, 2)> iMsg_base_def #x True END)%proto. iRewrite -"Hm2". simpl. iSplitL. @@ -210,8 +210,8 @@ Proof. iSpecialize ("Hm1" $!v (Next p1)). iRewrite -"Hm1" in "Hm1'". iDestruct "Hm1'" as (Heq) "[#Hm1' _]". - iSpecialize ("Hm2" $!v (Next (<Send 0> iMsg_base_def #x True END)))%proto. - iExists (<Send 0> iMsg_base_def #x True END)%proto. + iSpecialize ("Hm2" $!v (Next (<(Send, 0)> iMsg_base_def #x True END)))%proto. + iExists (<(Send, 0)> iMsg_base_def #x True END)%proto. iRewrite -"Hm2". simpl. iSplitL. @@ -377,12 +377,12 @@ Section example4. Context `{!heapGS Σ}. Definition iProto_example4 : gmap nat (iProto Σ) := - <[0 := (<(Send 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; - <(Recv 2)> MSG #() {{ l ↦ #(x+2) }} ; END)%proto]> - (<[1 := (<(Recv 0) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; - <(Send 2)> MSG #l {{ l ↦ #(x+1) }}; END)%proto]> - (<[2 := (<(Recv 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; - <(Send 0)> MSG #() {{ l ↦ #(x+1) }}; END)%proto]> + <[0 := (<(Send, 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; + <(Recv, 2)> MSG #() {{ l ↦ #(x+2) }} ; END)%proto]> + (<[1 := (<(Recv, 0) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; + <(Send, 2)> MSG #l {{ l ↦ #(x+1) }}; END)%proto]> + (<[2 := (<(Recv, 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; + <(Send, 0)> MSG #() {{ l ↦ #(x+1) }}; END)%proto]> ∅)). Lemma iProto_example4_consistent : @@ -434,10 +434,10 @@ Section example4. rewrite !iMsg_exist_eq. iRewrite -"Hm1" in "Hm1'". iDestruct "Hm1'" as (l x <-) "[#Hm1' Hl]". simpl. - iSpecialize ("Hm2" $!#l (Next (<Send 2> iMsg_base_def #l + iSpecialize ("Hm2" $!#l (Next (<(Send, 2)> iMsg_base_def #l (l ↦ #(x+1)) END)))%proto. Unshelve. 2-4: apply _. (* Why is this needed? *) - iExists (<Send 2> iMsg_base_def #l (l ↦ #(x+1)) END)%proto. + iExists (<(Send, 2)> iMsg_base_def #l (l ↦ #(x+1)) END)%proto. simpl. iSplitL. { iRewrite -"Hm2". iExists l, x. iSplit; [done|]. iFrame. done. } @@ -489,8 +489,8 @@ Section example4. iRewrite -"Hm1" in "Hm1'". iDestruct "Hm1'" as (<-) "[#Hm1' Hl]". simpl. - iSpecialize ("Hm2" $!#l (Next (<Send 0> iMsg_base_def #() (l ↦ #(x+1+1)) END)))%proto. - iExists (<Send 0> iMsg_base_def #() (l ↦ #(x+1+1)) END)%proto. + iSpecialize ("Hm2" $!#l (Next (<(Send, 0)> iMsg_base_def #() (l ↦ #(x+1+1)) END)))%proto. + iExists (<(Send, 0)> iMsg_base_def #() (l ↦ #(x+1+1)) END)%proto. Unshelve. 2-4: apply _. iRewrite -"Hm2". simpl. diff --git a/theories/channel/multi_proto_model.v b/theories/channel/multi_proto_model.v index 9c6c8be..e32bc26 100644 --- a/theories/channel/multi_proto_model.v +++ b/theories/channel/multi_proto_model.v @@ -39,12 +39,13 @@ From actris.utils Require Import cofe_solver_2. Set Default Proof Using "Type". Module Export action. - Inductive action := Send (n:nat) | Recv (n:nat). - Global Instance action_inhabited : Inhabited action := populate (Send 0). + Inductive tag := Send | Recv. + Definition action : Set := tag * nat. + Global Instance action_inhabited : Inhabited action := populate (Send,0). Definition action_dual (a : action) : action := - match a with Send n => Recv n | Recv n => Send n end. + match a with (Send, n) => (Recv, n) | (Recv, n) => (Send, n) end. Global Instance action_dual_involutive : Involutive (=) action_dual. - Proof. by intros []. Qed. + Proof. by intros [[]]. Qed. Canonical Structure actionO := leibnizO action. End action. -- GitLab From 652f69a544355b423f7ea9b26f29db2e38e6dd08 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Wed, 24 Jan 2024 14:28:33 +0100 Subject: [PATCH 18/81] Subprotocols done --- theories/channel/multi_channel.v | 156 +++++++++++-------------------- theories/channel/multi_proto.v | 78 +++++++++++----- 2 files changed, 109 insertions(+), 125 deletions(-) diff --git a/theories/channel/multi_channel.v b/theories/channel/multi_channel.v index 40f61e1..17cb350 100644 --- a/theories/channel/multi_channel.v +++ b/theories/channel/multi_channel.v @@ -104,7 +104,7 @@ Definition chan_inv `{!heapGS Σ, !chanG Σ} γ γE γt i j (l:loc) : iProp Σ : Definition iProto_pointsto_def `{!heapGS Σ, !chanG Σ} (c : val) (p : iProto Σ) : iProp Σ := - ∃ γ γE1 γE2 γt1 γt2 i (l:loc) ls, + ∃ γ γE1 γE2 γt1 γt2 i (l:loc) ls p', ⌜ c = PairV #(length ls) #l ⌠∗ inv (nroot.@"ctx") (iProto_ctx γ) ∗ l ↦∗ ls ∗ @@ -113,8 +113,9 @@ Definition iProto_pointsto_def `{!heapGS Σ, !chanG Σ} ⌜v = PairV #l1 #l2⌠∗ inv (nroot.@"p") (chan_inv γ γE1 γt1 i j l1) ∗ inv (nroot.@"p") (chan_inv γ γE2 γt2 j i l2)) ∗ - own γE1 (â—E (Next p)) ∗ own γE1 (â—¯E (Next p)) ∗ - iProto_own γ i p. + â–· (p' ⊑ p) ∗ + own γE1 (â—E (Next p')) ∗ own γE1 (â—¯E (Next p')) ∗ + iProto_own γ i p'. Definition iProto_pointsto_aux : seal (@iProto_pointsto_def). by eexists. Qed. Definition iProto_pointsto := iProto_pointsto_aux.(unseal). Definition iProto_pointsto_eq : @@ -134,12 +135,15 @@ Section channel. Global Instance iProto_pointsto_proper c : Proper ((≡) ==> (≡)) (iProto_pointsto c). Proof. apply (ne_proper _). Qed. - (* Lemma iProto_pointsto_le c p1 p2 : c ↣ p1 ⊢ â–· (p1 ⊑ p2) -∗ c ↣ p2. *) - (* Proof. *) - (* rewrite iProto_pointsto_eq. iDestruct 1 as (γ s l r lk ->) "[Hlk H]". *) - (* iIntros "Hle'". iExists γ, s, l, r, lk. iSplit; [done|]. iFrame "Hlk". *) - (* by iApply (iProto_own_le with "H"). *) - (* Qed. *) + Lemma iProto_pointsto_le c p1 p2 : c ↣ p1 ⊢ â–· (p1 ⊑ p2) -∗ c ↣ p2. + Proof. + rewrite iProto_pointsto_eq. + iDestruct 1 as + (γ γE1 γE2 γt1 γt2 i l ls p ->) "(#IH & Hl & Hls & Hle & Hâ— & Hâ—¯ & Hown)". + iIntros "Hle'". iExists γ, γE1, γE2, γt1, γt2, i, l, ls, p. + iSplit; [done|]. iFrame "#∗". + iApply (iProto_le_trans with "Hle Hle'"). + Qed. (** ** Specifications of [send] and [recv] *) Lemma new_chan_spec (n:nat) (ps:gmap nat (iProto Σ)) : @@ -155,7 +159,9 @@ Section channel. own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p1))) -∗ own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p2))) -∗ False. - Proof. Admitted. + Proof. + iIntros "Hi Hj". by iDestruct (own_prot_idx with "Hi Hj") as %Hneq. + Qed. Lemma send_spec c j v p : {{{ c ↣ <(Send, j)> MSG v; p }}} @@ -164,11 +170,11 @@ Section channel. Proof. rewrite iProto_pointsto_eq. iIntros (Φ) "Hc HΦ". wp_lam; wp_pures. iDestruct "Hc" as - (γ γE1 γE2 γt1 γt2 i l ls ->) "(#IH & Hl & Hls & Hâ— & Hâ—¯ & Hown)". + (γ γE1 γE2 γt1 γt2 i l ls p' ->) "(#IH & Hl & Hls & Hle & Hâ— & Hâ—¯ & Hown)". wp_pures. case_bool_decide; last first. { wp_pures. iClear "IH Hâ— Hâ—¯ Hown HΦ Hls Hl". - iLöb as "IH". wp_lam. iApply "IH". } + iLöb as "IH". wp_lam. iApply "IH". done. } assert (is_Some (ls !! j)) as [l' HSome]. { apply lookup_lt_is_Some_2. lia. } wp_pures. @@ -187,18 +193,24 @@ Section channel. - iDestruct "HIp" as (? m) "(>Hl' & Hown' & HIp)". wp_store. rewrite /iProto_own. + iDestruct "Hown" as (p'') "[Hle' Hown]". + iDestruct "Hown'" as (p''') "[Hle'' Hown']". iDestruct (own_prot_excl with "Hown Hown'") as "H". done. - - iDestruct "HIp" as (p') "(>Hl' & Hown' & HIp)". + - iDestruct "HIp" as (p'') "(>Hl' & Hown' & HIp)". wp_store. rewrite /iProto_own. + iDestruct "Hown" as (p''') "[Hle' Hown]". + iDestruct "Hown'" as (p'''') "[Hle'' Hown']". iDestruct (own_prot_excl with "Hown Hown'") as "H". done. } iDestruct "HIp" as "[>Hl' Htok]". wp_store. iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]". { apply excl_auth_update. } iModIntro. - iSplitL "Hl' Hâ— Hown". + iSplitL "Hl' Hâ— Hown Hle". { iRight. iLeft. iIntros "!>". iExists _, _. iFrame. + iSplitL "Hown Hle". + { iApply (iProto_own_le with "Hown Hle"). } iExists _. iFrame. rewrite iMsg_base_eq. simpl. done. } wp_pures. iLöb as "HL". @@ -214,7 +226,7 @@ Section channel. iSplitL "Hl' Hown HIp". { iRight. iLeft. iExists _, _. iFrame. } wp_pures. iApply ("HL" with "HΦ Hl Hls Htok Hâ—¯"). - - iDestruct "HIp" as (p') "(>Hl' & Hown & Hâ—)". + - iDestruct "HIp" as (p'') "(>Hl' & Hown & Hâ—)". wp_load. iModIntro. iSplitL "Hl' Htok". @@ -226,91 +238,28 @@ Section channel. { apply excl_auth_update. } iModIntro. iApply "HΦ". - iExists _, _, _, _, _, _, _, _. + iExists _,_, _, _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". - iRewrite -"Hagree'". done. + iRewrite -"Hagree'". iApply iProto_le_refl. Qed. - Lemma send_spec_tele {TT} c j (tt : TT) + Lemma send_spec_tele {TT} c i (tt : TT) (v : TT → val) (P : TT → iProp Σ) (p : TT → iProto Σ) : - {{{ c ↣ (<(Send, j) @.. x > MSG v x {{ P x }}; p x) ∗ P tt }}} - send c #j (v tt) + {{{ c ↣ (<(Send,i) @.. x > MSG v x {{ P x }}; p x) ∗ P tt }}} + send c #i (v tt) {{{ RET #(); c ↣ (p tt) }}}. Proof. - rewrite iProto_pointsto_eq. iIntros (Φ) "[Hc HP] HΦ". wp_lam; wp_pures. - iDestruct "Hc" as - (γ γE1 γE2 γt1 γt2 i l ls ->) "(#IH & Hl & Hls & Hâ— & Hâ—¯ & Hown)". - wp_pures. - case_bool_decide; last first. - { wp_pures. iClear "IH Hâ— Hâ—¯ Hown HΦ Hls Hl". - iLöb as "IH". wp_lam. iApply "IH". iFrame. } - assert (is_Some (ls !! j)) as [l' HSome]. - { apply lookup_lt_is_Some_2. lia. } - wp_pures. - wp_smart_apply (wp_load_offset with "Hl"). - { done. } - iIntros "Hl". wp_pures. - iDestruct (big_sepL_lookup_acc with "Hls") as "[Hj Hls]"; [done|]. - iDestruct "Hj" as (l1 l2 ->) "#[IHl1 IHl2]". - iDestruct ("Hls" with "[]") as "Hls". - { iExists _, _. iFrame "#". done. } - wp_pures. - wp_bind (Store _ _). - iInv "IHl1" as "HIp". - iDestruct "HIp" as "[HIp|HIp]"; last first. - { iDestruct "HIp" as "[HIp|HIp]". - - iDestruct "HIp" as (? m) "(>Hl' & Hown' & HIp)". - wp_store. - rewrite /iProto_own. - iDestruct (own_prot_excl with "Hown Hown'") as "H". done. - - iDestruct "HIp" as (p') "(>Hl' & Hown' & HIp)". - wp_store. - rewrite /iProto_own. - iDestruct (own_prot_excl with "Hown Hown'") as "H". done. } - iDestruct "HIp" as "[>Hl' Htok]". - wp_store. - iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]". - { apply excl_auth_update. } - iModIntro. - iSplitL "Hl' Hâ— Hown HP". - { iRight. iLeft. iIntros "!>". iExists _, _. iFrame. - iExists _. iFrame. rewrite iMsg_base_eq. simpl. - iApply iMsg_texist_exist. - simpl. iExists tt. - iSplit; [done|]. - iSplit; [done|]. - done. } - wp_pures. - iLöb as "HL". - wp_lam. - wp_bind (Load _). - iInv "IHl1" as "HIp". - iDestruct "HIp" as "[HIp|HIp]". - { iDestruct "HIp" as ">[Hl' Htok']". - iDestruct (own_valid_2 with "Htok Htok'") as %[]. } - iDestruct "HIp" as "[HIp|HIp]". - - iDestruct "HIp" as (? m) "(>Hl' & Hown & HIp)". - wp_load. iModIntro. - iSplitL "Hl' Hown HIp". - { iRight. iLeft. iExists _, _. iFrame. } - wp_pures. iApply ("HL" with "HΦ Hl Hls Htok Hâ—¯"). - - iDestruct "HIp" as (p') "(>Hl' & Hown & Hâ—)". - wp_load. - iModIntro. - iSplitL "Hl' Htok". - { iLeft. iFrame. } - iDestruct (own_valid_2 with "Hâ— Hâ—¯") as "#Hagree". - iDestruct (excl_auth_agreeI with "Hagree") as "Hagree'". - wp_pures. - iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]". - { apply excl_auth_update. } - iModIntro. - iApply "HΦ". - iExists _, _, _, _, _, _, _, _. - iSplit; [done|]. iFrame "#∗". - iRewrite -"Hagree'". done. + iIntros (Φ) "[Hc HP] HΦ". + iDestruct (iProto_pointsto_le _ _ (<(Send,i)> MSG v tt; p tt)%proto with "Hc [HP]") + as "Hc". + { iIntros "!>". + iApply iProto_le_trans. + iApply iProto_le_texist_intro_l. + by iApply iProto_le_payload_intro_l. } + by iApply (send_spec with "Hc"). Qed. + Lemma recv_spec {TT} c j (v : TT → val) (P : TT → iProp Σ) (p : TT → iProto Σ) : {{{ c ↣ <(Recv, j) @.. x> MSG v x {{ P x }}; p x }}} recv c #j @@ -319,11 +268,11 @@ Section channel. iIntros (Φ) "Hc HΦ". iLöb as "HL". wp_lam. rewrite iProto_pointsto_eq. iDestruct "Hc" as - (γ γE1 γE2 γt1 γt2 i l ls ->) "(#IH & Hl & Hls & Hâ— & Hâ—¯ & Hown)". + (γ γE1 γE2 γt1 γt2 i l ls p' ->) "(#IH & Hl & Hls & Hle & Hâ— & Hâ—¯ & Hown)". wp_pures. case_bool_decide; last first. { wp_pures. iClear "IH Hâ— Hâ—¯ Hown HΦ Hls Hl". - iLöb as "IH". wp_lam. iApply "IH". } + iLöb as "IH". wp_lam. iApply "IH". done. } wp_pures. assert (is_Some (ls !! j)) as [l' HSome]. { apply lookup_lt_is_Some_2. lia. } @@ -343,22 +292,23 @@ Section channel. iModIntro. iSplitL "Hl' Htok". { iLeft. iFrame. } - wp_pures. iApply ("HL" with "[Hâ— Hâ—¯ Hown Hls Hl] HΦ"). - iExists _, _, _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". } + wp_pures. iApply ("HL" with "[Hâ— Hâ—¯ Hown Hls Hl Hle] HΦ"). + iExists _, _, _, _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". } iDestruct "HIp" as "[HIp|HIp]"; last first. - { iDestruct "HIp" as (p') "[>Hl' [Hown' Hâ—¯']]". + { iDestruct "HIp" as (p'') "[>Hl' [Hown' Hâ—¯']]". wp_xchg. iModIntro. iSplitL "Hl' Hown' Hâ—¯'". { iRight. iRight. iExists _. iFrame. } - wp_pures. iApply ("HL" with "[Hâ— Hâ—¯ Hown Hls Hl] HΦ"). - iExists _, _, _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". } + wp_pures. iApply ("HL" with "[Hâ— Hâ—¯ Hown Hls Hl Hle] HΦ"). + iExists _, _, _, _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". } iDestruct "HIp" as (w m) "(>Hl' & Hown' & HIp)". - iDestruct "HIp" as (p') "[Hm Hp']". + iDestruct "HIp" as (p'') "[Hm Hp']". iInv "IH" as "Hctx". wp_xchg. + iDestruct (iProto_own_le with "Hown Hle") as "Hown". iMod (iProto_step with "Hctx Hown' Hown Hm") as - (p'') "(Hm & Hctx & Hown & Hown')". + (p''') "(Hm & Hctx & Hown & Hown')". iModIntro. iSplitL "Hctx"; [done|]. iModIntro. @@ -372,8 +322,8 @@ Section channel. { apply excl_auth_update. } iModIntro. iApply "HΦ". iFrame. - iExists _, _, _, _, _, _, _, _. iSplit; [done|]. - iRewrite "Hp". iFrame "#∗". + iExists _, _, _, _, _, _, _, _, _. iSplit; [done|]. + iRewrite "Hp". iFrame "#∗". iApply iProto_le_refl. Qed. End channel. diff --git a/theories/channel/multi_proto.v b/theories/channel/multi_proto.v index a24b94e..bf04ec6 100644 --- a/theories/channel/multi_proto.v +++ b/theories/channel/multi_proto.v @@ -633,7 +633,7 @@ Definition iProto_ctx `{protoG Σ V} (** * The connective for ownership of channel ends *) Definition iProto_own `{!protoG Σ V} (γ : gname) (i : nat) (p : iProto Σ V) : iProp Σ := - iProto_own_frag γ i p. + ∃ p', â–· (p' ⊑ p) ∗ iProto_own_frag γ i p'. Arguments iProto_own {_ _ _} _ _ _%proto. Global Instance: Params (@iProto_own) 3 := {}. @@ -757,13 +757,15 @@ Section proto. Qed. Lemma iProto_consistent_le ps i p1 p2 : - ps !! i = Some p1 → - p1 ⊑ p2 -∗ iProto_consistent ps -∗ + ps !!! i ≡ p1 -∗ + p1 ⊑ p2 -∗ iProto_consistent (<[i := p2]>ps). Proof. - iIntros (HSome) "Hle Hprot". - iLöb as "IH" forall (p1 p2 ps HSome). + iIntros "Hprot #HSome Hle". + iRevert "HSome". + iLöb as "IH" forall (p1 p2 ps). + iIntros "#HSome". rewrite !iProto_consistent_unfold. iIntros (i' j' m1 m2) "#Hm1 #Hm2". destruct (decide (i = i')) as [<-|Hneq]. @@ -788,13 +790,13 @@ Section proto. { rewrite lookup_total_insert. rewrite iProto_message_equivI. iDestruct "Hm2" as "[%Heq _]". done. } iDestruct ("Hprot" $!i j' with "[] [] H") as "Hprot". - { iRewrite -"Heq". rewrite !lookup_total_alt. rewrite HSome. done. } + { iRewrite -"Heq". rewrite !lookup_total_alt. iRewrite "HSome". done. } { rewrite lookup_total_insert_ne; [|done]. done. } iDestruct "Hprot" as (p'') "[Hm Hprot]". iExists p''. iFrame. iNext. - iDestruct ("IH" with "[] Hle Hprot") as "HI". - { iPureIntro. by rewrite lookup_insert. } + iDestruct ("IH" with "Hprot Hle [HSome]") as "HI". + { by rewrite lookup_total_insert. } iClear "IH Hm1 Hm2 Heq". rewrite insert_insert. rewrite (insert_commute _ j' i); [|done]. @@ -816,17 +818,16 @@ Section proto. iDestruct (iProto_le_recv_inv_r with "Hle") as "Hle". (* iRewrite -"Hm2" in "Hm2'". *) iDestruct "Hle" as (m') "[#Heq Hle]". - iDestruct ("Hprot" $!i' with "[] [] Hm1'") as "Hprot". { done. } - { rewrite !lookup_total_alt. rewrite HSome. done. } + { rewrite !lookup_total_alt. iRewrite "HSome". done. } iDestruct ("Hprot") as (p') "[Hm1' Hprot]". iDestruct ("Hle" with "Hm1'") as (p2') "[Hle Hm']". iSpecialize ("Hm2" $! v (Next p2')). iExists p2'. iRewrite -"Hm2". iFrame. - iDestruct ("IH" with "[] Hle Hprot") as "HI". - { iPureIntro. rewrite lookup_insert_ne; [|done]. rewrite lookup_insert. done. } + iDestruct ("IH" with "Hprot Hle []") as "HI". + { iPureIntro. rewrite lookup_total_insert_ne; [|done]. rewrite lookup_total_insert. done. } rewrite insert_commute; [|done]. rewrite !insert_insert. done. } rewrite lookup_total_insert_ne; [|done]. @@ -837,10 +838,9 @@ Section proto. iNext. rewrite (insert_commute _ j' i); [|done]. rewrite (insert_commute _ i' i); [|done]. - iApply ("IH" with "[] Hle Hprot"). - iPureIntro. - rewrite lookup_insert_ne; [|done]. - rewrite lookup_insert_ne; [|done]. + iApply ("IH" with "Hprot Hle []"). + rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. done. Qed. @@ -1060,7 +1060,16 @@ Section proto. - done. } iFrame. iModIntro. - rewrite -fmap_insert. iFrame. + rewrite -fmap_insert. + iFrame. + iExists _. iFrame. iApply iProto_le_refl. + Qed. + + Lemma iProto_own_le γ s p1 p2 : + iProto_own γ s p1 -∗ â–· (p1 ⊑ p2) -∗ iProto_own γ s p2. + Proof. + iDestruct 1 as (p1') "[Hle H]". iIntros "Hle'". + iExists p1'. iFrame "H". by iApply (iProto_le_trans with "Hle"). Qed. Lemma iProto_init ps : @@ -1072,6 +1081,12 @@ Section proto. iExists γ. iFrame. iExists _. iFrame. done. Qed. + Lemma own_prot_idx γ i j (p1 p2 : iProto Σ V) : + own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p1))) -∗ + own γ (gmap_view_frag j (DfracOwn 1) (Excl' (Next p2))) -∗ + ⌜i ≠jâŒ. + Proof. Admitted. + Lemma iProto_step γ i j m1 m2 p1 v : iProto_ctx γ -∗ iProto_own γ i (<(Send, j)> m1) -∗ @@ -1081,14 +1096,33 @@ Section proto. iProto_own γ i p1 ∗ iProto_own γ j p2. Proof. iIntros "Hctx Hi Hj Hm". + iDestruct "Hi" as (pi) "[Hile Hi]". + iDestruct "Hj" as (pj) "[Hjle Hj]". iDestruct "Hctx" as (ps) "[Hauth Hconsistent]". iDestruct (iProto_own_auth_agree with "Hauth Hi") as "#Hpi". iDestruct (iProto_own_auth_agree with "Hauth Hj") as "#Hpj". - iDestruct (iProto_consistent_step with "Hconsistent [//] [//] [Hm //]") as + iDestruct (own_prot_idx with "Hi Hj") as %Hneq. + iAssert (â–· (<[i:=<(Send, j)> m1]>ps !!! j ≡ pj))%I as "Hpj'". + { rewrite lookup_total_insert_ne; done. } + iDestruct (iProto_consistent_le with "Hconsistent Hpi Hile") as "Hconsistent". + iDestruct (iProto_consistent_le with "Hconsistent Hpj' Hjle") as "Hconsistent". + iDestruct (iProto_consistent_step _ _ _ i j with "Hconsistent [] [] [Hm //]") as (p2) "[Hm2 Hconsistent]". + { rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert_ne; [|done]. + iNext. rewrite lookup_total_insert. done. } + { rewrite lookup_total_insert_ne; [|done]. + iNext. rewrite lookup_total_insert. done. } iMod (iProto_own_auth_update _ _ _ _ p2 with "Hauth Hj") as "[Hauth Hj]". iMod (iProto_own_auth_update _ _ _ _ p1 with "Hauth Hi") as "[Hauth Hi]". - iIntros "!>!>". iExists p2. iFrame. iExists _. iFrame. + iIntros "!>!>". iExists p2. iFrame. + iSplitL "Hconsistent Hauth". + { iExists _. iFrame. rewrite insert_insert. + rewrite insert_commute; [|done]. rewrite insert_insert. + rewrite insert_commute; [|done]. done. } + iSplitL "Hi". + - iExists _. iFrame. iApply iProto_le_refl. + - iExists _. iFrame. iApply iProto_le_refl. Qed. (* (** The instances below make it possible to use the tactics [iIntros], *) @@ -1164,7 +1198,7 @@ Section proto. End proto. -(* Typeclasses Opaque iProto_ctx iProto_own. *) +Typeclasses Opaque iProto_ctx iProto_own. -(* Global Hint Extern 0 (environments.envs_entails _ (?x ⊑ ?y)) => *) -(* first [is_evar x; fail 1 | is_evar y; fail 1|iApply iProto_le_refl] : core. *) +Global Hint Extern 0 (environments.envs_entails _ (?x ⊑ ?y)) => + first [is_evar x; fail 1 | is_evar y; fail 1|iApply iProto_le_refl] : core. -- GitLab From 56371c908a45f88496c3bf2def2186d4234b8cb4 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Wed, 24 Jan 2024 15:54:13 +0100 Subject: [PATCH 19/81] Added get_chan primitive --- theories/channel/multi_channel.v | 48 ++++++++++- .../multi_proto_consistency_examples.v | 83 ++++++------------- 2 files changed, 70 insertions(+), 61 deletions(-) diff --git a/theories/channel/multi_channel.v b/theories/channel/multi_channel.v index 17cb350..d8d0f7d 100644 --- a/theories/channel/multi_channel.v +++ b/theories/channel/multi_channel.v @@ -41,6 +41,9 @@ Definition new_chan : val := ("xxs" +â‚— "i") <- "xs";; "go1" ("i"+#1)) #0;; "xxs". +Definition get_chan : val := + λ: "cs" "i", ! ("cs" +â‚— "i"). + Definition wait : val := rec: "go" "c" := match: !"c" with @@ -125,6 +128,13 @@ Global Instance: Params (@iProto_pointsto) 5 := {}. Notation "c ↣ p" := (iProto_pointsto c p) (at level 20, format "c ↣ p"). +Definition chan_pool `{!heapGS Σ, !chanG Σ} + (cs : val) (ps : gmap nat (iProto Σ)) : iProp Σ := + ∃ (l:loc) (ls : list val), + ⌜cs = #l⌠∗ ⌜∀ i, is_Some (ps !! i) → is_Some (ls !! i)⌠∗ + l ↦∗ ls ∗ + [∗list] i ↦ c ∈ ls, (∀ p, ⌜ps !! i = Some p⌠-∗ c ↣ p). + Section channel. Context `{!heapGS Σ, !chanG Σ}. Implicit Types p : iProto Σ. @@ -148,13 +158,45 @@ Section channel. (** ** Specifications of [send] and [recv] *) Lemma new_chan_spec (n:nat) (ps:gmap nat (iProto Σ)) : (∀ i, i < n → is_Some (ps !! i)) → + n = (size (dom ps)) → {{{ iProto_consistent ps }}} new_chan #n - {{{ cs ls, RET #cs; - ⌜length ls = n⌠∗ cs ↦∗ ls ∗ - [∗list] i ↦ l ∈ ls, l ↣ (ps !!! i) }}}. + {{{ cs, RET cs; chan_pool cs ps }}}. Proof. Admitted. + Lemma get_chan_spec cs (i:nat) ps p : + ps !! i = Some p → + {{{ chan_pool cs ps }}} + get_chan cs #i + {{{ c, RET c; c ↣ p ∗ chan_pool cs (delete i ps) }}}. + Proof. + iIntros (HSome Φ) "Hcs HΦ". + iDestruct "Hcs" as (l ls -> Hlen) "[Hl Hls]". + wp_lam. + assert (is_Some (ls !! i)) as [c HSome']. + { by apply Hlen. } + wp_smart_apply (wp_load_offset with "Hl"); [done|]. + iIntros "Hcs". + iApply "HΦ". + iDestruct (big_sepL_delete' _ _ i with "Hls") as "[Hc Hls]"; [set_solver|]. + iDestruct ("Hc" with "[//]") as "Hc". + iFrame. + iExists _, _. iSplit; [done|]. iFrame "Hcs". + iSplitR. + { iPureIntro. intros j HSome''. + destruct (decide (i=j)) as [<-|Hneq]. + { rewrite lookup_delete in HSome''. done. } + rewrite lookup_delete_ne in HSome''; [|done]. + by apply Hlen. } + iApply (big_sepL_impl with "Hls"). + iIntros "!>" (j v Hin) "H". + iIntros (p' HSome''). + destruct (decide (i=j)) as [<-|Hneq]. + { rewrite lookup_delete in HSome''. done. } + rewrite lookup_delete_ne in HSome''; [|done]. + by iApply "H". + Qed. + Lemma own_prot_excl γ i (p1 p2 : iProto Σ) : own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p1))) -∗ own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p2))) -∗ diff --git a/theories/channel/multi_proto_consistency_examples.v b/theories/channel/multi_proto_consistency_examples.v index f9ec28f..2031476 100644 --- a/theories/channel/multi_proto_consistency_examples.v +++ b/theories/channel/multi_proto_consistency_examples.v @@ -2,7 +2,6 @@ From iris.algebra Require Import gmap excl_auth gmap_view. From iris.proofmode Require Import proofmode. From iris.base_logic Require Export lib.iprop. From iris.base_logic Require Import lib.own. -From iris.program_logic Require Import language. From actris.channel Require Import multi_proto_model multi_proto multi_channel. Set Default Proof Using "Type". Export action. @@ -295,9 +294,9 @@ Qed. Definition roundtrip_prog : val := λ: <>, let: "cs" := new_chan #3 in - let: "c0" := ! ("cs" +â‚— #0) in - let: "c1" := ! ("cs" +â‚— #1) in - let: "c2" := ! ("cs" +â‚— #2) in + let: "c0" := get_chan "cs" #0 in + let: "c1" := get_chan "cs" #1 in + let: "c2" := get_chan "cs" #2 in Fork (let: "x" := recv "c1" #0 in send "c1" #2 "x");; Fork (let: "x" := recv "c2" #1 in send "c2" #0 "x");; send "c0" #1 #42;; recv "c0" #2. @@ -307,38 +306,23 @@ Section channel. Implicit Types p : iProto Σ. Implicit Types TT : tele. + + (* TODO: Fix nat/Z coercion. *) Lemma roundtrip_prog_spec : {{{ True }}} roundtrip_prog #() {{{ RET #42 ; True }}}. Proof using chanG0 heapGS0 Σ. iIntros (Φ) "_ HΦ". wp_lam. - wp_smart_apply (new_chan_spec 3 iProto_example3 with "[]"). + wp_smart_apply (new_chan_spec 3 iProto_example3). { intros i Hle. destruct i as [|[|[]]]; try set_solver. lia. } + { set_solver. } { iApply iProto_example3_consistent. } - iIntros (cs ls) "[%Hlen [Hcs Hls]]". - assert (is_Some (ls !! 0)) as [c0 HSome0]. - { apply lookup_lt_is_Some_2. lia. } - assert (is_Some (ls !! 1)) as [c1 HSome1]. - { apply lookup_lt_is_Some_2. lia. } - assert (is_Some (ls !! 2)) as [c2 HSome2]. - { apply lookup_lt_is_Some_2. lia. } - wp_smart_apply (wp_load_offset _ _ _ _ 0 with "Hcs"); [done|]. - iIntros "Hcs". - wp_smart_apply (wp_load_offset _ _ _ _ 1 with "Hcs"); [done|]. - iIntros "Hcs". - wp_smart_apply (wp_load_offset _ _ _ _ 2 with "Hcs"); [done|]. - iIntros "Hcs". - iDestruct (big_sepL_delete' _ _ 0 with "Hls") as "[Hc0 Hls]"; [set_solver|]. - iDestruct (big_sepL_delete' _ _ 1 with "Hls") as "[Hc1 Hls]"; [set_solver|]. - iDestruct (big_sepL_delete' _ _ 2 with "Hls") as "[Hc2 _]"; [set_solver|]. - iDestruct ("Hc1" with "[//]") as "Hc1". - iDestruct ("Hc2" with "[//] [//]") as "Hc2". - rewrite /iProto_example3. - rewrite !lookup_total_insert. - rewrite lookup_total_insert_ne; [|done]. - rewrite !lookup_total_insert. - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - rewrite !lookup_total_insert. + iIntros (cs) "Hcs". + wp_smart_apply (get_chan_spec _ 0 with "Hcs"); [set_solver|]. + iIntros (c0) "[Hc0 Hcs]". + wp_smart_apply (get_chan_spec _ 1 with "Hcs"); [set_solver|]. + iIntros (c1) "[Hc1 Hcs]". + wp_smart_apply (get_chan_spec _ 2 with "Hcs"); [set_solver|]. + iIntros (c2) "[Hc2 Hcs]". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". wp_smart_apply @@ -578,9 +562,9 @@ End example4. Definition roundtrip_ref_prog : val := λ: <>, let: "cs" := new_chan #3 in - let: "c0" := ! ("cs" +â‚— #0) in - let: "c1" := ! ("cs" +â‚— #1) in - let: "c2" := ! ("cs" +â‚— #2) in + let: "c0" := get_chan "cs" #0 in + let: "c1" := get_chan "cs" #1 in + let: "c2" := get_chan "cs" #2 in Fork (let: "l" := recv "c1" #0 in "l" <- !"l" + #1;; send "c1" #2 "l");; Fork (let: "l" := recv "c2" #1 in "l" <- !"l" + #1;; send "c2" #0 #());; let: "l" := ref #40 in send "c0" #1 "l";; recv "c0" #2;; !"l". @@ -596,32 +580,15 @@ Section proof. iIntros (Φ) "_ HΦ". wp_lam. wp_smart_apply (new_chan_spec 3 iProto_example4 with "[]"). { intros i Hle. destruct i as [|[|[]]]; try set_solver. lia. } + { set_solver. } { iApply iProto_example4_consistent. } - iIntros (cs ls) "[%Hlen [Hcs Hls]]". - assert (is_Some (ls !! 0)) as [c0 HSome0]. - { apply lookup_lt_is_Some_2. lia. } - assert (is_Some (ls !! 1)) as [c1 HSome1]. - { apply lookup_lt_is_Some_2. lia. } - assert (is_Some (ls !! 2)) as [c2 HSome2]. - { apply lookup_lt_is_Some_2. lia. } - wp_smart_apply (wp_load_offset _ _ _ _ 0 with "Hcs"); [done|]. - iIntros "Hcs". - wp_smart_apply (wp_load_offset _ _ _ _ 1 with "Hcs"); [done|]. - iIntros "Hcs". - wp_smart_apply (wp_load_offset _ _ _ _ 2 with "Hcs"); [done|]. - iIntros "Hcs". - iDestruct (big_sepL_delete' _ _ 0 with "Hls") as "[Hc0 Hls]"; [set_solver|]. - iDestruct (big_sepL_delete' _ _ 1 with "Hls") as "[Hc1 Hls]"; [set_solver|]. - iDestruct (big_sepL_delete' _ _ 2 with "Hls") as "[Hc2 _]"; [set_solver|]. - iDestruct ("Hc1" with "[//]") as "Hc1". - iDestruct ("Hc2" with "[//] [//]") as "Hc2". - rewrite /iProto_example4. - rewrite !lookup_total_insert. - rewrite lookup_total_insert_ne; [|done]. - rewrite !lookup_total_insert. - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - rewrite !lookup_total_insert. + iIntros (cs) "Hcs". + wp_smart_apply (get_chan_spec _ 0 with "Hcs"); [set_solver|]. + iIntros (c0) "[Hc0 Hcs]". + wp_smart_apply (get_chan_spec _ 1 with "Hcs"); [set_solver|]. + iIntros (c1) "[Hc1 Hcs]". + wp_smart_apply (get_chan_spec _ 2 with "Hcs"); [set_solver|]. + iIntros (c2) "[Hc2 Hcs]". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". wp_smart_apply -- GitLab From 3d41696f3d19c8f2fc64216ccb3f79f2fe37bec1 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Wed, 24 Jan 2024 16:02:03 +0100 Subject: [PATCH 20/81] Bumped subprotocol typeclasses --- theories/channel/multi_proto.v | 136 ++++++++++++++++----------------- 1 file changed, 68 insertions(+), 68 deletions(-) diff --git a/theories/channel/multi_proto.v b/theories/channel/multi_proto.v index bf04ec6..199c5f8 100644 --- a/theories/channel/multi_proto.v +++ b/theories/channel/multi_proto.v @@ -1127,74 +1127,74 @@ Section proto. (* (** The instances below make it possible to use the tactics [iIntros], *) (* [iExist], [iSplitL]/[iSplitR], [iFrame] and [iModIntro] on [iProto_le] goals. *) *) - (* Global Instance iProto_le_from_forall_l {A} a (m1 : A → iMsg Σ V) m2 name : *) - (* AsIdentName m1 name → *) - (* FromForall (iProto_message Recv (iMsg_exist m1) ⊑ (<a> m2)) *) - (* (λ x, (<?> m1 x) ⊑ (<a> m2))%I name | 10. *) - (* Proof. intros _. apply iProto_le_exist_elim_l. Qed. *) - (* Global Instance iProto_le_from_forall_r {A} a m1 (m2 : A → iMsg Σ V) name : *) - (* AsIdentName m2 name → *) - (* FromForall ((<a> m1) ⊑ iProto_message Send (iMsg_exist m2)) *) - (* (λ x, (<a> m1) ⊑ (<!> m2 x))%I name | 11. *) - (* Proof. intros _. apply iProto_le_exist_elim_r. Qed. *) - - (* Global Instance iProto_le_from_wand_l a m v P p : *) - (* TCIf (TCEq P True%I) False TCTrue → *) - (* FromWand ((<?> MSG v {{ P }}; p) ⊑ (<a> m)) P ((<?> MSG v; p) ⊑ (<a> m)) | 10. *) - (* Proof. intros _. apply iProto_le_payload_elim_l. Qed. *) - (* Global Instance iProto_le_from_wand_r a m v P p : *) - (* FromWand ((<a> m) ⊑ (<!> MSG v {{ P }}; p)) P ((<a> m) ⊑ (<!> MSG v; p)) | 11. *) - (* Proof. apply iProto_le_payload_elim_r. Qed. *) - - (* Global Instance iProto_le_from_exist_l {A} (m : A → iMsg Σ V) p : *) - (* FromExist ((<! x> m x) ⊑ p) (λ a, (<!> m a) ⊑ p)%I | 10. *) - (* Proof. *) - (* rewrite /FromExist. iDestruct 1 as (x) "H". *) - (* iApply (iProto_le_trans with "[] H"). iApply iProto_le_exist_intro_l. *) - (* Qed. *) - (* Global Instance iProto_le_from_exist_r {A} (m : A → iMsg Σ V) p : *) - (* FromExist (p ⊑ <? x> m x) (λ a, p ⊑ (<?> m a))%I | 11. *) - (* Proof. *) - (* rewrite /FromExist. iDestruct 1 as (x) "H". *) - (* iApply (iProto_le_trans with "H"). iApply iProto_le_exist_intro_r. *) - (* Qed. *) - - (* Global Instance iProto_le_from_sep_l m v P p : *) - (* FromSep ((<!> MSG v {{ P }}; p) ⊑ (<!> m)) P ((<!> MSG v; p) ⊑ (<!> m)) | 10. *) - (* Proof. *) - (* rewrite /FromSep. iIntros "[HP H]". *) - (* iApply (iProto_le_trans with "[HP] H"). by iApply iProto_le_payload_intro_l. *) - (* Qed. *) - (* Global Instance iProto_le_from_sep_r m v P p : *) - (* FromSep ((<?> m) ⊑ (<?> MSG v {{ P }}; p)) P ((<?> m) ⊑ (<?> MSG v; p)) | 11. *) - (* Proof. *) - (* rewrite /FromSep. iIntros "[HP H]". *) - (* iApply (iProto_le_trans with "H"). by iApply iProto_le_payload_intro_r. *) - (* Qed. *) - - (* Global Instance iProto_le_frame_l q m v R P Q p : *) - (* Frame q R P Q → *) - (* Frame q R ((<!> MSG v {{ P }}; p) ⊑ (<!> m)) *) - (* ((<!> MSG v {{ Q }}; p) ⊑ (<!> m)) | 10. *) - (* Proof. *) - (* rewrite /Frame /=. iIntros (HP) "[HR H]". *) - (* iApply (iProto_le_trans with "[HR] H"). iApply iProto_le_payload_elim_r. *) - (* iIntros "HQ". iApply iProto_le_payload_intro_l. iApply HP; iFrame. *) - (* Qed. *) - (* Global Instance iProto_le_frame_r q m v R P Q p : *) - (* Frame q R P Q → *) - (* Frame q R ((<?> m) ⊑ (<?> MSG v {{ P }}; p)) *) - (* ((<?> m) ⊑ (<?> MSG v {{ Q }}; p)) | 11. *) - (* Proof. *) - (* rewrite /Frame /=. iIntros (HP) "[HR H]". *) - (* iApply (iProto_le_trans with "H"). iApply iProto_le_payload_elim_l. *) - (* iIntros "HQ". iApply iProto_le_payload_intro_r. iApply HP; iFrame. *) - (* Qed. *) - - (* Global Instance iProto_le_from_modal a v p1 p2 : *) - (* FromModal True (modality_instances.modality_laterN 1) (p1 ⊑ p2) *) - (* ((<a> MSG v; p1) ⊑ (<a> MSG v; p2)) (p1 ⊑ p2). *) - (* Proof. intros _. apply iProto_le_base. Qed. *) + Global Instance iProto_le_from_forall_l {A} i (m1 : A → iMsg Σ V) m2 name : + AsIdentName m1 name → + FromForall (iProto_message (Recv,i) (iMsg_exist m1) ⊑ (<(Recv,i)> m2)) + (λ x, (<(Recv, i)> m1 x) ⊑ (<(Recv, i)> m2))%I name | 10. + Proof. intros _. apply iProto_le_exist_elim_l. Qed. + Global Instance iProto_le_from_forall_r {A} i m1 (m2 : A → iMsg Σ V) name : + AsIdentName m2 name → + FromForall ((<(Send,i)> m1) ⊑ iProto_message (Send,i) (iMsg_exist m2)) + (λ x, (<(Send,i)> m1) ⊑ (<(Send,i)> m2 x))%I name | 11. + Proof. intros _. apply iProto_le_exist_elim_r. Qed. + + Global Instance iProto_le_from_wand_l i m v P p : + TCIf (TCEq P True%I) False TCTrue → + FromWand ((<(Recv,i)> MSG v {{ P }}; p) ⊑ (<(Recv,i)> m)) P ((<(Recv,i)> MSG v; p) ⊑ (<(Recv,i)> m)) | 10. + Proof. intros _. apply iProto_le_payload_elim_l. Qed. + Global Instance iProto_le_from_wand_r i m v P p : + FromWand ((<(Send,i)> m) ⊑ (<(Send,i)> MSG v {{ P }}; p)) P ((<(Send,i)> m) ⊑ (<(Send,i)> MSG v; p)) | 11. + Proof. apply iProto_le_payload_elim_r. Qed. + + Global Instance iProto_le_from_exist_l {A} i (m : A → iMsg Σ V) p : + FromExist ((<(Send,i) @ x> m x) ⊑ p) (λ a, (<(Send,i)> m a) ⊑ p)%I | 10. + Proof. + rewrite /FromExist. iDestruct 1 as (x) "H". + iApply (iProto_le_trans with "[] H"). iApply iProto_le_exist_intro_l. + Qed. + Global Instance iProto_le_from_exist_r {A} i (m : A → iMsg Σ V) p : + FromExist (p ⊑ <(Recv,i) @ x> m x) (λ a, p ⊑ (<(Recv,i)> m a))%I | 11. + Proof. + rewrite /FromExist. iDestruct 1 as (x) "H". + iApply (iProto_le_trans with "H"). iApply iProto_le_exist_intro_r. + Qed. + + Global Instance iProto_le_from_sep_l i m v P p : + FromSep ((<(Send,i)> MSG v {{ P }}; p) ⊑ (<(Send,i)> m)) P ((<(Send,i)> MSG v; p) ⊑ (<(Send,i)> m)) | 10. + Proof. + rewrite /FromSep. iIntros "[HP H]". + iApply (iProto_le_trans with "[HP] H"). by iApply iProto_le_payload_intro_l. + Qed. + Global Instance iProto_le_from_sep_r i m v P p : + FromSep ((<(Recv,i)> m) ⊑ (<(Recv,i)> MSG v {{ P }}; p)) P ((<(Recv,i)> m) ⊑ (<(Recv,i)> MSG v; p)) | 11. + Proof. + rewrite /FromSep. iIntros "[HP H]". + iApply (iProto_le_trans with "H"). by iApply iProto_le_payload_intro_r. + Qed. + + Global Instance iProto_le_frame_l i q m v R P Q p : + Frame q R P Q → + Frame q R ((<(Send,i)> MSG v {{ P }}; p) ⊑ (<(Send,i)> m)) + ((<(Send,i)> MSG v {{ Q }}; p) ⊑ (<(Send,i)> m)) | 10. + Proof. + rewrite /Frame /=. iIntros (HP) "[HR H]". + iApply (iProto_le_trans with "[HR] H"). iApply iProto_le_payload_elim_r. + iIntros "HQ". iApply iProto_le_payload_intro_l. iApply HP; iFrame. + Qed. + Global Instance iProto_le_frame_r i q m v R P Q p : + Frame q R P Q → + Frame q R ((<(Recv,i)> m) ⊑ (<(Recv,i)> MSG v {{ P }}; p)) + ((<(Recv,i)> m) ⊑ (<(Recv,i)> MSG v {{ Q }}; p)) | 11. + Proof. + rewrite /Frame /=. iIntros (HP) "[HR H]". + iApply (iProto_le_trans with "H"). iApply iProto_le_payload_elim_l. + iIntros "HQ". iApply iProto_le_payload_intro_r. iApply HP; iFrame. + Qed. + + Global Instance iProto_le_from_modal a v p1 p2 : + FromModal True (modality_instances.modality_laterN 1) (p1 ⊑ p2) + ((<a> MSG v; p1) ⊑ (<a> MSG v; p2)) (p1 ⊑ p2). + Proof. intros _. iApply iProto_le_base. Qed. End proto. -- GitLab From de80a6934ad6b4a1f15648df482bcec818acbbd8 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Wed, 24 Jan 2024 16:42:20 +0100 Subject: [PATCH 21/81] Proofmode!! --- _CoqProject | 1 + theories/channel/multi_channel.v | 5 +- theories/channel/multi_proofmode.v | 382 ++++++++++++++++++ theories/channel/multi_proto.v | 55 +++ .../multi_proto_consistency_examples.v | 87 +--- 5 files changed, 460 insertions(+), 70 deletions(-) create mode 100644 theories/channel/multi_proofmode.v diff --git a/_CoqProject b/_CoqProject index 1b9d3a3..d5da394 100644 --- a/_CoqProject +++ b/_CoqProject @@ -20,6 +20,7 @@ theories/channel/channel.v theories/channel/multi_proto_model.v theories/channel/multi_proto.v theories/channel/multi_channel.v +theories/channel/multi_proofmode.v theories/channel/multi_proto_consistency_examples.v theories/channel/proofmode.v theories/examples/basics.v diff --git a/theories/channel/multi_channel.v b/theories/channel/multi_channel.v index d8d0f7d..0868ed6 100644 --- a/theories/channel/multi_channel.v +++ b/theories/channel/multi_channel.v @@ -24,7 +24,8 @@ the subprotocol relation [⊑] *) From iris.algebra Require Import gmap excl_auth gmap_view. From iris.base_logic.lib Require Import invariants. From iris.heap_lang Require Export primitive_laws notation proofmode. -From actris.channel Require Import multi_proto_model multi_proto. +From actris.channel Require Import multi_proto_model. +From actris.channel Require Export multi_proto. Set Default Proof Using "Type". (* TODO: Update new_chan definition to use pointers with offsets *) @@ -303,7 +304,7 @@ Section channel. Lemma recv_spec {TT} c j (v : TT → val) (P : TT → iProp Σ) (p : TT → iProto Σ) : - {{{ c ↣ <(Recv, j) @.. x> MSG v x {{ P x }}; p x }}} + {{{ c ↣ <(Recv, j) @.. x> MSG v x {{ â–· P x }}; p x }}} recv c #j {{{ x, RET v x; c ↣ p x ∗ P x }}}. Proof. diff --git a/theories/channel/multi_proofmode.v b/theories/channel/multi_proofmode.v new file mode 100644 index 0000000..937f4b8 --- /dev/null +++ b/theories/channel/multi_proofmode.v @@ -0,0 +1,382 @@ +(** This file contains the definitions of Actris's tactics for symbolic +execution of message-passing programs. The API of these tactics is documented +in the [README.md] file. The implementation follows the same pattern for the +implementation of these tactics that is used in Iris. In addition, it uses a +standard pattern using type classes to perform the normalization. + +In addition to the tactics for symbolic execution, this file defines the tactic +[solve_proto_contractive], which can be used to automatically prove that +recursive protocols are contractive. *) +From iris.proofmode Require Import coq_tactics reduction spec_patterns. +From iris.heap_lang Require Export proofmode notation. +From actris.channel Require Import multi_proto_model. +From actris Require Export multi_channel. +Export action. + +(** * Tactics for proving contractiveness of protocols *) +Ltac f_dist_le := + match goal with + | H : _ ≡{?n}≡ _ |- _ ≡{?n'}≡ _ => apply (dist_le n); [apply H|lia] + end. + +Ltac solve_proto_contractive := + solve_proper_core ltac:(fun _ => + first [f_contractive; simpl in * | f_equiv | f_dist_le]). + +(** * Normalization of protocols *) +Class ActionDualIf (d : bool) (a1 a2 : action) := + dual_action_if : a2 = if d then action_dual a1 else a1. +Global Hint Mode ActionDualIf ! ! - : typeclass_instances. + +Global Instance action_dual_if_false a : ActionDualIf false a a := eq_refl. +Global Instance action_dual_if_true_send i : ActionDualIf true (Send,i) (Recv,i) := eq_refl. +Global Instance action_dual_if_true_recv i : ActionDualIf true (Recv,i) (Send,i) := eq_refl. + +Class ProtoNormalize {Σ} (d : bool) (p : iProto Σ) + (pas : list (bool * iProto Σ)) (q : iProto Σ) := + proto_normalize : + ⊢ iProto_dual_if d p <++> + foldr (iProto_app ∘ uncurry iProto_dual_if) END%proto pas ⊑ q. +Global Hint Mode ProtoNormalize ! ! ! ! - : typeclass_instances. +Arguments ProtoNormalize {_} _ _%proto _%proto _%proto. + +Notation ProtoUnfold p1 p2 := (∀ d pas q, + ProtoNormalize d p2 pas q → ProtoNormalize d p1 pas q). + +Class MsgNormalize {Σ} (d : bool) (m1 : iMsg Σ) + (pas : list (bool * iProto Σ)) (m2 : iMsg Σ) := + msg_normalize a : + ProtoNormalize d (<a> m1) pas (<(if d then action_dual a else a)> m2). +Global Hint Mode MsgNormalize ! ! ! ! - : typeclass_instances. +Arguments MsgNormalize {_} _ _%msg _%msg _%msg. + +Section classes. + Context `{!chanG Σ, !heapGS Σ}. + Implicit Types TT : tele. + Implicit Types p : iProto Σ. + Implicit Types m : iMsg Σ. + Implicit Types P : iProp Σ. + + Lemma proto_unfold_eq p1 p2 : p1 ≡ p2 → ProtoUnfold p1 p2. + Proof. rewrite /ProtoNormalize=> Hp d pas q. by rewrite Hp. Qed. + + Global Instance proto_normalize_done p : ProtoNormalize false p [] p | 0. + Proof. rewrite /ProtoNormalize /= right_id. auto. Qed. + Global Instance proto_normalize_done_dual p : + ProtoNormalize true p [] (iProto_dual p) | 0. + Proof. rewrite /ProtoNormalize /= right_id. auto. Qed. + Global Instance proto_normalize_done_dual_end : + ProtoNormalize (Σ:=Σ) true END [] END | 0. + Proof. rewrite /ProtoNormalize /= right_id iProto_dual_end. auto. Qed. + + Global Instance proto_normalize_dual d p pas q : + ProtoNormalize (negb d) p pas q → + ProtoNormalize d (iProto_dual p) pas q. + Proof. rewrite /ProtoNormalize. by destruct d; rewrite /= ?involutive. Qed. + + Global Instance proto_normalize_app_l d p1 p2 pas q : + ProtoNormalize d p1 ((d,p2) :: pas) q → + ProtoNormalize d (p1 <++> p2) pas q. + Proof. + rewrite /ProtoNormalize /=. rewrite assoc. + by destruct d; by rewrite /= ?iProto_dual_app. + Qed. + + Global Instance proto_normalize_end d d' p pas q : + ProtoNormalize d p pas q → + ProtoNormalize d' END ((d,p) :: pas) q | 0. + Proof. + rewrite /ProtoNormalize /=. + destruct d'; by rewrite /= ?iProto_dual_end left_id. + Qed. + + Global Instance proto_normalize_app_r d p1 p2 pas q : + ProtoNormalize d p2 pas q → + ProtoNormalize false p1 ((d,p2) :: pas) (p1 <++> q) | 0. + Proof. rewrite /ProtoNormalize /= => H. by iApply iProto_le_app. Qed. + + Global Instance proto_normalize_app_r_dual d p1 p2 pas q : + ProtoNormalize d p2 pas q → + ProtoNormalize true p1 ((d,p2) :: pas) (iProto_dual p1 <++> q) | 0. + Proof. rewrite /ProtoNormalize /= => H. by iApply iProto_le_app. Qed. + + Global Instance msg_normalize_base d v P p q pas : + ProtoNormalize d p pas q → + MsgNormalize d (MSG v {{ P }}; p) pas (MSG v {{ P }}; q). + Proof. + rewrite /MsgNormalize /ProtoNormalize=> H a. + iApply iProto_le_trans; [|by iApply iProto_le_base]. + destruct d; by rewrite /= ?iProto_dual_message ?iMsg_dual_base + iProto_app_message iMsg_app_base. + Qed. + + Global Instance msg_normalize_exist {A} d (m1 m2 : A → iMsg Σ) pas : + (∀ x, MsgNormalize d (m1 x) pas (m2 x)) → + MsgNormalize d (∃ x, m1 x) pas (∃ x, m2 x). + Proof. + rewrite /MsgNormalize /ProtoNormalize=> H a. + destruct d, a as [[|]]; + simpl in *; rewrite ?iProto_dual_message ?iMsg_dual_exist + ?iProto_app_message ?iMsg_app_exist /=; iIntros (x); iExists x; first + [move: (H x (Send,n)); by rewrite ?iProto_dual_message ?iProto_app_message + |move: (H x (Recv,n)); by rewrite ?iProto_dual_message ?iProto_app_message]. + Qed. + + Global Instance proto_normalize_message d a1 a2 m1 m2 pas : + ActionDualIf d a1 a2 → + MsgNormalize d m1 pas m2 → + ProtoNormalize d (<a1> m1) pas (<a2> m2). + Proof. by rewrite /ActionDualIf /MsgNormalize /ProtoNormalize=> ->. Qed. + + (** Automatically perform normalization of protocols in the proof mode when + using [iAssumption] and [iFrame]. *) + Global Instance pointsto_proto_from_assumption q c p1 p2 : + ProtoNormalize false p1 [] p2 → + FromAssumption q (c ↣ p1) (c ↣ p2). + Proof. + rewrite /FromAssumption /ProtoNormalize /= right_id. + rewrite bi.intuitionistically_if_elim. + iIntros (?) "H". by iApply (iProto_pointsto_le with "H"). + Qed. + Global Instance pointsto_proto_from_frame q c p1 p2 : + ProtoNormalize false p1 [] p2 → + Frame q (c ↣ p1) (c ↣ p2) True. + Proof. + rewrite /Frame /ProtoNormalize /= right_id. + rewrite bi.intuitionistically_if_elim. + iIntros (?) "[H _]". by iApply (iProto_pointsto_le with "H"). + Qed. +End classes. + +(** * Symbolic execution tactics *) +(* TODO: Maybe strip laters from other hypotheses in the future? *) +Lemma tac_wp_recv `{!chanG Σ, !heapGS Σ} {TT : tele} Δ i j K v (n:nat) c p m tv tP tP' tp Φ : + v = #n → + envs_lookup i Δ = Some (false, c ↣ p)%I → + ProtoNormalize false p [] (<(Recv,n)> m) → + MsgTele m tv tP tp → + (∀.. x, MaybeIntoLaterN false 1 (tele_app tP x) (tele_app tP' x)) → + let Δ' := envs_delete false i false Δ in + (∀.. x : TT, + match envs_app false + (Esnoc (Esnoc Enil j (tele_app tP' x)) i (c ↣ tele_app tp x)) Δ' with + | Some Δ'' => envs_entails Δ'' (WP fill K (of_val (tele_app tv x)) {{ Φ }}) + | None => False + end) → + envs_entails Δ (WP fill K (recv c v) {{ Φ }}). +Proof. + intros ->. + rewrite envs_entails_unseal /ProtoNormalize /MsgTele /MaybeIntoLaterN /=. + rewrite !tforall_forall right_id. + intros ? Hp Hm HP HΦ. rewrite envs_lookup_sound //; simpl. + assert (c ↣ p ⊢ c ↣ <(Recv,n) @.. x> + MSG tele_app tv x {{ â–· tele_app tP' x }}; tele_app tp x) as ->. + { iIntros "Hc". iApply (iProto_pointsto_le with "Hc"). iIntros "!>". + iApply iProto_le_trans; [iApply Hp|rewrite Hm]. + iApply iProto_le_texist_elim_l; iIntros (x). + iApply iProto_le_trans; [|iApply (iProto_le_texist_intro_r _ _ x)]; simpl. + iIntros "H". by iDestruct (HP with "H") as "$". } + rewrite -wp_bind. eapply bi.wand_apply; + [by eapply bi.wand_entails, (recv_spec _ n (tele_app tv) (tele_app tP') (tele_app tp))|f_equiv; first done]. + rewrite -bi.later_intro; apply bi.forall_intro=> x. + specialize (HΦ x). destruct (envs_app _ _) as [Δ'|] eqn:HΔ'=> //. + rewrite envs_app_sound //; simpl. by rewrite right_id HΦ. +Qed. + +Tactic Notation "wp_recv_core" tactic3(tac_intros) "as" tactic3(tac) := + let solve_pointsto _ := + let c := match goal with |- _ = Some (_, (?c ↣ _)%I) => c end in + iAssumptionCore || fail "wp_recv: cannot find" c "↣ ? @ ?" in + wp_pures; + let Hnew := iFresh in + lazymatch goal with + | |- envs_entails _ (wp ?s ?E ?e ?Q) => + first + [reshape_expr e ltac:(fun K e' => eapply (tac_wp_recv _ _ Hnew K)) + |fail 1 "wp_recv: cannot find 'recv' in" e]; + [try done| + solve_pointsto () + |tc_solve || fail 1 "wp_recv: protocol not of the shape <?>" + |tc_solve || fail 1 "wp_recv: cannot convert to telescope" + |tc_solve + |pm_reduce; simpl; tac_intros; + tac Hnew; + wp_finish] + | _ => fail "wp_recv: not a 'wp'" + end. + +Tactic Notation "wp_recv" "as" constr(pat) := + wp_recv_core (idtac) as (fun H => iDestructHyp H as pat). + +Tactic Notation "wp_recv" "(" simple_intropattern_list(xs) ")" "as" constr(pat) := + wp_recv_core (intros xs) as (fun H => iDestructHyp H as pat). +Tactic Notation "wp_recv" "(" simple_intropattern_list(xs) ")" "as" "(" simple_intropattern(x1) ")" + constr(pat) := + wp_recv_core (intros xs) as (fun H => iDestructHyp H as ( x1 ) pat). +Tactic Notation "wp_recv" "(" simple_intropattern_list(xs) ")" "as" "(" simple_intropattern(x1) + simple_intropattern(x2) ")" constr(pat) := + wp_recv_core (intros xs) as (fun H => iDestructHyp H as ( x1 x2 ) pat). +Tactic Notation "wp_recv" "(" simple_intropattern_list(xs) ")" "as" "(" simple_intropattern(x1) + simple_intropattern(x2) simple_intropattern(x3) ")" constr(pat) := + wp_recv_core (intros xs) as (fun H => iDestructHyp H as ( x1 x2 x3 ) pat). +Tactic Notation "wp_recv" "(" simple_intropattern_list(xs) ")" "as" "(" simple_intropattern(x1) + simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4) ")" + constr(pat) := + wp_recv_core (intros xs) as (fun H => iDestructHyp H as ( x1 x2 x3 x4 ) pat). +Tactic Notation "wp_recv" "(" simple_intropattern_list(xs) ")" "as" "(" simple_intropattern(x1) + simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4) + simple_intropattern(x5) ")" constr(pat) := + wp_recv_core (intros xs) as (fun H => iDestructHyp H as ( x1 x2 x3 x4 x5 ) pat). +Tactic Notation "wp_recv" "(" simple_intropattern_list(xs) ")" "as" "(" simple_intropattern(x1) + simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4) + simple_intropattern(x5) simple_intropattern(x6) ")" constr(pat) := + wp_recv_core (intros xs) as (fun H => iDestructHyp H as ( x1 x2 x3 x4 x5 x6 ) pat). +Tactic Notation "wp_recv" "(" simple_intropattern_list(xs) ")" "as" "(" simple_intropattern(x1) + simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4) + simple_intropattern(x5) simple_intropattern(x6) simple_intropattern(x7) ")" + constr(pat) := + wp_recv_core (intros xs) as (fun H => iDestructHyp H as ( x1 x2 x3 x4 x5 x6 x7 ) pat). +Tactic Notation "wp_recv" "(" simple_intropattern_list(xs) ")" "as" "(" simple_intropattern(x1) + simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4) + simple_intropattern(x5) simple_intropattern(x6) simple_intropattern(x7) + simple_intropattern(x8) ")" constr(pat) := + wp_recv_core (intros xs) as (fun H => iDestructHyp H as ( x1 x2 x3 x4 x5 x6 x7 x8 ) pat). + +(* Section channel. *) +(* Context `{!heapGS Σ, !chanG Σ}. *) +(* Implicit Types p : iProto Σ. *) +(* Implicit Types TT : tele. *) + +(* Lemma recv_test c p : *) +(* {{{ c ↣ (<(Recv,0) @(x:Z)> MSG #x ; p) }}} *) +(* recv c #0 *) +(* {{{ x, RET #x; c ↣ p }}}. *) +(* Proof. *) +(* iIntros (Φ) "Hc HΦ". *) +(* wp_recv (x) as "_". *) + +Lemma tac_wp_send `{!chanG Σ, !heapGS Σ} {TT : tele} Δ neg i js K (n:nat) w c v p m tv tP tp Φ : + w = #n → + envs_lookup i Δ = Some (false, c ↣ p)%I → + ProtoNormalize false p [] (<(Send,n)> m) → + MsgTele m tv tP tp → + let Δ' := envs_delete false i false Δ in + (∃.. x : TT, + match envs_split (if neg is true then base.Right else base.Left) js Δ' with + | Some (Δ1,Δ2) => + match envs_app false (Esnoc Enil i (c ↣ tele_app tp x)) Δ2 with + | Some Δ2' => + v = tele_app tv x ∧ + envs_entails Δ1 (tele_app tP x) ∧ + envs_entails Δ2' (WP fill K (of_val #()) {{ Φ }}) + | None => False + end + | None => False + end) → + envs_entails Δ (WP fill K (send c w v) {{ Φ }}). +Proof. + intros ->. + rewrite envs_entails_unseal /ProtoNormalize /MsgTele /= right_id texist_exist. + intros ? Hp Hm [x HΦ]. rewrite envs_lookup_sound //; simpl. + destruct (envs_split _ _ _) as [[Δ1 Δ2]|] eqn:? => //. + destruct (envs_app _ _ _) as [Δ2'|] eqn:? => //. + rewrite envs_split_sound //; rewrite (envs_app_sound Δ2) //; simpl. + destruct HΦ as (-> & -> & ->). rewrite right_id assoc. + assert (c ↣ p ⊢ + c ↣ <(Send,n) @.. (x : TT)> MSG tele_app tv x {{ tele_app tP x }}; tele_app tp x) as ->. + { iIntros "Hc". iApply (iProto_pointsto_le with "Hc"); iIntros "!>". + iApply iProto_le_trans; [iApply Hp|]. by rewrite Hm. } + eapply bi.wand_apply; [rewrite -wp_bind; by eapply bi.wand_entails, send_spec_tele|]. + by rewrite -bi.later_intro. +Qed. + +Tactic Notation "wp_send_core" tactic3(tac_exist) "with" constr(pat) := + let solve_pointsto _ := + let c := match goal with |- _ = Some (_, (?c ↣ _)%I) => c end in + iAssumptionCore || fail "wp_send: cannot find" c "↣ ? @ ?" in + let solve_done d := + lazymatch d with + | true => + done || + let Q := match goal with |- envs_entails _ ?Q => Q end in + fail "wp_send: cannot solve" Q "using done" + | false => idtac + end in + lazymatch spec_pat.parse pat with + | [SGoal (SpecGoal GSpatial ?neg ?Hs_frame ?Hs ?d)] => + let Hs' := eval cbv in (if neg then Hs else Hs_frame ++ Hs) in + wp_pures; + lazymatch goal with + | |- envs_entails _ (wp ?s ?E ?e ?Q) => + first + [reshape_expr e ltac:(fun K e' => eapply (tac_wp_send _ neg _ Hs' K)) + |fail 1 "wp_send: cannot find 'send' in" e]; + [try done|solve_pointsto () + |tc_solve || fail 1 "wp_send: protocol not of the shape <!>" + |tc_solve || fail 1 "wp_send: cannot convert to telescope" + |pm_reduce; simpl; tac_exist; + repeat lazymatch goal with + | |- ∃ _, _ => eexists _ + end; + lazymatch goal with + | |- False => fail "wp_send:" Hs' "not found" + | _ => notypeclasses refine (conj (eq_refl _) (conj _ _)); + [iFrame Hs_frame; solve_done d + |wp_finish] + end] + | _ => fail "wp_send: not a 'wp'" + end + | _ => fail "wp_send: only a single goal spec pattern supported" + end. + +Tactic Notation "wp_send" "with" constr(pat) := + wp_send_core (idtac) with pat. +Tactic Notation "wp_send" "(" uconstr(x1) ")" "with" constr(pat) := + wp_send_core (eexists x1) with pat. +Tactic Notation "wp_send" "(" uconstr(x1) uconstr(x2) ")" "with" constr(pat) := + wp_send_core (eexists x1; eexists x2) with pat. +Tactic Notation "wp_send" "(" uconstr(x1) uconstr(x2) uconstr(x3) ")" + "with" constr(pat) := + wp_send_core (eexists x1; eexists x2; eexists x3) with pat. +Tactic Notation "wp_send" "(" uconstr(x1) uconstr(x2) uconstr(x3) uconstr(x4) ")" + "with" constr(pat) := + wp_send_core (eexists x1; eexists x2; eexists x3; eexists x4) with pat. +Tactic Notation "wp_send" "(" uconstr(x1) uconstr(x2) uconstr(x3) uconstr(x4) + uconstr(x5) ")" "with" constr(pat) := + wp_send_core (eexists x1; eexists x2; eexists x3; eexists x4; eexists x5) with pat. +Tactic Notation "wp_send" "(" uconstr(x1) uconstr(x2) uconstr(x3) uconstr(x4) ")" + uconstr(x5) uconstr(x6) ")" "with" constr(pat) := + wp_send_core (eexists x1; eexists x2; eexists x3; eexists x4; eexists x5; + eexists x6) with pat. +Tactic Notation "wp_send" "(" uconstr(x1) uconstr(x2) uconstr(x3) uconstr(x4) ")" + uconstr(x5) uconstr(x6) uconstr(x7) ")" "with" constr(pat) := + wp_send_core (eexists x1; eexists x2; eexists x3; eexists x4; eexists x5; + eexists x6; eexists x7) with pat. +Tactic Notation "wp_send" "(" uconstr(x1) uconstr(x2) uconstr(x3) uconstr(x4) ")" + uconstr(x5) uconstr(x6) uconstr(x7) uconstr(x8) ")" "with" constr(pat) := + wp_send_core (eexists x1; eexists x2; eexists x3; eexists x4; eexists x5; + eexists x6; eexists x7; eexists x8) with pat. + +(* Section channel. *) +(* Context `{!heapGS Σ, !chanG Σ}. *) +(* Implicit Types p : iProto Σ. *) +(* Implicit Types TT : tele. *) + +(* (* Lemma recv_test c p : *) *) +(* (* {{{ c ↣ (<(Recv,0) @(x:Z)> MSG #x ; p) }}} *) *) +(* (* recv c #0 *) *) +(* (* {{{ x, RET #x; c ↣ p }}}. *) *) +(* (* Proof. *) *) +(* (* iIntros (Φ) "Hc HΦ". *) *) +(* (* wp_recv (x) as "_". *) *) +(* (* { done. } *) *) +(* (* iApply "HΦ". *) *) + +(* Lemma send_test c p : *) +(* {{{ c ↣ (<(Send,0) @(x:Z)> MSG #x ; p) }}} *) +(* send c #0 #42 *) +(* {{{ x, RET #x; c ↣ p }}}. *) +(* Proof. *) +(* iIntros (Φ) "Hc HΦ". *) +(* wp_send (42%Z) with "[]". *) + diff --git a/theories/channel/multi_proto.v b/theories/channel/multi_proto.v index 199c5f8..898b391 100644 --- a/theories/channel/multi_proto.v +++ b/theories/channel/multi_proto.v @@ -903,6 +903,61 @@ Section proto. - iApply iProto_le_recv. auto 10 with iFrame. Qed. +Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. + Proof. + iIntros "H". iLöb as "IH" forall (p1 p2). + destruct (iProto_case p1) as [->|([]&i&m1&->)]. + - iDestruct (iProto_le_end_inv_l with "H") as "H". + iRewrite "H". iApply iProto_le_refl. + - iDestruct (iProto_le_send_inv with "H") as (m2) "[Hp2 H]". + iRewrite "Hp2"; clear p2. iEval (rewrite !iProto_dual_message). + iApply iProto_le_recv. iIntros (v p1d). + iDestruct 1 as (p1') "[Hm1 #Hp1d]". + iDestruct ("H" with "Hm1") as (p2') "[H Hm2]". + iDestruct ("IH" with "H") as "H". iExists (iProto_dual p2'). + iSplitL "H"; [iIntros "!>"; by iRewrite "Hp1d"|]. simpl; auto. + - iDestruct (iProto_le_recv_inv_r with "H") as (m2) "[Hp2 H]". + iRewrite "Hp2"; clear p2. iEval (rewrite !iProto_dual_message /=). + iApply iProto_le_send. iIntros (v p2d). + iDestruct 1 as (p2') "[Hm2 #Hp2d]". + iDestruct ("H" with "Hm2") as (p1') "[H Hm1]". + iDestruct ("IH" with "H") as "H". iExists (iProto_dual p1'). + iSplitL "H"; [iIntros "!>"; by iRewrite "Hp2d"|]. simpl; auto. + Qed. + + Lemma iProto_le_dual_l p1 p2 : iProto_dual p2 ⊑ p1 ⊢ iProto_dual p1 ⊑ p2. + Proof. + iIntros "H". iEval (rewrite -(involutive iProto_dual p2)). + by iApply iProto_le_dual. + Qed. + Lemma iProto_le_dual_r p1 p2 : p2 ⊑ iProto_dual p1 ⊢ p1 ⊑ iProto_dual p2. + Proof. + iIntros "H". iEval (rewrite -(involutive iProto_dual p1)). + by iApply iProto_le_dual. + Qed. + + Lemma iProto_le_app p1 p2 p3 p4 : + p1 ⊑ p2 -∗ p3 ⊑ p4 -∗ p1 <++> p3 ⊑ p2 <++> p4. + Proof. + iIntros "H1 H2". iLöb as "IH" forall (p1 p2 p3 p4). + destruct (iProto_case p2) as [->|([]&i&m2&->)]. + - iDestruct (iProto_le_end_inv_l with "H1") as "H1". + iRewrite "H1". by rewrite !left_id. + - iDestruct (iProto_le_send_inv with "H1") as (m1) "[Hp1 H1]". + iRewrite "Hp1"; clear p1. rewrite !iProto_app_message. + iApply iProto_le_send. iIntros (v p24). + iDestruct 1 as (p2') "[Hm2 #Hp24]". + iDestruct ("H1" with "Hm2") as (p1') "[H1 Hm1]". + iExists (p1' <++> p3). iSplitR "Hm1"; [|by simpl; eauto]. + iIntros "!>". iRewrite "Hp24". by iApply ("IH" with "H1"). + - iDestruct (iProto_le_recv_inv_r with "H1") as (m1) "[Hp1 H1]". + iRewrite "Hp1"; clear p1. rewrite !iProto_app_message. iApply iProto_le_recv. + iIntros (v p13). iDestruct 1 as (p1') "[Hm1 #Hp13]". + iDestruct ("H1" with "Hm1") as (p2'') "[H1 Hm2]". + iExists (p2'' <++> p4). iSplitR "Hm2"; [|by simpl; eauto]. + iIntros "!>". iRewrite "Hp13". by iApply ("IH" with "H1"). + Qed. + Lemma iProto_le_payload_elim_l i m v P p : (P -∗ (<(Recv,i)> MSG v; p) ⊑ (<(Recv,i)> m)) ⊢ (<(Recv,i)> MSG v {{ P }}; p) ⊑ <(Recv,i)> m. diff --git a/theories/channel/multi_proto_consistency_examples.v b/theories/channel/multi_proto_consistency_examples.v index 2031476..0c3b22a 100644 --- a/theories/channel/multi_proto_consistency_examples.v +++ b/theories/channel/multi_proto_consistency_examples.v @@ -2,7 +2,7 @@ From iris.algebra Require Import gmap excl_auth gmap_view. From iris.proofmode Require Import proofmode. From iris.base_logic Require Export lib.iprop. From iris.base_logic Require Import lib.own. -From actris.channel Require Import multi_proto_model multi_proto multi_channel. +From actris.channel Require Import multi_proto_model multi_proto multi_channel multi_proofmode. Set Default Proof Using "Type". Export action. @@ -325,34 +325,18 @@ Section channel. iIntros (c2) "[Hc2 Hcs]". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". - wp_smart_apply - (recv_spec (TT:=[tele Z]) c1 0 - (tele_app (λ (x:Z), #x)) (λ _, True)%I (tele_app (λ (x:Z), _)) - with "Hc1"). - iIntros (x') "[Hc1 _]". - epose proof (tele_arg_S_inv x') as [x [[] ->]]. simpl. - wp_smart_apply (send_spec c1 2 with "Hc1"). - by iIntros "_". } + (* TODO: Fix unification *) + wp_recv (x) as "_"; [done|]. + wp_send with "[//]"; [done|]. + done. } wp_smart_apply (wp_fork with "[Hc2]"). { iIntros "!>". - wp_smart_apply - (recv_spec (TT:=[tele Z]) c2 1 - (tele_app (λ (x:Z), #x)) (λ _, True)%I (tele_app (λ (x:Z), _)) - with "Hc2"). - iIntros (x') "[Hc1 _]". - epose proof (tele_arg_S_inv x') as [x [[] ->]]. simpl. - wp_smart_apply (send_spec c2 0 with "Hc1"). - by iIntros "_". } - wp_smart_apply - (send_spec_tele (TT:=[tele Z]) c0 1 ([tele_arg 42%Z]) - (tele_app (λ (x:Z), #x)) (λ _, True)%I - (tele_app (λ (x:Z), _)) - with "[Hc0]"). - { iSplitL; [|done]. simpl. iFrame "Hc0". } - iIntros "Hc0". - wp_smart_apply (recv_spec (TT:=[tele]) c0 2 (λ _, #42) (λ _, True)%I (λ _, _) - with "Hc0"). - iIntros (_) "Hc0". by iApply "HΦ". + wp_recv (x) as "_"; [done|]. + wp_send with "[//]"; [done|]. + done. } + wp_send with "[//]"; [done|]. + wp_recv as "_"; [done|]. + by iApply "HΦ". Qed. End channel. @@ -591,52 +575,19 @@ Section proof. iIntros (c2) "[Hc2 Hcs]". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". - wp_smart_apply - (recv_spec (TT:=[tele loc Z]) c1 0 - (tele_app (λ (l : loc) (x:Z), #l)) - (tele_app (λ (l : loc) (x:Z), l ↦ #x)%I) - (tele_app (λ (l : loc) (x:Z), _)) - with "Hc1"). - iIntros (x') "[Hc1 Hl]". - epose proof (tele_arg_S_inv x') as [l [y' ->]]. simpl. - epose proof (tele_arg_S_inv y') as [x [[] ->]]. simpl. + wp_recv (l x) as "Hl"; [done|]. wp_load. wp_store. - wp_smart_apply (send_spec_tele (TT:=[tele]) c1 2 - ([tele_arg]) - (λ _, #l) - (λ _, l ↦ #(x+1))%I - (λ _, _) with "[$Hc1 $Hl]"). - by iIntros "_". } + wp_send with "[$Hl]"; [done|]. + done. } wp_smart_apply (wp_fork with "[Hc2]"). { iIntros "!>". - wp_smart_apply - (recv_spec (TT:=[tele loc Z]) c2 1 - (tele_app (λ (l : loc) (x:Z), #l)) - (tele_app (λ (l : loc) (x:Z), l ↦ #x)%I) - (tele_app (λ (l : loc) (x:Z), _)) - with "Hc2"). - iIntros (x') "[Hc2 Hl]". - epose proof (tele_arg_S_inv x') as [l [y' ->]]. simpl. - epose proof (tele_arg_S_inv y') as [x [[] ->]]. simpl. + wp_recv (l x) as "Hl"; [done|]. wp_load. wp_store. - wp_smart_apply (send_spec_tele (TT:=[tele]) c2 0 - ([tele_arg]) - (λ _, #()) - (λ _, l ↦ #(x+1))%I - (λ _, _) with "[$Hc2 $Hl]"). - by iIntros "_". } + wp_send with "[$Hl]"; [done|]. + done. } wp_alloc l as "Hl". - wp_smart_apply - (send_spec_tele (TT:=[tele l Z]) c0 1 ([tele_arg l ; 40%Z]) - (tele_app (λ (l:loc) (x:Z), #l)) - (tele_app (λ (l:loc) (x:Z), l ↦ #x)%I) - (tele_app (λ (l:loc) (x:Z), _)) - with "[$Hc0 $Hl]"). - iIntros "Hc0". - wp_smart_apply (recv_spec (TT:=[tele]) c0 2 - (λ _, #()) (λ _, l ↦ #(40 + 2))%I (λ _, _) - with "Hc0"). - iIntros (_) "[Hc0 Hl]". wp_load. by iApply "HΦ". + wp_send with "[$Hl]"; [done|]. + wp_recv as "Hl"; [done|]. wp_load. by iApply "HΦ". Qed. End proof. -- GitLab From d88939b6aac91a33167d38737cea7da84f076d03 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Wed, 24 Jan 2024 17:11:26 +0100 Subject: [PATCH 22/81] Made proofmode resolve idx unification --- theories/channel/multi_proofmode.v | 55 ++++++++++--------- .../multi_proto_consistency_examples.v | 36 +++++------- 2 files changed, 42 insertions(+), 49 deletions(-) diff --git a/theories/channel/multi_proofmode.v b/theories/channel/multi_proofmode.v index 937f4b8..a86ad89 100644 --- a/theories/channel/multi_proofmode.v +++ b/theories/channel/multi_proofmode.v @@ -194,14 +194,13 @@ Tactic Notation "wp_recv_core" tactic3(tac_intros) "as" tactic3(tac) := first [reshape_expr e ltac:(fun K e' => eapply (tac_wp_recv _ _ Hnew K)) |fail 1 "wp_recv: cannot find 'recv' in" e]; - [try done| - solve_pointsto () + [|solve_pointsto () |tc_solve || fail 1 "wp_recv: protocol not of the shape <?>" |tc_solve || fail 1 "wp_recv: cannot convert to telescope" |tc_solve |pm_reduce; simpl; tac_intros; tac Hnew; - wp_finish] + wp_finish];[try done|] | _ => fail "wp_recv: not a 'wp'" end. @@ -311,7 +310,7 @@ Tactic Notation "wp_send_core" tactic3(tac_exist) "with" constr(pat) := first [reshape_expr e ltac:(fun K e' => eapply (tac_wp_send _ neg _ Hs' K)) |fail 1 "wp_send: cannot find 'send' in" e]; - [try done|solve_pointsto () + [|solve_pointsto () |tc_solve || fail 1 "wp_send: protocol not of the shape <!>" |tc_solve || fail 1 "wp_send: cannot convert to telescope" |pm_reduce; simpl; tac_exist; @@ -323,7 +322,7 @@ Tactic Notation "wp_send_core" tactic3(tac_exist) "with" constr(pat) := | _ => notypeclasses refine (conj (eq_refl _) (conj _ _)); [iFrame Hs_frame; solve_done d |wp_finish] - end] + end]; [try done|..] | _ => fail "wp_send: not a 'wp'" end | _ => fail "wp_send: only a single goal spec pattern supported" @@ -357,26 +356,28 @@ Tactic Notation "wp_send" "(" uconstr(x1) uconstr(x2) uconstr(x3) uconstr(x4) ") wp_send_core (eexists x1; eexists x2; eexists x3; eexists x4; eexists x5; eexists x6; eexists x7; eexists x8) with pat. -(* Section channel. *) -(* Context `{!heapGS Σ, !chanG Σ}. *) -(* Implicit Types p : iProto Σ. *) -(* Implicit Types TT : tele. *) +Section channel. + Context `{!heapGS Σ, !chanG Σ}. + Implicit Types p : iProto Σ. + Implicit Types TT : tele. -(* (* Lemma recv_test c p : *) *) -(* (* {{{ c ↣ (<(Recv,0) @(x:Z)> MSG #x ; p) }}} *) *) -(* (* recv c #0 *) *) -(* (* {{{ x, RET #x; c ↣ p }}}. *) *) -(* (* Proof. *) *) -(* (* iIntros (Φ) "Hc HΦ". *) *) -(* (* wp_recv (x) as "_". *) *) -(* (* { done. } *) *) -(* (* iApply "HΦ". *) *) - -(* Lemma send_test c p : *) -(* {{{ c ↣ (<(Send,0) @(x:Z)> MSG #x ; p) }}} *) -(* send c #0 #42 *) -(* {{{ x, RET #x; c ↣ p }}}. *) -(* Proof. *) -(* iIntros (Φ) "Hc HΦ". *) -(* wp_send (42%Z) with "[]". *) - + (* TODO: Why do the tactics not strip laters? *) + Lemma recv_test c p : + {{{ c ↣ (<(Recv,0) @(x:Z)> MSG #x ; p) }}} + recv c #0 + {{{ x, RET #x; c ↣ p }}}. + Proof. + iIntros (Φ) "Hc HΦ". + wp_recv (x) as "_". + Admitted. + + Lemma send_test c p : + {{{ c ↣ (<(Send,0) @(x:Z)> MSG #x ; p) }}} + send c #0 #42 + {{{ x, RET #x; c ↣ p }}}. + Proof. + iIntros (Φ) "Hc HΦ". + wp_send (42%Z) with "[//]". + Admitted. + +End channel. diff --git a/theories/channel/multi_proto_consistency_examples.v b/theories/channel/multi_proto_consistency_examples.v index 0c3b22a..bdbb9ae 100644 --- a/theories/channel/multi_proto_consistency_examples.v +++ b/theories/channel/multi_proto_consistency_examples.v @@ -325,17 +325,16 @@ Section channel. iIntros (c2) "[Hc2 Hcs]". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". - (* TODO: Fix unification *) - wp_recv (x) as "_"; [done|]. - wp_send with "[//]"; [done|]. + wp_recv (x) as "_". + wp_send with "[//]". done. } wp_smart_apply (wp_fork with "[Hc2]"). { iIntros "!>". - wp_recv (x) as "_"; [done|]. - wp_send with "[//]"; [done|]. + wp_recv (x) as "_". + wp_send with "[//]". done. } - wp_send with "[//]"; [done|]. - wp_recv as "_"; [done|]. + wp_send with "[//]". + wp_recv as "_". by iApply "HΦ". Qed. @@ -347,10 +346,10 @@ Section example4. Definition iProto_example4 : gmap nat (iProto Σ) := <[0 := (<(Send, 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; <(Recv, 2)> MSG #() {{ l ↦ #(x+2) }} ; END)%proto]> - (<[1 := (<(Recv, 0) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; - <(Send, 2)> MSG #l {{ l ↦ #(x+1) }}; END)%proto]> - (<[2 := (<(Recv, 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; - <(Send, 0)> MSG #() {{ l ↦ #(x+1) }}; END)%proto]> + (<[1 := (<(Recv, 0) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; + <(Send, 2)> MSG #l {{ l ↦ #(x+1) }}; END)%proto]> + (<[2 := (<(Recv, 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; + <(Send, 0)> MSG #() {{ l ↦ #(x+1) }}; END)%proto]> ∅)). Lemma iProto_example4_consistent : @@ -575,19 +574,12 @@ Section proof. iIntros (c2) "[Hc2 Hcs]". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". - wp_recv (l x) as "Hl"; [done|]. - wp_load. wp_store. - wp_send with "[$Hl]"; [done|]. - done. } + wp_recv (l x) as "Hl". wp_load. wp_store. by wp_send with "[$Hl]". } wp_smart_apply (wp_fork with "[Hc2]"). { iIntros "!>". - wp_recv (l x) as "Hl"; [done|]. - wp_load. wp_store. - wp_send with "[$Hl]"; [done|]. - done. } - wp_alloc l as "Hl". - wp_send with "[$Hl]"; [done|]. - wp_recv as "Hl"; [done|]. wp_load. by iApply "HΦ". + wp_recv (l x) as "Hl". wp_load. wp_store. by wp_send with "[$Hl]". } + wp_alloc l as "Hl". wp_send with "[$Hl]". wp_recv as "Hl". wp_load. + by iApply "HΦ". Qed. End proof. -- GitLab From f14549a3b64b9e83351962cb81d3f873b10f64f7 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Sat, 27 Jan 2024 03:15:57 +0100 Subject: [PATCH 23/81] Automation for proving consistency relation --- theories/channel/multi_proofmode.v | 57 +- .../multi_proto_consistency_examples.v | 549 ++++-------------- 2 files changed, 119 insertions(+), 487 deletions(-) diff --git a/theories/channel/multi_proofmode.v b/theories/channel/multi_proofmode.v index a86ad89..ff8c57b 100644 --- a/theories/channel/multi_proofmode.v +++ b/theories/channel/multi_proofmode.v @@ -240,19 +240,6 @@ Tactic Notation "wp_recv" "(" simple_intropattern_list(xs) ")" "as" "(" simple_i simple_intropattern(x5) simple_intropattern(x6) simple_intropattern(x7) simple_intropattern(x8) ")" constr(pat) := wp_recv_core (intros xs) as (fun H => iDestructHyp H as ( x1 x2 x3 x4 x5 x6 x7 x8 ) pat). - -(* Section channel. *) -(* Context `{!heapGS Σ, !chanG Σ}. *) -(* Implicit Types p : iProto Σ. *) -(* Implicit Types TT : tele. *) - -(* Lemma recv_test c p : *) -(* {{{ c ↣ (<(Recv,0) @(x:Z)> MSG #x ; p) }}} *) -(* recv c #0 *) -(* {{{ x, RET #x; c ↣ p }}}. *) -(* Proof. *) -(* iIntros (Φ) "Hc HΦ". *) -(* wp_recv (x) as "_". *) Lemma tac_wp_send `{!chanG Σ, !heapGS Σ} {TT : tele} Δ neg i js K (n:nat) w c v p m tv tP tp Φ : w = #n → @@ -356,28 +343,28 @@ Tactic Notation "wp_send" "(" uconstr(x1) uconstr(x2) uconstr(x3) uconstr(x4) ") wp_send_core (eexists x1; eexists x2; eexists x3; eexists x4; eexists x5; eexists x6; eexists x7; eexists x8) with pat. -Section channel. - Context `{!heapGS Σ, !chanG Σ}. - Implicit Types p : iProto Σ. - Implicit Types TT : tele. +(* Section channel. *) +(* Context `{!heapGS Σ, !chanG Σ}. *) +(* Implicit Types p : iProto Σ. *) +(* Implicit Types TT : tele. *) - (* TODO: Why do the tactics not strip laters? *) - Lemma recv_test c p : - {{{ c ↣ (<(Recv,0) @(x:Z)> MSG #x ; p) }}} - recv c #0 - {{{ x, RET #x; c ↣ p }}}. - Proof. - iIntros (Φ) "Hc HΦ". - wp_recv (x) as "_". - Admitted. +(* (* TODO: Why do the tactics not strip laters? *) *) +(* Lemma recv_test c p : *) +(* {{{ c ↣ (<(Recv,0) @(x:Z)> MSG #x ; p) }}} *) +(* recv c #0 *) +(* {{{ x, RET #x; c ↣ p }}}. *) +(* Proof. *) +(* iIntros (Φ) "Hc HΦ". *) +(* wp_recv (x) as "_". *) +(* Admitted. *) - Lemma send_test c p : - {{{ c ↣ (<(Send,0) @(x:Z)> MSG #x ; p) }}} - send c #0 #42 - {{{ x, RET #x; c ↣ p }}}. - Proof. - iIntros (Φ) "Hc HΦ". - wp_send (42%Z) with "[//]". - Admitted. +(* Lemma send_test c p : *) +(* {{{ c ↣ (<(Send,0) @(x:Z)> MSG #x ; p) }}} *) +(* send c #0 #42 *) +(* {{{ x, RET #x; c ↣ p }}}. *) +(* Proof. *) +(* iIntros (Φ) "Hc HΦ". *) +(* wp_send (42%Z) with "[//]". *) +(* Admitted. *) -End channel. +(* End channel. *) diff --git a/theories/channel/multi_proto_consistency_examples.v b/theories/channel/multi_proto_consistency_examples.v index bdbb9ae..7378b82 100644 --- a/theories/channel/multi_proto_consistency_examples.v +++ b/theories/channel/multi_proto_consistency_examples.v @@ -6,17 +6,83 @@ From actris.channel Require Import multi_proto_model multi_proto multi_channel m Set Default Proof Using "Type". Export action. +Lemma iProto_consistent_equiv_proof {Σ} (ps : gmap nat (iProto Σ)) : + (∀ i j m1 m2, + (ps !!! i ≡ (<(Send, j)> m1)%proto) -∗ + (ps !!! j ≡ (<(Recv, i)> m2)%proto) -∗ + ∃ m1' m2' (TT1:tele) (TT2:tele) tv1 tP1 tp1 tv2 tP2 tp2, + (<(Send, j)> m1')%proto ≡ (<(Send, j)> m1)%proto ∗ + (<(Recv, i)> m2')%proto ≡ (<(Recv, i)> m2)%proto ∗ + ⌜MsgTele (TT:=TT1) m1' tv1 tP1 tp1⌠∗ + ⌜MsgTele (TT:=TT2) m2' tv2 tP2 tp2⌠∗ + ∀.. (x : TT1), tele_app tP1 x -∗ + ∃.. (y : TT2), ⌜tele_app tv1 x = tele_app tv2 y⌠∗ + tele_app tP2 y ∗ + â–· (iProto_consistent + (<[i:=tele_app tp1 x]>(<[j:=tele_app tp2 y]>ps)))) -∗ + iProto_consistent ps. +Proof. + iIntros "H". + rewrite iProto_consistent_unfold. + iIntros (i j m1 m2) "Hm1 Hm2". + iDestruct ("H" with "Hm1 Hm2") + as (m1' m2' TT1 TT2 tv1 tP1 tp1 tv2 tP2 tp2) + "(Heq1 & Heq2& %Hm1' & %Hm2' & H)". + rewrite !iProto_message_equivI. + iDestruct "Heq1" as "[_ Heq1]". + iDestruct "Heq2" as "[_ Heq2]". + iIntros (v p1) "Hm1". + iSpecialize ("Heq1" $! v (Next p1)). + iRewrite -"Heq1" in "Hm1". + rewrite Hm1'. + rewrite iMsg_base_eq. rewrite iMsg_texist_exist. + iDestruct "Hm1" as (x Htv1) "[Hp1 HP1]". + iDestruct ("H" with "HP1") as (y Htv2) "[HP2 H]". + iExists (tele_app tp2 y). + iSpecialize ("Heq2" $! v (Next (tele_app tp2 y))). + iRewrite -"Heq2". + rewrite Hm2'. rewrite iMsg_base_eq. rewrite iMsg_texist_exist. + iSplitL "HP2". + { iExists y. iFrame. + iSplit; [|done]. + iPureIntro. subst. done. } + iNext. iRewrite -"Hp1". done. +Qed. + +Tactic Notation "iProto_consistent_take_step" := + let i := fresh in + let j := fresh in + let m1 := fresh in + let m2 := fresh in + iApply iProto_consistent_equiv_proof; + iIntros (i j m1 m2) "#Hm1 #Hm2"; + repeat (destruct i as [|i]; + [repeat (rewrite lookup_total_insert_ne; [|lia]); + try (by rewrite lookup_total_empty iProto_end_message_equivI); + try (rewrite lookup_total_insert; + try (by rewrite iProto_end_message_equivI); + iDestruct (iProto_message_equivI with "Hm1") + as "[%Heq1 Hm1']";simplify_eq)| + repeat (rewrite lookup_total_insert_ne; [|lia]); + try (by rewrite lookup_total_empty iProto_end_message_equivI)]); + repeat (rewrite lookup_total_insert_ne; [|lia]); + rewrite lookup_total_insert; + iDestruct (iProto_message_equivI with "Hm2") + as "[%Heq2 Hm2']";simplify_eq; + try (iClear "Hm1' Hm2'"; + iExists _,_,_,_,_,_,_,_,_,_; + iSplitL; [iFrame "#"|]; + iSplitL; [iFrame "#"|]; + iSplitL; [iPureIntro; tc_solve|]; + iSplitL; [iPureIntro; tc_solve|]; + simpl; iClear "#"; clear m1 m2). + Definition iProto_example1 {Σ} : gmap nat (iProto Σ) := ∅. Lemma iProto_example1_consistent {Σ} : ⊢ iProto_consistent (@iProto_example1 Σ). -Proof. - rewrite iProto_consistent_unfold. - iIntros (i j m1 m2) "Hi Hj". - rewrite lookup_total_empty. - by rewrite iProto_end_message_equivI. -Qed. +Proof. iProto_consistent_take_step. Qed. Definition iProto_example2 `{!invGS Σ} (P : iProp Σ) : gmap nat (iProto Σ) := <[0 := (<(Send, 1) @ (x:Z)> MSG #x {{ P }} ; END)%proto ]> @@ -26,79 +92,10 @@ Definition iProto_example2 `{!invGS Σ} (P : iProp Σ) : gmap nat (iProto Σ) := Lemma iProto_example2_consistent `{!invGS Σ} (P : iProp Σ) : ⊢ iProto_consistent (@iProto_example2 Σ invGS0 P). Proof. - rewrite iProto_consistent_unfold. rewrite /iProto_example2. - iIntros (i j m1 m2) "Hm1 Hm2". - destruct i, j. - - rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - iDestruct "Hm1" as (Hieq) "#Hm1". - iDestruct "Hm2" as (Hjeq) "#Hm2". - done. - - destruct j; last first. - { rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - iDestruct "Hm1" as (Hieq) "#Hm1". done. } - rewrite !lookup_total_insert. - rewrite lookup_total_insert_ne; [|done]. - rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - iDestruct "Hm1" as (Hieq) "#Hm1". - iDestruct "Hm2" as (Hjeq) "#Hm2". - iIntros (v p1) "Hm1'". - iSpecialize ("Hm1" $!v (Next p1)). - rewrite iMsg_base_eq. simpl. - rewrite iMsg_exist_eq. simpl. - iRewrite -"Hm1" in "Hm1'". - iExists END%proto. - iSpecialize ("Hm2" $!v (Next END%proto)). - iRewrite -"Hm2". - simpl. - iDestruct "Hm1'" as (x Heq) "[#Heq HP]". - iSplitL. - { iExists x. iSplit; [done|]. iFrame. done. } - iNext. iRewrite -"Heq". - rewrite insert_commute; [|done]. - rewrite insert_insert. - rewrite insert_commute; [|done]. - rewrite insert_insert. - rewrite iProto_consistent_unfold. - iIntros (i' j' m1' m2') "Hm1' Hm2'". - destruct i'. - { rewrite lookup_total_insert. - iDestruct (iProto_end_message_equivI with "Hm1'") as "H". - done. } - destruct i'. - { rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert. - iDestruct (iProto_end_message_equivI with "Hm1'") as "H". - done. } - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_alt. simpl. - iDestruct (iProto_end_message_equivI with "Hm1'") as "H". - done. - - destruct i; last first. - { rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_alt. simpl. - iDestruct (iProto_end_message_equivI with "Hm1") as "H". - done. } - rewrite !lookup_total_insert. - rewrite lookup_total_insert_ne; [|done]. - rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - iDestruct "Hm1" as (Hieq) "#Hm1". done. - - destruct i. - { rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert. - rewrite !iProto_message_equivI. - iDestruct "Hm1" as (Hieq) "#Hm1". done. } - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_alt. simpl. - iDestruct (iProto_end_message_equivI with "Hm1") as "H". - done. + iProto_consistent_take_step. + iIntros (x) "HP". iExists x. iSplit; [done|]. iFrame. iNext. + iProto_consistent_take_step. Qed. Definition iProto_example3 `{!invGS Σ} : gmap nat (iProto Σ) := @@ -110,192 +107,21 @@ Definition iProto_example3 `{!invGS Σ} : gmap nat (iProto Σ) := Lemma iProto_example3_consistent `{!invGS Σ} : ⊢ iProto_consistent (@iProto_example3 Σ invGS0). Proof. - rewrite iProto_consistent_unfold. rewrite /iProto_example3. - iIntros (i j m1 m2) "Hm1 Hm2". - destruct i; last first. - { destruct i. - { rewrite lookup_total_insert_ne; [|done]. - rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm1" as (Hieq) "#Hm1". } - destruct i. - { rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - rewrite !iProto_message_equivI. - by iDestruct "Hm1" as (Hieq) "#Hm1". } - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_empty=> /=. - by rewrite iProto_end_message_equivI. } - destruct j. - { rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm2" as (Hjeq) "#Hm2". } - destruct j; last first. - { rewrite !lookup_total_insert. - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - destruct j. - { rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm2" as (Hjeq) "#Hm2". } - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_empty=> /=. - by rewrite iProto_end_message_equivI. } - rewrite !lookup_total_insert. - rewrite lookup_total_insert_ne; [|done]. - rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - iDestruct "Hm1" as (Hieq) "#Hm1". - iDestruct "Hm2" as (Hjeq) "#Hm2". - iIntros (v p1) "Hm1'". - iSpecialize ("Hm1" $!v (Next p1)). - rewrite !iMsg_base_eq. - rewrite !iMsg_exist_eq. - iRewrite -"Hm1" in "Hm1'". - iDestruct "Hm1'" as (x Heq) "[#Hm1' _]". - iSpecialize ("Hm2" $!v (Next (<(Send, 2)> iMsg_base_def #x True END)))%proto. - iExists (<(Send, 2)> iMsg_base_def #x True END)%proto. - iRewrite -"Hm2". - simpl. - iSplitL. - { iExists x. iSplit; [done|]. iFrame. done. } - iNext. - rewrite iProto_consistent_unfold. - rewrite insert_commute; [|done]. - rewrite insert_insert. - rewrite insert_commute; [|done]. - rewrite insert_insert. - iRewrite -"Hm1'". - iClear (m1 m2 Hieq Hjeq p1 v Heq) "Hm1 Hm2 Hm1'". - iIntros (i j m1 m2) "Hm1 Hm2". - destruct i. - { rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm1" as (Hieq) "#Hm1". } - rewrite lookup_total_insert_ne; [|done]. - destruct i; last first. - { destruct i. - { rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm1" as (Hieq) "#Hm1". } - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_empty=> /=. - by rewrite iProto_end_message_equivI. } - rewrite lookup_total_insert. - destruct j. - { rewrite lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm2" as (Hjeq) "#Hm2". } - rewrite lookup_total_insert_ne; [|done]. - destruct j. - { rewrite lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm2" as (Hjeq) "#Hm2". } - rewrite lookup_total_insert_ne; [|done]. - destruct j; last first. - { rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_empty. - by rewrite iProto_end_message_equivI. } - rewrite lookup_total_insert. - rewrite !iProto_message_equivI. - iDestruct "Hm1" as (Hieq) "#Hm1". - iDestruct "Hm2" as (Hjeq) "#Hm2". - iIntros (v p1) "Hm1'". - iSpecialize ("Hm1" $!v (Next p1)). - iRewrite -"Hm1" in "Hm1'". - iDestruct "Hm1'" as (Heq) "[#Hm1' _]". - iSpecialize ("Hm2" $!v (Next (<(Send, 0)> iMsg_base_def #x True END)))%proto. - iExists (<(Send, 0)> iMsg_base_def #x True END)%proto. - iRewrite -"Hm2". - simpl. - iSplitL. - { iExists x. iSplit; [done|]. iFrame. done. } - iNext. - rewrite iProto_consistent_unfold. - rewrite (insert_commute _ 1 2); [|done]. - rewrite (insert_commute _ 1 0); [|done]. - rewrite insert_insert. - rewrite (insert_commute _ 2 0); [|done]. - rewrite (insert_commute _ 2 1); [|done]. - rewrite insert_insert. - iRewrite -"Hm1'". - iClear (m1 m2 Hieq Hjeq p1 v Heq) "Hm1 Hm2 Hm1'". - iIntros (i j m1 m2) "Hm1 Hm2". - destruct i. - { rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm1" as (Hieq) "#Hm1". } - rewrite lookup_total_insert_ne; [|done]. - destruct i. - { rewrite lookup_total_insert. - by rewrite iProto_end_message_equivI. } - rewrite lookup_total_insert_ne; [|done]. - destruct i; last first. - { rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_empty. - by rewrite iProto_end_message_equivI. } - rewrite lookup_total_insert. - destruct j; last first. - { rewrite lookup_total_insert_ne; [|done]. - destruct j. - { rewrite lookup_total_insert. - by rewrite iProto_end_message_equivI. } - rewrite lookup_total_insert_ne; [|done]. - destruct j. - { rewrite lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm2" as (Hjeq) "#Hm2". } - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_empty. - by rewrite iProto_end_message_equivI. } - rewrite lookup_total_insert. - rewrite !iProto_message_equivI. - iDestruct "Hm1" as (Hieq) "#Hm1". - iDestruct "Hm2" as (Hjeq) "#Hm2". - iIntros (v p1) "Hm1'". - iSpecialize ("Hm1" $!v (Next p1)). - iRewrite -"Hm1" in "Hm1'". - iDestruct "Hm1'" as (Heq) "[#Hm1' _]". - iSpecialize ("Hm2" $!v (Next END))%proto. - iExists END%proto. - iRewrite -"Hm2". - simpl. - iSplitL; [done|]. - rewrite insert_insert. - rewrite insert_commute; [|done]. - rewrite (insert_commute _ 2 1); [|done]. - rewrite insert_insert. - iNext. - iRewrite -"Hm1'". - rewrite iProto_consistent_unfold. - iClear (m1 m2 Hieq Hjeq p1 v Heq) "Hm1 Hm2 Hm1'". - iIntros (i j m1 m2) "Hm1 Hm2". - destruct i. - { rewrite lookup_total_insert. - by rewrite !iProto_end_message_equivI. } - rewrite lookup_total_insert_ne; [|done]. - destruct i. - { rewrite lookup_total_insert. - by rewrite !iProto_end_message_equivI. } - rewrite lookup_total_insert_ne; [|done]. - destruct i. - { rewrite lookup_total_insert. - by rewrite !iProto_end_message_equivI. } - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_empty. - by rewrite iProto_end_message_equivI. + iProto_consistent_take_step. + iIntros (x) "_". iExists x. iSplit; [done|]. iSplit; [done|]. iNext. + iProto_consistent_take_step. + iIntros "_". iExists x. iSplit; [done|]. iSplit; [done|]. iNext. + iProto_consistent_take_step. + iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. + iProto_consistent_take_step. Qed. Definition roundtrip_prog : val := λ: <>, let: "cs" := new_chan #3 in - let: "c0" := get_chan "cs" #0 in - let: "c1" := get_chan "cs" #1 in + let: "c0" := get_chan "cs" #0 in + let: "c1" := get_chan "cs" #1 in let: "c2" := get_chan "cs" #2 in Fork (let: "x" := recv "c1" #0 in send "c1" #2 "x");; Fork (let: "x" := recv "c2" #1 in send "c2" #0 "x");; @@ -306,7 +132,6 @@ Section channel. Implicit Types p : iProto Σ. Implicit Types TT : tele. - (* TODO: Fix nat/Z coercion. *) Lemma roundtrip_prog_spec : {{{ True }}} roundtrip_prog #() {{{ RET #42 ; True }}}. @@ -324,15 +149,9 @@ Section channel. wp_smart_apply (get_chan_spec _ 2 with "Hcs"); [set_solver|]. iIntros (c2) "[Hc2 Hcs]". wp_smart_apply (wp_fork with "[Hc1]"). - { iIntros "!>". - wp_recv (x) as "_". - wp_send with "[//]". - done. } + { iIntros "!>". wp_recv (x) as "_". wp_send with "[//]". done. } wp_smart_apply (wp_fork with "[Hc2]"). - { iIntros "!>". - wp_recv (x) as "_". - wp_send with "[//]". - done. } + { iIntros "!>". wp_recv (x) as "_". wp_send with "[//]". done. } wp_send with "[//]". wp_recv as "_". by iApply "HΦ". @@ -355,189 +174,15 @@ Section example4. Lemma iProto_example4_consistent : ⊢ iProto_consistent (iProto_example4). Proof. - rewrite iProto_consistent_unfold. rewrite /iProto_example4. - iIntros (i j m1 m2) "Hm1 Hm2". - destruct i; last first. - { destruct i. - { rewrite lookup_total_insert_ne; [|done]. - rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm1" as (Hieq) "#Hm1". } - destruct i. - { rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - rewrite !iProto_message_equivI. - by iDestruct "Hm1" as (Hieq) "#Hm1". } - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_empty=> /=. - by rewrite iProto_end_message_equivI. } - destruct j. - { rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm2" as (Hjeq) "#Hm2". } - destruct j; last first. - { rewrite !lookup_total_insert. - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - destruct j. - { rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm2" as (Hjeq) "#Hm2". } - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_empty=> /=. - by rewrite iProto_end_message_equivI. } - rewrite !lookup_total_insert. - rewrite lookup_total_insert_ne; [|done]. - rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - iDestruct "Hm1" as (Hieq) "#Hm1". - iDestruct "Hm2" as (Hjeq) "#Hm2". - iIntros (v p1) "Hm1'". - iSpecialize ("Hm1" $!v (Next p1)). - rewrite !iMsg_base_eq. - rewrite !iMsg_exist_eq. - iRewrite -"Hm1" in "Hm1'". - iDestruct "Hm1'" as (l x <-) "[#Hm1' Hl]". simpl. - iSpecialize ("Hm2" $!#l (Next (<(Send, 2)> iMsg_base_def #l - (l ↦ #(x+1)) END)))%proto. - Unshelve. 2-4: apply _. (* Why is this needed? *) - iExists (<(Send, 2)> iMsg_base_def #l (l ↦ #(x+1)) END)%proto. - simpl. - iSplitL. - { iRewrite -"Hm2". iExists l, x. iSplit; [done|]. iFrame. done. } - iNext. - rewrite iProto_consistent_unfold. - rewrite insert_commute; [|done]. - rewrite insert_insert. - rewrite insert_commute; [|done]. - rewrite insert_insert. - iRewrite -"Hm1'". - iClear (m1 m2 Hieq Hjeq p1) "Hm1 Hm2 Hm1'". - iIntros (i j m1 m2) "Hm1 Hm2". - destruct i. - { rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm1" as (Hieq) "#Hm1". } - rewrite lookup_total_insert_ne; [|done]. - destruct i; last first. - { destruct i. - { rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm1" as (Hieq) "#Hm1". } - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_empty=> /=. - by rewrite iProto_end_message_equivI. } - rewrite lookup_total_insert. - destruct j. - { rewrite lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm2" as (Hjeq) "#Hm2". } - rewrite lookup_total_insert_ne; [|done]. - destruct j. - { rewrite lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm2" as (Hjeq) "#Hm2". } - rewrite lookup_total_insert_ne; [|done]. - destruct j; last first. - { rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_empty. - by rewrite iProto_end_message_equivI. } - rewrite lookup_total_insert. - rewrite !iProto_message_equivI. - iDestruct "Hm1" as (Hieq) "#Hm1". - iDestruct "Hm2" as (Hjeq) "#Hm2". - iIntros (v p1) "Hm1'". - iSpecialize ("Hm1" $!v (Next p1)). - iRewrite -"Hm1" in "Hm1'". - iDestruct "Hm1'" as (<-) "[#Hm1' Hl]". - simpl. - iSpecialize ("Hm2" $!#l (Next (<(Send, 0)> iMsg_base_def #() (l ↦ #(x+1+1)) END)))%proto. - iExists (<(Send, 0)> iMsg_base_def #() (l ↦ #(x+1+1)) END)%proto. - Unshelve. 2-4: apply _. - iRewrite -"Hm2". - simpl. - iSplitL. - { iExists l, (x+1)%Z. iSplit; [done|]. iFrame. done. } - iNext. - rewrite iProto_consistent_unfold. - rewrite (insert_commute _ 1 2); [|done]. - rewrite (insert_commute _ 1 0); [|done]. - rewrite insert_insert. - rewrite (insert_commute _ 2 0); [|done]. - rewrite (insert_commute _ 2 1); [|done]. - rewrite insert_insert. - iRewrite -"Hm1'". - iClear (m1 m2 Hieq Hjeq p1) "Hm1 Hm2 Hm1'". - iIntros (i j m1 m2) "Hm1 Hm2". - destruct i. - { rewrite !lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm1" as (Hieq) "#Hm1". } - rewrite lookup_total_insert_ne; [|done]. - destruct i. - { rewrite lookup_total_insert. - by rewrite iProto_end_message_equivI. } - rewrite lookup_total_insert_ne; [|done]. - destruct i; last first. - { rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_empty. - by rewrite iProto_end_message_equivI. } - rewrite lookup_total_insert. - destruct j; last first. - { rewrite lookup_total_insert_ne; [|done]. - destruct j. - { rewrite lookup_total_insert. - by rewrite iProto_end_message_equivI. } - rewrite lookup_total_insert_ne; [|done]. - destruct j. - { rewrite lookup_total_insert. - rewrite !iProto_message_equivI. - by iDestruct "Hm2" as (Hjeq) "#Hm2". } - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_empty. - by rewrite iProto_end_message_equivI. } - rewrite lookup_total_insert. - rewrite !iProto_message_equivI. - iDestruct "Hm1" as (Hieq) "#Hm1". - iDestruct "Hm2" as (Hjeq) "#Hm2". - iIntros (v p1) "Hm1'". - iSpecialize ("Hm1" $!v (Next p1)). - iRewrite -"Hm1" in "Hm1'". - iDestruct "Hm1'" as (<-) "[#Hm1' Hl]". - iSpecialize ("Hm2" $!#() (Next END))%proto. - iExists END%proto. - iRewrite -"Hm2". - simpl. - replace (x + 1 + 1)%Z with (x+2)%Z by lia. - iSplitL; [iFrame;done|]. - rewrite insert_insert. - rewrite insert_commute; [|done]. - rewrite (insert_commute _ 2 1); [|done]. - rewrite insert_insert. - iNext. - iRewrite -"Hm1'". - rewrite iProto_consistent_unfold. - iClear (m1 m2 Hieq Hjeq p1) "Hm1 Hm2 Hm1'". - iIntros (i j m1 m2) "Hm1 Hm2". - destruct i. - { rewrite lookup_total_insert. - by rewrite !iProto_end_message_equivI. } - rewrite lookup_total_insert_ne; [|done]. - destruct i. - { rewrite lookup_total_insert. - by rewrite !iProto_end_message_equivI. } - rewrite lookup_total_insert_ne; [|done]. - destruct i. - { rewrite lookup_total_insert. - by rewrite !iProto_end_message_equivI. } - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_empty. - by rewrite iProto_end_message_equivI. + iProto_consistent_take_step. + iIntros (l x) "Hloc". iExists _, _. iSplit; [done|]. iFrame. iNext. + iProto_consistent_take_step. + iIntros "Hloc". iExists _, _. iSplit; [done|]. iFrame. iNext. + iProto_consistent_take_step. + iIntros "Hloc". iSplit; [done|]. + replace (x + 1 + 1)%Z with (x+2)%Z by lia. iFrame. iNext. + iProto_consistent_take_step. Qed. End example4. @@ -545,9 +190,9 @@ End example4. Definition roundtrip_ref_prog : val := λ: <>, let: "cs" := new_chan #3 in - let: "c0" := get_chan "cs" #0 in - let: "c1" := get_chan "cs" #1 in - let: "c2" := get_chan "cs" #2 in + let: "c0" := get_chan "cs" #0 in + let: "c1" := get_chan "cs" #1 in + let: "c2" := get_chan "cs" #2 in Fork (let: "l" := recv "c1" #0 in "l" <- !"l" + #1;; send "c1" #2 "l");; Fork (let: "l" := recv "c2" #1 in "l" <- !"l" + #1;; send "c2" #0 #());; let: "l" := ref #40 in send "c0" #1 "l";; recv "c0" #2;; !"l". -- GitLab From 25aa020787a63cf67d228a48419b10177621036a Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Sat, 27 Jan 2024 03:52:37 +0100 Subject: [PATCH 24/81] More automation --- .../multi_proto_consistency_examples.v | 102 +++++++++++++++++- 1 file changed, 101 insertions(+), 1 deletion(-) diff --git a/theories/channel/multi_proto_consistency_examples.v b/theories/channel/multi_proto_consistency_examples.v index 7378b82..af4b7af 100644 --- a/theories/channel/multi_proto_consistency_examples.v +++ b/theories/channel/multi_proto_consistency_examples.v @@ -49,6 +49,9 @@ Proof. iNext. iRewrite -"Hp1". done. Qed. +(* TODO: Improve automation *) +(* Could clean up repeated inserts to save traverses *) +(* Need bug fixin (see example 5) *) Tactic Notation "iProto_consistent_take_step" := let i := fresh in let j := fresh in @@ -66,6 +69,8 @@ Tactic Notation "iProto_consistent_take_step" := repeat (rewrite lookup_total_insert_ne; [|lia]); try (by rewrite lookup_total_empty iProto_end_message_equivI)]); repeat (rewrite lookup_total_insert_ne; [|lia]); + try rewrite lookup_total_empty; + try (by iProto_end_message_equivI); rewrite lookup_total_insert; iDestruct (iProto_message_equivI with "Hm2") as "[%Heq2 Hm2']";simplify_eq; @@ -75,7 +80,15 @@ Tactic Notation "iProto_consistent_take_step" := iSplitL; [iFrame "#"|]; iSplitL; [iPureIntro; tc_solve|]; iSplitL; [iPureIntro; tc_solve|]; - simpl; iClear "#"; clear m1 m2). + simpl; iClear "#"; clear m1 m2); + try (repeat (rewrite (insert_commute _ _ i); [|done]); + rewrite insert_insert; + repeat (rewrite (insert_commute _ _ j); [|done]); + rewrite insert_insert). + +Tactic Notation "clean_map" constr(i) := + repeat (rewrite (insert_commute _ _ i); [|done]); + rewrite (insert_insert _ i). Definition iProto_example1 {Σ} : gmap nat (iProto Σ) := ∅. @@ -228,3 +241,90 @@ Section proof. Qed. End proof. + +Section example5. + Context `{!heapGS Σ}. + + Definition iProto_example5 : gmap nat (iProto Σ) := + <[0 := (<(Send, 1) @ (x:Z)> MSG #x ; <(Send, 2)> MSG #x ; + <(Recv, 3)> MSG #x; <(Recv, 4)> MSG #x; END)%proto]> + (<[1 := (<(Recv, 0) @ (x:Z)> MSG #x ; + <(Send, 3)> MSG #x; END)%proto]> + (<[2 := (<(Recv, 0) @ (x:Z)> MSG #x ; + <(Send, 4)> MSG #x ; END)%proto]> + (<[3 := (<(Recv, 1) @ (x:Z)> MSG #x ; + <(Send, 0)> MSG #x; END)%proto]> + (<[4 := (<(Recv, 2) @ (x:Z)> MSG #x ; + <(Send, 0)> MSG #x ; END)%proto]> + ∅)))). + + Lemma iProto_example5_consistent : + ⊢ iProto_consistent iProto_example5. + Proof. + rewrite /iProto_example5. + iProto_consistent_take_step. + iIntros (x) "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 0. clean_map 1. + iProto_consistent_take_step. + - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 0. clean_map 2. + iProto_consistent_take_step. + + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 1. clean_map 3. + iProto_consistent_take_step. + * iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 2. clean_map 4. + iProto_consistent_take_step. + iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 0. clean_map 3. + iProto_consistent_take_step. + iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. + iProto_consistent_take_step. + * iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 3. clean_map 0. + iProto_consistent_take_step. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 2. clean_map 4. + iProto_consistent_take_step. + iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 4. clean_map 0. + iProto_consistent_take_step. + + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 2. clean_map 4. + iProto_consistent_take_step. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 1. clean_map 3. + iProto_consistent_take_step. + iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 3. clean_map 0. + iProto_consistent_take_step. + iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 4. clean_map 0. + iProto_consistent_take_step. + - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 1. clean_map 3. + iProto_consistent_take_step. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 0. clean_map 2. + iProto_consistent_take_step. + + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 2. clean_map 4. + iProto_consistent_take_step. + iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 3. clean_map 0. + iProto_consistent_take_step. + iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 4. clean_map 0. + iProto_consistent_take_step. + + iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 3. clean_map 0. + iProto_consistent_take_step. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 2. clean_map 4. + iProto_consistent_take_step. + iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 4. clean_map 0. + iProto_consistent_take_step. + Qed. + +End example5. -- GitLab From 66c993e21ae588d9789680f626709bc581102242 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Sat, 27 Jan 2024 17:48:57 +0100 Subject: [PATCH 25/81] Added two_buyer protocol --- .../multi_proto_consistency_examples.v | 76 +++++++++++++++++-- 1 file changed, 68 insertions(+), 8 deletions(-) diff --git a/theories/channel/multi_proto_consistency_examples.v b/theories/channel/multi_proto_consistency_examples.v index af4b7af..c5dcdfd 100644 --- a/theories/channel/multi_proto_consistency_examples.v +++ b/theories/channel/multi_proto_consistency_examples.v @@ -302,29 +302,89 @@ Section example5. clean_map 4. clean_map 0. iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. - clean_map 1. clean_map 3. + clean_map 1. clean_map 3. iProto_consistent_take_step. iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. - clean_map 0. clean_map 2. + clean_map 0. clean_map 2. iProto_consistent_take_step. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. - clean_map 2. clean_map 4. + clean_map 2. clean_map 4. iProto_consistent_take_step. iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. - clean_map 3. clean_map 0. + clean_map 3. clean_map 0. iProto_consistent_take_step. iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. - clean_map 4. clean_map 0. + clean_map 4. clean_map 0. iProto_consistent_take_step. + iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. - clean_map 3. clean_map 0. + clean_map 3. clean_map 0. iProto_consistent_take_step. iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. - clean_map 2. clean_map 4. + clean_map 2. clean_map 4. iProto_consistent_take_step. iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. - clean_map 4. clean_map 0. + clean_map 4. clean_map 0. iProto_consistent_take_step. Qed. End example5. + +Section two_buyer. + Context `{!heapGS Σ}. + + Definition two_buyer_prot : gmap nat (iProto Σ) := + <[0 := (<(Send, 1) @ (title:Z)> MSG #title ; + <(Recv, 1) @ (quote:Z)> MSG #quote ; + <(Send, 2) @ (contrib:Z)> MSG #contrib ; END)%proto]> + (<[1 := (<(Recv, 0) @ (title:Z)> MSG #title ; + <(Send, 0) @ (quote:Z)> MSG #quote ; + <(Send, 2)> MSG #quote ; + <(Recv, 2) @ (b:bool)> MSG #b ; + if b then + <(Recv, 2) @ (address:Z)> MSG #address ; + <(Send, 2) @ (date:Z)> MSG #date ; END + else END)%proto]> + (<[2 := (<(Recv, 1) @ (quote:Z)> MSG #quote ; + <(Recv, 0) @ (contrib:Z)> MSG #contrib ; + if bool_decide (contrib >= quote/2)%Z then + <(Send, 1)> MSG #true ; + <(Send, 1) @ (address:Z)> MSG #address ; + <(Recv, 1) @ (date:Z)> MSG #date ; END + else + <(Send, 1)> MSG #false ; END)%proto]> + ∅)). + + Lemma two_buyer_prot_consistent : + ⊢ iProto_consistent two_buyer_prot. + Proof. + rewrite /two_buyer_prot. + iProto_consistent_take_step. + iIntros (title) "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 0. clean_map 1. + iProto_consistent_take_step. + iIntros (quote) "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 0. clean_map 1. + iProto_consistent_take_step. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 1. clean_map 2. + iProto_consistent_take_step. + iIntros (contrib) "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 0. clean_map 2. + case_bool_decide. + - iProto_consistent_take_step. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 2. clean_map 1. + iProto_consistent_take_step. + iIntros (address) "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 2. clean_map 1. + iProto_consistent_take_step. + iIntros (date) "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 2. clean_map 1. + iProto_consistent_take_step. + - iProto_consistent_take_step. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + clean_map 2. clean_map 1. + iProto_consistent_take_step. + Qed. + +End two_buyer. -- GitLab From c741f4dd3d40a92d3fd4f9e40696efe2e0dcd2d2 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Sat, 27 Jan 2024 18:55:47 +0100 Subject: [PATCH 26/81] Two Buyer Ref protocol --- .../multi_proto_consistency_examples.v | 151 +++++++++++++----- 1 file changed, 114 insertions(+), 37 deletions(-) diff --git a/theories/channel/multi_proto_consistency_examples.v b/theories/channel/multi_proto_consistency_examples.v index c5dcdfd..27537f2 100644 --- a/theories/channel/multi_proto_consistency_examples.v +++ b/theories/channel/multi_proto_consistency_examples.v @@ -51,12 +51,12 @@ Qed. (* TODO: Improve automation *) (* Could clean up repeated inserts to save traverses *) -(* Need bug fixin (see example 5) *) Tactic Notation "iProto_consistent_take_step" := let i := fresh in let j := fresh in let m1 := fresh in let m2 := fresh in + try iNext; iApply iProto_consistent_equiv_proof; iIntros (i j m1 m2) "#Hm1 #Hm2"; repeat (destruct i as [|i]; @@ -107,7 +107,7 @@ Lemma iProto_example2_consistent `{!invGS Σ} (P : iProp Σ) : Proof. rewrite /iProto_example2. iProto_consistent_take_step. - iIntros (x) "HP". iExists x. iSplit; [done|]. iFrame. iNext. + iIntros (x) "HP". iExists x. iSplit; [done|]. iFrame. iProto_consistent_take_step. Qed. @@ -122,11 +122,11 @@ Lemma iProto_example3_consistent `{!invGS Σ} : Proof. rewrite /iProto_example3. iProto_consistent_take_step. - iIntros (x) "_". iExists x. iSplit; [done|]. iSplit; [done|]. iNext. + iIntros (x) "_". iExists x. iSplit; [done|]. iSplit; [done|]. iProto_consistent_take_step. - iIntros "_". iExists x. iSplit; [done|]. iSplit; [done|]. iNext. + iIntros "_". iExists x. iSplit; [done|]. iSplit; [done|]. iProto_consistent_take_step. - iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. + iIntros "_". iSplit; [done|]. iSplit; [done|]. iProto_consistent_take_step. Qed. @@ -189,12 +189,12 @@ Section example4. Proof. rewrite /iProto_example4. iProto_consistent_take_step. - iIntros (l x) "Hloc". iExists _, _. iSplit; [done|]. iFrame. iNext. + iIntros (l x) "Hloc". iExists _, _. iSplit; [done|]. iFrame. iProto_consistent_take_step. - iIntros "Hloc". iExists _, _. iSplit; [done|]. iFrame. iNext. + iIntros "Hloc". iExists _, _. iSplit; [done|]. iFrame. iProto_consistent_take_step. iIntros "Hloc". iSplit; [done|]. - replace (x + 1 + 1)%Z with (x+2)%Z by lia. iFrame. iNext. + replace (x + 1 + 1)%Z with (x+2)%Z by lia. iFrame. iProto_consistent_take_step. Qed. @@ -263,66 +263,66 @@ Section example5. Proof. rewrite /iProto_example5. iProto_consistent_take_step. - iIntros (x) "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + iIntros (x) "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 0. clean_map 1. iProto_consistent_take_step. - - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 0. clean_map 2. iProto_consistent_take_step. - + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 1. clean_map 3. iProto_consistent_take_step. - * iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + * iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 2. clean_map 4. iProto_consistent_take_step. - iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. + iIntros "_". iSplit; [done|]. iSplit; [done|]. clean_map 0. clean_map 3. iProto_consistent_take_step. - iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. + iIntros "_". iSplit; [done|]. iSplit; [done|]. iProto_consistent_take_step. - * iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. + * iIntros "_". iSplit; [done|]. iSplit; [done|]. clean_map 3. clean_map 0. iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 2. clean_map 4. iProto_consistent_take_step. - iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. + iIntros "_". iSplit; [done|]. iSplit; [done|]. clean_map 4. clean_map 0. iProto_consistent_take_step. - + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 2. clean_map 4. iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 1. clean_map 3. iProto_consistent_take_step. - iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. + iIntros "_". iSplit; [done|]. iSplit; [done|]. clean_map 3. clean_map 0. iProto_consistent_take_step. - iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. + iIntros "_". iSplit; [done|]. iSplit; [done|]. clean_map 4. clean_map 0. iProto_consistent_take_step. - - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 1. clean_map 3. iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 0. clean_map 2. iProto_consistent_take_step. - + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 2. clean_map 4. iProto_consistent_take_step. - iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. + iIntros "_". iSplit; [done|]. iSplit; [done|]. clean_map 3. clean_map 0. iProto_consistent_take_step. - iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. + iIntros "_". iSplit; [done|]. iSplit; [done|]. clean_map 4. clean_map 0. iProto_consistent_take_step. - + iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. + + iIntros "_". iSplit; [done|]. iSplit; [done|]. clean_map 3. clean_map 0. iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 2. clean_map 4. iProto_consistent_take_step. - iIntros "_". iSplit; [done|]. iSplit; [done|]. iNext. + iIntros "_". iSplit; [done|]. iSplit; [done|]. clean_map 4. clean_map 0. iProto_consistent_take_step. Qed. @@ -359,32 +359,109 @@ Section two_buyer. Proof. rewrite /two_buyer_prot. iProto_consistent_take_step. - iIntros (title) "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + iIntros (title) "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 0. clean_map 1. iProto_consistent_take_step. - iIntros (quote) "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + iIntros (quote) "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 0. clean_map 1. iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 1. clean_map 2. iProto_consistent_take_step. - iIntros (contrib) "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + iIntros (contrib) "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 0. clean_map 2. case_bool_decide. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 2. clean_map 1. iProto_consistent_take_step. - iIntros (address) "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + iIntros (address) "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 2. clean_map 1. iProto_consistent_take_step. - iIntros (date) "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + iIntros (date) "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 2. clean_map 1. iProto_consistent_take_step. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iNext. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 2. clean_map 1. iProto_consistent_take_step. Qed. End two_buyer. + +Section two_buyer_ref. + Context `{!heapGS Σ}. + + Definition two_buyer_ref_b1_prot : iProto Σ := + (<(Send, 1) @ (title:Z)> MSG #title ; + <(Recv, 1) @ (quote:Z)> MSG #quote ; + <(Send, 2) @ (l : loc) (amount:Z) (contrib:Z)> + MSG (#l,#contrib) {{ l ↦ #amount }} ; + <(Recv, 2) @ (b : bool)> + MSG #b {{ l ↦ #(if b then amount - contrib else amount) }}; + END)%proto. + + Definition two_buyer_ref_s_prot : iProto Σ := + (<(Recv, 0) @ (title:Z)> MSG #title ; + <(Send, 0) @ (quote:Z)> MSG #quote ; + <(Send, 2)> MSG #quote ; + <(Recv, 2) @ (b:bool)> MSG #b ; + if b then + <(Recv, 2) @ (l2 : loc) (amount2:Z) (address:Z)> + MSG (#l2,#address) {{ l2 ↦ #amount2 }} ; + <(Send, 2) @ (date : Z)> MSG #date {{ l2 ↦ #(amount2-quote) }}; END + else END)%proto. + + Definition two_buyer_ref_b2_prot : iProto Σ := + (<(Recv, 1) @ (quote:Z)> MSG #quote ; + <(Recv, 0) @ (l1 : loc) (amount1:Z) (contrib:Z)> + MSG (#l1,#contrib) {{ l1 ↦ #amount1 }}; + <(Send, 0) @ (b : bool)> + MSG #b {{ l1 ↦ #(if b then amount1 - contrib else amount1) }}; + <(Send, 1)> MSG #b; + if b then + <(Send, 1) @ (l2 : loc) (amount2:Z) (address:Z)> + MSG (#l2,#address) {{ l2 ↦ #amount2 }} ; + <(Recv, 1) @ (date : Z)> MSG #date {{ l2 ↦ #(amount2-quote) }}; + END + else END)%proto. + + Definition two_buyer_ref_prot : gmap nat (iProto Σ) := + <[0 := two_buyer_ref_b1_prot]> + (<[1 := two_buyer_ref_s_prot]> + (<[2 := two_buyer_ref_b2_prot]> + ∅)). + + Lemma two_buyer_ref_prot_consistent : + ⊢ iProto_consistent two_buyer_ref_prot. + Proof. + rewrite /two_buyer_prot. + iProto_consistent_take_step. + iIntros (title) "_". iExists _. iSplit; [done|]. iSplit; [done|]. + clean_map 0. clean_map 1. + iProto_consistent_take_step. + iIntros (quote) "_". iExists _. iSplit; [done|]. iSplit; [done|]. + clean_map 0. clean_map 1. + iProto_consistent_take_step. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. + clean_map 1. clean_map 2. + iProto_consistent_take_step. + iIntros (l1 amount1 contrib) "Hl1". iExists _,_,_. iSplit; [done|]. iFrame. + clean_map 0. clean_map 2. + iProto_consistent_take_step. + iIntros (b) "Hl1". iExists _. iSplit; [done|]. iFrame. + clean_map 0. clean_map 2. + iProto_consistent_take_step. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. + clean_map 1. clean_map 2. + destruct b. + - iProto_consistent_take_step. + iIntros (l2 amount2 address) "Hl2". iExists _,_,_. iSplit; [done|]. iFrame. + clean_map 2. clean_map 1. + iProto_consistent_take_step. + iIntros (date) "Hl2". iExists _. iSplit; [done|]. iFrame. + iProto_consistent_take_step. + - iProto_consistent_take_step. + Qed. + +End two_buyer_ref. -- GitLab From 34fe7e0d7a7a2a8082b671a3e3569aa585183ac1 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Sat, 27 Jan 2024 23:46:29 +0100 Subject: [PATCH 27/81] Added recursive example --- .../multi_proto_consistency_examples.v | 148 ++++++++++++++---- 1 file changed, 118 insertions(+), 30 deletions(-) diff --git a/theories/channel/multi_proto_consistency_examples.v b/theories/channel/multi_proto_consistency_examples.v index 27537f2..31aaa16 100644 --- a/theories/channel/multi_proto_consistency_examples.v +++ b/theories/channel/multi_proto_consistency_examples.v @@ -80,47 +80,46 @@ Tactic Notation "iProto_consistent_take_step" := iSplitL; [iFrame "#"|]; iSplitL; [iPureIntro; tc_solve|]; iSplitL; [iPureIntro; tc_solve|]; - simpl; iClear "#"; clear m1 m2); + simpl; iClear "Hm1 Hm2"; clear m1 m2); try (repeat (rewrite (insert_commute _ _ i); [|done]); rewrite insert_insert; repeat (rewrite (insert_commute _ _ j); [|done]); rewrite insert_insert). Tactic Notation "clean_map" constr(i) := - repeat (rewrite (insert_commute _ _ i); [|done]); + iEval (repeat (rewrite (insert_commute _ _ i); [|done])); rewrite (insert_insert _ i). -Definition iProto_example1 {Σ} : gmap nat (iProto Σ) := - ∅. +Definition iProto_empty {Σ} : gmap nat (iProto Σ) := ∅. -Lemma iProto_example1_consistent {Σ} : - ⊢ iProto_consistent (@iProto_example1 Σ). +Lemma iProto_consistent_empty {Σ} : + ⊢ iProto_consistent (@iProto_empty Σ). Proof. iProto_consistent_take_step. Qed. -Definition iProto_example2 `{!invGS Σ} (P : iProp Σ) : gmap nat (iProto Σ) := +Definition iProto_binary `{!invGS Σ} (P : iProp Σ) : gmap nat (iProto Σ) := <[0 := (<(Send, 1) @ (x:Z)> MSG #x {{ P }} ; END)%proto ]> (<[1 := (<(Recv, 0) @ (x:Z)> MSG #x {{ P }} ; END)%proto ]> ∅). -Lemma iProto_example2_consistent `{!invGS Σ} (P : iProp Σ) : - ⊢ iProto_consistent (@iProto_example2 Σ invGS0 P). +Lemma iProto_binary_consistent `{!invGS Σ} (P : iProp Σ) : + ⊢ iProto_consistent (@iProto_binary Σ invGS0 P). Proof. - rewrite /iProto_example2. + rewrite /iProto_binary. iProto_consistent_take_step. iIntros (x) "HP". iExists x. iSplit; [done|]. iFrame. iProto_consistent_take_step. Qed. -Definition iProto_example3 `{!invGS Σ} : gmap nat (iProto Σ) := +Definition iProto_roundtrip `{!invGS Σ} : gmap nat (iProto Σ) := <[0 := (<(Send, 1) @ (x:Z)> MSG #x ; <(Recv, 2)> MSG #x; END)%proto ]> (<[1 := (<(Recv, 0) @ (x:Z)> MSG #x ; <(Send, 2)> MSG #x; END)%proto ]> (<[2 := (<(Recv, 1) @ (x:Z)> MSG #x ; <(Send, 0)> MSG #x; END)%proto ]> ∅)). -Lemma iProto_example3_consistent `{!invGS Σ} : - ⊢ iProto_consistent (@iProto_example3 Σ invGS0). +Lemma iProto_roundtrip_consistent `{!invGS Σ} : + ⊢ iProto_consistent (@iProto_roundtrip Σ invGS0). Proof. - rewrite /iProto_example3. + rewrite /iProto_roundtrip. iProto_consistent_take_step. iIntros (x) "_". iExists x. iSplit; [done|]. iSplit; [done|]. iProto_consistent_take_step. @@ -150,10 +149,10 @@ Section channel. {{{ True }}} roundtrip_prog #() {{{ RET #42 ; True }}}. Proof using chanG0 heapGS0 Σ. iIntros (Φ) "_ HΦ". wp_lam. - wp_smart_apply (new_chan_spec 3 iProto_example3). + wp_smart_apply (new_chan_spec 3 iProto_roundtrip). { intros i Hle. destruct i as [|[|[]]]; try set_solver. lia. } { set_solver. } - { iApply iProto_example3_consistent. } + { iApply iProto_roundtrip_consistent. } iIntros (cs) "Hcs". wp_smart_apply (get_chan_spec _ 0 with "Hcs"); [set_solver|]. iIntros (c0) "[Hc0 Hcs]". @@ -172,10 +171,10 @@ Section channel. End channel. -Section example4. +Section roundtrip_ref. Context `{!heapGS Σ}. - Definition iProto_example4 : gmap nat (iProto Σ) := + Definition iProto_roundtrip_ref : gmap nat (iProto Σ) := <[0 := (<(Send, 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; <(Recv, 2)> MSG #() {{ l ↦ #(x+2) }} ; END)%proto]> (<[1 := (<(Recv, 0) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; @@ -184,10 +183,10 @@ Section example4. <(Send, 0)> MSG #() {{ l ↦ #(x+1) }}; END)%proto]> ∅)). - Lemma iProto_example4_consistent : - ⊢ iProto_consistent (iProto_example4). + Lemma iProto_roundtrip_ref_consistent : + ⊢ iProto_consistent iProto_roundtrip_ref. Proof. - rewrite /iProto_example4. + rewrite /iProto_roundtrip_ref. iProto_consistent_take_step. iIntros (l x) "Hloc". iExists _, _. iSplit; [done|]. iFrame. iProto_consistent_take_step. @@ -198,7 +197,7 @@ Section example4. iProto_consistent_take_step. Qed. -End example4. +End roundtrip_ref. Definition roundtrip_ref_prog : val := λ: <>, @@ -219,10 +218,10 @@ Section proof. {{{ True }}} roundtrip_ref_prog #() {{{ RET #42 ; True }}}. Proof using chanG0. iIntros (Φ) "_ HΦ". wp_lam. - wp_smart_apply (new_chan_spec 3 iProto_example4 with "[]"). + wp_smart_apply (new_chan_spec 3 iProto_roundtrip_ref with "[]"). { intros i Hle. destruct i as [|[|[]]]; try set_solver. lia. } { set_solver. } - { iApply iProto_example4_consistent. } + { iApply iProto_roundtrip_ref_consistent. } iIntros (cs) "Hcs". wp_smart_apply (get_chan_spec _ 0 with "Hcs"); [set_solver|]. iIntros (c0) "[Hc0 Hcs]". @@ -242,10 +241,99 @@ Section proof. End proof. -Section example5. +Section roundtrip_ref_rec. Context `{!heapGS Σ}. - Definition iProto_example5 : gmap nat (iProto Σ) := + Definition iProto_roundtrip_ref_rec1_aux (rec : iProto Σ) : iProto Σ := + (<(Send, 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; + <(Recv, 2)> MSG #() {{ l ↦ #(x+2) }} ; rec)%proto. + + Instance iProto_roundtrip_ref_rec1_contractive : + Contractive iProto_roundtrip_ref_rec1_aux. + Proof. solve_proto_contractive. Qed. + + Definition iProto_roundtrip_ref_rec1 := + fixpoint iProto_roundtrip_ref_rec1_aux. + + Lemma iProto_roundtrip_ref_rec1_unfold : + iProto_roundtrip_ref_rec1 ≡ + (iProto_roundtrip_ref_rec1_aux iProto_roundtrip_ref_rec1). + Proof. apply (fixpoint_unfold _). Qed. + + Definition iProto_roundtrip_ref_rec2_aux (rec : iProto Σ) : iProto Σ := + (<(Recv, 0) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; + <(Send, 2)> MSG #l {{ l ↦ #(x+1) }}; rec)%proto. + + Instance iProto_roundtrip_ref_rec2_contractive : + Contractive iProto_roundtrip_ref_rec2_aux. + Proof. solve_proto_contractive. Qed. + + Definition iProto_roundtrip_ref_rec2 := + fixpoint iProto_roundtrip_ref_rec2_aux. + + Lemma iProto_roundtrip_ref_rec2_unfold : + iProto_roundtrip_ref_rec2 ≡ + (iProto_roundtrip_ref_rec2_aux iProto_roundtrip_ref_rec2). + Proof. apply (fixpoint_unfold _). Qed. + + Definition iProto_roundtrip_ref_rec3_aux (rec : iProto Σ) : iProto Σ := + (<(Recv, 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; + <(Send, 0)> MSG #() {{ l ↦ #(x+1) }}; rec)%proto. + + Instance iProto_roundtrip_ref_rec3_contractive : + Contractive iProto_roundtrip_ref_rec3_aux. + Proof. solve_proto_contractive. Qed. + + Definition iProto_roundtrip_ref_rec3 := + fixpoint iProto_roundtrip_ref_rec3_aux. + + Lemma iProto_roundtrip_ref_rec3_unfold : + iProto_roundtrip_ref_rec3 ≡ + (iProto_roundtrip_ref_rec3_aux iProto_roundtrip_ref_rec3). + Proof. apply (fixpoint_unfold _). Qed. + + Definition iProto_roundtrip_ref_rec : gmap nat (iProto Σ) := + <[0 := iProto_roundtrip_ref_rec1]> + (<[1 := iProto_roundtrip_ref_rec2]> + (<[2 := iProto_roundtrip_ref_rec3]> ∅)). + + Lemma iProto_roundtrip_ref_rec_consistent : + ⊢ iProto_consistent iProto_roundtrip_ref_rec. + Proof. + iLöb as "IH". + rewrite /iProto_roundtrip_ref_rec. + iEval (rewrite iProto_roundtrip_ref_rec1_unfold). + iEval (rewrite iProto_roundtrip_ref_rec2_unfold). + iEval (rewrite iProto_roundtrip_ref_rec3_unfold). + iProto_consistent_take_step. + iIntros (l x) "Hloc". iExists _, _. iSplit; [done|]. iFrame. + iProto_consistent_take_step. + iIntros "Hloc". iExists _, _. iSplit; [done|]. iFrame. iNext. + rewrite iProto_roundtrip_ref_rec2_unfold. + iProto_consistent_take_step. + iIntros "Hloc". iSplit; [done|]. + replace (x + 1 + 1)%Z with (x+2)%Z by lia. iFrame. + rewrite -iProto_roundtrip_ref_rec2_unfold. + do 2 clean_map 0. do 2 clean_map 1. do 2 clean_map 2. + done. + Qed. + +End roundtrip_ref_rec. + +Section parallel. + Context `{!heapGS Σ}. + + (** + 0 + / \ + 1 2 + | | + 3 4 + \ / + 0 + *) + + Definition iProto_parallel : gmap nat (iProto Σ) := <[0 := (<(Send, 1) @ (x:Z)> MSG #x ; <(Send, 2)> MSG #x ; <(Recv, 3)> MSG #x; <(Recv, 4)> MSG #x; END)%proto]> (<[1 := (<(Recv, 0) @ (x:Z)> MSG #x ; @@ -258,10 +346,10 @@ Section example5. <(Send, 0)> MSG #x ; END)%proto]> ∅)))). - Lemma iProto_example5_consistent : - ⊢ iProto_consistent iProto_example5. + Lemma iProto_parallel_consistent : + ⊢ iProto_consistent iProto_parallel. Proof. - rewrite /iProto_example5. + rewrite /iProto_parallel. iProto_consistent_take_step. iIntros (x) "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 0. clean_map 1. @@ -327,7 +415,7 @@ Section example5. iProto_consistent_take_step. Qed. -End example5. +End parallel. Section two_buyer. Context `{!heapGS Σ}. -- GitLab From 7ed95fcd436ed61883afbfe625e553ccf73cb1b2 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Sun, 28 Jan 2024 11:31:50 +0100 Subject: [PATCH 28/81] Proved example using recursive protocol --- .../multi_proto_consistency_examples.v | 66 +++++++++++++++++++ 1 file changed, 66 insertions(+) diff --git a/theories/channel/multi_proto_consistency_examples.v b/theories/channel/multi_proto_consistency_examples.v index 31aaa16..ec48cd0 100644 --- a/theories/channel/multi_proto_consistency_examples.v +++ b/theories/channel/multi_proto_consistency_examples.v @@ -260,6 +260,11 @@ Section roundtrip_ref_rec. (iProto_roundtrip_ref_rec1_aux iProto_roundtrip_ref_rec1). Proof. apply (fixpoint_unfold _). Qed. + Global Instance iProto_roundtrip_ref_rec1_proto_unfold : + ProtoUnfold iProto_roundtrip_ref_rec1 + (iProto_roundtrip_ref_rec1_aux iProto_roundtrip_ref_rec1). + Proof. apply proto_unfold_eq, (fixpoint_unfold _). Qed. + Definition iProto_roundtrip_ref_rec2_aux (rec : iProto Σ) : iProto Σ := (<(Recv, 0) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; <(Send, 2)> MSG #l {{ l ↦ #(x+1) }}; rec)%proto. @@ -276,6 +281,11 @@ Section roundtrip_ref_rec. (iProto_roundtrip_ref_rec2_aux iProto_roundtrip_ref_rec2). Proof. apply (fixpoint_unfold _). Qed. + Global Instance iProto_roundtrip_ref_rec2_proto_unfold : + ProtoUnfold iProto_roundtrip_ref_rec2 + (iProto_roundtrip_ref_rec2_aux iProto_roundtrip_ref_rec2). + Proof. apply proto_unfold_eq, (fixpoint_unfold _). Qed. + Definition iProto_roundtrip_ref_rec3_aux (rec : iProto Σ) : iProto Σ := (<(Recv, 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; <(Send, 0)> MSG #() {{ l ↦ #(x+1) }}; rec)%proto. @@ -292,6 +302,11 @@ Section roundtrip_ref_rec. (iProto_roundtrip_ref_rec3_aux iProto_roundtrip_ref_rec3). Proof. apply (fixpoint_unfold _). Qed. + Global Instance iProto_roundtrip_ref_rec3_proto_unfold : + ProtoUnfold iProto_roundtrip_ref_rec3 + (iProto_roundtrip_ref_rec3_aux iProto_roundtrip_ref_rec3). + Proof. apply proto_unfold_eq, (fixpoint_unfold _). Qed. + Definition iProto_roundtrip_ref_rec : gmap nat (iProto Σ) := <[0 := iProto_roundtrip_ref_rec1]> (<[1 := iProto_roundtrip_ref_rec2]> @@ -320,6 +335,57 @@ Section roundtrip_ref_rec. End roundtrip_ref_rec. +Definition roundtrip_ref_rec_prog : val := + λ: <>, + let: "cs" := new_chan #3 in + let: "c0" := get_chan "cs" #0 in + let: "c1" := get_chan "cs" #1 in + let: "c2" := get_chan "cs" #2 in + Fork ((rec: "go" "c1" := + let: "l" := recv "c1" #0 in "l" <- !"l" + #1;; send "c1" #2 "l";; + "go" "c1") "c1");; + Fork ((rec: "go" "c2" := + let: "l" := recv "c2" #1 in "l" <- !"l" + #1;; send "c2" #0 #();; + "go" "c2") "c2");; + let: "l" := ref #38 in + send "c0" #1 "l";; recv "c0" #2;; + send "c0" #1 "l";; recv "c0" #2;; !"l". + +Section proof. + Context `{!heapGS Σ, !chanG Σ}. + Implicit Types p : iProto Σ. + Implicit Types TT : tele. + + Lemma roundtrip_ref_rec_prog_spec : + {{{ True }}} roundtrip_ref_rec_prog #() {{{ RET #42 ; True }}}. + Proof using chanG0. + iIntros (Φ) "_ HΦ". wp_lam. + wp_smart_apply (new_chan_spec 3 iProto_roundtrip_ref_rec with "[]"). + { intros i Hle. destruct i as [|[|[]]]; try set_solver. lia. } + { set_solver. } + { iApply iProto_roundtrip_ref_rec_consistent. } + iIntros (cs) "Hcs". + wp_smart_apply (get_chan_spec _ 0 with "Hcs"); [set_solver|]. + iIntros (c0) "[Hc0 Hcs]". + wp_smart_apply (get_chan_spec _ 1 with "Hcs"); [set_solver|]. + iIntros (c1) "[Hc1 Hcs]". + wp_smart_apply (get_chan_spec _ 2 with "Hcs"); [set_solver|]. + iIntros (c2) "[Hc2 Hcs]". + wp_smart_apply (wp_fork with "[Hc1]"). + { iIntros "!>". wp_pure _. iLöb as "IH". + wp_recv (l x) as "Hl". wp_load. wp_store. wp_send with "[$Hl]". + do 2 wp_pure _. by iApply "IH". } + wp_smart_apply (wp_fork with "[Hc2]"). + { iIntros "!>". wp_pure _. iLöb as "IH". + wp_recv (l x) as "Hl". wp_load. wp_store. wp_send with "[$Hl]". + do 2 wp_pure _. by iApply "IH". } + wp_alloc l as "Hl". wp_send with "[$Hl]". wp_recv as "Hl". + wp_send with "[$Hl]". wp_recv as "Hl". wp_load. + by iApply "HΦ". + Qed. + +End proof. + Section parallel. Context `{!heapGS Σ}. -- GitLab From 8a295caf8d31d23f629ecc7357de3a21165109a3 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Wed, 31 Jan 2024 02:17:21 +0100 Subject: [PATCH 29/81] Proved admitted ghost lemma --- theories/channel/multi_proto.v | 7 ++++++- 1 file changed, 6 insertions(+), 1 deletion(-) diff --git a/theories/channel/multi_proto.v b/theories/channel/multi_proto.v index 898b391..8cab819 100644 --- a/theories/channel/multi_proto.v +++ b/theories/channel/multi_proto.v @@ -1140,7 +1140,12 @@ Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p1))) -∗ own γ (gmap_view_frag j (DfracOwn 1) (Excl' (Next p2))) -∗ ⌜i ≠jâŒ. - Proof. Admitted. + Proof. + iIntros "Hown Hown'" (->). + iDestruct (own_valid_2 with "Hown Hown'") as "H". + rewrite uPred.cmra_valid_elim. + iDestruct "H" as %H%gmap_view_frag_op_validN. by destruct H. + Qed. Lemma iProto_step γ i j m1 m2 p1 v : iProto_ctx γ -∗ -- GitLab From b52cc463aa772e32a7c90d8722b8c705a4fb9e70 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Thu, 1 Feb 2024 03:02:27 +0100 Subject: [PATCH 30/81] Proved new_chan spec (but add a stronger side-condition) --- theories/channel/multi_channel.v | 407 +++++++++++++++++++++++-------- 1 file changed, 300 insertions(+), 107 deletions(-) diff --git a/theories/channel/multi_channel.v b/theories/channel/multi_channel.v index 0868ed6..590820f 100644 --- a/theories/channel/multi_channel.v +++ b/theories/channel/multi_channel.v @@ -28,22 +28,19 @@ From actris.channel Require Import multi_proto_model. From actris.channel Require Export multi_proto. Set Default Proof Using "Type". -(* TODO: Update new_chan definition to use pointers with offsets *) (** * The definition of the message-passing connectives *) Definition new_chan : val := - λ: "n", - let: "l" := AllocN ("n"*"n") NONEV in - let: "xxs" := AllocN "n" NONEV in - (rec: "go1" "i" := if: "i" = "n" then #() else - let: "xs" := AllocN "n" NONEV in - (rec: "go2" "j" := if: "j" = "n" then #() else - ("xs" +â‚— "j") <- ("l" +â‚— ("i"*"n"+"j"), "l" +â‚— ("j"*"n"+"i"));; - "go2" ("j"+#1)) #0;; - ("xxs" +â‚— "i") <- "xs";; - "go1" ("i"+#1)) #0;; "xxs". + λ: "n", (AllocN ("n"*"n") NONEV, "n"). Definition get_chan : val := - λ: "cs" "i", ! ("cs" +â‚— "i"). + λ: "cs" "i", ("cs","i"). + +Definition diverge : val := + λ: <>, (rec: "go" <> := "go" #())%V #(). + +Definition guard : val := + λ: "i" "n", + if: "i" < "n" then #() else diverge #(). Definition wait : val := rec: "go" "c" := @@ -52,27 +49,30 @@ Definition wait : val := | SOME <> => "go" "c" end. + +Definition pos (n i j : nat) : nat := i * n + j. +Definition vpos : val := λ: "n" "i" "j", "i"*"n" + "j". + Definition send : val := - λ: "c" "i" "v", - let: "len" := Fst "c" in - if: "i" < "len" then - let: "l" := Fst (! ((Snd "c") +â‚— "i")) in - "l" <- SOME "v";; wait "l" - (* OBS: Hacky *) - else (rec: "go" <> := "go" #())%V #(). + λ: "c" "j" "v", + let: "n" := Snd (Fst "c") in guard "j" "n";; + let: "ls" := Fst (Fst "c") in + let: "i" := Snd "c" in + let: "l" := "ls" +â‚— vpos "n" "i" "j" in + "l" <- SOME "v";; wait "l". +(* TODO: Move recursion further in *) Definition recv : val := - rec: "go" "c" "i" := - let: "len" := Fst "c" in - if: "i" < "len" then - let: "l" := Snd (! ((Snd "c") +â‚— "i")) in - let: "v" := Xchg "l" NONEV in - match: "v" with - NONE => "go" "c" "i" - | SOME "v" => "v" - end - (* OBS: Hacky *) - else (rec: "go" <> := "go" #())%V #(). + rec: "go" "c" "j" := + let: "n" := Snd (Fst "c") in guard "j" "n";; + let: "ls" := Fst (Fst "c") in + let: "i" := Snd "c" in + let: "l" := "ls" +â‚— vpos "n" "j" "i" in + let: "v" := Xchg "l" NONEV in + match: "v" with + NONE => "go" "c" "j" + | SOME "v" => "v" + end. (** * Setup of Iris's cameras *) Class proto_exclG Σ V := @@ -99,7 +99,7 @@ Notation iMsg Σ := (iMsg Σ val). Definition tok `{!chanG Σ} (γ : gname) : iProp Σ := own γ (Excl ()). Definition chan_inv `{!heapGS Σ, !chanG Σ} γ γE γt i j (l:loc) : iProp Σ := - (l ↦ NONEV ∗ tok γt) ∨ + (l ↦ NONEV ∗ tok γt)%I ∨ (∃ v m, l ↦ SOMEV v ∗ iProto_own γ i (<(Send, j)> m)%proto ∗ (∃ p, iMsg_car m v (Next p) ∗ own γE (â—E (Next p)))) ∨ @@ -108,15 +108,13 @@ Definition chan_inv `{!heapGS Σ, !chanG Σ} γ γE γt i j (l:loc) : iProp Σ : Definition iProto_pointsto_def `{!heapGS Σ, !chanG Σ} (c : val) (p : iProto Σ) : iProp Σ := - ∃ γ γE1 γE2 γt1 γt2 i (l:loc) ls p', - ⌜ c = PairV #(length ls) #l ⌠∗ + ∃ γ γE1 (l:loc) (i:nat) (n:nat) p', + ⌜ c = (#l,#n,#i)%V ⌠∗ inv (nroot.@"ctx") (iProto_ctx γ) ∗ - l ↦∗ ls ∗ - ([∗list] j ↦ v ∈ ls, - ∃ (l1 l2 : loc), - ⌜v = PairV #l1 #l2⌠∗ - inv (nroot.@"p") (chan_inv γ γE1 γt1 i j l1) ∗ - inv (nroot.@"p") (chan_inv γ γE2 γt2 j i l2)) ∗ + ([∗list] j ↦ _ ∈ replicate n (), + ∃ γE2 γt1 γt2, + inv (nroot.@"p") (chan_inv γ γE1 γt1 i j (l +â‚— (pos n i j))) ∗ + inv (nroot.@"p") (chan_inv γ γE2 γt2 j i (l +â‚— (pos n j i)))) ∗ â–· (p' ⊑ p) ∗ own γE1 (â—E (Next p')) ∗ own γE1 (â—¯E (Next p')) ∗ iProto_own γ i p'. @@ -131,10 +129,18 @@ Notation "c ↣ p" := (iProto_pointsto c p) Definition chan_pool `{!heapGS Σ, !chanG Σ} (cs : val) (ps : gmap nat (iProto Σ)) : iProp Σ := - ∃ (l:loc) (ls : list val), - ⌜cs = #l⌠∗ ⌜∀ i, is_Some (ps !! i) → is_Some (ls !! i)⌠∗ - l ↦∗ ls ∗ - [∗list] i ↦ c ∈ ls, (∀ p, ⌜ps !! i = Some p⌠-∗ c ↣ p). + ∃ γ (γEs : list gname) (l:loc) (n:nat), + ⌜cs = (#l,#n)%V⌠∗ ⌜∀ i, is_Some (ps !! i) → i < n⌠∗ + inv (nroot.@"ctx") (iProto_ctx γ) ∗ + [∗ list] i ↦ _ ∈ replicate n (), + (∀ p, ⌜ps !! i = Some p⌠-∗ + own (γEs !!! i) (â—E (Next p)) ∗ + own (γEs !!! i) (â—¯E (Next p)) ∗ + iProto_own γ i p) ∗ + [∗ list] j ↦ _ ∈ replicate n (), + ∃ γt1 γt2, + inv (nroot.@"p") (chan_inv γ (γEs !!! i) γt1 i j (l +â‚— (pos n i j))) ∗ + inv (nroot.@"p") (chan_inv γ (γEs !!! j) γt2 j i (l +â‚— (pos n j i))). Section channel. Context `{!heapGS Σ, !chanG Σ}. @@ -150,20 +156,185 @@ Section channel. Proof. rewrite iProto_pointsto_eq. iDestruct 1 as - (γ γE1 γE2 γt1 γt2 i l ls p ->) "(#IH & Hl & Hls & Hle & Hâ— & Hâ—¯ & Hown)". - iIntros "Hle'". iExists γ, γE1, γE2, γt1, γt2, i, l, ls, p. + (γ γE l n i p ->) "(#IH & Hls & Hle & Hâ— & Hâ—¯ & Hown)". + iIntros "Hle'". iExists γ, γE, l, n, i, p. iSplit; [done|]. iFrame "#∗". iApply (iProto_le_trans with "Hle Hle'"). Qed. + Lemma big_sepL_replicate {A B} n (x1 : A) (x2 : B) (P : nat → iProp Σ) : + ([∗ list] i ↦ _ ∈ replicate n x1, P i) ⊢ + ([∗ list] i ↦ _ ∈ replicate n x2, P i). + Proof. + iIntros "H". + iInduction n as [|n] "IHn". + { done. } + replace (S n) with (n + 1) by lia. + rewrite !replicate_add. + simpl. iDestruct "H" as "[H1 H2]". + iSplitL "H1". + { by iApply "IHn". } + simpl. rewrite !replicate_length. iFrame. + Qed. + + Lemma array_to_matrix_pre l n m v : + l ↦∗ replicate (n * m) v -∗ + [∗ list] i ↦ _ ∈ replicate n (), + (l +â‚— i*m) ↦∗ replicate m v. + Proof. + iIntros "Hl". + iInduction n as [|n] "IHn". + { done. } + replace (S n) with (n + 1) by lia. + replace ((n + 1) * m) with (n * m + m) by lia. + rewrite !replicate_add. simpl. + rewrite array_app. + iDestruct "Hl" as "[H1 H2]". + iDestruct ("IHn" with "H1") as "H1". + iFrame. + simpl. + rewrite Nat.add_0_r. + rewrite !replicate_length. + replace (Z.of_nat (n * m)) with (Z.of_nat n * Z.of_nat m)%Z by lia. + iFrame. + Qed. + + Lemma array_to_matrix l n v : + l ↦∗ replicate (n * n) v -∗ + [∗ list] i ↦ _ ∈ replicate n (), + [∗ list] j ↦ _ ∈ replicate n (), + (l +â‚— pos n i j) ↦ v. + Proof. + iIntros "H". + iDestruct (array_to_matrix_pre with "H") as "H". + iApply (big_sepL_impl with "H"). + iIntros "!>" (i ? HSome) "H". + clear HSome. + rewrite /array. + iApply big_sepL_replicate. + iApply (big_sepL_impl with "H"). + iIntros "!>" (j ? HSome) "Hl". + rewrite /pos. + rewrite Loc.add_assoc. + replace (Z.of_nat i * Z.of_nat n + Z.of_nat j)%Z with + (Z.of_nat (i * n + j))%Z by lia. + apply lookup_replicate in HSome as [-> _]. done. + Qed. + (** ** Specifications of [send] and [recv] *) Lemma new_chan_spec (n:nat) (ps:gmap nat (iProto Σ)) : - (∀ i, i < n → is_Some (ps !! i)) → - n = (size (dom ps)) → + 0 < n → + (∀ i, i < n ↔ is_Some (ps !! i)) → (* TODO: Weaken this! *) + (* n = (size (dom ps)) → *) {{{ iProto_consistent ps }}} new_chan #n {{{ cs, RET cs; chan_pool cs ps }}}. - Proof. Admitted. + Proof. + iIntros (Hle HSome Φ) "Hps HΦ". wp_lam. + wp_smart_apply wp_allocN; [lia|done|]. + iIntros (l) "[Hl _]". + iMod (iProto_init with "Hps") as (γ) "[Hps Hps']". + wp_pures. iApply "HΦ". + iAssert (|==> ∃ (γEs : list gname), + ⌜length γEs = n⌠∗ + [∗ list] i ↦ _ ∈ replicate n (), + own (γEs !!! i) (â—E (Next (ps !!! i))) ∗ + own (γEs !!! i) (â—¯E (Next (ps !!! i))) ∗ + iProto_own γ i (ps !!! i))%I with "[Hps']" as "H". + { clear Hle. + (* iInduction n as [|n] "IHn" forall (ps HSome Heq). *) + iInduction n as [|n] "IHn" forall (ps HSome). + { iExists []. iModIntro. simpl. done. } + assert (n < S n) by lia. + apply HSome in H. destruct H as [p ?]. + iDestruct (big_sepM_delete _ _ n with "Hps'") as "[Hp Hps']". + { done. } + iMod (own_alloc (â—E (Next p) â‹… â—¯E (Next p))) as (γE) "[Hauth Hfrag]". + { apply excl_auth_valid. } + iMod ("IHn" with "[] Hps'") as (γEs Hlen) "H". + { iPureIntro. + intros i. + split. + - intros Hle. + rewrite lookup_delete_ne; [|lia]. + apply HSome. lia. + - intros HSome'. + destruct (decide (i=n)). + + simplify_eq. rewrite lookup_delete in HSome'. by inversion HSome'. + + rewrite lookup_delete_ne in HSome'; [|lia]. apply HSome in HSome'. + lia. + } + iModIntro. iExists (γEs++[γE]). + replace (S n) with (n + 1) by lia. + rewrite replicate_add. + rewrite big_sepL_app. + rewrite app_length. + rewrite Hlen. iSplit; [done|]. + simpl. + iSplitL "H". + { iApply (big_sepL_impl with "H"). + iIntros "!>" (i ? HSome') "(Hauth & Hfrag & Hown)". + assert (i < n). + { by apply lookup_replicate_1 in HSome' as [? ?]. } + assert (delete n ps !!! i = ps !!! i) as Heq'. + { apply lookup_total_delete_ne. lia. } + rewrite Heq'. iFrame. + rewrite lookup_total_app_l; [|lia]. iFrame. } + simpl. rewrite replicate_length. rewrite Nat.add_0_r. + rewrite list_lookup_total_middle; [|done]. + rewrite lookup_total_alt. rewrite H. simpl. iFrame. } + iMod "H" as (γEs Hlen) "H". + iAssert (|={⊤}=> + [∗ list] i ↦ _ ∈ replicate n (), + [∗ list] j ↦ _ ∈ replicate n (), + ∃ γt, + inv (nroot.@"p") (chan_inv γ (γEs !!! i) γt i j + (l +â‚— (pos n i j))))%I with "[Hl]" as "IH". + { replace (Z.to_nat (Z.of_nat n * Z.of_nat n)) with (n*n) by lia. + iDestruct (array_to_matrix with "Hl") as "Hl". + iApply big_sepL_fupd. + iApply (big_sepL_impl with "Hl"). + iIntros "!>" (i ? HSome') "H1". + iApply big_sepL_fupd. + iApply (big_sepL_impl with "H1"). + iIntros "!>" (j ? HSome'') "H1". + iMod (own_alloc (Excl ())) as (γ') "Hγ'". + { done. } + iExists γ'. + iApply inv_alloc. + iLeft. iFrame. } + iMod "IH" as "#IH". + iMod (inv_alloc with "Hps") as "#IHp". + iExists _,_,_,_. + iModIntro. iSplit; [done|]. + iSplit. + { iPureIntro. intros. apply HSome in H. done. } + iFrame "IHp". + iApply (big_sepL_impl with "H"). + iIntros "!>" (i ? HSome') "(Hauth & Hfrag & Hown)". + iSplitL. + { iIntros (p HSome''). + rewrite lookup_total_alt. rewrite HSome''. + iFrame. } + iApply big_sepL_intro. + iIntros "!>" (j ? HSome''). + assert (i < n) as Hle'. + { apply lookup_replicate in HSome' as [_ Hle']. done. } + assert (j < n) as Hle''. + { apply lookup_replicate in HSome'' as [_ Hle'']. done. } + iDestruct (big_sepL_lookup _ _ i () with "IH") as "IH''". + { rewrite lookup_replicate. done. } + iDestruct (big_sepL_lookup _ _ j () with "IH''") as "IH'''". + { rewrite lookup_replicate. done. } + iFrame "#". + iDestruct (big_sepL_lookup _ _ j () with "IH") as "H". + { rewrite lookup_replicate. done. } + iDestruct (big_sepL_lookup _ _ i () with "H") as "H'". + { rewrite lookup_replicate. done. } + iDestruct "IH'''" as (γ1) "?". + iDestruct "H'" as (γ2) "?". + iExists _, _. iFrame "#". + Qed. Lemma get_chan_spec cs (i:nat) ps p : ps !! i = Some p → @@ -172,30 +343,37 @@ Section channel. {{{ c, RET c; c ↣ p ∗ chan_pool cs (delete i ps) }}}. Proof. iIntros (HSome Φ) "Hcs HΦ". - iDestruct "Hcs" as (l ls -> Hlen) "[Hl Hls]". - wp_lam. - assert (is_Some (ls !! i)) as [c HSome']. - { by apply Hlen. } - wp_smart_apply (wp_load_offset with "Hl"); [done|]. - iIntros "Hcs". + iDestruct "Hcs" as (γp γEs l n -> Hle) "[#IHp Hl]". + wp_lam. wp_pures. + assert (i < n). + { apply Hle. eexists _. done. } + iDestruct (big_sepL_delete _ _ i () with "Hl") as "[[Hi #IHs] H]". + { by apply lookup_replicate. } + iDestruct ("Hi" with "[//]") as "(Hauth & Hown & Hp)". + iModIntro. iApply "HΦ". - iDestruct (big_sepL_delete' _ _ i with "Hls") as "[Hc Hls]"; [set_solver|]. - iDestruct ("Hc" with "[//]") as "Hc". - iFrame. - iExists _, _. iSplit; [done|]. iFrame "Hcs". - iSplitR. - { iPureIntro. intros j HSome''. - destruct (decide (i=j)) as [<-|Hneq]. - { rewrite lookup_delete in HSome''. done. } - rewrite lookup_delete_ne in HSome''; [|done]. - by apply Hlen. } - iApply (big_sepL_impl with "Hls"). - iIntros "!>" (j v Hin) "H". - iIntros (p' HSome''). - destruct (decide (i=j)) as [<-|Hneq]. - { rewrite lookup_delete in HSome''. done. } - rewrite lookup_delete_ne in HSome''; [|done]. - by iApply "H". + iSplitR "H". + { rewrite iProto_pointsto_eq. + iExists _, _, _, _, _, _. + iSplit; [done|]. + iFrame "#∗". iSplit; [|iNext; done]. + iApply (big_sepL_impl with "IHs"). + iIntros "!>" (???). iDestruct 1 as (γt1 γt2) "[H1 H2]". + iExists _,_,_. iFrame. } + iExists _, _, _, _. + iSplit; [done|]. + iSplit. + { iPureIntro. intros i' HSome'. apply Hle. + assert (i ≠i'). + { intros ->. rewrite lookup_delete in HSome'. by inversion HSome'. } + rewrite lookup_delete_ne in HSome'; done. } + iFrame "#∗". + iApply (big_sepL_impl with "H"). + iIntros "!>" (i' ? HSome''). + case_decide. + { simplify_eq. iFrame "#". + iIntros "_" (p' Hin). simplify_eq. by rewrite lookup_delete in Hin. } + rewrite lookup_delete_ne; [|done]. eauto. Qed. Lemma own_prot_excl γ i (p1 p2 : iProto Σ) : @@ -206,6 +384,33 @@ Section channel. iIntros "Hi Hj". by iDestruct (own_prot_idx with "Hi Hj") as %Hneq. Qed. + Lemma diverge_spec P : + {{{ True }}} diverge #() {{{ RET #(); P }}}. + Proof. + iIntros (Φ) "_ HΦ". + wp_lam. iLöb as "IH". + wp_pures. by iApply "IH". + Qed. + + Lemma guard_spec (i n : nat) : + {{{ True }}} guard #i #n {{{ RET #(); ⌜i < n⌠}}}. + Proof. + iIntros (Φ) "_ HΦ". + wp_lam. wp_pures. + case_bool_decide. + - wp_pures. iApply "HΦ". iPureIntro. lia. + - by wp_smart_apply diverge_spec. + Qed. + + Lemma vpos_spec (n i j : nat) : + {{{ True }}} vpos #n #i #j {{{ RET #(pos n i j); True }}}. + Proof. + iIntros (Φ) "_ HΦ". wp_lam. wp_pures. rewrite /pos. + replace (Z.of_nat i * Z.of_nat n + Z.of_nat j)%Z with + (Z.of_nat (i * n + j)) by lia. + by iApply "HΦ". + Qed. + Lemma send_spec c j v p : {{{ c ↣ <(Send, j)> MSG v; p }}} send c #j v @@ -213,27 +418,21 @@ Section channel. Proof. rewrite iProto_pointsto_eq. iIntros (Φ) "Hc HΦ". wp_lam; wp_pures. iDestruct "Hc" as - (γ γE1 γE2 γt1 γt2 i l ls p' ->) "(#IH & Hl & Hls & Hle & Hâ— & Hâ—¯ & Hown)". + (γ γE l i n p' ->) "(#IH & #Hls & Hle & Hâ— & Hâ—¯ & Hown)". wp_pures. - case_bool_decide; last first. - { wp_pures. iClear "IH Hâ— Hâ—¯ Hown HΦ Hls Hl". - iLöb as "IH". wp_lam. iApply "IH". done. } - assert (is_Some (ls !! j)) as [l' HSome]. - { apply lookup_lt_is_Some_2. lia. } + wp_smart_apply guard_spec; [done|]. + iDestruct 1 as %Hle. wp_pures. - wp_smart_apply (wp_load_offset with "Hl"). - { done. } - iIntros "Hl". wp_pures. - iDestruct (big_sepL_lookup_acc with "Hls") as "[Hj Hls]"; [done|]. - iDestruct "Hj" as (l1 l2 ->) "#[IHl1 IHl2]". - iDestruct ("Hls" with "[]") as "Hls". - { iExists _, _. iFrame "#". done. } + wp_smart_apply (vpos_spec); [done|]. iIntros "_". + iDestruct (big_sepL_lookup_acc with "Hls") as "[Hj _]". + { by apply lookup_replicate_2. } + iDestruct "Hj" as (γE' γt γt') "#[IHl1 IHl2]". wp_pures. wp_bind (Store _ _). iInv "IHl1" as "HIp". iDestruct "HIp" as "[HIp|HIp]"; last first. { iDestruct "HIp" as "[HIp|HIp]". - - iDestruct "HIp" as (? m) "(>Hl' & Hown' & HIp)". + - iDestruct "HIp" as (? m) "(>Hl & Hown' & HIp)". wp_store. rewrite /iProto_own. iDestruct "Hown" as (p'') "[Hle' Hown]". @@ -268,7 +467,7 @@ Section channel. wp_load. iModIntro. iSplitL "Hl' Hown HIp". { iRight. iLeft. iExists _, _. iFrame. } - wp_pures. iApply ("HL" with "HΦ Hl Hls Htok Hâ—¯"). + wp_pures. iApply ("HL" with "HΦ Htok Hâ—¯"). - iDestruct "HIp" as (p'') "(>Hl' & Hown & Hâ—)". wp_load. iModIntro. @@ -281,11 +480,11 @@ Section channel. { apply excl_auth_update. } iModIntro. iApply "HΦ". - iExists _,_, _, _, _, _, _, _, _. + iExists _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". iRewrite -"Hagree'". iApply iProto_le_refl. Qed. - + Lemma send_spec_tele {TT} c i (tt : TT) (v : TT → val) (P : TT → iProp Σ) (p : TT → iProto Σ) : {{{ c ↣ (<(Send,i) @.. x > MSG v x {{ P x }}; p x) ∗ P tt }}} @@ -302,7 +501,6 @@ Section channel. by iApply (send_spec with "Hc"). Qed. - Lemma recv_spec {TT} c j (v : TT → val) (P : TT → iProp Σ) (p : TT → iProto Σ) : {{{ c ↣ <(Recv, j) @.. x> MSG v x {{ â–· P x }}; p x }}} recv c #j @@ -311,21 +509,16 @@ Section channel. iIntros (Φ) "Hc HΦ". iLöb as "HL". wp_lam. rewrite iProto_pointsto_eq. iDestruct "Hc" as - (γ γE1 γE2 γt1 γt2 i l ls p' ->) "(#IH & Hl & Hls & Hle & Hâ— & Hâ—¯ & Hown)". + (γ γE l i n p' ->) "(#IH & #Hls & Hle & Hâ— & Hâ—¯ & Hown)". + wp_pures. wp_pures. - case_bool_decide; last first. - { wp_pures. iClear "IH Hâ— Hâ—¯ Hown HΦ Hls Hl". - iLöb as "IH". wp_lam. iApply "IH". done. } + wp_smart_apply guard_spec; [done|]. + iDestruct 1 as %Hle. wp_pures. - assert (is_Some (ls !! j)) as [l' HSome]. - { apply lookup_lt_is_Some_2. lia. } - wp_smart_apply (wp_load_offset with "Hl"). - { done. } - iIntros "Hl". wp_pures. - iDestruct (big_sepL_lookup_acc with "Hls") as "[Hj Hls]"; [done|]. - iDestruct "Hj" as (l1 l2 ->) "#[IHl1 IHl2]". - iDestruct ("Hls" with "[]") as "Hls". - { iExists _, _. iFrame "#". done. } + wp_smart_apply (vpos_spec); [done|]. iIntros "_". + iDestruct (big_sepL_lookup_acc with "Hls") as "[Hj _]". + { by apply lookup_replicate_2. } + iDestruct "Hj" as (γE' γt γt') "#[IHl1 IHl2]". wp_pures. wp_bind (Xchg _ _). iInv "IHl2" as "HIp". @@ -335,16 +528,16 @@ Section channel. iModIntro. iSplitL "Hl' Htok". { iLeft. iFrame. } - wp_pures. iApply ("HL" with "[Hâ— Hâ—¯ Hown Hls Hl Hle] HΦ"). - iExists _, _, _, _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". } + wp_pures. iApply ("HL" with "[Hâ— Hâ—¯ Hown Hle] HΦ"). + iExists _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". } iDestruct "HIp" as "[HIp|HIp]"; last first. { iDestruct "HIp" as (p'') "[>Hl' [Hown' Hâ—¯']]". wp_xchg. iModIntro. iSplitL "Hl' Hown' Hâ—¯'". { iRight. iRight. iExists _. iFrame. } - wp_pures. iApply ("HL" with "[Hâ— Hâ—¯ Hown Hls Hl Hle] HΦ"). - iExists _, _, _, _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". } + wp_pures. iApply ("HL" with "[Hâ— Hâ—¯ Hown Hle] HΦ"). + iExists _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". } iDestruct "HIp" as (w m) "(>Hl' & Hown' & HIp)". iDestruct "HIp" as (p'') "[Hm Hp']". iInv "IH" as "Hctx". @@ -365,7 +558,7 @@ Section channel. { apply excl_auth_update. } iModIntro. iApply "HΦ". iFrame. - iExists _, _, _, _, _, _, _, _, _. iSplit; [done|]. + iExists _, _, _, _, _, _. iSplit; [done|]. iRewrite "Hp". iFrame "#∗". iApply iProto_le_refl. Qed. -- GitLab From 0677da80f695c4e53ccfd7f1a5e7a58d81be6c6b Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Thu, 1 Feb 2024 11:41:57 +0100 Subject: [PATCH 31/81] Improved new_chan spec and closed example proofs --- theories/channel/multi_channel.v | 88 ++++++++++-------- .../multi_proto_consistency_examples.v | 93 ++++++++++--------- 2 files changed, 98 insertions(+), 83 deletions(-) diff --git a/theories/channel/multi_channel.v b/theories/channel/multi_channel.v index 590820f..c5cc715 100644 --- a/theories/channel/multi_channel.v +++ b/theories/channel/multi_channel.v @@ -111,7 +111,7 @@ Definition iProto_pointsto_def `{!heapGS Σ, !chanG Σ} ∃ γ γE1 (l:loc) (i:nat) (n:nat) p', ⌜ c = (#l,#n,#i)%V ⌠∗ inv (nroot.@"ctx") (iProto_ctx γ) ∗ - ([∗list] j ↦ _ ∈ replicate n (), + ([∗list] j ↦ _ ∈ replicate n (), ∃ γE2 γt1 γt2, inv (nroot.@"p") (chan_inv γ γE1 γt1 i j (l +â‚— (pos n i j))) ∗ inv (nroot.@"p") (chan_inv γ γE2 γt2 j i (l +â‚— (pos n j i)))) ∗ @@ -136,7 +136,7 @@ Definition chan_pool `{!heapGS Σ, !chanG Σ} (∀ p, ⌜ps !! i = Some p⌠-∗ own (γEs !!! i) (â—E (Next p)) ∗ own (γEs !!! i) (â—¯E (Next p)) ∗ - iProto_own γ i p) ∗ + iProto_own γ i p) ∗ [∗ list] j ↦ _ ∈ replicate n (), ∃ γt1 γt2, inv (nroot.@"p") (chan_inv γ (γEs !!! i) γt1 i j (l +â‚— (pos n i j))) ∗ @@ -170,13 +170,13 @@ Section channel. iInduction n as [|n] "IHn". { done. } replace (S n) with (n + 1) by lia. - rewrite !replicate_add. + rewrite !replicate_add. simpl. iDestruct "H" as "[H1 H2]". iSplitL "H1". { by iApply "IHn". } simpl. rewrite !replicate_length. iFrame. Qed. - + Lemma array_to_matrix_pre l n m v : l ↦∗ replicate (n * m) v -∗ [∗ list] i ↦ _ ∈ replicate n (), @@ -192,7 +192,7 @@ Section channel. iDestruct "Hl" as "[H1 H2]". iDestruct ("IHn" with "H1") as "H1". iFrame. - simpl. + simpl. rewrite Nat.add_0_r. rewrite !replicate_length. replace (Z.of_nat (n * m)) with (Z.of_nat n * Z.of_nat m)%Z by lia. @@ -215,7 +215,7 @@ Section channel. iApply (big_sepL_impl with "H"). iIntros "!>" (j ? HSome) "Hl". rewrite /pos. - rewrite Loc.add_assoc. + rewrite Loc.add_assoc. replace (Z.of_nat i * Z.of_nat n + Z.of_nat j)%Z with (Z.of_nat (i * n + j))%Z by lia. apply lookup_replicate in HSome as [-> _]. done. @@ -224,46 +224,48 @@ Section channel. (** ** Specifications of [send] and [recv] *) Lemma new_chan_spec (n:nat) (ps:gmap nat (iProto Σ)) : 0 < n → - (∀ i, i < n ↔ is_Some (ps !! i)) → (* TODO: Weaken this! *) - (* n = (size (dom ps)) → *) + dom ps = list_to_set $ seq 0 n → (* TODO: Consider using [set_eq] instead *) {{{ iProto_consistent ps }}} new_chan #n {{{ cs, RET cs; chan_pool cs ps }}}. Proof. - iIntros (Hle HSome Φ) "Hps HΦ". wp_lam. + iIntros (Hle Hdom Φ) "Hps HΦ". wp_lam. wp_smart_apply wp_allocN; [lia|done|]. iIntros (l) "[Hl _]". iMod (iProto_init with "Hps") as (γ) "[Hps Hps']". wp_pures. iApply "HΦ". iAssert (|==> ∃ (γEs : list gname), - ⌜length γEs = n⌠∗ + ⌜length γEs = n⌠∗ [∗ list] i ↦ _ ∈ replicate n (), own (γEs !!! i) (â—E (Next (ps !!! i))) ∗ own (γEs !!! i) (â—¯E (Next (ps !!! i))) ∗ iProto_own γ i (ps !!! i))%I with "[Hps']" as "H". { clear Hle. (* iInduction n as [|n] "IHn" forall (ps HSome Heq). *) - iInduction n as [|n] "IHn" forall (ps HSome). + iInduction n as [|n] "IHn" forall (ps Hdom). { iExists []. iModIntro. simpl. done. } - assert (n < S n) by lia. - apply HSome in H. destruct H as [p ?]. + (* assert (n < S n) by lia. *) + assert (is_Some (ps !! n)) as [p ?]. + { apply elem_of_dom. + rewrite Hdom. + rewrite list_to_set_seq. + apply elem_of_set_seq. + lia. } iDestruct (big_sepM_delete _ _ n with "Hps'") as "[Hp Hps']". { done. } iMod (own_alloc (â—E (Next p) â‹… â—¯E (Next p))) as (γE) "[Hauth Hfrag]". { apply excl_auth_valid. } iMod ("IHn" with "[] Hps'") as (γEs Hlen) "H". { iPureIntro. - intros i. - split. - - intros Hle. - rewrite lookup_delete_ne; [|lia]. - apply HSome. lia. - - intros HSome'. - destruct (decide (i=n)). - + simplify_eq. rewrite lookup_delete in HSome'. by inversion HSome'. - + rewrite lookup_delete_ne in HSome'; [|lia]. apply HSome in HSome'. - lia. - } + rewrite dom_delete_L. rewrite Hdom. + replace (S n) with (n + 1) by lia. + rewrite seq_app. + simpl. rewrite list_to_set_app_L. simpl. + rewrite right_id_L. + rewrite difference_union_distr_l_L. + rewrite difference_diag_L. rewrite right_id_L. + assert (n ∉ seq 0 n); [|set_solver]. + intros Hin%elem_of_seq. lia. } iModIntro. iExists (γEs++[γE]). replace (S n) with (n + 1) by lia. rewrite replicate_add. @@ -277,7 +279,7 @@ Section channel. assert (i < n). { by apply lookup_replicate_1 in HSome' as [? ?]. } assert (delete n ps !!! i = ps !!! i) as Heq'. - { apply lookup_total_delete_ne. lia. } + { apply lookup_total_delete_ne. lia. } rewrite Heq'. iFrame. rewrite lookup_total_app_l; [|lia]. iFrame. } simpl. rewrite replicate_length. rewrite Nat.add_0_r. @@ -308,7 +310,11 @@ Section channel. iExists _,_,_,_. iModIntro. iSplit; [done|]. iSplit. - { iPureIntro. intros. apply HSome in H. done. } + { iPureIntro. intros. + apply elem_of_dom in H. + rewrite Hdom in H. + rewrite list_to_set_seq in H. + apply elem_of_set_seq in H. lia. } iFrame "IHp". iApply (big_sepL_impl with "H"). iIntros "!>" (i ? HSome') "(Hauth & Hfrag & Hown)". @@ -349,7 +355,7 @@ Section channel. { apply Hle. eexists _. done. } iDestruct (big_sepL_delete _ _ i () with "Hl") as "[[Hi #IHs] H]". { by apply lookup_replicate. } - iDestruct ("Hi" with "[//]") as "(Hauth & Hown & Hp)". + iDestruct ("Hi" with "[//]") as "(Hauth & Hown & Hp)". iModIntro. iApply "HΦ". iSplitR "H". @@ -365,7 +371,7 @@ Section channel. iSplit. { iPureIntro. intros i' HSome'. apply Hle. assert (i ≠i'). - { intros ->. rewrite lookup_delete in HSome'. by inversion HSome'. } + { intros ->. rewrite lookup_delete in HSome'. by inversion HSome'. } rewrite lookup_delete_ne in HSome'; done. } iFrame "#∗". iApply (big_sepL_impl with "H"). @@ -380,29 +386,29 @@ Section channel. own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p1))) -∗ own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p2))) -∗ False. - Proof. + Proof. iIntros "Hi Hj". by iDestruct (own_prot_idx with "Hi Hj") as %Hneq. Qed. - Lemma diverge_spec P : + Lemma diverge_spec P : {{{ True }}} diverge #() {{{ RET #(); P }}}. Proof. - iIntros (Φ) "_ HΦ". + iIntros (Φ) "_ HΦ". wp_lam. iLöb as "IH". wp_pures. by iApply "IH". Qed. - Lemma guard_spec (i n : nat) : + Lemma guard_spec (i n : nat) : {{{ True }}} guard #i #n {{{ RET #(); ⌜i < n⌠}}}. Proof. - iIntros (Φ) "_ HΦ". - wp_lam. wp_pures. + iIntros (Φ) "_ HΦ". + wp_lam. wp_pures. case_bool_decide. - - wp_pures. iApply "HΦ". iPureIntro. lia. - - by wp_smart_apply diverge_spec. + - wp_pures. iApply "HΦ". iPureIntro. lia. + - by wp_smart_apply diverge_spec. Qed. - Lemma vpos_spec (n i j : nat) : + Lemma vpos_spec (n i j : nat) : {{{ True }}} vpos #n #i #j {{{ RET #(pos n i j); True }}}. Proof. iIntros (Φ) "_ HΦ". wp_lam. wp_pures. rewrite /pos. @@ -426,13 +432,13 @@ Section channel. wp_smart_apply (vpos_spec); [done|]. iIntros "_". iDestruct (big_sepL_lookup_acc with "Hls") as "[Hj _]". { by apply lookup_replicate_2. } - iDestruct "Hj" as (γE' γt γt') "#[IHl1 IHl2]". + iDestruct "Hj" as (γE' γt γt') "#[IHl1 IHl2]". wp_pures. wp_bind (Store _ _). iInv "IHl1" as "HIp". iDestruct "HIp" as "[HIp|HIp]"; last first. { iDestruct "HIp" as "[HIp|HIp]". - - iDestruct "HIp" as (? m) "(>Hl & Hown' & HIp)". + - iDestruct "HIp" as (? m) "(>Hl & Hown' & HIp)". wp_store. rewrite /iProto_own. iDestruct "Hown" as (p'') "[Hle' Hown]". @@ -449,7 +455,7 @@ Section channel. iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]". { apply excl_auth_update. } iModIntro. - iSplitL "Hl' Hâ— Hown Hle". + iSplitL "Hl' Hâ— Hown Hle". { iRight. iLeft. iIntros "!>". iExists _, _. iFrame. iSplitL "Hown Hle". { iApply (iProto_own_le with "Hown Hle"). } @@ -518,7 +524,7 @@ Section channel. wp_smart_apply (vpos_spec); [done|]. iIntros "_". iDestruct (big_sepL_lookup_acc with "Hls") as "[Hj _]". { by apply lookup_replicate_2. } - iDestruct "Hj" as (γE' γt γt') "#[IHl1 IHl2]". + iDestruct "Hj" as (γE' γt γt') "#[IHl1 IHl2]". wp_pures. wp_bind (Xchg _ _). iInv "IHl2" as "HIp". diff --git a/theories/channel/multi_proto_consistency_examples.v b/theories/channel/multi_proto_consistency_examples.v index ec48cd0..78e48a6 100644 --- a/theories/channel/multi_proto_consistency_examples.v +++ b/theories/channel/multi_proto_consistency_examples.v @@ -76,10 +76,10 @@ Tactic Notation "iProto_consistent_take_step" := as "[%Heq2 Hm2']";simplify_eq; try (iClear "Hm1' Hm2'"; iExists _,_,_,_,_,_,_,_,_,_; - iSplitL; [iFrame "#"|]; - iSplitL; [iFrame "#"|]; - iSplitL; [iPureIntro; tc_solve|]; - iSplitL; [iPureIntro; tc_solve|]; + iSplitL "Hm1"; [iFrame "#"|]; + iSplitL "Hm2"; [iFrame "#"|]; + iSplit; [iPureIntro; tc_solve|]; + iSplit; [iPureIntro; tc_solve|]; simpl; iClear "Hm1 Hm2"; clear m1 m2); try (repeat (rewrite (insert_commute _ _ i); [|done]); rewrite insert_insert; @@ -96,13 +96,13 @@ Lemma iProto_consistent_empty {Σ} : ⊢ iProto_consistent (@iProto_empty Σ). Proof. iProto_consistent_take_step. Qed. -Definition iProto_binary `{!invGS Σ} (P : iProp Σ) : gmap nat (iProto Σ) := - <[0 := (<(Send, 1) @ (x:Z)> MSG #x {{ P }} ; END)%proto ]> - (<[1 := (<(Recv, 0) @ (x:Z)> MSG #x {{ P }} ; END)%proto ]> +Definition iProto_binary `{!invGS Σ} : gmap nat (iProto Σ) := + <[0 := (<(Send, 1) @ (x:Z)> MSG #x ; END)%proto ]> + (<[1 := (<(Recv, 0) @ (x:Z)> MSG #x ; END)%proto ]> ∅). -Lemma iProto_binary_consistent `{!invGS Σ} (P : iProp Σ) : - ⊢ iProto_consistent (@iProto_binary Σ invGS0 P). +Lemma iProto_binary_consistent `{!invGS Σ} : + ⊢ iProto_consistent (@iProto_binary Σ invGS0). Proof. rewrite /iProto_binary. iProto_consistent_take_step. @@ -113,8 +113,7 @@ Qed. Definition iProto_roundtrip `{!invGS Σ} : gmap nat (iProto Σ) := <[0 := (<(Send, 1) @ (x:Z)> MSG #x ; <(Recv, 2)> MSG #x; END)%proto ]> (<[1 := (<(Recv, 0) @ (x:Z)> MSG #x ; <(Send, 2)> MSG #x; END)%proto ]> - (<[2 := (<(Recv, 1) @ (x:Z)> MSG #x ; <(Send, 0)> MSG #x; END)%proto ]> - ∅)). + (<[2 := (<(Recv, 1) @ (x:Z)> MSG #x ; <(Send, 0)> MSG #x; END)%proto ]> ∅)). Lemma iProto_roundtrip_consistent `{!invGS Σ} : ⊢ iProto_consistent (@iProto_roundtrip Σ invGS0). @@ -150,7 +149,7 @@ Section channel. Proof using chanG0 heapGS0 Σ. iIntros (Φ) "_ HΦ". wp_lam. wp_smart_apply (new_chan_spec 3 iProto_roundtrip). - { intros i Hle. destruct i as [|[|[]]]; try set_solver. lia. } + { lia. } { set_solver. } { iApply iProto_roundtrip_consistent. } iIntros (cs) "Hcs". @@ -219,7 +218,7 @@ Section proof. Proof using chanG0. iIntros (Φ) "_ HΦ". wp_lam. wp_smart_apply (new_chan_spec 3 iProto_roundtrip_ref with "[]"). - { intros i Hle. destruct i as [|[|[]]]; try set_solver. lia. } + { lia. } { set_solver. } { iApply iProto_roundtrip_ref_consistent. } iIntros (cs) "Hcs". @@ -361,7 +360,7 @@ Section proof. Proof using chanG0. iIntros (Φ) "_ HΦ". wp_lam. wp_smart_apply (new_chan_spec 3 iProto_roundtrip_ref_rec with "[]"). - { intros i Hle. destruct i as [|[|[]]]; try set_solver. lia. } + { lia. } { set_solver. } { iApply iProto_roundtrip_ref_rec_consistent. } iIntros (cs) "Hcs". @@ -389,7 +388,16 @@ End proof. Section parallel. Context `{!heapGS Σ}. - (** + (** + + 0 -> 1 : (x1:Z) < x1 > . + 0 -> 2 : (x2:Z) < x2 > . + 2 -> 3 : (y1:Z) < x1+y1 > ; + 3 -> 4 : (y2:Z) < x2+y2 > ; + 3 -> 0 : < x1+y1 > ; + 4 -> 0 : < x2+y2 > ; + end + 0 / \ 1 2 @@ -400,83 +408,84 @@ Section parallel. *) Definition iProto_parallel : gmap nat (iProto Σ) := - <[0 := (<(Send, 1) @ (x:Z)> MSG #x ; <(Send, 2)> MSG #x ; - <(Recv, 3)> MSG #x; <(Recv, 4)> MSG #x; END)%proto]> + <[0 := (<(Send, 1) @ (x1:Z)> MSG #x1 ; + <(Send, 2) @ (x2:Z)> MSG #x2 ; + <(Recv, 3) @ (y1:Z)> MSG #(x1+y1); + <(Recv, 4) @ (y2:Z)> MSG #(x2+y2); END)%proto]> (<[1 := (<(Recv, 0) @ (x:Z)> MSG #x ; - <(Send, 3)> MSG #x; END)%proto]> + <(Send, 3) @ (y:Z)> MSG #(x+y); END)%proto]> (<[2 := (<(Recv, 0) @ (x:Z)> MSG #x ; - <(Send, 4)> MSG #x ; END)%proto]> + <(Send, 4) @ (y:Z)> MSG #(x+y) ; END)%proto]> (<[3 := (<(Recv, 1) @ (x:Z)> MSG #x ; <(Send, 0)> MSG #x; END)%proto]> (<[4 := (<(Recv, 2) @ (x:Z)> MSG #x ; - <(Send, 0)> MSG #x ; END)%proto]> - ∅)))). + <(Send, 0)> MSG #x ; END)%proto]> ∅)))). Lemma iProto_parallel_consistent : ⊢ iProto_consistent iProto_parallel. Proof. rewrite /iProto_parallel. iProto_consistent_take_step. - iIntros (x) "_". iExists _. iSplit; [done|]. iSplit; [done|]. + iIntros (x1) "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 0. clean_map 1. iProto_consistent_take_step. - - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. + - iIntros (x2) "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 0. clean_map 2. iProto_consistent_take_step. - + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. + + iIntros (y1) "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 1. clean_map 3. iProto_consistent_take_step. - * iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. + * iIntros (y2) "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 2. clean_map 4. iProto_consistent_take_step. - iIntros "_". iSplit; [done|]. iSplit; [done|]. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 0. clean_map 3. iProto_consistent_take_step. - iIntros "_". iSplit; [done|]. iSplit; [done|]. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. iProto_consistent_take_step. - * iIntros "_". iSplit; [done|]. iSplit; [done|]. + * iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 3. clean_map 0. iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. + iIntros (y2) "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 2. clean_map 4. iProto_consistent_take_step. - iIntros "_". iSplit; [done|]. iSplit; [done|]. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 4. clean_map 0. iProto_consistent_take_step. - + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. + + iIntros (y1) "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 2. clean_map 4. iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. + iIntros (y2) "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 1. clean_map 3. iProto_consistent_take_step. - iIntros "_". iSplit; [done|]. iSplit; [done|]. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 3. clean_map 0. iProto_consistent_take_step. - iIntros "_". iSplit; [done|]. iSplit; [done|]. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 4. clean_map 0. iProto_consistent_take_step. - - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. + - iIntros (y1) "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 1. clean_map 3. iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. + iIntros (x2) "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 0. clean_map 2. iProto_consistent_take_step. - + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. + + iIntros (y2) "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 2. clean_map 4. iProto_consistent_take_step. - iIntros "_". iSplit; [done|]. iSplit; [done|]. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 3. clean_map 0. iProto_consistent_take_step. - iIntros "_". iSplit; [done|]. iSplit; [done|]. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 4. clean_map 0. iProto_consistent_take_step. - + iIntros "_". iSplit; [done|]. iSplit; [done|]. + + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 3. clean_map 0. iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. + iIntros (z) "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 2. clean_map 4. iProto_consistent_take_step. - iIntros "_". iSplit; [done|]. iSplit; [done|]. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. clean_map 4. clean_map 0. iProto_consistent_take_step. Qed. -- GitLab From 9c584bab0ff4423f5ea9f156faecbfe16e53a00c Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Thu, 1 Feb 2024 11:43:01 +0100 Subject: [PATCH 32/81] Added forwarder example --- .../multi_proto_consistency_examples.v | 38 +++++++++++++++++++ 1 file changed, 38 insertions(+) diff --git a/theories/channel/multi_proto_consistency_examples.v b/theories/channel/multi_proto_consistency_examples.v index 78e48a6..826b9b4 100644 --- a/theories/channel/multi_proto_consistency_examples.v +++ b/theories/channel/multi_proto_consistency_examples.v @@ -628,3 +628,41 @@ Section two_buyer_ref. Qed. End two_buyer_ref. + + +Section fwd. + Context `{!heapGS Σ}. + + Definition iProto_fwd : gmap nat (iProto Σ) := + <[0 := (<(Send, 1) @ (x:Z)> MSG #x ; + <(Send, 1) @ (b:bool)> MSG #b ; + <(Send, if b then 2 else 3) > MSG #x ; END)%proto]> + (<[1 := (<(Recv, 0) @ (x:Z)> MSG #x ; + <(Recv, 0) @ (b:bool)> MSG #b; + if b + then <(Send,2)> MSG #x ; END + else <(Send,3)> MSG #x ; END)%proto]> + (<[2 := (<(Recv, 1) @ (x:Z)> MSG #x ; + <(Send, 0)> MSG #x ; END)%proto]> + (<[3 := (<(Recv, 1) @ (x:Z)> MSG #x ; + <(Send, 0)> MSG #x ; END)%proto]> ∅))). + + Lemma iProto_fwd_consistent : + ⊢ iProto_consistent iProto_fwd. + Proof. + rewrite /iProto_fwd. + iProto_consistent_take_step. + iIntros (x) "_". iExists _. iSplit; [done|]. iSplit; [done|]. + iProto_consistent_take_step. + iIntros ([]) "_". + - iExists _. iSplit; [done|]. iSplit; [done|]. + iProto_consistent_take_step. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. + iProto_consistent_take_step. + - iExists _. iSplit; [done|]. iSplit; [done|]. + iProto_consistent_take_step. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. + iProto_consistent_take_step. + Qed. + +End fwd. -- GitLab From 8db776a5c34fe1107a9fba5a63d1449266089902 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Thu, 1 Feb 2024 11:51:51 +0100 Subject: [PATCH 33/81] [list_to_set $ seq] -> [set_eq] --- theories/channel/multi_channel.v | 20 +++++++++----------- 1 file changed, 9 insertions(+), 11 deletions(-) diff --git a/theories/channel/multi_channel.v b/theories/channel/multi_channel.v index c5cc715..3164899 100644 --- a/theories/channel/multi_channel.v +++ b/theories/channel/multi_channel.v @@ -224,7 +224,7 @@ Section channel. (** ** Specifications of [send] and [recv] *) Lemma new_chan_spec (n:nat) (ps:gmap nat (iProto Σ)) : 0 < n → - dom ps = list_to_set $ seq 0 n → (* TODO: Consider using [set_eq] instead *) + dom ps = set_seq 0 n → {{{ iProto_consistent ps }}} new_chan #n {{{ cs, RET cs; chan_pool cs ps }}}. @@ -248,7 +248,6 @@ Section channel. assert (is_Some (ps !! n)) as [p ?]. { apply elem_of_dom. rewrite Hdom. - rewrite list_to_set_seq. apply elem_of_set_seq. lia. } iDestruct (big_sepM_delete _ _ n with "Hps'") as "[Hp Hps']". @@ -259,13 +258,13 @@ Section channel. { iPureIntro. rewrite dom_delete_L. rewrite Hdom. replace (S n) with (n + 1) by lia. - rewrite seq_app. - simpl. rewrite list_to_set_app_L. simpl. + rewrite set_seq_add_L. + simpl. rewrite right_id_L. rewrite difference_union_distr_l_L. rewrite difference_diag_L. rewrite right_id_L. - assert (n ∉ seq 0 n); [|set_solver]. - intros Hin%elem_of_seq. lia. } + assert (n ∉ (set_seq 0 n:gset _)); [|set_solver]. + intros Hin%elem_of_set_seq. lia. } iModIntro. iExists (γEs++[γE]). replace (S n) with (n + 1) by lia. rewrite replicate_add. @@ -310,11 +309,10 @@ Section channel. iExists _,_,_,_. iModIntro. iSplit; [done|]. iSplit. - { iPureIntro. intros. - apply elem_of_dom in H. - rewrite Hdom in H. - rewrite list_to_set_seq in H. - apply elem_of_set_seq in H. lia. } + { iPureIntro. intros i HSome. + apply elem_of_dom in HSome. + rewrite Hdom in HSome. + apply elem_of_set_seq in HSome. lia. } iFrame "IHp". iApply (big_sepL_impl with "H"). iIntros "!>" (i ? HSome') "(Hauth & Hfrag & Hown)". -- GitLab From cf65eca7abd24f5a261aa70ad864cdc37a6a1594 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Fri, 2 Feb 2024 15:12:27 +0100 Subject: [PATCH 34/81] Added valid targets to consistency relation and removed diverges --- theories/channel/multi_channel.v | 73 +++--- theories/channel/multi_proto.v | 215 ++++++++++++++---- .../multi_proto_consistency_examples.v | 38 +++- 3 files changed, 243 insertions(+), 83 deletions(-) diff --git a/theories/channel/multi_channel.v b/theories/channel/multi_channel.v index 3164899..aabf48d 100644 --- a/theories/channel/multi_channel.v +++ b/theories/channel/multi_channel.v @@ -35,13 +35,6 @@ Definition new_chan : val := Definition get_chan : val := λ: "cs" "i", ("cs","i"). -Definition diverge : val := - λ: <>, (rec: "go" <> := "go" #())%V #(). - -Definition guard : val := - λ: "i" "n", - if: "i" < "n" then #() else diverge #(). - Definition wait : val := rec: "go" "c" := match: !"c" with @@ -55,7 +48,7 @@ Definition vpos : val := λ: "n" "i" "j", "i"*"n" + "j". Definition send : val := λ: "c" "j" "v", - let: "n" := Snd (Fst "c") in guard "j" "n";; + let: "n" := Snd (Fst "c") in let: "ls" := Fst (Fst "c") in let: "i" := Snd "c" in let: "l" := "ls" +â‚— vpos "n" "i" "j" in @@ -64,7 +57,7 @@ Definition send : val := (* TODO: Move recursion further in *) Definition recv : val := rec: "go" "c" "j" := - let: "n" := Snd (Fst "c") in guard "j" "n";; + let: "n" := Snd (Fst "c") in let: "ls" := Fst (Fst "c") in let: "i" := Snd "c" in let: "l" := "ls" +â‚— vpos "n" "j" "i" in @@ -110,7 +103,7 @@ Definition iProto_pointsto_def `{!heapGS Σ, !chanG Σ} (c : val) (p : iProto Σ) : iProp Σ := ∃ γ γE1 (l:loc) (i:nat) (n:nat) p', ⌜ c = (#l,#n,#i)%V ⌠∗ - inv (nroot.@"ctx") (iProto_ctx γ) ∗ + inv (nroot.@"ctx") (iProto_ctx γ (set_seq 0 n)) ∗ ([∗list] j ↦ _ ∈ replicate n (), ∃ γE2 γt1 γt2, inv (nroot.@"p") (chan_inv γ γE1 γt1 i j (l +â‚— (pos n i j))) ∗ @@ -131,7 +124,7 @@ Definition chan_pool `{!heapGS Σ, !chanG Σ} (cs : val) (ps : gmap nat (iProto Σ)) : iProp Σ := ∃ γ (γEs : list gname) (l:loc) (n:nat), ⌜cs = (#l,#n)%V⌠∗ ⌜∀ i, is_Some (ps !! i) → i < n⌠∗ - inv (nroot.@"ctx") (iProto_ctx γ) ∗ + inv (nroot.@"ctx") (iProto_ctx γ (set_seq 0 n)) ∗ [∗ list] i ↦ _ ∈ replicate n (), (∀ p, ⌜ps !! i = Some p⌠-∗ own (γEs !!! i) (â—E (Next p)) ∗ @@ -313,7 +306,7 @@ Section channel. apply elem_of_dom in HSome. rewrite Hdom in HSome. apply elem_of_set_seq in HSome. lia. } - iFrame "IHp". + rewrite Hdom. iFrame "IHp". iApply (big_sepL_impl with "H"). iIntros "!>" (i ? HSome') "(Hauth & Hfrag & Hown)". iSplitL. @@ -388,24 +381,6 @@ Section channel. iIntros "Hi Hj". by iDestruct (own_prot_idx with "Hi Hj") as %Hneq. Qed. - Lemma diverge_spec P : - {{{ True }}} diverge #() {{{ RET #(); P }}}. - Proof. - iIntros (Φ) "_ HΦ". - wp_lam. iLöb as "IH". - wp_pures. by iApply "IH". - Qed. - - Lemma guard_spec (i n : nat) : - {{{ True }}} guard #i #n {{{ RET #(); ⌜i < n⌠}}}. - Proof. - iIntros (Φ) "_ HΦ". - wp_lam. wp_pures. - case_bool_decide. - - wp_pures. iApply "HΦ". iPureIntro. lia. - - by wp_smart_apply diverge_spec. - Qed. - Lemma vpos_spec (n i j : nat) : {{{ True }}} vpos #n #i #j {{{ RET #(pos n i j); True }}}. Proof. @@ -423,11 +398,24 @@ Section channel. rewrite iProto_pointsto_eq. iIntros (Φ) "Hc HΦ". wp_lam; wp_pures. iDestruct "Hc" as (γ γE l i n p' ->) "(#IH & #Hls & Hle & Hâ— & Hâ—¯ & Hown)". - wp_pures. - wp_smart_apply guard_spec; [done|]. - iDestruct 1 as %Hle. + wp_bind (Fst _). + iInv "IH" as "HI". + iDestruct (iProto_le_msg_inv_r with "Hle") as (m') "#Heq". + iRewrite "Heq" in "Hown". + iAssert (â–· (â–· ⌜j < n⌠∗ iProto_own γ i (<(Send, j)> m') ∗ + iProto_ctx γ (set_seq 0 n)))%I with "[HI Hown]" as "[HI [Hown Hctx]]". + { iNext. + iDestruct (iProto_target with "HI Hown") as "[H1 [H2 H3]]". + iFrame. + iNext. + iDestruct "H1" as %Hin. iPureIntro. + set_solver. } + iRewrite -"Heq" in "Hown". + wp_pures. iModIntro. + iFrame. wp_pures. wp_smart_apply (vpos_spec); [done|]. iIntros "_". + iDestruct "HI" as %Hle. iDestruct (big_sepL_lookup_acc with "Hls") as "[Hj _]". { by apply lookup_replicate_2. } iDestruct "Hj" as (γE' γt γt') "#[IHl1 IHl2]". @@ -514,12 +502,23 @@ Section channel. rewrite iProto_pointsto_eq. iDestruct "Hc" as (γ γE l i n p' ->) "(#IH & #Hls & Hle & Hâ— & Hâ—¯ & Hown)". + do 6 wp_pure _. wp_bind (Fst _). wp_pure _. + iInv "IH" as "HI". + iDestruct (iProto_le_msg_inv_r with "Hle") as (m') "#Heq". + iRewrite "Heq" in "Hown". + iAssert (â–· (â–· ⌜j < n⌠∗ iProto_own γ i (<(Recv, j)> m') ∗ + iProto_ctx γ (set_seq 0 n)))%I with "[HI Hown]" as "[HI [Hown Hctx]]". + { iNext. + iDestruct (iProto_target with "HI Hown") as "[H1 [H2 H3]]". + iFrame. + iNext. + iDestruct "H1" as %Hin. iPureIntro. + set_solver. } + iRewrite -"Heq" in "Hown". wp_pures. - wp_pures. - wp_smart_apply guard_spec; [done|]. - iDestruct 1 as %Hle. - wp_pures. + iModIntro. iFrame. wp_smart_apply (vpos_spec); [done|]. iIntros "_". + iDestruct "HI" as %Hle. iDestruct (big_sepL_lookup_acc with "Hls") as "[Hj _]". { by apply lookup_replicate_2. } iDestruct "Hj" as (γE' γt γt') "#[IHl1 IHl2]". diff --git a/theories/channel/multi_proto.v b/theories/channel/multi_proto.v index 8cab819..90368e6 100644 --- a/theories/channel/multi_proto.v +++ b/theories/channel/multi_proto.v @@ -515,19 +515,23 @@ Instance iProto_inhabited {Σ V} : Inhabited (iProto Σ V) := populate (END). Definition can_step {Σ V} (rec : gmap nat (iProto Σ V) → iProp Σ) (ps : gmap nat (iProto Σ V)) (i j : nat) : iProp Σ := ∀ m1 m2, - (ps !!! i ≡ <(Send, j)> m1) -∗ (ps !!! j ≡ <(Recv, i)> m2) -∗ + ((ps !!! i ≡ <(Send, j)> m1) -∗ (ps !!! j ≡ <(Recv, i)> m2) -∗ ∀ v p1, iMsg_car m1 v (Next p1) -∗ ∃ p2, iMsg_car m2 v (Next p2) ∗ - â–· (rec (<[i:=p1]>(<[j:=p2]>ps))). + â–· (rec (<[i:=p1]>(<[j:=p2]>ps)))). + +Definition valid_target {Σ V} (ps : gmap nat (iProto Σ V)) (i j : nat) : iProp Σ := + ∀ a m, ((ps !!! i ≡ <(a, j)> m) -∗ ⌜is_Some (ps !! j)âŒ). Definition iProto_consistent_pre {Σ V} (rec : gmap nat (iProto Σ V) → iProp Σ) (ps : gmap nat (iProto Σ V)) : iProp Σ := - ∀ i j, can_step rec ps i j. + (∀ i j, valid_target ps i j) ∗ + (∀ i j, can_step rec ps i j). Global Instance iProto_consistent_pre_ne {Σ V} (rec : gmapO natO (iProto Σ V) -n> iPropO Σ) : NonExpansive (iProto_consistent_pre rec). -Proof. rewrite /iProto_consistent_pre /can_step. solve_proper. Qed. +Proof. rewrite /iProto_consistent_pre /can_step /valid_target. solve_proper. Qed. Program Definition iProto_consistent_pre' {Σ V} (rec : gmapO natO (iProto Σ V) -n> iPropO Σ) : @@ -595,17 +599,6 @@ Proof. solve_proper. Qed. Global Instance iProto_le_proper {Σ V} : Proper ((≡) ==> (≡) ==> (⊣⊢)) (@iProto_le Σ V). Proof. solve_proper. Qed. -(* Definition iProto_le {Σ V} (i:nat) (p1 p2 : iProto Σ V) : iProp Σ := *) -(* ∀ ps, iProto_consistent (<[i:=p1]>ps) -∗ iProto_consistent (<[i:=p2]>ps). *) -(* Arguments iProto_le {_ _} _ _%proto _%proto. *) -(* Global Instance: Params (@iProto_le) 2 := {}. *) -(* Notation "p ⊑{ i } q" := (iProto_le i p q) (at level 20) : bi_scope. *) - -(* Global Instance iProto_le_ne {Σ V} n : Proper ((=) ==> (dist n) ==> (dist n) ==> (dist n)) (@iProto_le Σ V). *) -(* Proof. solve_proper. Qed. *) -(* Global Instance iProto_le_proper {Σ V} : Proper ((=) ==> (≡) ==> (≡) ==> (⊣⊢)) (@iProto_le Σ V). *) -(* Proof. solve_proper. Qed. *) - Record proto_name := ProtName { proto_names : gmap nat gname }. Global Instance proto_name_inhabited : Inhabited proto_name := populate (ProtName inhabitant). @@ -627,8 +620,8 @@ Definition iProto_own_auth `{!protoG Σ V} (γ : gname) own γ (gmap_view_auth (DfracOwn 1) (((λ p, Excl' (Next p)) <$> ps) : gmap _ _)). Definition iProto_ctx `{protoG Σ V} - (γ : gname) : iProp Σ := - ∃ ps, iProto_own_auth γ ps ∗ â–· iProto_consistent ps. + (γ : gname) (ps_dom : gset nat) : iProp Σ := + ∃ ps, ⌜dom ps = ps_dom⌠∗ iProto_own_auth γ ps ∗ â–· iProto_consistent ps. (** * The connective for ownership of channel ends *) Definition iProto_own `{!protoG Σ V} @@ -659,14 +652,14 @@ Section proto. Lemma iProto_le_end : ⊢ END ⊑ (END : iProto Σ V). Proof. rewrite iProto_le_unfold. iLeft. auto 10. Qed. - Lemma iProto_le_end_inv_l p : p ⊑ END -∗ (p ≡ END). + Lemma iProto_le_end_inv_r p : p ⊑ END -∗ (p ≡ END). Proof. rewrite iProto_le_unfold. iIntros "[[Hp _]|H] //". iDestruct "H" as (a1 a2 m1 m2) "(_ & Heq & _)". by iDestruct (iProto_end_message_equivI with "Heq") as %[]. Qed. - Lemma iProto_le_end_inv_r p : END ⊑ p -∗ (p ≡ END). + Lemma iProto_le_end_inv_l p : END ⊑ p -∗ (p ≡ END). Proof. rewrite iProto_le_unfold. iIntros "[[_ Hp]|H] //". iDestruct "H" as (a1 a2 m1 m2) "(Heq & _ & _)". @@ -679,7 +672,7 @@ Section proto. ∀ v p2', iMsg_car m2 v (Next p2') -∗ ∃ p1', â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1'). Proof. - rewrite iProto_le_unfold. + rewrite iProto_le_unfold. iIntros "[[_ Heq]|H]". { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } iDestruct "H" as (a1 a2 m1 m2') "(Hp1 & Hp2 & H)". @@ -703,14 +696,13 @@ Section proto. iRewrite -("Hm1" $! v (Next p1')) in "Hm". auto with iFrame. Qed. - Lemma iProto_le_recv_inv_l i m1 p2 : (<(Recv,i)> m1) ⊑ p2 -∗ ∃ m2, (p2 ≡ <(Recv,i)> m2) ∗ ∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). Proof. - rewrite iProto_le_unfold. + rewrite iProto_le_unfold. iIntros "[[Heq _]|H]". { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } iDestruct "H" as (a1 a2 m1' m2) "(Hp1 & Hp2 & H)". @@ -730,7 +722,7 @@ Section proto. ∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). Proof. - rewrite iProto_le_unfold. + rewrite iProto_le_unfold. iIntros "[[_ Heq]|H]". { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } iDestruct "H" as (a1 a2 m1 m2') "(Hp1 & Hp2 & H)". @@ -756,6 +748,79 @@ Section proto. by iRewrite -("Hm1" $! v (Next p1')). Qed. + Lemma iProto_le_msg_inv_l i a m1 p2 : + (<(a,i)> m1) ⊑ p2 -∗ ∃ m2, p2 ≡ <(a,i)> m2. + Proof. + rewrite iProto_le_unfold /iProto_le_pre. + iIntros "[[Heq _]|H]". + { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } + iDestruct "H" as (a1 a2 m1' m2) "(Hp1 & Hp2 & H)". + destruct a1 as [t1 ?], a2 as [t2 ?]. + destruct t1,t2; [|done|done|]. + - rewrite iProto_message_equivI. + iDestruct "Hp1" as (Heq) "Hp1". simplify_eq. + iDestruct "H" as (->) "H". + iExists _. done. + - rewrite iProto_message_equivI. + iDestruct "Hp1" as (Heq) "Hp1". simplify_eq. + iDestruct "H" as (->) "H". + iExists _. done. + Qed. + + Lemma iProto_le_msg_inv_r j a p1 m2 : + (p1 ⊑ <(a,j)> m2) -∗ ∃ m1, p1 ≡ <(a,j)> m1. + Proof. + rewrite iProto_le_unfold /iProto_le_pre. + iIntros "[[_ Heq]|H]". + { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } + iDestruct "H" as (a1 a2 m1 m2') "(Hp1 & Hp2 & H)". + destruct a1 as [t1 ?], a2 as [t2 ?]. + destruct t1,t2; [|done|done|]. + - rewrite iProto_message_equivI. + iDestruct "Hp2" as (Heq) "Hp2". simplify_eq. + iDestruct "H" as (->) "H". + iExists _. done. + - rewrite iProto_message_equivI. + iDestruct "Hp2" as (Heq) "Hp2". simplify_eq. + iDestruct "H" as (->) "H". + iExists _. done. + Qed. + + Lemma valid_target_le ps i p1 p2 : + (∀ i' j', valid_target ps i' j') -∗ + ps !!! i ≡ p1 -∗ + p1 ⊑ p2 -∗ + (∀ i' j', valid_target (<[i := p2]>ps) i' j') ∗ p1 ⊑ p2. + Proof. + iIntros "Hprot #HSome Hle". + pose proof (iProto_case p1) as [Hend|Hmsg]. + { rewrite Hend. iDestruct (iProto_le_end_inv_l with "Hle") as "#H". + iFrame "Hle". + iIntros (i' j' a m) "Hm". + destruct (decide (i = j')) as [->|Hneqj]. + { rewrite lookup_insert. done. } + rewrite lookup_insert_ne; [|done]. + destruct (decide (i = i')) as [->|Hneqi]. + { rewrite lookup_total_insert. iRewrite "H" in "Hm". + by iDestruct (iProto_end_message_equivI with "Hm") as "Hm". } + rewrite lookup_total_insert_ne; [|done]. + by iApply "Hprot". } + destruct Hmsg as (t & n & m & Hmsg). + setoid_rewrite Hmsg. + iDestruct (iProto_le_msg_inv_l with "Hle") as (m2) "#Heq". + iFrame. + iIntros (i' j' a m') "Hm". + destruct (decide (i = j')) as [->|Hneqj]. + { rewrite lookup_insert. done. } + rewrite lookup_insert_ne; [|done]. + destruct (decide (i = i')) as [->|Hneqi]. + { rewrite lookup_total_insert. iRewrite "Heq" in "Hm". + iDestruct (iProto_message_equivI with "Hm") as (Heq) "Hm". + simplify_eq. by iApply "Hprot". } + rewrite lookup_total_insert_ne; [|done]. + by iApply "Hprot". + Qed. + Lemma iProto_consistent_le ps i p1 p2 : iProto_consistent ps -∗ ps !!! i ≡ p1 -∗ @@ -767,6 +832,9 @@ Section proto. iLöb as "IH" forall (p1 p2 ps). iIntros "#HSome". rewrite !iProto_consistent_unfold. + iDestruct "Hprot" as "(Htar & Hprot)". + iDestruct (valid_target_le with "Htar HSome Hle") as "[Htar Hle]". + iFrame. iIntros (i' j' m1 m2) "#Hm1 #Hm2". destruct (decide (i = i')) as [<-|Hneq]. { rewrite lookup_total_insert. @@ -782,7 +850,7 @@ Section proto. inversion Htag. simplify_eq. iIntros (v p) "Hm1'". iSpecialize ("Hm1" $! v (Next p)). - iDestruct (iProto_le_send_inv with "Hle") as "Hle". + iDestruct (iProto_le_send_inv with "Hle") as "Hle". iRewrite -"Hm1" in "Hm1'". iDestruct "Hle" as (m') "[#Heq H]". iDestruct ("H" with "Hm1'") as (p') "[Hle H]". @@ -793,7 +861,7 @@ Section proto. { iRewrite -"Heq". rewrite !lookup_total_alt. iRewrite "HSome". done. } { rewrite lookup_total_insert_ne; [|done]. done. } iDestruct "Hprot" as (p'') "[Hm Hprot]". - iExists p''. iFrame. + iExists p''. iFrame. iNext. iDestruct ("IH" with "Hprot Hle [HSome]") as "HI". { by rewrite lookup_total_insert. } @@ -815,19 +883,20 @@ Section proto. iDestruct "Hm2" as "[%Htag Hm2]". inversion Htag. simplify_eq. iIntros (v p) "Hm1'". - iDestruct (iProto_le_recv_inv_r with "Hle") as "Hle". + iDestruct (iProto_le_recv_inv_r with "Hle") as "Hle". (* iRewrite -"Hm2" in "Hm2'". *) iDestruct "Hle" as (m') "[#Heq Hle]". iDestruct ("Hprot" $!i' with "[] [] Hm1'") as "Hprot". { done. } { rewrite !lookup_total_alt. iRewrite "HSome". done. } iDestruct ("Hprot") as (p') "[Hm1' Hprot]". - iDestruct ("Hle" with "Hm1'") as (p2') "[Hle Hm']". + iDestruct ("Hle" with "Hm1'") as (p2') "[Hle Hm']". iSpecialize ("Hm2" $! v (Next p2')). iExists p2'. iRewrite -"Hm2". iFrame. iDestruct ("IH" with "Hprot Hle []") as "HI". - { iPureIntro. rewrite lookup_total_insert_ne; [|done]. rewrite lookup_total_insert. done. } + { iPureIntro. rewrite lookup_total_insert_ne; [|done]. + rewrite lookup_total_insert. done. } rewrite insert_commute; [|done]. rewrite !insert_insert. done. } rewrite lookup_total_insert_ne; [|done]. @@ -874,10 +943,10 @@ Section proto. Qed. Lemma iProto_le_trans p1 p2 p3 : p1 ⊑ p2 -∗ p2 ⊑ p3 -∗ p1 ⊑ p3. - Proof. + Proof. iIntros "H1 H2". iLöb as "IH" forall (p1 p2 p3). destruct (iProto_case p3) as [->|([]&i&m3&->)]. - - iDestruct (iProto_le_end_inv_l with "H2") as "H2". by iRewrite "H2" in "H1". + - iDestruct (iProto_le_end_inv_r with "H2") as "H2". by iRewrite "H2" in "H1". - iDestruct (iProto_le_send_inv with "H2") as (m2) "[Hp2 H2]". iRewrite "Hp2" in "H1"; clear p2. iDestruct (iProto_le_send_inv with "H1") as (m1) "[Hp1 H1]". @@ -907,7 +976,7 @@ Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. Proof. iIntros "H". iLöb as "IH" forall (p1 p2). destruct (iProto_case p1) as [->|([]&i&m1&->)]. - - iDestruct (iProto_le_end_inv_l with "H") as "H". + - iDestruct (iProto_le_end_inv_r with "H") as "H". iRewrite "H". iApply iProto_le_refl. - iDestruct (iProto_le_send_inv with "H") as (m2) "[Hp2 H]". iRewrite "Hp2"; clear p2. iEval (rewrite !iProto_dual_message). @@ -941,7 +1010,7 @@ Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. Proof. iIntros "H1 H2". iLöb as "IH" forall (p1 p2 p3 p4). destruct (iProto_case p2) as [->|([]&i&m2&->)]. - - iDestruct (iProto_le_end_inv_l with "H1") as "H1". + - iDestruct (iProto_le_end_inv_r with "H1") as "H1". iRewrite "H1". by rewrite !left_id. - iDestruct (iProto_le_send_inv with "H1") as (m1) "[Hp1 H1]". iRewrite "Hp1"; clear p1. rewrite !iProto_app_message. @@ -962,7 +1031,7 @@ Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. (P -∗ (<(Recv,i)> MSG v; p) ⊑ (<(Recv,i)> m)) ⊢ (<(Recv,i)> MSG v {{ P }}; p) ⊑ <(Recv,i)> m. Proof. - rewrite iMsg_base_eq. iIntros "H". + rewrite iMsg_base_eq. iIntros "H". iApply iProto_le_recv. iIntros (v' p') "(->&Hp&HP)". iApply (iProto_le_recv_recv_inv with "(H HP)"); simpl; auto. Qed. @@ -971,7 +1040,7 @@ Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. (<(Send,i)> m) ⊑ (<(Send,i)> MSG v {{ P }}; p). Proof. rewrite iMsg_base_eq. iIntros "H". - iApply iProto_le_send. iIntros (v' p') "(->&Hp&HP)". + iApply iProto_le_send. iIntros (v' p') "(->&Hp&HP)". iApply (iProto_le_send_send_inv with "(H HP)"); simpl; auto. Qed. Lemma iProto_le_payload_intro_l i v P p : @@ -1049,6 +1118,16 @@ Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. iApply iProto_le_trans; [|by iApply iProto_le_exist_intro_r]. iApply IH. Qed. + Lemma iProto_consistent_target ps m a i j : + iProto_consistent ps -∗ + ps !!! i ≡ (<(a, j)> m) -∗ + ⌜is_Some (ps !! j)âŒ. + Proof. + rewrite iProto_consistent_unfold. + iDestruct 1 as "[Htar _]". + iApply "Htar". + Qed. + Lemma iProto_consistent_step ps m1 m2 i j v p1 : iProto_consistent ps -∗ ps !!! i ≡ (<(Send, j)> m1) -∗ @@ -1059,6 +1138,7 @@ Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. Proof. iIntros "Hprot #Hi #Hj Hm1". rewrite iProto_consistent_unfold /iProto_consistent_pre. + iDestruct "Hprot" as "[_ Hprot]". iDestruct ("Hprot" with "Hi Hj Hm1") as (p2) "[Hm2 Hprot]". iExists p2. iFrame. Qed. @@ -1129,7 +1209,7 @@ Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. Lemma iProto_init ps : â–· iProto_consistent ps -∗ - |==> ∃ γ, iProto_ctx γ ∗ [∗ map] i ↦p ∈ ps, iProto_own γ i p. + |==> ∃ γ, iProto_ctx γ (dom ps) ∗ [∗ map] i ↦p ∈ ps, iProto_own γ i p. Proof. iIntros "Hconsistnet". iMod iProto_own_auth_alloc as (γ) "[Hauth Hfrags]". @@ -1147,23 +1227,42 @@ Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. iDestruct "H" as %H%gmap_view_frag_op_validN. by destruct H. Qed. - Lemma iProto_step γ i j m1 m2 p1 v : - iProto_ctx γ -∗ + Lemma iProto_step γ ps_dom i j m1 m2 p1 v : + iProto_ctx γ ps_dom -∗ iProto_own γ i (<(Send, j)> m1) -∗ iProto_own γ j (<(Recv, i)> m2) -∗ iMsg_car m1 v (Next p1) ==∗ - â–· ∃ p2, iMsg_car m2 v (Next p2) ∗ iProto_ctx γ ∗ + â–· ∃ p2, iMsg_car m2 v (Next p2) ∗ iProto_ctx γ ps_dom ∗ iProto_own γ i p1 ∗ iProto_own γ j p2. Proof. iIntros "Hctx Hi Hj Hm". iDestruct "Hi" as (pi) "[Hile Hi]". iDestruct "Hj" as (pj) "[Hjle Hj]". - iDestruct "Hctx" as (ps) "[Hauth Hconsistent]". + iDestruct "Hctx" as (ps Hdom) "[Hauth Hconsistent]". iDestruct (iProto_own_auth_agree with "Hauth Hi") as "#Hpi". iDestruct (iProto_own_auth_agree with "Hauth Hj") as "#Hpj". - iDestruct (own_prot_idx with "Hi Hj") as %Hneq. + iDestruct (own_prot_idx with "Hi Hj") as %Hneq. iAssert (â–· (<[i:=<(Send, j)> m1]>ps !!! j ≡ pj))%I as "Hpj'". { rewrite lookup_total_insert_ne; done. } + + iAssert (â–· (⌜is_Some (ps !! i)⌠∗ (pi ⊑ (<(Send, j)> m1))))%I with "[Hile]" + as "[Hi' Hile]". + { iNext. + iDestruct (iProto_le_msg_inv_r with "Hile") as (m) "#H". + iFrame. + iRewrite "H" in "Hpi". + rewrite lookup_total_alt. simpl. + destruct (ps !! i); [done|]. + simpl. iDestruct (iProto_end_message_equivI with "Hpi") as "[]". } + iAssert (â–· (⌜is_Some (ps !! j)⌠∗ (pj ⊑ (<(Recv, i)> m2))))%I with "[Hjle]" + as "[Hj' Hjle]". + { iNext. + iDestruct (iProto_le_msg_inv_r with "Hjle") as (m) "#H". + iFrame. + iRewrite "H" in "Hpj". + rewrite !lookup_total_alt. simpl. + destruct (ps !! j); [done|]. + simpl. iDestruct (iProto_end_message_equivI with "Hpj") as "[]". } iDestruct (iProto_consistent_le with "Hconsistent Hpi Hile") as "Hconsistent". iDestruct (iProto_consistent_le with "Hconsistent Hpj' Hjle") as "Hconsistent". iDestruct (iProto_consistent_step _ _ _ i j with "Hconsistent [] [] [Hm //]") as @@ -1172,12 +1271,19 @@ Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. rewrite lookup_total_insert_ne; [|done]. iNext. rewrite lookup_total_insert. done. } { rewrite lookup_total_insert_ne; [|done]. - iNext. rewrite lookup_total_insert. done. } + iNext. rewrite lookup_total_insert. done. } iMod (iProto_own_auth_update _ _ _ _ p2 with "Hauth Hj") as "[Hauth Hj]". iMod (iProto_own_auth_update _ _ _ _ p1 with "Hauth Hi") as "[Hauth Hi]". iIntros "!>!>". iExists p2. iFrame. + iDestruct "Hi'" as %Hi. + iDestruct "Hj'" as %Hj. iSplitL "Hconsistent Hauth". - { iExists _. iFrame. rewrite insert_insert. + { iExists (<[i:=p1]> (<[j:=p2]> ps)). + iSplit. + { rewrite !dom_insert_lookup_L; [done|..]. + - done. + - by rewrite lookup_insert_ne. } + iFrame. rewrite insert_insert. rewrite insert_commute; [|done]. rewrite insert_insert. rewrite insert_commute; [|done]. done. } iSplitL "Hi". @@ -1185,6 +1291,31 @@ Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. - iExists _. iFrame. iApply iProto_le_refl. Qed. + Lemma iProto_target γ ps_dom i a j m : + iProto_ctx γ ps_dom -∗ + iProto_own γ i (<(a, j)> m) -∗ + â–· (⌜j ∈ ps_domâŒ) ∗ iProto_ctx γ ps_dom ∗ iProto_own γ i (<(a, j)> m). + Proof. + iIntros "Hctx Hown". + rewrite /iProto_ctx /iProto_own. + iDestruct "Hctx" as (ps Hdom) "[Hauth Hps]". + iDestruct "Hown" as (p') "[Hle Hown]". + iDestruct (iProto_own_auth_agree with "Hauth Hown") as "#H". + iDestruct (iProto_le_msg_inv_r with "Hle") as (m') "#Heq". + iAssert (â–· (⌜is_Some (ps !! j)⌠∗ iProto_consistent ps))%I + with "[Hps]" as "[H1 H2]". + { iNext. + iRewrite "Heq" in "H". + iDestruct (iProto_consistent_target with "Hps H") as "#H'". + iFrame. done. } + iSplitL "H1". + { iNext. iDestruct "H1" as %Heq. + iPureIntro. simplify_eq. apply elem_of_dom. done. } + iSplitL "Hauth H2". + { iExists _. iFrame. done. } + iExists _. iFrame. + Qed. + (* (** The instances below make it possible to use the tactics [iIntros], *) (* [iExist], [iSplitL]/[iSplitR], [iFrame] and [iModIntro] on [iProto_le] goals. *) *) Global Instance iProto_le_from_forall_l {A} i (m1 : A → iMsg Σ V) m2 name : diff --git a/theories/channel/multi_proto_consistency_examples.v b/theories/channel/multi_proto_consistency_examples.v index 826b9b4..a52088b 100644 --- a/theories/channel/multi_proto_consistency_examples.v +++ b/theories/channel/multi_proto_consistency_examples.v @@ -7,6 +7,7 @@ Set Default Proof Using "Type". Export action. Lemma iProto_consistent_equiv_proof {Σ} (ps : gmap nat (iProto Σ)) : + (∀ i j, valid_target ps i j) ∗ (∀ i j m1 m2, (ps !!! i ≡ (<(Send, j)> m1)%proto) -∗ (ps !!! j ≡ (<(Recv, i)> m2)%proto) -∗ @@ -22,8 +23,9 @@ Lemma iProto_consistent_equiv_proof {Σ} (ps : gmap nat (iProto Σ)) : (<[i:=tele_app tp1 x]>(<[j:=tele_app tp2 y]>ps)))) -∗ iProto_consistent ps. Proof. - iIntros "H". + iIntros "[H' H]". rewrite iProto_consistent_unfold. + iFrame. iIntros (i j m1 m2) "Hm1 Hm2". iDestruct ("H" with "Hm1 Hm2") as (m1' m2' TT1 TT2 tv1 tP1 tp1 tv2 tP2 tp2) @@ -51,13 +53,11 @@ Qed. (* TODO: Improve automation *) (* Could clean up repeated inserts to save traverses *) -Tactic Notation "iProto_consistent_take_step" := +Tactic Notation "iProto_consistent_take_step_step" := let i := fresh in let j := fresh in let m1 := fresh in let m2 := fresh in - try iNext; - iApply iProto_consistent_equiv_proof; iIntros (i j m1 m2) "#Hm1 #Hm2"; repeat (destruct i as [|i]; [repeat (rewrite lookup_total_insert_ne; [|lia]); @@ -86,6 +86,36 @@ Tactic Notation "iProto_consistent_take_step" := repeat (rewrite (insert_commute _ _ j); [|done]); rewrite insert_insert). +Tactic Notation "iProto_consistent_take_step_target" := + let i := fresh in + iIntros (i j a m); rewrite /valid_target; + iIntros "#Hm"; + repeat (destruct i as [|i]; + [repeat (rewrite lookup_total_insert_ne; [|lia]); + try (by rewrite lookup_total_empty iProto_end_message_equivI); + try (rewrite lookup_total_insert; + try (by rewrite iProto_end_message_equivI); + iDestruct (iProto_message_equivI with "Hm1") + as "[%Heq1 Hm1']";simplify_eq)| + repeat (rewrite lookup_total_insert_ne; [|lia]); + try (by rewrite lookup_total_empty iProto_end_message_equivI)]); + repeat (rewrite lookup_total_insert_ne; [|lia]); + try rewrite lookup_total_empty; + try (by iProto_end_message_equivI); + rewrite lookup_total_insert; + iDestruct (iProto_message_equivI with "Hm") + as "[%Heq Hm']";simplify_eq; + repeat (try rewrite lookup_empty; + try rewrite lookup_insert; + rewrite lookup_insert_ne; [|lia]); + try rewrite lookup_insert; try done. + +Tactic Notation "iProto_consistent_take_step" := + try iNext; + iApply iProto_consistent_equiv_proof; + iSplitR; [iProto_consistent_take_step_target| + iProto_consistent_take_step_step]. + Tactic Notation "clean_map" constr(i) := iEval (repeat (rewrite (insert_commute _ _ i); [|done])); rewrite (insert_insert _ i). -- GitLab From d74e871f5777052e02796f566e5d13a51c617289 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Fri, 2 Feb 2024 16:24:50 +0100 Subject: [PATCH 35/81] Some cleanup --- theories/channel/multi_channel.v | 114 +++++++++------------- theories/channel/multi_proto.v | 158 ++++++++++++------------------- 2 files changed, 103 insertions(+), 169 deletions(-) diff --git a/theories/channel/multi_channel.v b/theories/channel/multi_channel.v index aabf48d..2cf1f6d 100644 --- a/theories/channel/multi_channel.v +++ b/theories/channel/multi_channel.v @@ -148,9 +148,8 @@ Section channel. Lemma iProto_pointsto_le c p1 p2 : c ↣ p1 ⊢ â–· (p1 ⊑ p2) -∗ c ↣ p2. Proof. rewrite iProto_pointsto_eq. - iDestruct 1 as - (γ γE l n i p ->) "(#IH & Hls & Hle & Hâ— & Hâ—¯ & Hown)". - iIntros "Hle'". iExists γ, γE, l, n, i, p. + iDestruct 1 as (?????? ->) "(#IH & Hls & Hle & Hâ— & Hâ—¯ & Hown)". + iIntros "Hle'". iExists _,_,_,_,_,p'. iSplit; [done|]. iFrame "#∗". iApply (iProto_le_trans with "Hle Hle'"). Qed. @@ -160,36 +159,28 @@ Section channel. ([∗ list] i ↦ _ ∈ replicate n x2, P i). Proof. iIntros "H". - iInduction n as [|n] "IHn". - { done. } + iInduction n as [|n] "IHn"; [done|]. replace (S n) with (n + 1) by lia. - rewrite !replicate_add. - simpl. iDestruct "H" as "[H1 H2]". - iSplitL "H1". - { by iApply "IHn". } - simpl. rewrite !replicate_length. iFrame. + rewrite !replicate_add /=. iDestruct "H" as "[H1 H2]". + iSplitL "H1"; [by iApply "IHn"|]=> /=. + by rewrite !replicate_length. Qed. Lemma array_to_matrix_pre l n m v : l ↦∗ replicate (n * m) v -∗ - [∗ list] i ↦ _ ∈ replicate n (), - (l +â‚— i*m) ↦∗ replicate m v. + [∗ list] i ↦ _ ∈ replicate n (), (l +â‚— i*m) ↦∗ replicate m v. Proof. iIntros "Hl". - iInduction n as [|n] "IHn". - { done. } + iInduction n as [|n] "IHn"; [done|]. replace (S n) with (n + 1) by lia. replace ((n + 1) * m) with (n * m + m) by lia. - rewrite !replicate_add. simpl. - rewrite array_app. - iDestruct "Hl" as "[H1 H2]". - iDestruct ("IHn" with "H1") as "H1". - iFrame. - simpl. - rewrite Nat.add_0_r. - rewrite !replicate_length. + rewrite !replicate_add /= array_app=> /=. + iDestruct "Hl" as "[Hl Hls]". + iDestruct ("IHn" with "Hl") as "Hl". + iFrame=> /=. + rewrite Nat.add_0_r !replicate_length. replace (Z.of_nat (n * m)) with (Z.of_nat n * Z.of_nat m)%Z by lia. - iFrame. + by iFrame. Qed. Lemma array_to_matrix l n v : @@ -198,20 +189,17 @@ Section channel. [∗ list] j ↦ _ ∈ replicate n (), (l +â‚— pos n i j) ↦ v. Proof. - iIntros "H". - iDestruct (array_to_matrix_pre with "H") as "H". - iApply (big_sepL_impl with "H"). - iIntros "!>" (i ? HSome) "H". - clear HSome. - rewrite /array. + iIntros "Hl". + iDestruct (array_to_matrix_pre with "Hl") as "Hl". + iApply (big_sepL_impl with "Hl"). + iIntros "!>" (i ? _) "Hl". iApply big_sepL_replicate. - iApply (big_sepL_impl with "H"). + iApply (big_sepL_impl with "Hl"). iIntros "!>" (j ? HSome) "Hl". - rewrite /pos. rewrite Loc.add_assoc. replace (Z.of_nat i * Z.of_nat n + Z.of_nat j)%Z with (Z.of_nat (i * n + j))%Z by lia. - apply lookup_replicate in HSome as [-> _]. done. + by apply lookup_replicate in HSome as [-> _]. Qed. (** ** Specifications of [send] and [recv] *) @@ -234,49 +222,37 @@ Section channel. own (γEs !!! i) (â—¯E (Next (ps !!! i))) ∗ iProto_own γ i (ps !!! i))%I with "[Hps']" as "H". { clear Hle. - (* iInduction n as [|n] "IHn" forall (ps HSome Heq). *) iInduction n as [|n] "IHn" forall (ps Hdom). { iExists []. iModIntro. simpl. done. } - (* assert (n < S n) by lia. *) - assert (is_Some (ps !! n)) as [p ?]. - { apply elem_of_dom. - rewrite Hdom. - apply elem_of_set_seq. - lia. } - iDestruct (big_sepM_delete _ _ n with "Hps'") as "[Hp Hps']". - { done. } + assert (is_Some (ps !! n)) as [p HSome]. + { apply elem_of_dom. rewrite Hdom. apply elem_of_set_seq. lia. } + iDestruct (big_sepM_delete _ _ n with "Hps'") as "[Hp Hps']"; [done|]. iMod (own_alloc (â—E (Next p) â‹… â—¯E (Next p))) as (γE) "[Hauth Hfrag]". { apply excl_auth_valid. } iMod ("IHn" with "[] Hps'") as (γEs Hlen) "H". { iPureIntro. - rewrite dom_delete_L. rewrite Hdom. + rewrite dom_delete_L Hdom. replace (S n) with (n + 1) by lia. - rewrite set_seq_add_L. - simpl. - rewrite right_id_L. - rewrite difference_union_distr_l_L. - rewrite difference_diag_L. rewrite right_id_L. + rewrite set_seq_add_L /= right_id_L difference_union_distr_l_L + difference_diag_L right_id_L. assert (n ∉ (set_seq 0 n:gset _)); [|set_solver]. intros Hin%elem_of_set_seq. lia. } iModIntro. iExists (γEs++[γE]). replace (S n) with (n + 1) by lia. - rewrite replicate_add. - rewrite big_sepL_app. - rewrite app_length. - rewrite Hlen. iSplit; [done|]. - simpl. + rewrite replicate_add big_sepL_app app_length Hlen. + iSplit; [done|]=> /=. iSplitL "H". { iApply (big_sepL_impl with "H"). iIntros "!>" (i ? HSome') "(Hauth & Hfrag & Hown)". - assert (i < n). - { by apply lookup_replicate_1 in HSome' as [? ?]. } + assert (i < n) as Hle. + { by apply lookup_replicate_1 in HSome' as [??]. } assert (delete n ps !!! i = ps !!! i) as Heq'. { apply lookup_total_delete_ne. lia. } rewrite Heq'. iFrame. rewrite lookup_total_app_l; [|lia]. iFrame. } - simpl. rewrite replicate_length. rewrite Nat.add_0_r. + rewrite replicate_length Nat.add_0_r. rewrite list_lookup_total_middle; [|done]. - rewrite lookup_total_alt. rewrite H. simpl. iFrame. } + rewrite lookup_total_alt HSome. by iFrame. } iMod "H" as (γEs Hlen) "H". iAssert (|={⊤}=> [∗ list] i ↦ _ ∈ replicate n (), @@ -292,11 +268,8 @@ Section channel. iApply big_sepL_fupd. iApply (big_sepL_impl with "H1"). iIntros "!>" (j ? HSome'') "H1". - iMod (own_alloc (Excl ())) as (γ') "Hγ'". - { done. } - iExists γ'. - iApply inv_alloc. - iLeft. iFrame. } + iMod (own_alloc (Excl ())) as (γ') "Hγ'"; [done|]. + iExists γ'. iApply inv_alloc. iLeft. iFrame. } iMod "IH" as "#IH". iMod (inv_alloc with "Hps") as "#IHp". iExists _,_,_,_. @@ -319,16 +292,15 @@ Section channel. { apply lookup_replicate in HSome' as [_ Hle']. done. } assert (j < n) as Hle''. { apply lookup_replicate in HSome'' as [_ Hle'']. done. } - iDestruct (big_sepL_lookup _ _ i () with "IH") as "IH''". - { rewrite lookup_replicate. done. } - iDestruct (big_sepL_lookup _ _ j () with "IH''") as "IH'''". - { rewrite lookup_replicate. done. } - iFrame "#". - iDestruct (big_sepL_lookup _ _ j () with "IH") as "H". - { rewrite lookup_replicate. done. } - iDestruct (big_sepL_lookup _ _ i () with "H") as "H'". - { rewrite lookup_replicate. done. } - iDestruct "IH'''" as (γ1) "?". + iDestruct (big_sepL_lookup _ _ i () with "IH") as "IH'"; + [by rewrite lookup_replicate|]. + iDestruct (big_sepL_lookup _ _ j () with "IH'") as "IH''"; + [by rewrite lookup_replicate|]. + iDestruct (big_sepL_lookup _ _ j () with "IH") as "H"; + [by rewrite lookup_replicate|]. + iDestruct (big_sepL_lookup _ _ i () with "H") as "H'"; + [by rewrite lookup_replicate|]. + iDestruct "IH''" as (γ1) "?". iDestruct "H'" as (γ2) "?". iExists _, _. iFrame "#". Qed. diff --git a/theories/channel/multi_proto.v b/theories/channel/multi_proto.v index 90368e6..86f4505 100644 --- a/theories/channel/multi_proto.v +++ b/theories/channel/multi_proto.v @@ -510,23 +510,23 @@ Section proto. End proto. -Instance iProto_inhabited {Σ V} : Inhabited (iProto Σ V) := populate (END). +Global Instance iProto_inhabited {Σ V} : Inhabited (iProto Σ V) := populate END. Definition can_step {Σ V} (rec : gmap nat (iProto Σ V) → iProp Σ) (ps : gmap nat (iProto Σ V)) (i j : nat) : iProp Σ := ∀ m1 m2, - ((ps !!! i ≡ <(Send, j)> m1) -∗ (ps !!! j ≡ <(Recv, i)> m2) -∗ - ∀ v p1, iMsg_car m1 v (Next p1) -∗ - ∃ p2, iMsg_car m2 v (Next p2) ∗ - â–· (rec (<[i:=p1]>(<[j:=p2]>ps)))). + (ps !!! i ≡ <(Send, j)> m1) -∗ + (ps !!! j ≡ <(Recv, i)> m2) -∗ + ∀ v p1, iMsg_car m1 v (Next p1) -∗ + ∃ p2, iMsg_car m2 v (Next p2) ∗ + â–· (rec (<[i:=p1]>(<[j:=p2]>ps))). Definition valid_target {Σ V} (ps : gmap nat (iProto Σ V)) (i j : nat) : iProp Σ := - ∀ a m, ((ps !!! i ≡ <(a, j)> m) -∗ ⌜is_Some (ps !! j)âŒ). + ∀ a m, (ps !!! i ≡ <(a, j)> m) -∗ ⌜is_Some (ps !! j)âŒ. Definition iProto_consistent_pre {Σ V} (rec : gmap nat (iProto Σ V) → iProp Σ) (ps : gmap nat (iProto Σ V)) : iProp Σ := - (∀ i j, valid_target ps i j) ∗ - (∀ i j, can_step rec ps i j). + (∀ i j, valid_target ps i j) ∗ (∀ i j, can_step rec ps i j). Global Instance iProto_consistent_pre_ne {Σ V} (rec : gmapO natO (iProto Σ V) -n> iPropO Σ) : @@ -616,7 +616,6 @@ Definition iProto_own_frag `{!protoG Σ V} (γ : gname) Definition iProto_own_auth `{!protoG Σ V} (γ : gname) (ps : gmap nat (iProto Σ V)) : iProp Σ := - (* own γ (◠∅). *) own γ (gmap_view_auth (DfracOwn 1) (((λ p, Excl' (Next p)) <$> ps) : gmap _ _)). Definition iProto_ctx `{protoG Σ V} @@ -669,21 +668,18 @@ Section proto. Lemma iProto_le_send_inv i p1 m2 : p1 ⊑ (<(Send,i)> m2) -∗ ∃ m1, (p1 ≡ <(Send,i)> m1) ∗ - ∀ v p2', iMsg_car m2 v (Next p2') -∗ ∃ p1', - â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1'). + ∀ v p2', iMsg_car m2 v (Next p2') -∗ + ∃ p1', â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1'). Proof. rewrite iProto_le_unfold. iIntros "[[_ Heq]|H]". - { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } + { by iDestruct (iProto_message_end_equivI with "Heq") as %[]. } iDestruct "H" as (a1 a2 m1 m2') "(Hp1 & Hp2 & H)". rewrite iProto_message_equivI. iDestruct "Hp2" as "[%Heq Hm2]". simplify_eq. destruct a1 as [[]]; [|done]. - iDestruct "H" as (->) "H". - iExists m1. iFrame. - iIntros (v p2). - iSpecialize ("Hm2" $! v (Next p2)). - iRewrite "Hm2". iApply "H". + iDestruct "H" as (->) "H". iExists m1. iFrame "Hp1". + iIntros (v p2). iSpecialize ("Hm2" $! v (Next p2)). by iRewrite "Hm2". Qed. Lemma iProto_le_send_send_inv i m1 m2 v p2' : @@ -699,8 +695,8 @@ Section proto. Lemma iProto_le_recv_inv_l i m1 p2 : (<(Recv,i)> m1) ⊑ p2 -∗ ∃ m2, (p2 ≡ <(Recv,i)> m2) ∗ - ∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', - â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). + ∀ v p1', iMsg_car m1 v (Next p1') -∗ + ∃ p2', â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). Proof. rewrite iProto_le_unfold. iIntros "[[Heq _]|H]". @@ -709,18 +705,15 @@ Section proto. rewrite iProto_message_equivI. iDestruct "Hp1" as "[%Heq Hm1]". simplify_eq. destruct a2 as [[]]; [done|]. - iDestruct "H" as (->) "H". - iExists m2. iFrame. - iIntros (v p1). - iSpecialize ("Hm1" $! v (Next p1)). - iRewrite "Hm1". iApply "H". + iDestruct "H" as (->) "H". iExists m2. iFrame "Hp2". + iIntros (v p1). iSpecialize ("Hm1" $! v (Next p1)). by iRewrite "Hm1". Qed. Lemma iProto_le_recv_inv_r i p1 m2 : (p1 ⊑ <(Recv,i)> m2) -∗ ∃ m1, (p1 ≡ <(Recv,i)> m1) ∗ - ∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', - â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). + ∀ v p1', iMsg_car m1 v (Next p1') -∗ + ∃ p2', â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). Proof. rewrite iProto_le_unfold. iIntros "[[_ Heq]|H]". @@ -759,12 +752,10 @@ Section proto. destruct t1,t2; [|done|done|]. - rewrite iProto_message_equivI. iDestruct "Hp1" as (Heq) "Hp1". simplify_eq. - iDestruct "H" as (->) "H". - iExists _. done. + iDestruct "H" as (->) "H". by iExists _. - rewrite iProto_message_equivI. iDestruct "Hp1" as (Heq) "Hp1". simplify_eq. - iDestruct "H" as (->) "H". - iExists _. done. + iDestruct "H" as (->) "H". by iExists _. Qed. Lemma iProto_le_msg_inv_r j a p1 m2 : @@ -778,12 +769,10 @@ Section proto. destruct t1,t2; [|done|done|]. - rewrite iProto_message_equivI. iDestruct "Hp2" as (Heq) "Hp2". simplify_eq. - iDestruct "H" as (->) "H". - iExists _. done. + iDestruct "H" as (->) "H". by iExists _. - rewrite iProto_message_equivI. iDestruct "Hp2" as (Heq) "Hp2". simplify_eq. - iDestruct "H" as (->) "H". - iExists _. done. + iDestruct "H" as (->) "H". by iExists _. Qed. Lemma valid_target_le ps i p1 p2 : @@ -798,7 +787,7 @@ Section proto. iFrame "Hle". iIntros (i' j' a m) "Hm". destruct (decide (i = j')) as [->|Hneqj]. - { rewrite lookup_insert. done. } + { by rewrite lookup_insert. } rewrite lookup_insert_ne; [|done]. destruct (decide (i = i')) as [->|Hneqi]. { rewrite lookup_total_insert. iRewrite "H" in "Hm". @@ -807,8 +796,7 @@ Section proto. by iApply "Hprot". } destruct Hmsg as (t & n & m & Hmsg). setoid_rewrite Hmsg. - iDestruct (iProto_le_msg_inv_l with "Hle") as (m2) "#Heq". - iFrame. + iDestruct (iProto_le_msg_inv_l with "Hle") as (m2) "#Heq". iFrame "Hle". iIntros (i' j' a m') "Hm". destruct (decide (i = j')) as [->|Hneqj]. { rewrite lookup_insert. done. } @@ -1020,7 +1008,8 @@ Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. iExists (p1' <++> p3). iSplitR "Hm1"; [|by simpl; eauto]. iIntros "!>". iRewrite "Hp24". by iApply ("IH" with "H1"). - iDestruct (iProto_le_recv_inv_r with "H1") as (m1) "[Hp1 H1]". - iRewrite "Hp1"; clear p1. rewrite !iProto_app_message. iApply iProto_le_recv. + iRewrite "Hp1"; clear p1. rewrite !iProto_app_message. + iApply iProto_le_recv. iIntros (v p13). iDestruct 1 as (p1') "[Hm1 #Hp13]". iDestruct ("H1" with "Hm1") as (p2'') "[H1 Hm2]". iExists (p2'' <++> p4). iSplitR "Hm2"; [|by simpl; eauto]. @@ -1123,9 +1112,7 @@ Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. ps !!! i ≡ (<(a, j)> m) -∗ ⌜is_Some (ps !! j)âŒ. Proof. - rewrite iProto_consistent_unfold. - iDestruct 1 as "[Htar _]". - iApply "Htar". + rewrite iProto_consistent_unfold. iDestruct 1 as "[Htar _]". iApply "Htar". Qed. Lemma iProto_consistent_step ps m1 m2 i j v p1 : @@ -1153,13 +1140,10 @@ Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. iDestruct (own_valid_2 with "Hâ— Hâ—¯") as "H✓". rewrite gmap_view_both_validI. iDestruct "H✓" as "[_ [H1 H2]]". - rewrite lookup_total_alt. - rewrite lookup_fmap. + rewrite lookup_total_alt lookup_fmap. destruct (ps !! i); last first. - { simpl. rewrite !option_equivI. done. } - simpl. - rewrite !option_equivI excl_equivI. - by iNext. + { by rewrite !option_equivI. } + simpl. rewrite !option_equivI excl_equivI. by iNext. Qed. Lemma iProto_own_auth_update γ ps i p p' : @@ -1184,19 +1168,13 @@ Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. { apply gmap_view_auth_valid. } iExists γ. iInduction ps as [|i p ps Hin] "IH" using map_ind. - { iModIntro. iFrame. - by iApply big_sepM_empty. } + { iModIntro. iFrame. by iApply big_sepM_empty. } iMod ("IH" with "Hauth") as "[Hauth Hfrags]". rewrite big_sepM_insert; [|done]. iFrame "Hfrags". iMod (own_update with "Hauth") as "[Hauth Hfrag]". - { apply (gmap_view_alloc _ i (DfracOwn 1) (Excl' (Next p))). - - rewrite lookup_fmap. rewrite Hin. done. - - done. - - done. } - iFrame. - iModIntro. - rewrite -fmap_insert. - iFrame. + { apply (gmap_view_alloc _ i (DfracOwn 1) (Excl' (Next p))); [|done|done]. + by rewrite lookup_fmap Hin. } + iModIntro. rewrite -fmap_insert. iFrame. iExists _. iFrame. iApply iProto_le_refl. Qed. @@ -1211,9 +1189,9 @@ Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. â–· iProto_consistent ps -∗ |==> ∃ γ, iProto_ctx γ (dom ps) ∗ [∗ map] i ↦p ∈ ps, iProto_own γ i p. Proof. - iIntros "Hconsistnet". + iIntros "Hconsistent". iMod iProto_own_auth_alloc as (γ) "[Hauth Hfrags]". - iExists γ. iFrame. iExists _. iFrame. done. + iExists γ. iFrame. iExists _. by iFrame. Qed. Lemma own_prot_idx γ i j (p1 p2 : iProto Σ V) : @@ -1224,7 +1202,7 @@ Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. iIntros "Hown Hown'" (->). iDestruct (own_valid_2 with "Hown Hown'") as "H". rewrite uPred.cmra_valid_elim. - iDestruct "H" as %H%gmap_view_frag_op_validN. by destruct H. + by iDestruct "H" as %[]%gmap_view_frag_op_validN. Qed. Lemma iProto_step γ ps_dom i j m1 m2 p1 v : @@ -1243,52 +1221,40 @@ Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. iDestruct (iProto_own_auth_agree with "Hauth Hj") as "#Hpj". iDestruct (own_prot_idx with "Hi Hj") as %Hneq. iAssert (â–· (<[i:=<(Send, j)> m1]>ps !!! j ≡ pj))%I as "Hpj'". - { rewrite lookup_total_insert_ne; done. } - + { by rewrite lookup_total_insert_ne. } iAssert (â–· (⌜is_Some (ps !! i)⌠∗ (pi ⊑ (<(Send, j)> m1))))%I with "[Hile]" as "[Hi' Hile]". - { iNext. - iDestruct (iProto_le_msg_inv_r with "Hile") as (m) "#H". - iFrame. - iRewrite "H" in "Hpi". - rewrite lookup_total_alt. simpl. + { iNext. iDestruct (iProto_le_msg_inv_r with "Hile") as (m) "#Heq". + iFrame. iRewrite "Heq" in "Hpi". rewrite lookup_total_alt. destruct (ps !! i); [done|]. - simpl. iDestruct (iProto_end_message_equivI with "Hpi") as "[]". } + iDestruct (iProto_end_message_equivI with "Hpi") as "[]". } iAssert (â–· (⌜is_Some (ps !! j)⌠∗ (pj ⊑ (<(Recv, i)> m2))))%I with "[Hjle]" as "[Hj' Hjle]". - { iNext. - iDestruct (iProto_le_msg_inv_r with "Hjle") as (m) "#H". - iFrame. - iRewrite "H" in "Hpj". - rewrite !lookup_total_alt. simpl. + { iNext. iDestruct (iProto_le_msg_inv_r with "Hjle") as (m) "#Heq". + iFrame. iRewrite "Heq" in "Hpj". rewrite !lookup_total_alt. destruct (ps !! j); [done|]. - simpl. iDestruct (iProto_end_message_equivI with "Hpj") as "[]". } + iDestruct (iProto_end_message_equivI with "Hpj") as "[]". } iDestruct (iProto_consistent_le with "Hconsistent Hpi Hile") as "Hconsistent". iDestruct (iProto_consistent_le with "Hconsistent Hpj' Hjle") as "Hconsistent". iDestruct (iProto_consistent_step _ _ _ i j with "Hconsistent [] [] [Hm //]") as (p2) "[Hm2 Hconsistent]". { rewrite lookup_total_insert_ne; [|done]. rewrite lookup_total_insert_ne; [|done]. - iNext. rewrite lookup_total_insert. done. } + by rewrite lookup_total_insert. } { rewrite lookup_total_insert_ne; [|done]. - iNext. rewrite lookup_total_insert. done. } + by rewrite lookup_total_insert. } iMod (iProto_own_auth_update _ _ _ _ p2 with "Hauth Hj") as "[Hauth Hj]". iMod (iProto_own_auth_update _ _ _ _ p1 with "Hauth Hi") as "[Hauth Hi]". - iIntros "!>!>". iExists p2. iFrame. - iDestruct "Hi'" as %Hi. - iDestruct "Hj'" as %Hj. + iIntros "!>!>". iExists p2. iFrame "Hm2". + iDestruct "Hi'" as %Hi. iDestruct "Hj'" as %Hj. iSplitL "Hconsistent Hauth". { iExists (<[i:=p1]> (<[j:=p2]> ps)). iSplit. - { rewrite !dom_insert_lookup_L; [done|..]. - - done. - - by rewrite lookup_insert_ne. } + { rewrite !dom_insert_lookup_L; [done..|by rewrite lookup_insert_ne]. } iFrame. rewrite insert_insert. rewrite insert_commute; [|done]. rewrite insert_insert. - rewrite insert_commute; [|done]. done. } - iSplitL "Hi". - - iExists _. iFrame. iApply iProto_le_refl. - - iExists _. iFrame. iApply iProto_le_refl. + by rewrite insert_commute; [|done]. } + iSplitL "Hi"; iExists _; iFrame; iApply iProto_le_refl. Qed. Lemma iProto_target γ ps_dom i a j m : @@ -1300,20 +1266,16 @@ Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. rewrite /iProto_ctx /iProto_own. iDestruct "Hctx" as (ps Hdom) "[Hauth Hps]". iDestruct "Hown" as (p') "[Hle Hown]". - iDestruct (iProto_own_auth_agree with "Hauth Hown") as "#H". + iDestruct (iProto_own_auth_agree with "Hauth Hown") as "#Hi". iDestruct (iProto_le_msg_inv_r with "Hle") as (m') "#Heq". iAssert (â–· (⌜is_Some (ps !! j)⌠∗ iProto_consistent ps))%I - with "[Hps]" as "[H1 H2]". - { iNext. - iRewrite "Heq" in "H". - iDestruct (iProto_consistent_target with "Hps H") as "#H'". - iFrame. done. } - iSplitL "H1". - { iNext. iDestruct "H1" as %Heq. - iPureIntro. simplify_eq. apply elem_of_dom. done. } - iSplitL "Hauth H2". - { iExists _. iFrame. done. } - iExists _. iFrame. + with "[Hps]" as "[HSome Hps]". + { iNext. iRewrite "Heq" in "Hi". + iDestruct (iProto_consistent_target with "Hps Hi") as "#$". by iFrame. } + iSplitL "HSome". + { iNext. iDestruct "HSome" as %Heq. + iPureIntro. simplify_eq. by apply elem_of_dom. } + iSplitL "Hauth Hps"; iExists _; by iFrame. Qed. (* (** The instances below make it possible to use the tactics [iIntros], *) -- GitLab From 53a26ad32a0d2824def8fd227612c2cb58ed7637 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Mon, 5 Feb 2024 12:35:08 +0100 Subject: [PATCH 36/81] More cleanup --- theories/channel/multi_channel.v | 98 +++++++++++--------------------- theories/channel/multi_proto.v | 27 +++++---- 2 files changed, 50 insertions(+), 75 deletions(-) diff --git a/theories/channel/multi_channel.v b/theories/channel/multi_channel.v index 2cf1f6d..0b87509 100644 --- a/theories/channel/multi_channel.v +++ b/theories/channel/multi_channel.v @@ -314,43 +314,31 @@ Section channel. iIntros (HSome Φ) "Hcs HΦ". iDestruct "Hcs" as (γp γEs l n -> Hle) "[#IHp Hl]". wp_lam. wp_pures. - assert (i < n). - { apply Hle. eexists _. done. } - iDestruct (big_sepL_delete _ _ i () with "Hl") as "[[Hi #IHs] H]". - { by apply lookup_replicate. } + assert (i < n); [by apply Hle|]. + iDestruct (big_sepL_delete _ _ i () with "Hl") as "[[Hi #IHs] H]"; + [by apply lookup_replicate|]. iDestruct ("Hi" with "[//]") as "(Hauth & Hown & Hp)". iModIntro. iApply "HΦ". iSplitR "H". - { rewrite iProto_pointsto_eq. - iExists _, _, _, _, _, _. - iSplit; [done|]. - iFrame "#∗". iSplit; [|iNext; done]. + { rewrite iProto_pointsto_eq. iExists _, _, _, _, _, _. + iSplit; [done|]. iFrame "#∗". iSplit; [|iNext; done]. iApply (big_sepL_impl with "IHs"). - iIntros "!>" (???). iDestruct 1 as (γt1 γt2) "[H1 H2]". + iIntros "!>" (???). iDestruct 1 as (γt1 γt2) "[??]". iExists _,_,_. iFrame. } - iExists _, _, _, _. - iSplit; [done|]. + iExists _, _, _, _. iSplit; [done|]. iSplit. { iPureIntro. intros i' HSome'. apply Hle. assert (i ≠i'). { intros ->. rewrite lookup_delete in HSome'. by inversion HSome'. } - rewrite lookup_delete_ne in HSome'; done. } + by rewrite lookup_delete_ne in HSome'. } iFrame "#∗". iApply (big_sepL_impl with "H"). iIntros "!>" (i' ? HSome''). case_decide. { simplify_eq. iFrame "#". iIntros "_" (p' Hin). simplify_eq. by rewrite lookup_delete in Hin. } - rewrite lookup_delete_ne; [|done]. eauto. - Qed. - - Lemma own_prot_excl γ i (p1 p2 : iProto Σ) : - own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p1))) -∗ - own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p2))) -∗ - False. - Proof. - iIntros "Hi Hj". by iDestruct (own_prot_idx with "Hi Hj") as %Hneq. + rewrite lookup_delete_ne; [|done]. by eauto. Qed. Lemma vpos_spec (n i j : nat) : @@ -375,24 +363,17 @@ Section channel. iDestruct (iProto_le_msg_inv_r with "Hle") as (m') "#Heq". iRewrite "Heq" in "Hown". iAssert (â–· (â–· ⌜j < n⌠∗ iProto_own γ i (<(Send, j)> m') ∗ - iProto_ctx γ (set_seq 0 n)))%I with "[HI Hown]" as "[HI [Hown Hctx]]". - { iNext. - iDestruct (iProto_target with "HI Hown") as "[H1 [H2 H3]]". - iFrame. - iNext. - iDestruct "H1" as %Hin. iPureIntro. - set_solver. } - iRewrite -"Heq" in "Hown". - wp_pures. iModIntro. - iFrame. - wp_pures. - wp_smart_apply (vpos_spec); [done|]. iIntros "_". + iProto_ctx γ (set_seq 0 n)))%I with "[HI Hown]" + as "[HI [Hown Hctx]]". + { iNext. iDestruct (iProto_target with "HI Hown") as "[Hin [$ $]]". + iFrame. iNext. iDestruct "Hin" as %Hin. by set_solver. } + iRewrite -"Heq" in "Hown". wp_pures. iModIntro. iFrame. + wp_smart_apply (vpos_spec); [done|]; iIntros "_". iDestruct "HI" as %Hle. - iDestruct (big_sepL_lookup_acc with "Hls") as "[Hj _]". - { by apply lookup_replicate_2. } + iDestruct (big_sepL_lookup_acc with "Hls") as "[Hj _]"; + [by apply lookup_replicate_2|]. iDestruct "Hj" as (γE' γt γt') "#[IHl1 IHl2]". - wp_pures. - wp_bind (Store _ _). + wp_pures. wp_bind (Store _ _). iInv "IHl1" as "HIp". iDestruct "HIp" as "[HIp|HIp]"; last first. { iDestruct "HIp" as "[HIp|HIp]". @@ -410,14 +391,12 @@ Section channel. iDestruct (own_prot_excl with "Hown Hown'") as "H". done. } iDestruct "HIp" as "[>Hl' Htok]". wp_store. - iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]". - { apply excl_auth_update. } + iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]"; [by apply excl_auth_update|]. iModIntro. iSplitL "Hl' Hâ— Hown Hle". { iRight. iLeft. iIntros "!>". iExists _, _. iFrame. - iSplitL "Hown Hle". - { iApply (iProto_own_le with "Hown Hle"). } - iExists _. iFrame. rewrite iMsg_base_eq. simpl. done. } + iSplitL "Hown Hle"; [by iApply (iProto_own_le with "Hown Hle")|]. + iExists _. iFrame. by rewrite iMsg_base_eq. } wp_pures. iLöb as "HL". wp_lam. @@ -456,11 +435,9 @@ Section channel. {{{ RET #(); c ↣ (p tt) }}}. Proof. iIntros (Φ) "[Hc HP] HΦ". - iDestruct (iProto_pointsto_le _ _ (<(Send,i)> MSG v tt; p tt)%proto with "Hc [HP]") - as "Hc". - { iIntros "!>". - iApply iProto_le_trans. - iApply iProto_le_texist_intro_l. + iDestruct (iProto_pointsto_le _ _ (<(Send,i)> MSG v tt; p tt)%proto + with "Hc [HP]") as "Hc". + { iIntros "!>". iApply iProto_le_trans. iApply iProto_le_texist_intro_l. by iApply iProto_le_payload_intro_l. } by iApply (send_spec with "Hc"). Qed. @@ -480,27 +457,20 @@ Section channel. iRewrite "Heq" in "Hown". iAssert (â–· (â–· ⌜j < n⌠∗ iProto_own γ i (<(Recv, j)> m') ∗ iProto_ctx γ (set_seq 0 n)))%I with "[HI Hown]" as "[HI [Hown Hctx]]". - { iNext. - iDestruct (iProto_target with "HI Hown") as "[H1 [H2 H3]]". - iFrame. - iNext. - iDestruct "H1" as %Hin. iPureIntro. - set_solver. } - iRewrite -"Heq" in "Hown". - wp_pures. - iModIntro. iFrame. - wp_smart_apply (vpos_spec); [done|]. iIntros "_". + { iNext. iDestruct (iProto_target with "HI Hown") as "[Hin [$$]]". + iFrame. iNext. iDestruct "Hin" as %Hin. by set_solver. } + iRewrite -"Heq" in "Hown". wp_pures. iModIntro. iFrame. + wp_smart_apply (vpos_spec); [done|]; iIntros "_". iDestruct "HI" as %Hle. - iDestruct (big_sepL_lookup_acc with "Hls") as "[Hj _]". - { by apply lookup_replicate_2. } + iDestruct (big_sepL_lookup_acc with "Hls") as "[Hj _]"; + [by apply lookup_replicate_2|]. iDestruct "Hj" as (γE' γt γt') "#[IHl1 IHl2]". wp_pures. wp_bind (Xchg _ _). iInv "IHl2" as "HIp". iDestruct "HIp" as "[HIp|HIp]". { iDestruct "HIp" as ">[Hl' Htok]". - wp_xchg. - iModIntro. + wp_xchg. iModIntro. iSplitL "Hl' Htok". { iLeft. iFrame. } wp_pures. iApply ("HL" with "[Hâ— Hâ—¯ Hown Hle] HΦ"). @@ -529,11 +499,9 @@ Section channel. rewrite iMsg_base_eq. iDestruct (iMsg_texist_exist with "Hm") as (x <-) "[Hp HP]". wp_pures. - iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]". - { apply excl_auth_update. } + iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]";[apply excl_auth_update|]. iModIntro. iApply "HΦ". - iFrame. - iExists _, _, _, _, _, _. iSplit; [done|]. + iFrame. iExists _, _, _, _, _, _. iSplit; [done|]. iRewrite "Hp". iFrame "#∗". iApply iProto_le_refl. Qed. diff --git a/theories/channel/multi_proto.v b/theories/channel/multi_proto.v index 86f4505..e178523 100644 --- a/theories/channel/multi_proto.v +++ b/theories/channel/multi_proto.v @@ -644,6 +644,23 @@ Section proto. Implicit Types p pl pr : iProto Σ V. Implicit Types m : iMsg Σ V. + Lemma own_prot_idx γ i j (p1 p2 : iProto Σ V) : + own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p1))) -∗ + own γ (gmap_view_frag j (DfracOwn 1) (Excl' (Next p2))) -∗ + ⌜i ≠jâŒ. + Proof. + iIntros "Hown Hown'" (->). + iDestruct (own_valid_2 with "Hown Hown'") as "H". + rewrite uPred.cmra_valid_elim. + by iDestruct "H" as %[]%gmap_view_frag_op_validN. + Qed. + + Lemma own_prot_excl γ i (p1 p2 : iProto Σ V) : + own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p1))) -∗ + own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p2))) -∗ + False. + Proof. iIntros "Hi Hj". by iDestruct (own_prot_idx with "Hi Hj") as %?. Qed. + (** ** Protocol entailment **) Lemma iProto_le_unfold p1 p2 : iProto_le p1 p2 ≡ iProto_le_pre iProto_le p1 p2. Proof. apply: (fixpoint_unfold iProto_le_pre'). Qed. @@ -1194,16 +1211,6 @@ Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. iExists γ. iFrame. iExists _. by iFrame. Qed. - Lemma own_prot_idx γ i j (p1 p2 : iProto Σ V) : - own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p1))) -∗ - own γ (gmap_view_frag j (DfracOwn 1) (Excl' (Next p2))) -∗ - ⌜i ≠jâŒ. - Proof. - iIntros "Hown Hown'" (->). - iDestruct (own_valid_2 with "Hown Hown'") as "H". - rewrite uPred.cmra_valid_elim. - by iDestruct "H" as %[]%gmap_view_frag_op_validN. - Qed. Lemma iProto_step γ ps_dom i j m1 m2 p1 v : iProto_ctx γ ps_dom -∗ -- GitLab From 7132239ebf0f013ba6ebbb5344b2a6639b0511aa Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Tue, 6 Feb 2024 13:59:32 +0100 Subject: [PATCH 37/81] Refactoring --- _CoqProject | 13 +- .../channel/channel.v | 4 +- .../channel/proofmode.v | 4 +- .../channel/proto.v | 2 +- .../channel/proto_consistency_examples.v | 3 +- .../channel/proto_model.v | 2 +- multi_actris/utils/cofe_solver_2.v | 88 ++ theories/channel/channel.v | 487 +++--- theories/channel/proto.v | 1361 ++++++++--------- 9 files changed, 1025 insertions(+), 939 deletions(-) rename theories/channel/multi_channel.v => multi_actris/channel/channel.v (99%) rename theories/channel/multi_proofmode.v => multi_actris/channel/proofmode.v (99%) rename theories/channel/multi_proto.v => multi_actris/channel/proto.v (99%) rename theories/channel/multi_proto_consistency_examples.v => multi_actris/channel/proto_consistency_examples.v (99%) rename theories/channel/multi_proto_model.v => multi_actris/channel/proto_model.v (99%) create mode 100644 multi_actris/utils/cofe_solver_2.v diff --git a/_CoqProject b/_CoqProject index d5da394..98b145e 100644 --- a/_CoqProject +++ b/_CoqProject @@ -1,4 +1,5 @@ -Q theories actris +-Q multi_actris multi_actris # We sometimes want to locally override notation, and there is no good way to do that with scopes. -arg -w -arg -notation-overridden # Cannot use non-canonical projections as it causes massive unification failures @@ -17,11 +18,6 @@ theories/utils/switch.v theories/channel/proto_model.v theories/channel/proto.v theories/channel/channel.v -theories/channel/multi_proto_model.v -theories/channel/multi_proto.v -theories/channel/multi_channel.v -theories/channel/multi_proofmode.v -theories/channel/multi_proto_consistency_examples.v theories/channel/proofmode.v theories/examples/basics.v theories/examples/equivalence.v @@ -58,3 +54,10 @@ theories/logrel/examples/mapper.v theories/logrel/examples/mapper_list.v theories/logrel/examples/compute_service.v theories/logrel/examples/compute_client_list.v + +multi_actris/utils/cofe_solver_2.v +multi_actris/channel/proto_model.v +multi_actris/channel/proto.v +multi_actris/channel/channel.v +multi_actris/channel/proofmode.v +multi_actris/channel/proto_consistency_examples.v diff --git a/theories/channel/multi_channel.v b/multi_actris/channel/channel.v similarity index 99% rename from theories/channel/multi_channel.v rename to multi_actris/channel/channel.v index 0b87509..89937a2 100644 --- a/theories/channel/multi_channel.v +++ b/multi_actris/channel/channel.v @@ -24,8 +24,8 @@ the subprotocol relation [⊑] *) From iris.algebra Require Import gmap excl_auth gmap_view. From iris.base_logic.lib Require Import invariants. From iris.heap_lang Require Export primitive_laws notation proofmode. -From actris.channel Require Import multi_proto_model. -From actris.channel Require Export multi_proto. +From multi_actris.channel Require Import proto_model. +From multi_actris.channel Require Export proto. Set Default Proof Using "Type". (** * The definition of the message-passing connectives *) diff --git a/theories/channel/multi_proofmode.v b/multi_actris/channel/proofmode.v similarity index 99% rename from theories/channel/multi_proofmode.v rename to multi_actris/channel/proofmode.v index ff8c57b..33e7d5a 100644 --- a/theories/channel/multi_proofmode.v +++ b/multi_actris/channel/proofmode.v @@ -9,8 +9,8 @@ In addition to the tactics for symbolic execution, this file defines the tactic recursive protocols are contractive. *) From iris.proofmode Require Import coq_tactics reduction spec_patterns. From iris.heap_lang Require Export proofmode notation. -From actris.channel Require Import multi_proto_model. -From actris Require Export multi_channel. +From multi_actris.channel Require Import proto_model. +From multi_actris Require Export channel. Export action. (** * Tactics for proving contractiveness of protocols *) diff --git a/theories/channel/multi_proto.v b/multi_actris/channel/proto.v similarity index 99% rename from theories/channel/multi_proto.v rename to multi_actris/channel/proto.v index e178523..2575c7a 100644 --- a/theories/channel/multi_proto.v +++ b/multi_actris/channel/proto.v @@ -50,7 +50,7 @@ From iris.proofmode Require Import proofmode. From iris.base_logic Require Export lib.iprop. From iris.base_logic Require Import lib.own. From iris.program_logic Require Import language. -From actris.channel Require Import multi_proto_model. +From multi_actris.channel Require Import proto_model. Set Default Proof Using "Type". Export action. diff --git a/theories/channel/multi_proto_consistency_examples.v b/multi_actris/channel/proto_consistency_examples.v similarity index 99% rename from theories/channel/multi_proto_consistency_examples.v rename to multi_actris/channel/proto_consistency_examples.v index a52088b..14771d2 100644 --- a/theories/channel/multi_proto_consistency_examples.v +++ b/multi_actris/channel/proto_consistency_examples.v @@ -2,9 +2,8 @@ From iris.algebra Require Import gmap excl_auth gmap_view. From iris.proofmode Require Import proofmode. From iris.base_logic Require Export lib.iprop. From iris.base_logic Require Import lib.own. -From actris.channel Require Import multi_proto_model multi_proto multi_channel multi_proofmode. +From multi_actris.channel Require Import proofmode. Set Default Proof Using "Type". -Export action. Lemma iProto_consistent_equiv_proof {Σ} (ps : gmap nat (iProto Σ)) : (∀ i j, valid_target ps i j) ∗ diff --git a/theories/channel/multi_proto_model.v b/multi_actris/channel/proto_model.v similarity index 99% rename from theories/channel/multi_proto_model.v rename to multi_actris/channel/proto_model.v index e32bc26..8bd4bd3 100644 --- a/theories/channel/multi_proto_model.v +++ b/multi_actris/channel/proto_model.v @@ -35,7 +35,7 @@ The defined functions on the type [proto] are: all terminations [END] in [p1] with [p2]. *) From iris.base_logic Require Import base_logic. From iris.proofmode Require Import proofmode. -From actris.utils Require Import cofe_solver_2. +From multi_actris.utils Require Import cofe_solver_2. Set Default Proof Using "Type". Module Export action. diff --git a/multi_actris/utils/cofe_solver_2.v b/multi_actris/utils/cofe_solver_2.v new file mode 100644 index 0000000..7a2425b --- /dev/null +++ b/multi_actris/utils/cofe_solver_2.v @@ -0,0 +1,88 @@ +From iris.algebra Require Import cofe_solver. + +(** Version of the cofe_solver for a parametrized functor. Generalize and move +to Iris. *) +Record solution_2 (F : ofe → oFunctor) := Solution2 { + solution_2_car : ∀ An `{!Cofe An} A `{!Cofe A}, ofe; + solution_2_cofe `{!Cofe An, !Cofe A} : Cofe (solution_2_car An A); + solution_2_iso `{!Cofe An, !Cofe A} :> + ofe_iso (oFunctor_apply (F A) (solution_2_car A An)) (solution_2_car An A); +}. +Arguments solution_2_car {F}. +Global Existing Instance solution_2_cofe. + +Section cofe_solver_2. + Context (F : ofe → oFunctor). + Context `{Fcontr : !∀ T, oFunctorContractive (F T)}. + Context `{Fcofe : !∀ `{!Cofe T, Cofe Bn, !Cofe B}, Cofe (oFunctor_car (F T) Bn B)}. + Context `{Finh : !∀ `{!Cofe T, Cofe Bn, !Cofe B}, Inhabited (oFunctor_car (F T) Bn B)}. + Notation map := (oFunctor_map (F _)). + + Definition F_2 (An : ofe) `{!Cofe An} (A : ofe) `{!Cofe A} : oFunctor := + oFunctor_oFunctor_compose (F A) (F An). + + Definition T_result `{!Cofe An, !Cofe A} : solution (F_2 An A) := solver.result _. + Definition T (An : ofe) `{!Cofe An} (A : ofe) `{!Cofe A} : ofe := + T_result (An:=An) (A:=A). + Instance T_cofe `{!Cofe An, !Cofe A} : Cofe (T An A) := _. + Instance T_inhabited `{!Cofe An, !Cofe A} : Inhabited (T An A) := + populate (ofe_iso_1 T_result inhabitant). + + Definition T_iso_fun_aux `{!Cofe An, !Cofe A} + (rec : prodO (oFunctor_apply (F An) (T An A) -n> T A An) + (T A An -n> oFunctor_apply (F An) (T An A))) : + prodO (oFunctor_apply (F A) (T A An) -n> T An A) + (T An A -n> oFunctor_apply (F A) (T A An)) := + (ofe_iso_1 T_result â—Ž map (rec.1,rec.2), map (rec.2,rec.1) â—Ž ofe_iso_2 T_result). + Instance T_iso_aux_fun_contractive `{!Cofe An, !Cofe A} : + Contractive (T_iso_fun_aux (An:=An) (A:=A)). + Proof. solve_contractive. Qed. + + Definition T_iso_fun_aux_2 `{!Cofe An, !Cofe A} + (rec : prodO (oFunctor_apply (F A) (T A An) -n> T An A) + (T An A -n> oFunctor_apply (F A) (T A An))) : + prodO (oFunctor_apply (F A) (T A An) -n> T An A) + (T An A -n> oFunctor_apply (F A) (T A An)) := + T_iso_fun_aux (T_iso_fun_aux rec). + Instance T_iso_fun_aux_2_contractive `{!Cofe An, !Cofe A} : + Contractive (T_iso_fun_aux_2 (An:=An) (A:=A)). + Proof. + intros n rec1 rec2 Hrec. rewrite /T_iso_fun_aux_2. + f_equiv. by apply T_iso_aux_fun_contractive. + Qed. + + Definition T_iso_fun `{!Cofe An, !Cofe A} : + prodO (oFunctor_apply (F A) (T A An) -n> T An A) + (T An A -n> oFunctor_apply (F A) (T A An)) := + fixpoint T_iso_fun_aux_2. + Definition T_iso_fun_unfold_1 `{!Cofe An, !Cofe A} y : + T_iso_fun (An:=An) (A:=A).1 y ≡ (T_iso_fun_aux (T_iso_fun_aux T_iso_fun)).1 y. + Proof. apply (fixpoint_unfold T_iso_fun_aux_2). Qed. + Definition T_iso_fun_unfold_2 `{!Cofe An, !Cofe A} y : + T_iso_fun (An:=An) (A:=A).2 y ≡ (T_iso_fun_aux (T_iso_fun_aux T_iso_fun)).2 y. + Proof. apply (fixpoint_unfold T_iso_fun_aux_2). Qed. + + Lemma result_2 : solution_2 F. + Proof using Fcontr Fcofe Finh. + apply (Solution2 F T _)=> An Hcn A Hc. + apply (OfeIso (T_iso_fun.1) (T_iso_fun.2)). + - intros y. apply equiv_dist=> n. revert An Hcn A Hc y. + induction (lt_wf n) as [n _ IH]; intros An ? A ? y. + rewrite T_iso_fun_unfold_1 T_iso_fun_unfold_2 /=. + rewrite -{2}(ofe_iso_12 T_result y). f_equiv. + rewrite -(oFunctor_map_id (F _) (ofe_iso_2 T_result y)) + -!oFunctor_map_compose. + f_equiv; apply Fcontr; dist_later_intro as n' Hn'; split=> x /=; + rewrite ofe_iso_21 -{2}(oFunctor_map_id (F _) x) + -!oFunctor_map_compose; do 2 f_equiv; split=> z /=; auto. + - intros y. apply equiv_dist=> n. revert An Hcn A Hc y. + induction (lt_wf n) as [n _ IH]; intros An ? A ? y. + rewrite T_iso_fun_unfold_1 T_iso_fun_unfold_2 /= ofe_iso_21. + rewrite -(oFunctor_map_id (F _) y) -!oFunctor_map_compose. + f_equiv; apply Fcontr; dist_later_intro as n' Hn'; split=> x /=; + rewrite -{2}(ofe_iso_12 T_result x); f_equiv; + rewrite -{2}(oFunctor_map_id (F _) (ofe_iso_2 T_result x)) + -!oFunctor_map_compose; + do 2 f_equiv; split=> z /=; auto. + Qed. +End cofe_solver_2. diff --git a/theories/channel/channel.v b/theories/channel/channel.v index 1dbb8d8..e8d038e 100644 --- a/theories/channel/channel.v +++ b/theories/channel/channel.v @@ -21,85 +21,88 @@ In this file we define the three message-passing connectives: It is additionaly shown that the channel ownership [c ↣ prot] is closed under the subprotocol relation [⊑] *) -From iris.algebra Require Import gmap excl_auth gmap_view. From iris.heap_lang Require Export primitive_laws notation. From iris.heap_lang Require Export proofmode. From iris.heap_lang.lib Require Import spin_lock. -From actris.utils Require Export llist. -From actris.channel Require Import proto_model. From actris.channel Require Export proto. +From actris.utils Require Import llist skip. Set Default Proof Using "Type". +Local Existing Instance spin_lock. + (** * The definition of the message-passing connectives *) Definition new_chan : val := - λ: <>, let: "l1" := ref NONEV in - let: "l2" := ref NONEV in - (("l1","l2"), ("l2","l1")). + λ: <>, + let: "l" := lnil #() in + let: "r" := lnil #() in + let: "lk" := newlock #() in + ((("l","r"),"lk"), (("r","l"),"lk")). Definition fork_chan : val := λ: "f", let: "cc" := new_chan #() in Fork ("f" (Snd "cc"));; Fst "cc". -Definition wait : val := - rec: "go" "l" := - match: !"l" with - NONE => #() - | SOME <> => "go" "l" - end. - Definition send : val := λ: "c" "v", - let: "l" := Fst "c" in - "l" <- SOME "v";; wait "l". + let: "lk" := Snd "c" in + let: "l" := Fst (Fst "c") in + let: "r" := Snd (Fst "c") in + acquire "lk";; + lsnoc "l" "v";; + release "lk". -(* Definition recv : val := *) -(* rec: "go" "c" "i" := *) -(* let: "l" := Snd (llookup "c" "i") in *) -(* match: !"l" with *) -(* NONE => "go" "c" *) -(* | SOME "v" => "c" <- NONE;; "v" *) -(* end. *) +Definition try_recv : val := + λ: "c", + let: "lk" := Snd "c" in + acquire "lk";; + let: "l" := Snd (Fst "c") in + let: "ret" := if: lisnil "l" then NONE else SOME (lpop "l") in + release "lk";; "ret". Definition recv : val := rec: "go" "c" := - let: "l" := Snd "c" in - let: "v" := Xchg "l" NONEV in - match: "v" with - NONE => "go" "c" - | SOME "v" => "v" + match: try_recv "c" with + SOME "p" => "p" + | NONE => "go" "c" end. +(** * Setup of Iris's cameras *) Class chanG Σ := { - chanG_tokG :: inG Σ (exclR unitO); + chanG_lockG :: lockG Σ; chanG_protoG :: protoG Σ val; }. -Definition chanΣ : gFunctors := #[ protoΣ val; GFunctor (exclR unitO) ]. +Definition chanΣ : gFunctors := #[ spin_lockΣ; protoΣ val ]. Global Instance subG_chanΣ {Σ} : subG chanΣ Σ → chanG Σ. Proof. solve_inG. Qed. +Record chan_name := ChanName { + chan_lock_name : gname; + chan_proto_name : proto_name; +}. +Global Instance chan_name_inhabited : Inhabited chan_name := + populate (ChanName inhabitant inhabitant). +Global Instance chan_name_eq_dec : EqDecision chan_name. +Proof. solve_decision. Qed. +Global Instance chan_name_countable : Countable chan_name. +Proof. + refine (inj_countable (λ '(ChanName γl γr), (γl,γr)) + (λ '(γl, γr), Some (ChanName γl γr)) _); by intros []. +Qed. + (** * Definition of the pointsto connective *) Notation iProto Σ := (iProto Σ val). Notation iMsg Σ := (iMsg Σ val). -Definition tok `{!chanG Σ} (γ : gname) : iProp Σ := own γ (Excl ()). - -Definition chan_inv `{!heapGS Σ, !chanG Σ} γ γE γt s (l:loc) : iProp Σ := - (l ↦ NONEV ∗ tok γt) ∨ - (∃ v m, l ↦ SOMEV v ∗ - iProto_own γ s (<!> m)%proto ∗ - (∃ p, iMsg_car m v (Next p) ∗ own γE (â—E (Next p)))) ∨ - (∃ p, l ↦ NONEV ∗ - iProto_own γ s p ∗ own γE (â—E (Next p))). - Definition iProto_pointsto_def `{!heapGS Σ, !chanG Σ} (c : val) (p : iProto Σ) : iProp Σ := - ∃ γ γE1 γE2 γt1 γt2 s (l1 l2:loc), - ⌜ c = PairV #l1 #l2 ⌠∗ - inv (nroot.@"ctx") (iProto_ctx γ) ∗ - inv (nroot.@"p") (chan_inv γ γE1 γt1 s l1) ∗ - inv (nroot.@"p") (chan_inv γ γE2 γt2 (side_dual s) l2) ∗ - own γE1 (â—E (Next p)) ∗ own γE1 (â—¯E (Next p)) ∗ - iProto_own γ s p. + ∃ γ s (l r : loc) (lk : val), + ⌜ c = ((#(side_elim s l r), #(side_elim s r l)), lk)%V ⌠∗ + is_lock (chan_lock_name γ) lk (∃ vsl vsr, + llist internal_eq l vsl ∗ + llist internal_eq r vsr ∗ + steps_lb (length vsl) ∗ steps_lb (length vsr) ∗ + iProto_ctx (chan_proto_name γ) vsl vsr) ∗ + iProto_own (chan_proto_name γ) s p. Definition iProto_pointsto_aux : seal (@iProto_pointsto_def). by eexists. Qed. Definition iProto_pointsto := iProto_pointsto_aux.(unseal). Definition iProto_pointsto_eq : @@ -109,6 +112,25 @@ Global Instance: Params (@iProto_pointsto) 4 := {}. Notation "c ↣ p" := (iProto_pointsto c p) (at level 20, format "c ↣ p"). +Global Instance iProto_pointsto_contractive `{!heapGS Σ, !chanG Σ} c : + Contractive (iProto_pointsto c). +Proof. rewrite iProto_pointsto_eq. solve_contractive. Qed. + +Definition iProto_choice {Σ} (a : action) (P1 P2 : iProp Σ) + (p1 p2 : iProto Σ) : iProto Σ := + (<a @ (b : bool)> MSG #b {{ if b then P1 else P2 }}; if b then p1 else p2)%proto. +Global Typeclasses Opaque iProto_choice. +Arguments iProto_choice {_} _ _%I _%I _%proto _%proto. +Global Instance: Params (@iProto_choice) 2 := {}. +Infix "<{ P1 }+{ P2 }>" := (iProto_choice Send P1 P2) (at level 85) : proto_scope. +Infix "<{ P1 }&{ P2 }>" := (iProto_choice Recv P1 P2) (at level 85) : proto_scope. +Infix "<+{ P2 }>" := (iProto_choice Send True P2) (at level 85) : proto_scope. +Infix "<&{ P2 }>" := (iProto_choice Recv True P2) (at level 85) : proto_scope. +Infix "<{ P1 }+>" := (iProto_choice Send P1 True) (at level 85) : proto_scope. +Infix "<{ P1 }&>" := (iProto_choice Recv P1 True) (at level 85) : proto_scope. +Infix "<+>" := (iProto_choice Send True True) (at level 85) : proto_scope. +Infix "<&>" := (iProto_choice Recv True True) (at level 85) : proto_scope. + Section channel. Context `{!heapGS Σ, !chanG Σ}. Implicit Types p : iProto Σ. @@ -119,12 +141,64 @@ Section channel. Global Instance iProto_pointsto_proper c : Proper ((≡) ==> (≡)) (iProto_pointsto c). Proof. apply (ne_proper _). Qed. - (* Lemma iProto_pointsto_le c p1 p2 : c ↣ p1 ⊢ â–· (p1 ⊑ p2) -∗ c ↣ p2. *) - (* Proof. *) - (* rewrite iProto_pointsto_eq. iDestruct 1 as (γ s l r lk ->) "[Hlk H]". *) - (* iIntros "Hle'". iExists γ, s, l, r, lk. iSplit; [done|]. iFrame "Hlk". *) - (* by iApply (iProto_own_le with "H"). *) - (* Qed. *) + Lemma iProto_pointsto_le c p1 p2 : c ↣ p1 ⊢ â–· (p1 ⊑ p2) -∗ c ↣ p2. + Proof. + rewrite iProto_pointsto_eq. iDestruct 1 as (γ s l r lk ->) "[Hlk H]". + iIntros "Hle'". iExists γ, s, l, r, lk. iSplit; [done|]. iFrame "Hlk". + by iApply (iProto_own_le with "H"). + Qed. + + Global Instance iProto_choice_contractive n a : + Proper (dist n ==> dist n ==> + dist_later n ==> dist_later n ==> dist n) (@iProto_choice Σ a). + Proof. solve_contractive. Qed. + Global Instance iProto_choice_ne n a : + Proper (dist n ==> dist n ==> dist n ==> dist n ==> dist n) (@iProto_choice Σ a). + Proof. solve_proper. Qed. + Global Instance iProto_choice_proper a : + Proper ((≡) ==> (≡) ==> (≡) ==> (≡) ==> (≡)) (@iProto_choice Σ a). + Proof. solve_proper. Qed. + + Lemma iProto_choice_equiv a1 a2 (P11 P12 P21 P22 : iProp Σ) + (p11 p12 p21 p22 : iProto Σ) : + ⌜a1 = a2⌠-∗ ((P11 ≡ P12):iProp Σ) -∗ (P21 ≡ P22) -∗ + â–· (p11 ≡ p12) -∗ â–· (p21 ≡ p22) -∗ + iProto_choice a1 P11 P21 p11 p21 ≡ iProto_choice a2 P12 P22 p12 p22. + Proof. + iIntros (->) "#HP1 #HP2 #Hp1 #Hp2". + rewrite /iProto_choice. iApply iProto_message_equiv; [ eauto | | ]. + - iIntros "!>" (b) "H". iExists b. iSplit; [ done | ]. + destruct b; + [ iRewrite -"HP1"; iFrame "H Hp1" | iRewrite -"HP2"; iFrame "H Hp2" ]. + - iIntros "!>" (b) "H". iExists b. iSplit; [ done | ]. + destruct b; + [ iRewrite "HP1"; iFrame "H Hp1" | iRewrite "HP2"; iFrame "H Hp2" ]. + Qed. + + Lemma iProto_dual_choice a P1 P2 p1 p2 : + iProto_dual (iProto_choice a P1 P2 p1 p2) + ≡ iProto_choice (action_dual a) P1 P2 (iProto_dual p1) (iProto_dual p2). + Proof. + rewrite /iProto_choice iProto_dual_message /= iMsg_dual_exist. + f_equiv; f_equiv=> -[]; by rewrite iMsg_dual_base. + Qed. + + Lemma iProto_app_choice a P1 P2 p1 p2 q : + (iProto_choice a P1 P2 p1 p2 <++> q)%proto + ≡ (iProto_choice a P1 P2 (p1 <++> q) (p2 <++> q))%proto. + Proof. + rewrite /iProto_choice iProto_app_message /= iMsg_app_exist. + f_equiv; f_equiv=> -[]; by rewrite iMsg_app_base. + Qed. + + Lemma iProto_le_choice a P1 P2 p1 p2 p1' p2' : + (P1 -∗ P1 ∗ â–· (p1 ⊑ p1')) ∧ (P2 -∗ P2 ∗ â–· (p2 ⊑ p2')) -∗ + iProto_choice a P1 P2 p1 p2 ⊑ iProto_choice a P1 P2 p1' p2'. + Proof. + iIntros "H". rewrite /iProto_choice. destruct a; + iIntros (b) "HP"; iExists b; destruct b; + iDestruct ("H" with "HP") as "[$ ?]"; by iModIntro. + Qed. (** ** Specifications of [send] and [recv] *) Lemma new_chan_spec p : @@ -132,36 +206,19 @@ Section channel. new_chan #() {{{ c1 c2, RET (c1,c2); c1 ↣ p ∗ c2 ↣ iProto_dual p }}}. Proof. - iIntros (Φ _) "HΦ". wp_lam. - wp_alloc l1 as "Hl1". - wp_alloc l2 as "Hl2". + iIntros (Φ _) "HΦ". wp_lam. iMod (steps_lb_0) as "#Hlb". + wp_smart_apply (lnil_spec internal_eq with "[//]"); iIntros (l) "Hl". + wp_smart_apply (lnil_spec internal_eq with "[//]"); iIntros (r) "Hr". iMod (iProto_init p) as (γp) "(Hctx & Hcl & Hcr)". - iMod (own_alloc (â—E (Next p) â‹… â—¯E (Next p))) as (γl) "[Hâ—l Hâ—¯l]". - { by apply excl_auth_valid. } - iMod (own_alloc (â—E (Next (iProto_dual p)) â‹… â—¯E (Next (iProto_dual p)))) as (γr) "[Hâ—r Hâ—¯r]". - { by apply excl_auth_valid. } - iMod (own_alloc (Excl ())) as (γtl) "Htokl". - { done. } - iMod (own_alloc (Excl ())) as (γtr) "Htokr". - { done. } - wp_pures. - iMod (inv_alloc _ ⊤ (iProto_ctx γp) with "[Hctx]") - as "#IH". - { done. } - iMod (inv_alloc _ ⊤ (chan_inv γp γl γtl Left l1) with "[Hl1 Htokl]") - as "#IHl". - { iLeft. iFrame. } - iMod (inv_alloc _ ⊤ (chan_inv γp γr γtr Right l2) with "[Hl2 Htokr]") - as "#IHr". - { iLeft. iFrame. } - iModIntro. - iApply "HΦ". - iSplitL "Hcl Hâ—l Hâ—¯l". - - rewrite iProto_pointsto_eq. - iExists _, _, _, _, _, _, _, _. eauto with iFrame. - - rewrite iProto_pointsto_eq. - iExists _, _, _, _, _, _, _, _. - iSplit; [done|]. iFrame "#∗". + wp_smart_apply (newlock_spec (∃ vsl vsr, + llist internal_eq l vsl ∗ llist internal_eq r vsr ∗ + steps_lb (length vsl) ∗ steps_lb (length vsr) ∗ + iProto_ctx γp vsl vsr) with "[Hl Hr Hctx]"). + { iExists [], []. iFrame "#∗". } + iIntros (lk γlk) "#Hlk". wp_pures. iApply "HΦ". + set (γ := ChanName γlk γp). iSplitL "Hcl". + - rewrite iProto_pointsto_eq. iExists γ, Left, l, r, lk. by iFrame "Hcl #". + - rewrite iProto_pointsto_eq. iExists γ, Right, l, r, lk. by iFrame "Hcr #". Qed. Lemma fork_chan_spec p Φ (f : val) : @@ -176,159 +233,149 @@ Section channel. wp_pures. iApply ("HΦ" with "Hc1"). Qed. - Lemma own_prot_excl γ (p1 p2 : iProto Σ) : - own γ (â—¯E (Next p1)) -∗ own γ (â—¯E (Next p2)) -∗ False. - Proof. Admitted. - Lemma send_spec c v p : {{{ c ↣ <!> MSG v; p }}} send c v {{{ RET #(); c ↣ p }}}. Proof. rewrite iProto_pointsto_eq. iIntros (Φ) "Hc HΦ". wp_lam; wp_pures. - iDestruct "Hc" as (γ γE1 γE2 γt1 γt2 s l1 l2 ->) - "(#IH & #IHl & #IHr & Hâ— & Hâ—¯ & Hown)". - wp_pures. - wp_bind (Store _ _). - iInv "IHl" as "HIp". - iDestruct "HIp" as "[HIp|HIp]"; last first. - { iDestruct "HIp" as "[HIp|HIp]". - - iDestruct "HIp" as (? m) "(>Hl & Hown' & HIp)". - wp_store. - rewrite /iProto_own. - iDestruct "Hown" as (p') "[_ Hown]". - iDestruct "Hown'" as (p'') "[_ Hown']". - iDestruct (own_prot_excl with "Hown Hown'") as "H". done. - - iDestruct "HIp" as (p') "(>Hl & Hown' & HIp)". - wp_store. - rewrite /iProto_own. - iDestruct "Hown" as (p'') "[_ Hown]". - iDestruct "Hown'" as (p''') "[_ Hown']". - iDestruct (own_prot_excl with "Hown Hown'") as "H". done. } - iDestruct "HIp" as "[>Hl Htok]". - wp_store. - iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]". - { apply excl_auth_update. } - iModIntro. - iSplitL "Hl Hâ— Hown". - { iRight. iLeft. iIntros "!>". iExists _, _. iFrame. - iExists _. iFrame. rewrite iMsg_base_eq. simpl. done. } - wp_pures. - iLöb as "HL". - wp_lam. - wp_bind (Load _). - iInv "IHl" as "HIp". - iDestruct "HIp" as "[HIp|HIp]". - { iDestruct "HIp" as ">[Hl Htok']". - iDestruct (own_valid_2 with "Htok Htok'") as %H. done. } - iDestruct "HIp" as "[HIp|HIp]". - - iDestruct "HIp" as (? m) "(>Hl & Hown & HIp)". - wp_load. iModIntro. - iSplitL "Hl Hown HIp". - { iRight. iLeft. iExists _, _. iFrame. } - wp_pures. iApply ("HL" with "HΦ Htok Hâ—¯"). - - iDestruct "HIp" as (p') "(>Hl & Hown & Hâ—)". - wp_load. - iModIntro. - iSplitL "Hl Htok". - { iLeft. iFrame. } - iDestruct (own_valid_2 with "Hâ— Hâ—¯") as "#Hagree". - iDestruct (excl_auth_agreeI with "Hagree") as "Hagree'". - wp_pures. - iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]". - { apply excl_auth_update. } - iModIntro. - iApply "HΦ". - iExists _, _, _, _, _, _, _, _. - iSplit; [done|]. iFrame "#∗". - iRewrite -"Hagree'". done. + iDestruct "Hc" as (γ s l r lk ->) "[#Hlk H]"; wp_pures. + wp_smart_apply (acquire_spec with "Hlk"); iIntros "[Hlkd Hinv]". + iDestruct "Hinv" as (vsl vsr) "(Hl & Hr & #Hlbl & #Hlbr & Hctx)". + destruct s; simpl. + - wp_pures. wp_bind (lsnoc _ _). + iApply (wp_step_fupdN_lb with "Hlbr [Hctx H]"); [done| |]. + { iApply fupd_mask_intro; [set_solver|]. simpl. + iIntros "Hclose !>!>". + iMod (iProto_send_l with "Hctx H []") as "[Hctx H]". + { rewrite iMsg_base_eq /=; auto. } + iModIntro. + iApply step_fupdN_intro; [done|]. + iIntros "!>". iMod "Hclose". + iCombine ("Hctx H") as "H". + iExact "H". } + iApply (wp_lb_update with "Hlbl"). + wp_smart_apply (lsnoc_spec with "[$Hl //]"); iIntros "Hl". + iIntros "#Hlbl' [Hctx H] !>". + wp_smart_apply (release_spec with "[Hl Hr Hctx $Hlk $Hlkd]"). + { iExists (vsl ++ [v]), vsr. + rewrite app_length /=. + replace (length vsl + 1) with (S (length vsl)) by lia. + iFrame "#∗". } + iIntros "_". iApply "HΦ". iExists γ, Left, l, r, lk. eauto 10 with iFrame. + - wp_pures. wp_bind (lsnoc _ _). + iApply (wp_step_fupdN_lb with "Hlbl [Hctx H]"); [done| |]. + { iApply fupd_mask_intro; [set_solver|]. simpl. + iIntros "Hclose !>!>". + iMod (iProto_send_r with "Hctx H []") as "[Hctx H]". + { rewrite iMsg_base_eq /=; auto. } + iModIntro. + iApply step_fupdN_intro; [done|]. + iIntros "!>". iMod "Hclose". + iCombine ("Hctx H") as "H". + iExact "H". } + iApply (wp_lb_update with "Hlbr"). + wp_smart_apply (lsnoc_spec with "[$Hr //]"); iIntros "Hr". + iIntros "#Hlbr' [Hctx H] !>". + wp_smart_apply (release_spec with "[Hl Hr Hctx $Hlk $Hlkd]"). + { iExists vsl, (vsr ++ [v]). + rewrite app_length /=. + replace (length vsr + 1) with (S (length vsr)) by lia. + iFrame "#∗". } + iIntros "_". iApply "HΦ". iExists γ, Right, l, r, lk. eauto 10 with iFrame. + Qed. + + Lemma send_spec_tele {TT} c (tt : TT) + (v : TT → val) (P : TT → iProp Σ) (p : TT → iProto Σ) : + {{{ c ↣ (<!.. x > MSG v x {{ P x }}; p x) ∗ P tt }}} + send c (v tt) + {{{ RET #(); c ↣ (p tt) }}}. + Proof. + iIntros (Φ) "[Hc HP] HΦ". + iDestruct (iProto_pointsto_le _ _ (<!> MSG v tt; p tt)%proto with "Hc [HP]") + as "Hc". + { iIntros "!>". + iApply iProto_le_trans. + iApply iProto_le_texist_intro_l. + by iFrame "HP". } + by iApply (send_spec with "Hc"). Qed. - (* Lemma send_spec_tele {TT} c (tt : TT) *) - (* (v : TT → val) (P : TT → iProp Σ) (p : TT → iProto Σ) : *) - (* {{{ c ↣ (<!.. x > MSG v x {{ P x }}; p x) ∗ P tt }}} *) - (* send c (v tt) *) - (* {{{ RET #(); c ↣ (p tt) }}}. *) - (* Proof. *) - (* iIntros (Φ) "[Hc HP] HΦ". *) - (* iDestruct (iProto_pointsto_le _ _ (<!> MSG v tt; p tt)%proto with "Hc [HP]") *) - (* as "Hc". *) - (* { iIntros "!>". *) - (* iApply iProto_le_trans. *) - (* iApply iProto_le_texist_intro_l. *) - (* by iFrame "HP". } *) - (* by iApply (send_spec with "Hc"). *) - (* Qed. *) + Lemma try_recv_spec {TT} c (v : TT → val) (P : TT → iProp Σ) (p : TT → iProto Σ) : + {{{ c ↣ <?.. x> MSG v x {{ P x }}; p x }}} + try_recv c + {{{ w, RET w; (⌜w = NONEV⌠∗ c ↣ <?.. x> MSG v x {{ P x }}; p x) ∨ + (∃.. x, ⌜w = SOMEV (v x)⌠∗ c ↣ p x ∗ P x) }}}. + Proof. + rewrite iProto_pointsto_eq. iIntros (Φ) "Hc HΦ". wp_lam; wp_pures. + iDestruct "Hc" as (γ s l r lk ->) "[#Hlk H]"; wp_pures. + wp_smart_apply (acquire_spec with "Hlk"); iIntros "[Hlkd Hinv]". + iDestruct "Hinv" as (vsl vsr) "(Hl & Hr & #Hlbl & #Hlbr & Hctx)". destruct s; simpl. + - wp_smart_apply (lisnil_spec with "Hr"); iIntros "Hr". + destruct vsr as [|vr vsr]; wp_pures. + { wp_smart_apply (release_spec with "[Hl Hr Hctx $Hlk $Hlkd]"); [by eauto with iFrame|]. + iIntros "_". wp_pures. iModIntro. iApply "HΦ". iLeft. iSplit; [done|]. + iExists γ, Left, l, r, lk. eauto 10 with iFrame. } + wp_smart_apply (lpop_spec with "Hr"); iIntros (v') "[% Hr]"; simplify_eq/=. + iMod (iProto_recv_l with "Hctx H") as (q) "(Hctx & H & Hm)". wp_pures. + rewrite iMsg_base_eq. + iDestruct (iMsg_texist_exist with "Hm") as (x <-) "[Hp HP]". + iDestruct (steps_lb_le _ (length vsr) with "Hlbr") as "#Hlbr'"; [lia|]. + wp_smart_apply (release_spec with "[Hl Hr Hctx $Hlk $Hlkd]"); [by eauto with iFrame|]. + iIntros "_". wp_pures. iModIntro. iApply "HΦ". iRight. iExists x. iSplit; [done|]. + iFrame "HP". iExists γ, Left, l, r, lk. iSplit; [done|]. iFrame "Hlk". + by iRewrite "Hp". + - wp_smart_apply (lisnil_spec with "Hl"); iIntros "Hl". + destruct vsl as [|vl vsl]; wp_pures. + { wp_smart_apply (release_spec with "[Hl Hr Hctx $Hlk $Hlkd]"); [by eauto with iFrame|]. + iIntros "_". wp_pures. iModIntro. iApply "HΦ". iLeft. iSplit; [done|]. + iExists γ, Right, l, r, lk. eauto 10 with iFrame. } + wp_smart_apply (lpop_spec with "Hl"); iIntros (v') "[% Hl]"; simplify_eq/=. + iMod (iProto_recv_r with "Hctx H") as (q) "(Hctx & H & Hm)". wp_pures. + rewrite iMsg_base_eq. + iDestruct (iMsg_texist_exist with "Hm") as (x <-) "[Hp HP]". + iDestruct (steps_lb_le _ (length vsl) with "Hlbl") as "#Hlbl'"; [lia|]. + wp_smart_apply (release_spec with "[Hl Hr Hctx $Hlk $Hlkd]"); [by eauto with iFrame|]. + iIntros "_". wp_pures. iModIntro. iApply "HΦ". iRight. iExists x. iSplit; [done|]. + iFrame "HP". iExists γ, Right, l, r, lk. iSplit; [done|]. iFrame "Hlk". + by iRewrite "Hp". + Qed. Lemma recv_spec {TT} c (v : TT → val) (P : TT → iProp Σ) (p : TT → iProto Σ) : {{{ c ↣ <?.. x> MSG v x {{ â–· P x }}; p x }}} recv c {{{ x, RET v x; c ↣ p x ∗ P x }}}. Proof. - iIntros (Φ) "Hc HΦ". iLöb as "HL". wp_lam. - rewrite iProto_pointsto_eq. - iDestruct "Hc" as (γ E1 γE2 γt1 γt2 s l1 l2 ->) - "(#IH & #IHl & #IHr & Hâ— & Hâ—¯ & Hown)". - wp_pures. - wp_bind (Xchg _ _). - iInv "IHr" as "HIp". - iDestruct "HIp" as "[HIp|HIp]". - { iDestruct "HIp" as ">[Hl Htok]". - wp_xchg. - iModIntro. - iSplitL "Hl Htok". - { iLeft. iFrame. } - wp_pures. iApply ("HL" with "[Hâ— Hâ—¯ Hown] HΦ"). - iExists _, _, _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". } - iDestruct "HIp" as "[HIp|HIp]"; last first. - { iDestruct "HIp" as (p') "[>Hl [Hown' Hâ—¯']]". - wp_xchg. - iModIntro. - iSplitL "Hl Hown' Hâ—¯'". - { iRight. iRight. iExists _. iFrame. } - wp_pures. iApply ("HL" with "[Hâ— Hâ—¯ Hown] HΦ"). - iExists _, _, _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". } - iDestruct "HIp" as (w m) "(>Hl & Hown' & HIp)". - iDestruct "HIp" as (p') "[Hm Hp']". - iInv "IH" as "Hctx". - wp_xchg. - destruct s. - - simpl. - iMod (iProto_step_r with "Hctx Hown Hown' Hm") as - (p'') "(Hm & Hctx & Hown & Hown')". - iModIntro. - iSplitL "Hctx"; [done|]. - iModIntro. - iSplitL "Hl Hown' Hp'". - { iRight. iRight. iExists _. iFrame. } - wp_pure _. - rewrite iMsg_base_eq. - iDestruct (iMsg_texist_exist with "Hm") as (x <-) "[Hp HP]". - wp_pures. - iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]". - { apply excl_auth_update. } - iModIntro. iApply "HΦ". - iFrame. - iExists _, _, _, _, _, _, _, _. iSplit; [done|]. - iRewrite "Hp". iFrame "#∗". done. - - simpl. - iMod (iProto_step_l with "Hctx Hown' Hown Hm") as - (p'') "(Hm & Hctx & Hown & Hown')". - iModIntro. - iSplitL "Hctx"; [done|]. - iModIntro. - iSplitL "Hl Hown Hp'". - { iRight. iRight. iExists _. iFrame. } - wp_pure _. - rewrite iMsg_base_eq. - iDestruct (iMsg_texist_exist with "Hm") as (x <-) "[Hp HP]". - wp_pures. - iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]". - { apply excl_auth_update. } - iModIntro. iApply "HΦ". - iFrame. - iExists _, _, _, _, _, _, _, _. iSplit; [done|]. - iRewrite "Hp". iFrame "#∗". done. + iIntros (Φ) "Hc HΦ". iLöb as "IH". wp_lam. + wp_smart_apply (try_recv_spec with "Hc"); iIntros (w) "[[-> H]|H]". + { wp_pures. by iApply ("IH" with "[$]"). } + iDestruct "H" as (x ->) "[Hc HP]". wp_pures. iApply "HΦ". by iFrame. + Qed. + + (** ** Specifications for choice *) + Lemma select_spec c (b : bool) P1 P2 p1 p2 : + {{{ c ↣ (p1 <{P1}+{P2}> p2) ∗ if b then P1 else P2 }}} + send c #b + {{{ RET #(); c ↣ (if b then p1 else p2) }}}. + Proof. + rewrite /iProto_choice. iIntros (Φ) "[Hc HP] HΦ". + iApply (send_spec with "[Hc HP] HΦ"). + iApply (iProto_pointsto_le with "Hc"). + iIntros "!>". iExists b. by iFrame "HP". Qed. + Lemma branch_spec c P1 P2 p1 p2 : + {{{ c ↣ (p1 <{P1}&{P2}> p2) }}} + recv c + {{{ b, RET #b; c ↣ (if b : bool then p1 else p2) ∗ if b then P1 else P2 }}}. + Proof. + rewrite /iProto_choice. iIntros (Φ) "Hc HΦ". + iApply (recv_spec _ (tele_app _) + (tele_app (TT:=[tele _ : bool]) (λ b, if b then P1 else P2))%I + (tele_app _) with "[Hc]"). + { iApply (iProto_pointsto_le with "Hc"). + iIntros "!> /=" (b) "HP". iExists b. by iSplitL. } + rewrite -bi_tforall_forall. + iIntros "!>" (x) "[Hc H]". iApply "HΦ". iFrame. + Qed. End channel. diff --git a/theories/channel/proto.v b/theories/channel/proto.v index f7b54fe..92c4c00 100644 --- a/theories/channel/proto.v +++ b/theories/channel/proto.v @@ -245,53 +245,38 @@ Definition iProto_dual_if {Σ V} (d : bool) (p : iProto Σ V) : iProto Σ V := Arguments iProto_dual_if {_ _} _ _%proto. Global Instance: Params (@iProto_dual_if) 3 := {}. -Definition iProto_consistent_pre {Σ V} (rec : iProto Σ V → iProto Σ V → iProp Σ) - (pl pr : iProto Σ V) : iProp Σ := - (∀ a1 a2 m1 m2, - (pl ≡ <a1> m1) -∗ (pr ≡ <a2> m2) -∗ - match a1,a2 with - | Send,Recv => ∀ v p1, iMsg_car m1 v (Next p1) -∗ - ∃ p2, iMsg_car m2 v (Next p2) ∗ â–· (rec p1 p2) - | Recv,Send => ∀ v p2, iMsg_car m2 v (Next p2) -∗ - ∃ p1, iMsg_car m1 v (Next p1) ∗ â–· (rec p1 p2) - | _, _ => True - end). - -Global Instance iProto_consistent_pre_ne {Σ V} - (rec : iProto Σ V -n> iProto Σ V -n> iPropO Σ) : - NonExpansive2 (iProto_consistent_pre (λ p1 p2, rec p1 p2)). +(** * Protocol entailment *) +Definition iProto_le_pre {Σ V} + (rec : iProto Σ V → iProto Σ V → iProp Σ) (p1 p2 : iProto Σ V) : iProp Σ := + (p1 ≡ END ∗ p2 ≡ END) ∨ + ∃ a1 a2 m1 m2, + (p1 ≡ <a1> m1) ∗ (p2 ≡ <a2> m2) ∗ + match a1, a2 with + | Recv, Recv => ∀ v p1', + iMsg_car m1 v (Next p1') -∗ ∃ p2', â–· rec p1' p2' ∗ iMsg_car m2 v (Next p2') + | Send, Send => ∀ v p2', + iMsg_car m2 v (Next p2') -∗ ∃ p1', â–· rec p1' p2' ∗ iMsg_car m1 v (Next p1') + | Recv, Send => ∀ v1 v2 p1' p2', + iMsg_car m1 v1 (Next p1') -∗ iMsg_car m2 v2 (Next p2') -∗ ∃ pt, + â–· rec p1' (<!> MSG v2; pt) ∗ â–· rec (<?> MSG v1; pt) p2' + | Send, Recv => False + end. +Global Instance iProto_le_pre_ne {Σ V} (rec : iProto Σ V → iProto Σ V → iProp Σ) : + NonExpansive2 (iProto_le_pre rec). Proof. solve_proper. Qed. -Program Definition iProto_consistent_pre' {Σ V} - (rec : iProto Σ V -n> iProto Σ V -n> iPropO Σ) : - iProto Σ V -n> iProto Σ V -n> iPropO Σ := - λne p1 p2, iProto_consistent_pre (λ p1 p2, rec p1 p2) p1 p2. +Program Definition iProto_le_pre' {Σ V} + (rec : iProto Σ V -n> iProto Σ V -n> iPropO Σ) : + iProto Σ V -n> iProto Σ V -n> iPropO Σ := λne p1 p2, + iProto_le_pre (λ p1' p2', rec p1' p2') p1 p2. Solve Obligations with solve_proper. - -Local Instance iProto_consistent_pre_contractive {Σ V} : Contractive (@iProto_consistent_pre' Σ V). +Local Instance iProto_le_pre_contractive {Σ V} : Contractive (@iProto_le_pre' Σ V). Proof. - intros n rec1 rec2 Hrec pl pr. rewrite /iProto_consistent_pre' /iProto_consistent_pre /=. - repeat (f_contractive || f_equiv). done. done. -Qed. - -Definition iProto_consistent {Σ V} (pl pr : iProto Σ V) : iProp Σ := - fixpoint iProto_consistent_pre' pl pr. - -Arguments iProto_consistent {_ _} _%proto _%proto. -Global Instance: Params (@iProto_consistent) 2 := {}. - -Global Instance iProto_consistent_ne {Σ V} : NonExpansive2 (@iProto_consistent Σ V). -Proof. solve_proper. Qed. -Global Instance iProto_consistent_proper {Σ V} : Proper ((≡) ==> (≡) ==> (⊣⊢)) (@iProto_consistent Σ V). -Proof. solve_proper. Qed. - -Lemma iProto_consistent_unfold {Σ V} (pl pr : iProto Σ V) : - iProto_consistent pl pr ≡ (iProto_consistent_pre iProto_consistent) pl pr. -Proof. - apply: (fixpoint_unfold iProto_consistent_pre'). + intros n rec1 rec2 Hrec p1 p2. rewrite /iProto_le_pre' /iProto_le_pre /=. + by repeat (f_contractive || f_equiv). Qed. Definition iProto_le {Σ V} (p1 p2 : iProto Σ V) : iProp Σ := - ∀ p3, iProto_consistent p1 p3 -∗ iProto_consistent p2 p3. + fixpoint iProto_le_pre' p1 p2. Arguments iProto_le {_ _} _%proto _%proto. Global Instance: Params (@iProto_le) 2 := {}. Notation "p ⊑ q" := (iProto_le p q) : bi_scope. @@ -306,6 +291,8 @@ Fixpoint iProto_app_recvs {Σ V} (vs : list V) (p : iProto Σ V) : iProto Σ V : | [] => p | v :: vs => <?> MSG v; iProto_app_recvs vs p end. +Definition iProto_interp {Σ V} (vsl vsr : list V) (pl pr : iProto Σ V) : iProp Σ := + ∃ p, iProto_app_recvs vsr p ⊑ pl ∗ iProto_app_recvs vsl (iProto_dual p) ⊑ pr. Record proto_name := ProtName { proto_l_name : gname; proto_r_name : gname }. Global Instance proto_name_inhabited : Inhabited proto_name := @@ -338,11 +325,11 @@ Definition iProto_own_auth `{!protoG Σ V} (γ : proto_name) (s : side) own (side_elim s proto_l_name proto_r_name γ) (â—E (Next p)). Definition iProto_ctx `{protoG Σ V} - (γ : proto_name) : iProp Σ := + (γ : proto_name) (vsl vsr : list V) : iProp Σ := ∃ pl pr, iProto_own_auth γ Left pl ∗ iProto_own_auth γ Right pr ∗ - â–· iProto_consistent pl pr. + â–· iProto_interp vsl vsr pl pr. (** * The connective for ownership of channel ends *) Definition iProto_own `{!protoG Σ V} @@ -617,531 +604,413 @@ Section proto. iRewrite "Hp12". iIntros "!>". iRewrite "Hp1'". by iRewrite ("IH" $! p1d p2). Qed. - Lemma iProto_consistent_flip p1 p2 : - iProto_consistent p2 p1 -∗ - iProto_consistent p1 p2. - Proof. - iLöb as "IH" forall (p1 p2). - iIntros "Hprot". - rewrite iProto_consistent_unfold /iProto_consistent_pre. - rewrite iProto_consistent_unfold /iProto_consistent_pre. - iIntros (a1 a2 m1 m2) "Hm1 Hm2". - iDestruct ("Hprot" with "Hm2 Hm1") as "Hprot". - destruct a1, a2; [done| | |done]. - - iIntros (v p1') "Hm1". - iDestruct ("Hprot" with "Hm1") as (p2') "[Hprot Hrec]". - iExists p2'. iFrame. by iApply "IH". + (** ** Protocol entailment **) + Lemma iProto_le_unfold p1 p2 : iProto_le p1 p2 ≡ iProto_le_pre iProto_le p1 p2. + Proof. apply: (fixpoint_unfold iProto_le_pre'). Qed. + + Lemma iProto_le_end : ⊢ END ⊑ (END : iProto Σ V). + Proof. rewrite iProto_le_unfold. iLeft. auto 10. Qed. + + Lemma iProto_le_send m1 m2 : + (∀ v p2', iMsg_car m2 v (Next p2') -∗ ∃ p1', + â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1')) -∗ + (<!> m1) ⊑ (<!> m2). + Proof. rewrite iProto_le_unfold. iIntros "H". iRight. auto 10. Qed. + Lemma iProto_le_recv m1 m2 : + (∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', + â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2')) -∗ + (<?> m1) ⊑ (<?> m2). + Proof. rewrite iProto_le_unfold. iIntros "H". iRight. auto 10. Qed. + Lemma iProto_le_swap m1 m2 : + (∀ v1 v2 p1' p2', + iMsg_car m1 v1 (Next p1') -∗ iMsg_car m2 v2 (Next p2') -∗ ∃ pt, + â–· (p1' ⊑ <!> MSG v2; pt) ∗ â–· ((<?> MSG v1; pt) ⊑ p2')) -∗ + (<?> m1) ⊑ (<!> m2). + Proof. rewrite iProto_le_unfold. iIntros "H". iRight. auto 10. Qed. + + Lemma iProto_le_end_inv_l p : p ⊑ END -∗ (p ≡ END). + Proof. + rewrite iProto_le_unfold. iIntros "[[Hp _]|H] //". + iDestruct "H" as (a1 a2 m1 m2) "(_ & Heq & _)". + by iDestruct (iProto_end_message_equivI with "Heq") as %[]. + Qed. + + Lemma iProto_le_end_inv_r p : END ⊑ p -∗ (p ≡ END). + Proof. + rewrite iProto_le_unfold. iIntros "[[_ Hp]|H] //". + iDestruct "H" as (a1 a2 m1 m2) "(Heq & _ & _)". + iDestruct (iProto_end_message_equivI with "Heq") as %[]. + Qed. + + Lemma iProto_le_send_inv p1 m2 : + p1 ⊑ (<!> m2) -∗ ∃ a1 m1, + (p1 ≡ <a1> m1) ∗ + match a1 with + | Send => ∀ v p2', + iMsg_car m2 v (Next p2') -∗ ∃ p1', + â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1') + | Recv => ∀ v1 v2 p1' p2', + iMsg_car m1 v1 (Next p1') -∗ iMsg_car m2 v2 (Next p2') -∗ ∃ pt, + â–· (p1' ⊑ <!> MSG v2; pt) ∗ â–· ((<?> MSG v1; pt) ⊑ p2') + end. + Proof. + rewrite iProto_le_unfold. iIntros "[[_ Heq]|H]". + { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } + iDestruct "H" as (a1 a2 m1 m2') "(Hp1 & Hp2 & H)". + iExists _, _; iSplit; [done|]. destruct a1, a2. - iIntros (v p2') "Hm2". - iDestruct ("Hprot" with "Hm2") as (p1') "[Hprot Hrec]". - iExists p1'. iFrame. by iApply "IH". - Qed. - - Lemma iProto_consistent_step m1 m2 v p1 : - iProto_consistent (<!> m1) (<?> m2) -∗ - iMsg_car m1 v (Next p1) -∗ - ∃ p2, iMsg_car m2 v (Next p2) ∗ â–· iProto_consistent p1 p2. - Proof. - iIntros "Hprot Hm1". - rewrite iProto_consistent_unfold /iProto_consistent_pre. - iDestruct ("Hprot" with "[//] [//] Hm1") as (p2) "[Hm2 Hprot]". - iExists p2. iFrame. - Qed. - - Lemma iProto_consistent_dual p : - ⊢ iProto_consistent p (iProto_dual p). - Proof. - iLöb as "IH" forall (p). - rewrite iProto_consistent_unfold. - iIntros (a1 a2 m1 m2) "#Hmeq1 #Hmeq2". - destruct a1, a2; [done| | |done]. - - iIntros (v p') "Hm1". - iRewrite "Hmeq1" in "Hmeq2". - rewrite iProto_dual_message. simpl. - rewrite iProto_message_equivI. - iDestruct "Hmeq2" as (_) "Hmeq2". - iSpecialize ("Hmeq2" $! v (Next (iProto_dual p'))). - iExists (iProto_dual p'). - iRewrite -"Hmeq2". - iSplitL; [|iNext; by iApply "IH"]. - simpl. iExists p'. iFrame. done. - - iIntros (v p') "Hm2". - iRewrite "Hmeq1" in "Hmeq2". - rewrite iProto_dual_message. simpl. - rewrite iProto_message_equivI. - iDestruct "Hmeq2" as (_) "Hmeq2". - iSpecialize ("Hmeq2" $! v (Next p')). - iRewrite -"Hmeq2" in "Hm2". - simpl. - iDestruct "Hm2" as (p'') "[Hm2 Hmeq']". - iExists p''. iFrame. - iNext. iRewrite "Hmeq'". - iApply "IH". - Qed. - - Lemma iProto_consistent_le_l pl pl' pr : - iProto_consistent pl pr -∗ pl ⊑ pl' -∗ iProto_consistent pl' pr. - Proof. iIntros "Hprot Hle". by iApply "Hle". Qed. - Lemma iProto_consistent_le_r pl pr pr' : - iProto_consistent pl pr -∗ pr ⊑ pr' -∗ iProto_consistent pl pr'. - Proof. - iIntros "H Hle". iApply iProto_consistent_flip. - iApply "Hle". by iApply iProto_consistent_flip. + iDestruct (iProto_message_equivI with "Hp2") as (_) "{Hp2} #Hm". + iApply "H". by iRewrite -("Hm" $! v (Next p2')). + - done. + - iIntros (v1 v2 p1' p2') "Hm1 Hm2". + iDestruct (iProto_message_equivI with "Hp2") as (_) "{Hp2} #Hm". + iApply ("H" with "Hm1"). by iRewrite -("Hm" $! v2 (Next p2')). + - iDestruct (iProto_message_equivI with "Hp2") as ([=]) "_". + Qed. + Lemma iProto_le_send_send_inv m1 m2 v p2' : + (<!> m1) ⊑ (<!> m2) -∗ + iMsg_car m2 v (Next p2') -∗ ∃ p1', â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1'). + Proof. + iIntros "H Hm2". iDestruct (iProto_le_send_inv with "H") as (a m1') "[Hm1 H]". + iDestruct (iProto_message_equivI with "Hm1") as (<-) "Hm1". + iDestruct ("H" with "Hm2") as (p1') "[Hle Hm]". + iRewrite -("Hm1" $! v (Next p1')) in "Hm". auto with iFrame. + Qed. + Lemma iProto_le_recv_send_inv m1 m2 v1 v2 p1' p2' : + (<?> m1) ⊑ (<!> m2) -∗ + iMsg_car m1 v1 (Next p1') -∗ iMsg_car m2 v2 (Next p2') -∗ ∃ pt, + â–· (p1' ⊑ <!> MSG v2; pt) ∗ â–· ((<?> MSG v1; pt) ⊑ p2'). + Proof. + iIntros "H Hm1 Hm2". iDestruct (iProto_le_send_inv with "H") as (a m1') "[Hm H]". + iDestruct (iProto_message_equivI with "Hm") as (<-) "{Hm} #Hm". + iApply ("H" with "[Hm1] Hm2"). by iRewrite -("Hm" $! v1 (Next p1')). Qed. - Lemma iProto_le_refl p : ⊢ p ⊑ p. - Proof. iIntros (p') "$". Qed. + Lemma iProto_le_recv_inv p1 m2 : + p1 ⊑ (<?> m2) -∗ ∃ m1, + (p1 ≡ <?> m1) ∗ + ∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', + â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). + Proof. + rewrite iProto_le_unfold. iIntros "[[_ Heq]|H]". + { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } + iDestruct "H" as (a1 a2 m1 m2') "(Hp1 & Hp2 & H)". + iExists m1. + iDestruct (iProto_message_equivI with "Hp2") as (<-) "{Hp2} #Hm2". + destruct a1; [done|]. iSplit; [done|]. + iIntros (v p1') "Hm". iDestruct ("H" with "Hm") as (p2') "[Hle Hm]". + iExists p2'. iIntros "{$Hle}". by iRewrite ("Hm2" $! v (Next p2')). + Qed. + Lemma iProto_le_recv_recv_inv m1 m2 v p1' : + (<?> m1) ⊑ (<?> m2) -∗ + iMsg_car m1 v (Next p1') -∗ ∃ p2', â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). + Proof. + iIntros "H Hm2". iDestruct (iProto_le_recv_inv with "H") as (m1') "[Hm1 H]". + iApply "H". iDestruct (iProto_message_equivI with "Hm1") as (_) "Hm1". + by iRewrite -("Hm1" $! v (Next p1')). + Qed. - Lemma iProto_le_end : ⊢ END ⊑ (END : iProto Σ V). - Proof. iApply iProto_le_refl. Qed. - - (* Lemma iProto_le_send m1 m2 : *) - (* (∀ v p2', iMsg_car m2 v (Next p2') -∗ ∃ p1', *) - (* â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1')) -∗ *) - (* (<!> m1) ⊑ (<!> m2). *) - (* Proof. *) - (* iLöb as "IH" forall (m1 m2). *) - (* iIntros "Hle". *) - (* repeat (unfold iProto_le at 3). *) - (* iIntros (p) "Hprot". *) - (* repeat rewrite iProto_consistent_unfold; unfold iProto_consistent_pre. *) - (* iSplit; [iDestruct "Hprot" as "[Hprot _]"|iDestruct "Hprot" as "[_ Hprot]"]. *) - (* { iIntros (a m) "H". *) - (* iDestruct (iProto_message_equivI with "H") as (Heq) "{H} #Hmeq"; subst. *) - (* iIntros (v p') "Hm". *) - (* iSpecialize ("Hprot" with "[//]"). *) - (* iSpecialize ("Hmeq" $! v (Next p')). iRewrite -"Hmeq" in "Hm". *) - (* iSpecialize ("Hle" $! _ _ with "Hm"). iDestruct "Hle" as (p'') "[Hle Hm]". *) - (* iApply "Hle". iApply "Hprot". iApply "Hm". } *) - (* { iIntros (a m) "H". *) - (* iSpecialize ("Hprot" $! _ _ with "H"). *) - (* destruct a. *) - (* - iIntros (v p') "Hm". *) - (* iApply ("IH" with "[Hle]"). *) - (* { iNext. iIntros (v' p'') "Hm". *) - (* iSpecialize ("Hle" with "Hm"). *) - (* iDestruct "Hle" as (p''') "[Hle Hm]". *) - (* iExists p'''. iSplitR "Hm"; [iApply "Hle" | iApply "Hm"]. } *) - (* iApply "Hprot". iApply "Hm". *) - (* - iIntros (v vs Heq); subst. *) - (* iSpecialize ("Hprot" $! _ _ eq_refl). *) - (* iDestruct "Hprot" as (p') "[Hm Hprot]". *) - (* iExists (p'). iFrame "Hm". *) - (* iApply ("IH" with "[Hle] [Hprot]"); [|iApply "Hprot"]. *) - (* iNext. iIntros (v' p'') "Hm". *) - (* iSpecialize ("Hle" with "Hm"). iDestruct "Hle" as (p''') "[Hle Hm]". *) - (* iExists p'''; iSplitR "Hm"; [iApply "Hle" |iApply "Hm"]. } *) - (* Qed. *) - - (* Lemma iProto_le_recv m1 m2 : *) - (* (∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', *) - (* â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2')) -∗ *) - (* (<?> m1) ⊑ (<?> m2). *) - (* Proof. *) - (* iIntros "H" (vsl vsr p) "Hprot". *) - (* iLöb as "IH" forall (vsl vsr m2 p). *) - (* iEval (rewrite iProto_consistent_unfold). *) - (* rewrite /iProto_consistent_pre. *) - (* rewrite iProto_consistent_unfold /iProto_consistent_pre. *) - (* iSplit; last first. *) - (* { *) - (* iDestruct "Hprot" as "[_ Hprot]". *) - (* iIntros ([] m) "Heq". *) - (* { iIntros (v vs) "Hm". *) - (* iDestruct ("Hprot" with "Heq Hm") as "Hprot". *) - (* iIntros "!>". *) - (* by iApply ("IH" with "H"). } *) - (* { iIntros (v vs) "Hm". *) - (* iDestruct ("Hprot" with "Heq Hm") as (p') "[Hm Hprot]". *) - (* iExists p'. iFrame. *) - (* iIntros "!>". *) - (* by iApply ("IH" with "H"). } *) - (* } *) - (* destruct vsr. *) - (* { iIntros (a m) "Heq". rewrite iProto_message_equivI. *) - (* iDestruct "Heq" as (<-) "Heq". *) - (* iIntros (v vs Heq). done. } *) - (* iIntros (a m) "Heq". rewrite iProto_message_equivI. *) - (* iDestruct "Heq" as (<-) "Heq". *) - (* iIntros (w vs Heq). *) - (* assert (v = w) as <- by set_solver. *) - (* assert (vsr = vs) as <- by set_solver. *) - (* iDestruct "Hprot" as "[Hprot _]". *) - (* iDestruct ("Hprot" with "[//] [//]") as (p') "[Hm Hprot]". *) - (* iDestruct ("H" with "Hm") as (p'') "[Hle H]". *) - (* iExists p''. *) - (* iSpecialize ("Heq" $! v (Next p'')). iRewrite -"Heq". *) - (* iFrame. iIntros "!>". *) - (* iApply (iProto_consistent_le_l with "Hprot Hle"). *) - (* Qed. *) - - (* Lemma iProto_le_base a v P p1 p2 : *) - (* â–· (p1 ⊑ p2) -∗ *) - (* (<a> MSG v {{ P }}; p1) ⊑ (<a> MSG v {{ P }}; p2). *) - (* Proof. *) - (* rewrite iMsg_base_eq. iIntros "H". destruct a. *) - (* - iApply iProto_le_send. iIntros (v' p') "(->&Hp&$)". *) - (* iExists p1. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". *) - (* - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&$)". *) - (* iExists p2. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". *) - (* Qed. *) - - (* Lemma iProto_le_end_inv_l p : p ⊑ END -∗ (p ≡ END). *) - (* Proof. *) - (* rewrite iProto_le_unfold. iIntros "[[Hp _]|H] //". *) - (* iDestruct "H" as (a1 a2 m1 m2) "(_ & Heq & _)". *) - (* by iDestruct (iProto_end_message_equivI with "Heq") as %[]. *) - (* Qed. *) - - (* Lemma iProto_le_end_inv_r p : END ⊑ p -∗ (p ≡ END). *) - (* Proof. *) - (* rewrite iProto_le_unfold. iIntros "[[_ Hp]|H] //". *) - (* iDestruct "H" as (a1 a2 m1 m2) "(Heq & _ & _)". *) - (* iDestruct (iProto_end_message_equivI with "Heq") as %[]. *) - (* Qed. *) - - (* Lemma iProto_le_send_inv p1 m2 : *) - (* p1 ⊑ (<!> m2) -∗ ∃ a1 m1, *) - (* (p1 ≡ <a1> m1) ∗ *) - (* match a1 with *) - (* | Send => ∀ v p2', *) - (* iMsg_car m2 v (Next p2') -∗ ∃ p1', *) - (* â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1') *) - (* | Recv => ∀ v1 v2 p1' p2', *) - (* iMsg_car m1 v1 (Next p1') -∗ iMsg_car m2 v2 (Next p2') -∗ ∃ pt, *) - (* â–· (p1' ⊑ <!> MSG v2; pt) ∗ â–· ((<?> MSG v1; pt) ⊑ p2') *) - (* end. *) - (* Proof. *) - (* rewrite iProto_le_unfold. iIntros "[[_ Heq]|H]". *) - (* { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } *) - (* iDestruct "H" as (a1 a2 m1 m2') "(Hp1 & Hp2 & H)". *) - (* iExists _, _; iSplit; [done|]. destruct a1, a2. *) - (* - iIntros (v p2') "Hm2". *) - (* iDestruct (iProto_message_equivI with "Hp2") as (_) "{Hp2} #Hm". *) - (* iApply "H". by iRewrite -("Hm" $! v (Next p2')). *) - (* - done. *) - (* - iIntros (v1 v2 p1' p2') "Hm1 Hm2". *) - (* iDestruct (iProto_message_equivI with "Hp2") as (_) "{Hp2} #Hm". *) - (* iApply ("H" with "Hm1"). by iRewrite -("Hm" $! v2 (Next p2')). *) - (* - iDestruct (iProto_message_equivI with "Hp2") as ([=]) "_". *) - (* Qed. *) - (* Lemma iProto_le_send_send_inv m1 m2 v p2' : *) - (* (<!> m1) ⊑ (<!> m2) -∗ *) - (* iMsg_car m2 v (Next p2') -∗ ∃ p1', â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1'). *) - (* Proof. *) - (* iIntros "H Hm2". iDestruct (iProto_le_send_inv with "H") as (a m1') "[Hm1 H]". *) - (* iDestruct (iProto_message_equivI with "Hm1") as (<-) "Hm1". *) - (* iDestruct ("H" with "Hm2") as (p1') "[Hle Hm]". *) - (* iRewrite -("Hm1" $! v (Next p1')) in "Hm". auto with iFrame. *) - (* Qed. *) - (* Lemma iProto_le_recv_send_inv m1 m2 v1 v2 p1' p2' : *) - (* (<?> m1) ⊑ (<!> m2) -∗ *) - (* iMsg_car m1 v1 (Next p1') -∗ iMsg_car m2 v2 (Next p2') -∗ ∃ pt, *) - (* â–· (p1' ⊑ <!> MSG v2; pt) ∗ â–· ((<?> MSG v1; pt) ⊑ p2'). *) - (* Proof. *) - (* iIntros "H Hm1 Hm2". iDestruct (iProto_le_send_inv with "H") as (a m1') "[Hm H]". *) - (* iDestruct (iProto_message_equivI with "Hm") as (<-) "{Hm} #Hm". *) - (* iApply ("H" with "[Hm1] Hm2"). by iRewrite -("Hm" $! v1 (Next p1')). *) - (* Qed. *) - - (* Lemma iProto_le_recv_inv p1 m2 : *) - (* p1 ⊑ (<?> m2) -∗ ∃ m1, *) - (* (p1 ≡ <?> m1) ∗ *) - (* ∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', *) - (* â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). *) - (* Proof. *) - (* rewrite iProto_le_unfold. iIntros "[[_ Heq]|H]". *) - (* { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } *) - (* iDestruct "H" as (a1 a2 m1 m2') "(Hp1 & Hp2 & H)". *) - (* iExists m1. *) - (* iDestruct (iProto_message_equivI with "Hp2") as (<-) "{Hp2} #Hm2". *) - (* destruct a1; [done|]. iSplit; [done|]. *) - (* iIntros (v p1') "Hm". iDestruct ("H" with "Hm") as (p2') "[Hle Hm]". *) - (* iExists p2'. iIntros "{$Hle}". by iRewrite ("Hm2" $! v (Next p2')). *) - (* Qed. *) - (* Lemma iProto_le_recv_recv_inv m1 m2 v p1' : *) - (* (<?> m1) ⊑ (<?> m2) -∗ *) - (* iMsg_car m1 v (Next p1') -∗ ∃ p2', â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). *) - (* Proof. *) - (* iIntros "H Hm2". iDestruct (iProto_le_recv_inv with "H") as (m1') "[Hm1 H]". *) - (* iApply "H". iDestruct (iProto_message_equivI with "Hm1") as (_) "Hm1". *) - (* by iRewrite -("Hm1" $! v (Next p1')). *) - (* Qed. *) - - (* Lemma iProto_le_refl p : ⊢ p ⊑ p. *) - (* Proof. *) - (* iLöb as "IH" forall (p). destruct (iProto_case p) as [->|([]&m&->)]. *) - (* - iApply iProto_le_end. *) - (* - iApply iProto_le_send. auto 10 with iFrame. *) - (* - iApply iProto_le_recv. auto 10 with iFrame. *) - (* Qed. *) + Lemma iProto_le_refl p : ⊢ p ⊑ p. + Proof. + iLöb as "IH" forall (p). destruct (iProto_case p) as [->|([]&m&->)]. + - iApply iProto_le_end. + - iApply iProto_le_send. auto 10 with iFrame. + - iApply iProto_le_recv. auto 10 with iFrame. + Qed. Lemma iProto_le_trans p1 p2 p3 : p1 ⊑ p2 -∗ p2 ⊑ p3 -∗ p1 ⊑ p3. Proof. - iIntros "H1 H2" (p) "Hprot". - iApply "H2". iApply "H1". done. - Qed. - - (* Lemma iProto_le_payload_elim_l a m v P p : *) - (* (P -∗ (<?> MSG v; p) ⊑ (<a> m)) -∗ *) - (* (<?> MSG v {{ P }}; p) ⊑ (<a> m). *) - (* Proof. *) - (* rewrite iMsg_base_eq. iIntros "H". destruct a. *) - (* - iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= (#?&#?&HP) Hm2 /=". *) - (* iApply (iProto_le_recv_send_inv with "(H HP)"); simpl; auto. *) - (* - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&HP)". *) - (* iApply (iProto_le_recv_recv_inv with "(H HP)"); simpl; auto. *) - (* Qed. *) - (* Lemma iProto_le_payload_elim_r a m v P p : *) - (* (P -∗ (<a> m) ⊑ (<!> MSG v; p)) -∗ *) - (* (<a> m) ⊑ (<!> MSG v {{ P }}; p). *) - (* Proof. *) - (* rewrite iMsg_base_eq. iIntros "H". destruct a. *) - (* - iApply iProto_le_send. iIntros (v' p') "(->&Hp&HP)". *) - (* iApply (iProto_le_send_send_inv with "(H HP)"); simpl; auto. *) - (* - iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1 (->&#?&HP) /=". *) - (* iApply (iProto_le_recv_send_inv with "(H HP) Hm1"); simpl; auto. *) - (* Qed. *) - (* Lemma iProto_le_payload_intro_l v P p : *) - (* P -∗ (<!> MSG v {{ P }}; p) ⊑ (<!> MSG v; p). *) - (* Proof. *) - (* rewrite iMsg_base_eq. *) - (* iIntros "HP". iApply iProto_le_send. iIntros (v' p') "(->&Hp&_) /=". *) - (* iExists p'. iSplitR; [iApply iProto_le_refl|]. auto. *) - (* Qed. *) - (* Lemma iProto_le_payload_intro_r v P p : *) - (* P -∗ (<?> MSG v; p) ⊑ (<?> MSG v {{ P }}; p). *) - (* Proof. *) - (* rewrite iMsg_base_eq. *) - (* iIntros "HP". iApply iProto_le_recv. iIntros (v' p') "(->&Hp&_) /=". *) - (* iExists p'. iSplitR; [iApply iProto_le_refl|]. auto. *) - (* Qed. *) - - (* Lemma iProto_le_exist_elim_l {A} (m1 : A → iMsg Σ V) a m2 : *) - (* (∀ x, (<?> m1 x) ⊑ (<a> m2)) -∗ *) - (* (<? x> m1 x) ⊑ (<a> m2). *) - (* Proof. *) - (* rewrite iMsg_exist_eq. iIntros "H". destruct a. *) - (* - iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1 Hm2 /=". *) - (* iDestruct "Hm1" as (x) "Hm1". *) - (* iApply (iProto_le_recv_send_inv with "H Hm1 Hm2"). *) - (* - iApply iProto_le_recv. iIntros (v p1') "/=". iDestruct 1 as (x) "Hm". *) - (* by iApply (iProto_le_recv_recv_inv with "H"). *) - (* Qed. *) - - (* Lemma iProto_le_exist_elim_l_inhabited `{!Inhabited A} (m : A → iMsg Σ V) p : *) - (* (∀ x, (<?> m x) ⊑ p) -∗ *) - (* (<? x> m x) ⊑ p. *) - (* Proof. *) - (* rewrite iMsg_exist_eq. iIntros "H". *) - (* destruct (iProto_case p) as [Heq | [a [m' Heq]]]. *) - (* - unshelve iSpecialize ("H" $!inhabitant); first by apply _. *) - (* rewrite Heq. *) - (* iDestruct (iProto_le_end_inv_l with "H") as "H". *) - (* rewrite iProto_end_eq iProto_message_eq. *) - (* iDestruct (proto_message_end_equivI with "H") as "[]". *) - (* - iEval (rewrite Heq). destruct a. *) - (* + iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1 Hm2 /=". *) - (* iDestruct "Hm1" as (x) "Hm1". *) - (* iSpecialize ("H" $! x). rewrite Heq. *) - (* iApply (iProto_le_recv_send_inv with "H Hm1 Hm2"). *) - (* + iApply iProto_le_recv. iIntros (v p1') "/=". iDestruct 1 as (x) "Hm". *) - (* iSpecialize ("H" $! x). rewrite Heq. *) - (* by iApply (iProto_le_recv_recv_inv with "H"). *) - (* Qed. *) - - (* Lemma iProto_le_exist_elim_r {A} a m1 (m2 : A → iMsg Σ V) : *) - (* (∀ x, (<a> m1) ⊑ (<!> m2 x)) -∗ *) - (* (<a> m1) ⊑ (<! x> m2 x). *) - (* Proof. *) - (* rewrite iMsg_exist_eq. iIntros "H". destruct a. *) - (* - iApply iProto_le_send. iIntros (v p2'). iDestruct 1 as (x) "Hm". *) - (* by iApply (iProto_le_send_send_inv with "H"). *) - (* - iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1". *) - (* iDestruct 1 as (x) "Hm2". *) - (* iApply (iProto_le_recv_send_inv with "H Hm1 Hm2"). *) - (* Qed. *) - (* Lemma iProto_le_exist_elim_r_inhabited `{Hinh : Inhabited A} p (m : A → iMsg Σ V) : *) - (* (∀ x, p ⊑ (<!> m x)) -∗ *) - (* p ⊑ (<! x> m x). *) - (* Proof. *) - (* rewrite iMsg_exist_eq. iIntros "H". *) - (* destruct (iProto_case p) as [Heq | [a [m' Heq]]]. *) - (* - unshelve iSpecialize ("H" $!inhabitant); first by apply _. *) - (* rewrite Heq. *) - (* iDestruct (iProto_le_end_inv_r with "H") as "H". *) - (* rewrite iProto_end_eq iProto_message_eq. *) - (* iDestruct (proto_message_end_equivI with "H") as "[]". *) - (* - iEval (rewrite Heq). destruct a. *) - (* + iApply iProto_le_send. iIntros (v p2'). iDestruct 1 as (x) "Hm". *) - (* iSpecialize ("H" $! x). rewrite Heq. *) - (* by iApply (iProto_le_send_send_inv with "H"). *) - (* + iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1". *) - (* iDestruct 1 as (x) "Hm2". *) - (* iSpecialize ("H" $! x). rewrite Heq. *) - (* iApply (iProto_le_recv_send_inv with "H Hm1 Hm2"). *) - (* Qed. *) - (* Lemma iProto_le_exist_intro_l {A} (m : A → iMsg Σ V) a : *) - (* ⊢ (<! x> m x) ⊑ (<!> m a). *) - (* Proof. *) - (* rewrite iMsg_exist_eq. iApply iProto_le_send. iIntros (v p') "Hm /=". *) - (* iExists p'. iSplitR; last by auto. iApply iProto_le_refl. *) - (* Qed. *) - (* Lemma iProto_le_exist_intro_r {A} (m : A → iMsg Σ V) a : *) - (* ⊢ (<?> m a) ⊑ (<? x> m x). *) - (* Proof. *) - (* rewrite iMsg_exist_eq. iApply iProto_le_recv. iIntros (v p') "Hm /=". *) - (* iExists p'. iSplitR; last by auto. iApply iProto_le_refl. *) - (* Qed. *) - - (* Lemma iProto_le_texist_elim_l {TT : tele} (m1 : TT → iMsg Σ V) a m2 : *) - (* (∀ x, (<?> m1 x) ⊑ (<a> m2)) -∗ *) - (* (<?.. x> m1 x) ⊑ (<a> m2). *) - (* Proof. *) - (* iIntros "H". iInduction TT as [|T TT] "IH"; simpl; [done|]. *) - (* iApply iProto_le_exist_elim_l; iIntros (x). *) - (* iApply "IH". iIntros (xs). iApply "H". *) - (* Qed. *) - (* Lemma iProto_le_texist_elim_r {TT : tele} a m1 (m2 : TT → iMsg Σ V) : *) - (* (∀ x, (<a> m1) ⊑ (<!> m2 x)) -∗ *) - (* (<a> m1) ⊑ (<!.. x> m2 x). *) - (* Proof. *) - (* iIntros "H". iInduction TT as [|T TT] "IH"; simpl; [done|]. *) - (* iApply iProto_le_exist_elim_r; iIntros (x). *) - (* iApply "IH". iIntros (xs). iApply "H". *) - (* Qed. *) - - (* Lemma iProto_le_texist_intro_l {TT : tele} (m : TT → iMsg Σ V) x : *) - (* ⊢ (<!.. x> m x) ⊑ (<!> m x). *) - (* Proof. *) - (* induction x as [|T TT x xs IH] using tele_arg_ind; simpl. *) - (* { iApply iProto_le_refl. } *) - (* iApply iProto_le_trans; [by iApply iProto_le_exist_intro_l|]. iApply IH. *) - (* Qed. *) - (* Lemma iProto_le_texist_intro_r {TT : tele} (m : TT → iMsg Σ V) x : *) - (* ⊢ (<?> m x) ⊑ (<?.. x> m x). *) - (* Proof. *) - (* induction x as [|T TT x xs IH] using tele_arg_ind; simpl. *) - (* { iApply iProto_le_refl. } *) - (* iApply iProto_le_trans; [|by iApply iProto_le_exist_intro_r]. iApply IH. *) - (* Qed. *) - - (* Lemma iProto_le_base a v P p1 p2 : *) - (* â–· (p1 ⊑ p2) -∗ *) - (* (<a> MSG v {{ P }}; p1) ⊑ (<a> MSG v {{ P }}; p2). *) - (* Proof. *) - (* rewrite iMsg_base_eq. iIntros "H". destruct a. *) - (* - iApply iProto_le_send. iIntros (v' p') "(->&Hp&$)". *) - (* iExists p1. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". *) - (* - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&$)". *) - (* iExists p2. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". *) - (* Qed. *) - - (* Lemma iProto_le_base_swap v1 v2 P1 P2 p : *) - (* ⊢ (<?> MSG v1 {{ P1 }}; <!> MSG v2 {{ P2 }}; p) *) - (* ⊑ (<!> MSG v2 {{ P2 }}; <?> MSG v1 {{ P1 }}; p). *) - (* Proof. *) - (* rewrite {1 3}iMsg_base_eq. iApply iProto_le_swap. *) - (* iIntros (v1' v2' p1' p2') "/= (->&#Hp1&HP1) (->&#Hp2&HP2)". iExists p. *) - (* iSplitL "HP2". *) - (* - iIntros "!>". iRewrite -"Hp1". by iApply iProto_le_payload_intro_l. *) - (* - iIntros "!>". iRewrite -"Hp2". by iApply iProto_le_payload_intro_r. *) - (* Qed. *) - - (* Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. *) - (* Proof. *) - (* iIntros "H". iLöb as "IH" forall (p1 p2). *) - (* destruct (iProto_case p1) as [->|([]&m1&->)]. *) - (* - iDestruct (iProto_le_end_inv_l with "H") as "H". *) - (* iRewrite "H". iApply iProto_le_refl. *) - (* - iDestruct (iProto_le_send_inv with "H") as (a2 m2) "[Hp2 H]". *) - (* iRewrite "Hp2"; clear p2. iEval (rewrite !iProto_dual_message). *) - (* destruct a2; simpl. *) - (* + iApply iProto_le_recv. iIntros (v p1d). *) - (* iDestruct 1 as (p1') "[Hm1 #Hp1d]". *) - (* iDestruct ("H" with "Hm1") as (p2') "[H Hm2]". *) - (* iDestruct ("IH" with "H") as "H". iExists (iProto_dual p2'). *) - (* iSplitL "H"; [iIntros "!>"; by iRewrite "Hp1d"|]. simpl; auto. *) - (* + iApply iProto_le_swap. iIntros (v1 v2 p1d p2d). *) - (* iDestruct 1 as (p1') "[Hm1 #Hp1d]". iDestruct 1 as (p2') "[Hm2 #Hp2d]". *) - (* iDestruct ("H" with "Hm2 Hm1") as (pt) "[H1 H2]". *) - (* iDestruct ("IH" with "H1") as "H1". iDestruct ("IH" with "H2") as "H2 {IH}". *) - (* rewrite !iProto_dual_message /=. iExists (iProto_dual pt). iSplitL "H2". *) - (* * iIntros "!>". iRewrite "Hp1d". by rewrite -iMsg_dual_base. *) - (* * iIntros "!>". iRewrite "Hp2d". by rewrite -iMsg_dual_base. *) - (* - iDestruct (iProto_le_recv_inv with "H") as (m2) "[Hp2 H]". *) - (* iRewrite "Hp2"; clear p2. iEval (rewrite !iProto_dual_message /=). *) - (* iApply iProto_le_send. iIntros (v p2d). *) - (* iDestruct 1 as (p2') "[Hm2 #Hp2d]". *) - (* iDestruct ("H" with "Hm2") as (p1') "[H Hm1]". *) - (* iDestruct ("IH" with "H") as "H". iExists (iProto_dual p1'). *) - (* iSplitL "H"; [iIntros "!>"; by iRewrite "Hp2d"|]. simpl; auto. *) - (* Qed. *) - - (* Lemma iProto_le_amber_internal (p1 p2 : iProto Σ V → iProto Σ V) *) - (* `{Contractive p1, Contractive p2}: *) - (* â–¡ (∀ rec1 rec2, â–· (rec1 ⊑ rec2) → p1 rec1 ⊑ p2 rec2) -∗ *) - (* fixpoint p1 ⊑ fixpoint p2. *) - (* Proof. *) - (* iIntros "#H". iLöb as "IH". *) - (* iEval (rewrite (fixpoint_unfold p1)). *) - (* iEval (rewrite (fixpoint_unfold p2)). *) - (* iApply "H". iApply "IH". *) - (* Qed. *) - (* Lemma iProto_le_amber_external (p1 p2 : iProto Σ V → iProto Σ V) *) - (* `{Contractive p1, Contractive p2}: *) - (* (∀ rec1 rec2, (⊢ rec1 ⊑ rec2) → ⊢ p1 rec1 ⊑ p2 rec2) → *) - (* ⊢ fixpoint p1 ⊑ fixpoint p2. *) - (* Proof. *) - (* intros IH. apply fixpoint_ind. *) - (* - by intros p1' p2' -> ?. *) - (* - exists (fixpoint p2). iApply iProto_le_refl. *) - (* - intros p' ?. rewrite (fixpoint_unfold p2). by apply IH. *) - (* - apply bi.limit_preserving_entails; [done|solve_proper]. *) - (* Qed. *) - - (* Lemma iProto_le_dual_l p1 p2 : iProto_dual p2 ⊑ p1 -∗ iProto_dual p1 ⊑ p2. *) - (* Proof. *) - (* iIntros "H". iEval (rewrite -(involutive iProto_dual p2)). *) - (* by iApply iProto_le_dual. *) - (* Qed. *) - (* Lemma iProto_le_dual_r p1 p2 : p2 ⊑ iProto_dual p1 -∗ p1 ⊑ iProto_dual p2. *) - (* Proof. *) - (* iIntros "H". iEval (rewrite -(involutive iProto_dual p1)). *) - (* by iApply iProto_le_dual. *) - (* Qed. *) - - (* Lemma iProto_le_app p1 p2 p3 p4 : *) - (* p1 ⊑ p2 -∗ p3 ⊑ p4 -∗ p1 <++> p3 ⊑ p2 <++> p4. *) - (* Proof. *) - (* iIntros "H1 H2". iLöb as "IH" forall (p1 p2 p3 p4). *) - (* destruct (iProto_case p2) as [->|([]&m2&->)]. *) - (* - iDestruct (iProto_le_end_inv_l with "H1") as "H1". *) - (* iRewrite "H1". by rewrite !left_id. *) - (* - iDestruct (iProto_le_send_inv with "H1") as (a1 m1) "[Hp1 H1]". *) - (* iRewrite "Hp1"; clear p1. rewrite !iProto_app_message. destruct a1; simpl. *) - (* + iApply iProto_le_send. iIntros (v p24). *) - (* iDestruct 1 as (p2') "[Hm2 #Hp24]". *) - (* iDestruct ("H1" with "Hm2") as (p1') "[H1 Hm1]". *) - (* iExists (p1' <++> p3). iSplitR "Hm1"; [|by simpl; eauto]. *) - (* iIntros "!>". iRewrite "Hp24". by iApply ("IH" with "H1"). *) - (* + iApply iProto_le_swap. iIntros (v1 v2 p13 p24). *) - (* iDestruct 1 as (p1') "[Hm1 #Hp13]". iDestruct 1 as (p2') "[Hm2 #Hp24]". *) - (* iSpecialize ("H1" with "Hm1 Hm2"). *) - (* iDestruct "H1" as (pt) "[H1 H1']". *) - (* iExists (pt <++> p3). iSplitL "H1". *) - (* * iIntros "!>". iRewrite "Hp13". *) - (* rewrite /= -iMsg_app_base -iProto_app_message. *) - (* iApply ("IH" with "H1"). iApply iProto_le_refl. *) - (* * iIntros "!>". iRewrite "Hp24". *) - (* rewrite /= -iMsg_app_base -iProto_app_message. *) - (* iApply ("IH" with "H1' H2"). *) - (* - iDestruct (iProto_le_recv_inv with "H1") as (m1) "[Hp1 H1]". *) - (* iRewrite "Hp1"; clear p1. rewrite !iProto_app_message. iApply iProto_le_recv. *) - (* iIntros (v p13). iDestruct 1 as (p1') "[Hm1 #Hp13]". *) - (* iDestruct ("H1" with "Hm1") as (p2'') "[H1 Hm2]". *) - (* iExists (p2'' <++> p4). iSplitR "Hm2"; [|by simpl; eauto]. *) - (* iIntros "!>". iRewrite "Hp13". by iApply ("IH" with "H1"). *) - (* Qed. *) + iIntros "H1 H2". iLöb as "IH" forall (p1 p2 p3). + destruct (iProto_case p3) as [->|([]&m3&->)]. + - iDestruct (iProto_le_end_inv_l with "H2") as "H2". by iRewrite "H2" in "H1". + - iDestruct (iProto_le_send_inv with "H2") as (a2 m2) "[Hp2 H2]". + iRewrite "Hp2" in "H1"; clear p2. destruct a2. + + iDestruct (iProto_le_send_inv with "H1") as (a1 m1) "[Hp1 H1]". + iRewrite "Hp1"; clear p1. destruct a1. + * iApply iProto_le_send. iIntros (v p3') "Hm3". + iDestruct ("H2" with "Hm3") as (p2') "[Hle Hm2]". + iDestruct ("H1" with "Hm2") as (p1') "[Hle' Hm1]". + iExists p1'. iIntros "{$Hm1} !>". by iApply ("IH" with "Hle'"). + * iApply iProto_le_swap. iIntros (v1 v3 p1' p3') "Hm1 Hm3". + iDestruct ("H2" with "Hm3") as (p2') "[Hle Hm2]". + iDestruct ("H1" with "Hm1 Hm2") as (pt) "[Hp1' Hp2']". + iExists pt. iIntros "{$Hp1'} !>". by iApply ("IH" with "Hp2'"). + + iDestruct (iProto_le_recv_inv with "H1") as (m1) "[Hp1 H1]". + iRewrite "Hp1"; clear p1. iApply iProto_le_swap. + iIntros (v1 v3 p1' p3') "Hm1 Hm3". + iDestruct ("H1" with "Hm1") as (p2') "[Hle Hm2]". + iDestruct ("H2" with "Hm2 Hm3") as (pt) "[Hp2' Hp3']". + iExists pt. iIntros "{$Hp3'} !>". by iApply ("IH" with "Hle"). + - iDestruct (iProto_le_recv_inv with "H2") as (m2) "[Hp2 H3]". + iRewrite "Hp2" in "H1". + iDestruct (iProto_le_recv_inv with "H1") as (m1) "[Hp1 H2]". + iRewrite "Hp1". iApply iProto_le_recv. iIntros (v p1') "Hm1". + iDestruct ("H2" with "Hm1") as (p2') "[Hle Hm2]". + iDestruct ("H3" with "Hm2") as (p3') "[Hle' Hm3]". + iExists p3'. iIntros "{$Hm3} !>". by iApply ("IH" with "Hle"). + Qed. + + Lemma iProto_le_payload_elim_l a m v P p : + (P -∗ (<?> MSG v; p) ⊑ (<a> m)) ⊢ + (<?> MSG v {{ P }}; p) ⊑ (<a> m). + Proof. + rewrite iMsg_base_eq. iIntros "H". destruct a. + - iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= (#?&#?&HP) Hm2 /=". + iApply (iProto_le_recv_send_inv with "(H HP)"); simpl; auto. + - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&HP)". + iApply (iProto_le_recv_recv_inv with "(H HP)"); simpl; auto. + Qed. + Lemma iProto_le_payload_elim_r a m v P p : + (P -∗ (<a> m) ⊑ (<!> MSG v; p)) ⊢ + (<a> m) ⊑ (<!> MSG v {{ P }}; p). + Proof. + rewrite iMsg_base_eq. iIntros "H". destruct a. + - iApply iProto_le_send. iIntros (v' p') "(->&Hp&HP)". + iApply (iProto_le_send_send_inv with "(H HP)"); simpl; auto. + - iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1 (->&#?&HP) /=". + iApply (iProto_le_recv_send_inv with "(H HP) Hm1"); simpl; auto. + Qed. + Lemma iProto_le_payload_intro_l v P p : + P -∗ (<!> MSG v {{ P }}; p) ⊑ (<!> MSG v; p). + Proof. + rewrite iMsg_base_eq. + iIntros "HP". iApply iProto_le_send. iIntros (v' p') "(->&Hp&_) /=". + iExists p'. iSplitR; [iApply iProto_le_refl|]. auto. + Qed. + Lemma iProto_le_payload_intro_r v P p : + P -∗ (<?> MSG v; p) ⊑ (<?> MSG v {{ P }}; p). + Proof. + rewrite iMsg_base_eq. + iIntros "HP". iApply iProto_le_recv. iIntros (v' p') "(->&Hp&_) /=". + iExists p'. iSplitR; [iApply iProto_le_refl|]. auto. + Qed. + + Lemma iProto_le_exist_elim_l {A} (m1 : A → iMsg Σ V) a m2 : + (∀ x, (<?> m1 x) ⊑ (<a> m2)) ⊢ + (<? x> m1 x) ⊑ (<a> m2). + Proof. + rewrite iMsg_exist_eq. iIntros "H". destruct a. + - iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1 Hm2 /=". + iDestruct "Hm1" as (x) "Hm1". + iApply (iProto_le_recv_send_inv with "H Hm1 Hm2"). + - iApply iProto_le_recv. iIntros (v p1') "/=". iDestruct 1 as (x) "Hm". + by iApply (iProto_le_recv_recv_inv with "H"). + Qed. + + Lemma iProto_le_exist_elim_l_inhabited `{!Inhabited A} (m : A → iMsg Σ V) p : + (∀ x, (<?> m x) ⊑ p) ⊢ + (<? x> m x) ⊑ p. + Proof. + rewrite iMsg_exist_eq. iIntros "H". + destruct (iProto_case p) as [Heq | [a [m' Heq]]]. + - unshelve iSpecialize ("H" $!inhabitant); first by apply _. + rewrite Heq. + iDestruct (iProto_le_end_inv_l with "H") as "H". + rewrite iProto_end_eq iProto_message_eq. + iDestruct (proto_message_end_equivI with "H") as "[]". + - iEval (rewrite Heq). destruct a. + + iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1 Hm2 /=". + iDestruct "Hm1" as (x) "Hm1". + iSpecialize ("H" $! x). rewrite Heq. + iApply (iProto_le_recv_send_inv with "H Hm1 Hm2"). + + iApply iProto_le_recv. iIntros (v p1') "/=". iDestruct 1 as (x) "Hm". + iSpecialize ("H" $! x). rewrite Heq. + by iApply (iProto_le_recv_recv_inv with "H"). + Qed. + + Lemma iProto_le_exist_elim_r {A} a m1 (m2 : A → iMsg Σ V) : + (∀ x, (<a> m1) ⊑ (<!> m2 x)) ⊢ + (<a> m1) ⊑ (<! x> m2 x). + Proof. + rewrite iMsg_exist_eq. iIntros "H". destruct a. + - iApply iProto_le_send. iIntros (v p2'). iDestruct 1 as (x) "Hm". + by iApply (iProto_le_send_send_inv with "H"). + - iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1". + iDestruct 1 as (x) "Hm2". + iApply (iProto_le_recv_send_inv with "H Hm1 Hm2"). + Qed. + Lemma iProto_le_exist_elim_r_inhabited `{Hinh : Inhabited A} p (m : A → iMsg Σ V) : + (∀ x, p ⊑ (<!> m x)) ⊢ + p ⊑ (<! x> m x). + Proof. + rewrite iMsg_exist_eq. iIntros "H". + destruct (iProto_case p) as [Heq | [a [m' Heq]]]. + - unshelve iSpecialize ("H" $!inhabitant); first by apply _. + rewrite Heq. + iDestruct (iProto_le_end_inv_r with "H") as "H". + rewrite iProto_end_eq iProto_message_eq. + iDestruct (proto_message_end_equivI with "H") as "[]". + - iEval (rewrite Heq). destruct a. + + iApply iProto_le_send. iIntros (v p2'). iDestruct 1 as (x) "Hm". + iSpecialize ("H" $! x). rewrite Heq. + by iApply (iProto_le_send_send_inv with "H"). + + iApply iProto_le_swap. iIntros (v1 v2 p1' p2') "/= Hm1". + iDestruct 1 as (x) "Hm2". + iSpecialize ("H" $! x). rewrite Heq. + iApply (iProto_le_recv_send_inv with "H Hm1 Hm2"). + Qed. + Lemma iProto_le_exist_intro_l {A} (m : A → iMsg Σ V) a : + ⊢ (<! x> m x) ⊑ (<!> m a). + Proof. + rewrite iMsg_exist_eq. iApply iProto_le_send. iIntros (v p') "Hm /=". + iExists p'. iSplitR; last by auto. iApply iProto_le_refl. + Qed. + Lemma iProto_le_exist_intro_r {A} (m : A → iMsg Σ V) a : + ⊢ (<?> m a) ⊑ (<? x> m x). + Proof. + rewrite iMsg_exist_eq. iApply iProto_le_recv. iIntros (v p') "Hm /=". + iExists p'. iSplitR; last by auto. iApply iProto_le_refl. + Qed. + + Lemma iProto_le_texist_elim_l {TT : tele} (m1 : TT → iMsg Σ V) a m2 : + (∀ x, (<?> m1 x) ⊑ (<a> m2)) ⊢ + (<?.. x> m1 x) ⊑ (<a> m2). + Proof. + iIntros "H". iInduction TT as [|T TT] "IH"; simpl; [done|]. + iApply iProto_le_exist_elim_l; iIntros (x). + iApply "IH". iIntros (xs). iApply "H". + Qed. + Lemma iProto_le_texist_elim_r {TT : tele} a m1 (m2 : TT → iMsg Σ V) : + (∀ x, (<a> m1) ⊑ (<!> m2 x)) -∗ + (<a> m1) ⊑ (<!.. x> m2 x). + Proof. + iIntros "H". iInduction TT as [|T TT] "IH"; simpl; [done|]. + iApply iProto_le_exist_elim_r; iIntros (x). + iApply "IH". iIntros (xs). iApply "H". + Qed. + + Lemma iProto_le_texist_intro_l {TT : tele} (m : TT → iMsg Σ V) x : + ⊢ (<!.. x> m x) ⊑ (<!> m x). + Proof. + induction x as [|T TT x xs IH] using tele_arg_ind; simpl. + { iApply iProto_le_refl. } + iApply iProto_le_trans; [by iApply iProto_le_exist_intro_l|]. iApply IH. + Qed. + Lemma iProto_le_texist_intro_r {TT : tele} (m : TT → iMsg Σ V) x : + ⊢ (<?> m x) ⊑ (<?.. x> m x). + Proof. + induction x as [|T TT x xs IH] using tele_arg_ind; simpl. + { iApply iProto_le_refl. } + iApply iProto_le_trans; [|by iApply iProto_le_exist_intro_r]. iApply IH. + Qed. + + Lemma iProto_le_base a v P p1 p2 : + â–· (p1 ⊑ p2) ⊢ + (<a> MSG v {{ P }}; p1) ⊑ (<a> MSG v {{ P }}; p2). + Proof. + rewrite iMsg_base_eq. iIntros "H". destruct a. + - iApply iProto_le_send. iIntros (v' p') "(->&Hp&$)". + iExists p1. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". + - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&$)". + iExists p2. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". + Qed. + + Lemma iProto_le_base_swap v1 v2 P1 P2 p : + ⊢ (<?> MSG v1 {{ P1 }}; <!> MSG v2 {{ P2 }}; p) + ⊑ (<!> MSG v2 {{ P2 }}; <?> MSG v1 {{ P1 }}; p). + Proof. + rewrite {1 3}iMsg_base_eq. iApply iProto_le_swap. + iIntros (v1' v2' p1' p2') "/= (->&#Hp1&HP1) (->&#Hp2&HP2)". iExists p. + iSplitL "HP2". + - iIntros "!>". iRewrite -"Hp1". by iApply iProto_le_payload_intro_l. + - iIntros "!>". iRewrite -"Hp2". by iApply iProto_le_payload_intro_r. + Qed. + + Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. + Proof. + iIntros "H". iLöb as "IH" forall (p1 p2). + destruct (iProto_case p1) as [->|([]&m1&->)]. + - iDestruct (iProto_le_end_inv_l with "H") as "H". + iRewrite "H". iApply iProto_le_refl. + - iDestruct (iProto_le_send_inv with "H") as (a2 m2) "[Hp2 H]". + iRewrite "Hp2"; clear p2. iEval (rewrite !iProto_dual_message). + destruct a2; simpl. + + iApply iProto_le_recv. iIntros (v p1d). + iDestruct 1 as (p1') "[Hm1 #Hp1d]". + iDestruct ("H" with "Hm1") as (p2') "[H Hm2]". + iDestruct ("IH" with "H") as "H". iExists (iProto_dual p2'). + iSplitL "H"; [iIntros "!>"; by iRewrite "Hp1d"|]. simpl; auto. + + iApply iProto_le_swap. iIntros (v1 v2 p1d p2d). + iDestruct 1 as (p1') "[Hm1 #Hp1d]". iDestruct 1 as (p2') "[Hm2 #Hp2d]". + iDestruct ("H" with "Hm2 Hm1") as (pt) "[H1 H2]". + iDestruct ("IH" with "H1") as "H1". iDestruct ("IH" with "H2") as "H2 {IH}". + rewrite !iProto_dual_message /=. iExists (iProto_dual pt). iSplitL "H2". + * iIntros "!>". iRewrite "Hp1d". by rewrite -iMsg_dual_base. + * iIntros "!>". iRewrite "Hp2d". by rewrite -iMsg_dual_base. + - iDestruct (iProto_le_recv_inv with "H") as (m2) "[Hp2 H]". + iRewrite "Hp2"; clear p2. iEval (rewrite !iProto_dual_message /=). + iApply iProto_le_send. iIntros (v p2d). + iDestruct 1 as (p2') "[Hm2 #Hp2d]". + iDestruct ("H" with "Hm2") as (p1') "[H Hm1]". + iDestruct ("IH" with "H") as "H". iExists (iProto_dual p1'). + iSplitL "H"; [iIntros "!>"; by iRewrite "Hp2d"|]. simpl; auto. + Qed. + + Lemma iProto_le_amber_internal (p1 p2 : iProto Σ V → iProto Σ V) + `{Contractive p1, Contractive p2}: + â–¡ (∀ rec1 rec2, â–· (rec1 ⊑ rec2) → p1 rec1 ⊑ p2 rec2) ⊢ + fixpoint p1 ⊑ fixpoint p2. + Proof. + iIntros "#H". iLöb as "IH". + iEval (rewrite (fixpoint_unfold p1)). + iEval (rewrite (fixpoint_unfold p2)). + iApply "H". iApply "IH". + Qed. + Lemma iProto_le_amber_external (p1 p2 : iProto Σ V → iProto Σ V) + `{Contractive p1, Contractive p2}: + (∀ rec1 rec2, (⊢ rec1 ⊑ rec2) → ⊢ p1 rec1 ⊑ p2 rec2) → + ⊢ fixpoint p1 ⊑ fixpoint p2. + Proof. + intros IH. apply fixpoint_ind. + - by intros p1' p2' -> ?. + - exists (fixpoint p2). iApply iProto_le_refl. + - intros p' ?. rewrite (fixpoint_unfold p2). by apply IH. + - apply bi.limit_preserving_entails; [done|solve_proper]. + Qed. + + Lemma iProto_le_dual_l p1 p2 : iProto_dual p2 ⊑ p1 ⊢ iProto_dual p1 ⊑ p2. + Proof. + iIntros "H". iEval (rewrite -(involutive iProto_dual p2)). + by iApply iProto_le_dual. + Qed. + Lemma iProto_le_dual_r p1 p2 : p2 ⊑ iProto_dual p1 ⊢ p1 ⊑ iProto_dual p2. + Proof. + iIntros "H". iEval (rewrite -(involutive iProto_dual p1)). + by iApply iProto_le_dual. + Qed. + + Lemma iProto_le_app p1 p2 p3 p4 : + p1 ⊑ p2 -∗ p3 ⊑ p4 -∗ p1 <++> p3 ⊑ p2 <++> p4. + Proof. + iIntros "H1 H2". iLöb as "IH" forall (p1 p2 p3 p4). + destruct (iProto_case p2) as [->|([]&m2&->)]. + - iDestruct (iProto_le_end_inv_l with "H1") as "H1". + iRewrite "H1". by rewrite !left_id. + - iDestruct (iProto_le_send_inv with "H1") as (a1 m1) "[Hp1 H1]". + iRewrite "Hp1"; clear p1. rewrite !iProto_app_message. destruct a1; simpl. + + iApply iProto_le_send. iIntros (v p24). + iDestruct 1 as (p2') "[Hm2 #Hp24]". + iDestruct ("H1" with "Hm2") as (p1') "[H1 Hm1]". + iExists (p1' <++> p3). iSplitR "Hm1"; [|by simpl; eauto]. + iIntros "!>". iRewrite "Hp24". by iApply ("IH" with "H1"). + + iApply iProto_le_swap. iIntros (v1 v2 p13 p24). + iDestruct 1 as (p1') "[Hm1 #Hp13]". iDestruct 1 as (p2') "[Hm2 #Hp24]". + iSpecialize ("H1" with "Hm1 Hm2"). + iDestruct "H1" as (pt) "[H1 H1']". + iExists (pt <++> p3). iSplitL "H1". + * iIntros "!>". iRewrite "Hp13". + rewrite /= -iMsg_app_base -iProto_app_message. + iApply ("IH" with "H1"). iApply iProto_le_refl. + * iIntros "!>". iRewrite "Hp24". + rewrite /= -iMsg_app_base -iProto_app_message. + iApply ("IH" with "H1' H2"). + - iDestruct (iProto_le_recv_inv with "H1") as (m1) "[Hp1 H1]". + iRewrite "Hp1"; clear p1. rewrite !iProto_app_message. iApply iProto_le_recv. + iIntros (v p13). iDestruct 1 as (p1') "[Hm1 #Hp13]". + iDestruct ("H1" with "Hm1") as (p2'') "[H1 Hm2]". + iExists (p2'' <++> p4). iSplitR "Hm2"; [|by simpl; eauto]. + iIntros "!>". iRewrite "Hp13". by iApply ("IH" with "H1"). + Qed. (** ** Lemmas about the auxiliary definitions and invariants *) Global Instance iProto_app_recvs_ne vs : @@ -1150,6 +1019,12 @@ Section proto. Global Instance iProto_app_recvs_proper vs : Proper ((≡) ==> (≡)) (iProto_app_recvs (Σ:=Σ) (V:=V) vs). Proof. induction vs; solve_proper. Qed. + Global Instance iProto_interp_ne vsl vsr : + NonExpansive2 (iProto_interp (Σ:=Σ) (V:=V) vsl vsr). + Proof. solve_proper. Qed. + Global Instance iProto_interp_proper vsl vsr : + Proper ((≡) ==> (≡) ==> (≡)) (iProto_interp (Σ:=Σ) (V:=V) vsl vsr). + Proof. apply (ne_proper_2 _). Qed. Global Instance iProto_own_frag_ne γ s : NonExpansive (iProto_own_frag γ s). Proof. solve_proper. Qed. @@ -1157,8 +1032,8 @@ Section proto. Lemma iProto_own_auth_agree γ s p p' : iProto_own_auth γ s p -∗ iProto_own_frag γ s p' -∗ â–· (p ≡ p'). Proof. - iIntros "Hâ— Hâ—¯". iDestruct (own_valid_2 with "Hâ— Hâ—¯") as "H✓". - iDestruct (excl_auth_agreeI with "H✓") as "H✓". + iIntros "Hâ— Hâ—¯". iCombine "Hâ— Hâ—¯" gives "H✓". + iDestruct (excl_auth_agreeI with "H✓") as "{H✓} H✓". iApply (later_equivI_1 with "H✓"). Qed. @@ -1171,6 +1046,67 @@ Section proto. by rewrite own_op. Qed. + Lemma iProto_interp_nil p : ⊢ iProto_interp [] [] p (iProto_dual p). + Proof. iExists p; simpl. iSplitL; iApply iProto_le_refl. Qed. + + Lemma iProto_interp_flip vsl vsr pl pr : + iProto_interp vsl vsr pl pr -∗ iProto_interp vsr vsl pr pl. + Proof. + iDestruct 1 as (p) "[Hp Hdp]". iExists (iProto_dual p). + rewrite (involutive iProto_dual). iFrame. + Qed. + + Lemma iProto_interp_le_l vsl vsr pl pl' pr : + iProto_interp vsl vsr pl pr -∗ pl ⊑ pl' -∗ iProto_interp vsl vsr pl' pr. + Proof. + iDestruct 1 as (p) "[Hp Hdp]". iIntros "Hle". iExists p. + iFrame "Hdp". by iApply (iProto_le_trans with "Hp"). + Qed. + Lemma iProto_interp_le_r vsl vsr pl pr pr' : + iProto_interp vsl vsr pl pr -∗ pr ⊑ pr' -∗ iProto_interp vsl vsr pl pr'. + Proof. + iIntros "H Hle". iApply iProto_interp_flip. + iApply (iProto_interp_le_l with "[H] Hle"). by iApply iProto_interp_flip. + Qed. + + Lemma iProto_interp_send vl ml vsl vsr pr pl' : + iProto_interp vsl vsr (<!> ml) pr -∗ + iMsg_car ml vl (Next pl') -∗ + â–·^(length vsr) iProto_interp (vsl ++ [vl]) vsr pl' pr. + Proof. + iDestruct 1 as (p) "[Hp Hdp] /="; iIntros "Hml". + iDestruct (iProto_le_trans _ _ (<!> MSG vl; pl') with "Hp [Hml]") as "Hp". + { iApply iProto_le_send. rewrite iMsg_base_eq. iIntros (v' p') "(->&Hp&_) /=". + iExists p'. iSplitR; [iApply iProto_le_refl|]. by iRewrite -"Hp". } + iInduction vsr as [|vr vsr] "IH" forall (pl'); simpl. + { iExists pl'; simpl. iSplitR; [iApply iProto_le_refl|]. + iApply (iProto_le_trans with "[Hp] Hdp"). + iInduction vsl as [|vl' vsl] "IH"; simpl. + { iApply iProto_le_dual_r. rewrite iProto_dual_message iMsg_dual_base /=. + by rewrite involutive. } + iApply iProto_le_base; iIntros "!>". by iApply "IH". } + iDestruct (iProto_le_recv_send_inv _ _ vr vl + (iProto_app_recvs vsr p) pl' with "Hp [] []") as (p') "[H1 H2]"; + [rewrite iMsg_base_eq; by auto..|]. + iIntros "!>". iSpecialize ("IH" with "Hdp H1"). iIntros "!>". + iDestruct "IH" as (p'') "[Hp'' Hdp'']". iExists p''. iFrame "Hdp''". + iApply (iProto_le_trans with "[Hp''] H2"); simpl. by iApply iProto_le_base. + Qed. + + Lemma iProto_interp_recv vl vsl vsr pl mr : + iProto_interp (vl :: vsl) vsr pl (<?> mr) -∗ + ∃ pr, iMsg_car mr vl (Next pr) ∗ â–· iProto_interp vsl vsr pl pr. + Proof. + iDestruct 1 as (p) "[Hp Hdp] /=". + iDestruct (iProto_le_recv_inv with "Hdp") as (m) "[#Hm Hpr]". + iDestruct (iProto_message_equivI with "Hm") as (_) "{Hm} #Hm". + iDestruct ("Hpr" $! vl (iProto_app_recvs vsl (iProto_dual p)) with "[]") + as (pr'') "[Hler Hpr]". + { iRewrite -("Hm" $! vl (Next (iProto_app_recvs vsl (iProto_dual p)))). + rewrite iMsg_base_eq; auto. } + iExists pr''. iIntros "{$Hpr} !>". iExists p. iFrame. + Qed. + Global Instance iProto_own_ne γ s : NonExpansive (iProto_own γ s). Proof. solve_proper. Qed. Global Instance iProto_own_proper γ s : Proper ((≡) ==> (≡)) (iProto_own γ s). @@ -1185,7 +1121,8 @@ Section proto. Lemma iProto_init p : ⊢ |==> ∃ γ, - iProto_ctx γ ∗ iProto_own γ Left p ∗ iProto_own γ Right (iProto_dual p). + iProto_ctx γ [] [] ∗ + iProto_own γ Left p ∗ iProto_own γ Right (iProto_dual p). Proof. iMod (own_alloc (â—E (Next p) â‹… â—¯E (Next p))) as (lγ) "[Hâ—l Hâ—¯l]". { by apply excl_auth_valid. } @@ -1193,154 +1130,166 @@ Section proto. â—¯E (Next (iProto_dual p)))) as (rγ) "[Hâ—r Hâ—¯r]". { by apply excl_auth_valid. } pose (ProtName lγ rγ) as γ. iModIntro. iExists γ. iSplitL "Hâ—l Hâ—r". - { iExists p, (iProto_dual p). iFrame. iApply iProto_consistent_dual. } + { iExists p, (iProto_dual p). iFrame. iApply iProto_interp_nil. } iSplitL "Hâ—¯l"; iExists _; iFrame; iApply iProto_le_refl. Qed. - Definition side_dual s := - match s with - | Left => Right - | Right => Left - end. - - Lemma iProto_step_l γ m1 m2 p1 v : - iProto_ctx γ -∗ iProto_own γ Left (<!> m1) -∗ iProto_own γ Right (<?> m2) -∗ - iMsg_car m1 v (Next p1) ==∗ - â–· ∃ p2, iMsg_car m2 v (Next p2) ∗ iProto_ctx γ ∗ - iProto_own γ Left p1 ∗ iProto_own γ Right p2. - Proof. - iDestruct 1 as (pl pr) "(Hâ—l & Hâ—r & Hconsistent)". - iDestruct 1 as (pl') "[Hlel Hâ—¯l]". - iDestruct 1 as (pr') "[Hler Hâ—¯r]". - iIntros "Hm". - iDestruct (iProto_own_auth_agree with "Hâ—l Hâ—¯l") as "#Hpl". - iDestruct (iProto_own_auth_agree with "Hâ—r Hâ—¯r") as "#Hpr". - iAssert (â–· (pl ⊑ <!> m1))%I - with "[Hlel]" as "{Hpl} Hlel"; first (iNext; by iRewrite "Hpl"). - iAssert (â–· (pr ⊑ <?> m2))%I - with "[Hler]" as "{Hpr} Hler"; first (iNext; by iRewrite "Hpr"). - iDestruct (iProto_consistent_le_l with "Hconsistent Hlel") as "Hconsistent". - iDestruct (iProto_consistent_le_r with "Hconsistent Hler") as "Hconsistent". - iDestruct (iProto_consistent_step with "Hconsistent [Hm //]") as - (p2) "[Hm2 Hconsistent]". - iMod (iProto_own_auth_update _ _ _ _ p1 with "Hâ—l Hâ—¯l") as "[Hâ—l Hâ—¯l]". - iMod (iProto_own_auth_update _ _ _ _ p2 with "Hâ—r Hâ—¯r") as "[Hâ—r Hâ—¯r]". - iIntros "!>!>". - iExists p2. iFrame. - iSplitL "Hconsistent Hâ—l Hâ—r". - - iExists _, _. iFrame. - - iSplitL "Hâ—¯l". - + iExists _. iFrame. iApply iProto_le_refl. - + iExists _. iFrame. iApply iProto_le_refl. - Qed. - - Lemma iProto_step_r γ m1 m2 p2 v : - iProto_ctx γ -∗ iProto_own γ Left (<?> m1) -∗ iProto_own γ Right (<!> m2) -∗ - iMsg_car m2 v (Next p2) ==∗ - â–· ∃ p1, iMsg_car m1 v (Next p1) ∗ iProto_ctx γ ∗ - iProto_own γ Left p1 ∗ iProto_own γ Right p2. - Proof. - iDestruct 1 as (pl pr) "(Hâ—l & Hâ—r & Hconsistent)". - iDestruct 1 as (pl') "[Hlel Hâ—¯l]". - iDestruct 1 as (pr') "[Hler Hâ—¯r]". - iIntros "Hm". - iDestruct (iProto_own_auth_agree with "Hâ—l Hâ—¯l") as "#Hpl". - iDestruct (iProto_own_auth_agree with "Hâ—r Hâ—¯r") as "#Hpr". - iAssert (â–· (pl ⊑ <?> m1))%I - with "[Hlel]" as "{Hpl} Hlel"; first (iNext; by iRewrite "Hpl"). - iAssert (â–· (pr ⊑ <!> m2))%I - with "[Hler]" as "{Hpr} Hler"; first (iNext; by iRewrite "Hpr"). - iDestruct (iProto_consistent_le_l with "Hconsistent Hlel") as "Hconsistent". - iDestruct (iProto_consistent_le_r with "Hconsistent Hler") as "Hconsistent". - iDestruct (iProto_consistent_flip with "Hconsistent") as - "Hconsistent". - iDestruct (iProto_consistent_step with "Hconsistent [Hm //]") as - (p1) "[Hm1 Hconsistent]". - iMod (iProto_own_auth_update _ _ _ _ p1 with "Hâ—l Hâ—¯l") as "[Hâ—l Hâ—¯l]". - iMod (iProto_own_auth_update _ _ _ _ p2 with "Hâ—r Hâ—¯r") as "[Hâ—r Hâ—¯r]". - iIntros "!>!>". - iExists p1. iFrame. - iSplitL "Hconsistent Hâ—l Hâ—r". - - iExists _, _. iFrame. iApply iProto_consistent_flip. iFrame. - - iSplitL "Hâ—¯l". - + iExists _. iFrame. iApply iProto_le_refl. - + iExists _. iFrame. iApply iProto_le_refl. - Qed. - - (* (** The instances below make it possible to use the tactics [iIntros], *) - (* [iExist], [iSplitL]/[iSplitR], [iFrame] and [iModIntro] on [iProto_le] goals. *) *) - (* Global Instance iProto_le_from_forall_l {A} a (m1 : A → iMsg Σ V) m2 name : *) - (* AsIdentName m1 name → *) - (* FromForall (iProto_message Recv (iMsg_exist m1) ⊑ (<a> m2)) *) - (* (λ x, (<?> m1 x) ⊑ (<a> m2))%I name | 10. *) - (* Proof. intros _. apply iProto_le_exist_elim_l. Qed. *) - (* Global Instance iProto_le_from_forall_r {A} a m1 (m2 : A → iMsg Σ V) name : *) - (* AsIdentName m2 name → *) - (* FromForall ((<a> m1) ⊑ iProto_message Send (iMsg_exist m2)) *) - (* (λ x, (<a> m1) ⊑ (<!> m2 x))%I name | 11. *) - (* Proof. intros _. apply iProto_le_exist_elim_r. Qed. *) - - (* Global Instance iProto_le_from_wand_l a m v P p : *) - (* TCIf (TCEq P True%I) False TCTrue → *) - (* FromWand ((<?> MSG v {{ P }}; p) ⊑ (<a> m)) P ((<?> MSG v; p) ⊑ (<a> m)) | 10. *) - (* Proof. intros _. apply iProto_le_payload_elim_l. Qed. *) - (* Global Instance iProto_le_from_wand_r a m v P p : *) - (* FromWand ((<a> m) ⊑ (<!> MSG v {{ P }}; p)) P ((<a> m) ⊑ (<!> MSG v; p)) | 11. *) - (* Proof. apply iProto_le_payload_elim_r. Qed. *) - - (* Global Instance iProto_le_from_exist_l {A} (m : A → iMsg Σ V) p : *) - (* FromExist ((<! x> m x) ⊑ p) (λ a, (<!> m a) ⊑ p)%I | 10. *) - (* Proof. *) - (* rewrite /FromExist. iDestruct 1 as (x) "H". *) - (* iApply (iProto_le_trans with "[] H"). iApply iProto_le_exist_intro_l. *) - (* Qed. *) - (* Global Instance iProto_le_from_exist_r {A} (m : A → iMsg Σ V) p : *) - (* FromExist (p ⊑ <? x> m x) (λ a, p ⊑ (<?> m a))%I | 11. *) - (* Proof. *) - (* rewrite /FromExist. iDestruct 1 as (x) "H". *) - (* iApply (iProto_le_trans with "H"). iApply iProto_le_exist_intro_r. *) - (* Qed. *) - - (* Global Instance iProto_le_from_sep_l m v P p : *) - (* FromSep ((<!> MSG v {{ P }}; p) ⊑ (<!> m)) P ((<!> MSG v; p) ⊑ (<!> m)) | 10. *) - (* Proof. *) - (* rewrite /FromSep. iIntros "[HP H]". *) - (* iApply (iProto_le_trans with "[HP] H"). by iApply iProto_le_payload_intro_l. *) - (* Qed. *) - (* Global Instance iProto_le_from_sep_r m v P p : *) - (* FromSep ((<?> m) ⊑ (<?> MSG v {{ P }}; p)) P ((<?> m) ⊑ (<?> MSG v; p)) | 11. *) - (* Proof. *) - (* rewrite /FromSep. iIntros "[HP H]". *) - (* iApply (iProto_le_trans with "H"). by iApply iProto_le_payload_intro_r. *) - (* Qed. *) - - (* Global Instance iProto_le_frame_l q m v R P Q p : *) - (* Frame q R P Q → *) - (* Frame q R ((<!> MSG v {{ P }}; p) ⊑ (<!> m)) *) - (* ((<!> MSG v {{ Q }}; p) ⊑ (<!> m)) | 10. *) - (* Proof. *) - (* rewrite /Frame /=. iIntros (HP) "[HR H]". *) - (* iApply (iProto_le_trans with "[HR] H"). iApply iProto_le_payload_elim_r. *) - (* iIntros "HQ". iApply iProto_le_payload_intro_l. iApply HP; iFrame. *) - (* Qed. *) - (* Global Instance iProto_le_frame_r q m v R P Q p : *) - (* Frame q R P Q → *) - (* Frame q R ((<?> m) ⊑ (<?> MSG v {{ P }}; p)) *) - (* ((<?> m) ⊑ (<?> MSG v {{ Q }}; p)) | 11. *) - (* Proof. *) - (* rewrite /Frame /=. iIntros (HP) "[HR H]". *) - (* iApply (iProto_le_trans with "H"). iApply iProto_le_payload_elim_l. *) - (* iIntros "HQ". iApply iProto_le_payload_intro_r. iApply HP; iFrame. *) - (* Qed. *) - - (* Global Instance iProto_le_from_modal a v p1 p2 : *) - (* FromModal True (modality_instances.modality_laterN 1) (p1 ⊑ p2) *) - (* ((<a> MSG v; p1) ⊑ (<a> MSG v; p2)) (p1 ⊑ p2). *) - (* Proof. intros _. apply iProto_le_base. Qed. *) + Lemma iProto_send_l γ m vsr vsl vl p : + iProto_ctx γ vsl vsr -∗ + iProto_own γ Left (<!> m) -∗ + iMsg_car m vl (Next p) ==∗ + â–·^(length vsr) iProto_ctx γ (vsl ++ [vl]) vsr ∗ + iProto_own γ Left p. + Proof. + iDestruct 1 as (pl pr) "(Hâ—l & Hâ—r & Hinterp)". + iDestruct 1 as (pl') "[Hle Hâ—¯]". iIntros "Hm". + iDestruct (iProto_own_auth_agree with "Hâ—l Hâ—¯") as "#Hp". + iAssert (â–· (pl ⊑ <!> m))%I + with "[Hle]" as "{Hp} Hle"; first (iNext; by iRewrite "Hp"). + iDestruct (iProto_interp_le_l with "Hinterp Hle") as "Hinterp". + iDestruct (iProto_interp_send with "Hinterp [Hm //]") as "Hinterp". + iMod (iProto_own_auth_update _ _ _ _ p with "Hâ—l Hâ—¯") as "[Hâ—l Hâ—¯]". + iIntros "!>". iSplitR "Hâ—¯". + - iIntros "!>". iExists p, pr. iFrame. + - iExists p. iFrame. iApply iProto_le_refl. + Qed. + + Lemma iProto_send_r γ m vsr vsl vr p : + iProto_ctx γ vsl vsr -∗ + iProto_own γ Right (<!> m) -∗ + iMsg_car m vr (Next p) ==∗ + â–·^(length vsl) iProto_ctx γ vsl (vsr ++ [vr]) ∗ + iProto_own γ Right p. + Proof. + iDestruct 1 as (pl pr) "(Hâ—l & Hâ—r & Hinterp)". + iDestruct 1 as (pr') "[Hle Hâ—¯]". iIntros "Hm". + iDestruct (iProto_own_auth_agree with "Hâ—r Hâ—¯") as "#Hp". + iAssert (â–· (pr ⊑ <!> m))%I + with "[Hle]" as "{Hp} Hle"; first (iNext; by iRewrite "Hp"). + iDestruct (iProto_interp_flip with "Hinterp") as "Hinterp". + iDestruct (iProto_interp_le_l with "Hinterp Hle") as "Hinterp". + iDestruct (iProto_interp_send with "Hinterp [Hm //]") as "Hinterp". + iMod (iProto_own_auth_update _ _ _ _ p with "Hâ—r Hâ—¯") as "[Hâ—r Hâ—¯]". + iIntros "!>". iSplitR "Hâ—¯". + - iIntros "!>". iExists pl, p. iFrame. by iApply iProto_interp_flip. + - iExists p. iFrame. iApply iProto_le_refl. + Qed. + + Lemma iProto_recv_l γ m vr vsr vsl : + iProto_ctx γ vsl (vr :: vsr) -∗ + iProto_own γ Left (<?> m) ==∗ + â–· ∃ p, + iProto_ctx γ vsl vsr ∗ + iProto_own γ Left p ∗ + iMsg_car m vr (Next p). + Proof. + iDestruct 1 as (pl pr) "(Hâ—l & Hâ—r & Hinterp)". + iDestruct 1 as (p) "[Hle Hâ—¯]". + iDestruct (iProto_own_auth_agree with "Hâ—l Hâ—¯") as "#Hp". + iDestruct (iProto_interp_le_l with "Hinterp [Hle]") as "Hinterp". + { iIntros "!>". by iRewrite "Hp". } + iDestruct (iProto_interp_flip with "Hinterp") as "Hinterp". + iDestruct (iProto_interp_recv with "Hinterp") as (q) "[Hm Hinterp]". + iMod (iProto_own_auth_update _ _ _ _ q with "Hâ—l Hâ—¯") as "[Hâ—l Hâ—¯]". + iIntros "!> !> /=". iExists q. iFrame "Hm". iSplitR "Hâ—¯". + - iExists q, pr. iFrame. by iApply iProto_interp_flip. + - iExists q. iIntros "{$Hâ—¯} !>". iApply iProto_le_refl. + Qed. + + Lemma iProto_recv_r γ m vl vsr vsl : + iProto_ctx γ (vl :: vsl) vsr -∗ + iProto_own γ Right (<?> m) ==∗ + â–· ∃ p, + iProto_ctx γ vsl vsr ∗ + iProto_own γ Right p ∗ + iMsg_car m vl (Next p). + Proof. + iDestruct 1 as (pl pr) "(Hâ—l & Hâ—r & Hinterp)". + iDestruct 1 as (p) "[Hle Hâ—¯]". + iDestruct (iProto_own_auth_agree with "Hâ—r Hâ—¯") as "#Hp". + iDestruct (iProto_interp_le_r with "Hinterp [Hle]") as "Hinterp". + { iIntros "!>". by iRewrite "Hp". } + iDestruct (iProto_interp_recv with "Hinterp") as (q) "[Hm Hinterp]". + iMod (iProto_own_auth_update _ _ _ _ q with "Hâ—r Hâ—¯") as "[Hâ—r Hâ—¯]". + iIntros "!> !> /=". iExists q. iFrame "Hm". iSplitR "Hâ—¯". + - iExists pl, q. iFrame. + - iExists q. iIntros "{$Hâ—¯} !>". iApply iProto_le_refl. + Qed. + + (** The instances below make it possible to use the tactics [iIntros], + [iExist], [iSplitL]/[iSplitR], [iFrame] and [iModIntro] on [iProto_le] goals. *) + Global Instance iProto_le_from_forall_l {A} a (m1 : A → iMsg Σ V) m2 name : + AsIdentName m1 name → + FromForall (iProto_message Recv (iMsg_exist m1) ⊑ (<a> m2)) + (λ x, (<?> m1 x) ⊑ (<a> m2))%I name | 10. + Proof. intros _. apply iProto_le_exist_elim_l. Qed. + Global Instance iProto_le_from_forall_r {A} a m1 (m2 : A → iMsg Σ V) name : + AsIdentName m2 name → + FromForall ((<a> m1) ⊑ iProto_message Send (iMsg_exist m2)) + (λ x, (<a> m1) ⊑ (<!> m2 x))%I name | 11. + Proof. intros _. apply iProto_le_exist_elim_r. Qed. + + Global Instance iProto_le_from_wand_l a m v P p : + TCIf (TCEq P True%I) False TCTrue → + FromWand ((<?> MSG v {{ P }}; p) ⊑ (<a> m)) P ((<?> MSG v; p) ⊑ (<a> m)) | 10. + Proof. intros _. apply iProto_le_payload_elim_l. Qed. + Global Instance iProto_le_from_wand_r a m v P p : + FromWand ((<a> m) ⊑ (<!> MSG v {{ P }}; p)) P ((<a> m) ⊑ (<!> MSG v; p)) | 11. + Proof. apply iProto_le_payload_elim_r. Qed. + + Global Instance iProto_le_from_exist_l {A} (m : A → iMsg Σ V) p : + FromExist ((<! x> m x) ⊑ p) (λ a, (<!> m a) ⊑ p)%I | 10. + Proof. + rewrite /FromExist. iDestruct 1 as (x) "H". + iApply (iProto_le_trans with "[] H"). iApply iProto_le_exist_intro_l. + Qed. + Global Instance iProto_le_from_exist_r {A} (m : A → iMsg Σ V) p : + FromExist (p ⊑ <? x> m x) (λ a, p ⊑ (<?> m a))%I | 11. + Proof. + rewrite /FromExist. iDestruct 1 as (x) "H". + iApply (iProto_le_trans with "H"). iApply iProto_le_exist_intro_r. + Qed. + + Global Instance iProto_le_from_sep_l m v P p : + FromSep ((<!> MSG v {{ P }}; p) ⊑ (<!> m)) P ((<!> MSG v; p) ⊑ (<!> m)) | 10. + Proof. + rewrite /FromSep. iIntros "[HP H]". + iApply (iProto_le_trans with "[HP] H"). by iApply iProto_le_payload_intro_l. + Qed. + Global Instance iProto_le_from_sep_r m v P p : + FromSep ((<?> m) ⊑ (<?> MSG v {{ P }}; p)) P ((<?> m) ⊑ (<?> MSG v; p)) | 11. + Proof. + rewrite /FromSep. iIntros "[HP H]". + iApply (iProto_le_trans with "H"). by iApply iProto_le_payload_intro_r. + Qed. + + Global Instance iProto_le_frame_l q m v R P Q p : + Frame q R P Q → + Frame q R ((<!> MSG v {{ P }}; p) ⊑ (<!> m)) + ((<!> MSG v {{ Q }}; p) ⊑ (<!> m)) | 10. + Proof. + rewrite /Frame /=. iIntros (HP) "[HR H]". + iApply (iProto_le_trans with "[HR] H"). iApply iProto_le_payload_elim_r. + iIntros "HQ". iApply iProto_le_payload_intro_l. iApply HP; iFrame. + Qed. + Global Instance iProto_le_frame_r q m v R P Q p : + Frame q R P Q → + Frame q R ((<?> m) ⊑ (<?> MSG v {{ P }}; p)) + ((<?> m) ⊑ (<?> MSG v {{ Q }}; p)) | 11. + Proof. + rewrite /Frame /=. iIntros (HP) "[HR H]". + iApply (iProto_le_trans with "H"). iApply iProto_le_payload_elim_l. + iIntros "HQ". iApply iProto_le_payload_intro_r. iApply HP; iFrame. + Qed. + + Global Instance iProto_le_from_modal a v p1 p2 : + FromModal True (modality_instances.modality_laterN 1) (p1 ⊑ p2) + ((<a> MSG v; p1) ⊑ (<a> MSG v; p2)) (p1 ⊑ p2). + Proof. intros _. apply iProto_le_base. Qed. End proto. -Typeclasses Opaque iProto_ctx iProto_own. +Global Typeclasses Opaque iProto_ctx iProto_own. Global Hint Extern 0 (environments.envs_entails _ (?x ⊑ ?y)) => first [is_evar x; fail 1 | is_evar y; fail 1|iApply iProto_le_refl] : core. -- GitLab From 16764081e3f24f903d82c1a5d5449a989e9afc14 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Tue, 6 Feb 2024 14:13:15 +0100 Subject: [PATCH 38/81] More refactoring --- multi_actris/channel/proofmode.v | 134 +++++++++++++++--- .../channel/proto_consistency_examples.v | 118 --------------- 2 files changed, 111 insertions(+), 141 deletions(-) diff --git a/multi_actris/channel/proofmode.v b/multi_actris/channel/proofmode.v index 33e7d5a..cc8cc20 100644 --- a/multi_actris/channel/proofmode.v +++ b/multi_actris/channel/proofmode.v @@ -7,11 +7,12 @@ standard pattern using type classes to perform the normalization. In addition to the tactics for symbolic execution, this file defines the tactic [solve_proto_contractive], which can be used to automatically prove that recursive protocols are contractive. *) +From iris.algebra Require Import gmap. From iris.proofmode Require Import coq_tactics reduction spec_patterns. From iris.heap_lang Require Export proofmode notation. From multi_actris.channel Require Import proto_model. From multi_actris Require Export channel. -Export action. +Set Default Proof Using "Type". (** * Tactics for proving contractiveness of protocols *) Ltac f_dist_le := @@ -343,28 +344,115 @@ Tactic Notation "wp_send" "(" uconstr(x1) uconstr(x2) uconstr(x3) uconstr(x4) ") wp_send_core (eexists x1; eexists x2; eexists x3; eexists x4; eexists x5; eexists x6; eexists x7; eexists x8) with pat. -(* Section channel. *) -(* Context `{!heapGS Σ, !chanG Σ}. *) -(* Implicit Types p : iProto Σ. *) -(* Implicit Types TT : tele. *) +Lemma iProto_consistent_equiv_proof {Σ} (ps : gmap nat (iProto Σ)) : + (∀ i j, valid_target ps i j) ∗ + (∀ i j m1 m2, + (ps !!! i ≡ (<(Send, j)> m1)%proto) -∗ + (ps !!! j ≡ (<(Recv, i)> m2)%proto) -∗ + ∃ m1' m2' (TT1:tele) (TT2:tele) tv1 tP1 tp1 tv2 tP2 tp2, + (<(Send, j)> m1')%proto ≡ (<(Send, j)> m1)%proto ∗ + (<(Recv, i)> m2')%proto ≡ (<(Recv, i)> m2)%proto ∗ + ⌜MsgTele (TT:=TT1) m1' tv1 tP1 tp1⌠∗ + ⌜MsgTele (TT:=TT2) m2' tv2 tP2 tp2⌠∗ + ∀.. (x : TT1), tele_app tP1 x -∗ + ∃.. (y : TT2), ⌜tele_app tv1 x = tele_app tv2 y⌠∗ + tele_app tP2 y ∗ + â–· (iProto_consistent + (<[i:=tele_app tp1 x]>(<[j:=tele_app tp2 y]>ps)))) -∗ + iProto_consistent ps. +Proof. + iIntros "[H' H]". + rewrite iProto_consistent_unfold. + iFrame. + iIntros (i j m1 m2) "Hm1 Hm2". + iDestruct ("H" with "Hm1 Hm2") + as (m1' m2' TT1 TT2 tv1 tP1 tp1 tv2 tP2 tp2) + "(Heq1 & Heq2& %Hm1' & %Hm2' & H)". + rewrite !iProto_message_equivI. + iDestruct "Heq1" as "[_ Heq1]". + iDestruct "Heq2" as "[_ Heq2]". + iIntros (v p1) "Hm1". + iSpecialize ("Heq1" $! v (Next p1)). + iRewrite -"Heq1" in "Hm1". + rewrite Hm1'. + rewrite iMsg_base_eq. rewrite iMsg_texist_exist. + iDestruct "Hm1" as (x Htv1) "[Hp1 HP1]". + iDestruct ("H" with "HP1") as (y Htv2) "[HP2 H]". + iExists (tele_app tp2 y). + iSpecialize ("Heq2" $! v (Next (tele_app tp2 y))). + iRewrite -"Heq2". + rewrite Hm2'. rewrite iMsg_base_eq. rewrite iMsg_texist_exist. + iSplitL "HP2". + { iExists y. iFrame. + iSplit; [|done]. + iPureIntro. subst. done. } + iNext. iRewrite -"Hp1". done. +Qed. + +(* TODO: Improve automation *) +Tactic Notation "iProto_consistent_take_step_step" := + let i := fresh in + let j := fresh in + let m1 := fresh in + let m2 := fresh in + iIntros (i j m1 m2) "#Hm1 #Hm2"; + repeat (destruct i as [|i]; + [repeat (rewrite lookup_total_insert_ne; [|lia]); + try (by rewrite lookup_total_empty iProto_end_message_equivI); + try (rewrite lookup_total_insert; + try (by rewrite iProto_end_message_equivI); + iDestruct (iProto_message_equivI with "Hm1") + as "[%Heq1 Hm1']";simplify_eq)| + repeat (rewrite lookup_total_insert_ne; [|lia]); + try (by rewrite lookup_total_empty iProto_end_message_equivI)]); + repeat (rewrite lookup_total_insert_ne; [|lia]); + try rewrite lookup_total_empty; + try (by iProto_end_message_equivI); + rewrite lookup_total_insert; + iDestruct (iProto_message_equivI with "Hm2") + as "[%Heq2 Hm2']";simplify_eq; + try (iClear "Hm1' Hm2'"; + iExists _,_,_,_,_,_,_,_,_,_; + iSplitL "Hm1"; [iFrame "#"|]; + iSplitL "Hm2"; [iFrame "#"|]; + iSplit; [iPureIntro; tc_solve|]; + iSplit; [iPureIntro; tc_solve|]; + simpl; iClear "Hm1 Hm2"; clear m1 m2); + try (repeat (rewrite (insert_commute _ _ i); [|done]); + rewrite insert_insert; + repeat (rewrite (insert_commute _ _ j); [|done]); + rewrite insert_insert). -(* (* TODO: Why do the tactics not strip laters? *) *) -(* Lemma recv_test c p : *) -(* {{{ c ↣ (<(Recv,0) @(x:Z)> MSG #x ; p) }}} *) -(* recv c #0 *) -(* {{{ x, RET #x; c ↣ p }}}. *) -(* Proof. *) -(* iIntros (Φ) "Hc HΦ". *) -(* wp_recv (x) as "_". *) -(* Admitted. *) +Tactic Notation "iProto_consistent_take_step_target" := + let i := fresh in + iIntros (i j a m); rewrite /valid_target; + iIntros "#Hm"; + repeat (destruct i as [|i]; + [repeat (rewrite lookup_total_insert_ne; [|lia]); + try (by rewrite lookup_total_empty iProto_end_message_equivI); + try (rewrite lookup_total_insert; + try (by rewrite iProto_end_message_equivI); + iDestruct (iProto_message_equivI with "Hm1") + as "[%Heq1 Hm1']";simplify_eq)| + repeat (rewrite lookup_total_insert_ne; [|lia]); + try (by rewrite lookup_total_empty iProto_end_message_equivI)]); + repeat (rewrite lookup_total_insert_ne; [|lia]); + try rewrite lookup_total_empty; + try (by iProto_end_message_equivI); + rewrite lookup_total_insert; + iDestruct (iProto_message_equivI with "Hm") + as "[%Heq Hm']";simplify_eq; + repeat (try rewrite lookup_empty; + try rewrite lookup_insert; + rewrite lookup_insert_ne; [|lia]); + try rewrite lookup_insert; try done. -(* Lemma send_test c p : *) -(* {{{ c ↣ (<(Send,0) @(x:Z)> MSG #x ; p) }}} *) -(* send c #0 #42 *) -(* {{{ x, RET #x; c ↣ p }}}. *) -(* Proof. *) -(* iIntros (Φ) "Hc HΦ". *) -(* wp_send (42%Z) with "[//]". *) -(* Admitted. *) +Tactic Notation "iProto_consistent_take_step" := + try iNext; + iApply iProto_consistent_equiv_proof; + iSplitR; [iProto_consistent_take_step_target| + iProto_consistent_take_step_step]. -(* End channel. *) +Tactic Notation "clean_map" constr(i) := + iEval (repeat (rewrite (insert_commute _ _ i); [|done])); + rewrite (insert_insert _ i). diff --git a/multi_actris/channel/proto_consistency_examples.v b/multi_actris/channel/proto_consistency_examples.v index 14771d2..0583fb2 100644 --- a/multi_actris/channel/proto_consistency_examples.v +++ b/multi_actris/channel/proto_consistency_examples.v @@ -1,124 +1,6 @@ -From iris.algebra Require Import gmap excl_auth gmap_view. -From iris.proofmode Require Import proofmode. -From iris.base_logic Require Export lib.iprop. -From iris.base_logic Require Import lib.own. From multi_actris.channel Require Import proofmode. Set Default Proof Using "Type". -Lemma iProto_consistent_equiv_proof {Σ} (ps : gmap nat (iProto Σ)) : - (∀ i j, valid_target ps i j) ∗ - (∀ i j m1 m2, - (ps !!! i ≡ (<(Send, j)> m1)%proto) -∗ - (ps !!! j ≡ (<(Recv, i)> m2)%proto) -∗ - ∃ m1' m2' (TT1:tele) (TT2:tele) tv1 tP1 tp1 tv2 tP2 tp2, - (<(Send, j)> m1')%proto ≡ (<(Send, j)> m1)%proto ∗ - (<(Recv, i)> m2')%proto ≡ (<(Recv, i)> m2)%proto ∗ - ⌜MsgTele (TT:=TT1) m1' tv1 tP1 tp1⌠∗ - ⌜MsgTele (TT:=TT2) m2' tv2 tP2 tp2⌠∗ - ∀.. (x : TT1), tele_app tP1 x -∗ - ∃.. (y : TT2), ⌜tele_app tv1 x = tele_app tv2 y⌠∗ - tele_app tP2 y ∗ - â–· (iProto_consistent - (<[i:=tele_app tp1 x]>(<[j:=tele_app tp2 y]>ps)))) -∗ - iProto_consistent ps. -Proof. - iIntros "[H' H]". - rewrite iProto_consistent_unfold. - iFrame. - iIntros (i j m1 m2) "Hm1 Hm2". - iDestruct ("H" with "Hm1 Hm2") - as (m1' m2' TT1 TT2 tv1 tP1 tp1 tv2 tP2 tp2) - "(Heq1 & Heq2& %Hm1' & %Hm2' & H)". - rewrite !iProto_message_equivI. - iDestruct "Heq1" as "[_ Heq1]". - iDestruct "Heq2" as "[_ Heq2]". - iIntros (v p1) "Hm1". - iSpecialize ("Heq1" $! v (Next p1)). - iRewrite -"Heq1" in "Hm1". - rewrite Hm1'. - rewrite iMsg_base_eq. rewrite iMsg_texist_exist. - iDestruct "Hm1" as (x Htv1) "[Hp1 HP1]". - iDestruct ("H" with "HP1") as (y Htv2) "[HP2 H]". - iExists (tele_app tp2 y). - iSpecialize ("Heq2" $! v (Next (tele_app tp2 y))). - iRewrite -"Heq2". - rewrite Hm2'. rewrite iMsg_base_eq. rewrite iMsg_texist_exist. - iSplitL "HP2". - { iExists y. iFrame. - iSplit; [|done]. - iPureIntro. subst. done. } - iNext. iRewrite -"Hp1". done. -Qed. - -(* TODO: Improve automation *) -(* Could clean up repeated inserts to save traverses *) -Tactic Notation "iProto_consistent_take_step_step" := - let i := fresh in - let j := fresh in - let m1 := fresh in - let m2 := fresh in - iIntros (i j m1 m2) "#Hm1 #Hm2"; - repeat (destruct i as [|i]; - [repeat (rewrite lookup_total_insert_ne; [|lia]); - try (by rewrite lookup_total_empty iProto_end_message_equivI); - try (rewrite lookup_total_insert; - try (by rewrite iProto_end_message_equivI); - iDestruct (iProto_message_equivI with "Hm1") - as "[%Heq1 Hm1']";simplify_eq)| - repeat (rewrite lookup_total_insert_ne; [|lia]); - try (by rewrite lookup_total_empty iProto_end_message_equivI)]); - repeat (rewrite lookup_total_insert_ne; [|lia]); - try rewrite lookup_total_empty; - try (by iProto_end_message_equivI); - rewrite lookup_total_insert; - iDestruct (iProto_message_equivI with "Hm2") - as "[%Heq2 Hm2']";simplify_eq; - try (iClear "Hm1' Hm2'"; - iExists _,_,_,_,_,_,_,_,_,_; - iSplitL "Hm1"; [iFrame "#"|]; - iSplitL "Hm2"; [iFrame "#"|]; - iSplit; [iPureIntro; tc_solve|]; - iSplit; [iPureIntro; tc_solve|]; - simpl; iClear "Hm1 Hm2"; clear m1 m2); - try (repeat (rewrite (insert_commute _ _ i); [|done]); - rewrite insert_insert; - repeat (rewrite (insert_commute _ _ j); [|done]); - rewrite insert_insert). - -Tactic Notation "iProto_consistent_take_step_target" := - let i := fresh in - iIntros (i j a m); rewrite /valid_target; - iIntros "#Hm"; - repeat (destruct i as [|i]; - [repeat (rewrite lookup_total_insert_ne; [|lia]); - try (by rewrite lookup_total_empty iProto_end_message_equivI); - try (rewrite lookup_total_insert; - try (by rewrite iProto_end_message_equivI); - iDestruct (iProto_message_equivI with "Hm1") - as "[%Heq1 Hm1']";simplify_eq)| - repeat (rewrite lookup_total_insert_ne; [|lia]); - try (by rewrite lookup_total_empty iProto_end_message_equivI)]); - repeat (rewrite lookup_total_insert_ne; [|lia]); - try rewrite lookup_total_empty; - try (by iProto_end_message_equivI); - rewrite lookup_total_insert; - iDestruct (iProto_message_equivI with "Hm") - as "[%Heq Hm']";simplify_eq; - repeat (try rewrite lookup_empty; - try rewrite lookup_insert; - rewrite lookup_insert_ne; [|lia]); - try rewrite lookup_insert; try done. - -Tactic Notation "iProto_consistent_take_step" := - try iNext; - iApply iProto_consistent_equiv_proof; - iSplitR; [iProto_consistent_take_step_target| - iProto_consistent_take_step_step]. - -Tactic Notation "clean_map" constr(i) := - iEval (repeat (rewrite (insert_commute _ _ i); [|done])); - rewrite (insert_insert _ i). - Definition iProto_empty {Σ} : gmap nat (iProto Σ) := ∅. Lemma iProto_consistent_empty {Σ} : -- GitLab From fd433c0e8c0dba0fd3f48394c8761b9dda7b990a Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Tue, 6 Feb 2024 14:30:05 +0100 Subject: [PATCH 39/81] Refactoring of examples --- .../basics.v} | 172 ++++++++---------- 1 file changed, 71 insertions(+), 101 deletions(-) rename multi_actris/{channel/proto_consistency_examples.v => examples/basics.v} (91%) diff --git a/multi_actris/channel/proto_consistency_examples.v b/multi_actris/examples/basics.v similarity index 91% rename from multi_actris/channel/proto_consistency_examples.v rename to multi_actris/examples/basics.v index 0583fb2..69a94fb 100644 --- a/multi_actris/channel/proto_consistency_examples.v +++ b/multi_actris/examples/basics.v @@ -7,13 +7,13 @@ Lemma iProto_consistent_empty {Σ} : ⊢ iProto_consistent (@iProto_empty Σ). Proof. iProto_consistent_take_step. Qed. -Definition iProto_binary `{!invGS Σ} : gmap nat (iProto Σ) := +Definition iProto_binary `{!heapGS Σ} : gmap nat (iProto Σ) := <[0 := (<(Send, 1) @ (x:Z)> MSG #x ; END)%proto ]> (<[1 := (<(Recv, 0) @ (x:Z)> MSG #x ; END)%proto ]> ∅). -Lemma iProto_binary_consistent `{!invGS Σ} : - ⊢ iProto_consistent (@iProto_binary Σ invGS0). +Lemma iProto_binary_consistent `{!heapGS Σ} : + ⊢ iProto_consistent iProto_binary. Proof. rewrite /iProto_binary. iProto_consistent_take_step. @@ -21,24 +21,6 @@ Proof. iProto_consistent_take_step. Qed. -Definition iProto_roundtrip `{!invGS Σ} : gmap nat (iProto Σ) := - <[0 := (<(Send, 1) @ (x:Z)> MSG #x ; <(Recv, 2)> MSG #x; END)%proto ]> - (<[1 := (<(Recv, 0) @ (x:Z)> MSG #x ; <(Send, 2)> MSG #x; END)%proto ]> - (<[2 := (<(Recv, 1) @ (x:Z)> MSG #x ; <(Send, 0)> MSG #x; END)%proto ]> ∅)). - -Lemma iProto_roundtrip_consistent `{!invGS Σ} : - ⊢ iProto_consistent (@iProto_roundtrip Σ invGS0). -Proof. - rewrite /iProto_roundtrip. - iProto_consistent_take_step. - iIntros (x) "_". iExists x. iSplit; [done|]. iSplit; [done|]. - iProto_consistent_take_step. - iIntros "_". iExists x. iSplit; [done|]. iSplit; [done|]. - iProto_consistent_take_step. - iIntros "_". iSplit; [done|]. iSplit; [done|]. - iProto_consistent_take_step. -Qed. - Definition roundtrip_prog : val := λ: <>, let: "cs" := new_chan #3 in @@ -51,18 +33,32 @@ Definition roundtrip_prog : val := Section channel. Context `{!heapGS Σ, !chanG Σ}. - Implicit Types p : iProto Σ. - Implicit Types TT : tele. + + Definition iProto_roundtrip : gmap nat (iProto Σ) := + <[0 := (<(Send, 1) @ (x:Z)> MSG #x ; <(Recv, 2)> MSG #x; END)%proto ]> + (<[1 := (<(Recv, 0) @ (x:Z)> MSG #x ; <(Send, 2)> MSG #x; END)%proto ]> + (<[2 := (<(Recv, 1) @ (x:Z)> MSG #x ; <(Send, 0)> MSG #x; END)%proto ]> ∅)). + + Lemma iProto_roundtrip_consistent : + ⊢ iProto_consistent iProto_roundtrip. + Proof. + rewrite /iProto_roundtrip. + iProto_consistent_take_step. + iIntros (x) "_". iExists x. iSplit; [done|]. iSplit; [done|]. + iProto_consistent_take_step. + iIntros "_". iExists x. iSplit; [done|]. iSplit; [done|]. + iProto_consistent_take_step. + iIntros "_". iSplit; [done|]. iSplit; [done|]. + iProto_consistent_take_step. + Qed. (* TODO: Fix nat/Z coercion. *) Lemma roundtrip_prog_spec : {{{ True }}} roundtrip_prog #() {{{ RET #42 ; True }}}. Proof using chanG0 heapGS0 Σ. iIntros (Φ) "_ HΦ". wp_lam. - wp_smart_apply (new_chan_spec 3 iProto_roundtrip). - { lia. } - { set_solver. } - { iApply iProto_roundtrip_consistent. } + wp_smart_apply (new_chan_spec 3 iProto_roundtrip); + [lia|set_solver|iApply iProto_roundtrip_consistent|]. iIntros (cs) "Hcs". wp_smart_apply (get_chan_spec _ 0 with "Hcs"); [set_solver|]. iIntros (c0) "[Hc0 Hcs]". @@ -74,15 +70,23 @@ Section channel. { iIntros "!>". wp_recv (x) as "_". wp_send with "[//]". done. } wp_smart_apply (wp_fork with "[Hc2]"). { iIntros "!>". wp_recv (x) as "_". wp_send with "[//]". done. } - wp_send with "[//]". - wp_recv as "_". - by iApply "HΦ". + wp_send with "[//]". wp_recv as "_". by iApply "HΦ". Qed. End channel. +Definition roundtrip_ref_prog : val := + λ: <>, + let: "cs" := new_chan #3 in + let: "c0" := get_chan "cs" #0 in + let: "c1" := get_chan "cs" #1 in + let: "c2" := get_chan "cs" #2 in + Fork (let: "l" := recv "c1" #0 in "l" <- !"l" + #1;; send "c1" #2 "l");; + Fork (let: "l" := recv "c2" #1 in "l" <- !"l" + #1;; send "c2" #0 #());; + let: "l" := ref #40 in send "c0" #1 "l";; recv "c0" #2;; !"l". + Section roundtrip_ref. - Context `{!heapGS Σ}. + Context `{!heapGS Σ, !chanG Σ}. Definition iProto_roundtrip_ref : gmap nat (iProto Σ) := <[0 := (<(Send, 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; @@ -107,23 +111,6 @@ Section roundtrip_ref. iProto_consistent_take_step. Qed. -End roundtrip_ref. - -Definition roundtrip_ref_prog : val := - λ: <>, - let: "cs" := new_chan #3 in - let: "c0" := get_chan "cs" #0 in - let: "c1" := get_chan "cs" #1 in - let: "c2" := get_chan "cs" #2 in - Fork (let: "l" := recv "c1" #0 in "l" <- !"l" + #1;; send "c1" #2 "l");; - Fork (let: "l" := recv "c2" #1 in "l" <- !"l" + #1;; send "c2" #0 #());; - let: "l" := ref #40 in send "c0" #1 "l";; recv "c0" #2;; !"l". - -Section proof. - Context `{!heapGS Σ, !chanG Σ}. - Implicit Types p : iProto Σ. - Implicit Types TT : tele. - Lemma roundtrip_ref_prog_spec : {{{ True }}} roundtrip_ref_prog #() {{{ RET #42 ; True }}}. Proof using chanG0. @@ -149,27 +136,39 @@ Section proof. by iApply "HΦ". Qed. -End proof. +End roundtrip_ref. + +Definition roundtrip_ref_rec_prog : val := + λ: <>, + let: "cs" := new_chan #3 in + let: "c0" := get_chan "cs" #0 in + let: "c1" := get_chan "cs" #1 in + let: "c2" := get_chan "cs" #2 in + Fork ((rec: "go" "c1" := + let: "l" := recv "c1" #0 in "l" <- !"l" + #1;; send "c1" #2 "l";; + "go" "c1") "c1");; + Fork ((rec: "go" "c2" := + let: "l" := recv "c2" #1 in "l" <- !"l" + #1;; send "c2" #0 #();; + "go" "c2") "c2");; + let: "l" := ref #38 in + send "c0" #1 "l";; recv "c0" #2;; + send "c0" #1 "l";; recv "c0" #2;; !"l". Section roundtrip_ref_rec. - Context `{!heapGS Σ}. + Context `{!heapGS Σ, !chanG Σ}. Definition iProto_roundtrip_ref_rec1_aux (rec : iProto Σ) : iProto Σ := (<(Send, 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; <(Recv, 2)> MSG #() {{ l ↦ #(x+2) }} ; rec)%proto. - Instance iProto_roundtrip_ref_rec1_contractive : Contractive iProto_roundtrip_ref_rec1_aux. Proof. solve_proto_contractive. Qed. - Definition iProto_roundtrip_ref_rec1 := fixpoint iProto_roundtrip_ref_rec1_aux. - Lemma iProto_roundtrip_ref_rec1_unfold : iProto_roundtrip_ref_rec1 ≡ (iProto_roundtrip_ref_rec1_aux iProto_roundtrip_ref_rec1). Proof. apply (fixpoint_unfold _). Qed. - Global Instance iProto_roundtrip_ref_rec1_proto_unfold : ProtoUnfold iProto_roundtrip_ref_rec1 (iProto_roundtrip_ref_rec1_aux iProto_roundtrip_ref_rec1). @@ -178,14 +177,11 @@ Section roundtrip_ref_rec. Definition iProto_roundtrip_ref_rec2_aux (rec : iProto Σ) : iProto Σ := (<(Recv, 0) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; <(Send, 2)> MSG #l {{ l ↦ #(x+1) }}; rec)%proto. - Instance iProto_roundtrip_ref_rec2_contractive : Contractive iProto_roundtrip_ref_rec2_aux. Proof. solve_proto_contractive. Qed. - Definition iProto_roundtrip_ref_rec2 := fixpoint iProto_roundtrip_ref_rec2_aux. - Lemma iProto_roundtrip_ref_rec2_unfold : iProto_roundtrip_ref_rec2 ≡ (iProto_roundtrip_ref_rec2_aux iProto_roundtrip_ref_rec2). @@ -195,23 +191,18 @@ Section roundtrip_ref_rec. ProtoUnfold iProto_roundtrip_ref_rec2 (iProto_roundtrip_ref_rec2_aux iProto_roundtrip_ref_rec2). Proof. apply proto_unfold_eq, (fixpoint_unfold _). Qed. - Definition iProto_roundtrip_ref_rec3_aux (rec : iProto Σ) : iProto Σ := (<(Recv, 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; <(Send, 0)> MSG #() {{ l ↦ #(x+1) }}; rec)%proto. - Instance iProto_roundtrip_ref_rec3_contractive : Contractive iProto_roundtrip_ref_rec3_aux. Proof. solve_proto_contractive. Qed. - Definition iProto_roundtrip_ref_rec3 := fixpoint iProto_roundtrip_ref_rec3_aux. - Lemma iProto_roundtrip_ref_rec3_unfold : iProto_roundtrip_ref_rec3 ≡ (iProto_roundtrip_ref_rec3_aux iProto_roundtrip_ref_rec3). Proof. apply (fixpoint_unfold _). Qed. - Global Instance iProto_roundtrip_ref_rec3_proto_unfold : ProtoUnfold iProto_roundtrip_ref_rec3 (iProto_roundtrip_ref_rec3_aux iProto_roundtrip_ref_rec3). @@ -243,29 +234,6 @@ Section roundtrip_ref_rec. done. Qed. -End roundtrip_ref_rec. - -Definition roundtrip_ref_rec_prog : val := - λ: <>, - let: "cs" := new_chan #3 in - let: "c0" := get_chan "cs" #0 in - let: "c1" := get_chan "cs" #1 in - let: "c2" := get_chan "cs" #2 in - Fork ((rec: "go" "c1" := - let: "l" := recv "c1" #0 in "l" <- !"l" + #1;; send "c1" #2 "l";; - "go" "c1") "c1");; - Fork ((rec: "go" "c2" := - let: "l" := recv "c2" #1 in "l" <- !"l" + #1;; send "c2" #0 #();; - "go" "c2") "c2");; - let: "l" := ref #38 in - send "c0" #1 "l";; recv "c0" #2;; - send "c0" #1 "l";; recv "c0" #2;; !"l". - -Section proof. - Context `{!heapGS Σ, !chanG Σ}. - Implicit Types p : iProto Σ. - Implicit Types TT : tele. - Lemma roundtrip_ref_rec_prog_spec : {{{ True }}} roundtrip_ref_rec_prog #() {{{ RET #42 ; True }}}. Proof using chanG0. @@ -294,21 +262,12 @@ Section proof. by iApply "HΦ". Qed. -End proof. +End roundtrip_ref_rec. Section parallel. Context `{!heapGS Σ}. (** - - 0 -> 1 : (x1:Z) < x1 > . - 0 -> 2 : (x2:Z) < x2 > . - 2 -> 3 : (y1:Z) < x1+y1 > ; - 3 -> 4 : (y2:Z) < x2+y2 > ; - 3 -> 0 : < x1+y1 > ; - 4 -> 0 : < x2+y2 > ; - end - 0 / \ 1 2 @@ -541,10 +500,21 @@ Section two_buyer_ref. End two_buyer_ref. -Section fwd. +Section forwarder. Context `{!heapGS Σ}. - Definition iProto_fwd : gmap nat (iProto Σ) := + (** + + 0 + / | \ + / | \ + | 1 | + \ / \ / + 2 3 + + *) + + Definition iProto_forwarder : gmap nat (iProto Σ) := <[0 := (<(Send, 1) @ (x:Z)> MSG #x ; <(Send, 1) @ (b:bool)> MSG #b ; <(Send, if b then 2 else 3) > MSG #x ; END)%proto]> @@ -558,10 +528,10 @@ Section fwd. (<[3 := (<(Recv, 1) @ (x:Z)> MSG #x ; <(Send, 0)> MSG #x ; END)%proto]> ∅))). - Lemma iProto_fwd_consistent : - ⊢ iProto_consistent iProto_fwd. + Lemma iProto_forwarder_consistent : + ⊢ iProto_consistent iProto_forwarder. Proof. - rewrite /iProto_fwd. + rewrite /iProto_forwarder. iProto_consistent_take_step. iIntros (x) "_". iExists _. iSplit; [done|]. iSplit; [done|]. iProto_consistent_take_step. @@ -576,4 +546,4 @@ Section fwd. iProto_consistent_take_step. Qed. -End fwd. +End forwarder. -- GitLab From 22e341f421adc9d9fbb7d798617e3193f7168848 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Wed, 7 Feb 2024 13:16:26 +0100 Subject: [PATCH 40/81] Fixed _CoqProject bug --- _CoqProject | 2 +- multi_actris/channel/proofmode.v | 2 +- 2 files changed, 2 insertions(+), 2 deletions(-) diff --git a/_CoqProject b/_CoqProject index 98b145e..4489ccb 100644 --- a/_CoqProject +++ b/_CoqProject @@ -60,4 +60,4 @@ multi_actris/channel/proto_model.v multi_actris/channel/proto.v multi_actris/channel/channel.v multi_actris/channel/proofmode.v -multi_actris/channel/proto_consistency_examples.v +multi_actris/examples/basics.v diff --git a/multi_actris/channel/proofmode.v b/multi_actris/channel/proofmode.v index cc8cc20..9e729ed 100644 --- a/multi_actris/channel/proofmode.v +++ b/multi_actris/channel/proofmode.v @@ -11,7 +11,7 @@ From iris.algebra Require Import gmap. From iris.proofmode Require Import coq_tactics reduction spec_patterns. From iris.heap_lang Require Export proofmode notation. From multi_actris.channel Require Import proto_model. -From multi_actris Require Export channel. +From multi_actris.channel Require Export channel. Set Default Proof Using "Type". (** * Tactics for proving contractiveness of protocols *) -- GitLab From d8122f698b622427381e50fa81bf088ed10d3abf Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Thu, 8 Feb 2024 14:46:22 +0100 Subject: [PATCH 41/81] Minor cleanup --- multi_actris/channel/proto.v | 13 ++++++------- 1 file changed, 6 insertions(+), 7 deletions(-) diff --git a/multi_actris/channel/proto.v b/multi_actris/channel/proto.v index 2575c7a..dbb1b16 100644 --- a/multi_actris/channel/proto.v +++ b/multi_actris/channel/proto.v @@ -636,6 +636,11 @@ Proof. solve_contractive. Qed. Global Instance iProto_own_contractive `{protoG Σ V} γ i : Contractive (iProto_own γ i). Proof. solve_contractive. Qed. +Global Instance iProto_own_ne `{protoG Σ V} γ s : NonExpansive (iProto_own γ s). +Proof. solve_proper. Qed. +Global Instance iProto_own_proper `{protoG Σ V} γ s : + Proper ((≡) ==> (≡)) (iProto_own γ s). +Proof. apply (ne_proper _). Qed. (** * Proofs *) Section proto. @@ -889,7 +894,6 @@ Section proto. inversion Htag. simplify_eq. iIntros (v p) "Hm1'". iDestruct (iProto_le_recv_inv_r with "Hle") as "Hle". - (* iRewrite -"Hm2" in "Hm2'". *) iDestruct "Hle" as (m') "[#Heq Hle]". iDestruct ("Hprot" $!i' with "[] [] Hm1'") as "Hprot". { done. } @@ -977,7 +981,7 @@ Section proto. - iApply iProto_le_recv. auto 10 with iFrame. Qed. -Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. + Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. Proof. iIntros "H". iLöb as "IH" forall (p1 p2). destruct (iProto_case p1) as [->|([]&i&m1&->)]. @@ -1173,11 +1177,6 @@ Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. iFrame. rewrite -fmap_insert. done. Qed. - Global Instance iProto_own_ne γ s : NonExpansive (iProto_own γ s). - Proof. solve_proper. Qed. - Global Instance iProto_own_proper γ s : Proper ((≡) ==> (≡)) (iProto_own γ s). - Proof. apply (ne_proper _). Qed. - Lemma iProto_own_auth_alloc ps : ⊢ |==> ∃ γ, iProto_own_auth γ ps ∗ [∗ map] i ↦p ∈ ps, iProto_own γ i p. Proof. -- GitLab From 19bea05beea55fde69c491004405aaac198f0e11 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Thu, 8 Feb 2024 15:59:42 +0100 Subject: [PATCH 42/81] Added another example --- multi_actris/channel/proofmode.v | 7 ++ multi_actris/examples/basics.v | 144 +++++++++++++++++++++++++------ 2 files changed, 125 insertions(+), 26 deletions(-) diff --git a/multi_actris/channel/proofmode.v b/multi_actris/channel/proofmode.v index 9e729ed..07c0449 100644 --- a/multi_actris/channel/proofmode.v +++ b/multi_actris/channel/proofmode.v @@ -456,3 +456,10 @@ Tactic Notation "iProto_consistent_take_step" := Tactic Notation "clean_map" constr(i) := iEval (repeat (rewrite (insert_commute _ _ i); [|done])); rewrite (insert_insert _ i). + +Tactic Notation "iProto_consistent_resolve_step" := + repeat iIntros (?); repeat iIntros "?"; + repeat iExists _; repeat (iSplit; [done|]); try iFrame. + +Tactic Notation "iProto_consistent_take_steps" := + repeat (iProto_consistent_take_step; iProto_consistent_resolve_step). diff --git a/multi_actris/examples/basics.v b/multi_actris/examples/basics.v index 69a94fb..8ff18fa 100644 --- a/multi_actris/examples/basics.v +++ b/multi_actris/examples/basics.v @@ -14,12 +14,7 @@ Definition iProto_binary `{!heapGS Σ} : gmap nat (iProto Σ) := Lemma iProto_binary_consistent `{!heapGS Σ} : ⊢ iProto_consistent iProto_binary. -Proof. - rewrite /iProto_binary. - iProto_consistent_take_step. - iIntros (x) "HP". iExists x. iSplit; [done|]. iFrame. - iProto_consistent_take_step. -Qed. +Proof. rewrite /iProto_binary. iProto_consistent_take_steps. Qed. Definition roundtrip_prog : val := λ: <>, @@ -41,16 +36,7 @@ Section channel. Lemma iProto_roundtrip_consistent : ⊢ iProto_consistent iProto_roundtrip. - Proof. - rewrite /iProto_roundtrip. - iProto_consistent_take_step. - iIntros (x) "_". iExists x. iSplit; [done|]. iSplit; [done|]. - iProto_consistent_take_step. - iIntros "_". iExists x. iSplit; [done|]. iSplit; [done|]. - iProto_consistent_take_step. - iIntros "_". iSplit; [done|]. iSplit; [done|]. - iProto_consistent_take_step. - Qed. + Proof. rewrite /iProto_roundtrip. iProto_consistent_take_steps. Qed. (* TODO: Fix nat/Z coercion. *) Lemma roundtrip_prog_spec : @@ -101,13 +87,8 @@ Section roundtrip_ref. ⊢ iProto_consistent iProto_roundtrip_ref. Proof. rewrite /iProto_roundtrip_ref. - iProto_consistent_take_step. - iIntros (l x) "Hloc". iExists _, _. iSplit; [done|]. iFrame. - iProto_consistent_take_step. - iIntros "Hloc". iExists _, _. iSplit; [done|]. iFrame. - iProto_consistent_take_step. - iIntros "Hloc". iSplit; [done|]. - replace (x + 1 + 1)%Z with (x+2)%Z by lia. iFrame. + iProto_consistent_take_steps. + replace (x0 + 1 + 1)%Z with (x0+2)%Z by lia. iFrame. iProto_consistent_take_step. Qed. @@ -231,7 +212,7 @@ Section roundtrip_ref_rec. replace (x + 1 + 1)%Z with (x+2)%Z by lia. iFrame. rewrite -iProto_roundtrip_ref_rec2_unfold. do 2 clean_map 0. do 2 clean_map 1. do 2 clean_map 2. - done. + iApply "IH". Qed. Lemma roundtrip_ref_rec_prog_spec : @@ -517,7 +498,7 @@ Section forwarder. Definition iProto_forwarder : gmap nat (iProto Σ) := <[0 := (<(Send, 1) @ (x:Z)> MSG #x ; <(Send, 1) @ (b:bool)> MSG #b ; - <(Send, if b then 2 else 3) > MSG #x ; END)%proto]> + <(Recv, if b then 2 else 3) > MSG #x ; END)%proto]> (<[1 := (<(Recv, 0) @ (x:Z)> MSG #x ; <(Recv, 0) @ (b:bool)> MSG #b; if b @@ -539,11 +520,122 @@ Section forwarder. - iExists _. iSplit; [done|]. iSplit; [done|]. iProto_consistent_take_step. iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. + repeat clean_map 0. repeat clean_map 1. + iProto_consistent_take_steps. + - iExists _. iSplit; [done|]. iSplit; [done|]. + iProto_consistent_take_step. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. + repeat clean_map 0. repeat clean_map 1. + iProto_consistent_take_steps. + Qed. + +End forwarder. + +Section forwarder_rec. + Context `{!heapGS Σ}. + + (** + + 0 + / | \ + / | \ + | 1 | + \ / \ / + 2 3 + + *) + + Definition iProto_forwarder_rec_0_aux (rec : iProto Σ) : iProto Σ := + (<(Send, 1) @ (x:Z)> MSG #x ; + <(Send, 1) @ (b:bool)> MSG #b ; + <(Recv, if b then 2 else 3) > MSG #x ; rec)%proto. + Instance iProto_forwarder_rec_0_contractive : + Contractive iProto_forwarder_rec_0_aux. + Proof. solve_proto_contractive. Qed. + Definition iProto_forwarder_rec_0 := + fixpoint iProto_forwarder_rec_0_aux. + Lemma iProto_forwarder_rec_0_unfold : + iProto_forwarder_rec_0 ≡ + (iProto_forwarder_rec_0_aux iProto_forwarder_rec_0). + Proof. apply (fixpoint_unfold _). Qed. + + Definition iProto_forwarder_rec_1_aux (rec : iProto Σ) : iProto Σ := + (<(Recv, 0) @ (x:Z)> MSG #x ; + <(Recv, 0) @ (b:bool)> MSG #b; + if b + then <(Send,2)> MSG #x ; rec + else <(Send,3)> MSG #x ; rec)%proto. + Instance iProto_forwarder_rec_1_contractive : + Contractive iProto_forwarder_rec_1_aux. + Proof. solve_proto_contractive. Qed. + Definition iProto_forwarder_rec_1 := + fixpoint iProto_forwarder_rec_1_aux. + Lemma iProto_forwarder_rec_1_unfold : + iProto_forwarder_rec_1 ≡ + (iProto_forwarder_rec_1_aux iProto_forwarder_rec_1). + Proof. apply (fixpoint_unfold _). Qed. + + Definition iProto_forwarder_rec_n_aux (rec : iProto Σ) : iProto Σ := + (<(Recv, 1) @ (x:Z)> MSG #x ; + <(Send, 0)> MSG #x ; rec)%proto. + Instance iProto_forwarder_rec_n_contractive : + Contractive iProto_forwarder_rec_n_aux. + Proof. solve_proto_contractive. Qed. + Definition iProto_forwarder_rec_n := + fixpoint iProto_forwarder_rec_n_aux. + Lemma iProto_forwarder_rec_n_unfold : + iProto_forwarder_rec_n ≡ + (iProto_forwarder_rec_n_aux iProto_forwarder_rec_n). + Proof. apply (fixpoint_unfold _). Qed. + + Definition iProto_forwarder_rec : gmap nat (iProto Σ) := + <[0 := iProto_forwarder_rec_0]> + (<[1 := iProto_forwarder_rec_1]> + (<[2 := iProto_forwarder_rec_n]> + (<[3 := iProto_forwarder_rec_n]>∅))). + + Lemma iProto_forwarder_rec_consistent : + ⊢ iProto_consistent iProto_forwarder_rec. + Proof. + iLöb as "IH". + rewrite /iProto_forwarder_rec. + iEval (rewrite iProto_forwarder_rec_0_unfold). + iEval (rewrite iProto_forwarder_rec_1_unfold). + iEval (rewrite iProto_forwarder_rec_n_unfold). + iProto_consistent_take_step. + iIntros (x) "_". iExists _. iSplit; [done|]. iSplit; [done|]. + iProto_consistent_take_step. + iIntros ([]) "_". + - iExists _. iSplit; [done|]. iSplit; [done|]. + iProto_consistent_take_step. + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. + repeat clean_map 1. repeat clean_map 2. repeat clean_map 0. + iNext. + iEval (rewrite iProto_forwarder_rec_1_unfold). + iEval (rewrite iProto_forwarder_rec_n_unfold). iProto_consistent_take_step. + iIntros "_". iSplit; [done|]. iSplit; [done|]. + repeat clean_map 1. repeat clean_map 2. repeat clean_map 0. + rewrite (insert_commute _ 2 1); [|done]. + iEval (rewrite -iProto_forwarder_rec_1_unfold). + iEval (rewrite -iProto_forwarder_rec_n_unfold). + iEval (rewrite -iProto_forwarder_rec_n_unfold). + iApply "IH". - iExists _. iSplit; [done|]. iSplit; [done|]. iProto_consistent_take_step. iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. + repeat clean_map 1. repeat clean_map 2. repeat clean_map 0. + iNext. + iEval (rewrite iProto_forwarder_rec_1_unfold). + iEval (rewrite iProto_forwarder_rec_n_unfold). iProto_consistent_take_step. + iIntros "_". iSplit; [done|]. iSplit; [done|]. + repeat clean_map 1. repeat clean_map 3. repeat clean_map 0. + rewrite (insert_commute _ 3 1); [|done]. + iEval (rewrite -iProto_forwarder_rec_1_unfold). + iEval (rewrite -iProto_forwarder_rec_n_unfold). + iEval (rewrite -iProto_forwarder_rec_n_unfold). + iApply "IH". Qed. -End forwarder. +End forwarder_rec. -- GitLab From e90d9120ead5358968f45c60f6e866a972c32417 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Mon, 19 Feb 2024 20:34:15 +0100 Subject: [PATCH 43/81] WIP: Leader electon --- multi_actris/examples/leader_election.v | 210 ++++++++++++++++++++++++ 1 file changed, 210 insertions(+) create mode 100644 multi_actris/examples/leader_election.v diff --git a/multi_actris/examples/leader_election.v b/multi_actris/examples/leader_election.v new file mode 100644 index 0000000..0c1f3da --- /dev/null +++ b/multi_actris/examples/leader_election.v @@ -0,0 +1,210 @@ +From multi_actris.channel Require Import proofmode. +From iris.heap_lang.lib Require Import assert. + +(** Inspired by https://en.wikipedia.org/wiki/Chang_and_Roberts_algorithm *) + +Definition process : val := + rec: "go" "c" "idl" "idr" "id" "id_max" "is_participant" := + if: recv "c" "idr" + then let: "id'" := recv "c" "idr" in + if: "id_max" < "id'" (** Case 1 *) + then send "c" "idl" #true;; send "c" "idl" "id'";; + "go" "c" "idl" "idr" "id" "id'" "is_participant" + else if: "id_max" = "id'" (** Case 4 *) + then send "c" "idl" #false;; send "c" "idl" "id_max";; + "go" "c" "idl" "idr" "id" "id_max" #false + else if: "is_participant" (** Case 3 *) + then "go" "c" "idl" "idr" "id" "id_max" "is_participant" (** Case 4 *) + else send "c" "idl" #true;; send "c" "idl" "id_max";; + "go" "c" "idl" "idr" "id" "id_max" #true + else let: "id'" := recv "c" "idr" in + assert: ("id_max" = "id'");; + if: "id" = "id'" then "id'" + else send "c" "idl" #false;; send "c" "idl" "id_max";; "id_max". + +Definition init : val := + λ: "c" "idl" "idr" "id", + (* Notice missing leader *) + send "c" "idl" #true;; send "c" "idl" "id";; + process "c" "idl" "idr" "id" "id" #true. + +Definition program : val := + λ: <>, + let: "cs" := new_chan #4 in + let: "c0" := get_chan "cs" #0 in + let: "c1" := get_chan "cs" #1 in + let: "c2" := get_chan "cs" #2 in + let: "c3" := get_chan "cs" #3 in + Fork (let: "id_max" := init "c1" #3 #2 #1 in send "c1" #0 "id_max");; + Fork (let: "id_max" := process "c2" #1 #3 #2 #2 #false in + send "c2" #0 "id_max");; + Fork (let: "id_max" := init "c3" #2 #1 #3 in send "c3" #0 "id_max");; + let: "res1" := recv "c0" #1 in + let: "res2" := recv "c0" #2 in + let: "res3" := recv "c0" #3 in + assert: ("res1" = "res2");; + assert: ("res2" = "res3"). + +Notation iProto_choice a p1 p2 := + (<a @ (b : bool)> MSG #b; if b then p1 else p2)%proto. + +Section ring_leader_election_example. + Context `{!heapGS Σ, chanG Σ, spawnG Σ, mono_natG Σ}. + + Definition my_send_prot (il ir i : nat) i' + (rec : bool -d> nat -d> iProto Σ) : bool -d> nat -d> iProto Σ := + λ (isp:bool) (i_max:nat), + iProto_choice (Send,il) + (<(Send,il)> MSG #(max i' i_max) ; rec isp (max i' i_max))%proto + (<(Send,il)> MSG #i_max ; rec isp i_max)%proto. + + Definition my_recv_prot (il ir i : nat) (p : nat → iProto Σ) + (rec : bool -d> nat -d> iProto Σ) : bool -d> nat -d> iProto Σ := + λ (isp:bool) (i_max:nat), + iProto_choice (Recv,ir) + (<(Recv,ir) @ (i':nat)> MSG #i' ; + if bool_decide (i_max < i') + then my_send_prot il ir i i' rec isp i_max + else if bool_decide (i_max = i') + then my_send_prot il ir i i' rec false i_max + else if isp then rec isp i_max + else my_send_prot il ir i i' rec true i_max)%proto + (<(Recv,ir)> MSG #i_max ; + if (bool_decide (i = i_max)) then p i_max + else <(Send,il)> MSG #false; <(Send,il)> MSG #i_max; p i_max)%proto. + + Instance rle_prot_aux_contractive il ir i p : Contractive (my_recv_prot il ir i p). + Proof. rewrite /my_recv_prot /my_send_prot. solve_proto_contractive. Qed. + Definition rle_prot il ir i p := fixpoint (my_recv_prot il ir i p). + Instance rle_prot_unfold il ir i isp max_id p : + ProtoUnfold (rle_prot il ir i p isp max_id) (my_recv_prot il ir i p (rle_prot il ir i p) isp max_id). + Proof. apply proto_unfold_eq, (fixpoint_unfold (my_recv_prot il ir i p)). Qed. + + Definition rle_preprot (il ir i : nat) p : iProto Σ := + (<(Send, il)> MSG #true; <(Send, il)> MSG #i ; + rle_prot il ir i p true i)%proto. + + Lemma process_spec il ir i p i_max c (isp:bool) : + i <= i_max → + {{{ c ↣ (rle_prot il ir i p isp i_max) }}} + process c #il #ir #i #i_max #isp + {{{ i', RET #i'; c ↣ p i' }}}. + Proof. + iIntros (Hle Φ) "Hc HΦ". + iLöb as "IH" forall (Φ isp i_max Hle). + wp_lam. wp_recv (b) as "_". + destruct b. + - wp_pures. wp_recv (i') as "_". + wp_pures. + case_bool_decide as Hlt. + { case_bool_decide; [|lia]. + wp_pures. wp_send with "[//]". + iEval (replace i' with (max i' i_max) by lia). + wp_send with "[//]". wp_pures. + iApply ("IH" with "[] Hc HΦ"). iPureIntro. lia. } + case_bool_decide as Hlt2. + { case_bool_decide; [lia|]. + wp_pures. case_bool_decide; [|simplify_eq;lia]. + wp_send with "[//]". + iEval (replace i' with (max i' i_max) by lia). + wp_send with "[//]". wp_pures. + iApply ("IH" with "[] Hc HΦ"). iPureIntro. lia. } + case_bool_decide; [lia|]. + wp_pures. + case_bool_decide; [simplify_eq;lia|]. + wp_pures. + destruct isp. + { wp_pures. iApply ("IH" with "[] Hc HΦ"). iPureIntro. lia. } + wp_pures. + rewrite /my_send_prot. + wp_send with "[//]". + iEval (replace i_max with (max i' i_max) by lia). + wp_send with "[//]". + wp_pures. iApply ("IH" with "[] Hc HΦ"). iPureIntro. lia. + - wp_pures. + wp_recv as "_". + wp_smart_apply wp_assert. + wp_pures. iModIntro. + iSplitR. + { iPureIntro. f_equiv. f_equiv. apply bool_decide_eq_true_2. done. } + iNext. + wp_pures. + case_bool_decide. + { wp_pures. simplify_eq. + case_bool_decide; [|simplify_eq;lia]. wp_pures. + iModIntro. by iApply "HΦ". } + wp_pures. + case_bool_decide; [simplify_eq;lia|]. wp_pures. + wp_send with "[//]". + wp_send with "[//]". + wp_pures. by iApply "HΦ". + Qed. + + Lemma init_spec c il ir i p : + {{{ c ↣ rle_preprot il ir i p }}} + init c #il #ir #i + {{{ res, RET #res; c ↣ p res }}}. + Proof. + iIntros (Φ) "Hc HΦ". wp_lam. wp_send with "[//]". wp_send with "[//]". + wp_pures. by iApply (process_spec with "Hc HΦ"). + Qed. + + Definition prot_tail (i_max : nat) : iProto Σ := + (<(Send,0)> MSG #i_max; END)%proto. + + Definition prot_pool : gmap nat (iProto Σ) := + <[0 := (<(Recv,1) @ (id_max : nat)> MSG #id_max ; + <(Recv,2)> MSG #id_max ; + <(Recv,3)> MSG #id_max ; + END)%proto ]> + (<[1 := rle_preprot 3 2 1 prot_tail ]> + (<[2 := rle_prot 1 3 2 prot_tail false 2 ]> + (<[3 := rle_preprot 2 1 3 prot_tail ]> ∅))). + + Lemma rle_prot_unfold' il ir i isp max_id p : + (rle_prot il ir i p isp max_id) ≡ (my_recv_prot il ir i p (rle_prot il ir i p) isp max_id). + Proof. apply (fixpoint_unfold (my_recv_prot il ir i p)). Qed. + + Lemma prot_pool_consistent : ⊢ iProto_consistent prot_pool. + Proof. Admitted. + + Lemma program_spec : + {{{ True }}} + program #() + {{{ RET #(); True }}}. + Proof. + iIntros (Φ) "_ HΦ". + wp_lam. + wp_smart_apply (new_chan_spec 4 prot_pool); + [lia|set_solver|iApply prot_pool_consistent|]. + iIntros (cs) "Hcs". + wp_smart_apply (get_chan_spec _ 0 with "Hcs"); [done|]. + iIntros (c0) "[Hc0 Hcs]". + wp_smart_apply (get_chan_spec _ 1 with "Hcs"); [done|]. + iIntros (c1) "[Hc1 Hcs]". + wp_smart_apply (get_chan_spec _ 2 with "Hcs"); [done|]. + iIntros (c2) "[Hc2 Hcs]". + wp_smart_apply (get_chan_spec _ 3 with "Hcs"); [done|]. + iIntros (c3) "[Hc3 Hcs]". + wp_smart_apply (wp_fork with "[Hc1]"). + { iIntros "!>". wp_smart_apply (init_spec with "Hc1"). + iIntros (i') "Hc1". wp_send with "[//]". done. } + wp_smart_apply (wp_fork with "[Hc2]"). + { iIntros "!>". wp_smart_apply (process_spec with "Hc2"); [lia|]. + iIntros (i') "Hc2". wp_pures. wp_send with "[//]". done. } + wp_smart_apply (wp_fork with "[Hc3]"). + { iIntros "!>". wp_smart_apply (init_spec with "Hc3"). + iIntros (i') "Hc3". wp_send with "[//]". done. } + wp_recv (id_max) as "_". + wp_recv as "_". + wp_recv as "_". + wp_smart_apply wp_assert. wp_pures. iModIntro. iSplitR; [iPureIntro|]. + { do 2 f_equiv. apply bool_decide_eq_true_2. done. } + iIntros "!>". + wp_smart_apply wp_assert. wp_pures. iModIntro. iSplitR; [iPureIntro|]. + { do 2 f_equiv. apply bool_decide_eq_true_2. done. } + iIntros "!>". + by iApply "HΦ". + Qed. + +End ring_leader_election_example. -- GitLab From c70cb821355bed1b82a83d4181dfd1600a38a8f5 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Tue, 20 Feb 2024 11:14:12 +0100 Subject: [PATCH 44/81] Closed proof --- multi_actris/examples/leader_election.v | 80 ++++++++++++++++++++----- 1 file changed, 65 insertions(+), 15 deletions(-) diff --git a/multi_actris/examples/leader_election.v b/multi_actris/examples/leader_election.v index 0c1f3da..9a21b17 100644 --- a/multi_actris/examples/leader_election.v +++ b/multi_actris/examples/leader_election.v @@ -51,30 +51,25 @@ Notation iProto_choice a p1 p2 := Section ring_leader_election_example. Context `{!heapGS Σ, chanG Σ, spawnG Σ, mono_natG Σ}. - Definition my_send_prot (il ir i : nat) i' - (rec : bool -d> nat -d> iProto Σ) : bool -d> nat -d> iProto Σ := - λ (isp:bool) (i_max:nat), - iProto_choice (Send,il) - (<(Send,il)> MSG #(max i' i_max) ; rec isp (max i' i_max))%proto - (<(Send,il)> MSG #i_max ; rec isp i_max)%proto. - Definition my_recv_prot (il ir i : nat) (p : nat → iProto Σ) (rec : bool -d> nat -d> iProto Σ) : bool -d> nat -d> iProto Σ := λ (isp:bool) (i_max:nat), iProto_choice (Recv,ir) (<(Recv,ir) @ (i':nat)> MSG #i' ; if bool_decide (i_max < i') - then my_send_prot il ir i i' rec isp i_max + then <(Send,il)> MSG #true ; <(Send,il)> MSG #(max i' i_max) ; + rec isp (max i' i_max) else if bool_decide (i_max = i') - then my_send_prot il ir i i' rec false i_max + then <(Send,il)> MSG #false ; <(Send, il)> MSG #i_max ; rec false i_max else if isp then rec isp i_max - else my_send_prot il ir i i' rec true i_max)%proto + else <(Send,il)> MSG #true ; <(Send,il)> MSG #(max i' i_max) ; + rec true (max i' i_max))%proto (<(Recv,ir)> MSG #i_max ; if (bool_decide (i = i_max)) then p i_max else <(Send,il)> MSG #false; <(Send,il)> MSG #i_max; p i_max)%proto. Instance rle_prot_aux_contractive il ir i p : Contractive (my_recv_prot il ir i p). - Proof. rewrite /my_recv_prot /my_send_prot. solve_proto_contractive. Qed. + Proof. rewrite /my_recv_prot. solve_proto_contractive. Qed. Definition rle_prot il ir i p := fixpoint (my_recv_prot il ir i p). Instance rle_prot_unfold il ir i isp max_id p : ProtoUnfold (rle_prot il ir i p isp max_id) (my_recv_prot il ir i p (rle_prot il ir i p) isp max_id). @@ -100,7 +95,7 @@ Section ring_leader_election_example. { case_bool_decide; [|lia]. wp_pures. wp_send with "[//]". iEval (replace i' with (max i' i_max) by lia). - wp_send with "[//]". wp_pures. + wp_send with "[//]". wp_pures. iApply ("IH" with "[] Hc HΦ"). iPureIntro. lia. } case_bool_decide as Hlt2. { case_bool_decide; [lia|]. @@ -116,7 +111,6 @@ Section ring_leader_election_example. destruct isp. { wp_pures. iApply ("IH" with "[] Hc HΦ"). iPureIntro. lia. } wp_pures. - rewrite /my_send_prot. wp_send with "[//]". iEval (replace i_max with (max i' i_max) by lia). wp_send with "[//]". @@ -166,8 +160,64 @@ Section ring_leader_election_example. Proof. apply (fixpoint_unfold (my_recv_prot il ir i p)). Qed. Lemma prot_pool_consistent : ⊢ iProto_consistent prot_pool. - Proof. Admitted. - + Proof. + rewrite /prot_pool /rle_preprot. + rewrite !rle_prot_unfold'. + iProto_consistent_take_step. + iProto_consistent_resolve_step. + iProto_consistent_take_step. + iProto_consistent_resolve_step. + repeat clean_map 0. repeat clean_map 1. + repeat clean_map 2. repeat clean_map 3. + iProto_consistent_take_step. + iProto_consistent_resolve_step. + iProto_consistent_take_step. + iProto_consistent_resolve_step. + repeat clean_map 0. repeat clean_map 1. + repeat clean_map 2. repeat clean_map 3. + rewrite !rle_prot_unfold'. + iProto_consistent_take_step. + iProto_consistent_resolve_step. + iProto_consistent_take_step. + iProto_consistent_resolve_step. + repeat clean_map 0. repeat clean_map 1. + repeat clean_map 2. repeat clean_map 3. + iProto_consistent_take_step. + iProto_consistent_resolve_step. + iProto_consistent_take_step. + iProto_consistent_resolve_step. + repeat clean_map 0. repeat clean_map 1. + repeat clean_map 2. repeat clean_map 3. + rewrite !rle_prot_unfold'. + iProto_consistent_take_step. + iProto_consistent_resolve_step. + iProto_consistent_take_step. + iProto_consistent_resolve_step. + repeat clean_map 0. repeat clean_map 1. + repeat clean_map 2. repeat clean_map 3. + iProto_consistent_take_step. + iProto_consistent_resolve_step. + iProto_consistent_take_step. + iProto_consistent_resolve_step. + repeat clean_map 0. repeat clean_map 1. + repeat clean_map 2. repeat clean_map 3. + iProto_consistent_take_step. + iProto_consistent_resolve_step. + iProto_consistent_take_step. + iProto_consistent_resolve_step. + repeat clean_map 0. repeat clean_map 1. + repeat clean_map 2. repeat clean_map 3. + iProto_consistent_take_step. + iProto_consistent_resolve_step. + iProto_consistent_take_step. + iProto_consistent_resolve_step. + repeat clean_map 0. repeat clean_map 1. + repeat clean_map 2. repeat clean_map 3. + iProto_consistent_take_step. + iProto_consistent_resolve_step. + iProto_consistent_take_step. + Qed. + Lemma program_spec : {{{ True }}} program #() -- GitLab From b265d4e15de48a20f18fd829aa8c3ea6ee2f415f Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Tue, 20 Feb 2024 11:21:07 +0100 Subject: [PATCH 45/81] Typo --- multi_actris/examples/leader_election.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/multi_actris/examples/leader_election.v b/multi_actris/examples/leader_election.v index 9a21b17..05ea6a3 100644 --- a/multi_actris/examples/leader_election.v +++ b/multi_actris/examples/leader_election.v @@ -14,7 +14,7 @@ Definition process : val := then send "c" "idl" #false;; send "c" "idl" "id_max";; "go" "c" "idl" "idr" "id" "id_max" #false else if: "is_participant" (** Case 3 *) - then "go" "c" "idl" "idr" "id" "id_max" "is_participant" (** Case 4 *) + then "go" "c" "idl" "idr" "id" "id_max" "is_participant" (** Case 2 *) else send "c" "idl" #true;; send "c" "idl" "id_max";; "go" "c" "idl" "idr" "id" "id_max" #true else let: "id'" := recv "c" "idr" in -- GitLab From 206ac1d7ef042ae8b417761d8f5ac26bb0de0279 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Tue, 20 Feb 2024 14:18:37 +0100 Subject: [PATCH 46/81] Refactoring --- multi_actris/examples/leader_election.v | 68 ++++++++++++------------- 1 file changed, 34 insertions(+), 34 deletions(-) diff --git a/multi_actris/examples/leader_election.v b/multi_actris/examples/leader_election.v index 05ea6a3..d881432 100644 --- a/multi_actris/examples/leader_election.v +++ b/multi_actris/examples/leader_election.v @@ -74,6 +74,10 @@ Section ring_leader_election_example. Instance rle_prot_unfold il ir i isp max_id p : ProtoUnfold (rle_prot il ir i p isp max_id) (my_recv_prot il ir i p (rle_prot il ir i p) isp max_id). Proof. apply proto_unfold_eq, (fixpoint_unfold (my_recv_prot il ir i p)). Qed. + Lemma rle_prot_unfold' il ir i isp max_id p : + (rle_prot il ir i p isp max_id) ≡ + (my_recv_prot il ir i p (rle_prot il ir i p) isp max_id). + Proof. apply (fixpoint_unfold (my_recv_prot il ir i p)). Qed. Definition rle_preprot (il ir i : nat) p : iProto Σ := (<(Send, il)> MSG #true; <(Send, il)> MSG #i ; @@ -127,11 +131,8 @@ Section ring_leader_election_example. { wp_pures. simplify_eq. case_bool_decide; [|simplify_eq;lia]. wp_pures. iModIntro. by iApply "HΦ". } - wp_pures. - case_bool_decide; [simplify_eq;lia|]. wp_pures. - wp_send with "[//]". - wp_send with "[//]". - wp_pures. by iApply "HΦ". + wp_pures. case_bool_decide; [simplify_eq;lia|]. wp_pures. + wp_send with "[//]". wp_send with "[//]". wp_pures. by iApply "HΦ". Qed. Lemma init_spec c il ir i p : @@ -146,6 +147,19 @@ Section ring_leader_election_example. Definition prot_tail (i_max : nat) : iProto Σ := (<(Send,0)> MSG #i_max; END)%proto. + Definition pre_prot_pool id_max : gmap nat (iProto Σ) := + <[0 := (<(Recv,1) @ (id_max : nat)> MSG #id_max ; + <(Recv,2)> MSG #id_max ; + <(Recv,3)> MSG #id_max ; + END)%proto ]> + (<[1 := prot_tail id_max ]> + (<[2 := prot_tail id_max ]> + (<[3 := prot_tail id_max ]> ∅))). + + Lemma pre_prot_pool_consistent id_max : + ⊢ iProto_consistent (pre_prot_pool id_max). + Proof. rewrite /pre_prot_pool. iProto_consistent_take_steps. Qed. + Definition prot_pool : gmap nat (iProto Σ) := <[0 := (<(Recv,1) @ (id_max : nat)> MSG #id_max ; <(Recv,2)> MSG #id_max ; @@ -155,12 +169,8 @@ Section ring_leader_election_example. (<[2 := rle_prot 1 3 2 prot_tail false 2 ]> (<[3 := rle_preprot 2 1 3 prot_tail ]> ∅))). - Lemma rle_prot_unfold' il ir i isp max_id p : - (rle_prot il ir i p isp max_id) ≡ (my_recv_prot il ir i p (rle_prot il ir i p) isp max_id). - Proof. apply (fixpoint_unfold (my_recv_prot il ir i p)). Qed. - Lemma prot_pool_consistent : ⊢ iProto_consistent prot_pool. - Proof. + Proof. rewrite /prot_pool /rle_preprot. rewrite !rle_prot_unfold'. iProto_consistent_take_step. @@ -207,24 +217,17 @@ Section ring_leader_election_example. iProto_consistent_resolve_step. repeat clean_map 0. repeat clean_map 1. repeat clean_map 2. repeat clean_map 3. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - repeat clean_map 0. repeat clean_map 1. - repeat clean_map 2. repeat clean_map 3. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - iProto_consistent_take_step. + repeat (rewrite (insert_commute _ _ 3); [|lia]). + repeat (rewrite (insert_commute _ _ 2); [|lia]). + repeat (rewrite (insert_commute _ _ 1); [|lia]). + repeat (rewrite (insert_commute _ _ 0); [|lia]). + iApply pre_prot_pool_consistent. Qed. Lemma program_spec : - {{{ True }}} - program #() - {{{ RET #(); True }}}. + {{{ True }}} program #() {{{ RET #(); True }}}. Proof. - iIntros (Φ) "_ HΦ". - wp_lam. + iIntros (Φ) "_ HΦ". wp_lam. wp_smart_apply (new_chan_spec 4 prot_pool); [lia|set_solver|iApply prot_pool_consistent|]. iIntros (cs) "Hcs". @@ -238,23 +241,20 @@ Section ring_leader_election_example. iIntros (c3) "[Hc3 Hcs]". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". wp_smart_apply (init_spec with "Hc1"). - iIntros (i') "Hc1". wp_send with "[//]". done. } + iIntros (i') "Hc1". by wp_send with "[//]". } wp_smart_apply (wp_fork with "[Hc2]"). { iIntros "!>". wp_smart_apply (process_spec with "Hc2"); [lia|]. - iIntros (i') "Hc2". wp_pures. wp_send with "[//]". done. } + iIntros (i') "Hc2". by wp_send with "[//]". } wp_smart_apply (wp_fork with "[Hc3]"). { iIntros "!>". wp_smart_apply (init_spec with "Hc3"). - iIntros (i') "Hc3". wp_send with "[//]". done. } - wp_recv (id_max) as "_". - wp_recv as "_". - wp_recv as "_". + iIntros (i') "Hc3". by wp_send with "[//]". } + wp_recv (id_max) as "_". wp_recv as "_". wp_recv as "_". wp_smart_apply wp_assert. wp_pures. iModIntro. iSplitR; [iPureIntro|]. - { do 2 f_equiv. apply bool_decide_eq_true_2. done. } + { do 2 f_equiv. by apply bool_decide_eq_true_2. } iIntros "!>". wp_smart_apply wp_assert. wp_pures. iModIntro. iSplitR; [iPureIntro|]. - { do 2 f_equiv. apply bool_decide_eq_true_2. done. } - iIntros "!>". - by iApply "HΦ". + { do 2 f_equiv. by apply bool_decide_eq_true_2. } + iIntros "!>". by iApply "HΦ". Qed. End ring_leader_election_example. -- GitLab From 7b66c5329fc412a683f53ffaf34a2c55a87deb7b Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Tue, 20 Feb 2024 14:23:24 +0100 Subject: [PATCH 47/81] Refactoring --- multi_actris/examples/leader_election.v | 64 ++++++++++++------------- 1 file changed, 32 insertions(+), 32 deletions(-) diff --git a/multi_actris/examples/leader_election.v b/multi_actris/examples/leader_election.v index d881432..c0a16a8 100644 --- a/multi_actris/examples/leader_election.v +++ b/multi_actris/examples/leader_election.v @@ -4,29 +4,29 @@ From iris.heap_lang.lib Require Import assert. (** Inspired by https://en.wikipedia.org/wiki/Chang_and_Roberts_algorithm *) Definition process : val := - rec: "go" "c" "idl" "idr" "id" "id_max" "is_participant" := + rec: "go" "c" "idl" "id" "idr" "isp" "id_max" := if: recv "c" "idr" then let: "id'" := recv "c" "idr" in if: "id_max" < "id'" (** Case 1 *) then send "c" "idl" #true;; send "c" "idl" "id'";; - "go" "c" "idl" "idr" "id" "id'" "is_participant" + "go" "c" "idl" "id" "idr" "isp" "id'" else if: "id_max" = "id'" (** Case 4 *) then send "c" "idl" #false;; send "c" "idl" "id_max";; - "go" "c" "idl" "idr" "id" "id_max" #false - else if: "is_participant" (** Case 3 *) - then "go" "c" "idl" "idr" "id" "id_max" "is_participant" (** Case 2 *) + "go" "c" "idl" "id" "idr" #false "id_max" + else if: "isp" (** Case 3 *) + then "go" "c" "idl" "id" "idr" "isp" "id_max" (** Case 2 *) else send "c" "idl" #true;; send "c" "idl" "id_max";; - "go" "c" "idl" "idr" "id" "id_max" #true + "go" "c" "idl" "id" "idr" #true "id_max" else let: "id'" := recv "c" "idr" in assert: ("id_max" = "id'");; if: "id" = "id'" then "id'" else send "c" "idl" #false;; send "c" "idl" "id_max";; "id_max". Definition init : val := - λ: "c" "idl" "idr" "id", + λ: "c" "idl" "id" "idr", (* Notice missing leader *) send "c" "idl" #true;; send "c" "idl" "id";; - process "c" "idl" "idr" "id" "id" #true. + process "c" "idl" "id" "idr" #true "id". Definition program : val := λ: <>, @@ -35,10 +35,10 @@ Definition program : val := let: "c1" := get_chan "cs" #1 in let: "c2" := get_chan "cs" #2 in let: "c3" := get_chan "cs" #3 in - Fork (let: "id_max" := init "c1" #3 #2 #1 in send "c1" #0 "id_max");; - Fork (let: "id_max" := process "c2" #1 #3 #2 #2 #false in + Fork (let: "id_max" := init "c1" #3 #1 #2 in send "c1" #0 "id_max");; + Fork (let: "id_max" := process "c2" #1 #2 #3 #false #2 in send "c2" #0 "id_max");; - Fork (let: "id_max" := init "c3" #2 #1 #3 in send "c3" #0 "id_max");; + Fork (let: "id_max" := init "c3" #2 #3 #1 in send "c3" #0 "id_max");; let: "res1" := recv "c0" #1 in let: "res2" := recv "c0" #2 in let: "res3" := recv "c0" #3 in @@ -51,7 +51,7 @@ Notation iProto_choice a p1 p2 := Section ring_leader_election_example. Context `{!heapGS Σ, chanG Σ, spawnG Σ, mono_natG Σ}. - Definition my_recv_prot (il ir i : nat) (p : nat → iProto Σ) + Definition my_recv_prot (il i ir : nat) (p : nat → iProto Σ) (rec : bool -d> nat -d> iProto Σ) : bool -d> nat -d> iProto Σ := λ (isp:bool) (i_max:nat), iProto_choice (Recv,ir) @@ -68,25 +68,25 @@ Section ring_leader_election_example. if (bool_decide (i = i_max)) then p i_max else <(Send,il)> MSG #false; <(Send,il)> MSG #i_max; p i_max)%proto. - Instance rle_prot_aux_contractive il ir i p : Contractive (my_recv_prot il ir i p). + Instance rle_prot_aux_contractive il i ir p : Contractive (my_recv_prot il i ir p). Proof. rewrite /my_recv_prot. solve_proto_contractive. Qed. - Definition rle_prot il ir i p := fixpoint (my_recv_prot il ir i p). - Instance rle_prot_unfold il ir i isp max_id p : - ProtoUnfold (rle_prot il ir i p isp max_id) (my_recv_prot il ir i p (rle_prot il ir i p) isp max_id). - Proof. apply proto_unfold_eq, (fixpoint_unfold (my_recv_prot il ir i p)). Qed. - Lemma rle_prot_unfold' il ir i isp max_id p : - (rle_prot il ir i p isp max_id) ≡ - (my_recv_prot il ir i p (rle_prot il ir i p) isp max_id). - Proof. apply (fixpoint_unfold (my_recv_prot il ir i p)). Qed. + Definition rle_prot il i ir p := fixpoint (my_recv_prot il i ir p). + Instance rle_prot_unfold il i ir isp max_id p : + ProtoUnfold (rle_prot il i ir p isp max_id) (my_recv_prot il i ir p (rle_prot il i ir p) isp max_id). + Proof. apply proto_unfold_eq, (fixpoint_unfold (my_recv_prot il i ir p)). Qed. + Lemma rle_prot_unfold' il i ir isp max_id p : + (rle_prot il i ir p isp max_id) ≡ + (my_recv_prot il i ir p (rle_prot il i ir p) isp max_id). + Proof. apply (fixpoint_unfold (my_recv_prot il i ir p)). Qed. - Definition rle_preprot (il ir i : nat) p : iProto Σ := + Definition rle_preprot (il i ir : nat) p : iProto Σ := (<(Send, il)> MSG #true; <(Send, il)> MSG #i ; - rle_prot il ir i p true i)%proto. + rle_prot il i ir p true i)%proto. - Lemma process_spec il ir i p i_max c (isp:bool) : + Lemma process_spec il i ir p i_max c (isp:bool) : i <= i_max → - {{{ c ↣ (rle_prot il ir i p isp i_max) }}} - process c #il #ir #i #i_max #isp + {{{ c ↣ (rle_prot il i ir p isp i_max) }}} + process c #il #i #ir #isp #i_max {{{ i', RET #i'; c ↣ p i' }}}. Proof. iIntros (Hle Φ) "Hc HΦ". @@ -135,9 +135,9 @@ Section ring_leader_election_example. wp_send with "[//]". wp_send with "[//]". wp_pures. by iApply "HΦ". Qed. - Lemma init_spec c il ir i p : - {{{ c ↣ rle_preprot il ir i p }}} - init c #il #ir #i + Lemma init_spec c il i ir p : + {{{ c ↣ rle_preprot il i ir p }}} + init c #il #i #ir {{{ res, RET #res; c ↣ p res }}}. Proof. iIntros (Φ) "Hc HΦ". wp_lam. wp_send with "[//]". wp_send with "[//]". @@ -165,9 +165,9 @@ Section ring_leader_election_example. <(Recv,2)> MSG #id_max ; <(Recv,3)> MSG #id_max ; END)%proto ]> - (<[1 := rle_preprot 3 2 1 prot_tail ]> - (<[2 := rle_prot 1 3 2 prot_tail false 2 ]> - (<[3 := rle_preprot 2 1 3 prot_tail ]> ∅))). + (<[1 := rle_preprot 3 1 2 prot_tail ]> + (<[2 := rle_prot 1 2 3 prot_tail false 2 ]> + (<[3 := rle_preprot 2 3 1 prot_tail ]> ∅))). Lemma prot_pool_consistent : ⊢ iProto_consistent prot_pool. Proof. -- GitLab From 1833321542e204ae50829280b0b228a11b1d03d4 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Tue, 20 Feb 2024 16:05:50 +0100 Subject: [PATCH 48/81] Simplified lemma --- multi_actris/examples/leader_election.v | 15 +++++++-------- 1 file changed, 7 insertions(+), 8 deletions(-) diff --git a/multi_actris/examples/leader_election.v b/multi_actris/examples/leader_election.v index c0a16a8..a464555 100644 --- a/multi_actris/examples/leader_election.v +++ b/multi_actris/examples/leader_election.v @@ -84,13 +84,12 @@ Section ring_leader_election_example. rle_prot il i ir p true i)%proto. Lemma process_spec il i ir p i_max c (isp:bool) : - i <= i_max → {{{ c ↣ (rle_prot il i ir p isp i_max) }}} process c #il #i #ir #isp #i_max {{{ i', RET #i'; c ↣ p i' }}}. Proof. - iIntros (Hle Φ) "Hc HΦ". - iLöb as "IH" forall (Φ isp i_max Hle). + iIntros (Φ) "Hc HΦ". + iLöb as "IH" forall (Φ isp i_max). wp_lam. wp_recv (b) as "_". destruct b. - wp_pures. wp_recv (i') as "_". @@ -100,25 +99,25 @@ Section ring_leader_election_example. wp_pures. wp_send with "[//]". iEval (replace i' with (max i' i_max) by lia). wp_send with "[//]". wp_pures. - iApply ("IH" with "[] Hc HΦ"). iPureIntro. lia. } + iApply ("IH" with "Hc HΦ"). } case_bool_decide as Hlt2. { case_bool_decide; [lia|]. wp_pures. case_bool_decide; [|simplify_eq;lia]. wp_send with "[//]". iEval (replace i' with (max i' i_max) by lia). wp_send with "[//]". wp_pures. - iApply ("IH" with "[] Hc HΦ"). iPureIntro. lia. } + iApply ("IH" with "Hc HΦ"). } case_bool_decide; [lia|]. wp_pures. case_bool_decide; [simplify_eq;lia|]. wp_pures. destruct isp. - { wp_pures. iApply ("IH" with "[] Hc HΦ"). iPureIntro. lia. } + { wp_pures. iApply ("IH" with "Hc HΦ"). } wp_pures. wp_send with "[//]". iEval (replace i_max with (max i' i_max) by lia). wp_send with "[//]". - wp_pures. iApply ("IH" with "[] Hc HΦ"). iPureIntro. lia. + wp_pures. iApply ("IH" with "Hc HΦ"). - wp_pures. wp_recv as "_". wp_smart_apply wp_assert. @@ -243,7 +242,7 @@ Section ring_leader_election_example. { iIntros "!>". wp_smart_apply (init_spec with "Hc1"). iIntros (i') "Hc1". by wp_send with "[//]". } wp_smart_apply (wp_fork with "[Hc2]"). - { iIntros "!>". wp_smart_apply (process_spec with "Hc2"); [lia|]. + { iIntros "!>". wp_smart_apply (process_spec with "Hc2"). iIntros (i') "Hc2". by wp_send with "[//]". } wp_smart_apply (wp_fork with "[Hc3]"). { iIntros "!>". wp_smart_apply (init_spec with "Hc3"). -- GitLab From 179c1a13f7665d8ec7bccb9a9ee8b93c5ce69f8e Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Tue, 20 Feb 2024 23:11:56 +0100 Subject: [PATCH 49/81] Added smuggle example --- multi_actris/examples/basics.v | 61 ++++++++++++++++++++++++++++++---- 1 file changed, 54 insertions(+), 7 deletions(-) diff --git a/multi_actris/examples/basics.v b/multi_actris/examples/basics.v index 8ff18fa..e505c2e 100644 --- a/multi_actris/examples/basics.v +++ b/multi_actris/examples/basics.v @@ -80,8 +80,7 @@ Section roundtrip_ref. (<[1 := (<(Recv, 0) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; <(Send, 2)> MSG #l {{ l ↦ #(x+1) }}; END)%proto]> (<[2 := (<(Recv, 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; - <(Send, 0)> MSG #() {{ l ↦ #(x+1) }}; END)%proto]> - ∅)). + <(Send, 0)> MSG #() {{ l ↦ #(x+1) }}; END)%proto]> ∅)). Lemma iProto_roundtrip_ref_consistent : ⊢ iProto_consistent iProto_roundtrip_ref. @@ -245,6 +244,58 @@ Section roundtrip_ref_rec. End roundtrip_ref_rec. +Definition smuggle_ref_prog : val := + λ: <>, + let: "cs" := new_chan #3 in + let: "c0" := get_chan "cs" #0 in + let: "c1" := get_chan "cs" #1 in + let: "c2" := get_chan "cs" #2 in + Fork (send "c1" #2 (recv "c1" #0));; + Fork (let: "l" := recv "c2" #1 in "l" <- !"l" + #2;; send "c2" #0 #());; + let: "l" := ref #40 in + send "c0" #1 "l";; recv "c0" #2;; !"l". + +Section smuggle_ref. + Context `{!heapGS Σ, !chanG Σ}. + + Definition iProto_smuggle_ref : gmap nat (iProto Σ) := + <[0 := (<(Send, 1) @ (l:loc) (x:Z)> MSG #l {{ l ↦ #x }} ; + <(Recv,2)> MSG #() {{ l ↦ #(x+2) }} ; END)%proto]> + (<[1 := (<(Recv, 0) @ (v:val)> MSG v ; <(Send,2)> MSG v ; END)%proto]> + (<[2 := (<(Recv, 1) @ (l:loc) (x:Z)> MSG #l {{ l ↦ #x }} ; + <(Send,0)> MSG #() {{ l ↦ #(x+2) }} ; END)%proto]> ∅)). + + Lemma iProto_smuggle_ref_consistent : + ⊢ iProto_consistent iProto_smuggle_ref. + Proof. rewrite /iProto_smuggle_ref. iProto_consistent_take_steps. Qed. + + Lemma smuggle_ref_spec : + {{{ True }}} smuggle_ref_prog #() {{{ RET #42 ; True }}}. + Proof using chanG0 heapGS0 Σ. + iIntros (Φ) "_ HΦ". wp_lam. + wp_smart_apply (new_chan_spec 3 iProto_smuggle_ref with "[]"). + { lia. } + { set_solver. } + { iApply iProto_smuggle_ref_consistent. } + iIntros (cs) "Hcs". + wp_pures. + wp_smart_apply (get_chan_spec _ 0 with "Hcs"); [set_solver|]. + iIntros (c0) "[Hc0 Hcs]". + wp_smart_apply (get_chan_spec _ 1 with "Hcs"); [set_solver|]. + iIntros (c1) "[Hc1 Hcs]". + wp_smart_apply (get_chan_spec _ 2 with "Hcs"); [set_solver|]. + iIntros (c2) "[Hc2 Hcs]". + wp_smart_apply (wp_fork with "[Hc1]"). + { iIntros "!>". wp_recv (v) as "_". by wp_send with "[//]". } + wp_smart_apply (wp_fork with "[Hc2]"). + { iIntros "!>". wp_recv (l x) as "Hl". wp_load. wp_store. + by wp_send with "[$Hl]". } + wp_alloc l as "Hl". wp_send with "[$Hl]". wp_recv as "Hl". wp_load. + by iApply "HΦ". + Qed. + +End smuggle_ref. + Section parallel. Context `{!heapGS Σ}. @@ -501,9 +552,7 @@ Section forwarder. <(Recv, if b then 2 else 3) > MSG #x ; END)%proto]> (<[1 := (<(Recv, 0) @ (x:Z)> MSG #x ; <(Recv, 0) @ (b:bool)> MSG #b; - if b - then <(Send,2)> MSG #x ; END - else <(Send,3)> MSG #x ; END)%proto]> + <(Send, if b then 2 else 3)> MSG #x ; END)%proto]> (<[2 := (<(Recv, 1) @ (x:Z)> MSG #x ; <(Send, 0)> MSG #x ; END)%proto]> (<[3 := (<(Recv, 1) @ (x:Z)> MSG #x ; @@ -535,14 +584,12 @@ Section forwarder_rec. Context `{!heapGS Σ}. (** - 0 / | \ / | \ | 1 | \ / \ / 2 3 - *) Definition iProto_forwarder_rec_0_aux (rec : iProto Σ) : iProto Σ := -- GitLab From 88379326564689bb431f64cdd29477b7ebbfe34f Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Tue, 20 Feb 2024 23:14:15 +0100 Subject: [PATCH 50/81] Also smuggle on return --- multi_actris/examples/basics.v | 16 +++++++++------- 1 file changed, 9 insertions(+), 7 deletions(-) diff --git a/multi_actris/examples/basics.v b/multi_actris/examples/basics.v index e505c2e..9efa72b 100644 --- a/multi_actris/examples/basics.v +++ b/multi_actris/examples/basics.v @@ -250,20 +250,21 @@ Definition smuggle_ref_prog : val := let: "c0" := get_chan "cs" #0 in let: "c1" := get_chan "cs" #1 in let: "c2" := get_chan "cs" #2 in - Fork (send "c1" #2 (recv "c1" #0));; - Fork (let: "l" := recv "c2" #1 in "l" <- !"l" + #2;; send "c2" #0 #());; + Fork (send "c1" #2 (recv "c1" #0);; send "c1" #0 (recv "c1" #2));; + Fork (let: "l" := recv "c2" #1 in "l" <- !"l" + #2;; send "c2" #1 #());; let: "l" := ref #40 in - send "c0" #1 "l";; recv "c0" #2;; !"l". + send "c0" #1 "l";; recv "c0" #1;; !"l". Section smuggle_ref. Context `{!heapGS Σ, !chanG Σ}. Definition iProto_smuggle_ref : gmap nat (iProto Σ) := <[0 := (<(Send, 1) @ (l:loc) (x:Z)> MSG #l {{ l ↦ #x }} ; - <(Recv,2)> MSG #() {{ l ↦ #(x+2) }} ; END)%proto]> - (<[1 := (<(Recv, 0) @ (v:val)> MSG v ; <(Send,2)> MSG v ; END)%proto]> + <(Recv,1)> MSG #() {{ l ↦ #(x+2) }} ; END)%proto]> + (<[1 := (<(Recv, 0) @ (v:val)> MSG v ; <(Send,2)> MSG v ; + <(Recv, 2)> MSG #(); <(Send,0)> MSG #() ; END)%proto]> (<[2 := (<(Recv, 1) @ (l:loc) (x:Z)> MSG #l {{ l ↦ #x }} ; - <(Send,0)> MSG #() {{ l ↦ #(x+2) }} ; END)%proto]> ∅)). + <(Send,1)> MSG #() {{ l ↦ #(x+2) }} ; END)%proto]> ∅)). Lemma iProto_smuggle_ref_consistent : ⊢ iProto_consistent iProto_smuggle_ref. @@ -286,7 +287,8 @@ Section smuggle_ref. wp_smart_apply (get_chan_spec _ 2 with "Hcs"); [set_solver|]. iIntros (c2) "[Hc2 Hcs]". wp_smart_apply (wp_fork with "[Hc1]"). - { iIntros "!>". wp_recv (v) as "_". by wp_send with "[//]". } + { iIntros "!>". wp_recv (v) as "_". wp_send with "[//]". + wp_recv as "_". by wp_send with "[//]". } wp_smart_apply (wp_fork with "[Hc2]"). { iIntros "!>". wp_recv (l x) as "Hl". wp_load. wp_store. by wp_send with "[$Hl]". } -- GitLab From afe9f282945a0cbfeb59928857a3932eb32cc04d Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Mon, 11 Mar 2024 18:33:46 +0100 Subject: [PATCH 51/81] Improved leader election --- multi_actris/examples/leader_election.v | 95 +++++++++++-------------- 1 file changed, 41 insertions(+), 54 deletions(-) diff --git a/multi_actris/examples/leader_election.v b/multi_actris/examples/leader_election.v index a464555..cb21e82 100644 --- a/multi_actris/examples/leader_election.v +++ b/multi_actris/examples/leader_election.v @@ -4,29 +4,28 @@ From iris.heap_lang.lib Require Import assert. (** Inspired by https://en.wikipedia.org/wiki/Chang_and_Roberts_algorithm *) Definition process : val := - rec: "go" "c" "idl" "id" "idr" "isp" "id_max" := + rec: "go" "c" "idl" "id" "idr" "isp" := if: recv "c" "idr" then let: "id'" := recv "c" "idr" in - if: "id_max" < "id'" (** Case 1 *) + if: "id" < "id'" (** Case 1 *) then send "c" "idl" #true;; send "c" "idl" "id'";; - "go" "c" "idl" "id" "idr" "isp" "id'" - else if: "id_max" = "id'" (** Case 4 *) - then send "c" "idl" #false;; send "c" "idl" "id_max";; - "go" "c" "idl" "id" "idr" #false "id_max" + "go" "c" "idl" "id" "idr" #true + else if: "id" = "id'" (** Case 4 *) + then send "c" "idl" #false;; send "c" "idl" "id";; + "go" "c" "idl" "id" "idr" #false else if: "isp" (** Case 3 *) - then "go" "c" "idl" "id" "idr" "isp" "id_max" (** Case 2 *) - else send "c" "idl" #true;; send "c" "idl" "id_max";; - "go" "c" "idl" "id" "idr" #true "id_max" + then "go" "c" "idl" "id" "idr" "isp" (** Case 2 *) + else send "c" "idl" #true;; send "c" "idl" "id";; + "go" "c" "idl" "id" "idr" #true else let: "id'" := recv "c" "idr" in - assert: ("id_max" = "id'");; if: "id" = "id'" then "id'" - else send "c" "idl" #false;; send "c" "idl" "id_max";; "id_max". + else send "c" "idl" #false;; send "c" "idl" "id'";; "id'". Definition init : val := λ: "c" "idl" "id" "idr", (* Notice missing leader *) send "c" "idl" #true;; send "c" "idl" "id";; - process "c" "idl" "id" "idr" #true "id". + process "c" "idl" "id" "idr" #true. Definition program : val := λ: <>, @@ -36,7 +35,7 @@ Definition program : val := let: "c2" := get_chan "cs" #2 in let: "c3" := get_chan "cs" #3 in Fork (let: "id_max" := init "c1" #3 #1 #2 in send "c1" #0 "id_max");; - Fork (let: "id_max" := process "c2" #1 #2 #3 #false #2 in + Fork (let: "id_max" := process "c2" #1 #2 #3 #false in send "c2" #0 "id_max");; Fork (let: "id_max" := init "c3" #2 #3 #1 in send "c3" #0 "id_max");; let: "res1" := recv "c0" #1 in @@ -52,44 +51,42 @@ Section ring_leader_election_example. Context `{!heapGS Σ, chanG Σ, spawnG Σ, mono_natG Σ}. Definition my_recv_prot (il i ir : nat) (p : nat → iProto Σ) - (rec : bool -d> nat -d> iProto Σ) : bool -d> nat -d> iProto Σ := - λ (isp:bool) (i_max:nat), + (rec : bool -d> iProto Σ) : bool -d> iProto Σ := + λ (isp:bool), iProto_choice (Recv,ir) (<(Recv,ir) @ (i':nat)> MSG #i' ; - if bool_decide (i_max < i') - then <(Send,il)> MSG #true ; <(Send,il)> MSG #(max i' i_max) ; - rec isp (max i' i_max) - else if bool_decide (i_max = i') - then <(Send,il)> MSG #false ; <(Send, il)> MSG #i_max ; rec false i_max - else if isp then rec isp i_max - else <(Send,il)> MSG #true ; <(Send,il)> MSG #(max i' i_max) ; - rec true (max i' i_max))%proto - (<(Recv,ir)> MSG #i_max ; - if (bool_decide (i = i_max)) then p i_max - else <(Send,il)> MSG #false; <(Send,il)> MSG #i_max; p i_max)%proto. + if bool_decide (i < i') + then <(Send,il)> MSG #true ; <(Send,il)> MSG #i' ; rec true + else if bool_decide (i = i') + then <(Send,il)> MSG #false ; <(Send, il)> MSG #i ; rec false + else if isp then rec isp + else <(Send,il)> MSG #true ; <(Send,il)> MSG #i ; rec true)%proto + (<(Recv,ir) @ (i':nat)> MSG #i' ; + if (bool_decide (i = i')) then p i + else <(Send,il)> MSG #false; <(Send,il)> MSG #i'; p i')%proto. Instance rle_prot_aux_contractive il i ir p : Contractive (my_recv_prot il i ir p). Proof. rewrite /my_recv_prot. solve_proto_contractive. Qed. Definition rle_prot il i ir p := fixpoint (my_recv_prot il i ir p). - Instance rle_prot_unfold il i ir isp max_id p : - ProtoUnfold (rle_prot il i ir p isp max_id) (my_recv_prot il i ir p (rle_prot il i ir p) isp max_id). + Instance rle_prot_unfold il i ir isp p : + ProtoUnfold (rle_prot il i ir p isp) (my_recv_prot il i ir p (rle_prot il i ir p) isp). Proof. apply proto_unfold_eq, (fixpoint_unfold (my_recv_prot il i ir p)). Qed. - Lemma rle_prot_unfold' il i ir isp max_id p : - (rle_prot il i ir p isp max_id) ≡ - (my_recv_prot il i ir p (rle_prot il i ir p) isp max_id). + Lemma rle_prot_unfold' il i ir isp p : + (rle_prot il i ir p isp) ≡ + (my_recv_prot il i ir p (rle_prot il i ir p) isp). Proof. apply (fixpoint_unfold (my_recv_prot il i ir p)). Qed. Definition rle_preprot (il i ir : nat) p : iProto Σ := (<(Send, il)> MSG #true; <(Send, il)> MSG #i ; - rle_prot il i ir p true i)%proto. + rle_prot il i ir p true)%proto. - Lemma process_spec il i ir p i_max c (isp:bool) : - {{{ c ↣ (rle_prot il i ir p isp i_max) }}} - process c #il #i #ir #isp #i_max + Lemma process_spec il i ir p c (isp:bool) : + {{{ c ↣ (rle_prot il i ir p isp) }}} + process c #il #i #ir #isp {{{ i', RET #i'; c ↣ p i' }}}. Proof. iIntros (Φ) "Hc HΦ". - iLöb as "IH" forall (Φ isp i_max). + iLöb as "IH" forall (Φ isp). wp_lam. wp_recv (b) as "_". destruct b. - wp_pures. wp_recv (i') as "_". @@ -97,40 +94,30 @@ Section ring_leader_election_example. case_bool_decide as Hlt. { case_bool_decide; [|lia]. wp_pures. wp_send with "[//]". - iEval (replace i' with (max i' i_max) by lia). - wp_send with "[//]". wp_pures. + wp_send with "[//]". wp_pures. iApply ("IH" with "Hc HΦ"). } case_bool_decide as Hlt2. { case_bool_decide; [lia|]. wp_pures. case_bool_decide; [|simplify_eq;lia]. wp_send with "[//]". - iEval (replace i' with (max i' i_max) by lia). wp_send with "[//]". wp_pures. iApply ("IH" with "Hc HΦ"). } case_bool_decide; [lia|]. wp_pures. case_bool_decide; [simplify_eq;lia|]. - wp_pures. + wp_pures. destruct isp. { wp_pures. iApply ("IH" with "Hc HΦ"). } wp_pures. wp_send with "[//]". - iEval (replace i_max with (max i' i_max) by lia). wp_send with "[//]". wp_pures. iApply ("IH" with "Hc HΦ"). - wp_pures. - wp_recv as "_". - wp_smart_apply wp_assert. - wp_pures. iModIntro. - iSplitR. - { iPureIntro. f_equiv. f_equiv. apply bool_decide_eq_true_2. done. } - iNext. - wp_pures. - case_bool_decide. - { wp_pures. simplify_eq. - case_bool_decide; [|simplify_eq;lia]. wp_pures. - iModIntro. by iApply "HΦ". } - wp_pures. case_bool_decide; [simplify_eq;lia|]. wp_pures. + wp_recv (id') as "_". wp_pures. + case_bool_decide as Hlt. + { case_bool_decide; [|simplify_eq;lia]. + wp_pures. subst. by iApply "HΦ". } + case_bool_decide; [simplify_eq;lia|]. wp_send with "[//]". wp_send with "[//]". wp_pures. by iApply "HΦ". Qed. @@ -165,7 +152,7 @@ Section ring_leader_election_example. <(Recv,3)> MSG #id_max ; END)%proto ]> (<[1 := rle_preprot 3 1 2 prot_tail ]> - (<[2 := rle_prot 1 2 3 prot_tail false 2 ]> + (<[2 := rle_prot 1 2 3 prot_tail false ]> (<[3 := rle_preprot 2 3 1 prot_tail ]> ∅))). Lemma prot_pool_consistent : ⊢ iProto_consistent prot_pool. -- GitLab From 4e7fc218ec1ad468e9929fc5ead65d2ca4bfe4dd Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Thu, 21 Mar 2024 17:46:37 +0100 Subject: [PATCH 52/81] WIP: Used lists in iProto_consistent instead of maps --- multi_actris/channel/channel.v | 184 +++++++++++++++---------- multi_actris/channel/proto.v | 237 +++++++++++++++------------------ 2 files changed, 222 insertions(+), 199 deletions(-) diff --git a/multi_actris/channel/channel.v b/multi_actris/channel/channel.v index 89937a2..0a1248a 100644 --- a/multi_actris/channel/channel.v +++ b/multi_actris/channel/channel.v @@ -103,7 +103,7 @@ Definition iProto_pointsto_def `{!heapGS Σ, !chanG Σ} (c : val) (p : iProto Σ) : iProp Σ := ∃ γ γE1 (l:loc) (i:nat) (n:nat) p', ⌜ c = (#l,#n,#i)%V ⌠∗ - inv (nroot.@"ctx") (iProto_ctx γ (set_seq 0 n)) ∗ + inv (nroot.@"ctx") (iProto_ctx γ n) ∗ ([∗list] j ↦ _ ∈ replicate n (), ∃ γE2 γt1 γt2, inv (nroot.@"p") (chan_inv γ γE1 γt1 i j (l +â‚— (pos n i j))) ∗ @@ -121,12 +121,12 @@ Notation "c ↣ p" := (iProto_pointsto c p) (at level 20, format "c ↣ p"). Definition chan_pool `{!heapGS Σ, !chanG Σ} - (cs : val) (ps : gmap nat (iProto Σ)) : iProp Σ := - ∃ γ (γEs : list gname) (l:loc) (n:nat), - ⌜cs = (#l,#n)%V⌠∗ ⌜∀ i, is_Some (ps !! i) → i < n⌠∗ - inv (nroot.@"ctx") (iProto_ctx γ (set_seq 0 n)) ∗ + (cs : val) (i':nat) (ps : list (iProto Σ)) : iProp Σ := + ∃ γ (γEs : list gname) (n:nat) (l:loc), + ⌜cs = (#l,#n)%V⌠∗ ⌜(i' + length ps = n)%nat⌠∗ + inv (nroot.@"ctx") (iProto_ctx γ n) ∗ [∗ list] i ↦ _ ∈ replicate n (), - (∀ p, ⌜ps !! i = Some p⌠-∗ + (∀ p, ⌜i' <= i⌠-∗ ⌜ps !! (i - i') = Some p⌠-∗ own (γEs !!! i) (â—E (Next p)) ∗ own (γEs !!! i) (â—¯E (Next p)) ∗ iProto_own γ i p) ∗ @@ -203,63 +203,53 @@ Section channel. Qed. (** ** Specifications of [send] and [recv] *) - Lemma new_chan_spec (n:nat) (ps:gmap nat (iProto Σ)) : - 0 < n → - dom ps = set_seq 0 n → + Lemma new_chan_spec (ps:list (iProto Σ)) : + 0 < length ps → {{{ iProto_consistent ps }}} - new_chan #n - {{{ cs, RET cs; chan_pool cs ps }}}. + new_chan #(length ps) + {{{ cs, RET cs; chan_pool cs 0 ps }}}. Proof. - iIntros (Hle Hdom Φ) "Hps HΦ". wp_lam. + iIntros (Hle Φ) "Hps HΦ". wp_lam. wp_smart_apply wp_allocN; [lia|done|]. iIntros (l) "[Hl _]". iMod (iProto_init with "Hps") as (γ) "[Hps Hps']". wp_pures. iApply "HΦ". iAssert (|==> ∃ (γEs : list gname), - ⌜length γEs = n⌠∗ - [∗ list] i ↦ _ ∈ replicate n (), + ⌜length γEs = length ps⌠∗ + [∗ list] i ↦ _ ∈ replicate (length ps) (), own (γEs !!! i) (â—E (Next (ps !!! i))) ∗ own (γEs !!! i) (â—¯E (Next (ps !!! i))) ∗ iProto_own γ i (ps !!! i))%I with "[Hps']" as "H". { clear Hle. - iInduction n as [|n] "IHn" forall (ps Hdom). + iInduction ps as [|p ps] "IHn" using rev_ind. { iExists []. iModIntro. simpl. done. } - assert (is_Some (ps !! n)) as [p HSome]. - { apply elem_of_dom. rewrite Hdom. apply elem_of_set_seq. lia. } - iDestruct (big_sepM_delete _ _ n with "Hps'") as "[Hp Hps']"; [done|]. + iDestruct "Hps'" as "[Hps' Hp]". iMod (own_alloc (â—E (Next p) â‹… â—¯E (Next p))) as (γE) "[Hauth Hfrag]". { apply excl_auth_valid. } - iMod ("IHn" with "[] Hps'") as (γEs Hlen) "H". - { iPureIntro. - rewrite dom_delete_L Hdom. - replace (S n) with (n + 1) by lia. - rewrite set_seq_add_L /= right_id_L difference_union_distr_l_L - difference_diag_L right_id_L. - assert (n ∉ (set_seq 0 n:gset _)); [|set_solver]. - intros Hin%elem_of_set_seq. lia. } + iMod ("IHn" with "Hps'") as (γEs Hlen) "H". iModIntro. iExists (γEs++[γE]). - replace (S n) with (n + 1) by lia. - rewrite replicate_add big_sepL_app app_length Hlen. - iSplit; [done|]=> /=. + rewrite !app_length Hlen. + iSplit; [iPureIntro=>/=;lia|]=> /=. + rewrite replicate_add. iSplitL "H". { iApply (big_sepL_impl with "H"). iIntros "!>" (i ? HSome') "(Hauth & Hfrag & Hown)". - assert (i < n) as Hle. + assert (i < length ps) as Hle. { by apply lookup_replicate_1 in HSome' as [??]. } - assert (delete n ps !!! i = ps !!! i) as Heq'. - { apply lookup_total_delete_ne. lia. } - rewrite Heq'. iFrame. - rewrite lookup_total_app_l; [|lia]. iFrame. } + rewrite !lookup_total_app_l; [|lia..]. iFrame. } rewrite replicate_length Nat.add_0_r. - rewrite list_lookup_total_middle; [|done]. - rewrite lookup_total_alt HSome. by iFrame. } + simpl. rewrite right_id_L. + rewrite !lookup_total_app_r; [|lia..]. rewrite !Hlen. + rewrite Nat.sub_diag. simpl. iFrame. + iDestruct "Hp" as "[$ _]". } iMod "H" as (γEs Hlen) "H". + set n := length ps. iAssert (|={⊤}=> - [∗ list] i ↦ _ ∈ replicate n (), - [∗ list] j ↦ _ ∈ replicate n (), + [∗ list] i ↦ _ ∈ replicate (length ps) (), + [∗ list] j ↦ _ ∈ replicate (length ps) (), ∃ γt, inv (nroot.@"p") (chan_inv γ (γEs !!! i) γt i j - (l +â‚— (pos n i j))))%I with "[Hl]" as "IH". + (l +â‚— (pos (length ps) i j))))%I with "[Hl]" as "IH". { replace (Z.to_nat (Z.of_nat n * Z.of_nat n)) with (n*n) by lia. iDestruct (array_to_matrix with "Hl") as "Hl". iApply big_sepL_fupd. @@ -273,19 +263,19 @@ Section channel. iMod "IH" as "#IH". iMod (inv_alloc with "Hps") as "#IHp". iExists _,_,_,_. - iModIntro. iSplit; [done|]. + iModIntro. + iSplit. + { iPureIntro. done. } iSplit. - { iPureIntro. intros i HSome. - apply elem_of_dom in HSome. - rewrite Hdom in HSome. - apply elem_of_set_seq in HSome. lia. } - rewrite Hdom. iFrame "IHp". + { done. } + iFrame "IHp". iApply (big_sepL_impl with "H"). iIntros "!>" (i ? HSome') "(Hauth & Hfrag & Hown)". iSplitL. - { iIntros (p HSome''). - rewrite lookup_total_alt. rewrite HSome''. - iFrame. } + { iIntros (p Hle' HSome''). + iFrame. rewrite right_id_L in HSome''. + rewrite (list_lookup_total_alt ps). + rewrite HSome''. simpl. iFrame. } iApply big_sepL_intro. iIntros "!>" (j ? HSome''). assert (i < n) as Hle'. @@ -305,40 +295,64 @@ Section channel. iExists _, _. iFrame "#". Qed. - Lemma get_chan_spec cs (i:nat) ps p : - ps !! i = Some p → - {{{ chan_pool cs ps }}} + Lemma get_chan_spec cs i ps p : + {{{ chan_pool cs i (p::ps) }}} get_chan cs #i - {{{ c, RET c; c ↣ p ∗ chan_pool cs (delete i ps) }}}. + {{{ c, RET c; c ↣ p ∗ chan_pool cs (i+1) ps }}}. Proof. - iIntros (HSome Φ) "Hcs HΦ". - iDestruct "Hcs" as (γp γEs l n -> Hle) "[#IHp Hl]". + iIntros (Φ) "Hcs HΦ". + iDestruct "Hcs" as (γp γEs n l -> <-) "[#IHp Hl]". wp_lam. wp_pures. - assert (i < n); [by apply Hle|]. - iDestruct (big_sepL_delete _ _ i () with "Hl") as "[[Hi #IHs] H]"; - [by apply lookup_replicate|]. - iDestruct ("Hi" with "[//]") as "(Hauth & Hown & Hp)". + simpl. + rewrite replicate_add. simpl. + iDestruct "Hl" as "[Hl1 [[Hi #IHs] Hl3]]". simpl. + iDestruct ("Hi" with "[] []") as "(Hauth & Hown & Hp)". + { rewrite right_id_L. rewrite replicate_length. iPureIntro. lia. } + { rewrite right_id_L. rewrite replicate_length. + rewrite Nat.sub_diag. simpl. done. } iModIntro. iApply "HΦ". - iSplitR "H". + iSplitR "Hl1 Hl3". { rewrite iProto_pointsto_eq. iExists _, _, _, _, _, _. - iSplit; [done|]. iFrame "#∗". iSplit; [|iNext; done]. + iSplit; [done|]. + rewrite replicate_length. rewrite right_id_L. + iFrame "#∗". + iSplit; [|iNext; done]. + rewrite replicate_add. iApply (big_sepL_impl with "IHs"). iIntros "!>" (???). iDestruct 1 as (γt1 γt2) "[??]". iExists _,_,_. iFrame. } iExists _, _, _, _. iSplit; [done|]. iSplit. - { iPureIntro. intros i' HSome'. apply Hle. - assert (i ≠i'). - { intros ->. rewrite lookup_delete in HSome'. by inversion HSome'. } - by rewrite lookup_delete_ne in HSome'. } + { iPureIntro. lia. } iFrame "#∗". - iApply (big_sepL_impl with "H"). + rewrite replicate_add. + simpl. + iSplitL "Hl1". + { iApply (big_sepL_impl with "Hl1"). + iIntros "!>" (i' ? HSome''). + assert (i' < i). + { rewrite lookup_replicate in HSome''. lia. } + iIntros "[H $]" (p' Hle'). lia. } + simpl. + iFrame "#∗". + iSplitR. + { iIntros (p' Hle'). rewrite right_id_L in Hle'. + rewrite replicate_length in Hle'. lia. } + iApply (big_sepL_impl with "Hl3"). iIntros "!>" (i' ? HSome''). - case_decide. - { simplify_eq. iFrame "#". - iIntros "_" (p' Hin). simplify_eq. by rewrite lookup_delete in Hin. } - rewrite lookup_delete_ne; [|done]. by eauto. + assert (i' < length ps). + { rewrite lookup_replicate in HSome''. lia. } + iIntros "[H $]" (p' Hle' HSome). + iApply "H". + { iPureIntro. lia. } + iPureIntro. + rewrite replicate_length. + rewrite replicate_length in HSome. + replace (i + S i' - i) with (S i') by lia. + simpl. + replace (i + S i' - (i+1)) with (i') in HSome by lia. + done. Qed. Lemma vpos_spec (n i j : nat) : @@ -363,10 +377,10 @@ Section channel. iDestruct (iProto_le_msg_inv_r with "Hle") as (m') "#Heq". iRewrite "Heq" in "Hown". iAssert (â–· (â–· ⌜j < n⌠∗ iProto_own γ i (<(Send, j)> m') ∗ - iProto_ctx γ (set_seq 0 n)))%I with "[HI Hown]" + iProto_ctx γ n))%I with "[HI Hown]" as "[HI [Hown Hctx]]". { iNext. iDestruct (iProto_target with "HI Hown") as "[Hin [$ $]]". - iFrame. iNext. iDestruct "Hin" as %Hin. by set_solver. } + iFrame. } iRewrite -"Heq" in "Hown". wp_pures. iModIntro. iFrame. wp_smart_apply (vpos_spec); [done|]; iIntros "_". iDestruct "HI" as %Hle. @@ -456,9 +470,9 @@ Section channel. iDestruct (iProto_le_msg_inv_r with "Hle") as (m') "#Heq". iRewrite "Heq" in "Hown". iAssert (â–· (â–· ⌜j < n⌠∗ iProto_own γ i (<(Recv, j)> m') ∗ - iProto_ctx γ (set_seq 0 n)))%I with "[HI Hown]" as "[HI [Hown Hctx]]". + iProto_ctx γ n))%I with "[HI Hown]" as "[HI [Hown Hctx]]". { iNext. iDestruct (iProto_target with "HI Hown") as "[Hin [$$]]". - iFrame. iNext. iDestruct "Hin" as %Hin. by set_solver. } + iFrame. } iRewrite -"Heq" in "Hown". wp_pures. iModIntro. iFrame. wp_smart_apply (vpos_spec); [done|]; iIntros "_". iDestruct "HI" as %Hle. @@ -505,4 +519,28 @@ Section channel. iRewrite "Hp". iFrame "#∗". iApply iProto_le_refl. Qed. + Lemma iProto_le_select_l {TT1 TT2:tele} j + (v1 : TT1 → val) (v2 : TT2 → val) P1 P2 (p1 : TT1 → iProto Σ) (p2 : TT2 → iProto Σ) : + ⊢ (iProto_choice (Send, j) v1 v2 P1 P2 p1 p2) ⊑ + (<(Send,j) @.. (tt:TT1)> MSG (InjLV (v1 tt)) {{ P1 tt }} ; p1 tt). + Proof. + rewrite /iProto_choice. + iApply iProto_le_trans; last first. + { iApply iProto_le_texist_elim_r. iIntros (x). iExists x. + iApply iProto_le_refl. } + iIntros (tt). by iExists (inl tt). + Qed. + + Lemma iProto_le_select_r {TT1 TT2:tele} j + (v1 : TT1 → val) (v2 : TT2 → val) P1 P2 (p1 : TT1 → iProto Σ) (p2 : TT2 → iProto Σ) : + ⊢ (iProto_choice (Send, j) v1 v2 P1 P2 p1 p2) ⊑ + (<(Send,j) @.. (tt:TT2)> MSG (InjRV (v2 tt)) {{ P2 tt }} ; p2 tt). + Proof. + rewrite /iProto_choice. + iApply iProto_le_trans; last first. + { iApply iProto_le_texist_elim_r. iIntros (x). iExists x. + iApply iProto_le_refl. } + iIntros (tt). by iExists (inr tt). + Qed. + End channel. diff --git a/multi_actris/channel/proto.v b/multi_actris/channel/proto.v index dbb1b16..e5a2b78 100644 --- a/multi_actris/channel/proto.v +++ b/multi_actris/channel/proto.v @@ -173,22 +173,6 @@ Notation "< a @.. x1 .. xn > m" := (iProto_message a (∃.. x1, .. (∃.. xn, m) (at level 200, a at level 10, x1 closed binder, xn closed binder, m at level 200, format "< a @.. x1 .. xn > m") : proto_scope. -Notation "<!> m" := (<Send> m) (at level 200, m at level 200) : proto_scope. -Notation "<! x1 .. xn > m" := (<!> ∃ x1, .. (∃ xn, m) ..) - (at level 200, x1 closed binder, xn closed binder, m at level 200, - format "<! x1 .. xn > m") : proto_scope. -Notation "<!.. x1 .. xn > m" := (<!> ∃.. x1, .. (∃.. xn, m) ..) - (at level 200, x1 closed binder, xn closed binder, m at level 200, - format "<!.. x1 .. xn > m") : proto_scope. - -Notation "<?> m" := (<Recv> m) (at level 200, m at level 200) : proto_scope. -Notation "<? x1 .. xn > m" := (<?> ∃ x1, .. (∃ xn, m) ..) - (at level 200, x1 closed binder, xn closed binder, m at level 200, - format "<? x1 .. xn > m") : proto_scope. -Notation "<?.. x1 .. xn > m" := (<?> ∃.. x1, .. (∃.. xn, m) ..) - (at level 200, x1 closed binder, xn closed binder, m at level 200, - format "<?.. x1 .. xn > m") : proto_scope. - Class MsgTele {Σ V} {TT : tele} (m : iMsg Σ V) (tv : TT -t> V) (tP : TT -t> iProp Σ) (tp : TT -t> iProto Σ V) := msg_tele : m ≡ (∃.. x, MSG tele_app tv x {{ tele_app tP x }}; tele_app tp x)%msg. @@ -512,8 +496,8 @@ End proto. Global Instance iProto_inhabited {Σ V} : Inhabited (iProto Σ V) := populate END. -Definition can_step {Σ V} (rec : gmap nat (iProto Σ V) → iProp Σ) - (ps : gmap nat (iProto Σ V)) (i j : nat) : iProp Σ := +Definition can_step {Σ V} (rec : list (iProto Σ V) → iProp Σ) + (ps : list (iProto Σ V)) (i j : nat) : iProp Σ := ∀ m1 m2, (ps !!! i ≡ <(Send, j)> m1) -∗ (ps !!! j ≡ <(Recv, i)> m2) -∗ @@ -521,21 +505,21 @@ Definition can_step {Σ V} (rec : gmap nat (iProto Σ V) → iProp Σ) ∃ p2, iMsg_car m2 v (Next p2) ∗ â–· (rec (<[i:=p1]>(<[j:=p2]>ps))). -Definition valid_target {Σ V} (ps : gmap nat (iProto Σ V)) (i j : nat) : iProp Σ := +Definition valid_target {Σ V} (ps : list (iProto Σ V)) (i j : nat) : iProp Σ := ∀ a m, (ps !!! i ≡ <(a, j)> m) -∗ ⌜is_Some (ps !! j)âŒ. -Definition iProto_consistent_pre {Σ V} (rec : gmap nat (iProto Σ V) → iProp Σ) - (ps : gmap nat (iProto Σ V)) : iProp Σ := +Definition iProto_consistent_pre {Σ V} (rec : list (iProto Σ V) → iProp Σ) + (ps : list (iProto Σ V)) : iProp Σ := (∀ i j, valid_target ps i j) ∗ (∀ i j, can_step rec ps i j). Global Instance iProto_consistent_pre_ne {Σ V} - (rec : gmapO natO (iProto Σ V) -n> iPropO Σ) : + (rec : listO (iProto Σ V) -n> iPropO Σ) : NonExpansive (iProto_consistent_pre rec). Proof. rewrite /iProto_consistent_pre /can_step /valid_target. solve_proper. Qed. Program Definition iProto_consistent_pre' {Σ V} - (rec : gmapO natO (iProto Σ V) -n> iPropO Σ) : - gmapO natO (iProto Σ V) -n> iPropO Σ := + (rec : listO (iProto Σ V) -n> iPropO Σ) : + listO (iProto Σ V) -n> iPropO Σ := λne ps, iProto_consistent_pre (λ ps, rec ps) ps. Local Instance iProto_consistent_pre_contractive {Σ V} : Contractive (@iProto_consistent_pre' Σ V). @@ -544,7 +528,7 @@ Proof. solve_contractive. Qed. -Definition iProto_consistent {Σ V} (ps : gmap nat (iProto Σ V)) : iProp Σ := +Definition iProto_consistent {Σ V} (ps : list (iProto Σ V)) : iProp Σ := fixpoint iProto_consistent_pre' ps. Arguments iProto_consistent {_ _} _%proto. @@ -555,7 +539,7 @@ Proof. solve_proper. Qed. Global Instance iProto_consistent_proper {Σ V} : Proper ((≡) ==> (⊣⊢)) (@iProto_consistent Σ V). Proof. solve_proper. Qed. -Lemma iProto_consistent_unfold {Σ V} (ps : gmap nat (iProto Σ V)) : +Lemma iProto_consistent_unfold {Σ V} (ps : list (iProto Σ V)) : iProto_consistent ps ≡ (iProto_consistent_pre iProto_consistent) ps. Proof. apply: (fixpoint_unfold iProto_consistent_pre'). @@ -615,12 +599,13 @@ Definition iProto_own_frag `{!protoG Σ V} (γ : gname) own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p))). Definition iProto_own_auth `{!protoG Σ V} (γ : gname) - (ps : gmap nat (iProto Σ V)) : iProp Σ := - own γ (gmap_view_auth (DfracOwn 1) (((λ p, Excl' (Next p)) <$> ps) : gmap _ _)). + (ps : list (iProto Σ V)) : iProp Σ := + own γ (gmap_view_auth (DfracOwn 1) (((λ p, Excl' (Next p)) <$> + (list_to_map (zip (seq 0 (length ps)) ps))) : gmap _ _)). Definition iProto_ctx `{protoG Σ V} - (γ : gname) (ps_dom : gset nat) : iProp Σ := - ∃ ps, ⌜dom ps = ps_dom⌠∗ iProto_own_auth γ ps ∗ â–· iProto_consistent ps. + (γ : gname) (ps_len : nat) : iProp Σ := + ∃ ps, ⌜length ps = ps_len⌠∗ iProto_own_auth γ ps ∗ â–· iProto_consistent ps. (** * The connective for ownership of channel ends *) Definition iProto_own `{!protoG Σ V} @@ -802,34 +787,34 @@ Section proto. ps !!! i ≡ p1 -∗ p1 ⊑ p2 -∗ (∀ i' j', valid_target (<[i := p2]>ps) i' j') ∗ p1 ⊑ p2. - Proof. - iIntros "Hprot #HSome Hle". - pose proof (iProto_case p1) as [Hend|Hmsg]. - { rewrite Hend. iDestruct (iProto_le_end_inv_l with "Hle") as "#H". - iFrame "Hle". - iIntros (i' j' a m) "Hm". - destruct (decide (i = j')) as [->|Hneqj]. - { by rewrite lookup_insert. } - rewrite lookup_insert_ne; [|done]. - destruct (decide (i = i')) as [->|Hneqi]. - { rewrite lookup_total_insert. iRewrite "H" in "Hm". - by iDestruct (iProto_end_message_equivI with "Hm") as "Hm". } - rewrite lookup_total_insert_ne; [|done]. - by iApply "Hprot". } - destruct Hmsg as (t & n & m & Hmsg). - setoid_rewrite Hmsg. - iDestruct (iProto_le_msg_inv_l with "Hle") as (m2) "#Heq". iFrame "Hle". - iIntros (i' j' a m') "Hm". - destruct (decide (i = j')) as [->|Hneqj]. - { rewrite lookup_insert. done. } - rewrite lookup_insert_ne; [|done]. - destruct (decide (i = i')) as [->|Hneqi]. - { rewrite lookup_total_insert. iRewrite "Heq" in "Hm". - iDestruct (iProto_message_equivI with "Hm") as (Heq) "Hm". - simplify_eq. by iApply "Hprot". } - rewrite lookup_total_insert_ne; [|done]. - by iApply "Hprot". - Qed. + Proof. Admitted. + (* iIntros "Hprot #HSome Hle". *) + (* pose proof (iProto_case p1) as [Hend|Hmsg]. *) + (* { rewrite Hend. iDestruct (iProto_le_end_inv_l with "Hle") as "#H". *) + (* iFrame "Hle". *) + (* iIntros (i' j' a m) "Hm". *) + (* destruct (decide (i = j')) as [->|Hneqj]. *) + (* { Search list_lookup_total insert. rewrite list_lookup_total_insert. ; [done|]. lia. done. } *) + (* rewrite lookup_insert_ne; [|done]. *) + (* destruct (decide (i = i')) as [->|Hneqi]. *) + (* { rewrite lookup_total_insert. iRewrite "H" in "Hm". *) + (* by iDestruct (iProto_end_message_equivI with "Hm") as "Hm". } *) + (* rewrite lookup_total_insert_ne; [|done]. *) + (* by iApply "Hprot". } *) + (* destruct Hmsg as (t & n & m & Hmsg). *) + (* setoid_rewrite Hmsg. *) + (* iDestruct (iProto_le_msg_inv_l with "Hle") as (m2) "#Heq". iFrame "Hle". *) + (* iIntros (i' j' a m') "Hm". *) + (* destruct (decide (i = j')) as [->|Hneqj]. *) + (* { rewrite lookup_insert. done. } *) + (* rewrite lookup_insert_ne; [|done]. *) + (* destruct (decide (i = i')) as [->|Hneqi]. *) + (* { rewrite lookup_total_insert. iRewrite "Heq" in "Hm". *) + (* iDestruct (iProto_message_equivI with "Hm") as (Heq) "Hm". *) + (* simplify_eq. by iApply "Hprot". } *) + (* rewrite lookup_total_insert_ne; [|done]. *) + (* by iApply "Hprot". *) + (* Qed. *) Lemma iProto_consistent_le ps i p1 p2 : iProto_consistent ps -∗ @@ -847,7 +832,7 @@ Section proto. iFrame. iIntros (i' j' m1 m2) "#Hm1 #Hm2". destruct (decide (i = i')) as [<-|Hneq]. - { rewrite lookup_total_insert. + { rewrite list_lookup_total_insert; [|admit]. pose proof (iProto_case p2) as [Hend|Hmsg]. { setoid_rewrite Hend. rewrite iProto_end_message_equivI. done. } destruct Hmsg as (a&?&m&Hmsg). @@ -865,23 +850,23 @@ Section proto. iDestruct "Hle" as (m') "[#Heq H]". iDestruct ("H" with "Hm1'") as (p') "[Hle H]". destruct (decide (i = j')) as [<-|Hneq]. - { rewrite lookup_total_insert. rewrite iProto_message_equivI. - iDestruct "Hm2" as "[%Heq _]". done. } + { rewrite list_lookup_total_insert. rewrite iProto_message_equivI. + iDestruct "Hm2" as "[%Heq _]". done. admit. } iDestruct ("Hprot" $!i j' with "[] [] H") as "Hprot". - { iRewrite -"Heq". rewrite !lookup_total_alt. iRewrite "HSome". done. } - { rewrite lookup_total_insert_ne; [|done]. done. } + { iRewrite -"Heq". rewrite !list_lookup_total_alt. iRewrite "HSome". done. } + { rewrite list_lookup_total_insert_ne; [|done]. done. } iDestruct "Hprot" as (p'') "[Hm Hprot]". iExists p''. iFrame. iNext. iDestruct ("IH" with "Hprot Hle [HSome]") as "HI". - { by rewrite lookup_total_insert. } + { rewrite list_lookup_total_insert; [done|]. admit. } iClear "IH Hm1 Hm2 Heq". - rewrite insert_insert. - rewrite (insert_commute _ j' i); [|done]. - rewrite insert_insert. done. } - rewrite lookup_total_insert_ne; [|done]. + rewrite list_insert_insert. + rewrite (list_insert_commute _ j' i); [|done]. + rewrite list_insert_insert. done. } + rewrite list_lookup_total_insert_ne; [|done]. destruct (decide (i = j')) as [<-|Hneq']. - { rewrite lookup_total_insert. + { rewrite list_lookup_total_insert. pose proof (iProto_case p2) as [Hend|Hmsg]. { setoid_rewrite Hend. rewrite iProto_end_message_equivI. done. } destruct Hmsg as (a&?&m&Hmsg). @@ -897,31 +882,31 @@ Section proto. iDestruct "Hle" as (m') "[#Heq Hle]". iDestruct ("Hprot" $!i' with "[] [] Hm1'") as "Hprot". { done. } - { rewrite !lookup_total_alt. iRewrite "HSome". done. } + { rewrite !list_lookup_total_alt. iRewrite "HSome". done. } iDestruct ("Hprot") as (p') "[Hm1' Hprot]". iDestruct ("Hle" with "Hm1'") as (p2') "[Hle Hm']". iSpecialize ("Hm2" $! v (Next p2')). iExists p2'. iRewrite -"Hm2". iFrame. iDestruct ("IH" with "Hprot Hle []") as "HI". - { iPureIntro. rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert. done. } - rewrite insert_commute; [|done]. - rewrite !insert_insert. done. } - rewrite lookup_total_insert_ne; [|done]. + { iPureIntro. rewrite list_lookup_total_insert_ne; [|done]. + rewrite list_lookup_total_insert. done. admit. } + rewrite list_insert_commute; [|done]. + rewrite !list_insert_insert. done. admit. } + rewrite list_lookup_total_insert_ne; [|done]. iIntros (v p) "Hm1'". iDestruct ("Hprot" $!i' j' with "[//] [//] Hm1'") as "Hprot". iDestruct "Hprot" as (p') "[Hm2' Hprot]". iExists p'. iFrame. iNext. - rewrite (insert_commute _ j' i); [|done]. - rewrite (insert_commute _ i' i); [|done]. + rewrite (list_insert_commute _ j' i); [|done]. + rewrite (list_insert_commute _ i' i); [|done]. iApply ("IH" with "Hprot Hle []"). - rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. + rewrite list_lookup_total_insert_ne; [|done]. + rewrite list_lookup_total_insert_ne; [|done]. done. - Qed. - + Admitted. + Lemma iProto_le_send i m1 m2 : (∀ v p2', iMsg_car m2 v (Next p2') -∗ ∃ p1', â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1')) -∗ @@ -1156,16 +1141,16 @@ Section proto. Lemma iProto_own_auth_agree γ ps i p : iProto_own_auth γ ps -∗ iProto_own_frag γ i p -∗ â–· (ps !!! i ≡ p). - Proof. - iIntros "Hâ— Hâ—¯". - iDestruct (own_valid_2 with "Hâ— Hâ—¯") as "H✓". - rewrite gmap_view_both_validI. - iDestruct "H✓" as "[_ [H1 H2]]". - rewrite lookup_total_alt lookup_fmap. - destruct (ps !! i); last first. - { by rewrite !option_equivI. } - simpl. rewrite !option_equivI excl_equivI. by iNext. - Qed. + Proof. Admitted. + (* iIntros "Hâ— Hâ—¯". *) + (* iDestruct (own_valid_2 with "Hâ— Hâ—¯") as "H✓". *) + (* rewrite gmap_view_both_validI. *) + (* iDestruct "H✓" as "[_ [H1 H2]]". *) + (* rewrite list_lookup_total_alt lookup_fmap. *) + (* destruct (ps !! i); last first. *) + (* { rewrite !option_equivI. } *) + (* simpl. rewrite !option_equivI excl_equivI. by iNext. *) + (* Qed. *) Lemma iProto_own_auth_update γ ps i p p' : iProto_own_auth γ ps -∗ iProto_own_frag γ i p ==∗ @@ -1174,25 +1159,24 @@ Section proto. iIntros "Hâ— Hâ—¯". iMod (own_update_2 with "Hâ— Hâ—¯") as "[H1 H2]"; [|iModIntro]. { eapply (gmap_view_replace _ _ _ (Excl' (Next p'))). done. } - iFrame. rewrite -fmap_insert. done. - Qed. - + iFrame. rewrite -fmap_insert. Admitted. + Lemma iProto_own_auth_alloc ps : - ⊢ |==> ∃ γ, iProto_own_auth γ ps ∗ [∗ map] i ↦p ∈ ps, iProto_own γ i p. - Proof. - iMod (own_alloc (gmap_view_auth (DfracOwn 1) ∅)) as (γ) "Hauth". - { apply gmap_view_auth_valid. } - iExists γ. - iInduction ps as [|i p ps Hin] "IH" using map_ind. - { iModIntro. iFrame. by iApply big_sepM_empty. } - iMod ("IH" with "Hauth") as "[Hauth Hfrags]". - rewrite big_sepM_insert; [|done]. iFrame "Hfrags". - iMod (own_update with "Hauth") as "[Hauth Hfrag]". - { apply (gmap_view_alloc _ i (DfracOwn 1) (Excl' (Next p))); [|done|done]. - by rewrite lookup_fmap Hin. } - iModIntro. rewrite -fmap_insert. iFrame. - iExists _. iFrame. iApply iProto_le_refl. - Qed. + ⊢ |==> ∃ γ, iProto_own_auth γ ps ∗ [∗ list] i ↦p ∈ ps, iProto_own γ i p. + Proof. Admitted. + (* iMod (own_alloc (gmap_view_auth (DfracOwn 1) ∅)) as (γ) "Hauth". *) + (* { apply gmap_view_auth_valid. } *) + (* iExists γ. *) + (* iInduction ps as [|i p ps Hin] "IH" using map_ind. *) + (* { iModIntro. iFrame. by iApply big_sepM_empty. } *) + (* iMod ("IH" with "Hauth") as "[Hauth Hfrags]". *) + (* rewrite big_sepM_insert; [|done]. iFrame "Hfrags". *) + (* iMod (own_update with "Hauth") as "[Hauth Hfrag]". *) + (* { apply (gmap_view_alloc _ i (DfracOwn 1) (Excl' (Next p))); [|done|done]. *) + (* by rewrite lookup_fmap Hin. } *) + (* iModIntro. rewrite -fmap_insert. iFrame. *) + (* iExists _. iFrame. iApply iProto_le_refl. *) + (* Qed. *) Lemma iProto_own_le γ s p1 p2 : iProto_own γ s p1 -∗ â–· (p1 ⊑ p2) -∗ iProto_own γ s p2. @@ -1203,7 +1187,7 @@ Section proto. Lemma iProto_init ps : â–· iProto_consistent ps -∗ - |==> ∃ γ, iProto_ctx γ (dom ps) ∗ [∗ map] i ↦p ∈ ps, iProto_own γ i p. + |==> ∃ γ, iProto_ctx γ (length ps) ∗ [∗ list] i ↦p ∈ ps, iProto_own γ i p. Proof. iIntros "Hconsistent". iMod iProto_own_auth_alloc as (γ) "[Hauth Hfrags]". @@ -1227,28 +1211,28 @@ Section proto. iDestruct (iProto_own_auth_agree with "Hauth Hj") as "#Hpj". iDestruct (own_prot_idx with "Hi Hj") as %Hneq. iAssert (â–· (<[i:=<(Send, j)> m1]>ps !!! j ≡ pj))%I as "Hpj'". - { by rewrite lookup_total_insert_ne. } + { by rewrite list_lookup_total_insert_ne. } iAssert (â–· (⌜is_Some (ps !! i)⌠∗ (pi ⊑ (<(Send, j)> m1))))%I with "[Hile]" as "[Hi' Hile]". { iNext. iDestruct (iProto_le_msg_inv_r with "Hile") as (m) "#Heq". - iFrame. iRewrite "Heq" in "Hpi". rewrite lookup_total_alt. + iFrame. iRewrite "Heq" in "Hpi". rewrite list_lookup_total_alt. destruct (ps !! i); [done|]. iDestruct (iProto_end_message_equivI with "Hpi") as "[]". } iAssert (â–· (⌜is_Some (ps !! j)⌠∗ (pj ⊑ (<(Recv, i)> m2))))%I with "[Hjle]" as "[Hj' Hjle]". { iNext. iDestruct (iProto_le_msg_inv_r with "Hjle") as (m) "#Heq". - iFrame. iRewrite "Heq" in "Hpj". rewrite !lookup_total_alt. + iFrame. iRewrite "Heq" in "Hpj". rewrite !list_lookup_total_alt. destruct (ps !! j); [done|]. iDestruct (iProto_end_message_equivI with "Hpj") as "[]". } iDestruct (iProto_consistent_le with "Hconsistent Hpi Hile") as "Hconsistent". iDestruct (iProto_consistent_le with "Hconsistent Hpj' Hjle") as "Hconsistent". iDestruct (iProto_consistent_step _ _ _ i j with "Hconsistent [] [] [Hm //]") as (p2) "[Hm2 Hconsistent]". - { rewrite lookup_total_insert_ne; [|done]. - rewrite lookup_total_insert_ne; [|done]. - by rewrite lookup_total_insert. } - { rewrite lookup_total_insert_ne; [|done]. - by rewrite lookup_total_insert. } + { rewrite list_lookup_total_insert_ne; [|done]. + rewrite list_lookup_total_insert_ne; [|done]. + rewrite list_lookup_total_insert; [done|]. admit. } + { rewrite list_lookup_total_insert_ne; [|done]. + rewrite list_lookup_total_insert; [done|]. admit. } iMod (iProto_own_auth_update _ _ _ _ p2 with "Hauth Hj") as "[Hauth Hj]". iMod (iProto_own_auth_update _ _ _ _ p1 with "Hauth Hi") as "[Hauth Hi]". iIntros "!>!>". iExists p2. iFrame "Hm2". @@ -1256,17 +1240,18 @@ Section proto. iSplitL "Hconsistent Hauth". { iExists (<[i:=p1]> (<[j:=p2]> ps)). iSplit. - { rewrite !dom_insert_lookup_L; [done..|by rewrite lookup_insert_ne]. } - iFrame. rewrite insert_insert. - rewrite insert_commute; [|done]. rewrite insert_insert. - by rewrite insert_commute; [|done]. } + { admit. + (* rewrite !dom_insert_lookup_L; [done..|by rewrite lookup_insert_ne]. *)} + iFrame. rewrite list_insert_insert. + rewrite list_insert_commute; [|done]. rewrite list_insert_insert. + by rewrite list_insert_commute; [|done]. } iSplitL "Hi"; iExists _; iFrame; iApply iProto_le_refl. - Qed. + Admitted. Lemma iProto_target γ ps_dom i a j m : iProto_ctx γ ps_dom -∗ iProto_own γ i (<(a, j)> m) -∗ - â–· (⌜j ∈ ps_domâŒ) ∗ iProto_ctx γ ps_dom ∗ iProto_own γ i (<(a, j)> m). + â–· (⌜j < ps_domâŒ) ∗ iProto_ctx γ ps_dom ∗ iProto_own γ i (<(a, j)> m). Proof. iIntros "Hctx Hown". rewrite /iProto_ctx /iProto_own. @@ -1280,9 +1265,9 @@ Section proto. iDestruct (iProto_consistent_target with "Hps Hi") as "#$". by iFrame. } iSplitL "HSome". { iNext. iDestruct "HSome" as %Heq. - iPureIntro. simplify_eq. by apply elem_of_dom. } + iPureIntro. simplify_eq. admit. } iSplitL "Hauth Hps"; iExists _; by iFrame. - Qed. + Admitted. (* (** The instances below make it possible to use the tactics [iIntros], *) (* [iExist], [iSplitL]/[iSplitR], [iFrame] and [iModIntro] on [iProto_le] goals. *) *) -- GitLab From c914ef04e2c79d3f52d40ace2a9c6faea0ad63e5 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Thu, 21 Mar 2024 20:27:47 +0100 Subject: [PATCH 53/81] WIP: Computable consistency --- multi_actris/channel/proto_alt.v | 1408 ++++++++++++++++++++++++++++++ 1 file changed, 1408 insertions(+) create mode 100644 multi_actris/channel/proto_alt.v diff --git a/multi_actris/channel/proto_alt.v b/multi_actris/channel/proto_alt.v new file mode 100644 index 0000000..29d8603 --- /dev/null +++ b/multi_actris/channel/proto_alt.v @@ -0,0 +1,1408 @@ +(** This file defines the core of the Actris logic: It defines dependent +separation protocols and the various operations on it like dual, append, and +branching. + +Dependent separation protocols [iProto] are defined by instantiating the +parameterized version in [proto_model] with the type of propositions [iProp] of Iris. +We define ways of constructing instances of the instantiated type via two +constructors: +- [iProto_end], which is identical to [proto_end]. +- [iProto_message], which takes an [action] and an [iMsg]. The type [iMsg] is a + sequence of binders [iMsg_exist], terminated by the payload constructed with + [iMsg_base] based on arguments [v], [P] and [prot], which are the value, the + carried proposition and the [iProto] tail, respectively. + +For convenience sake, we provide the following notations: +- [END], which is simply [iProto_end]. +- [∃ x, m], which is [iMsg_exist] with argument [m]. +- [MSG v {{ P }}; prot], which is [iMsg_Base] with arguments [v], [P] and [prot]. +- [<a> m], which is [iProto_message] with arguments [a] and [m]. + +We also include custom notation to more easily construct complete constructions: +- [<a x1 .. xn> m], which is [<a> ∃ x1, .. ∃ xn, m]. +- [<a x1 .. xn> MSG v; {{ P }}; prot], which constructs a full protocol. + +Futhermore, we define the following operations: +- [iProto_dual], which turns all [Send] of a protocol into [Recv] and vice-versa. +- [iProto_app], which appends two protocols. + +In addition we define the subprotocol relation [iProto_le] [⊑], which generalises +the following inductive definition for asynchronous subtyping on session types: + + p1 <: p2 p1 <: p2 p1 <: !B.p3 ?A.p3 <: p2 +---------- ---------------- ---------------- ---------------------------- +end <: end !A.p1 <: !A.p2 ?A.p1 <: ?A.p2 ?A.p1 <: !B.p2 + +Example: + +!R <: !R ?Q <: ?Q ?Q <: ?Q +------------------- -------------- +?Q.!R <: !R.?Q ?P.?Q <: ?P.?Q +------------------------------------ + ?P.?Q.!R <: !R.?P.?Q + +Lastly, relevant type classes instances are defined for each of the above +notions, such as contractiveness and non-expansiveness, after which the +specifications of the message-passing primitives are defined in terms of the +protocol connectives. *) +From iris.algebra Require Import gmap excl_auth gmap_view. +From iris.proofmode Require Import proofmode. +From iris.base_logic Require Export lib.iprop. +From iris.base_logic Require Import lib.own. +From iris.program_logic Require Import language. +From multi_actris.channel Require Import proto_model. +Set Default Proof Using "Type". +Export action. + +(** * Setup of Iris's cameras *) +Class protoG Σ V := + protoG_authG :: + inG Σ (gmap_viewR natO + (optionUR (exclR (laterO (proto (leibnizO V) (iPropO Σ) (iPropO Σ)))))). + +Definition protoΣ V := #[ + GFunctor ((gmap_viewRF natO (optionRF (exclRF (laterOF (protoOF (leibnizO V) idOF idOF)))))) +]. +Global Instance subG_chanΣ {Σ V} : subG (protoΣ V) Σ → protoG Σ V. +Proof. solve_inG. Qed. + +(** * Types *) +Definition iProto Σ V := proto V (iPropO Σ) (iPropO Σ). +Declare Scope proto_scope. +Delimit Scope proto_scope with proto. +Bind Scope proto_scope with iProto. +Local Open Scope proto. + +(** * Messages *) +Section iMsg. + Set Primitive Projections. + Record iMsg Σ V := IMsg { iMsg_car : V → laterO (iProto Σ V) -n> iPropO Σ }. +End iMsg. +Arguments IMsg {_ _} _. +Arguments iMsg_car {_ _} _. + +Declare Scope msg_scope. +Delimit Scope msg_scope with msg. +Bind Scope msg_scope with iMsg. +Global Instance iMsg_inhabited {Σ V} : Inhabited (iMsg Σ V) := populate (IMsg inhabitant). + +Section imsg_ofe. + Context {Σ : gFunctors} {V : Type}. + + Instance iMsg_equiv : Equiv (iMsg Σ V) := λ m1 m2, + ∀ w p, iMsg_car m1 w p ≡ iMsg_car m2 w p. + Instance iMsg_dist : Dist (iMsg Σ V) := λ n m1 m2, + ∀ w p, iMsg_car m1 w p ≡{n}≡ iMsg_car m2 w p. + + Lemma iMsg_ofe_mixin : OfeMixin (iMsg Σ V). + Proof. by apply (iso_ofe_mixin (iMsg_car : _ → V -d> _ -n> _)). Qed. + Canonical Structure iMsgO := Ofe (iMsg Σ V) iMsg_ofe_mixin. + + Global Instance iMsg_cofe : Cofe iMsgO. + Proof. by apply (iso_cofe (IMsg : (V -d> _ -n> _) → _) iMsg_car). Qed. +End imsg_ofe. + +Program Definition iMsg_base_def {Σ V} + (v : V) (P : iProp Σ) (p : iProto Σ V) : iMsg Σ V := + IMsg (λ v', λne p', ⌜ v = v' ⌠∗ Next p ≡ p' ∗ P)%I. +Next Obligation. solve_proper. Qed. +Definition iMsg_base_aux : seal (@iMsg_base_def). by eexists. Qed. +Definition iMsg_base := iMsg_base_aux.(unseal). +Definition iMsg_base_eq : @iMsg_base = @iMsg_base_def := iMsg_base_aux.(seal_eq). +Arguments iMsg_base {_ _} _%V _%I _%proto. +Global Instance: Params (@iMsg_base) 3 := {}. + +Program Definition iMsg_exist_def {Σ V A} (m : A → iMsg Σ V) : iMsg Σ V := + IMsg (λ v', λne p', ∃ x, iMsg_car (m x) v' p')%I. +Next Obligation. solve_proper. Qed. +Definition iMsg_exist_aux : seal (@iMsg_exist_def). by eexists. Qed. +Definition iMsg_exist := iMsg_exist_aux.(unseal). +Definition iMsg_exist_eq : @iMsg_exist = @iMsg_exist_def := iMsg_exist_aux.(seal_eq). +Arguments iMsg_exist {_ _ _} _%msg. +Global Instance: Params (@iMsg_exist) 3 := {}. + +Definition iMsg_texist {Σ V} {TT : tele} (m : TT → iMsg Σ V) : iMsg Σ V := + tele_fold (@iMsg_exist Σ V) (λ x, x) (tele_bind m). +Arguments iMsg_texist {_ _ !_} _%msg /. + +Notation "'MSG' v {{ P } } ; p" := (iMsg_base v P p) + (at level 200, v at level 20, right associativity, + format "MSG v {{ P } } ; p") : msg_scope. +Notation "'MSG' v ; p" := (iMsg_base v True p) + (at level 200, v at level 20, right associativity, + format "MSG v ; p") : msg_scope. +Notation "∃ x .. y , m" := + (iMsg_exist (λ x, .. (iMsg_exist (λ y, m)) ..)%msg) : msg_scope. +Notation "'∃..' x .. y , m" := + (iMsg_texist (λ x, .. (iMsg_texist (λ y, m)) .. )%msg) + (at level 200, x binder, y binder, right associativity, + format "∃.. x .. y , m") : msg_scope. + +Lemma iMsg_texist_exist {Σ V} {TT : tele} w lp (m : TT → iMsg Σ V) : + iMsg_car (∃.. x, m x)%msg w lp ⊣⊢ (∃.. x, iMsg_car (m x) w lp). +Proof. + rewrite /iMsg_texist iMsg_exist_eq. + induction TT as [|T TT IH]; simpl; [done|]. f_equiv=> x. apply IH. +Qed. + +(** * Operators *) +Definition iProto_end_def {Σ V} : iProto Σ V := proto_end. +Definition iProto_end_aux : seal (@iProto_end_def). by eexists. Qed. +Definition iProto_end := iProto_end_aux.(unseal). +Definition iProto_end_eq : @iProto_end = @iProto_end_def := iProto_end_aux.(seal_eq). +Arguments iProto_end {_ _}. + +Definition iProto_message_def {Σ V} (a : action) (m : iMsg Σ V) : iProto Σ V := + proto_message a (iMsg_car m). +Definition iProto_message_aux : seal (@iProto_message_def). by eexists. Qed. +Definition iProto_message := iProto_message_aux.(unseal). +Definition iProto_message_eq : + @iProto_message = @iProto_message_def := iProto_message_aux.(seal_eq). +Arguments iProto_message {_ _} _ _%msg. +Global Instance: Params (@iProto_message) 3 := {}. + +Notation "'END'" := iProto_end : proto_scope. + +Notation "< a > m" := (iProto_message a m) + (at level 200, a at level 10, m at level 200, + format "< a > m") : proto_scope. +Notation "< a @ x1 .. xn > m" := (iProto_message a (∃ x1, .. (∃ xn, m) ..)) + (at level 200, a at level 10, x1 closed binder, xn closed binder, m at level 200, + format "< a @ x1 .. xn > m") : proto_scope. +Notation "< a @.. x1 .. xn > m" := (iProto_message a (∃.. x1, .. (∃.. xn, m) ..)) + (at level 200, a at level 10, x1 closed binder, xn closed binder, m at level 200, + format "< a @.. x1 .. xn > m") : proto_scope. + +Class MsgTele {Σ V} {TT : tele} (m : iMsg Σ V) + (tv : TT -t> V) (tP : TT -t> iProp Σ) (tp : TT -t> iProto Σ V) := + msg_tele : m ≡ (∃.. x, MSG tele_app tv x {{ tele_app tP x }}; tele_app tp x)%msg. +Global Hint Mode MsgTele ! ! - ! - - - : typeclass_instances. + +(** * Operations *) +Program Definition iMsg_map {Σ V} + (rec : iProto Σ V → iProto Σ V) (m : iMsg Σ V) : iMsg Σ V := + IMsg (λ v, λne p1', ∃ p1, iMsg_car m v (Next p1) ∗ p1' ≡ Next (rec p1))%I. +Next Obligation. solve_proper. Qed. + +Program Definition iProto_map_app_aux {Σ V} + (f : action → action) (p2 : iProto Σ V) + (rec : iProto Σ V -n> iProto Σ V) : iProto Σ V -n> iProto Σ V := λne p, + proto_elim p2 (λ a m, + proto_message (f a) (iMsg_car (iMsg_map rec (IMsg m)))) p. +Next Obligation. + intros Σ V f p2 rec n p1 p1' Hp. apply proto_elim_ne=> // a m1 m2 Hm. + apply proto_message_ne=> v p' /=. by repeat f_equiv. +Qed. + +Global Instance iProto_map_app_aux_contractive {Σ V} f (p2 : iProto Σ V) : + Contractive (iProto_map_app_aux f p2). +Proof. + intros n rec1 rec2 Hrec p1; simpl. apply proto_elim_ne=> // a m1 m2 Hm. + apply proto_message_ne=> v p' /=. by repeat (f_contractive || f_equiv). +Qed. + +Definition iProto_map_app {Σ V} (f : action → action) + (p2 : iProto Σ V) : iProto Σ V -n> iProto Σ V := + fixpoint (iProto_map_app_aux f p2). + +Definition iProto_app_def {Σ V} (p1 p2 : iProto Σ V) : iProto Σ V := + iProto_map_app id p2 p1. +Definition iProto_app_aux : seal (@iProto_app_def). Proof. by eexists. Qed. +Definition iProto_app := iProto_app_aux.(unseal). +Definition iProto_app_eq : @iProto_app = @iProto_app_def := iProto_app_aux.(seal_eq). +Arguments iProto_app {_ _} _%proto _%proto. +Global Instance: Params (@iProto_app) 2 := {}. +Infix "<++>" := iProto_app (at level 60) : proto_scope. +Notation "m <++> p" := (iMsg_map (flip iProto_app p) m) : msg_scope. + +Definition iProto_dual_def {Σ V} (p : iProto Σ V) : iProto Σ V := + iProto_map_app action_dual proto_end p. +Definition iProto_dual_aux : seal (@iProto_dual_def). Proof. by eexists. Qed. +Definition iProto_dual := iProto_dual_aux.(unseal). +Definition iProto_dual_eq : + @iProto_dual = @iProto_dual_def := iProto_dual_aux.(seal_eq). +Arguments iProto_dual {_ _} _%proto. +Global Instance: Params (@iProto_dual) 2 := {}. +Notation iMsg_dual := (iMsg_map iProto_dual). + +Definition iProto_dual_if {Σ V} (d : bool) (p : iProto Σ V) : iProto Σ V := + if d then iProto_dual p else p. +Arguments iProto_dual_if {_ _} _ _%proto. +Global Instance: Params (@iProto_dual_if) 3 := {}. + +(** * Proofs *) +Section proto. + Context `{!protoG Σ V}. + Implicit Types v : V. + Implicit Types p pl pr : iProto Σ V. + Implicit Types m : iMsg Σ V. + + (** ** Equality *) + Lemma iProto_case p : p ≡ END ∨ ∃ t n m, p ≡ <(t,n)> m. + Proof. + rewrite iProto_message_eq iProto_end_eq. + destruct (proto_case p) as [|([a n]&m&?)]; [by left|right]. + by exists a, n, (IMsg m). + Qed. + Lemma iProto_message_equivI `{!BiInternalEq SPROP} a1 a2 m1 m2 : + (<a1> m1) ≡ (<a2> m2) ⊣⊢@{SPROP} ⌜ a1 = a2 ⌠∧ + (∀ v lp, iMsg_car m1 v lp ≡ iMsg_car m2 v lp). + Proof. rewrite iProto_message_eq. apply proto_message_equivI. Qed. + + Lemma iProto_message_end_equivI `{!BiInternalEq SPROP} a m : + (<a> m) ≡ END ⊢@{SPROP} False. + Proof. rewrite iProto_message_eq iProto_end_eq. apply proto_message_end_equivI. Qed. + Lemma iProto_end_message_equivI `{!BiInternalEq SPROP} a m : + END ≡ (<a> m) ⊢@{SPROP} False. + Proof. by rewrite internal_eq_sym iProto_message_end_equivI. Qed. + + (** ** Non-expansiveness of operators *) + Global Instance iMsg_car_proper : + Proper (iMsg_equiv ==> (=) ==> (≡) ==> (≡)) (iMsg_car (Σ:=Σ) (V:=V)). + Proof. + intros m1 m2 meq v1 v2 veq p1 p2 peq. rewrite meq. + f_equiv; [ by f_equiv | done ]. + Qed. + Global Instance iMsg_car_ne n : + Proper (iMsg_dist n ==> (=) ==> (dist n) ==> (dist n)) (iMsg_car (Σ:=Σ) (V:=V)). + Proof. + intros m1 m2 meq v1 v2 veq p1 p2 peq. rewrite meq. + f_equiv; [ by f_equiv | done ]. + Qed. + + Global Instance iMsg_contractive v n : + Proper (dist n ==> dist_later n ==> dist n) (iMsg_base (Σ:=Σ) (V:=V) v). + Proof. rewrite iMsg_base_eq=> P1 P2 HP p1 p2 Hp w q /=. solve_contractive. Qed. + Global Instance iMsg_ne v : NonExpansive2 (iMsg_base (Σ:=Σ) (V:=V) v). + Proof. rewrite iMsg_base_eq=> P1 P2 HP p1 p2 Hp w q /=. solve_proper. Qed. + Global Instance iMsg_proper v : + Proper ((≡) ==> (≡) ==> (≡)) (iMsg_base (Σ:=Σ) (V:=V) v). + Proof. apply (ne_proper_2 _). Qed. + + Global Instance iMsg_exist_ne A n : + Proper (pointwise_relation _ (dist n) ==> (dist n)) (@iMsg_exist Σ V A). + Proof. rewrite iMsg_exist_eq=> m1 m2 Hm v p /=. f_equiv=> x. apply Hm. Qed. + Global Instance iMsg_exist_proper A : + Proper (pointwise_relation _ (≡) ==> (≡)) (@iMsg_exist Σ V A). + Proof. rewrite iMsg_exist_eq=> m1 m2 Hm v p /=. f_equiv=> x. apply Hm. Qed. + + Global Instance msg_tele_base (v:V) (P : iProp Σ) (p : iProto Σ V) : + MsgTele (TT:=TeleO) (MSG v {{ P }}; p) v P p. + Proof. done. Qed. + Global Instance msg_tele_exist {A} {TT : A → tele} (m : A → iMsg Σ V) tv tP tp : + (∀ x, MsgTele (TT:=TT x) (m x) (tv x) (tP x) (tp x)) → + MsgTele (TT:=TeleS TT) (∃ x, m x) tv tP tp. + Proof. intros Hm. rewrite /MsgTele /=. f_equiv=> x. apply Hm. Qed. + + Global Instance iProto_message_ne a : + NonExpansive (iProto_message (Σ:=Σ) (V:=V) a). + Proof. rewrite iProto_message_eq. solve_proper. Qed. + Global Instance iProto_message_proper a : + Proper ((≡) ==> (≡)) (iProto_message (Σ:=Σ) (V:=V) a). + Proof. apply (ne_proper _). Qed. + + Lemma iProto_message_equiv {TT1 TT2 : tele} a1 a2 + (m1 m2 : iMsg Σ V) + (v1 : TT1 -t> V) (v2 : TT2 -t> V) + (P1 : TT1 -t> iProp Σ) (P2 : TT2 -t> iProp Σ) + (prot1 : TT1 -t> iProto Σ V) (prot2 : TT2 -t> iProto Σ V) : + MsgTele m1 v1 P1 prot1 → + MsgTele m2 v2 P2 prot2 → + ⌜ a1 = a2 ⌠-∗ + (■∀.. (xs1 : TT1), tele_app P1 xs1 -∗ + ∃.. (xs2 : TT2), ⌜tele_app v1 xs1 = tele_app v2 xs2⌠∗ + â–· (tele_app prot1 xs1 ≡ tele_app prot2 xs2) ∗ + tele_app P2 xs2) -∗ + (■∀.. (xs2 : TT2), tele_app P2 xs2 -∗ + ∃.. (xs1 : TT1), ⌜tele_app v1 xs1 = tele_app v2 xs2⌠∗ + â–· (tele_app prot1 xs1 ≡ tele_app prot2 xs2) ∗ + tele_app P1 xs1) -∗ + (<a1> m1) ≡ (<a2> m2). + Proof. + iIntros (Hm1 Hm2 Heq) "#Heq1 #Heq2". + unfold MsgTele in Hm1. rewrite Hm1. clear Hm1. + unfold MsgTele in Hm2. rewrite Hm2. clear Hm2. + rewrite iProto_message_eq proto_message_equivI. + iSplit; [ done | ]. + iIntros (v p'). + do 2 rewrite iMsg_texist_exist. + rewrite iMsg_base_eq /=. + iApply prop_ext. + iIntros "!>". iSplit. + - iDestruct 1 as (xs1 Hveq1) "[Hrec1 HP1]". + iDestruct ("Heq1" with "HP1") as (xs2 Hveq2) "[Hrec2 HP2]". + iExists xs2. rewrite -Hveq1 Hveq2. + iSplitR; [ done | ]. iSplitR "HP2"; [ | done ]. + iRewrite -"Hrec1". iApply later_equivI. iIntros "!>". by iRewrite "Hrec2". + - iDestruct 1 as (xs2 Hveq2) "[Hrec2 HP2]". + iDestruct ("Heq2" with "HP2") as (xs1 Hveq1) "[Hrec1 HP1]". + iExists xs1. rewrite -Hveq2 Hveq1. + iSplitR; [ done | ]. iSplitR "HP1"; [ | done ]. + iRewrite -"Hrec2". iApply later_equivI. iIntros "!>". by iRewrite "Hrec1". + Qed. + + (** Helpers *) + Lemma iMsg_map_base f v P p : + NonExpansive f → + iMsg_map f (MSG v {{ P }}; p) ≡ (MSG v {{ P }}; f p)%msg. + Proof. + rewrite iMsg_base_eq. intros ? v' p'; simpl. iSplit. + - iDestruct 1 as (p'') "[(->&Hp&$) Hp']". iSplit; [done|]. + iRewrite "Hp'". iIntros "!>". by iRewrite "Hp". + - iIntros "(->&Hp'&$)". iExists p. iRewrite -"Hp'". auto. + Qed. + Lemma iMsg_map_exist {A} f (m : A → iMsg Σ V) : + iMsg_map f (∃ x, m x) ≡ (∃ x, iMsg_map f (m x))%msg. + Proof. + rewrite iMsg_exist_eq. intros v' p'; simpl. iSplit. + - iDestruct 1 as (p'') "[H Hp']". iDestruct "H" as (x) "H"; auto. + - iDestruct 1 as (x p'') "[Hm Hp']". auto. + Qed. + + (** ** Dual *) + Global Instance iProto_dual_ne : NonExpansive (@iProto_dual Σ V). + Proof. rewrite iProto_dual_eq. solve_proper. Qed. + Global Instance iProto_dual_proper : Proper ((≡) ==> (≡)) (@iProto_dual Σ V). + Proof. apply (ne_proper _). Qed. + Global Instance iProto_dual_if_ne d : NonExpansive (@iProto_dual_if Σ V d). + Proof. solve_proper. Qed. + Global Instance iProto_dual_if_proper d : + Proper ((≡) ==> (≡)) (@iProto_dual_if Σ V d). + Proof. apply (ne_proper _). Qed. + + Lemma iProto_dual_end : iProto_dual (Σ:=Σ) (V:=V) END ≡ END. + Proof. + rewrite iProto_end_eq iProto_dual_eq /iProto_dual_def /iProto_map_app. + etrans; [apply (fixpoint_unfold (iProto_map_app_aux _ _))|]; simpl. + by rewrite proto_elim_end. + Qed. + Lemma iProto_dual_message a m : + iProto_dual (<a> m) ≡ <action_dual a> iMsg_dual m. + Proof. + rewrite iProto_message_eq iProto_dual_eq /iProto_dual_def /iProto_map_app. + etrans; [apply (fixpoint_unfold (iProto_map_app_aux _ _))|]; simpl. + rewrite /iProto_message_def. rewrite ->proto_elim_message; [done|]. + intros a' m1 m2 Hm; f_equiv; solve_proper. + Qed. + Lemma iMsg_dual_base v P p : + iMsg_dual (MSG v {{ P }}; p) ≡ (MSG v {{ P }}; iProto_dual p)%msg. + Proof. apply iMsg_map_base, _. Qed. + Lemma iMsg_dual_exist {A} (m : A → iMsg Σ V) : + iMsg_dual (∃ x, m x) ≡ (∃ x, iMsg_dual (m x))%msg. + Proof. apply iMsg_map_exist. Qed. + + Global Instance iProto_dual_involutive : Involutive (≡) (@iProto_dual Σ V). + Proof. + intros p. apply (uPred.internal_eq_soundness (M:=iResUR Σ)). + iLöb as "IH" forall (p). destruct (iProto_case p) as [->|(a&n&m&->)]. + { by rewrite !iProto_dual_end. } + rewrite !iProto_dual_message involutive. + iApply iProto_message_equivI; iSplit; [done|]; iIntros (v p') "/=". + iApply prop_ext; iIntros "!>"; iSplit. + - iDestruct 1 as (pd) "[H Hp']". iRewrite "Hp'". + iDestruct "H" as (pdd) "[H #Hpd]". + iApply (internal_eq_rewrite); [|done]; iIntros "!>". + iRewrite "Hpd". by iRewrite ("IH" $! pdd). + - iIntros "H". destruct (Next_uninj p') as [p'' Hp']. iExists _. + rewrite Hp'. iSplitL; [by auto|]. iIntros "!>". by iRewrite ("IH" $! p''). + Qed. + + (** ** Append *) + Global Instance iProto_app_end_l : LeftId (≡) END (@iProto_app Σ V). + Proof. + intros p. rewrite iProto_end_eq iProto_app_eq /iProto_app_def /iProto_map_app. + etrans; [apply (fixpoint_unfold (iProto_map_app_aux _ _))|]; simpl. + by rewrite proto_elim_end. + Qed. + Lemma iProto_app_message a m p2 : (<a> m) <++> p2 ≡ <a> m <++> p2. + Proof. + rewrite iProto_message_eq iProto_app_eq /iProto_app_def /iProto_map_app. + etrans; [apply (fixpoint_unfold (iProto_map_app_aux _ _))|]; simpl. + rewrite /iProto_message_def. rewrite ->proto_elim_message; [done|]. + intros a' m1 m2 Hm. f_equiv; solve_proper. + Qed. + + Global Instance iProto_app_ne : NonExpansive2 (@iProto_app Σ V). + Proof. + assert (∀ n, Proper (dist n ==> (=) ==> dist n) (@iProto_app Σ V)). + { intros n p1 p1' Hp1 p2 p2' <-. by rewrite iProto_app_eq /iProto_app_def Hp1. } + assert (Proper ((≡) ==> (=) ==> (≡)) (@iProto_app Σ V)). + { intros p1 p1' Hp1 p2 p2' <-. by rewrite iProto_app_eq /iProto_app_def Hp1. } + intros n p1 p1' Hp1 p2 p2' Hp2. rewrite Hp1. clear p1 Hp1. + revert p1'. induction (lt_wf n) as [n _ IH]; intros p1. + destruct (iProto_case p1) as [->|(a&i&m&->)]. + { by rewrite !left_id. } + rewrite !iProto_app_message. f_equiv=> v p' /=. do 4 f_equiv. + f_contractive. apply IH; eauto using dist_le with lia. + Qed. + Global Instance iProto_app_proper : Proper ((≡) ==> (≡) ==> (≡)) (@iProto_app Σ V). + Proof. apply (ne_proper_2 _). Qed. + + Lemma iMsg_app_base v P p1 p2 : + ((MSG v {{ P }}; p1) <++> p2)%msg ≡ (MSG v {{ P }}; p1 <++> p2)%msg. + Proof. apply: iMsg_map_base. Qed. + Lemma iMsg_app_exist {A} (m : A → iMsg Σ V) p2 : + ((∃ x, m x) <++> p2)%msg ≡ (∃ x, m x <++> p2)%msg. + Proof. apply iMsg_map_exist. Qed. + + Global Instance iProto_app_end_r : RightId (≡) END (@iProto_app Σ V). + Proof. + intros p. apply (uPred.internal_eq_soundness (M:=iResUR Σ)). + iLöb as "IH" forall (p). destruct (iProto_case p) as [->|(a&i&m&->)]. + { by rewrite left_id. } + rewrite iProto_app_message. + iApply iProto_message_equivI; iSplit; [done|]; iIntros (v p') "/=". + iApply prop_ext; iIntros "!>". iSplit. + - iDestruct 1 as (p1') "[H Hp']". iRewrite "Hp'". + iApply (internal_eq_rewrite); [|done]; iIntros "!>". + by iRewrite ("IH" $! p1'). + - iIntros "H". destruct (Next_uninj p') as [p'' Hp']. iExists p''. + rewrite Hp'. iSplitL; [by auto|]. iIntros "!>". by iRewrite ("IH" $! p''). + Qed. + Global Instance iProto_app_assoc : Assoc (≡) (@iProto_app Σ V). + Proof. + intros p1 p2 p3. apply (uPred.internal_eq_soundness (M:=iResUR Σ)). + iLöb as "IH" forall (p1). destruct (iProto_case p1) as [->|(a&i&m&->)]. + { by rewrite !left_id. } + rewrite !iProto_app_message. + iApply iProto_message_equivI; iSplit; [done|]; iIntros (v p123) "/=". + iApply prop_ext; iIntros "!>". iSplit. + - iDestruct 1 as (p1') "[H #Hp']". + iExists (p1' <++> p2). iSplitL; [by auto|]. + iRewrite "Hp'". iIntros "!>". iApply "IH". + - iDestruct 1 as (p12) "[H #Hp123]". iDestruct "H" as (p1') "[H #Hp12]". + iExists p1'. iFrame "H". iRewrite "Hp123". + iIntros "!>". iRewrite "Hp12". by iRewrite ("IH" $! p1'). + Qed. + + Lemma iProto_dual_app p1 p2 : + iProto_dual (p1 <++> p2) ≡ iProto_dual p1 <++> iProto_dual p2. + Proof. + apply (uPred.internal_eq_soundness (M:=iResUR Σ)). + iLöb as "IH" forall (p1 p2). destruct (iProto_case p1) as [->|(a&i&m&->)]. + { by rewrite iProto_dual_end !left_id. } + rewrite iProto_dual_message !iProto_app_message iProto_dual_message /=. + iApply iProto_message_equivI; iSplit; [done|]; iIntros (v p12) "/=". + iApply prop_ext; iIntros "!>". iSplit. + - iDestruct 1 as (p12d) "[H #Hp12]". iDestruct "H" as (p1') "[H #Hp12d]". + iExists (iProto_dual p1'). iSplitL; [by auto|]. + iRewrite "Hp12". iIntros "!>". iRewrite "Hp12d". iApply "IH". + - iDestruct 1 as (p1') "[H #Hp12]". iDestruct "H" as (p1d) "[H #Hp1']". + iExists (p1d <++> p2). iSplitL; [by auto|]. + iRewrite "Hp12". iIntros "!>". iRewrite "Hp1'". by iRewrite ("IH" $! p1d p2). + Qed. + +End proto. + +Global Instance iProto_inhabited {Σ V} : Inhabited (iProto Σ V) := populate END. + +Program Definition iProto_elim {Σ V A} + (x : A) (f : action → iMsg Σ V -> A) (p : iProto Σ V) : A := + proto_elim x (λ a m, f a (IMsg (λ v, λne p, m v p)))%I p. +Next Obligation. solve_proper. Qed. + +Lemma iProto_elim_message {Σ V} {A:ofe} + (x : A) (f : action → iMsg Σ V -> A) a m + (* `{Hf : ∀ a, Proper ((≡) ==> (≡)) (f a)} : *) + : + iProto_elim x f (iProto_message a m) ≡ f a m. +Proof. + rewrite /iProto_elim. + rewrite iProto_message_eq /iProto_message_def. simpl. + setoid_rewrite proto_elim_message. + { f_equiv. destruct m. f_equiv. simpl. + admit. } + intros a'. + intros f1 f2 Hf'. f_equiv. f_equiv. +Admitted. + +Definition nat_beq := Eval compute in Nat.eqb. + +Definition find_recv {Σ V} (i:nat) (j:nat) (ps : list (iProto Σ V)) : + option $ iMsg Σ V := + iProto_elim None (λ a m, + match a with + | (Recv, i') => if nat_beq i i' then Some m else None + | (Send, _) => None + end) (ps !!! j). + +Fixpoint sync_pairs_aux {Σ V} (i : nat) (ps_full ps : list (iProto Σ V)) : + list (nat * nat * iMsg Σ V * iMsg Σ V) := + match ps with + | [] => [] + | p :: ps => + iProto_elim (sync_pairs_aux (S i) ps_full ps) + (λ a mi, match a with + | (Recv,_) => sync_pairs_aux (S i) ps_full ps + | (Send,j) => match find_recv i j ps_full with + | None => sync_pairs_aux (S i) ps_full ps + | Some mj => (i,j,mi,mj) :: + sync_pairs_aux (S i) ps_full ps + end + end) p + end. + +Notation sync_pairs ps := (sync_pairs_aux 0 ps ps). + +Definition can_step {Σ V} (rec : list (iProto Σ V) → iProp Σ) + (ps : list (iProto Σ V)) : iProp Σ := + [∧ list] '(i,j,m1,m2) ∈ sync_pairs ps, + ∀ v p1, iMsg_car m1 v (Next p1) -∗ + ∃ p2, iMsg_car m2 v (Next p2) ∗ + â–· (rec (<[i:=p1]>(<[j:=p2]>ps))). + +From iris.heap_lang Require Import notation. + +Definition iProto_binary `{!heapGS Σ} : list (iProto Σ val) := + [(<(Send, 1) @ (x:Z)> MSG #x ; END)%proto; + (<(Recv, 0) @ (x:Z)> MSG #x ; END)%proto]. + +Lemma iProto_binary_consistent `{!heapGS Σ} : + ⊢ can_step (λ _, True) (@iProto_binary _ Σ heapGS). +Proof. rewrite /iProto_binary /can_step /iProto_elim. simpl. + rewrite /find_recv. simpl. + Fail rewrite iProto_elim_message. + (* OBS: Break here *) + + +Definition valid_target {Σ V} (ps : list (iProto Σ V)) (i j : nat) : iProp Σ := + ∀ a m, (ps !!! i ≡ <(a, j)> m) -∗ ⌜is_Some (ps !! j)âŒ. + +Definition iProto_consistent_pre {Σ V} (rec : list (iProto Σ V) → iProp Σ) + (ps : list (iProto Σ V)) : iProp Σ := + (∀ i j, valid_target ps i j) ∗ (can_step rec ps). + +Global Instance iProto_consistent_pre_ne {Σ V} + (rec : listO (iProto Σ V) -n> iPropO Σ) : + NonExpansive (iProto_consistent_pre rec). +Proof. rewrite /iProto_consistent_pre /can_step /valid_target. solve_proper. Qed. + +Program Definition iProto_consistent_pre' {Σ V} + (rec : listO (iProto Σ V) -n> iPropO Σ) : + listO (iProto Σ V) -n> iPropO Σ := + λne ps, iProto_consistent_pre (λ ps, rec ps) ps. + +Local Instance iProto_consistent_pre_contractive {Σ V} : Contractive (@iProto_consistent_pre' Σ V). +Proof. + rewrite /iProto_consistent_pre' /iProto_consistent_pre /can_step. + solve_contractive. +Qed. + +Definition iProto_consistent {Σ V} (ps : list (iProto Σ V)) : iProp Σ := + fixpoint iProto_consistent_pre' ps. + +Arguments iProto_consistent {_ _} _%proto. +Global Instance: Params (@iProto_consistent) 1 := {}. + +Global Instance iProto_consistent_ne {Σ V} : NonExpansive (@iProto_consistent Σ V). +Proof. solve_proper. Qed. +Global Instance iProto_consistent_proper {Σ V} : Proper ((≡) ==> (⊣⊢)) (@iProto_consistent Σ V). +Proof. solve_proper. Qed. + +Lemma iProto_consistent_unfold {Σ V} (ps : list (iProto Σ V)) : + iProto_consistent ps ≡ (iProto_consistent_pre iProto_consistent) ps. +Proof. + apply: (fixpoint_unfold iProto_consistent_pre'). +Qed. + +(** * Protocol entailment *) +Definition iProto_le_pre {Σ V} + (rec : iProto Σ V → iProto Σ V → iProp Σ) (p1 p2 : iProto Σ V) : iProp Σ := + (p1 ≡ END ∗ p2 ≡ END) ∨ + ∃ a1 a2 m1 m2, + (p1 ≡ <a1> m1) ∗ (p2 ≡ <a2> m2) ∗ + match a1, a2 with + | (Recv,i), (Recv,j) => ⌜i = j⌠∗ ∀ v p1', + iMsg_car m1 v (Next p1') -∗ ∃ p2', â–· rec p1' p2' ∗ iMsg_car m2 v (Next p2') + | (Send,i), (Send,j) => ⌜i = j⌠∗ ∀ v p2', + iMsg_car m2 v (Next p2') -∗ ∃ p1', â–· rec p1' p2' ∗ iMsg_car m1 v (Next p1') + | _, _ => False + end. +Global Instance iProto_le_pre_ne {Σ V} (rec : iProto Σ V → iProto Σ V → iProp Σ) : + NonExpansive2 (iProto_le_pre rec). +Proof. solve_proper. Qed. + +Program Definition iProto_le_pre' {Σ V} + (rec : iProto Σ V -n> iProto Σ V -n> iPropO Σ) : + iProto Σ V -n> iProto Σ V -n> iPropO Σ := λne p1 p2, + iProto_le_pre (λ p1' p2', rec p1' p2') p1 p2. +Solve Obligations with solve_proper. +Local Instance iProto_le_pre_contractive {Σ V} : Contractive (@iProto_le_pre' Σ V). +Proof. + intros n rec1 rec2 Hrec p1 p2. rewrite /iProto_le_pre' /iProto_le_pre /=. + by repeat (f_contractive || f_equiv). +Qed. +Definition iProto_le {Σ V} (p1 p2 : iProto Σ V) : iProp Σ := + fixpoint iProto_le_pre' p1 p2. +Arguments iProto_le {_ _} _%proto _%proto. +Global Instance: Params (@iProto_le) 2 := {}. +Notation "p ⊑ q" := (iProto_le p q) : bi_scope. + +Global Instance iProto_le_ne {Σ V} : NonExpansive2 (@iProto_le Σ V). +Proof. solve_proper. Qed. +Global Instance iProto_le_proper {Σ V} : Proper ((≡) ==> (≡) ==> (⊣⊢)) (@iProto_le Σ V). +Proof. solve_proper. Qed. + +Record proto_name := ProtName { proto_names : gmap nat gname }. +Global Instance proto_name_inhabited : Inhabited proto_name := + populate (ProtName inhabitant). +Global Instance proto_name_eq_dec : EqDecision proto_name. +Proof. solve_decision. Qed. +Global Instance proto_name_countable : Countable proto_name. +Proof. + refine (inj_countable (λ '(ProtName γs), (γs)) + (λ '(γs), Some (ProtName γs)) _); by intros []. +Qed. + +Definition iProto_own_frag `{!protoG Σ V} (γ : gname) + (i : nat) (p : iProto Σ V) : iProp Σ := + own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p))). + +Definition iProto_own_auth `{!protoG Σ V} (γ : gname) + (ps : list (iProto Σ V)) : iProp Σ := + own γ (gmap_view_auth (DfracOwn 1) (((λ p, Excl' (Next p)) <$> + (list_to_map (zip (seq 0 (length ps)) ps))) : gmap _ _)). + +Definition iProto_ctx `{protoG Σ V} + (γ : gname) (ps_len : nat) : iProp Σ := + ∃ ps, ⌜length ps = ps_len⌠∗ iProto_own_auth γ ps ∗ â–· iProto_consistent ps. + +(** * The connective for ownership of channel ends *) +Definition iProto_own `{!protoG Σ V} + (γ : gname) (i : nat) (p : iProto Σ V) : iProp Σ := + ∃ p', â–· (p' ⊑ p) ∗ iProto_own_frag γ i p'. +Arguments iProto_own {_ _ _} _ _ _%proto. +Global Instance: Params (@iProto_own) 3 := {}. + +Global Instance iProto_own_frag_contractive `{protoG Σ V} γ i : + Contractive (iProto_own_frag γ i). +Proof. solve_contractive. Qed. + +Global Instance iProto_own_contractive `{protoG Σ V} γ i : + Contractive (iProto_own γ i). +Proof. solve_contractive. Qed. +Global Instance iProto_own_ne `{protoG Σ V} γ s : NonExpansive (iProto_own γ s). +Proof. solve_proper. Qed. +Global Instance iProto_own_proper `{protoG Σ V} γ s : + Proper ((≡) ==> (≡)) (iProto_own γ s). +Proof. apply (ne_proper _). Qed. + +(** * Proofs *) +Section proto. + Context `{!protoG Σ V}. + Implicit Types v : V. + Implicit Types p pl pr : iProto Σ V. + Implicit Types m : iMsg Σ V. + + Lemma own_prot_idx γ i j (p1 p2 : iProto Σ V) : + own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p1))) -∗ + own γ (gmap_view_frag j (DfracOwn 1) (Excl' (Next p2))) -∗ + ⌜i ≠jâŒ. + Proof. + iIntros "Hown Hown'" (->). + iDestruct (own_valid_2 with "Hown Hown'") as "H". + rewrite uPred.cmra_valid_elim. + by iDestruct "H" as %[]%gmap_view_frag_op_validN. + Qed. + + Lemma own_prot_excl γ i (p1 p2 : iProto Σ V) : + own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p1))) -∗ + own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p2))) -∗ + False. + Proof. iIntros "Hi Hj". by iDestruct (own_prot_idx with "Hi Hj") as %?. Qed. + + (** ** Protocol entailment **) + Lemma iProto_le_unfold p1 p2 : iProto_le p1 p2 ≡ iProto_le_pre iProto_le p1 p2. + Proof. apply: (fixpoint_unfold iProto_le_pre'). Qed. + + Lemma iProto_le_end : ⊢ END ⊑ (END : iProto Σ V). + Proof. rewrite iProto_le_unfold. iLeft. auto 10. Qed. + + Lemma iProto_le_end_inv_r p : p ⊑ END -∗ (p ≡ END). + Proof. + rewrite iProto_le_unfold. iIntros "[[Hp _]|H] //". + iDestruct "H" as (a1 a2 m1 m2) "(_ & Heq & _)". + by iDestruct (iProto_end_message_equivI with "Heq") as %[]. + Qed. + + Lemma iProto_le_end_inv_l p : END ⊑ p -∗ (p ≡ END). + Proof. + rewrite iProto_le_unfold. iIntros "[[_ Hp]|H] //". + iDestruct "H" as (a1 a2 m1 m2) "(Heq & _ & _)". + iDestruct (iProto_end_message_equivI with "Heq") as %[]. + Qed. + + Lemma iProto_le_send_inv i p1 m2 : + p1 ⊑ (<(Send,i)> m2) -∗ ∃ m1, + (p1 ≡ <(Send,i)> m1) ∗ + ∀ v p2', iMsg_car m2 v (Next p2') -∗ + ∃ p1', â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1'). + Proof. + rewrite iProto_le_unfold. + iIntros "[[_ Heq]|H]". + { by iDestruct (iProto_message_end_equivI with "Heq") as %[]. } + iDestruct "H" as (a1 a2 m1 m2') "(Hp1 & Hp2 & H)". + rewrite iProto_message_equivI. iDestruct "Hp2" as "[%Heq Hm2]". + simplify_eq. + destruct a1 as [[]]; [|done]. + iDestruct "H" as (->) "H". iExists m1. iFrame "Hp1". + iIntros (v p2). iSpecialize ("Hm2" $! v (Next p2)). by iRewrite "Hm2". + Qed. + + Lemma iProto_le_send_send_inv i m1 m2 v p2' : + (<(Send,i)> m1) ⊑ (<(Send,i)> m2) -∗ + iMsg_car m2 v (Next p2') -∗ ∃ p1', â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1'). + Proof. + iIntros "H Hm2". iDestruct (iProto_le_send_inv with "H") as (m1') "[Hm1 H]". + iDestruct (iProto_message_equivI with "Hm1") as (Heq) "Hm1". + iDestruct ("H" with "Hm2") as (p1') "[Hle Hm]". + iRewrite -("Hm1" $! v (Next p1')) in "Hm". auto with iFrame. + Qed. + + Lemma iProto_le_recv_inv_l i m1 p2 : + (<(Recv,i)> m1) ⊑ p2 -∗ ∃ m2, + (p2 ≡ <(Recv,i)> m2) ∗ + ∀ v p1', iMsg_car m1 v (Next p1') -∗ + ∃ p2', â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). + Proof. + rewrite iProto_le_unfold. + iIntros "[[Heq _]|H]". + { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } + iDestruct "H" as (a1 a2 m1' m2) "(Hp1 & Hp2 & H)". + rewrite iProto_message_equivI. iDestruct "Hp1" as "[%Heq Hm1]". + simplify_eq. + destruct a2 as [[]]; [done|]. + iDestruct "H" as (->) "H". iExists m2. iFrame "Hp2". + iIntros (v p1). iSpecialize ("Hm1" $! v (Next p1)). by iRewrite "Hm1". + Qed. + + Lemma iProto_le_recv_inv_r i p1 m2 : + (p1 ⊑ <(Recv,i)> m2) -∗ ∃ m1, + (p1 ≡ <(Recv,i)> m1) ∗ + ∀ v p1', iMsg_car m1 v (Next p1') -∗ + ∃ p2', â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). + Proof. + rewrite iProto_le_unfold. + iIntros "[[_ Heq]|H]". + { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } + iDestruct "H" as (a1 a2 m1 m2') "(Hp1 & Hp2 & H)". + rewrite iProto_message_equivI. + iDestruct "Hp2" as "[%Heq Hm2]". + simplify_eq. + destruct a1 as [[]]; [done|]. + iDestruct "H" as (->) "H". + iExists m1. iFrame. + iIntros (v p2). + iIntros "Hm1". iDestruct ("H" with "Hm1") as (p2') "[Hle H]". + iSpecialize ("Hm2" $! v (Next p2')). + iExists p2'. iFrame. + iRewrite "Hm2". iApply "H". + Qed. + + Lemma iProto_le_recv_recv_inv i m1 m2 v p1' : + (<(Recv, i)> m1) ⊑ (<(Recv, i)> m2) -∗ + iMsg_car m1 v (Next p1') -∗ ∃ p2', â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). + Proof. + iIntros "H Hm2". iDestruct (iProto_le_recv_inv_r with "H") as (m1') "[Hm1 H]". + iApply "H". iDestruct (iProto_message_equivI with "Hm1") as (_) "Hm1". + by iRewrite -("Hm1" $! v (Next p1')). + Qed. + + Lemma iProto_le_msg_inv_l i a m1 p2 : + (<(a,i)> m1) ⊑ p2 -∗ ∃ m2, p2 ≡ <(a,i)> m2. + Proof. + rewrite iProto_le_unfold /iProto_le_pre. + iIntros "[[Heq _]|H]". + { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } + iDestruct "H" as (a1 a2 m1' m2) "(Hp1 & Hp2 & H)". + destruct a1 as [t1 ?], a2 as [t2 ?]. + destruct t1,t2; [|done|done|]. + - rewrite iProto_message_equivI. + iDestruct "Hp1" as (Heq) "Hp1". simplify_eq. + iDestruct "H" as (->) "H". by iExists _. + - rewrite iProto_message_equivI. + iDestruct "Hp1" as (Heq) "Hp1". simplify_eq. + iDestruct "H" as (->) "H". by iExists _. + Qed. + + Lemma iProto_le_msg_inv_r j a p1 m2 : + (p1 ⊑ <(a,j)> m2) -∗ ∃ m1, p1 ≡ <(a,j)> m1. + Proof. + rewrite iProto_le_unfold /iProto_le_pre. + iIntros "[[_ Heq]|H]". + { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } + iDestruct "H" as (a1 a2 m1 m2') "(Hp1 & Hp2 & H)". + destruct a1 as [t1 ?], a2 as [t2 ?]. + destruct t1,t2; [|done|done|]. + - rewrite iProto_message_equivI. + iDestruct "Hp2" as (Heq) "Hp2". simplify_eq. + iDestruct "H" as (->) "H". by iExists _. + - rewrite iProto_message_equivI. + iDestruct "Hp2" as (Heq) "Hp2". simplify_eq. + iDestruct "H" as (->) "H". by iExists _. + Qed. + + Lemma valid_target_le ps i p1 p2 : + (∀ i' j', valid_target ps i' j') -∗ + ps !!! i ≡ p1 -∗ + p1 ⊑ p2 -∗ + (∀ i' j', valid_target (<[i := p2]>ps) i' j') ∗ p1 ⊑ p2. + Proof. Admitted. + (* iIntros "Hprot #HSome Hle". *) + (* pose proof (iProto_case p1) as [Hend|Hmsg]. *) + (* { rewrite Hend. iDestruct (iProto_le_end_inv_l with "Hle") as "#H". *) + (* iFrame "Hle". *) + (* iIntros (i' j' a m) "Hm". *) + (* destruct (decide (i = j')) as [->|Hneqj]. *) + (* { Search list_lookup_total insert. rewrite list_lookup_total_insert. ; [done|]. lia. done. } *) + (* rewrite lookup_insert_ne; [|done]. *) + (* destruct (decide (i = i')) as [->|Hneqi]. *) + (* { rewrite lookup_total_insert. iRewrite "H" in "Hm". *) + (* by iDestruct (iProto_end_message_equivI with "Hm") as "Hm". } *) + (* rewrite lookup_total_insert_ne; [|done]. *) + (* by iApply "Hprot". } *) + (* destruct Hmsg as (t & n & m & Hmsg). *) + (* setoid_rewrite Hmsg. *) + (* iDestruct (iProto_le_msg_inv_l with "Hle") as (m2) "#Heq". iFrame "Hle". *) + (* iIntros (i' j' a m') "Hm". *) + (* destruct (decide (i = j')) as [->|Hneqj]. *) + (* { rewrite lookup_insert. done. } *) + (* rewrite lookup_insert_ne; [|done]. *) + (* destruct (decide (i = i')) as [->|Hneqi]. *) + (* { rewrite lookup_total_insert. iRewrite "Heq" in "Hm". *) + (* iDestruct (iProto_message_equivI with "Hm") as (Heq) "Hm". *) + (* simplify_eq. by iApply "Hprot". } *) + (* rewrite lookup_total_insert_ne; [|done]. *) + (* by iApply "Hprot". *) + (* Qed. *) + + Lemma iProto_consistent_le ps i p1 p2 : + iProto_consistent ps -∗ + ps !!! i ≡ p1 -∗ + p1 ⊑ p2 -∗ + iProto_consistent (<[i := p2]>ps). + Proof. + iIntros "Hprot #HSome Hle". + iRevert "HSome". + iLöb as "IH" forall (p1 p2 ps). + iIntros "#HSome". + rewrite !iProto_consistent_unfold. + iDestruct "Hprot" as "(Htar & Hprot)". + iDestruct (valid_target_le with "Htar HSome Hle") as "[Htar Hle]". + iFrame. + iIntros (i' j' m1 m2) "#Hm1 #Hm2". + destruct (decide (i = i')) as [<-|Hneq]. + { rewrite list_lookup_total_insert; [|admit]. + pose proof (iProto_case p2) as [Hend|Hmsg]. + { setoid_rewrite Hend. rewrite iProto_end_message_equivI. done. } + destruct Hmsg as (a&?&m&Hmsg). + setoid_rewrite Hmsg. + destruct a; last first. + { rewrite iProto_message_equivI. + iDestruct "Hm1" as "[%Htag Hm1]". done. } + rewrite iProto_message_equivI. + iDestruct "Hm1" as "[%Htag Hm1]". + inversion Htag. simplify_eq. + iIntros (v p) "Hm1'". + iSpecialize ("Hm1" $! v (Next p)). + iDestruct (iProto_le_send_inv with "Hle") as "Hle". + iRewrite -"Hm1" in "Hm1'". + iDestruct "Hle" as (m') "[#Heq H]". + iDestruct ("H" with "Hm1'") as (p') "[Hle H]". + destruct (decide (i = j')) as [<-|Hneq]. + { rewrite list_lookup_total_insert. rewrite iProto_message_equivI. + iDestruct "Hm2" as "[%Heq _]". done. admit. } + iDestruct ("Hprot" $!i j' with "[] [] H") as "Hprot". + { iRewrite -"Heq". rewrite !list_lookup_total_alt. iRewrite "HSome". done. } + { rewrite list_lookup_total_insert_ne; [|done]. done. } + iDestruct "Hprot" as (p'') "[Hm Hprot]". + iExists p''. iFrame. + iNext. + iDestruct ("IH" with "Hprot Hle [HSome]") as "HI". + { rewrite list_lookup_total_insert; [done|]. admit. } + iClear "IH Hm1 Hm2 Heq". + rewrite list_insert_insert. + rewrite (list_insert_commute _ j' i); [|done]. + rewrite list_insert_insert. done. } + rewrite list_lookup_total_insert_ne; [|done]. + destruct (decide (i = j')) as [<-|Hneq']. + { rewrite list_lookup_total_insert. + pose proof (iProto_case p2) as [Hend|Hmsg]. + { setoid_rewrite Hend. rewrite iProto_end_message_equivI. done. } + destruct Hmsg as (a&?&m&Hmsg). + setoid_rewrite Hmsg. + destruct a. + { rewrite iProto_message_equivI. + iDestruct "Hm2" as "[%Htag Hm2]". done. } + rewrite iProto_message_equivI. + iDestruct "Hm2" as "[%Htag Hm2]". + inversion Htag. simplify_eq. + iIntros (v p) "Hm1'". + iDestruct (iProto_le_recv_inv_r with "Hle") as "Hle". + iDestruct "Hle" as (m') "[#Heq Hle]". + iDestruct ("Hprot" $!i' with "[] [] Hm1'") as "Hprot". + { done. } + { rewrite !list_lookup_total_alt. iRewrite "HSome". done. } + iDestruct ("Hprot") as (p') "[Hm1' Hprot]". + iDestruct ("Hle" with "Hm1'") as (p2') "[Hle Hm']". + iSpecialize ("Hm2" $! v (Next p2')). + iExists p2'. + iRewrite -"Hm2". iFrame. + iDestruct ("IH" with "Hprot Hle []") as "HI". + { iPureIntro. rewrite list_lookup_total_insert_ne; [|done]. + rewrite list_lookup_total_insert. done. admit. } + rewrite list_insert_commute; [|done]. + rewrite !list_insert_insert. done. admit. } + rewrite list_lookup_total_insert_ne; [|done]. + iIntros (v p) "Hm1'". + iDestruct ("Hprot" $!i' j' with "[//] [//] Hm1'") as "Hprot". + iDestruct "Hprot" as (p') "[Hm2' Hprot]". + iExists p'. iFrame. + iNext. + rewrite (list_insert_commute _ j' i); [|done]. + rewrite (list_insert_commute _ i' i); [|done]. + iApply ("IH" with "Hprot Hle []"). + rewrite list_lookup_total_insert_ne; [|done]. + rewrite list_lookup_total_insert_ne; [|done]. + done. + Admitted. + + Lemma iProto_le_send i m1 m2 : + (∀ v p2', iMsg_car m2 v (Next p2') -∗ ∃ p1', + â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1')) -∗ + (<(Send,i)> m1) ⊑ (<(Send,i)> m2). + Proof. + iIntros "Hle". rewrite iProto_le_unfold. + iRight. iExists (Send, i), (Send, i), m1, m2. by eauto. + Qed. + + Lemma iProto_le_recv i m1 m2 : + (∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', + â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2')) -∗ + (<(Recv,i)> m1) ⊑ (<(Recv,i)> m2). + Proof. + iIntros "Hle". rewrite iProto_le_unfold. + iRight. iExists (Recv, i), (Recv, i), m1, m2. by eauto. + Qed. + + Lemma iProto_le_base a v P p1 p2 : + â–· (p1 ⊑ p2) -∗ + (<a> MSG v {{ P }}; p1) ⊑ (<a> MSG v {{ P }}; p2). + Proof. + rewrite iMsg_base_eq. iIntros "H". destruct a as [[]]. + - iApply iProto_le_send. iIntros (v' p') "(->&Hp&$)". + iExists p1. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". + - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&$)". + iExists p2. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". + Qed. + + Lemma iProto_le_trans p1 p2 p3 : p1 ⊑ p2 -∗ p2 ⊑ p3 -∗ p1 ⊑ p3. + Proof. + iIntros "H1 H2". iLöb as "IH" forall (p1 p2 p3). + destruct (iProto_case p3) as [->|([]&i&m3&->)]. + - iDestruct (iProto_le_end_inv_r with "H2") as "H2". by iRewrite "H2" in "H1". + - iDestruct (iProto_le_send_inv with "H2") as (m2) "[Hp2 H2]". + iRewrite "Hp2" in "H1"; clear p2. + iDestruct (iProto_le_send_inv with "H1") as (m1) "[Hp1 H1]". + iRewrite "Hp1"; clear p1. + iApply iProto_le_send. iIntros (v p3') "Hm3". + iDestruct ("H2" with "Hm3") as (p2') "[Hle Hm2]". + iDestruct ("H1" with "Hm2") as (p1') "[Hle' Hm1]". + iExists p1'. iIntros "{$Hm1} !>". by iApply ("IH" with "Hle'"). + - iDestruct (iProto_le_recv_inv_r with "H2") as (m2) "[Hp2 H3]". + iRewrite "Hp2" in "H1". + iDestruct (iProto_le_recv_inv_r with "H1") as (m1) "[Hp1 H2]". + iRewrite "Hp1". iApply iProto_le_recv. iIntros (v p1') "Hm1". + iDestruct ("H2" with "Hm1") as (p2') "[Hle Hm2]". + iDestruct ("H3" with "Hm2") as (p3') "[Hle' Hm3]". + iExists p3'. iIntros "{$Hm3} !>". by iApply ("IH" with "Hle"). + Qed. + + Lemma iProto_le_refl p : ⊢ p ⊑ p. + Proof. + iLöb as "IH" forall (p). destruct (iProto_case p) as [->|([]&i&m&->)]. + - iApply iProto_le_end. + - iApply iProto_le_send. auto 10 with iFrame. + - iApply iProto_le_recv. auto 10 with iFrame. + Qed. + + Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. + Proof. + iIntros "H". iLöb as "IH" forall (p1 p2). + destruct (iProto_case p1) as [->|([]&i&m1&->)]. + - iDestruct (iProto_le_end_inv_r with "H") as "H". + iRewrite "H". iApply iProto_le_refl. + - iDestruct (iProto_le_send_inv with "H") as (m2) "[Hp2 H]". + iRewrite "Hp2"; clear p2. iEval (rewrite !iProto_dual_message). + iApply iProto_le_recv. iIntros (v p1d). + iDestruct 1 as (p1') "[Hm1 #Hp1d]". + iDestruct ("H" with "Hm1") as (p2') "[H Hm2]". + iDestruct ("IH" with "H") as "H". iExists (iProto_dual p2'). + iSplitL "H"; [iIntros "!>"; by iRewrite "Hp1d"|]. simpl; auto. + - iDestruct (iProto_le_recv_inv_r with "H") as (m2) "[Hp2 H]". + iRewrite "Hp2"; clear p2. iEval (rewrite !iProto_dual_message /=). + iApply iProto_le_send. iIntros (v p2d). + iDestruct 1 as (p2') "[Hm2 #Hp2d]". + iDestruct ("H" with "Hm2") as (p1') "[H Hm1]". + iDestruct ("IH" with "H") as "H". iExists (iProto_dual p1'). + iSplitL "H"; [iIntros "!>"; by iRewrite "Hp2d"|]. simpl; auto. + Qed. + + Lemma iProto_le_dual_l p1 p2 : iProto_dual p2 ⊑ p1 ⊢ iProto_dual p1 ⊑ p2. + Proof. + iIntros "H". iEval (rewrite -(involutive iProto_dual p2)). + by iApply iProto_le_dual. + Qed. + Lemma iProto_le_dual_r p1 p2 : p2 ⊑ iProto_dual p1 ⊢ p1 ⊑ iProto_dual p2. + Proof. + iIntros "H". iEval (rewrite -(involutive iProto_dual p1)). + by iApply iProto_le_dual. + Qed. + + Lemma iProto_le_app p1 p2 p3 p4 : + p1 ⊑ p2 -∗ p3 ⊑ p4 -∗ p1 <++> p3 ⊑ p2 <++> p4. + Proof. + iIntros "H1 H2". iLöb as "IH" forall (p1 p2 p3 p4). + destruct (iProto_case p2) as [->|([]&i&m2&->)]. + - iDestruct (iProto_le_end_inv_r with "H1") as "H1". + iRewrite "H1". by rewrite !left_id. + - iDestruct (iProto_le_send_inv with "H1") as (m1) "[Hp1 H1]". + iRewrite "Hp1"; clear p1. rewrite !iProto_app_message. + iApply iProto_le_send. iIntros (v p24). + iDestruct 1 as (p2') "[Hm2 #Hp24]". + iDestruct ("H1" with "Hm2") as (p1') "[H1 Hm1]". + iExists (p1' <++> p3). iSplitR "Hm1"; [|by simpl; eauto]. + iIntros "!>". iRewrite "Hp24". by iApply ("IH" with "H1"). + - iDestruct (iProto_le_recv_inv_r with "H1") as (m1) "[Hp1 H1]". + iRewrite "Hp1"; clear p1. rewrite !iProto_app_message. + iApply iProto_le_recv. + iIntros (v p13). iDestruct 1 as (p1') "[Hm1 #Hp13]". + iDestruct ("H1" with "Hm1") as (p2'') "[H1 Hm2]". + iExists (p2'' <++> p4). iSplitR "Hm2"; [|by simpl; eauto]. + iIntros "!>". iRewrite "Hp13". by iApply ("IH" with "H1"). + Qed. + + Lemma iProto_le_payload_elim_l i m v P p : + (P -∗ (<(Recv,i)> MSG v; p) ⊑ (<(Recv,i)> m)) ⊢ + (<(Recv,i)> MSG v {{ P }}; p) ⊑ <(Recv,i)> m. + Proof. + rewrite iMsg_base_eq. iIntros "H". + iApply iProto_le_recv. iIntros (v' p') "(->&Hp&HP)". + iApply (iProto_le_recv_recv_inv with "(H HP)"); simpl; auto. + Qed. + Lemma iProto_le_payload_elim_r i m v P p : + (P -∗ (<(Send, i)> m) ⊑ (<(Send, i)> MSG v; p)) ⊢ + (<(Send,i)> m) ⊑ (<(Send,i)> MSG v {{ P }}; p). + Proof. + rewrite iMsg_base_eq. iIntros "H". + iApply iProto_le_send. iIntros (v' p') "(->&Hp&HP)". + iApply (iProto_le_send_send_inv with "(H HP)"); simpl; auto. + Qed. + Lemma iProto_le_payload_intro_l i v P p : + P -∗ (<(Send,i)> MSG v {{ P }}; p) ⊑ (<(Send,i)> MSG v; p). + Proof. + rewrite iMsg_base_eq. + iIntros "HP". iApply iProto_le_send. iIntros (v' p') "(->&Hp&_) /=". + iExists p'. iSplitR; [iApply iProto_le_refl|]. auto. + Qed. + Lemma iProto_le_payload_intro_r i v P p : + P -∗ (<(Recv,i)> MSG v; p) ⊑ (<(Recv,i)> MSG v {{ P }}; p). + Proof. + rewrite iMsg_base_eq. + iIntros "HP". iApply iProto_le_recv. iIntros (v' p') "(->&Hp&_) /=". + iExists p'. iSplitR; [iApply iProto_le_refl|]. auto. + Qed. + Lemma iProto_le_exist_elim_l {A} i (m1 : A → iMsg Σ V) m2 : + (∀ x, (<(Recv,i)> m1 x) ⊑ (<(Recv,i)> m2)) ⊢ + (<(Recv,i) @ x> m1 x) ⊑ (<(Recv,i)> m2). + Proof. + rewrite iMsg_exist_eq. iIntros "H". + iApply iProto_le_recv. iIntros (v p1') "/=". iDestruct 1 as (x) "Hm". + by iApply (iProto_le_recv_recv_inv with "H"). + Qed. + Lemma iProto_le_exist_elim_r {A} i m1 (m2 : A → iMsg Σ V) : + (∀ x, (<(Send,i)> m1) ⊑ (<(Send,i)> m2 x)) ⊢ + (<(Send,i)> m1) ⊑ (<(Send,i) @ x> m2 x). + Proof. + rewrite iMsg_exist_eq. iIntros "H". + iApply iProto_le_send. iIntros (v p2'). iDestruct 1 as (x) "Hm". + by iApply (iProto_le_send_send_inv with "H"). + Qed. + Lemma iProto_le_exist_intro_l {A} i (m : A → iMsg Σ V) a : + ⊢ (<(Send,i) @ x> m x) ⊑ (<(Send,i)> m a). + Proof. + rewrite iMsg_exist_eq. iApply iProto_le_send. iIntros (v p') "Hm /=". + iExists p'. iSplitR; last by auto. iApply iProto_le_refl. + Qed. + Lemma iProto_le_exist_intro_r {A} i (m : A → iMsg Σ V) a : + ⊢ (<(Recv,i)> m a) ⊑ (<(Recv,i) @ x> m x). + Proof. + rewrite iMsg_exist_eq. iApply iProto_le_recv. iIntros (v p') "Hm /=". + iExists p'. iSplitR; last by auto. iApply iProto_le_refl. + Qed. + + Lemma iProto_le_texist_elim_l {TT : tele} i (m1 : TT → iMsg Σ V) m2 : + (∀ x, (<(Recv,i)> m1 x) ⊑ (<(Recv,i)> m2)) ⊢ + (<(Recv,i) @.. x> m1 x) ⊑ (<(Recv,i)> m2). + Proof. + iIntros "H". iInduction TT as [|T TT] "IH"; simpl; [done|]. + iApply iProto_le_exist_elim_l; iIntros (x). + iApply "IH". iIntros (xs). iApply "H". + Qed. + Lemma iProto_le_texist_elim_r {TT : tele} i m1 (m2 : TT → iMsg Σ V) : + (∀ x, (<(Send,i)> m1) ⊑ (<(Send,i)> m2 x)) -∗ + (<(Send,i)> m1) ⊑ (<(Send,i) @.. x> m2 x). + Proof. + iIntros "H". iInduction TT as [|T TT] "IH"; simpl; [done|]. + iApply iProto_le_exist_elim_r; iIntros (x). + iApply "IH". iIntros (xs). iApply "H". + Qed. + + Lemma iProto_le_texist_intro_l {TT : tele} i (m : TT → iMsg Σ V) x : + ⊢ (<(Send,i) @.. x> m x) ⊑ (<(Send,i)> m x). + Proof. + induction x as [|T TT x xs IH] using tele_arg_ind; simpl. + { iApply iProto_le_refl. } + iApply iProto_le_trans; [by iApply iProto_le_exist_intro_l|]. iApply IH. + Qed. + Lemma iProto_le_texist_intro_r {TT : tele} i (m : TT → iMsg Σ V) x : + ⊢ (<(Recv,i)> m x) ⊑ (<(Recv,i) @.. x> m x). + Proof. + induction x as [|T TT x xs IH] using tele_arg_ind; simpl. + { iApply iProto_le_refl. } + iApply iProto_le_trans; [|by iApply iProto_le_exist_intro_r]. iApply IH. + Qed. + + Lemma iProto_consistent_target ps m a i j : + iProto_consistent ps -∗ + ps !!! i ≡ (<(a, j)> m) -∗ + ⌜is_Some (ps !! j)âŒ. + Proof. + rewrite iProto_consistent_unfold. iDestruct 1 as "[Htar _]". iApply "Htar". + Qed. + + Lemma iProto_consistent_step ps m1 m2 i j v p1 : + iProto_consistent ps -∗ + ps !!! i ≡ (<(Send, j)> m1) -∗ + ps !!! j ≡ (<(Recv, i)> m2) -∗ + iMsg_car m1 v (Next p1) -∗ + ∃ p2, iMsg_car m2 v (Next p2) ∗ + â–· iProto_consistent (<[i := p1]>(<[j := p2]>ps)). + Proof. + iIntros "Hprot #Hi #Hj Hm1". + rewrite iProto_consistent_unfold /iProto_consistent_pre. + iDestruct "Hprot" as "[_ Hprot]". + iDestruct ("Hprot" with "Hi Hj Hm1") as (p2) "[Hm2 Hprot]". + iExists p2. iFrame. + Qed. + + Global Instance iProto_own_frag_ne γ s : NonExpansive (iProto_own_frag γ s). + Proof. solve_proper. Qed. + + Lemma iProto_own_auth_agree γ ps i p : + iProto_own_auth γ ps -∗ iProto_own_frag γ i p -∗ â–· (ps !!! i ≡ p). + Proof. Admitted. + (* iIntros "Hâ— Hâ—¯". *) + (* iDestruct (own_valid_2 with "Hâ— Hâ—¯") as "H✓". *) + (* rewrite gmap_view_both_validI. *) + (* iDestruct "H✓" as "[_ [H1 H2]]". *) + (* rewrite list_lookup_total_alt lookup_fmap. *) + (* destruct (ps !! i); last first. *) + (* { rewrite !option_equivI. } *) + (* simpl. rewrite !option_equivI excl_equivI. by iNext. *) + (* Qed. *) + + Lemma iProto_own_auth_update γ ps i p p' : + iProto_own_auth γ ps -∗ iProto_own_frag γ i p ==∗ + iProto_own_auth γ (<[i := p']>ps) ∗ iProto_own_frag γ i p'. + Proof. + iIntros "Hâ— Hâ—¯". + iMod (own_update_2 with "Hâ— Hâ—¯") as "[H1 H2]"; [|iModIntro]. + { eapply (gmap_view_replace _ _ _ (Excl' (Next p'))). done. } + iFrame. rewrite -fmap_insert. Admitted. + + Lemma iProto_own_auth_alloc ps : + ⊢ |==> ∃ γ, iProto_own_auth γ ps ∗ [∗ list] i ↦p ∈ ps, iProto_own γ i p. + Proof. Admitted. + (* iMod (own_alloc (gmap_view_auth (DfracOwn 1) ∅)) as (γ) "Hauth". *) + (* { apply gmap_view_auth_valid. } *) + (* iExists γ. *) + (* iInduction ps as [|i p ps Hin] "IH" using map_ind. *) + (* { iModIntro. iFrame. by iApply big_sepM_empty. } *) + (* iMod ("IH" with "Hauth") as "[Hauth Hfrags]". *) + (* rewrite big_sepM_insert; [|done]. iFrame "Hfrags". *) + (* iMod (own_update with "Hauth") as "[Hauth Hfrag]". *) + (* { apply (gmap_view_alloc _ i (DfracOwn 1) (Excl' (Next p))); [|done|done]. *) + (* by rewrite lookup_fmap Hin. } *) + (* iModIntro. rewrite -fmap_insert. iFrame. *) + (* iExists _. iFrame. iApply iProto_le_refl. *) + (* Qed. *) + + Lemma iProto_own_le γ s p1 p2 : + iProto_own γ s p1 -∗ â–· (p1 ⊑ p2) -∗ iProto_own γ s p2. + Proof. + iDestruct 1 as (p1') "[Hle H]". iIntros "Hle'". + iExists p1'. iFrame "H". by iApply (iProto_le_trans with "Hle"). + Qed. + + Lemma iProto_init ps : + â–· iProto_consistent ps -∗ + |==> ∃ γ, iProto_ctx γ (length ps) ∗ [∗ list] i ↦p ∈ ps, iProto_own γ i p. + Proof. + iIntros "Hconsistent". + iMod iProto_own_auth_alloc as (γ) "[Hauth Hfrags]". + iExists γ. iFrame. iExists _. by iFrame. + Qed. + + + Lemma iProto_step γ ps_dom i j m1 m2 p1 v : + iProto_ctx γ ps_dom -∗ + iProto_own γ i (<(Send, j)> m1) -∗ + iProto_own γ j (<(Recv, i)> m2) -∗ + iMsg_car m1 v (Next p1) ==∗ + â–· ∃ p2, iMsg_car m2 v (Next p2) ∗ iProto_ctx γ ps_dom ∗ + iProto_own γ i p1 ∗ iProto_own γ j p2. + Proof. + iIntros "Hctx Hi Hj Hm". + iDestruct "Hi" as (pi) "[Hile Hi]". + iDestruct "Hj" as (pj) "[Hjle Hj]". + iDestruct "Hctx" as (ps Hdom) "[Hauth Hconsistent]". + iDestruct (iProto_own_auth_agree with "Hauth Hi") as "#Hpi". + iDestruct (iProto_own_auth_agree with "Hauth Hj") as "#Hpj". + iDestruct (own_prot_idx with "Hi Hj") as %Hneq. + iAssert (â–· (<[i:=<(Send, j)> m1]>ps !!! j ≡ pj))%I as "Hpj'". + { by rewrite list_lookup_total_insert_ne. } + iAssert (â–· (⌜is_Some (ps !! i)⌠∗ (pi ⊑ (<(Send, j)> m1))))%I with "[Hile]" + as "[Hi' Hile]". + { iNext. iDestruct (iProto_le_msg_inv_r with "Hile") as (m) "#Heq". + iFrame. iRewrite "Heq" in "Hpi". rewrite list_lookup_total_alt. + destruct (ps !! i); [done|]. + iDestruct (iProto_end_message_equivI with "Hpi") as "[]". } + iAssert (â–· (⌜is_Some (ps !! j)⌠∗ (pj ⊑ (<(Recv, i)> m2))))%I with "[Hjle]" + as "[Hj' Hjle]". + { iNext. iDestruct (iProto_le_msg_inv_r with "Hjle") as (m) "#Heq". + iFrame. iRewrite "Heq" in "Hpj". rewrite !list_lookup_total_alt. + destruct (ps !! j); [done|]. + iDestruct (iProto_end_message_equivI with "Hpj") as "[]". } + iDestruct (iProto_consistent_le with "Hconsistent Hpi Hile") as "Hconsistent". + iDestruct (iProto_consistent_le with "Hconsistent Hpj' Hjle") as "Hconsistent". + iDestruct (iProto_consistent_step _ _ _ i j with "Hconsistent [] [] [Hm //]") as + (p2) "[Hm2 Hconsistent]". + { rewrite list_lookup_total_insert_ne; [|done]. + rewrite list_lookup_total_insert_ne; [|done]. + rewrite list_lookup_total_insert; [done|]. admit. } + { rewrite list_lookup_total_insert_ne; [|done]. + rewrite list_lookup_total_insert; [done|]. admit. } + iMod (iProto_own_auth_update _ _ _ _ p2 with "Hauth Hj") as "[Hauth Hj]". + iMod (iProto_own_auth_update _ _ _ _ p1 with "Hauth Hi") as "[Hauth Hi]". + iIntros "!>!>". iExists p2. iFrame "Hm2". + iDestruct "Hi'" as %Hi. iDestruct "Hj'" as %Hj. + iSplitL "Hconsistent Hauth". + { iExists (<[i:=p1]> (<[j:=p2]> ps)). + iSplit. + { admit. + (* rewrite !dom_insert_lookup_L; [done..|by rewrite lookup_insert_ne]. *)} + iFrame. rewrite list_insert_insert. + rewrite list_insert_commute; [|done]. rewrite list_insert_insert. + by rewrite list_insert_commute; [|done]. } + iSplitL "Hi"; iExists _; iFrame; iApply iProto_le_refl. + Admitted. + + Lemma iProto_target γ ps_dom i a j m : + iProto_ctx γ ps_dom -∗ + iProto_own γ i (<(a, j)> m) -∗ + â–· (⌜j < ps_domâŒ) ∗ iProto_ctx γ ps_dom ∗ iProto_own γ i (<(a, j)> m). + Proof. + iIntros "Hctx Hown". + rewrite /iProto_ctx /iProto_own. + iDestruct "Hctx" as (ps Hdom) "[Hauth Hps]". + iDestruct "Hown" as (p') "[Hle Hown]". + iDestruct (iProto_own_auth_agree with "Hauth Hown") as "#Hi". + iDestruct (iProto_le_msg_inv_r with "Hle") as (m') "#Heq". + iAssert (â–· (⌜is_Some (ps !! j)⌠∗ iProto_consistent ps))%I + with "[Hps]" as "[HSome Hps]". + { iNext. iRewrite "Heq" in "Hi". + iDestruct (iProto_consistent_target with "Hps Hi") as "#$". by iFrame. } + iSplitL "HSome". + { iNext. iDestruct "HSome" as %Heq. + iPureIntro. simplify_eq. admit. } + iSplitL "Hauth Hps"; iExists _; by iFrame. + Admitted. + + (* (** The instances below make it possible to use the tactics [iIntros], *) + (* [iExist], [iSplitL]/[iSplitR], [iFrame] and [iModIntro] on [iProto_le] goals. *) *) + Global Instance iProto_le_from_forall_l {A} i (m1 : A → iMsg Σ V) m2 name : + AsIdentName m1 name → + FromForall (iProto_message (Recv,i) (iMsg_exist m1) ⊑ (<(Recv,i)> m2)) + (λ x, (<(Recv, i)> m1 x) ⊑ (<(Recv, i)> m2))%I name | 10. + Proof. intros _. apply iProto_le_exist_elim_l. Qed. + Global Instance iProto_le_from_forall_r {A} i m1 (m2 : A → iMsg Σ V) name : + AsIdentName m2 name → + FromForall ((<(Send,i)> m1) ⊑ iProto_message (Send,i) (iMsg_exist m2)) + (λ x, (<(Send,i)> m1) ⊑ (<(Send,i)> m2 x))%I name | 11. + Proof. intros _. apply iProto_le_exist_elim_r. Qed. + + Global Instance iProto_le_from_wand_l i m v P p : + TCIf (TCEq P True%I) False TCTrue → + FromWand ((<(Recv,i)> MSG v {{ P }}; p) ⊑ (<(Recv,i)> m)) P ((<(Recv,i)> MSG v; p) ⊑ (<(Recv,i)> m)) | 10. + Proof. intros _. apply iProto_le_payload_elim_l. Qed. + Global Instance iProto_le_from_wand_r i m v P p : + FromWand ((<(Send,i)> m) ⊑ (<(Send,i)> MSG v {{ P }}; p)) P ((<(Send,i)> m) ⊑ (<(Send,i)> MSG v; p)) | 11. + Proof. apply iProto_le_payload_elim_r. Qed. + + Global Instance iProto_le_from_exist_l {A} i (m : A → iMsg Σ V) p : + FromExist ((<(Send,i) @ x> m x) ⊑ p) (λ a, (<(Send,i)> m a) ⊑ p)%I | 10. + Proof. + rewrite /FromExist. iDestruct 1 as (x) "H". + iApply (iProto_le_trans with "[] H"). iApply iProto_le_exist_intro_l. + Qed. + Global Instance iProto_le_from_exist_r {A} i (m : A → iMsg Σ V) p : + FromExist (p ⊑ <(Recv,i) @ x> m x) (λ a, p ⊑ (<(Recv,i)> m a))%I | 11. + Proof. + rewrite /FromExist. iDestruct 1 as (x) "H". + iApply (iProto_le_trans with "H"). iApply iProto_le_exist_intro_r. + Qed. + + Global Instance iProto_le_from_sep_l i m v P p : + FromSep ((<(Send,i)> MSG v {{ P }}; p) ⊑ (<(Send,i)> m)) P ((<(Send,i)> MSG v; p) ⊑ (<(Send,i)> m)) | 10. + Proof. + rewrite /FromSep. iIntros "[HP H]". + iApply (iProto_le_trans with "[HP] H"). by iApply iProto_le_payload_intro_l. + Qed. + Global Instance iProto_le_from_sep_r i m v P p : + FromSep ((<(Recv,i)> m) ⊑ (<(Recv,i)> MSG v {{ P }}; p)) P ((<(Recv,i)> m) ⊑ (<(Recv,i)> MSG v; p)) | 11. + Proof. + rewrite /FromSep. iIntros "[HP H]". + iApply (iProto_le_trans with "H"). by iApply iProto_le_payload_intro_r. + Qed. + + Global Instance iProto_le_frame_l i q m v R P Q p : + Frame q R P Q → + Frame q R ((<(Send,i)> MSG v {{ P }}; p) ⊑ (<(Send,i)> m)) + ((<(Send,i)> MSG v {{ Q }}; p) ⊑ (<(Send,i)> m)) | 10. + Proof. + rewrite /Frame /=. iIntros (HP) "[HR H]". + iApply (iProto_le_trans with "[HR] H"). iApply iProto_le_payload_elim_r. + iIntros "HQ". iApply iProto_le_payload_intro_l. iApply HP; iFrame. + Qed. + Global Instance iProto_le_frame_r i q m v R P Q p : + Frame q R P Q → + Frame q R ((<(Recv,i)> m) ⊑ (<(Recv,i)> MSG v {{ P }}; p)) + ((<(Recv,i)> m) ⊑ (<(Recv,i)> MSG v {{ Q }}; p)) | 11. + Proof. + rewrite /Frame /=. iIntros (HP) "[HR H]". + iApply (iProto_le_trans with "H"). iApply iProto_le_payload_elim_l. + iIntros "HQ". iApply iProto_le_payload_intro_r. iApply HP; iFrame. + Qed. + + Global Instance iProto_le_from_modal a v p1 p2 : + FromModal True (modality_instances.modality_laterN 1) (p1 ⊑ p2) + ((<a> MSG v; p1) ⊑ (<a> MSG v; p2)) (p1 ⊑ p2). + Proof. intros _. iApply iProto_le_base. Qed. + +End proto. + +Typeclasses Opaque iProto_ctx iProto_own. + +Global Hint Extern 0 (environments.envs_entails _ (?x ⊑ ?y)) => + first [is_evar x; fail 1 | is_evar y; fail 1|iApply iProto_le_refl] : core. -- GitLab From 20dcdc7175abba7a8e4fe0860810686fb03ce225 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Sat, 23 Mar 2024 18:20:27 +0100 Subject: [PATCH 54/81] Improved defs, closed admits, bumped proofmode to lists --- multi_actris/channel/channel.v | 160 ++++----- multi_actris/channel/proofmode.v | 59 ++-- multi_actris/channel/proto.v | 432 ++++++++++++++---------- multi_actris/examples/basics.v | 337 +++++------------- multi_actris/examples/leader_election.v | 94 ++---- 5 files changed, 449 insertions(+), 633 deletions(-) diff --git a/multi_actris/channel/channel.v b/multi_actris/channel/channel.v index 0a1248a..11c498b 100644 --- a/multi_actris/channel/channel.v +++ b/multi_actris/channel/channel.v @@ -93,23 +93,21 @@ Definition tok `{!chanG Σ} (γ : gname) : iProp Σ := own γ (Excl ()). Definition chan_inv `{!heapGS Σ, !chanG Σ} γ γE γt i j (l:loc) : iProp Σ := (l ↦ NONEV ∗ tok γt)%I ∨ - (∃ v m, l ↦ SOMEV v ∗ + (∃ v m p, l ↦ SOMEV v ∗ iProto_own γ i (<(Send, j)> m)%proto ∗ - (∃ p, iMsg_car m v (Next p) ∗ own γE (â—E (Next p)))) ∨ + iMsg_car m v (Next p) ∗ own γE (â—E (Next p))) ∨ (∃ p, l ↦ NONEV ∗ iProto_own γ i p ∗ own γE (â—E (Next p))). Definition iProto_pointsto_def `{!heapGS Σ, !chanG Σ} (c : val) (p : iProto Σ) : iProp Σ := - ∃ γ γE1 (l:loc) (i:nat) (n:nat) p', + ∃ γ (γEs : list gname) (l:loc) (i:nat) (n:nat) p', ⌜ c = (#l,#n,#i)%V ⌠∗ inv (nroot.@"ctx") (iProto_ctx γ n) ∗ - ([∗list] j ↦ _ ∈ replicate n (), - ∃ γE2 γt1 γt2, - inv (nroot.@"p") (chan_inv γ γE1 γt1 i j (l +â‚— (pos n i j))) ∗ - inv (nroot.@"p") (chan_inv γ γE2 γt2 j i (l +â‚— (pos n j i)))) ∗ + (∀ i j, ⌜i < n ⌠-∗ ⌜j < n⌠-∗ + ∃ γt, inv (nroot.@"p") (chan_inv γ (γEs !!! i) γt i j (l +â‚— (pos n i j)))) ∗ â–· (p' ⊑ p) ∗ - own γE1 (â—E (Next p')) ∗ own γE1 (â—¯E (Next p')) ∗ + own (γEs !!! i) (â—E (Next p')) ∗ own (γEs !!! i) (â—¯E (Next p')) ∗ iProto_own γ i p'. Definition iProto_pointsto_aux : seal (@iProto_pointsto_def). by eexists. Qed. Definition iProto_pointsto := iProto_pointsto_aux.(unseal). @@ -125,15 +123,13 @@ Definition chan_pool `{!heapGS Σ, !chanG Σ} ∃ γ (γEs : list gname) (n:nat) (l:loc), ⌜cs = (#l,#n)%V⌠∗ ⌜(i' + length ps = n)%nat⌠∗ inv (nroot.@"ctx") (iProto_ctx γ n) ∗ + (∀ i j, ⌜i < n ⌠-∗ ⌜j < n⌠-∗ + ∃ γt, inv (nroot.@"p") (chan_inv γ (γEs !!! i) γt i j (l +â‚— (pos n i j)))) ∗ [∗ list] i ↦ _ ∈ replicate n (), (∀ p, ⌜i' <= i⌠-∗ ⌜ps !! (i - i') = Some p⌠-∗ own (γEs !!! i) (â—E (Next p)) ∗ own (γEs !!! i) (â—¯E (Next p)) ∗ - iProto_own γ i p) ∗ - [∗ list] j ↦ _ ∈ replicate n (), - ∃ γt1 γt2, - inv (nroot.@"p") (chan_inv γ (γEs !!! i) γt1 i j (l +â‚— (pos n i j))) ∗ - inv (nroot.@"p") (chan_inv γ (γEs !!! j) γt2 j i (l +â‚— (pos n j i))). + iProto_own γ i p). Section channel. Context `{!heapGS Σ, !chanG Σ}. @@ -269,30 +265,17 @@ Section channel. iSplit. { done. } iFrame "IHp". + iSplit. + { iIntros (i j Hi Hj). + rewrite (big_sepL_lookup _ _ i); [|by rewrite lookup_replicate]. + rewrite (big_sepL_lookup _ _ j); [|by rewrite lookup_replicate]. + iApply "IH". } iApply (big_sepL_impl with "H"). iIntros "!>" (i ? HSome') "(Hauth & Hfrag & Hown)". - iSplitL. - { iIntros (p Hle' HSome''). - iFrame. rewrite right_id_L in HSome''. - rewrite (list_lookup_total_alt ps). - rewrite HSome''. simpl. iFrame. } - iApply big_sepL_intro. - iIntros "!>" (j ? HSome''). - assert (i < n) as Hle'. - { apply lookup_replicate in HSome' as [_ Hle']. done. } - assert (j < n) as Hle''. - { apply lookup_replicate in HSome'' as [_ Hle'']. done. } - iDestruct (big_sepL_lookup _ _ i () with "IH") as "IH'"; - [by rewrite lookup_replicate|]. - iDestruct (big_sepL_lookup _ _ j () with "IH'") as "IH''"; - [by rewrite lookup_replicate|]. - iDestruct (big_sepL_lookup _ _ j () with "IH") as "H"; - [by rewrite lookup_replicate|]. - iDestruct (big_sepL_lookup _ _ i () with "H") as "H'"; - [by rewrite lookup_replicate|]. - iDestruct "IH''" as (γ1) "?". - iDestruct "H'" as (γ2) "?". - iExists _, _. iFrame "#". + iIntros (p Hle' HSome''). + iFrame. rewrite right_id_L in HSome''. + rewrite (list_lookup_total_alt ps). + rewrite HSome''. simpl. iFrame. Qed. Lemma get_chan_spec cs i ps p : @@ -301,11 +284,10 @@ Section channel. {{{ c, RET c; c ↣ p ∗ chan_pool cs (i+1) ps }}}. Proof. iIntros (Φ) "Hcs HΦ". - iDestruct "Hcs" as (γp γEs n l -> <-) "[#IHp Hl]". - wp_lam. wp_pures. - simpl. + iDestruct "Hcs" as (γp γEs n l -> <-) "(#IHp & #IHl & Hl)". + wp_lam. wp_pures. rewrite replicate_add. simpl. - iDestruct "Hl" as "[Hl1 [[Hi #IHs] Hl3]]". simpl. + iDestruct "Hl" as "[Hl1 [Hi Hl3]]". iDestruct ("Hi" with "[] []") as "(Hauth & Hown & Hp)". { rewrite right_id_L. rewrite replicate_length. iPureIntro. lia. } { rewrite right_id_L. rewrite replicate_length. @@ -316,16 +298,11 @@ Section channel. { rewrite iProto_pointsto_eq. iExists _, _, _, _, _, _. iSplit; [done|]. rewrite replicate_length. rewrite right_id_L. - iFrame "#∗". - iSplit; [|iNext; done]. - rewrite replicate_add. - iApply (big_sepL_impl with "IHs"). - iIntros "!>" (???). iDestruct 1 as (γt1 γt2) "[??]". - iExists _,_,_. iFrame. } - iExists _, _, _, _. iSplit; [done|]. + iFrame. iFrame "#∗". iNext. done. } + iExists γp, γEs, _, _. iSplit; [done|]. iSplit. { iPureIntro. lia. } - iFrame "#∗". + iFrame. iFrame "#∗". rewrite replicate_add. simpl. iSplitL "Hl1". @@ -333,7 +310,7 @@ Section channel. iIntros "!>" (i' ? HSome''). assert (i' < i). { rewrite lookup_replicate in HSome''. lia. } - iIntros "[H $]" (p' Hle'). lia. } + iIntros "H" (p' Hle'). lia. } simpl. iFrame "#∗". iSplitR. @@ -343,7 +320,7 @@ Section channel. iIntros "!>" (i' ? HSome''). assert (i' < length ps). { rewrite lookup_replicate in HSome''. lia. } - iIntros "[H $]" (p' Hle' HSome). + iIntros "H" (p' Hle' HSome). iApply "H". { iPureIntro. lia. } iPureIntro. @@ -373,44 +350,42 @@ Section channel. iDestruct "Hc" as (γ γE l i n p' ->) "(#IH & #Hls & Hle & Hâ— & Hâ—¯ & Hown)". wp_bind (Fst _). - iInv "IH" as "HI". + iInv "IH" as "Hctx". iDestruct (iProto_le_msg_inv_r with "Hle") as (m') "#Heq". iRewrite "Heq" in "Hown". + iAssert (â–· (⌜i < n⌠∗ iProto_own γ i (<(Send, j)> m') ∗ + iProto_ctx γ n))%I with "[Hctx Hown]" + as "[Hi [Hown Hctx]]". + { iNext. iDestruct (iProto_ctx_agree with "Hctx Hown") as %Hi. + iFrame. done. } iAssert (â–· (â–· ⌜j < n⌠∗ iProto_own γ i (<(Send, j)> m') ∗ - iProto_ctx γ n))%I with "[HI Hown]" - as "[HI [Hown Hctx]]". - { iNext. iDestruct (iProto_target with "HI Hown") as "[Hin [$ $]]". + iProto_ctx γ n))%I with "[Hctx Hown]" + as "[Hj [Hown Hctx]]". + { iNext. iDestruct (iProto_target with "Hctx Hown") as "[Hin [$ $]]". iFrame. } iRewrite -"Heq" in "Hown". wp_pures. iModIntro. iFrame. wp_smart_apply (vpos_spec); [done|]; iIntros "_". - iDestruct "HI" as %Hle. - iDestruct (big_sepL_lookup_acc with "Hls") as "[Hj _]"; - [by apply lookup_replicate_2|]. - iDestruct "Hj" as (γE' γt γt') "#[IHl1 IHl2]". + iDestruct "Hi" as %Hi. + iDestruct "Hj" as %Hj. + iDestruct ("Hls" $! i j with "[//] [//]") as (γt) "IHl1". wp_pures. wp_bind (Store _ _). iInv "IHl1" as "HIp". iDestruct "HIp" as "[HIp|HIp]"; last first. { iDestruct "HIp" as "[HIp|HIp]". - - iDestruct "HIp" as (? m) "(>Hl & Hown' & HIp)". + - iDestruct "HIp" as (? m p'') "(>Hl & Hown' & HIp)". wp_store. - rewrite /iProto_own. - iDestruct "Hown" as (p'') "[Hle' Hown]". - iDestruct "Hown'" as (p''') "[Hle'' Hown']". - iDestruct (own_prot_excl with "Hown Hown'") as "H". done. + iDestruct (iProto_own_excl with "Hown Hown'") as %[]. - iDestruct "HIp" as (p'') "(>Hl' & Hown' & HIp)". wp_store. - rewrite /iProto_own. - iDestruct "Hown" as (p''') "[Hle' Hown]". - iDestruct "Hown'" as (p'''') "[Hle'' Hown']". - iDestruct (own_prot_excl with "Hown Hown'") as "H". done. } + iDestruct (iProto_own_excl with "Hown Hown'") as %[]. } iDestruct "HIp" as "[>Hl' Htok]". wp_store. iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]"; [by apply excl_auth_update|]. iModIntro. iSplitL "Hl' Hâ— Hown Hle". - { iRight. iLeft. iIntros "!>". iExists _, _. iFrame. + { iRight. iLeft. iIntros "!>". iExists _, _, _. iFrame. iSplitL "Hown Hle"; [by iApply (iProto_own_le with "Hown Hle")|]. - iExists _. iFrame. by rewrite iMsg_base_eq. } + by rewrite iMsg_base_eq. } wp_pures. iLöb as "HL". wp_lam. @@ -420,10 +395,10 @@ Section channel. { iDestruct "HIp" as ">[Hl' Htok']". iDestruct (own_valid_2 with "Htok Htok'") as %[]. } iDestruct "HIp" as "[HIp|HIp]". - - iDestruct "HIp" as (? m) "(>Hl' & Hown & HIp)". + - iDestruct "HIp" as (? m p'') "(>Hl' & Hown & HIp)". wp_load. iModIntro. iSplitL "Hl' Hown HIp". - { iRight. iLeft. iExists _, _. iFrame. } + { iRight. iLeft. iExists _, _,_. iFrame. } wp_pures. iApply ("HL" with "HΦ Htok Hâ—¯"). - iDestruct "HIp" as (p'') "(>Hl' & Hown & Hâ—)". wp_load. @@ -466,19 +441,23 @@ Section channel. iDestruct "Hc" as (γ γE l i n p' ->) "(#IH & #Hls & Hle & Hâ— & Hâ—¯ & Hown)". do 6 wp_pure _. wp_bind (Fst _). wp_pure _. - iInv "IH" as "HI". + iInv "IH" as "Hctx". iDestruct (iProto_le_msg_inv_r with "Hle") as (m') "#Heq". iRewrite "Heq" in "Hown". + iAssert (â–· (⌜i < n⌠∗ iProto_own γ i (<(Recv, j)> m') ∗ + iProto_ctx γ n))%I with "[Hctx Hown]" + as "[Hi [Hown Hctx]]". + { iNext. iDestruct (iProto_ctx_agree with "Hctx Hown") as %Hi. + iFrame. done. } iAssert (â–· (â–· ⌜j < n⌠∗ iProto_own γ i (<(Recv, j)> m') ∗ - iProto_ctx γ n))%I with "[HI Hown]" as "[HI [Hown Hctx]]". - { iNext. iDestruct (iProto_target with "HI Hown") as "[Hin [$$]]". + iProto_ctx γ n))%I with "[Hctx Hown]" as "[Hj [Hown Hctx]]". + { iNext. iDestruct (iProto_target with "Hctx Hown") as "[Hin [$$]]". iFrame. } iRewrite -"Heq" in "Hown". wp_pures. iModIntro. iFrame. wp_smart_apply (vpos_spec); [done|]; iIntros "_". - iDestruct "HI" as %Hle. - iDestruct (big_sepL_lookup_acc with "Hls") as "[Hj _]"; - [by apply lookup_replicate_2|]. - iDestruct "Hj" as (γE' γt γt') "#[IHl1 IHl2]". + iDestruct "Hi" as %Hi. + iDestruct "Hj" as %Hj. + iDestruct ("Hls" $! j i with "[//] [//]") as (γt) "IHl2". wp_pures. wp_bind (Xchg _ _). iInv "IHl2" as "HIp". @@ -497,8 +476,7 @@ Section channel. { iRight. iRight. iExists _. iFrame. } wp_pures. iApply ("HL" with "[Hâ— Hâ—¯ Hown Hle] HΦ"). iExists _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". } - iDestruct "HIp" as (w m) "(>Hl' & Hown' & HIp)". - iDestruct "HIp" as (p'') "[Hm Hp']". + iDestruct "HIp" as (w m p'') "(>Hl' & Hown' & Hm & Hp')". iInv "IH" as "Hctx". wp_xchg. iDestruct (iProto_own_le with "Hown Hle") as "Hown". @@ -519,28 +497,4 @@ Section channel. iRewrite "Hp". iFrame "#∗". iApply iProto_le_refl. Qed. - Lemma iProto_le_select_l {TT1 TT2:tele} j - (v1 : TT1 → val) (v2 : TT2 → val) P1 P2 (p1 : TT1 → iProto Σ) (p2 : TT2 → iProto Σ) : - ⊢ (iProto_choice (Send, j) v1 v2 P1 P2 p1 p2) ⊑ - (<(Send,j) @.. (tt:TT1)> MSG (InjLV (v1 tt)) {{ P1 tt }} ; p1 tt). - Proof. - rewrite /iProto_choice. - iApply iProto_le_trans; last first. - { iApply iProto_le_texist_elim_r. iIntros (x). iExists x. - iApply iProto_le_refl. } - iIntros (tt). by iExists (inl tt). - Qed. - - Lemma iProto_le_select_r {TT1 TT2:tele} j - (v1 : TT1 → val) (v2 : TT2 → val) P1 P2 (p1 : TT1 → iProto Σ) (p2 : TT2 → iProto Σ) : - ⊢ (iProto_choice (Send, j) v1 v2 P1 P2 p1 p2) ⊑ - (<(Send,j) @.. (tt:TT2)> MSG (InjRV (v2 tt)) {{ P2 tt }} ; p2 tt). - Proof. - rewrite /iProto_choice. - iApply iProto_le_trans; last first. - { iApply iProto_le_texist_elim_r. iIntros (x). iExists x. - iApply iProto_le_refl. } - iIntros (tt). by iExists (inr tt). - Qed. - End channel. diff --git a/multi_actris/channel/proofmode.v b/multi_actris/channel/proofmode.v index 07c0449..a125a6f 100644 --- a/multi_actris/channel/proofmode.v +++ b/multi_actris/channel/proofmode.v @@ -344,7 +344,7 @@ Tactic Notation "wp_send" "(" uconstr(x1) uconstr(x2) uconstr(x3) uconstr(x4) ") wp_send_core (eexists x1; eexists x2; eexists x3; eexists x4; eexists x5; eexists x6; eexists x7; eexists x8) with pat. -Lemma iProto_consistent_equiv_proof {Σ} (ps : gmap nat (iProto Σ)) : +Lemma iProto_consistent_equiv_proof {Σ} (ps : list (iProto Σ)) : (∀ i j, valid_target ps i j) ∗ (∀ i j m1 m2, (ps !!! i ≡ (<(Send, j)> m1)%proto) -∗ @@ -396,56 +396,32 @@ Tactic Notation "iProto_consistent_take_step_step" := let m1 := fresh in let m2 := fresh in iIntros (i j m1 m2) "#Hm1 #Hm2"; - repeat (destruct i as [|i]; - [repeat (rewrite lookup_total_insert_ne; [|lia]); - try (by rewrite lookup_total_empty iProto_end_message_equivI); - try (rewrite lookup_total_insert; - try (by rewrite iProto_end_message_equivI); - iDestruct (iProto_message_equivI with "Hm1") - as "[%Heq1 Hm1']";simplify_eq)| - repeat (rewrite lookup_total_insert_ne; [|lia]); - try (by rewrite lookup_total_empty iProto_end_message_equivI)]); - repeat (rewrite lookup_total_insert_ne; [|lia]); - try rewrite lookup_total_empty; - try (by iProto_end_message_equivI); - rewrite lookup_total_insert; + repeat (destruct i as [|i]=> /=; + [try (rewrite lookup_total_nil); try (by rewrite iProto_end_message_equivI); + iDestruct (iProto_message_equivI with "Hm1") + as "[%Heq1 Hm1']";simplify_eq=> /=| + try (rewrite lookup_total_nil); try (by rewrite iProto_end_message_equivI)]); + try (rewrite lookup_total_nil); + try (by rewrite iProto_end_message_equivI); iDestruct (iProto_message_equivI with "Hm2") - as "[%Heq2 Hm2']";simplify_eq; + as "[%Heq2 Hm2']";simplify_eq=> /=; try (iClear "Hm1' Hm2'"; iExists _,_,_,_,_,_,_,_,_,_; iSplitL "Hm1"; [iFrame "#"|]; iSplitL "Hm2"; [iFrame "#"|]; iSplit; [iPureIntro; tc_solve|]; iSplit; [iPureIntro; tc_solve|]; - simpl; iClear "Hm1 Hm2"; clear m1 m2); - try (repeat (rewrite (insert_commute _ _ i); [|done]); - rewrite insert_insert; - repeat (rewrite (insert_commute _ _ j); [|done]); - rewrite insert_insert). + simpl; iClear "Hm1 Hm2"; clear m1 m2). Tactic Notation "iProto_consistent_take_step_target" := let i := fresh in iIntros (i j a m); rewrite /valid_target; - iIntros "#Hm"; - repeat (destruct i as [|i]; - [repeat (rewrite lookup_total_insert_ne; [|lia]); - try (by rewrite lookup_total_empty iProto_end_message_equivI); - try (rewrite lookup_total_insert; - try (by rewrite iProto_end_message_equivI); - iDestruct (iProto_message_equivI with "Hm1") - as "[%Heq1 Hm1']";simplify_eq)| - repeat (rewrite lookup_total_insert_ne; [|lia]); - try (by rewrite lookup_total_empty iProto_end_message_equivI)]); - repeat (rewrite lookup_total_insert_ne; [|lia]); - try rewrite lookup_total_empty; - try (by iProto_end_message_equivI); - rewrite lookup_total_insert; - iDestruct (iProto_message_equivI with "Hm") - as "[%Heq Hm']";simplify_eq; - repeat (try rewrite lookup_empty; - try rewrite lookup_insert; - rewrite lookup_insert_ne; [|lia]); - try rewrite lookup_insert; try done. + iIntros "#Hm1"; + repeat (destruct i as [|i]=> /=; + [try (rewrite lookup_total_nil); try (by rewrite iProto_end_message_equivI); + by iDestruct (iProto_message_equivI with "Hm1") + as "[%Heq1 Hm1']" ; simplify_eq=> /=| + try (rewrite lookup_total_nil); try (by rewrite iProto_end_message_equivI)]). Tactic Notation "iProto_consistent_take_step" := try iNext; @@ -463,3 +439,6 @@ Tactic Notation "iProto_consistent_resolve_step" := Tactic Notation "iProto_consistent_take_steps" := repeat (iProto_consistent_take_step; iProto_consistent_resolve_step). + +Tactic Notation "wp_get_chan" "(" simple_intropattern(c) ")" constr(pat) := + wp_smart_apply (get_chan_spec with "[$]"); iIntros (c); iIntros pat. diff --git a/multi_actris/channel/proto.v b/multi_actris/channel/proto.v index e5a2b78..aa0826a 100644 --- a/multi_actris/channel/proto.v +++ b/multi_actris/channel/proto.v @@ -499,14 +499,14 @@ Global Instance iProto_inhabited {Σ V} : Inhabited (iProto Σ V) := populate EN Definition can_step {Σ V} (rec : list (iProto Σ V) → iProp Σ) (ps : list (iProto Σ V)) (i j : nat) : iProp Σ := ∀ m1 m2, - (ps !!! i ≡ <(Send, j)> m1) -∗ - (ps !!! j ≡ <(Recv, i)> m2) -∗ + (ps !!! i ≡ (<(Send, j)> m1)) -∗ + (ps !!! j ≡ (<(Recv, i)> m2)) -∗ ∀ v p1, iMsg_car m1 v (Next p1) -∗ ∃ p2, iMsg_car m2 v (Next p2) ∗ â–· (rec (<[i:=p1]>(<[j:=p2]>ps))). Definition valid_target {Σ V} (ps : list (iProto Σ V)) (i j : nat) : iProp Σ := - ∀ a m, (ps !!! i ≡ <(a, j)> m) -∗ ⌜is_Some (ps !! j)âŒ. + ∀ a m, (ps !!! i ≡ (<(a, j)> m)) -∗ ⌜is_Some (ps !! j)âŒ. Definition iProto_consistent_pre {Σ V} (rec : list (iProto Σ V) → iProp Σ) (ps : list (iProto Σ V)) : iProp Σ := @@ -600,8 +600,7 @@ Definition iProto_own_frag `{!protoG Σ V} (γ : gname) Definition iProto_own_auth `{!protoG Σ V} (γ : gname) (ps : list (iProto Σ V)) : iProp Σ := - own γ (gmap_view_auth (DfracOwn 1) (((λ p, Excl' (Next p)) <$> - (list_to_map (zip (seq 0 (length ps)) ps))) : gmap _ _)). + own γ (gmap_view_auth (DfracOwn 1) ((λ p, Excl' (Next p)) <$> map_seq 0 ps)). Definition iProto_ctx `{protoG Σ V} (γ : gname) (ps_len : nat) : iProp Σ := @@ -781,44 +780,191 @@ Section proto. iDestruct "Hp2" as (Heq) "Hp2". simplify_eq. iDestruct "H" as (->) "H". by iExists _. Qed. + + Lemma iProto_le_send i m1 m2 : + (∀ v p2', iMsg_car m2 v (Next p2') -∗ ∃ p1', + â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1')) -∗ + (<(Send,i)> m1) ⊑ (<(Send,i)> m2). + Proof. + iIntros "Hle". rewrite iProto_le_unfold. + iRight. iExists (Send, i), (Send, i), m1, m2. by eauto. + Qed. + + Lemma iProto_le_recv i m1 m2 : + (∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', + â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2')) -∗ + (<(Recv,i)> m1) ⊑ (<(Recv,i)> m2). + Proof. + iIntros "Hle". rewrite iProto_le_unfold. + iRight. iExists (Recv, i), (Recv, i), m1, m2. by eauto. + Qed. + + Lemma iProto_le_base a v P p1 p2 : + â–· (p1 ⊑ p2) -∗ + (<a> MSG v {{ P }}; p1) ⊑ (<a> MSG v {{ P }}; p2). + Proof. + rewrite iMsg_base_eq. iIntros "H". destruct a as [[]]. + - iApply iProto_le_send. iIntros (v' p') "(->&Hp&$)". + iExists p1. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". + - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&$)". + iExists p2. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". + Qed. + + Lemma iProto_le_trans p1 p2 p3 : p1 ⊑ p2 -∗ p2 ⊑ p3 -∗ p1 ⊑ p3. + Proof. + iIntros "H1 H2". iLöb as "IH" forall (p1 p2 p3). + destruct (iProto_case p3) as [->|([]&i&m3&->)]. + - iDestruct (iProto_le_end_inv_r with "H2") as "H2". by iRewrite "H2" in "H1". + - iDestruct (iProto_le_send_inv with "H2") as (m2) "[Hp2 H2]". + iRewrite "Hp2" in "H1"; clear p2. + iDestruct (iProto_le_send_inv with "H1") as (m1) "[Hp1 H1]". + iRewrite "Hp1"; clear p1. + iApply iProto_le_send. iIntros (v p3') "Hm3". + iDestruct ("H2" with "Hm3") as (p2') "[Hle Hm2]". + iDestruct ("H1" with "Hm2") as (p1') "[Hle' Hm1]". + iExists p1'. iIntros "{$Hm1} !>". by iApply ("IH" with "Hle'"). + - iDestruct (iProto_le_recv_inv_r with "H2") as (m2) "[Hp2 H3]". + iRewrite "Hp2" in "H1". + iDestruct (iProto_le_recv_inv_r with "H1") as (m1) "[Hp1 H2]". + iRewrite "Hp1". iApply iProto_le_recv. iIntros (v p1') "Hm1". + iDestruct ("H2" with "Hm1") as (p2') "[Hle Hm2]". + iDestruct ("H3" with "Hm2") as (p3') "[Hle' Hm3]". + iExists p3'. iIntros "{$Hm3} !>". by iApply ("IH" with "Hle"). + Qed. + + Lemma iProto_le_refl p : ⊢ p ⊑ p. + Proof. + iLöb as "IH" forall (p). destruct (iProto_case p) as [->|([]&i&m&->)]. + - iApply iProto_le_end. + - iApply iProto_le_send. auto 10 with iFrame. + - iApply iProto_le_recv. auto 10 with iFrame. + Qed. + + + Global Instance iProto_own_frag_ne γ s : NonExpansive (iProto_own_frag γ s). + Proof. solve_proper. Qed. + + Lemma iProto_own_auth_agree γ ps i p : + iProto_own_auth γ ps -∗ iProto_own_frag γ i p -∗ â–· (ps !! i ≡ Some p). + Proof. + iIntros "Hâ— Hâ—¯". + iDestruct (own_valid_2 with "Hâ— Hâ—¯") as "H✓". + rewrite gmap_view_both_validI. + iDestruct "H✓" as "[_ [H1 H2]]". + rewrite lookup_fmap. + simpl. + rewrite lookup_map_seq_0. + destruct (ps !! i) eqn:Heqn; last first. + { rewrite Heqn. rewrite !option_equivI. done. } + rewrite Heqn. + simpl. rewrite !option_equivI excl_equivI. by iNext. + Qed. + + Lemma iProto_own_auth_agree_Some γ ps i p : + iProto_own_auth γ ps -∗ iProto_own_frag γ i p -∗ ⌜is_Some (ps !! i)âŒ. + Proof. + iIntros "Hâ— Hâ—¯". + iDestruct (own_valid_2 with "Hâ— Hâ—¯") as "H✓". + rewrite gmap_view_both_validI. + iDestruct "H✓" as "[_ [H1 H2]]". + rewrite lookup_fmap. + simpl. + rewrite lookup_map_seq_0. + destruct (ps !! i) eqn:Heqn; last first. + { rewrite Heqn. rewrite !option_equivI. done. } + rewrite Heqn. + simpl. rewrite !option_equivI excl_equivI. done. + Qed. + + Lemma iProto_own_auth_update γ ps i p p' : + iProto_own_auth γ ps -∗ iProto_own_frag γ i p ==∗ + iProto_own_auth γ (<[i := p']>ps) ∗ iProto_own_frag γ i p'. + Proof. + iIntros "Hâ— Hâ—¯". + iDestruct (iProto_own_auth_agree_Some with "Hâ— Hâ—¯") as %HSome. + iMod (own_update_2 with "Hâ— Hâ—¯") as "[H1 H2]"; [|iModIntro]. + { eapply (gmap_view_replace _ _ _ (Excl' (Next p'))). done. } + iFrame. rewrite -fmap_insert. + rewrite /iProto_own_auth. + rewrite insert_map_seq_0; [done|]. + by apply lookup_lt_is_Some_1. + Qed. + + Lemma iProto_own_auth_alloc ps : + ⊢ |==> ∃ γ, iProto_own_auth γ ps ∗ [∗ list] i ↦p ∈ ps, iProto_own γ i p. + Proof. + iMod (own_alloc (gmap_view_auth (DfracOwn 1) ∅)) as (γ) "Hauth". + { apply gmap_view_auth_valid. } + iExists γ. + iInduction ps as [|p ps] "IH" using rev_ind. + { iModIntro. iFrame. done. } + iMod ("IH" with "Hauth") as "[Hauth Hfrags]". + iFrame "Hfrags". + iMod (own_update with "Hauth") as "[Hauth Hfrag]". + { apply (gmap_view_alloc _ (length ps) (DfracOwn 1) (Excl' (Next p))); [|done|done]. + rewrite fmap_map_seq. + rewrite lookup_map_seq_0. + apply lookup_ge_None_2. rewrite fmap_length. done. } + simpl. + iModIntro. + rewrite right_id_L. + rewrite -fmap_insert. iFrame. + iSplitL "Hauth". + - rewrite /iProto_own_auth. + rewrite map_seq_snoc. simpl. done. + - iSplit; [|done]. + iExists _. iFrame. by iApply iProto_le_refl. + Qed. + + Lemma list_lookup_Some_le (ps : list $ iProto Σ V) (i : nat) (p1 : iProto Σ V) : + ⊢@{iProp Σ} ps !! i ≡ Some p1 -∗ ⌜i < length psâŒ. + Proof. + iIntros "HSome". + rewrite option_equivI. + destruct (ps !! i) eqn:Heqn; [|done]. + iPureIntro. + by apply lookup_lt_is_Some_1. + Qed. Lemma valid_target_le ps i p1 p2 : (∀ i' j', valid_target ps i' j') -∗ - ps !!! i ≡ p1 -∗ + ps !! i ≡ Some p1 -∗ p1 ⊑ p2 -∗ (∀ i' j', valid_target (<[i := p2]>ps) i' j') ∗ p1 ⊑ p2. - Proof. Admitted. - (* iIntros "Hprot #HSome Hle". *) - (* pose proof (iProto_case p1) as [Hend|Hmsg]. *) - (* { rewrite Hend. iDestruct (iProto_le_end_inv_l with "Hle") as "#H". *) - (* iFrame "Hle". *) - (* iIntros (i' j' a m) "Hm". *) - (* destruct (decide (i = j')) as [->|Hneqj]. *) - (* { Search list_lookup_total insert. rewrite list_lookup_total_insert. ; [done|]. lia. done. } *) - (* rewrite lookup_insert_ne; [|done]. *) - (* destruct (decide (i = i')) as [->|Hneqi]. *) - (* { rewrite lookup_total_insert. iRewrite "H" in "Hm". *) - (* by iDestruct (iProto_end_message_equivI with "Hm") as "Hm". } *) - (* rewrite lookup_total_insert_ne; [|done]. *) - (* by iApply "Hprot". } *) - (* destruct Hmsg as (t & n & m & Hmsg). *) - (* setoid_rewrite Hmsg. *) - (* iDestruct (iProto_le_msg_inv_l with "Hle") as (m2) "#Heq". iFrame "Hle". *) - (* iIntros (i' j' a m') "Hm". *) - (* destruct (decide (i = j')) as [->|Hneqj]. *) - (* { rewrite lookup_insert. done. } *) - (* rewrite lookup_insert_ne; [|done]. *) - (* destruct (decide (i = i')) as [->|Hneqi]. *) - (* { rewrite lookup_total_insert. iRewrite "Heq" in "Hm". *) - (* iDestruct (iProto_message_equivI with "Hm") as (Heq) "Hm". *) - (* simplify_eq. by iApply "Hprot". } *) - (* rewrite lookup_total_insert_ne; [|done]. *) - (* by iApply "Hprot". *) - (* Qed. *) + Proof. + iIntros "Hprot #HSome Hle". + iDestruct (list_lookup_Some_le with "HSome") as %Hi. + pose proof (iProto_case p1) as [Hend|Hmsg]. + { rewrite Hend. iDestruct (iProto_le_end_inv_l with "Hle") as "#H". + iFrame "Hle". + iIntros (i' j' a m) "Hm". + destruct (decide (i = j')) as [->|Hneqj]. + { rewrite list_lookup_insert; [done|]. done. } + rewrite (list_lookup_insert_ne _ i j'); [|done]. + destruct (decide (i = i')) as [->|Hneqi]. + { rewrite list_lookup_total_insert; [|done]. iRewrite "H" in "Hm". + by iDestruct (iProto_end_message_equivI with "Hm") as "Hm". } + rewrite list_lookup_total_insert_ne; [|done]. + by iApply "Hprot". } + destruct Hmsg as (t & n & m & Hmsg). + setoid_rewrite Hmsg. + iDestruct (iProto_le_msg_inv_l with "Hle") as (m2) "#Heq". iFrame "Hle". + iIntros (i' j' a m') "Hm". + destruct (decide (i = j')) as [->|Hneqj]. + { by rewrite list_lookup_insert. } + rewrite (list_lookup_insert_ne _ i j'); [|done]. + destruct (decide (i = i')) as [->|Hneqi]. + { rewrite list_lookup_total_insert; [|done]. iRewrite "Heq" in "Hm". + iDestruct (iProto_message_equivI with "Hm") as (Heq) "Hm". + simplify_eq. iApply ("Hprot" $! i'). + rewrite list_lookup_total_alt. iRewrite "HSome". done. } + rewrite list_lookup_total_insert_ne; [|done]. + by iApply "Hprot". + Qed. Lemma iProto_consistent_le ps i p1 p2 : iProto_consistent ps -∗ - ps !!! i ≡ p1 -∗ + ps !! i ≡ Some p1 -∗ p1 ⊑ p2 -∗ iProto_consistent (<[i := p2]>ps). Proof. @@ -826,19 +972,21 @@ Section proto. iRevert "HSome". iLöb as "IH" forall (p1 p2 ps). iIntros "#HSome". + iDestruct (list_lookup_Some_le with "HSome") as %Hi. rewrite !iProto_consistent_unfold. iDestruct "Hprot" as "(Htar & Hprot)". iDestruct (valid_target_le with "Htar HSome Hle") as "[Htar Hle]". iFrame. iIntros (i' j' m1 m2) "#Hm1 #Hm2". destruct (decide (i = i')) as [<-|Hneq]. - { rewrite list_lookup_total_insert; [|admit]. + { rewrite list_lookup_total_insert; [|done]. pose proof (iProto_case p2) as [Hend|Hmsg]. - { setoid_rewrite Hend. rewrite iProto_end_message_equivI. done. } + { setoid_rewrite Hend. + rewrite !option_equivI. rewrite iProto_end_message_equivI. done. } destruct Hmsg as (a&?&m&Hmsg). setoid_rewrite Hmsg. destruct a; last first. - { rewrite iProto_message_equivI. + { rewrite !option_equivI. rewrite iProto_message_equivI. iDestruct "Hm1" as "[%Htag Hm1]". done. } rewrite iProto_message_equivI. iDestruct "Hm1" as "[%Htag Hm1]". @@ -850,29 +998,35 @@ Section proto. iDestruct "Hle" as (m') "[#Heq H]". iDestruct ("H" with "Hm1'") as (p') "[Hle H]". destruct (decide (i = j')) as [<-|Hneq]. - { rewrite list_lookup_total_insert. rewrite iProto_message_equivI. - iDestruct "Hm2" as "[%Heq _]". done. admit. } + { rewrite list_lookup_total_insert; [|done]. + rewrite iProto_message_equivI. + iDestruct "Hm2" as "[%Heq _]". done. } iDestruct ("Hprot" $!i j' with "[] [] H") as "Hprot". - { iRewrite -"Heq". rewrite !list_lookup_total_alt. iRewrite "HSome". done. } + { iRewrite -"Heq". iEval (rewrite list_lookup_total_alt). + iRewrite "HSome". done. } { rewrite list_lookup_total_insert_ne; [|done]. done. } iDestruct "Hprot" as (p'') "[Hm Hprot]". iExists p''. iFrame. iNext. iDestruct ("IH" with "Hprot Hle [HSome]") as "HI". - { rewrite list_lookup_total_insert; [done|]. admit. } + { rewrite list_lookup_insert; [done|]. + by rewrite insert_length. } iClear "IH Hm1 Hm2 Heq". rewrite list_insert_insert. rewrite (list_insert_commute _ j' i); [|done]. rewrite list_insert_insert. done. } rewrite list_lookup_total_insert_ne; [|done]. destruct (decide (i = j')) as [<-|Hneq']. - { rewrite list_lookup_total_insert. + { rewrite list_lookup_total_insert; [|done]. pose proof (iProto_case p2) as [Hend|Hmsg]. - { setoid_rewrite Hend. rewrite iProto_end_message_equivI. done. } + { setoid_rewrite Hend. + rewrite !option_equivI. + rewrite iProto_end_message_equivI. done. } destruct Hmsg as (a&?&m&Hmsg). setoid_rewrite Hmsg. destruct a. - { rewrite iProto_message_equivI. + { rewrite !option_equivI. + rewrite iProto_message_equivI. iDestruct "Hm2" as "[%Htag Hm2]". done. } rewrite iProto_message_equivI. iDestruct "Hm2" as "[%Htag Hm2]". @@ -882,17 +1036,17 @@ Section proto. iDestruct "Hle" as (m') "[#Heq Hle]". iDestruct ("Hprot" $!i' with "[] [] Hm1'") as "Hprot". { done. } - { rewrite !list_lookup_total_alt. iRewrite "HSome". done. } + { iEval (rewrite list_lookup_total_alt). iRewrite "HSome". done. } iDestruct ("Hprot") as (p') "[Hm1' Hprot]". iDestruct ("Hle" with "Hm1'") as (p2') "[Hle Hm']". iSpecialize ("Hm2" $! v (Next p2')). iExists p2'. iRewrite -"Hm2". iFrame. iDestruct ("IH" with "Hprot Hle []") as "HI". - { iPureIntro. rewrite list_lookup_total_insert_ne; [|done]. - rewrite list_lookup_total_insert. done. admit. } + { iPureIntro. rewrite list_lookup_insert_ne; [|done]. + by rewrite list_lookup_insert. } rewrite list_insert_commute; [|done]. - rewrite !list_insert_insert. done. admit. } + rewrite !list_insert_insert. done. } rewrite list_lookup_total_insert_ne; [|done]. iIntros (v p) "Hm1'". iDestruct ("Hprot" $!i' j' with "[//] [//] Hm1'") as "Hprot". @@ -902,68 +1056,9 @@ Section proto. rewrite (list_insert_commute _ j' i); [|done]. rewrite (list_insert_commute _ i' i); [|done]. iApply ("IH" with "Hprot Hle []"). - rewrite list_lookup_total_insert_ne; [|done]. - rewrite list_lookup_total_insert_ne; [|done]. + rewrite list_lookup_insert_ne; [|done]. + rewrite list_lookup_insert_ne; [|done]. done. - Admitted. - - Lemma iProto_le_send i m1 m2 : - (∀ v p2', iMsg_car m2 v (Next p2') -∗ ∃ p1', - â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1')) -∗ - (<(Send,i)> m1) ⊑ (<(Send,i)> m2). - Proof. - iIntros "Hle". rewrite iProto_le_unfold. - iRight. iExists (Send, i), (Send, i), m1, m2. by eauto. - Qed. - - Lemma iProto_le_recv i m1 m2 : - (∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', - â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2')) -∗ - (<(Recv,i)> m1) ⊑ (<(Recv,i)> m2). - Proof. - iIntros "Hle". rewrite iProto_le_unfold. - iRight. iExists (Recv, i), (Recv, i), m1, m2. by eauto. - Qed. - - Lemma iProto_le_base a v P p1 p2 : - â–· (p1 ⊑ p2) -∗ - (<a> MSG v {{ P }}; p1) ⊑ (<a> MSG v {{ P }}; p2). - Proof. - rewrite iMsg_base_eq. iIntros "H". destruct a as [[]]. - - iApply iProto_le_send. iIntros (v' p') "(->&Hp&$)". - iExists p1. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". - - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&$)". - iExists p2. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". - Qed. - - Lemma iProto_le_trans p1 p2 p3 : p1 ⊑ p2 -∗ p2 ⊑ p3 -∗ p1 ⊑ p3. - Proof. - iIntros "H1 H2". iLöb as "IH" forall (p1 p2 p3). - destruct (iProto_case p3) as [->|([]&i&m3&->)]. - - iDestruct (iProto_le_end_inv_r with "H2") as "H2". by iRewrite "H2" in "H1". - - iDestruct (iProto_le_send_inv with "H2") as (m2) "[Hp2 H2]". - iRewrite "Hp2" in "H1"; clear p2. - iDestruct (iProto_le_send_inv with "H1") as (m1) "[Hp1 H1]". - iRewrite "Hp1"; clear p1. - iApply iProto_le_send. iIntros (v p3') "Hm3". - iDestruct ("H2" with "Hm3") as (p2') "[Hle Hm2]". - iDestruct ("H1" with "Hm2") as (p1') "[Hle' Hm1]". - iExists p1'. iIntros "{$Hm1} !>". by iApply ("IH" with "Hle'"). - - iDestruct (iProto_le_recv_inv_r with "H2") as (m2) "[Hp2 H3]". - iRewrite "Hp2" in "H1". - iDestruct (iProto_le_recv_inv_r with "H1") as (m1) "[Hp1 H2]". - iRewrite "Hp1". iApply iProto_le_recv. iIntros (v p1') "Hm1". - iDestruct ("H2" with "Hm1") as (p2') "[Hle Hm2]". - iDestruct ("H3" with "Hm2") as (p3') "[Hle' Hm3]". - iExists p3'. iIntros "{$Hm3} !>". by iApply ("IH" with "Hle"). - Qed. - - Lemma iProto_le_refl p : ⊢ p ⊑ p. - Proof. - iLöb as "IH" forall (p). destruct (iProto_case p) as [->|([]&i&m&->)]. - - iApply iProto_le_end. - - iApply iProto_le_send. auto 10 with iFrame. - - iApply iProto_le_recv. auto 10 with iFrame. Qed. Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. @@ -1115,16 +1210,18 @@ Section proto. Lemma iProto_consistent_target ps m a i j : iProto_consistent ps -∗ - ps !!! i ≡ (<(a, j)> m) -∗ + ps !! i ≡ Some (<(a, j)> m) -∗ ⌜is_Some (ps !! j)âŒ. Proof. - rewrite iProto_consistent_unfold. iDestruct 1 as "[Htar _]". iApply "Htar". + rewrite iProto_consistent_unfold. iDestruct 1 as "[Htar _]". + iIntros "H". iApply ("Htar" $! i). + rewrite list_lookup_total_alt. iRewrite "H". done. Qed. Lemma iProto_consistent_step ps m1 m2 i j v p1 : iProto_consistent ps -∗ - ps !!! i ≡ (<(Send, j)> m1) -∗ - ps !!! j ≡ (<(Recv, i)> m2) -∗ + ps !! i ≡ Some (<(Send, j)> m1) -∗ + ps !! j ≡ Some (<(Recv, i)> m2) -∗ iMsg_car m1 v (Next p1) -∗ ∃ p2, iMsg_car m2 v (Next p2) ∗ â–· iProto_consistent (<[i := p1]>(<[j := p2]>ps)). @@ -1132,52 +1229,12 @@ Section proto. iIntros "Hprot #Hi #Hj Hm1". rewrite iProto_consistent_unfold /iProto_consistent_pre. iDestruct "Hprot" as "[_ Hprot]". - iDestruct ("Hprot" with "Hi Hj Hm1") as (p2) "[Hm2 Hprot]". + iDestruct ("Hprot" $! i j with "[Hi] [Hj] Hm1") as (p2) "[Hm2 Hprot]". + { rewrite list_lookup_total_alt. iRewrite "Hi". done. } + { rewrite list_lookup_total_alt. iRewrite "Hj". done. } iExists p2. iFrame. Qed. - Global Instance iProto_own_frag_ne γ s : NonExpansive (iProto_own_frag γ s). - Proof. solve_proper. Qed. - - Lemma iProto_own_auth_agree γ ps i p : - iProto_own_auth γ ps -∗ iProto_own_frag γ i p -∗ â–· (ps !!! i ≡ p). - Proof. Admitted. - (* iIntros "Hâ— Hâ—¯". *) - (* iDestruct (own_valid_2 with "Hâ— Hâ—¯") as "H✓". *) - (* rewrite gmap_view_both_validI. *) - (* iDestruct "H✓" as "[_ [H1 H2]]". *) - (* rewrite list_lookup_total_alt lookup_fmap. *) - (* destruct (ps !! i); last first. *) - (* { rewrite !option_equivI. } *) - (* simpl. rewrite !option_equivI excl_equivI. by iNext. *) - (* Qed. *) - - Lemma iProto_own_auth_update γ ps i p p' : - iProto_own_auth γ ps -∗ iProto_own_frag γ i p ==∗ - iProto_own_auth γ (<[i := p']>ps) ∗ iProto_own_frag γ i p'. - Proof. - iIntros "Hâ— Hâ—¯". - iMod (own_update_2 with "Hâ— Hâ—¯") as "[H1 H2]"; [|iModIntro]. - { eapply (gmap_view_replace _ _ _ (Excl' (Next p'))). done. } - iFrame. rewrite -fmap_insert. Admitted. - - Lemma iProto_own_auth_alloc ps : - ⊢ |==> ∃ γ, iProto_own_auth γ ps ∗ [∗ list] i ↦p ∈ ps, iProto_own γ i p. - Proof. Admitted. - (* iMod (own_alloc (gmap_view_auth (DfracOwn 1) ∅)) as (γ) "Hauth". *) - (* { apply gmap_view_auth_valid. } *) - (* iExists γ. *) - (* iInduction ps as [|i p ps Hin] "IH" using map_ind. *) - (* { iModIntro. iFrame. by iApply big_sepM_empty. } *) - (* iMod ("IH" with "Hauth") as "[Hauth Hfrags]". *) - (* rewrite big_sepM_insert; [|done]. iFrame "Hfrags". *) - (* iMod (own_update with "Hauth") as "[Hauth Hfrag]". *) - (* { apply (gmap_view_alloc _ i (DfracOwn 1) (Excl' (Next p))); [|done|done]. *) - (* by rewrite lookup_fmap Hin. } *) - (* iModIntro. rewrite -fmap_insert. iFrame. *) - (* iExists _. iFrame. iApply iProto_le_refl. *) - (* Qed. *) - Lemma iProto_own_le γ s p1 p2 : iProto_own γ s p1 -∗ â–· (p1 ⊑ p2) -∗ iProto_own γ s p2. Proof. @@ -1185,6 +1242,29 @@ Section proto. iExists p1'. iFrame "H". by iApply (iProto_le_trans with "Hle"). Qed. + Lemma iProto_own_excl γ i (p1 p2 : iProto Σ V) : + iProto_own γ i p1 -∗ iProto_own γ i p2 -∗ False. + Proof. + rewrite /iProto_own. + iDestruct 1 as (p1') "[_ Hp1]". + iDestruct 1 as (p2') "[_ Hp2]". + iDestruct (own_prot_excl with "Hp1 Hp2") as %[]. + Qed. + + Lemma iProto_ctx_agree γ n i p : + iProto_ctx γ n -∗ + iProto_own γ i p -∗ + ⌜i < nâŒ. + Proof. + iIntros "Hctx Hown". + rewrite /iProto_ctx /iProto_own. + iDestruct "Hctx" as (ps <-) "[Hauth Hps]". + iDestruct "Hown" as (p') "[Hle Hown]". + iDestruct (iProto_own_auth_agree_Some with "Hauth Hown") as %HSome. + iPureIntro. + by apply lookup_lt_is_Some_1. + Qed. + Lemma iProto_init ps : â–· iProto_consistent ps -∗ |==> ∃ γ, iProto_ctx γ (length ps) ∗ [∗ list] i ↦p ∈ ps, iProto_own γ i p. @@ -1194,7 +1274,6 @@ Section proto. iExists γ. iFrame. iExists _. by iFrame. Qed. - Lemma iProto_step γ ps_dom i j m1 m2 p1 v : iProto_ctx γ ps_dom -∗ iProto_own γ i (<(Send, j)> m1) -∗ @@ -1204,49 +1283,48 @@ Section proto. iProto_own γ i p1 ∗ iProto_own γ j p2. Proof. iIntros "Hctx Hi Hj Hm". + iDestruct (iProto_ctx_agree with "Hctx Hi") as %Hi. + iDestruct (iProto_ctx_agree with "Hctx Hj") as %Hij. iDestruct "Hi" as (pi) "[Hile Hi]". iDestruct "Hj" as (pj) "[Hjle Hj]". iDestruct "Hctx" as (ps Hdom) "[Hauth Hconsistent]". iDestruct (iProto_own_auth_agree with "Hauth Hi") as "#Hpi". iDestruct (iProto_own_auth_agree with "Hauth Hj") as "#Hpj". iDestruct (own_prot_idx with "Hi Hj") as %Hneq. - iAssert (â–· (<[i:=<(Send, j)> m1]>ps !!! j ≡ pj))%I as "Hpj'". - { by rewrite list_lookup_total_insert_ne. } + iAssert (â–· (<[i:=<(Send, j)> m1]>ps !! j ≡ Some pj))%I as "Hpj'". + { by rewrite list_lookup_insert_ne. } iAssert (â–· (⌜is_Some (ps !! i)⌠∗ (pi ⊑ (<(Send, j)> m1))))%I with "[Hile]" as "[Hi' Hile]". { iNext. iDestruct (iProto_le_msg_inv_r with "Hile") as (m) "#Heq". - iFrame. iRewrite "Heq" in "Hpi". rewrite list_lookup_total_alt. - destruct (ps !! i); [done|]. - iDestruct (iProto_end_message_equivI with "Hpi") as "[]". } + iFrame. iRewrite "Heq" in "Hpi". destruct (ps !! i); [done|]. + by rewrite option_equivI. } iAssert (â–· (⌜is_Some (ps !! j)⌠∗ (pj ⊑ (<(Recv, i)> m2))))%I with "[Hjle]" as "[Hj' Hjle]". { iNext. iDestruct (iProto_le_msg_inv_r with "Hjle") as (m) "#Heq". - iFrame. iRewrite "Heq" in "Hpj". rewrite !list_lookup_total_alt. - destruct (ps !! j); [done|]. - iDestruct (iProto_end_message_equivI with "Hpj") as "[]". } + iFrame. iRewrite "Heq" in "Hpj". + destruct (ps !! j); [done|]. by rewrite !option_equivI. } iDestruct (iProto_consistent_le with "Hconsistent Hpi Hile") as "Hconsistent". iDestruct (iProto_consistent_le with "Hconsistent Hpj' Hjle") as "Hconsistent". iDestruct (iProto_consistent_step _ _ _ i j with "Hconsistent [] [] [Hm //]") as (p2) "[Hm2 Hconsistent]". - { rewrite list_lookup_total_insert_ne; [|done]. - rewrite list_lookup_total_insert_ne; [|done]. - rewrite list_lookup_total_insert; [done|]. admit. } - { rewrite list_lookup_total_insert_ne; [|done]. - rewrite list_lookup_total_insert; [done|]. admit. } + { rewrite list_lookup_insert_ne; [|done]. + rewrite list_lookup_insert_ne; [|done]. + rewrite list_lookup_insert; [done|]. lia. } + { rewrite list_lookup_insert_ne; [|done]. + rewrite list_lookup_insert; [done|]. rewrite insert_length. lia. } iMod (iProto_own_auth_update _ _ _ _ p2 with "Hauth Hj") as "[Hauth Hj]". iMod (iProto_own_auth_update _ _ _ _ p1 with "Hauth Hi") as "[Hauth Hi]". iIntros "!>!>". iExists p2. iFrame "Hm2". - iDestruct "Hi'" as %Hi. iDestruct "Hj'" as %Hj. + iDestruct "Hi'" as %Hi'. iDestruct "Hj'" as %Hj'. iSplitL "Hconsistent Hauth". { iExists (<[i:=p1]> (<[j:=p2]> ps)). iSplit. - { admit. - (* rewrite !dom_insert_lookup_L; [done..|by rewrite lookup_insert_ne]. *)} + { iPureIntro. rewrite !insert_length. done. } iFrame. rewrite list_insert_insert. rewrite list_insert_commute; [|done]. rewrite list_insert_insert. by rewrite list_insert_commute; [|done]. } iSplitL "Hi"; iExists _; iFrame; iApply iProto_le_refl. - Admitted. + Qed. Lemma iProto_target γ ps_dom i a j m : iProto_ctx γ ps_dom -∗ @@ -1265,9 +1343,9 @@ Section proto. iDestruct (iProto_consistent_target with "Hps Hi") as "#$". by iFrame. } iSplitL "HSome". { iNext. iDestruct "HSome" as %Heq. - iPureIntro. simplify_eq. admit. } + iPureIntro. simplify_eq. by apply lookup_lt_is_Some_1. } iSplitL "Hauth Hps"; iExists _; by iFrame. - Admitted. + Qed. (* (** The instances below make it possible to use the tactics [iIntros], *) (* [iExist], [iSplitL]/[iSplitR], [iFrame] and [iModIntro] on [iProto_le] goals. *) *) diff --git a/multi_actris/examples/basics.v b/multi_actris/examples/basics.v index 9efa72b..e6b46d3 100644 --- a/multi_actris/examples/basics.v +++ b/multi_actris/examples/basics.v @@ -1,16 +1,15 @@ From multi_actris.channel Require Import proofmode. Set Default Proof Using "Type". -Definition iProto_empty {Σ} : gmap nat (iProto Σ) := ∅. +Definition iProto_empty {Σ} : list (iProto Σ) := []. Lemma iProto_consistent_empty {Σ} : ⊢ iProto_consistent (@iProto_empty Σ). -Proof. iProto_consistent_take_step. Qed. +Proof. iProto_consistent_take_steps. Qed. -Definition iProto_binary `{!heapGS Σ} : gmap nat (iProto Σ) := - <[0 := (<(Send, 1) @ (x:Z)> MSG #x ; END)%proto ]> - (<[1 := (<(Recv, 0) @ (x:Z)> MSG #x ; END)%proto ]> - ∅). +Definition iProto_binary `{!heapGS Σ} : list (iProto Σ) := + [(<(Send, 1) @ (x:Z)> MSG #x ; END)%proto; + (<(Recv, 0) @ (x:Z)> MSG #x ; END)%proto]. Lemma iProto_binary_consistent `{!heapGS Σ} : ⊢ iProto_consistent iProto_binary. @@ -29,10 +28,10 @@ Definition roundtrip_prog : val := Section channel. Context `{!heapGS Σ, !chanG Σ}. - Definition iProto_roundtrip : gmap nat (iProto Σ) := - <[0 := (<(Send, 1) @ (x:Z)> MSG #x ; <(Recv, 2)> MSG #x; END)%proto ]> - (<[1 := (<(Recv, 0) @ (x:Z)> MSG #x ; <(Send, 2)> MSG #x; END)%proto ]> - (<[2 := (<(Recv, 1) @ (x:Z)> MSG #x ; <(Send, 0)> MSG #x; END)%proto ]> ∅)). + Definition iProto_roundtrip : list (iProto Σ) := + [(<(Send, 1) @ (x:Z)> MSG #x ; <(Recv, 2)> MSG #x; END)%proto; + (<(Recv, 0) @ (x:Z)> MSG #x ; <(Send, 2)> MSG #x; END)%proto; + (<(Recv, 1) @ (x:Z)> MSG #x ; <(Send, 0)> MSG #x; END)%proto]. Lemma iProto_roundtrip_consistent : ⊢ iProto_consistent iProto_roundtrip. @@ -43,15 +42,12 @@ Section channel. {{{ True }}} roundtrip_prog #() {{{ RET #42 ; True }}}. Proof using chanG0 heapGS0 Σ. iIntros (Φ) "_ HΦ". wp_lam. - wp_smart_apply (new_chan_spec 3 iProto_roundtrip); - [lia|set_solver|iApply iProto_roundtrip_consistent|]. + wp_smart_apply (new_chan_spec iProto_roundtrip); + [set_solver|iApply iProto_roundtrip_consistent|]. iIntros (cs) "Hcs". - wp_smart_apply (get_chan_spec _ 0 with "Hcs"); [set_solver|]. - iIntros (c0) "[Hc0 Hcs]". - wp_smart_apply (get_chan_spec _ 1 with "Hcs"); [set_solver|]. - iIntros (c1) "[Hc1 Hcs]". - wp_smart_apply (get_chan_spec _ 2 with "Hcs"); [set_solver|]. - iIntros (c2) "[Hc2 Hcs]". + wp_get_chan (c0) "[Hc0 Hcs]". + wp_get_chan (c1) "[Hc1 Hcs]". + wp_get_chan (c2) "[Hc2 Hcs]". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". wp_recv (x) as "_". wp_send with "[//]". done. } wp_smart_apply (wp_fork with "[Hc2]"). @@ -74,13 +70,13 @@ Definition roundtrip_ref_prog : val := Section roundtrip_ref. Context `{!heapGS Σ, !chanG Σ}. - Definition iProto_roundtrip_ref : gmap nat (iProto Σ) := - <[0 := (<(Send, 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; - <(Recv, 2)> MSG #() {{ l ↦ #(x+2) }} ; END)%proto]> - (<[1 := (<(Recv, 0) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; - <(Send, 2)> MSG #l {{ l ↦ #(x+1) }}; END)%proto]> - (<[2 := (<(Recv, 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; - <(Send, 0)> MSG #() {{ l ↦ #(x+1) }}; END)%proto]> ∅)). + Definition iProto_roundtrip_ref : list (iProto Σ) := + [(<(Send, 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; + <(Recv, 2)> MSG #() {{ l ↦ #(x+2) }} ; END)%proto; + (<(Recv, 0) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; + <(Send, 2)> MSG #l {{ l ↦ #(x+1) }}; END)%proto; + (<(Recv, 1) @ (l:loc) (x:Z)> MSG #l {{ (l ↦ #x)%I }} ; + <(Send, 0)> MSG #() {{ l ↦ #(x+1) }}; END)%proto]. Lemma iProto_roundtrip_ref_consistent : ⊢ iProto_consistent iProto_roundtrip_ref. @@ -95,17 +91,12 @@ Section roundtrip_ref. {{{ True }}} roundtrip_ref_prog #() {{{ RET #42 ; True }}}. Proof using chanG0. iIntros (Φ) "_ HΦ". wp_lam. - wp_smart_apply (new_chan_spec 3 iProto_roundtrip_ref with "[]"). - { lia. } - { set_solver. } - { iApply iProto_roundtrip_ref_consistent. } + wp_smart_apply (new_chan_spec iProto_roundtrip_ref); + [set_solver|iApply iProto_roundtrip_ref_consistent|]. iIntros (cs) "Hcs". - wp_smart_apply (get_chan_spec _ 0 with "Hcs"); [set_solver|]. - iIntros (c0) "[Hc0 Hcs]". - wp_smart_apply (get_chan_spec _ 1 with "Hcs"); [set_solver|]. - iIntros (c1) "[Hc1 Hcs]". - wp_smart_apply (get_chan_spec _ 2 with "Hcs"); [set_solver|]. - iIntros (c2) "[Hc2 Hcs]". + wp_get_chan (c0) "[Hc0 Hcs]". + wp_get_chan (c1) "[Hc1 Hcs]". + wp_get_chan (c2) "[Hc2 Hcs]". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". wp_recv (l x) as "Hl". wp_load. wp_store. by wp_send with "[$Hl]". } @@ -188,10 +179,10 @@ Section roundtrip_ref_rec. (iProto_roundtrip_ref_rec3_aux iProto_roundtrip_ref_rec3). Proof. apply proto_unfold_eq, (fixpoint_unfold _). Qed. - Definition iProto_roundtrip_ref_rec : gmap nat (iProto Σ) := - <[0 := iProto_roundtrip_ref_rec1]> - (<[1 := iProto_roundtrip_ref_rec2]> - (<[2 := iProto_roundtrip_ref_rec3]> ∅)). + Definition iProto_roundtrip_ref_rec : list (iProto Σ) := + [iProto_roundtrip_ref_rec1; + iProto_roundtrip_ref_rec2; + iProto_roundtrip_ref_rec3]. Lemma iProto_roundtrip_ref_rec_consistent : ⊢ iProto_consistent iProto_roundtrip_ref_rec. @@ -210,7 +201,6 @@ Section roundtrip_ref_rec. iIntros "Hloc". iSplit; [done|]. replace (x + 1 + 1)%Z with (x+2)%Z by lia. iFrame. rewrite -iProto_roundtrip_ref_rec2_unfold. - do 2 clean_map 0. do 2 clean_map 1. do 2 clean_map 2. iApply "IH". Qed. @@ -218,17 +208,12 @@ Section roundtrip_ref_rec. {{{ True }}} roundtrip_ref_rec_prog #() {{{ RET #42 ; True }}}. Proof using chanG0. iIntros (Φ) "_ HΦ". wp_lam. - wp_smart_apply (new_chan_spec 3 iProto_roundtrip_ref_rec with "[]"). - { lia. } - { set_solver. } - { iApply iProto_roundtrip_ref_rec_consistent. } + wp_smart_apply (new_chan_spec iProto_roundtrip_ref_rec); + [set_solver|iApply iProto_roundtrip_ref_rec_consistent|]. iIntros (cs) "Hcs". - wp_smart_apply (get_chan_spec _ 0 with "Hcs"); [set_solver|]. - iIntros (c0) "[Hc0 Hcs]". - wp_smart_apply (get_chan_spec _ 1 with "Hcs"); [set_solver|]. - iIntros (c1) "[Hc1 Hcs]". - wp_smart_apply (get_chan_spec _ 2 with "Hcs"); [set_solver|]. - iIntros (c2) "[Hc2 Hcs]". + wp_get_chan (c0) "[Hc0 Hcs]". + wp_get_chan (c1) "[Hc1 Hcs]". + wp_get_chan (c2) "[Hc2 Hcs]". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". wp_pure _. iLöb as "IH". wp_recv (l x) as "Hl". wp_load. wp_store. wp_send with "[$Hl]". @@ -258,13 +243,13 @@ Definition smuggle_ref_prog : val := Section smuggle_ref. Context `{!heapGS Σ, !chanG Σ}. - Definition iProto_smuggle_ref : gmap nat (iProto Σ) := - <[0 := (<(Send, 1) @ (l:loc) (x:Z)> MSG #l {{ l ↦ #x }} ; - <(Recv,1)> MSG #() {{ l ↦ #(x+2) }} ; END)%proto]> - (<[1 := (<(Recv, 0) @ (v:val)> MSG v ; <(Send,2)> MSG v ; - <(Recv, 2)> MSG #(); <(Send,0)> MSG #() ; END)%proto]> - (<[2 := (<(Recv, 1) @ (l:loc) (x:Z)> MSG #l {{ l ↦ #x }} ; - <(Send,1)> MSG #() {{ l ↦ #(x+2) }} ; END)%proto]> ∅)). + Definition iProto_smuggle_ref : list (iProto Σ) := + [(<(Send, 1) @ (l:loc) (x:Z)> MSG #l {{ l ↦ #x }} ; + <(Recv,1)> MSG #() {{ l ↦ #(x+2) }} ; END)%proto; + (<(Recv, 0) @ (v:val)> MSG v ; <(Send,2)> MSG v ; + <(Recv, 2)> MSG #(); <(Send,0)> MSG #() ; END)%proto; + (<(Recv, 1) @ (l:loc) (x:Z)> MSG #l {{ l ↦ #x }} ; + <(Send,1)> MSG #() {{ l ↦ #(x+2) }} ; END)%proto]. Lemma iProto_smuggle_ref_consistent : ⊢ iProto_consistent iProto_smuggle_ref. @@ -274,18 +259,13 @@ Section smuggle_ref. {{{ True }}} smuggle_ref_prog #() {{{ RET #42 ; True }}}. Proof using chanG0 heapGS0 Σ. iIntros (Φ) "_ HΦ". wp_lam. - wp_smart_apply (new_chan_spec 3 iProto_smuggle_ref with "[]"). - { lia. } - { set_solver. } - { iApply iProto_smuggle_ref_consistent. } + wp_smart_apply (new_chan_spec iProto_smuggle_ref); + [set_solver|iApply iProto_smuggle_ref_consistent|]. iIntros (cs) "Hcs". wp_pures. - wp_smart_apply (get_chan_spec _ 0 with "Hcs"); [set_solver|]. - iIntros (c0) "[Hc0 Hcs]". - wp_smart_apply (get_chan_spec _ 1 with "Hcs"); [set_solver|]. - iIntros (c1) "[Hc1 Hcs]". - wp_smart_apply (get_chan_spec _ 2 with "Hcs"); [set_solver|]. - iIntros (c2) "[Hc2 Hcs]". + wp_get_chan (c0) "[Hc0 Hcs]". + wp_get_chan (c1) "[Hc1 Hcs]". + wp_get_chan (c2) "[Hc2 Hcs]". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". wp_recv (v) as "_". wp_send with "[//]". wp_recv as "_". by wp_send with "[//]". } @@ -311,149 +291,57 @@ Section parallel. 0 *) - Definition iProto_parallel : gmap nat (iProto Σ) := - <[0 := (<(Send, 1) @ (x1:Z)> MSG #x1 ; + Definition iProto_parallel : list (iProto Σ) := + [(<(Send, 1) @ (x1:Z)> MSG #x1 ; <(Send, 2) @ (x2:Z)> MSG #x2 ; <(Recv, 3) @ (y1:Z)> MSG #(x1+y1); - <(Recv, 4) @ (y2:Z)> MSG #(x2+y2); END)%proto]> - (<[1 := (<(Recv, 0) @ (x:Z)> MSG #x ; - <(Send, 3) @ (y:Z)> MSG #(x+y); END)%proto]> - (<[2 := (<(Recv, 0) @ (x:Z)> MSG #x ; - <(Send, 4) @ (y:Z)> MSG #(x+y) ; END)%proto]> - (<[3 := (<(Recv, 1) @ (x:Z)> MSG #x ; - <(Send, 0)> MSG #x; END)%proto]> - (<[4 := (<(Recv, 2) @ (x:Z)> MSG #x ; - <(Send, 0)> MSG #x ; END)%proto]> ∅)))). + <(Recv, 4) @ (y2:Z)> MSG #(x2+y2); END)%proto; + (<(Recv, 0) @ (x:Z)> MSG #x ; + <(Send, 3) @ (y:Z)> MSG #(x+y); END)%proto; + (<(Recv, 0) @ (x:Z)> MSG #x ; + <(Send, 4) @ (y:Z)> MSG #(x+y) ; END)%proto; + (<(Recv, 1) @ (x:Z)> MSG #x ; + <(Send, 0)> MSG #x; END)%proto; + (<(Recv, 2) @ (x:Z)> MSG #x ; + <(Send, 0)> MSG #x ; END)%proto]. Lemma iProto_parallel_consistent : ⊢ iProto_consistent iProto_parallel. - Proof. - rewrite /iProto_parallel. - iProto_consistent_take_step. - iIntros (x1) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 0. clean_map 1. - iProto_consistent_take_step. - - iIntros (x2) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 0. clean_map 2. - iProto_consistent_take_step. - + iIntros (y1) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 1. clean_map 3. - iProto_consistent_take_step. - * iIntros (y2) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 2. clean_map 4. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 0. clean_map 3. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - iProto_consistent_take_step. - * iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 3. clean_map 0. - iProto_consistent_take_step. - iIntros (y2) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 2. clean_map 4. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 4. clean_map 0. - iProto_consistent_take_step. - + iIntros (y1) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 2. clean_map 4. - iProto_consistent_take_step. - iIntros (y2) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 1. clean_map 3. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 3. clean_map 0. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 4. clean_map 0. - iProto_consistent_take_step. - - iIntros (y1) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 1. clean_map 3. - iProto_consistent_take_step. - iIntros (x2) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 0. clean_map 2. - iProto_consistent_take_step. - + iIntros (y2) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 2. clean_map 4. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 3. clean_map 0. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 4. clean_map 0. - iProto_consistent_take_step. - + iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 3. clean_map 0. - iProto_consistent_take_step. - iIntros (z) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 2. clean_map 4. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 4. clean_map 0. - iProto_consistent_take_step. - Qed. + Proof. rewrite /iProto_parallel. iProto_consistent_take_steps. Qed. End parallel. Section two_buyer. Context `{!heapGS Σ}. - Definition two_buyer_prot : gmap nat (iProto Σ) := - <[0 := (<(Send, 1) @ (title:Z)> MSG #title ; + Definition two_buyer_prot : list (iProto Σ) := + [(<(Send, 1) @ (title:Z)> MSG #title ; <(Recv, 1) @ (quote:Z)> MSG #quote ; - <(Send, 2) @ (contrib:Z)> MSG #contrib ; END)%proto]> - (<[1 := (<(Recv, 0) @ (title:Z)> MSG #title ; + <(Send, 2) @ (contrib:Z)> MSG #contrib ; END)%proto; + (<(Recv, 0) @ (title:Z)> MSG #title ; <(Send, 0) @ (quote:Z)> MSG #quote ; <(Send, 2)> MSG #quote ; <(Recv, 2) @ (b:bool)> MSG #b ; if b then <(Recv, 2) @ (address:Z)> MSG #address ; <(Send, 2) @ (date:Z)> MSG #date ; END - else END)%proto]> - (<[2 := (<(Recv, 1) @ (quote:Z)> MSG #quote ; + else END)%proto; + (<(Recv, 1) @ (quote:Z)> MSG #quote ; <(Recv, 0) @ (contrib:Z)> MSG #contrib ; if bool_decide (contrib >= quote/2)%Z then <(Send, 1)> MSG #true ; <(Send, 1) @ (address:Z)> MSG #address ; <(Recv, 1) @ (date:Z)> MSG #date ; END else - <(Send, 1)> MSG #false ; END)%proto]> - ∅)). + <(Send, 1)> MSG #false ; END)%proto]. Lemma two_buyer_prot_consistent : ⊢ iProto_consistent two_buyer_prot. Proof. - rewrite /two_buyer_prot. - iProto_consistent_take_step. - iIntros (title) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 0. clean_map 1. - iProto_consistent_take_step. - iIntros (quote) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 0. clean_map 1. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 1. clean_map 2. - iProto_consistent_take_step. - iIntros (contrib) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 0. clean_map 2. - case_bool_decide. - - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 2. clean_map 1. - iProto_consistent_take_step. - iIntros (address) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 2. clean_map 1. - iProto_consistent_take_step. - iIntros (date) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 2. clean_map 1. - iProto_consistent_take_step. - - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 2. clean_map 1. - iProto_consistent_take_step. + rewrite /two_buyer_prot. iProto_consistent_take_steps. + case_bool_decide; iProto_consistent_take_steps. Qed. - + End two_buyer. Section two_buyer_ref. @@ -493,47 +381,21 @@ Section two_buyer_ref. END else END)%proto. - Definition two_buyer_ref_prot : gmap nat (iProto Σ) := - <[0 := two_buyer_ref_b1_prot]> - (<[1 := two_buyer_ref_s_prot]> - (<[2 := two_buyer_ref_b2_prot]> - ∅)). + Definition two_buyer_ref_prot : list (iProto Σ) := + [two_buyer_ref_b1_prot; + two_buyer_ref_s_prot; + two_buyer_ref_b2_prot]. + (* TODO: Anonymous variable in this is unsatisfactory *) Lemma two_buyer_ref_prot_consistent : ⊢ iProto_consistent two_buyer_ref_prot. Proof. - rewrite /two_buyer_prot. - iProto_consistent_take_step. - iIntros (title) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 0. clean_map 1. - iProto_consistent_take_step. - iIntros (quote) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 0. clean_map 1. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 1. clean_map 2. - iProto_consistent_take_step. - iIntros (l1 amount1 contrib) "Hl1". iExists _,_,_. iSplit; [done|]. iFrame. - clean_map 0. clean_map 2. - iProto_consistent_take_step. - iIntros (b) "Hl1". iExists _. iSplit; [done|]. iFrame. - clean_map 0. clean_map 2. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - clean_map 1. clean_map 2. - destruct b. - - iProto_consistent_take_step. - iIntros (l2 amount2 address) "Hl2". iExists _,_,_. iSplit; [done|]. iFrame. - clean_map 2. clean_map 1. - iProto_consistent_take_step. - iIntros (date) "Hl2". iExists _. iSplit; [done|]. iFrame. - iProto_consistent_take_step. - - iProto_consistent_take_step. + rewrite /two_buyer_prot. iProto_consistent_take_steps. + iNext. destruct x4; iProto_consistent_take_steps. Qed. End two_buyer_ref. - Section forwarder. Context `{!heapGS Σ}. @@ -548,36 +410,25 @@ Section forwarder. *) - Definition iProto_forwarder : gmap nat (iProto Σ) := - <[0 := (<(Send, 1) @ (x:Z)> MSG #x ; + Definition iProto_forwarder : list (iProto Σ) := + [(<(Send, 1) @ (x:Z)> MSG #x ; <(Send, 1) @ (b:bool)> MSG #b ; - <(Recv, if b then 2 else 3) > MSG #x ; END)%proto]> - (<[1 := (<(Recv, 0) @ (x:Z)> MSG #x ; + <(Recv, if b then 2 else 3) > MSG #x ; END)%proto; + (<(Recv, 0) @ (x:Z)> MSG #x ; <(Recv, 0) @ (b:bool)> MSG #b; - <(Send, if b then 2 else 3)> MSG #x ; END)%proto]> - (<[2 := (<(Recv, 1) @ (x:Z)> MSG #x ; - <(Send, 0)> MSG #x ; END)%proto]> - (<[3 := (<(Recv, 1) @ (x:Z)> MSG #x ; - <(Send, 0)> MSG #x ; END)%proto]> ∅))). + <(Send, if b then 2 else 3)> MSG #x ; END)%proto; + (<(Recv, 1) @ (x:Z)> MSG #x ; + <(Send, 0)> MSG #x ; END)%proto; + (<(Recv, 1) @ (x:Z)> MSG #x ; + <(Send, 0)> MSG #x ; END)%proto]. + (* TODO: Anonymous variable in this is unsatisfactory *) Lemma iProto_forwarder_consistent : ⊢ iProto_consistent iProto_forwarder. Proof. rewrite /iProto_forwarder. - iProto_consistent_take_step. - iIntros (x) "_". iExists _. iSplit; [done|]. iSplit; [done|]. - iProto_consistent_take_step. - iIntros ([]) "_". - - iExists _. iSplit; [done|]. iSplit; [done|]. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - repeat clean_map 0. repeat clean_map 1. - iProto_consistent_take_steps. - - iExists _. iSplit; [done|]. iSplit; [done|]. - iProto_consistent_take_step. - iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - repeat clean_map 0. repeat clean_map 1. - iProto_consistent_take_steps. + iProto_consistent_take_steps. + destruct x0; iProto_consistent_take_steps. Qed. End forwarder. @@ -637,11 +488,11 @@ Section forwarder_rec. (iProto_forwarder_rec_n_aux iProto_forwarder_rec_n). Proof. apply (fixpoint_unfold _). Qed. - Definition iProto_forwarder_rec : gmap nat (iProto Σ) := - <[0 := iProto_forwarder_rec_0]> - (<[1 := iProto_forwarder_rec_1]> - (<[2 := iProto_forwarder_rec_n]> - (<[3 := iProto_forwarder_rec_n]>∅))). + Definition iProto_forwarder_rec : list (iProto Σ) := + [iProto_forwarder_rec_0; + iProto_forwarder_rec_1; + iProto_forwarder_rec_n; + iProto_forwarder_rec_n]. Lemma iProto_forwarder_rec_consistent : ⊢ iProto_consistent iProto_forwarder_rec. @@ -665,7 +516,6 @@ Section forwarder_rec. iProto_consistent_take_step. iIntros "_". iSplit; [done|]. iSplit; [done|]. repeat clean_map 1. repeat clean_map 2. repeat clean_map 0. - rewrite (insert_commute _ 2 1); [|done]. iEval (rewrite -iProto_forwarder_rec_1_unfold). iEval (rewrite -iProto_forwarder_rec_n_unfold). iEval (rewrite -iProto_forwarder_rec_n_unfold). @@ -680,7 +530,6 @@ Section forwarder_rec. iProto_consistent_take_step. iIntros "_". iSplit; [done|]. iSplit; [done|]. repeat clean_map 1. repeat clean_map 3. repeat clean_map 0. - rewrite (insert_commute _ 3 1); [|done]. iEval (rewrite -iProto_forwarder_rec_1_unfold). iEval (rewrite -iProto_forwarder_rec_n_unfold). iEval (rewrite -iProto_forwarder_rec_n_unfold). diff --git a/multi_actris/examples/leader_election.v b/multi_actris/examples/leader_election.v index cb21e82..0a3bcff 100644 --- a/multi_actris/examples/leader_election.v +++ b/multi_actris/examples/leader_election.v @@ -133,98 +133,54 @@ Section ring_leader_election_example. Definition prot_tail (i_max : nat) : iProto Σ := (<(Send,0)> MSG #i_max; END)%proto. - Definition pre_prot_pool id_max : gmap nat (iProto Σ) := - <[0 := (<(Recv,1) @ (id_max : nat)> MSG #id_max ; + Definition pre_prot_pool id_max : list (iProto Σ) := + [(<(Recv,1) @ (id_max : nat)> MSG #id_max ; <(Recv,2)> MSG #id_max ; <(Recv,3)> MSG #id_max ; - END)%proto ]> - (<[1 := prot_tail id_max ]> - (<[2 := prot_tail id_max ]> - (<[3 := prot_tail id_max ]> ∅))). + END)%proto; + prot_tail id_max; + prot_tail id_max; + prot_tail id_max]. Lemma pre_prot_pool_consistent id_max : ⊢ iProto_consistent (pre_prot_pool id_max). Proof. rewrite /pre_prot_pool. iProto_consistent_take_steps. Qed. - Definition prot_pool : gmap nat (iProto Σ) := - <[0 := (<(Recv,1) @ (id_max : nat)> MSG #id_max ; + Definition prot_pool : list (iProto Σ) := + [(<(Recv,1) @ (id_max : nat)> MSG #id_max ; <(Recv,2)> MSG #id_max ; <(Recv,3)> MSG #id_max ; - END)%proto ]> - (<[1 := rle_preprot 3 1 2 prot_tail ]> - (<[2 := rle_prot 1 2 3 prot_tail false ]> - (<[3 := rle_preprot 2 3 1 prot_tail ]> ∅))). + END)%proto; + rle_preprot 3 1 2 prot_tail; + rle_prot 1 2 3 prot_tail false; + rle_preprot 2 3 1 prot_tail]. Lemma prot_pool_consistent : ⊢ iProto_consistent prot_pool. Proof. rewrite /prot_pool /rle_preprot. rewrite !rle_prot_unfold'. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - repeat clean_map 0. repeat clean_map 1. - repeat clean_map 2. repeat clean_map 3. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - repeat clean_map 0. repeat clean_map 1. - repeat clean_map 2. repeat clean_map 3. + iProto_consistent_take_steps. + case_bool_decide; try lia. + case_bool_decide; try lia. rewrite !rle_prot_unfold'. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - repeat clean_map 0. repeat clean_map 1. - repeat clean_map 2. repeat clean_map 3. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - repeat clean_map 0. repeat clean_map 1. - repeat clean_map 2. repeat clean_map 3. + iProto_consistent_take_steps. + case_bool_decide; try lia. + case_bool_decide; try lia. rewrite !rle_prot_unfold'. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - repeat clean_map 0. repeat clean_map 1. - repeat clean_map 2. repeat clean_map 3. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - repeat clean_map 0. repeat clean_map 1. - repeat clean_map 2. repeat clean_map 3. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - iProto_consistent_take_step. - iProto_consistent_resolve_step. - repeat clean_map 0. repeat clean_map 1. - repeat clean_map 2. repeat clean_map 3. - repeat (rewrite (insert_commute _ _ 3); [|lia]). - repeat (rewrite (insert_commute _ _ 2); [|lia]). - repeat (rewrite (insert_commute _ _ 1); [|lia]). - repeat (rewrite (insert_commute _ _ 0); [|lia]). - iApply pre_prot_pool_consistent. + iProto_consistent_take_steps. Qed. Lemma program_spec : {{{ True }}} program #() {{{ RET #(); True }}}. Proof. iIntros (Φ) "_ HΦ". wp_lam. - wp_smart_apply (new_chan_spec 4 prot_pool); - [lia|set_solver|iApply prot_pool_consistent|]. + wp_smart_apply (new_chan_spec prot_pool); + [set_solver|iApply prot_pool_consistent|]. iIntros (cs) "Hcs". - wp_smart_apply (get_chan_spec _ 0 with "Hcs"); [done|]. - iIntros (c0) "[Hc0 Hcs]". - wp_smart_apply (get_chan_spec _ 1 with "Hcs"); [done|]. - iIntros (c1) "[Hc1 Hcs]". - wp_smart_apply (get_chan_spec _ 2 with "Hcs"); [done|]. - iIntros (c2) "[Hc2 Hcs]". - wp_smart_apply (get_chan_spec _ 3 with "Hcs"); [done|]. - iIntros (c3) "[Hc3 Hcs]". + wp_get_chan (c0) "[Hc0 Hcs]". + wp_get_chan (c1) "[Hc1 Hcs]". + wp_get_chan (c2) "[Hc2 Hcs]". + wp_get_chan (c3) "[Hc3 Hcs]". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". wp_smart_apply (init_spec with "Hc1"). iIntros (i') "Hc1". by wp_send with "[//]". } -- GitLab From 19d0c303fd32e32922582ac383f8b3dc2c5aa5af Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Sat, 23 Mar 2024 18:31:53 +0100 Subject: [PATCH 55/81] Further clean up of proofmode --- multi_actris/channel/proofmode.v | 11 +++++------ 1 file changed, 5 insertions(+), 6 deletions(-) diff --git a/multi_actris/channel/proofmode.v b/multi_actris/channel/proofmode.v index a125a6f..f0fff71 100644 --- a/multi_actris/channel/proofmode.v +++ b/multi_actris/channel/proofmode.v @@ -397,11 +397,10 @@ Tactic Notation "iProto_consistent_take_step_step" := let m2 := fresh in iIntros (i j m1 m2) "#Hm1 #Hm2"; repeat (destruct i as [|i]=> /=; - [try (rewrite lookup_total_nil); try (by rewrite iProto_end_message_equivI); + [try (by rewrite iProto_end_message_equivI); iDestruct (iProto_message_equivI with "Hm1") as "[%Heq1 Hm1']";simplify_eq=> /=| - try (rewrite lookup_total_nil); try (by rewrite iProto_end_message_equivI)]); - try (rewrite lookup_total_nil); + try (by rewrite iProto_end_message_equivI)]); try (by rewrite iProto_end_message_equivI); iDestruct (iProto_message_equivI with "Hm2") as "[%Heq2 Hm2']";simplify_eq=> /=; @@ -418,10 +417,10 @@ Tactic Notation "iProto_consistent_take_step_target" := iIntros (i j a m); rewrite /valid_target; iIntros "#Hm1"; repeat (destruct i as [|i]=> /=; - [try (rewrite lookup_total_nil); try (by rewrite iProto_end_message_equivI); + [try (by rewrite iProto_end_message_equivI); by iDestruct (iProto_message_equivI with "Hm1") - as "[%Heq1 Hm1']" ; simplify_eq=> /=| - try (rewrite lookup_total_nil); try (by rewrite iProto_end_message_equivI)]). + as "[%Heq1 _]" ; simplify_eq=> /=| + try (by rewrite iProto_end_message_equivI)]). Tactic Notation "iProto_consistent_take_step" := try iNext; -- GitLab From 59850ce17ac2fbe2f31df917eca9bba4d1f83cc0 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Sat, 23 Mar 2024 18:32:00 +0100 Subject: [PATCH 56/81] Added leader election to _CoqProject --- _CoqProject | 1 + 1 file changed, 1 insertion(+) diff --git a/_CoqProject b/_CoqProject index 4489ccb..496c210 100644 --- a/_CoqProject +++ b/_CoqProject @@ -61,3 +61,4 @@ multi_actris/channel/proto.v multi_actris/channel/channel.v multi_actris/channel/proofmode.v multi_actris/examples/basics.v +multi_actris/examples/leader_election.v -- GitLab From 834704aba336443335543e09a1ad645288350f60 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Sun, 24 Mar 2024 16:09:57 +0100 Subject: [PATCH 57/81] Improved proofmode --- multi_actris/channel/proofmode.v | 87 +++++++------------------ multi_actris/examples/basics.v | 33 +++------- multi_actris/examples/leader_election.v | 9 +-- 3 files changed, 35 insertions(+), 94 deletions(-) diff --git a/multi_actris/channel/proofmode.v b/multi_actris/channel/proofmode.v index f0fff71..5f9d7b8 100644 --- a/multi_actris/channel/proofmode.v +++ b/multi_actris/channel/proofmode.v @@ -205,42 +205,14 @@ Tactic Notation "wp_recv_core" tactic3(tac_intros) "as" tactic3(tac) := | _ => fail "wp_recv: not a 'wp'" end. + Tactic Notation "wp_recv" "as" constr(pat) := wp_recv_core (idtac) as (fun H => iDestructHyp H as pat). - Tactic Notation "wp_recv" "(" simple_intropattern_list(xs) ")" "as" constr(pat) := wp_recv_core (intros xs) as (fun H => iDestructHyp H as pat). -Tactic Notation "wp_recv" "(" simple_intropattern_list(xs) ")" "as" "(" simple_intropattern(x1) ")" - constr(pat) := - wp_recv_core (intros xs) as (fun H => iDestructHyp H as ( x1 ) pat). -Tactic Notation "wp_recv" "(" simple_intropattern_list(xs) ")" "as" "(" simple_intropattern(x1) - simple_intropattern(x2) ")" constr(pat) := - wp_recv_core (intros xs) as (fun H => iDestructHyp H as ( x1 x2 ) pat). -Tactic Notation "wp_recv" "(" simple_intropattern_list(xs) ")" "as" "(" simple_intropattern(x1) - simple_intropattern(x2) simple_intropattern(x3) ")" constr(pat) := - wp_recv_core (intros xs) as (fun H => iDestructHyp H as ( x1 x2 x3 ) pat). -Tactic Notation "wp_recv" "(" simple_intropattern_list(xs) ")" "as" "(" simple_intropattern(x1) - simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4) ")" - constr(pat) := - wp_recv_core (intros xs) as (fun H => iDestructHyp H as ( x1 x2 x3 x4 ) pat). -Tactic Notation "wp_recv" "(" simple_intropattern_list(xs) ")" "as" "(" simple_intropattern(x1) - simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4) - simple_intropattern(x5) ")" constr(pat) := - wp_recv_core (intros xs) as (fun H => iDestructHyp H as ( x1 x2 x3 x4 x5 ) pat). -Tactic Notation "wp_recv" "(" simple_intropattern_list(xs) ")" "as" "(" simple_intropattern(x1) - simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4) - simple_intropattern(x5) simple_intropattern(x6) ")" constr(pat) := - wp_recv_core (intros xs) as (fun H => iDestructHyp H as ( x1 x2 x3 x4 x5 x6 ) pat). -Tactic Notation "wp_recv" "(" simple_intropattern_list(xs) ")" "as" "(" simple_intropattern(x1) - simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4) - simple_intropattern(x5) simple_intropattern(x6) simple_intropattern(x7) ")" - constr(pat) := - wp_recv_core (intros xs) as (fun H => iDestructHyp H as ( x1 x2 x3 x4 x5 x6 x7 ) pat). -Tactic Notation "wp_recv" "(" simple_intropattern_list(xs) ")" "as" "(" simple_intropattern(x1) - simple_intropattern(x2) simple_intropattern(x3) simple_intropattern(x4) - simple_intropattern(x5) simple_intropattern(x6) simple_intropattern(x7) - simple_intropattern(x8) ")" constr(pat) := - wp_recv_core (intros xs) as (fun H => iDestructHyp H as ( x1 x2 x3 x4 x5 x6 x7 x8 ) pat). +Tactic Notation "wp_recv" "(" simple_intropattern_list(xs) ")" "as" + "(" ne_simple_intropattern_list(ys) ")" constr(pat) := + wp_recv_core (intros xs) as (fun H => _iDestructHyp H ys pat). Lemma tac_wp_send `{!chanG Σ, !heapGS Σ} {TT : tele} Δ neg i js K (n:nat) w c v p m tv tP tp Φ : w = #n → @@ -318,31 +290,8 @@ Tactic Notation "wp_send_core" tactic3(tac_exist) "with" constr(pat) := Tactic Notation "wp_send" "with" constr(pat) := wp_send_core (idtac) with pat. -Tactic Notation "wp_send" "(" uconstr(x1) ")" "with" constr(pat) := - wp_send_core (eexists x1) with pat. -Tactic Notation "wp_send" "(" uconstr(x1) uconstr(x2) ")" "with" constr(pat) := - wp_send_core (eexists x1; eexists x2) with pat. -Tactic Notation "wp_send" "(" uconstr(x1) uconstr(x2) uconstr(x3) ")" - "with" constr(pat) := - wp_send_core (eexists x1; eexists x2; eexists x3) with pat. -Tactic Notation "wp_send" "(" uconstr(x1) uconstr(x2) uconstr(x3) uconstr(x4) ")" - "with" constr(pat) := - wp_send_core (eexists x1; eexists x2; eexists x3; eexists x4) with pat. -Tactic Notation "wp_send" "(" uconstr(x1) uconstr(x2) uconstr(x3) uconstr(x4) - uconstr(x5) ")" "with" constr(pat) := - wp_send_core (eexists x1; eexists x2; eexists x3; eexists x4; eexists x5) with pat. -Tactic Notation "wp_send" "(" uconstr(x1) uconstr(x2) uconstr(x3) uconstr(x4) ")" - uconstr(x5) uconstr(x6) ")" "with" constr(pat) := - wp_send_core (eexists x1; eexists x2; eexists x3; eexists x4; eexists x5; - eexists x6) with pat. -Tactic Notation "wp_send" "(" uconstr(x1) uconstr(x2) uconstr(x3) uconstr(x4) ")" - uconstr(x5) uconstr(x6) uconstr(x7) ")" "with" constr(pat) := - wp_send_core (eexists x1; eexists x2; eexists x3; eexists x4; eexists x5; - eexists x6; eexists x7) with pat. -Tactic Notation "wp_send" "(" uconstr(x1) uconstr(x2) uconstr(x3) uconstr(x4) ")" - uconstr(x5) uconstr(x6) uconstr(x7) uconstr(x8) ")" "with" constr(pat) := - wp_send_core (eexists x1; eexists x2; eexists x3; eexists x4; eexists x5; - eexists x6; eexists x7; eexists x8) with pat. +Tactic Notation "wp_send" "(" ne_uconstr_list(xs) ")" "with" constr(pat) := + wp_send_core (ltac1_list_iter ltac:(fun x => eexists x) xs) with pat. Lemma iProto_consistent_equiv_proof {Σ} (ps : list (iProto Σ)) : (∀ i j, valid_target ps i j) ∗ @@ -428,10 +377,6 @@ Tactic Notation "iProto_consistent_take_step" := iSplitR; [iProto_consistent_take_step_target| iProto_consistent_take_step_step]. -Tactic Notation "clean_map" constr(i) := - iEval (repeat (rewrite (insert_commute _ _ i); [|done])); - rewrite (insert_insert _ i). - Tactic Notation "iProto_consistent_resolve_step" := repeat iIntros (?); repeat iIntros "?"; repeat iExists _; repeat (iSplit; [done|]); try iFrame. @@ -440,4 +385,22 @@ Tactic Notation "iProto_consistent_take_steps" := repeat (iProto_consistent_take_step; iProto_consistent_resolve_step). Tactic Notation "wp_get_chan" "(" simple_intropattern(c) ")" constr(pat) := - wp_smart_apply (get_chan_spec with "[$]"); iIntros (c); iIntros pat. + wp_smart_apply (get_chan_spec with "[$]"); iIntros (c) "[_TMP ?]"; + iRevert "_TMP"; iIntros pat. + +Ltac ltac1_list_iter2 tac l1 l2 := + let go := ltac2:(tac l1 l2 |- List.iter2 (ltac1:(tac x y |- tac x y) tac) + (of_ltac1_list l1) (of_ltac1_list l2)) in + go tac l1 l2. + +Tactic Notation "wp_new_chan" constr(prot) "as" + "(" simple_intropattern_list(xs) ")" constr_list(pats) := + wp_smart_apply (new_chan_spec prot); + [set_solver| |iIntros (?) "?"]; + [|ltac1_list_iter2 ltac:(fun x y => wp_get_chan (x) y) xs pats]. + +Tactic Notation "wp_new_chan" constr(prot) "with" constr(lem) "as" + "(" simple_intropattern_list(xs) ")" constr_list(pats) := + wp_smart_apply (new_chan_spec prot); + [set_solver|by iApply lem|iIntros (?) "?"]; + ltac1_list_iter2 ltac:(fun x y => wp_get_chan (x) y) xs pats. diff --git a/multi_actris/examples/basics.v b/multi_actris/examples/basics.v index e6b46d3..4656fbd 100644 --- a/multi_actris/examples/basics.v +++ b/multi_actris/examples/basics.v @@ -42,12 +42,8 @@ Section channel. {{{ True }}} roundtrip_prog #() {{{ RET #42 ; True }}}. Proof using chanG0 heapGS0 Σ. iIntros (Φ) "_ HΦ". wp_lam. - wp_smart_apply (new_chan_spec iProto_roundtrip); - [set_solver|iApply iProto_roundtrip_consistent|]. - iIntros (cs) "Hcs". - wp_get_chan (c0) "[Hc0 Hcs]". - wp_get_chan (c1) "[Hc1 Hcs]". - wp_get_chan (c2) "[Hc2 Hcs]". + wp_new_chan iProto_roundtrip with iProto_roundtrip_consistent + as (c0 c1 c2) "Hc0" "Hc1" "Hc2". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". wp_recv (x) as "_". wp_send with "[//]". done. } wp_smart_apply (wp_fork with "[Hc2]"). @@ -91,12 +87,8 @@ Section roundtrip_ref. {{{ True }}} roundtrip_ref_prog #() {{{ RET #42 ; True }}}. Proof using chanG0. iIntros (Φ) "_ HΦ". wp_lam. - wp_smart_apply (new_chan_spec iProto_roundtrip_ref); - [set_solver|iApply iProto_roundtrip_ref_consistent|]. - iIntros (cs) "Hcs". - wp_get_chan (c0) "[Hc0 Hcs]". - wp_get_chan (c1) "[Hc1 Hcs]". - wp_get_chan (c2) "[Hc2 Hcs]". + wp_new_chan iProto_roundtrip_ref with iProto_roundtrip_ref_consistent + as (c0 c1 c2) "Hc0" "Hc1" "Hc2". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". wp_recv (l x) as "Hl". wp_load. wp_store. by wp_send with "[$Hl]". } @@ -208,12 +200,8 @@ Section roundtrip_ref_rec. {{{ True }}} roundtrip_ref_rec_prog #() {{{ RET #42 ; True }}}. Proof using chanG0. iIntros (Φ) "_ HΦ". wp_lam. - wp_smart_apply (new_chan_spec iProto_roundtrip_ref_rec); - [set_solver|iApply iProto_roundtrip_ref_rec_consistent|]. - iIntros (cs) "Hcs". - wp_get_chan (c0) "[Hc0 Hcs]". - wp_get_chan (c1) "[Hc1 Hcs]". - wp_get_chan (c2) "[Hc2 Hcs]". + wp_new_chan iProto_roundtrip_ref_rec with iProto_roundtrip_ref_rec_consistent + as (c0 c1 c2) "Hc0" "Hc1" "Hc2". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". wp_pure _. iLöb as "IH". wp_recv (l x) as "Hl". wp_load. wp_store. wp_send with "[$Hl]". @@ -259,13 +247,8 @@ Section smuggle_ref. {{{ True }}} smuggle_ref_prog #() {{{ RET #42 ; True }}}. Proof using chanG0 heapGS0 Σ. iIntros (Φ) "_ HΦ". wp_lam. - wp_smart_apply (new_chan_spec iProto_smuggle_ref); - [set_solver|iApply iProto_smuggle_ref_consistent|]. - iIntros (cs) "Hcs". - wp_pures. - wp_get_chan (c0) "[Hc0 Hcs]". - wp_get_chan (c1) "[Hc1 Hcs]". - wp_get_chan (c2) "[Hc2 Hcs]". + wp_new_chan iProto_smuggle_ref with iProto_smuggle_ref_consistent + as (c0 c1 c2) "Hc0" "Hc1" "Hc2". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". wp_recv (v) as "_". wp_send with "[//]". wp_recv as "_". by wp_send with "[//]". } diff --git a/multi_actris/examples/leader_election.v b/multi_actris/examples/leader_election.v index 0a3bcff..e8810fc 100644 --- a/multi_actris/examples/leader_election.v +++ b/multi_actris/examples/leader_election.v @@ -174,13 +174,8 @@ Section ring_leader_election_example. {{{ True }}} program #() {{{ RET #(); True }}}. Proof. iIntros (Φ) "_ HΦ". wp_lam. - wp_smart_apply (new_chan_spec prot_pool); - [set_solver|iApply prot_pool_consistent|]. - iIntros (cs) "Hcs". - wp_get_chan (c0) "[Hc0 Hcs]". - wp_get_chan (c1) "[Hc1 Hcs]". - wp_get_chan (c2) "[Hc2 Hcs]". - wp_get_chan (c3) "[Hc3 Hcs]". + wp_new_chan prot_pool with prot_pool_consistent + as (c0 c1 c2 c3) "Hc0" "Hc1" "Hc2" "Hc3". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". wp_smart_apply (init_spec with "Hc1"). iIntros (i') "Hc1". by wp_send with "[//]". } -- GitLab From 0c291783639d509056445050de9b21790753f5eb Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Sun, 24 Mar 2024 16:11:31 +0100 Subject: [PATCH 58/81] Removed deprecated tactic calls --- multi_actris/examples/basics.v | 4 ---- 1 file changed, 4 deletions(-) diff --git a/multi_actris/examples/basics.v b/multi_actris/examples/basics.v index 4656fbd..36c500f 100644 --- a/multi_actris/examples/basics.v +++ b/multi_actris/examples/basics.v @@ -492,13 +492,11 @@ Section forwarder_rec. - iExists _. iSplit; [done|]. iSplit; [done|]. iProto_consistent_take_step. iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - repeat clean_map 1. repeat clean_map 2. repeat clean_map 0. iNext. iEval (rewrite iProto_forwarder_rec_1_unfold). iEval (rewrite iProto_forwarder_rec_n_unfold). iProto_consistent_take_step. iIntros "_". iSplit; [done|]. iSplit; [done|]. - repeat clean_map 1. repeat clean_map 2. repeat clean_map 0. iEval (rewrite -iProto_forwarder_rec_1_unfold). iEval (rewrite -iProto_forwarder_rec_n_unfold). iEval (rewrite -iProto_forwarder_rec_n_unfold). @@ -506,13 +504,11 @@ Section forwarder_rec. - iExists _. iSplit; [done|]. iSplit; [done|]. iProto_consistent_take_step. iIntros "_". iExists _. iSplit; [done|]. iSplit; [done|]. - repeat clean_map 1. repeat clean_map 2. repeat clean_map 0. iNext. iEval (rewrite iProto_forwarder_rec_1_unfold). iEval (rewrite iProto_forwarder_rec_n_unfold). iProto_consistent_take_step. iIntros "_". iSplit; [done|]. iSplit; [done|]. - repeat clean_map 1. repeat clean_map 3. repeat clean_map 0. iEval (rewrite -iProto_forwarder_rec_1_unfold). iEval (rewrite -iProto_forwarder_rec_n_unfold). iEval (rewrite -iProto_forwarder_rec_n_unfold). -- GitLab From 684acce4f53d825005517e610b3720b7176a5b70 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Sun, 24 Mar 2024 16:25:25 +0100 Subject: [PATCH 59/81] Added last attempted changes of proto_alt --- multi_actris/channel/proto_alt.v | 21 +++++++++++++++------ 1 file changed, 15 insertions(+), 6 deletions(-) diff --git a/multi_actris/channel/proto_alt.v b/multi_actris/channel/proto_alt.v index 29d8603..8e27dbf 100644 --- a/multi_actris/channel/proto_alt.v +++ b/multi_actris/channel/proto_alt.v @@ -502,19 +502,28 @@ Program Definition iProto_elim {Σ V A} Next Obligation. solve_proper. Qed. Lemma iProto_elim_message {Σ V} {A:ofe} + (x : A) (f : action → iMsg Σ V -> A) a m + `{Hf : ∀ a, Proper ((≡) ==> (≡)) (f a)} : + iProto_elim x f (iProto_message a m) ≡ f a m. + Proof. + rewrite /iProto_elim iProto_message_eq /iProto_message_def /=. + setoid_rewrite proto_elim_message. + { apply Hf. done. } + intros a' f1 f2 Hf'. apply Hf=> v p /=. apply Hf'. +Qed. + +Lemma iProto_elim_message' {Σ V} {A:ofe} (x : A) (f : action → iMsg Σ V -> A) a m - (* `{Hf : ∀ a, Proper ((≡) ==> (≡)) (f a)} : *) - : + `{Hf : ∀ a, Proper ((≡) ==> (≡)) (f a)} : iProto_elim x f (iProto_message a m) ≡ f a m. Proof. rewrite /iProto_elim. rewrite iProto_message_eq /iProto_message_def. simpl. setoid_rewrite proto_elim_message. - { f_equiv. destruct m. f_equiv. simpl. - admit. } + { f_equiv. done. } intros a'. - intros f1 f2 Hf'. f_equiv. f_equiv. -Admitted. + intros f1 f2 Hf'. f_equiv. done. +Qed. Definition nat_beq := Eval compute in Nat.eqb. -- GitLab From 3cfa16786c293611ee9c2584e4c83945609fcc2c Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Sun, 24 Mar 2024 17:35:05 +0100 Subject: [PATCH 60/81] Simplified invariant --- multi_actris/channel/channel.v | 20 ++++++++++---------- 1 file changed, 10 insertions(+), 10 deletions(-) diff --git a/multi_actris/channel/channel.v b/multi_actris/channel/channel.v index 11c498b..4ed4b47 100644 --- a/multi_actris/channel/channel.v +++ b/multi_actris/channel/channel.v @@ -93,9 +93,8 @@ Definition tok `{!chanG Σ} (γ : gname) : iProp Σ := own γ (Excl ()). Definition chan_inv `{!heapGS Σ, !chanG Σ} γ γE γt i j (l:loc) : iProp Σ := (l ↦ NONEV ∗ tok γt)%I ∨ - (∃ v m p, l ↦ SOMEV v ∗ - iProto_own γ i (<(Send, j)> m)%proto ∗ - iMsg_car m v (Next p) ∗ own γE (â—E (Next p))) ∨ + (∃ v p, l ↦ SOMEV v ∗ + iProto_own γ i (<(Send, j)> MSG v ; p)%proto ∗ own γE (â—E (Next p))) ∨ (∃ p, l ↦ NONEV ∗ iProto_own γ i p ∗ own γE (â—E (Next p))). @@ -372,7 +371,7 @@ Section channel. iInv "IHl1" as "HIp". iDestruct "HIp" as "[HIp|HIp]"; last first. { iDestruct "HIp" as "[HIp|HIp]". - - iDestruct "HIp" as (? m p'') "(>Hl & Hown' & HIp)". + - iDestruct "HIp" as (? p'') "(>Hl & Hown' & HIp)". wp_store. iDestruct (iProto_own_excl with "Hown Hown'") as %[]. - iDestruct "HIp" as (p'') "(>Hl' & Hown' & HIp)". @@ -383,8 +382,8 @@ Section channel. iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]"; [by apply excl_auth_update|]. iModIntro. iSplitL "Hl' Hâ— Hown Hle". - { iRight. iLeft. iIntros "!>". iExists _, _, _. iFrame. - iSplitL "Hown Hle"; [by iApply (iProto_own_le with "Hown Hle")|]. + { iRight. iLeft. iIntros "!>". iExists _, _. iFrame. + iDestruct (iProto_own_le with "Hown Hle") as "Hown". by rewrite iMsg_base_eq. } wp_pures. iLöb as "HL". @@ -395,10 +394,10 @@ Section channel. { iDestruct "HIp" as ">[Hl' Htok']". iDestruct (own_valid_2 with "Htok Htok'") as %[]. } iDestruct "HIp" as "[HIp|HIp]". - - iDestruct "HIp" as (? m p'') "(>Hl' & Hown & HIp)". + - iDestruct "HIp" as (? p'') "(>Hl' & Hown & HIp)". wp_load. iModIntro. iSplitL "Hl' Hown HIp". - { iRight. iLeft. iExists _, _,_. iFrame. } + { iRight. iLeft. iExists _,_. iFrame. } wp_pures. iApply ("HL" with "HΦ Htok Hâ—¯"). - iDestruct "HIp" as (p'') "(>Hl' & Hown & Hâ—)". wp_load. @@ -476,12 +475,13 @@ Section channel. { iRight. iRight. iExists _. iFrame. } wp_pures. iApply ("HL" with "[Hâ— Hâ—¯ Hown Hle] HΦ"). iExists _, _, _, _, _, _. iSplit; [done|]. iFrame "#∗". } - iDestruct "HIp" as (w m p'') "(>Hl' & Hown' & Hm & Hp')". + iDestruct "HIp" as (w p'') "(>Hl' & Hown' & Hp')". iInv "IH" as "Hctx". wp_xchg. iDestruct (iProto_own_le with "Hown Hle") as "Hown". - iMod (iProto_step with "Hctx Hown' Hown Hm") as + iMod (iProto_step with "Hctx Hown' Hown []") as (p''') "(Hm & Hctx & Hown & Hown')". + { by rewrite iMsg_base_eq. } iModIntro. iSplitL "Hctx"; [done|]. iModIntro. -- GitLab From 4ca1240f5d6ec0a0a6e905e7286ada31baf71767 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Sun, 24 Mar 2024 18:01:56 +0100 Subject: [PATCH 61/81] Simplified iProto_target lemma --- multi_actris/channel/channel.v | 23 ++++------------------- multi_actris/channel/proto.v | 14 +++++--------- 2 files changed, 9 insertions(+), 28 deletions(-) diff --git a/multi_actris/channel/channel.v b/multi_actris/channel/channel.v index 11c498b..09d12c0 100644 --- a/multi_actris/channel/channel.v +++ b/multi_actris/channel/channel.v @@ -353,16 +353,8 @@ Section channel. iInv "IH" as "Hctx". iDestruct (iProto_le_msg_inv_r with "Hle") as (m') "#Heq". iRewrite "Heq" in "Hown". - iAssert (â–· (⌜i < n⌠∗ iProto_own γ i (<(Send, j)> m') ∗ - iProto_ctx γ n))%I with "[Hctx Hown]" - as "[Hi [Hown Hctx]]". - { iNext. iDestruct (iProto_ctx_agree with "Hctx Hown") as %Hi. - iFrame. done. } - iAssert (â–· (â–· ⌜j < n⌠∗ iProto_own γ i (<(Send, j)> m') ∗ - iProto_ctx γ n))%I with "[Hctx Hown]" - as "[Hj [Hown Hctx]]". - { iNext. iDestruct (iProto_target with "Hctx Hown") as "[Hin [$ $]]". - iFrame. } + iDestruct (iProto_ctx_agree with "Hctx [Hown//]") as "#Hi". + iDestruct (iProto_target with "Hctx [Hown//]") as "#Hj". iRewrite -"Heq" in "Hown". wp_pures. iModIntro. iFrame. wp_smart_apply (vpos_spec); [done|]; iIntros "_". iDestruct "Hi" as %Hi. @@ -444,15 +436,8 @@ Section channel. iInv "IH" as "Hctx". iDestruct (iProto_le_msg_inv_r with "Hle") as (m') "#Heq". iRewrite "Heq" in "Hown". - iAssert (â–· (⌜i < n⌠∗ iProto_own γ i (<(Recv, j)> m') ∗ - iProto_ctx γ n))%I with "[Hctx Hown]" - as "[Hi [Hown Hctx]]". - { iNext. iDestruct (iProto_ctx_agree with "Hctx Hown") as %Hi. - iFrame. done. } - iAssert (â–· (â–· ⌜j < n⌠∗ iProto_own γ i (<(Recv, j)> m') ∗ - iProto_ctx γ n))%I with "[Hctx Hown]" as "[Hj [Hown Hctx]]". - { iNext. iDestruct (iProto_target with "Hctx Hown") as "[Hin [$$]]". - iFrame. } + iDestruct (iProto_ctx_agree with "Hctx [Hown//]") as "#Hi". + iDestruct (iProto_target with "Hctx [Hown//]") as "#Hj". iRewrite -"Heq" in "Hown". wp_pures. iModIntro. iFrame. wp_smart_apply (vpos_spec); [done|]; iIntros "_". iDestruct "Hi" as %Hi. diff --git a/multi_actris/channel/proto.v b/multi_actris/channel/proto.v index aa0826a..dea40ec 100644 --- a/multi_actris/channel/proto.v +++ b/multi_actris/channel/proto.v @@ -1329,7 +1329,7 @@ Section proto. Lemma iProto_target γ ps_dom i a j m : iProto_ctx γ ps_dom -∗ iProto_own γ i (<(a, j)> m) -∗ - â–· (⌜j < ps_domâŒ) ∗ iProto_ctx γ ps_dom ∗ iProto_own γ i (<(a, j)> m). + â–· (⌜j < ps_domâŒ). Proof. iIntros "Hctx Hown". rewrite /iProto_ctx /iProto_own. @@ -1337,14 +1337,10 @@ Section proto. iDestruct "Hown" as (p') "[Hle Hown]". iDestruct (iProto_own_auth_agree with "Hauth Hown") as "#Hi". iDestruct (iProto_le_msg_inv_r with "Hle") as (m') "#Heq". - iAssert (â–· (⌜is_Some (ps !! j)⌠∗ iProto_consistent ps))%I - with "[Hps]" as "[HSome Hps]". - { iNext. iRewrite "Heq" in "Hi". - iDestruct (iProto_consistent_target with "Hps Hi") as "#$". by iFrame. } - iSplitL "HSome". - { iNext. iDestruct "HSome" as %Heq. - iPureIntro. simplify_eq. by apply lookup_lt_is_Some_1. } - iSplitL "Hauth Hps"; iExists _; by iFrame. + iDestruct (iProto_consistent_target with "Hps []") as "#H". + { iNext. iRewrite "Heq" in "Hi". done. } + iNext. iDestruct "H" as %HSome. + iPureIntro. subst. by apply lookup_lt_is_Some_1. Qed. (* (** The instances below make it possible to use the tactics [iIntros], *) -- GitLab From dc5a67a07094fcccb3816bdea1777ad54e24adc6 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Tue, 26 Mar 2024 00:07:57 +0100 Subject: [PATCH 62/81] Added matrix abstraction and redefined channels using it --- _CoqProject | 1 + multi_actris/channel/channel.v | 227 +++++++++------------------------ multi_actris/channel/matrix.v | 221 ++++++++++++++++++++++++++++++++ 3 files changed, 284 insertions(+), 165 deletions(-) create mode 100644 multi_actris/channel/matrix.v diff --git a/_CoqProject b/_CoqProject index 496c210..e46494d 100644 --- a/_CoqProject +++ b/_CoqProject @@ -56,6 +56,7 @@ theories/logrel/examples/compute_service.v theories/logrel/examples/compute_client_list.v multi_actris/utils/cofe_solver_2.v +multi_actris/utils/matrix.v multi_actris/channel/proto_model.v multi_actris/channel/proto.v multi_actris/channel/channel.v diff --git a/multi_actris/channel/channel.v b/multi_actris/channel/channel.v index dec9b4d..f8dc651 100644 --- a/multi_actris/channel/channel.v +++ b/multi_actris/channel/channel.v @@ -24,13 +24,14 @@ the subprotocol relation [⊑] *) From iris.algebra Require Import gmap excl_auth gmap_view. From iris.base_logic.lib Require Import invariants. From iris.heap_lang Require Export primitive_laws notation proofmode. +From multi_actris.utils Require Import matrix. From multi_actris.channel Require Import proto_model. From multi_actris.channel Require Export proto. Set Default Proof Using "Type". (** * The definition of the message-passing connectives *) Definition new_chan : val := - λ: "n", (AllocN ("n"*"n") NONEV, "n"). + λ: "n", new_matrix "n" "n" NONEV. Definition get_chan : val := λ: "cs" "i", ("cs","i"). @@ -42,25 +43,19 @@ Definition wait : val := | SOME <> => "go" "c" end. - -Definition pos (n i j : nat) : nat := i * n + j. -Definition vpos : val := λ: "n" "i" "j", "i"*"n" + "j". - Definition send : val := λ: "c" "j" "v", - let: "n" := Snd (Fst "c") in - let: "ls" := Fst (Fst "c") in + let: "m" := Fst "c" in let: "i" := Snd "c" in - let: "l" := "ls" +â‚— vpos "n" "i" "j" in + let: "l" := matrix_get "m" "i" "j" in "l" <- SOME "v";; wait "l". (* TODO: Move recursion further in *) Definition recv : val := rec: "go" "c" "j" := - let: "n" := Snd (Fst "c") in - let: "ls" := Fst (Fst "c") in + let: "m" := Fst "c" in let: "i" := Snd "c" in - let: "l" := "ls" +â‚— vpos "n" "j" "i" in + let: "l" := matrix_get "m" "j" "i" in let: "v" := Xchg "l" NONEV in match: "v" with NONE => "go" "c" "j" @@ -100,11 +95,11 @@ Definition chan_inv `{!heapGS Σ, !chanG Σ} γ γE γt i j (l:loc) : iProp Σ : Definition iProto_pointsto_def `{!heapGS Σ, !chanG Σ} (c : val) (p : iProto Σ) : iProp Σ := - ∃ γ (γEs : list gname) (l:loc) (i:nat) (n:nat) p', - ⌜ c = (#l,#n,#i)%V ⌠∗ + ∃ γ (γEs : list gname) (m:val) (i:nat) (n:nat) p', + ⌜ c = (m,#i)%V ⌠∗ inv (nroot.@"ctx") (iProto_ctx γ n) ∗ - (∀ i j, ⌜i < n ⌠-∗ ⌜j < n⌠-∗ - ∃ γt, inv (nroot.@"p") (chan_inv γ (γEs !!! i) γt i j (l +â‚— (pos n i j)))) ∗ + is_matrix m n n + (λ i j l, ∃ γt, inv (nroot.@"p") (chan_inv γ (γEs !!! i) γt i j l)) ∗ â–· (p' ⊑ p) ∗ own (γEs !!! i) (â—E (Next p')) ∗ own (γEs !!! i) (â—¯E (Next p')) ∗ iProto_own γ i p'. @@ -118,17 +113,16 @@ Notation "c ↣ p" := (iProto_pointsto c p) (at level 20, format "c ↣ p"). Definition chan_pool `{!heapGS Σ, !chanG Σ} - (cs : val) (i':nat) (ps : list (iProto Σ)) : iProp Σ := - ∃ γ (γEs : list gname) (n:nat) (l:loc), - ⌜cs = (#l,#n)%V⌠∗ ⌜(i' + length ps = n)%nat⌠∗ + (m : val) (i':nat) (ps : list (iProto Σ)) : iProp Σ := + ∃ γ (γEs : list gname) (n:nat), + ⌜(i' + length ps = n)%nat⌠∗ inv (nroot.@"ctx") (iProto_ctx γ n) ∗ - (∀ i j, ⌜i < n ⌠-∗ ⌜j < n⌠-∗ - ∃ γt, inv (nroot.@"p") (chan_inv γ (γEs !!! i) γt i j (l +â‚— (pos n i j)))) ∗ - [∗ list] i ↦ _ ∈ replicate n (), - (∀ p, ⌜i' <= i⌠-∗ ⌜ps !! (i - i') = Some p⌠-∗ - own (γEs !!! i) (â—E (Next p)) ∗ - own (γEs !!! i) (â—¯E (Next p)) ∗ - iProto_own γ i p). + is_matrix m n n (λ i j l, + ∃ γt, inv (nroot.@"p") (chan_inv γ (γEs !!! i) γt i j l)) ∗ + [∗ list] i ↦ p ∈ ps, + (own (γEs !!! (i+i')) (â—E (Next p)) ∗ + own (γEs !!! (i+i')) (â—¯E (Next p)) ∗ + iProto_own γ (i+i') p). Section channel. Context `{!heapGS Σ, !chanG Σ}. @@ -149,54 +143,6 @@ Section channel. iApply (iProto_le_trans with "Hle Hle'"). Qed. - Lemma big_sepL_replicate {A B} n (x1 : A) (x2 : B) (P : nat → iProp Σ) : - ([∗ list] i ↦ _ ∈ replicate n x1, P i) ⊢ - ([∗ list] i ↦ _ ∈ replicate n x2, P i). - Proof. - iIntros "H". - iInduction n as [|n] "IHn"; [done|]. - replace (S n) with (n + 1) by lia. - rewrite !replicate_add /=. iDestruct "H" as "[H1 H2]". - iSplitL "H1"; [by iApply "IHn"|]=> /=. - by rewrite !replicate_length. - Qed. - - Lemma array_to_matrix_pre l n m v : - l ↦∗ replicate (n * m) v -∗ - [∗ list] i ↦ _ ∈ replicate n (), (l +â‚— i*m) ↦∗ replicate m v. - Proof. - iIntros "Hl". - iInduction n as [|n] "IHn"; [done|]. - replace (S n) with (n + 1) by lia. - replace ((n + 1) * m) with (n * m + m) by lia. - rewrite !replicate_add /= array_app=> /=. - iDestruct "Hl" as "[Hl Hls]". - iDestruct ("IHn" with "Hl") as "Hl". - iFrame=> /=. - rewrite Nat.add_0_r !replicate_length. - replace (Z.of_nat (n * m)) with (Z.of_nat n * Z.of_nat m)%Z by lia. - by iFrame. - Qed. - - Lemma array_to_matrix l n v : - l ↦∗ replicate (n * n) v -∗ - [∗ list] i ↦ _ ∈ replicate n (), - [∗ list] j ↦ _ ∈ replicate n (), - (l +â‚— pos n i j) ↦ v. - Proof. - iIntros "Hl". - iDestruct (array_to_matrix_pre with "Hl") as "Hl". - iApply (big_sepL_impl with "Hl"). - iIntros "!>" (i ? _) "Hl". - iApply big_sepL_replicate. - iApply (big_sepL_impl with "Hl"). - iIntros "!>" (j ? HSome) "Hl". - rewrite Loc.add_assoc. - replace (Z.of_nat i * Z.of_nat n + Z.of_nat j)%Z with - (Z.of_nat (i * n + j))%Z by lia. - by apply lookup_replicate in HSome as [-> _]. - Qed. - (** ** Specifications of [send] and [recv] *) Lemma new_chan_spec (ps:list (iProto Σ)) : 0 < length ps → @@ -205,13 +151,10 @@ Section channel. {{{ cs, RET cs; chan_pool cs 0 ps }}}. Proof. iIntros (Hle Φ) "Hps HΦ". wp_lam. - wp_smart_apply wp_allocN; [lia|done|]. - iIntros (l) "[Hl _]". iMod (iProto_init with "Hps") as (γ) "[Hps Hps']". - wp_pures. iApply "HΦ". iAssert (|==> ∃ (γEs : list gname), ⌜length γEs = length ps⌠∗ - [∗ list] i ↦ _ ∈ replicate (length ps) (), + [∗ list] i ↦ p ∈ ps, own (γEs !!! i) (â—E (Next (ps !!! i))) ∗ own (γEs !!! i) (â—¯E (Next (ps !!! i))) ∗ iProto_own γ i (ps !!! i))%I with "[Hps']" as "H". @@ -225,56 +168,42 @@ Section channel. iModIntro. iExists (γEs++[γE]). rewrite !app_length Hlen. iSplit; [iPureIntro=>/=;lia|]=> /=. - rewrite replicate_add. iSplitL "H". { iApply (big_sepL_impl with "H"). iIntros "!>" (i ? HSome') "(Hauth & Hfrag & Hown)". assert (i < length ps) as Hle. - { by apply lookup_replicate_1 in HSome' as [??]. } + { by apply lookup_lt_is_Some_1. } rewrite !lookup_total_app_l; [|lia..]. iFrame. } - rewrite replicate_length Nat.add_0_r. + rewrite Nat.add_0_r. simpl. rewrite right_id_L. rewrite !lookup_total_app_r; [|lia..]. rewrite !Hlen. rewrite Nat.sub_diag. simpl. iFrame. iDestruct "Hp" as "[$ _]". } iMod "H" as (γEs Hlen) "H". - set n := length ps. - iAssert (|={⊤}=> - [∗ list] i ↦ _ ∈ replicate (length ps) (), - [∗ list] j ↦ _ ∈ replicate (length ps) (), - ∃ γt, + iMod (inv_alloc with "Hps") as "#IHp". + wp_smart_apply (new_matrix_spec _ _ _ _ + (λ i j l, ∃ γt, inv (nroot.@"p") (chan_inv γ (γEs !!! i) γt i j - (l +â‚— (pos (length ps) i j))))%I with "[Hl]" as "IH". - { replace (Z.to_nat (Z.of_nat n * Z.of_nat n)) with (n*n) by lia. - iDestruct (array_to_matrix with "Hl") as "Hl". - iApply big_sepL_fupd. - iApply (big_sepL_impl with "Hl"). - iIntros "!>" (i ? HSome') "H1". - iApply big_sepL_fupd. - iApply (big_sepL_impl with "H1"). - iIntros "!>" (j ? HSome'') "H1". + l))%I); [done..| |]. + { iApply (big_sepL_intro). + iIntros "!>" (k tt Hin). + iApply (big_sepL_intro). + iIntros "!>" (k' tt' Hin'). + iIntros (l) "Hl". iMod (own_alloc (Excl ())) as (γ') "Hγ'"; [done|]. - iExists γ'. iApply inv_alloc. iLeft. iFrame. } - iMod "IH" as "#IH". - iMod (inv_alloc with "Hps") as "#IHp". - iExists _,_,_,_. - iModIntro. - iSplit. - { iPureIntro. done. } - iSplit. - { done. } - iFrame "IHp". - iSplit. - { iIntros (i j Hi Hj). - rewrite (big_sepL_lookup _ _ i); [|by rewrite lookup_replicate]. - rewrite (big_sepL_lookup _ _ j); [|by rewrite lookup_replicate]. - iApply "IH". } + iExists γ'. + iApply inv_alloc. + iNext. + iLeft. iFrame. } + iIntros (mat) "Hmat". + iApply "HΦ". + iExists _, _, _. iFrame "#∗". + rewrite left_id. iSplit; [done|]. iApply (big_sepL_impl with "H"). iIntros "!>" (i ? HSome') "(Hauth & Hfrag & Hown)". - iIntros (p Hle' HSome''). - iFrame. rewrite right_id_L in HSome''. + iFrame. rewrite (list_lookup_total_alt ps). - rewrite HSome''. simpl. iFrame. + simpl. rewrite right_id_L. rewrite HSome'. iFrame. Qed. Lemma get_chan_spec cs i ps p : @@ -283,61 +212,24 @@ Section channel. {{{ c, RET c; c ↣ p ∗ chan_pool cs (i+1) ps }}}. Proof. iIntros (Φ) "Hcs HΦ". - iDestruct "Hcs" as (γp γEs n l -> <-) "(#IHp & #IHl & Hl)". - wp_lam. wp_pures. - rewrite replicate_add. simpl. - iDestruct "Hl" as "[Hl1 [Hi Hl3]]". - iDestruct ("Hi" with "[] []") as "(Hauth & Hown & Hp)". - { rewrite right_id_L. rewrite replicate_length. iPureIntro. lia. } - { rewrite right_id_L. rewrite replicate_length. - rewrite Nat.sub_diag. simpl. done. } + iDestruct "Hcs" as (γp γEs n <-) "(#IHp & #Hm & Hl)". + wp_lam. wp_pures. + iDestruct "Hl" as "[Hl Hls]". iModIntro. iApply "HΦ". - iSplitR "Hl1 Hl3". + iSplitL "Hl". { rewrite iProto_pointsto_eq. iExists _, _, _, _, _, _. iSplit; [done|]. - rewrite replicate_length. rewrite right_id_L. iFrame. iFrame "#∗". iNext. done. } - iExists γp, γEs, _, _. iSplit; [done|]. - iSplit. - { iPureIntro. lia. } + iExists γp, γEs, _. iSplit; [done|]. iFrame. iFrame "#∗". - rewrite replicate_add. - simpl. - iSplitL "Hl1". - { iApply (big_sepL_impl with "Hl1"). - iIntros "!>" (i' ? HSome''). - assert (i' < i). - { rewrite lookup_replicate in HSome''. lia. } - iIntros "H" (p' Hle'). lia. } simpl. - iFrame "#∗". - iSplitR. - { iIntros (p' Hle'). rewrite right_id_L in Hle'. - rewrite replicate_length in Hle'. lia. } - iApply (big_sepL_impl with "Hl3"). - iIntros "!>" (i' ? HSome''). - assert (i' < length ps). - { rewrite lookup_replicate in HSome''. lia. } - iIntros "H" (p' Hle' HSome). - iApply "H". - { iPureIntro. lia. } - iPureIntro. - rewrite replicate_length. - rewrite replicate_length in HSome. - replace (i + S i' - i) with (S i') by lia. - simpl. - replace (i + S i' - (i+1)) with (i') in HSome by lia. - done. - Qed. - - Lemma vpos_spec (n i j : nat) : - {{{ True }}} vpos #n #i #j {{{ RET #(pos n i j); True }}}. - Proof. - iIntros (Φ) "_ HΦ". wp_lam. wp_pures. rewrite /pos. - replace (Z.of_nat i * Z.of_nat n + Z.of_nat j)%Z with - (Z.of_nat (i * n + j)) by lia. - by iApply "HΦ". + replace (i + 1 + length ps) with (i + (S $ length ps))%nat by lia. + iFrame "#". + iApply (big_sepL_impl with "Hls"). + iIntros "!>" (k x HSome) "(H1 & H2 & H3)". + replace (S (k + i)) with (k + (i + 1)) by lia. + iFrame. Qed. Lemma send_spec c j v p : @@ -355,10 +247,12 @@ Section channel. iDestruct (iProto_ctx_agree with "Hctx [Hown//]") as "#Hi". iDestruct (iProto_target with "Hctx [Hown//]") as "#Hj". iRewrite -"Heq" in "Hown". wp_pures. iModIntro. iFrame. - wp_smart_apply (vpos_spec); [done|]; iIntros "_". + wp_pures. iDestruct "Hi" as %Hi. iDestruct "Hj" as %Hj. - iDestruct ("Hls" $! i j with "[//] [//]") as (γt) "IHl1". + wp_smart_apply (matrix_get_spec_pers with "Hls"); [done..|]. + iIntros (l') "Hl'". + iDestruct "Hl'" as (γt) "#IHl1". wp_pures. wp_bind (Store _ _). iInv "IHl1" as "HIp". iDestruct "HIp" as "[HIp|HIp]"; last first. @@ -431,17 +325,20 @@ Section channel. rewrite iProto_pointsto_eq. iDestruct "Hc" as (γ γE l i n p' ->) "(#IH & #Hls & Hle & Hâ— & Hâ—¯ & Hown)". - do 6 wp_pure _. wp_bind (Fst _). wp_pure _. + do 5 wp_pure _. + wp_bind (Snd _). iInv "IH" as "Hctx". iDestruct (iProto_le_msg_inv_r with "Hle") as (m') "#Heq". iRewrite "Heq" in "Hown". iDestruct (iProto_ctx_agree with "Hctx [Hown//]") as "#Hi". iDestruct (iProto_target with "Hctx [Hown//]") as "#Hj". iRewrite -"Heq" in "Hown". wp_pures. iModIntro. iFrame. - wp_smart_apply (vpos_spec); [done|]; iIntros "_". + wp_pure _. iDestruct "Hi" as %Hi. iDestruct "Hj" as %Hj. - iDestruct ("Hls" $! j i with "[//] [//]") as (γt) "IHl2". + wp_smart_apply (matrix_get_spec_pers with "Hls"); [done..|]. + iIntros (l') "Hl'". + iDestruct "Hl'" as (γt) "#IHl2". wp_pures. wp_bind (Xchg _ _). iInv "IHl2" as "HIp". diff --git a/multi_actris/channel/matrix.v b/multi_actris/channel/matrix.v new file mode 100644 index 0000000..9b642da --- /dev/null +++ b/multi_actris/channel/matrix.v @@ -0,0 +1,221 @@ +From iris.bi Require Import big_op. +From iris.heap_lang Require Export primitive_laws notation proofmode. + +Definition new_matrix : val := + λ: "m" "n" "v", (AllocN ("m"*"n") "v","m","n"). + +Definition pos (n i j : nat) : nat := i * n + j. +Definition vpos : val := λ: "n" "i" "j", "i"*"n" + "j". + +Definition matrix_get : val := + λ: "m" "i" "j", + let: "l" := Fst (Fst "m") in + let: "n" := Snd "m" in + "l" +â‚— vpos "n" "i" "j". + +Section with_Σ. + Context `{heapGS Σ}. + + Definition is_matrix (v : val) (m n : nat) + (P : nat → nat → loc → iProp Σ) : iProp Σ := + ∃ (l:loc), + ⌜v = PairV (PairV #l #m) #n⌠∗ + [∗ list] i ↦ _ ∈ replicate m (), + [∗ list] j ↦ _ ∈ replicate n (), + P i j (l +â‚— pos n i j). + + + Lemma array_to_matrix_pre l n m v : + l ↦∗ replicate (n * m) v -∗ + [∗ list] i ↦ _ ∈ replicate n (), (l +â‚— i*m) ↦∗ replicate m v. + Proof. + iIntros "Hl". + iInduction n as [|n] "IHn"; [done|]. + replace (S n) with (n + 1) by lia. + replace ((n + 1) * m) with (n * m + m) by lia. + rewrite !replicate_add /= array_app=> /=. + iDestruct "Hl" as "[Hl Hls]". + iDestruct ("IHn" with "Hl") as "Hl". + iFrame=> /=. + rewrite Nat.add_0_r !replicate_length. + replace (Z.of_nat (n * m)) with (Z.of_nat n * Z.of_nat m)%Z by lia. + by iFrame. + Qed. + + (* TODO: rename *) + Lemma big_sepL_replicate_type {A B} n (x1 : A) (x2 : B) (P : nat → iProp Σ) : + ([∗ list] i ↦ _ ∈ replicate n x1, P i) ⊢ + ([∗ list] i ↦ _ ∈ replicate n x2, P i). + Proof. + iIntros "H". + iInduction n as [|n] "IHn"; [done|]. + replace (S n) with (n + 1) by lia. + rewrite !replicate_add /=. iDestruct "H" as "[H1 H2]". + iSplitL "H1"; [by iApply "IHn"|]=> /=. + by rewrite !replicate_length. + Qed. + + Lemma array_to_matrix l m n v : + l ↦∗ replicate (m * n) v -∗ + [∗ list] i ↦ _ ∈ replicate m (), + [∗ list] j ↦ _ ∈ replicate n (), + (l +â‚— pos n i j) ↦ v. + Proof. + iIntros "Hl". + iDestruct (array_to_matrix_pre with "Hl") as "Hl". + iApply (big_sepL_impl with "Hl"). + iIntros "!>" (i ? _) "Hl". + iApply big_sepL_replicate_type. + iApply (big_sepL_impl with "Hl"). + iIntros "!>" (j ? HSome) "Hl". + rewrite Loc.add_assoc. + replace (Z.of_nat i * Z.of_nat n + Z.of_nat j)%Z with + (Z.of_nat (i * n + j))%Z by lia. + by apply lookup_replicate in HSome as [-> _]. + Qed. + + Lemma new_matrix_spec m n v P : + 0 < m → 0 < n → + {{{ [∗ list] i ↦ _ ∈ replicate m (), [∗ list] j ↦ _ ∈ replicate n (), + ∀ l, l ↦ v -∗ P i j l }}} + (* {{{ (∀ i j l, ⌜i < m⌠-∗ ⌜j < n⌠-∗ l ↦ v -∗ P i j l) }}} *) + new_matrix #m #n v + {{{ mat, RET mat; is_matrix mat m n P }}}. + Proof. + iIntros (Hm Hn Φ) "HP HΦ". + wp_lam. wp_pures. + wp_smart_apply wp_allocN; [lia|done|]. + iIntros (l) "[Hl _]". + wp_pures. iApply "HΦ". + iModIntro. + iExists _. iSplit; [done|]. + replace (Z.to_nat (Z.of_nat m * Z.of_nat n)) with (m * n) by lia. + iDestruct (array_to_matrix with "Hl") as "Hl". + iCombine "HP Hl" as "HPl". + rewrite -big_sepL_sep. + iApply (big_sepL_impl with "HPl"). + iIntros "!>" (k x Hin) "H". + rewrite -big_sepL_sep. + iApply (big_sepL_impl with "H"). + iIntros "!>" (k' x' Hin') "[HP Hl]". + by iApply "HP". + Qed. + + Lemma matrix_sep m n l (P : _ → _ → _ → iProp Σ) i j : + i < m → j < n → + ([∗ list] i ↦ _ ∈ replicate m (), + [∗ list] j ↦ _ ∈ replicate n (), + P i j (l +â‚— pos n i j)) -∗ + (P i j (l +â‚— pos n i j) ∗ + (P i j (l +â‚— pos n i j) -∗ + ([∗ list] i ↦ _ ∈ replicate m (), + [∗ list] j ↦ _ ∈ replicate n (), + P i j (l +â‚— pos n i j)))). + Proof. + iIntros (Hm Hn) "Hm". + iDestruct (big_sepL_impl with "Hm []") as "Hm". + { iIntros "!>" (k x HIn). + iApply (big_sepL_lookup_acc _ _ j ()). + by apply lookup_replicate. } + simpl. + rewrite {1}(big_sepL_lookup_acc _ (replicate m ()) i ()); + [|by apply lookup_replicate]. + iDestruct "Hm" as "[[Hij Hi] Hm]". + iFrame. + iIntros "HP". + iDestruct ("Hm" with "[$Hi $HP]") as "Hm". + iApply (big_sepL_impl with "Hm"). + iIntros "!>" (k x Hin). + iIntros "[Hm Hm']". iApply "Hm'". done. + Qed. + + Lemma matrix_sep_pers m n l (P : _ → _ → _ → iProp Σ) i j : + (∀ i j l, Persistent (P i j l)) → + i < m → j < n → + ([∗ list] i ↦ _ ∈ replicate m (), + [∗ list] j ↦ _ ∈ replicate n (), + P i j (l +â‚— pos n i j)) -∗ + (P i j (l +â‚— pos n i j) ∗ + (([∗ list] i ↦ _ ∈ replicate m (), + [∗ list] j ↦ _ ∈ replicate n (), + P i j (l +â‚— pos n i j)))). + Proof. + iIntros (Hpers Hm Hn) "Hm". + iDestruct (matrix_sep with "Hm") as "[#HP Hm]"; [done..|]. + iFrame "#". by iApply "Hm". + Qed. + + (* Lemma is_matrix_sep m n v P i j : *) + (* i < m → j < n → *) + (* is_matrix v m n P -∗ (∃ l, P i j l ∗ (P i j l -∗ is_matrix v m n P)). *) + (* Proof. *) + (* iDestruct 1 as (l ->) "Hm". *) + (* iDestruct (big_sepL_impl with "Hm []") as "Hm". *) + (* { iIntros "!>" (k x HIn). *) + (* iApply (big_sepL_lookup_acc _ _ j ()). *) + (* by apply lookup_replicate. } *) + (* simpl. *) + (* rewrite {1}(big_sepL_lookup_acc _ (replicate m ()) i ()); *) + (* [|by apply lookup_replicate]. *) + (* iDestruct "Hm" as "[[Hij Hi] Hm]". *) + (* iExists (l +â‚— pos n i j). *) + (* iFrame. *) + (* iIntros "HP". *) + (* iExists _. iSplit; [done|]. *) + (* iDestruct ("Hm" with "[$Hi $HP]") as "Hm". *) + (* iApply (big_sepL_impl with "Hm"). *) + (* iIntros "!>" (k x Hin). *) + (* iIntros "[Hm Hm']". iApply "Hm'". done. *) + (* Qed. *) + + Lemma vpos_spec (n i j : nat) : + {{{ True }}} vpos #n #i #j {{{ RET #(pos n i j); True }}}. + Proof. + iIntros (Φ) "_ HΦ". wp_lam. wp_pures. rewrite /pos. + replace (Z.of_nat i * Z.of_nat n + Z.of_nat j)%Z with + (Z.of_nat (i * n + j)) by lia. + by iApply "HΦ". + Qed. + + Lemma matrix_get_spec m n i j v P : + i < m → j < n → + {{{ is_matrix v m n P }}} + matrix_get v #i #j + {{{ l, RET #l; P i j l ∗ (P i j l -∗ is_matrix v m n P) }}}. + Proof. + iIntros (Hm Hn Φ) "Hm HΦ". + wp_lam. + iDestruct "Hm" as (l ->) "Hm". + wp_pures. + wp_smart_apply vpos_spec; [done|]. + iIntros "_". + wp_pures. + iApply "HΦ". + iModIntro. + iDestruct (matrix_sep _ _ _ _ i j with "Hm") as "[$ Hm]"; [lia|lia|]. + iIntros "H". + iDestruct ("Hm" with "H") as "Hm". + iExists _. iSplit; [done|]. iFrame. + Qed. + + Lemma matrix_get_spec_pers m n i j v P : + (∀ i j l, Persistent (P i j l)) → + i < m → j < n → + {{{ is_matrix v m n P }}} + matrix_get v #i #j + {{{ l, RET #l; P i j l ∗ is_matrix v m n P }}}. + Proof. + iIntros (Hpers Hm Hn Φ) "Hm HΦ". + wp_lam. + iDestruct "Hm" as (l ->) "Hm". + wp_pures. + wp_smart_apply vpos_spec; [done|]. + iIntros "_". + wp_pures. + iApply "HΦ". + iModIntro. + iDestruct (matrix_sep_pers _ _ _ _ i j with "Hm") as "[$ Hm]"; [lia|lia|]. + iExists _. iSplit; [done|]. iFrame. + Qed. + +End with_Σ. -- GitLab From 528c968c5948cfa626ff7b78affbe7996c1d538f Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Wed, 27 Mar 2024 19:05:03 +0100 Subject: [PATCH 63/81] Removed redundant _pers specs, and added missing matrix file --- multi_actris/channel/channel.v | 8 +- multi_actris/utils/matrix.v | 162 +++++++++++++++++++++++++++++++++ 2 files changed, 166 insertions(+), 4 deletions(-) create mode 100644 multi_actris/utils/matrix.v diff --git a/multi_actris/channel/channel.v b/multi_actris/channel/channel.v index f8dc651..8573236 100644 --- a/multi_actris/channel/channel.v +++ b/multi_actris/channel/channel.v @@ -250,8 +250,8 @@ Section channel. wp_pures. iDestruct "Hi" as %Hi. iDestruct "Hj" as %Hj. - wp_smart_apply (matrix_get_spec_pers with "Hls"); [done..|]. - iIntros (l') "Hl'". + wp_smart_apply (matrix_get_spec with "Hls"); [done..|]. + iIntros (l') "[Hl' _]". iDestruct "Hl'" as (γt) "#IHl1". wp_pures. wp_bind (Store _ _). iInv "IHl1" as "HIp". @@ -336,8 +336,8 @@ Section channel. wp_pure _. iDestruct "Hi" as %Hi. iDestruct "Hj" as %Hj. - wp_smart_apply (matrix_get_spec_pers with "Hls"); [done..|]. - iIntros (l') "Hl'". + wp_smart_apply (matrix_get_spec with "Hls"); [done..|]. + iIntros (l') "[Hl' _]". iDestruct "Hl'" as (γt) "#IHl2". wp_pures. wp_bind (Xchg _ _). diff --git a/multi_actris/utils/matrix.v b/multi_actris/utils/matrix.v new file mode 100644 index 0000000..066da4f --- /dev/null +++ b/multi_actris/utils/matrix.v @@ -0,0 +1,162 @@ +From iris.bi Require Import big_op. +From iris.heap_lang Require Export primitive_laws notation proofmode. + +Definition new_matrix : val := + λ: "m" "n" "v", (AllocN ("m"*"n") "v","m","n"). + +Definition pos (n i j : nat) : nat := i * n + j. +Definition vpos : val := λ: "n" "i" "j", "i"*"n" + "j". + +Definition matrix_get : val := + λ: "m" "i" "j", + let: "l" := Fst (Fst "m") in + let: "n" := Snd "m" in + "l" +â‚— vpos "n" "i" "j". + +Section with_Σ. + Context `{heapGS Σ}. + + Definition is_matrix (v : val) (m n : nat) + (P : nat → nat → loc → iProp Σ) : iProp Σ := + ∃ (l:loc), + ⌜v = PairV (PairV #l #m) #n⌠∗ + [∗ list] i ↦ _ ∈ replicate m (), + [∗ list] j ↦ _ ∈ replicate n (), + P i j (l +â‚— pos n i j). + + + Lemma array_to_matrix_pre l n m v : + l ↦∗ replicate (n * m) v -∗ + [∗ list] i ↦ _ ∈ replicate n (), (l +â‚— i*m) ↦∗ replicate m v. + Proof. + iIntros "Hl". + iInduction n as [|n] "IHn"; [done|]. + replace (S n) with (n + 1) by lia. + replace ((n + 1) * m) with (n * m + m) by lia. + rewrite !replicate_add /= array_app=> /=. + iDestruct "Hl" as "[Hl Hls]". + iDestruct ("IHn" with "Hl") as "Hl". + iFrame=> /=. + rewrite Nat.add_0_r !replicate_length. + replace (Z.of_nat (n * m)) with (Z.of_nat n * Z.of_nat m)%Z by lia. + by iFrame. + Qed. + + (* TODO: rename *) + Lemma big_sepL_replicate_type {A B} n (x1 : A) (x2 : B) (P : nat → iProp Σ) : + ([∗ list] i ↦ _ ∈ replicate n x1, P i) ⊢ + ([∗ list] i ↦ _ ∈ replicate n x2, P i). + Proof. + iIntros "H". + iInduction n as [|n] "IHn"; [done|]. + replace (S n) with (n + 1) by lia. + rewrite !replicate_add /=. iDestruct "H" as "[H1 H2]". + iSplitL "H1"; [by iApply "IHn"|]=> /=. + by rewrite !replicate_length. + Qed. + + Lemma array_to_matrix l m n v : + l ↦∗ replicate (m * n) v -∗ + [∗ list] i ↦ _ ∈ replicate m (), + [∗ list] j ↦ _ ∈ replicate n (), + (l +â‚— pos n i j) ↦ v. + Proof. + iIntros "Hl". + iDestruct (array_to_matrix_pre with "Hl") as "Hl". + iApply (big_sepL_impl with "Hl"). + iIntros "!>" (i ? _) "Hl". + iApply big_sepL_replicate_type. + iApply (big_sepL_impl with "Hl"). + iIntros "!>" (j ? HSome) "Hl". + rewrite Loc.add_assoc. + replace (Z.of_nat i * Z.of_nat n + Z.of_nat j)%Z with + (Z.of_nat (i * n + j))%Z by lia. + by apply lookup_replicate in HSome as [-> _]. + Qed. + + Lemma new_matrix_spec E m n v P : + 0 < m → 0 < n → + {{{ [∗ list] i ↦ _ ∈ replicate m (), [∗ list] j ↦ _ ∈ replicate n (), + ∀ l, l ↦ v ={E}=∗ P i j l }}} + new_matrix #m #n v @ E + {{{ mat, RET mat; is_matrix mat m n P }}}. + Proof. + iIntros (Hm Hn Φ) "HP HΦ". + wp_lam. wp_pures. + wp_smart_apply wp_allocN; [lia|done|]. + iIntros (l) "[Hl _]". + wp_pures. iApply "HΦ". + iExists _. iSplitR; [done|]. + replace (Z.to_nat (Z.of_nat m * Z.of_nat n)) with (m * n) by lia. + iDestruct (array_to_matrix with "Hl") as "Hl". + iCombine "HP Hl" as "HPl". + rewrite -big_sepL_sep. + iApply big_sepL_fupd. + iApply (big_sepL_impl with "HPl"). + iIntros "!>" (k x Hin) "H". + rewrite -big_sepL_sep. + iApply big_sepL_fupd. + iApply (big_sepL_impl with "H"). + iIntros "!>" (k' x' Hin') "[HP Hl]". + by iApply "HP". + Qed. + + Lemma matrix_sep m n l (P : _ → _ → _ → iProp Σ) i j : + i < m → j < n → + ([∗ list] i ↦ _ ∈ replicate m (), + [∗ list] j ↦ _ ∈ replicate n (), + P i j (l +â‚— pos n i j)) -∗ + (P i j (l +â‚— pos n i j) ∗ + (P i j (l +â‚— pos n i j) -∗ + ([∗ list] i ↦ _ ∈ replicate m (), + [∗ list] j ↦ _ ∈ replicate n (), + P i j (l +â‚— pos n i j)))). + Proof. + iIntros (Hm Hn) "Hm". + iDestruct (big_sepL_impl with "Hm []") as "Hm". + { iIntros "!>" (k x HIn). + iApply (big_sepL_lookup_acc _ _ j ()). + by apply lookup_replicate. } + simpl. + rewrite {1}(big_sepL_lookup_acc _ (replicate m ()) i ()); + [|by apply lookup_replicate]. + iDestruct "Hm" as "[[Hij Hi] Hm]". + iFrame. + iIntros "HP". + iDestruct ("Hm" with "[$Hi $HP]") as "Hm". + iApply (big_sepL_impl with "Hm"). + iIntros "!>" (k x Hin). + iIntros "[Hm Hm']". iApply "Hm'". done. + Qed. + + Lemma vpos_spec (n i j : nat) : + {{{ True }}} vpos #n #i #j {{{ RET #(pos n i j); True }}}. + Proof. + iIntros (Φ) "_ HΦ". wp_lam. wp_pures. rewrite /pos. + replace (Z.of_nat i * Z.of_nat n + Z.of_nat j)%Z with + (Z.of_nat (i * n + j)) by lia. + by iApply "HΦ". + Qed. + + Lemma matrix_get_spec m n i j v P : + i < m → j < n → + {{{ is_matrix v m n P }}} + matrix_get v #i #j + {{{ l, RET #l; P i j l ∗ (P i j l -∗ is_matrix v m n P) }}}. + Proof. + iIntros (Hm Hn Φ) "Hm HΦ". + wp_lam. + iDestruct "Hm" as (l ->) "Hm". + wp_pures. + wp_smart_apply vpos_spec; [done|]. + iIntros "_". + wp_pures. + iApply "HΦ". + iModIntro. + iDestruct (matrix_sep _ _ _ _ i j with "Hm") as "[$ Hm]"; [lia|lia|]. + iIntros "H". + iDestruct ("Hm" with "H") as "Hm". + iExists _. iSplit; [done|]. iFrame. + Qed. + +End with_Σ. -- GitLab From d4b1567add962b76e5247055772ad92f86c0293a Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Wed, 27 Mar 2024 19:49:17 +0100 Subject: [PATCH 64/81] Added adequacy proof for ring leader election --- multi_actris/examples/leader_election.v | 11 ++++++++++- 1 file changed, 10 insertions(+), 1 deletion(-) diff --git a/multi_actris/examples/leader_election.v b/multi_actris/examples/leader_election.v index e8810fc..0fa6ef5 100644 --- a/multi_actris/examples/leader_election.v +++ b/multi_actris/examples/leader_election.v @@ -1,5 +1,6 @@ -From multi_actris.channel Require Import proofmode. +From iris.heap_lang Require Import adequacy. From iris.heap_lang.lib Require Import assert. +From multi_actris.channel Require Import proofmode. (** Inspired by https://en.wikipedia.org/wiki/Chang_and_Roberts_algorithm *) @@ -195,3 +196,11 @@ Section ring_leader_election_example. Qed. End ring_leader_election_example. + +Lemma program_spec_adequate : + adequate NotStuck (program #()) ({|heap := ∅; used_proph_id := ∅|}) + (λ _ _, True). +Proof. + apply (heap_adequacy #[heapΣ; chanΣ]). + iIntros (Σ) "H". by iApply program_spec. +Qed. -- GitLab From 621aa83c54244f41796a80d93eff63be930487d7 Mon Sep 17 00:00:00 2001 From: Robbert Krebbers <mail@robbertkrebbers.nl> Date: Sun, 11 Feb 2024 18:04:51 +0100 Subject: [PATCH 65/81] Some clean up. --- _CoqProject | 1 - theories/channel/channel.v | 2 +- theories/channel/proto.v | 16 ++++++++-------- theories/utils/skip.v | 11 ----------- 4 files changed, 9 insertions(+), 21 deletions(-) delete mode 100644 theories/utils/skip.v diff --git a/_CoqProject b/_CoqProject index e46494d..843f066 100644 --- a/_CoqProject +++ b/_CoqProject @@ -8,7 +8,6 @@ # Fixing this one requires Coq 8.19 -arg -w -arg -argument-scope-delimiter -theories/utils/skip.v theories/utils/llist.v theories/utils/compare.v theories/utils/contribution.v diff --git a/theories/channel/channel.v b/theories/channel/channel.v index e8d038e..2c53a5e 100644 --- a/theories/channel/channel.v +++ b/theories/channel/channel.v @@ -25,7 +25,7 @@ From iris.heap_lang Require Export primitive_laws notation. From iris.heap_lang Require Export proofmode. From iris.heap_lang.lib Require Import spin_lock. From actris.channel Require Export proto. -From actris.utils Require Import llist skip. +From actris.utils Require Import llist. Set Default Proof Using "Type". Local Existing Instance spin_lock. diff --git a/theories/channel/proto.v b/theories/channel/proto.v index 92c4c00..9e9dfcf 100644 --- a/theories/channel/proto.v +++ b/theories/channel/proto.v @@ -54,10 +54,17 @@ From actris.channel Require Import proto_model. Set Default Proof Using "Type". Export action. +(** * Types *) +Definition iProto Σ V := proto V (iPropO Σ) (iPropO Σ). +Declare Scope proto_scope. +Delimit Scope proto_scope with proto. +Bind Scope proto_scope with iProto. +Open Scope proto. + (** * Setup of Iris's cameras *) Class protoG Σ V := protoG_authG :: - inG Σ (excl_authR (laterO (proto (leibnizO V) (iPropO Σ) (iPropO Σ)))). + inG Σ (excl_authR (laterO (iProto Σ V))). Definition protoΣ V := #[ GFunctor (authRF (optionURF (exclRF (laterOF (protoOF (leibnizO V) idOF idOF))))) @@ -65,13 +72,6 @@ Definition protoΣ V := #[ Global Instance subG_chanΣ {Σ V} : subG (protoΣ V) Σ → protoG Σ V. Proof. solve_inG. Qed. -(** * Types *) -Definition iProto Σ V := proto V (iPropO Σ) (iPropO Σ). -Declare Scope proto_scope. -Delimit Scope proto_scope with proto. -Bind Scope proto_scope with iProto. -Open Scope proto. - (** * Messages *) Section iMsg. Set Primitive Projections. diff --git a/theories/utils/skip.v b/theories/utils/skip.v deleted file mode 100644 index 76c38b3..0000000 --- a/theories/utils/skip.v +++ /dev/null @@ -1,11 +0,0 @@ -From iris.heap_lang Require Export lang. -From iris.heap_lang Require Import proofmode notation. - -Definition skipN : val := - rec: "go" "n" := if: #0 < "n" then "go" ("n" - #1) else #(). - -Lemma skipN_spec `{heapGS Σ} Φ (n : nat) : â–·^n (Φ #()) -∗ WP skipN #n {{ Φ }}. -Proof. - iIntros "HΦ". iInduction n as [|n] "IH"; wp_rec; wp_pures; [done|]. - rewrite Z.sub_1_r Nat2Z.inj_succ Z.pred_succ. by iApply "IH". -Qed. -- GitLab From d6fb18080144aae68794dc90bcedc4d60d7a5b78 Mon Sep 17 00:00:00 2001 From: Robbert Krebbers <mail@robbertkrebbers.nl> Date: Fri, 16 Feb 2024 19:59:22 +0100 Subject: [PATCH 66/81] Bump Iris (iFrame). --- coq-actris.opam | 2 +- theories/examples/list_rev.v | 3 +-- 2 files changed, 2 insertions(+), 3 deletions(-) diff --git a/coq-actris.opam b/coq-actris.opam index d79e967..0a6f3ae 100644 --- a/coq-actris.opam +++ b/coq-actris.opam @@ -8,7 +8,7 @@ bug-reports: "https://gitlab.mpi-sws.org/iris/actris/issues" dev-repo: "git+https://gitlab.mpi-sws.org/iris/actris.git" depends: [ - "coq-iris-heap-lang" { (= "dev.2024-02-02.1.b30c53e2") | (= "dev") } + "coq-iris-heap-lang" { (= "dev.2024-02-16.1.06f499e0") | (= "dev") } ] build: [make "-j%{jobs}%"] diff --git a/theories/examples/list_rev.v b/theories/examples/list_rev.v index 70d0ade..17cbdcd 100644 --- a/theories/examples/list_rev.v +++ b/theories/examples/list_rev.v @@ -39,8 +39,7 @@ Section list_rev_example. + iExists []. eauto. + iDestruct "Hl" as (v lcons) "[HIT [Hlcons Hrec]]". iDestruct ("IH" with "Hrec") as (vs) "[Hvs H]". - iExists (v::vs). iFrame. - iExists v, lcons. eauto with iFrame. + iExists (v::vs). by iFrame. - iDestruct 1 as (vs) "[Hvs HIT]". iInduction xs as [|x xs] "IH" forall (l vs). + by iDestruct (big_sepL2_nil_inv_l with "HIT") as %->. -- GitLab From 9e2761b026b9331a90c84f39f878af16ad2245c3 Mon Sep 17 00:00:00 2001 From: Robbert Krebbers <mail@robbertkrebbers.nl> Date: Tue, 12 Mar 2024 07:29:15 +0100 Subject: [PATCH 67/81] Bump Iris. --- coq-actris.opam | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/coq-actris.opam b/coq-actris.opam index 0a6f3ae..08e7857 100644 --- a/coq-actris.opam +++ b/coq-actris.opam @@ -8,7 +8,7 @@ bug-reports: "https://gitlab.mpi-sws.org/iris/actris/issues" dev-repo: "git+https://gitlab.mpi-sws.org/iris/actris.git" depends: [ - "coq-iris-heap-lang" { (= "dev.2024-02-16.1.06f499e0") | (= "dev") } + "coq-iris-heap-lang" { (= "dev.2024-03-12.0.c1e15cdc") | (= "dev") } ] build: [make "-j%{jobs}%"] -- GitLab From 9c0c5bb3e562feb65b63580fa2d9bcfa76dde350 Mon Sep 17 00:00:00 2001 From: Robbert Krebbers <mail@robbertkrebbers.nl> Date: Tue, 19 Mar 2024 08:20:59 +0100 Subject: [PATCH 68/81] Remove old comment, minor tweak. --- theories/channel/proto_model.v | 8 +++----- 1 file changed, 3 insertions(+), 5 deletions(-) diff --git a/theories/channel/proto_model.v b/theories/channel/proto_model.v index c6f6855..7305940 100644 --- a/theories/channel/proto_model.v +++ b/theories/channel/proto_model.v @@ -123,8 +123,6 @@ Lemma proto_end_message_equivI `{!BiInternalEq SPROP} {V} `{!Cofe PROPn, !Cofe P proto_end ≡ proto_message (V:=V) (PROPn:=PROPn) (PROP:=PROP) a m ⊢@{SPROP} False. Proof. by rewrite internal_eq_sym proto_message_end_equivI. Qed. -(** The eliminator [proto_elim x f p] is only well-behaved if the function [f] -is contractive *) Definition proto_elim {V} `{!Cofe PROPn, !Cofe PROP} {A} (x : A) (f : action → (V → laterO (proto V PROP PROPn) -n> PROP) → A) (p : proto V PROPn PROP) : A := @@ -151,11 +149,11 @@ Proof. by destruct (proto_unfold (proto_fold None)) as [[??]|] eqn:E; inversion Hfold. Qed. Lemma proto_elim_message {V} `{!Cofe PROPn, !Cofe PROP} {A : ofe} - (x : A) (f : action → (V → laterO (proto V PROP PROPn) -n> PROP) → A) - `{Hf : ∀ a, Proper (pointwise_relation _ (≡) ==> (≡)) (f a)} a m : + (x : A) (f : action → (V → laterO (proto V PROP PROPn) -n> PROP) → A) a m : + (∀ a, Proper (pointwise_relation _ (≡) ==> (≡)) (f a)) → proto_elim x f (proto_message a m) ≡ f a m. Proof. - rewrite /proto_elim /proto_message /=. + intros. rewrite /proto_elim /proto_message /=. pose proof (proto_unfold_fold (V:=V) (PROPn:=PROPn) (PROP:=PROP) (Some (a, m))) as Hfold. destruct (proto_unfold (proto_fold (Some (a, m)))) -- GitLab From 9b32c5b01575662249c04bbadef4b6bf3749270c Mon Sep 17 00:00:00 2001 From: Robbert Krebbers <mail@robbertkrebbers.nl> Date: Wed, 3 Apr 2024 14:45:34 +0200 Subject: [PATCH 69/81] Bump Iris. --- multi_actris/channel/channel.v | 8 ++++---- multi_actris/channel/proto.v | 7 +++---- 2 files changed, 7 insertions(+), 8 deletions(-) diff --git a/multi_actris/channel/channel.v b/multi_actris/channel/channel.v index 8573236..7e68d81 100644 --- a/multi_actris/channel/channel.v +++ b/multi_actris/channel/channel.v @@ -373,10 +373,10 @@ Section channel. rewrite iMsg_base_eq. iDestruct (iMsg_texist_exist with "Hm") as (x <-) "[Hp HP]". wp_pures. - iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]";[apply excl_auth_update|]. - iModIntro. iApply "HΦ". - iFrame. iExists _, _, _, _, _, _. iSplit; [done|]. - iRewrite "Hp". iFrame "#∗". iApply iProto_le_refl. + iMod (own_update_2 with "Hâ— Hâ—¯") as "[Hâ— Hâ—¯]"; + [apply (excl_auth_update _ _ (Next p'''))|]. + iModIntro. iApply "HΦ". rewrite /iProto_pointsto_def. iFrame "IH Hls ∗". + iSplit; [done|]. iRewrite "Hp". iApply iProto_le_refl. Qed. End channel. diff --git a/multi_actris/channel/proto.v b/multi_actris/channel/proto.v index dea40ec..fdd225f 100644 --- a/multi_actris/channel/proto.v +++ b/multi_actris/channel/proto.v @@ -912,8 +912,7 @@ Section proto. iSplitL "Hauth". - rewrite /iProto_own_auth. rewrite map_seq_snoc. simpl. done. - - iSplit; [|done]. - iExists _. iFrame. by iApply iProto_le_refl. + - by iApply iProto_le_refl. Qed. Lemma list_lookup_Some_le (ps : list $ iProto Σ V) (i : nat) (p1 : iProto Σ V) : @@ -1271,7 +1270,7 @@ Section proto. Proof. iIntros "Hconsistent". iMod iProto_own_auth_alloc as (γ) "[Hauth Hfrags]". - iExists γ. iFrame. iExists _. by iFrame. + iExists γ. by iFrame. Qed. Lemma iProto_step γ ps_dom i j m1 m2 p1 v : @@ -1416,7 +1415,7 @@ Section proto. End proto. -Typeclasses Opaque iProto_ctx iProto_own. +Global Typeclasses Opaque iProto_ctx iProto_own. Global Hint Extern 0 (environments.envs_entails _ (?x ⊑ ?y)) => first [is_evar x; fail 1 | is_evar y; fail 1|iApply iProto_le_refl] : core. -- GitLab From 8e968ab284f344b0b180b19fd181579e111b9e08 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Wed, 3 Apr 2024 16:49:51 +0200 Subject: [PATCH 70/81] Added delegation variant of leader election --- multi_actris/examples/leader_election_del.v | 277 ++++++++++++++++++++ 1 file changed, 277 insertions(+) create mode 100644 multi_actris/examples/leader_election_del.v diff --git a/multi_actris/examples/leader_election_del.v b/multi_actris/examples/leader_election_del.v new file mode 100644 index 0000000..7c40b66 --- /dev/null +++ b/multi_actris/examples/leader_election_del.v @@ -0,0 +1,277 @@ +From iris.heap_lang Require Import adequacy. +From iris.heap_lang.lib Require Import assert. +From multi_actris.channel Require Import proofmode. + +(** Inspired by https://en.wikipedia.org/wiki/Chang_and_Roberts_algorithm *) + +Definition process : val := + rec: "go" "c" "idl" "id" "idr" "isp" := + if: recv "c" "idr" + then let: "id'" := recv "c" "idr" in + if: "id" < "id'" (** Case 1 *) + then send "c" "idl" #true;; send "c" "idl" "id'";; + "go" "c" "idl" "id" "idr" #true + else if: "id" = "id'" (** Case 4 *) + then send "c" "idl" #false;; send "c" "idl" "id";; + "go" "c" "idl" "id" "idr" #false + else if: "isp" (** Case 3 *) + then "go" "c" "idl" "id" "idr" "isp" (** Case 2 *) + else send "c" "idl" #true;; send "c" "idl" "id";; + "go" "c" "idl" "id" "idr" #true + else let: "id'" := recv "c" "idr" in + if: "id" = "id'" then "id'" + else send "c" "idl" #false;; send "c" "idl" "id'";; "id'". + +Definition init : val := + λ: "c" "idl" "id" "idr", + (* Notice missing leader *) + send "c" "idl" #true;; send "c" "idl" "id";; + process "c" "idl" "id" "idr" #true. + +Definition forward : val := + λ: "c" "idl" "id" "idr" "id_max", + if: "id" = "id_max" then + let: "cs'" := new_chan #2 in + let: "c0" := get_chan "cs'" #0 in + let: "c1" := get_chan "cs'" #1 in + send "c1" #0 "id_max";; + send "c" "idl" "c1";; + Fork ((rec: "f" <> := + let: "id'" := recv "c0" #1 in + assert: ("id'" = "id_max");; "f" #()) #());; + recv "c" "idr";; #() + else + let: "c1" := recv "c" "idr" in + send "c1" #0 "id_max";; + send "c" "idl" "c1". + +Definition program : val := + λ: <>, + let: "cs" := new_chan #4 in + let: "c0" := get_chan "cs" #0 in + let: "c1" := get_chan "cs" #1 in + let: "c2" := get_chan "cs" #2 in + let: "c3" := get_chan "cs" #3 in + Fork (let: "id_max" := process "c1" #0 #1 #2 #false in + forward "c1" #0 #1 #2 "id_max");; + Fork (let: "id_max" := process "c2" #1 #2 #3 #false in + forward "c2" #1 #2 #3 "id_max");; + Fork (let: "id_max" := process "c3" #2 #3 #0 #false in + forward "c3" #2 #3 #0 "id_max");; + let: "id_max" := init "c0" #3 #0 #1 in + forward "c0" #3 #0 #1 "id_max". + +Notation iProto_choice a p1 p2 := + (<a @ (b : bool)> MSG #b; if b then p1 else p2)%proto. + +Section ring_leader_election_example. + Context `{!heapGS Σ, chanG Σ, spawnG Σ, mono_natG Σ}. + + Definition my_recv_prot (il i ir : nat) (p : nat → iProto Σ) + (rec : bool -d> iProto Σ) : bool -d> iProto Σ := + λ (isp:bool), + iProto_choice (Recv,ir) + (<(Recv,ir) @ (i':nat)> MSG #i' ; + if bool_decide (i < i') + then <(Send,il)> MSG #true ; <(Send,il)> MSG #i' ; rec true + else if bool_decide (i = i') + then <(Send,il)> MSG #false ; <(Send, il)> MSG #i ; rec false + else if isp then rec isp + else <(Send,il)> MSG #true ; <(Send,il)> MSG #i ; rec true)%proto + (<(Recv,ir) @ (i':nat)> MSG #i' ; + if (bool_decide (i = i')) then p i + else <(Send,il)> MSG #false; <(Send,il)> MSG #i'; p i')%proto. + + Instance rle_prot_aux_contractive il i ir p : Contractive (my_recv_prot il i ir p). + Proof. rewrite /my_recv_prot. solve_proto_contractive. Qed. + Definition rle_prot il i ir p := fixpoint (my_recv_prot il i ir p). + Instance rle_prot_unfold il i ir isp p : + ProtoUnfold (rle_prot il i ir p isp) (my_recv_prot il i ir p (rle_prot il i ir p) isp). + Proof. apply proto_unfold_eq, (fixpoint_unfold (my_recv_prot il i ir p)). Qed. + Lemma rle_prot_unfold' il i ir isp p : + (rle_prot il i ir p isp) ≡ + (my_recv_prot il i ir p (rle_prot il i ir p) isp). + Proof. apply (fixpoint_unfold (my_recv_prot il i ir p)). Qed. + + Definition rle_preprot (il i ir : nat) p : iProto Σ := + (<(Send, il)> MSG #true; <(Send, il)> MSG #i ; + rle_prot il i ir p true)%proto. + + Lemma process_spec il i ir p c (isp:bool) : + {{{ c ↣ (rle_prot il i ir p isp) }}} + process c #il #i #ir #isp + {{{ i', RET #i'; c ↣ p i' }}}. + Proof. + iIntros (Φ) "Hc HΦ". + iLöb as "IH" forall (Φ isp). + wp_lam. wp_recv (b) as "_". + destruct b. + - wp_pures. wp_recv (i') as "_". + wp_pures. + case_bool_decide as Hlt. + { case_bool_decide; [|lia]. + wp_pures. wp_send with "[//]". + wp_send with "[//]". wp_pures. + iApply ("IH" with "Hc HΦ"). } + case_bool_decide as Hlt2. + { case_bool_decide; [lia|]. + wp_pures. case_bool_decide; [|simplify_eq;lia]. + wp_send with "[//]". + wp_send with "[//]". wp_pures. + iApply ("IH" with "Hc HΦ"). } + case_bool_decide; [lia|]. + wp_pures. + case_bool_decide; [simplify_eq;lia|]. + wp_pures. + destruct isp. + { wp_pures. iApply ("IH" with "Hc HΦ"). } + wp_pures. + wp_send with "[//]". + wp_send with "[//]". + wp_pures. iApply ("IH" with "Hc HΦ"). + - wp_pures. + wp_recv (id') as "_". wp_pures. + case_bool_decide as Hlt. + { case_bool_decide; [|simplify_eq;lia]. + wp_pures. subst. by iApply "HΦ". } + case_bool_decide; [simplify_eq;lia|]. + wp_send with "[//]". wp_send with "[//]". wp_pures. by iApply "HΦ". + Qed. + + Lemma init_spec c il i ir p : + {{{ c ↣ rle_preprot il i ir p }}} + init c #il #i #ir + {{{ res, RET #res; c ↣ p res }}}. + Proof. + iIntros (Φ) "Hc HΦ". wp_lam. wp_send with "[//]". wp_send with "[//]". + wp_pures. by iApply (process_spec with "Hc HΦ"). + Qed. + + Definition forward_prot (p : iProto Σ) (il i ir i_max : nat) : iProto Σ := + if bool_decide (i = i_max) then + (<(Send,il) @ (c:val)> MSG c {{ c ↣ p }} ; <(Recv,ir)> MSG c {{ c ↣ p }}; END)%proto + else + (<(Recv,ir) @ (c:val)> MSG c {{ c ↣ p }} ; <(Send,il)> MSG c {{ c ↣ p }}; END)%proto. + + Definition relay_send_aux (id : nat) (rec : iProto Σ) : iProto Σ := + (<(Send,0)> MSG #id ; rec)%proto. + Instance relay_send_aux_contractive i : Contractive (relay_send_aux i). + Proof. solve_proto_contractive. Qed. + Definition relay_send_prot i := fixpoint (relay_send_aux i). + Instance relay_send_prot_unfold i : + ProtoUnfold (relay_send_prot i) (relay_send_aux i (relay_send_prot i)). + Proof. apply proto_unfold_eq, (fixpoint_unfold (relay_send_aux i)). Qed. + Lemma relay_send_prot_unfold' i : + (relay_send_prot i) ≡ + (relay_send_aux i (relay_send_prot i)). + Proof. apply (fixpoint_unfold (relay_send_aux i)). Qed. + + Definition relay_recv_aux (id : nat) (rec : iProto Σ) : iProto Σ := + (<(Recv,1)> MSG #id ; rec)%proto. + Instance relay_recv_aux_contractive i : Contractive (relay_recv_aux i). + Proof. solve_proto_contractive. Qed. + Definition relay_recv_prot i := fixpoint (relay_recv_aux i). + Instance relay_recv_prot_unfold i : + ProtoUnfold (relay_recv_prot i) (relay_recv_aux i (relay_recv_prot i)). + Proof. apply proto_unfold_eq, (fixpoint_unfold (relay_recv_aux i)). Qed. + Lemma relay_recv_prot_unfold' i : + (relay_recv_prot i) ≡ + (relay_recv_aux i (relay_recv_prot i)). + Proof. apply (fixpoint_unfold (relay_recv_aux i)). Qed. + + + Definition prot_pool : list (iProto Σ) := + [rle_preprot 3 0 1 (λ id_max, forward_prot (relay_send_prot id_max) 3 0 1 id_max); + rle_prot 0 1 2 (λ id_max, forward_prot (relay_send_prot id_max) 0 1 2 id_max) false; + rle_prot 1 2 3 (λ id_max, forward_prot (relay_send_prot id_max) 1 2 3 id_max) false; + rle_prot 2 3 0 (λ id_max, forward_prot (relay_send_prot id_max) 2 3 0 id_max) false]. + + Lemma prot_pool_consistent : ⊢ iProto_consistent prot_pool. + Proof. + rewrite /prot_pool /rle_preprot. + rewrite !rle_prot_unfold'. + iProto_consistent_take_steps. + case_bool_decide; try lia. + rewrite !rle_prot_unfold'. + iProto_consistent_take_steps. + case_bool_decide; try lia. + rewrite !rle_prot_unfold'. + iProto_consistent_take_steps. + case_bool_decide; try lia. + rewrite !rle_prot_unfold'. + iProto_consistent_take_steps. + Qed. + + Definition prot_pool' (i:nat) : list (iProto Σ) := + [relay_recv_prot i; + relay_send_prot i]. + + Lemma prot_pool_consistent' i : ⊢ iProto_consistent (prot_pool' i). + Proof. + rewrite /prot_pool'. + iLöb as "IH". + iEval (rewrite relay_recv_prot_unfold'). + iEval (rewrite relay_send_prot_unfold'). + iProto_consistent_take_steps. + done. + Qed. + + Lemma forward_spec c il i ir i_max : + {{{ c ↣ forward_prot (relay_send_prot i_max) il i ir i_max }}} + forward c #il #i #ir #i_max + {{{ RET #(); True }}}. + Proof. + iIntros (Φ) "Hc HΦ". wp_lam. + rewrite /forward_prot. + wp_pures. case_bool_decide. + - simplify_eq. wp_pures. + case_bool_decide; [|simplify_eq;lia]. + wp_new_chan (prot_pool' i_max) with (prot_pool_consistent' i_max) + as (c0 c1) "Hc0" "Hc1". + wp_send with "[//]". + wp_send with "[Hc1//]". + wp_smart_apply (wp_fork with "[Hc0]"). + { iIntros "!>". + wp_pures. + iLöb as "IH". + wp_recv as "_". + wp_smart_apply wp_assert. + wp_pures. iModIntro. iSplit; [iPureIntro; f_equiv; by case_bool_decide|]. + iIntros "!>". wp_pures. by iApply "IH". } + wp_recv as "Hc'". wp_pures. by iApply "HΦ". + - case_bool_decide; [simplify_eq;lia|]. + wp_pures. + wp_recv (c') as "Hc'". + wp_send with "[//]". + wp_send with "[Hc'//]". + by iApply "HΦ". + Qed. + + Lemma program_spec : + {{{ True }}} program #() {{{ RET #(); True }}}. + Proof. + iIntros (Φ) "_ HΦ". wp_lam. + wp_new_chan prot_pool with prot_pool_consistent + as (c0 c1 c2 c3) "Hc0" "Hc1" "Hc2" "Hc3". + wp_smart_apply (wp_fork with "[Hc1]"). + { iIntros "!>". wp_smart_apply (process_spec with "Hc1"). + iIntros (i') "Hc1". by wp_smart_apply (forward_spec with "Hc1"). } + wp_smart_apply (wp_fork with "[Hc2]"). + { iIntros "!>". wp_smart_apply (process_spec with "Hc2"). + iIntros (i') "Hc2". by wp_smart_apply (forward_spec with "Hc2"). } + wp_smart_apply (wp_fork with "[Hc3]"). + { iIntros "!>". wp_smart_apply (process_spec with "Hc3"). + iIntros (i') "Hc3". by wp_smart_apply (forward_spec with "Hc3"). } + wp_smart_apply (init_spec with "Hc0"). + iIntros (i') "Hc0". by wp_smart_apply (forward_spec with "Hc0"). + Qed. + +End ring_leader_election_example. + +Lemma program_spec_adequate : + adequate NotStuck (program #()) ({|heap := ∅; used_proph_id := ∅|}) + (λ _ _, True). +Proof. + apply (heap_adequacy #[heapΣ; chanΣ]). + iIntros (Σ) "H". by iApply program_spec. +Qed. -- GitLab From 8b313fcad5a3ffbb7ab413d96455e1725c920c0f Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Thu, 4 Apr 2024 16:15:44 +0200 Subject: [PATCH 71/81] Added implementation for two-buyer and refactored --- _CoqProject | 1 + multi_actris/channel/proofmode.v | 4 +- multi_actris/examples/basics.v | 85 ------------------- multi_actris/examples/two_buyer.v | 133 ++++++++++++++++++++++++++++++ 4 files changed, 136 insertions(+), 87 deletions(-) create mode 100644 multi_actris/examples/two_buyer.v diff --git a/_CoqProject b/_CoqProject index 843f066..f187ca7 100644 --- a/_CoqProject +++ b/_CoqProject @@ -61,4 +61,5 @@ multi_actris/channel/proto.v multi_actris/channel/channel.v multi_actris/channel/proofmode.v multi_actris/examples/basics.v +multi_actris/examples/two_buyer.v multi_actris/examples/leader_election.v diff --git a/multi_actris/channel/proofmode.v b/multi_actris/channel/proofmode.v index 5f9d7b8..319de3c 100644 --- a/multi_actris/channel/proofmode.v +++ b/multi_actris/channel/proofmode.v @@ -396,11 +396,11 @@ Ltac ltac1_list_iter2 tac l1 l2 := Tactic Notation "wp_new_chan" constr(prot) "as" "(" simple_intropattern_list(xs) ")" constr_list(pats) := wp_smart_apply (new_chan_spec prot); - [set_solver| |iIntros (?) "?"]; + [set_solver| |iIntros (_cs) "?"]; [|ltac1_list_iter2 ltac:(fun x y => wp_get_chan (x) y) xs pats]. Tactic Notation "wp_new_chan" constr(prot) "with" constr(lem) "as" "(" simple_intropattern_list(xs) ")" constr_list(pats) := wp_smart_apply (new_chan_spec prot); - [set_solver|by iApply lem|iIntros (?) "?"]; + [set_solver|by iApply lem|iIntros (_cs) "?"]; ltac1_list_iter2 ltac:(fun x y => wp_get_chan (x) y) xs pats. diff --git a/multi_actris/examples/basics.v b/multi_actris/examples/basics.v index 36c500f..82a0173 100644 --- a/multi_actris/examples/basics.v +++ b/multi_actris/examples/basics.v @@ -294,91 +294,6 @@ Section parallel. End parallel. -Section two_buyer. - Context `{!heapGS Σ}. - - Definition two_buyer_prot : list (iProto Σ) := - [(<(Send, 1) @ (title:Z)> MSG #title ; - <(Recv, 1) @ (quote:Z)> MSG #quote ; - <(Send, 2) @ (contrib:Z)> MSG #contrib ; END)%proto; - (<(Recv, 0) @ (title:Z)> MSG #title ; - <(Send, 0) @ (quote:Z)> MSG #quote ; - <(Send, 2)> MSG #quote ; - <(Recv, 2) @ (b:bool)> MSG #b ; - if b then - <(Recv, 2) @ (address:Z)> MSG #address ; - <(Send, 2) @ (date:Z)> MSG #date ; END - else END)%proto; - (<(Recv, 1) @ (quote:Z)> MSG #quote ; - <(Recv, 0) @ (contrib:Z)> MSG #contrib ; - if bool_decide (contrib >= quote/2)%Z then - <(Send, 1)> MSG #true ; - <(Send, 1) @ (address:Z)> MSG #address ; - <(Recv, 1) @ (date:Z)> MSG #date ; END - else - <(Send, 1)> MSG #false ; END)%proto]. - - Lemma two_buyer_prot_consistent : - ⊢ iProto_consistent two_buyer_prot. - Proof. - rewrite /two_buyer_prot. iProto_consistent_take_steps. - case_bool_decide; iProto_consistent_take_steps. - Qed. - -End two_buyer. - -Section two_buyer_ref. - Context `{!heapGS Σ}. - - Definition two_buyer_ref_b1_prot : iProto Σ := - (<(Send, 1) @ (title:Z)> MSG #title ; - <(Recv, 1) @ (quote:Z)> MSG #quote ; - <(Send, 2) @ (l : loc) (amount:Z) (contrib:Z)> - MSG (#l,#contrib) {{ l ↦ #amount }} ; - <(Recv, 2) @ (b : bool)> - MSG #b {{ l ↦ #(if b then amount - contrib else amount) }}; - END)%proto. - - Definition two_buyer_ref_s_prot : iProto Σ := - (<(Recv, 0) @ (title:Z)> MSG #title ; - <(Send, 0) @ (quote:Z)> MSG #quote ; - <(Send, 2)> MSG #quote ; - <(Recv, 2) @ (b:bool)> MSG #b ; - if b then - <(Recv, 2) @ (l2 : loc) (amount2:Z) (address:Z)> - MSG (#l2,#address) {{ l2 ↦ #amount2 }} ; - <(Send, 2) @ (date : Z)> MSG #date {{ l2 ↦ #(amount2-quote) }}; END - else END)%proto. - - Definition two_buyer_ref_b2_prot : iProto Σ := - (<(Recv, 1) @ (quote:Z)> MSG #quote ; - <(Recv, 0) @ (l1 : loc) (amount1:Z) (contrib:Z)> - MSG (#l1,#contrib) {{ l1 ↦ #amount1 }}; - <(Send, 0) @ (b : bool)> - MSG #b {{ l1 ↦ #(if b then amount1 - contrib else amount1) }}; - <(Send, 1)> MSG #b; - if b then - <(Send, 1) @ (l2 : loc) (amount2:Z) (address:Z)> - MSG (#l2,#address) {{ l2 ↦ #amount2 }} ; - <(Recv, 1) @ (date : Z)> MSG #date {{ l2 ↦ #(amount2-quote) }}; - END - else END)%proto. - - Definition two_buyer_ref_prot : list (iProto Σ) := - [two_buyer_ref_b1_prot; - two_buyer_ref_s_prot; - two_buyer_ref_b2_prot]. - - (* TODO: Anonymous variable in this is unsatisfactory *) - Lemma two_buyer_ref_prot_consistent : - ⊢ iProto_consistent two_buyer_ref_prot. - Proof. - rewrite /two_buyer_prot. iProto_consistent_take_steps. - iNext. destruct x4; iProto_consistent_take_steps. - Qed. - -End two_buyer_ref. - Section forwarder. Context `{!heapGS Σ}. diff --git a/multi_actris/examples/two_buyer.v b/multi_actris/examples/two_buyer.v new file mode 100644 index 0000000..15eef77 --- /dev/null +++ b/multi_actris/examples/two_buyer.v @@ -0,0 +1,133 @@ +From multi_actris.channel Require Import proofmode. +Set Default Proof Using "Type". + +Definition buyer1_prog : val := + λ: "cb1" "s" "b2", + send "cb1" "s" #0;; + let: "quote" := recv "cb1" "s" in + send "cb1" "b2" ("quote" `rem` #2);; + recv "cb1" "b2";; #(). + +Definition seller_prog : val := + λ: "cs" "b1" "b2", + let: "title" := recv "cs" "b1" in + send "cs" "b1" #42;; send "cs" "b2" #42;; + let: "b" := recv "cs" "b2" in + if: "b" then + send "cs" "b2" #0 + else #(). + +Definition buyer2_prog : val := + λ: "cb2" "b1" "s", + let: "quote" := recv "cb2" "s" in + let: "contrib" := recv "cb2" "b1" in + if: ("quote" - "contrib" < #100) + then send "cb2" "s" #true;; send "cb2" "b1" #true;; + recv "cb2" "s" + else #(). + +Definition two_buyer_prog : val := + λ: <>, + let: "cs" := new_chan #3 in + let: "b1" := get_chan "cs" #0 in + let: "s" := get_chan "cs" #1 in + let: "b2" := get_chan "cs" #2 in + Fork (seller_prog "s" #0 #2);; + Fork (buyer2_prog "b2" #0 #1);; + buyer1_prog "b1" #1 #2. + +Section two_buyer. + Context `{!heapGS Σ, !chanG Σ}. + + Definition buyer1_prot s b2 : iProto Σ := + (<(Send, s) @ (title:Z)> MSG #title ; + <(Recv, s) @ (quote:Z)> MSG #quote ; + <(Send, b2) @ (contrib:Z)> MSG #contrib ; + <(Recv, b2) @ (b:bool)> MSG #b; END)%proto. + + Lemma buyer1_spec c s b2 : + {{{ c ↣ buyer1_prot s b2 }}} + buyer1_prog c #s #b2 + {{{ RET #(); True }}}. + Proof. + iIntros (Φ) "Hc HΦ". wp_lam. + wp_send with "[//]". + wp_recv (quote) as "_". + wp_send with "[//]". + wp_recv (b) as "_". wp_pures. + by iApply "HΦ". + Qed. + + Definition seller_prot b1 b2 : iProto Σ := + (<(Recv, b1) @ (title:Z)> MSG #title ; + <(Send, b1) @ (quote:Z)> MSG #quote ; + <(Send, b2)> MSG #quote ; + <(Recv, b2) @ (b:bool)> MSG #b ; + if b then + <(Send, b2) @ (date:Z)> MSG #date ; END + else END)%proto. + + Lemma seller_spec c b1 b2 : + {{{ c ↣ seller_prot b1 b2 }}} + seller_prog c #b1 #b2 + {{{ b, RET #b; True }}}. + Proof. + iIntros (Φ) "Hc HΦ". wp_lam. + wp_recv (title) as "_". + wp_send with "[//]". + wp_send with "[//]". + wp_recv (b) as "_". + destruct b. + - wp_pures. wp_send with "[//]". by iApply "HΦ". + - wp_pures. by iApply "HΦ". + Qed. + + Definition buyer2_prot b1 s : iProto Σ := + (<(Recv, s) @ (quote:Z)> MSG #quote ; + <(Recv, b1) @ (contrib:Z)> MSG #contrib ; + <(Send, s) @ (b:bool)> MSG #b ; + <(Send, b1)> MSG #b ; + if b then <(Recv, s) @ (date:Z)> MSG #date ; END + else <(Send, s)> MSG #false ; END)%proto. + + Lemma buyer2_spec c b1 s : + {{{ c ↣ buyer2_prot b1 s }}} + buyer2_prog c #b1 #s + {{{ b, RET #b; True }}}. + Proof. + iIntros (Φ) "Hc HΦ". wp_lam. + wp_recv (quote) as "_". + wp_recv (contrib) as "_". + wp_pures. case_bool_decide. + - wp_send with "[//]". wp_send with "[//]". wp_recv (date) as "_". + by iApply "HΦ". + - wp_pures. by iApply "HΦ". + Qed. + + Definition two_buyer_prot : list (iProto Σ) := + [buyer1_prot 1 2 ; seller_prot 0 2; buyer2_prot 0 1]. + + Lemma two_buyer_prot_consistent : + ⊢ iProto_consistent two_buyer_prot. + Proof. + rewrite /two_buyer_prot. iProto_consistent_take_steps. + destruct x2; iProto_consistent_take_steps. + Qed. + + Lemma two_buyer_spec : + {{{ True }}} + two_buyer_prog #() + {{{ RET #(); True }}}. + Proof using chanG0 heapGS0 Σ. + iIntros (Φ) "Hc HΦ". wp_lam. + wp_new_chan two_buyer_prot with two_buyer_prot_consistent + as (???) "Hcb1" "Hcs" "Hcb2". + wp_smart_apply (wp_fork with "[Hcs]"). + { by iApply (seller_spec with "Hcs"). } + wp_smart_apply (wp_fork with "[Hcb2]"). + { by iApply (buyer2_spec with "Hcb2"). } + wp_smart_apply (buyer1_spec with "Hcb1"). + by iApply "HΦ". + Qed. + +End two_buyer. -- GitLab From 32e16a45718a540c15cac881251ed2d84f25d5c5 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Thu, 4 Apr 2024 17:22:38 +0200 Subject: [PATCH 72/81] Added three buyer example --- _CoqProject | 1 + multi_actris/channel/proofmode.v | 4 +- multi_actris/examples/three_buyer.v | 198 ++++++++++++++++++++++++++++ 3 files changed, 201 insertions(+), 2 deletions(-) create mode 100644 multi_actris/examples/three_buyer.v diff --git a/_CoqProject b/_CoqProject index f187ca7..f25af83 100644 --- a/_CoqProject +++ b/_CoqProject @@ -62,4 +62,5 @@ multi_actris/channel/channel.v multi_actris/channel/proofmode.v multi_actris/examples/basics.v multi_actris/examples/two_buyer.v +multi_actris/examples/three_buyer.v multi_actris/examples/leader_election.v diff --git a/multi_actris/channel/proofmode.v b/multi_actris/channel/proofmode.v index 319de3c..5f9d7b8 100644 --- a/multi_actris/channel/proofmode.v +++ b/multi_actris/channel/proofmode.v @@ -396,11 +396,11 @@ Ltac ltac1_list_iter2 tac l1 l2 := Tactic Notation "wp_new_chan" constr(prot) "as" "(" simple_intropattern_list(xs) ")" constr_list(pats) := wp_smart_apply (new_chan_spec prot); - [set_solver| |iIntros (_cs) "?"]; + [set_solver| |iIntros (?) "?"]; [|ltac1_list_iter2 ltac:(fun x y => wp_get_chan (x) y) xs pats]. Tactic Notation "wp_new_chan" constr(prot) "with" constr(lem) "as" "(" simple_intropattern_list(xs) ")" constr_list(pats) := wp_smart_apply (new_chan_spec prot); - [set_solver|by iApply lem|iIntros (_cs) "?"]; + [set_solver|by iApply lem|iIntros (?) "?"]; ltac1_list_iter2 ltac:(fun x y => wp_get_chan (x) y) xs pats. diff --git a/multi_actris/examples/three_buyer.v b/multi_actris/examples/three_buyer.v new file mode 100644 index 0000000..aa4c6e8 --- /dev/null +++ b/multi_actris/examples/three_buyer.v @@ -0,0 +1,198 @@ +From multi_actris.channel Require Import proofmode. +Set Default Proof Using "Type". + +Definition buyer1_prog : val := + λ: "cb1" "s" "b2", + send "cb1" "s" #0;; + let: "quote" := recv "cb1" "s" in + send "cb1" "b2" ("quote" `rem` #2);; + recv "cb1" "b2";; #(). + +Definition seller_prog : val := + λ: "cs" "b1" "b2", + let: "title" := recv "cs" "b1" in + send "cs" "b1" #42;; send "cs" "b2" #42;; + let: "b" := recv "cs" "b2" in + if: "b" then + send "cs" "b2" #0 + else #(). + +Definition buyer2_prog : val := + λ: "cb2" "b1" "s" "cb2'" "b3", + let: "quote" := recv "cb2" "s" in + let: "contrib" := recv "cb2" "b1" in + if: ("quote" - "contrib" < #100) + then send "cb2" "s" #true;; send "cb2" "b1" #true;; + recv "cb2" "s" + else send "cb2'" "b3" "quote";; + send "cb2'" "b3" ("contrib" + #100);; + send "cb2'" "b3" "cb2". + +Definition buyer3_prog : val := + λ: "cb3" "b1" "b2'" "s", + let: "quote" := recv "cb3" "b2'" in + let: "contrib" := recv "cb3" "b2'" in + let: "cb2" := recv "cb3" "b2'" in + if: ("quote" - "contrib" < #100) + then send "cb2" "s" #true;; send "cb2" "b1" #true;; + recv "cb2" "s" + else #(). + +Definition three_buyer_prog : val := + λ: <>, + let: "cs" := new_chan #3 in + let: "b1" := get_chan "cs" #0 in + let: "s" := get_chan "cs" #1 in + let: "b2" := get_chan "cs" #2 in + let: "cs'" := new_chan #2 in + let: "b2'" := get_chan "cs'" #0 in + let: "b3" := get_chan "cs'" #1 in + Fork (seller_prog "s" #0 #2);; + Fork (buyer2_prog "b2" #0 #1 "b2'" #1);; + Fork (buyer3_prog "b3" #0 #1 #0);; + buyer1_prog "b1" #1 #2. + +Section three_buyer. + Context `{!heapGS Σ, !chanG Σ}. + + Definition buyer1_prot s b2 : iProto Σ := + (<(Send, s) @ (title:Z)> MSG #title ; + <(Recv, s) @ (quote:Z)> MSG #quote ; + <(Send, b2) @ (contrib:Z)> MSG #contrib ; + <(Recv, b2) @ (b:bool)> MSG #b; END)%proto. + + Lemma buyer1_spec c s b2 : + {{{ c ↣ buyer1_prot s b2 }}} + buyer1_prog c #s #b2 + {{{ RET #(); True }}}. + Proof. + iIntros (Φ) "Hc HΦ". wp_lam. + wp_send with "[//]". + wp_recv (quote) as "_". + wp_send with "[//]". + wp_recv (b) as "_". wp_pures. + by iApply "HΦ". + Qed. + + Definition seller_prot b1 b2 : iProto Σ := + (<(Recv, b1) @ (title:Z)> MSG #title ; + <(Send, b1) @ (quote:Z)> MSG #quote ; + <(Send, b2)> MSG #quote ; + <(Recv, b2) @ (b:bool)> MSG #b ; + if b then + <(Send, b2) @ (date:Z)> MSG #date ; END + else END)%proto. + + Lemma seller_spec c b1 b2 : + {{{ c ↣ seller_prot b1 b2 }}} + seller_prog c #b1 #b2 + {{{ b, RET #b; True }}}. + Proof. + iIntros (Φ) "Hc HΦ". wp_lam. + wp_recv (title) as "_". + wp_send with "[//]". + wp_send with "[//]". + wp_recv (b) as "_". + destruct b. + - wp_pures. wp_send with "[//]". by iApply "HΦ". + - wp_pures. by iApply "HΦ". + Qed. + + Definition buyer2_prot b1 s : iProto Σ := + (<(Recv, s) @ (quote:Z)> MSG #quote ; + <(Recv, b1) @ (contrib:Z)> MSG #contrib ; + <(Send, s) @ (b:bool)> MSG #b ; + <(Send, b1)> MSG #b ; + if b then <(Recv, s) @ (date:Z)> MSG #date ; END + else <(Send, s)> MSG #false ; END)%proto. + + Definition buyer2_prot' b3 b1 s : iProto Σ := + (<(Send, b3) @ (quote:Z)> MSG #quote ; + <(Send, b3) @ (contrib:Z)> MSG #contrib ; + <(Send, b3) @ (c:val)> MSG c + {{ c ↣ (<(Send, s) @ (b:bool)> MSG #b ; + <(Send, b1)> MSG #b ; + if b then <(Recv, s) @ (date:Z)> MSG #date ; END + else <(Send, s)> MSG #false ; END)%proto }} ; + END)%proto. + + Lemma buyer2_spec c b1 s c' b3 : + {{{ c ↣ buyer2_prot b1 s ∗ c' ↣ buyer2_prot' b3 b1 s }}} + buyer2_prog c #b1 #s c' #b3 + {{{ b, RET #b; True }}}. + Proof. + iIntros (Φ) "[Hc Hc'] HΦ". wp_lam. + wp_recv (quote) as "_". + wp_recv (contrib) as "_". + wp_pures. case_bool_decide. + - wp_send with "[//]". wp_send with "[//]". wp_recv (date) as "_". + by iApply "HΦ". + - wp_send with "[//]". wp_send with "[//]". wp_send with "[Hc//]". + by iApply "HΦ". + Qed. + + Definition buyer3_prot b1 b2' s : iProto Σ := + (<(Recv, b2') @ (quote:Z)> MSG #quote ; + <(Recv, b2') @ (contrib:Z)> MSG #contrib ; + <(Recv, b2') @ (c:val)> MSG c + {{ c ↣ (<(Send, s) @ (b:bool)> MSG #b ; + <(Send, b1)> MSG #b ; + if b then <(Recv, s) @ (date:Z)> MSG #date ; END + else <(Send, s)> MSG #false ; END)%proto }} ; + END)%proto. + + Lemma buyer3_spec c b1 b2' s : + {{{ c ↣ buyer3_prot b1 b2' s }}} + buyer3_prog c #b1 #b2' #s + {{{ b, RET #b; True }}}. + Proof. + iIntros (Φ) "Hc HΦ". wp_lam. wp_pures. + wp_recv (quote) as "_". + wp_recv (contrib) as "_". + wp_recv (c') as "Hc'". + wp_pures. + case_bool_decide. + - wp_send with "[//]". wp_send with "[//]". wp_recv (date) as "_". + by iApply "HΦ". + - wp_pures. by iApply "HΦ". + Qed. + + Definition two_buyer_prot : list (iProto Σ) := + [buyer1_prot 1 2 ; seller_prot 0 2; buyer2_prot 0 1]. + + Lemma two_buyer_prot_consistent : + ⊢ iProto_consistent two_buyer_prot. + Proof. + rewrite /two_buyer_prot. iProto_consistent_take_steps. + destruct x2; iProto_consistent_take_steps. + Qed. + + Definition three_buyer_prot : list (iProto Σ) := + [buyer2_prot' 1 0 1; buyer3_prot 0 1 0]. + + Lemma three_buyer_prot_consistent : + ⊢ iProto_consistent three_buyer_prot. + Proof. rewrite /three_buyer_prot. iProto_consistent_take_steps. Qed. + + Lemma three_buyer_spec : + {{{ True }}} + three_buyer_prog #() + {{{ RET #(); True }}}. + Proof using chanG0 heapGS0 Σ. + iIntros (Φ) "Hc HΦ". wp_lam. + wp_new_chan two_buyer_prot with two_buyer_prot_consistent + as (???) "Hcb1" "Hcs" "Hcb2". + wp_pures. + wp_new_chan three_buyer_prot with three_buyer_prot_consistent + as (??) "Hcb2'" "Hcb3". + wp_smart_apply (wp_fork with "[Hcs]"). + { by iApply (seller_spec with "Hcs"). } + wp_smart_apply (wp_fork with "[Hcb2 Hcb2']"). + { by iApply (buyer2_spec with "[$Hcb2 $Hcb2']"). } + wp_smart_apply (wp_fork with "[Hcb3]"). + { by iApply (buyer3_spec with "Hcb3"). } + wp_smart_apply (buyer1_spec with "Hcb1"). + by iApply "HΦ". + Qed. + +End three_buyer. -- GitLab From df6f77e7de7bd88477c7c71e06cd0b61453bbc3d Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Thu, 4 Apr 2024 20:34:28 +0200 Subject: [PATCH 73/81] Attempt at better delegation example --- .../examples/leader_election_del_alt.v | 291 ++++++++++++++++++ 1 file changed, 291 insertions(+) create mode 100644 multi_actris/examples/leader_election_del_alt.v diff --git a/multi_actris/examples/leader_election_del_alt.v b/multi_actris/examples/leader_election_del_alt.v new file mode 100644 index 0000000..a5b63aa --- /dev/null +++ b/multi_actris/examples/leader_election_del_alt.v @@ -0,0 +1,291 @@ +From iris.heap_lang Require Import adequacy. +From iris.heap_lang.lib Require Import assert. +From multi_actris.channel Require Import proofmode. + +(** Inspired by https://en.wikipedia.org/wiki/Chang_and_Roberts_algorithm *) + +Definition process : val := + rec: "go" "c" "idl" "id" "idr" "isp" := + if: recv "c" "idr" + then let: "id'" := recv "c" "idr" in + if: "id" < "id'" (** Case 1 *) + then send "c" "idl" #true;; send "c" "idl" "id'";; + "go" "c" "idl" "id" "idr" #true + else if: "id" = "id'" (** Case 4 *) + then send "c" "idl" #false;; send "c" "idl" "id";; + "go" "c" "idl" "id" "idr" #false + else if: "isp" (** Case 3 *) + then "go" "c" "idl" "id" "idr" "isp" (** Case 2 *) + else send "c" "idl" #true;; send "c" "idl" "id";; + "go" "c" "idl" "id" "idr" #true + else let: "id'" := recv "c" "idr" in + if: "id" = "id'" then "id'" + else send "c" "idl" #false;; send "c" "idl" "id'";; "id'". + +Definition init : val := + λ: "c" "idl" "id" "idr", + (* Notice missing leader *) + send "c" "idl" #true;; send "c" "idl" "id";; + process "c" "idl" "id" "idr" #true. + +Definition forward : val := + λ: "c" "c'" "idl" "id" "idr" "id_max", + if: "id" = "id_max" then + send "c'" #0 "id_max";; + send "c" "idl" #();; + recv "c" "idr" + else + recv "c" "idr";; + send "c'" #0 "id_max";; + send "c" "idl" #(). + +Definition program : val := + λ: <>, + let: "cs'" := new_chan #2 in + let: "c0'" := get_chan "cs'" #0 in + let: "c1'" := get_chan "cs'" #1 in + Fork (let: "id_max" := recv "c0'" #1 in + (rec: "f" <> := + let: "id'" := recv "c0'" #1 in + assert: ("id'" = "id_max");; "f" #()) #());; + let: "cs" := new_chan #4 in + let: "c0" := get_chan "cs" #0 in + let: "c1" := get_chan "cs" #1 in + let: "c2" := get_chan "cs" #2 in + let: "c3" := get_chan "cs" #3 in + Fork (let: "id_max" := process "c1" #0 #1 #2 #false in + forward "c1" "c1'" #0 #1 #2 "id_max");; + Fork (let: "id_max" := process "c2" #1 #2 #3 #false in + forward "c2" "c1'" #1 #2 #3 "id_max");; + Fork (let: "id_max" := process "c3" #2 #3 #0 #false in + forward "c3" "c1'" #2 #3 #0 "id_max");; + let: "id_max" := init "c0" #3 #0 #1 in + forward "c0" "c1'" #3 #0 #1 "id_max". + +Notation iProto_choice a p1 p2 := + (<a @ (b : bool)> MSG #b; if b then p1 else p2)%proto. + +Section ring_leader_election_example. + Context `{!heapGS Σ, chanG Σ, spawnG Σ, mono_natG Σ}. + + Definition my_recv_prot (il i ir : nat) (P : iProp Σ) (p : nat → iProto Σ) + (rec : bool -d> iProto Σ) : bool -d> iProto Σ := + λ (isp:bool), + iProto_choice (Recv,ir) + (<(Recv,ir) @ (i':nat)> MSG #i'; + if bool_decide (i < i') + then <(Send,il)> MSG #true ; <(Send,il)> MSG #i' ; rec true + else if bool_decide (i = i') + then <(Send,il)> MSG #false ; <(Send, il)> MSG #i ; rec false + else if isp then rec isp + else <(Send,il)> MSG #true ; <(Send,il)> MSG #i ; rec true)%proto + (<(Recv,ir) @ (i':nat)> MSG #i' {{ if bool_decide (i = i') then P else True%I }}; + if (bool_decide (i = i')) then p i + else <(Send,il)> MSG #false; <(Send,il)> MSG #i'; p i')%proto. + + Instance rle_prot_aux_contractive il i ir P p : Contractive (my_recv_prot il i ir P p). + Proof. rewrite /my_recv_prot. solve_proto_contractive. Qed. + Definition rle_prot il i ir P p := fixpoint (my_recv_prot il i ir P p). + Instance rle_prot_unfold il i ir isp P p : + ProtoUnfold (rle_prot il i ir P p isp) (my_recv_prot il i ir P p (rle_prot il i ir P p) isp). + Proof. apply proto_unfold_eq, (fixpoint_unfold (my_recv_prot il i ir P p)). Qed. + Lemma rle_prot_unfold' il i ir isp P p : + (rle_prot il i ir P p isp) ≡ + (my_recv_prot il i ir P p (rle_prot il i ir P p) isp). + Proof. apply (fixpoint_unfold (my_recv_prot il i ir P p)). Qed. + + Definition rle_preprot (il i ir : nat) P p : iProto Σ := + (<(Send, il)> MSG #true; <(Send, il)> MSG #i {{ P }} ; + rle_prot il i ir P p true)%proto. + + Lemma process_spec il i ir P p c (isp:bool) : + {{{ c ↣ (rle_prot il i ir P p isp) }}} + process c #il #i #ir #isp + {{{ i', RET #i'; c ↣ p i' ∗ if (bool_decide (i = i')) then P else True%I }}}. + Proof. + iIntros (Φ) "Hc HΦ". + iLöb as "IH" forall (Φ isp). + wp_lam. wp_recv (b) as "_". + destruct b. + - wp_pures. wp_recv (i') as "_". + wp_pures. + case_bool_decide as Hlt. + { case_bool_decide; [|lia]. + wp_pures. wp_send with "[//]". + wp_send with "[//]". wp_pures. + iApply ("IH" with "Hc HΦ"). } + case_bool_decide as Hlt2. + { case_bool_decide; [lia|]. + wp_pures. case_bool_decide; [|simplify_eq;lia]. + wp_send with "[//]". + wp_send with "[//]". wp_pures. + iApply ("IH" with "Hc HΦ"). } + case_bool_decide; [lia|]. + wp_pures. + case_bool_decide; [simplify_eq;lia|]. + wp_pures. + destruct isp. + { wp_pures. iApply ("IH" with "Hc HΦ"). } + wp_pures. + wp_send with "[//]". + wp_send with "[//]". + wp_pures. iApply ("IH" with "Hc HΦ"). + - wp_pures. + wp_recv (id') as "HP". wp_pures. + case_bool_decide as Hlt. + { case_bool_decide; [|simplify_eq;lia]. + wp_pures. subst. iApply "HΦ". + iFrame. + case_bool_decide; [|lia]. + done. } + case_bool_decide; [simplify_eq;lia|]. + wp_send with "[//]". wp_send with "[//]". wp_pures. iApply "HΦ". + iFrame. + case_bool_decide; [simplify_eq;lia|]. + done. + Qed. + + Lemma init_spec c il i ir p P : + {{{ c ↣ rle_preprot il i ir P p ∗ P }}} + init c #il #i #ir + {{{ i', RET #i'; c ↣ p i' ∗ if (bool_decide (i = i')) then P else True%I }}}. + Proof. + iIntros (Φ) "[Hc HP] HΦ". wp_lam. wp_send with "[//]". wp_send with "[HP//]". + wp_pures. iApply (process_spec with "Hc HΦ"). + Qed. + + Definition forward_prot (c : val) (p : iProto Σ) (il i ir i_max : nat) : iProto Σ := + if bool_decide (i = i_max) then + (<(Send,il)> MSG #() {{ c ↣ p }} ; <(Recv,ir)> MSG #() {{ c ↣ p }}; END)%proto + else + (<(Recv,ir)> MSG #() {{ c ↣ p }} ; <(Send,il)> MSG #() {{ c ↣ p }}; END)%proto. + + Definition relay_send_aux (id : nat) (rec : iProto Σ) : iProto Σ := + (<(Send,0)> MSG #id ; rec)%proto. + Instance relay_send_aux_contractive i : Contractive (relay_send_aux i). + Proof. solve_proto_contractive. Qed. + Definition relay_send_prot i := fixpoint (relay_send_aux i). + Instance relay_send_prot_unfold i : + ProtoUnfold (relay_send_prot i) (relay_send_aux i (relay_send_prot i)). + Proof. apply proto_unfold_eq, (fixpoint_unfold (relay_send_aux i)). Qed. + Lemma relay_send_prot_unfold' i : + (relay_send_prot i) ≡ + (relay_send_aux i (relay_send_prot i)). + Proof. apply (fixpoint_unfold (relay_send_aux i)). Qed. + + Definition relay_send_preprot : iProto Σ := + (<(Send,0) @ (i_max:nat)> MSG #i_max ; relay_send_prot i_max)%proto. + + Definition relay_recv_aux (id : nat) (rec : iProto Σ) : iProto Σ := + (<(Recv,1)> MSG #id ; rec)%proto. + Instance relay_recv_aux_contractive i : Contractive (relay_recv_aux i). + Proof. solve_proto_contractive. Qed. + Definition relay_recv_prot i := fixpoint (relay_recv_aux i). + Instance relay_recv_prot_unfold i : + ProtoUnfold (relay_recv_prot i) (relay_recv_aux i (relay_recv_prot i)). + Proof. apply proto_unfold_eq, (fixpoint_unfold (relay_recv_aux i)). Qed. + Lemma relay_recv_prot_unfold' i : + (relay_recv_prot i) ≡ + (relay_recv_aux i (relay_recv_prot i)). + Proof. apply (fixpoint_unfold (relay_recv_aux i)). Qed. + Definition relay_recv_preprot : iProto Σ := + (<(Recv,1) @ (i_max:nat)> MSG #i_max ; relay_recv_prot i_max)%proto. + + Definition prot_pool c P : list (iProto Σ) := + [rle_preprot 3 0 1 P (λ id_max, forward_prot c (relay_send_prot id_max) 3 0 1 id_max); + rle_prot 0 1 2 P (λ id_max, forward_prot c (relay_send_prot id_max) 0 1 2 id_max) false; + rle_prot 1 2 3 P (λ id_max, forward_prot c (relay_send_prot id_max) 1 2 3 id_max) false; + rle_prot 2 3 0 P (λ id_max, forward_prot c (relay_send_prot id_max) 2 3 0 id_max) false]. + + Lemma prot_pool_consistent c P : ⊢ iProto_consistent (prot_pool c P). + Proof. + rewrite /prot_pool /rle_preprot. + rewrite !rle_prot_unfold'. + iProto_consistent_take_steps. + case_bool_decide; try lia. + rewrite !rle_prot_unfold'. + iProto_consistent_take_steps. + case_bool_decide; try lia. + rewrite !rle_prot_unfold'. + iProto_consistent_take_steps. + case_bool_decide; try lia. + rewrite !rle_prot_unfold'. + iProto_consistent_take_steps. + Qed. + + Definition prot_pool' : list (iProto Σ) := + [relay_recv_preprot; + relay_send_preprot]. + + Lemma prot_pool_consistent' : ⊢ iProto_consistent (prot_pool'). + Proof. + rewrite /prot_pool'. + iProto_consistent_take_steps. + iLöb as "IH". + iEval (rewrite relay_recv_prot_unfold'). + iEval (rewrite relay_send_prot_unfold'). + iProto_consistent_take_steps. + done. + Qed. + + Lemma forward_spec c c' il i ir i_max : + {{{ c ↣ forward_prot c' (relay_send_prot i_max) il i ir i_max ∗ + if (bool_decide (i = i_max)) then c' ↣ relay_send_preprot else True%I }}} + forward c c' #il #i #ir #i_max + {{{ RET #(); True }}}. + Proof. + iIntros (Φ) "[Hc Hc'] HΦ". wp_lam. + rewrite /forward_prot. + wp_pures. case_bool_decide. + - simplify_eq. wp_pures. + case_bool_decide; [|simplify_eq;lia]. + wp_send with "[//]". + wp_pures. wp_send with "[Hc'//]". + wp_recv as "Hc'". by iApply "HΦ". + - case_bool_decide; [simplify_eq;lia|]. + wp_pures. iClear "Hc'". + wp_recv as "Hc'". + wp_send with "[//]". + wp_send with "[Hc'//]". + by iApply "HΦ". + Qed. + + Lemma program_spec : + {{{ True }}} program #() {{{ RET #(); True }}}. + Proof. + iIntros (Φ) "_ HΦ". wp_lam. + wp_new_chan (prot_pool') with (prot_pool_consistent') + as (c0' c1') "Hc0'" "Hc1'". + wp_smart_apply (wp_fork with "[Hc0']"). + { iIntros "!>". + wp_pures. + wp_recv (i_max) as "_". wp_pures. + iLöb as "IH". + wp_recv as "_". wp_pures. + wp_smart_apply wp_assert. + wp_pures. iModIntro. iSplit; [iPureIntro; f_equiv; by case_bool_decide|]. + iIntros "!>". wp_pures. by iApply "IH". } + wp_new_chan (prot_pool c1' (c1' ↣ relay_send_preprot)) with prot_pool_consistent + as (c0 c1 c2 c3) "Hc0" "Hc1" "Hc2" "Hc3". + wp_smart_apply (wp_fork with "[Hc1]"). + { iIntros "!>". wp_smart_apply (process_spec with "Hc1"). + iIntros (i') "Hc1". by wp_smart_apply (forward_spec with "[$Hc1]"). } + wp_smart_apply (wp_fork with "[Hc2]"). + { iIntros "!>". wp_smart_apply (process_spec with "Hc2"). + iIntros (i') "Hc2". by wp_smart_apply (forward_spec with "[$Hc2]"). } + wp_smart_apply (wp_fork with "[Hc3]"). + { iIntros "!>". wp_smart_apply (process_spec with "Hc3"). + iIntros (i') "Hc3". by wp_smart_apply (forward_spec with "[$Hc3]"). } + wp_smart_apply (init_spec with "[$Hc0 $Hc1']"). + iIntros (i') "Hc0". by wp_smart_apply (forward_spec with "[$Hc0]"). + Qed. + +End ring_leader_election_example. + +Lemma program_spec_adequate : + adequate NotStuck (program #()) ({|heap := ∅; used_proph_id := ∅|}) + (λ _ _, True). +Proof. + apply (heap_adequacy #[heapΣ; chanΣ]). + iIntros (Σ) "H". by iApply program_spec. +Qed. -- GitLab From 4dd80e836173024c6330893ddcc5d100e64aea71 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Thu, 4 Apr 2024 20:41:25 +0200 Subject: [PATCH 74/81] Generalised resources further --- .../examples/leader_election_del_alt.v | 22 +++++++++---------- 1 file changed, 11 insertions(+), 11 deletions(-) diff --git a/multi_actris/examples/leader_election_del_alt.v b/multi_actris/examples/leader_election_del_alt.v index a5b63aa..4495041 100644 --- a/multi_actris/examples/leader_election_del_alt.v +++ b/multi_actris/examples/leader_election_del_alt.v @@ -154,11 +154,11 @@ Section ring_leader_election_example. wp_pures. iApply (process_spec with "Hc HΦ"). Qed. - Definition forward_prot (c : val) (p : iProto Σ) (il i ir i_max : nat) : iProto Σ := + Definition forward_prot (P:iProp Σ) (il i ir i_max : nat) : iProto Σ := if bool_decide (i = i_max) then - (<(Send,il)> MSG #() {{ c ↣ p }} ; <(Recv,ir)> MSG #() {{ c ↣ p }}; END)%proto + (<(Send,il)> MSG #() {{ P }} ; <(Recv,ir)> MSG #() {{ P }}; END)%proto else - (<(Recv,ir)> MSG #() {{ c ↣ p }} ; <(Send,il)> MSG #() {{ c ↣ p }}; END)%proto. + (<(Recv,ir)> MSG #() {{ P }} ; <(Send,il)> MSG #() {{ P }}; END)%proto. Definition relay_send_aux (id : nat) (rec : iProto Σ) : iProto Σ := (<(Send,0)> MSG #id ; rec)%proto. @@ -191,13 +191,13 @@ Section ring_leader_election_example. Definition relay_recv_preprot : iProto Σ := (<(Recv,1) @ (i_max:nat)> MSG #i_max ; relay_recv_prot i_max)%proto. - Definition prot_pool c P : list (iProto Σ) := - [rle_preprot 3 0 1 P (λ id_max, forward_prot c (relay_send_prot id_max) 3 0 1 id_max); - rle_prot 0 1 2 P (λ id_max, forward_prot c (relay_send_prot id_max) 0 1 2 id_max) false; - rle_prot 1 2 3 P (λ id_max, forward_prot c (relay_send_prot id_max) 1 2 3 id_max) false; - rle_prot 2 3 0 P (λ id_max, forward_prot c (relay_send_prot id_max) 2 3 0 id_max) false]. + Definition prot_pool P Φ : list (iProto Σ) := + [rle_preprot 3 0 1 P (λ id_max, forward_prot (Φ id_max) 3 0 1 id_max); + rle_prot 0 1 2 P (λ id_max, forward_prot (Φ id_max) 0 1 2 id_max) false; + rle_prot 1 2 3 P (λ id_max, forward_prot (Φ id_max) 1 2 3 id_max) false; + rle_prot 2 3 0 P (λ id_max, forward_prot (Φ id_max) 2 3 0 id_max) false]. - Lemma prot_pool_consistent c P : ⊢ iProto_consistent (prot_pool c P). + Lemma prot_pool_consistent P Φ : ⊢ iProto_consistent (prot_pool P Φ). Proof. rewrite /prot_pool /rle_preprot. rewrite !rle_prot_unfold'. @@ -229,7 +229,7 @@ Section ring_leader_election_example. Qed. Lemma forward_spec c c' il i ir i_max : - {{{ c ↣ forward_prot c' (relay_send_prot i_max) il i ir i_max ∗ + {{{ c ↣ forward_prot (c' ↣ relay_send_prot i_max) il i ir i_max ∗ if (bool_decide (i = i_max)) then c' ↣ relay_send_preprot else True%I }}} forward c c' #il #i #ir #i_max {{{ RET #(); True }}}. @@ -265,7 +265,7 @@ Section ring_leader_election_example. wp_smart_apply wp_assert. wp_pures. iModIntro. iSplit; [iPureIntro; f_equiv; by case_bool_decide|]. iIntros "!>". wp_pures. by iApply "IH". } - wp_new_chan (prot_pool c1' (c1' ↣ relay_send_preprot)) with prot_pool_consistent + wp_new_chan (prot_pool (c1' ↣ relay_send_preprot) (λ i_max, c1' ↣ relay_send_prot i_max)) with prot_pool_consistent as (c0 c1 c2 c3) "Hc0" "Hc1" "Hc2" "Hc3". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". wp_smart_apply (process_spec with "Hc1"). -- GitLab From afed8a408f7a802a63687fb8c1201b358c37d32b Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Fri, 5 Apr 2024 01:17:22 +0200 Subject: [PATCH 75/81] Clean up --- .../examples/leader_election_del_alt.v | 93 +++++++------------ 1 file changed, 36 insertions(+), 57 deletions(-) diff --git a/multi_actris/examples/leader_election_del_alt.v b/multi_actris/examples/leader_election_del_alt.v index 4495041..e7429c2 100644 --- a/multi_actris/examples/leader_election_del_alt.v +++ b/multi_actris/examples/leader_election_del_alt.v @@ -24,7 +24,6 @@ Definition process : val := Definition init : val := λ: "c" "idl" "id" "idr", - (* Notice missing leader *) send "c" "idl" #true;; send "c" "idl" "id";; process "c" "idl" "id" "idr" #true. @@ -53,11 +52,11 @@ Definition program : val := let: "c1" := get_chan "cs" #1 in let: "c2" := get_chan "cs" #2 in let: "c3" := get_chan "cs" #3 in - Fork (let: "id_max" := process "c1" #0 #1 #2 #false in + Fork (let: "id_max" := process "c1" #0 #1 #2 #false in forward "c1" "c1'" #0 #1 #2 "id_max");; - Fork (let: "id_max" := process "c2" #1 #2 #3 #false in + Fork (let: "id_max" := process "c2" #1 #2 #3 #false in forward "c2" "c1'" #1 #2 #3 "id_max");; - Fork (let: "id_max" := process "c3" #2 #3 #0 #false in + Fork (let: "id_max" := process "c3" #2 #3 #0 #false in forward "c3" "c1'" #2 #3 #0 "id_max");; let: "id_max" := init "c0" #3 #0 #1 in forward "c0" "c1'" #3 #0 #1 "id_max". @@ -79,15 +78,18 @@ Section ring_leader_election_example. then <(Send,il)> MSG #false ; <(Send, il)> MSG #i ; rec false else if isp then rec isp else <(Send,il)> MSG #true ; <(Send,il)> MSG #i ; rec true)%proto - (<(Recv,ir) @ (i':nat)> MSG #i' {{ if bool_decide (i = i') then P else True%I }}; + (<(Recv,ir) @ (i':nat)> MSG #i' + {{ if bool_decide (i = i') then P else True%I }}; if (bool_decide (i = i')) then p i else <(Send,il)> MSG #false; <(Send,il)> MSG #i'; p i')%proto. - Instance rle_prot_aux_contractive il i ir P p : Contractive (my_recv_prot il i ir P p). + Instance rle_prot_aux_contractive il i ir P p : + Contractive (my_recv_prot il i ir P p). Proof. rewrite /my_recv_prot. solve_proto_contractive. Qed. Definition rle_prot il i ir P p := fixpoint (my_recv_prot il i ir P p). Instance rle_prot_unfold il i ir isp P p : - ProtoUnfold (rle_prot il i ir P p isp) (my_recv_prot il i ir P p (rle_prot il i ir P p) isp). + ProtoUnfold (rle_prot il i ir P p isp) + (my_recv_prot il i ir P p (rle_prot il i ir P p) isp). Proof. apply proto_unfold_eq, (fixpoint_unfold (my_recv_prot il i ir P p)). Qed. Lemma rle_prot_unfold' il i ir isp P p : (rle_prot il i ir P p isp) ≡ @@ -107,42 +109,27 @@ Section ring_leader_election_example. iLöb as "IH" forall (Φ isp). wp_lam. wp_recv (b) as "_". destruct b. - - wp_pures. wp_recv (i') as "_". - wp_pures. - case_bool_decide as Hlt. + - wp_recv (i') as "_". + wp_pures. case_bool_decide as Hlt. { case_bool_decide; [|lia]. - wp_pures. wp_send with "[//]". - wp_send with "[//]". wp_pures. - iApply ("IH" with "Hc HΦ"). } + wp_send with "[//]". wp_send with "[//]". + wp_smart_apply ("IH" with "Hc HΦ"). } case_bool_decide as Hlt2. - { case_bool_decide; [lia|]. - wp_pures. case_bool_decide; [|simplify_eq;lia]. - wp_send with "[//]". - wp_send with "[//]". wp_pures. - iApply ("IH" with "Hc HΦ"). } + { case_bool_decide; [lia|]. wp_pures. case_bool_decide; [|simplify_eq;lia]. + wp_send with "[//]". wp_send with "[//]". + wp_smart_apply ("IH" with "Hc HΦ"). } case_bool_decide; [lia|]. - wp_pures. - case_bool_decide; [simplify_eq;lia|]. - wp_pures. - destruct isp. - { wp_pures. iApply ("IH" with "Hc HΦ"). } - wp_pures. - wp_send with "[//]". - wp_send with "[//]". - wp_pures. iApply ("IH" with "Hc HΦ"). - - wp_pures. - wp_recv (id') as "HP". wp_pures. + wp_pures. case_bool_decide; [simplify_eq;lia|]. wp_pures. + destruct isp; [by wp_smart_apply ("IH" with "Hc HΦ")|]. + wp_send with "[//]". wp_send with "[//]". + wp_smart_apply ("IH" with "Hc HΦ"). + - wp_recv (id') as "HP". wp_pures. case_bool_decide as Hlt. { case_bool_decide; [|simplify_eq;lia]. - wp_pures. subst. iApply "HΦ". - iFrame. - case_bool_decide; [|lia]. - done. } + wp_pures. subst. iApply "HΦ". iFrame. by case_bool_decide. } case_bool_decide; [simplify_eq;lia|]. wp_send with "[//]". wp_send with "[//]". wp_pures. iApply "HΦ". - iFrame. - case_bool_decide; [simplify_eq;lia|]. - done. + iFrame. by case_bool_decide; [simplify_eq;lia|]. Qed. Lemma init_spec c il i ir p P : @@ -151,7 +138,7 @@ Section ring_leader_election_example. {{{ i', RET #i'; c ↣ p i' ∗ if (bool_decide (i = i')) then P else True%I }}}. Proof. iIntros (Φ) "[Hc HP] HΦ". wp_lam. wp_send with "[//]". wp_send with "[HP//]". - wp_pures. iApply (process_spec with "Hc HΦ"). + wp_smart_apply (process_spec with "Hc HΦ"). Qed. Definition forward_prot (P:iProp Σ) (il i ir i_max : nat) : iProto Σ := @@ -214,8 +201,7 @@ Section ring_leader_election_example. Qed. Definition prot_pool' : list (iProto Σ) := - [relay_recv_preprot; - relay_send_preprot]. + [relay_recv_preprot; relay_send_preprot]. Lemma prot_pool_consistent' : ⊢ iProto_consistent (prot_pool'). Proof. @@ -229,7 +215,7 @@ Section ring_leader_election_example. Qed. Lemma forward_spec c c' il i ir i_max : - {{{ c ↣ forward_prot (c' ↣ relay_send_prot i_max) il i ir i_max ∗ + {{{ c ↣ forward_prot (c' ↣ relay_send_prot i_max) il i ir i_max ∗ if (bool_decide (i = i_max)) then c' ↣ relay_send_preprot else True%I }}} forward c c' #il #i #ir #i_max {{{ RET #(); True }}}. @@ -239,34 +225,27 @@ Section ring_leader_election_example. wp_pures. case_bool_decide. - simplify_eq. wp_pures. case_bool_decide; [|simplify_eq;lia]. - wp_send with "[//]". - wp_pures. wp_send with "[Hc'//]". - wp_recv as "Hc'". by iApply "HΦ". - - case_bool_decide; [simplify_eq;lia|]. - wp_pures. iClear "Hc'". - wp_recv as "Hc'". - wp_send with "[//]". - wp_send with "[Hc'//]". + wp_send with "[//]". wp_send with "[Hc'//]". wp_recv as "Hc'". by iApply "HΦ". + - case_bool_decide; [simplify_eq;lia|]. + iClear "Hc'". wp_recv as "Hc'". + wp_send with "[//]". wp_send with "[Hc'//]". by iApply "HΦ". Qed. Lemma program_spec : {{{ True }}} program #() {{{ RET #(); True }}}. - Proof. + Proof. iIntros (Φ) "_ HΦ". wp_lam. wp_new_chan (prot_pool') with (prot_pool_consistent') as (c0' c1') "Hc0'" "Hc1'". wp_smart_apply (wp_fork with "[Hc0']"). - { iIntros "!>". - wp_pures. - wp_recv (i_max) as "_". wp_pures. - iLöb as "IH". - wp_recv as "_". wp_pures. - wp_smart_apply wp_assert. + { iIntros "!>". wp_recv (i_max) as "_". wp_pures. + iLöb as "IH". wp_recv as "_". wp_smart_apply wp_assert. wp_pures. iModIntro. iSplit; [iPureIntro; f_equiv; by case_bool_decide|]. iIntros "!>". wp_pures. by iApply "IH". } - wp_new_chan (prot_pool (c1' ↣ relay_send_preprot) (λ i_max, c1' ↣ relay_send_prot i_max)) with prot_pool_consistent - as (c0 c1 c2 c3) "Hc0" "Hc1" "Hc2" "Hc3". + wp_new_chan (prot_pool (c1' ↣ relay_send_preprot) + (λ i_max, c1' ↣ relay_send_prot i_max)) + with prot_pool_consistent as (c0 c1 c2 c3) "Hc0" "Hc1" "Hc2" "Hc3". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". wp_smart_apply (process_spec with "Hc1"). iIntros (i') "Hc1". by wp_smart_apply (forward_spec with "[$Hc1]"). } -- GitLab From 5e0b1ef1bbe822af0f624f9869b1a8a8140f9d7f Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Fri, 5 Apr 2024 01:31:54 +0200 Subject: [PATCH 76/81] Mechanised rle_ref_prog --- .../examples/leader_election_del_alt.v | 74 ++++++++++++++++++- 1 file changed, 73 insertions(+), 1 deletion(-) diff --git a/multi_actris/examples/leader_election_del_alt.v b/multi_actris/examples/leader_election_del_alt.v index e7429c2..e3445c3 100644 --- a/multi_actris/examples/leader_election_del_alt.v +++ b/multi_actris/examples/leader_election_del_alt.v @@ -65,7 +65,7 @@ Notation iProto_choice a p1 p2 := (<a @ (b : bool)> MSG #b; if b then p1 else p2)%proto. Section ring_leader_election_example. - Context `{!heapGS Σ, chanG Σ, spawnG Σ, mono_natG Σ}. + Context `{!heapGS Σ, chanG Σ}. Definition my_recv_prot (il i ir : nat) (P : iProp Σ) (p : nat → iProto Σ) (rec : bool -d> iProto Σ) : bool -d> iProto Σ := @@ -268,3 +268,75 @@ Proof. apply (heap_adequacy #[heapΣ; chanΣ]). iIntros (Σ) "H". by iApply program_spec. Qed. + +Definition program2 : val := + λ: <>, + let: "l" := ref #42 in + let: "cs" := new_chan #4 in + let: "c0" := get_chan "cs" #0 in + let: "c1" := get_chan "cs" #1 in + let: "c2" := get_chan "cs" #2 in + let: "c3" := get_chan "cs" #3 in + Fork (let: "id_max" := process "c1" #0 #1 #2 #false in + if: "id_max" = #1 then Free "l" else #());; + Fork (let: "id_max" := process "c2" #1 #2 #3 #false in + if: "id_max" = #2 then Free "l" else #());; + Fork (let: "id_max" := process "c3" #2 #3 #0 #false in + if: "id_max" = #3 then Free "l" else #());; + let: "id_max" := init "c0" #3 #0 #1 in + if: "id_max" = #0 then Free "l" else #(). + +Section ring_leader_election_example. + Context `{!heapGS Σ, chanG Σ}. + + Definition prot_pool'' P : list (iProto Σ) := + [rle_preprot 3 0 1 P (λ id_max, END)%proto; + rle_prot 0 1 2 P (λ id_max, END)%proto false; + rle_prot 1 2 3 P (λ id_max, END)%proto false; + rle_prot 2 3 0 P (λ id_max, END)%proto false]. + + Lemma prot_pool_consistent'' P : ⊢ iProto_consistent (prot_pool'' P). + Proof. + rewrite /prot_pool'' /rle_preprot. + rewrite !rle_prot_unfold'. + iProto_consistent_take_steps. + case_bool_decide; try lia. + rewrite !rle_prot_unfold'. + iProto_consistent_take_steps. + case_bool_decide; try lia. + rewrite !rle_prot_unfold'. + iProto_consistent_take_steps. + case_bool_decide; try lia. + rewrite !rle_prot_unfold'. + iProto_consistent_take_steps. + Qed. + + Lemma program_spec'' : + {{{ True }}} program2 #() {{{ RET #(); True }}}. + Proof. + iIntros (Φ) "_ HΦ". wp_lam. + wp_alloc l as "Hl". + wp_new_chan (prot_pool'' (l ↦ #42)) + with prot_pool_consistent'' as (c0 c1 c2 c3) "Hc0" "Hc1" "Hc2" "Hc3". + wp_smart_apply (wp_fork with "[Hc1]"). + { iIntros "!>". wp_smart_apply (process_spec with "Hc1"). + iIntros (i') "[_ Hl]". wp_pures. case_bool_decide. + - case_bool_decide; [|simplify_eq; lia]. by wp_free. + - case_bool_decide; [simplify_eq; lia|]. by wp_pures. } + wp_smart_apply (wp_fork with "[Hc2]"). + { iIntros "!>". wp_smart_apply (process_spec with "Hc2"). + iIntros (i') "[_ Hl]". wp_pures. case_bool_decide. + - case_bool_decide; [|simplify_eq; lia]. by wp_free. + - case_bool_decide; [simplify_eq; lia|]. by wp_pures. } + wp_smart_apply (wp_fork with "[Hc3]"). + { iIntros "!>". wp_smart_apply (process_spec with "Hc3"). + iIntros (i') "[_ Hl]". wp_pures. case_bool_decide. + - case_bool_decide; [|simplify_eq; lia]. by wp_free. + - case_bool_decide; [simplify_eq; lia|]. by wp_pures. } + wp_smart_apply (init_spec with "[$Hc0 $Hl]"). + iIntros (i') "[_ Hl]". wp_pures. case_bool_decide. + - case_bool_decide; [|simplify_eq; lia]. wp_free. by iApply "HΦ". + - case_bool_decide; [simplify_eq; lia|]. wp_pures. by iApply "HΦ". + Qed. + +End ring_leader_election_example. -- GitLab From d470e20005e4e98cca8c9bc23c93ddcd8d57159b Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Fri, 5 Apr 2024 01:41:52 +0200 Subject: [PATCH 77/81] Refactored leader election example --- .../examples/leader_election_del_alt.v | 284 +++++++++--------- 1 file changed, 149 insertions(+), 135 deletions(-) diff --git a/multi_actris/examples/leader_election_del_alt.v b/multi_actris/examples/leader_election_del_alt.v index e3445c3..093badc 100644 --- a/multi_actris/examples/leader_election_del_alt.v +++ b/multi_actris/examples/leader_election_del_alt.v @@ -27,44 +27,10 @@ Definition init : val := send "c" "idl" #true;; send "c" "idl" "id";; process "c" "idl" "id" "idr" #true. -Definition forward : val := - λ: "c" "c'" "idl" "id" "idr" "id_max", - if: "id" = "id_max" then - send "c'" #0 "id_max";; - send "c" "idl" #();; - recv "c" "idr" - else - recv "c" "idr";; - send "c'" #0 "id_max";; - send "c" "idl" #(). - -Definition program : val := - λ: <>, - let: "cs'" := new_chan #2 in - let: "c0'" := get_chan "cs'" #0 in - let: "c1'" := get_chan "cs'" #1 in - Fork (let: "id_max" := recv "c0'" #1 in - (rec: "f" <> := - let: "id'" := recv "c0'" #1 in - assert: ("id'" = "id_max");; "f" #()) #());; - let: "cs" := new_chan #4 in - let: "c0" := get_chan "cs" #0 in - let: "c1" := get_chan "cs" #1 in - let: "c2" := get_chan "cs" #2 in - let: "c3" := get_chan "cs" #3 in - Fork (let: "id_max" := process "c1" #0 #1 #2 #false in - forward "c1" "c1'" #0 #1 #2 "id_max");; - Fork (let: "id_max" := process "c2" #1 #2 #3 #false in - forward "c2" "c1'" #1 #2 #3 "id_max");; - Fork (let: "id_max" := process "c3" #2 #3 #0 #false in - forward "c3" "c1'" #2 #3 #0 "id_max");; - let: "id_max" := init "c0" #3 #0 #1 in - forward "c0" "c1'" #3 #0 #1 "id_max". - Notation iProto_choice a p1 p2 := (<a @ (b : bool)> MSG #b; if b then p1 else p2)%proto. -Section ring_leader_election_example. +Section ring_leader_election_protocols. Context `{!heapGS Σ, chanG Σ}. Definition my_recv_prot (il i ir : nat) (P : iProp Σ) (p : nat → iProto Σ) @@ -141,7 +107,124 @@ Section ring_leader_election_example. wp_smart_apply (process_spec with "Hc HΦ"). Qed. - Definition forward_prot (P:iProp Σ) (il i ir i_max : nat) : iProto Σ := +End ring_leader_election_protocols. + +Definition rle_ref_prog : val := + λ: <>, + let: "l" := ref #42 in + let: "cs" := new_chan #4 in + let: "c0" := get_chan "cs" #0 in + let: "c1" := get_chan "cs" #1 in + let: "c2" := get_chan "cs" #2 in + let: "c3" := get_chan "cs" #3 in + Fork (let: "id_max" := process "c1" #0 #1 #2 #false in + if: "id_max" = #1 then Free "l" else #());; + Fork (let: "id_max" := process "c2" #1 #2 #3 #false in + if: "id_max" = #2 then Free "l" else #());; + Fork (let: "id_max" := process "c3" #2 #3 #0 #false in + if: "id_max" = #3 then Free "l" else #());; + let: "id_max" := init "c0" #3 #0 #1 in + if: "id_max" = #0 then Free "l" else #(). + +Section rle_ref_prog_proof. + Context `{!heapGS Σ, chanG Σ}. + + Definition rle_ref_prog_prot_pool P : list (iProto Σ) := + [rle_preprot 3 0 1 P (λ id_max, END)%proto; + rle_prot 0 1 2 P (λ id_max, END)%proto false; + rle_prot 1 2 3 P (λ id_max, END)%proto false; + rle_prot 2 3 0 P (λ id_max, END)%proto false]. + + Lemma rle_ref_prog_prot_pool_consistent P : + ⊢ iProto_consistent (rle_ref_prog_prot_pool P). + Proof. + rewrite /rle_ref_prog_prot_pool /rle_preprot. + rewrite !rle_prot_unfold'. + iProto_consistent_take_steps. + case_bool_decide; try lia. + rewrite !rle_prot_unfold'. + iProto_consistent_take_steps. + case_bool_decide; try lia. + rewrite !rle_prot_unfold'. + iProto_consistent_take_steps. + case_bool_decide; try lia. + rewrite !rle_prot_unfold'. + iProto_consistent_take_steps. + Qed. + + Lemma rle_ref_prog_spec : + {{{ True }}} rle_ref_prog #() {{{ RET #(); True }}}. + Proof. + iIntros (Φ) "_ HΦ". wp_lam. + wp_alloc l as "Hl". + wp_new_chan (rle_ref_prog_prot_pool (l ↦ #42)) + with rle_ref_prog_prot_pool_consistent + as (c0 c1 c2 c3) "Hc0" "Hc1" "Hc2" "Hc3". + wp_smart_apply (wp_fork with "[Hc1]"). + { iIntros "!>". wp_smart_apply (process_spec with "Hc1"). + iIntros (i') "[_ Hl]". wp_pures. case_bool_decide. + - case_bool_decide; [|simplify_eq; lia]. by wp_free. + - case_bool_decide; [simplify_eq; lia|]. by wp_pures. } + wp_smart_apply (wp_fork with "[Hc2]"). + { iIntros "!>". wp_smart_apply (process_spec with "Hc2"). + iIntros (i') "[_ Hl]". wp_pures. case_bool_decide. + - case_bool_decide; [|simplify_eq; lia]. by wp_free. + - case_bool_decide; [simplify_eq; lia|]. by wp_pures. } + wp_smart_apply (wp_fork with "[Hc3]"). + { iIntros "!>". wp_smart_apply (process_spec with "Hc3"). + iIntros (i') "[_ Hl]". wp_pures. case_bool_decide. + - case_bool_decide; [|simplify_eq; lia]. by wp_free. + - case_bool_decide; [simplify_eq; lia|]. by wp_pures. } + wp_smart_apply (init_spec with "[$Hc0 $Hl]"). + iIntros (i') "[_ Hl]". wp_pures. case_bool_decide. + - case_bool_decide; [|simplify_eq; lia]. wp_free. by iApply "HΦ". + - case_bool_decide; [simplify_eq; lia|]. wp_pures. by iApply "HΦ". + Qed. + +End rle_ref_prog_proof. + +Lemma rle_ref_prog_adequate : + adequate NotStuck (rle_ref_prog #()) ({|heap := ∅; used_proph_id := ∅|}) + (λ _ _, True). +Proof. + apply (heap_adequacy #[heapΣ; chanΣ]). + iIntros (Σ) "H". by iApply rle_ref_prog_spec. +Qed. + +Definition relay : val := + λ: "c" "c'" "idl" "id" "idr" "id_max", + if: "id" = "id_max" then + send "c'" #0 "id_max";; send "c" "idl" #();; recv "c" "idr" + else + recv "c" "idr";; send "c'" #0 "id_max";; send "c" "idl" #(). + +Definition rle_del_prog : val := + λ: <>, + let: "cs'" := new_chan #2 in + let: "c0'" := get_chan "cs'" #0 in + let: "c1'" := get_chan "cs'" #1 in + Fork (let: "id_max" := recv "c0'" #1 in + (rec: "f" <> := + let: "id'" := recv "c0'" #1 in + assert: ("id'" = "id_max");; "f" #()) #());; + let: "cs" := new_chan #4 in + let: "c0" := get_chan "cs" #0 in + let: "c1" := get_chan "cs" #1 in + let: "c2" := get_chan "cs" #2 in + let: "c3" := get_chan "cs" #3 in + Fork (let: "id_max" := process "c1" #0 #1 #2 #false in + relay "c1" "c1'" #0 #1 #2 "id_max");; + Fork (let: "id_max" := process "c2" #1 #2 #3 #false in + relay "c2" "c1'" #1 #2 #3 "id_max");; + Fork (let: "id_max" := process "c3" #2 #3 #0 #false in + relay "c3" "c1'" #2 #3 #0 "id_max");; + let: "id_max" := init "c0" #3 #0 #1 in + relay "c0" "c1'" #3 #0 #1 "id_max". + +Section rle_del_prog_proof. + Context `{!heapGS Σ, chanG Σ}. + + Definition relay_del_prot (P:iProp Σ) (il i ir i_max : nat) : iProto Σ := if bool_decide (i = i_max) then (<(Send,il)> MSG #() {{ P }} ; <(Recv,ir)> MSG #() {{ P }}; END)%proto else @@ -178,15 +261,16 @@ Section ring_leader_election_example. Definition relay_recv_preprot : iProto Σ := (<(Recv,1) @ (i_max:nat)> MSG #i_max ; relay_recv_prot i_max)%proto. - Definition prot_pool P Φ : list (iProto Σ) := - [rle_preprot 3 0 1 P (λ id_max, forward_prot (Φ id_max) 3 0 1 id_max); - rle_prot 0 1 2 P (λ id_max, forward_prot (Φ id_max) 0 1 2 id_max) false; - rle_prot 1 2 3 P (λ id_max, forward_prot (Φ id_max) 1 2 3 id_max) false; - rle_prot 2 3 0 P (λ id_max, forward_prot (Φ id_max) 2 3 0 id_max) false]. + Definition rle_del_prog_prot_pool P Φ : list (iProto Σ) := + [rle_preprot 3 0 1 P (λ id_max, relay_del_prot (Φ id_max) 3 0 1 id_max); + rle_prot 0 1 2 P (λ id_max, relay_del_prot (Φ id_max) 0 1 2 id_max) false; + rle_prot 1 2 3 P (λ id_max, relay_del_prot (Φ id_max) 1 2 3 id_max) false; + rle_prot 2 3 0 P (λ id_max, relay_del_prot (Φ id_max) 2 3 0 id_max) false]. - Lemma prot_pool_consistent P Φ : ⊢ iProto_consistent (prot_pool P Φ). + Lemma rle_del_prog_prot_pool_consistent P Φ : + ⊢ iProto_consistent (rle_del_prog_prot_pool P Φ). Proof. - rewrite /prot_pool /rle_preprot. + rewrite /rle_del_prog_prot_pool /rle_preprot. rewrite !rle_prot_unfold'. iProto_consistent_take_steps. case_bool_decide; try lia. @@ -200,12 +284,13 @@ Section ring_leader_election_example. iProto_consistent_take_steps. Qed. - Definition prot_pool' : list (iProto Σ) := + Definition rle_del_prog_prot_pool' : list (iProto Σ) := [relay_recv_preprot; relay_send_preprot]. - Lemma prot_pool_consistent' : ⊢ iProto_consistent (prot_pool'). + Lemma rle_del_prog_prot_pool_consistent' : + ⊢ iProto_consistent (rle_del_prog_prot_pool'). Proof. - rewrite /prot_pool'. + rewrite /rle_del_prog_prot_pool'. iProto_consistent_take_steps. iLöb as "IH". iEval (rewrite relay_recv_prot_unfold'). @@ -214,14 +299,14 @@ Section ring_leader_election_example. done. Qed. - Lemma forward_spec c c' il i ir i_max : - {{{ c ↣ forward_prot (c' ↣ relay_send_prot i_max) il i ir i_max ∗ + Lemma relay_spec c c' il i ir i_max : + {{{ c ↣ relay_del_prot (c' ↣ relay_send_prot i_max) il i ir i_max ∗ if (bool_decide (i = i_max)) then c' ↣ relay_send_preprot else True%I }}} - forward c c' #il #i #ir #i_max + relay c c' #il #i #ir #i_max {{{ RET #(); True }}}. Proof. iIntros (Φ) "[Hc Hc'] HΦ". wp_lam. - rewrite /forward_prot. + rewrite /relay_del_prot. wp_pures. case_bool_decide. - simplify_eq. wp_pures. case_bool_decide; [|simplify_eq;lia]. @@ -232,111 +317,40 @@ Section ring_leader_election_example. wp_send with "[//]". wp_send with "[Hc'//]". by iApply "HΦ". Qed. - Lemma program_spec : - {{{ True }}} program #() {{{ RET #(); True }}}. + Lemma rle_del_prog_spec : + {{{ True }}} rle_del_prog #() {{{ RET #(); True }}}. Proof. iIntros (Φ) "_ HΦ". wp_lam. - wp_new_chan (prot_pool') with (prot_pool_consistent') + wp_new_chan rle_del_prog_prot_pool' with rle_del_prog_prot_pool_consistent' as (c0' c1') "Hc0'" "Hc1'". wp_smart_apply (wp_fork with "[Hc0']"). { iIntros "!>". wp_recv (i_max) as "_". wp_pures. iLöb as "IH". wp_recv as "_". wp_smart_apply wp_assert. wp_pures. iModIntro. iSplit; [iPureIntro; f_equiv; by case_bool_decide|]. iIntros "!>". wp_pures. by iApply "IH". } - wp_new_chan (prot_pool (c1' ↣ relay_send_preprot) + wp_new_chan (rle_del_prog_prot_pool (c1' ↣ relay_send_preprot) (λ i_max, c1' ↣ relay_send_prot i_max)) - with prot_pool_consistent as (c0 c1 c2 c3) "Hc0" "Hc1" "Hc2" "Hc3". + with rle_del_prog_prot_pool_consistent + as (c0 c1 c2 c3) "Hc0" "Hc1" "Hc2" "Hc3". wp_smart_apply (wp_fork with "[Hc1]"). { iIntros "!>". wp_smart_apply (process_spec with "Hc1"). - iIntros (i') "Hc1". by wp_smart_apply (forward_spec with "[$Hc1]"). } + iIntros (i') "Hc1". by wp_smart_apply (relay_spec with "[$Hc1]"). } wp_smart_apply (wp_fork with "[Hc2]"). { iIntros "!>". wp_smart_apply (process_spec with "Hc2"). - iIntros (i') "Hc2". by wp_smart_apply (forward_spec with "[$Hc2]"). } + iIntros (i') "Hc2". by wp_smart_apply (relay_spec with "[$Hc2]"). } wp_smart_apply (wp_fork with "[Hc3]"). { iIntros "!>". wp_smart_apply (process_spec with "Hc3"). - iIntros (i') "Hc3". by wp_smart_apply (forward_spec with "[$Hc3]"). } + iIntros (i') "Hc3". by wp_smart_apply (relay_spec with "[$Hc3]"). } wp_smart_apply (init_spec with "[$Hc0 $Hc1']"). - iIntros (i') "Hc0". by wp_smart_apply (forward_spec with "[$Hc0]"). + iIntros (i') "Hc0". by wp_smart_apply (relay_spec with "[$Hc0]"). Qed. -End ring_leader_election_example. +End rle_del_prog_proof. -Lemma program_spec_adequate : - adequate NotStuck (program #()) ({|heap := ∅; used_proph_id := ∅|}) +Lemma rle_del_prog_adequate : + adequate NotStuck (rle_del_prog #()) ({|heap := ∅; used_proph_id := ∅|}) (λ _ _, True). Proof. apply (heap_adequacy #[heapΣ; chanΣ]). - iIntros (Σ) "H". by iApply program_spec. + iIntros (Σ) "H". by iApply rle_del_prog_spec. Qed. - -Definition program2 : val := - λ: <>, - let: "l" := ref #42 in - let: "cs" := new_chan #4 in - let: "c0" := get_chan "cs" #0 in - let: "c1" := get_chan "cs" #1 in - let: "c2" := get_chan "cs" #2 in - let: "c3" := get_chan "cs" #3 in - Fork (let: "id_max" := process "c1" #0 #1 #2 #false in - if: "id_max" = #1 then Free "l" else #());; - Fork (let: "id_max" := process "c2" #1 #2 #3 #false in - if: "id_max" = #2 then Free "l" else #());; - Fork (let: "id_max" := process "c3" #2 #3 #0 #false in - if: "id_max" = #3 then Free "l" else #());; - let: "id_max" := init "c0" #3 #0 #1 in - if: "id_max" = #0 then Free "l" else #(). - -Section ring_leader_election_example. - Context `{!heapGS Σ, chanG Σ}. - - Definition prot_pool'' P : list (iProto Σ) := - [rle_preprot 3 0 1 P (λ id_max, END)%proto; - rle_prot 0 1 2 P (λ id_max, END)%proto false; - rle_prot 1 2 3 P (λ id_max, END)%proto false; - rle_prot 2 3 0 P (λ id_max, END)%proto false]. - - Lemma prot_pool_consistent'' P : ⊢ iProto_consistent (prot_pool'' P). - Proof. - rewrite /prot_pool'' /rle_preprot. - rewrite !rle_prot_unfold'. - iProto_consistent_take_steps. - case_bool_decide; try lia. - rewrite !rle_prot_unfold'. - iProto_consistent_take_steps. - case_bool_decide; try lia. - rewrite !rle_prot_unfold'. - iProto_consistent_take_steps. - case_bool_decide; try lia. - rewrite !rle_prot_unfold'. - iProto_consistent_take_steps. - Qed. - - Lemma program_spec'' : - {{{ True }}} program2 #() {{{ RET #(); True }}}. - Proof. - iIntros (Φ) "_ HΦ". wp_lam. - wp_alloc l as "Hl". - wp_new_chan (prot_pool'' (l ↦ #42)) - with prot_pool_consistent'' as (c0 c1 c2 c3) "Hc0" "Hc1" "Hc2" "Hc3". - wp_smart_apply (wp_fork with "[Hc1]"). - { iIntros "!>". wp_smart_apply (process_spec with "Hc1"). - iIntros (i') "[_ Hl]". wp_pures. case_bool_decide. - - case_bool_decide; [|simplify_eq; lia]. by wp_free. - - case_bool_decide; [simplify_eq; lia|]. by wp_pures. } - wp_smart_apply (wp_fork with "[Hc2]"). - { iIntros "!>". wp_smart_apply (process_spec with "Hc2"). - iIntros (i') "[_ Hl]". wp_pures. case_bool_decide. - - case_bool_decide; [|simplify_eq; lia]. by wp_free. - - case_bool_decide; [simplify_eq; lia|]. by wp_pures. } - wp_smart_apply (wp_fork with "[Hc3]"). - { iIntros "!>". wp_smart_apply (process_spec with "Hc3"). - iIntros (i') "[_ Hl]". wp_pures. case_bool_decide. - - case_bool_decide; [|simplify_eq; lia]. by wp_free. - - case_bool_decide; [simplify_eq; lia|]. by wp_pures. } - wp_smart_apply (init_spec with "[$Hc0 $Hl]"). - iIntros (i') "[_ Hl]". wp_pures. case_bool_decide. - - case_bool_decide; [|simplify_eq; lia]. wp_free. by iApply "HΦ". - - case_bool_decide; [simplify_eq; lia|]. wp_pures. by iApply "HΦ". - Qed. - -End ring_leader_election_example. -- GitLab From 15a1c8011eff36a3299701e38419f3c999f3f2ad Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Fri, 5 Apr 2024 02:00:14 +0200 Subject: [PATCH 78/81] Misc changes --- multi_actris/channel/channel.v | 44 +++++++----------- multi_actris/channel/proofmode.v | 1 - multi_actris/channel/proto.v | 50 ++------------------- multi_actris/channel/proto_model.v | 5 ++- multi_actris/examples/basics.v | 2 - multi_actris/examples/leader_election_del.v | 1 - multi_actris/utils/matrix.v | 3 -- 7 files changed, 22 insertions(+), 84 deletions(-) diff --git a/multi_actris/channel/channel.v b/multi_actris/channel/channel.v index 7e68d81..9fa9113 100644 --- a/multi_actris/channel/channel.v +++ b/multi_actris/channel/channel.v @@ -1,23 +1,20 @@ -(** This file contains the definition of the channels, encoded as a pair of -lock-protected buffers, and their primitive proof rules. Moreover: +(** This file contains the definition of the channels, +and their primitive proof rules. Moreover: - It defines the connective [c ↣ prot] for ownership of channel endpoints, which describes that channel endpoint [c] adheres to protocol [prot]. -- It proves Actris's specifications of [send] and [recv] w.r.t. dependent - separation protocols. - -An encoding of the usual (binary) choice connectives [prot1 <{Q1}+{Q2}> prot2] -and [prot1 <{Q1}&{Q2}> prot2], inspired by session types, is also included in -this file. +- It proves Actris's specifications of [send] and [recv] w.r.t. + multiparty dependent separation protocols. In this file we define the three message-passing connectives: -- [new_chan] creates references to two empty buffers and a lock, and returns a - pair of endpoints, where the order of the two references determines the - polarity of the endpoints. -- [send] takes an endpoint and adds an element to the first buffer. -- [recv] performs a busy loop until there is something in the second buffer, - which it pops and returns, locking during each peek. +- [new_chan] creates an n*n matrix of references where [i,j] is the singleton + buffer from participant i to participant j +- [send] takes an endpoint, a participant id, and a value, and puts the value in + the reference corresponding to the participant id, and waits until recv takes + it out. +- [recv] takes an endpoint, and a participant id, and waits until a value is put + into the corresponding reference. It is additionaly shown that the channel ownership [c ↣ prot] is closed under the subprotocol relation [⊑] *) @@ -50,7 +47,6 @@ Definition send : val := let: "l" := matrix_get "m" "i" "j" in "l" <- SOME "v";; wait "l". -(* TODO: Move recursion further in *) Definition recv : val := rec: "go" "c" "j" := let: "m" := Fst "c" in @@ -185,23 +181,15 @@ Section channel. (λ i j l, ∃ γt, inv (nroot.@"p") (chan_inv γ (γEs !!! i) γt i j l))%I); [done..| |]. - { iApply (big_sepL_intro). - iIntros "!>" (k tt Hin). - iApply (big_sepL_intro). - iIntros "!>" (k' tt' Hin'). - iIntros (l) "Hl". + { iApply (big_sepL_intro). iIntros "!>" (k tt Hin). iApply (big_sepL_intro). + iIntros "!>" (k' tt' Hin'). iIntros (l) "Hl". iMod (own_alloc (Excl ())) as (γ') "Hγ'"; [done|]. - iExists γ'. - iApply inv_alloc. - iNext. - iLeft. iFrame. } - iIntros (mat) "Hmat". - iApply "HΦ". + iExists γ'. iApply inv_alloc. iNext. iLeft. iFrame. } + iIntros (mat) "Hmat". iApply "HΦ". iExists _, _, _. iFrame "#∗". rewrite left_id. iSplit; [done|]. iApply (big_sepL_impl with "H"). - iIntros "!>" (i ? HSome') "(Hauth & Hfrag & Hown)". - iFrame. + iIntros "!>" (i ? HSome') "(Hauth & Hfrag & Hown)". iFrame. rewrite (list_lookup_total_alt ps). simpl. rewrite right_id_L. rewrite HSome'. iFrame. Qed. diff --git a/multi_actris/channel/proofmode.v b/multi_actris/channel/proofmode.v index 5f9d7b8..81ca7e6 100644 --- a/multi_actris/channel/proofmode.v +++ b/multi_actris/channel/proofmode.v @@ -338,7 +338,6 @@ Proof. iNext. iRewrite -"Hp1". done. Qed. -(* TODO: Improve automation *) Tactic Notation "iProto_consistent_take_step_step" := let i := fresh in let j := fresh in diff --git a/multi_actris/channel/proto.v b/multi_actris/channel/proto.v index fdd225f..75efec7 100644 --- a/multi_actris/channel/proto.v +++ b/multi_actris/channel/proto.v @@ -1,50 +1,6 @@ -(** This file defines the core of the Actris logic: It defines dependent -separation protocols and the various operations on it like dual, append, and -branching. - -Dependent separation protocols [iProto] are defined by instantiating the -parameterized version in [proto_model] with the type of propositions [iProp] of Iris. -We define ways of constructing instances of the instantiated type via two -constructors: -- [iProto_end], which is identical to [proto_end]. -- [iProto_message], which takes an [action] and an [iMsg]. The type [iMsg] is a - sequence of binders [iMsg_exist], terminated by the payload constructed with - [iMsg_base] based on arguments [v], [P] and [prot], which are the value, the - carried proposition and the [iProto] tail, respectively. - -For convenience sake, we provide the following notations: -- [END], which is simply [iProto_end]. -- [∃ x, m], which is [iMsg_exist] with argument [m]. -- [MSG v {{ P }}; prot], which is [iMsg_Base] with arguments [v], [P] and [prot]. -- [<a> m], which is [iProto_message] with arguments [a] and [m]. - -We also include custom notation to more easily construct complete constructions: -- [<a x1 .. xn> m], which is [<a> ∃ x1, .. ∃ xn, m]. -- [<a x1 .. xn> MSG v; {{ P }}; prot], which constructs a full protocol. - -Futhermore, we define the following operations: -- [iProto_dual], which turns all [Send] of a protocol into [Recv] and vice-versa. -- [iProto_app], which appends two protocols. - -In addition we define the subprotocol relation [iProto_le] [⊑], which generalises -the following inductive definition for asynchronous subtyping on session types: - - p1 <: p2 p1 <: p2 p1 <: !B.p3 ?A.p3 <: p2 ----------- ---------------- ---------------- ---------------------------- -end <: end !A.p1 <: !A.p2 ?A.p1 <: ?A.p2 ?A.p1 <: !B.p2 - -Example: - -!R <: !R ?Q <: ?Q ?Q <: ?Q -------------------- -------------- -?Q.!R <: !R.?Q ?P.?Q <: ?P.?Q ------------------------------------- - ?P.?Q.!R <: !R.?P.?Q - -Lastly, relevant type classes instances are defined for each of the above -notions, such as contractiveness and non-expansiveness, after which the -specifications of the message-passing primitives are defined in terms of the -protocol connectives. *) +(** This file defines the core of the Multiparty Actris logic: +It defines multiparty dependent separation protocols, and the Multiparty Actris +ghost theory. *) From iris.algebra Require Import gmap excl_auth gmap_view. From iris.proofmode Require Import proofmode. From iris.base_logic Require Export lib.iprop. diff --git a/multi_actris/channel/proto_model.v b/multi_actris/channel/proto_model.v index 8bd4bd3..007f8ef 100644 --- a/multi_actris/channel/proto_model.v +++ b/multi_actris/channel/proto_model.v @@ -12,8 +12,9 @@ recursive domain equation: Here, the left-hand side of the sum is used for the terminated process, while the right-hand side is used for the communication constructors. The type -[action] is an inductively defined datatype with two constructors [Send] and -[Recv]. Compared to having an additional sum in [proto], this makes it +[action] is a pair of a an inductively defined datatype [tag] with two +constructors [Send] and [Recv], and a synchronosation id of [nat]. +Compared to having an additional sum in [proto], this makes it possible to factorize the code in a better way. The remainder [V → â–¶ proto → PROP] is a predicate that ranges over the diff --git a/multi_actris/examples/basics.v b/multi_actris/examples/basics.v index 82a0173..80169be 100644 --- a/multi_actris/examples/basics.v +++ b/multi_actris/examples/basics.v @@ -37,7 +37,6 @@ Section channel. ⊢ iProto_consistent iProto_roundtrip. Proof. rewrite /iProto_roundtrip. iProto_consistent_take_steps. Qed. - (* TODO: Fix nat/Z coercion. *) Lemma roundtrip_prog_spec : {{{ True }}} roundtrip_prog #() {{{ RET #42 ; True }}}. Proof using chanG0 heapGS0 Σ. @@ -320,7 +319,6 @@ Section forwarder. (<(Recv, 1) @ (x:Z)> MSG #x ; <(Send, 0)> MSG #x ; END)%proto]. - (* TODO: Anonymous variable in this is unsatisfactory *) Lemma iProto_forwarder_consistent : ⊢ iProto_consistent iProto_forwarder. Proof. diff --git a/multi_actris/examples/leader_election_del.v b/multi_actris/examples/leader_election_del.v index 7c40b66..1ff5abd 100644 --- a/multi_actris/examples/leader_election_del.v +++ b/multi_actris/examples/leader_election_del.v @@ -179,7 +179,6 @@ Section ring_leader_election_example. (relay_recv_aux i (relay_recv_prot i)). Proof. apply (fixpoint_unfold (relay_recv_aux i)). Qed. - Definition prot_pool : list (iProto Σ) := [rle_preprot 3 0 1 (λ id_max, forward_prot (relay_send_prot id_max) 3 0 1 id_max); rle_prot 0 1 2 (λ id_max, forward_prot (relay_send_prot id_max) 0 1 2 id_max) false; diff --git a/multi_actris/utils/matrix.v b/multi_actris/utils/matrix.v index 066da4f..166b7d2 100644 --- a/multi_actris/utils/matrix.v +++ b/multi_actris/utils/matrix.v @@ -24,7 +24,6 @@ Section with_Σ. [∗ list] j ↦ _ ∈ replicate n (), P i j (l +â‚— pos n i j). - Lemma array_to_matrix_pre l n m v : l ↦∗ replicate (n * m) v -∗ [∗ list] i ↦ _ ∈ replicate n (), (l +â‚— i*m) ↦∗ replicate m v. @@ -42,7 +41,6 @@ Section with_Σ. by iFrame. Qed. - (* TODO: rename *) Lemma big_sepL_replicate_type {A B} n (x1 : A) (x2 : B) (P : nat → iProp Σ) : ([∗ list] i ↦ _ ∈ replicate n x1, P i) ⊢ ([∗ list] i ↦ _ ∈ replicate n x2, P i). @@ -117,7 +115,6 @@ Section with_Σ. { iIntros "!>" (k x HIn). iApply (big_sepL_lookup_acc _ _ j ()). by apply lookup_replicate. } - simpl. rewrite {1}(big_sepL_lookup_acc _ (replicate m ()) i ()); [|by apply lookup_replicate]. iDestruct "Hm" as "[[Hij Hi] Hm]". -- GitLab From a1e24f8c217eff05e97fdc9bef152520b48806a1 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Thu, 23 Jan 2025 16:10:06 +0100 Subject: [PATCH 79/81] Bumped multiparty branch --- multi_actris/channel/channel.v | 2 +- multi_actris/channel/proofmode.v | 4 ++-- multi_actris/channel/proto.v | 28 ++++++++++++++-------------- multi_actris/channel/proto_alt.v | 20 ++++++++++---------- multi_actris/utils/matrix.v | 4 ++-- 5 files changed, 29 insertions(+), 29 deletions(-) diff --git a/multi_actris/channel/channel.v b/multi_actris/channel/channel.v index 9fa9113..c0a0fbb 100644 --- a/multi_actris/channel/channel.v +++ b/multi_actris/channel/channel.v @@ -103,7 +103,7 @@ Definition iProto_pointsto_aux : seal (@iProto_pointsto_def). by eexists. Qed. Definition iProto_pointsto := iProto_pointsto_aux.(unseal). Definition iProto_pointsto_eq : @iProto_pointsto = @iProto_pointsto_def := iProto_pointsto_aux.(seal_eq). -Arguments iProto_pointsto {_ _ _} _ _%proto. +Arguments iProto_pointsto {_ _ _} _ _%_proto. Global Instance: Params (@iProto_pointsto) 5 := {}. Notation "c ↣ p" := (iProto_pointsto c p) (at level 20, format "c ↣ p"). diff --git a/multi_actris/channel/proofmode.v b/multi_actris/channel/proofmode.v index 81ca7e6..2c02a5c 100644 --- a/multi_actris/channel/proofmode.v +++ b/multi_actris/channel/proofmode.v @@ -39,7 +39,7 @@ Class ProtoNormalize {Σ} (d : bool) (p : iProto Σ) ⊢ iProto_dual_if d p <++> foldr (iProto_app ∘ uncurry iProto_dual_if) END%proto pas ⊑ q. Global Hint Mode ProtoNormalize ! ! ! ! - : typeclass_instances. -Arguments ProtoNormalize {_} _ _%proto _%proto _%proto. +Arguments ProtoNormalize {_} _ _%_proto _%_proto _%_proto. Notation ProtoUnfold p1 p2 := (∀ d pas q, ProtoNormalize d p2 pas q → ProtoNormalize d p1 pas q). @@ -49,7 +49,7 @@ Class MsgNormalize {Σ} (d : bool) (m1 : iMsg Σ) msg_normalize a : ProtoNormalize d (<a> m1) pas (<(if d then action_dual a else a)> m2). Global Hint Mode MsgNormalize ! ! ! ! - : typeclass_instances. -Arguments MsgNormalize {_} _ _%msg _%msg _%msg. +Arguments MsgNormalize {_} _ _%_msg _%_msg _%_msg. Section classes. Context `{!chanG Σ, !heapGS Σ}. diff --git a/multi_actris/channel/proto.v b/multi_actris/channel/proto.v index 75efec7..aaae475 100644 --- a/multi_actris/channel/proto.v +++ b/multi_actris/channel/proto.v @@ -65,7 +65,7 @@ Next Obligation. solve_proper. Qed. Definition iMsg_base_aux : seal (@iMsg_base_def). by eexists. Qed. Definition iMsg_base := iMsg_base_aux.(unseal). Definition iMsg_base_eq : @iMsg_base = @iMsg_base_def := iMsg_base_aux.(seal_eq). -Arguments iMsg_base {_ _} _%V _%I _%proto. +Arguments iMsg_base {_ _} _%_V _%_I _%_proto. Global Instance: Params (@iMsg_base) 3 := {}. Program Definition iMsg_exist_def {Σ V A} (m : A → iMsg Σ V) : iMsg Σ V := @@ -74,12 +74,12 @@ Next Obligation. solve_proper. Qed. Definition iMsg_exist_aux : seal (@iMsg_exist_def). by eexists. Qed. Definition iMsg_exist := iMsg_exist_aux.(unseal). Definition iMsg_exist_eq : @iMsg_exist = @iMsg_exist_def := iMsg_exist_aux.(seal_eq). -Arguments iMsg_exist {_ _ _} _%msg. +Arguments iMsg_exist {_ _ _} _%_msg. Global Instance: Params (@iMsg_exist) 3 := {}. Definition iMsg_texist {Σ V} {TT : tele} (m : TT → iMsg Σ V) : iMsg Σ V := tele_fold (@iMsg_exist Σ V) (λ x, x) (tele_bind m). -Arguments iMsg_texist {_ _ !_} _%msg /. +Arguments iMsg_texist {_ _ !_} _%_msg /. Notation "'MSG' v {{ P } } ; p" := (iMsg_base v P p) (at level 200, v at level 20, right associativity, @@ -114,7 +114,7 @@ Definition iProto_message_aux : seal (@iProto_message_def). by eexists. Qed. Definition iProto_message := iProto_message_aux.(unseal). Definition iProto_message_eq : @iProto_message = @iProto_message_def := iProto_message_aux.(seal_eq). -Arguments iProto_message {_ _} _ _%msg. +Arguments iProto_message {_ _} _ _%_msg. Global Instance: Params (@iProto_message) 3 := {}. Notation "'END'" := iProto_end : proto_scope. @@ -166,7 +166,7 @@ Definition iProto_app_def {Σ V} (p1 p2 : iProto Σ V) : iProto Σ V := Definition iProto_app_aux : seal (@iProto_app_def). Proof. by eexists. Qed. Definition iProto_app := iProto_app_aux.(unseal). Definition iProto_app_eq : @iProto_app = @iProto_app_def := iProto_app_aux.(seal_eq). -Arguments iProto_app {_ _} _%proto _%proto. +Arguments iProto_app {_ _} _%_proto _%_proto. Global Instance: Params (@iProto_app) 2 := {}. Infix "<++>" := iProto_app (at level 60) : proto_scope. Notation "m <++> p" := (iMsg_map (flip iProto_app p) m) : msg_scope. @@ -177,13 +177,13 @@ Definition iProto_dual_aux : seal (@iProto_dual_def). Proof. by eexists. Qed. Definition iProto_dual := iProto_dual_aux.(unseal). Definition iProto_dual_eq : @iProto_dual = @iProto_dual_def := iProto_dual_aux.(seal_eq). -Arguments iProto_dual {_ _} _%proto. +Arguments iProto_dual {_ _} _%_proto. Global Instance: Params (@iProto_dual) 2 := {}. Notation iMsg_dual := (iMsg_map iProto_dual). Definition iProto_dual_if {Σ V} (d : bool) (p : iProto Σ V) : iProto Σ V := if d then iProto_dual p else p. -Arguments iProto_dual_if {_ _} _ _%proto. +Arguments iProto_dual_if {_ _} _ _%_proto. Global Instance: Params (@iProto_dual_if) 3 := {}. (** * Proofs *) @@ -487,7 +487,7 @@ Qed. Definition iProto_consistent {Σ V} (ps : list (iProto Σ V)) : iProp Σ := fixpoint iProto_consistent_pre' ps. -Arguments iProto_consistent {_ _} _%proto. +Arguments iProto_consistent {_ _} _%_proto. Global Instance: Params (@iProto_consistent) 1 := {}. Global Instance iProto_consistent_ne {Σ V} : NonExpansive (@iProto_consistent Σ V). @@ -530,7 +530,7 @@ Proof. Qed. Definition iProto_le {Σ V} (p1 p2 : iProto Σ V) : iProp Σ := fixpoint iProto_le_pre' p1 p2. -Arguments iProto_le {_ _} _%proto _%proto. +Arguments iProto_le {_ _} _%_proto _%_proto. Global Instance: Params (@iProto_le) 2 := {}. Notation "p ⊑ q" := (iProto_le p q) : bi_scope. @@ -566,7 +566,7 @@ Definition iProto_ctx `{protoG Σ V} Definition iProto_own `{!protoG Σ V} (γ : gname) (i : nat) (p : iProto Σ V) : iProp Σ := ∃ p', â–· (p' ⊑ p) ∗ iProto_own_frag γ i p'. -Arguments iProto_own {_ _ _} _ _ _%proto. +Arguments iProto_own {_ _ _} _ _ _%_proto. Global Instance: Params (@iProto_own) 3 := {}. Global Instance iProto_own_frag_contractive `{protoG Σ V} γ i : @@ -860,7 +860,7 @@ Section proto. { apply (gmap_view_alloc _ (length ps) (DfracOwn 1) (Excl' (Next p))); [|done|done]. rewrite fmap_map_seq. rewrite lookup_map_seq_0. - apply lookup_ge_None_2. rewrite fmap_length. done. } + apply lookup_ge_None_2. rewrite length_fmap. done. } simpl. iModIntro. rewrite right_id_L. @@ -965,7 +965,7 @@ Section proto. iNext. iDestruct ("IH" with "Hprot Hle [HSome]") as "HI". { rewrite list_lookup_insert; [done|]. - by rewrite insert_length. } + by rewrite length_insert. } iClear "IH Hm1 Hm2 Heq". rewrite list_insert_insert. rewrite (list_insert_commute _ j' i); [|done]. @@ -1266,7 +1266,7 @@ Section proto. rewrite list_lookup_insert_ne; [|done]. rewrite list_lookup_insert; [done|]. lia. } { rewrite list_lookup_insert_ne; [|done]. - rewrite list_lookup_insert; [done|]. rewrite insert_length. lia. } + rewrite list_lookup_insert; [done|]. rewrite length_insert. lia. } iMod (iProto_own_auth_update _ _ _ _ p2 with "Hauth Hj") as "[Hauth Hj]". iMod (iProto_own_auth_update _ _ _ _ p1 with "Hauth Hi") as "[Hauth Hi]". iIntros "!>!>". iExists p2. iFrame "Hm2". @@ -1274,7 +1274,7 @@ Section proto. iSplitL "Hconsistent Hauth". { iExists (<[i:=p1]> (<[j:=p2]> ps)). iSplit. - { iPureIntro. rewrite !insert_length. done. } + { iPureIntro. rewrite !length_insert. done. } iFrame. rewrite list_insert_insert. rewrite list_insert_commute; [|done]. rewrite list_insert_insert. by rewrite list_insert_commute; [|done]. } diff --git a/multi_actris/channel/proto_alt.v b/multi_actris/channel/proto_alt.v index 8e27dbf..1b4f380 100644 --- a/multi_actris/channel/proto_alt.v +++ b/multi_actris/channel/proto_alt.v @@ -109,7 +109,7 @@ Next Obligation. solve_proper. Qed. Definition iMsg_base_aux : seal (@iMsg_base_def). by eexists. Qed. Definition iMsg_base := iMsg_base_aux.(unseal). Definition iMsg_base_eq : @iMsg_base = @iMsg_base_def := iMsg_base_aux.(seal_eq). -Arguments iMsg_base {_ _} _%V _%I _%proto. +Arguments iMsg_base {_ _} _%_V _%_I _%_proto. Global Instance: Params (@iMsg_base) 3 := {}. Program Definition iMsg_exist_def {Σ V A} (m : A → iMsg Σ V) : iMsg Σ V := @@ -118,12 +118,12 @@ Next Obligation. solve_proper. Qed. Definition iMsg_exist_aux : seal (@iMsg_exist_def). by eexists. Qed. Definition iMsg_exist := iMsg_exist_aux.(unseal). Definition iMsg_exist_eq : @iMsg_exist = @iMsg_exist_def := iMsg_exist_aux.(seal_eq). -Arguments iMsg_exist {_ _ _} _%msg. +Arguments iMsg_exist {_ _ _} _%_msg. Global Instance: Params (@iMsg_exist) 3 := {}. Definition iMsg_texist {Σ V} {TT : tele} (m : TT → iMsg Σ V) : iMsg Σ V := tele_fold (@iMsg_exist Σ V) (λ x, x) (tele_bind m). -Arguments iMsg_texist {_ _ !_} _%msg /. +Arguments iMsg_texist {_ _ !_} _%_msg /. Notation "'MSG' v {{ P } } ; p" := (iMsg_base v P p) (at level 200, v at level 20, right associativity, @@ -158,7 +158,7 @@ Definition iProto_message_aux : seal (@iProto_message_def). by eexists. Qed. Definition iProto_message := iProto_message_aux.(unseal). Definition iProto_message_eq : @iProto_message = @iProto_message_def := iProto_message_aux.(seal_eq). -Arguments iProto_message {_ _} _ _%msg. +Arguments iProto_message {_ _} _ _%_msg. Global Instance: Params (@iProto_message) 3 := {}. Notation "'END'" := iProto_end : proto_scope. @@ -210,7 +210,7 @@ Definition iProto_app_def {Σ V} (p1 p2 : iProto Σ V) : iProto Σ V := Definition iProto_app_aux : seal (@iProto_app_def). Proof. by eexists. Qed. Definition iProto_app := iProto_app_aux.(unseal). Definition iProto_app_eq : @iProto_app = @iProto_app_def := iProto_app_aux.(seal_eq). -Arguments iProto_app {_ _} _%proto _%proto. +Arguments iProto_app {_ _} _%_proto _%_proto. Global Instance: Params (@iProto_app) 2 := {}. Infix "<++>" := iProto_app (at level 60) : proto_scope. Notation "m <++> p" := (iMsg_map (flip iProto_app p) m) : msg_scope. @@ -221,13 +221,13 @@ Definition iProto_dual_aux : seal (@iProto_dual_def). Proof. by eexists. Qed. Definition iProto_dual := iProto_dual_aux.(unseal). Definition iProto_dual_eq : @iProto_dual = @iProto_dual_def := iProto_dual_aux.(seal_eq). -Arguments iProto_dual {_ _} _%proto. +Arguments iProto_dual {_ _} _%_proto. Global Instance: Params (@iProto_dual) 2 := {}. Notation iMsg_dual := (iMsg_map iProto_dual). Definition iProto_dual_if {Σ V} (d : bool) (p : iProto Σ V) : iProto Σ V := if d then iProto_dual p else p. -Arguments iProto_dual_if {_ _} _ _%proto. +Arguments iProto_dual_if {_ _} _ _%_proto. Global Instance: Params (@iProto_dual_if) 3 := {}. (** * Proofs *) @@ -600,7 +600,7 @@ Qed. Definition iProto_consistent {Σ V} (ps : list (iProto Σ V)) : iProp Σ := fixpoint iProto_consistent_pre' ps. -Arguments iProto_consistent {_ _} _%proto. +Arguments iProto_consistent {_ _} _%_proto. Global Instance: Params (@iProto_consistent) 1 := {}. Global Instance iProto_consistent_ne {Σ V} : NonExpansive (@iProto_consistent Σ V). @@ -643,7 +643,7 @@ Proof. Qed. Definition iProto_le {Σ V} (p1 p2 : iProto Σ V) : iProp Σ := fixpoint iProto_le_pre' p1 p2. -Arguments iProto_le {_ _} _%proto _%proto. +Arguments iProto_le {_ _} _%_proto _%_proto. Global Instance: Params (@iProto_le) 2 := {}. Notation "p ⊑ q" := (iProto_le p q) : bi_scope. @@ -680,7 +680,7 @@ Definition iProto_ctx `{protoG Σ V} Definition iProto_own `{!protoG Σ V} (γ : gname) (i : nat) (p : iProto Σ V) : iProp Σ := ∃ p', â–· (p' ⊑ p) ∗ iProto_own_frag γ i p'. -Arguments iProto_own {_ _ _} _ _ _%proto. +Arguments iProto_own {_ _ _} _ _ _%_proto. Global Instance: Params (@iProto_own) 3 := {}. Global Instance iProto_own_frag_contractive `{protoG Σ V} γ i : diff --git a/multi_actris/utils/matrix.v b/multi_actris/utils/matrix.v index 166b7d2..6a64360 100644 --- a/multi_actris/utils/matrix.v +++ b/multi_actris/utils/matrix.v @@ -36,7 +36,7 @@ Section with_Σ. iDestruct "Hl" as "[Hl Hls]". iDestruct ("IHn" with "Hl") as "Hl". iFrame=> /=. - rewrite Nat.add_0_r !replicate_length. + rewrite Nat.add_0_r !length_replicate. replace (Z.of_nat (n * m)) with (Z.of_nat n * Z.of_nat m)%Z by lia. by iFrame. Qed. @@ -50,7 +50,7 @@ Section with_Σ. replace (S n) with (n + 1) by lia. rewrite !replicate_add /=. iDestruct "H" as "[H1 H2]". iSplitL "H1"; [by iApply "IHn"|]=> /=. - by rewrite !replicate_length. + by rewrite !length_replicate. Qed. Lemma array_to_matrix l m n v : -- GitLab From a041e7dfd5868a89d6f874914f0477c51956445c Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Thu, 23 Jan 2025 16:43:04 +0100 Subject: [PATCH 80/81] Renamed theories -> actris and let .opam package only publish actris --- _CoqProject | 92 +++++++++---------- {theories => actris}/channel/channel.v | 0 {theories => actris}/channel/proofmode.v | 0 {theories => actris}/channel/proto.v | 0 {theories => actris}/channel/proto_model.v | 0 {theories => actris}/examples/basics.v | 0 {theories => actris}/examples/equivalence.v | 0 {theories => actris}/examples/list_rev.v | 0 {theories => actris}/examples/map_reduce.v | 0 {theories => actris}/examples/par_map.v | 0 {theories => actris}/examples/pizza.v | 0 {theories => actris}/examples/rpc.v | 0 {theories => actris}/examples/sort.v | 0 {theories => actris}/examples/sort_br_del.v | 0 {theories => actris}/examples/sort_fg.v | 0 {theories => actris}/examples/subprotocols.v | 0 {theories => actris}/examples/swap_mapper.v | 0 {theories => actris}/logrel/contexts.v | 0 .../logrel/examples/choice_subtyping.v | 0 .../logrel/examples/compute_client_list.v | 0 .../logrel/examples/compute_service.v | 0 {theories => actris}/logrel/examples/mapper.v | 0 .../logrel/examples/mapper_list.v | 0 {theories => actris}/logrel/examples/pair.v | 0 .../logrel/examples/par_recv.v | 0 .../logrel/examples/rec_subtyping.v | 0 {theories => actris}/logrel/lib/list.v | 0 {theories => actris}/logrel/lib/mutex.v | 0 {theories => actris}/logrel/lib/par_start.v | 0 {theories => actris}/logrel/model.v | 0 {theories => actris}/logrel/napp.v | 0 {theories => actris}/logrel/operators.v | 0 {theories => actris}/logrel/session_types.v | 0 .../logrel/session_typing_rules.v | 0 {theories => actris}/logrel/subtyping.v | 0 {theories => actris}/logrel/subtyping_rules.v | 0 {theories => actris}/logrel/telescopes.v | 0 {theories => actris}/logrel/term_types.v | 0 .../logrel/term_typing_judgment.v | 0 .../logrel/term_typing_rules.v | 0 {theories => actris}/utils/cofe_solver_2.v | 0 {theories => actris}/utils/compare.v | 0 {theories => actris}/utils/contribution.v | 0 {theories => actris}/utils/group.v | 0 {theories => actris}/utils/llist.v | 0 {theories => actris}/utils/switch.v | 0 coq-actris.opam | 4 +- make-package | 32 +++++++ 48 files changed, 80 insertions(+), 48 deletions(-) rename {theories => actris}/channel/channel.v (100%) rename {theories => actris}/channel/proofmode.v (100%) rename {theories => actris}/channel/proto.v (100%) rename {theories => actris}/channel/proto_model.v (100%) rename {theories => actris}/examples/basics.v (100%) rename {theories => actris}/examples/equivalence.v (100%) rename {theories => actris}/examples/list_rev.v (100%) rename {theories => actris}/examples/map_reduce.v (100%) rename {theories => actris}/examples/par_map.v (100%) rename {theories => actris}/examples/pizza.v (100%) rename {theories => actris}/examples/rpc.v (100%) rename {theories => actris}/examples/sort.v (100%) rename {theories => actris}/examples/sort_br_del.v (100%) rename {theories => actris}/examples/sort_fg.v (100%) rename {theories => actris}/examples/subprotocols.v (100%) rename {theories => actris}/examples/swap_mapper.v (100%) rename {theories => actris}/logrel/contexts.v (100%) rename {theories => actris}/logrel/examples/choice_subtyping.v (100%) rename {theories => actris}/logrel/examples/compute_client_list.v (100%) rename {theories => actris}/logrel/examples/compute_service.v (100%) rename {theories => actris}/logrel/examples/mapper.v (100%) rename {theories => actris}/logrel/examples/mapper_list.v (100%) rename {theories => actris}/logrel/examples/pair.v (100%) rename {theories => actris}/logrel/examples/par_recv.v (100%) rename {theories => actris}/logrel/examples/rec_subtyping.v (100%) rename {theories => actris}/logrel/lib/list.v (100%) rename {theories => actris}/logrel/lib/mutex.v (100%) rename {theories => actris}/logrel/lib/par_start.v (100%) rename {theories => actris}/logrel/model.v (100%) rename {theories => actris}/logrel/napp.v (100%) rename {theories => actris}/logrel/operators.v (100%) rename {theories => actris}/logrel/session_types.v (100%) rename {theories => actris}/logrel/session_typing_rules.v (100%) rename {theories => actris}/logrel/subtyping.v (100%) rename {theories => actris}/logrel/subtyping_rules.v (100%) rename {theories => actris}/logrel/telescopes.v (100%) rename {theories => actris}/logrel/term_types.v (100%) rename {theories => actris}/logrel/term_typing_judgment.v (100%) rename {theories => actris}/logrel/term_typing_rules.v (100%) rename {theories => actris}/utils/cofe_solver_2.v (100%) rename {theories => actris}/utils/compare.v (100%) rename {theories => actris}/utils/contribution.v (100%) rename {theories => actris}/utils/group.v (100%) rename {theories => actris}/utils/llist.v (100%) rename {theories => actris}/utils/switch.v (100%) create mode 100755 make-package diff --git a/_CoqProject b/_CoqProject index 77a40ad..5baef5e 100644 --- a/_CoqProject +++ b/_CoqProject @@ -1,4 +1,4 @@ --Q theories actris +-Q actris actris -Q multi_actris multi_actris # We sometimes want to locally override notation, and there is no good way to do that with scopes. -arg -w -arg -notation-overridden @@ -6,51 +6,51 @@ # (https://github.com/coq/coq/issues/6294). -arg -w -arg -redundant-canonical-projection -theories/utils/llist.v -theories/utils/compare.v -theories/utils/contribution.v -theories/utils/group.v -theories/utils/cofe_solver_2.v -theories/utils/switch.v -theories/channel/proto_model.v -theories/channel/proto.v -theories/channel/channel.v -theories/channel/proofmode.v -theories/examples/basics.v -theories/examples/equivalence.v -theories/examples/sort.v -theories/examples/sort_br_del.v -theories/examples/sort_fg.v -theories/examples/par_map.v -theories/examples/map_reduce.v -theories/examples/swap_mapper.v -theories/examples/subprotocols.v -theories/examples/list_rev.v -theories/examples/rpc.v -theories/examples/pizza.v -theories/logrel/model.v -theories/logrel/telescopes.v -theories/logrel/subtyping.v -theories/logrel/contexts.v -theories/logrel/term_types.v -theories/logrel/session_types.v -theories/logrel/operators.v -theories/logrel/term_typing_judgment.v -theories/logrel/subtyping_rules.v -theories/logrel/term_typing_rules.v -theories/logrel/session_typing_rules.v -theories/logrel/napp.v -theories/logrel/lib/mutex.v -theories/logrel/lib/list.v -theories/logrel/lib/par_start.v -theories/logrel/examples/pair.v -theories/logrel/examples/par_recv.v -theories/logrel/examples/rec_subtyping.v -theories/logrel/examples/choice_subtyping.v -theories/logrel/examples/mapper.v -theories/logrel/examples/mapper_list.v -theories/logrel/examples/compute_service.v -theories/logrel/examples/compute_client_list.v +actris/utils/llist.v +actris/utils/compare.v +actris/utils/contribution.v +actris/utils/group.v +actris/utils/cofe_solver_2.v +actris/utils/switch.v +actris/channel/proto_model.v +actris/channel/proto.v +actris/channel/channel.v +actris/channel/proofmode.v +actris/examples/basics.v +actris/examples/equivalence.v +actris/examples/sort.v +actris/examples/sort_br_del.v +actris/examples/sort_fg.v +actris/examples/par_map.v +actris/examples/map_reduce.v +actris/examples/swap_mapper.v +actris/examples/subprotocols.v +actris/examples/list_rev.v +actris/examples/rpc.v +actris/examples/pizza.v +actris/logrel/model.v +actris/logrel/telescopes.v +actris/logrel/subtyping.v +actris/logrel/contexts.v +actris/logrel/term_types.v +actris/logrel/session_types.v +actris/logrel/operators.v +actris/logrel/term_typing_judgment.v +actris/logrel/subtyping_rules.v +actris/logrel/term_typing_rules.v +actris/logrel/session_typing_rules.v +actris/logrel/napp.v +actris/logrel/lib/mutex.v +actris/logrel/lib/list.v +actris/logrel/lib/par_start.v +actris/logrel/examples/pair.v +actris/logrel/examples/par_recv.v +actris/logrel/examples/rec_subtyping.v +actris/logrel/examples/choice_subtyping.v +actris/logrel/examples/mapper.v +actris/logrel/examples/mapper_list.v +actris/logrel/examples/compute_service.v +actris/logrel/examples/compute_client_list.v multi_actris/utils/cofe_solver_2.v multi_actris/utils/matrix.v diff --git a/theories/channel/channel.v b/actris/channel/channel.v similarity index 100% rename from theories/channel/channel.v rename to actris/channel/channel.v diff --git a/theories/channel/proofmode.v b/actris/channel/proofmode.v similarity index 100% rename from theories/channel/proofmode.v rename to actris/channel/proofmode.v diff --git a/theories/channel/proto.v b/actris/channel/proto.v similarity index 100% rename from theories/channel/proto.v rename to actris/channel/proto.v diff --git a/theories/channel/proto_model.v b/actris/channel/proto_model.v similarity index 100% rename from theories/channel/proto_model.v rename to actris/channel/proto_model.v diff --git a/theories/examples/basics.v b/actris/examples/basics.v similarity index 100% rename from theories/examples/basics.v rename to actris/examples/basics.v diff --git a/theories/examples/equivalence.v b/actris/examples/equivalence.v similarity index 100% rename from theories/examples/equivalence.v rename to actris/examples/equivalence.v diff --git a/theories/examples/list_rev.v b/actris/examples/list_rev.v similarity index 100% rename from theories/examples/list_rev.v rename to actris/examples/list_rev.v diff --git a/theories/examples/map_reduce.v b/actris/examples/map_reduce.v similarity index 100% rename from theories/examples/map_reduce.v rename to actris/examples/map_reduce.v diff --git a/theories/examples/par_map.v b/actris/examples/par_map.v similarity index 100% rename from theories/examples/par_map.v rename to actris/examples/par_map.v diff --git a/theories/examples/pizza.v b/actris/examples/pizza.v similarity index 100% rename from theories/examples/pizza.v rename to actris/examples/pizza.v diff --git a/theories/examples/rpc.v b/actris/examples/rpc.v similarity index 100% rename from theories/examples/rpc.v rename to actris/examples/rpc.v diff --git a/theories/examples/sort.v b/actris/examples/sort.v similarity index 100% rename from theories/examples/sort.v rename to actris/examples/sort.v diff --git a/theories/examples/sort_br_del.v b/actris/examples/sort_br_del.v similarity index 100% rename from theories/examples/sort_br_del.v rename to actris/examples/sort_br_del.v diff --git a/theories/examples/sort_fg.v b/actris/examples/sort_fg.v similarity index 100% rename from theories/examples/sort_fg.v rename to actris/examples/sort_fg.v diff --git a/theories/examples/subprotocols.v b/actris/examples/subprotocols.v similarity index 100% rename from theories/examples/subprotocols.v rename to actris/examples/subprotocols.v diff --git a/theories/examples/swap_mapper.v b/actris/examples/swap_mapper.v similarity index 100% rename from theories/examples/swap_mapper.v rename to actris/examples/swap_mapper.v diff --git a/theories/logrel/contexts.v b/actris/logrel/contexts.v similarity index 100% rename from theories/logrel/contexts.v rename to actris/logrel/contexts.v diff --git a/theories/logrel/examples/choice_subtyping.v b/actris/logrel/examples/choice_subtyping.v similarity index 100% rename from theories/logrel/examples/choice_subtyping.v rename to actris/logrel/examples/choice_subtyping.v diff --git a/theories/logrel/examples/compute_client_list.v b/actris/logrel/examples/compute_client_list.v similarity index 100% rename from theories/logrel/examples/compute_client_list.v rename to actris/logrel/examples/compute_client_list.v diff --git a/theories/logrel/examples/compute_service.v b/actris/logrel/examples/compute_service.v similarity index 100% rename from theories/logrel/examples/compute_service.v rename to actris/logrel/examples/compute_service.v diff --git a/theories/logrel/examples/mapper.v b/actris/logrel/examples/mapper.v similarity index 100% rename from theories/logrel/examples/mapper.v rename to actris/logrel/examples/mapper.v diff --git a/theories/logrel/examples/mapper_list.v b/actris/logrel/examples/mapper_list.v similarity index 100% rename from theories/logrel/examples/mapper_list.v rename to actris/logrel/examples/mapper_list.v diff --git a/theories/logrel/examples/pair.v b/actris/logrel/examples/pair.v similarity index 100% rename from theories/logrel/examples/pair.v rename to actris/logrel/examples/pair.v diff --git a/theories/logrel/examples/par_recv.v b/actris/logrel/examples/par_recv.v similarity index 100% rename from theories/logrel/examples/par_recv.v rename to actris/logrel/examples/par_recv.v diff --git a/theories/logrel/examples/rec_subtyping.v b/actris/logrel/examples/rec_subtyping.v similarity index 100% rename from theories/logrel/examples/rec_subtyping.v rename to actris/logrel/examples/rec_subtyping.v diff --git a/theories/logrel/lib/list.v b/actris/logrel/lib/list.v similarity index 100% rename from theories/logrel/lib/list.v rename to actris/logrel/lib/list.v diff --git a/theories/logrel/lib/mutex.v b/actris/logrel/lib/mutex.v similarity index 100% rename from theories/logrel/lib/mutex.v rename to actris/logrel/lib/mutex.v diff --git a/theories/logrel/lib/par_start.v b/actris/logrel/lib/par_start.v similarity index 100% rename from theories/logrel/lib/par_start.v rename to actris/logrel/lib/par_start.v diff --git a/theories/logrel/model.v b/actris/logrel/model.v similarity index 100% rename from theories/logrel/model.v rename to actris/logrel/model.v diff --git a/theories/logrel/napp.v b/actris/logrel/napp.v similarity index 100% rename from theories/logrel/napp.v rename to actris/logrel/napp.v diff --git a/theories/logrel/operators.v b/actris/logrel/operators.v similarity index 100% rename from theories/logrel/operators.v rename to actris/logrel/operators.v diff --git a/theories/logrel/session_types.v b/actris/logrel/session_types.v similarity index 100% rename from theories/logrel/session_types.v rename to actris/logrel/session_types.v diff --git a/theories/logrel/session_typing_rules.v b/actris/logrel/session_typing_rules.v similarity index 100% rename from theories/logrel/session_typing_rules.v rename to actris/logrel/session_typing_rules.v diff --git a/theories/logrel/subtyping.v b/actris/logrel/subtyping.v similarity index 100% rename from theories/logrel/subtyping.v rename to actris/logrel/subtyping.v diff --git a/theories/logrel/subtyping_rules.v b/actris/logrel/subtyping_rules.v similarity index 100% rename from theories/logrel/subtyping_rules.v rename to actris/logrel/subtyping_rules.v diff --git a/theories/logrel/telescopes.v b/actris/logrel/telescopes.v similarity index 100% rename from theories/logrel/telescopes.v rename to actris/logrel/telescopes.v diff --git a/theories/logrel/term_types.v b/actris/logrel/term_types.v similarity index 100% rename from theories/logrel/term_types.v rename to actris/logrel/term_types.v diff --git a/theories/logrel/term_typing_judgment.v b/actris/logrel/term_typing_judgment.v similarity index 100% rename from theories/logrel/term_typing_judgment.v rename to actris/logrel/term_typing_judgment.v diff --git a/theories/logrel/term_typing_rules.v b/actris/logrel/term_typing_rules.v similarity index 100% rename from theories/logrel/term_typing_rules.v rename to actris/logrel/term_typing_rules.v diff --git a/theories/utils/cofe_solver_2.v b/actris/utils/cofe_solver_2.v similarity index 100% rename from theories/utils/cofe_solver_2.v rename to actris/utils/cofe_solver_2.v diff --git a/theories/utils/compare.v b/actris/utils/compare.v similarity index 100% rename from theories/utils/compare.v rename to actris/utils/compare.v diff --git a/theories/utils/contribution.v b/actris/utils/contribution.v similarity index 100% rename from theories/utils/contribution.v rename to actris/utils/contribution.v diff --git a/theories/utils/group.v b/actris/utils/group.v similarity index 100% rename from theories/utils/group.v rename to actris/utils/group.v diff --git a/theories/utils/llist.v b/actris/utils/llist.v similarity index 100% rename from theories/utils/llist.v rename to actris/utils/llist.v diff --git a/theories/utils/switch.v b/actris/utils/switch.v similarity index 100% rename from theories/utils/switch.v rename to actris/utils/switch.v diff --git a/coq-actris.opam b/coq-actris.opam index 0f0153a..c94b261 100644 --- a/coq-actris.opam +++ b/coq-actris.opam @@ -11,5 +11,5 @@ depends: [ "coq-iris-heap-lang" { (= "dev.2024-10-30.0.27f837c1") | (= "dev") } ] -build: [make "-j%{jobs}%"] -install: [make "install"] +build: ["./make-package" "actris" "-j%{jobs}%"] +install: ["./make-package" "actris" "install"] diff --git a/make-package b/make-package new file mode 100755 index 0000000..5bf541b --- /dev/null +++ b/make-package @@ -0,0 +1,32 @@ +#!/bin/bash +set -e +# Helper script to build and/or install just one package out of this repository. +# Assumes that all the other packages it depends on have been installed already! + +PROJECT="$1" +shift + +COQFILE="_CoqProject.$PROJECT" +MAKEFILE="Makefile.package.$PROJECT" + +if ! grep -E -q "^$PROJECT/" _CoqProject; then + echo "No files in $PROJECT/ found in _CoqProject; this does not seem to be a valid project name." + exit 1 +fi + +# Generate _CoqProject file and Makefile +rm -f "$COQFILE" +# Get the right "-Q" line. +grep -E "^-Q $PROJECT[ /]" _CoqProject >> "$COQFILE" +# Get everything until the first empty line except for the "-Q" lines. +sed -n '/./!q;p' _CoqProject | grep -E -v "^-Q " >> "$COQFILE" +# Get the files. +grep -E "^$PROJECT/" _CoqProject >> "$COQFILE" +# Now we can run coq_makefile. +"${COQBIN}coq_makefile" -f "$COQFILE" -o "$MAKEFILE" + +# Run build +make -f "$MAKEFILE" "$@" + +# Cleanup +rm -f ".$MAKEFILE.d" "$MAKEFILE"* -- GitLab From 90edf73754629cef47a65a25c28f03dc0fa42364 Mon Sep 17 00:00:00 2001 From: Jonas Kastberg Hinrichsen <jihgfee@gmail.com> Date: Thu, 23 Jan 2025 17:02:45 +0100 Subject: [PATCH 81/81] Renamed multi_actris into multris --- _CoqProject | 22 +- multi_actris/channel/matrix.v | 221 --- multi_actris/channel/proto_alt.v | 1417 ----------------- multi_actris/examples/leader_election.v | 206 --- multi_actris/examples/leader_election_del.v | 276 ---- {multi_actris => multris}/channel/channel.v | 8 +- {multi_actris => multris}/channel/proofmode.v | 4 +- {multi_actris => multris}/channel/proto.v | 2 +- .../channel/proto_model.v | 2 +- {multi_actris => multris}/examples/basics.v | 2 +- .../examples/leader_election.v | 2 +- .../examples/three_buyer.v | 2 +- .../examples/two_buyer.v | 2 +- .../utils/cofe_solver_2.v | 0 {multi_actris => multris}/utils/matrix.v | 0 15 files changed, 23 insertions(+), 2143 deletions(-) delete mode 100644 multi_actris/channel/matrix.v delete mode 100644 multi_actris/channel/proto_alt.v delete mode 100644 multi_actris/examples/leader_election.v delete mode 100644 multi_actris/examples/leader_election_del.v rename {multi_actris => multris}/channel/channel.v (98%) rename {multi_actris => multris}/channel/proofmode.v (99%) rename {multi_actris => multris}/channel/proto.v (99%) rename {multi_actris => multris}/channel/proto_model.v (99%) rename {multi_actris => multris}/examples/basics.v (99%) rename multi_actris/examples/leader_election_del_alt.v => multris/examples/leader_election.v (99%) rename {multi_actris => multris}/examples/three_buyer.v (99%) rename {multi_actris => multris}/examples/two_buyer.v (98%) rename {multi_actris => multris}/utils/cofe_solver_2.v (100%) rename {multi_actris => multris}/utils/matrix.v (100%) diff --git a/_CoqProject b/_CoqProject index 5baef5e..1041167 100644 --- a/_CoqProject +++ b/_CoqProject @@ -1,5 +1,5 @@ -Q actris actris --Q multi_actris multi_actris +-Q multris multris # We sometimes want to locally override notation, and there is no good way to do that with scopes. -arg -w -arg -notation-overridden # Cannot use non-canonical projections as it causes massive unification failures @@ -52,13 +52,13 @@ actris/logrel/examples/mapper_list.v actris/logrel/examples/compute_service.v actris/logrel/examples/compute_client_list.v -multi_actris/utils/cofe_solver_2.v -multi_actris/utils/matrix.v -multi_actris/channel/proto_model.v -multi_actris/channel/proto.v -multi_actris/channel/channel.v -multi_actris/channel/proofmode.v -multi_actris/examples/basics.v -multi_actris/examples/two_buyer.v -multi_actris/examples/three_buyer.v -multi_actris/examples/leader_election.v +multris/utils/cofe_solver_2.v +multris/utils/matrix.v +multris/channel/proto_model.v +multris/channel/proto.v +multris/channel/channel.v +multris/channel/proofmode.v +multris/examples/basics.v +multris/examples/two_buyer.v +multris/examples/three_buyer.v +multris/examples/leader_election.v diff --git a/multi_actris/channel/matrix.v b/multi_actris/channel/matrix.v deleted file mode 100644 index 9b642da..0000000 --- a/multi_actris/channel/matrix.v +++ /dev/null @@ -1,221 +0,0 @@ -From iris.bi Require Import big_op. -From iris.heap_lang Require Export primitive_laws notation proofmode. - -Definition new_matrix : val := - λ: "m" "n" "v", (AllocN ("m"*"n") "v","m","n"). - -Definition pos (n i j : nat) : nat := i * n + j. -Definition vpos : val := λ: "n" "i" "j", "i"*"n" + "j". - -Definition matrix_get : val := - λ: "m" "i" "j", - let: "l" := Fst (Fst "m") in - let: "n" := Snd "m" in - "l" +â‚— vpos "n" "i" "j". - -Section with_Σ. - Context `{heapGS Σ}. - - Definition is_matrix (v : val) (m n : nat) - (P : nat → nat → loc → iProp Σ) : iProp Σ := - ∃ (l:loc), - ⌜v = PairV (PairV #l #m) #n⌠∗ - [∗ list] i ↦ _ ∈ replicate m (), - [∗ list] j ↦ _ ∈ replicate n (), - P i j (l +â‚— pos n i j). - - - Lemma array_to_matrix_pre l n m v : - l ↦∗ replicate (n * m) v -∗ - [∗ list] i ↦ _ ∈ replicate n (), (l +â‚— i*m) ↦∗ replicate m v. - Proof. - iIntros "Hl". - iInduction n as [|n] "IHn"; [done|]. - replace (S n) with (n + 1) by lia. - replace ((n + 1) * m) with (n * m + m) by lia. - rewrite !replicate_add /= array_app=> /=. - iDestruct "Hl" as "[Hl Hls]". - iDestruct ("IHn" with "Hl") as "Hl". - iFrame=> /=. - rewrite Nat.add_0_r !replicate_length. - replace (Z.of_nat (n * m)) with (Z.of_nat n * Z.of_nat m)%Z by lia. - by iFrame. - Qed. - - (* TODO: rename *) - Lemma big_sepL_replicate_type {A B} n (x1 : A) (x2 : B) (P : nat → iProp Σ) : - ([∗ list] i ↦ _ ∈ replicate n x1, P i) ⊢ - ([∗ list] i ↦ _ ∈ replicate n x2, P i). - Proof. - iIntros "H". - iInduction n as [|n] "IHn"; [done|]. - replace (S n) with (n + 1) by lia. - rewrite !replicate_add /=. iDestruct "H" as "[H1 H2]". - iSplitL "H1"; [by iApply "IHn"|]=> /=. - by rewrite !replicate_length. - Qed. - - Lemma array_to_matrix l m n v : - l ↦∗ replicate (m * n) v -∗ - [∗ list] i ↦ _ ∈ replicate m (), - [∗ list] j ↦ _ ∈ replicate n (), - (l +â‚— pos n i j) ↦ v. - Proof. - iIntros "Hl". - iDestruct (array_to_matrix_pre with "Hl") as "Hl". - iApply (big_sepL_impl with "Hl"). - iIntros "!>" (i ? _) "Hl". - iApply big_sepL_replicate_type. - iApply (big_sepL_impl with "Hl"). - iIntros "!>" (j ? HSome) "Hl". - rewrite Loc.add_assoc. - replace (Z.of_nat i * Z.of_nat n + Z.of_nat j)%Z with - (Z.of_nat (i * n + j))%Z by lia. - by apply lookup_replicate in HSome as [-> _]. - Qed. - - Lemma new_matrix_spec m n v P : - 0 < m → 0 < n → - {{{ [∗ list] i ↦ _ ∈ replicate m (), [∗ list] j ↦ _ ∈ replicate n (), - ∀ l, l ↦ v -∗ P i j l }}} - (* {{{ (∀ i j l, ⌜i < m⌠-∗ ⌜j < n⌠-∗ l ↦ v -∗ P i j l) }}} *) - new_matrix #m #n v - {{{ mat, RET mat; is_matrix mat m n P }}}. - Proof. - iIntros (Hm Hn Φ) "HP HΦ". - wp_lam. wp_pures. - wp_smart_apply wp_allocN; [lia|done|]. - iIntros (l) "[Hl _]". - wp_pures. iApply "HΦ". - iModIntro. - iExists _. iSplit; [done|]. - replace (Z.to_nat (Z.of_nat m * Z.of_nat n)) with (m * n) by lia. - iDestruct (array_to_matrix with "Hl") as "Hl". - iCombine "HP Hl" as "HPl". - rewrite -big_sepL_sep. - iApply (big_sepL_impl with "HPl"). - iIntros "!>" (k x Hin) "H". - rewrite -big_sepL_sep. - iApply (big_sepL_impl with "H"). - iIntros "!>" (k' x' Hin') "[HP Hl]". - by iApply "HP". - Qed. - - Lemma matrix_sep m n l (P : _ → _ → _ → iProp Σ) i j : - i < m → j < n → - ([∗ list] i ↦ _ ∈ replicate m (), - [∗ list] j ↦ _ ∈ replicate n (), - P i j (l +â‚— pos n i j)) -∗ - (P i j (l +â‚— pos n i j) ∗ - (P i j (l +â‚— pos n i j) -∗ - ([∗ list] i ↦ _ ∈ replicate m (), - [∗ list] j ↦ _ ∈ replicate n (), - P i j (l +â‚— pos n i j)))). - Proof. - iIntros (Hm Hn) "Hm". - iDestruct (big_sepL_impl with "Hm []") as "Hm". - { iIntros "!>" (k x HIn). - iApply (big_sepL_lookup_acc _ _ j ()). - by apply lookup_replicate. } - simpl. - rewrite {1}(big_sepL_lookup_acc _ (replicate m ()) i ()); - [|by apply lookup_replicate]. - iDestruct "Hm" as "[[Hij Hi] Hm]". - iFrame. - iIntros "HP". - iDestruct ("Hm" with "[$Hi $HP]") as "Hm". - iApply (big_sepL_impl with "Hm"). - iIntros "!>" (k x Hin). - iIntros "[Hm Hm']". iApply "Hm'". done. - Qed. - - Lemma matrix_sep_pers m n l (P : _ → _ → _ → iProp Σ) i j : - (∀ i j l, Persistent (P i j l)) → - i < m → j < n → - ([∗ list] i ↦ _ ∈ replicate m (), - [∗ list] j ↦ _ ∈ replicate n (), - P i j (l +â‚— pos n i j)) -∗ - (P i j (l +â‚— pos n i j) ∗ - (([∗ list] i ↦ _ ∈ replicate m (), - [∗ list] j ↦ _ ∈ replicate n (), - P i j (l +â‚— pos n i j)))). - Proof. - iIntros (Hpers Hm Hn) "Hm". - iDestruct (matrix_sep with "Hm") as "[#HP Hm]"; [done..|]. - iFrame "#". by iApply "Hm". - Qed. - - (* Lemma is_matrix_sep m n v P i j : *) - (* i < m → j < n → *) - (* is_matrix v m n P -∗ (∃ l, P i j l ∗ (P i j l -∗ is_matrix v m n P)). *) - (* Proof. *) - (* iDestruct 1 as (l ->) "Hm". *) - (* iDestruct (big_sepL_impl with "Hm []") as "Hm". *) - (* { iIntros "!>" (k x HIn). *) - (* iApply (big_sepL_lookup_acc _ _ j ()). *) - (* by apply lookup_replicate. } *) - (* simpl. *) - (* rewrite {1}(big_sepL_lookup_acc _ (replicate m ()) i ()); *) - (* [|by apply lookup_replicate]. *) - (* iDestruct "Hm" as "[[Hij Hi] Hm]". *) - (* iExists (l +â‚— pos n i j). *) - (* iFrame. *) - (* iIntros "HP". *) - (* iExists _. iSplit; [done|]. *) - (* iDestruct ("Hm" with "[$Hi $HP]") as "Hm". *) - (* iApply (big_sepL_impl with "Hm"). *) - (* iIntros "!>" (k x Hin). *) - (* iIntros "[Hm Hm']". iApply "Hm'". done. *) - (* Qed. *) - - Lemma vpos_spec (n i j : nat) : - {{{ True }}} vpos #n #i #j {{{ RET #(pos n i j); True }}}. - Proof. - iIntros (Φ) "_ HΦ". wp_lam. wp_pures. rewrite /pos. - replace (Z.of_nat i * Z.of_nat n + Z.of_nat j)%Z with - (Z.of_nat (i * n + j)) by lia. - by iApply "HΦ". - Qed. - - Lemma matrix_get_spec m n i j v P : - i < m → j < n → - {{{ is_matrix v m n P }}} - matrix_get v #i #j - {{{ l, RET #l; P i j l ∗ (P i j l -∗ is_matrix v m n P) }}}. - Proof. - iIntros (Hm Hn Φ) "Hm HΦ". - wp_lam. - iDestruct "Hm" as (l ->) "Hm". - wp_pures. - wp_smart_apply vpos_spec; [done|]. - iIntros "_". - wp_pures. - iApply "HΦ". - iModIntro. - iDestruct (matrix_sep _ _ _ _ i j with "Hm") as "[$ Hm]"; [lia|lia|]. - iIntros "H". - iDestruct ("Hm" with "H") as "Hm". - iExists _. iSplit; [done|]. iFrame. - Qed. - - Lemma matrix_get_spec_pers m n i j v P : - (∀ i j l, Persistent (P i j l)) → - i < m → j < n → - {{{ is_matrix v m n P }}} - matrix_get v #i #j - {{{ l, RET #l; P i j l ∗ is_matrix v m n P }}}. - Proof. - iIntros (Hpers Hm Hn Φ) "Hm HΦ". - wp_lam. - iDestruct "Hm" as (l ->) "Hm". - wp_pures. - wp_smart_apply vpos_spec; [done|]. - iIntros "_". - wp_pures. - iApply "HΦ". - iModIntro. - iDestruct (matrix_sep_pers _ _ _ _ i j with "Hm") as "[$ Hm]"; [lia|lia|]. - iExists _. iSplit; [done|]. iFrame. - Qed. - -End with_Σ. diff --git a/multi_actris/channel/proto_alt.v b/multi_actris/channel/proto_alt.v deleted file mode 100644 index 1b4f380..0000000 --- a/multi_actris/channel/proto_alt.v +++ /dev/null @@ -1,1417 +0,0 @@ -(** This file defines the core of the Actris logic: It defines dependent -separation protocols and the various operations on it like dual, append, and -branching. - -Dependent separation protocols [iProto] are defined by instantiating the -parameterized version in [proto_model] with the type of propositions [iProp] of Iris. -We define ways of constructing instances of the instantiated type via two -constructors: -- [iProto_end], which is identical to [proto_end]. -- [iProto_message], which takes an [action] and an [iMsg]. The type [iMsg] is a - sequence of binders [iMsg_exist], terminated by the payload constructed with - [iMsg_base] based on arguments [v], [P] and [prot], which are the value, the - carried proposition and the [iProto] tail, respectively. - -For convenience sake, we provide the following notations: -- [END], which is simply [iProto_end]. -- [∃ x, m], which is [iMsg_exist] with argument [m]. -- [MSG v {{ P }}; prot], which is [iMsg_Base] with arguments [v], [P] and [prot]. -- [<a> m], which is [iProto_message] with arguments [a] and [m]. - -We also include custom notation to more easily construct complete constructions: -- [<a x1 .. xn> m], which is [<a> ∃ x1, .. ∃ xn, m]. -- [<a x1 .. xn> MSG v; {{ P }}; prot], which constructs a full protocol. - -Futhermore, we define the following operations: -- [iProto_dual], which turns all [Send] of a protocol into [Recv] and vice-versa. -- [iProto_app], which appends two protocols. - -In addition we define the subprotocol relation [iProto_le] [⊑], which generalises -the following inductive definition for asynchronous subtyping on session types: - - p1 <: p2 p1 <: p2 p1 <: !B.p3 ?A.p3 <: p2 ----------- ---------------- ---------------- ---------------------------- -end <: end !A.p1 <: !A.p2 ?A.p1 <: ?A.p2 ?A.p1 <: !B.p2 - -Example: - -!R <: !R ?Q <: ?Q ?Q <: ?Q -------------------- -------------- -?Q.!R <: !R.?Q ?P.?Q <: ?P.?Q ------------------------------------- - ?P.?Q.!R <: !R.?P.?Q - -Lastly, relevant type classes instances are defined for each of the above -notions, such as contractiveness and non-expansiveness, after which the -specifications of the message-passing primitives are defined in terms of the -protocol connectives. *) -From iris.algebra Require Import gmap excl_auth gmap_view. -From iris.proofmode Require Import proofmode. -From iris.base_logic Require Export lib.iprop. -From iris.base_logic Require Import lib.own. -From iris.program_logic Require Import language. -From multi_actris.channel Require Import proto_model. -Set Default Proof Using "Type". -Export action. - -(** * Setup of Iris's cameras *) -Class protoG Σ V := - protoG_authG :: - inG Σ (gmap_viewR natO - (optionUR (exclR (laterO (proto (leibnizO V) (iPropO Σ) (iPropO Σ)))))). - -Definition protoΣ V := #[ - GFunctor ((gmap_viewRF natO (optionRF (exclRF (laterOF (protoOF (leibnizO V) idOF idOF)))))) -]. -Global Instance subG_chanΣ {Σ V} : subG (protoΣ V) Σ → protoG Σ V. -Proof. solve_inG. Qed. - -(** * Types *) -Definition iProto Σ V := proto V (iPropO Σ) (iPropO Σ). -Declare Scope proto_scope. -Delimit Scope proto_scope with proto. -Bind Scope proto_scope with iProto. -Local Open Scope proto. - -(** * Messages *) -Section iMsg. - Set Primitive Projections. - Record iMsg Σ V := IMsg { iMsg_car : V → laterO (iProto Σ V) -n> iPropO Σ }. -End iMsg. -Arguments IMsg {_ _} _. -Arguments iMsg_car {_ _} _. - -Declare Scope msg_scope. -Delimit Scope msg_scope with msg. -Bind Scope msg_scope with iMsg. -Global Instance iMsg_inhabited {Σ V} : Inhabited (iMsg Σ V) := populate (IMsg inhabitant). - -Section imsg_ofe. - Context {Σ : gFunctors} {V : Type}. - - Instance iMsg_equiv : Equiv (iMsg Σ V) := λ m1 m2, - ∀ w p, iMsg_car m1 w p ≡ iMsg_car m2 w p. - Instance iMsg_dist : Dist (iMsg Σ V) := λ n m1 m2, - ∀ w p, iMsg_car m1 w p ≡{n}≡ iMsg_car m2 w p. - - Lemma iMsg_ofe_mixin : OfeMixin (iMsg Σ V). - Proof. by apply (iso_ofe_mixin (iMsg_car : _ → V -d> _ -n> _)). Qed. - Canonical Structure iMsgO := Ofe (iMsg Σ V) iMsg_ofe_mixin. - - Global Instance iMsg_cofe : Cofe iMsgO. - Proof. by apply (iso_cofe (IMsg : (V -d> _ -n> _) → _) iMsg_car). Qed. -End imsg_ofe. - -Program Definition iMsg_base_def {Σ V} - (v : V) (P : iProp Σ) (p : iProto Σ V) : iMsg Σ V := - IMsg (λ v', λne p', ⌜ v = v' ⌠∗ Next p ≡ p' ∗ P)%I. -Next Obligation. solve_proper. Qed. -Definition iMsg_base_aux : seal (@iMsg_base_def). by eexists. Qed. -Definition iMsg_base := iMsg_base_aux.(unseal). -Definition iMsg_base_eq : @iMsg_base = @iMsg_base_def := iMsg_base_aux.(seal_eq). -Arguments iMsg_base {_ _} _%_V _%_I _%_proto. -Global Instance: Params (@iMsg_base) 3 := {}. - -Program Definition iMsg_exist_def {Σ V A} (m : A → iMsg Σ V) : iMsg Σ V := - IMsg (λ v', λne p', ∃ x, iMsg_car (m x) v' p')%I. -Next Obligation. solve_proper. Qed. -Definition iMsg_exist_aux : seal (@iMsg_exist_def). by eexists. Qed. -Definition iMsg_exist := iMsg_exist_aux.(unseal). -Definition iMsg_exist_eq : @iMsg_exist = @iMsg_exist_def := iMsg_exist_aux.(seal_eq). -Arguments iMsg_exist {_ _ _} _%_msg. -Global Instance: Params (@iMsg_exist) 3 := {}. - -Definition iMsg_texist {Σ V} {TT : tele} (m : TT → iMsg Σ V) : iMsg Σ V := - tele_fold (@iMsg_exist Σ V) (λ x, x) (tele_bind m). -Arguments iMsg_texist {_ _ !_} _%_msg /. - -Notation "'MSG' v {{ P } } ; p" := (iMsg_base v P p) - (at level 200, v at level 20, right associativity, - format "MSG v {{ P } } ; p") : msg_scope. -Notation "'MSG' v ; p" := (iMsg_base v True p) - (at level 200, v at level 20, right associativity, - format "MSG v ; p") : msg_scope. -Notation "∃ x .. y , m" := - (iMsg_exist (λ x, .. (iMsg_exist (λ y, m)) ..)%msg) : msg_scope. -Notation "'∃..' x .. y , m" := - (iMsg_texist (λ x, .. (iMsg_texist (λ y, m)) .. )%msg) - (at level 200, x binder, y binder, right associativity, - format "∃.. x .. y , m") : msg_scope. - -Lemma iMsg_texist_exist {Σ V} {TT : tele} w lp (m : TT → iMsg Σ V) : - iMsg_car (∃.. x, m x)%msg w lp ⊣⊢ (∃.. x, iMsg_car (m x) w lp). -Proof. - rewrite /iMsg_texist iMsg_exist_eq. - induction TT as [|T TT IH]; simpl; [done|]. f_equiv=> x. apply IH. -Qed. - -(** * Operators *) -Definition iProto_end_def {Σ V} : iProto Σ V := proto_end. -Definition iProto_end_aux : seal (@iProto_end_def). by eexists. Qed. -Definition iProto_end := iProto_end_aux.(unseal). -Definition iProto_end_eq : @iProto_end = @iProto_end_def := iProto_end_aux.(seal_eq). -Arguments iProto_end {_ _}. - -Definition iProto_message_def {Σ V} (a : action) (m : iMsg Σ V) : iProto Σ V := - proto_message a (iMsg_car m). -Definition iProto_message_aux : seal (@iProto_message_def). by eexists. Qed. -Definition iProto_message := iProto_message_aux.(unseal). -Definition iProto_message_eq : - @iProto_message = @iProto_message_def := iProto_message_aux.(seal_eq). -Arguments iProto_message {_ _} _ _%_msg. -Global Instance: Params (@iProto_message) 3 := {}. - -Notation "'END'" := iProto_end : proto_scope. - -Notation "< a > m" := (iProto_message a m) - (at level 200, a at level 10, m at level 200, - format "< a > m") : proto_scope. -Notation "< a @ x1 .. xn > m" := (iProto_message a (∃ x1, .. (∃ xn, m) ..)) - (at level 200, a at level 10, x1 closed binder, xn closed binder, m at level 200, - format "< a @ x1 .. xn > m") : proto_scope. -Notation "< a @.. x1 .. xn > m" := (iProto_message a (∃.. x1, .. (∃.. xn, m) ..)) - (at level 200, a at level 10, x1 closed binder, xn closed binder, m at level 200, - format "< a @.. x1 .. xn > m") : proto_scope. - -Class MsgTele {Σ V} {TT : tele} (m : iMsg Σ V) - (tv : TT -t> V) (tP : TT -t> iProp Σ) (tp : TT -t> iProto Σ V) := - msg_tele : m ≡ (∃.. x, MSG tele_app tv x {{ tele_app tP x }}; tele_app tp x)%msg. -Global Hint Mode MsgTele ! ! - ! - - - : typeclass_instances. - -(** * Operations *) -Program Definition iMsg_map {Σ V} - (rec : iProto Σ V → iProto Σ V) (m : iMsg Σ V) : iMsg Σ V := - IMsg (λ v, λne p1', ∃ p1, iMsg_car m v (Next p1) ∗ p1' ≡ Next (rec p1))%I. -Next Obligation. solve_proper. Qed. - -Program Definition iProto_map_app_aux {Σ V} - (f : action → action) (p2 : iProto Σ V) - (rec : iProto Σ V -n> iProto Σ V) : iProto Σ V -n> iProto Σ V := λne p, - proto_elim p2 (λ a m, - proto_message (f a) (iMsg_car (iMsg_map rec (IMsg m)))) p. -Next Obligation. - intros Σ V f p2 rec n p1 p1' Hp. apply proto_elim_ne=> // a m1 m2 Hm. - apply proto_message_ne=> v p' /=. by repeat f_equiv. -Qed. - -Global Instance iProto_map_app_aux_contractive {Σ V} f (p2 : iProto Σ V) : - Contractive (iProto_map_app_aux f p2). -Proof. - intros n rec1 rec2 Hrec p1; simpl. apply proto_elim_ne=> // a m1 m2 Hm. - apply proto_message_ne=> v p' /=. by repeat (f_contractive || f_equiv). -Qed. - -Definition iProto_map_app {Σ V} (f : action → action) - (p2 : iProto Σ V) : iProto Σ V -n> iProto Σ V := - fixpoint (iProto_map_app_aux f p2). - -Definition iProto_app_def {Σ V} (p1 p2 : iProto Σ V) : iProto Σ V := - iProto_map_app id p2 p1. -Definition iProto_app_aux : seal (@iProto_app_def). Proof. by eexists. Qed. -Definition iProto_app := iProto_app_aux.(unseal). -Definition iProto_app_eq : @iProto_app = @iProto_app_def := iProto_app_aux.(seal_eq). -Arguments iProto_app {_ _} _%_proto _%_proto. -Global Instance: Params (@iProto_app) 2 := {}. -Infix "<++>" := iProto_app (at level 60) : proto_scope. -Notation "m <++> p" := (iMsg_map (flip iProto_app p) m) : msg_scope. - -Definition iProto_dual_def {Σ V} (p : iProto Σ V) : iProto Σ V := - iProto_map_app action_dual proto_end p. -Definition iProto_dual_aux : seal (@iProto_dual_def). Proof. by eexists. Qed. -Definition iProto_dual := iProto_dual_aux.(unseal). -Definition iProto_dual_eq : - @iProto_dual = @iProto_dual_def := iProto_dual_aux.(seal_eq). -Arguments iProto_dual {_ _} _%_proto. -Global Instance: Params (@iProto_dual) 2 := {}. -Notation iMsg_dual := (iMsg_map iProto_dual). - -Definition iProto_dual_if {Σ V} (d : bool) (p : iProto Σ V) : iProto Σ V := - if d then iProto_dual p else p. -Arguments iProto_dual_if {_ _} _ _%_proto. -Global Instance: Params (@iProto_dual_if) 3 := {}. - -(** * Proofs *) -Section proto. - Context `{!protoG Σ V}. - Implicit Types v : V. - Implicit Types p pl pr : iProto Σ V. - Implicit Types m : iMsg Σ V. - - (** ** Equality *) - Lemma iProto_case p : p ≡ END ∨ ∃ t n m, p ≡ <(t,n)> m. - Proof. - rewrite iProto_message_eq iProto_end_eq. - destruct (proto_case p) as [|([a n]&m&?)]; [by left|right]. - by exists a, n, (IMsg m). - Qed. - Lemma iProto_message_equivI `{!BiInternalEq SPROP} a1 a2 m1 m2 : - (<a1> m1) ≡ (<a2> m2) ⊣⊢@{SPROP} ⌜ a1 = a2 ⌠∧ - (∀ v lp, iMsg_car m1 v lp ≡ iMsg_car m2 v lp). - Proof. rewrite iProto_message_eq. apply proto_message_equivI. Qed. - - Lemma iProto_message_end_equivI `{!BiInternalEq SPROP} a m : - (<a> m) ≡ END ⊢@{SPROP} False. - Proof. rewrite iProto_message_eq iProto_end_eq. apply proto_message_end_equivI. Qed. - Lemma iProto_end_message_equivI `{!BiInternalEq SPROP} a m : - END ≡ (<a> m) ⊢@{SPROP} False. - Proof. by rewrite internal_eq_sym iProto_message_end_equivI. Qed. - - (** ** Non-expansiveness of operators *) - Global Instance iMsg_car_proper : - Proper (iMsg_equiv ==> (=) ==> (≡) ==> (≡)) (iMsg_car (Σ:=Σ) (V:=V)). - Proof. - intros m1 m2 meq v1 v2 veq p1 p2 peq. rewrite meq. - f_equiv; [ by f_equiv | done ]. - Qed. - Global Instance iMsg_car_ne n : - Proper (iMsg_dist n ==> (=) ==> (dist n) ==> (dist n)) (iMsg_car (Σ:=Σ) (V:=V)). - Proof. - intros m1 m2 meq v1 v2 veq p1 p2 peq. rewrite meq. - f_equiv; [ by f_equiv | done ]. - Qed. - - Global Instance iMsg_contractive v n : - Proper (dist n ==> dist_later n ==> dist n) (iMsg_base (Σ:=Σ) (V:=V) v). - Proof. rewrite iMsg_base_eq=> P1 P2 HP p1 p2 Hp w q /=. solve_contractive. Qed. - Global Instance iMsg_ne v : NonExpansive2 (iMsg_base (Σ:=Σ) (V:=V) v). - Proof. rewrite iMsg_base_eq=> P1 P2 HP p1 p2 Hp w q /=. solve_proper. Qed. - Global Instance iMsg_proper v : - Proper ((≡) ==> (≡) ==> (≡)) (iMsg_base (Σ:=Σ) (V:=V) v). - Proof. apply (ne_proper_2 _). Qed. - - Global Instance iMsg_exist_ne A n : - Proper (pointwise_relation _ (dist n) ==> (dist n)) (@iMsg_exist Σ V A). - Proof. rewrite iMsg_exist_eq=> m1 m2 Hm v p /=. f_equiv=> x. apply Hm. Qed. - Global Instance iMsg_exist_proper A : - Proper (pointwise_relation _ (≡) ==> (≡)) (@iMsg_exist Σ V A). - Proof. rewrite iMsg_exist_eq=> m1 m2 Hm v p /=. f_equiv=> x. apply Hm. Qed. - - Global Instance msg_tele_base (v:V) (P : iProp Σ) (p : iProto Σ V) : - MsgTele (TT:=TeleO) (MSG v {{ P }}; p) v P p. - Proof. done. Qed. - Global Instance msg_tele_exist {A} {TT : A → tele} (m : A → iMsg Σ V) tv tP tp : - (∀ x, MsgTele (TT:=TT x) (m x) (tv x) (tP x) (tp x)) → - MsgTele (TT:=TeleS TT) (∃ x, m x) tv tP tp. - Proof. intros Hm. rewrite /MsgTele /=. f_equiv=> x. apply Hm. Qed. - - Global Instance iProto_message_ne a : - NonExpansive (iProto_message (Σ:=Σ) (V:=V) a). - Proof. rewrite iProto_message_eq. solve_proper. Qed. - Global Instance iProto_message_proper a : - Proper ((≡) ==> (≡)) (iProto_message (Σ:=Σ) (V:=V) a). - Proof. apply (ne_proper _). Qed. - - Lemma iProto_message_equiv {TT1 TT2 : tele} a1 a2 - (m1 m2 : iMsg Σ V) - (v1 : TT1 -t> V) (v2 : TT2 -t> V) - (P1 : TT1 -t> iProp Σ) (P2 : TT2 -t> iProp Σ) - (prot1 : TT1 -t> iProto Σ V) (prot2 : TT2 -t> iProto Σ V) : - MsgTele m1 v1 P1 prot1 → - MsgTele m2 v2 P2 prot2 → - ⌜ a1 = a2 ⌠-∗ - (■∀.. (xs1 : TT1), tele_app P1 xs1 -∗ - ∃.. (xs2 : TT2), ⌜tele_app v1 xs1 = tele_app v2 xs2⌠∗ - â–· (tele_app prot1 xs1 ≡ tele_app prot2 xs2) ∗ - tele_app P2 xs2) -∗ - (■∀.. (xs2 : TT2), tele_app P2 xs2 -∗ - ∃.. (xs1 : TT1), ⌜tele_app v1 xs1 = tele_app v2 xs2⌠∗ - â–· (tele_app prot1 xs1 ≡ tele_app prot2 xs2) ∗ - tele_app P1 xs1) -∗ - (<a1> m1) ≡ (<a2> m2). - Proof. - iIntros (Hm1 Hm2 Heq) "#Heq1 #Heq2". - unfold MsgTele in Hm1. rewrite Hm1. clear Hm1. - unfold MsgTele in Hm2. rewrite Hm2. clear Hm2. - rewrite iProto_message_eq proto_message_equivI. - iSplit; [ done | ]. - iIntros (v p'). - do 2 rewrite iMsg_texist_exist. - rewrite iMsg_base_eq /=. - iApply prop_ext. - iIntros "!>". iSplit. - - iDestruct 1 as (xs1 Hveq1) "[Hrec1 HP1]". - iDestruct ("Heq1" with "HP1") as (xs2 Hveq2) "[Hrec2 HP2]". - iExists xs2. rewrite -Hveq1 Hveq2. - iSplitR; [ done | ]. iSplitR "HP2"; [ | done ]. - iRewrite -"Hrec1". iApply later_equivI. iIntros "!>". by iRewrite "Hrec2". - - iDestruct 1 as (xs2 Hveq2) "[Hrec2 HP2]". - iDestruct ("Heq2" with "HP2") as (xs1 Hveq1) "[Hrec1 HP1]". - iExists xs1. rewrite -Hveq2 Hveq1. - iSplitR; [ done | ]. iSplitR "HP1"; [ | done ]. - iRewrite -"Hrec2". iApply later_equivI. iIntros "!>". by iRewrite "Hrec1". - Qed. - - (** Helpers *) - Lemma iMsg_map_base f v P p : - NonExpansive f → - iMsg_map f (MSG v {{ P }}; p) ≡ (MSG v {{ P }}; f p)%msg. - Proof. - rewrite iMsg_base_eq. intros ? v' p'; simpl. iSplit. - - iDestruct 1 as (p'') "[(->&Hp&$) Hp']". iSplit; [done|]. - iRewrite "Hp'". iIntros "!>". by iRewrite "Hp". - - iIntros "(->&Hp'&$)". iExists p. iRewrite -"Hp'". auto. - Qed. - Lemma iMsg_map_exist {A} f (m : A → iMsg Σ V) : - iMsg_map f (∃ x, m x) ≡ (∃ x, iMsg_map f (m x))%msg. - Proof. - rewrite iMsg_exist_eq. intros v' p'; simpl. iSplit. - - iDestruct 1 as (p'') "[H Hp']". iDestruct "H" as (x) "H"; auto. - - iDestruct 1 as (x p'') "[Hm Hp']". auto. - Qed. - - (** ** Dual *) - Global Instance iProto_dual_ne : NonExpansive (@iProto_dual Σ V). - Proof. rewrite iProto_dual_eq. solve_proper. Qed. - Global Instance iProto_dual_proper : Proper ((≡) ==> (≡)) (@iProto_dual Σ V). - Proof. apply (ne_proper _). Qed. - Global Instance iProto_dual_if_ne d : NonExpansive (@iProto_dual_if Σ V d). - Proof. solve_proper. Qed. - Global Instance iProto_dual_if_proper d : - Proper ((≡) ==> (≡)) (@iProto_dual_if Σ V d). - Proof. apply (ne_proper _). Qed. - - Lemma iProto_dual_end : iProto_dual (Σ:=Σ) (V:=V) END ≡ END. - Proof. - rewrite iProto_end_eq iProto_dual_eq /iProto_dual_def /iProto_map_app. - etrans; [apply (fixpoint_unfold (iProto_map_app_aux _ _))|]; simpl. - by rewrite proto_elim_end. - Qed. - Lemma iProto_dual_message a m : - iProto_dual (<a> m) ≡ <action_dual a> iMsg_dual m. - Proof. - rewrite iProto_message_eq iProto_dual_eq /iProto_dual_def /iProto_map_app. - etrans; [apply (fixpoint_unfold (iProto_map_app_aux _ _))|]; simpl. - rewrite /iProto_message_def. rewrite ->proto_elim_message; [done|]. - intros a' m1 m2 Hm; f_equiv; solve_proper. - Qed. - Lemma iMsg_dual_base v P p : - iMsg_dual (MSG v {{ P }}; p) ≡ (MSG v {{ P }}; iProto_dual p)%msg. - Proof. apply iMsg_map_base, _. Qed. - Lemma iMsg_dual_exist {A} (m : A → iMsg Σ V) : - iMsg_dual (∃ x, m x) ≡ (∃ x, iMsg_dual (m x))%msg. - Proof. apply iMsg_map_exist. Qed. - - Global Instance iProto_dual_involutive : Involutive (≡) (@iProto_dual Σ V). - Proof. - intros p. apply (uPred.internal_eq_soundness (M:=iResUR Σ)). - iLöb as "IH" forall (p). destruct (iProto_case p) as [->|(a&n&m&->)]. - { by rewrite !iProto_dual_end. } - rewrite !iProto_dual_message involutive. - iApply iProto_message_equivI; iSplit; [done|]; iIntros (v p') "/=". - iApply prop_ext; iIntros "!>"; iSplit. - - iDestruct 1 as (pd) "[H Hp']". iRewrite "Hp'". - iDestruct "H" as (pdd) "[H #Hpd]". - iApply (internal_eq_rewrite); [|done]; iIntros "!>". - iRewrite "Hpd". by iRewrite ("IH" $! pdd). - - iIntros "H". destruct (Next_uninj p') as [p'' Hp']. iExists _. - rewrite Hp'. iSplitL; [by auto|]. iIntros "!>". by iRewrite ("IH" $! p''). - Qed. - - (** ** Append *) - Global Instance iProto_app_end_l : LeftId (≡) END (@iProto_app Σ V). - Proof. - intros p. rewrite iProto_end_eq iProto_app_eq /iProto_app_def /iProto_map_app. - etrans; [apply (fixpoint_unfold (iProto_map_app_aux _ _))|]; simpl. - by rewrite proto_elim_end. - Qed. - Lemma iProto_app_message a m p2 : (<a> m) <++> p2 ≡ <a> m <++> p2. - Proof. - rewrite iProto_message_eq iProto_app_eq /iProto_app_def /iProto_map_app. - etrans; [apply (fixpoint_unfold (iProto_map_app_aux _ _))|]; simpl. - rewrite /iProto_message_def. rewrite ->proto_elim_message; [done|]. - intros a' m1 m2 Hm. f_equiv; solve_proper. - Qed. - - Global Instance iProto_app_ne : NonExpansive2 (@iProto_app Σ V). - Proof. - assert (∀ n, Proper (dist n ==> (=) ==> dist n) (@iProto_app Σ V)). - { intros n p1 p1' Hp1 p2 p2' <-. by rewrite iProto_app_eq /iProto_app_def Hp1. } - assert (Proper ((≡) ==> (=) ==> (≡)) (@iProto_app Σ V)). - { intros p1 p1' Hp1 p2 p2' <-. by rewrite iProto_app_eq /iProto_app_def Hp1. } - intros n p1 p1' Hp1 p2 p2' Hp2. rewrite Hp1. clear p1 Hp1. - revert p1'. induction (lt_wf n) as [n _ IH]; intros p1. - destruct (iProto_case p1) as [->|(a&i&m&->)]. - { by rewrite !left_id. } - rewrite !iProto_app_message. f_equiv=> v p' /=. do 4 f_equiv. - f_contractive. apply IH; eauto using dist_le with lia. - Qed. - Global Instance iProto_app_proper : Proper ((≡) ==> (≡) ==> (≡)) (@iProto_app Σ V). - Proof. apply (ne_proper_2 _). Qed. - - Lemma iMsg_app_base v P p1 p2 : - ((MSG v {{ P }}; p1) <++> p2)%msg ≡ (MSG v {{ P }}; p1 <++> p2)%msg. - Proof. apply: iMsg_map_base. Qed. - Lemma iMsg_app_exist {A} (m : A → iMsg Σ V) p2 : - ((∃ x, m x) <++> p2)%msg ≡ (∃ x, m x <++> p2)%msg. - Proof. apply iMsg_map_exist. Qed. - - Global Instance iProto_app_end_r : RightId (≡) END (@iProto_app Σ V). - Proof. - intros p. apply (uPred.internal_eq_soundness (M:=iResUR Σ)). - iLöb as "IH" forall (p). destruct (iProto_case p) as [->|(a&i&m&->)]. - { by rewrite left_id. } - rewrite iProto_app_message. - iApply iProto_message_equivI; iSplit; [done|]; iIntros (v p') "/=". - iApply prop_ext; iIntros "!>". iSplit. - - iDestruct 1 as (p1') "[H Hp']". iRewrite "Hp'". - iApply (internal_eq_rewrite); [|done]; iIntros "!>". - by iRewrite ("IH" $! p1'). - - iIntros "H". destruct (Next_uninj p') as [p'' Hp']. iExists p''. - rewrite Hp'. iSplitL; [by auto|]. iIntros "!>". by iRewrite ("IH" $! p''). - Qed. - Global Instance iProto_app_assoc : Assoc (≡) (@iProto_app Σ V). - Proof. - intros p1 p2 p3. apply (uPred.internal_eq_soundness (M:=iResUR Σ)). - iLöb as "IH" forall (p1). destruct (iProto_case p1) as [->|(a&i&m&->)]. - { by rewrite !left_id. } - rewrite !iProto_app_message. - iApply iProto_message_equivI; iSplit; [done|]; iIntros (v p123) "/=". - iApply prop_ext; iIntros "!>". iSplit. - - iDestruct 1 as (p1') "[H #Hp']". - iExists (p1' <++> p2). iSplitL; [by auto|]. - iRewrite "Hp'". iIntros "!>". iApply "IH". - - iDestruct 1 as (p12) "[H #Hp123]". iDestruct "H" as (p1') "[H #Hp12]". - iExists p1'. iFrame "H". iRewrite "Hp123". - iIntros "!>". iRewrite "Hp12". by iRewrite ("IH" $! p1'). - Qed. - - Lemma iProto_dual_app p1 p2 : - iProto_dual (p1 <++> p2) ≡ iProto_dual p1 <++> iProto_dual p2. - Proof. - apply (uPred.internal_eq_soundness (M:=iResUR Σ)). - iLöb as "IH" forall (p1 p2). destruct (iProto_case p1) as [->|(a&i&m&->)]. - { by rewrite iProto_dual_end !left_id. } - rewrite iProto_dual_message !iProto_app_message iProto_dual_message /=. - iApply iProto_message_equivI; iSplit; [done|]; iIntros (v p12) "/=". - iApply prop_ext; iIntros "!>". iSplit. - - iDestruct 1 as (p12d) "[H #Hp12]". iDestruct "H" as (p1') "[H #Hp12d]". - iExists (iProto_dual p1'). iSplitL; [by auto|]. - iRewrite "Hp12". iIntros "!>". iRewrite "Hp12d". iApply "IH". - - iDestruct 1 as (p1') "[H #Hp12]". iDestruct "H" as (p1d) "[H #Hp1']". - iExists (p1d <++> p2). iSplitL; [by auto|]. - iRewrite "Hp12". iIntros "!>". iRewrite "Hp1'". by iRewrite ("IH" $! p1d p2). - Qed. - -End proto. - -Global Instance iProto_inhabited {Σ V} : Inhabited (iProto Σ V) := populate END. - -Program Definition iProto_elim {Σ V A} - (x : A) (f : action → iMsg Σ V -> A) (p : iProto Σ V) : A := - proto_elim x (λ a m, f a (IMsg (λ v, λne p, m v p)))%I p. -Next Obligation. solve_proper. Qed. - -Lemma iProto_elim_message {Σ V} {A:ofe} - (x : A) (f : action → iMsg Σ V -> A) a m - `{Hf : ∀ a, Proper ((≡) ==> (≡)) (f a)} : - iProto_elim x f (iProto_message a m) ≡ f a m. - Proof. - rewrite /iProto_elim iProto_message_eq /iProto_message_def /=. - setoid_rewrite proto_elim_message. - { apply Hf. done. } - intros a' f1 f2 Hf'. apply Hf=> v p /=. apply Hf'. -Qed. - -Lemma iProto_elim_message' {Σ V} {A:ofe} - (x : A) (f : action → iMsg Σ V -> A) a m - `{Hf : ∀ a, Proper ((≡) ==> (≡)) (f a)} : - iProto_elim x f (iProto_message a m) ≡ f a m. -Proof. - rewrite /iProto_elim. - rewrite iProto_message_eq /iProto_message_def. simpl. - setoid_rewrite proto_elim_message. - { f_equiv. done. } - intros a'. - intros f1 f2 Hf'. f_equiv. done. -Qed. - -Definition nat_beq := Eval compute in Nat.eqb. - -Definition find_recv {Σ V} (i:nat) (j:nat) (ps : list (iProto Σ V)) : - option $ iMsg Σ V := - iProto_elim None (λ a m, - match a with - | (Recv, i') => if nat_beq i i' then Some m else None - | (Send, _) => None - end) (ps !!! j). - -Fixpoint sync_pairs_aux {Σ V} (i : nat) (ps_full ps : list (iProto Σ V)) : - list (nat * nat * iMsg Σ V * iMsg Σ V) := - match ps with - | [] => [] - | p :: ps => - iProto_elim (sync_pairs_aux (S i) ps_full ps) - (λ a mi, match a with - | (Recv,_) => sync_pairs_aux (S i) ps_full ps - | (Send,j) => match find_recv i j ps_full with - | None => sync_pairs_aux (S i) ps_full ps - | Some mj => (i,j,mi,mj) :: - sync_pairs_aux (S i) ps_full ps - end - end) p - end. - -Notation sync_pairs ps := (sync_pairs_aux 0 ps ps). - -Definition can_step {Σ V} (rec : list (iProto Σ V) → iProp Σ) - (ps : list (iProto Σ V)) : iProp Σ := - [∧ list] '(i,j,m1,m2) ∈ sync_pairs ps, - ∀ v p1, iMsg_car m1 v (Next p1) -∗ - ∃ p2, iMsg_car m2 v (Next p2) ∗ - â–· (rec (<[i:=p1]>(<[j:=p2]>ps))). - -From iris.heap_lang Require Import notation. - -Definition iProto_binary `{!heapGS Σ} : list (iProto Σ val) := - [(<(Send, 1) @ (x:Z)> MSG #x ; END)%proto; - (<(Recv, 0) @ (x:Z)> MSG #x ; END)%proto]. - -Lemma iProto_binary_consistent `{!heapGS Σ} : - ⊢ can_step (λ _, True) (@iProto_binary _ Σ heapGS). -Proof. rewrite /iProto_binary /can_step /iProto_elim. simpl. - rewrite /find_recv. simpl. - Fail rewrite iProto_elim_message. - (* OBS: Break here *) - - -Definition valid_target {Σ V} (ps : list (iProto Σ V)) (i j : nat) : iProp Σ := - ∀ a m, (ps !!! i ≡ <(a, j)> m) -∗ ⌜is_Some (ps !! j)âŒ. - -Definition iProto_consistent_pre {Σ V} (rec : list (iProto Σ V) → iProp Σ) - (ps : list (iProto Σ V)) : iProp Σ := - (∀ i j, valid_target ps i j) ∗ (can_step rec ps). - -Global Instance iProto_consistent_pre_ne {Σ V} - (rec : listO (iProto Σ V) -n> iPropO Σ) : - NonExpansive (iProto_consistent_pre rec). -Proof. rewrite /iProto_consistent_pre /can_step /valid_target. solve_proper. Qed. - -Program Definition iProto_consistent_pre' {Σ V} - (rec : listO (iProto Σ V) -n> iPropO Σ) : - listO (iProto Σ V) -n> iPropO Σ := - λne ps, iProto_consistent_pre (λ ps, rec ps) ps. - -Local Instance iProto_consistent_pre_contractive {Σ V} : Contractive (@iProto_consistent_pre' Σ V). -Proof. - rewrite /iProto_consistent_pre' /iProto_consistent_pre /can_step. - solve_contractive. -Qed. - -Definition iProto_consistent {Σ V} (ps : list (iProto Σ V)) : iProp Σ := - fixpoint iProto_consistent_pre' ps. - -Arguments iProto_consistent {_ _} _%_proto. -Global Instance: Params (@iProto_consistent) 1 := {}. - -Global Instance iProto_consistent_ne {Σ V} : NonExpansive (@iProto_consistent Σ V). -Proof. solve_proper. Qed. -Global Instance iProto_consistent_proper {Σ V} : Proper ((≡) ==> (⊣⊢)) (@iProto_consistent Σ V). -Proof. solve_proper. Qed. - -Lemma iProto_consistent_unfold {Σ V} (ps : list (iProto Σ V)) : - iProto_consistent ps ≡ (iProto_consistent_pre iProto_consistent) ps. -Proof. - apply: (fixpoint_unfold iProto_consistent_pre'). -Qed. - -(** * Protocol entailment *) -Definition iProto_le_pre {Σ V} - (rec : iProto Σ V → iProto Σ V → iProp Σ) (p1 p2 : iProto Σ V) : iProp Σ := - (p1 ≡ END ∗ p2 ≡ END) ∨ - ∃ a1 a2 m1 m2, - (p1 ≡ <a1> m1) ∗ (p2 ≡ <a2> m2) ∗ - match a1, a2 with - | (Recv,i), (Recv,j) => ⌜i = j⌠∗ ∀ v p1', - iMsg_car m1 v (Next p1') -∗ ∃ p2', â–· rec p1' p2' ∗ iMsg_car m2 v (Next p2') - | (Send,i), (Send,j) => ⌜i = j⌠∗ ∀ v p2', - iMsg_car m2 v (Next p2') -∗ ∃ p1', â–· rec p1' p2' ∗ iMsg_car m1 v (Next p1') - | _, _ => False - end. -Global Instance iProto_le_pre_ne {Σ V} (rec : iProto Σ V → iProto Σ V → iProp Σ) : - NonExpansive2 (iProto_le_pre rec). -Proof. solve_proper. Qed. - -Program Definition iProto_le_pre' {Σ V} - (rec : iProto Σ V -n> iProto Σ V -n> iPropO Σ) : - iProto Σ V -n> iProto Σ V -n> iPropO Σ := λne p1 p2, - iProto_le_pre (λ p1' p2', rec p1' p2') p1 p2. -Solve Obligations with solve_proper. -Local Instance iProto_le_pre_contractive {Σ V} : Contractive (@iProto_le_pre' Σ V). -Proof. - intros n rec1 rec2 Hrec p1 p2. rewrite /iProto_le_pre' /iProto_le_pre /=. - by repeat (f_contractive || f_equiv). -Qed. -Definition iProto_le {Σ V} (p1 p2 : iProto Σ V) : iProp Σ := - fixpoint iProto_le_pre' p1 p2. -Arguments iProto_le {_ _} _%_proto _%_proto. -Global Instance: Params (@iProto_le) 2 := {}. -Notation "p ⊑ q" := (iProto_le p q) : bi_scope. - -Global Instance iProto_le_ne {Σ V} : NonExpansive2 (@iProto_le Σ V). -Proof. solve_proper. Qed. -Global Instance iProto_le_proper {Σ V} : Proper ((≡) ==> (≡) ==> (⊣⊢)) (@iProto_le Σ V). -Proof. solve_proper. Qed. - -Record proto_name := ProtName { proto_names : gmap nat gname }. -Global Instance proto_name_inhabited : Inhabited proto_name := - populate (ProtName inhabitant). -Global Instance proto_name_eq_dec : EqDecision proto_name. -Proof. solve_decision. Qed. -Global Instance proto_name_countable : Countable proto_name. -Proof. - refine (inj_countable (λ '(ProtName γs), (γs)) - (λ '(γs), Some (ProtName γs)) _); by intros []. -Qed. - -Definition iProto_own_frag `{!protoG Σ V} (γ : gname) - (i : nat) (p : iProto Σ V) : iProp Σ := - own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p))). - -Definition iProto_own_auth `{!protoG Σ V} (γ : gname) - (ps : list (iProto Σ V)) : iProp Σ := - own γ (gmap_view_auth (DfracOwn 1) (((λ p, Excl' (Next p)) <$> - (list_to_map (zip (seq 0 (length ps)) ps))) : gmap _ _)). - -Definition iProto_ctx `{protoG Σ V} - (γ : gname) (ps_len : nat) : iProp Σ := - ∃ ps, ⌜length ps = ps_len⌠∗ iProto_own_auth γ ps ∗ â–· iProto_consistent ps. - -(** * The connective for ownership of channel ends *) -Definition iProto_own `{!protoG Σ V} - (γ : gname) (i : nat) (p : iProto Σ V) : iProp Σ := - ∃ p', â–· (p' ⊑ p) ∗ iProto_own_frag γ i p'. -Arguments iProto_own {_ _ _} _ _ _%_proto. -Global Instance: Params (@iProto_own) 3 := {}. - -Global Instance iProto_own_frag_contractive `{protoG Σ V} γ i : - Contractive (iProto_own_frag γ i). -Proof. solve_contractive. Qed. - -Global Instance iProto_own_contractive `{protoG Σ V} γ i : - Contractive (iProto_own γ i). -Proof. solve_contractive. Qed. -Global Instance iProto_own_ne `{protoG Σ V} γ s : NonExpansive (iProto_own γ s). -Proof. solve_proper. Qed. -Global Instance iProto_own_proper `{protoG Σ V} γ s : - Proper ((≡) ==> (≡)) (iProto_own γ s). -Proof. apply (ne_proper _). Qed. - -(** * Proofs *) -Section proto. - Context `{!protoG Σ V}. - Implicit Types v : V. - Implicit Types p pl pr : iProto Σ V. - Implicit Types m : iMsg Σ V. - - Lemma own_prot_idx γ i j (p1 p2 : iProto Σ V) : - own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p1))) -∗ - own γ (gmap_view_frag j (DfracOwn 1) (Excl' (Next p2))) -∗ - ⌜i ≠jâŒ. - Proof. - iIntros "Hown Hown'" (->). - iDestruct (own_valid_2 with "Hown Hown'") as "H". - rewrite uPred.cmra_valid_elim. - by iDestruct "H" as %[]%gmap_view_frag_op_validN. - Qed. - - Lemma own_prot_excl γ i (p1 p2 : iProto Σ V) : - own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p1))) -∗ - own γ (gmap_view_frag i (DfracOwn 1) (Excl' (Next p2))) -∗ - False. - Proof. iIntros "Hi Hj". by iDestruct (own_prot_idx with "Hi Hj") as %?. Qed. - - (** ** Protocol entailment **) - Lemma iProto_le_unfold p1 p2 : iProto_le p1 p2 ≡ iProto_le_pre iProto_le p1 p2. - Proof. apply: (fixpoint_unfold iProto_le_pre'). Qed. - - Lemma iProto_le_end : ⊢ END ⊑ (END : iProto Σ V). - Proof. rewrite iProto_le_unfold. iLeft. auto 10. Qed. - - Lemma iProto_le_end_inv_r p : p ⊑ END -∗ (p ≡ END). - Proof. - rewrite iProto_le_unfold. iIntros "[[Hp _]|H] //". - iDestruct "H" as (a1 a2 m1 m2) "(_ & Heq & _)". - by iDestruct (iProto_end_message_equivI with "Heq") as %[]. - Qed. - - Lemma iProto_le_end_inv_l p : END ⊑ p -∗ (p ≡ END). - Proof. - rewrite iProto_le_unfold. iIntros "[[_ Hp]|H] //". - iDestruct "H" as (a1 a2 m1 m2) "(Heq & _ & _)". - iDestruct (iProto_end_message_equivI with "Heq") as %[]. - Qed. - - Lemma iProto_le_send_inv i p1 m2 : - p1 ⊑ (<(Send,i)> m2) -∗ ∃ m1, - (p1 ≡ <(Send,i)> m1) ∗ - ∀ v p2', iMsg_car m2 v (Next p2') -∗ - ∃ p1', â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1'). - Proof. - rewrite iProto_le_unfold. - iIntros "[[_ Heq]|H]". - { by iDestruct (iProto_message_end_equivI with "Heq") as %[]. } - iDestruct "H" as (a1 a2 m1 m2') "(Hp1 & Hp2 & H)". - rewrite iProto_message_equivI. iDestruct "Hp2" as "[%Heq Hm2]". - simplify_eq. - destruct a1 as [[]]; [|done]. - iDestruct "H" as (->) "H". iExists m1. iFrame "Hp1". - iIntros (v p2). iSpecialize ("Hm2" $! v (Next p2)). by iRewrite "Hm2". - Qed. - - Lemma iProto_le_send_send_inv i m1 m2 v p2' : - (<(Send,i)> m1) ⊑ (<(Send,i)> m2) -∗ - iMsg_car m2 v (Next p2') -∗ ∃ p1', â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1'). - Proof. - iIntros "H Hm2". iDestruct (iProto_le_send_inv with "H") as (m1') "[Hm1 H]". - iDestruct (iProto_message_equivI with "Hm1") as (Heq) "Hm1". - iDestruct ("H" with "Hm2") as (p1') "[Hle Hm]". - iRewrite -("Hm1" $! v (Next p1')) in "Hm". auto with iFrame. - Qed. - - Lemma iProto_le_recv_inv_l i m1 p2 : - (<(Recv,i)> m1) ⊑ p2 -∗ ∃ m2, - (p2 ≡ <(Recv,i)> m2) ∗ - ∀ v p1', iMsg_car m1 v (Next p1') -∗ - ∃ p2', â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). - Proof. - rewrite iProto_le_unfold. - iIntros "[[Heq _]|H]". - { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } - iDestruct "H" as (a1 a2 m1' m2) "(Hp1 & Hp2 & H)". - rewrite iProto_message_equivI. iDestruct "Hp1" as "[%Heq Hm1]". - simplify_eq. - destruct a2 as [[]]; [done|]. - iDestruct "H" as (->) "H". iExists m2. iFrame "Hp2". - iIntros (v p1). iSpecialize ("Hm1" $! v (Next p1)). by iRewrite "Hm1". - Qed. - - Lemma iProto_le_recv_inv_r i p1 m2 : - (p1 ⊑ <(Recv,i)> m2) -∗ ∃ m1, - (p1 ≡ <(Recv,i)> m1) ∗ - ∀ v p1', iMsg_car m1 v (Next p1') -∗ - ∃ p2', â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). - Proof. - rewrite iProto_le_unfold. - iIntros "[[_ Heq]|H]". - { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } - iDestruct "H" as (a1 a2 m1 m2') "(Hp1 & Hp2 & H)". - rewrite iProto_message_equivI. - iDestruct "Hp2" as "[%Heq Hm2]". - simplify_eq. - destruct a1 as [[]]; [done|]. - iDestruct "H" as (->) "H". - iExists m1. iFrame. - iIntros (v p2). - iIntros "Hm1". iDestruct ("H" with "Hm1") as (p2') "[Hle H]". - iSpecialize ("Hm2" $! v (Next p2')). - iExists p2'. iFrame. - iRewrite "Hm2". iApply "H". - Qed. - - Lemma iProto_le_recv_recv_inv i m1 m2 v p1' : - (<(Recv, i)> m1) ⊑ (<(Recv, i)> m2) -∗ - iMsg_car m1 v (Next p1') -∗ ∃ p2', â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2'). - Proof. - iIntros "H Hm2". iDestruct (iProto_le_recv_inv_r with "H") as (m1') "[Hm1 H]". - iApply "H". iDestruct (iProto_message_equivI with "Hm1") as (_) "Hm1". - by iRewrite -("Hm1" $! v (Next p1')). - Qed. - - Lemma iProto_le_msg_inv_l i a m1 p2 : - (<(a,i)> m1) ⊑ p2 -∗ ∃ m2, p2 ≡ <(a,i)> m2. - Proof. - rewrite iProto_le_unfold /iProto_le_pre. - iIntros "[[Heq _]|H]". - { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } - iDestruct "H" as (a1 a2 m1' m2) "(Hp1 & Hp2 & H)". - destruct a1 as [t1 ?], a2 as [t2 ?]. - destruct t1,t2; [|done|done|]. - - rewrite iProto_message_equivI. - iDestruct "Hp1" as (Heq) "Hp1". simplify_eq. - iDestruct "H" as (->) "H". by iExists _. - - rewrite iProto_message_equivI. - iDestruct "Hp1" as (Heq) "Hp1". simplify_eq. - iDestruct "H" as (->) "H". by iExists _. - Qed. - - Lemma iProto_le_msg_inv_r j a p1 m2 : - (p1 ⊑ <(a,j)> m2) -∗ ∃ m1, p1 ≡ <(a,j)> m1. - Proof. - rewrite iProto_le_unfold /iProto_le_pre. - iIntros "[[_ Heq]|H]". - { iDestruct (iProto_message_end_equivI with "Heq") as %[]. } - iDestruct "H" as (a1 a2 m1 m2') "(Hp1 & Hp2 & H)". - destruct a1 as [t1 ?], a2 as [t2 ?]. - destruct t1,t2; [|done|done|]. - - rewrite iProto_message_equivI. - iDestruct "Hp2" as (Heq) "Hp2". simplify_eq. - iDestruct "H" as (->) "H". by iExists _. - - rewrite iProto_message_equivI. - iDestruct "Hp2" as (Heq) "Hp2". simplify_eq. - iDestruct "H" as (->) "H". by iExists _. - Qed. - - Lemma valid_target_le ps i p1 p2 : - (∀ i' j', valid_target ps i' j') -∗ - ps !!! i ≡ p1 -∗ - p1 ⊑ p2 -∗ - (∀ i' j', valid_target (<[i := p2]>ps) i' j') ∗ p1 ⊑ p2. - Proof. Admitted. - (* iIntros "Hprot #HSome Hle". *) - (* pose proof (iProto_case p1) as [Hend|Hmsg]. *) - (* { rewrite Hend. iDestruct (iProto_le_end_inv_l with "Hle") as "#H". *) - (* iFrame "Hle". *) - (* iIntros (i' j' a m) "Hm". *) - (* destruct (decide (i = j')) as [->|Hneqj]. *) - (* { Search list_lookup_total insert. rewrite list_lookup_total_insert. ; [done|]. lia. done. } *) - (* rewrite lookup_insert_ne; [|done]. *) - (* destruct (decide (i = i')) as [->|Hneqi]. *) - (* { rewrite lookup_total_insert. iRewrite "H" in "Hm". *) - (* by iDestruct (iProto_end_message_equivI with "Hm") as "Hm". } *) - (* rewrite lookup_total_insert_ne; [|done]. *) - (* by iApply "Hprot". } *) - (* destruct Hmsg as (t & n & m & Hmsg). *) - (* setoid_rewrite Hmsg. *) - (* iDestruct (iProto_le_msg_inv_l with "Hle") as (m2) "#Heq". iFrame "Hle". *) - (* iIntros (i' j' a m') "Hm". *) - (* destruct (decide (i = j')) as [->|Hneqj]. *) - (* { rewrite lookup_insert. done. } *) - (* rewrite lookup_insert_ne; [|done]. *) - (* destruct (decide (i = i')) as [->|Hneqi]. *) - (* { rewrite lookup_total_insert. iRewrite "Heq" in "Hm". *) - (* iDestruct (iProto_message_equivI with "Hm") as (Heq) "Hm". *) - (* simplify_eq. by iApply "Hprot". } *) - (* rewrite lookup_total_insert_ne; [|done]. *) - (* by iApply "Hprot". *) - (* Qed. *) - - Lemma iProto_consistent_le ps i p1 p2 : - iProto_consistent ps -∗ - ps !!! i ≡ p1 -∗ - p1 ⊑ p2 -∗ - iProto_consistent (<[i := p2]>ps). - Proof. - iIntros "Hprot #HSome Hle". - iRevert "HSome". - iLöb as "IH" forall (p1 p2 ps). - iIntros "#HSome". - rewrite !iProto_consistent_unfold. - iDestruct "Hprot" as "(Htar & Hprot)". - iDestruct (valid_target_le with "Htar HSome Hle") as "[Htar Hle]". - iFrame. - iIntros (i' j' m1 m2) "#Hm1 #Hm2". - destruct (decide (i = i')) as [<-|Hneq]. - { rewrite list_lookup_total_insert; [|admit]. - pose proof (iProto_case p2) as [Hend|Hmsg]. - { setoid_rewrite Hend. rewrite iProto_end_message_equivI. done. } - destruct Hmsg as (a&?&m&Hmsg). - setoid_rewrite Hmsg. - destruct a; last first. - { rewrite iProto_message_equivI. - iDestruct "Hm1" as "[%Htag Hm1]". done. } - rewrite iProto_message_equivI. - iDestruct "Hm1" as "[%Htag Hm1]". - inversion Htag. simplify_eq. - iIntros (v p) "Hm1'". - iSpecialize ("Hm1" $! v (Next p)). - iDestruct (iProto_le_send_inv with "Hle") as "Hle". - iRewrite -"Hm1" in "Hm1'". - iDestruct "Hle" as (m') "[#Heq H]". - iDestruct ("H" with "Hm1'") as (p') "[Hle H]". - destruct (decide (i = j')) as [<-|Hneq]. - { rewrite list_lookup_total_insert. rewrite iProto_message_equivI. - iDestruct "Hm2" as "[%Heq _]". done. admit. } - iDestruct ("Hprot" $!i j' with "[] [] H") as "Hprot". - { iRewrite -"Heq". rewrite !list_lookup_total_alt. iRewrite "HSome". done. } - { rewrite list_lookup_total_insert_ne; [|done]. done. } - iDestruct "Hprot" as (p'') "[Hm Hprot]". - iExists p''. iFrame. - iNext. - iDestruct ("IH" with "Hprot Hle [HSome]") as "HI". - { rewrite list_lookup_total_insert; [done|]. admit. } - iClear "IH Hm1 Hm2 Heq". - rewrite list_insert_insert. - rewrite (list_insert_commute _ j' i); [|done]. - rewrite list_insert_insert. done. } - rewrite list_lookup_total_insert_ne; [|done]. - destruct (decide (i = j')) as [<-|Hneq']. - { rewrite list_lookup_total_insert. - pose proof (iProto_case p2) as [Hend|Hmsg]. - { setoid_rewrite Hend. rewrite iProto_end_message_equivI. done. } - destruct Hmsg as (a&?&m&Hmsg). - setoid_rewrite Hmsg. - destruct a. - { rewrite iProto_message_equivI. - iDestruct "Hm2" as "[%Htag Hm2]". done. } - rewrite iProto_message_equivI. - iDestruct "Hm2" as "[%Htag Hm2]". - inversion Htag. simplify_eq. - iIntros (v p) "Hm1'". - iDestruct (iProto_le_recv_inv_r with "Hle") as "Hle". - iDestruct "Hle" as (m') "[#Heq Hle]". - iDestruct ("Hprot" $!i' with "[] [] Hm1'") as "Hprot". - { done. } - { rewrite !list_lookup_total_alt. iRewrite "HSome". done. } - iDestruct ("Hprot") as (p') "[Hm1' Hprot]". - iDestruct ("Hle" with "Hm1'") as (p2') "[Hle Hm']". - iSpecialize ("Hm2" $! v (Next p2')). - iExists p2'. - iRewrite -"Hm2". iFrame. - iDestruct ("IH" with "Hprot Hle []") as "HI". - { iPureIntro. rewrite list_lookup_total_insert_ne; [|done]. - rewrite list_lookup_total_insert. done. admit. } - rewrite list_insert_commute; [|done]. - rewrite !list_insert_insert. done. admit. } - rewrite list_lookup_total_insert_ne; [|done]. - iIntros (v p) "Hm1'". - iDestruct ("Hprot" $!i' j' with "[//] [//] Hm1'") as "Hprot". - iDestruct "Hprot" as (p') "[Hm2' Hprot]". - iExists p'. iFrame. - iNext. - rewrite (list_insert_commute _ j' i); [|done]. - rewrite (list_insert_commute _ i' i); [|done]. - iApply ("IH" with "Hprot Hle []"). - rewrite list_lookup_total_insert_ne; [|done]. - rewrite list_lookup_total_insert_ne; [|done]. - done. - Admitted. - - Lemma iProto_le_send i m1 m2 : - (∀ v p2', iMsg_car m2 v (Next p2') -∗ ∃ p1', - â–· (p1' ⊑ p2') ∗ iMsg_car m1 v (Next p1')) -∗ - (<(Send,i)> m1) ⊑ (<(Send,i)> m2). - Proof. - iIntros "Hle". rewrite iProto_le_unfold. - iRight. iExists (Send, i), (Send, i), m1, m2. by eauto. - Qed. - - Lemma iProto_le_recv i m1 m2 : - (∀ v p1', iMsg_car m1 v (Next p1') -∗ ∃ p2', - â–· (p1' ⊑ p2') ∗ iMsg_car m2 v (Next p2')) -∗ - (<(Recv,i)> m1) ⊑ (<(Recv,i)> m2). - Proof. - iIntros "Hle". rewrite iProto_le_unfold. - iRight. iExists (Recv, i), (Recv, i), m1, m2. by eauto. - Qed. - - Lemma iProto_le_base a v P p1 p2 : - â–· (p1 ⊑ p2) -∗ - (<a> MSG v {{ P }}; p1) ⊑ (<a> MSG v {{ P }}; p2). - Proof. - rewrite iMsg_base_eq. iIntros "H". destruct a as [[]]. - - iApply iProto_le_send. iIntros (v' p') "(->&Hp&$)". - iExists p1. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". - - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&$)". - iExists p2. iSplit; [|by auto]. iIntros "!>". by iRewrite -"Hp". - Qed. - - Lemma iProto_le_trans p1 p2 p3 : p1 ⊑ p2 -∗ p2 ⊑ p3 -∗ p1 ⊑ p3. - Proof. - iIntros "H1 H2". iLöb as "IH" forall (p1 p2 p3). - destruct (iProto_case p3) as [->|([]&i&m3&->)]. - - iDestruct (iProto_le_end_inv_r with "H2") as "H2". by iRewrite "H2" in "H1". - - iDestruct (iProto_le_send_inv with "H2") as (m2) "[Hp2 H2]". - iRewrite "Hp2" in "H1"; clear p2. - iDestruct (iProto_le_send_inv with "H1") as (m1) "[Hp1 H1]". - iRewrite "Hp1"; clear p1. - iApply iProto_le_send. iIntros (v p3') "Hm3". - iDestruct ("H2" with "Hm3") as (p2') "[Hle Hm2]". - iDestruct ("H1" with "Hm2") as (p1') "[Hle' Hm1]". - iExists p1'. iIntros "{$Hm1} !>". by iApply ("IH" with "Hle'"). - - iDestruct (iProto_le_recv_inv_r with "H2") as (m2) "[Hp2 H3]". - iRewrite "Hp2" in "H1". - iDestruct (iProto_le_recv_inv_r with "H1") as (m1) "[Hp1 H2]". - iRewrite "Hp1". iApply iProto_le_recv. iIntros (v p1') "Hm1". - iDestruct ("H2" with "Hm1") as (p2') "[Hle Hm2]". - iDestruct ("H3" with "Hm2") as (p3') "[Hle' Hm3]". - iExists p3'. iIntros "{$Hm3} !>". by iApply ("IH" with "Hle"). - Qed. - - Lemma iProto_le_refl p : ⊢ p ⊑ p. - Proof. - iLöb as "IH" forall (p). destruct (iProto_case p) as [->|([]&i&m&->)]. - - iApply iProto_le_end. - - iApply iProto_le_send. auto 10 with iFrame. - - iApply iProto_le_recv. auto 10 with iFrame. - Qed. - - Lemma iProto_le_dual p1 p2 : p2 ⊑ p1 -∗ iProto_dual p1 ⊑ iProto_dual p2. - Proof. - iIntros "H". iLöb as "IH" forall (p1 p2). - destruct (iProto_case p1) as [->|([]&i&m1&->)]. - - iDestruct (iProto_le_end_inv_r with "H") as "H". - iRewrite "H". iApply iProto_le_refl. - - iDestruct (iProto_le_send_inv with "H") as (m2) "[Hp2 H]". - iRewrite "Hp2"; clear p2. iEval (rewrite !iProto_dual_message). - iApply iProto_le_recv. iIntros (v p1d). - iDestruct 1 as (p1') "[Hm1 #Hp1d]". - iDestruct ("H" with "Hm1") as (p2') "[H Hm2]". - iDestruct ("IH" with "H") as "H". iExists (iProto_dual p2'). - iSplitL "H"; [iIntros "!>"; by iRewrite "Hp1d"|]. simpl; auto. - - iDestruct (iProto_le_recv_inv_r with "H") as (m2) "[Hp2 H]". - iRewrite "Hp2"; clear p2. iEval (rewrite !iProto_dual_message /=). - iApply iProto_le_send. iIntros (v p2d). - iDestruct 1 as (p2') "[Hm2 #Hp2d]". - iDestruct ("H" with "Hm2") as (p1') "[H Hm1]". - iDestruct ("IH" with "H") as "H". iExists (iProto_dual p1'). - iSplitL "H"; [iIntros "!>"; by iRewrite "Hp2d"|]. simpl; auto. - Qed. - - Lemma iProto_le_dual_l p1 p2 : iProto_dual p2 ⊑ p1 ⊢ iProto_dual p1 ⊑ p2. - Proof. - iIntros "H". iEval (rewrite -(involutive iProto_dual p2)). - by iApply iProto_le_dual. - Qed. - Lemma iProto_le_dual_r p1 p2 : p2 ⊑ iProto_dual p1 ⊢ p1 ⊑ iProto_dual p2. - Proof. - iIntros "H". iEval (rewrite -(involutive iProto_dual p1)). - by iApply iProto_le_dual. - Qed. - - Lemma iProto_le_app p1 p2 p3 p4 : - p1 ⊑ p2 -∗ p3 ⊑ p4 -∗ p1 <++> p3 ⊑ p2 <++> p4. - Proof. - iIntros "H1 H2". iLöb as "IH" forall (p1 p2 p3 p4). - destruct (iProto_case p2) as [->|([]&i&m2&->)]. - - iDestruct (iProto_le_end_inv_r with "H1") as "H1". - iRewrite "H1". by rewrite !left_id. - - iDestruct (iProto_le_send_inv with "H1") as (m1) "[Hp1 H1]". - iRewrite "Hp1"; clear p1. rewrite !iProto_app_message. - iApply iProto_le_send. iIntros (v p24). - iDestruct 1 as (p2') "[Hm2 #Hp24]". - iDestruct ("H1" with "Hm2") as (p1') "[H1 Hm1]". - iExists (p1' <++> p3). iSplitR "Hm1"; [|by simpl; eauto]. - iIntros "!>". iRewrite "Hp24". by iApply ("IH" with "H1"). - - iDestruct (iProto_le_recv_inv_r with "H1") as (m1) "[Hp1 H1]". - iRewrite "Hp1"; clear p1. rewrite !iProto_app_message. - iApply iProto_le_recv. - iIntros (v p13). iDestruct 1 as (p1') "[Hm1 #Hp13]". - iDestruct ("H1" with "Hm1") as (p2'') "[H1 Hm2]". - iExists (p2'' <++> p4). iSplitR "Hm2"; [|by simpl; eauto]. - iIntros "!>". iRewrite "Hp13". by iApply ("IH" with "H1"). - Qed. - - Lemma iProto_le_payload_elim_l i m v P p : - (P -∗ (<(Recv,i)> MSG v; p) ⊑ (<(Recv,i)> m)) ⊢ - (<(Recv,i)> MSG v {{ P }}; p) ⊑ <(Recv,i)> m. - Proof. - rewrite iMsg_base_eq. iIntros "H". - iApply iProto_le_recv. iIntros (v' p') "(->&Hp&HP)". - iApply (iProto_le_recv_recv_inv with "(H HP)"); simpl; auto. - Qed. - Lemma iProto_le_payload_elim_r i m v P p : - (P -∗ (<(Send, i)> m) ⊑ (<(Send, i)> MSG v; p)) ⊢ - (<(Send,i)> m) ⊑ (<(Send,i)> MSG v {{ P }}; p). - Proof. - rewrite iMsg_base_eq. iIntros "H". - iApply iProto_le_send. iIntros (v' p') "(->&Hp&HP)". - iApply (iProto_le_send_send_inv with "(H HP)"); simpl; auto. - Qed. - Lemma iProto_le_payload_intro_l i v P p : - P -∗ (<(Send,i)> MSG v {{ P }}; p) ⊑ (<(Send,i)> MSG v; p). - Proof. - rewrite iMsg_base_eq. - iIntros "HP". iApply iProto_le_send. iIntros (v' p') "(->&Hp&_) /=". - iExists p'. iSplitR; [iApply iProto_le_refl|]. auto. - Qed. - Lemma iProto_le_payload_intro_r i v P p : - P -∗ (<(Recv,i)> MSG v; p) ⊑ (<(Recv,i)> MSG v {{ P }}; p). - Proof. - rewrite iMsg_base_eq. - iIntros "HP". iApply iProto_le_recv. iIntros (v' p') "(->&Hp&_) /=". - iExists p'. iSplitR; [iApply iProto_le_refl|]. auto. - Qed. - Lemma iProto_le_exist_elim_l {A} i (m1 : A → iMsg Σ V) m2 : - (∀ x, (<(Recv,i)> m1 x) ⊑ (<(Recv,i)> m2)) ⊢ - (<(Recv,i) @ x> m1 x) ⊑ (<(Recv,i)> m2). - Proof. - rewrite iMsg_exist_eq. iIntros "H". - iApply iProto_le_recv. iIntros (v p1') "/=". iDestruct 1 as (x) "Hm". - by iApply (iProto_le_recv_recv_inv with "H"). - Qed. - Lemma iProto_le_exist_elim_r {A} i m1 (m2 : A → iMsg Σ V) : - (∀ x, (<(Send,i)> m1) ⊑ (<(Send,i)> m2 x)) ⊢ - (<(Send,i)> m1) ⊑ (<(Send,i) @ x> m2 x). - Proof. - rewrite iMsg_exist_eq. iIntros "H". - iApply iProto_le_send. iIntros (v p2'). iDestruct 1 as (x) "Hm". - by iApply (iProto_le_send_send_inv with "H"). - Qed. - Lemma iProto_le_exist_intro_l {A} i (m : A → iMsg Σ V) a : - ⊢ (<(Send,i) @ x> m x) ⊑ (<(Send,i)> m a). - Proof. - rewrite iMsg_exist_eq. iApply iProto_le_send. iIntros (v p') "Hm /=". - iExists p'. iSplitR; last by auto. iApply iProto_le_refl. - Qed. - Lemma iProto_le_exist_intro_r {A} i (m : A → iMsg Σ V) a : - ⊢ (<(Recv,i)> m a) ⊑ (<(Recv,i) @ x> m x). - Proof. - rewrite iMsg_exist_eq. iApply iProto_le_recv. iIntros (v p') "Hm /=". - iExists p'. iSplitR; last by auto. iApply iProto_le_refl. - Qed. - - Lemma iProto_le_texist_elim_l {TT : tele} i (m1 : TT → iMsg Σ V) m2 : - (∀ x, (<(Recv,i)> m1 x) ⊑ (<(Recv,i)> m2)) ⊢ - (<(Recv,i) @.. x> m1 x) ⊑ (<(Recv,i)> m2). - Proof. - iIntros "H". iInduction TT as [|T TT] "IH"; simpl; [done|]. - iApply iProto_le_exist_elim_l; iIntros (x). - iApply "IH". iIntros (xs). iApply "H". - Qed. - Lemma iProto_le_texist_elim_r {TT : tele} i m1 (m2 : TT → iMsg Σ V) : - (∀ x, (<(Send,i)> m1) ⊑ (<(Send,i)> m2 x)) -∗ - (<(Send,i)> m1) ⊑ (<(Send,i) @.. x> m2 x). - Proof. - iIntros "H". iInduction TT as [|T TT] "IH"; simpl; [done|]. - iApply iProto_le_exist_elim_r; iIntros (x). - iApply "IH". iIntros (xs). iApply "H". - Qed. - - Lemma iProto_le_texist_intro_l {TT : tele} i (m : TT → iMsg Σ V) x : - ⊢ (<(Send,i) @.. x> m x) ⊑ (<(Send,i)> m x). - Proof. - induction x as [|T TT x xs IH] using tele_arg_ind; simpl. - { iApply iProto_le_refl. } - iApply iProto_le_trans; [by iApply iProto_le_exist_intro_l|]. iApply IH. - Qed. - Lemma iProto_le_texist_intro_r {TT : tele} i (m : TT → iMsg Σ V) x : - ⊢ (<(Recv,i)> m x) ⊑ (<(Recv,i) @.. x> m x). - Proof. - induction x as [|T TT x xs IH] using tele_arg_ind; simpl. - { iApply iProto_le_refl. } - iApply iProto_le_trans; [|by iApply iProto_le_exist_intro_r]. iApply IH. - Qed. - - Lemma iProto_consistent_target ps m a i j : - iProto_consistent ps -∗ - ps !!! i ≡ (<(a, j)> m) -∗ - ⌜is_Some (ps !! j)âŒ. - Proof. - rewrite iProto_consistent_unfold. iDestruct 1 as "[Htar _]". iApply "Htar". - Qed. - - Lemma iProto_consistent_step ps m1 m2 i j v p1 : - iProto_consistent ps -∗ - ps !!! i ≡ (<(Send, j)> m1) -∗ - ps !!! j ≡ (<(Recv, i)> m2) -∗ - iMsg_car m1 v (Next p1) -∗ - ∃ p2, iMsg_car m2 v (Next p2) ∗ - â–· iProto_consistent (<[i := p1]>(<[j := p2]>ps)). - Proof. - iIntros "Hprot #Hi #Hj Hm1". - rewrite iProto_consistent_unfold /iProto_consistent_pre. - iDestruct "Hprot" as "[_ Hprot]". - iDestruct ("Hprot" with "Hi Hj Hm1") as (p2) "[Hm2 Hprot]". - iExists p2. iFrame. - Qed. - - Global Instance iProto_own_frag_ne γ s : NonExpansive (iProto_own_frag γ s). - Proof. solve_proper. Qed. - - Lemma iProto_own_auth_agree γ ps i p : - iProto_own_auth γ ps -∗ iProto_own_frag γ i p -∗ â–· (ps !!! i ≡ p). - Proof. Admitted. - (* iIntros "Hâ— Hâ—¯". *) - (* iDestruct (own_valid_2 with "Hâ— Hâ—¯") as "H✓". *) - (* rewrite gmap_view_both_validI. *) - (* iDestruct "H✓" as "[_ [H1 H2]]". *) - (* rewrite list_lookup_total_alt lookup_fmap. *) - (* destruct (ps !! i); last first. *) - (* { rewrite !option_equivI. } *) - (* simpl. rewrite !option_equivI excl_equivI. by iNext. *) - (* Qed. *) - - Lemma iProto_own_auth_update γ ps i p p' : - iProto_own_auth γ ps -∗ iProto_own_frag γ i p ==∗ - iProto_own_auth γ (<[i := p']>ps) ∗ iProto_own_frag γ i p'. - Proof. - iIntros "Hâ— Hâ—¯". - iMod (own_update_2 with "Hâ— Hâ—¯") as "[H1 H2]"; [|iModIntro]. - { eapply (gmap_view_replace _ _ _ (Excl' (Next p'))). done. } - iFrame. rewrite -fmap_insert. Admitted. - - Lemma iProto_own_auth_alloc ps : - ⊢ |==> ∃ γ, iProto_own_auth γ ps ∗ [∗ list] i ↦p ∈ ps, iProto_own γ i p. - Proof. Admitted. - (* iMod (own_alloc (gmap_view_auth (DfracOwn 1) ∅)) as (γ) "Hauth". *) - (* { apply gmap_view_auth_valid. } *) - (* iExists γ. *) - (* iInduction ps as [|i p ps Hin] "IH" using map_ind. *) - (* { iModIntro. iFrame. by iApply big_sepM_empty. } *) - (* iMod ("IH" with "Hauth") as "[Hauth Hfrags]". *) - (* rewrite big_sepM_insert; [|done]. iFrame "Hfrags". *) - (* iMod (own_update with "Hauth") as "[Hauth Hfrag]". *) - (* { apply (gmap_view_alloc _ i (DfracOwn 1) (Excl' (Next p))); [|done|done]. *) - (* by rewrite lookup_fmap Hin. } *) - (* iModIntro. rewrite -fmap_insert. iFrame. *) - (* iExists _. iFrame. iApply iProto_le_refl. *) - (* Qed. *) - - Lemma iProto_own_le γ s p1 p2 : - iProto_own γ s p1 -∗ â–· (p1 ⊑ p2) -∗ iProto_own γ s p2. - Proof. - iDestruct 1 as (p1') "[Hle H]". iIntros "Hle'". - iExists p1'. iFrame "H". by iApply (iProto_le_trans with "Hle"). - Qed. - - Lemma iProto_init ps : - â–· iProto_consistent ps -∗ - |==> ∃ γ, iProto_ctx γ (length ps) ∗ [∗ list] i ↦p ∈ ps, iProto_own γ i p. - Proof. - iIntros "Hconsistent". - iMod iProto_own_auth_alloc as (γ) "[Hauth Hfrags]". - iExists γ. iFrame. iExists _. by iFrame. - Qed. - - - Lemma iProto_step γ ps_dom i j m1 m2 p1 v : - iProto_ctx γ ps_dom -∗ - iProto_own γ i (<(Send, j)> m1) -∗ - iProto_own γ j (<(Recv, i)> m2) -∗ - iMsg_car m1 v (Next p1) ==∗ - â–· ∃ p2, iMsg_car m2 v (Next p2) ∗ iProto_ctx γ ps_dom ∗ - iProto_own γ i p1 ∗ iProto_own γ j p2. - Proof. - iIntros "Hctx Hi Hj Hm". - iDestruct "Hi" as (pi) "[Hile Hi]". - iDestruct "Hj" as (pj) "[Hjle Hj]". - iDestruct "Hctx" as (ps Hdom) "[Hauth Hconsistent]". - iDestruct (iProto_own_auth_agree with "Hauth Hi") as "#Hpi". - iDestruct (iProto_own_auth_agree with "Hauth Hj") as "#Hpj". - iDestruct (own_prot_idx with "Hi Hj") as %Hneq. - iAssert (â–· (<[i:=<(Send, j)> m1]>ps !!! j ≡ pj))%I as "Hpj'". - { by rewrite list_lookup_total_insert_ne. } - iAssert (â–· (⌜is_Some (ps !! i)⌠∗ (pi ⊑ (<(Send, j)> m1))))%I with "[Hile]" - as "[Hi' Hile]". - { iNext. iDestruct (iProto_le_msg_inv_r with "Hile") as (m) "#Heq". - iFrame. iRewrite "Heq" in "Hpi". rewrite list_lookup_total_alt. - destruct (ps !! i); [done|]. - iDestruct (iProto_end_message_equivI with "Hpi") as "[]". } - iAssert (â–· (⌜is_Some (ps !! j)⌠∗ (pj ⊑ (<(Recv, i)> m2))))%I with "[Hjle]" - as "[Hj' Hjle]". - { iNext. iDestruct (iProto_le_msg_inv_r with "Hjle") as (m) "#Heq". - iFrame. iRewrite "Heq" in "Hpj". rewrite !list_lookup_total_alt. - destruct (ps !! j); [done|]. - iDestruct (iProto_end_message_equivI with "Hpj") as "[]". } - iDestruct (iProto_consistent_le with "Hconsistent Hpi Hile") as "Hconsistent". - iDestruct (iProto_consistent_le with "Hconsistent Hpj' Hjle") as "Hconsistent". - iDestruct (iProto_consistent_step _ _ _ i j with "Hconsistent [] [] [Hm //]") as - (p2) "[Hm2 Hconsistent]". - { rewrite list_lookup_total_insert_ne; [|done]. - rewrite list_lookup_total_insert_ne; [|done]. - rewrite list_lookup_total_insert; [done|]. admit. } - { rewrite list_lookup_total_insert_ne; [|done]. - rewrite list_lookup_total_insert; [done|]. admit. } - iMod (iProto_own_auth_update _ _ _ _ p2 with "Hauth Hj") as "[Hauth Hj]". - iMod (iProto_own_auth_update _ _ _ _ p1 with "Hauth Hi") as "[Hauth Hi]". - iIntros "!>!>". iExists p2. iFrame "Hm2". - iDestruct "Hi'" as %Hi. iDestruct "Hj'" as %Hj. - iSplitL "Hconsistent Hauth". - { iExists (<[i:=p1]> (<[j:=p2]> ps)). - iSplit. - { admit. - (* rewrite !dom_insert_lookup_L; [done..|by rewrite lookup_insert_ne]. *)} - iFrame. rewrite list_insert_insert. - rewrite list_insert_commute; [|done]. rewrite list_insert_insert. - by rewrite list_insert_commute; [|done]. } - iSplitL "Hi"; iExists _; iFrame; iApply iProto_le_refl. - Admitted. - - Lemma iProto_target γ ps_dom i a j m : - iProto_ctx γ ps_dom -∗ - iProto_own γ i (<(a, j)> m) -∗ - â–· (⌜j < ps_domâŒ) ∗ iProto_ctx γ ps_dom ∗ iProto_own γ i (<(a, j)> m). - Proof. - iIntros "Hctx Hown". - rewrite /iProto_ctx /iProto_own. - iDestruct "Hctx" as (ps Hdom) "[Hauth Hps]". - iDestruct "Hown" as (p') "[Hle Hown]". - iDestruct (iProto_own_auth_agree with "Hauth Hown") as "#Hi". - iDestruct (iProto_le_msg_inv_r with "Hle") as (m') "#Heq". - iAssert (â–· (⌜is_Some (ps !! j)⌠∗ iProto_consistent ps))%I - with "[Hps]" as "[HSome Hps]". - { iNext. iRewrite "Heq" in "Hi". - iDestruct (iProto_consistent_target with "Hps Hi") as "#$". by iFrame. } - iSplitL "HSome". - { iNext. iDestruct "HSome" as %Heq. - iPureIntro. simplify_eq. admit. } - iSplitL "Hauth Hps"; iExists _; by iFrame. - Admitted. - - (* (** The instances below make it possible to use the tactics [iIntros], *) - (* [iExist], [iSplitL]/[iSplitR], [iFrame] and [iModIntro] on [iProto_le] goals. *) *) - Global Instance iProto_le_from_forall_l {A} i (m1 : A → iMsg Σ V) m2 name : - AsIdentName m1 name → - FromForall (iProto_message (Recv,i) (iMsg_exist m1) ⊑ (<(Recv,i)> m2)) - (λ x, (<(Recv, i)> m1 x) ⊑ (<(Recv, i)> m2))%I name | 10. - Proof. intros _. apply iProto_le_exist_elim_l. Qed. - Global Instance iProto_le_from_forall_r {A} i m1 (m2 : A → iMsg Σ V) name : - AsIdentName m2 name → - FromForall ((<(Send,i)> m1) ⊑ iProto_message (Send,i) (iMsg_exist m2)) - (λ x, (<(Send,i)> m1) ⊑ (<(Send,i)> m2 x))%I name | 11. - Proof. intros _. apply iProto_le_exist_elim_r. Qed. - - Global Instance iProto_le_from_wand_l i m v P p : - TCIf (TCEq P True%I) False TCTrue → - FromWand ((<(Recv,i)> MSG v {{ P }}; p) ⊑ (<(Recv,i)> m)) P ((<(Recv,i)> MSG v; p) ⊑ (<(Recv,i)> m)) | 10. - Proof. intros _. apply iProto_le_payload_elim_l. Qed. - Global Instance iProto_le_from_wand_r i m v P p : - FromWand ((<(Send,i)> m) ⊑ (<(Send,i)> MSG v {{ P }}; p)) P ((<(Send,i)> m) ⊑ (<(Send,i)> MSG v; p)) | 11. - Proof. apply iProto_le_payload_elim_r. Qed. - - Global Instance iProto_le_from_exist_l {A} i (m : A → iMsg Σ V) p : - FromExist ((<(Send,i) @ x> m x) ⊑ p) (λ a, (<(Send,i)> m a) ⊑ p)%I | 10. - Proof. - rewrite /FromExist. iDestruct 1 as (x) "H". - iApply (iProto_le_trans with "[] H"). iApply iProto_le_exist_intro_l. - Qed. - Global Instance iProto_le_from_exist_r {A} i (m : A → iMsg Σ V) p : - FromExist (p ⊑ <(Recv,i) @ x> m x) (λ a, p ⊑ (<(Recv,i)> m a))%I | 11. - Proof. - rewrite /FromExist. iDestruct 1 as (x) "H". - iApply (iProto_le_trans with "H"). iApply iProto_le_exist_intro_r. - Qed. - - Global Instance iProto_le_from_sep_l i m v P p : - FromSep ((<(Send,i)> MSG v {{ P }}; p) ⊑ (<(Send,i)> m)) P ((<(Send,i)> MSG v; p) ⊑ (<(Send,i)> m)) | 10. - Proof. - rewrite /FromSep. iIntros "[HP H]". - iApply (iProto_le_trans with "[HP] H"). by iApply iProto_le_payload_intro_l. - Qed. - Global Instance iProto_le_from_sep_r i m v P p : - FromSep ((<(Recv,i)> m) ⊑ (<(Recv,i)> MSG v {{ P }}; p)) P ((<(Recv,i)> m) ⊑ (<(Recv,i)> MSG v; p)) | 11. - Proof. - rewrite /FromSep. iIntros "[HP H]". - iApply (iProto_le_trans with "H"). by iApply iProto_le_payload_intro_r. - Qed. - - Global Instance iProto_le_frame_l i q m v R P Q p : - Frame q R P Q → - Frame q R ((<(Send,i)> MSG v {{ P }}; p) ⊑ (<(Send,i)> m)) - ((<(Send,i)> MSG v {{ Q }}; p) ⊑ (<(Send,i)> m)) | 10. - Proof. - rewrite /Frame /=. iIntros (HP) "[HR H]". - iApply (iProto_le_trans with "[HR] H"). iApply iProto_le_payload_elim_r. - iIntros "HQ". iApply iProto_le_payload_intro_l. iApply HP; iFrame. - Qed. - Global Instance iProto_le_frame_r i q m v R P Q p : - Frame q R P Q → - Frame q R ((<(Recv,i)> m) ⊑ (<(Recv,i)> MSG v {{ P }}; p)) - ((<(Recv,i)> m) ⊑ (<(Recv,i)> MSG v {{ Q }}; p)) | 11. - Proof. - rewrite /Frame /=. iIntros (HP) "[HR H]". - iApply (iProto_le_trans with "H"). iApply iProto_le_payload_elim_l. - iIntros "HQ". iApply iProto_le_payload_intro_r. iApply HP; iFrame. - Qed. - - Global Instance iProto_le_from_modal a v p1 p2 : - FromModal True (modality_instances.modality_laterN 1) (p1 ⊑ p2) - ((<a> MSG v; p1) ⊑ (<a> MSG v; p2)) (p1 ⊑ p2). - Proof. intros _. iApply iProto_le_base. Qed. - -End proto. - -Typeclasses Opaque iProto_ctx iProto_own. - -Global Hint Extern 0 (environments.envs_entails _ (?x ⊑ ?y)) => - first [is_evar x; fail 1 | is_evar y; fail 1|iApply iProto_le_refl] : core. diff --git a/multi_actris/examples/leader_election.v b/multi_actris/examples/leader_election.v deleted file mode 100644 index 0fa6ef5..0000000 --- a/multi_actris/examples/leader_election.v +++ /dev/null @@ -1,206 +0,0 @@ -From iris.heap_lang Require Import adequacy. -From iris.heap_lang.lib Require Import assert. -From multi_actris.channel Require Import proofmode. - -(** Inspired by https://en.wikipedia.org/wiki/Chang_and_Roberts_algorithm *) - -Definition process : val := - rec: "go" "c" "idl" "id" "idr" "isp" := - if: recv "c" "idr" - then let: "id'" := recv "c" "idr" in - if: "id" < "id'" (** Case 1 *) - then send "c" "idl" #true;; send "c" "idl" "id'";; - "go" "c" "idl" "id" "idr" #true - else if: "id" = "id'" (** Case 4 *) - then send "c" "idl" #false;; send "c" "idl" "id";; - "go" "c" "idl" "id" "idr" #false - else if: "isp" (** Case 3 *) - then "go" "c" "idl" "id" "idr" "isp" (** Case 2 *) - else send "c" "idl" #true;; send "c" "idl" "id";; - "go" "c" "idl" "id" "idr" #true - else let: "id'" := recv "c" "idr" in - if: "id" = "id'" then "id'" - else send "c" "idl" #false;; send "c" "idl" "id'";; "id'". - -Definition init : val := - λ: "c" "idl" "id" "idr", - (* Notice missing leader *) - send "c" "idl" #true;; send "c" "idl" "id";; - process "c" "idl" "id" "idr" #true. - -Definition program : val := - λ: <>, - let: "cs" := new_chan #4 in - let: "c0" := get_chan "cs" #0 in - let: "c1" := get_chan "cs" #1 in - let: "c2" := get_chan "cs" #2 in - let: "c3" := get_chan "cs" #3 in - Fork (let: "id_max" := init "c1" #3 #1 #2 in send "c1" #0 "id_max");; - Fork (let: "id_max" := process "c2" #1 #2 #3 #false in - send "c2" #0 "id_max");; - Fork (let: "id_max" := init "c3" #2 #3 #1 in send "c3" #0 "id_max");; - let: "res1" := recv "c0" #1 in - let: "res2" := recv "c0" #2 in - let: "res3" := recv "c0" #3 in - assert: ("res1" = "res2");; - assert: ("res2" = "res3"). - -Notation iProto_choice a p1 p2 := - (<a @ (b : bool)> MSG #b; if b then p1 else p2)%proto. - -Section ring_leader_election_example. - Context `{!heapGS Σ, chanG Σ, spawnG Σ, mono_natG Σ}. - - Definition my_recv_prot (il i ir : nat) (p : nat → iProto Σ) - (rec : bool -d> iProto Σ) : bool -d> iProto Σ := - λ (isp:bool), - iProto_choice (Recv,ir) - (<(Recv,ir) @ (i':nat)> MSG #i' ; - if bool_decide (i < i') - then <(Send,il)> MSG #true ; <(Send,il)> MSG #i' ; rec true - else if bool_decide (i = i') - then <(Send,il)> MSG #false ; <(Send, il)> MSG #i ; rec false - else if isp then rec isp - else <(Send,il)> MSG #true ; <(Send,il)> MSG #i ; rec true)%proto - (<(Recv,ir) @ (i':nat)> MSG #i' ; - if (bool_decide (i = i')) then p i - else <(Send,il)> MSG #false; <(Send,il)> MSG #i'; p i')%proto. - - Instance rle_prot_aux_contractive il i ir p : Contractive (my_recv_prot il i ir p). - Proof. rewrite /my_recv_prot. solve_proto_contractive. Qed. - Definition rle_prot il i ir p := fixpoint (my_recv_prot il i ir p). - Instance rle_prot_unfold il i ir isp p : - ProtoUnfold (rle_prot il i ir p isp) (my_recv_prot il i ir p (rle_prot il i ir p) isp). - Proof. apply proto_unfold_eq, (fixpoint_unfold (my_recv_prot il i ir p)). Qed. - Lemma rle_prot_unfold' il i ir isp p : - (rle_prot il i ir p isp) ≡ - (my_recv_prot il i ir p (rle_prot il i ir p) isp). - Proof. apply (fixpoint_unfold (my_recv_prot il i ir p)). Qed. - - Definition rle_preprot (il i ir : nat) p : iProto Σ := - (<(Send, il)> MSG #true; <(Send, il)> MSG #i ; - rle_prot il i ir p true)%proto. - - Lemma process_spec il i ir p c (isp:bool) : - {{{ c ↣ (rle_prot il i ir p isp) }}} - process c #il #i #ir #isp - {{{ i', RET #i'; c ↣ p i' }}}. - Proof. - iIntros (Φ) "Hc HΦ". - iLöb as "IH" forall (Φ isp). - wp_lam. wp_recv (b) as "_". - destruct b. - - wp_pures. wp_recv (i') as "_". - wp_pures. - case_bool_decide as Hlt. - { case_bool_decide; [|lia]. - wp_pures. wp_send with "[//]". - wp_send with "[//]". wp_pures. - iApply ("IH" with "Hc HΦ"). } - case_bool_decide as Hlt2. - { case_bool_decide; [lia|]. - wp_pures. case_bool_decide; [|simplify_eq;lia]. - wp_send with "[//]". - wp_send with "[//]". wp_pures. - iApply ("IH" with "Hc HΦ"). } - case_bool_decide; [lia|]. - wp_pures. - case_bool_decide; [simplify_eq;lia|]. - wp_pures. - destruct isp. - { wp_pures. iApply ("IH" with "Hc HΦ"). } - wp_pures. - wp_send with "[//]". - wp_send with "[//]". - wp_pures. iApply ("IH" with "Hc HΦ"). - - wp_pures. - wp_recv (id') as "_". wp_pures. - case_bool_decide as Hlt. - { case_bool_decide; [|simplify_eq;lia]. - wp_pures. subst. by iApply "HΦ". } - case_bool_decide; [simplify_eq;lia|]. - wp_send with "[//]". wp_send with "[//]". wp_pures. by iApply "HΦ". - Qed. - - Lemma init_spec c il i ir p : - {{{ c ↣ rle_preprot il i ir p }}} - init c #il #i #ir - {{{ res, RET #res; c ↣ p res }}}. - Proof. - iIntros (Φ) "Hc HΦ". wp_lam. wp_send with "[//]". wp_send with "[//]". - wp_pures. by iApply (process_spec with "Hc HΦ"). - Qed. - - Definition prot_tail (i_max : nat) : iProto Σ := - (<(Send,0)> MSG #i_max; END)%proto. - - Definition pre_prot_pool id_max : list (iProto Σ) := - [(<(Recv,1) @ (id_max : nat)> MSG #id_max ; - <(Recv,2)> MSG #id_max ; - <(Recv,3)> MSG #id_max ; - END)%proto; - prot_tail id_max; - prot_tail id_max; - prot_tail id_max]. - - Lemma pre_prot_pool_consistent id_max : - ⊢ iProto_consistent (pre_prot_pool id_max). - Proof. rewrite /pre_prot_pool. iProto_consistent_take_steps. Qed. - - Definition prot_pool : list (iProto Σ) := - [(<(Recv,1) @ (id_max : nat)> MSG #id_max ; - <(Recv,2)> MSG #id_max ; - <(Recv,3)> MSG #id_max ; - END)%proto; - rle_preprot 3 1 2 prot_tail; - rle_prot 1 2 3 prot_tail false; - rle_preprot 2 3 1 prot_tail]. - - Lemma prot_pool_consistent : ⊢ iProto_consistent prot_pool. - Proof. - rewrite /prot_pool /rle_preprot. - rewrite !rle_prot_unfold'. - iProto_consistent_take_steps. - case_bool_decide; try lia. - case_bool_decide; try lia. - rewrite !rle_prot_unfold'. - iProto_consistent_take_steps. - case_bool_decide; try lia. - case_bool_decide; try lia. - rewrite !rle_prot_unfold'. - iProto_consistent_take_steps. - Qed. - - Lemma program_spec : - {{{ True }}} program #() {{{ RET #(); True }}}. - Proof. - iIntros (Φ) "_ HΦ". wp_lam. - wp_new_chan prot_pool with prot_pool_consistent - as (c0 c1 c2 c3) "Hc0" "Hc1" "Hc2" "Hc3". - wp_smart_apply (wp_fork with "[Hc1]"). - { iIntros "!>". wp_smart_apply (init_spec with "Hc1"). - iIntros (i') "Hc1". by wp_send with "[//]". } - wp_smart_apply (wp_fork with "[Hc2]"). - { iIntros "!>". wp_smart_apply (process_spec with "Hc2"). - iIntros (i') "Hc2". by wp_send with "[//]". } - wp_smart_apply (wp_fork with "[Hc3]"). - { iIntros "!>". wp_smart_apply (init_spec with "Hc3"). - iIntros (i') "Hc3". by wp_send with "[//]". } - wp_recv (id_max) as "_". wp_recv as "_". wp_recv as "_". - wp_smart_apply wp_assert. wp_pures. iModIntro. iSplitR; [iPureIntro|]. - { do 2 f_equiv. by apply bool_decide_eq_true_2. } - iIntros "!>". - wp_smart_apply wp_assert. wp_pures. iModIntro. iSplitR; [iPureIntro|]. - { do 2 f_equiv. by apply bool_decide_eq_true_2. } - iIntros "!>". by iApply "HΦ". - Qed. - -End ring_leader_election_example. - -Lemma program_spec_adequate : - adequate NotStuck (program #()) ({|heap := ∅; used_proph_id := ∅|}) - (λ _ _, True). -Proof. - apply (heap_adequacy #[heapΣ; chanΣ]). - iIntros (Σ) "H". by iApply program_spec. -Qed. diff --git a/multi_actris/examples/leader_election_del.v b/multi_actris/examples/leader_election_del.v deleted file mode 100644 index 1ff5abd..0000000 --- a/multi_actris/examples/leader_election_del.v +++ /dev/null @@ -1,276 +0,0 @@ -From iris.heap_lang Require Import adequacy. -From iris.heap_lang.lib Require Import assert. -From multi_actris.channel Require Import proofmode. - -(** Inspired by https://en.wikipedia.org/wiki/Chang_and_Roberts_algorithm *) - -Definition process : val := - rec: "go" "c" "idl" "id" "idr" "isp" := - if: recv "c" "idr" - then let: "id'" := recv "c" "idr" in - if: "id" < "id'" (** Case 1 *) - then send "c" "idl" #true;; send "c" "idl" "id'";; - "go" "c" "idl" "id" "idr" #true - else if: "id" = "id'" (** Case 4 *) - then send "c" "idl" #false;; send "c" "idl" "id";; - "go" "c" "idl" "id" "idr" #false - else if: "isp" (** Case 3 *) - then "go" "c" "idl" "id" "idr" "isp" (** Case 2 *) - else send "c" "idl" #true;; send "c" "idl" "id";; - "go" "c" "idl" "id" "idr" #true - else let: "id'" := recv "c" "idr" in - if: "id" = "id'" then "id'" - else send "c" "idl" #false;; send "c" "idl" "id'";; "id'". - -Definition init : val := - λ: "c" "idl" "id" "idr", - (* Notice missing leader *) - send "c" "idl" #true;; send "c" "idl" "id";; - process "c" "idl" "id" "idr" #true. - -Definition forward : val := - λ: "c" "idl" "id" "idr" "id_max", - if: "id" = "id_max" then - let: "cs'" := new_chan #2 in - let: "c0" := get_chan "cs'" #0 in - let: "c1" := get_chan "cs'" #1 in - send "c1" #0 "id_max";; - send "c" "idl" "c1";; - Fork ((rec: "f" <> := - let: "id'" := recv "c0" #1 in - assert: ("id'" = "id_max");; "f" #()) #());; - recv "c" "idr";; #() - else - let: "c1" := recv "c" "idr" in - send "c1" #0 "id_max";; - send "c" "idl" "c1". - -Definition program : val := - λ: <>, - let: "cs" := new_chan #4 in - let: "c0" := get_chan "cs" #0 in - let: "c1" := get_chan "cs" #1 in - let: "c2" := get_chan "cs" #2 in - let: "c3" := get_chan "cs" #3 in - Fork (let: "id_max" := process "c1" #0 #1 #2 #false in - forward "c1" #0 #1 #2 "id_max");; - Fork (let: "id_max" := process "c2" #1 #2 #3 #false in - forward "c2" #1 #2 #3 "id_max");; - Fork (let: "id_max" := process "c3" #2 #3 #0 #false in - forward "c3" #2 #3 #0 "id_max");; - let: "id_max" := init "c0" #3 #0 #1 in - forward "c0" #3 #0 #1 "id_max". - -Notation iProto_choice a p1 p2 := - (<a @ (b : bool)> MSG #b; if b then p1 else p2)%proto. - -Section ring_leader_election_example. - Context `{!heapGS Σ, chanG Σ, spawnG Σ, mono_natG Σ}. - - Definition my_recv_prot (il i ir : nat) (p : nat → iProto Σ) - (rec : bool -d> iProto Σ) : bool -d> iProto Σ := - λ (isp:bool), - iProto_choice (Recv,ir) - (<(Recv,ir) @ (i':nat)> MSG #i' ; - if bool_decide (i < i') - then <(Send,il)> MSG #true ; <(Send,il)> MSG #i' ; rec true - else if bool_decide (i = i') - then <(Send,il)> MSG #false ; <(Send, il)> MSG #i ; rec false - else if isp then rec isp - else <(Send,il)> MSG #true ; <(Send,il)> MSG #i ; rec true)%proto - (<(Recv,ir) @ (i':nat)> MSG #i' ; - if (bool_decide (i = i')) then p i - else <(Send,il)> MSG #false; <(Send,il)> MSG #i'; p i')%proto. - - Instance rle_prot_aux_contractive il i ir p : Contractive (my_recv_prot il i ir p). - Proof. rewrite /my_recv_prot. solve_proto_contractive. Qed. - Definition rle_prot il i ir p := fixpoint (my_recv_prot il i ir p). - Instance rle_prot_unfold il i ir isp p : - ProtoUnfold (rle_prot il i ir p isp) (my_recv_prot il i ir p (rle_prot il i ir p) isp). - Proof. apply proto_unfold_eq, (fixpoint_unfold (my_recv_prot il i ir p)). Qed. - Lemma rle_prot_unfold' il i ir isp p : - (rle_prot il i ir p isp) ≡ - (my_recv_prot il i ir p (rle_prot il i ir p) isp). - Proof. apply (fixpoint_unfold (my_recv_prot il i ir p)). Qed. - - Definition rle_preprot (il i ir : nat) p : iProto Σ := - (<(Send, il)> MSG #true; <(Send, il)> MSG #i ; - rle_prot il i ir p true)%proto. - - Lemma process_spec il i ir p c (isp:bool) : - {{{ c ↣ (rle_prot il i ir p isp) }}} - process c #il #i #ir #isp - {{{ i', RET #i'; c ↣ p i' }}}. - Proof. - iIntros (Φ) "Hc HΦ". - iLöb as "IH" forall (Φ isp). - wp_lam. wp_recv (b) as "_". - destruct b. - - wp_pures. wp_recv (i') as "_". - wp_pures. - case_bool_decide as Hlt. - { case_bool_decide; [|lia]. - wp_pures. wp_send with "[//]". - wp_send with "[//]". wp_pures. - iApply ("IH" with "Hc HΦ"). } - case_bool_decide as Hlt2. - { case_bool_decide; [lia|]. - wp_pures. case_bool_decide; [|simplify_eq;lia]. - wp_send with "[//]". - wp_send with "[//]". wp_pures. - iApply ("IH" with "Hc HΦ"). } - case_bool_decide; [lia|]. - wp_pures. - case_bool_decide; [simplify_eq;lia|]. - wp_pures. - destruct isp. - { wp_pures. iApply ("IH" with "Hc HΦ"). } - wp_pures. - wp_send with "[//]". - wp_send with "[//]". - wp_pures. iApply ("IH" with "Hc HΦ"). - - wp_pures. - wp_recv (id') as "_". wp_pures. - case_bool_decide as Hlt. - { case_bool_decide; [|simplify_eq;lia]. - wp_pures. subst. by iApply "HΦ". } - case_bool_decide; [simplify_eq;lia|]. - wp_send with "[//]". wp_send with "[//]". wp_pures. by iApply "HΦ". - Qed. - - Lemma init_spec c il i ir p : - {{{ c ↣ rle_preprot il i ir p }}} - init c #il #i #ir - {{{ res, RET #res; c ↣ p res }}}. - Proof. - iIntros (Φ) "Hc HΦ". wp_lam. wp_send with "[//]". wp_send with "[//]". - wp_pures. by iApply (process_spec with "Hc HΦ"). - Qed. - - Definition forward_prot (p : iProto Σ) (il i ir i_max : nat) : iProto Σ := - if bool_decide (i = i_max) then - (<(Send,il) @ (c:val)> MSG c {{ c ↣ p }} ; <(Recv,ir)> MSG c {{ c ↣ p }}; END)%proto - else - (<(Recv,ir) @ (c:val)> MSG c {{ c ↣ p }} ; <(Send,il)> MSG c {{ c ↣ p }}; END)%proto. - - Definition relay_send_aux (id : nat) (rec : iProto Σ) : iProto Σ := - (<(Send,0)> MSG #id ; rec)%proto. - Instance relay_send_aux_contractive i : Contractive (relay_send_aux i). - Proof. solve_proto_contractive. Qed. - Definition relay_send_prot i := fixpoint (relay_send_aux i). - Instance relay_send_prot_unfold i : - ProtoUnfold (relay_send_prot i) (relay_send_aux i (relay_send_prot i)). - Proof. apply proto_unfold_eq, (fixpoint_unfold (relay_send_aux i)). Qed. - Lemma relay_send_prot_unfold' i : - (relay_send_prot i) ≡ - (relay_send_aux i (relay_send_prot i)). - Proof. apply (fixpoint_unfold (relay_send_aux i)). Qed. - - Definition relay_recv_aux (id : nat) (rec : iProto Σ) : iProto Σ := - (<(Recv,1)> MSG #id ; rec)%proto. - Instance relay_recv_aux_contractive i : Contractive (relay_recv_aux i). - Proof. solve_proto_contractive. Qed. - Definition relay_recv_prot i := fixpoint (relay_recv_aux i). - Instance relay_recv_prot_unfold i : - ProtoUnfold (relay_recv_prot i) (relay_recv_aux i (relay_recv_prot i)). - Proof. apply proto_unfold_eq, (fixpoint_unfold (relay_recv_aux i)). Qed. - Lemma relay_recv_prot_unfold' i : - (relay_recv_prot i) ≡ - (relay_recv_aux i (relay_recv_prot i)). - Proof. apply (fixpoint_unfold (relay_recv_aux i)). Qed. - - Definition prot_pool : list (iProto Σ) := - [rle_preprot 3 0 1 (λ id_max, forward_prot (relay_send_prot id_max) 3 0 1 id_max); - rle_prot 0 1 2 (λ id_max, forward_prot (relay_send_prot id_max) 0 1 2 id_max) false; - rle_prot 1 2 3 (λ id_max, forward_prot (relay_send_prot id_max) 1 2 3 id_max) false; - rle_prot 2 3 0 (λ id_max, forward_prot (relay_send_prot id_max) 2 3 0 id_max) false]. - - Lemma prot_pool_consistent : ⊢ iProto_consistent prot_pool. - Proof. - rewrite /prot_pool /rle_preprot. - rewrite !rle_prot_unfold'. - iProto_consistent_take_steps. - case_bool_decide; try lia. - rewrite !rle_prot_unfold'. - iProto_consistent_take_steps. - case_bool_decide; try lia. - rewrite !rle_prot_unfold'. - iProto_consistent_take_steps. - case_bool_decide; try lia. - rewrite !rle_prot_unfold'. - iProto_consistent_take_steps. - Qed. - - Definition prot_pool' (i:nat) : list (iProto Σ) := - [relay_recv_prot i; - relay_send_prot i]. - - Lemma prot_pool_consistent' i : ⊢ iProto_consistent (prot_pool' i). - Proof. - rewrite /prot_pool'. - iLöb as "IH". - iEval (rewrite relay_recv_prot_unfold'). - iEval (rewrite relay_send_prot_unfold'). - iProto_consistent_take_steps. - done. - Qed. - - Lemma forward_spec c il i ir i_max : - {{{ c ↣ forward_prot (relay_send_prot i_max) il i ir i_max }}} - forward c #il #i #ir #i_max - {{{ RET #(); True }}}. - Proof. - iIntros (Φ) "Hc HΦ". wp_lam. - rewrite /forward_prot. - wp_pures. case_bool_decide. - - simplify_eq. wp_pures. - case_bool_decide; [|simplify_eq;lia]. - wp_new_chan (prot_pool' i_max) with (prot_pool_consistent' i_max) - as (c0 c1) "Hc0" "Hc1". - wp_send with "[//]". - wp_send with "[Hc1//]". - wp_smart_apply (wp_fork with "[Hc0]"). - { iIntros "!>". - wp_pures. - iLöb as "IH". - wp_recv as "_". - wp_smart_apply wp_assert. - wp_pures. iModIntro. iSplit; [iPureIntro; f_equiv; by case_bool_decide|]. - iIntros "!>". wp_pures. by iApply "IH". } - wp_recv as "Hc'". wp_pures. by iApply "HΦ". - - case_bool_decide; [simplify_eq;lia|]. - wp_pures. - wp_recv (c') as "Hc'". - wp_send with "[//]". - wp_send with "[Hc'//]". - by iApply "HΦ". - Qed. - - Lemma program_spec : - {{{ True }}} program #() {{{ RET #(); True }}}. - Proof. - iIntros (Φ) "_ HΦ". wp_lam. - wp_new_chan prot_pool with prot_pool_consistent - as (c0 c1 c2 c3) "Hc0" "Hc1" "Hc2" "Hc3". - wp_smart_apply (wp_fork with "[Hc1]"). - { iIntros "!>". wp_smart_apply (process_spec with "Hc1"). - iIntros (i') "Hc1". by wp_smart_apply (forward_spec with "Hc1"). } - wp_smart_apply (wp_fork with "[Hc2]"). - { iIntros "!>". wp_smart_apply (process_spec with "Hc2"). - iIntros (i') "Hc2". by wp_smart_apply (forward_spec with "Hc2"). } - wp_smart_apply (wp_fork with "[Hc3]"). - { iIntros "!>". wp_smart_apply (process_spec with "Hc3"). - iIntros (i') "Hc3". by wp_smart_apply (forward_spec with "Hc3"). } - wp_smart_apply (init_spec with "Hc0"). - iIntros (i') "Hc0". by wp_smart_apply (forward_spec with "Hc0"). - Qed. - -End ring_leader_election_example. - -Lemma program_spec_adequate : - adequate NotStuck (program #()) ({|heap := ∅; used_proph_id := ∅|}) - (λ _ _, True). -Proof. - apply (heap_adequacy #[heapΣ; chanΣ]). - iIntros (Σ) "H". by iApply program_spec. -Qed. diff --git a/multi_actris/channel/channel.v b/multris/channel/channel.v similarity index 98% rename from multi_actris/channel/channel.v rename to multris/channel/channel.v index c0a0fbb..46083fd 100644 --- a/multi_actris/channel/channel.v +++ b/multris/channel/channel.v @@ -21,9 +21,9 @@ the subprotocol relation [⊑] *) From iris.algebra Require Import gmap excl_auth gmap_view. From iris.base_logic.lib Require Import invariants. From iris.heap_lang Require Export primitive_laws notation proofmode. -From multi_actris.utils Require Import matrix. -From multi_actris.channel Require Import proto_model. -From multi_actris.channel Require Export proto. +From multris.utils Require Import matrix. +From multris.channel Require Import proto_model. +From multris.channel Require Export proto. Set Default Proof Using "Type". (** * The definition of the message-passing connectives *) @@ -162,7 +162,7 @@ Section channel. { apply excl_auth_valid. } iMod ("IHn" with "Hps'") as (γEs Hlen) "H". iModIntro. iExists (γEs++[γE]). - rewrite !app_length Hlen. + rewrite !length_app Hlen. iSplit; [iPureIntro=>/=;lia|]=> /=. iSplitL "H". { iApply (big_sepL_impl with "H"). diff --git a/multi_actris/channel/proofmode.v b/multris/channel/proofmode.v similarity index 99% rename from multi_actris/channel/proofmode.v rename to multris/channel/proofmode.v index 2c02a5c..d3da00d 100644 --- a/multi_actris/channel/proofmode.v +++ b/multris/channel/proofmode.v @@ -10,8 +10,8 @@ recursive protocols are contractive. *) From iris.algebra Require Import gmap. From iris.proofmode Require Import coq_tactics reduction spec_patterns. From iris.heap_lang Require Export proofmode notation. -From multi_actris.channel Require Import proto_model. -From multi_actris.channel Require Export channel. +From multris.channel Require Import proto_model. +From multris.channel Require Export channel. Set Default Proof Using "Type". (** * Tactics for proving contractiveness of protocols *) diff --git a/multi_actris/channel/proto.v b/multris/channel/proto.v similarity index 99% rename from multi_actris/channel/proto.v rename to multris/channel/proto.v index aaae475..56522bb 100644 --- a/multi_actris/channel/proto.v +++ b/multris/channel/proto.v @@ -6,7 +6,7 @@ From iris.proofmode Require Import proofmode. From iris.base_logic Require Export lib.iprop. From iris.base_logic Require Import lib.own. From iris.program_logic Require Import language. -From multi_actris.channel Require Import proto_model. +From multris.channel Require Import proto_model. Set Default Proof Using "Type". Export action. diff --git a/multi_actris/channel/proto_model.v b/multris/channel/proto_model.v similarity index 99% rename from multi_actris/channel/proto_model.v rename to multris/channel/proto_model.v index 007f8ef..3dc4bf0 100644 --- a/multi_actris/channel/proto_model.v +++ b/multris/channel/proto_model.v @@ -36,7 +36,7 @@ The defined functions on the type [proto] are: all terminations [END] in [p1] with [p2]. *) From iris.base_logic Require Import base_logic. From iris.proofmode Require Import proofmode. -From multi_actris.utils Require Import cofe_solver_2. +From multris.utils Require Import cofe_solver_2. Set Default Proof Using "Type". Module Export action. diff --git a/multi_actris/examples/basics.v b/multris/examples/basics.v similarity index 99% rename from multi_actris/examples/basics.v rename to multris/examples/basics.v index 80169be..53ad44a 100644 --- a/multi_actris/examples/basics.v +++ b/multris/examples/basics.v @@ -1,4 +1,4 @@ -From multi_actris.channel Require Import proofmode. +From multris.channel Require Import proofmode. Set Default Proof Using "Type". Definition iProto_empty {Σ} : list (iProto Σ) := []. diff --git a/multi_actris/examples/leader_election_del_alt.v b/multris/examples/leader_election.v similarity index 99% rename from multi_actris/examples/leader_election_del_alt.v rename to multris/examples/leader_election.v index 093badc..c833695 100644 --- a/multi_actris/examples/leader_election_del_alt.v +++ b/multris/examples/leader_election.v @@ -1,6 +1,6 @@ From iris.heap_lang Require Import adequacy. From iris.heap_lang.lib Require Import assert. -From multi_actris.channel Require Import proofmode. +From multris.channel Require Import proofmode. (** Inspired by https://en.wikipedia.org/wiki/Chang_and_Roberts_algorithm *) diff --git a/multi_actris/examples/three_buyer.v b/multris/examples/three_buyer.v similarity index 99% rename from multi_actris/examples/three_buyer.v rename to multris/examples/three_buyer.v index aa4c6e8..a4d24c4 100644 --- a/multi_actris/examples/three_buyer.v +++ b/multris/examples/three_buyer.v @@ -1,4 +1,4 @@ -From multi_actris.channel Require Import proofmode. +From multris.channel Require Import proofmode. Set Default Proof Using "Type". Definition buyer1_prog : val := diff --git a/multi_actris/examples/two_buyer.v b/multris/examples/two_buyer.v similarity index 98% rename from multi_actris/examples/two_buyer.v rename to multris/examples/two_buyer.v index 15eef77..a317924 100644 --- a/multi_actris/examples/two_buyer.v +++ b/multris/examples/two_buyer.v @@ -1,4 +1,4 @@ -From multi_actris.channel Require Import proofmode. +From multris.channel Require Import proofmode. Set Default Proof Using "Type". Definition buyer1_prog : val := diff --git a/multi_actris/utils/cofe_solver_2.v b/multris/utils/cofe_solver_2.v similarity index 100% rename from multi_actris/utils/cofe_solver_2.v rename to multris/utils/cofe_solver_2.v diff --git a/multi_actris/utils/matrix.v b/multris/utils/matrix.v similarity index 100% rename from multi_actris/utils/matrix.v rename to multris/utils/matrix.v -- GitLab