aneris.examples.ping_pong.proof
From aneris Require Import tactics proofmode notation adequacy network.
From aneris.examples.library Require Import frac_auth.
From aneris.examples.ping_pong Require Export code.
From stdpp Require Import base.
Set Default Proof Using "Type".
Import Network.
Import String.
Import uPred.
Section pong.
Context `{dG : distG Σ}
`{node : Network.node}.
(* The socket interpretation is abstracting over any message received through
the socket and as such a predicate on all received messages. *)
Definition pong_si : socket_interp Σ :=
(λ msg,
(* client is governed by a protocol and the message body should be PING *)
∃ ϕ, ms_sender msg ⤇ ϕ ∗ ⌜ms_body msg = "PING"⌝
(* client should be fine with getting a PONG response *)
∗ ▷ (∀ m, ⌜ms_body m = "PONG"⌝ -∗ ϕ m))%I.
Arguments pong_si : simpl never.
Lemma pong_spec (A : gset socket_address) (addr : socket_address) :
(* addr is static *)
addr ∈ A ->
{{{ (* the address/socket is governed by the pong_si socket predicate *)
addr ⤇ pong_si
(* A should contain static addresses & the port should be free *)
∗ Fixed A
∗ FreePorts (ip_of_address addr) {[port_of_address addr]}
(* the node is valid *)
∗ IsNode node }}}
⟨node; pong #addr⟩
(* pong does not terminate, hence False *)
{{{ v, RET 〈node;v〉; False }}}.
Proof.
iIntros (Haddr HΦ) "(#Hsi & #Hfixed & Hip & Hn) Hcont".
wp_lam.
(* allocate socket, use that the node is valid. *)
wp_socket h as "Hsocket". wp_let. iDestruct "Hn" as "#Hn".
(* static binding requires that the adress is static (fixed A), the port is
free (FreePorts ...) and the socket is allocated (h s↦node ...) *)
wp_apply (wp_socketbind_static_2 with "[$]"); try done.
iDestruct 1 as (M) "[Hip Hbind]".
(* the port is no longer available, the message soup g contains the set of
messages that have been received at the socket (addr r↦{ 1 / 2} M). *)
wp_pures. iLöb as "IH" forall (M HΦ).
(* the listen spec consumes the unbound socket and the message soup. *)
iApply (listen_spec True%I with "[] [$] [$]"); first done.
iIntros (m rm') "!# % (_ & Hsocket & Hrm & Hm & _ & Hpsi) HΦ /=".
iDestruct "Hpsi" as (si) "(#Hsp & % & Hwand)". iSimplifyEq.
(* due to the socket interpration we know the conditional will always be true *)
rewrite H. wp_pures.
(* Send the message. *)
wp_apply (wp_send_to_bound with "[Hwand $Hsocket $Hsp]"); first done.
- (* We prove that the message adheres to the socket interpretation *)
by iApply "Hwand".
- (* The recursive call goes well c.f. the IH *)
iIntros "[Hsocket _] /=". wp_pures.
by iApply ("IH" with "HΦ Hsocket Hrm").
Qed.
End pong.
Section ping.
Context `{dG : distG Σ}
`{node : Network.node}.
Definition ping_si : socket_interp Σ :=
(λ msg, ⌜ms_body msg = "PONG"⌝)%I.
Lemma ping_spec (A : gset socket_address) (addr pong : socket_address):
(* the pong address is static *)
pong ∈ A ->
(* the ping is not*)
addr ∉ A ->
{{{ (* the pong server satisfies its socket interpretation *)
pong ⤇ pong_si
(* A should contain static addresses & the port should be free *)
∗ Fixed A
∗ FreePorts (ip_of_address addr) {[port_of_address addr]}
(* the node is valid *)
∗ IsNode node }}}
⟨node; ping #addr #pong⟩
{{{ v, RET ⟨node; v⟩; ⌜v = #"PONG"⌝ }}}.
Proof.
iIntros (Hserver Haddr Φ) "(#Hsi & #Hfixed & Hip & Hn) Hcont".
wp_lam. wp_let.
(* Allocate socket, use that the node is valid. *)
wp_socket h as "Hsocket". wp_let.
(* Dynamic binding requires that the port is free, the socket is allocated
and a socket interpretation *)
wp_apply (wp_socketbind_dynamic _ _ _ _ _ _ _ (ping_si) with "[$]"); try done.
iIntros "H //=". iDestruct "H" as (g) "(Hsocket & Hbind & #Hpingsi)".
(* Send the ping message *)
wp_apply (wp_send_to_bound with "[$Hsi $Hsocket]"); first done.
- (* The message adheres to the socket interpretation *)
rewrite /pong_si /=. iExists ping_si. simpl. by iSplit.
- (* The rest of the progrma goes well c.f. listen_wait_spec *)
iIntros "[Hsocket _]". wp_seq. wp_bind (listen_wait _)%E.
iApply (listen_wait_spec with "[$]"); try done.
iNext. iIntros (m mId) "(Hsocket & Hbind & Hm & _ & H)".
wp_proj. rewrite /ping_si /=.
iDestruct "H" as "->". by iApply "Hcont".
Qed.
End ping.
Section ping_pong_runner.
Context `{dG : distG Σ}.
(* the set of ips used *)
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) :
(* the pong address is static *)
pong_addr ∈ A ->
(* the ping adress is not *)
ping_addr ∉ A ->
{{{ (* the pong server satisfies its socket interpretation *)
pong_addr ⤇ pong_si
(* A contain static addresses, and the ips we use are free *)
∗ Fixed A
∗ [∗ set] ip ∈ ips, FreeIP ip }}}
ping_pong_runner
{{{ v, RET v; True }}}.
Proof.
iIntros (server client Φ) "(#Hsi & #Hfix & Hips) Hcont". rewrite /ping_pong_runner.
(* Split up the free ips in separate resources *)
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.
(* start the nodes *)
wp_apply (wp_start with "[-]"); first done. iFrame. simpl. iSplitL; last first.
- iNext. iIntros "Hγs Hn".
by iApply (pong_spec _ _ server with "[$]").
- iNext. wp_seq.
wp_apply (wp_start with "[-]"); first done. iFrame. iSplitL.
+ by iApply "Hcont".
+ iNext. iIntros "Hγs Hn".
(* we provide the predicate as Coq is apparently not able to resolve it
through unification ... *)
by iApply (ping_spec _ _ _ server client (λ _, True%I) with "[$]").
Qed.
End ping_pong_runner.
Definition ping_pong_is :=
{|
state_heaps := ∅;
state_sockets := ∅;
state_ports_in_use :=
<["0.0.0.0" := ∅ ]> $ <["0.0.0.1" := ∅ ]> $ ∅;
state_ms := ∅;
|}.
(* 'Global' socket interepration, i.e. all statically bound socket interpretions *)
Definition socket_interp `{distG} sa : socket_interp Σ :=
(match sa with
| SocketAddressInet "0.0.0.0" 80 => pong_si (* pong *)
| _ => λ msg, ⌜True⌝
end)%I.
(* The static/fixed domain contains only the adress of the pong server *)
Definition fixed_dom : gset socket_address := {[ pong_addr ]}.
Theorem ping_pong_safe : adequate NotStuck ping_pong_runner ping_pong_is (λ _ _, True).
Proof.
set (Σ := #[distΣ]).
eapply (@dist_adequacy Σ _ ips fixed_dom); try done; last first.
{ set_solver. }
{ intros ?. 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). iModIntro. iExists socket_interp.
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.
From aneris.examples.library Require Import frac_auth.
From aneris.examples.ping_pong Require Export code.
From stdpp Require Import base.
Set Default Proof Using "Type".
Import Network.
Import String.
Import uPred.
Section pong.
Context `{dG : distG Σ}
`{node : Network.node}.
(* The socket interpretation is abstracting over any message received through
the socket and as such a predicate on all received messages. *)
Definition pong_si : socket_interp Σ :=
(λ msg,
(* client is governed by a protocol and the message body should be PING *)
∃ ϕ, ms_sender msg ⤇ ϕ ∗ ⌜ms_body msg = "PING"⌝
(* client should be fine with getting a PONG response *)
∗ ▷ (∀ m, ⌜ms_body m = "PONG"⌝ -∗ ϕ m))%I.
Arguments pong_si : simpl never.
Lemma pong_spec (A : gset socket_address) (addr : socket_address) :
(* addr is static *)
addr ∈ A ->
{{{ (* the address/socket is governed by the pong_si socket predicate *)
addr ⤇ pong_si
(* A should contain static addresses & the port should be free *)
∗ Fixed A
∗ FreePorts (ip_of_address addr) {[port_of_address addr]}
(* the node is valid *)
∗ IsNode node }}}
⟨node; pong #addr⟩
(* pong does not terminate, hence False *)
{{{ v, RET 〈node;v〉; False }}}.
Proof.
iIntros (Haddr HΦ) "(#Hsi & #Hfixed & Hip & Hn) Hcont".
wp_lam.
(* allocate socket, use that the node is valid. *)
wp_socket h as "Hsocket". wp_let. iDestruct "Hn" as "#Hn".
(* static binding requires that the adress is static (fixed A), the port is
free (FreePorts ...) and the socket is allocated (h s↦node ...) *)
wp_apply (wp_socketbind_static_2 with "[$]"); try done.
iDestruct 1 as (M) "[Hip Hbind]".
(* the port is no longer available, the message soup g contains the set of
messages that have been received at the socket (addr r↦{ 1 / 2} M). *)
wp_pures. iLöb as "IH" forall (M HΦ).
(* the listen spec consumes the unbound socket and the message soup. *)
iApply (listen_spec True%I with "[] [$] [$]"); first done.
iIntros (m rm') "!# % (_ & Hsocket & Hrm & Hm & _ & Hpsi) HΦ /=".
iDestruct "Hpsi" as (si) "(#Hsp & % & Hwand)". iSimplifyEq.
(* due to the socket interpration we know the conditional will always be true *)
rewrite H. wp_pures.
(* Send the message. *)
wp_apply (wp_send_to_bound with "[Hwand $Hsocket $Hsp]"); first done.
- (* We prove that the message adheres to the socket interpretation *)
by iApply "Hwand".
- (* The recursive call goes well c.f. the IH *)
iIntros "[Hsocket _] /=". wp_pures.
by iApply ("IH" with "HΦ Hsocket Hrm").
Qed.
End pong.
Section ping.
Context `{dG : distG Σ}
`{node : Network.node}.
Definition ping_si : socket_interp Σ :=
(λ msg, ⌜ms_body msg = "PONG"⌝)%I.
Lemma ping_spec (A : gset socket_address) (addr pong : socket_address):
(* the pong address is static *)
pong ∈ A ->
(* the ping is not*)
addr ∉ A ->
{{{ (* the pong server satisfies its socket interpretation *)
pong ⤇ pong_si
(* A should contain static addresses & the port should be free *)
∗ Fixed A
∗ FreePorts (ip_of_address addr) {[port_of_address addr]}
(* the node is valid *)
∗ IsNode node }}}
⟨node; ping #addr #pong⟩
{{{ v, RET ⟨node; v⟩; ⌜v = #"PONG"⌝ }}}.
Proof.
iIntros (Hserver Haddr Φ) "(#Hsi & #Hfixed & Hip & Hn) Hcont".
wp_lam. wp_let.
(* Allocate socket, use that the node is valid. *)
wp_socket h as "Hsocket". wp_let.
(* Dynamic binding requires that the port is free, the socket is allocated
and a socket interpretation *)
wp_apply (wp_socketbind_dynamic _ _ _ _ _ _ _ (ping_si) with "[$]"); try done.
iIntros "H //=". iDestruct "H" as (g) "(Hsocket & Hbind & #Hpingsi)".
(* Send the ping message *)
wp_apply (wp_send_to_bound with "[$Hsi $Hsocket]"); first done.
- (* The message adheres to the socket interpretation *)
rewrite /pong_si /=. iExists ping_si. simpl. by iSplit.
- (* The rest of the progrma goes well c.f. listen_wait_spec *)
iIntros "[Hsocket _]". wp_seq. wp_bind (listen_wait _)%E.
iApply (listen_wait_spec with "[$]"); try done.
iNext. iIntros (m mId) "(Hsocket & Hbind & Hm & _ & H)".
wp_proj. rewrite /ping_si /=.
iDestruct "H" as "->". by iApply "Hcont".
Qed.
End ping.
Section ping_pong_runner.
Context `{dG : distG Σ}.
(* the set of ips used *)
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) :
(* the pong address is static *)
pong_addr ∈ A ->
(* the ping adress is not *)
ping_addr ∉ A ->
{{{ (* the pong server satisfies its socket interpretation *)
pong_addr ⤇ pong_si
(* A contain static addresses, and the ips we use are free *)
∗ Fixed A
∗ [∗ set] ip ∈ ips, FreeIP ip }}}
ping_pong_runner
{{{ v, RET v; True }}}.
Proof.
iIntros (server client Φ) "(#Hsi & #Hfix & Hips) Hcont". rewrite /ping_pong_runner.
(* Split up the free ips in separate resources *)
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.
(* start the nodes *)
wp_apply (wp_start with "[-]"); first done. iFrame. simpl. iSplitL; last first.
- iNext. iIntros "Hγs Hn".
by iApply (pong_spec _ _ server with "[$]").
- iNext. wp_seq.
wp_apply (wp_start with "[-]"); first done. iFrame. iSplitL.
+ by iApply "Hcont".
+ iNext. iIntros "Hγs Hn".
(* we provide the predicate as Coq is apparently not able to resolve it
through unification ... *)
by iApply (ping_spec _ _ _ server client (λ _, True%I) with "[$]").
Qed.
End ping_pong_runner.
Definition ping_pong_is :=
{|
state_heaps := ∅;
state_sockets := ∅;
state_ports_in_use :=
<["0.0.0.0" := ∅ ]> $ <["0.0.0.1" := ∅ ]> $ ∅;
state_ms := ∅;
|}.
(* 'Global' socket interepration, i.e. all statically bound socket interpretions *)
Definition socket_interp `{distG} sa : socket_interp Σ :=
(match sa with
| SocketAddressInet "0.0.0.0" 80 => pong_si (* pong *)
| _ => λ msg, ⌜True⌝
end)%I.
(* The static/fixed domain contains only the adress of the pong server *)
Definition fixed_dom : gset socket_address := {[ pong_addr ]}.
Theorem ping_pong_safe : adequate NotStuck ping_pong_runner ping_pong_is (λ _ _, True).
Proof.
set (Σ := #[distΣ]).
eapply (@dist_adequacy Σ _ ips fixed_dom); try done; last first.
{ set_solver. }
{ intros ?. 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). iModIntro. iExists socket_interp.
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.