RunST.rec_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 rec_hoisting.rec_hoisting
rec_hoisting.rec_hoisting_part1.
From iris.proofmode Require Import tactics.
Hint Resolve to_of_val.
Lemma rec_eta_LR1 `{ICG_ST Σ} `{LogRelG Σ} Γ e e' τ τ' :
Γ ⊨ e ≤log≤ e' : TArrow τ τ' →
Γ ⊨ e ≤log≤ App (Rec (Rec (App (Var 3) (Var 1)))).[ren (+1)] e'
: TArrow τ τ'.
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'', (RecV _); 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 (ww) "#Hww".
iIntros (γh3 γh4 σ3') "Hσ3' /=".
rewrite ic_eq /ic_def /=. iIntros (σ3 σ4 u k') "[Hrd3 How]".
iDestruct "Hrd3" as %Hrd3.
iSpecialize ("Hrl1" $! _ with "[]"); iFrame "#"; simpl.
iSpecialize ("Hrl1" $! _ _ _ with "[$Hσ3']"); simpl.
rewrite ic_eq /ic_def /=.
iSpecialize ("Hrl1" $! _ _ _ _ with "[$How]"); eauto.
iMod "Hrl1" as "[He HPh1]".
iDestruct "He" as (σ3'' u'') "(Hrd4 & Ho3'' & #Hrl2)".
iDestruct "Hrd4" as %Hrd4.
iModIntro; iFrame.
iExists _, _; iFrame; iFrame "#".
{ iPureIntro.
eapply (rtc_l _ (_, _) (_, _)); simpl.
{ apply head_prim_step. econstructor; eauto. }
asimpl. eauto. }
Qed.
Lemma rec_eta_LR3 `{ICG_ST Σ} `{LogRelG Σ} Γ e' τ τ' :
Γ ⊢ₜ e' : TArrow τ τ' →
Γ ⊨ Rec (App (Rec (App (Var 1) (Var 3))) e'.[ren (+2)])
≤log≤
Rec (App e'.[ren (+2)] (Var 1)) : TArrow τ τ'.
Proof.
iIntros (Ht Δ vvs HΔ) "#(Hinv & Hvvs)".
iDestruct (interp_env_length with "[]") as %Hvvsl; iFrame "#".
assert (TArrow τ τ' :: τ :: Γ ⊢ₜ
(App (Rec (App (Var 1) (Var 3))) e'.[ren (+2)]) : τ') as Ht1.
{ econstructor; last by eapply (context_weakening [_; _]).
repeat econstructor. }
assert (TArrow τ τ' :: τ :: Γ ⊢ₜ App e'.[ren (+2)] (Var 1) : τ') as Ht2.
{ econstructor; first by eapply (context_weakening [_; _]).
by econstructor. }
unshelve (iApply weak_bin_log_related_rec; last iFrame "#"); auto.
iIntros (ff vv) "[#Hinv #Hvvs]". cbn [fmap list_fmap].
do 2 (erewrite <- typed_subst_head_simpl_2; eauto;
last by rewrite /= fmap_length Hvvsl).
asimpl.
iIntros (γh γh' σ1') "Hσ1' /=".
rewrite ic_eq /ic_def /=. iIntros (σ1 σ2 v n) "[Hrd1 How]".
iDestruct "Hrd1" as %Hrd1.
iPoseProof (binary_fundamental _ _ _ Ht1 with "[]") as "He"; iFrame "#".
cbn [fmap list_fmap].
do 2 (erewrite <- typed_subst_head_simpl_2; eauto;
last by rewrite /= fmap_length Hvvsl).
asimpl.
iSpecialize ("He" $! _ _ _ with "[$Hσ1']").
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.
apply rtc_nsteps in Hrd2. destruct Hrd2 as [k Hrd2].
apply (@nsteps_bind det_lang (fill_item (AppRCtx (RecV _))) _) in Hrd2.
destruct Hrd2 as (k1 & σ3 & u & HK1 & Hrd21 & Hrd22). simpl in *.
eapply det_rec_red_step in Hrd22. destruct Hrd22 as (m & Hm & Hrd22).
asimpl in Hrd22.
iModIntro; iFrame.
iExists σ2'', _; iSplit; iFrame; iFrame "#".
{ iPureIntro. etrans.
- eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (AppLCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
- by eapply (nsteps_rtc _). }
Qed.
Lemma rec_eta_CR1 Γ e e' τ τ' :
Γ ⊢ₜ e : TArrow τ τ' →
Γ ⊢ₜ e' : TArrow τ τ' →
(∀ `{ICG_ST Σ} `{LogRelG Σ}, Γ ⊨ e ≤log≤ e' : TArrow τ τ') →
Γ ⊨ e ≤ctx≤
App (Rec (Rec (App (Var 3) (Var 1)))).[ren (+1)] e'
: TArrow τ τ'.
Proof.
intros Ht Ht' HLR.
apply binary_soundness; auto.
- econstructor; last eauto. econstructor. asimpl.
econstructor. econstructor.
constructor; simpl; eauto.
by constructor; simpl.
- intros. eapply rec_eta_LR1; eauto.
Qed.
Lemma rec_eta_CR2 Γ e' τ τ' :
Γ ⊢ₜ e' : TArrow τ τ' →
Γ ⊨ App (Rec (Rec (App (Var 3) (Var 1)))).[ren (+1)] e'
≤ctx≤
Rec (App (Rec (App (Var 1) (Var 3))) e'.[ren (+2)])
: TArrow τ τ'.
Proof.
intros Ht'.
pose proof (Rec_Hoisting Γ e' (TArrow τ τ') (App (Var 2) (Var 1))
τ τ' Ht') as Hrh.
asimpl in *.
eapply ctx_refines_trans; first apply Hrh.
{ repeat econstructor. }
assert (Γ ⊢ₜ Rec (App (Rec (App (Var 1) (Var 3))) e'.[ren (+2)])
: TArrow τ τ') as Htp.
{ econstructor; last eauto.
econstructor; last by eapply (context_weakening [_; _]).
repeat econstructor. }
apply binary_soundness; eauto.
intros. by apply binary_fundamental.
Qed.
Lemma rec_eta_CR3 Γ e' τ τ' :
Γ ⊢ₜ e' : TArrow τ τ' →
Γ ⊨ Rec (App (Rec (App (Var 1) (Var 3))) e'.[ren (+2)])
≤ctx≤
Rec (App e'.[ren (+2)] (Var 1))
: TArrow τ τ'.
Proof.
intros Ht.
apply binary_soundness; eauto.
- econstructor; eauto.
econstructor; last by eapply (context_weakening [_; _]).
repeat econstructor; simpl.
- econstructor. econstructor; first by eapply (context_weakening [_; _]).
by econstructor.
- intros. eapply rec_eta_LR3; eauto.
Qed.
Lemma rec_eta_basic Γ e e' τ τ' :
Γ ⊢ₜ e : TArrow τ τ' → Γ ⊢ₜ e' : TArrow τ τ' →
(∀ `{ICG_ST Σ} `{LogRelG Σ}, Γ ⊨ e ≤log≤ e' : TArrow τ τ') →
Γ ⊨ e ≤ctx≤ Rec (App e'.[ren (+2)] (Var 1)) : TArrow τ τ'.
Proof.
intros Ht1 Ht2 HLR.
eapply ctx_refines_trans.
eapply rec_eta_CR1; eauto.
eapply ctx_refines_trans.
eapply rec_eta_CR2; eauto.
eapply rec_eta_CR3; eauto.
Qed.
Lemma rec_eta Γ e τ τ' :
Γ ⊢ₜ e : TArrow τ τ' →
Γ ⊨ e ≤ctx≤ Rec (App e.[ren (+2)] (Var 1)) : TArrow τ τ'.
Proof.
intros.
eapply rec_eta_basic; eauto.
intros; apply binary_fundamental; auto.
Qed.
future contextual_refinement saved_pred ST_Lang_reduction
logrel_shared soundness logrel fundamental rec_hoisting.rec_hoisting
rec_hoisting.rec_hoisting_part1.
From iris.proofmode Require Import tactics.
Hint Resolve to_of_val.
Lemma rec_eta_LR1 `{ICG_ST Σ} `{LogRelG Σ} Γ e e' τ τ' :
Γ ⊨ e ≤log≤ e' : TArrow τ τ' →
Γ ⊨ e ≤log≤ App (Rec (Rec (App (Var 3) (Var 1)))).[ren (+1)] e'
: TArrow τ τ'.
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'', (RecV _); 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 (ww) "#Hww".
iIntros (γh3 γh4 σ3') "Hσ3' /=".
rewrite ic_eq /ic_def /=. iIntros (σ3 σ4 u k') "[Hrd3 How]".
iDestruct "Hrd3" as %Hrd3.
iSpecialize ("Hrl1" $! _ with "[]"); iFrame "#"; simpl.
iSpecialize ("Hrl1" $! _ _ _ with "[$Hσ3']"); simpl.
rewrite ic_eq /ic_def /=.
iSpecialize ("Hrl1" $! _ _ _ _ with "[$How]"); eauto.
iMod "Hrl1" as "[He HPh1]".
iDestruct "He" as (σ3'' u'') "(Hrd4 & Ho3'' & #Hrl2)".
iDestruct "Hrd4" as %Hrd4.
iModIntro; iFrame.
iExists _, _; iFrame; iFrame "#".
{ iPureIntro.
eapply (rtc_l _ (_, _) (_, _)); simpl.
{ apply head_prim_step. econstructor; eauto. }
asimpl. eauto. }
Qed.
Lemma rec_eta_LR3 `{ICG_ST Σ} `{LogRelG Σ} Γ e' τ τ' :
Γ ⊢ₜ e' : TArrow τ τ' →
Γ ⊨ Rec (App (Rec (App (Var 1) (Var 3))) e'.[ren (+2)])
≤log≤
Rec (App e'.[ren (+2)] (Var 1)) : TArrow τ τ'.
Proof.
iIntros (Ht Δ vvs HΔ) "#(Hinv & Hvvs)".
iDestruct (interp_env_length with "[]") as %Hvvsl; iFrame "#".
assert (TArrow τ τ' :: τ :: Γ ⊢ₜ
(App (Rec (App (Var 1) (Var 3))) e'.[ren (+2)]) : τ') as Ht1.
{ econstructor; last by eapply (context_weakening [_; _]).
repeat econstructor. }
assert (TArrow τ τ' :: τ :: Γ ⊢ₜ App e'.[ren (+2)] (Var 1) : τ') as Ht2.
{ econstructor; first by eapply (context_weakening [_; _]).
by econstructor. }
unshelve (iApply weak_bin_log_related_rec; last iFrame "#"); auto.
iIntros (ff vv) "[#Hinv #Hvvs]". cbn [fmap list_fmap].
do 2 (erewrite <- typed_subst_head_simpl_2; eauto;
last by rewrite /= fmap_length Hvvsl).
asimpl.
iIntros (γh γh' σ1') "Hσ1' /=".
rewrite ic_eq /ic_def /=. iIntros (σ1 σ2 v n) "[Hrd1 How]".
iDestruct "Hrd1" as %Hrd1.
iPoseProof (binary_fundamental _ _ _ Ht1 with "[]") as "He"; iFrame "#".
cbn [fmap list_fmap].
do 2 (erewrite <- typed_subst_head_simpl_2; eauto;
last by rewrite /= fmap_length Hvvsl).
asimpl.
iSpecialize ("He" $! _ _ _ with "[$Hσ1']").
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.
apply rtc_nsteps in Hrd2. destruct Hrd2 as [k Hrd2].
apply (@nsteps_bind det_lang (fill_item (AppRCtx (RecV _))) _) in Hrd2.
destruct Hrd2 as (k1 & σ3 & u & HK1 & Hrd21 & Hrd22). simpl in *.
eapply det_rec_red_step in Hrd22. destruct Hrd22 as (m & Hm & Hrd22).
asimpl in Hrd22.
iModIntro; iFrame.
iExists σ2'', _; iSplit; iFrame; iFrame "#".
{ iPureIntro. etrans.
- eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (AppLCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
- by eapply (nsteps_rtc _). }
Qed.
Lemma rec_eta_CR1 Γ e e' τ τ' :
Γ ⊢ₜ e : TArrow τ τ' →
Γ ⊢ₜ e' : TArrow τ τ' →
(∀ `{ICG_ST Σ} `{LogRelG Σ}, Γ ⊨ e ≤log≤ e' : TArrow τ τ') →
Γ ⊨ e ≤ctx≤
App (Rec (Rec (App (Var 3) (Var 1)))).[ren (+1)] e'
: TArrow τ τ'.
Proof.
intros Ht Ht' HLR.
apply binary_soundness; auto.
- econstructor; last eauto. econstructor. asimpl.
econstructor. econstructor.
constructor; simpl; eauto.
by constructor; simpl.
- intros. eapply rec_eta_LR1; eauto.
Qed.
Lemma rec_eta_CR2 Γ e' τ τ' :
Γ ⊢ₜ e' : TArrow τ τ' →
Γ ⊨ App (Rec (Rec (App (Var 3) (Var 1)))).[ren (+1)] e'
≤ctx≤
Rec (App (Rec (App (Var 1) (Var 3))) e'.[ren (+2)])
: TArrow τ τ'.
Proof.
intros Ht'.
pose proof (Rec_Hoisting Γ e' (TArrow τ τ') (App (Var 2) (Var 1))
τ τ' Ht') as Hrh.
asimpl in *.
eapply ctx_refines_trans; first apply Hrh.
{ repeat econstructor. }
assert (Γ ⊢ₜ Rec (App (Rec (App (Var 1) (Var 3))) e'.[ren (+2)])
: TArrow τ τ') as Htp.
{ econstructor; last eauto.
econstructor; last by eapply (context_weakening [_; _]).
repeat econstructor. }
apply binary_soundness; eauto.
intros. by apply binary_fundamental.
Qed.
Lemma rec_eta_CR3 Γ e' τ τ' :
Γ ⊢ₜ e' : TArrow τ τ' →
Γ ⊨ Rec (App (Rec (App (Var 1) (Var 3))) e'.[ren (+2)])
≤ctx≤
Rec (App e'.[ren (+2)] (Var 1))
: TArrow τ τ'.
Proof.
intros Ht.
apply binary_soundness; eauto.
- econstructor; eauto.
econstructor; last by eapply (context_weakening [_; _]).
repeat econstructor; simpl.
- econstructor. econstructor; first by eapply (context_weakening [_; _]).
by econstructor.
- intros. eapply rec_eta_LR3; eauto.
Qed.
Lemma rec_eta_basic Γ e e' τ τ' :
Γ ⊢ₜ e : TArrow τ τ' → Γ ⊢ₜ e' : TArrow τ τ' →
(∀ `{ICG_ST Σ} `{LogRelG Σ}, Γ ⊨ e ≤log≤ e' : TArrow τ τ') →
Γ ⊨ e ≤ctx≤ Rec (App e'.[ren (+2)] (Var 1)) : TArrow τ τ'.
Proof.
intros Ht1 Ht2 HLR.
eapply ctx_refines_trans.
eapply rec_eta_CR1; eauto.
eapply ctx_refines_trans.
eapply rec_eta_CR2; eauto.
eapply rec_eta_CR3; eauto.
Qed.
Lemma rec_eta Γ e τ τ' :
Γ ⊢ₜ e : TArrow τ τ' →
Γ ⊨ e ≤ctx≤ Rec (App e.[ren (+2)] (Var 1)) : TArrow τ τ'.
Proof.
intros.
eapply rec_eta_basic; eauto.
intros; apply binary_fundamental; auto.
Qed.