RunST.effectful_equations
From RunST Require Import IC base lang rules typing reduction regions
contextual_refinement saved_pred ST_Lang_reduction logrel_shared soundness
logrel fundamental rec_hoisting.rec_hoisting basic_equations
rec_hoisting.rec_hoisting_part1 future.
From iris.proofmode Require Import tactics.
Hint Resolve to_of_val.
Lemma ST_left_id_LR1 `{ICG_ST Σ} `{LogRelG Σ} Γ e e' ρ τ :
Γ ⊨ e ≤log≤ e' : (TST ρ τ) →
Γ ⊨ (Bind e (Rec (Return (Var 1)))) ≤log≤ e' : (TST ρ τ).
Proof.
iIntros (Hee Δ vvs HΔ) "[#Hinv #Hvvs]". asimpl.
iIntros (γh γh' σ1') "Hσ1'". simpl.
iPoseProof (Hee Δ vvs HΔ with "[]") as "Hee"; first iFrame "#".
iSpecialize ("Hee" $! _ _ _ with "[$Hσ1']").
iApply (ic_bind (fill_item (BindLCtx _)) _ with "[-]"); simpl.
iApply ic_wand_r; iSplitL; first iApply "Hee". simpl.
iIntros (v _) "H". iDestruct "H" as (σ2' v') "(Hrd & How & #Hvv)".
iDestruct "Hrd" as %Hrd.
iApply ic_value; first by simpl; rewrite to_of_val; eauto.
iExists _, _; iSplit; iFrame; auto.
iAlways. iIntros (γh3 γh4 σ3) "[Ho Hreg]". simpl.
change (Rec (Return (ids 1))) with (of_val (RecV (Return (ids 1)))).
iApply (ic_runst_to_bind with "[Ho Hreg]").
iApply ic_wand_r; iSplitL.
{ iApply "Hvv"; iFrame. }
simpl. iIntros (w _); iDestruct 1 as (σ3' w') "(Hrd & Ho & Hww & Hreg)".
iDestruct "Hrd" as %Hrd'.
iApply (ic_bind (fill_item (RunSTCtx))).
iApply ic_rec; eauto.
iNext. asimpl.
iApply ic_value; first by simpl; rewrite to_of_val.
simpl. iApply ic_return.
iNext. by iExists _, _; iFrame.
Qed.
Lemma ST_left_id_LR2 `{ICG_ST Σ} `{LogRelG Σ} Γ e e' ρ τ :
Γ ⊨ e ≤log≤ e' : (TST ρ τ) →
Γ ⊨ e ≤log≤ (Bind e' (Rec (Return (Var 1)))) : (TST ρ τ).
Proof.
iIntros (Hee Δ vvs HΔ) "[#Hinv #Hvvs]". asimpl.
iIntros (γh γh' σ1') "Hσ1'". simpl.
iPoseProof (Hee Δ vvs HΔ with "[]") as "Hee"; first iFrame "#".
iSpecialize ("Hee" $! _ _ _ with "[$Hσ1']").
iApply ic_wand_r; iSplitL; first iApply "Hee". simpl.
iIntros (v _) "H". iDestruct "H" as (σ2' v') "(Hrd & How & #Hvv)".
iDestruct "Hrd" as %Hrd.
apply rtc_nsteps in Hrd; destruct Hrd as [n Hrd].
iExists σ2', (BindV v' (RecV (Return (ids 1)))); iSplit; iFrame; auto.
{ iPureIntro.
apply (nsteps_rtc n).
eapply (@nsteps_ctx det_lang (fill_item (BindLCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
iAlways. iIntros (γh3 γh4 σ3) "[Ho Hreg]". simpl.
iApply ic_wand_r; iSplitL.
{ iApply "Hvv"; iFrame. }
simpl. iIntros (w _); iDestruct 1 as (σ3' w') "(Hrd & Ho & Hww & Hreg)".
iDestruct "Hrd" as %Hrd'.
iExists _, _; iFrame. iPureIntro.
apply rtc_nsteps in Hrd' as [n' Hrd'].
apply (nsteps_rtc (n' + 2)).
change (Rec (Return (ids 1))) with (of_val (RecV (Return (ids 1)))).
eapply red_RunST_Bind'; eauto.
eapply (nsteps_l _ _ _ (_, _) (_, _)); simpl; eauto.
eapply (Ectx_step _ _ _ _ _ [RunSTCtx]); simpl; eauto.
econstructor; eauto.
asimpl.
eapply (nsteps_l _ _ _ (_, _) (_, _)); simpl; eauto; last econstructor.
apply head_prim_step. econstructor; eauto.
Qed.
Lemma ST_left_id Γ e ρ τ :
Γ ⊢ₜ e : TST ρ τ → Γ ⊨ (Bind e (Rec (Return (Var 1)))) ≈ctx≈ e : (TST ρ τ).
Proof.
split.
- apply binary_soundness; eauto.
+ repeat econstructor; eauto.
+ intros; apply ST_left_id_LR1. by apply binary_fundamental.
- apply binary_soundness; eauto.
+ repeat econstructor; eauto.
+ intros; apply ST_left_id_LR2. by apply binary_fundamental.
Qed.
Lemma ST_right_id_LR1 `{ICG_ST Σ} `{LogRelG Σ} Γ e1 e1' τ e2 e2' ρ τ' :
Γ ⊨ e1 ≤log≤ e1' : τ →
Γ ⊨ e2 ≤log≤ e2' : TArrow τ (TST ρ τ') ->
Γ ⊨ App e2 e1 ≤log≤
App (App (Rec (Rec (Bind (Return (Var 1))
(App (Rec (Rec (Var 3))) (App (Var 3) (Var 1))))))
e2')
e1'
: TST ρ τ'.
Proof.
iIntros (He1 He2 Δ vvs HΔ) "[#Hinv #Hvvs]". asimpl.
iIntros (γh γh' σ1') "Hσ1'". simpl.
iPoseProof (He2 Δ vvs HΔ with "[]") as "He2"; first iFrame "#".
iSpecialize ("He2" $! _ _ _ with "[$Hσ1']").
iApply (ic_bind (fill_item (AppLCtx _)) _ with "[-]"); simpl.
iApply ic_wand_r; iSplitL; first iApply "He2". simpl.
iIntros (v _) "H". iDestruct "H" as (σ2' v') "(#Hrd & How & #Hvv)".
iDestruct "Hrd" as %Hrd2.
apply rtc_nsteps in Hrd2; destruct Hrd2 as [k2 Hrd2].
iApply (ic_bind (fill_item (AppRCtx _)) _ with "[-]"); simpl.
iPoseProof (He1 Δ vvs HΔ with "[]") as "He1"; first iFrame "#".
asimpl. iSpecialize ("He1" $! _ _ _ with "[$How]").
iApply ic_wand_r; iSplitL; first iApply "He1". simpl.
iIntros (v0 _) "H". iDestruct "H" as (σ3 v3) "(Hrd & Hσ3 & #H)".
iDestruct "Hrd" as %Hrd1.
apply rtc_nsteps in Hrd1; destruct Hrd1 as [k1 Hrd1].
iSpecialize ("Hvv" $! (v0, v3) with "H"). unfold interp_expr.
iApply ic_wand_r; iSplitL; first iApply "Hvv"; simpl; eauto.
iClear "Hvv".
iIntros (w _) "HH". iDestruct "HH" as (σ4 v4) "(Hrd3 & Hσ4 & #HH)".
iDestruct "Hrd3" as %Hrd3.
apply rtc_nsteps in Hrd3; destruct Hrd3 as [k3 Hrd3].
iExists _, (BindV _ (RecV _)); iFrame. iSplit.
{ 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.
{ simpl.
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. }
asimpl.
etrans.
{ eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (AppRCtx (RecV _))));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
asimpl.
eapply (rtc_l _ (_, _) (_, _)); simpl.
{ apply head_prim_step. econstructor; eauto. }
asimpl.
eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (BindRCtx (ReturnV _))));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
eapply nsteps_trans.
eapply (@nsteps_ctx det_lang (fill_item (AppRCtx (RecV _))));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
asimpl.
eapply (nsteps_l _ _ (_, _) (_, _)); simpl.
{ apply head_prim_step. econstructor; eauto. }
econstructor.
}
asimpl. iAlways.
iIntros (γh3 γh4 σ3') "[Hσ3' Hreg]".
iSpecialize ("HH" $! _ _ _ with "[$Hσ3' $Hreg]").
iApply ic_mono; last iExact "HH".
simpl. iIntros (u _). iDestruct 1 as (σ4' u') "(Hrd4 & Hσ4' & #Huu & Hreg)".
iDestruct "Hrd4" as %Hrd4.
iExists _, _; iFrame; iFrame "#".
iPureIntro.
{ etrans.
- eapply (rtc_l _ (_, _) (_, _)); simpl.
{ apply head_prim_step. econstructor; eauto.
apply head_prim_step. econstructor; eauto. }
eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (RunSTCtx)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
eapply (nsteps_l _ _ (_, _) (_, _)); simpl.
{ apply head_prim_step. econstructor; eauto. }
asimpl. econstructor.
- by asimpl.
}
Qed.
Lemma ctx_typed Γ e1 e2 ρ τ τ':
Γ ⊢ₜ e1 : τ →
Γ ⊢ₜ e2 : TArrow τ (TST ρ τ') ->
typed_ctx
((CTX_AppL e1) :: (CTX_AppL e2) :: CTX_Rec ::
CTX_Rec :: (CTX_BindR (Return (Var 1))) :: [])
(TArrow τ (TST ρ τ') :: τ ::
TArrow (TArrow τ (TST ρ τ')) (TArrow τ (TST ρ τ')) ::
TArrow τ (TST ρ τ') :: Γ)
(TArrow τ (TST ρ τ'))
Γ (TST ρ τ').
Proof.
intros Ht1 Ht2.
econstructor; first econstructor; eauto.
econstructor; first econstructor; eauto.
econstructor; first econstructor.
econstructor; first econstructor.
econstructor; first repeat econstructor.
econstructor.
Qed.
Lemma ST_right_id_CR2 Γ e1' τ e2' ρ τ' :
Γ ⊢ₜ e1' : τ →
Γ ⊢ₜ e2' : TArrow τ (TST ρ τ') ->
Γ ⊨ App (App (Rec (Rec (Bind (Return (Var 1))
(App (Rec (Rec (Var 3))) (App (Var 3) (Var 1)))))) e2') e1'
≤ctx≤
App (App (Rec (Rec (Bind (Return (Var 1))
(Rec (App (Rec (Var 1)) (App (Var 5) (Var 3))))))) e2') e1'
: TST ρ τ'.
Proof.
intros Ht1 Ht2.
eapply (ctx_refines_ctx_closed _ _ _ _
((CTX_AppL e1') :: (CTX_AppL e2') :: CTX_Rec ::
CTX_Rec :: (CTX_BindR (Return (Var 1))) :: []));
eauto using ctx_typed.
eapply (Rec_Hoisting _ (App (Var 3) (Var 1)) _ (Var 2)).
- repeat econstructor.
- by econstructor.
Qed.
Lemma nstep_RunST_ret k u v σ σ':
nsteps (pstep' ST_lang) k (RunST (Return (of_val u)), σ)
(of_val v, σ') →
k = 1 ∧ σ' = σ ∧ v = u.
Proof.
intros Hrd.
inversion Hrd as [Hs11 Hs12 Hs13 Hs14|k' Hs12 [e12 σ12] Hs14 Hs15 Hs16];
subst; first by destruct v; inversion Hs14.
asimpl in *.
apply head_reducible_prim_step in Hs15; last first.
{ eexists _, _, _; eapply RunRet; eauto. }
inversion Hs15; subst.
- simpl in *.
match goal with
H : prim_step (Return (of_val u)) σ e' σ12 [] |- _ =>
apply val_prim_stuck in H; rewrite /= to_of_val /= in H; inversion H
end.
- match goal with
H : to_val (of_val u) = Some _ |- _ =>
rewrite to_of_val in H; inversion H; subst
end.
eapply (nsteps_val ST_lang) in Hs16;
destruct Hs16 as [? [?%of_val_inj ?]]; subst.
auto.
Qed.
Lemma ST_right_id_LR3 `{ICG_ST Σ} `{LogRelG Σ} Γ τ ρ τ' :
(TArrow τ (TST ρ τ') :: τ ::
TArrow (TArrow τ (TST ρ τ')) (TArrow τ (TST ρ τ')) ::
TArrow τ (TST ρ τ') :: Γ)
⊨
Bind (Return (Var 1)) (Rec (App (Rec (Var 1)) (App (Var 5) (Var 3))))
≤log≤
Bind (Return (Var 1)) (Var 3)
: TST ρ τ'.
Proof.
iIntros (Δ vvs HΔ) "[#Hinv #Hvvs]".
iDestruct (interp_env_length with "[]") as %Hvvsl; first iFrame "#".
destruct vvs as [|vv0 [|vv1 [|vv2 [|vv3 vvs]]]]; inversion Hvvsl as [Hvvsl'].
assert (TArrow τ (TST ρ τ') :: τ ::
TArrow (TArrow τ (TST ρ τ')) (TArrow τ (TST ρ τ')) ::
TArrow τ (TST ρ τ') :: Γ
⊢ₜ Bind (Return (Var 1))
(Rec (App (Rec (Var 1)) (App (Var 5) (Var 3)))) : TST ρ τ')
as Ht1 by repeat econstructor.
assert (TArrow τ (TST ρ τ') :: τ ::
TArrow (TArrow τ (TST ρ τ')) (TArrow τ (TST ρ τ')) ::
TArrow τ (TST ρ τ') :: Γ
⊢ₜ (Bind (Return (Var 1)) (Var 3)) : TST ρ τ')
as Ht2 by repeat econstructor.
cbn [fmap list_fmap].
erewrite <- typed_subst_head_simpl_4; eauto;
last by rewrite /= map_length Hvvsl'.
erewrite <- typed_subst_head_simpl_4; eauto;
last by rewrite /= map_length Hvvsl'.
asimpl.
rewrite ?interp_env_cons.
iDestruct "Hvvs" as "(_ & Hvv1 & _ & Hvv3 & Hvvs)".
iIntros (γh γh' σ1') "Hσ1'". simpl.
iApply ic_value; first rewrite /= to_of_val //.
iExists _, (BindV (ReturnV _) _); iFrame.
iSplit; first by iPureIntro; econstructor.
iAlways. clear γh γh' σ1'.
iIntros (γh γh' σ1') "[Hσ1' Hreg]". simpl.
rewrite ic_eq /ic_def. iIntros (σ1 σ2 v n) "(Hrd & Hσ1)". asimpl.
iDestruct "Hrd" as %Hrd.
apply (red_RunST_Bind _ _ (RecV _)) in Hrd.
destruct Hrd as (k & u' & σ3 & Hk & Hrd1 & Hrd2).
asimpl in Hrd1; asimpl in Hrd2.
apply nstep_RunST_ret in Hrd1; destruct Hrd1 as [? [? ?]]; subst.
apply (nsteps_bind (fill_item RunSTCtx)) in Hrd2. simpl in *.
destruct Hrd2 as (k' & σ4 & u & Hk' & Hrd21 & Hrd22).
apply rec_red_step in Hrd21. asimpl in Hrd21.
destruct Hrd21 as (m & Hm & Hrd21).
apply (nsteps_bind (fill_item (AppRCtx (RecV _)))) in Hrd21. simpl in *.
destruct Hrd21 as (k'' & σ5 & w & Hk'' & Hrd211 & Hrd212).
apply rec_red_step in Hrd212. asimpl in Hrd212.
destruct Hrd212 as (m' & Hm' & Hrd212).
eapply (nsteps_val ST_lang) in Hrd212.
destruct Hrd212 as [? [?%of_val_inj ?]]; subst.
iSpecialize ("Hvv3" $! vv1 with "[]"); first iFrame "#".
iSpecialize ("Hvv3" $! _ _ _ with "[$Hσ1']"). simpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hvv3" $! _ _ _ _ with "[$Hσ1]"); auto.
iMod "Hvv3" as "[Hvv3 Hσ4]".
iDestruct "Hvv3" as (σ4' w') "(Hrd4 & Hσ4' & #Hrl)"; simpl.
iDestruct "Hrd4" as %Hrd4.
apply rtc_nsteps in Hrd4; destruct Hrd4 as [k4 Hrd4].
iSpecialize ("Hrl" $! _ _ _ with "[$Hσ4' $Hreg]").
iSpecialize ("Hrl" $! _ _ _ _ with "[$Hσ4]"); auto.
iMod "Hrl" as "[Hrl Hσ2]".
iDestruct "Hrl" as (σ2' v') "(Hrd5 & Hσ2' & Hvv & Hreg)".
iDestruct "Hrd5" as %Hrd5.
iModIntro; iFrame.
iExists _, _; iFrame.
iPureIntro.
eapply (rtc_l _ (_, _) (_, _)); simpl.
{ apply head_prim_step; econstructor.
apply head_prim_step; econstructor; eauto. }
etrans.
- eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (RunSTCtx)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
- eauto.
Qed.
Lemma ctx_typed2 Γ e1 e2 ρ τ τ':
Γ ⊢ₜ e1 : τ →
Γ ⊢ₜ e2 : TArrow τ (TST ρ τ') ->
typed_ctx
((CTX_AppL e1) :: (CTX_AppL e2) :: CTX_Rec ::
CTX_Rec :: [])
(TArrow τ (TST ρ τ') :: τ ::
TArrow (TArrow τ (TST ρ τ')) (TArrow τ (TST ρ τ')) ::
TArrow τ (TST ρ τ') :: Γ)
(TST ρ τ')
Γ (TST ρ τ').
Proof.
intros Ht1 Ht2.
econstructor; first econstructor; eauto.
econstructor; first econstructor; eauto.
econstructor; first econstructor.
econstructor; first econstructor.
econstructor.
Qed.
Lemma ST_right_id_CR4 Γ e1' τ e2' ρ τ' :
Γ ⊢ₜ e1' : τ →
Γ ⊢ₜ e2' : TArrow τ (TST ρ τ') ->
Γ ⊨ App (App (Rec (Rec (Bind (Return (Var 1))
(Rec (App (Rec (Var 1)) (App (Var 5) (Var 3))))))) e2') e1'
≤ctx≤
App (App (Rec (Rec (Bind (Return (Var 1)) (Var 3)))) e2') e1'
: TST ρ τ'.
Proof.
intros Ht1 Ht2.
eapply (ctx_refines_ctx_closed _ _ _ _
((CTX_AppL e1') :: (CTX_AppL e2') :: CTX_Rec ::
CTX_Rec :: []));
eauto using ctx_typed2.
apply binary_soundness.
- repeat econstructor.
- repeat econstructor.
- intros. apply ST_right_id_LR3.
Qed.
Lemma ST_right_id_CR5 Γ e1' τ e2' ρ τ' :
Γ ⊢ₜ e1' : τ →
Γ ⊢ₜ e2' : TArrow τ (TST ρ τ') ->
Γ ⊨ App (App (Rec (Rec (Bind (Return (Var 1)) (Var 3)))) e2') e1'
≤ctx≤
App (App (Rec (Rec (Bind (Return (Var 3)) (Var 1)))) e1') e2'
: TST ρ τ'.
Proof.
intros Ht1 Ht2.
eapply (Commutativity_alt _ _ _ _ _ (Bind (Return (Var 0)) (Var 1))); eauto.
repeat econstructor.
Qed.
Lemma ST_right_id_LR6 `{ICG_ST Σ} `{LogRelG Σ} Γ e1' τ e2' ρ τ' :
Γ ⊢ₜ e1' : τ →
Γ ⊢ₜ e2' : TArrow τ (TST ρ τ') ->
Γ ⊨ App (App (Rec (Rec (Bind (Return (Var 3)) (Var 1)))) e1') e2'
≤log≤ Bind (Return e1') e2' : TST ρ τ'.
Proof.
iIntros (Ht1 Ht2 Δ vvs HΔ) "[#Hinv #Hvvs]". asimpl.
iApply (logrel_bind [AppLCtx _; AppRCtx (RecV _)]
[BindLCtx _; ReturnCtx] (TST _ _)).
iSplitL; first by iApply (binary_fundamental _ _ _ Ht1); iFrame "#".
iIntros (vv) "#Hvv /=".
iIntros (γh γh' σ1') "Hσ1' /=".
iApply (ic_bind (fill_item (AppLCtx _))).
iApply ic_rec; auto.
iNext. asimpl.
iApply ic_value; auto.
simpl.
iRevert (γh γh' σ1') "Hσ1'".
iApply (logrel_bind [AppRCtx (RecV _)]
[BindRCtx (ReturnV _)] (TST _ _)).
iSplitL; first by iApply (binary_fundamental _ _ _ Ht2); iFrame "#".
iIntros (ww) "#Hww /=".
iIntros (γh γh' σ1') "Hσ1'". simpl.
iApply ic_rec; auto.
iNext. asimpl.
iApply ic_value; first by rewrite /= ?to_of_val /=.
iExists _, (BindV (ReturnV _) _); iFrame; iSplit.
{ iPureIntro; econstructor. }
iAlways. simpl.
clear γh γh' σ1'. iIntros (γh γh' σ1') "[Hσ1' Hreg] /=".
iApply (ic_runst_to_bind (ReturnV _)).
iApply ic_return. iNext.
iSpecialize ("Hww" $! _ with "[]"); first iFrame "#".
iApply (ic_bind (fill_item RunSTCtx)).
iApply ic_wand_r; iSplitL "Hσ1' Hww"; first by iApply "Hww".
simpl.
iIntros (u _). iDestruct 1 as (σ2' u') "(Hrd1 & Hσ2' & #Huu)".
iDestruct "Hrd1" as %Hrd1.
apply rtc_nsteps in Hrd1; destruct Hrd1 as [k1 Hrd1].
iApply ic_wand_r; iSplitL; first iApply ("Huu" $! _ _ _ with "[$Hreg $Hσ2']").
iClear "Huu".
simpl.
iIntros (w _). iDestruct 1 as (σ3' w') "(Hrd2 & Hσ3' & Hww & Hreg)".
iDestruct "Hrd2" as %Hrd2.
iExists _, _; iFrame.
iPureIntro.
eapply (rtc_l _ (_, _) (_, _)); simpl.
{ apply head_prim_step; econstructor.
apply head_prim_step; econstructor; eauto. }
etrans.
- eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (RunSTCtx)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
- auto.
Qed.
contextual_refinement saved_pred ST_Lang_reduction logrel_shared soundness
logrel fundamental rec_hoisting.rec_hoisting basic_equations
rec_hoisting.rec_hoisting_part1 future.
From iris.proofmode Require Import tactics.
Hint Resolve to_of_val.
Lemma ST_left_id_LR1 `{ICG_ST Σ} `{LogRelG Σ} Γ e e' ρ τ :
Γ ⊨ e ≤log≤ e' : (TST ρ τ) →
Γ ⊨ (Bind e (Rec (Return (Var 1)))) ≤log≤ e' : (TST ρ τ).
Proof.
iIntros (Hee Δ vvs HΔ) "[#Hinv #Hvvs]". asimpl.
iIntros (γh γh' σ1') "Hσ1'". simpl.
iPoseProof (Hee Δ vvs HΔ with "[]") as "Hee"; first iFrame "#".
iSpecialize ("Hee" $! _ _ _ with "[$Hσ1']").
iApply (ic_bind (fill_item (BindLCtx _)) _ with "[-]"); simpl.
iApply ic_wand_r; iSplitL; first iApply "Hee". simpl.
iIntros (v _) "H". iDestruct "H" as (σ2' v') "(Hrd & How & #Hvv)".
iDestruct "Hrd" as %Hrd.
iApply ic_value; first by simpl; rewrite to_of_val; eauto.
iExists _, _; iSplit; iFrame; auto.
iAlways. iIntros (γh3 γh4 σ3) "[Ho Hreg]". simpl.
change (Rec (Return (ids 1))) with (of_val (RecV (Return (ids 1)))).
iApply (ic_runst_to_bind with "[Ho Hreg]").
iApply ic_wand_r; iSplitL.
{ iApply "Hvv"; iFrame. }
simpl. iIntros (w _); iDestruct 1 as (σ3' w') "(Hrd & Ho & Hww & Hreg)".
iDestruct "Hrd" as %Hrd'.
iApply (ic_bind (fill_item (RunSTCtx))).
iApply ic_rec; eauto.
iNext. asimpl.
iApply ic_value; first by simpl; rewrite to_of_val.
simpl. iApply ic_return.
iNext. by iExists _, _; iFrame.
Qed.
Lemma ST_left_id_LR2 `{ICG_ST Σ} `{LogRelG Σ} Γ e e' ρ τ :
Γ ⊨ e ≤log≤ e' : (TST ρ τ) →
Γ ⊨ e ≤log≤ (Bind e' (Rec (Return (Var 1)))) : (TST ρ τ).
Proof.
iIntros (Hee Δ vvs HΔ) "[#Hinv #Hvvs]". asimpl.
iIntros (γh γh' σ1') "Hσ1'". simpl.
iPoseProof (Hee Δ vvs HΔ with "[]") as "Hee"; first iFrame "#".
iSpecialize ("Hee" $! _ _ _ with "[$Hσ1']").
iApply ic_wand_r; iSplitL; first iApply "Hee". simpl.
iIntros (v _) "H". iDestruct "H" as (σ2' v') "(Hrd & How & #Hvv)".
iDestruct "Hrd" as %Hrd.
apply rtc_nsteps in Hrd; destruct Hrd as [n Hrd].
iExists σ2', (BindV v' (RecV (Return (ids 1)))); iSplit; iFrame; auto.
{ iPureIntro.
apply (nsteps_rtc n).
eapply (@nsteps_ctx det_lang (fill_item (BindLCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
iAlways. iIntros (γh3 γh4 σ3) "[Ho Hreg]". simpl.
iApply ic_wand_r; iSplitL.
{ iApply "Hvv"; iFrame. }
simpl. iIntros (w _); iDestruct 1 as (σ3' w') "(Hrd & Ho & Hww & Hreg)".
iDestruct "Hrd" as %Hrd'.
iExists _, _; iFrame. iPureIntro.
apply rtc_nsteps in Hrd' as [n' Hrd'].
apply (nsteps_rtc (n' + 2)).
change (Rec (Return (ids 1))) with (of_val (RecV (Return (ids 1)))).
eapply red_RunST_Bind'; eauto.
eapply (nsteps_l _ _ _ (_, _) (_, _)); simpl; eauto.
eapply (Ectx_step _ _ _ _ _ [RunSTCtx]); simpl; eauto.
econstructor; eauto.
asimpl.
eapply (nsteps_l _ _ _ (_, _) (_, _)); simpl; eauto; last econstructor.
apply head_prim_step. econstructor; eauto.
Qed.
Lemma ST_left_id Γ e ρ τ :
Γ ⊢ₜ e : TST ρ τ → Γ ⊨ (Bind e (Rec (Return (Var 1)))) ≈ctx≈ e : (TST ρ τ).
Proof.
split.
- apply binary_soundness; eauto.
+ repeat econstructor; eauto.
+ intros; apply ST_left_id_LR1. by apply binary_fundamental.
- apply binary_soundness; eauto.
+ repeat econstructor; eauto.
+ intros; apply ST_left_id_LR2. by apply binary_fundamental.
Qed.
Lemma ST_right_id_LR1 `{ICG_ST Σ} `{LogRelG Σ} Γ e1 e1' τ e2 e2' ρ τ' :
Γ ⊨ e1 ≤log≤ e1' : τ →
Γ ⊨ e2 ≤log≤ e2' : TArrow τ (TST ρ τ') ->
Γ ⊨ App e2 e1 ≤log≤
App (App (Rec (Rec (Bind (Return (Var 1))
(App (Rec (Rec (Var 3))) (App (Var 3) (Var 1))))))
e2')
e1'
: TST ρ τ'.
Proof.
iIntros (He1 He2 Δ vvs HΔ) "[#Hinv #Hvvs]". asimpl.
iIntros (γh γh' σ1') "Hσ1'". simpl.
iPoseProof (He2 Δ vvs HΔ with "[]") as "He2"; first iFrame "#".
iSpecialize ("He2" $! _ _ _ with "[$Hσ1']").
iApply (ic_bind (fill_item (AppLCtx _)) _ with "[-]"); simpl.
iApply ic_wand_r; iSplitL; first iApply "He2". simpl.
iIntros (v _) "H". iDestruct "H" as (σ2' v') "(#Hrd & How & #Hvv)".
iDestruct "Hrd" as %Hrd2.
apply rtc_nsteps in Hrd2; destruct Hrd2 as [k2 Hrd2].
iApply (ic_bind (fill_item (AppRCtx _)) _ with "[-]"); simpl.
iPoseProof (He1 Δ vvs HΔ with "[]") as "He1"; first iFrame "#".
asimpl. iSpecialize ("He1" $! _ _ _ with "[$How]").
iApply ic_wand_r; iSplitL; first iApply "He1". simpl.
iIntros (v0 _) "H". iDestruct "H" as (σ3 v3) "(Hrd & Hσ3 & #H)".
iDestruct "Hrd" as %Hrd1.
apply rtc_nsteps in Hrd1; destruct Hrd1 as [k1 Hrd1].
iSpecialize ("Hvv" $! (v0, v3) with "H"). unfold interp_expr.
iApply ic_wand_r; iSplitL; first iApply "Hvv"; simpl; eauto.
iClear "Hvv".
iIntros (w _) "HH". iDestruct "HH" as (σ4 v4) "(Hrd3 & Hσ4 & #HH)".
iDestruct "Hrd3" as %Hrd3.
apply rtc_nsteps in Hrd3; destruct Hrd3 as [k3 Hrd3].
iExists _, (BindV _ (RecV _)); iFrame. iSplit.
{ 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.
{ simpl.
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. }
asimpl.
etrans.
{ eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (AppRCtx (RecV _))));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
asimpl.
eapply (rtc_l _ (_, _) (_, _)); simpl.
{ apply head_prim_step. econstructor; eauto. }
asimpl.
eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (BindRCtx (ReturnV _))));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
eapply nsteps_trans.
eapply (@nsteps_ctx det_lang (fill_item (AppRCtx (RecV _))));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
asimpl.
eapply (nsteps_l _ _ (_, _) (_, _)); simpl.
{ apply head_prim_step. econstructor; eauto. }
econstructor.
}
asimpl. iAlways.
iIntros (γh3 γh4 σ3') "[Hσ3' Hreg]".
iSpecialize ("HH" $! _ _ _ with "[$Hσ3' $Hreg]").
iApply ic_mono; last iExact "HH".
simpl. iIntros (u _). iDestruct 1 as (σ4' u') "(Hrd4 & Hσ4' & #Huu & Hreg)".
iDestruct "Hrd4" as %Hrd4.
iExists _, _; iFrame; iFrame "#".
iPureIntro.
{ etrans.
- eapply (rtc_l _ (_, _) (_, _)); simpl.
{ apply head_prim_step. econstructor; eauto.
apply head_prim_step. econstructor; eauto. }
eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (RunSTCtx)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
eapply (nsteps_l _ _ (_, _) (_, _)); simpl.
{ apply head_prim_step. econstructor; eauto. }
asimpl. econstructor.
- by asimpl.
}
Qed.
Lemma ctx_typed Γ e1 e2 ρ τ τ':
Γ ⊢ₜ e1 : τ →
Γ ⊢ₜ e2 : TArrow τ (TST ρ τ') ->
typed_ctx
((CTX_AppL e1) :: (CTX_AppL e2) :: CTX_Rec ::
CTX_Rec :: (CTX_BindR (Return (Var 1))) :: [])
(TArrow τ (TST ρ τ') :: τ ::
TArrow (TArrow τ (TST ρ τ')) (TArrow τ (TST ρ τ')) ::
TArrow τ (TST ρ τ') :: Γ)
(TArrow τ (TST ρ τ'))
Γ (TST ρ τ').
Proof.
intros Ht1 Ht2.
econstructor; first econstructor; eauto.
econstructor; first econstructor; eauto.
econstructor; first econstructor.
econstructor; first econstructor.
econstructor; first repeat econstructor.
econstructor.
Qed.
Lemma ST_right_id_CR2 Γ e1' τ e2' ρ τ' :
Γ ⊢ₜ e1' : τ →
Γ ⊢ₜ e2' : TArrow τ (TST ρ τ') ->
Γ ⊨ App (App (Rec (Rec (Bind (Return (Var 1))
(App (Rec (Rec (Var 3))) (App (Var 3) (Var 1)))))) e2') e1'
≤ctx≤
App (App (Rec (Rec (Bind (Return (Var 1))
(Rec (App (Rec (Var 1)) (App (Var 5) (Var 3))))))) e2') e1'
: TST ρ τ'.
Proof.
intros Ht1 Ht2.
eapply (ctx_refines_ctx_closed _ _ _ _
((CTX_AppL e1') :: (CTX_AppL e2') :: CTX_Rec ::
CTX_Rec :: (CTX_BindR (Return (Var 1))) :: []));
eauto using ctx_typed.
eapply (Rec_Hoisting _ (App (Var 3) (Var 1)) _ (Var 2)).
- repeat econstructor.
- by econstructor.
Qed.
Lemma nstep_RunST_ret k u v σ σ':
nsteps (pstep' ST_lang) k (RunST (Return (of_val u)), σ)
(of_val v, σ') →
k = 1 ∧ σ' = σ ∧ v = u.
Proof.
intros Hrd.
inversion Hrd as [Hs11 Hs12 Hs13 Hs14|k' Hs12 [e12 σ12] Hs14 Hs15 Hs16];
subst; first by destruct v; inversion Hs14.
asimpl in *.
apply head_reducible_prim_step in Hs15; last first.
{ eexists _, _, _; eapply RunRet; eauto. }
inversion Hs15; subst.
- simpl in *.
match goal with
H : prim_step (Return (of_val u)) σ e' σ12 [] |- _ =>
apply val_prim_stuck in H; rewrite /= to_of_val /= in H; inversion H
end.
- match goal with
H : to_val (of_val u) = Some _ |- _ =>
rewrite to_of_val in H; inversion H; subst
end.
eapply (nsteps_val ST_lang) in Hs16;
destruct Hs16 as [? [?%of_val_inj ?]]; subst.
auto.
Qed.
Lemma ST_right_id_LR3 `{ICG_ST Σ} `{LogRelG Σ} Γ τ ρ τ' :
(TArrow τ (TST ρ τ') :: τ ::
TArrow (TArrow τ (TST ρ τ')) (TArrow τ (TST ρ τ')) ::
TArrow τ (TST ρ τ') :: Γ)
⊨
Bind (Return (Var 1)) (Rec (App (Rec (Var 1)) (App (Var 5) (Var 3))))
≤log≤
Bind (Return (Var 1)) (Var 3)
: TST ρ τ'.
Proof.
iIntros (Δ vvs HΔ) "[#Hinv #Hvvs]".
iDestruct (interp_env_length with "[]") as %Hvvsl; first iFrame "#".
destruct vvs as [|vv0 [|vv1 [|vv2 [|vv3 vvs]]]]; inversion Hvvsl as [Hvvsl'].
assert (TArrow τ (TST ρ τ') :: τ ::
TArrow (TArrow τ (TST ρ τ')) (TArrow τ (TST ρ τ')) ::
TArrow τ (TST ρ τ') :: Γ
⊢ₜ Bind (Return (Var 1))
(Rec (App (Rec (Var 1)) (App (Var 5) (Var 3)))) : TST ρ τ')
as Ht1 by repeat econstructor.
assert (TArrow τ (TST ρ τ') :: τ ::
TArrow (TArrow τ (TST ρ τ')) (TArrow τ (TST ρ τ')) ::
TArrow τ (TST ρ τ') :: Γ
⊢ₜ (Bind (Return (Var 1)) (Var 3)) : TST ρ τ')
as Ht2 by repeat econstructor.
cbn [fmap list_fmap].
erewrite <- typed_subst_head_simpl_4; eauto;
last by rewrite /= map_length Hvvsl'.
erewrite <- typed_subst_head_simpl_4; eauto;
last by rewrite /= map_length Hvvsl'.
asimpl.
rewrite ?interp_env_cons.
iDestruct "Hvvs" as "(_ & Hvv1 & _ & Hvv3 & Hvvs)".
iIntros (γh γh' σ1') "Hσ1'". simpl.
iApply ic_value; first rewrite /= to_of_val //.
iExists _, (BindV (ReturnV _) _); iFrame.
iSplit; first by iPureIntro; econstructor.
iAlways. clear γh γh' σ1'.
iIntros (γh γh' σ1') "[Hσ1' Hreg]". simpl.
rewrite ic_eq /ic_def. iIntros (σ1 σ2 v n) "(Hrd & Hσ1)". asimpl.
iDestruct "Hrd" as %Hrd.
apply (red_RunST_Bind _ _ (RecV _)) in Hrd.
destruct Hrd as (k & u' & σ3 & Hk & Hrd1 & Hrd2).
asimpl in Hrd1; asimpl in Hrd2.
apply nstep_RunST_ret in Hrd1; destruct Hrd1 as [? [? ?]]; subst.
apply (nsteps_bind (fill_item RunSTCtx)) in Hrd2. simpl in *.
destruct Hrd2 as (k' & σ4 & u & Hk' & Hrd21 & Hrd22).
apply rec_red_step in Hrd21. asimpl in Hrd21.
destruct Hrd21 as (m & Hm & Hrd21).
apply (nsteps_bind (fill_item (AppRCtx (RecV _)))) in Hrd21. simpl in *.
destruct Hrd21 as (k'' & σ5 & w & Hk'' & Hrd211 & Hrd212).
apply rec_red_step in Hrd212. asimpl in Hrd212.
destruct Hrd212 as (m' & Hm' & Hrd212).
eapply (nsteps_val ST_lang) in Hrd212.
destruct Hrd212 as [? [?%of_val_inj ?]]; subst.
iSpecialize ("Hvv3" $! vv1 with "[]"); first iFrame "#".
iSpecialize ("Hvv3" $! _ _ _ with "[$Hσ1']"). simpl.
rewrite ic_eq /ic_def.
iSpecialize ("Hvv3" $! _ _ _ _ with "[$Hσ1]"); auto.
iMod "Hvv3" as "[Hvv3 Hσ4]".
iDestruct "Hvv3" as (σ4' w') "(Hrd4 & Hσ4' & #Hrl)"; simpl.
iDestruct "Hrd4" as %Hrd4.
apply rtc_nsteps in Hrd4; destruct Hrd4 as [k4 Hrd4].
iSpecialize ("Hrl" $! _ _ _ with "[$Hσ4' $Hreg]").
iSpecialize ("Hrl" $! _ _ _ _ with "[$Hσ4]"); auto.
iMod "Hrl" as "[Hrl Hσ2]".
iDestruct "Hrl" as (σ2' v') "(Hrd5 & Hσ2' & Hvv & Hreg)".
iDestruct "Hrd5" as %Hrd5.
iModIntro; iFrame.
iExists _, _; iFrame.
iPureIntro.
eapply (rtc_l _ (_, _) (_, _)); simpl.
{ apply head_prim_step; econstructor.
apply head_prim_step; econstructor; eauto. }
etrans.
- eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (RunSTCtx)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
- eauto.
Qed.
Lemma ctx_typed2 Γ e1 e2 ρ τ τ':
Γ ⊢ₜ e1 : τ →
Γ ⊢ₜ e2 : TArrow τ (TST ρ τ') ->
typed_ctx
((CTX_AppL e1) :: (CTX_AppL e2) :: CTX_Rec ::
CTX_Rec :: [])
(TArrow τ (TST ρ τ') :: τ ::
TArrow (TArrow τ (TST ρ τ')) (TArrow τ (TST ρ τ')) ::
TArrow τ (TST ρ τ') :: Γ)
(TST ρ τ')
Γ (TST ρ τ').
Proof.
intros Ht1 Ht2.
econstructor; first econstructor; eauto.
econstructor; first econstructor; eauto.
econstructor; first econstructor.
econstructor; first econstructor.
econstructor.
Qed.
Lemma ST_right_id_CR4 Γ e1' τ e2' ρ τ' :
Γ ⊢ₜ e1' : τ →
Γ ⊢ₜ e2' : TArrow τ (TST ρ τ') ->
Γ ⊨ App (App (Rec (Rec (Bind (Return (Var 1))
(Rec (App (Rec (Var 1)) (App (Var 5) (Var 3))))))) e2') e1'
≤ctx≤
App (App (Rec (Rec (Bind (Return (Var 1)) (Var 3)))) e2') e1'
: TST ρ τ'.
Proof.
intros Ht1 Ht2.
eapply (ctx_refines_ctx_closed _ _ _ _
((CTX_AppL e1') :: (CTX_AppL e2') :: CTX_Rec ::
CTX_Rec :: []));
eauto using ctx_typed2.
apply binary_soundness.
- repeat econstructor.
- repeat econstructor.
- intros. apply ST_right_id_LR3.
Qed.
Lemma ST_right_id_CR5 Γ e1' τ e2' ρ τ' :
Γ ⊢ₜ e1' : τ →
Γ ⊢ₜ e2' : TArrow τ (TST ρ τ') ->
Γ ⊨ App (App (Rec (Rec (Bind (Return (Var 1)) (Var 3)))) e2') e1'
≤ctx≤
App (App (Rec (Rec (Bind (Return (Var 3)) (Var 1)))) e1') e2'
: TST ρ τ'.
Proof.
intros Ht1 Ht2.
eapply (Commutativity_alt _ _ _ _ _ (Bind (Return (Var 0)) (Var 1))); eauto.
repeat econstructor.
Qed.
Lemma ST_right_id_LR6 `{ICG_ST Σ} `{LogRelG Σ} Γ e1' τ e2' ρ τ' :
Γ ⊢ₜ e1' : τ →
Γ ⊢ₜ e2' : TArrow τ (TST ρ τ') ->
Γ ⊨ App (App (Rec (Rec (Bind (Return (Var 3)) (Var 1)))) e1') e2'
≤log≤ Bind (Return e1') e2' : TST ρ τ'.
Proof.
iIntros (Ht1 Ht2 Δ vvs HΔ) "[#Hinv #Hvvs]". asimpl.
iApply (logrel_bind [AppLCtx _; AppRCtx (RecV _)]
[BindLCtx _; ReturnCtx] (TST _ _)).
iSplitL; first by iApply (binary_fundamental _ _ _ Ht1); iFrame "#".
iIntros (vv) "#Hvv /=".
iIntros (γh γh' σ1') "Hσ1' /=".
iApply (ic_bind (fill_item (AppLCtx _))).
iApply ic_rec; auto.
iNext. asimpl.
iApply ic_value; auto.
simpl.
iRevert (γh γh' σ1') "Hσ1'".
iApply (logrel_bind [AppRCtx (RecV _)]
[BindRCtx (ReturnV _)] (TST _ _)).
iSplitL; first by iApply (binary_fundamental _ _ _ Ht2); iFrame "#".
iIntros (ww) "#Hww /=".
iIntros (γh γh' σ1') "Hσ1'". simpl.
iApply ic_rec; auto.
iNext. asimpl.
iApply ic_value; first by rewrite /= ?to_of_val /=.
iExists _, (BindV (ReturnV _) _); iFrame; iSplit.
{ iPureIntro; econstructor. }
iAlways. simpl.
clear γh γh' σ1'. iIntros (γh γh' σ1') "[Hσ1' Hreg] /=".
iApply (ic_runst_to_bind (ReturnV _)).
iApply ic_return. iNext.
iSpecialize ("Hww" $! _ with "[]"); first iFrame "#".
iApply (ic_bind (fill_item RunSTCtx)).
iApply ic_wand_r; iSplitL "Hσ1' Hww"; first by iApply "Hww".
simpl.
iIntros (u _). iDestruct 1 as (σ2' u') "(Hrd1 & Hσ2' & #Huu)".
iDestruct "Hrd1" as %Hrd1.
apply rtc_nsteps in Hrd1; destruct Hrd1 as [k1 Hrd1].
iApply ic_wand_r; iSplitL; first iApply ("Huu" $! _ _ _ with "[$Hreg $Hσ2']").
iClear "Huu".
simpl.
iIntros (w _). iDestruct 1 as (σ3' w') "(Hrd2 & Hσ3' & Hww & Hreg)".
iDestruct "Hrd2" as %Hrd2.
iExists _, _; iFrame.
iPureIntro.
eapply (rtc_l _ (_, _) (_, _)); simpl.
{ apply head_prim_step; econstructor.
apply head_prim_step; econstructor; eauto. }
etrans.
- eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (RunSTCtx)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
- auto.
Qed.
Theorem ST_right_id Γ e1 τ e2 ρ τ' :
Γ ⊢ₜ e1 : τ → Γ ⊢ₜ e2 : (TArrow τ (TST ρ τ')) →
Γ ⊨ (App e2 e1) ≤ctx≤ (Bind (Return e1) e2) : (TST ρ τ').
Proof.
intros Ht1 Ht2.
eapply ctx_refines_trans.
{ apply binary_soundness;
last by intros; eapply ST_right_id_LR1; eauto using binary_fundamental.
- econstructor; eauto.
- repeat econstructor; eauto. }
eapply ctx_refines_trans; first eapply ST_right_id_CR2; eauto.
eapply ctx_refines_trans; first eapply ST_right_id_CR4; eauto.
eapply ctx_refines_trans; first eapply ST_right_id_CR5; eauto.
apply binary_soundness.
- repeat econstructor; eauto.
- repeat econstructor; eauto.
- intros; eapply ST_right_id_LR6; eauto.
Qed.
Lemma ST_associativity_LR1
`{ICG_ST Σ} `{LogRelG Σ} Γ e1 e1' ρ τ1 e2 e2' τ2 e3 e3' τ3 :
Γ ⊨ e1 ≤log≤ e1' : TST ρ τ1 →
Γ ⊨ e2 ≤log≤ e2' : (TArrow τ1 (TST ρ τ2)) →
Γ ⊨ e3 ≤log≤ e3' : (TArrow τ2 (TST ρ τ3)) →
Γ ⊨ (Bind (Bind e1 e2) e3) ≤log≤
(Bind e1'
(App
(Rec (Rec (Bind (App (Fst (Var 3)) (Var 1)) (Snd (Var 3)))))
(Pair e2' e3')))
: (TST ρ τ3).
Proof.
iIntros (LR1 LR2 LR3 Δ vvs HΔ) "[#Hinv #Hvvs]". asimpl.
iApply (logrel_bind [BindLCtx _; BindLCtx _] [BindLCtx _] (TST _ _)).
iSplitL; first by iApply LR1; iFrame "#".
iIntros (vv1) "#Hvv1 /=".
iApply (logrel_bind [BindLCtx _; BindRCtx _]
[BindRCtx _; AppRCtx (RecV _); PairLCtx _] (TST _ _)).
iSplitL; first by iApply LR2; iFrame "#".
iIntros (vv2) "#Hvv2 /=".
iApply (logrel_bind [BindRCtx (BindV _ _)]
[BindRCtx _; AppRCtx (RecV _); PairRCtx _] (TST _ _)).
iSplitL; first by iApply LR3; iFrame "#".
iIntros (vv3) "#Hvv3 /=".
iIntros (γh γh' σ1') "Hσ1' /=".
iApply ic_value; first by rewrite /= ?to_of_val.
iExists _, (BindV _ (RecV _)); iFrame. iSplit.
{ iPureIntro.
eapply (nsteps_rtc _). simpl.
eapply (@nsteps_ctx det_lang (fill_item (BindRCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
eapply (nsteps_l _ _ (_, _) (_, _)); simpl.
{ apply head_prim_step; econstructor; by rewrite /= ?to_of_val. }
asimpl. econstructor. }
iAlways. asimpl.
clear γh γh' σ1'. iIntros (γh γh' σ1') "[Hσ1' Hreg] /=".
iApply (ic_runst_to_bind (BindV _ _)).
iApply ic_runst_to_bind.
iApply ic_wand_r; iSplitL.
{ iApply ("Hvv1" $! _ _ _ with "[-]"); first iFrame. }
simpl. iIntros (w1 _).
iDestruct 1 as (σ2' w1') "(Hrd1 & Hσ2' & #Hww1 & Hreg)".
iDestruct "Hrd1" as %Hrd1.
apply rtc_nsteps in Hrd1; destruct Hrd1 as [k1 Hrd1].
iSpecialize ("Hvv2" $! (w1, w1') with "[]"); first iFrame "#".
iApply (ic_bind (fill_item RunSTCtx)).
iApply ic_wand_r; iSplitL "Hσ2' Hvv2".
{ iApply ("Hvv2" $! _ _ _ with "[-]"); first iFrame. }
simpl. iIntros (w2 _).
iDestruct 1 as (σ3' w2') "(Hrd2 & Hσ3' & #Hww2)".
iDestruct "Hrd2" as %Hrd2.
apply rtc_nsteps in Hrd2; destruct Hrd2 as [k2 Hrd2].
iApply ic_wand_r; iSplitL.
{ iApply ("Hww2" $! _ _ _ with "[-]"); first iFrame. }
simpl. iIntros (u _).
iDestruct 1 as (σ4' u') "(Hrd3 & Hσ4' & #Huu & Hreg)".
iDestruct "Hrd3" as %Hrd3.
apply rtc_nsteps in Hrd3; destruct Hrd3 as [k3 Hrd3].
iSpecialize ("Hvv3" $! (u, u') with "[]"); first iFrame "#".
iApply (ic_bind (fill_item RunSTCtx)).
iApply ic_wand_r; iSplitL "Hσ4' Hvv3".
{ iApply ("Hvv3" $! _ _ _ with "[-]"); first iFrame. }
simpl. iIntros (w3 _).
iDestruct 1 as (σ5' w3') "(Hrd4 & Hσ5' & #Hww3)".
iDestruct "Hrd4" as %Hrd4.
apply rtc_nsteps in Hrd4; destruct Hrd4 as [k4 Hrd4].
iApply ic_wand_r; iSplitL.
{ iApply ("Hww3" $! _ _ _ with "[-]"); first iFrame. }
simpl. iIntros (u2 _).
iDestruct 1 as (σ6' u2') "(Hrd5 & Hσ6' & #Huu2 & Hreg)".
iDestruct "Hrd5" as %Hrd5.
apply rtc_nsteps in Hrd5; destruct Hrd5 as [k5 Hrd5].
iExists _, _; iFrame; iFrame "#".
iPureIntro.
etrans.
- eapply nsteps_rtc.
eapply (red_RunST_Bind' _ _ (RecV _)); eauto.
eapply nsteps_trans.
{ eapply (@nsteps_ctx det_lang (fill_item (RunSTCtx)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
eapply (nsteps_l _ _ (_, _) (_, _)); simpl.
{ apply head_prim_step; econstructor; eauto. }
asimpl.
eapply nsteps_trans.
- eapply (@nsteps_ctx det_lang (fill_item (BindLCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
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. }
econstructor.
- eapply (@nsteps_ctx det_lang (fill_item (BindLCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
eapply nsteps_trans.
{ eapply (@nsteps_ctx det_lang (fill_item (RunSTCtx)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
eapply (@nsteps_ctx det_lang (fill_item (BindRCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
eapply (nsteps_l _ _ (_, _) (_, _)); simpl.
{ apply head_prim_step; econstructor; eauto. }
econstructor. }
simpl.
eapply red_RunST_Bind'; eauto.
eapply nsteps_trans.
{ eapply (@nsteps_ctx det_lang (fill_item (RunSTCtx)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
simpl. eauto.
- econstructor.
Qed.
Lemma ST_associativity_CR2 Γ e1' ρ τ1 e2' τ2 e3' τ3 :
Γ ⊢ₜ e1' : TST ρ τ1 → Γ ⊢ₜ e2' : (TArrow τ1 (TST ρ τ2)) →
Γ ⊢ₜ e3' : (TArrow τ2 (TST ρ τ3)) →
Γ ⊨ (Bind e1'
(App
(Rec (Rec (Bind (App (Fst (Var 3)) (Var 1)) (Snd (Var 3)))))
(Pair e2' e3'))) ≤ctx≤
(Bind e1'
(Rec (App (Rec (Bind (App (Fst (Var 1)) (Var 3)) (Snd (Var 1))))
(Pair e2'.[ren (+2)] e3'.[ren (+2)]))))
: (TST ρ τ3).
Proof.
intros Ht1 Ht2 Ht3.
eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_BindR _]).
{ repeat econstructor; eauto. }
eapply (Rec_Hoisting _ (Pair e2' e3') _
(Bind (App (Fst (Var 2)) (Var 1)) (Snd (Var 2)))).
- repeat econstructor; eauto; eapply (context_weakening [ _ ; _]); eauto.
- repeat econstructor.
Qed.
Lemma ST_associativity_LR3 `{ICG_ST Σ} `{LogRelG Σ} Γ ρ τ1 e2' τ2 e3' τ3 :
Γ ⊢ₜ e2' : (TArrow τ1 (TST ρ τ2)) →
Γ ⊢ₜ e3' : (TArrow τ2 (TST ρ τ3)) →
Γ ⊨ Rec (App (Rec (Bind (App (Fst (Var 1)) (Var 3)) (Snd (Var 1))))
(Pair e2'.[ren (+2)] e3'.[ren (+2)]))
≤log≤
Rec (App (Rec (App (App (Rec (Rec (Bind (Var 1) (Var 3))))
e3'.[ren (+4)])
(App (Var 1) (Var 3)))) e2'.[ren (+2)])
: TArrow τ1 (TST ρ τ3).
Proof.
iIntros (Ht2 Ht3 Δ vvs HΔ) "[#Hinv #Hvvs]".
iDestruct (interp_env_length with "[]") as %Hvvsl; first iFrame "#".
assert ( TArrow τ1 (TST ρ τ3) :: τ1 :: Γ
⊢ₜ App (Rec (Bind (App (Fst (Var 1)) (Var 3)) (Snd (Var 1))))
(Pair e2'.[ren (+2)] e3'.[ren (+2)])
: TST ρ τ3) as Ht4.
{ repeat econstructor; eapply (context_weakening [ _ ; _]); eauto. }
assert (TArrow τ1 (TST ρ τ3) :: τ1 :: Γ ⊢ₜ
(App
(Rec
(App (App (Rec (Rec (Bind (Var 1) (Var 3)))) e3'.[ren (+4)])
(App (Var 1) (Var 3)))) e2'.[ren (+2)]) : TST ρ τ3)
as Ht5.
{ repeat econstructor; last eapply (context_weakening [ _ ; _]); eauto.
by eapply (context_weakening [_; _; _; _]). }
iApply (weak_bin_log_related_rec _ _ _ _ _ _ _ HΔ); eauto.
iIntros (vv1 vv2) "[#Hinv #Hvvs]".
cbn [fmap list_fmap].
erewrite <- typed_subst_head_simpl_2; eauto;
last by rewrite /= map_length Hvvsl.
erewrite <- typed_subst_head_simpl_2; eauto;
last by rewrite /= map_length Hvvsl.
asimpl.
rewrite ?interp_env_cons.
iDestruct "Hvvs" as "(_ & #Hvv1 & #Hvvs)".
iApply (logrel_bind [AppRCtx (RecV _); PairLCtx _]
[AppRCtx (RecV _)] (TST _ _)).
iSplitL; first by iApply (binary_fundamental _ _ _ Ht2); auto.
iIntros (ww2) "#Hww2 /=".
iIntros (γh γh' σ1') "Hσ1' /=".
iApply (ic_bind (fill_item (AppRCtx (RecV _)))).
iApply (ic_bind (fill_item (PairRCtx _))).
iPoseProof (binary_fundamental _ _ _ Ht3 with "[]") as "He3"; first
iFrame "#".
iApply ic_wand_r; iSplitL; first iApply ("He3" $! _ _ _ with "[$Hσ1']").
simpl.
iIntros (v3 _). iDestruct 1 as (σ2' v3') "(Hrd1 & Hσ2' & #Hvv3)".
iDestruct "Hrd1" as %Hrd1.
apply rtc_nsteps in Hrd1; destruct Hrd1 as [k1 Hrd1].
iApply ic_value; first by rewrite /= ?to_of_val.
iApply ic_rec; auto. iNext.
asimpl.
iApply (ic_bind (fill_item (BindLCtx _))).
iApply (ic_bind (fill_item (AppLCtx _))).
iApply ic_fst; auto.
iNext; iModIntro. asimpl.
iSpecialize ("Hww2" $! vv1 with "[]"); first iFrame "#".
iApply ic_wand_r; iSplitL; first iApply ("Hww2" $! _ _ _ with "[$Hσ2']").
simpl.
iIntros (v4 _). iDestruct 1 as (σ3' v4') "(Hrd2 & Hσ3' & #Hvv4)".
iDestruct "Hrd2" as %Hrd2.
apply rtc_nsteps in Hrd2; destruct Hrd2 as [k2 Hrd2].
iApply (ic_bind (fill_item (BindRCtx _))).
iApply ic_snd; auto.
iNext; iModIntro. asimpl.
iApply ic_value; first by rewrite /= ?to_of_val.
iExists _, (BindV _ _); iFrame. iSplit.
{ iPureIntro. simpl.
etrans.
{ eapply (rtc_l _ (_, _) (_, _)); simpl.
{ apply head_prim_step; econstructor; eauto. }
asimpl.
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).
- 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. }
asimpl.
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.
clear dependent γh γh' σ1' σ2' σ3'.
iIntros (γh γh' σ1') "[Hσ1' Hreg]".
iApply ic_runst_to_bind.
iApply ic_wand_r; iSplitL; first iApply ("Hvv4" $! _ _ _ with "[-]");
first iFrame. simpl.
iIntros (w1 _). iDestruct 1 as (σ2' w1') "(Hrd4 & Hσ2' & #Hww1 & Hreg)".
iDestruct "Hrd4" as %Hrd4.
apply rtc_nsteps in Hrd4; destruct Hrd4 as [k4 Hrd4].
iApply (ic_bind (fill_item RunSTCtx)).
iSpecialize ("Hvv3" $! (w1, w1') with "[]"); first iFrame "#".
iApply ic_wand_r; iSplitL "Hσ2' Hvv3";
first iApply ("Hvv3" $! _ _ _ with "[$Hσ2']"). simpl.
iIntros (w2 _). iDestruct 1 as (σ3' w2') "(Hrd5 & Hσ3' & #Hww2)".
iDestruct "Hrd5" as %Hrd5.
apply rtc_nsteps in Hrd5; destruct Hrd5 as [k5 Hrd5].
iApply ic_wand_r; iSplitL;
first iApply ("Hww2" $! _ _ _ with "[$Hσ3' $Hreg]"). simpl.
iIntros (w3 _). iDestruct 1 as (σ4' w3') "(Hrd6 & Hσ4' & #Hww3 & Hreg)".
iDestruct "Hrd6" as %Hrd6.
apply rtc_nsteps in Hrd6; destruct Hrd6 as [k6 Hrd6].
iExists _, _; iFrame; iFrame "#".
iPureIntro.
eapply nsteps_rtc.
eapply red_RunST_Bind'; eauto.
eapply nsteps_trans.
{ eapply (@nsteps_ctx det_lang (fill_item (RunSTCtx)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
eauto.
Qed.
Lemma ST_associativity_CR4 Γ e1' ρ τ1 e2' τ2 e3' τ3 :
Γ ⊢ₜ e1' : TST ρ τ1 → Γ ⊢ₜ e2' : (TArrow τ1 (TST ρ τ2)) →
Γ ⊢ₜ e3' : (TArrow τ2 (TST ρ τ3)) →
Γ ⊨ Rec (App (Rec (App (App (Rec (Rec (Bind (Var 1) (Var 3))))
e3'.[ren (+4)])
(App (Var 1) (Var 3)))) e2'.[ren (+2)])
≤ctx≤
Rec (App (Rec (App (App (Rec (Rec (Bind (Var 3) (Var 1))))
(App (Var 1) (Var 3)))
e3'.[ren (+4)])) e2'.[ren (+2)])
: TArrow τ1 (TST ρ τ3).
Proof.
intros Ht1 Ht2 Ht3.
eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_Rec; CTX_AppL _; CTX_Rec]).
{ repeat econstructor; eauto. by eapply (context_weakening [ _ ; _]). }
eapply (Commutativity_alt _ _ _ _ _ (Bind (Var 0) (Var 1))); eauto.
- repeat econstructor.
- eapply (context_weakening [_; _; _ ; _]); eauto.
- repeat econstructor.
Qed.
Lemma ST_associativity_LR5 `{ICG_ST Σ} `{LogRelG Σ} Γ ρ τ1 e2' τ2 e3' τ3 :
Γ ⊢ₜ e2' : (TArrow τ1 (TST ρ τ2)) →
Γ ⊢ₜ e3' : (TArrow τ2 (TST ρ τ3)) →
Γ ⊨ Rec (App (Rec (App (App (Rec (Rec (Bind (Var 3) (Var 1))))
(App (Var 1) (Var 3)))
e3'.[ren (+4)])) e2'.[ren (+2)])
≤log≤
(Rec (Bind (App e2'.[ren (+2)] (Var 1)) e3'.[ren (+2)]))
: TArrow τ1 (TST ρ τ3).
Proof.
iIntros (Ht2 Ht3 Δ vvs HΔ) "[#Hinv #Hvvs]".
iDestruct (interp_env_length with "[]") as %Hvvsl; first iFrame "#".
assert ( TArrow τ1 (TST ρ τ3) :: τ1 :: Γ
⊢ₜ (App (Rec (App (App (Rec (Rec (Bind (Var 3) (Var 1))))
(App (Var 1) (Var 3)))
e3'.[ren (+4)])) e2'.[ren (+2)])
: TST ρ τ3) as Ht4.
{ repeat econstructor; last eapply (context_weakening [ _ ; _]); eauto.
by eapply (context_weakening [_; _; _; _]). }
assert (TArrow τ1 (TST ρ τ3) :: τ1 :: Γ ⊢ₜ
(Bind (App e2'.[ren (+2)] (Var 1)) e3'.[ren (+2)]) : TST ρ τ3)
as Ht5.
{ repeat econstructor; eapply (context_weakening [ _ ; _]); eauto. }
iApply (weak_bin_log_related_rec _ _ _ _ _ _ _ HΔ); eauto.
iIntros (vv1 vv2) "[#Hinv #Hvvs]".
cbn [fmap list_fmap].
erewrite <- typed_subst_head_simpl_2; eauto;
last by rewrite /= map_length Hvvsl.
erewrite <- typed_subst_head_simpl_2; eauto;
last by rewrite /= map_length Hvvsl.
asimpl.
rewrite ?interp_env_cons.
iDestruct "Hvvs" as "(_ & #Hvv1 & #Hvvs)".
iApply (logrel_bind [AppRCtx (RecV _)]
[BindLCtx _; AppLCtx _] (TST _ _)).
iSplitL; first by iApply (binary_fundamental _ _ _ Ht2); auto.
iIntros (ww2) "#Hww2 /=".
iIntros (γh γh' σ1') "Hσ1' /=".
iApply ic_rec; eauto.
iNext. asimpl.
iRevert (γh γh' σ1') "Hσ1'".
iSpecialize ("Hww2" $! vv1 with "[]"); first iFrame "#".
iApply (@logrel_bind _ _ _ [AppLCtx _; AppRCtx (RecV _)]
[BindLCtx _] (TST _ _) _ _ (TST _ _)).
iSplitL; first iExact "Hww2".
iIntros (ww3) "#Hww3 /=".
iIntros (γh γh' σ1') "Hσ1' /=".
iApply (ic_bind (fill_item (AppLCtx _))).
iApply ic_rec; eauto.
iNext. asimpl. iApply ic_value; auto.
iRevert (γh γh' σ1') "Hσ1'". simpl.
iApply (@logrel_bind _ _ _ [AppRCtx (RecV _)]
[BindRCtx _] (TST _ _) _ _ (TArrow _ (TST _ _))).
iSplitL; first by iApply (binary_fundamental _ _ _ Ht3); auto.
iIntros (ww4) "#Hww4 /=".
iIntros (γh γh' σ1') "Hσ1' /=".
iApply ic_rec; eauto.
iNext. asimpl.
iApply ic_value; first by rewrite /= ?to_of_val.
iExists _, (BindV _ _); iFrame. iSplit.
{ iPureIntro. econstructor. }
asimpl.
iAlways.
clear γh γh' σ1'.
iIntros (γh γh' σ1') "[Hσ1' Hreg]".
iApply ic_runst_to_bind.
iApply ic_wand_r; iSplitL; first iApply ("Hww3" $! _ _ _ with "[-]");
first iFrame. simpl.
iIntros (w1 _). iDestruct 1 as (σ2' w1') "(Hrd4 & Hσ2' & #Hww1 & Hreg)".
iDestruct "Hrd4" as %Hrd4.
apply rtc_nsteps in Hrd4; destruct Hrd4 as [k4 Hrd4].
iApply (ic_bind (fill_item RunSTCtx)).
iSpecialize ("Hww4" $! (w1, w1') with "[]"); first iFrame "#".
iApply ic_wand_r; iSplitL "Hσ2' Hww4";
first iApply ("Hww4" $! _ _ _ with "[$Hσ2']"). simpl.
iIntros (w2 _). iDestruct 1 as (σ3' w2') "(Hrd5 & Hσ3' & #Hww2)".
iDestruct "Hrd5" as %Hrd5.
apply rtc_nsteps in Hrd5; destruct Hrd5 as [k5 Hrd5].
iApply ic_wand_r; iSplitL;
first iApply ("Hww2" $! _ _ _ with "[$Hσ3' $Hreg]"). simpl.
iIntros (w3 _). iDestruct 1 as (σ4' w3') "(Hrd6 & Hσ4' & #Hww5 & Hreg)".
iDestruct "Hrd6" as %Hrd6.
apply rtc_nsteps in Hrd6; destruct Hrd6 as [k6 Hrd6].
iExists _, _; iFrame; iFrame "#".
iPureIntro.
eapply nsteps_rtc.
eapply red_RunST_Bind'; eauto.
eapply nsteps_trans.
{ eapply (@nsteps_ctx det_lang (fill_item (RunSTCtx)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
eauto.
Qed.
Γ ⊢ₜ e1 : τ → Γ ⊢ₜ e2 : (TArrow τ (TST ρ τ')) →
Γ ⊨ (App e2 e1) ≤ctx≤ (Bind (Return e1) e2) : (TST ρ τ').
Proof.
intros Ht1 Ht2.
eapply ctx_refines_trans.
{ apply binary_soundness;
last by intros; eapply ST_right_id_LR1; eauto using binary_fundamental.
- econstructor; eauto.
- repeat econstructor; eauto. }
eapply ctx_refines_trans; first eapply ST_right_id_CR2; eauto.
eapply ctx_refines_trans; first eapply ST_right_id_CR4; eauto.
eapply ctx_refines_trans; first eapply ST_right_id_CR5; eauto.
apply binary_soundness.
- repeat econstructor; eauto.
- repeat econstructor; eauto.
- intros; eapply ST_right_id_LR6; eauto.
Qed.
Lemma ST_associativity_LR1
`{ICG_ST Σ} `{LogRelG Σ} Γ e1 e1' ρ τ1 e2 e2' τ2 e3 e3' τ3 :
Γ ⊨ e1 ≤log≤ e1' : TST ρ τ1 →
Γ ⊨ e2 ≤log≤ e2' : (TArrow τ1 (TST ρ τ2)) →
Γ ⊨ e3 ≤log≤ e3' : (TArrow τ2 (TST ρ τ3)) →
Γ ⊨ (Bind (Bind e1 e2) e3) ≤log≤
(Bind e1'
(App
(Rec (Rec (Bind (App (Fst (Var 3)) (Var 1)) (Snd (Var 3)))))
(Pair e2' e3')))
: (TST ρ τ3).
Proof.
iIntros (LR1 LR2 LR3 Δ vvs HΔ) "[#Hinv #Hvvs]". asimpl.
iApply (logrel_bind [BindLCtx _; BindLCtx _] [BindLCtx _] (TST _ _)).
iSplitL; first by iApply LR1; iFrame "#".
iIntros (vv1) "#Hvv1 /=".
iApply (logrel_bind [BindLCtx _; BindRCtx _]
[BindRCtx _; AppRCtx (RecV _); PairLCtx _] (TST _ _)).
iSplitL; first by iApply LR2; iFrame "#".
iIntros (vv2) "#Hvv2 /=".
iApply (logrel_bind [BindRCtx (BindV _ _)]
[BindRCtx _; AppRCtx (RecV _); PairRCtx _] (TST _ _)).
iSplitL; first by iApply LR3; iFrame "#".
iIntros (vv3) "#Hvv3 /=".
iIntros (γh γh' σ1') "Hσ1' /=".
iApply ic_value; first by rewrite /= ?to_of_val.
iExists _, (BindV _ (RecV _)); iFrame. iSplit.
{ iPureIntro.
eapply (nsteps_rtc _). simpl.
eapply (@nsteps_ctx det_lang (fill_item (BindRCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
eapply (nsteps_l _ _ (_, _) (_, _)); simpl.
{ apply head_prim_step; econstructor; by rewrite /= ?to_of_val. }
asimpl. econstructor. }
iAlways. asimpl.
clear γh γh' σ1'. iIntros (γh γh' σ1') "[Hσ1' Hreg] /=".
iApply (ic_runst_to_bind (BindV _ _)).
iApply ic_runst_to_bind.
iApply ic_wand_r; iSplitL.
{ iApply ("Hvv1" $! _ _ _ with "[-]"); first iFrame. }
simpl. iIntros (w1 _).
iDestruct 1 as (σ2' w1') "(Hrd1 & Hσ2' & #Hww1 & Hreg)".
iDestruct "Hrd1" as %Hrd1.
apply rtc_nsteps in Hrd1; destruct Hrd1 as [k1 Hrd1].
iSpecialize ("Hvv2" $! (w1, w1') with "[]"); first iFrame "#".
iApply (ic_bind (fill_item RunSTCtx)).
iApply ic_wand_r; iSplitL "Hσ2' Hvv2".
{ iApply ("Hvv2" $! _ _ _ with "[-]"); first iFrame. }
simpl. iIntros (w2 _).
iDestruct 1 as (σ3' w2') "(Hrd2 & Hσ3' & #Hww2)".
iDestruct "Hrd2" as %Hrd2.
apply rtc_nsteps in Hrd2; destruct Hrd2 as [k2 Hrd2].
iApply ic_wand_r; iSplitL.
{ iApply ("Hww2" $! _ _ _ with "[-]"); first iFrame. }
simpl. iIntros (u _).
iDestruct 1 as (σ4' u') "(Hrd3 & Hσ4' & #Huu & Hreg)".
iDestruct "Hrd3" as %Hrd3.
apply rtc_nsteps in Hrd3; destruct Hrd3 as [k3 Hrd3].
iSpecialize ("Hvv3" $! (u, u') with "[]"); first iFrame "#".
iApply (ic_bind (fill_item RunSTCtx)).
iApply ic_wand_r; iSplitL "Hσ4' Hvv3".
{ iApply ("Hvv3" $! _ _ _ with "[-]"); first iFrame. }
simpl. iIntros (w3 _).
iDestruct 1 as (σ5' w3') "(Hrd4 & Hσ5' & #Hww3)".
iDestruct "Hrd4" as %Hrd4.
apply rtc_nsteps in Hrd4; destruct Hrd4 as [k4 Hrd4].
iApply ic_wand_r; iSplitL.
{ iApply ("Hww3" $! _ _ _ with "[-]"); first iFrame. }
simpl. iIntros (u2 _).
iDestruct 1 as (σ6' u2') "(Hrd5 & Hσ6' & #Huu2 & Hreg)".
iDestruct "Hrd5" as %Hrd5.
apply rtc_nsteps in Hrd5; destruct Hrd5 as [k5 Hrd5].
iExists _, _; iFrame; iFrame "#".
iPureIntro.
etrans.
- eapply nsteps_rtc.
eapply (red_RunST_Bind' _ _ (RecV _)); eauto.
eapply nsteps_trans.
{ eapply (@nsteps_ctx det_lang (fill_item (RunSTCtx)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
eapply (nsteps_l _ _ (_, _) (_, _)); simpl.
{ apply head_prim_step; econstructor; eauto. }
asimpl.
eapply nsteps_trans.
- eapply (@nsteps_ctx det_lang (fill_item (BindLCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
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. }
econstructor.
- eapply (@nsteps_ctx det_lang (fill_item (BindLCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
eapply nsteps_trans.
{ eapply (@nsteps_ctx det_lang (fill_item (RunSTCtx)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
eapply (@nsteps_ctx det_lang (fill_item (BindRCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
eapply (nsteps_l _ _ (_, _) (_, _)); simpl.
{ apply head_prim_step; econstructor; eauto. }
econstructor. }
simpl.
eapply red_RunST_Bind'; eauto.
eapply nsteps_trans.
{ eapply (@nsteps_ctx det_lang (fill_item (RunSTCtx)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
simpl. eauto.
- econstructor.
Qed.
Lemma ST_associativity_CR2 Γ e1' ρ τ1 e2' τ2 e3' τ3 :
Γ ⊢ₜ e1' : TST ρ τ1 → Γ ⊢ₜ e2' : (TArrow τ1 (TST ρ τ2)) →
Γ ⊢ₜ e3' : (TArrow τ2 (TST ρ τ3)) →
Γ ⊨ (Bind e1'
(App
(Rec (Rec (Bind (App (Fst (Var 3)) (Var 1)) (Snd (Var 3)))))
(Pair e2' e3'))) ≤ctx≤
(Bind e1'
(Rec (App (Rec (Bind (App (Fst (Var 1)) (Var 3)) (Snd (Var 1))))
(Pair e2'.[ren (+2)] e3'.[ren (+2)]))))
: (TST ρ τ3).
Proof.
intros Ht1 Ht2 Ht3.
eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_BindR _]).
{ repeat econstructor; eauto. }
eapply (Rec_Hoisting _ (Pair e2' e3') _
(Bind (App (Fst (Var 2)) (Var 1)) (Snd (Var 2)))).
- repeat econstructor; eauto; eapply (context_weakening [ _ ; _]); eauto.
- repeat econstructor.
Qed.
Lemma ST_associativity_LR3 `{ICG_ST Σ} `{LogRelG Σ} Γ ρ τ1 e2' τ2 e3' τ3 :
Γ ⊢ₜ e2' : (TArrow τ1 (TST ρ τ2)) →
Γ ⊢ₜ e3' : (TArrow τ2 (TST ρ τ3)) →
Γ ⊨ Rec (App (Rec (Bind (App (Fst (Var 1)) (Var 3)) (Snd (Var 1))))
(Pair e2'.[ren (+2)] e3'.[ren (+2)]))
≤log≤
Rec (App (Rec (App (App (Rec (Rec (Bind (Var 1) (Var 3))))
e3'.[ren (+4)])
(App (Var 1) (Var 3)))) e2'.[ren (+2)])
: TArrow τ1 (TST ρ τ3).
Proof.
iIntros (Ht2 Ht3 Δ vvs HΔ) "[#Hinv #Hvvs]".
iDestruct (interp_env_length with "[]") as %Hvvsl; first iFrame "#".
assert ( TArrow τ1 (TST ρ τ3) :: τ1 :: Γ
⊢ₜ App (Rec (Bind (App (Fst (Var 1)) (Var 3)) (Snd (Var 1))))
(Pair e2'.[ren (+2)] e3'.[ren (+2)])
: TST ρ τ3) as Ht4.
{ repeat econstructor; eapply (context_weakening [ _ ; _]); eauto. }
assert (TArrow τ1 (TST ρ τ3) :: τ1 :: Γ ⊢ₜ
(App
(Rec
(App (App (Rec (Rec (Bind (Var 1) (Var 3)))) e3'.[ren (+4)])
(App (Var 1) (Var 3)))) e2'.[ren (+2)]) : TST ρ τ3)
as Ht5.
{ repeat econstructor; last eapply (context_weakening [ _ ; _]); eauto.
by eapply (context_weakening [_; _; _; _]). }
iApply (weak_bin_log_related_rec _ _ _ _ _ _ _ HΔ); eauto.
iIntros (vv1 vv2) "[#Hinv #Hvvs]".
cbn [fmap list_fmap].
erewrite <- typed_subst_head_simpl_2; eauto;
last by rewrite /= map_length Hvvsl.
erewrite <- typed_subst_head_simpl_2; eauto;
last by rewrite /= map_length Hvvsl.
asimpl.
rewrite ?interp_env_cons.
iDestruct "Hvvs" as "(_ & #Hvv1 & #Hvvs)".
iApply (logrel_bind [AppRCtx (RecV _); PairLCtx _]
[AppRCtx (RecV _)] (TST _ _)).
iSplitL; first by iApply (binary_fundamental _ _ _ Ht2); auto.
iIntros (ww2) "#Hww2 /=".
iIntros (γh γh' σ1') "Hσ1' /=".
iApply (ic_bind (fill_item (AppRCtx (RecV _)))).
iApply (ic_bind (fill_item (PairRCtx _))).
iPoseProof (binary_fundamental _ _ _ Ht3 with "[]") as "He3"; first
iFrame "#".
iApply ic_wand_r; iSplitL; first iApply ("He3" $! _ _ _ with "[$Hσ1']").
simpl.
iIntros (v3 _). iDestruct 1 as (σ2' v3') "(Hrd1 & Hσ2' & #Hvv3)".
iDestruct "Hrd1" as %Hrd1.
apply rtc_nsteps in Hrd1; destruct Hrd1 as [k1 Hrd1].
iApply ic_value; first by rewrite /= ?to_of_val.
iApply ic_rec; auto. iNext.
asimpl.
iApply (ic_bind (fill_item (BindLCtx _))).
iApply (ic_bind (fill_item (AppLCtx _))).
iApply ic_fst; auto.
iNext; iModIntro. asimpl.
iSpecialize ("Hww2" $! vv1 with "[]"); first iFrame "#".
iApply ic_wand_r; iSplitL; first iApply ("Hww2" $! _ _ _ with "[$Hσ2']").
simpl.
iIntros (v4 _). iDestruct 1 as (σ3' v4') "(Hrd2 & Hσ3' & #Hvv4)".
iDestruct "Hrd2" as %Hrd2.
apply rtc_nsteps in Hrd2; destruct Hrd2 as [k2 Hrd2].
iApply (ic_bind (fill_item (BindRCtx _))).
iApply ic_snd; auto.
iNext; iModIntro. asimpl.
iApply ic_value; first by rewrite /= ?to_of_val.
iExists _, (BindV _ _); iFrame. iSplit.
{ iPureIntro. simpl.
etrans.
{ eapply (rtc_l _ (_, _) (_, _)); simpl.
{ apply head_prim_step; econstructor; eauto. }
asimpl.
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).
- 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. }
asimpl.
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.
clear dependent γh γh' σ1' σ2' σ3'.
iIntros (γh γh' σ1') "[Hσ1' Hreg]".
iApply ic_runst_to_bind.
iApply ic_wand_r; iSplitL; first iApply ("Hvv4" $! _ _ _ with "[-]");
first iFrame. simpl.
iIntros (w1 _). iDestruct 1 as (σ2' w1') "(Hrd4 & Hσ2' & #Hww1 & Hreg)".
iDestruct "Hrd4" as %Hrd4.
apply rtc_nsteps in Hrd4; destruct Hrd4 as [k4 Hrd4].
iApply (ic_bind (fill_item RunSTCtx)).
iSpecialize ("Hvv3" $! (w1, w1') with "[]"); first iFrame "#".
iApply ic_wand_r; iSplitL "Hσ2' Hvv3";
first iApply ("Hvv3" $! _ _ _ with "[$Hσ2']"). simpl.
iIntros (w2 _). iDestruct 1 as (σ3' w2') "(Hrd5 & Hσ3' & #Hww2)".
iDestruct "Hrd5" as %Hrd5.
apply rtc_nsteps in Hrd5; destruct Hrd5 as [k5 Hrd5].
iApply ic_wand_r; iSplitL;
first iApply ("Hww2" $! _ _ _ with "[$Hσ3' $Hreg]"). simpl.
iIntros (w3 _). iDestruct 1 as (σ4' w3') "(Hrd6 & Hσ4' & #Hww3 & Hreg)".
iDestruct "Hrd6" as %Hrd6.
apply rtc_nsteps in Hrd6; destruct Hrd6 as [k6 Hrd6].
iExists _, _; iFrame; iFrame "#".
iPureIntro.
eapply nsteps_rtc.
eapply red_RunST_Bind'; eauto.
eapply nsteps_trans.
{ eapply (@nsteps_ctx det_lang (fill_item (RunSTCtx)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
eauto.
Qed.
Lemma ST_associativity_CR4 Γ e1' ρ τ1 e2' τ2 e3' τ3 :
Γ ⊢ₜ e1' : TST ρ τ1 → Γ ⊢ₜ e2' : (TArrow τ1 (TST ρ τ2)) →
Γ ⊢ₜ e3' : (TArrow τ2 (TST ρ τ3)) →
Γ ⊨ Rec (App (Rec (App (App (Rec (Rec (Bind (Var 1) (Var 3))))
e3'.[ren (+4)])
(App (Var 1) (Var 3)))) e2'.[ren (+2)])
≤ctx≤
Rec (App (Rec (App (App (Rec (Rec (Bind (Var 3) (Var 1))))
(App (Var 1) (Var 3)))
e3'.[ren (+4)])) e2'.[ren (+2)])
: TArrow τ1 (TST ρ τ3).
Proof.
intros Ht1 Ht2 Ht3.
eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_Rec; CTX_AppL _; CTX_Rec]).
{ repeat econstructor; eauto. by eapply (context_weakening [ _ ; _]). }
eapply (Commutativity_alt _ _ _ _ _ (Bind (Var 0) (Var 1))); eauto.
- repeat econstructor.
- eapply (context_weakening [_; _; _ ; _]); eauto.
- repeat econstructor.
Qed.
Lemma ST_associativity_LR5 `{ICG_ST Σ} `{LogRelG Σ} Γ ρ τ1 e2' τ2 e3' τ3 :
Γ ⊢ₜ e2' : (TArrow τ1 (TST ρ τ2)) →
Γ ⊢ₜ e3' : (TArrow τ2 (TST ρ τ3)) →
Γ ⊨ Rec (App (Rec (App (App (Rec (Rec (Bind (Var 3) (Var 1))))
(App (Var 1) (Var 3)))
e3'.[ren (+4)])) e2'.[ren (+2)])
≤log≤
(Rec (Bind (App e2'.[ren (+2)] (Var 1)) e3'.[ren (+2)]))
: TArrow τ1 (TST ρ τ3).
Proof.
iIntros (Ht2 Ht3 Δ vvs HΔ) "[#Hinv #Hvvs]".
iDestruct (interp_env_length with "[]") as %Hvvsl; first iFrame "#".
assert ( TArrow τ1 (TST ρ τ3) :: τ1 :: Γ
⊢ₜ (App (Rec (App (App (Rec (Rec (Bind (Var 3) (Var 1))))
(App (Var 1) (Var 3)))
e3'.[ren (+4)])) e2'.[ren (+2)])
: TST ρ τ3) as Ht4.
{ repeat econstructor; last eapply (context_weakening [ _ ; _]); eauto.
by eapply (context_weakening [_; _; _; _]). }
assert (TArrow τ1 (TST ρ τ3) :: τ1 :: Γ ⊢ₜ
(Bind (App e2'.[ren (+2)] (Var 1)) e3'.[ren (+2)]) : TST ρ τ3)
as Ht5.
{ repeat econstructor; eapply (context_weakening [ _ ; _]); eauto. }
iApply (weak_bin_log_related_rec _ _ _ _ _ _ _ HΔ); eauto.
iIntros (vv1 vv2) "[#Hinv #Hvvs]".
cbn [fmap list_fmap].
erewrite <- typed_subst_head_simpl_2; eauto;
last by rewrite /= map_length Hvvsl.
erewrite <- typed_subst_head_simpl_2; eauto;
last by rewrite /= map_length Hvvsl.
asimpl.
rewrite ?interp_env_cons.
iDestruct "Hvvs" as "(_ & #Hvv1 & #Hvvs)".
iApply (logrel_bind [AppRCtx (RecV _)]
[BindLCtx _; AppLCtx _] (TST _ _)).
iSplitL; first by iApply (binary_fundamental _ _ _ Ht2); auto.
iIntros (ww2) "#Hww2 /=".
iIntros (γh γh' σ1') "Hσ1' /=".
iApply ic_rec; eauto.
iNext. asimpl.
iRevert (γh γh' σ1') "Hσ1'".
iSpecialize ("Hww2" $! vv1 with "[]"); first iFrame "#".
iApply (@logrel_bind _ _ _ [AppLCtx _; AppRCtx (RecV _)]
[BindLCtx _] (TST _ _) _ _ (TST _ _)).
iSplitL; first iExact "Hww2".
iIntros (ww3) "#Hww3 /=".
iIntros (γh γh' σ1') "Hσ1' /=".
iApply (ic_bind (fill_item (AppLCtx _))).
iApply ic_rec; eauto.
iNext. asimpl. iApply ic_value; auto.
iRevert (γh γh' σ1') "Hσ1'". simpl.
iApply (@logrel_bind _ _ _ [AppRCtx (RecV _)]
[BindRCtx _] (TST _ _) _ _ (TArrow _ (TST _ _))).
iSplitL; first by iApply (binary_fundamental _ _ _ Ht3); auto.
iIntros (ww4) "#Hww4 /=".
iIntros (γh γh' σ1') "Hσ1' /=".
iApply ic_rec; eauto.
iNext. asimpl.
iApply ic_value; first by rewrite /= ?to_of_val.
iExists _, (BindV _ _); iFrame. iSplit.
{ iPureIntro. econstructor. }
asimpl.
iAlways.
clear γh γh' σ1'.
iIntros (γh γh' σ1') "[Hσ1' Hreg]".
iApply ic_runst_to_bind.
iApply ic_wand_r; iSplitL; first iApply ("Hww3" $! _ _ _ with "[-]");
first iFrame. simpl.
iIntros (w1 _). iDestruct 1 as (σ2' w1') "(Hrd4 & Hσ2' & #Hww1 & Hreg)".
iDestruct "Hrd4" as %Hrd4.
apply rtc_nsteps in Hrd4; destruct Hrd4 as [k4 Hrd4].
iApply (ic_bind (fill_item RunSTCtx)).
iSpecialize ("Hww4" $! (w1, w1') with "[]"); first iFrame "#".
iApply ic_wand_r; iSplitL "Hσ2' Hww4";
first iApply ("Hww4" $! _ _ _ with "[$Hσ2']"). simpl.
iIntros (w2 _). iDestruct 1 as (σ3' w2') "(Hrd5 & Hσ3' & #Hww2)".
iDestruct "Hrd5" as %Hrd5.
apply rtc_nsteps in Hrd5; destruct Hrd5 as [k5 Hrd5].
iApply ic_wand_r; iSplitL;
first iApply ("Hww2" $! _ _ _ with "[$Hσ3' $Hreg]"). simpl.
iIntros (w3 _). iDestruct 1 as (σ4' w3') "(Hrd6 & Hσ4' & #Hww5 & Hreg)".
iDestruct "Hrd6" as %Hrd6.
apply rtc_nsteps in Hrd6; destruct Hrd6 as [k6 Hrd6].
iExists _, _; iFrame; iFrame "#".
iPureIntro.
eapply nsteps_rtc.
eapply red_RunST_Bind'; eauto.
eapply nsteps_trans.
{ eapply (@nsteps_ctx det_lang (fill_item (RunSTCtx)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang). }
eauto.
Qed.
Theorem ST_associativity Γ e1 ρ τ1 e2 τ2 e3 τ3 :
Γ ⊢ₜ e1 : TST ρ τ1 → Γ ⊢ₜ e2 : (TArrow τ1 (TST ρ τ2)) →
Γ ⊢ₜ e3 : (TArrow τ2 (TST ρ τ3)) →
Γ ⊨ (Bind (Bind e1 e2) e3) ≤ctx≤
(Bind e1 (Rec (Bind (App e2.[ren (+2)] (Var 1)) e3.[ren (+2)])))
: (TST ρ τ3).
Proof.
intros Ht1 Ht2 Ht3.
eapply ctx_refines_trans.
{ eapply binary_soundness; last (intros; eapply ST_associativity_LR1;
eauto 2 using binary_fundamental).
- repeat econstructor; eauto.
- repeat econstructor; eauto. }
eapply ctx_refines_trans; first eapply ST_associativity_CR2; eauto.
eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_BindR _]).
{ repeat econstructor; eauto. }
eapply ctx_refines_trans.
{ eapply binary_soundness; last (intros; eapply ST_associativity_LR3; eauto).
- repeat econstructor; eapply (context_weakening [ _ ; _]); eauto.
- repeat econstructor; last eapply (context_weakening [ _ ; _]); eauto.
eapply (context_weakening [_; _; _ ; _]); eauto. }
eapply ctx_refines_trans; first eapply ST_associativity_CR4; eauto.
eapply binary_soundness; last (intros; eapply ST_associativity_LR5; eauto).
- repeat econstructor; last eapply (context_weakening [ _ ; _]); eauto.
eapply (context_weakening [_; _; _ ; _]); eauto.
- repeat econstructor; eapply (context_weakening [ _ ; _]); eauto.
Qed.
Γ ⊢ₜ e1 : TST ρ τ1 → Γ ⊢ₜ e2 : (TArrow τ1 (TST ρ τ2)) →
Γ ⊢ₜ e3 : (TArrow τ2 (TST ρ τ3)) →
Γ ⊨ (Bind (Bind e1 e2) e3) ≤ctx≤
(Bind e1 (Rec (Bind (App e2.[ren (+2)] (Var 1)) e3.[ren (+2)])))
: (TST ρ τ3).
Proof.
intros Ht1 Ht2 Ht3.
eapply ctx_refines_trans.
{ eapply binary_soundness; last (intros; eapply ST_associativity_LR1;
eauto 2 using binary_fundamental).
- repeat econstructor; eauto.
- repeat econstructor; eauto. }
eapply ctx_refines_trans; first eapply ST_associativity_CR2; eauto.
eapply (ctx_refines_ctx_closed _ _ _ _ [CTX_BindR _]).
{ repeat econstructor; eauto. }
eapply ctx_refines_trans.
{ eapply binary_soundness; last (intros; eapply ST_associativity_LR3; eauto).
- repeat econstructor; eapply (context_weakening [ _ ; _]); eauto.
- repeat econstructor; last eapply (context_weakening [ _ ; _]); eauto.
eapply (context_weakening [_; _; _ ; _]); eauto. }
eapply ctx_refines_trans; first eapply ST_associativity_CR4; eauto.
eapply binary_soundness; last (intros; eapply ST_associativity_LR5; eauto).
- repeat econstructor; last eapply (context_weakening [ _ ; _]); eauto.
eapply (context_weakening [_; _; _ ; _]); eauto.
- repeat econstructor; eapply (context_weakening [ _ ; _]); eauto.
Qed.