aneris.dist_lang.resources

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.
From stdpp Require Import fin_maps gmap.
Set Default Proof Using "Type".

Import uPred.
Import Network.

Record node_gnames := Node_gname {
  heap_name : gname;
  heap_meta_name : gname;
  socket_name : gname;
  socket_meta_name : gname;
}.

Definition node_gnamesO := leibnizO node_gnames.
Definition socket_received := gmap socket_address message_soup.
Definition system_state_mapUR : ucmraT := gmapUR node (agreeR node_gnamesO).

Definition socket_interpR : cmraT := authR (gmapUR socket_address
                                                   (agreeR (leibnizO gname))).

Instance system_state_mapUR_unit : Unit (gmap node (agree node_gnames)) :=
  ( : gmap node (agree node_gnames)).
Global Instance system_state_core_id (x : system_state_mapUR) : CoreId x.
Proof. apply _. Qed.

Definition socket_interp Σ := message_stable -d> iProp Σ.

The system CMRA we need.
Class distG Σ := DistG {
  dist_invG :> invG Σ;
  dist_mapG :> inG Σ (authR system_state_mapUR);
  dist_map_name : gname;
  dist_heapG :> gen_heapPreG loc ground_lang.val Σ;
  (* network *)
  dist_socketG :> gen_heapPreG socket_handle socket Σ;
  dist_messagesG :> gen_heapG message_id message Σ;
  dist_messagesSentG :> gen_heapG message_id message_stable Σ;
  dist_messagesReceivedG :> gen_heapG socket_address message_soup Σ;
  dist_freeipsG :> inG Σ (authUR (gset_disjUR ip_address));
  dist_freeips_name : gname;
  dist_freeportsG :> inG Σ (authUR (gmapUR ip_address (gset_disjUR port)));
  dist_freeports_name : gname;
  dist_siG :> inG Σ socket_interpR;
  dist_si_name : gname;
  dist_fixedG :> inG Σ (agreeR (gsetUR socket_address));
  dist_fixed_name : gname;
  dist_savedPredG :> savedPredG Σ message_stable
}.

Arguments gen_heap_name {_ _ _ _ _} _ : assert.
Arguments gen_meta_name {_ _ _ _ _} _ : assert.

Definition gen_heap_loc_instance γ γ' `{distG Σ} :=
  GenHeapG loc ground_lang.val _ _ _ _ _ _ γ γ'.

Definition gen_heap_soc_instance γ γ' `{distG Σ} :=
  GenHeapG socket_handle socket _ _ _ _ _ _ γ γ'.

Section Definitions.
  Context `{dG : distG Σ}.

  Definition mapsto_node_def (n : node) (x : node_gnames) :=
    own (dist_map_name) ( {[ n := to_agree x ]} : auth system_state_mapUR).
  Definition mapsto_node_aux : seal (@mapsto_node_def). by eexists. Qed.
  Definition mapsto_node := unseal mapsto_node_aux.
  Definition mapsto_node_eq : @mapsto_node = @mapsto_node_def :=
    seal_eq mapsto_node_aux.

  Global Instance mapsto_node_persistent n x : Persistent (mapsto_node n x).
  Proof. rewrite mapsto_node_eq /mapsto_node_def. apply _. Qed.

  Definition mapsto_l n l q v :=
    ( γ's, mapsto_node n γ's
            mapsto (hG := gen_heap_loc_instance (heap_name γ's) (heap_meta_name γ's)) l q v)%I.
  Definition mapsto_s n h q s :=
    ( γ's, mapsto_node n γ's
            mapsto (hG := gen_heap_soc_instance (socket_name γ's) (socket_meta_name γ's)) h q s)%I.

End Definitions.

For some reason, Coq cannot find these RAs, so we define them explicitly
Definition ownS γ (m : gmap node node_gnames) `{distG Σ} :=
  own (A := authR system_state_mapUR) γ ( (to_agree <$> m)).

Definition ownF (f : gset socket_address) `{distG} :=
  own dist_fixed_name (to_agree f).

Notation "n n↦ x" := (mapsto_node n x)
                       (at level 20, format "n n↦ x") : uPred_scope.

Open Scope uPred_scope.

Global Instance mapsto_n_timeless `{distG Σ} (n : node) v : Timeless (n n v).
Proof. rewrite mapsto_node_eq /mapsto_node_def. apply _. Qed.

Global Instance mapsto_l_timeless `{distG Σ} (n : node) l q v : Timeless (mapsto_l n l q v).
Proof. rewrite /mapsto_l /Timeless.
       iIntros ">H".
       iDestruct "H" as (γ's) "(H1 & H2)".
       iExists γ's. iFrame.
Qed.

Definition IsNode `{distG Σ} n := ( γs, n n γs)%I.
Global Instance IsNode_persistent `{distG Σ} n : Persistent (IsNode n).
Proof. apply _. Qed.

Global Opaque IsNode.

Fixed socket end-points in the system
Definition Fixed `{distG Σ} A := ownF A.

Lemma ownF_agree `{distG Σ} A B : Fixed A -∗ Fixed B -∗ A = B.
Proof.
  iIntros "HA HB".
  by iDestruct (own_valid_2 with "HA HB") as %?%agree_op_invL'.
Qed.

Points-to locations on nodes

Notation "l ↦[ n ]{ q } v" := (mapsto_l n l q v)
  (at level 20, q at level 50, format "l ↦[ n ]{ q } v") : uPred_scope.
Notation "l ↦[ n ] v" := (l ↦[n]{1} v) (at level 20) : uPred_scope.
Notation "l ↦[ n ]{ q } -" := ( v, l ↦[n]{q} v)%I
  (at level 20, q at level 50, format "l ↦[ n ]{ q } -") : uPred_scope.
Notation "l ↦[ n ] -" := (l ↦[n]{1} -)%I (at level 20) : uPred_scope.

Sockets s↦
Notation "h s↦[ n ]{ q } s" := (mapsto_s n h q s)
  (at level 20, q at level 50, format "h s↦[ n ]{ q } s") : uPred_scope.
Notation "h s↦[ n ] s" := (mapsto_s n h 1 s) (at level 20) : uPred_scope.

Lookup stored socket interpretation γ↦
Notation "a γ↦ γ" :=
  (own (A:=socket_interpR)
       dist_si_name ( {[ a := (to_agree γ) ]})) (at level 20) : uPred_scope.

Messages m↦
Notation "i m↦{ q } m" := (mapsto (L:=message_id) (V:=message) i q m)
  (at level 20, q at level 50, format "i m↦{ q } m") : uPred_scope.
Notation "i m↦ m" :=
  (mapsto (L:=message_id) (V:=message) i 1 m) (at level 20) : uPred_scope.

Sent messages c↦
Notation "i o↦{ q } s" := (mapsto (L:=positive) (V:=message_stable) i q s)
  (at level 20, q at level 50, format "i o↦{ q } s") : uPred_scope.
Notation "i o↦ s" :=
  (mapsto (L:=positive) (V:=message_stable) i 1 s) (at level 20) : uPred_scope.

Received messages r↦
Notation "a r↦{ q } g" := (mapsto (L:=socket_address) (V:=message_soup) a q g)
  (at level 20) : uPred_scope.
Notation "a r↦ g" := (mapsto (L:=socket_address) (V:=message_soup) a 1 g)
  (at level 20) : uPred_scope.

Socket interpretations ρ ⤇

Global Instance saved_pred_proper `{savedPredG Σ A} n γ:
  Proper ((dist n) ==> (dist n)) (@saved_pred_own Σ A _ γ : (A -d> iProp Σ) -d> iProp Σ).
Proof.
  intros Φ Ψ Hps.
  f_equiv.
  destruct n; first done.
    by apply dist_S.
Qed.

Global Instance saved_pred_proper' `{savedPredG Σ A} γ:
  Proper ((≡) ==> (≡)) (@saved_pred_own Σ A _ γ : (A -d> iProp Σ) -d> iProp Σ).
Proof.
  apply ne_proper => n. apply _.
Qed.

Definition si_pred `{distG Σ} a (Φ : socket_interp Σ) : iProp Σ :=
  ( γ, a γ γ saved_pred_own (A:=message_stable) γ Φ)%I.

Global Instance si_pred_prop `{distG Σ} a : Proper ((≡) ==> (≡)) (si_pred a).
Proof.
  intros x y Heq.
  apply exist_proper => z; f_equiv.
  by rewrite Heq.
Qed.

Notation "a ⤇ Φ" := (si_pred a Φ) (at level 20).

Definition FreeIPs_ctx `{distG Σ} (F : gset ip_address) :=
  own dist_freeips_name ( GSet F).

Definition FreeIP `{distG Σ} (ip : ip_address) :=
  own dist_freeips_name ( GSet {[ ip ]}).

Lemma is_FreeIP `{distG Σ} F ip :
  FreeIPs_ctx F -∗ FreeIP ip -∗ ip F.
Proof.
  iIntros "HF Hip". iDestruct (own_valid_2 with "HF Hip") as %[_ Hi].
  iPureIntro.
  move: (Hi 0%nat). rewrite /= left_id.
  move => [? [/to_agree_injN /discrete /leibniz_equiv_iff <- [/gset_disj_included ? _]]].
  by apply elem_of_subseteq_singleton.
Qed.

Lemma Use_FreeIP `{distG Σ} F ip :
  FreeIPs_ctx F -∗ FreeIP ip ==∗ FreeIPs_ctx (F {[ ip ]}).
Proof.
  iIntros "HF Hip".
  iDestruct (is_FreeIP with "HF Hip") as %Hip.
  replace F with ({[ ip ]} (F {[ ip ]})) at 1; last first.
  { rewrite (comm_L _ {[ _ ]}) difference_union_L
    -(comm_L _ {[ _ ]}) subseteq_union_1_L //.
    by apply elem_of_subseteq_singleton. }
  iCombine "HF" "Hip" as "H".
  iMod (own_update with "H") as "H"; last by iFrame "H".
  apply auth_update_dealloc.
  rewrite -gset_disj_union; last by set_solver.
  by apply gset_disj_dealloc_empty_local_update.
Qed.

Lemma FreeIps_alloc `{inG Σ (authUR (gset_disjUR ip_address))}
      (F : gset ip_address) :
  (|==> γ, own γ ( GSet F) [∗ set] ip F, own γ ( GSet {[ ip ]}))%I.
Proof.
  iMod (own_alloc ( GSet )) as (γ) "HM"; first by apply auth_auth_valid.
  iAssert (|==>
              M : gset ip_address,
               (elements M ≡ₚ elements F)
                  own γ ( GSet M) [∗ set] ip M, own γ ( GSet {[ ip ]}))%I
    with "[HM]" as "HF".
  { pose proof (NoDup_elements F) as Hnd.
    iInduction (elements F) as [|a l] "IHl".
    - iModIntro. iExists .
      rewrite big_sepS_empty. iFrame.
      by iPureIntro.
    - inversion Hnd as [|? ? ? Hrd']; subst.
      iMod ("IHl" $! Hrd' with "HM") as (M HMl) "[HM HML]"; iFrame.
      assert (a M) as Hnm.
      { by rewrite -elem_of_elements HMl. }
      iMod (own_update _ _ ( GSet ({[a]} M) GSet {[a]}) with "HM")
        as "[HM Ha]".
      { apply auth_update_alloc.
        apply gset_disj_alloc_empty_local_update; last by set_solver. }
      iModIntro.
      iExists ({[a]} M); iFrame.
      iSplit; first by iPureIntro; rewrite elements_union_singleton // HMl.
      rewrite big_sepS_insert; last done. iFrame. }
  iMod "HF" as (M HMF) "[? ?]".
  replace M with F; first by iModIntro; iExists _; iFrame.
  apply elem_of_equiv_L => x.
  rewrite -!elem_of_elements.
  rewrite -elem_of_list_permutation_proper; eauto.
Qed.

Definition FreePorts_ctx `{distG Σ} (P : gmap ip_address (gset_disjUR port)) :=
  own dist_freeports_name ( P).

Definition FreePorts `{distG Σ} (ip : ip_address) (ports : gset port) :=
  own dist_freeports_name ( ({[ ip := (GSet ports)]})).

Lemma FreePorts_included `{distG Σ} P ip ports :
  FreePorts_ctx P -∗ FreePorts ip ports -∗
     ports', P !! ip = Some (GSet ports') ports ports'.
Proof.
  iIntros "HP Hip"; rewrite /FreePorts_ctx /FreePorts.
  iCombine "HP" "Hip" as "HPip". iDestruct (own_valid with "HPip") as %[_ Hvld].
  iPureIntro.
  move: Hvld; rewrite //= left_id_L => HP. move: (HP 0%nat).
  move=> [a [/to_agree_injN /discrete /leibniz_equiv_iff -> [Hvld Hv]]].
  apply cmra_discrete_included_iff, singleton_included in Hvld.
  destruct Hvld as [ports' [Hports' HS]].
  move: HS; rewrite Some_included_total => HS.
  destruct ports' as [ports'|].
  - apply gset_disj_included in HS.
    apply leibniz_equiv_iff in Hports'. by eexists.
  - by pose proof (lookup_valid_Some _ _ _ Hv Hports').
Qed.

Lemma FreePorts_distribute `{distG Σ} ip ports ports' :
  ports ## ports'
  FreePorts ip (ports ports') ⊣⊢ FreePorts ip ports FreePorts ip ports'.
Proof.
  intros ?.
  by rewrite /FreePorts -gset_disj_union // -own_op -auth_frag_op op_singleton.
Qed.

Lemma FreePorts_alloc `{distG Σ} P ip ports :
  ip (dom (gset _) P)
  FreePorts_ctx P ==∗ FreePorts_ctx (<[ ip := GSet ports ]>P) FreePorts ip ports.
Proof.
  iIntros (?) "HP"; rewrite /FreePorts_ctx /FreePorts.
  iMod (own_update _ _ ( _ {[ ip := (GSet ports)]}) with "HP")
    as "[HP Hip]"; last by iFrame.
  apply auth_update_alloc, alloc_singleton_local_update; last done.
  by eapply not_elem_of_dom.
Qed.

Lemma FreePorts_dealloc `{distG Σ} P ip ports :
  FreePorts_ctx P -∗ FreePorts ip ports ==∗
     ports', P !! ip = Some (GSet ports') ports ports'
      FreePorts_ctx (<[ip := GSet (ports' ports)]> P).
Proof.
  iIntros "HP Hip"; rewrite /FreePorts_ctx /FreePorts.
  iDestruct (own_valid with "HP") as %[_ HP].
  iCombine "HP" "Hip" as "HPip".
  iDestruct (own_valid with "HPip") as %[_ Hvld].
  move: Hvld; rewrite //= left_id_L => HQ. move: (HQ 0%nat).
  move=> [a [/to_agree_injN /discrete /leibniz_equiv_iff -> [Hvld Hv]]].
  apply cmra_discrete_included_iff, singleton_included in Hvld.
  destruct Hvld as [ports' [Hports' HS]].
  revert HS; rewrite Some_included_total => HS.
  destruct ports' as [ports'|];
    last by pose proof (lookup_valid_Some _ _ _ Hv Hports').
  apply gset_disj_included in HS.
  apply leibniz_equiv in Hports'.
  iMod (own_update _ _ ( _ {[ ip := (GSet )]}) with "HPip")
    as "[? ?]"; last by iExists _; iFrame; eauto.
  eapply auth_update, singleton_local_update; eauto.
  assert (ports' = ports' ports ports) as HPeq.
  { by rewrite difference_union_L (comm_L _ _ ports) subseteq_union_1_L. }
  rewrite {1}HPeq.
  rewrite -gset_disj_union; last set_solver.
  rewrite (comm_L (⋅)).
  apply gset_disj_dealloc_empty_local_update.
Qed.

Definition local_state_coherence (γmap : gmap node node_gnames)
           (heaps : gmap node heap) (sockets : gmap node sockets) :=
  dom (gset node) γmap = dom (gset node) heaps
  dom (gset node) γmap = dom (gset node) sockets.

Definition network_coherence (rec : socket_received) (ips : ports_in_use) :=
   a P, is_Some (rec !! a)
         ips !! ip_of_address a = Some P
         (port_of_address a) P.

Definition local_state_i `{distG Σ} (σ : state) (n : node) (x : node_gnames) :=
  ( (h : heap) (S : sockets),
      state_heaps σ !! n = Some h
      state_sockets σ !! n = Some S
      n n x
      gen_heap_ctx (hG := gen_heap_loc_instance (heap_name x) (heap_meta_name x)) h
      gen_heap_ctx (hG := gen_heap_soc_instance (socket_name x) (socket_meta_name x)) S
  )%I.

Definition IPs (p : ports_in_use) := dom (gset ip_address) p.

Definition socket_interp_coherence `{distG Σ} (ips : ports_in_use) :=
  ( si A,
      (* the addresses in A are also in the domain of ips *)
       a, a A ip_of_address a dom (gset ip_address) ips
      (* for all addresses in the system, being in si means either having a
         Fixed interpretation (a ∈ A) or being dynamically bound. *)

       a, (ip_of_address a dom (gset ip_address) ips)
            (a dom (gset socket_address) si
               a A (a A P, ips !! ip_of_address a = Some P port_of_address a P))
       Fixed A
       own (A:=socket_interpR) dist_si_name ( si)
       ([∗ set] s (dom (gset socket_address) si), φ, s φ))%I.

Definition messages_received_coherence (rec : socket_received) (M : message_soup) :=
   i M', rec !! i = Some M'
          messages_received_at i M = M'.

Definition messages_sent_coherence (sent : messages_stable) (M : message_soup) :=
  dom (gset message_id) sent = dom (gset message_id) M
   mId, sent !! mId = message_stable_from_message <$> (M !! mId).

Definition messages_coherence (rec : socket_received) (sent : messages_stable)
           (M : message_soup) (P : ports_in_use) :=
  network_coherence rec P
  messages_received_coherence rec M
  messages_sent_coherence sent M.

Definition messages {Σ} `{distG Σ} (M : message_soup) (P : ports_in_use) :=
  (
     sent rec,
      messages_coherence rec sent M P
      gen_heap_ctx M
      gen_heap_ctx sent
      gen_heap_ctx rec
      ([∗ map] sag rec,
       g = messages_received_at sa M
       sa r↦{½} g)
      ([∗ map] mIdm M,
        π (Φ : message_stable -d> iProp Σ), (m_state m = MS_RECEIVED
                     π = 1%Qp (m_destination m) Φ
                      Φ (message_stable_from_message m))
       mId m↦{π} m))%I.

Global Instance distG_irisG `{distG Σ} :
  irisG dist_lang Σ :=
  {|
    iris_invG := _;
    state_interp σ κ _ :=
      (
         (m : gmap node node_gnames),
          local_state_coherence m (state_heaps σ) (state_sockets σ)⌝
           socket_interp_coherence (state_ports_in_use σ)
           ownS dist_map_name m
           ([∗ map] n γs m, local_state_i σ n γs)
           (∃ Fip Piu, (⌜dom (gset _) Piu ## Fip
                        (∀ ip, ip Fip state_ports_in_use σ !! ip = Some ∅)
                        (∀ ip P, Piu !! ip = Some (GSet P)
                            Q, (state_ports_in_use σ) !! ip = Some Q P ## Q)⌝)
                          FreeIPs_ctx Fip FreePorts_ctx Piu)
            messages (state_ms σ) (state_ports_in_use σ)
      )%I;
    fork_post _ := True%I;
  |}.

Global Opaque iris_invG.

Local Hint Extern 0 (head_reducible _ _) => eexists _, _, _; simpl.

Local Hint Constructors head_step.
Local Hint Resolve alloc_fresh.
Local Hint Resolve to_of_val.

Lemma network_coherence_insert a P rec ips M :
  ips !! ip_of_address a = Some P
  network_coherence rec ips
  network_coherence (<[a:=M]> rec) (<[ip_of_address a:={[port_of_address a]} P]> ips).
Proof.
  rewrite /network_coherence; simpl.
  intros Hlookup H a' P' Hlookup'; simpl in *.
  destruct (decide (a = a')); simplify_eq.
  - rewrite lookup_insert. intros. simplify_eq.
      by apply elem_of_union_l,elem_of_singleton.
  - destruct (decide (ip_of_address a = ip_of_address a')) as [Heq|Hneq].
    + rewrite lookup_insert_ne in Hlookup'; auto.
      rewrite Heq lookup_insert. intro; simplify_eq.
      rewrite elem_of_union; right.
      apply H; eauto. by rewrite -Heq.
    + rewrite lookup_insert_ne; auto.
      rewrite lookup_insert_ne in Hlookup'; try done.
        by apply H.
Qed.

Section GhostStateLemmas.
  Context `{distG Σ}.

Definition message_soup_info M sent :=
  (messages_sent_coherence sent M
    gen_heap_ctx M gen_heap_ctx sent)%I.

Lemma mapsto_node_agree n γs γs' :
  n n γs -∗ n n γs' -∗ γs = γs'.
Proof.
  apply wand_intro_r.
  rewrite /ownS mapsto_node_eq -own_op own_valid discrete_valid.
  f_equiv=> /auth_frag_proj_valid /=. rewrite op_singleton singleton_valid.
  apply (agree_op_invL' (A := node_gnamesO)).
Qed.

Lemma node_in_map n γ's (m : gmap node node_gnames) :
  (n n γ's -∗
     ownS dist_map_name m -∗
     m !! n = Some γ's)%I.
Proof.
  iIntros "H1 H2".
  iCombine "H2" "H1" as "H".
  rewrite /ownS mapsto_node_eq -own_op own_valid.
  iDestruct "H" as %HvalidR. iPureIntro.
  revert HvalidR.
  rewrite auth_both_valid.
  rewrite (singleton_included)=> -[[y [Hlookup Hless]] Hvalid].
  assert (Hvalidy := lookup_valid_Some _ n y Hvalid Hlookup).
  revert Hlookup.
  rewrite lookup_fmap fmap_Some_equiv=> -[v' [Hl Heq]]. revert Hless Heq.
  rewrite Some_included_total.
  destruct (to_agree_uninj y Hvalidy) as [y' <-].
  rewrite to_agree_included.
  intros Heq Heq'%(to_agree_inj y' v').
  apply leibniz_equiv in Heq.
  apply leibniz_equiv in Heq'.
    by simplify_eq.
Qed.

Lemma node_in_state_heaps n γ's (m : gmap node node_gnames) σ :
  m !! n = Some γ's
  (([∗ map] n' γ's m, local_state_i σ n' γ's) -∗
   x, (state_heaps σ) !! n = Some x)%I.
Proof.
  iIntros (Hninm) "Hmap".
  iDestruct (big_sepM_lookup (local_state_i σ) m n _ with "[Hmap]") as "Hl";
    first done; iFrame.
  iDestruct "Hl" as (h s p) "(% & _)". iPureIntro; by exists h.
Qed.

Lemma node_in_state_sockets n γ's (m : gmap node node_gnames) σ :
  m !! n = Some γ's
  (([∗ map] n' γ's m, local_state_i σ n' γ's) -∗
    x, (state_sockets σ) !! n = Some x)%I.
Proof.
  iIntros (Hninm) "Hmap".
  iDestruct (big_sepM_lookup (local_state_i σ) m n _ with "[Hmap]") as "Hl";
    first done; iFrame.
  iDestruct "Hl" as (h s) "(_ & % & _)". iPureIntro; by exists s.
Qed.

Lemma node_local_state n γ's m (σ : state) :
  m !! n = Some γ's
  (([∗ map] n' x m, local_state_i σ n' x) -∗
  local_state_i σ n γ's
  [∗ map] n' x (delete n m), local_state_i σ n' x)%I.
Proof.
  iIntros (Hinm) "Hmap".
  by rewrite -(big_sepM_delete (local_state_i σ) m n).
Qed.

Lemma map_local_state_i_update_heaps n m h (σ : state) :
  (([∗ map] n' x (delete n m), local_state_i σ n' x) -∗
  [∗ map] n' x (delete n m), local_state_i {|
                                     state_heaps := <[n:=h]>(state_heaps σ);
                                     state_sockets := state_sockets σ;
                                     state_ports_in_use := state_ports_in_use σ;
                                     state_ms := state_ms σ;
                                   |} n' x)%I.
Proof.
  erewrite (big_sepM_mono (local_state_i σ) (local_state_i _)).
  - iIntros "Hmap"; iFrame.
  - intros k x Hdel.
    destruct (decide (k = n)); rewrite /local_state_i; subst.
    + by rewrite lookup_delete in Hdel.
    + by rewrite lookup_insert_ne; last done.
Qed.

Lemma map_local_state_i_update_sockets n m S (σ : state) :
  (([∗ map] n' x (delete n m), local_state_i σ n' x) -∗
  [∗ map] n' x (delete n m), local_state_i {|
                                     state_heaps := state_heaps σ;
                                     state_sockets := <[n:=S]> (state_sockets σ);
                                     state_ports_in_use := state_ports_in_use σ;
                                     state_ms := state_ms σ;
                                    |} n' x)%I.
Proof.
  erewrite (big_sepM_mono (local_state_i σ) (local_state_i _)).
  - iIntros "Hmap"; iFrame.
  - intros k x Hdel.
    destruct (decide (k = n)); rewrite /local_state_i; subst.
    + by rewrite lookup_delete in Hdel.
    + by rewrite lookup_insert_ne; last done.
Qed.

Lemma map_local_state_i_update m σ P M :
  (([∗ map] n x m, local_state_i σ n x) -∗
  [∗ map] n x m, local_state_i {|
                          state_heaps := state_heaps σ;
                          state_sockets := state_sockets σ;
                          state_ports_in_use := P;
                          state_ms := M;
                        |} n x)%I.
Proof.
  rewrite /local_state_i; simpl. iIntros "H". iFrame.
Qed.

Lemma node_local_state_rev n γ's m (σ : state) :
  m !! n = Some γ's
  (local_state_i σ n γ's -∗
  ([∗ map] n' x (delete n m), local_state_i σ n' x) -∗
  [∗ map] n' x m, local_state_i σ n' x )%I.
Proof.
  iIntros (Hinm) "Hl Hmap".
  iDestruct (big_sepM_insert (local_state_i σ)
                             (delete n m) n γ's with "[Hl Hmap]") as "HP".
  - apply lookup_delete.
  - iFrame.
  - apply insert_id in Hinm.
      by rewrite insert_delete Hinm.
Qed.

Lemma si_pred_agree a Φ Ψ x :
  a Φ -∗ a Ψ -∗ (Φ x Ψ x).
Proof.
  iIntros "#H1 #H2".
  iDestruct "H1" as (γ) "[H1 H1']".
  iDestruct "H2" as (γ') "[H2 H2']".
  iDestruct (own_valid_2 with "H1 H2") as %Hvalid.
  rewrite -auth_frag_op op_singleton in Hvalid.
  apply singleton_valid in Hvalid.
  apply (agree_op_invL' γ γ') in Hvalid.
  rewrite Hvalid.
  iDestruct (saved_pred_agree _ _ _ x with "H1' H2'") as "H".
  iExact "H".
Qed.

Lemma si_pred_equiv a Φ Ψ :
  a Φ -∗ a Ψ -∗ (Φ Ψ).
Proof.
  iIntros "#H1 #H2".
  iDestruct "H1" as (γ) "[H1 H1']".
  iDestruct "H2" as (γ') "[H2 H2']".
  iDestruct (own_valid_2 with "H1 H2") as %Hvalid.
  rewrite -auth_frag_op op_singleton in Hvalid.
  apply singleton_valid in Hvalid.
  apply (agree_op_invL' γ γ') in Hvalid.
  rewrite Hvalid discrete_fun_equivI. iIntros (?).
  by iDestruct (saved_pred_agree with "H1' H2'") as "H".
Qed.

Lemma find_fixed_socket_interp a A ips :
  a A
  Fixed A -∗ socket_interp_coherence ips -∗ Φ, a Φ.
Proof.
  iIntros (HaA) "#HA Hsoccoh".
  iDestruct "Hsoccoh" as (si fx HA Hdom) "(#Hsi & Hfix & #Hdyn)".
  iDestruct (ownF_agree with "HA Hsi") as %?; subst.
  iDestruct (big_sepS_elem_of (A:=socket_address) with "Hdyn") as "$".
  by rewrite (Hdom _ (HA _ HaA)); left.
Qed.

Lemma message_sepM_later M :
  (([∗ map] mIdm M,
     π (Φ : message_stable -d> iProp Σ), (m_state m = MS_RECEIVED
        π = 1%Qp (m_destination m) Φ Φ (message_stable_from_message m))
        mId m↦{π} m) -∗
  ([∗ map] mIdm M,
     π (Φ : message_stable -d> iProp Σ), (m_state m = MS_RECEIVED
        π = 1%Qp (m_destination m) Φ Φ (message_stable_from_message m))
        mId m↦{π} m))%I.
Proof.
  iIntros "H".
  rewrite big_sepM_mono.
  - iFrame.
  - iIntros (k x H') "H".
    iDestruct "H" as (π φ) "([% | (H1 & H1' & H1'')] & H2)"; iExists π, φ; iFrame.
    + by iLeft.
    + iRight. iFrame.
Qed.

Lemma message_state_later M P :
  (( sent rec,
      messages_coherence rec sent M P
      gen_heap_ctx M
      gen_heap_ctx sent
      gen_heap_ctx rec
      ([∗ map] sag rec,
       g = messages_received_at sa M
       sa r↦{½} g)
      ([∗ map] mIdm M,
        π (Φ : message_stable -d> iProp Σ), (m_state m = MS_RECEIVED
                     (π = 1%Qp (m_destination m) Φ
                     Φ (message_stable_from_message m)))
       mId m↦{π} m)) -∗ messages M P)%I.
Proof.
  iIntros "H". iDestruct "H" as (sent rec) "(H1 & H2 & H3 & H4 & H5 & H6)".
  iDestruct (message_sepM_later with "H6") as "H6".
  iExists sent, rec. iFrame.
Qed.

Lemma messages_received_id M sa mId m :
  M !! mId = None
  m_state m = MS_SENT
  messages_received_at sa (<[mId:=m]> M) =
  messages_received_at sa M.
Proof.
  intros.
  rewrite /messages_received_at map_filter_lookup_None_insert; auto.
  rewrite map_filter_lookup_None. by left.
  intros []; simpl in *. by rewrite H1 in H3.
Qed.

Lemma messages_insert M P mId msg :
  M !! mId = None
  m_state msg = MS_SENT
  (messages M P ==∗
  mId m msg
  mId o message_stable_from_message msg
   sent rec,
    messages_coherence
     rec
     (<[mId:=(message_stable_from_message msg)]>sent)
     (<[mId:=msg]>M) P
    gen_heap_ctx (<[mId:=msg]>M)
    gen_heap_ctx (<[mId:=message_stable_from_message msg]>sent)
    gen_heap_ctx rec
    ([∗ map] sag rec, g = messages_received_at sa M sa r↦{½} g)
    ([∗ map] mIdm M,
      π (Φ : message_stable -d> iProp Σ),
       (m_state m = MS_RECEIVED
       π = 1%Qp (m_destination m) Φ
        Φ (message_stable_from_message m))
       mId m↦{π} m))%I.
Proof.
  iIntros (Hlookup Hsent) "Hmessages".
  iDestruct "Hmessages" as (sent rec [Hnetcoh [Hreccoh [Hdom Hsentcoh]]])
                             "(Hm & Hsent & Hrec & Hrecs & Hmsgs)".
  iMod (gen_heap_alloc M mId _ with "Hm") as "(Hms & HmId & _)"; first done.
  iMod (gen_heap_alloc sent mId (message_stable_from_message msg)
          with "Hsent") as "(Hsent & HsentId & _)".
  { by rewrite -(not_elem_of_dom (D:=gset message_id)) Hdom not_elem_of_dom. }
  iModIntro; iFrame. iExists _, _; iFrame. iPureIntro.
  rewrite /messages_coherence /messages_received_coherence /messages_sent_coherence /=.
  repeat split.
  - done.
  - intros. rewrite messages_received_id //=. by apply Hreccoh.
  - by rewrite !dom_insert_L Hdom.
  - intro. destruct (decide (mId = mId0)); subst.
    + rewrite !lookup_insert //=.
    + rewrite !lookup_insert_ne //=.
Qed.

Lemma message_split π1 π2 π m v :
  (π = π1 + π2)%Qp
  m m↦{π1} v m m↦{π2} v ⊣⊢ m m↦{π} v.
Proof.
  intros ->.
  iSplit; iIntros "H";
  iDestruct "H" as "[H1 H2]"; iFrame.
Qed.

Lemma messages_update_received M mId m m' :
  M !! mId = Some m
  m_state m = MS_SENT
  m' = {| 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
       |}
  (gen_heap_ctx M
  ([∗ map] kx M,
    ( (π : Qp) (φ : message_stable -d> iProp Σ),
        (m_state x = MS_RECEIVED
         π = 1%Qp m_destination x φ
          φ (message_stable_from_message x)) k m↦{π} x)) ==∗
   φ, (m_destination m φ φ (message_stable_from_message m)) mId m↦{¾} m'
  gen_heap_ctx (<[mId:=m']>M) ([∗ map] kx (<[mId:=m']>M),
       ( (π : Qp) (φ : message_stable -d> iProp Σ),
        (m_state x = MS_RECEIVED
         π = 1%Qp m_destination x φ
          φ (message_stable_from_message x)) k m↦{π} x)))%I.
Proof.
  iIntros (Hlookup Hsent Hm') "(Hms & Hmsg)".
  iDestruct (big_sepM_delete _ M with "Hmsg") as "[Hm Hmsg]". done.
  iDestruct "Hm" as (π φ) "([% | (% & Hsi & HP)] & Hm)".
  by rewrite H0 in Hsent.
  subst.
  iMod (gen_heap_update M 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 Hm")
    as "[Hms Hm]".
  iExists φ. iFrame.
  rewrite -(message_split ¼ ¾ 1%Qp mId);
    last by rewrite Qp_quarter_three_quarter.
  iDestruct "Hm" as "[Hm Hm']". iSplitL "Hm'"; first by iFrame.
  iModIntro.
  iDestruct (big_sepM_insert _ (delete mId M) mId _ with "[Hmsg Hm]") as "H";
  first by rewrite lookup_delete.
  - iFrame. iExists ¼,φ; iFrame; eauto.
  - by rewrite insert_delete; iFrame.
Qed.

Lemma messages_received_at_update M d msg m rec mId :
  M !! mId = Some m
  m_state m MS_RECEIVED
  msg = {| m_from_node := m_from_node m;
           m_sender := m_sender m;
           m_destination := d;
           m_protocol := m_protocol m;
           m_body := m_body m;
           m_state := MS_RECEIVED;
        |}
  (([∗ map] ky (delete d rec),
   y = messages_received_at k M k r↦{ ½} y) -∗
  d r↦{ ½} <[mId:=msg]> (messages_received_at d M) -∗
  ([∗ map] ky (<[d:=<[mId:=msg]>(messages_received_at d M)]> rec),
   y = messages_received_at k (<[mId:=msg]>M) k r↦{ ½} y))%I.
Proof.
  iIntros (Hlookup Hstate Hmsg) "Hrecs Hrec".
  iDestruct (big_sepM_mono
               (λ k y, (y = messages_received_at k M k r↦{ ½} y))%I
               (λ k y, (y = messages_received_at k (<[mId:=msg]>M)
                                                   k r↦{ ½} y))%I
               (delete d rec)
               with "[Hrecs]") as "Hrecs"; auto.
  iIntros (k M' Hlookup') "(% & H)".
  - iFrame.
  destruct (decide (d = k)).
  { rewrite e in Hlookup'. rewrite lookup_delete in Hlookup'.
    inversion Hlookup'. }
  iPureIntro.
  rewrite H0 /messages_received_at map_filter_lookup_None_insert //=.
  { rewrite map_filter_lookup_None. right. intros; simplify_eq.
    rewrite Hlookup in H1; simplify_eq.
    intros [H2 H3]. simpl in *. by destruct Hstate. }
  { intros [H3 H4]. rewrite Hmsg in H3. simpl in *. by destruct n. }
  - iDestruct (big_sepM_insert _ (delete d rec) d _ with "[Hrecs Hrec]") as "H".
    by rewrite lookup_delete. iFrame. iFrame.
    iPureIntro. rewrite /messages_received_at map_filter_insert Hmsg //=.
    by rewrite insert_delete.
Qed.

End GhostStateLemmas.