LogrelCC.examples.refinement.server.refinement_1
From iris.algebra Require Import agree auth.
From iris.proofmode Require Import tactics.
From LogrelCC Require Export fundamental_binary logrel_binary rules_unary
rules_binary.
From LogrelCC.examples Require Import list_rev
list_assoc list_basics lock lock_unary_spec list_assoc_binary_spec.
From LogrelCC.examples.refinement.server Require Import server_nat server_cont.
Section server_refinement.
Context `{heapG Σ, cfgSG Σ}.
Context `{inG Σ (exclR unitR)}. (* for the locks *)
(* The continuation stored by the continuation-based server. N :
the sum so far; assoc: the associate function applied to the table;
K: the context that we are working under. K is existentially
quantified below. *)
Definition stored_cont N assoc K :=
ContV
(LetInCtx
(App
(Rec
(LetIn
(Fst (ids 1))
(LetIn
(Snd (ids 2))
(Seq
(App (ids 0) (InjR (ids 1)))
(LetIn
(Callcc
(Seq
(LetIn
(InjL
(App
(of_val assoc)
(ids 0))) (App (ids 2) (ids 0)))
(Throw Unit (Cont []))))
(App (ids 3)
(Pair
(BinOp Add (ids 2) (Fst (ids 0)))
(Snd (ids 0)))))))))
(Pair (BinOp Add (#n N) (Fst (ids 0))) (Snd (ids 0)))) :: K).
(* The relation relating the two stored entities: a continuation and a number. *)
Program Definition rel_cont_nat assoc vv : iProp Σ :=
(∃ K n, ⌜vv = (stored_cont n assoc K, #nv n)⌝)%I.
Lemma server_refinement_cont_nat : [] ⊨ server_cont ≤log≤ server_nat :
TArrow (TProd (TArrow TUnit (TProd (TSum TNat TUnit) TNat))
(TArrow (TSum TNat TNat) TUnit)) TUnit.
Proof.
iIntros (Δ vvs ρ HΔ) "[#Hspc #Hvvs] /=".
iDestruct (interp_env_length with "Hvvs") as %Heq; destruct vvs; last done.
iClear "Hvvs".
asimpl.
iIntros ([K K'] j) "[Hj #HKK]"; simpl.
remember (∀ _, _)%I as HKK.
iApply (wp_tapp (LetInCtx _:: _)). simpl. iNext.
iMod (step_tlam _ _ _ (LetInCtx _:: _) with "[$Hj]") as "Hj"; eauto.
iApply (wp_alloc (LetInCtx _ :: LetInCtx _ :: _)); eauto.
iNext. iIntros (lst) "Hl /=".
iMod (step_alloc _ _ _ (LetInCtx _ :: LetInCtx _ :: _)
with "[$Hj]") as (lst') "[Hj Hl'] /="; eauto.
iApply (wp_LetIn (LetInCtx _ :: _)); eauto. iNext.
simpl. asimpl.
iMod (step_LetIn _ _ _ (LetInCtx _ :: _) with "[$Hj]") as "Hj"; eauto.
simpl. asimpl.
iMod (steps_newlock _ _ _ (LetInCtx _ :: LetInCtx _ :: _)
with "[$Hj]") as (lock') "[Hj Hlst]"; eauto.
iApply (clwp_cl (LetInCtx _ :: LetInCtx _ :: _)
with "[] [-]"); first by iApply (@clwp_newlock).
iIntros (l); iDestruct 1 as (lock γ) "[% Hlk]"; subst l.
iApply (wp_LetIn (LetInCtx _ :: _)); eauto. iNext.
simpl. asimpl.
iMod (step_LetIn _ _ _ (LetInCtx _ :: _) with "[$Hj]") as "Hj"; eauto; eauto.
simpl. asimpl.
iApply (wp_LetIn); eauto. iNext.
simpl. asimpl.
iMod (step_LetIn with "[$Hj]") as "Hj"; eauto; eauto.
simpl. asimpl.
iApply (wp_fst (LetInCtx _ :: _)); eauto.
iNext. simpl.
iMod (step_fst _ _ _ (LetInCtx _ :: _) with "[$Hj]") as "Hj"; eauto; simpl.
iApply wp_LetIn; eauto.
iNext.
simpl. asimpl.
iMod (step_LetIn with "[$Hj]") as "Hj"; eauto.
simpl. asimpl.
iApply (wp_snd (LetInCtx _ :: _)); eauto.
iNext; simpl.
iMod (step_snd _ _ _ (LetInCtx _ :: _) with "[$Hj]") as "Hj"; eauto; simpl.
iApply wp_LetIn; eauto.
iNext. simpl. asimpl.
iMod (step_LetIn with "[$Hj]") as "Hj"; eauto.
simpl. asimpl.
iAssert ((|={⊤}=> assoc_invariant
γ lock lock' lst lst'
(rel_cont_nat (associateV (Loc lock) (Loc lst))))%I)
with "[Hl Hl' Hlst Hlk]" as "Hinv" .
{ iMod ("Hlk" $! (lock' ↦ₛ (#♭v false) ∗
∃ ls, lst ↦ᵢ embed_list (map fst ls) ∗
lst' ↦ₛ embed_list (map snd ls) ∗
[∗ list] x ∈ ls, (rel_cont_nat (associateV (Loc lock) (Loc lst)) x))%I
with "[Hl Hl' Hlst]") as "#Hlk'".
{ iFrame. iExists []; simpl; iFrame. }
auto. }
iMod "Hinv" as "#Hinv".
rewrite HeqHKK.
iApply ("HKK" $! (LamV _, LamV _) with "[$Hj]").
iClear "HKK". clear HeqHKK HKK K K' j.
iAlways.
iIntros ([chan chan']) "Hchan". iIntros ([K K'] j) "[Hj #HKK] /=".
iDestruct "Hchan" as ([reader reader'] [writer writer'])
"[% [#Hreaders #Hwriters]]"; iSimplifyEq.
remember (∀ _, _)%I as HKK in |- *.
remember (∀ _, _)%I as Hreaders in |- *.
remember (∀ _, _)%I as Hwriters in |- *.
iApply wp_Lam; eauto.
iNext; simpl; asimpl.
iMod (step_Lam with "[$Hj]") as "Hj"; eauto.
simpl. asimpl.
iApply (wp_fst (LetInCtx _ :: _)); eauto.
iNext; simpl.
iMod (step_fst _ _ _ (LetInCtx _ :: _) with "[$Hj]") as "Hj"; eauto; simpl.
iApply wp_LetIn; eauto.
iNext. simpl. asimpl.
iMod (step_LetIn with "[$Hj]") as "Hj"; eauto.
simpl. asimpl.
iApply (wp_snd (LetInCtx _ :: _)); eauto. iNext; simpl.
iMod (step_snd _ _ _ (LetInCtx _ :: _) with "[$Hj]") as "Hj"; eauto; simpl.
iApply wp_LetIn; eauto.
iNext. simpl. asimpl.
iMod (step_LetIn with "[$Hj]") as "Hj"; eauto.
simpl. asimpl.
rewrite server_nat_main_eq server_cont_main_eq.
asimpl.
rewrite HeqHreaders.
iPoseProof ("Hreaders" $! (UnitV, UnitV) with "[]") as "Hrd"; first done.
rewrite -HeqHreaders.
iApply ("Hrd" $! (LetInCtx _ :: _, LetInCtx _ :: _) with "[$Hj]").
clear j.
iAlways. iIntros ([v1 v2] j) "[Hj #Hvv]".
iDestruct "Hvv" as ([rid rid'] [n n']) "[Heq1 [Hrid Hn]]";
iDestruct "Heq1" as %Heq1; inversion Heq1; subst v1 v2; clear Heq1.
iDestruct "Hn" as (N) "/= [% %]"; subst n n'.
iDestruct "Hrid" as "[Hrid|Hrid]".
- (* rid provided *)
iDestruct "Hrid" as ([v1 w2]) "[Heq1 Hrid]";
iDestruct "Heq1" as %Heq1; inversion Heq1; subst rid rid'; clear Heq1.
iDestruct "Hrid" as (resid) "/= [% %]"; subst v1 w2.
iApply wp_LetIn; eauto. iNext. simpl. asimpl.
iMod (step_LetIn with "[$Hj]") as "Hj"; eauto. simpl. asimpl.
iApply (wp_fst (CaseCtx _ _ :: _)); eauto.
simpl. iNext.
iMod (step_fst _ _ _ (CaseCtx _ _ :: _) with "[Hj]") as "Hj";
eauto; simpl; eauto.
iApply wp_case_injl; eauto. simpl; iNext. asimpl.
iMod (step_case_inl with "[$Hj]") as "Hj"; eauto. simpl. asimpl.
iApply (clwp_cl (CaseCtx _ _ :: _) with "[Hj]").
{ iApply (get_rel _ _ _ (CaseCtx _ _ :: _));
try iFrame "#"; try iFrame; eauto.
apply _. }
iIntros (v) "[[Heq Hj]|HC]".
+ iDestruct "Heq" as %->.
simpl.
iApply wp_case_injl; eauto. simpl; iNext.
iMod (step_case_inl with "[$Hj]") as "Hj"; eauto.
simpl.
rewrite HeqHKK.
iApply ("HKK" $! (UnitV, UnitV)); auto.
+ iDestruct "HC" as (w w') "(#Hww & Heq & Hj)".
iDestruct "Heq" as %->. simpl.
iApply wp_case_injr; eauto. simpl; iNext.
iMod (step_case_inr with "[$Hj]") as "Hj"; eauto.
simpl. asimpl.
iDestruct "Hww" as (Kex n) "Heq1"; iDestruct "Heq1" as %Heq1;
inversion Heq1; subst w w'.
iApply (wp_snd (PairLCtx _ :: ThrowLCtx _ :: _)); eauto.
simpl. iNext.
iApply wp_throw; eauto; simpl.
iNext.
iApply wp_LetIn; eauto. iNext.
simpl. asimpl.
iApply (wp_fst (BinOpRCtx _ (#nv _) ::
PairLCtx _ :: AppRCtx (RecV _) :: _)); eauto.
simpl. iNext.
iApply (wp_bin_op (PairLCtx _ :: AppRCtx (RecV _) :: _));
eauto. simpl. iNext.
iMod (step_snd _ _ _ (BinOpRCtx _ (#nv _) :: LetInCtx _ :: _)
with "[$Hj]") as "Hj"; eauto. simpl.
iMod (step_nat_binop _ _ _ (LetInCtx _ :: _) with "[$Hj]")
as "Hj"; eauto.
iApply (wp_snd (PairRCtx (#nv _) :: AppRCtx (RecV _) :: _));
eauto. simpl. iNext.
rewrite associate_of_val; asimpl.
iApply wp_rec; eauto. iNext.
match goal with |- context[_.[?A, _/]] => remember A as Orec end.
simpl. asimpl.
iApply (wp_fst (LetInCtx _ :: _)); eauto; simpl.
iNext.
iApply wp_LetIn; eauto. iNext. simpl. asimpl.
iApply (wp_snd (LetInCtx _ :: _)); eauto; simpl. iNext.
iApply wp_LetIn; eauto. iNext. simpl. asimpl.
iMod (step_LetIn with "[$Hj]") as "Hj"; eauto. asimpl.
rewrite HeqHwriters.
iPoseProof ("Hwriters" $! (InjRV (#nv (n + N)), InjRV (#nv (n + N)))
with "[]") as "Hrw";
first by iRight; iExists (_, _); eauto.
rewrite -HeqHwriters.
simpl.
iApply ("Hrw" $! (SeqCtx _ :: _, SeqCtx _ :: _)); iFrame.
iAlways. clear j.
iIntros ([u1 u2] j) "/= [Hj [% %]]"; subst u1 u2.
iApply wp_Seq; eauto. simpl. iNext.
iMod (step_Seq with "[$Hj]") as "Hj"; eauto.
iApply (wp_callcc (LetInCtx _ :: _)). iNext. simpl. asimpl.
iApply (clwp_cl (InjLCtx :: LetInCtx _:: SeqCtx _
:: LetInCtx _ :: _) with "[Hj]").
{ rewrite HeqOrec.
iApply (associate_rel _ _ _ (InjLCtx :: AppRCtx _:: SeqCtx _ :: _)); eauto;
try iFrame "#"; try iFrame; eauto. }
iIntros (v).
iDestruct 1 as (newresid) "[Heq Hj]"; iDestruct "Heq" as %-> ; simpl.
iApply (wp_LetIn (SeqCtx _ :: LetInCtx _ :: _)); eauto. simpl. iNext.
asimpl.
rewrite HeqHwriters.
iPoseProof ("Hwriters" $! (InjLV (#nv newresid), InjLV (#nv newresid))
with "[]") as "Hrw";
first by iLeft; iExists (_, _); eauto. simpl.
rewrite -HeqHwriters.
iApply ("Hrw" $! (SeqCtx _ :: LetInCtx _ :: _, SeqCtx _ :: _)); iFrame.
iAlways. clear j.
iIntros ([w1 w2] j) "/= [Hj [% %]]"; subst w1 w2.
iApply (wp_Seq (LetInCtx _ :: _)); eauto.
iNext. simpl.
iApply (wp_throw (LetInCtx _ :: _)); eauto; simpl. iNext.
iMod (step_Seq with "[$Hj]") as "Hj"; eauto.
iMod (step_throw with "[$Hj]") as "Hj"; eauto using to_of_val.
iApply wp_value; eauto. iExists UnitV; eauto.
- (* rid not provided *)
iDestruct "Hrid" as ([v1 w2]) "/=[Heq1 [% %]]"; iDestruct "Heq1" as %Heq1;
inversion Heq1; subst rid rid' v1 w2; clear Heq1.
iApply wp_LetIn; eauto.
iNext. simpl. asimpl.
iApply (wp_fst (CaseCtx _ _ :: _)); eauto.
simpl. iNext.
iMod (step_LetIn with "[$Hj]") as "Hj"; eauto. simpl. asimpl.
iMod (step_fst _ _ _ (CaseCtx _ _ :: _) with "[Hj]") as "Hj";
eauto; simpl; eauto.
iApply wp_case_injr; eauto. simpl; iNext.
asimpl.
iMod (step_case_inr with "[$Hj]") as "Hj"; eauto.
asimpl.
iApply (wp_snd (PairLCtx _ :: AppRCtx (RecV _) :: _)); eauto.
simpl. iNext.
iMod (step_snd _ _ _ (InjRCtx :: AppRCtx _ :: SeqCtx _ :: _)
with "[$Hj]") as "Hj"; eauto. simpl.
iApply wp_rec; eauto. iNext.
match goal with |- context[_.[?A, _/]] => remember A as O end.
simpl. asimpl.
iApply (wp_fst (LetInCtx _ :: _)); eauto. iNext. simpl.
iApply wp_LetIn; eauto. iNext.
asimpl.
iApply (wp_snd (LetInCtx _ :: _)); eauto.
iNext. simpl.
iApply wp_LetIn; eauto. iNext. simpl. asimpl.
rewrite HeqHwriters.
iPoseProof ("Hwriters" $! (InjRV (#nv N), InjRV (#nv N)) with "[]")
as "Hrw"; first by iRight; iExists (_, _); eauto.
rewrite -HeqHwriters.
simpl.
iApply ("Hrw" $! (SeqCtx _ :: _, SeqCtx _ :: _)); iFrame.
iAlways. clear j.
iIntros ([u1 u2] j) "[Hj [Heq1 Heq2]] /=".
iDestruct "Heq1" as %->. iDestruct "Heq2" as %->.
iApply wp_Seq; eauto. iNext.
iMod (step_Seq with "[$Hj]") as "Hj"; eauto. simpl.
iApply (wp_callcc (LetInCtx _ :: _)). iNext. simpl. asimpl.
iMod (step_snd _ _ _ (AppRCtx (associateV (Loc _) (Loc _)) :: InjLCtx ::
(AppRCtx _) :: SeqCtx _:: _) with "[Hj]")
as "Hj"; rewrite /= ?associate_of_val; eauto; eauto.
iApply (clwp_cl
(InjLCtx :: LetInCtx _ :: SeqCtx _ :: LetInCtx _ :: _)
with "[Hj]").
{ iApply (associate_rel _ _ _ (InjLCtx :: AppRCtx _ :: SeqCtx _ :: _));
try iFrame "#"; try iFrame; eauto; eauto.
iExists _, _; iPureIntro.
f_equal. rewrite /stored_cont HeqO; trivial. }
iIntros (v).
iDestruct 1 as (resid) "[Heq Hj]"; iDestruct "Heq" as %-> ;simpl.
iApply (wp_LetIn (SeqCtx _ :: LetInCtx _ :: _));
eauto.
iNext. simpl. asimpl.
rewrite HeqHwriters.
iPoseProof ("Hwriters" $! (InjLV (#nv resid), InjLV (#nv resid)) with "[]")
as "Hrw"; first by iLeft; iExists (_, _); eauto.
rewrite -HeqHwriters.
simpl.
iApply ("Hrw" $! (SeqCtx _ :: LetInCtx _ :: _, SeqCtx _ :: _)); iFrame.
simpl.
iAlways. clear j.
iIntros ([w1 w2] j) "[Hj %]"; simpl.
iApply (wp_Seq (LetInCtx _ :: _)); eauto. simpl. iNext.
iApply (wp_throw (LetInCtx _ :: _)); eauto. simpl. iNext.
iMod (step_Seq with "[$Hj]") as "Hj"; eauto. simpl.
iMod (step_throw with "[$Hj]") as "Hj"; eauto.
iApply wp_value; eauto. iExists UnitV; eauto.
Qed.
End server_refinement.
Require Import LogrelCC.soundness_binary.
Instance subG_lockΣ {Σ} : subG (GFunctor (exclR unitR)) Σ → inG Σ (exclR unitR).
Proof. solve_inG. Qed.
Theorem server_ctx_refinement_cont_nat :
[] ⊨ server_cont ≤ctx≤ server_nat :
TArrow (TProd (TArrow TUnit (TProd (TSum TNat TUnit) TNat))
(TArrow (TSum TNat TNat) TUnit)) TUnit.
Proof.
set (Σ := #[invΣ ; gen_heapΣ loc val ;
GFunctor (authR cfgUR) ; GFunctor (exclR unitR) ]).
set (HG := soundness_unary.HeapPreIG Σ _ _).
eapply (binary_soundness Σ); auto.
intros. eapply server_refinement_cont_nat.
Qed.
From iris.proofmode Require Import tactics.
From LogrelCC Require Export fundamental_binary logrel_binary rules_unary
rules_binary.
From LogrelCC.examples Require Import list_rev
list_assoc list_basics lock lock_unary_spec list_assoc_binary_spec.
From LogrelCC.examples.refinement.server Require Import server_nat server_cont.
Section server_refinement.
Context `{heapG Σ, cfgSG Σ}.
Context `{inG Σ (exclR unitR)}. (* for the locks *)
(* The continuation stored by the continuation-based server. N :
the sum so far; assoc: the associate function applied to the table;
K: the context that we are working under. K is existentially
quantified below. *)
Definition stored_cont N assoc K :=
ContV
(LetInCtx
(App
(Rec
(LetIn
(Fst (ids 1))
(LetIn
(Snd (ids 2))
(Seq
(App (ids 0) (InjR (ids 1)))
(LetIn
(Callcc
(Seq
(LetIn
(InjL
(App
(of_val assoc)
(ids 0))) (App (ids 2) (ids 0)))
(Throw Unit (Cont []))))
(App (ids 3)
(Pair
(BinOp Add (ids 2) (Fst (ids 0)))
(Snd (ids 0)))))))))
(Pair (BinOp Add (#n N) (Fst (ids 0))) (Snd (ids 0)))) :: K).
(* The relation relating the two stored entities: a continuation and a number. *)
Program Definition rel_cont_nat assoc vv : iProp Σ :=
(∃ K n, ⌜vv = (stored_cont n assoc K, #nv n)⌝)%I.
Lemma server_refinement_cont_nat : [] ⊨ server_cont ≤log≤ server_nat :
TArrow (TProd (TArrow TUnit (TProd (TSum TNat TUnit) TNat))
(TArrow (TSum TNat TNat) TUnit)) TUnit.
Proof.
iIntros (Δ vvs ρ HΔ) "[#Hspc #Hvvs] /=".
iDestruct (interp_env_length with "Hvvs") as %Heq; destruct vvs; last done.
iClear "Hvvs".
asimpl.
iIntros ([K K'] j) "[Hj #HKK]"; simpl.
remember (∀ _, _)%I as HKK.
iApply (wp_tapp (LetInCtx _:: _)). simpl. iNext.
iMod (step_tlam _ _ _ (LetInCtx _:: _) with "[$Hj]") as "Hj"; eauto.
iApply (wp_alloc (LetInCtx _ :: LetInCtx _ :: _)); eauto.
iNext. iIntros (lst) "Hl /=".
iMod (step_alloc _ _ _ (LetInCtx _ :: LetInCtx _ :: _)
with "[$Hj]") as (lst') "[Hj Hl'] /="; eauto.
iApply (wp_LetIn (LetInCtx _ :: _)); eauto. iNext.
simpl. asimpl.
iMod (step_LetIn _ _ _ (LetInCtx _ :: _) with "[$Hj]") as "Hj"; eauto.
simpl. asimpl.
iMod (steps_newlock _ _ _ (LetInCtx _ :: LetInCtx _ :: _)
with "[$Hj]") as (lock') "[Hj Hlst]"; eauto.
iApply (clwp_cl (LetInCtx _ :: LetInCtx _ :: _)
with "[] [-]"); first by iApply (@clwp_newlock).
iIntros (l); iDestruct 1 as (lock γ) "[% Hlk]"; subst l.
iApply (wp_LetIn (LetInCtx _ :: _)); eauto. iNext.
simpl. asimpl.
iMod (step_LetIn _ _ _ (LetInCtx _ :: _) with "[$Hj]") as "Hj"; eauto; eauto.
simpl. asimpl.
iApply (wp_LetIn); eauto. iNext.
simpl. asimpl.
iMod (step_LetIn with "[$Hj]") as "Hj"; eauto; eauto.
simpl. asimpl.
iApply (wp_fst (LetInCtx _ :: _)); eauto.
iNext. simpl.
iMod (step_fst _ _ _ (LetInCtx _ :: _) with "[$Hj]") as "Hj"; eauto; simpl.
iApply wp_LetIn; eauto.
iNext.
simpl. asimpl.
iMod (step_LetIn with "[$Hj]") as "Hj"; eauto.
simpl. asimpl.
iApply (wp_snd (LetInCtx _ :: _)); eauto.
iNext; simpl.
iMod (step_snd _ _ _ (LetInCtx _ :: _) with "[$Hj]") as "Hj"; eauto; simpl.
iApply wp_LetIn; eauto.
iNext. simpl. asimpl.
iMod (step_LetIn with "[$Hj]") as "Hj"; eauto.
simpl. asimpl.
iAssert ((|={⊤}=> assoc_invariant
γ lock lock' lst lst'
(rel_cont_nat (associateV (Loc lock) (Loc lst))))%I)
with "[Hl Hl' Hlst Hlk]" as "Hinv" .
{ iMod ("Hlk" $! (lock' ↦ₛ (#♭v false) ∗
∃ ls, lst ↦ᵢ embed_list (map fst ls) ∗
lst' ↦ₛ embed_list (map snd ls) ∗
[∗ list] x ∈ ls, (rel_cont_nat (associateV (Loc lock) (Loc lst)) x))%I
with "[Hl Hl' Hlst]") as "#Hlk'".
{ iFrame. iExists []; simpl; iFrame. }
auto. }
iMod "Hinv" as "#Hinv".
rewrite HeqHKK.
iApply ("HKK" $! (LamV _, LamV _) with "[$Hj]").
iClear "HKK". clear HeqHKK HKK K K' j.
iAlways.
iIntros ([chan chan']) "Hchan". iIntros ([K K'] j) "[Hj #HKK] /=".
iDestruct "Hchan" as ([reader reader'] [writer writer'])
"[% [#Hreaders #Hwriters]]"; iSimplifyEq.
remember (∀ _, _)%I as HKK in |- *.
remember (∀ _, _)%I as Hreaders in |- *.
remember (∀ _, _)%I as Hwriters in |- *.
iApply wp_Lam; eauto.
iNext; simpl; asimpl.
iMod (step_Lam with "[$Hj]") as "Hj"; eauto.
simpl. asimpl.
iApply (wp_fst (LetInCtx _ :: _)); eauto.
iNext; simpl.
iMod (step_fst _ _ _ (LetInCtx _ :: _) with "[$Hj]") as "Hj"; eauto; simpl.
iApply wp_LetIn; eauto.
iNext. simpl. asimpl.
iMod (step_LetIn with "[$Hj]") as "Hj"; eauto.
simpl. asimpl.
iApply (wp_snd (LetInCtx _ :: _)); eauto. iNext; simpl.
iMod (step_snd _ _ _ (LetInCtx _ :: _) with "[$Hj]") as "Hj"; eauto; simpl.
iApply wp_LetIn; eauto.
iNext. simpl. asimpl.
iMod (step_LetIn with "[$Hj]") as "Hj"; eauto.
simpl. asimpl.
rewrite server_nat_main_eq server_cont_main_eq.
asimpl.
rewrite HeqHreaders.
iPoseProof ("Hreaders" $! (UnitV, UnitV) with "[]") as "Hrd"; first done.
rewrite -HeqHreaders.
iApply ("Hrd" $! (LetInCtx _ :: _, LetInCtx _ :: _) with "[$Hj]").
clear j.
iAlways. iIntros ([v1 v2] j) "[Hj #Hvv]".
iDestruct "Hvv" as ([rid rid'] [n n']) "[Heq1 [Hrid Hn]]";
iDestruct "Heq1" as %Heq1; inversion Heq1; subst v1 v2; clear Heq1.
iDestruct "Hn" as (N) "/= [% %]"; subst n n'.
iDestruct "Hrid" as "[Hrid|Hrid]".
- (* rid provided *)
iDestruct "Hrid" as ([v1 w2]) "[Heq1 Hrid]";
iDestruct "Heq1" as %Heq1; inversion Heq1; subst rid rid'; clear Heq1.
iDestruct "Hrid" as (resid) "/= [% %]"; subst v1 w2.
iApply wp_LetIn; eauto. iNext. simpl. asimpl.
iMod (step_LetIn with "[$Hj]") as "Hj"; eauto. simpl. asimpl.
iApply (wp_fst (CaseCtx _ _ :: _)); eauto.
simpl. iNext.
iMod (step_fst _ _ _ (CaseCtx _ _ :: _) with "[Hj]") as "Hj";
eauto; simpl; eauto.
iApply wp_case_injl; eauto. simpl; iNext. asimpl.
iMod (step_case_inl with "[$Hj]") as "Hj"; eauto. simpl. asimpl.
iApply (clwp_cl (CaseCtx _ _ :: _) with "[Hj]").
{ iApply (get_rel _ _ _ (CaseCtx _ _ :: _));
try iFrame "#"; try iFrame; eauto.
apply _. }
iIntros (v) "[[Heq Hj]|HC]".
+ iDestruct "Heq" as %->.
simpl.
iApply wp_case_injl; eauto. simpl; iNext.
iMod (step_case_inl with "[$Hj]") as "Hj"; eauto.
simpl.
rewrite HeqHKK.
iApply ("HKK" $! (UnitV, UnitV)); auto.
+ iDestruct "HC" as (w w') "(#Hww & Heq & Hj)".
iDestruct "Heq" as %->. simpl.
iApply wp_case_injr; eauto. simpl; iNext.
iMod (step_case_inr with "[$Hj]") as "Hj"; eauto.
simpl. asimpl.
iDestruct "Hww" as (Kex n) "Heq1"; iDestruct "Heq1" as %Heq1;
inversion Heq1; subst w w'.
iApply (wp_snd (PairLCtx _ :: ThrowLCtx _ :: _)); eauto.
simpl. iNext.
iApply wp_throw; eauto; simpl.
iNext.
iApply wp_LetIn; eauto. iNext.
simpl. asimpl.
iApply (wp_fst (BinOpRCtx _ (#nv _) ::
PairLCtx _ :: AppRCtx (RecV _) :: _)); eauto.
simpl. iNext.
iApply (wp_bin_op (PairLCtx _ :: AppRCtx (RecV _) :: _));
eauto. simpl. iNext.
iMod (step_snd _ _ _ (BinOpRCtx _ (#nv _) :: LetInCtx _ :: _)
with "[$Hj]") as "Hj"; eauto. simpl.
iMod (step_nat_binop _ _ _ (LetInCtx _ :: _) with "[$Hj]")
as "Hj"; eauto.
iApply (wp_snd (PairRCtx (#nv _) :: AppRCtx (RecV _) :: _));
eauto. simpl. iNext.
rewrite associate_of_val; asimpl.
iApply wp_rec; eauto. iNext.
match goal with |- context[_.[?A, _/]] => remember A as Orec end.
simpl. asimpl.
iApply (wp_fst (LetInCtx _ :: _)); eauto; simpl.
iNext.
iApply wp_LetIn; eauto. iNext. simpl. asimpl.
iApply (wp_snd (LetInCtx _ :: _)); eauto; simpl. iNext.
iApply wp_LetIn; eauto. iNext. simpl. asimpl.
iMod (step_LetIn with "[$Hj]") as "Hj"; eauto. asimpl.
rewrite HeqHwriters.
iPoseProof ("Hwriters" $! (InjRV (#nv (n + N)), InjRV (#nv (n + N)))
with "[]") as "Hrw";
first by iRight; iExists (_, _); eauto.
rewrite -HeqHwriters.
simpl.
iApply ("Hrw" $! (SeqCtx _ :: _, SeqCtx _ :: _)); iFrame.
iAlways. clear j.
iIntros ([u1 u2] j) "/= [Hj [% %]]"; subst u1 u2.
iApply wp_Seq; eauto. simpl. iNext.
iMod (step_Seq with "[$Hj]") as "Hj"; eauto.
iApply (wp_callcc (LetInCtx _ :: _)). iNext. simpl. asimpl.
iApply (clwp_cl (InjLCtx :: LetInCtx _:: SeqCtx _
:: LetInCtx _ :: _) with "[Hj]").
{ rewrite HeqOrec.
iApply (associate_rel _ _ _ (InjLCtx :: AppRCtx _:: SeqCtx _ :: _)); eauto;
try iFrame "#"; try iFrame; eauto. }
iIntros (v).
iDestruct 1 as (newresid) "[Heq Hj]"; iDestruct "Heq" as %-> ; simpl.
iApply (wp_LetIn (SeqCtx _ :: LetInCtx _ :: _)); eauto. simpl. iNext.
asimpl.
rewrite HeqHwriters.
iPoseProof ("Hwriters" $! (InjLV (#nv newresid), InjLV (#nv newresid))
with "[]") as "Hrw";
first by iLeft; iExists (_, _); eauto. simpl.
rewrite -HeqHwriters.
iApply ("Hrw" $! (SeqCtx _ :: LetInCtx _ :: _, SeqCtx _ :: _)); iFrame.
iAlways. clear j.
iIntros ([w1 w2] j) "/= [Hj [% %]]"; subst w1 w2.
iApply (wp_Seq (LetInCtx _ :: _)); eauto.
iNext. simpl.
iApply (wp_throw (LetInCtx _ :: _)); eauto; simpl. iNext.
iMod (step_Seq with "[$Hj]") as "Hj"; eauto.
iMod (step_throw with "[$Hj]") as "Hj"; eauto using to_of_val.
iApply wp_value; eauto. iExists UnitV; eauto.
- (* rid not provided *)
iDestruct "Hrid" as ([v1 w2]) "/=[Heq1 [% %]]"; iDestruct "Heq1" as %Heq1;
inversion Heq1; subst rid rid' v1 w2; clear Heq1.
iApply wp_LetIn; eauto.
iNext. simpl. asimpl.
iApply (wp_fst (CaseCtx _ _ :: _)); eauto.
simpl. iNext.
iMod (step_LetIn with "[$Hj]") as "Hj"; eauto. simpl. asimpl.
iMod (step_fst _ _ _ (CaseCtx _ _ :: _) with "[Hj]") as "Hj";
eauto; simpl; eauto.
iApply wp_case_injr; eauto. simpl; iNext.
asimpl.
iMod (step_case_inr with "[$Hj]") as "Hj"; eauto.
asimpl.
iApply (wp_snd (PairLCtx _ :: AppRCtx (RecV _) :: _)); eauto.
simpl. iNext.
iMod (step_snd _ _ _ (InjRCtx :: AppRCtx _ :: SeqCtx _ :: _)
with "[$Hj]") as "Hj"; eauto. simpl.
iApply wp_rec; eauto. iNext.
match goal with |- context[_.[?A, _/]] => remember A as O end.
simpl. asimpl.
iApply (wp_fst (LetInCtx _ :: _)); eauto. iNext. simpl.
iApply wp_LetIn; eauto. iNext.
asimpl.
iApply (wp_snd (LetInCtx _ :: _)); eauto.
iNext. simpl.
iApply wp_LetIn; eauto. iNext. simpl. asimpl.
rewrite HeqHwriters.
iPoseProof ("Hwriters" $! (InjRV (#nv N), InjRV (#nv N)) with "[]")
as "Hrw"; first by iRight; iExists (_, _); eauto.
rewrite -HeqHwriters.
simpl.
iApply ("Hrw" $! (SeqCtx _ :: _, SeqCtx _ :: _)); iFrame.
iAlways. clear j.
iIntros ([u1 u2] j) "[Hj [Heq1 Heq2]] /=".
iDestruct "Heq1" as %->. iDestruct "Heq2" as %->.
iApply wp_Seq; eauto. iNext.
iMod (step_Seq with "[$Hj]") as "Hj"; eauto. simpl.
iApply (wp_callcc (LetInCtx _ :: _)). iNext. simpl. asimpl.
iMod (step_snd _ _ _ (AppRCtx (associateV (Loc _) (Loc _)) :: InjLCtx ::
(AppRCtx _) :: SeqCtx _:: _) with "[Hj]")
as "Hj"; rewrite /= ?associate_of_val; eauto; eauto.
iApply (clwp_cl
(InjLCtx :: LetInCtx _ :: SeqCtx _ :: LetInCtx _ :: _)
with "[Hj]").
{ iApply (associate_rel _ _ _ (InjLCtx :: AppRCtx _ :: SeqCtx _ :: _));
try iFrame "#"; try iFrame; eauto; eauto.
iExists _, _; iPureIntro.
f_equal. rewrite /stored_cont HeqO; trivial. }
iIntros (v).
iDestruct 1 as (resid) "[Heq Hj]"; iDestruct "Heq" as %-> ;simpl.
iApply (wp_LetIn (SeqCtx _ :: LetInCtx _ :: _));
eauto.
iNext. simpl. asimpl.
rewrite HeqHwriters.
iPoseProof ("Hwriters" $! (InjLV (#nv resid), InjLV (#nv resid)) with "[]")
as "Hrw"; first by iLeft; iExists (_, _); eauto.
rewrite -HeqHwriters.
simpl.
iApply ("Hrw" $! (SeqCtx _ :: LetInCtx _ :: _, SeqCtx _ :: _)); iFrame.
simpl.
iAlways. clear j.
iIntros ([w1 w2] j) "[Hj %]"; simpl.
iApply (wp_Seq (LetInCtx _ :: _)); eauto. simpl. iNext.
iApply (wp_throw (LetInCtx _ :: _)); eauto. simpl. iNext.
iMod (step_Seq with "[$Hj]") as "Hj"; eauto. simpl.
iMod (step_throw with "[$Hj]") as "Hj"; eauto.
iApply wp_value; eauto. iExists UnitV; eauto.
Qed.
End server_refinement.
Require Import LogrelCC.soundness_binary.
Instance subG_lockΣ {Σ} : subG (GFunctor (exclR unitR)) Σ → inG Σ (exclR unitR).
Proof. solve_inG. Qed.
Theorem server_ctx_refinement_cont_nat :
[] ⊨ server_cont ≤ctx≤ server_nat :
TArrow (TProd (TArrow TUnit (TProd (TSum TNat TUnit) TNat))
(TArrow (TSum TNat TNat) TUnit)) TUnit.
Proof.
set (Σ := #[invΣ ; gen_heapΣ loc val ;
GFunctor (authR cfgUR) ; GFunctor (exclR unitR) ]).
set (HG := soundness_unary.HeapPreIG Σ _ _).
eapply (binary_soundness Σ); auto.
intros. eapply server_refinement_cont_nat.
Qed.