RunST.tlam_eta
From RunST Require Import IC base lang rules typing reduction regions
future contextual_refinement saved_pred ST_Lang_reduction
logrel_shared soundness logrel fundamental
tlam_hoisting.
From iris.proofmode Require Import tactics.
Hint Resolve to_of_val.
Lemma tlam_eta_LR1 `{ICG_ST Σ} `{LogRelG Σ} Γ e e' τ :
Γ ⊨ e ≤log≤ e' : TForall τ →
Γ ⊨ e ≤log≤ App (Rec (TLam (TApp (Var 1)))) e'
: TForall τ.
Proof.
iIntros (HLR Δ vvs HΔ) "[#Hinv #Hvvs]".
iIntros (γh γh' σ1') "Hσ1' /=". asimpl.
rewrite ic_eq /ic_def /=. iIntros (σ1 σ2 v n) "[Hrd1 How]".
iDestruct "Hrd1" as %Hrd1.
iPoseProof (HLR with "[]") as "He"; iFrame "#".
iSpecialize ("He" $! _ _ _ with "[$Hσ1']"); simpl.
rewrite ic_eq /ic_def /=.
iSpecialize ("He" $! _ _ _ _ with "[$How]"); eauto.
iMod "He" as "[He HPh1]".
iDestruct "He" as (σ2'' v'' ) "(Hrd2 & Ho1 & #Hrl1)".
iDestruct "Hrd2" as %Hrd2.
iModIntro; iFrame.
apply rtc_nsteps in Hrd2. destruct Hrd2 as [k Hrd2].
iExists σ2'', (TLamV _); iSplit; iFrame.
{ 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. econstructor. }
asimpl. iAlways.
iIntros (τi rn Hwwp).
iSpecialize ("Hrl1" $! τi rn with "[]"); auto.
iIntros (γh3 γh4 σ3') "Hσ3' /=".
iSpecialize ("Hrl1" $! γh3 _ _ with "[$Hσ3']").
iApply ic_mono; last trivial. simpl.
iIntros (f _); iDestruct 1 as (σ3 u) "(Hrd3 & Hσ3 & #Hrl)".
iDestruct "Hrd3" as %Hrd3.
iExists _, _; iFrame; iFrame "#".
iPureIntro.
etrans; last apply Hrd3.
eapply (rtc_l _ (_, _) (_, _)); simpl.
{ apply head_prim_step. econstructor; eauto. }
asimpl. econstructor.
Qed.
Lemma tlam_eta_LR3 `{ICG_ST Σ} `{LogRelG Σ} Γ e' τ :
Γ ⊢ₜ e' : TForall τ →
Γ ⊨ TLam (App (Rec (TApp (Var 1))) e') ≤log≤ TLam (TApp e') : TForall τ.
Proof.
iIntros (Ht Δ vvs HΔ) "#(Hinv & Hvvs)".
iDestruct (interp_env_length with "[]") as %Hvvsl; iFrame "#".
iIntros (γh γh' σ1') "Hσ1' /=".
rewrite ic_eq /ic_def /=. iIntros (σ1 σ2 v n) "[Hrd1 How]".
iDestruct "Hrd1" as %Hrd1.
eapply (nsteps_val _ _ (TLamV _)) in Hrd1.
destruct Hrd1 as (Hneq & Hveq & Hσeq); subst.
iModIntro; iFrame.
iExists _, (TLamV _); iFrame; iSplit.
{ iPureIntro. econstructor. }
rewrite Hveq; asimpl.
iAlways. iIntros (τi rn Hτi).
iApply logrel_pure_det_head_step.
{ intros σ; econstructor; eauto. }
{ intros ? ? ? Hrd. inversion Hrd; subst; auto. }
{ intros σ. econstructor; eauto. }
iNext.
iApply (logrel_bind [AppRCtx (RecV _)] [TAppCtx]);
iSplitL; first by iApply (binary_fundamental _ _ _ Ht with "[]"); iFrame "#".
iIntros (VV) "#Hvv"; simpl.
iSpecialize ("Hvv" $! τi rn with "[]"); auto.
iIntros (γh3 γh4 σ3') "Hσ3' /=".
iApply ic_rec; auto. iNext.
iSpecialize ("Hvv" $! γh3 γh4 σ3' with "[$Hσ3']").
by asimpl.
Qed.
Lemma tlam_eta_CR1 Γ e e' τ :
Γ ⊢ₜ e : TForall τ →
Γ ⊢ₜ e' : TForall τ →
(∀ `{ICG_ST Σ} `{LogRelG Σ}, Γ ⊨ e ≤log≤ e' : TForall τ) →
Γ ⊨ e ≤ctx≤ App (Rec (TLam (TApp (Var 1)))) e'
: TForall τ.
Proof.
intros Ht Ht' HLR.
apply binary_soundness; auto.
- econstructor; eauto. econstructor. econstructor.
replace τ with (τ.[up (ren (+1))].[TVar 0/]) at 4 by by asimpl.
econstructor. by econstructor.
- intros. eapply tlam_eta_LR1; eauto.
Qed.
Lemma tlam_eta_CR2 Γ e' τ :
Γ ⊢ₜ e' : TForall τ →
Γ ⊨ App (Rec (TLam (TApp (Var 1)))) e'
≤ctx≤
TLam (App (Rec (TApp (Var 1))) e') : TForall τ.
Proof.
intros Ht'.
pose proof (TLam_Hoisting Γ e' (TForall τ) (TApp (Var 0)) τ) as Hrh.
asimpl in *.
eapply ctx_refines_trans; first apply Hrh; eauto.
{ replace τ with (τ.[up (ren (+1))].[TVar 0/]) at 2 by by asimpl.
econstructor. by econstructor; asimpl. }
assert (Γ ⊢ₜ TLam (App (Rec (TApp (ids 1))) e') : TForall τ) as Htp.
{ econstructor; last eauto.
econstructor; last by apply context_rename.
econstructor.
replace τ with (τ.[up (ren (+1))].[TVar 0/]) at 4 by by asimpl.
repeat econstructor. }
apply binary_soundness; eauto.
intros. by apply binary_fundamental.
Qed.
Lemma tlam_eta_CR3 Γ e' τ :
Γ ⊢ₜ e' : TForall τ →
Γ ⊨ TLam (App (Rec (TApp (Var 1))) e') ≤ctx≤ TLam (TApp e') : TForall τ.
Proof.
intros Ht.
apply binary_soundness; eauto.
- econstructor. econstructor; last by apply context_rename.
econstructor.
replace τ with (τ.[up (ren (+1))].[TVar 0/]) at 4 by by asimpl.
econstructor. by econstructor; asimpl.
- econstructor. replace τ with (τ.[up (ren (+1))].[TVar 0/]) by by asimpl.
econstructor.
change (TForall τ.[up (ren (+1))]) with (TForall τ).[ren (+1)].
by apply context_rename.
- intros. eapply tlam_eta_LR3; eauto.
Qed.
Lemma tlam_eta_basic Γ e e' τ :
Γ ⊢ₜ e : TForall τ → Γ ⊢ₜ e' : TForall τ →
(∀ `{ICG_ST Σ} `{LogRelG Σ}, Γ ⊨ e ≤log≤ e' : TForall τ) →
Γ ⊨ e ≤ctx≤ TLam (TApp e') : TForall τ.
Proof.
intros Ht1 Ht2 HLR.
eapply ctx_refines_trans.
eapply tlam_eta_CR1; eauto.
eapply ctx_refines_trans.
eapply tlam_eta_CR2; eauto.
eapply tlam_eta_CR3; eauto.
Qed.
Lemma tlam_eta Γ e τ :
Γ ⊢ₜ e : TForall τ → Γ ⊨ e ≤ctx≤ TLam (TApp e) : TForall τ.
Proof.
intros.
eapply tlam_eta_basic; eauto.
intros; apply binary_fundamental; auto.
Qed.
future contextual_refinement saved_pred ST_Lang_reduction
logrel_shared soundness logrel fundamental
tlam_hoisting.
From iris.proofmode Require Import tactics.
Hint Resolve to_of_val.
Lemma tlam_eta_LR1 `{ICG_ST Σ} `{LogRelG Σ} Γ e e' τ :
Γ ⊨ e ≤log≤ e' : TForall τ →
Γ ⊨ e ≤log≤ App (Rec (TLam (TApp (Var 1)))) e'
: TForall τ.
Proof.
iIntros (HLR Δ vvs HΔ) "[#Hinv #Hvvs]".
iIntros (γh γh' σ1') "Hσ1' /=". asimpl.
rewrite ic_eq /ic_def /=. iIntros (σ1 σ2 v n) "[Hrd1 How]".
iDestruct "Hrd1" as %Hrd1.
iPoseProof (HLR with "[]") as "He"; iFrame "#".
iSpecialize ("He" $! _ _ _ with "[$Hσ1']"); simpl.
rewrite ic_eq /ic_def /=.
iSpecialize ("He" $! _ _ _ _ with "[$How]"); eauto.
iMod "He" as "[He HPh1]".
iDestruct "He" as (σ2'' v'' ) "(Hrd2 & Ho1 & #Hrl1)".
iDestruct "Hrd2" as %Hrd2.
iModIntro; iFrame.
apply rtc_nsteps in Hrd2. destruct Hrd2 as [k Hrd2].
iExists σ2'', (TLamV _); iSplit; iFrame.
{ 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. econstructor. }
asimpl. iAlways.
iIntros (τi rn Hwwp).
iSpecialize ("Hrl1" $! τi rn with "[]"); auto.
iIntros (γh3 γh4 σ3') "Hσ3' /=".
iSpecialize ("Hrl1" $! γh3 _ _ with "[$Hσ3']").
iApply ic_mono; last trivial. simpl.
iIntros (f _); iDestruct 1 as (σ3 u) "(Hrd3 & Hσ3 & #Hrl)".
iDestruct "Hrd3" as %Hrd3.
iExists _, _; iFrame; iFrame "#".
iPureIntro.
etrans; last apply Hrd3.
eapply (rtc_l _ (_, _) (_, _)); simpl.
{ apply head_prim_step. econstructor; eauto. }
asimpl. econstructor.
Qed.
Lemma tlam_eta_LR3 `{ICG_ST Σ} `{LogRelG Σ} Γ e' τ :
Γ ⊢ₜ e' : TForall τ →
Γ ⊨ TLam (App (Rec (TApp (Var 1))) e') ≤log≤ TLam (TApp e') : TForall τ.
Proof.
iIntros (Ht Δ vvs HΔ) "#(Hinv & Hvvs)".
iDestruct (interp_env_length with "[]") as %Hvvsl; iFrame "#".
iIntros (γh γh' σ1') "Hσ1' /=".
rewrite ic_eq /ic_def /=. iIntros (σ1 σ2 v n) "[Hrd1 How]".
iDestruct "Hrd1" as %Hrd1.
eapply (nsteps_val _ _ (TLamV _)) in Hrd1.
destruct Hrd1 as (Hneq & Hveq & Hσeq); subst.
iModIntro; iFrame.
iExists _, (TLamV _); iFrame; iSplit.
{ iPureIntro. econstructor. }
rewrite Hveq; asimpl.
iAlways. iIntros (τi rn Hτi).
iApply logrel_pure_det_head_step.
{ intros σ; econstructor; eauto. }
{ intros ? ? ? Hrd. inversion Hrd; subst; auto. }
{ intros σ. econstructor; eauto. }
iNext.
iApply (logrel_bind [AppRCtx (RecV _)] [TAppCtx]);
iSplitL; first by iApply (binary_fundamental _ _ _ Ht with "[]"); iFrame "#".
iIntros (VV) "#Hvv"; simpl.
iSpecialize ("Hvv" $! τi rn with "[]"); auto.
iIntros (γh3 γh4 σ3') "Hσ3' /=".
iApply ic_rec; auto. iNext.
iSpecialize ("Hvv" $! γh3 γh4 σ3' with "[$Hσ3']").
by asimpl.
Qed.
Lemma tlam_eta_CR1 Γ e e' τ :
Γ ⊢ₜ e : TForall τ →
Γ ⊢ₜ e' : TForall τ →
(∀ `{ICG_ST Σ} `{LogRelG Σ}, Γ ⊨ e ≤log≤ e' : TForall τ) →
Γ ⊨ e ≤ctx≤ App (Rec (TLam (TApp (Var 1)))) e'
: TForall τ.
Proof.
intros Ht Ht' HLR.
apply binary_soundness; auto.
- econstructor; eauto. econstructor. econstructor.
replace τ with (τ.[up (ren (+1))].[TVar 0/]) at 4 by by asimpl.
econstructor. by econstructor.
- intros. eapply tlam_eta_LR1; eauto.
Qed.
Lemma tlam_eta_CR2 Γ e' τ :
Γ ⊢ₜ e' : TForall τ →
Γ ⊨ App (Rec (TLam (TApp (Var 1)))) e'
≤ctx≤
TLam (App (Rec (TApp (Var 1))) e') : TForall τ.
Proof.
intros Ht'.
pose proof (TLam_Hoisting Γ e' (TForall τ) (TApp (Var 0)) τ) as Hrh.
asimpl in *.
eapply ctx_refines_trans; first apply Hrh; eauto.
{ replace τ with (τ.[up (ren (+1))].[TVar 0/]) at 2 by by asimpl.
econstructor. by econstructor; asimpl. }
assert (Γ ⊢ₜ TLam (App (Rec (TApp (ids 1))) e') : TForall τ) as Htp.
{ econstructor; last eauto.
econstructor; last by apply context_rename.
econstructor.
replace τ with (τ.[up (ren (+1))].[TVar 0/]) at 4 by by asimpl.
repeat econstructor. }
apply binary_soundness; eauto.
intros. by apply binary_fundamental.
Qed.
Lemma tlam_eta_CR3 Γ e' τ :
Γ ⊢ₜ e' : TForall τ →
Γ ⊨ TLam (App (Rec (TApp (Var 1))) e') ≤ctx≤ TLam (TApp e') : TForall τ.
Proof.
intros Ht.
apply binary_soundness; eauto.
- econstructor. econstructor; last by apply context_rename.
econstructor.
replace τ with (τ.[up (ren (+1))].[TVar 0/]) at 4 by by asimpl.
econstructor. by econstructor; asimpl.
- econstructor. replace τ with (τ.[up (ren (+1))].[TVar 0/]) by by asimpl.
econstructor.
change (TForall τ.[up (ren (+1))]) with (TForall τ).[ren (+1)].
by apply context_rename.
- intros. eapply tlam_eta_LR3; eauto.
Qed.
Lemma tlam_eta_basic Γ e e' τ :
Γ ⊢ₜ e : TForall τ → Γ ⊢ₜ e' : TForall τ →
(∀ `{ICG_ST Σ} `{LogRelG Σ}, Γ ⊨ e ≤log≤ e' : TForall τ) →
Γ ⊨ e ≤ctx≤ TLam (TApp e') : TForall τ.
Proof.
intros Ht1 Ht2 HLR.
eapply ctx_refines_trans.
eapply tlam_eta_CR1; eauto.
eapply ctx_refines_trans.
eapply tlam_eta_CR2; eauto.
eapply tlam_eta_CR3; eauto.
Qed.
Lemma tlam_eta Γ e τ :
Γ ⊢ₜ e : TForall τ → Γ ⊨ e ≤ctx≤ TLam (TApp e) : TForall τ.
Proof.
intros.
eapply tlam_eta_basic; eauto.
intros; apply binary_fundamental; auto.
Qed.