aneris.dist_lang.lifting
From iris.algebra Require Import auth gmap frac agree coPset gset frac_auth ofe.
From iris.base_logic Require Export own gen_heap.
From iris.bi.lib Require Import fractional.
From iris.base_logic.lib Require Import viewshifts saved_prop gen_heap.
From iris.program_logic Require Export weakestpre.
From iris.program_logic Require Import ectx_lifting total_ectx_lifting.
From iris.proofmode Require Import tactics.
From aneris Require Export helpers lang notation tactics network resources.
From stdpp Require Import fin_maps gmap.
Set Default Proof Using "Type".
Import uPred.
Import Network.
From iris.base_logic Require Export own gen_heap.
From iris.bi.lib Require Import fractional.
From iris.base_logic.lib Require Import viewshifts saved_prop gen_heap.
From iris.program_logic Require Export weakestpre.
From iris.program_logic Require Import ectx_lifting total_ectx_lifting.
From iris.proofmode Require Import tactics.
From aneris Require Export helpers lang notation tactics network resources.
From stdpp Require Import fin_maps gmap.
Set Default Proof Using "Type".
Import uPred.
Import Network.
The tactic inv_head_step performs inversion on hypotheses of the shape
head_step. The tactic will discharge head-reductions starting from values, and
simplifies hypothesis related to conversions from and to values, and finite map
operations. This tactic is slightly ad-hoc and tuned for proving our lifting
lemmas.
Ltac inv_head_step :=
repeat match goal with
| _ => progress simplify_map_eq/= (* simplify memory stuff *)
| H : to_val _ = Some _ |- _ => apply to_base_val in H
| H : ground_lang.head_step ?e _ _ _ _ _ |- _ =>
try (is_var e; fail 1); (* inversion yields many goals if e is a variable
and can thus better be avoided. *)
inversion H; subst; clear H
| H : head_step ?e _ _ _ _ _ |- _ =>
try (is_var e; fail 1);
inversion H; subst; clear H
| H: socket_step _ ?e _ _ _ _ _ _ _ |- _ =>
try (is_var e; fail 1);
inversion H; subst; clear H
end.
Instance into_val_val n v : IntoVal ⟨n;Val v⟩ 〈n;v〉.
Proof. done. Qed.
Instance as_val_val n v : AsVal ⟨n;Val v⟩.
Proof. by exists 〈n;v〉. Qed.
Local Ltac solve_atomic :=
apply strongly_atomic_atomic, ectx_language_atomic;
[inversion 1; inv_head_step; naive_solver
|apply ectxi_language_sub_redexes_are_values; intros [] **; inv_head_step;
rewrite /to_val /is_Some /=; eexists; by
match goal with
| H: _ = _ |- _ => rewrite -H
end
].
Instance binop_atomic n s op v1 v2 : Atomic s ⟨n;BinOp op (Val v1) (Val v2)⟩.
Proof. solve_atomic. Qed.
Instance alloc_atomic n s v : Atomic s ⟨n;Alloc (Val v)⟩.
Proof. solve_atomic. Qed.
Instance load_atomic n s v : Atomic s ⟨n;Load (Val v)⟩.
Proof. solve_atomic. Qed.
Instance store_atomic n s v1 v2 : Atomic s ⟨n;Store (Val v1) (Val v2)⟩.
Proof. solve_atomic. Qed.
Instance cmpxchg_atomic n s v0 v1 v2 : Atomic s ⟨n;CAS (Val v0) (Val v1) (Val v2)⟩.
Proof. solve_atomic. Qed.
Instance fork_atomic n s e : Atomic s ⟨n;Fork e⟩.
Proof. solve_atomic. Qed.
Instance skip_atomic n s : Atomic s ⟨n;Skip⟩.
Proof. solve_atomic. Qed.
Instance newsocket_atomic n v0 v1 v2 s : Atomic s ⟨n; NewSocket (Val v0) (Val v1) (Val v2)⟩.
Proof. solve_atomic. Qed.
Instance socketbind_atomic n v0 v1 s : Atomic s ⟨n; SocketBind (Val v0) (Val v1)⟩.
Proof. solve_atomic. Qed.
Instance sendto_atomic n v0 v1 v2 s : Atomic s ⟨n; SendTo (Val v0) (Val v1) (Val v2)⟩.
Proof. solve_atomic. Qed.
Instance receivefrom_atomic n v s : Atomic s ⟨n; ReceiveFrom (Val v)⟩.
Proof. solve_atomic. Qed.
Local Ltac solve_exec_safe :=
intros;
repeat match goal with
| H: _ ∧ _ |- _ => destruct H as [??]
end;
simplify_eq;
do 3 eexists; eapply (LocalStepPureS _ ∅); econstructor; eauto.
Local Ltac solve_exec_puredet :=
simpl; intros; inv_head_step;
first (by repeat match goal with
| H: _ ∧ _ |- _ => destruct H as [??]; simplify_eq
| H : ground_lang.to_val _ = Some _ |- _ =>
rewrite ground_lang.to_of_val in H; simplify_eq
end);
try by match goal with
| H : socket_step _ _ _ _ _ _ _ _ _ _ _ |- _ =>
inversion H
end.
Local Ltac solve_pure_exec :=
simplify_eq; rewrite /PureExec; intros; apply nsteps_once, pure_head_step_pure_step;
constructor; [solve_exec_safe | solve_exec_puredet].
Class AsRecV (v : ground_lang.val) (f x : binder) (erec : ground_lang.expr) :=
as_recv : v = RecV f x erec.
Hint Mode AsRecV ! - - - : typeclass_instances.
Definition AsRecV_recv f x e : AsRecV (RecV f x e) f x e := eq_refl.
Hint Extern 0 (AsRecV (RecV _ _ _) _ _ _) =>
apply AsRecV_recv : typeclass_instances.
Instance pure_rec f x (erec : ground_lang.expr) :
PureExec True 1 ⟨n;Rec f x erec⟩ ⟨n;Val $ RecV f x erec⟩.
Proof. solve_pure_exec. Qed.
Instance pure_pairc (v1 v2 : ground_lang.val) :
PureExec True 1 ⟨n;(Pair (Val v1) (Val v2))⟩ ⟨n;(Val $ PairV v1 v2)⟩.
Proof. solve_pure_exec. Qed.
Instance pure_injlc (v : ground_lang.val) :
PureExec True 1 ⟨n;(InjL $ Val v)⟩ ⟨n;(Val $ InjLV v)⟩.
Proof. solve_pure_exec. Qed.
Instance pure_injrc (v : ground_lang.val) :
PureExec True 1 ⟨n;(InjR $ Val v)⟩ ⟨n;(Val $ InjRV v)⟩.
Proof. solve_pure_exec. Qed.
Instance pure_beta f x erec (v1 v2 : ground_lang.val) `{!AsRecV v1 f x erec} :
PureExec True 1 ⟨n;(App (Val v1) (Val v2))⟩ ⟨n;(subst' x v2 (subst' f v1 erec))⟩.
Proof. unfold AsRecV in *. solve_pure_exec. Qed.
Instance pure_unop op v v' :
PureExec (un_op_eval op v = Some v') 1 ⟨n;UnOp op (Val v)⟩ ⟨n;ground_lang.of_val v'⟩.
Proof. solve_pure_exec. Qed.
Instance pure_binop op v1 v2 v' :
PureExec (bin_op_eval op v1 v2 = Some v') 1
⟨n;BinOp op (Val v1) (Val v2)⟩ ⟨n;ground_lang.of_val v'⟩.
Proof. solve_pure_exec. Qed.
Instance pure_if_true e1 e2 : PureExec True 1 ⟨n;If (Val $ LitV $ LitBool true) e1 e2⟩ ⟨n;e1⟩.
Proof. solve_pure_exec. Qed.
Instance pure_if_false e1 e2 : PureExec True 1 ⟨n;If (Val $ LitV $ LitBool false) e1 e2⟩ ⟨n;e2⟩.
Proof. solve_pure_exec. Qed.
Instance pure_fst v1 v2 :
PureExec True 1 ⟨n;Fst (PairV v1 v2)⟩ ⟨n;Val v1⟩.
Proof. solve_pure_exec. Qed.
Instance pure_snd v1 v2 :
PureExec True 1 ⟨n;Snd (PairV v1 v2)⟩ ⟨n;Val v2⟩.
Proof. solve_pure_exec. Qed.
Instance pure_find_from v0 v1 n1 v2 v' :
PureExec (index n1 v2 v0 = v' ∧ Z.of_nat n1 = v1) 1
⟨n;FindFrom (Val $ LitV $ LitString v0)
(Val $ LitV $ LitInt v1)
(Val $ LitV $ LitString v2)⟩
⟨n;ground_lang.of_val (option_nat_to_val v')⟩.
Proof. solve_pure_exec. Qed.
Instance pure_substring v0 v1 n1 v2 n2 v' :
PureExec (substring n1 n2 v0 = v' ∧ Z.of_nat n1 = v1 ∧ Z.of_nat n2 = v2) 1
⟨n;Substring (LitV $ LitString v0) (LitV $ LitInt v1) (LitV $ LitInt v2)⟩
⟨n;ground_lang.of_val (LitV $ LitString v')⟩.
Proof. solve_pure_exec. Qed.
Instance pure_case_inl v e1 e2 :
PureExec True 1 ⟨n;Case (Val $ InjLV v) e1 e2⟩ ⟨n;App e1 (Val v)⟩.
Proof. solve_pure_exec. Qed.
Instance pure_case_inr v e1 e2 :
PureExec True 1 ⟨n;Case (Val $ InjRV v) e1 e2⟩ ⟨n;App e2 (Val v)⟩.
Proof. solve_pure_exec. Qed.
Instance pure_make_address v1 v2 :
PureExec True 1
⟨n;MakeAddress (LitV (LitString v1)) (LitV (LitInt (v2)))⟩
⟨n;LitV (LitSocketAddress (SocketAddressInet v1 (Z.to_pos v2)))⟩.
Proof. solve_pure_exec. Qed.
Local Hint Constructors head_step.
Local Hint Resolve alloc_fresh.
Local Hint Resolve to_of_val.
Section lifting.
Context `{distG Σ}.
Implicit Types P Q : iProp Σ.
Implicit Types Φ : val → iProp Σ.
Implicit Types efs : list expr.
Implicit Types σ : state.
repeat match goal with
| _ => progress simplify_map_eq/= (* simplify memory stuff *)
| H : to_val _ = Some _ |- _ => apply to_base_val in H
| H : ground_lang.head_step ?e _ _ _ _ _ |- _ =>
try (is_var e; fail 1); (* inversion yields many goals if e is a variable
and can thus better be avoided. *)
inversion H; subst; clear H
| H : head_step ?e _ _ _ _ _ |- _ =>
try (is_var e; fail 1);
inversion H; subst; clear H
| H: socket_step _ ?e _ _ _ _ _ _ _ |- _ =>
try (is_var e; fail 1);
inversion H; subst; clear H
end.
Instance into_val_val n v : IntoVal ⟨n;Val v⟩ 〈n;v〉.
Proof. done. Qed.
Instance as_val_val n v : AsVal ⟨n;Val v⟩.
Proof. by exists 〈n;v〉. Qed.
Local Ltac solve_atomic :=
apply strongly_atomic_atomic, ectx_language_atomic;
[inversion 1; inv_head_step; naive_solver
|apply ectxi_language_sub_redexes_are_values; intros [] **; inv_head_step;
rewrite /to_val /is_Some /=; eexists; by
match goal with
| H: _ = _ |- _ => rewrite -H
end
].
Instance binop_atomic n s op v1 v2 : Atomic s ⟨n;BinOp op (Val v1) (Val v2)⟩.
Proof. solve_atomic. Qed.
Instance alloc_atomic n s v : Atomic s ⟨n;Alloc (Val v)⟩.
Proof. solve_atomic. Qed.
Instance load_atomic n s v : Atomic s ⟨n;Load (Val v)⟩.
Proof. solve_atomic. Qed.
Instance store_atomic n s v1 v2 : Atomic s ⟨n;Store (Val v1) (Val v2)⟩.
Proof. solve_atomic. Qed.
Instance cmpxchg_atomic n s v0 v1 v2 : Atomic s ⟨n;CAS (Val v0) (Val v1) (Val v2)⟩.
Proof. solve_atomic. Qed.
Instance fork_atomic n s e : Atomic s ⟨n;Fork e⟩.
Proof. solve_atomic. Qed.
Instance skip_atomic n s : Atomic s ⟨n;Skip⟩.
Proof. solve_atomic. Qed.
Instance newsocket_atomic n v0 v1 v2 s : Atomic s ⟨n; NewSocket (Val v0) (Val v1) (Val v2)⟩.
Proof. solve_atomic. Qed.
Instance socketbind_atomic n v0 v1 s : Atomic s ⟨n; SocketBind (Val v0) (Val v1)⟩.
Proof. solve_atomic. Qed.
Instance sendto_atomic n v0 v1 v2 s : Atomic s ⟨n; SendTo (Val v0) (Val v1) (Val v2)⟩.
Proof. solve_atomic. Qed.
Instance receivefrom_atomic n v s : Atomic s ⟨n; ReceiveFrom (Val v)⟩.
Proof. solve_atomic. Qed.
Local Ltac solve_exec_safe :=
intros;
repeat match goal with
| H: _ ∧ _ |- _ => destruct H as [??]
end;
simplify_eq;
do 3 eexists; eapply (LocalStepPureS _ ∅); econstructor; eauto.
Local Ltac solve_exec_puredet :=
simpl; intros; inv_head_step;
first (by repeat match goal with
| H: _ ∧ _ |- _ => destruct H as [??]; simplify_eq
| H : ground_lang.to_val _ = Some _ |- _ =>
rewrite ground_lang.to_of_val in H; simplify_eq
end);
try by match goal with
| H : socket_step _ _ _ _ _ _ _ _ _ _ _ |- _ =>
inversion H
end.
Local Ltac solve_pure_exec :=
simplify_eq; rewrite /PureExec; intros; apply nsteps_once, pure_head_step_pure_step;
constructor; [solve_exec_safe | solve_exec_puredet].
Class AsRecV (v : ground_lang.val) (f x : binder) (erec : ground_lang.expr) :=
as_recv : v = RecV f x erec.
Hint Mode AsRecV ! - - - : typeclass_instances.
Definition AsRecV_recv f x e : AsRecV (RecV f x e) f x e := eq_refl.
Hint Extern 0 (AsRecV (RecV _ _ _) _ _ _) =>
apply AsRecV_recv : typeclass_instances.
Instance pure_rec f x (erec : ground_lang.expr) :
PureExec True 1 ⟨n;Rec f x erec⟩ ⟨n;Val $ RecV f x erec⟩.
Proof. solve_pure_exec. Qed.
Instance pure_pairc (v1 v2 : ground_lang.val) :
PureExec True 1 ⟨n;(Pair (Val v1) (Val v2))⟩ ⟨n;(Val $ PairV v1 v2)⟩.
Proof. solve_pure_exec. Qed.
Instance pure_injlc (v : ground_lang.val) :
PureExec True 1 ⟨n;(InjL $ Val v)⟩ ⟨n;(Val $ InjLV v)⟩.
Proof. solve_pure_exec. Qed.
Instance pure_injrc (v : ground_lang.val) :
PureExec True 1 ⟨n;(InjR $ Val v)⟩ ⟨n;(Val $ InjRV v)⟩.
Proof. solve_pure_exec. Qed.
Instance pure_beta f x erec (v1 v2 : ground_lang.val) `{!AsRecV v1 f x erec} :
PureExec True 1 ⟨n;(App (Val v1) (Val v2))⟩ ⟨n;(subst' x v2 (subst' f v1 erec))⟩.
Proof. unfold AsRecV in *. solve_pure_exec. Qed.
Instance pure_unop op v v' :
PureExec (un_op_eval op v = Some v') 1 ⟨n;UnOp op (Val v)⟩ ⟨n;ground_lang.of_val v'⟩.
Proof. solve_pure_exec. Qed.
Instance pure_binop op v1 v2 v' :
PureExec (bin_op_eval op v1 v2 = Some v') 1
⟨n;BinOp op (Val v1) (Val v2)⟩ ⟨n;ground_lang.of_val v'⟩.
Proof. solve_pure_exec. Qed.
Instance pure_if_true e1 e2 : PureExec True 1 ⟨n;If (Val $ LitV $ LitBool true) e1 e2⟩ ⟨n;e1⟩.
Proof. solve_pure_exec. Qed.
Instance pure_if_false e1 e2 : PureExec True 1 ⟨n;If (Val $ LitV $ LitBool false) e1 e2⟩ ⟨n;e2⟩.
Proof. solve_pure_exec. Qed.
Instance pure_fst v1 v2 :
PureExec True 1 ⟨n;Fst (PairV v1 v2)⟩ ⟨n;Val v1⟩.
Proof. solve_pure_exec. Qed.
Instance pure_snd v1 v2 :
PureExec True 1 ⟨n;Snd (PairV v1 v2)⟩ ⟨n;Val v2⟩.
Proof. solve_pure_exec. Qed.
Instance pure_find_from v0 v1 n1 v2 v' :
PureExec (index n1 v2 v0 = v' ∧ Z.of_nat n1 = v1) 1
⟨n;FindFrom (Val $ LitV $ LitString v0)
(Val $ LitV $ LitInt v1)
(Val $ LitV $ LitString v2)⟩
⟨n;ground_lang.of_val (option_nat_to_val v')⟩.
Proof. solve_pure_exec. Qed.
Instance pure_substring v0 v1 n1 v2 n2 v' :
PureExec (substring n1 n2 v0 = v' ∧ Z.of_nat n1 = v1 ∧ Z.of_nat n2 = v2) 1
⟨n;Substring (LitV $ LitString v0) (LitV $ LitInt v1) (LitV $ LitInt v2)⟩
⟨n;ground_lang.of_val (LitV $ LitString v')⟩.
Proof. solve_pure_exec. Qed.
Instance pure_case_inl v e1 e2 :
PureExec True 1 ⟨n;Case (Val $ InjLV v) e1 e2⟩ ⟨n;App e1 (Val v)⟩.
Proof. solve_pure_exec. Qed.
Instance pure_case_inr v e1 e2 :
PureExec True 1 ⟨n;Case (Val $ InjRV v) e1 e2⟩ ⟨n;App e2 (Val v)⟩.
Proof. solve_pure_exec. Qed.
Instance pure_make_address v1 v2 :
PureExec True 1
⟨n;MakeAddress (LitV (LitString v1)) (LitV (LitInt (v2)))⟩
⟨n;LitV (LitSocketAddress (SocketAddressInet v1 (Z.to_pos v2)))⟩.
Proof. solve_pure_exec. Qed.
Local Hint Constructors head_step.
Local Hint Resolve alloc_fresh.
Local Hint Resolve to_of_val.
Section lifting.
Context `{distG Σ}.
Implicit Types P Q : iProp Σ.
Implicit Types Φ : val → iProp Σ.
Implicit Types efs : list expr.
Implicit Types σ : state.
Base axioms for core primitives of the language: Stateless reductions
Lemma wp_fork n s E e Φ :
▷ Φ 〈n;LitV LitUnit〉 ∗ ▷ WP ⟨n;e⟩ @ s; ⊤ {{ _, True }} ⊢ WP ⟨n;Fork e⟩ @ s; E {{ Φ }}.
Proof.
iIntros "[HΦ He]". iApply wp_lift_atomic_head_step; [done|].
iIntros (σ1 κ κs m) "Hσ !>". iSplit.
- iPureIntro. eexists; solve_exec_safe.
- (* Should be enough with "iFrame", but "iNext" steps too agressively in Hσ *)
iModIntro. iIntros (v2 σ2 efs Hstep). iModIntro. inv_head_step. iFrame.
iDestruct "Hσ" as (?) "(?&?&?&?&?&Hσ)".
iDestruct "Hσ" as (? ?) "(% &?&?&?&?&?)".
iExists _; iFrame; iExists _, _; iFrame.
iSplitR; first by iPureIntro.
iApply big_sepM_mono; last done.
iIntros (? ?) "_ H".
iDestruct "H" as (π Φ0) "([? | (?&?&?)] & ?)";
iExists _, Φ0; iFrame; auto.
Qed.
Local Transparent IsNode.
Lemma wp_start (n : node) (i : ip_address) (P : gset port) s E e Φ :
n ≠ "system" →
▷ FreeIP i ∗
▷ Φ 〈"system";LitV LitUnit〉 ∗
▷ (IsNode n -∗ FreePorts i P -∗ WP ⟨n;e⟩ @ s; ⊤ {{ _, True }}) ⊢
WP ⟨"system";Start (LitString n) (LitString i) e⟩ @ s; E {{ Φ }}.
Proof.
iIntros (Hneq) "(>Hip & HΦ & Hwp)".
iApply (wp_lift_head_step with "[-]"); first auto.
- iIntros (σ1 κ κs m) "H".
iMod (fupd_intro_mask _ ∅ True%I with "[]") as "H'";
first set_solver; auto.
iDestruct "H" as (x [Hdomheap Hdomsock])
"(Hsi & Hγs & HlocS & Hips & Hmsg)".
iDestruct "Hips" as (Fip Piu (Hdsj & HFip & HPiu)) "[HFip HPiu]".
iDestruct (is_FreeIP with "HFip Hip") as %Hip.
iMod (Use_FreeIP with "HFip Hip") as "HFip".
iMod (FreePorts_alloc _ i P with "HPiu") as "[HPiu Hports]";
first set_solver.
iModIntro; iSplit; first iPureIntro.
+ do 4 eexists. apply StartStepS; eauto.
+ iNext. iIntros (e2 σ2 efs Hrd). iMod "H'" as "_".
iDestruct (message_state_later with "Hmsg") as "Hmsg".
inversion Hrd; last inversion SocketStep; subst; try inversion BaseStep.
destruct (state_heaps σ1 !! n) eqn:Heq1;
destruct (state_sockets σ1 !! n) eqn:Heq2; destruct σ1; simpl in *;
rewrite Heq1 Heq2 /=.
* assert (Hin : is_Some (x !! n)).
{ rewrite -(elem_of_dom (D:=gset node)) Hdomheap elem_of_dom; eauto. }
destruct Hin.
iDestruct (node_local_state with "HlocS") as "(Hloc & HlocS)";
first done.
iDestruct "Hloc" as (heap sockets Hheap Hsockets)
"(Hγs' & Hheap & Hsockets)".
rewrite /IsNode {2}mapsto_node_eq /mapsto_node_def.
iDestruct "Hγs'" as "#Hγs'".
iDestruct (node_local_state_rev with "[Hheap Hsockets] HlocS") as
"HlocS"; first done.
-- iExists heap,sockets; iFrame; simpl. repeat iSplitR; auto.
rewrite mapsto_node_eq /mapsto_node_def. iFrame "#".
-- iSplitR "HΦ Hwp Hports".
++ iExists (<[n:=x0]>x).
rewrite ![<[ _ := _]> state_sockets0](insert_id) //.
rewrite ![<[ _ := _]> state_heaps0](insert_id) //.
rewrite ![<[ _ := _]> x](insert_id) //.
iFrame; auto.
simpl in *.
iModIntro; repeat iSplit; auto.
iExists _, _; iFrame. iPureIntro; repeat split.
{ rewrite dom_insert. set_solver. }
{ intros ? ?; apply HFip; set_solver. }
{ intros ip Q HipQ.
destruct (decide (i = ip)); subst.
- rewrite lookup_insert in HipQ; simplify_eq.
rewrite HFip //. eexists; split; eauto.
set_solver.
- rewrite lookup_insert_ne // in HipQ.
eapply HPiu; eauto. }
++ iSplitL "HΦ". iApply wp_value; eauto.
iSplitL; last done.
iApply ("Hwp" with "[]"); last done.
rewrite mapsto_node_eq /mapsto_node_def; iFrame "#"; iFrame; eauto.
* revert Heq1 Heq2.
rewrite -(not_elem_of_dom (D:=gset node)).
rewrite -Hdomsock Hdomheap not_elem_of_dom.
intros Hsh1 Hsh2; rewrite Hsh1 in Hsh2; simplify_eq.
* revert Heq1 Heq2.
rewrite -(not_elem_of_dom (D:=gset node)).
rewrite -Hdomheap Hdomsock not_elem_of_dom.
intros Hsh1 Hsh2. rewrite Hsh1 in Hsh2. simplify_eq.
* iMod (own_alloc (A:=authR (gen_heapUR loc ground_lang.val))
(● to_gen_heap ∅)) as (γh) "Hheap".
{ apply/auth_auth_valid. exact: to_gen_heap_valid. }
iMod (own_alloc (A:=authR (gen_metaUR loc))
(● to_gen_meta ∅)) as (γhm) "Hheapm".
{ rewrite auth_auth_valid. exact: to_gen_meta_valid. }
iMod (own_alloc (A:=authR (gen_heapUR socket_handle socket))
(● to_gen_heap ∅)) as (γs) "Hsocket".
{ apply/auth_auth_valid. exact: to_gen_heap_valid. }
iMod (own_alloc (A:=authR (gen_metaUR socket_handle))
(● to_gen_meta ∅)) as (γsm) "Hsocketm".
{ rewrite auth_auth_valid. exact: to_gen_meta_valid. }
set (γn := {| heap_name := γh; heap_meta_name := γhm;
socket_name := γs; socket_meta_name := γsm |} : node_gnamesO).
iMod (own_update (A:=authR (gmapUR node (agreeR node_gnamesO)))
_
(● (to_agree <$> x))
(● (to_agree <$> (<[n:=γn]> x)) ⋅
(◯ {[ n := to_agree γn ]} : auth system_state_mapUR))
with "Hγs") as "[Hγs #Hγl]".
{
rewrite fmap_insert.
apply auth_update_alloc.
apply (alloc_singleton_local_update _ _ (to_agree γn)); last done.
by rewrite -(not_elem_of_dom (D:=gset node))
dom_fmap_L Hdomheap not_elem_of_dom. }
iSplitR "HΦ Hwp Hports".
-- iExists (<[n:=γn]> x).
iFrame. iSplitR.
{ iPureIntro. rewrite /local_state_coherence.
rewrite !dom_insert_L. rewrite -Hdomheap -Hdomsock /= //. }
assert (x !! n = None).
{ by rewrite -(not_elem_of_dom (D:=gset node))
Hdomheap not_elem_of_dom. }
rewrite big_sepM_insert; auto. rewrite -{1}(delete_notin x n); auto.
iDestruct (map_local_state_i_update_heaps with "HlocS") as "HlocS".
iDestruct (map_local_state_i_update_sockets with "HlocS") as "HlocS".
rewrite /= delete_notin; auto. iFrame.
iSplitL "Hheap Hheapm Hsocket Hsocketm".
++ iExists _,_.
rewrite !lookup_insert mapsto_node_eq /mapsto_node_def /= //.
iFrame; iFrame "#". iModIntro. do 2 (iSplit; first done).
iSplitL "Hheapm"; iExists _; by iFrame.
++ iModIntro; iExists _, _; iFrame. iPureIntro; repeat split.
{ rewrite dom_insert. set_solver. }
{ intros ? ?; apply HFip; set_solver. }
{ intros ip Q HipQ.
destruct (decide (i = ip)); subst.
- rewrite lookup_insert in HipQ; simplify_eq.
rewrite HFip //. eexists; split; eauto.
set_solver.
- rewrite lookup_insert_ne // in HipQ.
eapply HPiu; eauto. }
-- iSplitL "HΦ". iApply wp_value; rewrite /IntoVal /=; eauto; done.
iSplitL; last done.
iApply ("Hwp" with "[]"); last done.
rewrite /IsNode mapsto_node_eq /mapsto_node_def; iFrame "#"; iFrame; eauto.
Qed.
▷ Φ 〈n;LitV LitUnit〉 ∗ ▷ WP ⟨n;e⟩ @ s; ⊤ {{ _, True }} ⊢ WP ⟨n;Fork e⟩ @ s; E {{ Φ }}.
Proof.
iIntros "[HΦ He]". iApply wp_lift_atomic_head_step; [done|].
iIntros (σ1 κ κs m) "Hσ !>". iSplit.
- iPureIntro. eexists; solve_exec_safe.
- (* Should be enough with "iFrame", but "iNext" steps too agressively in Hσ *)
iModIntro. iIntros (v2 σ2 efs Hstep). iModIntro. inv_head_step. iFrame.
iDestruct "Hσ" as (?) "(?&?&?&?&?&Hσ)".
iDestruct "Hσ" as (? ?) "(% &?&?&?&?&?)".
iExists _; iFrame; iExists _, _; iFrame.
iSplitR; first by iPureIntro.
iApply big_sepM_mono; last done.
iIntros (? ?) "_ H".
iDestruct "H" as (π Φ0) "([? | (?&?&?)] & ?)";
iExists _, Φ0; iFrame; auto.
Qed.
Local Transparent IsNode.
Lemma wp_start (n : node) (i : ip_address) (P : gset port) s E e Φ :
n ≠ "system" →
▷ FreeIP i ∗
▷ Φ 〈"system";LitV LitUnit〉 ∗
▷ (IsNode n -∗ FreePorts i P -∗ WP ⟨n;e⟩ @ s; ⊤ {{ _, True }}) ⊢
WP ⟨"system";Start (LitString n) (LitString i) e⟩ @ s; E {{ Φ }}.
Proof.
iIntros (Hneq) "(>Hip & HΦ & Hwp)".
iApply (wp_lift_head_step with "[-]"); first auto.
- iIntros (σ1 κ κs m) "H".
iMod (fupd_intro_mask _ ∅ True%I with "[]") as "H'";
first set_solver; auto.
iDestruct "H" as (x [Hdomheap Hdomsock])
"(Hsi & Hγs & HlocS & Hips & Hmsg)".
iDestruct "Hips" as (Fip Piu (Hdsj & HFip & HPiu)) "[HFip HPiu]".
iDestruct (is_FreeIP with "HFip Hip") as %Hip.
iMod (Use_FreeIP with "HFip Hip") as "HFip".
iMod (FreePorts_alloc _ i P with "HPiu") as "[HPiu Hports]";
first set_solver.
iModIntro; iSplit; first iPureIntro.
+ do 4 eexists. apply StartStepS; eauto.
+ iNext. iIntros (e2 σ2 efs Hrd). iMod "H'" as "_".
iDestruct (message_state_later with "Hmsg") as "Hmsg".
inversion Hrd; last inversion SocketStep; subst; try inversion BaseStep.
destruct (state_heaps σ1 !! n) eqn:Heq1;
destruct (state_sockets σ1 !! n) eqn:Heq2; destruct σ1; simpl in *;
rewrite Heq1 Heq2 /=.
* assert (Hin : is_Some (x !! n)).
{ rewrite -(elem_of_dom (D:=gset node)) Hdomheap elem_of_dom; eauto. }
destruct Hin.
iDestruct (node_local_state with "HlocS") as "(Hloc & HlocS)";
first done.
iDestruct "Hloc" as (heap sockets Hheap Hsockets)
"(Hγs' & Hheap & Hsockets)".
rewrite /IsNode {2}mapsto_node_eq /mapsto_node_def.
iDestruct "Hγs'" as "#Hγs'".
iDestruct (node_local_state_rev with "[Hheap Hsockets] HlocS") as
"HlocS"; first done.
-- iExists heap,sockets; iFrame; simpl. repeat iSplitR; auto.
rewrite mapsto_node_eq /mapsto_node_def. iFrame "#".
-- iSplitR "HΦ Hwp Hports".
++ iExists (<[n:=x0]>x).
rewrite ![<[ _ := _]> state_sockets0](insert_id) //.
rewrite ![<[ _ := _]> state_heaps0](insert_id) //.
rewrite ![<[ _ := _]> x](insert_id) //.
iFrame; auto.
simpl in *.
iModIntro; repeat iSplit; auto.
iExists _, _; iFrame. iPureIntro; repeat split.
{ rewrite dom_insert. set_solver. }
{ intros ? ?; apply HFip; set_solver. }
{ intros ip Q HipQ.
destruct (decide (i = ip)); subst.
- rewrite lookup_insert in HipQ; simplify_eq.
rewrite HFip //. eexists; split; eauto.
set_solver.
- rewrite lookup_insert_ne // in HipQ.
eapply HPiu; eauto. }
++ iSplitL "HΦ". iApply wp_value; eauto.
iSplitL; last done.
iApply ("Hwp" with "[]"); last done.
rewrite mapsto_node_eq /mapsto_node_def; iFrame "#"; iFrame; eauto.
* revert Heq1 Heq2.
rewrite -(not_elem_of_dom (D:=gset node)).
rewrite -Hdomsock Hdomheap not_elem_of_dom.
intros Hsh1 Hsh2; rewrite Hsh1 in Hsh2; simplify_eq.
* revert Heq1 Heq2.
rewrite -(not_elem_of_dom (D:=gset node)).
rewrite -Hdomheap Hdomsock not_elem_of_dom.
intros Hsh1 Hsh2. rewrite Hsh1 in Hsh2. simplify_eq.
* iMod (own_alloc (A:=authR (gen_heapUR loc ground_lang.val))
(● to_gen_heap ∅)) as (γh) "Hheap".
{ apply/auth_auth_valid. exact: to_gen_heap_valid. }
iMod (own_alloc (A:=authR (gen_metaUR loc))
(● to_gen_meta ∅)) as (γhm) "Hheapm".
{ rewrite auth_auth_valid. exact: to_gen_meta_valid. }
iMod (own_alloc (A:=authR (gen_heapUR socket_handle socket))
(● to_gen_heap ∅)) as (γs) "Hsocket".
{ apply/auth_auth_valid. exact: to_gen_heap_valid. }
iMod (own_alloc (A:=authR (gen_metaUR socket_handle))
(● to_gen_meta ∅)) as (γsm) "Hsocketm".
{ rewrite auth_auth_valid. exact: to_gen_meta_valid. }
set (γn := {| heap_name := γh; heap_meta_name := γhm;
socket_name := γs; socket_meta_name := γsm |} : node_gnamesO).
iMod (own_update (A:=authR (gmapUR node (agreeR node_gnamesO)))
_
(● (to_agree <$> x))
(● (to_agree <$> (<[n:=γn]> x)) ⋅
(◯ {[ n := to_agree γn ]} : auth system_state_mapUR))
with "Hγs") as "[Hγs #Hγl]".
{
rewrite fmap_insert.
apply auth_update_alloc.
apply (alloc_singleton_local_update _ _ (to_agree γn)); last done.
by rewrite -(not_elem_of_dom (D:=gset node))
dom_fmap_L Hdomheap not_elem_of_dom. }
iSplitR "HΦ Hwp Hports".
-- iExists (<[n:=γn]> x).
iFrame. iSplitR.
{ iPureIntro. rewrite /local_state_coherence.
rewrite !dom_insert_L. rewrite -Hdomheap -Hdomsock /= //. }
assert (x !! n = None).
{ by rewrite -(not_elem_of_dom (D:=gset node))
Hdomheap not_elem_of_dom. }
rewrite big_sepM_insert; auto. rewrite -{1}(delete_notin x n); auto.
iDestruct (map_local_state_i_update_heaps with "HlocS") as "HlocS".
iDestruct (map_local_state_i_update_sockets with "HlocS") as "HlocS".
rewrite /= delete_notin; auto. iFrame.
iSplitL "Hheap Hheapm Hsocket Hsocketm".
++ iExists _,_.
rewrite !lookup_insert mapsto_node_eq /mapsto_node_def /= //.
iFrame; iFrame "#". iModIntro. do 2 (iSplit; first done).
iSplitL "Hheapm"; iExists _; by iFrame.
++ iModIntro; iExists _, _; iFrame. iPureIntro; repeat split.
{ rewrite dom_insert. set_solver. }
{ intros ? ?; apply HFip; set_solver. }
{ intros ip Q HipQ.
destruct (decide (i = ip)); subst.
- rewrite lookup_insert in HipQ; simplify_eq.
rewrite HFip //. eexists; split; eauto.
set_solver.
- rewrite lookup_insert_ne // in HipQ.
eapply HPiu; eauto. }
-- iSplitL "HΦ". iApply wp_value; rewrite /IntoVal /=; eauto; done.
iSplitL; last done.
iApply ("Hwp" with "[]"); last done.
rewrite /IsNode mapsto_node_eq /mapsto_node_def; iFrame "#"; iFrame; eauto.
Qed.
Heap
Lemma wp_alloc n s E v :
{{{ ▷ IsNode n }}} ⟨n; Alloc (Val v)⟩ @ s; E {{{ l, RET 〈n; #l〉; l ↦[n] v }}}.
Proof.
iIntros (Φ) ">Hn HΦ".
iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 ? ? ?) "H !>"; simpl in *.
iDestruct "H" as (m Hloch) "(Hsicoh & Hmaps & HlocS & Hrest)".
rewrite /IsNode. iDestruct "Hn" as (γ's) "Hn".
iDestruct (node_in_map with "Hn Hmaps") as %Hninm.
iDestruct (node_in_state_heaps with "HlocS") as %[h Hninss];
first by apply Hninm.
iSplitR.
{ iPureIntro. do 4 eexists. eapply LocalStepS; eauto. }
iIntros (v2 σ2 efs Hstep); inv_head_step.
rewrite /messages /messages_coherence //=. iFrame.
destruct γ's as [γh γs]; iSplitR; auto; iFrame. iNext.
iDestruct (node_local_state with "[HlocS]") as "(Hl & HlocS)";
first done; iFrame.
iDestruct "Hl" as (h' S' Hh Hs) "(Hn' & Hheap & Hsockets)";
rewrite Hh in Hninss; simplify_eq; simpl.
iMod (@gen_heap_alloc with "[Hheap]")
as "(Hheap & Hl & Hlm)"; eauto.
iDestruct (map_local_state_i_update_heaps with "HlocS") as "HlocS".
iModIntro. iSplitR "HΦ Hl Hn".
+ iExists m; iFrame.
iDestruct (node_local_state_rev with "[Hn' Hheap Hsockets] HlocS")
as "HX"; first done; simpl.
{ iExists _,_. simpl in *; iFrame; simpl. iPureIntro.
by rewrite lookup_insert. }
iFrame. rewrite /local_state_coherence.
rewrite (dom_insert_Some (D:=gset node) _ _ h) /= //.
+ iApply "HΦ". iExists _. iFrame.
Qed.
Lemma wp_load n s E l q v :
{{{ ▷ l ↦[n]{q} v }}}
⟨n; Load #l⟩ @ s; E
{{{ RET 〈n;v〉; l ↦[n]{q} v }}}.
Proof.
iIntros (Φ) ">Hl HΦ".
iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs n') "Hσ !>".
iDestruct "Hσ" as (m Hlcoh) "(Hsic & Hmaps & HlocS & Hrest)".
iDestruct "Hl" as ([γh γs]) "(Hn & Hl)".
iDestruct (node_in_map with "Hn Hmaps") as %Hninm.
iDestruct (node_in_state_heaps _ _ _ _ Hninm with "HlocS") as %[h Hninss].
iDestruct (node_local_state with "[HlocS]") as "(Hloc & HlocS)";
first done; iFrame.
iDestruct "Hloc" as (h' S' Hheap Hsockets) "(Hn' & Hheap & Hsockets)".
iDestruct (@gen_heap_valid with "Hheap Hl") as %?. iSplit.
{ iPureIntro. do 4 eexists. eapply LocalStepS; eauto. eapply LoadS; eauto. }
iIntros (e2 σ2 efs Hstep). inv_head_step.
iFrame. iNext; iModIntro; iSplit => //=.
iSplitR "HΦ Hl Hn".
+ iExists m; iFrame.
iDestruct (node_local_state_rev with "[Hn' Hheap Hsockets] HlocS")
as "HX"; first done; eauto.
{ iExists h',S'; iFrame; auto. }
{ rewrite insert_id; last auto. iFrame. auto. }
+ iApply "HΦ". iExists _; iFrame.
Qed.
Lemma wp_store n s E l v' v :
{{{ ▷ l ↦[n] v' }}}
⟨n; Store #l (Val v)⟩ @ s; E
{{{ RET 〈n; #()〉; l ↦[n] v }}}.
Proof.
iIntros (Φ) ">Hl HΦ". iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs n') "Hσ !>".
iDestruct "Hσ" as (m Hlcoh) "(Hsic & Hmaps & HlocS & Hrest)".
iDestruct "Hl" as ([γh γs]) "(Hn & Hl)".
iDestruct (node_in_map with "Hn Hmaps") as %Hninm.
iDestruct (node_in_state_heaps _ _ _ _ Hninm with "HlocS") as %[h Hninss].
iDestruct (node_local_state with "[HlocS]") as "(Hloc & HlocS) ";
first done; iFrame.
iDestruct "Hloc" as (h' S' Hheap Hsockets) "(Hn' & Hheap & Hsockets)".
rewrite Hheap in Hninss; simplify_eq.
iDestruct (@gen_heap_valid with "Hheap Hl") as %?.
iSplit.
{ iPureIntro; do 4 eexists. eapply LocalStepS; eauto.
eapply StoreS; eauto. }
iIntros (v2 σ2 efs Hstep); inv_head_step. iFrame; iNext.
iMod (gen_heap_update with "Hheap Hl")
as "(Hheap & Hl)".
iFrame. iModIntro. iSplit=>//.
iSplitR "HΦ Hl Hn".
+ iExists m; iFrame.
iDestruct (map_local_state_i_update_heaps with "HlocS") as "HlocS".
iDestruct (node_local_state_rev with "[Hn' Hheap Hsockets] HlocS")
as "HX"; first done; eauto.
iExists (<[l:=_]>h),S'.
iFrame; iSplit; iPureIntro;
[ apply lookup_insert | auto ].
iFrame. rewrite /local_state_coherence.
rewrite (dom_insert_Some (D:=gset node) _ _ h) /= //.
+ iApply "HΦ". iExists _. iFrame.
Qed.
Lemma wp_cas_fail n s E l q v' v1 v2 :
v' ≠ v1 →
{{{ ▷ l ↦[n]{q} v' }}}
⟨n;CAS #l (Val v1) (Val v2)⟩ @ s; E
{{{ RET 〈n; #false〉; l ↦[n]{q} v' }}}.
Proof.
iIntros (Heq Φ) ">Hl HΦ".
iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs n') "Hσ !>".
iDestruct "Hσ" as (m Hlcoh) "(Hsic & Hmaps & HlocS & Hrest)".
iDestruct "Hl" as ([γh γs]) "(Hn & Hl)".
iDestruct (node_in_map with "Hn Hmaps") as %Hninm.
iDestruct (node_in_state_heaps _ _ _ _ Hninm with "HlocS") as %[h Hninss].
iDestruct (node_local_state with "[HlocS]") as "(Hloc & HlocS) ";
first done; iFrame.
iDestruct "Hloc" as (h' S' Hheap Hsockets) "(Hn' & Hheap & Hsockets)".
iDestruct (@gen_heap_valid with "Hheap Hl") as %?. iSplit.
{ iPureIntro; do 4 eexists. eapply LocalStepS; eauto. eapply CasFailS; eauto. }
iIntros (v2' σ2 efs Hstep); inv_head_step.
iFrame; iNext; iModIntro; iSplit=>//; iSplitR "HΦ Hl Hn".
+ iExists m; iFrame.
iDestruct (node_local_state_rev with "[Hn' Hheap Hsockets] HlocS")
as "HX"; first done; eauto.
{ iExists h',S'; iFrame; auto. }
{ rewrite insert_id; last auto. iFrame. auto. }
+ iApply "HΦ". iExists _. iFrame.
Qed.
Lemma wp_cas_suc n s E l v1 v2 :
{{{ ▷ l ↦[n] v1 }}}
⟨n;CAS #l (Val v1) (Val v2)⟩ @ s; E
{{{ RET 〈n; #true〉; l ↦[n] v2 }}}.
Proof.
iIntros (Φ) ">Hl HΦ".
iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs n') "Hσ !>".
iDestruct "Hσ" as (m Hlcoh) "(Hsic & Hmaps & HlocS & Hrest)".
iDestruct "Hl" as ([γh γs]) "(Hn & Hl)".
iDestruct (node_in_map with "Hn Hmaps") as %Hninm.
iDestruct (node_in_state_heaps _ _ _ _ Hninm with "HlocS") as %[h Hninss].
iDestruct (node_local_state with "[HlocS]") as "(Hloc & HlocS) ";
first done; iFrame.
iDestruct "Hloc" as (h' S' Hheap Hsockets) "(Hn' & Hheap & Hsockets)".
iDestruct (@gen_heap_valid with "Hheap Hl") as %?. iSplit.
{ iPureIntro; do 4 eexists. eapply LocalStepS; eauto. eapply CasSucS; eauto. }
iIntros (v2' σ2 efs Hstep); inv_head_step. iFrame; iNext.
iMod (gen_heap_update with "Hheap Hl") as "(Hheap & Hl)".
iFrame. iModIntro. iSplit=>//.
iSplitR "HΦ Hl Hn".
- iExists m; iFrame.
iDestruct (map_local_state_i_update_heaps with "HlocS") as "HlocS".
iDestruct (node_local_state_rev with "[Hn' Hheap Hsockets] HlocS")
as "HX"; first done; eauto.
iExists (<[l:=_]>h),S'.
iFrame; iSplit; iPureIntro;
[ apply lookup_insert | auto ].
iFrame. rewrite /local_state_coherence.
rewrite (dom_insert_Some (D:=gset node) _ _ h) /= //.
- iApply "HΦ". iExists _. iFrame.
Qed.
(* Network *)
Lemma wp_new_socket n s E v1 v2 v3 :
{{{ ▷ IsNode n }}}
⟨n;NewSocket (Val $ LitV $ LitAddressFamily v1)
(Val $ LitV $ LitSocketType v2)
(Val $ LitV $ LitProtocol v3)⟩ @ s; E
{{{ h, RET 〈n;LitV (LitSocket h)〉;
h s↦[n] {| Network.sfamily := v1;
Network.stype := v2;
Network.sprotocol := v3;
Network.saddress := None |} }}}.
Proof.
iIntros (Φ) ">Hn HΦ".
iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs n') "H !>"; simpl in *.
iDestruct "H" as (m Hlcoh) "(Hsicoh & Hmaps & HlocS & Hrest)".
rewrite /IsNode. iDestruct "Hn" as (γ's) "Hn".
iDestruct (node_in_map with "Hn Hmaps") as %Hninm.
iDestruct (node_in_state_sockets _ _ _ _ Hninm with "HlocS") as %[S Hnins].
iDestruct (node_local_state with "[HlocS]") as "(Hs & HlocS) ";
first done; iFrame.
iSplitR.
{ iPureIntro; do 4 eexists; apply SocketStepS with (S := S) ; try auto; subst.
apply newsocket_fresh; by eauto. }
iIntros (v2' σ2 efs Hstep); inv_head_step.
destruct γ's as [γh γs]; iSplitR; auto; iFrame.
iDestruct "Hs" as (h' S'' Hh Hs) "(Hn' & Hheap & Hsockets)";
simplify_eq; simpl. iNext.
iMod ((gen_heap_alloc S handle) with "[Hsockets]") as "(Hsockets & Hs & _)"; auto.
iDestruct (map_local_state_i_update_sockets with "HlocS") as "HlocS".
iModIntro. iSplitR "HΦ Hs Hn"; iFrame.
+ iExists m. iFrame.
iDestruct (node_local_state_rev with "[Hn' Hheap Hsockets] HlocS")
as "HX"; first done; simpl.
{ iExists h',(<[handle:=_]> S).
simpl in *. iFrame. iPureIntro.
split; try auto. apply lookup_insert. }
try rewrite H6.
iFrame. rewrite /local_state_coherence.
rewrite (dom_insert_Some (D:=gset node) _ _ S) /= //.
+ iApply "HΦ". iExists _. iFrame.
Qed.
Lemma wp_socketbind_static n s A E sh v a :
saddress v = None →
a ∈ A →
{{{ Fixed A ∗ ▷ FreePorts (ip_of_address a) {[(port_of_address a)]} ∗ ▷ sh s↦[n] v }}}
⟨n; SocketBind (Val $ LitV $ LitSocket sh)
(Val $ LitV $ LitSocketAddress a)⟩ @ s; E
{{{ RET 〈n; #0〉;
∃ g, sh s↦[n] {| sfamily := sfamily v;
stype := stype v;
sprotocol := sprotocol v;
saddress := Some a |} ∗
a r↦{½} g ∗ ∃ φ, a ⤇ φ }}}.
Proof.
iIntros (Hnone HaA Φ) "(#Hf & >Hip & >Hs) HΦ".
iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs n') "Hσ".
iDestruct "Hσ" as (m Hlcoh)
"(Hsicoh & Hmaps & HlocS & Hips & Hms)".
iDestruct "Hips" as (Fip Piu (Hdsj & HFip & HPiu)) "[HFip HPiu]".
iMod (FreePorts_dealloc with "HPiu Hip")
as (ports [Hports Hportsa%elem_of_subseteq_singleton]) "HPiu".
iModIntro.
iDestruct "Hs" as ([γh γs]) "(Hn & Hs)".
iDestruct (node_in_map with "Hn Hmaps") as %Hninm.
iDestruct (node_in_state_sockets _ _ _ _ Hninm with "HlocS") as %[S Hsin].
iDestruct (node_local_state with "[HlocS]") as "(Hloc & HlocS) "; [done|iFrame|].
iDestruct "Hloc" as (h' S' Hh Hs) "(Hn' & Hheap & Hsockets)".
rewrite Hs in Hsin; simplify_eq; simpl.
iDestruct (@gen_heap_valid with "Hsockets Hs") as %Hlookup.
iDestruct (find_fixed_socket_interp with "Hf Hsicoh") as "#Haφ"; eauto.
iSplit.
- destruct (HPiu _ _ Hports) as [Q [Ha Hdisj]].
iPureIntro; do 4 eexists; eapply SocketStepS; eauto.
eapply (SocketBindSucS n sh a v S (state_ports_in_use σ1) _); try eauto.
- iIntros (v2 σ2 efs Hstep); inv_head_step.
iDestruct "Hms" as (seen rec [Hnetcoh [Hrec_coh Hsent_coh]])
"(Hms & Hsent & Hrec & Hrecs & Hmsgs)". iFrame. iNext.
iMod (gen_heap_update with "Hsockets Hs") as "(Hsockets & Hs)".
assert (rec !! a = None) as Hrecnone.
{ rewrite /network_coherence in Hnetcoh.
destruct (rec !! a) eqn:Heqrec; [|done].
match goal with
| H: port_of_address a ∉ ?P |- _ => cut (port_of_address a ∈ P)
end; first tauto.
apply Hnetcoh; eauto. }
iMod (@gen_heap_alloc _ _ _ _ _ _ rec a
(messages_received_at a (state_ms σ1)) with "Hrec")
as "[Hrec [Hreca _]]"; first done.
iDestruct (fractional_half with "Hreca") as "[Hreca Hreca']".
iModIntro. iSplit=>//.
iSplitR "HΦ Hs Hn Hreca'".
+ iExists m. iFrame.
iDestruct (map_local_state_i_update_sockets with "HlocS") as "HlocS".
iDestruct (node_local_state_rev with "[Hn' Hheap Hsockets] HlocS")
as "HlocS"; first done.
{ iExists _,_. iFrame; repeat iSplit; iPureIntro; eauto; apply lookup_insert. }
iFrame. iSplitR.
{ rewrite /local_state_coherence.
rewrite (dom_insert_Some (D:=gset node) _ _ S) /= //. }
iSplitL "Hsicoh". unfold network_coherence in *.
{ iDestruct "Hsicoh" as (si fx Hfx Hdms) "(#Hfx & Hsi & Hsis)".
iDestruct (ownF_agree with "Hf Hfx") as %?; subst A.
rewrite /socket_interp_coherence.
iExists _,_. iFrame; iFrame "#". iPureIntro. split; intros b.
{ intros Hb. specialize (Hfx _ Hb). by rewrite dom_insert elem_of_union; right. }
rewrite dom_insert elem_of_union. intros Hb. rewrite Hdms; last first.
{ destruct (decide (a = b)); subst; first by apply Hfx.
destruct Hb as [Hb|]; last done. apply elem_of_singleton in Hb.
rewrite Hb. by apply Hfx. }
split; intros [|[Hnb Hdm]]; auto; right; (split; first done).
* destruct (decide (ip_of_address a = ip_of_address b)) as [Heq|].
-- intros P. rewrite Heq. rewrite lookup_insert. intros; simplify_eq.
apply elem_of_union_r. apply Hdm. by rewrite -Heq.
-- intros P. rewrite lookup_insert_ne //. apply Hdm.
* destruct (decide (ip_of_address a = ip_of_address b)) as [Heq|].
-- intros P. rewrite Heq in Hdm. rewrite lookup_insert in Hdm.
match goal with
| H: port_of_address a ∉ ?P |- _ =>
specialize (Hdm ({[port_of_address a]} ∪ P) eq_refl)
end.
apply elem_of_union in Hdm. destruct Hdm as [Hdm | Hdm].
++ apply elem_of_singleton in Hdm. destruct a, b. simpl in *; by subst.
++ rewrite -Heq. by intros; simplify_eq.
-- rewrite lookup_insert_ne in Hdm; done. }
iSplitL "HFip HPiu".
{ iExists _, _. iFrame; iFrame "#". iPureIntro; repeat split.
- assert (ip_of_address a ∈ dom (gset _) Piu).
{ rewrite elem_of_dom; eauto. }
rewrite dom_insert. set_solver.
- intros ip Hip.
destruct (decide (ip = ip_of_address a)); subst.
+ exfalso; revert Hip. eapply elem_of_disjoint; eauto.
apply elem_of_dom; eauto.
+ rewrite lookup_insert_ne //; by apply HFip.
- intros ip Q.
destruct (decide (ip = ip_of_address a)); subst.
+ rewrite lookup_insert => ?; simplify_eq.
destruct (HPiu _ _ Hports) as [Q [Ha HQ]].
match goal with
Ha : state_ports_in_use σ1 !! ip_of_address a = Some _,
Hb : state_ports_in_use σ1 !! ip_of_address a = Some _ |- _ =>
rewrite Ha in Hb; simplify_eq
end.
exists ({[port_of_address a]} ∪ Q).
rewrite lookup_insert; split; first done.
set_solver.
+ rewrite !lookup_insert_ne //. apply HPiu. }
iExists seen,(<[a:=_]>rec). iSplitR.
* iPureIntro. split; simpl; first by apply network_coherence_insert.
split; last done. rewrite /messages_received_coherence. intros i M Hrec.
destruct (decide (a = i)); subst.
-- by rewrite lookup_insert in Hrec; simplify_eq.
-- rewrite lookup_insert_ne in Hrec; last done.
rewrite /messages_received_coherence in Hrec_coh.
by specialize (Hrec_coh i M Hrec).
* iFrame. rewrite big_sepM_insert; last done. iFrame. auto.
+ iApply "HΦ". iExists _. iFrame; iFrame "#". iExists _; iFrame.
Qed.
Lemma wp_socketbind_static_2 n s A E sh v a :
saddress v = None →
a ∈ A →
{{{ Fixed A ∗ ▷ FreePorts (ip_of_address a) {[(port_of_address a)]} ∗ ▷ sh s↦[n] v }}}
⟨n; SocketBind (Val $ LitV $ LitSocket sh)
(Val $ LitV $ LitSocketAddress a)⟩ @ s; E
{{{ RET 〈n; #0〉;
∃ g, sh s↦[n] {| sfamily := sfamily v;
stype := stype v;
sprotocol := sprotocol v;
saddress := Some a |} ∗
a r↦{½} g }}}.
Proof.
iIntros (? ? ?) "(#HA & Hports & Hsock) HΦ".
iApply (wp_socketbind_static with "[$HA Hports $Hsock]"); try auto.
iNext. iDestruct 1 as (M) "(Hsock & HM & Hψ)".
iDestruct "Hψ" as (Ψ) "#HΨ". iApply "HΦ".
iExists _. by iFrame.
Qed.
Lemma wp_socketbind_dynamic n s A E sh v a φ :
saddress v = None →
a ∉ A →
{{{ ▷ Fixed A ∗ ▷ FreePorts (ip_of_address a) {[port_of_address a]} ∗ ▷ sh s↦[n] v }}}
⟨n; SocketBind (Val $ LitV $ LitSocket sh)
(Val $ LitV $ LitSocketAddress a)⟩ @ s; E
{{{ RET 〈n; #0〉;
∃ g, sh s↦[n] {| sfamily := sfamily v;
stype := stype v;
sprotocol := sprotocol v;
saddress := Some a |} ∗
a r↦{½} g ∗ a ⤇ φ
}}}.
Proof.
iIntros (Hnone Hafix Φ) "(>#HA & >Hip & >Hs) HΦ".
iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs n') "Hσ".
iDestruct "Hσ" as (m Hncoh)
"(Hsicoh & Hmaps & HlocS & Hips & Hms)".
iDestruct "Hips" as (Fip Piu (Hdsj & HFip & HPiu)) "[HFip HPiu]".
iMod (FreePorts_dealloc with "HPiu Hip")
as (ports [Hports Hportsa%elem_of_subseteq_singleton]) "HPiu".
iModIntro.
iDestruct "Hs" as ([γh γs]) "(Hn & Hs)".
iDestruct (node_in_map with "Hn Hmaps") as %Hninm.
iDestruct (node_in_state_sockets _ _ _ _ Hninm with "HlocS") as %[S Hsin].
iDestruct (node_local_state with "[HlocS]") as "(Hloc & HlocS) "; [done|iFrame|].
iDestruct "Hloc" as (h' S' Hh Hs) "(Hn' & Hheap & Hsockets)".
rewrite Hs in Hsin; simplify_eq; simpl.
iDestruct (@gen_heap_valid with "Hsockets Hs") as %Hlookup.
iSplit.
- destruct (HPiu _ _ Hports) as [Q [Ha Hdisj]].
iPureIntro; do 4 eexists; eapply SocketStepS; eauto.
eapply (SocketBindSucS n sh a v S (state_ports_in_use σ1)); eauto.
- iIntros (v2 σ2 efs Hstep); inv_head_step.
iDestruct "Hms" as (seen rec [Hnetcoh [Hrec_coh Hsent_coh]])
"(Hms & Hsent & Hrec & Hrecs & Hmsgs)".
assert (rec !! a = None) as Hrecnone.
{ rewrite /network_coherence in Hnetcoh.
destruct (rec !! a) eqn:Heqrec; [|done].
match goal with
| _: port_of_address a ∉ ?P |- _ => cut (port_of_address a ∈ P)
end; first tauto.
apply Hnetcoh; eauto. }
rewrite /socket_interp_coherence.
iDestruct "Hsicoh" as (si fx Hfxsub Hdms) "(#Hfix & Hsi & #Hall)".
iDestruct (ownF_agree with "HA Hfix") as %?; subst.
assert (Hnotin': si !! a = None).
{ destruct (si !! a) eqn:Heq; last done.
destruct (Hdms a) as [[|[Hfx HP]] _]; try done.
- by eapply elem_of_dom_2.
- rewrite elem_of_dom. eauto.
- match goal with
| H: state_ports_in_use σ1 !! ip_of_address a = Some _ |- _ =>
by apply HP in H
end. }
iFrame. iNext.
iMod (saved_pred_alloc φ) as (γsi) "#Hsipred".
iMod (own_update _ _ (● <[a:=(to_agree γsi)]>si ⋅ ◯ {[ a:=(to_agree γsi)]})
with "Hsi") as "[Hsi #Hsip]".
{ apply auth_update_alloc. by apply (alloc_singleton_local_update si). }
iMod (gen_heap_update with "Hsockets Hs") as "(Hsockets & Hs)".
iMod (@gen_heap_alloc _ _ _ _ _ _ rec a
(messages_received_at a (state_ms σ1)) with "Hrec")
as "[Hrec [Hreca _]]"; first done.
iDestruct (fractional_half with "Hreca") as "[Hreca Hreca']".
iModIntro. iSplit=>//.
iSplitR "HΦ Hs Hn Hreca'".
+ iExists m. iFrame.
iDestruct (map_local_state_i_update_sockets with "HlocS") as "HlocS".
iDestruct (node_local_state_rev with "[Hn' Hheap Hsockets] HlocS")
as "HlocS"; first done.
{ iExists _,_. iFrame; repeat iSplit; iPureIntro; eauto; apply lookup_insert. }
iFrame. iSplitR.
{ iPureIntro. rewrite /local_state_coherence.
rewrite (dom_insert_Some (D:=gset node) _ _ S) /= //. }
iSplitL "Hsi Hall".
iExists (<[a:=(to_agree γsi)]>si), _; iFrame. iFrame "#".
iSplit.
{ iPureIntro. intros b Hb.
destruct (decide (a = b)); subst; first done.
apply dom_insert, elem_of_union. by auto. }
iSplitR.
{ iPureIntro. intros b Hbdom.
rewrite dom_insert elem_of_union elem_of_singleton Hdms; last first.
{ destruct (decide (a = b)).
- subst b. by eapply elem_of_dom_2.
- apply dom_insert, elem_of_union in Hbdom.
destruct Hbdom as [Hbdom|]; last done. apply elem_of_singleton in Hbdom.
rewrite Hbdom. by eapply elem_of_dom_2. }
split.
- intros [Hb | Hdom]; subst.
+ right. split; first done. rewrite lookup_insert.
intros; simplify_eq. by apply elem_of_union_l, elem_of_singleton.
+ destruct Hdom as [? | [Hb HP]]; first by left.
right. split; first done.
destruct (decide (ip_of_address a = ip_of_address b)) as [Heq |Hneq].
* rewrite Heq lookup_insert. intros; simplify_eq.
apply elem_of_union_r. apply HP. by rewrite -Heq.
* by rewrite lookup_insert_ne.
- intros [Hb | [Hb Hdom]]; first by auto.
destruct (decide (a = b)) as [Heq | Hneq]; first by auto.
right; right. split; first done.
destruct (decide (ip_of_address a = ip_of_address b)) as [Heq' |Hneq'].
+ intros P.
rewrite Heq' in Hdom.
rewrite lookup_insert in Hdom.
match goal with
| _: port_of_address a ∉ ?P |- _ =>
specialize (Hdom ({[port_of_address a]} ∪ P) eq_refl)
end.
apply elem_of_union in Hdom. destruct Hdom as [Hdom | Hport].
* apply elem_of_singleton in Hdom. destruct a,b; simpl in *; simplify_eq.
* intros.
match goal with
| HP: state_ports_in_use σ1 !! ip_of_address b = Some P |- _ =>
rewrite -Heq' in HP; by simplify_eq
end.
+ intros. apply Hdom. rewrite lookup_insert_ne; done. }
rewrite dom_insert_L big_sepS_insert.
iFrame; iFrame "#". iExists _, _; iFrame "#".
{ by rewrite not_elem_of_dom. }
iSplitL "HFip HPiu".
{ iExists _, _; iFrame. iPureIntro; repeat split.
- assert (ip_of_address a ∈ dom (gset _) Piu).
{ rewrite elem_of_dom; eauto. }
rewrite dom_insert. set_solver.
- intros ip Hip.
destruct (decide (ip = ip_of_address a)); subst.
+ exfalso; revert Hip. eapply elem_of_disjoint; eauto.
apply elem_of_dom; eauto.
+ rewrite lookup_insert_ne //; by apply HFip.
- intros ip Q. destruct (decide (ip = ip_of_address a)); subst.
+ rewrite lookup_insert => ?; simplify_eq.
destruct (HPiu _ _ Hports) as [Q [Ha HQ]].
match goal with
Ha : state_ports_in_use σ1 !! ip_of_address a = Some _,
Hb : state_ports_in_use σ1 !! ip_of_address a = Some _ |- _ =>
rewrite Ha in Hb; simplify_eq
end.
exists ({[port_of_address a]} ∪ Q).
rewrite lookup_insert; split; first done.
set_solver.
+ rewrite !lookup_insert_ne //. apply HPiu. }
iExists seen,(<[a:=_]>rec). iSplitR. iPureIntro. split; simpl.
-- by apply network_coherence_insert.
-- split; last done. rewrite /messages_received_coherence. intros.
destruct (decide (a = i)); subst.
++ by rewrite lookup_insert in H0; simplify_eq.
++ rewrite lookup_insert_ne in H0; last done.
rewrite /messages_received_coherence in Hrec_coh.
by specialize (Hrec_coh i M' H0).
-- iFrame. rewrite big_sepM_insert //=. by iFrame.
+ iApply "HΦ". iExists _. iFrame. iSplitL "Hn Hs".
iExists _; iFrame. iFrame. iExists _. iFrame; iFrame "#".
Qed.
Lemma wp_send_to_bound φ m sh a f n s E v π :
saddress v = Some f ->
let msg := {|
m_from_node := n;
m_sender := f;
m_destination := a;
m_protocol := sprotocol v;
m_state := MS_SENT;
m_body := m;
|} in
{{{ ▷ sh s↦[n]{π} v ∗ a ⤇ φ ∗ φ (message_stable_from_message msg) }}}
⟨n; SendTo (Val $ LitV $ LitSocket sh) #m #a⟩ @ s; E
{{{ RET 〈n; #(String.length m)〉;
sh s↦[n]{π} v ∗ ∃ mId, mId o↦{¾} (message_stable_from_message (msg)) }}}.
Proof.
iIntros (Hsome Hmsg Φ) "(>Hsocket & #Hsip & Hsipred) HΦ"; simpl.
iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs n') "Hσ !>".
iDestruct "Hσ" as (γmap Hncoh)
"(Hsi & Hmaps & HlocS & Hips & Hms)".
iDestruct "Hsocket" as ([γh γs]) "(Hn & Hs)".
iDestruct (node_in_map with "Hn Hmaps") as %Hninm.
iDestruct (node_in_state_sockets _ _ _ _ Hninm with "HlocS") as %[S Hnins].
iDestruct (node_local_state with "[HlocS]") as "(Hloc & HlocS) ";
first done; iFrame.
iDestruct "Hloc" as (h' S' Hheaps Hsockets) "(Hmap & Hheap & Hsockets)".
rewrite Hsockets in Hnins; simplify_eq; simpl.
iDestruct (@gen_heap_valid with "Hsockets Hs") as %Hv.
iSplitR.
{ iPureIntro; do 4 eexists; eapply SocketStepS; eauto.
eapply message_send_bound_fresh; eauto. }
iIntros (v2' σ2 efs Hstep); inv_head_step;
iSplitR; auto; subst. iNext.
iDestruct (message_state_later with "Hms") as "Hms".
iMod (messages_insert _ _ mId {|
m_from_node := n;
m_sender := f;
m_destination := a;
m_protocol := sprotocol v;
m_body := m;
m_state := MS_SENT |} with "Hms") as "(HmId & Hstable & Hms)";
try done.
iDestruct "Hms" as (sent rec [Hnetcoh [Hreccoh Hsentcoh]])
"(Hms & Hsent & Hrec & Hrecs & Hmsgs)".
iFrame. iSplitR "HΦ Hn Hs Hstable".
- iModIntro. iExists γmap. iFrame. iSplitR.
{ rewrite /local_state_coherence.
rewrite (dom_insert_Some (D:=gset node) _ _ S') /= //. }
iSplitL "Hmap Hheap Hsockets HlocS"; auto.
+ iDestruct (node_local_state_rev with "[Hmap Hheap Hsockets] HlocS")
as "HlocS"; first done.
{ iExists _,_. iFrame. auto. }
rewrite /local_state_i /= insert_id //.
+ iExists _,_. iFrame. iSplitR; first by auto.
iSplitL "Hrecs".
{ rewrite big_sepM_mono //. iIntros (k x Hl) "(% & H)". iFrame.
by rewrite messages_received_id. }
rewrite big_sepM_insert //. iFrame.
iExists 1%Qp,φ. iFrame. iRight.
rewrite /Hmsg /=. iFrame; iFrame "#". done.
- iModIntro. simpl. iApply "HΦ"; iFrame.
rewrite -{1}Qp_quarter_three_quarter.
iDestruct "Hstable" as "[_ Hstable]".
iSplitR "Hstable"; by (iExists _; iFrame).
Qed.
Lemma wp_send_to_unbound P Q φ m a n s E sh v π ip :
saddress v = None ->
let msg M f := {| m_from_node := n;
m_sender := f;
m_destination := a;
m_protocol := sprotocol v;
m_state := MS_SENT;
m_body := m; |} in
{{{ ▷ sh s↦[n]{π} v ∗ ▷ P ∗ a ⤇ φ ∗ FreePorts ip ∅ ∗
(∀ M sent mId f,
message_soup_info M sent ∗
mId o↦ (message_stable_from_message (msg M f)) ∗ P ={E}=∗
message_soup_info M sent ∗
▷ φ (message_stable_from_message (msg M f)) ∗ Q
)}}}
⟨n;SendTo (Val $ LitV $ LitSocket sh) #m #a⟩ @ s; E
{{{ RET 〈n;LitV (LitInt (String.length m))〉;
sh s↦[n]{π} v ∗ Q
}}}.
Proof.
iIntros (Hnone Hmsg Φ) "(>Hsocket & Hp & #Hsip & Hfp & Hvs) HΦ"; simpl.
iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs n') "Hσ !>".
iDestruct "Hσ" as (γmap Hncoh)
"(Hsi & Hmaps & HlocS & Hips & Hms)".
iDestruct "Hips" as (Fip Piu (Hdsj & HFip & HPiu)) "[HFip HPiu]".
iDestruct (FreePorts_included with "HPiu Hfp") as %[Prts [Hprts _]].
destruct (HPiu _ _ Hprts) as [Prts' [Hptrs' Hdsj']].
iDestruct "Hsocket" as ([γh γs]) "(Hn & Hs)".
iDestruct (node_in_map with "Hn Hmaps") as %Hninm.
iDestruct (node_in_state_sockets _ _ _ _ Hninm with "HlocS") as %[S Hnins].
iDestruct (node_local_state with "[HlocS]") as "(Hloc & HlocS) ";
first done; iFrame.
iDestruct "Hloc" as (h' S' Hheaps Hsockets) "(Hmap & Hheap & Hsockets)".
rewrite Hsockets in Hnins; simplify_eq; simpl.
iDestruct (@gen_heap_valid with "Hsockets Hs") as %Hv.
iSplitR.
{ iPureIntro. do 4 eexists; eapply SocketStepS; eauto.
eapply message_send_unbound_fresh; eauto. }
iIntros (v2' σ2 efs Hstep); inv_head_step;
iSplitR; auto; subst; iNext.
iDestruct (message_state_later with "Hms") as "Hms".
iMod (messages_insert _ _ mId {|
m_from_node := n;
m_sender := f;
m_destination := a;
m_protocol := sprotocol v;
m_body := m;
m_state := MS_SENT; |} with "Hms") as "(HmId & Hstable & Hms)";
try done.
iDestruct "Hms" as (sent rec [Hsoccoh [Hreccoh Hsentcoh]])
"(Hms & Hsent & Hrec & Hrecs & Hmsgs)".
iMod ("Hvs" $! _ _ mId f with "[$Hms $Hsent Hstable $Hp]")
as "((_ & Hms & Hsent) & Hsipred & HQ)".
{ iSplitR; first done.
done.
}
iFrame. iSplitR "HΦ HQ Hn Hs".
- iModIntro. iExists γmap. iFrame. iSplitR.
{ rewrite /local_state_coherence.
rewrite (dom_insert_Some (D:=gset node) _ _ S') /= //. }
iSplitL "Hheap Hmap Hsockets HlocS".
+ iDestruct (node_local_state_rev with "[Hmap Hheap Hsockets] HlocS")
as "HlocS"; first done.
{ iExists _,_. iFrame. auto. }
rewrite /local_state_i /= insert_id //.
+ iFrame. iSplitL "HFip HPiu".
{ iExists _, _; iFrame; eauto. }
iExists _,_. iFrame. iSplitR. done. iSplitL "Hrecs".
rewrite big_sepM_mono //.
iIntros (k x Hl) "(% & H)". iFrame. by rewrite messages_received_id.
rewrite big_sepM_insert //. iFrame. iExists 1%Qp,φ. iFrame. iRight.
rewrite /Hmsg /=. iFrame; iFrame "#". done.
- iModIntro. simpl.
iApply "HΦ"; iFrame. by iExists _; iFrame.
Qed.
Definition received_message_info a m :=
m_destination m = a ∧ m_state m = MS_RECEIVED.
Lemma wp_receive_from n s a E sh v π rm :
saddress v = Some a →
{{{ ▷ sh s↦[n]{π} v ∗ ▷ a r↦{½} rm }}}
⟨n;ReceiveFrom (Val $ LitV $ LitSocket sh)⟩ @ s; E
{{{ r, RET 〈n;r〉; ∃ rm',
sh s↦[n]{π} v ∗ a r↦{½} rm' ∗
((∃ mId m φ, ⌜received_message_info a m⌝ ∗
⌜r = SOMEV (PairV (LitV $ LitString (m_body m))
(LitV $ LitSocketAddress (m_sender m)))⌝ ∗
⌜rm' = <[mId:=m]>rm⌝ ∗
mId m↦{¾} m ∗ a ⤇ φ ∗ φ (message_stable_from_message m)) ∨
(⌜r = NONEV⌝ ∗ ⌜rm' = rm⌝))
}}}.
Proof.
iIntros (Hla Φ) "(>Hs & >Ha) HΦ".
iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs n') "Hσ !>".
iDestruct "Hσ" as (γmap Hncoh)
"(Hsi & Hmaps & HlocS & Hips & Hms)".
iDestruct "Hs" as ([γh γs]) "(Hn & Hs)".
iDestruct (node_in_map with "Hn Hmaps") as %Hninm.
iDestruct (node_in_state_sockets _ _ _ _ Hninm with "HlocS") as %[S Hnins].
iDestruct (node_local_state with "[HlocS]") as "(Hloc & HlocS) ";
first done; iFrame.
iDestruct "Hloc" as (h' S' Hheap Hsocket) "(Hn' & Hheap & Hsockets)".
rewrite Hsocket in Hnins; simplify_eq; simpl.
iDestruct (@gen_heap_valid with "Hsockets Hs") as %Hsh.
iSplitR.
{ iPureIntro; do 4 eexists; eapply SocketStepS; eauto.
by eapply ReceiveFromNoneS. }
iNext; iIntros (v2' σ2 efs Hstep); inv_head_step;
iSplitR; auto; subst; last first.
- iModIntro. iSplitR "HΦ Hs Hn Ha"; iFrame.
iExists _; iFrame.
iDestruct (node_local_state_rev with "[Hn' Hheap Hsockets] HlocS")
as "HX"; first done.
{ iExists h', S'. iFrame; by iPureIntro. }
iDestruct (message_state_later with "Hms") as "Hms".
rewrite insert_id //=; iFrame; auto.
iApply "HΦ". iExists _. iFrame. iSplitL. iExists _. iFrame. by iRight.
- iDestruct "Hms" as (sent rec [Hmsgsi [Hrec Hsent]])
"(Hms & Hsent & Hrec & Hrecs & Hmsg)".
iDestruct (@gen_heap_valid with "Hrec Ha") as %Hreca.
iDestruct (big_sepM_delete _ rec with "Hrecs") as "([% Hreca] & Hrecs)";
first done. iCombine "Ha" "Hreca" as "Ha".
iMod (gen_heap_update rec a _ (<[mId:={|
m_from_node := m_from_node m;
m_sender := m_sender m;
m_destination := m_destination m;
m_protocol := m_protocol m;
m_body := m_body m;
m_state := MS_RECEIVED; |}]>
(messages_received_at a (state_ms σ1)))
with "Hrec Ha") as "(Hrec & [Ha Ha'])".
revert H3. rewrite map_filter_lookup_Some.
intros [Hlookup [Hp1 Hp2]]; simpl in *; subst.
iMod (messages_update_received
(state_ms σ1) mId _
{|
m_from_node := m_from_node m;
m_sender := m_sender m;
m_destination := m_destination m;
m_protocol := m_protocol m;
m_body := m_body m;
m_state := MS_RECEIVED; |}
with "[Hms Hmsg]")
as (φ) "(Hsi' & Hp & Hms & Hmsg)"; try done; iFrame.
iDestruct (message_sepM_later with "Hmsg") as "Hmsg". iFrame.
iSplitR "HΦ Hs Hn Hsi' Hp Ha'".
* iExists γmap. iFrame. iSplitR.
{ rewrite /local_state_coherence.
rewrite (dom_insert_Some (D:=gset node) _ _ S') /= //. }
{ iDestruct (node_local_state_rev with "[Hn' Hheap Hsockets] HlocS")
as "HlocS"; first done.
{ iExists h', S'. iFrame; by iPureIntro. }
rewrite (insert_id (state_sockets σ1)); last done. iFrame.
iExists _,_. iFrame. iSplitR. iPureIntro.
rewrite /messages_coherence /network_coherence. auto. split.
{ rewrite /network_coherence in Hmsgsi.
intros ? ? Hrecl ?. apply Hmsgsi; last done.
rewrite -(insert_id rec (m_destination m)
(messages_received_at (m_destination m) (state_ms σ1))); last done.
apply lookup_insert_is_Some in Hrecl. destruct Hrecl as [Ha | [Ha Hrecl]].
- subst. rewrite lookup_insert. by eapply mk_is_Some.
- by rewrite lookup_insert_ne. }
split. unfold messages_received_coherence, messages_received_at in *.
intros i M'.
rewrite lookup_insert_Some. intros [[Heq Hinsert] | [Hneq Hlook]].
- rewrite -Hinsert -map_filter_insert Heq //.
- rewrite map_filter_lookup_None_insert. by apply Hrec.
rewrite map_filter_lookup_None. right. intros; simplify_eq.
intros [Hh1 Hh2]; simpl in *. rewrite Hp2 in Hh2. inversion Hh2.
intros [Hh1 Hh2]. simpl in *. by destruct Hneq.
- rewrite /messages_sent_coherence //=. split.
rewrite (dom_insert_Some (state_ms σ1) mId m) //=. by apply Hsent.
destruct Hsent as [Hsent1 Hsent2].
intros. destruct (decide (mId = mId0)); trivial; subst.
+ destruct m; simpl in *. specialize (Hsent2 mId0).
rewrite lookup_insert Hsent2 Hlookup Hp2
/message_stable_from_message //=.
+ rewrite lookup_insert_ne //=.
- iFrame.
iDestruct (messages_received_at_update with "Hrecs Ha") as "Hx";
auto; try done.
intro. rewrite Hp2 //= in H0; done. }
* iApply "HΦ". iModIntro.
iExists _. iFrame. iSplitL "Hs Hn". iExists _. iFrame.
iLeft. iExists mId,{|
m_from_node := m_from_node m;
m_sender := m_sender m;
m_destination := m_destination m;
m_protocol := m_protocol m;
m_body := m_body m;
m_state := MS_RECEIVED; |},φ.
iFrame. iSplitR; first done. iSplitR. iPureIntro; first done.
iSplitR; first done.
rewrite /message_stable_from_message Hp2 /=. iFrame.
Qed.
Lemma wp_receive_from_2 n s a E sh v π rm φ :
saddress v = Some a →
{{{ ▷ sh s↦[n]{π} v ∗ ▷ a r↦{½} rm ∗ ▷ a ⤇ φ }}}
⟨n;ReceiveFrom (Val $ LitV $ LitSocket sh)⟩ @ s; E
{{{ r, RET 〈n;r〉; sh s↦[n]{π} v ∗
(* a message was received *)
((∃ m i, ⌜r = SOMEV (#(m_body m), #(m_sender m))⌝
(* meta information about the message *)
∗ ⌜received_message_info a m⌝
(* extended list of received messages *)
∗ a r↦{½} <[i:=m]>rm
(* certificate that m was actually received *)
∗ i m↦{¾} m
(* the message satisfies the socket protocol *)
∗ ▷ φ (message_stable_from_message m))
(* no message received *)
∨ (⌜r = NONEV⌝ ∗ a r↦{½} rm))
}}}.
Proof.
iIntros (Hbound Φ) "(>Hs & Hr & #Hsi) HΦ".
iApply (wp_receive_from with "[$Hs $Hr]"); first done.
iNext. iDestruct 1 as (r) "H".
iDestruct "H" as "(Hsocket & Hr & [H | [-> ->]])".
- iDestruct "H" as (? m ψ) "(?&?&%&?&Hsi'&?)". subst.
iApply "HΦ". iFrame. iLeft. iExists _, _. iFrame.
iDestruct (si_pred_agree _ _ _ (message_stable_from_message m)
with "Hsi Hsi'") as "Hsiagree".
iNext. by iRewrite "Hsiagree".
- iApply "HΦ". iFrame. iRight. by iFrame.
Qed.
End lifting.
{{{ ▷ IsNode n }}} ⟨n; Alloc (Val v)⟩ @ s; E {{{ l, RET 〈n; #l〉; l ↦[n] v }}}.
Proof.
iIntros (Φ) ">Hn HΦ".
iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 ? ? ?) "H !>"; simpl in *.
iDestruct "H" as (m Hloch) "(Hsicoh & Hmaps & HlocS & Hrest)".
rewrite /IsNode. iDestruct "Hn" as (γ's) "Hn".
iDestruct (node_in_map with "Hn Hmaps") as %Hninm.
iDestruct (node_in_state_heaps with "HlocS") as %[h Hninss];
first by apply Hninm.
iSplitR.
{ iPureIntro. do 4 eexists. eapply LocalStepS; eauto. }
iIntros (v2 σ2 efs Hstep); inv_head_step.
rewrite /messages /messages_coherence //=. iFrame.
destruct γ's as [γh γs]; iSplitR; auto; iFrame. iNext.
iDestruct (node_local_state with "[HlocS]") as "(Hl & HlocS)";
first done; iFrame.
iDestruct "Hl" as (h' S' Hh Hs) "(Hn' & Hheap & Hsockets)";
rewrite Hh in Hninss; simplify_eq; simpl.
iMod (@gen_heap_alloc with "[Hheap]")
as "(Hheap & Hl & Hlm)"; eauto.
iDestruct (map_local_state_i_update_heaps with "HlocS") as "HlocS".
iModIntro. iSplitR "HΦ Hl Hn".
+ iExists m; iFrame.
iDestruct (node_local_state_rev with "[Hn' Hheap Hsockets] HlocS")
as "HX"; first done; simpl.
{ iExists _,_. simpl in *; iFrame; simpl. iPureIntro.
by rewrite lookup_insert. }
iFrame. rewrite /local_state_coherence.
rewrite (dom_insert_Some (D:=gset node) _ _ h) /= //.
+ iApply "HΦ". iExists _. iFrame.
Qed.
Lemma wp_load n s E l q v :
{{{ ▷ l ↦[n]{q} v }}}
⟨n; Load #l⟩ @ s; E
{{{ RET 〈n;v〉; l ↦[n]{q} v }}}.
Proof.
iIntros (Φ) ">Hl HΦ".
iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs n') "Hσ !>".
iDestruct "Hσ" as (m Hlcoh) "(Hsic & Hmaps & HlocS & Hrest)".
iDestruct "Hl" as ([γh γs]) "(Hn & Hl)".
iDestruct (node_in_map with "Hn Hmaps") as %Hninm.
iDestruct (node_in_state_heaps _ _ _ _ Hninm with "HlocS") as %[h Hninss].
iDestruct (node_local_state with "[HlocS]") as "(Hloc & HlocS)";
first done; iFrame.
iDestruct "Hloc" as (h' S' Hheap Hsockets) "(Hn' & Hheap & Hsockets)".
iDestruct (@gen_heap_valid with "Hheap Hl") as %?. iSplit.
{ iPureIntro. do 4 eexists. eapply LocalStepS; eauto. eapply LoadS; eauto. }
iIntros (e2 σ2 efs Hstep). inv_head_step.
iFrame. iNext; iModIntro; iSplit => //=.
iSplitR "HΦ Hl Hn".
+ iExists m; iFrame.
iDestruct (node_local_state_rev with "[Hn' Hheap Hsockets] HlocS")
as "HX"; first done; eauto.
{ iExists h',S'; iFrame; auto. }
{ rewrite insert_id; last auto. iFrame. auto. }
+ iApply "HΦ". iExists _; iFrame.
Qed.
Lemma wp_store n s E l v' v :
{{{ ▷ l ↦[n] v' }}}
⟨n; Store #l (Val v)⟩ @ s; E
{{{ RET 〈n; #()〉; l ↦[n] v }}}.
Proof.
iIntros (Φ) ">Hl HΦ". iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs n') "Hσ !>".
iDestruct "Hσ" as (m Hlcoh) "(Hsic & Hmaps & HlocS & Hrest)".
iDestruct "Hl" as ([γh γs]) "(Hn & Hl)".
iDestruct (node_in_map with "Hn Hmaps") as %Hninm.
iDestruct (node_in_state_heaps _ _ _ _ Hninm with "HlocS") as %[h Hninss].
iDestruct (node_local_state with "[HlocS]") as "(Hloc & HlocS) ";
first done; iFrame.
iDestruct "Hloc" as (h' S' Hheap Hsockets) "(Hn' & Hheap & Hsockets)".
rewrite Hheap in Hninss; simplify_eq.
iDestruct (@gen_heap_valid with "Hheap Hl") as %?.
iSplit.
{ iPureIntro; do 4 eexists. eapply LocalStepS; eauto.
eapply StoreS; eauto. }
iIntros (v2 σ2 efs Hstep); inv_head_step. iFrame; iNext.
iMod (gen_heap_update with "Hheap Hl")
as "(Hheap & Hl)".
iFrame. iModIntro. iSplit=>//.
iSplitR "HΦ Hl Hn".
+ iExists m; iFrame.
iDestruct (map_local_state_i_update_heaps with "HlocS") as "HlocS".
iDestruct (node_local_state_rev with "[Hn' Hheap Hsockets] HlocS")
as "HX"; first done; eauto.
iExists (<[l:=_]>h),S'.
iFrame; iSplit; iPureIntro;
[ apply lookup_insert | auto ].
iFrame. rewrite /local_state_coherence.
rewrite (dom_insert_Some (D:=gset node) _ _ h) /= //.
+ iApply "HΦ". iExists _. iFrame.
Qed.
Lemma wp_cas_fail n s E l q v' v1 v2 :
v' ≠ v1 →
{{{ ▷ l ↦[n]{q} v' }}}
⟨n;CAS #l (Val v1) (Val v2)⟩ @ s; E
{{{ RET 〈n; #false〉; l ↦[n]{q} v' }}}.
Proof.
iIntros (Heq Φ) ">Hl HΦ".
iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs n') "Hσ !>".
iDestruct "Hσ" as (m Hlcoh) "(Hsic & Hmaps & HlocS & Hrest)".
iDestruct "Hl" as ([γh γs]) "(Hn & Hl)".
iDestruct (node_in_map with "Hn Hmaps") as %Hninm.
iDestruct (node_in_state_heaps _ _ _ _ Hninm with "HlocS") as %[h Hninss].
iDestruct (node_local_state with "[HlocS]") as "(Hloc & HlocS) ";
first done; iFrame.
iDestruct "Hloc" as (h' S' Hheap Hsockets) "(Hn' & Hheap & Hsockets)".
iDestruct (@gen_heap_valid with "Hheap Hl") as %?. iSplit.
{ iPureIntro; do 4 eexists. eapply LocalStepS; eauto. eapply CasFailS; eauto. }
iIntros (v2' σ2 efs Hstep); inv_head_step.
iFrame; iNext; iModIntro; iSplit=>//; iSplitR "HΦ Hl Hn".
+ iExists m; iFrame.
iDestruct (node_local_state_rev with "[Hn' Hheap Hsockets] HlocS")
as "HX"; first done; eauto.
{ iExists h',S'; iFrame; auto. }
{ rewrite insert_id; last auto. iFrame. auto. }
+ iApply "HΦ". iExists _. iFrame.
Qed.
Lemma wp_cas_suc n s E l v1 v2 :
{{{ ▷ l ↦[n] v1 }}}
⟨n;CAS #l (Val v1) (Val v2)⟩ @ s; E
{{{ RET 〈n; #true〉; l ↦[n] v2 }}}.
Proof.
iIntros (Φ) ">Hl HΦ".
iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs n') "Hσ !>".
iDestruct "Hσ" as (m Hlcoh) "(Hsic & Hmaps & HlocS & Hrest)".
iDestruct "Hl" as ([γh γs]) "(Hn & Hl)".
iDestruct (node_in_map with "Hn Hmaps") as %Hninm.
iDestruct (node_in_state_heaps _ _ _ _ Hninm with "HlocS") as %[h Hninss].
iDestruct (node_local_state with "[HlocS]") as "(Hloc & HlocS) ";
first done; iFrame.
iDestruct "Hloc" as (h' S' Hheap Hsockets) "(Hn' & Hheap & Hsockets)".
iDestruct (@gen_heap_valid with "Hheap Hl") as %?. iSplit.
{ iPureIntro; do 4 eexists. eapply LocalStepS; eauto. eapply CasSucS; eauto. }
iIntros (v2' σ2 efs Hstep); inv_head_step. iFrame; iNext.
iMod (gen_heap_update with "Hheap Hl") as "(Hheap & Hl)".
iFrame. iModIntro. iSplit=>//.
iSplitR "HΦ Hl Hn".
- iExists m; iFrame.
iDestruct (map_local_state_i_update_heaps with "HlocS") as "HlocS".
iDestruct (node_local_state_rev with "[Hn' Hheap Hsockets] HlocS")
as "HX"; first done; eauto.
iExists (<[l:=_]>h),S'.
iFrame; iSplit; iPureIntro;
[ apply lookup_insert | auto ].
iFrame. rewrite /local_state_coherence.
rewrite (dom_insert_Some (D:=gset node) _ _ h) /= //.
- iApply "HΦ". iExists _. iFrame.
Qed.
(* Network *)
Lemma wp_new_socket n s E v1 v2 v3 :
{{{ ▷ IsNode n }}}
⟨n;NewSocket (Val $ LitV $ LitAddressFamily v1)
(Val $ LitV $ LitSocketType v2)
(Val $ LitV $ LitProtocol v3)⟩ @ s; E
{{{ h, RET 〈n;LitV (LitSocket h)〉;
h s↦[n] {| Network.sfamily := v1;
Network.stype := v2;
Network.sprotocol := v3;
Network.saddress := None |} }}}.
Proof.
iIntros (Φ) ">Hn HΦ".
iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs n') "H !>"; simpl in *.
iDestruct "H" as (m Hlcoh) "(Hsicoh & Hmaps & HlocS & Hrest)".
rewrite /IsNode. iDestruct "Hn" as (γ's) "Hn".
iDestruct (node_in_map with "Hn Hmaps") as %Hninm.
iDestruct (node_in_state_sockets _ _ _ _ Hninm with "HlocS") as %[S Hnins].
iDestruct (node_local_state with "[HlocS]") as "(Hs & HlocS) ";
first done; iFrame.
iSplitR.
{ iPureIntro; do 4 eexists; apply SocketStepS with (S := S) ; try auto; subst.
apply newsocket_fresh; by eauto. }
iIntros (v2' σ2 efs Hstep); inv_head_step.
destruct γ's as [γh γs]; iSplitR; auto; iFrame.
iDestruct "Hs" as (h' S'' Hh Hs) "(Hn' & Hheap & Hsockets)";
simplify_eq; simpl. iNext.
iMod ((gen_heap_alloc S handle) with "[Hsockets]") as "(Hsockets & Hs & _)"; auto.
iDestruct (map_local_state_i_update_sockets with "HlocS") as "HlocS".
iModIntro. iSplitR "HΦ Hs Hn"; iFrame.
+ iExists m. iFrame.
iDestruct (node_local_state_rev with "[Hn' Hheap Hsockets] HlocS")
as "HX"; first done; simpl.
{ iExists h',(<[handle:=_]> S).
simpl in *. iFrame. iPureIntro.
split; try auto. apply lookup_insert. }
try rewrite H6.
iFrame. rewrite /local_state_coherence.
rewrite (dom_insert_Some (D:=gset node) _ _ S) /= //.
+ iApply "HΦ". iExists _. iFrame.
Qed.
Lemma wp_socketbind_static n s A E sh v a :
saddress v = None →
a ∈ A →
{{{ Fixed A ∗ ▷ FreePorts (ip_of_address a) {[(port_of_address a)]} ∗ ▷ sh s↦[n] v }}}
⟨n; SocketBind (Val $ LitV $ LitSocket sh)
(Val $ LitV $ LitSocketAddress a)⟩ @ s; E
{{{ RET 〈n; #0〉;
∃ g, sh s↦[n] {| sfamily := sfamily v;
stype := stype v;
sprotocol := sprotocol v;
saddress := Some a |} ∗
a r↦{½} g ∗ ∃ φ, a ⤇ φ }}}.
Proof.
iIntros (Hnone HaA Φ) "(#Hf & >Hip & >Hs) HΦ".
iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs n') "Hσ".
iDestruct "Hσ" as (m Hlcoh)
"(Hsicoh & Hmaps & HlocS & Hips & Hms)".
iDestruct "Hips" as (Fip Piu (Hdsj & HFip & HPiu)) "[HFip HPiu]".
iMod (FreePorts_dealloc with "HPiu Hip")
as (ports [Hports Hportsa%elem_of_subseteq_singleton]) "HPiu".
iModIntro.
iDestruct "Hs" as ([γh γs]) "(Hn & Hs)".
iDestruct (node_in_map with "Hn Hmaps") as %Hninm.
iDestruct (node_in_state_sockets _ _ _ _ Hninm with "HlocS") as %[S Hsin].
iDestruct (node_local_state with "[HlocS]") as "(Hloc & HlocS) "; [done|iFrame|].
iDestruct "Hloc" as (h' S' Hh Hs) "(Hn' & Hheap & Hsockets)".
rewrite Hs in Hsin; simplify_eq; simpl.
iDestruct (@gen_heap_valid with "Hsockets Hs") as %Hlookup.
iDestruct (find_fixed_socket_interp with "Hf Hsicoh") as "#Haφ"; eauto.
iSplit.
- destruct (HPiu _ _ Hports) as [Q [Ha Hdisj]].
iPureIntro; do 4 eexists; eapply SocketStepS; eauto.
eapply (SocketBindSucS n sh a v S (state_ports_in_use σ1) _); try eauto.
- iIntros (v2 σ2 efs Hstep); inv_head_step.
iDestruct "Hms" as (seen rec [Hnetcoh [Hrec_coh Hsent_coh]])
"(Hms & Hsent & Hrec & Hrecs & Hmsgs)". iFrame. iNext.
iMod (gen_heap_update with "Hsockets Hs") as "(Hsockets & Hs)".
assert (rec !! a = None) as Hrecnone.
{ rewrite /network_coherence in Hnetcoh.
destruct (rec !! a) eqn:Heqrec; [|done].
match goal with
| H: port_of_address a ∉ ?P |- _ => cut (port_of_address a ∈ P)
end; first tauto.
apply Hnetcoh; eauto. }
iMod (@gen_heap_alloc _ _ _ _ _ _ rec a
(messages_received_at a (state_ms σ1)) with "Hrec")
as "[Hrec [Hreca _]]"; first done.
iDestruct (fractional_half with "Hreca") as "[Hreca Hreca']".
iModIntro. iSplit=>//.
iSplitR "HΦ Hs Hn Hreca'".
+ iExists m. iFrame.
iDestruct (map_local_state_i_update_sockets with "HlocS") as "HlocS".
iDestruct (node_local_state_rev with "[Hn' Hheap Hsockets] HlocS")
as "HlocS"; first done.
{ iExists _,_. iFrame; repeat iSplit; iPureIntro; eauto; apply lookup_insert. }
iFrame. iSplitR.
{ rewrite /local_state_coherence.
rewrite (dom_insert_Some (D:=gset node) _ _ S) /= //. }
iSplitL "Hsicoh". unfold network_coherence in *.
{ iDestruct "Hsicoh" as (si fx Hfx Hdms) "(#Hfx & Hsi & Hsis)".
iDestruct (ownF_agree with "Hf Hfx") as %?; subst A.
rewrite /socket_interp_coherence.
iExists _,_. iFrame; iFrame "#". iPureIntro. split; intros b.
{ intros Hb. specialize (Hfx _ Hb). by rewrite dom_insert elem_of_union; right. }
rewrite dom_insert elem_of_union. intros Hb. rewrite Hdms; last first.
{ destruct (decide (a = b)); subst; first by apply Hfx.
destruct Hb as [Hb|]; last done. apply elem_of_singleton in Hb.
rewrite Hb. by apply Hfx. }
split; intros [|[Hnb Hdm]]; auto; right; (split; first done).
* destruct (decide (ip_of_address a = ip_of_address b)) as [Heq|].
-- intros P. rewrite Heq. rewrite lookup_insert. intros; simplify_eq.
apply elem_of_union_r. apply Hdm. by rewrite -Heq.
-- intros P. rewrite lookup_insert_ne //. apply Hdm.
* destruct (decide (ip_of_address a = ip_of_address b)) as [Heq|].
-- intros P. rewrite Heq in Hdm. rewrite lookup_insert in Hdm.
match goal with
| H: port_of_address a ∉ ?P |- _ =>
specialize (Hdm ({[port_of_address a]} ∪ P) eq_refl)
end.
apply elem_of_union in Hdm. destruct Hdm as [Hdm | Hdm].
++ apply elem_of_singleton in Hdm. destruct a, b. simpl in *; by subst.
++ rewrite -Heq. by intros; simplify_eq.
-- rewrite lookup_insert_ne in Hdm; done. }
iSplitL "HFip HPiu".
{ iExists _, _. iFrame; iFrame "#". iPureIntro; repeat split.
- assert (ip_of_address a ∈ dom (gset _) Piu).
{ rewrite elem_of_dom; eauto. }
rewrite dom_insert. set_solver.
- intros ip Hip.
destruct (decide (ip = ip_of_address a)); subst.
+ exfalso; revert Hip. eapply elem_of_disjoint; eauto.
apply elem_of_dom; eauto.
+ rewrite lookup_insert_ne //; by apply HFip.
- intros ip Q.
destruct (decide (ip = ip_of_address a)); subst.
+ rewrite lookup_insert => ?; simplify_eq.
destruct (HPiu _ _ Hports) as [Q [Ha HQ]].
match goal with
Ha : state_ports_in_use σ1 !! ip_of_address a = Some _,
Hb : state_ports_in_use σ1 !! ip_of_address a = Some _ |- _ =>
rewrite Ha in Hb; simplify_eq
end.
exists ({[port_of_address a]} ∪ Q).
rewrite lookup_insert; split; first done.
set_solver.
+ rewrite !lookup_insert_ne //. apply HPiu. }
iExists seen,(<[a:=_]>rec). iSplitR.
* iPureIntro. split; simpl; first by apply network_coherence_insert.
split; last done. rewrite /messages_received_coherence. intros i M Hrec.
destruct (decide (a = i)); subst.
-- by rewrite lookup_insert in Hrec; simplify_eq.
-- rewrite lookup_insert_ne in Hrec; last done.
rewrite /messages_received_coherence in Hrec_coh.
by specialize (Hrec_coh i M Hrec).
* iFrame. rewrite big_sepM_insert; last done. iFrame. auto.
+ iApply "HΦ". iExists _. iFrame; iFrame "#". iExists _; iFrame.
Qed.
Lemma wp_socketbind_static_2 n s A E sh v a :
saddress v = None →
a ∈ A →
{{{ Fixed A ∗ ▷ FreePorts (ip_of_address a) {[(port_of_address a)]} ∗ ▷ sh s↦[n] v }}}
⟨n; SocketBind (Val $ LitV $ LitSocket sh)
(Val $ LitV $ LitSocketAddress a)⟩ @ s; E
{{{ RET 〈n; #0〉;
∃ g, sh s↦[n] {| sfamily := sfamily v;
stype := stype v;
sprotocol := sprotocol v;
saddress := Some a |} ∗
a r↦{½} g }}}.
Proof.
iIntros (? ? ?) "(#HA & Hports & Hsock) HΦ".
iApply (wp_socketbind_static with "[$HA Hports $Hsock]"); try auto.
iNext. iDestruct 1 as (M) "(Hsock & HM & Hψ)".
iDestruct "Hψ" as (Ψ) "#HΨ". iApply "HΦ".
iExists _. by iFrame.
Qed.
Lemma wp_socketbind_dynamic n s A E sh v a φ :
saddress v = None →
a ∉ A →
{{{ ▷ Fixed A ∗ ▷ FreePorts (ip_of_address a) {[port_of_address a]} ∗ ▷ sh s↦[n] v }}}
⟨n; SocketBind (Val $ LitV $ LitSocket sh)
(Val $ LitV $ LitSocketAddress a)⟩ @ s; E
{{{ RET 〈n; #0〉;
∃ g, sh s↦[n] {| sfamily := sfamily v;
stype := stype v;
sprotocol := sprotocol v;
saddress := Some a |} ∗
a r↦{½} g ∗ a ⤇ φ
}}}.
Proof.
iIntros (Hnone Hafix Φ) "(>#HA & >Hip & >Hs) HΦ".
iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs n') "Hσ".
iDestruct "Hσ" as (m Hncoh)
"(Hsicoh & Hmaps & HlocS & Hips & Hms)".
iDestruct "Hips" as (Fip Piu (Hdsj & HFip & HPiu)) "[HFip HPiu]".
iMod (FreePorts_dealloc with "HPiu Hip")
as (ports [Hports Hportsa%elem_of_subseteq_singleton]) "HPiu".
iModIntro.
iDestruct "Hs" as ([γh γs]) "(Hn & Hs)".
iDestruct (node_in_map with "Hn Hmaps") as %Hninm.
iDestruct (node_in_state_sockets _ _ _ _ Hninm with "HlocS") as %[S Hsin].
iDestruct (node_local_state with "[HlocS]") as "(Hloc & HlocS) "; [done|iFrame|].
iDestruct "Hloc" as (h' S' Hh Hs) "(Hn' & Hheap & Hsockets)".
rewrite Hs in Hsin; simplify_eq; simpl.
iDestruct (@gen_heap_valid with "Hsockets Hs") as %Hlookup.
iSplit.
- destruct (HPiu _ _ Hports) as [Q [Ha Hdisj]].
iPureIntro; do 4 eexists; eapply SocketStepS; eauto.
eapply (SocketBindSucS n sh a v S (state_ports_in_use σ1)); eauto.
- iIntros (v2 σ2 efs Hstep); inv_head_step.
iDestruct "Hms" as (seen rec [Hnetcoh [Hrec_coh Hsent_coh]])
"(Hms & Hsent & Hrec & Hrecs & Hmsgs)".
assert (rec !! a = None) as Hrecnone.
{ rewrite /network_coherence in Hnetcoh.
destruct (rec !! a) eqn:Heqrec; [|done].
match goal with
| _: port_of_address a ∉ ?P |- _ => cut (port_of_address a ∈ P)
end; first tauto.
apply Hnetcoh; eauto. }
rewrite /socket_interp_coherence.
iDestruct "Hsicoh" as (si fx Hfxsub Hdms) "(#Hfix & Hsi & #Hall)".
iDestruct (ownF_agree with "HA Hfix") as %?; subst.
assert (Hnotin': si !! a = None).
{ destruct (si !! a) eqn:Heq; last done.
destruct (Hdms a) as [[|[Hfx HP]] _]; try done.
- by eapply elem_of_dom_2.
- rewrite elem_of_dom. eauto.
- match goal with
| H: state_ports_in_use σ1 !! ip_of_address a = Some _ |- _ =>
by apply HP in H
end. }
iFrame. iNext.
iMod (saved_pred_alloc φ) as (γsi) "#Hsipred".
iMod (own_update _ _ (● <[a:=(to_agree γsi)]>si ⋅ ◯ {[ a:=(to_agree γsi)]})
with "Hsi") as "[Hsi #Hsip]".
{ apply auth_update_alloc. by apply (alloc_singleton_local_update si). }
iMod (gen_heap_update with "Hsockets Hs") as "(Hsockets & Hs)".
iMod (@gen_heap_alloc _ _ _ _ _ _ rec a
(messages_received_at a (state_ms σ1)) with "Hrec")
as "[Hrec [Hreca _]]"; first done.
iDestruct (fractional_half with "Hreca") as "[Hreca Hreca']".
iModIntro. iSplit=>//.
iSplitR "HΦ Hs Hn Hreca'".
+ iExists m. iFrame.
iDestruct (map_local_state_i_update_sockets with "HlocS") as "HlocS".
iDestruct (node_local_state_rev with "[Hn' Hheap Hsockets] HlocS")
as "HlocS"; first done.
{ iExists _,_. iFrame; repeat iSplit; iPureIntro; eauto; apply lookup_insert. }
iFrame. iSplitR.
{ iPureIntro. rewrite /local_state_coherence.
rewrite (dom_insert_Some (D:=gset node) _ _ S) /= //. }
iSplitL "Hsi Hall".
iExists (<[a:=(to_agree γsi)]>si), _; iFrame. iFrame "#".
iSplit.
{ iPureIntro. intros b Hb.
destruct (decide (a = b)); subst; first done.
apply dom_insert, elem_of_union. by auto. }
iSplitR.
{ iPureIntro. intros b Hbdom.
rewrite dom_insert elem_of_union elem_of_singleton Hdms; last first.
{ destruct (decide (a = b)).
- subst b. by eapply elem_of_dom_2.
- apply dom_insert, elem_of_union in Hbdom.
destruct Hbdom as [Hbdom|]; last done. apply elem_of_singleton in Hbdom.
rewrite Hbdom. by eapply elem_of_dom_2. }
split.
- intros [Hb | Hdom]; subst.
+ right. split; first done. rewrite lookup_insert.
intros; simplify_eq. by apply elem_of_union_l, elem_of_singleton.
+ destruct Hdom as [? | [Hb HP]]; first by left.
right. split; first done.
destruct (decide (ip_of_address a = ip_of_address b)) as [Heq |Hneq].
* rewrite Heq lookup_insert. intros; simplify_eq.
apply elem_of_union_r. apply HP. by rewrite -Heq.
* by rewrite lookup_insert_ne.
- intros [Hb | [Hb Hdom]]; first by auto.
destruct (decide (a = b)) as [Heq | Hneq]; first by auto.
right; right. split; first done.
destruct (decide (ip_of_address a = ip_of_address b)) as [Heq' |Hneq'].
+ intros P.
rewrite Heq' in Hdom.
rewrite lookup_insert in Hdom.
match goal with
| _: port_of_address a ∉ ?P |- _ =>
specialize (Hdom ({[port_of_address a]} ∪ P) eq_refl)
end.
apply elem_of_union in Hdom. destruct Hdom as [Hdom | Hport].
* apply elem_of_singleton in Hdom. destruct a,b; simpl in *; simplify_eq.
* intros.
match goal with
| HP: state_ports_in_use σ1 !! ip_of_address b = Some P |- _ =>
rewrite -Heq' in HP; by simplify_eq
end.
+ intros. apply Hdom. rewrite lookup_insert_ne; done. }
rewrite dom_insert_L big_sepS_insert.
iFrame; iFrame "#". iExists _, _; iFrame "#".
{ by rewrite not_elem_of_dom. }
iSplitL "HFip HPiu".
{ iExists _, _; iFrame. iPureIntro; repeat split.
- assert (ip_of_address a ∈ dom (gset _) Piu).
{ rewrite elem_of_dom; eauto. }
rewrite dom_insert. set_solver.
- intros ip Hip.
destruct (decide (ip = ip_of_address a)); subst.
+ exfalso; revert Hip. eapply elem_of_disjoint; eauto.
apply elem_of_dom; eauto.
+ rewrite lookup_insert_ne //; by apply HFip.
- intros ip Q. destruct (decide (ip = ip_of_address a)); subst.
+ rewrite lookup_insert => ?; simplify_eq.
destruct (HPiu _ _ Hports) as [Q [Ha HQ]].
match goal with
Ha : state_ports_in_use σ1 !! ip_of_address a = Some _,
Hb : state_ports_in_use σ1 !! ip_of_address a = Some _ |- _ =>
rewrite Ha in Hb; simplify_eq
end.
exists ({[port_of_address a]} ∪ Q).
rewrite lookup_insert; split; first done.
set_solver.
+ rewrite !lookup_insert_ne //. apply HPiu. }
iExists seen,(<[a:=_]>rec). iSplitR. iPureIntro. split; simpl.
-- by apply network_coherence_insert.
-- split; last done. rewrite /messages_received_coherence. intros.
destruct (decide (a = i)); subst.
++ by rewrite lookup_insert in H0; simplify_eq.
++ rewrite lookup_insert_ne in H0; last done.
rewrite /messages_received_coherence in Hrec_coh.
by specialize (Hrec_coh i M' H0).
-- iFrame. rewrite big_sepM_insert //=. by iFrame.
+ iApply "HΦ". iExists _. iFrame. iSplitL "Hn Hs".
iExists _; iFrame. iFrame. iExists _. iFrame; iFrame "#".
Qed.
Lemma wp_send_to_bound φ m sh a f n s E v π :
saddress v = Some f ->
let msg := {|
m_from_node := n;
m_sender := f;
m_destination := a;
m_protocol := sprotocol v;
m_state := MS_SENT;
m_body := m;
|} in
{{{ ▷ sh s↦[n]{π} v ∗ a ⤇ φ ∗ φ (message_stable_from_message msg) }}}
⟨n; SendTo (Val $ LitV $ LitSocket sh) #m #a⟩ @ s; E
{{{ RET 〈n; #(String.length m)〉;
sh s↦[n]{π} v ∗ ∃ mId, mId o↦{¾} (message_stable_from_message (msg)) }}}.
Proof.
iIntros (Hsome Hmsg Φ) "(>Hsocket & #Hsip & Hsipred) HΦ"; simpl.
iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs n') "Hσ !>".
iDestruct "Hσ" as (γmap Hncoh)
"(Hsi & Hmaps & HlocS & Hips & Hms)".
iDestruct "Hsocket" as ([γh γs]) "(Hn & Hs)".
iDestruct (node_in_map with "Hn Hmaps") as %Hninm.
iDestruct (node_in_state_sockets _ _ _ _ Hninm with "HlocS") as %[S Hnins].
iDestruct (node_local_state with "[HlocS]") as "(Hloc & HlocS) ";
first done; iFrame.
iDestruct "Hloc" as (h' S' Hheaps Hsockets) "(Hmap & Hheap & Hsockets)".
rewrite Hsockets in Hnins; simplify_eq; simpl.
iDestruct (@gen_heap_valid with "Hsockets Hs") as %Hv.
iSplitR.
{ iPureIntro; do 4 eexists; eapply SocketStepS; eauto.
eapply message_send_bound_fresh; eauto. }
iIntros (v2' σ2 efs Hstep); inv_head_step;
iSplitR; auto; subst. iNext.
iDestruct (message_state_later with "Hms") as "Hms".
iMod (messages_insert _ _ mId {|
m_from_node := n;
m_sender := f;
m_destination := a;
m_protocol := sprotocol v;
m_body := m;
m_state := MS_SENT |} with "Hms") as "(HmId & Hstable & Hms)";
try done.
iDestruct "Hms" as (sent rec [Hnetcoh [Hreccoh Hsentcoh]])
"(Hms & Hsent & Hrec & Hrecs & Hmsgs)".
iFrame. iSplitR "HΦ Hn Hs Hstable".
- iModIntro. iExists γmap. iFrame. iSplitR.
{ rewrite /local_state_coherence.
rewrite (dom_insert_Some (D:=gset node) _ _ S') /= //. }
iSplitL "Hmap Hheap Hsockets HlocS"; auto.
+ iDestruct (node_local_state_rev with "[Hmap Hheap Hsockets] HlocS")
as "HlocS"; first done.
{ iExists _,_. iFrame. auto. }
rewrite /local_state_i /= insert_id //.
+ iExists _,_. iFrame. iSplitR; first by auto.
iSplitL "Hrecs".
{ rewrite big_sepM_mono //. iIntros (k x Hl) "(% & H)". iFrame.
by rewrite messages_received_id. }
rewrite big_sepM_insert //. iFrame.
iExists 1%Qp,φ. iFrame. iRight.
rewrite /Hmsg /=. iFrame; iFrame "#". done.
- iModIntro. simpl. iApply "HΦ"; iFrame.
rewrite -{1}Qp_quarter_three_quarter.
iDestruct "Hstable" as "[_ Hstable]".
iSplitR "Hstable"; by (iExists _; iFrame).
Qed.
Lemma wp_send_to_unbound P Q φ m a n s E sh v π ip :
saddress v = None ->
let msg M f := {| m_from_node := n;
m_sender := f;
m_destination := a;
m_protocol := sprotocol v;
m_state := MS_SENT;
m_body := m; |} in
{{{ ▷ sh s↦[n]{π} v ∗ ▷ P ∗ a ⤇ φ ∗ FreePorts ip ∅ ∗
(∀ M sent mId f,
message_soup_info M sent ∗
mId o↦ (message_stable_from_message (msg M f)) ∗ P ={E}=∗
message_soup_info M sent ∗
▷ φ (message_stable_from_message (msg M f)) ∗ Q
)}}}
⟨n;SendTo (Val $ LitV $ LitSocket sh) #m #a⟩ @ s; E
{{{ RET 〈n;LitV (LitInt (String.length m))〉;
sh s↦[n]{π} v ∗ Q
}}}.
Proof.
iIntros (Hnone Hmsg Φ) "(>Hsocket & Hp & #Hsip & Hfp & Hvs) HΦ"; simpl.
iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs n') "Hσ !>".
iDestruct "Hσ" as (γmap Hncoh)
"(Hsi & Hmaps & HlocS & Hips & Hms)".
iDestruct "Hips" as (Fip Piu (Hdsj & HFip & HPiu)) "[HFip HPiu]".
iDestruct (FreePorts_included with "HPiu Hfp") as %[Prts [Hprts _]].
destruct (HPiu _ _ Hprts) as [Prts' [Hptrs' Hdsj']].
iDestruct "Hsocket" as ([γh γs]) "(Hn & Hs)".
iDestruct (node_in_map with "Hn Hmaps") as %Hninm.
iDestruct (node_in_state_sockets _ _ _ _ Hninm with "HlocS") as %[S Hnins].
iDestruct (node_local_state with "[HlocS]") as "(Hloc & HlocS) ";
first done; iFrame.
iDestruct "Hloc" as (h' S' Hheaps Hsockets) "(Hmap & Hheap & Hsockets)".
rewrite Hsockets in Hnins; simplify_eq; simpl.
iDestruct (@gen_heap_valid with "Hsockets Hs") as %Hv.
iSplitR.
{ iPureIntro. do 4 eexists; eapply SocketStepS; eauto.
eapply message_send_unbound_fresh; eauto. }
iIntros (v2' σ2 efs Hstep); inv_head_step;
iSplitR; auto; subst; iNext.
iDestruct (message_state_later with "Hms") as "Hms".
iMod (messages_insert _ _ mId {|
m_from_node := n;
m_sender := f;
m_destination := a;
m_protocol := sprotocol v;
m_body := m;
m_state := MS_SENT; |} with "Hms") as "(HmId & Hstable & Hms)";
try done.
iDestruct "Hms" as (sent rec [Hsoccoh [Hreccoh Hsentcoh]])
"(Hms & Hsent & Hrec & Hrecs & Hmsgs)".
iMod ("Hvs" $! _ _ mId f with "[$Hms $Hsent Hstable $Hp]")
as "((_ & Hms & Hsent) & Hsipred & HQ)".
{ iSplitR; first done.
done.
}
iFrame. iSplitR "HΦ HQ Hn Hs".
- iModIntro. iExists γmap. iFrame. iSplitR.
{ rewrite /local_state_coherence.
rewrite (dom_insert_Some (D:=gset node) _ _ S') /= //. }
iSplitL "Hheap Hmap Hsockets HlocS".
+ iDestruct (node_local_state_rev with "[Hmap Hheap Hsockets] HlocS")
as "HlocS"; first done.
{ iExists _,_. iFrame. auto. }
rewrite /local_state_i /= insert_id //.
+ iFrame. iSplitL "HFip HPiu".
{ iExists _, _; iFrame; eauto. }
iExists _,_. iFrame. iSplitR. done. iSplitL "Hrecs".
rewrite big_sepM_mono //.
iIntros (k x Hl) "(% & H)". iFrame. by rewrite messages_received_id.
rewrite big_sepM_insert //. iFrame. iExists 1%Qp,φ. iFrame. iRight.
rewrite /Hmsg /=. iFrame; iFrame "#". done.
- iModIntro. simpl.
iApply "HΦ"; iFrame. by iExists _; iFrame.
Qed.
Definition received_message_info a m :=
m_destination m = a ∧ m_state m = MS_RECEIVED.
Lemma wp_receive_from n s a E sh v π rm :
saddress v = Some a →
{{{ ▷ sh s↦[n]{π} v ∗ ▷ a r↦{½} rm }}}
⟨n;ReceiveFrom (Val $ LitV $ LitSocket sh)⟩ @ s; E
{{{ r, RET 〈n;r〉; ∃ rm',
sh s↦[n]{π} v ∗ a r↦{½} rm' ∗
((∃ mId m φ, ⌜received_message_info a m⌝ ∗
⌜r = SOMEV (PairV (LitV $ LitString (m_body m))
(LitV $ LitSocketAddress (m_sender m)))⌝ ∗
⌜rm' = <[mId:=m]>rm⌝ ∗
mId m↦{¾} m ∗ a ⤇ φ ∗ φ (message_stable_from_message m)) ∨
(⌜r = NONEV⌝ ∗ ⌜rm' = rm⌝))
}}}.
Proof.
iIntros (Hla Φ) "(>Hs & >Ha) HΦ".
iApply wp_lift_atomic_head_step_no_fork; auto.
iIntros (σ1 κ κs n') "Hσ !>".
iDestruct "Hσ" as (γmap Hncoh)
"(Hsi & Hmaps & HlocS & Hips & Hms)".
iDestruct "Hs" as ([γh γs]) "(Hn & Hs)".
iDestruct (node_in_map with "Hn Hmaps") as %Hninm.
iDestruct (node_in_state_sockets _ _ _ _ Hninm with "HlocS") as %[S Hnins].
iDestruct (node_local_state with "[HlocS]") as "(Hloc & HlocS) ";
first done; iFrame.
iDestruct "Hloc" as (h' S' Hheap Hsocket) "(Hn' & Hheap & Hsockets)".
rewrite Hsocket in Hnins; simplify_eq; simpl.
iDestruct (@gen_heap_valid with "Hsockets Hs") as %Hsh.
iSplitR.
{ iPureIntro; do 4 eexists; eapply SocketStepS; eauto.
by eapply ReceiveFromNoneS. }
iNext; iIntros (v2' σ2 efs Hstep); inv_head_step;
iSplitR; auto; subst; last first.
- iModIntro. iSplitR "HΦ Hs Hn Ha"; iFrame.
iExists _; iFrame.
iDestruct (node_local_state_rev with "[Hn' Hheap Hsockets] HlocS")
as "HX"; first done.
{ iExists h', S'. iFrame; by iPureIntro. }
iDestruct (message_state_later with "Hms") as "Hms".
rewrite insert_id //=; iFrame; auto.
iApply "HΦ". iExists _. iFrame. iSplitL. iExists _. iFrame. by iRight.
- iDestruct "Hms" as (sent rec [Hmsgsi [Hrec Hsent]])
"(Hms & Hsent & Hrec & Hrecs & Hmsg)".
iDestruct (@gen_heap_valid with "Hrec Ha") as %Hreca.
iDestruct (big_sepM_delete _ rec with "Hrecs") as "([% Hreca] & Hrecs)";
first done. iCombine "Ha" "Hreca" as "Ha".
iMod (gen_heap_update rec a _ (<[mId:={|
m_from_node := m_from_node m;
m_sender := m_sender m;
m_destination := m_destination m;
m_protocol := m_protocol m;
m_body := m_body m;
m_state := MS_RECEIVED; |}]>
(messages_received_at a (state_ms σ1)))
with "Hrec Ha") as "(Hrec & [Ha Ha'])".
revert H3. rewrite map_filter_lookup_Some.
intros [Hlookup [Hp1 Hp2]]; simpl in *; subst.
iMod (messages_update_received
(state_ms σ1) mId _
{|
m_from_node := m_from_node m;
m_sender := m_sender m;
m_destination := m_destination m;
m_protocol := m_protocol m;
m_body := m_body m;
m_state := MS_RECEIVED; |}
with "[Hms Hmsg]")
as (φ) "(Hsi' & Hp & Hms & Hmsg)"; try done; iFrame.
iDestruct (message_sepM_later with "Hmsg") as "Hmsg". iFrame.
iSplitR "HΦ Hs Hn Hsi' Hp Ha'".
* iExists γmap. iFrame. iSplitR.
{ rewrite /local_state_coherence.
rewrite (dom_insert_Some (D:=gset node) _ _ S') /= //. }
{ iDestruct (node_local_state_rev with "[Hn' Hheap Hsockets] HlocS")
as "HlocS"; first done.
{ iExists h', S'. iFrame; by iPureIntro. }
rewrite (insert_id (state_sockets σ1)); last done. iFrame.
iExists _,_. iFrame. iSplitR. iPureIntro.
rewrite /messages_coherence /network_coherence. auto. split.
{ rewrite /network_coherence in Hmsgsi.
intros ? ? Hrecl ?. apply Hmsgsi; last done.
rewrite -(insert_id rec (m_destination m)
(messages_received_at (m_destination m) (state_ms σ1))); last done.
apply lookup_insert_is_Some in Hrecl. destruct Hrecl as [Ha | [Ha Hrecl]].
- subst. rewrite lookup_insert. by eapply mk_is_Some.
- by rewrite lookup_insert_ne. }
split. unfold messages_received_coherence, messages_received_at in *.
intros i M'.
rewrite lookup_insert_Some. intros [[Heq Hinsert] | [Hneq Hlook]].
- rewrite -Hinsert -map_filter_insert Heq //.
- rewrite map_filter_lookup_None_insert. by apply Hrec.
rewrite map_filter_lookup_None. right. intros; simplify_eq.
intros [Hh1 Hh2]; simpl in *. rewrite Hp2 in Hh2. inversion Hh2.
intros [Hh1 Hh2]. simpl in *. by destruct Hneq.
- rewrite /messages_sent_coherence //=. split.
rewrite (dom_insert_Some (state_ms σ1) mId m) //=. by apply Hsent.
destruct Hsent as [Hsent1 Hsent2].
intros. destruct (decide (mId = mId0)); trivial; subst.
+ destruct m; simpl in *. specialize (Hsent2 mId0).
rewrite lookup_insert Hsent2 Hlookup Hp2
/message_stable_from_message //=.
+ rewrite lookup_insert_ne //=.
- iFrame.
iDestruct (messages_received_at_update with "Hrecs Ha") as "Hx";
auto; try done.
intro. rewrite Hp2 //= in H0; done. }
* iApply "HΦ". iModIntro.
iExists _. iFrame. iSplitL "Hs Hn". iExists _. iFrame.
iLeft. iExists mId,{|
m_from_node := m_from_node m;
m_sender := m_sender m;
m_destination := m_destination m;
m_protocol := m_protocol m;
m_body := m_body m;
m_state := MS_RECEIVED; |},φ.
iFrame. iSplitR; first done. iSplitR. iPureIntro; first done.
iSplitR; first done.
rewrite /message_stable_from_message Hp2 /=. iFrame.
Qed.
Lemma wp_receive_from_2 n s a E sh v π rm φ :
saddress v = Some a →
{{{ ▷ sh s↦[n]{π} v ∗ ▷ a r↦{½} rm ∗ ▷ a ⤇ φ }}}
⟨n;ReceiveFrom (Val $ LitV $ LitSocket sh)⟩ @ s; E
{{{ r, RET 〈n;r〉; sh s↦[n]{π} v ∗
(* a message was received *)
((∃ m i, ⌜r = SOMEV (#(m_body m), #(m_sender m))⌝
(* meta information about the message *)
∗ ⌜received_message_info a m⌝
(* extended list of received messages *)
∗ a r↦{½} <[i:=m]>rm
(* certificate that m was actually received *)
∗ i m↦{¾} m
(* the message satisfies the socket protocol *)
∗ ▷ φ (message_stable_from_message m))
(* no message received *)
∨ (⌜r = NONEV⌝ ∗ a r↦{½} rm))
}}}.
Proof.
iIntros (Hbound Φ) "(>Hs & Hr & #Hsi) HΦ".
iApply (wp_receive_from with "[$Hs $Hr]"); first done.
iNext. iDestruct 1 as (r) "H".
iDestruct "H" as "(Hsocket & Hr & [H | [-> ->]])".
- iDestruct "H" as (? m ψ) "(?&?&%&?&Hsi'&?)". subst.
iApply "HΦ". iFrame. iLeft. iExists _, _. iFrame.
iDestruct (si_pred_agree _ _ _ (message_stable_from_message m)
with "Hsi Hsi'") as "Hsiagree".
iNext. by iRewrite "Hsiagree".
- iApply "HΦ". iFrame. iRight. by iFrame.
Qed.
End lifting.