aneris.examples.ping_pong_done.proof

From aneris Require Import tactics proofmode notation adequacy network.
From aneris.examples.ping_pong_done Require Export code.
From iris.algebra Require Import cmra agree frac .
From stdpp Require Import base.

Set Default Proof Using "Type".

Import Network.
Import String.
Import uPred.

(* Last message received *)
Inductive msg_received := NONE | PING | PONG | DONE.

Definition ppCmra : cmraT := (prodR fracR (agreeR (leibnizO msg_received))).

Class ppdG Σ := PpdG {
  pongG_inG :> inG Σ ppCmra;
  pingG_inG :> inG Σ ppCmra
}.

Definition pingpongΣ : gFunctors := #[GFunctor ppCmra].

Instance subG_ppdΣ {Σ} : subG pingpongΣ Σ ppdG Σ.
Proof. solve_inG. Qed.

Class pingpongG Σ := MppdG {
  pong_inG :> inG Σ ppCmra;
  ping_inG :> inG Σ ppCmra;
  pong_name : gname;
  ping_name : gname
}.

Section ghost_state.
  Context `{!pingpongG Σ}.

  (* Corresponds to a view of what the last message pong received is *)
  Definition ownPong (q : Qp) (s : msg_received) :=
    own (A := ppCmra) pong_name (q, to_agree s).

  (* Corresponds to a view of what the last message ping received is *)
  Definition ownPing (q : Qp) (s : msg_received) :=
    own (A := ppCmra) ping_name (q, to_agree s).

  Lemma ownPong_eq p q s1 s2:
    ownPong p s1 -∗ ownPong q s2 -∗ s1 = s2.
  Proof.
    iIntros "H1 H2". rewrite /ownPong.
    iCombine "H1" "H2" as "H".
    iDestruct (own_valid with "H") as %HValid.
    destruct HValid as [_ H2]; simpl in H2;
    iIntros "!%". by apply (agree_op_invL' (A:=leibnizO msg_received)).
  Qed.

  Lemma ownPong_update s1 s2:
    ownPong (1/2) s1 -∗ ownPong (1/2) s1 ==∗ ownPong (1/2) s2 ownPong (1/2) s2.
  Proof.
    iIntros "H1 H2". rewrite /ownPong. iCombine "H1 H2" as "H".
    rewrite -own_op -pair_op frac_op' Qp_half_half agree_idemp.
    iApply (own_update with "H"). by apply cmra_update_exclusive.
  Qed.

  Lemma ownPing_update s1 s2:
    ownPing (1/2) s1 -∗ ownPing (1/2) s1 ==∗ ownPing (1/2) s2 ownPing (1/2) s2.
  Proof.
    iIntros "H1 H2". rewrite /ownPing. iCombine "H1 H2" as "H".
    rewrite -own_op -pair_op frac_op' Qp_half_half agree_idemp.
    iApply (own_update with "H"). by apply cmra_update_exclusive.
  Qed.

End ghost_state.

Section pong.
  Context `{ppd : !pingpongG Σ}
          `{dG : !distG Σ}
          `{node : Network.node}.

  (* The protocol and specification is to be read as follows:

       1. ping and pong each start with two 1/2 fraction resources corresponding
          to a view of the last received message by the node, that is, initially
          neither ping nor pong have received any. The 1/2 fractions ensure that
          neither of them can update the ghost state without receiving the other
          half.

       2. To send a "PING" message, ping has to send her pong fraction.

       3. pong now has boths fractions which allows her to update the ghost
          state corresponding to the last message she has seen.

       4. To send a "PONG" response, pong has to send an updated 1/2 fraction
          resource acknowledging that the "PING" message was received. pong also
          has to send her ping fraction.

       5. To send a "DONE" response, ping has to send an updated 1/2 fraction
          resource acknowledging that the "PONG" message was received. ping also
          sends back the pong resource such that pong can update her state to
          DONE even though this is not strictly necessary.

     Note that this protocol & specification does not allow you to conclude that
     the sequence of messages were actually exchanged; a pong implementation
     that simply receives a "PING" message and then terminate will also fulfil
     the specification. This is not too satisfying...

   *)

  Definition pong_si : socket_interp Σ :=
    (λ msg, ϕ, ms_sender msg ϕ
     ((ms_body msg = "PING" ownPong (1/2) NONE
          ( m, (ms_body m = "PONG" ownPong (1/2) PING ownPing (1/2) NONE) -∗ ϕ m))
     (ms_body msg = "DONE" ownPong (1/2) PING ownPing (1/2) PONG ))
    )%I.

  Arguments pong_si : simpl never.

  Lemma pong_spec (A : gset socket_address) (addr : socket_address) :
    addr A ->
    {{{ addr pong_si
         Fixed A
         FreePorts (ip_of_address addr) {[port_of_address addr]}
         IsNode node
         ownPong (1/2) NONE
         ownPing (1/2) NONE }}}
      node; pong #addr
    {{{ v, RET node;v; ownPong 1 DONE }}}.
  Proof.
    iIntros (Haddr ) "(#Hsi & #Hfixed & Hip & Hn & Hpong & Hping) Hcont".
    wp_lam.
    wp_socket h as "Hsocket". wp_let.
    wp_apply (wp_socketbind_static with "[$]"); try done.
    iDestruct 1 as (g) "(Hsocket & Hbind & _) /=".
    wp_pures. wp_bind (listen_wait _)%E.
    iApply (listen_wait_spec with "[$]"); try done.
    iNext. iIntros (m1 m1Id) "(Hsocket & Hsoup & Hm1 & _ & Hsi')".
    iDestruct "Hsi'" as (ψ) "[Hsi' [[Heq [Hpong' Hpsi]] | (_ & Hpong' & Hping' & _)]]"; last first.
    - (* Received a DONE, which contradicts the protocol at this point *)
      iDestruct (ownPong_eq with "Hpong Hpong'") as %Hcontr. inversion Hcontr.
    - (* Received a PING *)
      simpl. iDestruct "Heq" as %->. wp_pures. iSimplifyEq.
      iDestruct (ownPong_update _ PING with "Hpong' Hpong") as ">[Hpong' Hpong]".
      wp_apply (wp_send_to_bound with "[Hm1 Hsi $Hsocket $Hsi' Hpsi Hpong' Hping]"); first done.
      + (* We prove that the message is well-formed and transfer ownership of PONG's state *)
        iApply "Hpsi". by iFrame.
      + (* Sucessfully sent PONG *)
        iIntros "[Hsocket _]". wp_pures. wp_bind (listen_wait _)%E.
        iApply (listen_wait_spec with "[$]"); try done.
        iNext. iIntros (m2 m2Id) "(Hsocket & Hsoup & Hm2 & _ & Hsi')".
        iDestruct "Hsi'" as (ψ') "[Hsi' [[% [Hpong' Hpsi']] | (Heq & Hpong' & Hping')]] /=".
        * (* Received PING, which contradicts the protocol at this point *)
          iDestruct (ownPong_eq with "Hpong Hpong'") as %Hcontr. inversion Hcontr.
        * (* DONE *)
          iDestruct (ownPong_update _ DONE with "Hpong' Hpong") as ">Hpong".
          rewrite -own_op -pair_op frac_op' Qp_half_half agree_idemp.
          simpl. iDestruct "Heq" as %->. wp_pures. iSimplifyEq.
          iApply ("Hcont" with "Hpong").
  Qed.

End pong.

Section ping.
  Context `{ppd : !pingpongG Σ}
          `{dG : distG Σ}
          `{node : Network.node}.

  Definition ping_si : socket_interp Σ :=
    (λ msg, ms_body msg = "PONG" ownPong (1/2) PING ownPing (1/2) NONE)%I.

  Lemma ping_spec (A : gset socket_address) (addr pong : socket_address):
    pong A ->
    addr A ->
    {{{ pong pong_si
         Fixed A
         FreePorts (ip_of_address addr) {[port_of_address addr]}
         IsNode node
         ownPong (1/2) NONE
         ownPing (1/2) NONE }}}
      node; ping #addr #pong
    {{{ v, RET node; v; True }}}.
  Proof.
    iIntros (Hserver Haddr Φ) "(#Hsi & #Hfixed & Hip & Hn & Hpong & Hping) Hcont".
    wp_lam. wp_let.
    wp_socket h as "Hsocket". wp_pures.
    wp_apply (wp_socketbind_dynamic _ _ _ _ _ _ _ (ping_si) with "[$]"); try done.
    iIntros "H //=". iDestruct "H" as (g) "(Hsocket & Hbind & #Haddr)". wp_seq.
    (* Sending "PING" *)
    wp_apply (wp_send_to_bound with "[$Hsi $Hsocket Hpong]"); first done.
    - (* The message is well-formed and adheres to the socket interpretation *)
      rewrite /pong_si. iExists ping_si. simpl. iFrame "#". iLeft.
      iSplit; first done. iFrame. iNext. iIntros (m) "[Hbody Hpong]". by iFrame.
    - (* Sucesfully sent PING *)
      iIntros "[Hsocket _]". wp_pures. wp_bind (listen_wait _)%E.
      iApply (listen_wait_spec with "[$]"); try done.
      iNext. iIntros (m mId) "(Hsocket & Hbind & Hm & _ & H)".
      (* Receiveid PONG *)
      iDestruct "H" as "(Heq & Hpong' & Hping')". simpl. iDestruct "Heq" as %->.
      iDestruct (ownPing_update _ PONG with "Hping' Hping") as ">[Hping' Hping]".
      wp_pures.
      wp_apply (wp_send_to_bound with "[$Hsi $Hsocket Hpong' Hping']"); first done.
      + rewrite /pong_si /=. iExists ping_si. iFrame "#".
        iRight. iSplit; first done; iFrame.
      + iIntros "?". by iApply "Hcont".
  Qed.
End ping.

Section ping_pong_runner.
  Context `{dG : distG Σ}
          `{!pingpongG Σ}.

  Definition ips : gset string := {[ "0.0.0.0" ; "0.0.0.1"]}.
  Definition pong_addr := SocketAddressInet "0.0.0.0" 80.
  Definition ping_addr := SocketAddressInet "0.0.0.1" 80.

  Lemma ping_pong_runner_spec (A : gset socket_address) :
    pong_addr A ->
    ping_addr A ->
    {{{ pong_addr pong_si
         Fixed A
         ([∗ set] ip ips, FreeIP ip)
         ownPong 1 NONE
         ownPing 1 NONE
    }}}
      ping_pong_runner
    {{{ v, RET v; True }}}.
  Proof.
    iIntros (server client Φ) "(#Hsi & #Hfix & Hips & HownPo & HownPi) Hcont".
    rewrite /ping_pong_runner.
    iDestruct (big_sepS_delete _ _ "0.0.0.0" with "Hips") as "(Hpong & Hips)";
      first set_solver.
    iDestruct (big_sepS_delete _ _ "0.0.0.1" with "Hips") as "(Hping & Hips)";
      first set_solver.
    wp_pures.
    rewrite /ownPong /ownPing -Qp_half_half -frac_op'
            -(agree_idemp (@to_agree (leibnizO _) _))
            (pair_op _ _ (@to_agree (leibnizO _ ) _) _ ) !own_op.
    iDestruct "HownPo" as "[Hpong1 Hpong2]".
    iDestruct "HownPi" as "[Hping1 Hping2]".
    (* start the nodes *)
    wp_apply (wp_start with "[-]"); first done. iFrame. simpl.
    iSplitR "Hpong1 Hping1"; last first.
    - iNext. iIntros "Hγs Hn".
      iApply (pong_spec _ _ server with "[$]").
      iNext. by iIntros (?) "_".
    - iNext. wp_seq.
      wp_apply (wp_start with "[-]"); first done. iFrame. iSplitL "Hcont".
      + by iApply "Hcont".
      + iNext. iIntros "Hγs Hn".
        by iApply (ping_spec _ _ _ server client (λ _, True%I) with "[$]").
  Qed.

End ping_pong_runner.

Lemma make_pingpongG `{inG Σ (prodR fracR (agreeR (leibnizO msg_received)))} :
  (|==> _ : pingpongG Σ, ownPong 1 NONE ownPing 1 NONE)%I.
Proof.
  iMod (own_alloc ((1%Qp, (@to_agree (leibnizO _) NONE)))) as (γ) "Hγ"; first done.
  iMod (own_alloc ((1%Qp, (@to_agree (leibnizO _) NONE)))) as (γ') "Hγ'"; first done.
    by iExists {|pong_name := γ; ping_name := γ' |}; iFrame.
Qed.

Definition ping_pong_is :=
  {|
    state_heaps := ∅;
    state_sockets := ∅;
    state_ports_in_use :=
      <["0.0.0.0" := ]> $ <["0.0.0.1" := ]> $ ∅;
    state_ms := ∅;
  |}.

Definition socket_interp `{distG Σ, pingpongG Σ} sa : socket_interp Σ :=
  (match sa with
  | SocketAddressInet "0.0.0.0" 80 => pong_si
  | _ => λ msg, True
   end)%I.

Definition fixed_dom : gset socket_address := {[ pong_addr ]}.

Theorem ping_pong_safe : adequate NotStuck ping_pong_runner ping_pong_is (λ _ _, True).
Proof.
  set (Σ := #[distΣ; pingpongΣ]).
  eapply (@dist_adequacy Σ _ ips fixed_dom); try done; last first.
  { set_solver. }
  { intros i. rewrite /ips !elem_of_union !elem_of_singleton.
    intros [|]; subst; simpl; set_solver. }
  { rewrite /ips /= !dom_insert_L dom_empty_L right_id_L //. }
  iIntros (dinvG).
  iMod (@make_pingpongG Σ) as (?) "[Hpong Hping]".
  iExists (@socket_interp Σ dinvG H).
  iModIntro. iIntros "Hsi #Hsc Hips".
  iApply (ping_pong_runner_spec with "[-] []"); eauto;
    rewrite /fixed_dom; try iFrame; try set_solver.
  rewrite big_sepS_singleton /=. iFrame "#".
Qed.