RunST.basic_equations
From RunST Require Import IC base lang rules typing reduction regions
contextual_refinement saved_pred ST_Lang_reduction logrel_shared soundness
logrel fundamental future.
From iris.proofmode Require Import tactics.
Set Bullet Behavior "Strict Subproofs".
Hint Resolve to_of_val.
Lemma reduction_on_right `{ICG_ST Σ} `{LogRelG Σ} Δ e e' τ σ1' σ2' w' :
rtc (pstep' det_lang) (e', σ1') (of_val w', σ2') →
⟦ τ ⟧ₑ Δ (e, e') ⊢ ⟦ τ ⟧ₑ Δ (e, of_val w').
Proof.
iIntros (Hrd) "Hee". apply rtc_nsteps in Hrd; destruct Hrd as [n Hrd].
iIntros (γh γh' σ1'') "Hσ1''".
iMod (allocate_full_state σ1') as (γh'') "[Hfσ1' Hoσ1']".
iSpecialize ("Hee" $! γh γh'' σ1' with "[$Hfσ1']").
iApply (@ic_wand_r); iSplitL "Hee"; first iApply "Hee". simpl.
iIntros (v _); iDestruct 1 as (σ3' v') "(Hrd & HPh & Hvv)".
iDestruct "Hrd" as %Hrd'.
apply rtc_nsteps in Hrd'. destruct Hrd' as [n' Hrd'].
destruct (nsteps_deterministic' Hrd Hrd') as [? [? ?]]; subst.
by iExists _, _; iSplit; iFrame.
Qed.
Lemma Neutrality_LR `{ICG_ST Σ} `{LogRelG Σ} Γ e e' :
Γ ⊨ e ≤log≤ e' : TUnit → Γ ⊨ e ≤log≤ (Lit LitUnit) : TUnit.
Proof.
iIntros (Hee Δ vvs HΔ) "[#Hinv #Hvvs]".
iIntros (γh γh' σ1') "Hσ1'". asimpl.
iPoseProof (Hee Δ vvs HΔ with "[]") as "Hee"; first iFrame "#".
(* We have to allocate a new heap! *)
iApply fupd_ic.
iMod (allocate_full_state σ1') as (γh4) "[Hnfσ1 Hnoσ1]".
iSpecialize ("Hee" $! _ _ _ with "[$Hnfσ1]").
iApply ic_wand_r; iSplitR "Hσ1'"; first iApply "Hee". asimpl.
iIntros (v _) "H". iDestruct "H" as (σ2' v') "(Hrd & How & #[Hvv Hvv'])".
iDestruct "Hrd" as %Hrd.
apply rtc_nsteps in Hrd; destruct Hrd as [n Hrd].
iExists σ1', v'; iSplit; iFrame; auto. iDestruct "Hvv'" as %Hv'.
iPureIntro. rewrite Hv' /=. eapply rtc_reflexive.
Qed.
Theorem Neutrality Γ e :
Γ ⊢ₜ e : TUnit → Γ ⊨ e ≤ctx≤ (Lit LitUnit) : TUnit.
Proof.
intros Ht.
apply binary_soundness; eauto.
- constructor.
- intros; eapply Neutrality_LR; eauto; by apply binary_fundamental.
Qed.
Lemma Commutativity_LR1 `{ICG_ST Σ} `{LogRelG Σ} Γ e1 e1' τ e2 e2' τ' :
Γ ⊨ e1 ≤log≤ e1' : τ → Γ ⊨ e2 ≤log≤ e2' : τ' →
Γ ⊨ (App (Rec (Pair e1.[ren (+2)] (Var 1))) e2) ≤log≤ (Pair e1' e2')
: (TProd τ τ').
Proof.
iIntros (Hee1 Hee2 Δ vvs HΔ) "[#Hinv #Hvvs]". asimpl.
iIntros (γh γh' σ1') "Hσ'". asimpl.
rewrite ic_eq /ic_def /=. iIntros (σ1 σ2 w n) "[H Hσ]".
iDestruct "H" as %Hrd1.
apply (nsteps_bind (fill_item (AppRCtx (RecV _)))) in Hrd1.
destruct Hrd1 as (k & σ3 & u & Hk & Hrd11 & Hrd12). simpl in *.
apply rec_red_step in Hrd12. destruct Hrd12 as (m & Hm & Hrd12). asimpl in *.
apply (nsteps_bind (fill_item (PairLCtx _))) in Hrd12.
destruct Hrd12 as (k' & σ4 & v & Hk' & Hrd121 & Hrd122). simpl in *.
apply (nsteps_val _ _ (PairV _ _)) in Hrd122.
destruct Hrd122 as (Heq31 & Heq32%of_val_inj & Heq33); subst.
set (Q := (|={⊤}=> Nat.iter k' (λ Q, ▷ Q) (◇ (∃ (σ2' : state) (v' : val),
■ (rtc (pstep' det_lang)
(e1'.[env_subst (vvs.*2)], σ1') (of_val v', σ2')))))%I : iProp Σ).
iAssert Q as "HQ".
{ unfold Q.
iMod (allocate_full_state σ3) as (γh3) "[Hfσ3 Hnσ3]".
iMod (allocate_full_state σ1') as (γh4) "[Hfσ1' Hnσ1']".
iPoseProof (Hee1 Δ vvs HΔ with "[]") as "Hee1"; first iFrame "#".
iSpecialize ("Hee1" $! _ _ _ with "[$Hfσ1']"). asimpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hee1" $! _ _ _ _ with "[$Hfσ3]"); auto.
iApply (future_plain with "[-]").
iMod "Hee1" as "[Hee1 _]". iModIntro.
iDestruct "Hee1" as (σ3' ?) "(Hrd2&Hσ3'&?)".
iDestruct "Hrd2" as %Hrd2.
iExists _, _; iPureIntro; repeat split; eauto. }
iMod "HQ" as "HQ". iModIntro. clear Q.
rewrite later_n_except_0_exist. iDestruct "HQ" as (σ3') "HQ".
rewrite later_n_except_0_exist. iDestruct "HQ" as (v') "HQ".
iMod (allocate_full_state σ3') as (γh3) "[Hfσ3' Hnσ3']".
iPoseProof (Hee2 Δ vvs HΔ with "[]") as "Hee2"; first iFrame "#".
iSpecialize ("Hee2" $! _ _ _ with "[$Hfσ3']"). asimpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hee2" $! _ _ _ _ with "[$Hσ]"); auto.
iMod "Hee2" as "[Hee2 Hσ]".
iDestruct "Hee2" as (σ4' v2') "(Hrd2 & Hσ3' & #Hvv)".
iDestruct "Hrd2" as %Hrd2.
apply rtc_nsteps in Hrd2; destruct Hrd2 as [k'' Hrd2].
iPoseProof (Hee1 Δ vvs HΔ with "[]") as "Hee1"; first iFrame "#".
iSpecialize ("Hee1" $! _ _ _ with "[$Hσ']"). asimpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hee1" $! _ _ _ _ with "[$Hσ]"); auto.
iDestruct (later_n_except_0_future with "[$HQ]") as "HQ".
iCombine "HQ" "Hee1" as "Hee1".
iMod "Hee1" as "(HQ & Hee1 & Hσ)".
iDestruct "Hee1" as (σ2' u') "(Hrd3 & Hσ2' & #Huu)".
iDestruct "HQ" as %HQrd.
iDestruct "Hrd3" as %Hrd3.
apply rtc_nsteps in HQrd; destruct HQrd as [s HQrd].
apply rtc_nsteps in Hrd3; destruct Hrd3 as [s' Hrd3].
destruct (nsteps_deterministic' HQrd Hrd3) as [? [? ?]]; subst.
iDestruct (heap_included with "[$Hnσ3' $Hσ3']") as %[Hinc Hvl].
iMod (heap_catch_up with "[$Hσ2']") as "Hfσ3'"; eauto.
iModIntro; iFrame. iExists _, (PairV _ _); iFrame.
iSplit.
{ iPureIntro. etrans.
- eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (PairLCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
- eapply (nsteps_rtc _). simpl.
eapply (@nsteps_ctx det_lang (fill_item (PairRCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
iExists (_, _), (_, _); repeat iSplit; simpl; eauto.
Qed.
Lemma Commutativity_LR2 `{ICG_ST Σ} `{LogRelG Σ} Γ e1 e1' τ e2 e2' τ' :
Γ ⊨ e1 ≤log≤ e1' : τ → Γ ⊨ e2 ≤log≤ e2' : τ' →
Γ ⊨ (Pair e1 e2) ≤log≤ (App (Rec (Pair e1'.[ren (+2)] (Var 1))) e2')
: (TProd τ τ').
Proof.
iIntros (Hee1 Hee2 Δ vvs HΔ) "[#Hinv #Hvvs]". asimpl.
iIntros (γh γh' σ1') "Hσ'". asimpl.
rewrite ic_eq /ic_def /=. iIntros (σ1 σ2 w n) "[H Hσ]".
iDestruct "H" as %Hrd1.
apply (nsteps_bind (fill_item (PairLCtx _))) in Hrd1.
destruct Hrd1 as (k & σ3 & v1 & Hk & Hrd11 & Hrd12). simpl in *.
apply (nsteps_bind (fill_item (PairRCtx _))) in Hrd12.
destruct Hrd12 as (k' & σ4 & v2 & Hk' & Hrd121 & Hrd122). simpl in *.
apply (nsteps_val _ _ (PairV _ _)) in Hrd122.
destruct Hrd122 as (Heq31 & Heq32%of_val_inj & Heq33); subst.
set (Q := (|={⊤}=> Nat.iter k' (λ Q, ▷ Q) (◇ (∃ (σ2' : state) (v' : val),
■ (rtc (pstep' det_lang)
(e2'.[env_subst (vvs.*2)], σ1') (of_val v', σ2')))))%I : iProp Σ).
iAssert Q as "HQ".
{ unfold Q.
iMod (allocate_full_state σ3) as (γh3) "[Hfσ3 Hnσ3]".
iMod (allocate_full_state σ1') as (γh4) "[Hfσ1' Hnσ1']".
iPoseProof (Hee2 Δ vvs HΔ with "[]") as "Hee2"; first iFrame "#".
iSpecialize ("Hee2" $! _ _ _ with "[$Hfσ1']"). asimpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hee2" $! _ _ _ _ with "[$Hfσ3]"); auto.
iApply (future_plain with "[-]").
iMod "Hee2" as "[Hee2 _]". iModIntro.
iDestruct "Hee2" as (σ3' ?) "(Hrd2&Hσ3'&?)".
iDestruct "Hrd2" as %Hrd2.
iExists _, _; iPureIntro; repeat split; eauto. }
iMod "HQ" as "HQ". iModIntro. clear Q.
rewrite later_n_except_0_exist. iDestruct "HQ" as (σ3') "HQ".
rewrite later_n_except_0_exist. iDestruct "HQ" as (v') "HQ".
iMod (allocate_full_state σ3') as (γh3) "[Hfσ3' Hnσ3']".
iPoseProof (Hee1 Δ vvs HΔ with "[]") as "Hee1"; first iFrame "#".
iSpecialize ("Hee1" $! _ _ _ with "[$Hfσ3']"). asimpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hee1" $! _ _ _ _ with "[$Hσ]"); auto.
iMod "Hee1" as "[Hee1 Hσ]".
iDestruct "Hee1" as (σ4' v2') "(Hrd2 & Hσ3' & #Hvv)".
iDestruct "Hrd2" as %Hrd2.
apply rtc_nsteps in Hrd2; destruct Hrd2 as [k'' Hrd2].
iPoseProof (Hee2 Δ vvs HΔ with "[]") as "Hee2"; first iFrame "#".
iSpecialize ("Hee2" $! _ _ _ with "[$Hσ']"). asimpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hee2" $! _ _ _ _ with "[$Hσ]"); auto.
iDestruct (later_n_except_0_future with "[$HQ]") as "HQ".
iCombine "HQ" "Hee2" as "Hee2".
iMod "Hee2" as "(HQ & Hee2 & Hσ)".
iDestruct "Hee2" as (σ2' u') "(Hrd3 & Hσ2' & #Huu)".
iDestruct "HQ" as %HQrd.
iDestruct "Hrd3" as %Hrd3.
apply rtc_nsteps in HQrd; destruct HQrd as [s HQrd].
apply rtc_nsteps in Hrd3; destruct Hrd3 as [s' Hrd3].
destruct (nsteps_deterministic' HQrd Hrd3) as [? [? ?]]; subst.
iDestruct (heap_included with "[$Hnσ3' $Hσ3']") as %[Hinc Hvl].
iMod (heap_catch_up with "[$Hσ2']") as "Hfσ3'"; eauto.
iModIntro; iFrame. iExists _, (PairV _ _); iFrame.
iSplit.
{ iPureIntro. etrans.
- eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (AppRCtx (RecV _))));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
- eapply (rtc_l _ (_, _) (_, _)); simpl.
{ apply head_prim_step. econstructor; eauto. }
asimpl. eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (PairLCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
iExists (_, _), (_, _); repeat iSplit; simpl; eauto.
Qed.
Theorem Commutativity Γ e1 τ e2 τ' :
Γ ⊢ₜ e1 : τ → Γ ⊢ₜ e2 : τ' →
Γ ⊨ (App (Rec (Pair e1.[ren (+2)] (Var 1))) e2) ≈ctx≈ (Pair e1 e2)
: (TProd τ τ').
Proof.
split.
- apply binary_soundness; eauto.
+ repeat econstructor; eauto.
by apply (context_weakening [ _ ; _]).
+ repeat econstructor; eauto.
+ intros; eapply Commutativity_LR1; eauto; by apply binary_fundamental.
- apply binary_soundness; eauto.
+ repeat econstructor; eauto.
+ repeat econstructor; eauto.
by apply (context_weakening [ _ ; _]).
+ intros; eapply Commutativity_LR2; by apply binary_fundamental.
Qed.
Lemma Commutativity_alt_LR `{ICG_ST Σ} `{LogRelG Σ} Γ e1 e1' τ e2 e2' τ' e τ'' :
Γ ⊨ e1 ≤log≤ e1' : τ → Γ ⊨ e2 ≤log≤ e2' : τ' →
τ :: τ' :: Γ ⊢ₜ e : τ'' →
Γ ⊨ App (App (Rec (Rec e.[up (ren (+1))].[ren (+1)])) e2) e1 ≤log≤
App (App
(Rec (Rec e.[ren (1 .: 0 .: (+2))]
.[up (ren (+1))].[ren (+1)]))
e1') e2'
: τ''.
Proof.
iIntros (Hee1 Hee2 Hte Δ vvs HΔ) "[#Hinv #Hvvs]". asimpl.
iIntros (γh γh' σ1') "Hσ'". asimpl.
rewrite ic_eq /ic_def /=. iIntros (σ1 σ2 w n) "[H Hσ]".
iDestruct "H" as %Hrd1.
apply (nsteps_bind (fill_item (AppLCtx _))) in Hrd1.
destruct Hrd1 as (k & σ3 & u & Hk & Hrd11 & Hrd12). asimpl in *.
apply (nsteps_bind (fill_item (AppRCtx (RecV _)))) in Hrd11.
destruct Hrd11 as (k1 & σ4 & u2 & Hk1 & Hrd111 & Hrd112). asimpl in *.
apply (nsteps_bind (fill_item (AppRCtx _))) in Hrd12.
destruct Hrd12 as (k2 & σ5 & u1 & Hk2 & Hrd121 & Hrd122). asimpl in *.
apply rec_red_step in Hrd112. destruct Hrd112 as (m & Hm & Hrd112).
asimpl in *.
apply (nsteps_val _ _ (RecV _)) in Hrd112.
destruct Hrd112 as (Heq31 & Heq32%of_val_inj & Heq33); subst.
asimpl in *.
apply rec_red_step in Hrd122. destruct Hrd122 as (m' & Hm' & Hrd122).
asimpl in Hrd122. unfold upn in Hrd122. asimpl.
cbn [scons] in Hrd122. asimpl in Hrd122.
set (Q := (|={⊤}=> Nat.iter k2 (λ Q, ▷ Q) (◇ (∃ (σ2' : state) (v' : val),
■ (rtc (pstep' det_lang)
(e1'.[env_subst (vvs.*2)], σ1') (of_val v', σ2')))))%I : iProp Σ).
iAssert Q as "HQ".
{ unfold Q.
iMod (allocate_full_state σ3) as (γh3) "[Hfσ3 Hnσ3]".
iMod (allocate_full_state σ1') as (γh4) "[Hfσ1' Hnσ1']".
iPoseProof (Hee1 Δ vvs HΔ with "[]") as "Hee1"; first iFrame "#".
iSpecialize ("Hee1" $! _ _ _ with "[$Hfσ1']"). asimpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hee1" $! _ _ _ _ with "[$Hfσ3]"); auto.
iApply (future_plain with "[-]").
iMod "Hee1" as "[Hee1 _]". iModIntro.
iDestruct "Hee1" as (σ3' ?) "(Hrd2&Hσ3'&?)".
iDestruct "Hrd2" as %Hrd2.
iExists _, _; iPureIntro; repeat split; eauto. }
iMod "HQ" as "HQ". iModIntro. clear Q.
rewrite later_n_except_0_exist. iDestruct "HQ" as (σ3') "HQ".
rewrite later_n_except_0_exist. iDestruct "HQ" as (v') "HQ".
iMod (allocate_full_state σ3') as (γh3) "[Hfσ3' Hnσ3']".
iPoseProof (Hee2 Δ vvs HΔ with "[]") as "Hee2"; first iFrame "#".
iSpecialize ("Hee2" $! _ _ _ with "[$Hfσ3']"). asimpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hee2" $! _ _ _ _ with "[$Hσ]"); auto.
iMod "Hee2" as "[Hee2 Hσ]".
iDestruct "Hee2" as (σ4' v2') "(Hrd2 & Hσ3' & #Hvv)".
iDestruct "Hrd2" as %Hrd2.
apply rtc_nsteps in Hrd2; destruct Hrd2 as [k'' Hrd2].
iPoseProof (Hee1 Δ vvs HΔ with "[]") as "Hee1"; first iFrame "#".
iSpecialize ("Hee1" $! _ _ _ with "[$Hσ']"). asimpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hee1" $! _ _ _ _ with "[$Hσ]"); auto.
iDestruct (later_n_except_0_future with "[$HQ]") as "HQ".
iCombine "HQ" "Hee1" as "Hee1".
iMod "Hee1" as "(HQ & Hee1 & Hσ)".
iDestruct "Hee1" as (σ2' u') "(Hrd3 & Hσ2' & #Huu)".
iDestruct "HQ" as %HQrd.
iDestruct "Hrd3" as %Hrd3.
apply rtc_nsteps in HQrd; destruct HQrd as [s HQrd].
apply rtc_nsteps in Hrd3; destruct Hrd3 as [s' Hrd3].
destruct (nsteps_deterministic' HQrd Hrd3) as [? [? ?]]; subst.
iDestruct (heap_included with "[$Hnσ3' $Hσ3']") as %[Hinc Hvl].
iMod (heap_catch_up with "[$Hσ2']") as "Hfσ3'"; eauto.
iDestruct (interp_env_length with "[]") as %Hvvsl; first iFrame "#".
iPoseProof (binary_fundamental _ _ _ Hte _ ((u1, u') :: (u2, v2') :: vvs)
with "[]") as "Hrl";
first by rewrite ?interp_env_cons; iFrame "#".
cbn [fmap list_fmap].
erewrite <- typed_subst_head_simpl_2; eauto;
last by rewrite /= fmap_length Hvvsl.
erewrite <- typed_subst_head_simpl_2; eauto;
last by rewrite /= fmap_length Hvvsl.
simpl.
iSpecialize ("Hrl" $! _ _ _ with "[$Hfσ3']"). simpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hrl" $! _ _ _ _ with "[$Hσ]"); auto.
iMod "Hrl" as "[Hrl Hσ]".
iDestruct "Hrl" as (σ5' v3') "(Hrd4 & Hσ5' & #Hww)".
iDestruct "Hrd4" as %Hrd4.
apply rtc_nsteps in Hrd4; destruct Hrd4 as [k4 Hrd4].
iModIntro; iFrame. iExists _, _; iFrame; iFrame "#".
iPureIntro. etrans.
{ eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (AppLCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
eapply (@nsteps_ctx det_lang (fill_item (AppRCtx (RecV _))));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
etrans.
{ asimpl. eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (AppLCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
eapply (nsteps_l _ _ (_, _) (_, _)); simpl.
{ apply head_prim_step. econstructor; eauto. }
asimpl. econstructor. }
simpl.
etrans.
{ eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (AppRCtx (RecV _))));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
simpl.
etrans.
{ eapply (nsteps_rtc _). eapply (nsteps_l _ _ (_, _) (_, _)); simpl.
{ apply head_prim_step. econstructor; eauto. }
asimpl; eauto. }
econstructor.
Qed.
Require Import Coq.Logic.FunctionalExtensionality.
Lemma lookup_swap_2 : (lookup_n (swap_list 1 1)) = (1 .: 0 .: (+2)).
Proof.
extensionality i; destruct i as [|[|?]]; by cbv.
Qed.
Theorem Commutativity_alt Γ e1 τ e2 τ' e τ'' :
Γ ⊢ₜ e1 : τ → Γ ⊢ₜ e2 : τ' →
τ :: τ' :: Γ ⊢ₜ e : τ'' →
Γ ⊨ App (App (Rec (Rec e.[up (ren (+1))].[ren (+1)])) e2) e1 ≤ctx≤
App (App
(Rec (Rec e.[ren (1 .: 0 .: (+2))]
.[up (ren (+1))].[ren (+1)]))
e1) e2
: τ''.
Proof.
intros Ht1 Ht2 Ht3.
apply binary_soundness; eauto.
- repeat econstructor; eauto.
apply (context_weakening [_]).
by apply (context_gen_weakening [_] [_]).
- repeat econstructor; eauto.
apply (context_weakening [_]).
apply (context_gen_weakening [_] [_]).
rewrite -lookup_swap_2.
by apply (context_swap_typed [] [_] [_]).
- intros; eapply Commutativity_alt_LR; eauto using binary_fundamental.
Qed.
Lemma Idempotency_LR1 `{ICG_ST Σ} `{LogRelG Σ} Γ e e' τ :
Γ ⊨ e ≤log≤ e' : τ →
Γ ⊨ (App (Rec (Pair (Var 1) (Var 1))) e) ≤log≤ (Pair e' e') : (TProd τ τ).
Proof.
iIntros (Hrel Δ vvs HΔ) "[#Hinv #Hvvs]". asimpl.
iIntros (γh γh' σ1') "Hσ1'". asimpl.
rewrite ic_eq /ic_def /=. iIntros (σ1 σ2 w n) "[H Hσ1]".
iDestruct "H" as %Hredv.
apply (nsteps_bind (fill_item (AppRCtx (RecV _)))) in Hredv.
destruct Hredv as (k & σ3 & u & Hk & Hredv1 & Hredv2). simpl in *.
apply rec_red_step in Hredv2.
destruct Hredv2 as (m & Hm & Hredv2). asimpl in *.
apply (nsteps_val _ _ (PairV _ _)) in Hredv2.
destruct Hredv2 as (Heq31 & Heq32%of_val_inj & Heq33); subst.
set (Q := (|={⊤}=> Nat.iter k (λ Q, ▷ Q) (◇ (∃ (σ2' : state) (v' : val),
■ (rtc (pstep' det_lang)
(e'.[env_subst (vvs.*2)], σ1') (of_val v', σ2')))))%I : iProp Σ).
iAssert Q as "HQ".
{ unfold Q.
iMod (allocate_full_state σ1) as (γh3) "[Hfσ1 Hnσ1]".
iMod (allocate_full_state σ1') as (γh4) "[Hfσ1' Hnσ1']".
iPoseProof (Hrel Δ vvs HΔ with "[]") as "Hrel"; first iFrame "#".
iSpecialize ("Hrel" $! _ _ _ with "[$Hfσ1']"). asimpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hrel" $! _ _ _ _ with "[$Hfσ1]"); auto.
iApply (future_plain with "[-]").
iMod "Hrel" as "[Hrel _]". iModIntro.
iDestruct "Hrel" as (σ3' ?) "(Hrd2&Hσ3'&?)".
iDestruct "Hrd2" as %Hrd2.
iExists _, _; iPureIntro; repeat split; eauto. }
iMod "HQ" as "HQ". iModIntro. clear Q.
rewrite later_n_except_0_exist. iDestruct "HQ" as (σ3') "HQ".
rewrite later_n_except_0_exist. iDestruct "HQ" as (v') "HQ".
iPoseProof (Hrel Δ vvs HΔ with "[]") as "Hrel"; first iFrame "#".
iSpecialize ("Hrel" $! _ _ _ with "[$Hσ1']"). asimpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hrel" $! _ _ _ _ with "[$Hσ1]"); auto.
iDestruct (later_n_except_0_future with "[$HQ]") as "HQ".
iCombine "HQ" "Hrel" as "Hrel".
iMod (allocate_full_state σ3') as (γh3) "[Hfσ3' Hnσ3']".
iMod (allocate_full_state σ1) as (γh4) "[Hfσ1 Hnσ1]".
iPoseProof (Hrel Δ vvs HΔ with "[]") as "Hrel'"; first iFrame "#".
iSpecialize ("Hrel'" $! _ _ _ with "[$Hfσ3']"). asimpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hrel'" $! _ _ _ _ with "[$Hfσ1]"); auto.
iCombine "Hrel" "Hrel'" as "Hrel'".
iMod "Hrel'" as "[[HQ [Hrel Hσ2]] [Hrel' Hfσ2]]".
iDestruct "HQ" as %HQrd.
iDestruct "Hrel" as (σ4' v2') "(Hrd2 & Hσ4' & #Hvv)".
iDestruct "Hrel'" as (σ5' v3') "(Hrd2' & Hfσ5' & #Hvv')".
iDestruct "Hrd2" as %Hrd2. iDestruct "Hrd2'" as %Hrd2'.
apply rtc_nsteps in HQrd; destruct HQrd as [s HQrd].
apply rtc_nsteps in Hrd2; destruct Hrd2 as [s' Hrd2].
apply rtc_nsteps in Hrd2'; destruct Hrd2' as [s'' Hrd2'].
destruct (nsteps_deterministic' HQrd Hrd2) as [? [? ?]]; subst.
iDestruct (heap_included with "[$Hnσ3' $Hfσ5']") as %[Hinc Hvl].
iMod (heap_catch_up with "[$Hσ4']") as "Hσ5'"; eauto.
iModIntro; iFrame.
iExists _, (PairV _ _); iFrame. iSplit.
{ iPureIntro. etrans.
- eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (PairLCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
- eapply (nsteps_rtc _). simpl.
eapply (@nsteps_ctx det_lang (fill_item (PairRCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
iExists (_, _), (_, _); repeat iSplit; simpl; eauto.
Qed.
Lemma Idempotency_LR2 `{ICG_ST Σ} `{LogRelG Σ} Γ e e' τ :
Γ ⊨ e ≤log≤ e' : τ →
Γ ⊨ (Pair e e) ≤log≤ (App (Rec (Pair (Var 1) (Var 1))) e') : (TProd τ τ).
Proof.
iIntros (Hrel Δ vvs HΔ) "[#Hinv #Hvvs]". asimpl.
iIntros (γ γ' σ1') "Hσ1'". asimpl.
rewrite ic_eq /ic_def /=.
iIntros (σ1 σ2 w n) "[H Hσ1]". iDestruct "H" as %Hredv.
apply (nsteps_bind (fill_item (PairLCtx _))) in Hredv. simpl in *.
destruct Hredv as (m & σ3 & v2 & Hmn & [Hredleft Hredpair]).
apply (nsteps_bind (fill_item (PairRCtx _))) in Hredpair. simpl in *.
destruct Hredpair as (z & 市σ3 & v1 & Hznm & [Hredright Hredval]).
change (Pair (of_val v2) (of_val v1)) with (of_val (PairV v2 v1)) in Hredval.
destruct (nsteps_val _ _ _ _ _ _ Hredval) as (Hnmz0 & Hvall & Hstate).
have X: z = n-m by omega. simpl in *. subst. move=> {Hredval Hnmz0 Hznm}.
have X: w = PairV v2 v1 by eauto using of_val_inj. subst =>{Hvall}.
iMod (allocate_full_state σ1') as (γ3) "[Hnfσ1' Hnoσ1']".
iPoseProof (Hrel with "[Hinv Hvvs]") as "Hrel". by iSplit; auto.
iSpecialize ("Hrel" $! _ _ _ with "Hnfσ1'").
rewrite ic_eq /ic_def. asimpl.
iSpecialize ("Hrel" $! _ _ _ _ with "[Hσ1]"); iFrame. done.
iMod "Hrel". iDestruct "Hrel" as "[Hrel Hσ3]".
iDestruct "Hrel" as (σ2' v') "(Hrede' & Hσ2' & Hrel')".
iDestruct "Hrede'" as %Hrede'.
iPoseProof (Hrel with "[Hinv Hvvs]") as "Hrel". by iSplit; auto.
iSpecialize ("Hrel" $! _ _ _ with "Hσ1'").
rewrite ic_eq /ic_def. asimpl.
iSpecialize ("Hrel" $! _ _ _ _ with "[Hσ3]"); iFrame. done.
iMod "Hrel". iDestruct "Hrel" as "[Hrel Hσ2]".
iDestruct "Hrel" as (σ2'' v'') "(Hrede'' & Hσ2'' & Hrel'')".
iDestruct "Hrede''" as %Hrede''.
iModIntro. iFrame.
iExists _, (PairV _ _). iFrame. iSplit.
- iPureIntro. simpl. etrans.
apply (@rtc_ctx det_lang (fill_item (AppRCtx (RecV _))));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
eapply (rtc_l _ (_, _) (_, _)); simpl.
{ apply head_prim_step. econstructor; eauto. }
econstructor.
- iExists (v2,_), (v1,_). iSplit; eauto.
apply rtc_nsteps in Hrede'; destruct Hrede' as [s Hrede'].
apply rtc_nsteps in Hrede''; destruct Hrede'' as [s' Hrede''].
destruct (nsteps_deterministic' Hrede' Hrede'') as [? [? ?]]; subst.
iFrame.
Qed.
Theorem Idempotency Γ e τ :
Γ ⊢ₜ e : τ →
Γ ⊨ (App (Rec (Pair (Var 1) (Var 1))) e) ≈ctx≈ (Pair e e) : (TProd τ τ).
Proof.
split.
- apply binary_soundness; eauto.
+ repeat econstructor; eauto.
+ repeat econstructor; eauto.
+ intros; eapply Idempotency_LR1; eauto; by apply binary_fundamental.
- apply binary_soundness; eauto.
+ repeat econstructor; eauto.
+ repeat econstructor; eauto.
+ intros; eapply Idempotency_LR2; by apply binary_fundamental.
Qed.
contextual_refinement saved_pred ST_Lang_reduction logrel_shared soundness
logrel fundamental future.
From iris.proofmode Require Import tactics.
Set Bullet Behavior "Strict Subproofs".
Hint Resolve to_of_val.
Lemma reduction_on_right `{ICG_ST Σ} `{LogRelG Σ} Δ e e' τ σ1' σ2' w' :
rtc (pstep' det_lang) (e', σ1') (of_val w', σ2') →
⟦ τ ⟧ₑ Δ (e, e') ⊢ ⟦ τ ⟧ₑ Δ (e, of_val w').
Proof.
iIntros (Hrd) "Hee". apply rtc_nsteps in Hrd; destruct Hrd as [n Hrd].
iIntros (γh γh' σ1'') "Hσ1''".
iMod (allocate_full_state σ1') as (γh'') "[Hfσ1' Hoσ1']".
iSpecialize ("Hee" $! γh γh'' σ1' with "[$Hfσ1']").
iApply (@ic_wand_r); iSplitL "Hee"; first iApply "Hee". simpl.
iIntros (v _); iDestruct 1 as (σ3' v') "(Hrd & HPh & Hvv)".
iDestruct "Hrd" as %Hrd'.
apply rtc_nsteps in Hrd'. destruct Hrd' as [n' Hrd'].
destruct (nsteps_deterministic' Hrd Hrd') as [? [? ?]]; subst.
by iExists _, _; iSplit; iFrame.
Qed.
Lemma Neutrality_LR `{ICG_ST Σ} `{LogRelG Σ} Γ e e' :
Γ ⊨ e ≤log≤ e' : TUnit → Γ ⊨ e ≤log≤ (Lit LitUnit) : TUnit.
Proof.
iIntros (Hee Δ vvs HΔ) "[#Hinv #Hvvs]".
iIntros (γh γh' σ1') "Hσ1'". asimpl.
iPoseProof (Hee Δ vvs HΔ with "[]") as "Hee"; first iFrame "#".
(* We have to allocate a new heap! *)
iApply fupd_ic.
iMod (allocate_full_state σ1') as (γh4) "[Hnfσ1 Hnoσ1]".
iSpecialize ("Hee" $! _ _ _ with "[$Hnfσ1]").
iApply ic_wand_r; iSplitR "Hσ1'"; first iApply "Hee". asimpl.
iIntros (v _) "H". iDestruct "H" as (σ2' v') "(Hrd & How & #[Hvv Hvv'])".
iDestruct "Hrd" as %Hrd.
apply rtc_nsteps in Hrd; destruct Hrd as [n Hrd].
iExists σ1', v'; iSplit; iFrame; auto. iDestruct "Hvv'" as %Hv'.
iPureIntro. rewrite Hv' /=. eapply rtc_reflexive.
Qed.
Theorem Neutrality Γ e :
Γ ⊢ₜ e : TUnit → Γ ⊨ e ≤ctx≤ (Lit LitUnit) : TUnit.
Proof.
intros Ht.
apply binary_soundness; eauto.
- constructor.
- intros; eapply Neutrality_LR; eauto; by apply binary_fundamental.
Qed.
Lemma Commutativity_LR1 `{ICG_ST Σ} `{LogRelG Σ} Γ e1 e1' τ e2 e2' τ' :
Γ ⊨ e1 ≤log≤ e1' : τ → Γ ⊨ e2 ≤log≤ e2' : τ' →
Γ ⊨ (App (Rec (Pair e1.[ren (+2)] (Var 1))) e2) ≤log≤ (Pair e1' e2')
: (TProd τ τ').
Proof.
iIntros (Hee1 Hee2 Δ vvs HΔ) "[#Hinv #Hvvs]". asimpl.
iIntros (γh γh' σ1') "Hσ'". asimpl.
rewrite ic_eq /ic_def /=. iIntros (σ1 σ2 w n) "[H Hσ]".
iDestruct "H" as %Hrd1.
apply (nsteps_bind (fill_item (AppRCtx (RecV _)))) in Hrd1.
destruct Hrd1 as (k & σ3 & u & Hk & Hrd11 & Hrd12). simpl in *.
apply rec_red_step in Hrd12. destruct Hrd12 as (m & Hm & Hrd12). asimpl in *.
apply (nsteps_bind (fill_item (PairLCtx _))) in Hrd12.
destruct Hrd12 as (k' & σ4 & v & Hk' & Hrd121 & Hrd122). simpl in *.
apply (nsteps_val _ _ (PairV _ _)) in Hrd122.
destruct Hrd122 as (Heq31 & Heq32%of_val_inj & Heq33); subst.
set (Q := (|={⊤}=> Nat.iter k' (λ Q, ▷ Q) (◇ (∃ (σ2' : state) (v' : val),
■ (rtc (pstep' det_lang)
(e1'.[env_subst (vvs.*2)], σ1') (of_val v', σ2')))))%I : iProp Σ).
iAssert Q as "HQ".
{ unfold Q.
iMod (allocate_full_state σ3) as (γh3) "[Hfσ3 Hnσ3]".
iMod (allocate_full_state σ1') as (γh4) "[Hfσ1' Hnσ1']".
iPoseProof (Hee1 Δ vvs HΔ with "[]") as "Hee1"; first iFrame "#".
iSpecialize ("Hee1" $! _ _ _ with "[$Hfσ1']"). asimpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hee1" $! _ _ _ _ with "[$Hfσ3]"); auto.
iApply (future_plain with "[-]").
iMod "Hee1" as "[Hee1 _]". iModIntro.
iDestruct "Hee1" as (σ3' ?) "(Hrd2&Hσ3'&?)".
iDestruct "Hrd2" as %Hrd2.
iExists _, _; iPureIntro; repeat split; eauto. }
iMod "HQ" as "HQ". iModIntro. clear Q.
rewrite later_n_except_0_exist. iDestruct "HQ" as (σ3') "HQ".
rewrite later_n_except_0_exist. iDestruct "HQ" as (v') "HQ".
iMod (allocate_full_state σ3') as (γh3) "[Hfσ3' Hnσ3']".
iPoseProof (Hee2 Δ vvs HΔ with "[]") as "Hee2"; first iFrame "#".
iSpecialize ("Hee2" $! _ _ _ with "[$Hfσ3']"). asimpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hee2" $! _ _ _ _ with "[$Hσ]"); auto.
iMod "Hee2" as "[Hee2 Hσ]".
iDestruct "Hee2" as (σ4' v2') "(Hrd2 & Hσ3' & #Hvv)".
iDestruct "Hrd2" as %Hrd2.
apply rtc_nsteps in Hrd2; destruct Hrd2 as [k'' Hrd2].
iPoseProof (Hee1 Δ vvs HΔ with "[]") as "Hee1"; first iFrame "#".
iSpecialize ("Hee1" $! _ _ _ with "[$Hσ']"). asimpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hee1" $! _ _ _ _ with "[$Hσ]"); auto.
iDestruct (later_n_except_0_future with "[$HQ]") as "HQ".
iCombine "HQ" "Hee1" as "Hee1".
iMod "Hee1" as "(HQ & Hee1 & Hσ)".
iDestruct "Hee1" as (σ2' u') "(Hrd3 & Hσ2' & #Huu)".
iDestruct "HQ" as %HQrd.
iDestruct "Hrd3" as %Hrd3.
apply rtc_nsteps in HQrd; destruct HQrd as [s HQrd].
apply rtc_nsteps in Hrd3; destruct Hrd3 as [s' Hrd3].
destruct (nsteps_deterministic' HQrd Hrd3) as [? [? ?]]; subst.
iDestruct (heap_included with "[$Hnσ3' $Hσ3']") as %[Hinc Hvl].
iMod (heap_catch_up with "[$Hσ2']") as "Hfσ3'"; eauto.
iModIntro; iFrame. iExists _, (PairV _ _); iFrame.
iSplit.
{ iPureIntro. etrans.
- eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (PairLCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
- eapply (nsteps_rtc _). simpl.
eapply (@nsteps_ctx det_lang (fill_item (PairRCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
iExists (_, _), (_, _); repeat iSplit; simpl; eauto.
Qed.
Lemma Commutativity_LR2 `{ICG_ST Σ} `{LogRelG Σ} Γ e1 e1' τ e2 e2' τ' :
Γ ⊨ e1 ≤log≤ e1' : τ → Γ ⊨ e2 ≤log≤ e2' : τ' →
Γ ⊨ (Pair e1 e2) ≤log≤ (App (Rec (Pair e1'.[ren (+2)] (Var 1))) e2')
: (TProd τ τ').
Proof.
iIntros (Hee1 Hee2 Δ vvs HΔ) "[#Hinv #Hvvs]". asimpl.
iIntros (γh γh' σ1') "Hσ'". asimpl.
rewrite ic_eq /ic_def /=. iIntros (σ1 σ2 w n) "[H Hσ]".
iDestruct "H" as %Hrd1.
apply (nsteps_bind (fill_item (PairLCtx _))) in Hrd1.
destruct Hrd1 as (k & σ3 & v1 & Hk & Hrd11 & Hrd12). simpl in *.
apply (nsteps_bind (fill_item (PairRCtx _))) in Hrd12.
destruct Hrd12 as (k' & σ4 & v2 & Hk' & Hrd121 & Hrd122). simpl in *.
apply (nsteps_val _ _ (PairV _ _)) in Hrd122.
destruct Hrd122 as (Heq31 & Heq32%of_val_inj & Heq33); subst.
set (Q := (|={⊤}=> Nat.iter k' (λ Q, ▷ Q) (◇ (∃ (σ2' : state) (v' : val),
■ (rtc (pstep' det_lang)
(e2'.[env_subst (vvs.*2)], σ1') (of_val v', σ2')))))%I : iProp Σ).
iAssert Q as "HQ".
{ unfold Q.
iMod (allocate_full_state σ3) as (γh3) "[Hfσ3 Hnσ3]".
iMod (allocate_full_state σ1') as (γh4) "[Hfσ1' Hnσ1']".
iPoseProof (Hee2 Δ vvs HΔ with "[]") as "Hee2"; first iFrame "#".
iSpecialize ("Hee2" $! _ _ _ with "[$Hfσ1']"). asimpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hee2" $! _ _ _ _ with "[$Hfσ3]"); auto.
iApply (future_plain with "[-]").
iMod "Hee2" as "[Hee2 _]". iModIntro.
iDestruct "Hee2" as (σ3' ?) "(Hrd2&Hσ3'&?)".
iDestruct "Hrd2" as %Hrd2.
iExists _, _; iPureIntro; repeat split; eauto. }
iMod "HQ" as "HQ". iModIntro. clear Q.
rewrite later_n_except_0_exist. iDestruct "HQ" as (σ3') "HQ".
rewrite later_n_except_0_exist. iDestruct "HQ" as (v') "HQ".
iMod (allocate_full_state σ3') as (γh3) "[Hfσ3' Hnσ3']".
iPoseProof (Hee1 Δ vvs HΔ with "[]") as "Hee1"; first iFrame "#".
iSpecialize ("Hee1" $! _ _ _ with "[$Hfσ3']"). asimpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hee1" $! _ _ _ _ with "[$Hσ]"); auto.
iMod "Hee1" as "[Hee1 Hσ]".
iDestruct "Hee1" as (σ4' v2') "(Hrd2 & Hσ3' & #Hvv)".
iDestruct "Hrd2" as %Hrd2.
apply rtc_nsteps in Hrd2; destruct Hrd2 as [k'' Hrd2].
iPoseProof (Hee2 Δ vvs HΔ with "[]") as "Hee2"; first iFrame "#".
iSpecialize ("Hee2" $! _ _ _ with "[$Hσ']"). asimpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hee2" $! _ _ _ _ with "[$Hσ]"); auto.
iDestruct (later_n_except_0_future with "[$HQ]") as "HQ".
iCombine "HQ" "Hee2" as "Hee2".
iMod "Hee2" as "(HQ & Hee2 & Hσ)".
iDestruct "Hee2" as (σ2' u') "(Hrd3 & Hσ2' & #Huu)".
iDestruct "HQ" as %HQrd.
iDestruct "Hrd3" as %Hrd3.
apply rtc_nsteps in HQrd; destruct HQrd as [s HQrd].
apply rtc_nsteps in Hrd3; destruct Hrd3 as [s' Hrd3].
destruct (nsteps_deterministic' HQrd Hrd3) as [? [? ?]]; subst.
iDestruct (heap_included with "[$Hnσ3' $Hσ3']") as %[Hinc Hvl].
iMod (heap_catch_up with "[$Hσ2']") as "Hfσ3'"; eauto.
iModIntro; iFrame. iExists _, (PairV _ _); iFrame.
iSplit.
{ iPureIntro. etrans.
- eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (AppRCtx (RecV _))));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
- eapply (rtc_l _ (_, _) (_, _)); simpl.
{ apply head_prim_step. econstructor; eauto. }
asimpl. eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (PairLCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
iExists (_, _), (_, _); repeat iSplit; simpl; eauto.
Qed.
Theorem Commutativity Γ e1 τ e2 τ' :
Γ ⊢ₜ e1 : τ → Γ ⊢ₜ e2 : τ' →
Γ ⊨ (App (Rec (Pair e1.[ren (+2)] (Var 1))) e2) ≈ctx≈ (Pair e1 e2)
: (TProd τ τ').
Proof.
split.
- apply binary_soundness; eauto.
+ repeat econstructor; eauto.
by apply (context_weakening [ _ ; _]).
+ repeat econstructor; eauto.
+ intros; eapply Commutativity_LR1; eauto; by apply binary_fundamental.
- apply binary_soundness; eauto.
+ repeat econstructor; eauto.
+ repeat econstructor; eauto.
by apply (context_weakening [ _ ; _]).
+ intros; eapply Commutativity_LR2; by apply binary_fundamental.
Qed.
Lemma Commutativity_alt_LR `{ICG_ST Σ} `{LogRelG Σ} Γ e1 e1' τ e2 e2' τ' e τ'' :
Γ ⊨ e1 ≤log≤ e1' : τ → Γ ⊨ e2 ≤log≤ e2' : τ' →
τ :: τ' :: Γ ⊢ₜ e : τ'' →
Γ ⊨ App (App (Rec (Rec e.[up (ren (+1))].[ren (+1)])) e2) e1 ≤log≤
App (App
(Rec (Rec e.[ren (1 .: 0 .: (+2))]
.[up (ren (+1))].[ren (+1)]))
e1') e2'
: τ''.
Proof.
iIntros (Hee1 Hee2 Hte Δ vvs HΔ) "[#Hinv #Hvvs]". asimpl.
iIntros (γh γh' σ1') "Hσ'". asimpl.
rewrite ic_eq /ic_def /=. iIntros (σ1 σ2 w n) "[H Hσ]".
iDestruct "H" as %Hrd1.
apply (nsteps_bind (fill_item (AppLCtx _))) in Hrd1.
destruct Hrd1 as (k & σ3 & u & Hk & Hrd11 & Hrd12). asimpl in *.
apply (nsteps_bind (fill_item (AppRCtx (RecV _)))) in Hrd11.
destruct Hrd11 as (k1 & σ4 & u2 & Hk1 & Hrd111 & Hrd112). asimpl in *.
apply (nsteps_bind (fill_item (AppRCtx _))) in Hrd12.
destruct Hrd12 as (k2 & σ5 & u1 & Hk2 & Hrd121 & Hrd122). asimpl in *.
apply rec_red_step in Hrd112. destruct Hrd112 as (m & Hm & Hrd112).
asimpl in *.
apply (nsteps_val _ _ (RecV _)) in Hrd112.
destruct Hrd112 as (Heq31 & Heq32%of_val_inj & Heq33); subst.
asimpl in *.
apply rec_red_step in Hrd122. destruct Hrd122 as (m' & Hm' & Hrd122).
asimpl in Hrd122. unfold upn in Hrd122. asimpl.
cbn [scons] in Hrd122. asimpl in Hrd122.
set (Q := (|={⊤}=> Nat.iter k2 (λ Q, ▷ Q) (◇ (∃ (σ2' : state) (v' : val),
■ (rtc (pstep' det_lang)
(e1'.[env_subst (vvs.*2)], σ1') (of_val v', σ2')))))%I : iProp Σ).
iAssert Q as "HQ".
{ unfold Q.
iMod (allocate_full_state σ3) as (γh3) "[Hfσ3 Hnσ3]".
iMod (allocate_full_state σ1') as (γh4) "[Hfσ1' Hnσ1']".
iPoseProof (Hee1 Δ vvs HΔ with "[]") as "Hee1"; first iFrame "#".
iSpecialize ("Hee1" $! _ _ _ with "[$Hfσ1']"). asimpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hee1" $! _ _ _ _ with "[$Hfσ3]"); auto.
iApply (future_plain with "[-]").
iMod "Hee1" as "[Hee1 _]". iModIntro.
iDestruct "Hee1" as (σ3' ?) "(Hrd2&Hσ3'&?)".
iDestruct "Hrd2" as %Hrd2.
iExists _, _; iPureIntro; repeat split; eauto. }
iMod "HQ" as "HQ". iModIntro. clear Q.
rewrite later_n_except_0_exist. iDestruct "HQ" as (σ3') "HQ".
rewrite later_n_except_0_exist. iDestruct "HQ" as (v') "HQ".
iMod (allocate_full_state σ3') as (γh3) "[Hfσ3' Hnσ3']".
iPoseProof (Hee2 Δ vvs HΔ with "[]") as "Hee2"; first iFrame "#".
iSpecialize ("Hee2" $! _ _ _ with "[$Hfσ3']"). asimpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hee2" $! _ _ _ _ with "[$Hσ]"); auto.
iMod "Hee2" as "[Hee2 Hσ]".
iDestruct "Hee2" as (σ4' v2') "(Hrd2 & Hσ3' & #Hvv)".
iDestruct "Hrd2" as %Hrd2.
apply rtc_nsteps in Hrd2; destruct Hrd2 as [k'' Hrd2].
iPoseProof (Hee1 Δ vvs HΔ with "[]") as "Hee1"; first iFrame "#".
iSpecialize ("Hee1" $! _ _ _ with "[$Hσ']"). asimpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hee1" $! _ _ _ _ with "[$Hσ]"); auto.
iDestruct (later_n_except_0_future with "[$HQ]") as "HQ".
iCombine "HQ" "Hee1" as "Hee1".
iMod "Hee1" as "(HQ & Hee1 & Hσ)".
iDestruct "Hee1" as (σ2' u') "(Hrd3 & Hσ2' & #Huu)".
iDestruct "HQ" as %HQrd.
iDestruct "Hrd3" as %Hrd3.
apply rtc_nsteps in HQrd; destruct HQrd as [s HQrd].
apply rtc_nsteps in Hrd3; destruct Hrd3 as [s' Hrd3].
destruct (nsteps_deterministic' HQrd Hrd3) as [? [? ?]]; subst.
iDestruct (heap_included with "[$Hnσ3' $Hσ3']") as %[Hinc Hvl].
iMod (heap_catch_up with "[$Hσ2']") as "Hfσ3'"; eauto.
iDestruct (interp_env_length with "[]") as %Hvvsl; first iFrame "#".
iPoseProof (binary_fundamental _ _ _ Hte _ ((u1, u') :: (u2, v2') :: vvs)
with "[]") as "Hrl";
first by rewrite ?interp_env_cons; iFrame "#".
cbn [fmap list_fmap].
erewrite <- typed_subst_head_simpl_2; eauto;
last by rewrite /= fmap_length Hvvsl.
erewrite <- typed_subst_head_simpl_2; eauto;
last by rewrite /= fmap_length Hvvsl.
simpl.
iSpecialize ("Hrl" $! _ _ _ with "[$Hfσ3']"). simpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hrl" $! _ _ _ _ with "[$Hσ]"); auto.
iMod "Hrl" as "[Hrl Hσ]".
iDestruct "Hrl" as (σ5' v3') "(Hrd4 & Hσ5' & #Hww)".
iDestruct "Hrd4" as %Hrd4.
apply rtc_nsteps in Hrd4; destruct Hrd4 as [k4 Hrd4].
iModIntro; iFrame. iExists _, _; iFrame; iFrame "#".
iPureIntro. etrans.
{ eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (AppLCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
eapply (@nsteps_ctx det_lang (fill_item (AppRCtx (RecV _))));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
etrans.
{ asimpl. eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (AppLCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
eapply (nsteps_l _ _ (_, _) (_, _)); simpl.
{ apply head_prim_step. econstructor; eauto. }
asimpl. econstructor. }
simpl.
etrans.
{ eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (AppRCtx (RecV _))));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
simpl.
etrans.
{ eapply (nsteps_rtc _). eapply (nsteps_l _ _ (_, _) (_, _)); simpl.
{ apply head_prim_step. econstructor; eauto. }
asimpl; eauto. }
econstructor.
Qed.
Require Import Coq.Logic.FunctionalExtensionality.
Lemma lookup_swap_2 : (lookup_n (swap_list 1 1)) = (1 .: 0 .: (+2)).
Proof.
extensionality i; destruct i as [|[|?]]; by cbv.
Qed.
Theorem Commutativity_alt Γ e1 τ e2 τ' e τ'' :
Γ ⊢ₜ e1 : τ → Γ ⊢ₜ e2 : τ' →
τ :: τ' :: Γ ⊢ₜ e : τ'' →
Γ ⊨ App (App (Rec (Rec e.[up (ren (+1))].[ren (+1)])) e2) e1 ≤ctx≤
App (App
(Rec (Rec e.[ren (1 .: 0 .: (+2))]
.[up (ren (+1))].[ren (+1)]))
e1) e2
: τ''.
Proof.
intros Ht1 Ht2 Ht3.
apply binary_soundness; eauto.
- repeat econstructor; eauto.
apply (context_weakening [_]).
by apply (context_gen_weakening [_] [_]).
- repeat econstructor; eauto.
apply (context_weakening [_]).
apply (context_gen_weakening [_] [_]).
rewrite -lookup_swap_2.
by apply (context_swap_typed [] [_] [_]).
- intros; eapply Commutativity_alt_LR; eauto using binary_fundamental.
Qed.
Lemma Idempotency_LR1 `{ICG_ST Σ} `{LogRelG Σ} Γ e e' τ :
Γ ⊨ e ≤log≤ e' : τ →
Γ ⊨ (App (Rec (Pair (Var 1) (Var 1))) e) ≤log≤ (Pair e' e') : (TProd τ τ).
Proof.
iIntros (Hrel Δ vvs HΔ) "[#Hinv #Hvvs]". asimpl.
iIntros (γh γh' σ1') "Hσ1'". asimpl.
rewrite ic_eq /ic_def /=. iIntros (σ1 σ2 w n) "[H Hσ1]".
iDestruct "H" as %Hredv.
apply (nsteps_bind (fill_item (AppRCtx (RecV _)))) in Hredv.
destruct Hredv as (k & σ3 & u & Hk & Hredv1 & Hredv2). simpl in *.
apply rec_red_step in Hredv2.
destruct Hredv2 as (m & Hm & Hredv2). asimpl in *.
apply (nsteps_val _ _ (PairV _ _)) in Hredv2.
destruct Hredv2 as (Heq31 & Heq32%of_val_inj & Heq33); subst.
set (Q := (|={⊤}=> Nat.iter k (λ Q, ▷ Q) (◇ (∃ (σ2' : state) (v' : val),
■ (rtc (pstep' det_lang)
(e'.[env_subst (vvs.*2)], σ1') (of_val v', σ2')))))%I : iProp Σ).
iAssert Q as "HQ".
{ unfold Q.
iMod (allocate_full_state σ1) as (γh3) "[Hfσ1 Hnσ1]".
iMod (allocate_full_state σ1') as (γh4) "[Hfσ1' Hnσ1']".
iPoseProof (Hrel Δ vvs HΔ with "[]") as "Hrel"; first iFrame "#".
iSpecialize ("Hrel" $! _ _ _ with "[$Hfσ1']"). asimpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hrel" $! _ _ _ _ with "[$Hfσ1]"); auto.
iApply (future_plain with "[-]").
iMod "Hrel" as "[Hrel _]". iModIntro.
iDestruct "Hrel" as (σ3' ?) "(Hrd2&Hσ3'&?)".
iDestruct "Hrd2" as %Hrd2.
iExists _, _; iPureIntro; repeat split; eauto. }
iMod "HQ" as "HQ". iModIntro. clear Q.
rewrite later_n_except_0_exist. iDestruct "HQ" as (σ3') "HQ".
rewrite later_n_except_0_exist. iDestruct "HQ" as (v') "HQ".
iPoseProof (Hrel Δ vvs HΔ with "[]") as "Hrel"; first iFrame "#".
iSpecialize ("Hrel" $! _ _ _ with "[$Hσ1']"). asimpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hrel" $! _ _ _ _ with "[$Hσ1]"); auto.
iDestruct (later_n_except_0_future with "[$HQ]") as "HQ".
iCombine "HQ" "Hrel" as "Hrel".
iMod (allocate_full_state σ3') as (γh3) "[Hfσ3' Hnσ3']".
iMod (allocate_full_state σ1) as (γh4) "[Hfσ1 Hnσ1]".
iPoseProof (Hrel Δ vvs HΔ with "[]") as "Hrel'"; first iFrame "#".
iSpecialize ("Hrel'" $! _ _ _ with "[$Hfσ3']"). asimpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hrel'" $! _ _ _ _ with "[$Hfσ1]"); auto.
iCombine "Hrel" "Hrel'" as "Hrel'".
iMod "Hrel'" as "[[HQ [Hrel Hσ2]] [Hrel' Hfσ2]]".
iDestruct "HQ" as %HQrd.
iDestruct "Hrel" as (σ4' v2') "(Hrd2 & Hσ4' & #Hvv)".
iDestruct "Hrel'" as (σ5' v3') "(Hrd2' & Hfσ5' & #Hvv')".
iDestruct "Hrd2" as %Hrd2. iDestruct "Hrd2'" as %Hrd2'.
apply rtc_nsteps in HQrd; destruct HQrd as [s HQrd].
apply rtc_nsteps in Hrd2; destruct Hrd2 as [s' Hrd2].
apply rtc_nsteps in Hrd2'; destruct Hrd2' as [s'' Hrd2'].
destruct (nsteps_deterministic' HQrd Hrd2) as [? [? ?]]; subst.
iDestruct (heap_included with "[$Hnσ3' $Hfσ5']") as %[Hinc Hvl].
iMod (heap_catch_up with "[$Hσ4']") as "Hσ5'"; eauto.
iModIntro; iFrame.
iExists _, (PairV _ _); iFrame. iSplit.
{ iPureIntro. etrans.
- eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (PairLCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
- eapply (nsteps_rtc _). simpl.
eapply (@nsteps_ctx det_lang (fill_item (PairRCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
iExists (_, _), (_, _); repeat iSplit; simpl; eauto.
Qed.
Lemma Idempotency_LR2 `{ICG_ST Σ} `{LogRelG Σ} Γ e e' τ :
Γ ⊨ e ≤log≤ e' : τ →
Γ ⊨ (Pair e e) ≤log≤ (App (Rec (Pair (Var 1) (Var 1))) e') : (TProd τ τ).
Proof.
iIntros (Hrel Δ vvs HΔ) "[#Hinv #Hvvs]". asimpl.
iIntros (γ γ' σ1') "Hσ1'". asimpl.
rewrite ic_eq /ic_def /=.
iIntros (σ1 σ2 w n) "[H Hσ1]". iDestruct "H" as %Hredv.
apply (nsteps_bind (fill_item (PairLCtx _))) in Hredv. simpl in *.
destruct Hredv as (m & σ3 & v2 & Hmn & [Hredleft Hredpair]).
apply (nsteps_bind (fill_item (PairRCtx _))) in Hredpair. simpl in *.
destruct Hredpair as (z & 市σ3 & v1 & Hznm & [Hredright Hredval]).
change (Pair (of_val v2) (of_val v1)) with (of_val (PairV v2 v1)) in Hredval.
destruct (nsteps_val _ _ _ _ _ _ Hredval) as (Hnmz0 & Hvall & Hstate).
have X: z = n-m by omega. simpl in *. subst. move=> {Hredval Hnmz0 Hznm}.
have X: w = PairV v2 v1 by eauto using of_val_inj. subst =>{Hvall}.
iMod (allocate_full_state σ1') as (γ3) "[Hnfσ1' Hnoσ1']".
iPoseProof (Hrel with "[Hinv Hvvs]") as "Hrel". by iSplit; auto.
iSpecialize ("Hrel" $! _ _ _ with "Hnfσ1'").
rewrite ic_eq /ic_def. asimpl.
iSpecialize ("Hrel" $! _ _ _ _ with "[Hσ1]"); iFrame. done.
iMod "Hrel". iDestruct "Hrel" as "[Hrel Hσ3]".
iDestruct "Hrel" as (σ2' v') "(Hrede' & Hσ2' & Hrel')".
iDestruct "Hrede'" as %Hrede'.
iPoseProof (Hrel with "[Hinv Hvvs]") as "Hrel". by iSplit; auto.
iSpecialize ("Hrel" $! _ _ _ with "Hσ1'").
rewrite ic_eq /ic_def. asimpl.
iSpecialize ("Hrel" $! _ _ _ _ with "[Hσ3]"); iFrame. done.
iMod "Hrel". iDestruct "Hrel" as "[Hrel Hσ2]".
iDestruct "Hrel" as (σ2'' v'') "(Hrede'' & Hσ2'' & Hrel'')".
iDestruct "Hrede''" as %Hrede''.
iModIntro. iFrame.
iExists _, (PairV _ _). iFrame. iSplit.
- iPureIntro. simpl. etrans.
apply (@rtc_ctx det_lang (fill_item (AppRCtx (RecV _))));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
eapply (rtc_l _ (_, _) (_, _)); simpl.
{ apply head_prim_step. econstructor; eauto. }
econstructor.
- iExists (v2,_), (v1,_). iSplit; eauto.
apply rtc_nsteps in Hrede'; destruct Hrede' as [s Hrede'].
apply rtc_nsteps in Hrede''; destruct Hrede'' as [s' Hrede''].
destruct (nsteps_deterministic' Hrede' Hrede'') as [? [? ?]]; subst.
iFrame.
Qed.
Theorem Idempotency Γ e τ :
Γ ⊢ₜ e : τ →
Γ ⊨ (App (Rec (Pair (Var 1) (Var 1))) e) ≈ctx≈ (Pair e e) : (TProd τ τ).
Proof.
split.
- apply binary_soundness; eauto.
+ repeat econstructor; eauto.
+ repeat econstructor; eauto.
+ intros; eapply Idempotency_LR1; eauto; by apply binary_fundamental.
- apply binary_soundness; eauto.
+ repeat econstructor; eauto.
+ repeat econstructor; eauto.
+ intros; eapply Idempotency_LR2; by apply binary_fundamental.
Qed.