RunST.lam_beta.lam_beta_part1
From RunST Require Import IC base lang rules typing reduction regions
future contextual_refinement saved_pred ST_Lang_reduction
logrel_shared soundness logrel fundamental.
From iris.proofmode Require Import tactics.
Hint Resolve to_of_val.
Lemma lam_beta_var ξ τ τ' Γ x e2' :
(ξ ++ τ :: Γ) !! x = Some τ' ->
ξ ++ Γ ⊢ₜ e2'.[ren (+length ξ)] : τ →
ξ ++ Γ
⊨ App (Rec (Var x).[ren (lookup_n (swap_list 1 (length ξ)))].[ren (+1)])
e2'.[ren (+length ξ)]
≤ctx≤ (Var x).[upn (length ξ) (e2' .: ids)] : τ'.
Proof.
iIntros (Hctxlu Ht'). asimpl.
destruct (decide (x = length ξ)) as [|Hx]; first subst.
- replace (upn (length ξ) (e2' .: ids) (length ξ)) with
e2'.[ren (+length ξ)]; last first.
{ rewrite iter_up; destruct lt_dec; first omega.
replace (length ξ - length ξ) with 0 by omega.
by asimpl. }
replace (ids (S (lookup_n (swap_list 1 (length ξ)) (length ξ)))) with
(Var 1); last first.
{ rewrite /lookup_n /swap_list lookup_app_r seq_length //=.
by replace (length ξ - length ξ) with 0 by omega. }
rewrite lookup_app_r //= in Hctxlu.
replace (length ξ - length ξ) with 0 in Hctxlu by omega.
inversion Hctxlu; subst.
apply binary_soundness; eauto.
+ repeat econstructor; eauto.
+ intros.
iIntros (Δ vvs HΔ) "[#Hinv #Hvvs]".
iApply (logrel_bind [AppRCtx (RecV _)] []). iSplitL.
iApply binary_fundamental; eauto.
iIntros ([v1 v2]) "#Hvv". asimpl.
iIntros (γh γh' σ1') "Hσ'". asimpl.
iApply ic_rec; eauto.
iNext; asimpl.
iApply ic_value. eauto. iExists _,_. iFrame; iFrame "#".
iPureIntro. constructor.
- unfold var in *.
assert (∃ y, (upn (length ξ) (e2' .: ids) x) = (Var y) ∧
(ids (S (lookup_n (swap_list 1 (length ξ)) x)))
= (Var (S (S y))) ∧
(ξ ++ Γ) !! y = Some τ') as Hvar.
{ rewrite iter_up; destruct lt_dec.
- rewrite lookup_app_l in Hctxlu; auto.
exists x. rewrite /lookup_n /swap_list !lookup_app_l ?seq_length; try omega.
rewrite lookup_seq; auto.
- rewrite lookup_app_r in Hctxlu; auto with omega.
rewrite /lookup_n /swap_list lookup_app_r ?seq_length; try omega.
destruct (x - length ξ) as [|z] eqn:Heqx; first omega.
asimpl.
exists (length ξ + z).
replace x with (length ξ + S z) by omega.
repeat split; auto.
rewrite lookup_app_r; auto with omega.
by replace (length ξ + z - length ξ) with z by omega. }
clear Hctxlu Hx. destruct Hvar as (y & Heq1 & Heq2 & Hctxlu).
rewrite Heq1 Heq2.
apply binary_soundness; eauto.
+ repeat econstructor; eauto.
+ by econstructor.
+ intros.
iIntros (Δ vvs HΔ) "[#Hinv #Hvvs]". asimpl.
iIntros (γh γh' σ1') "Hσ'". asimpl.
rewrite ic_eq /ic_def.
iIntros (σ1 σ2 v n) "(Hrd & Hσ1)".
iDestruct "Hrd" as %Hrd. simpl in *.
iMod (allocate_full_state σ1') as (γh'') "(Hσ1' & Hfσ1')".
apply (nsteps_bind (fill_item (AppRCtx (RecV _)))) in Hrd
as [k [σ3 [w [Hkltn [He2 Hectx]]]]].
iPoseProof (binary_fundamental _ _ _ Ht' with "[]") as "Hx";
first iFrame "#".
iSpecialize ("Hx" $! γh γh'' σ1' with "Hσ1'").
rewrite ic_eq /ic_def. asimpl.
iSpecialize ("Hx" $! _ _ _ _ with "[$Hσ1]"); eauto.
iMod "Hx" as "(Hx & Hσ2)".
iDestruct "Hx" as (σ2' v') "(Hrtc & Hσ2' & Hrel)".
iModIntro; iFrame.
simpl in Hectx.
apply rec_red_step in Hectx; asimpl in Hectx.
destruct Hectx as [m [HnkSm Hnsteps]].
iDestruct (interp_env_Some_l with "[]") as ([v1 v2]) "[Hvveq Hvv]"; eauto.
iDestruct "Hvveq" as %Hvveq.
rewrite /env_subst list_lookup_fmap in Hnsteps.
rewrite /env_subst list_lookup_fmap.
rewrite Hvveq in Hnsteps.
apply (nsteps_val ST_lang) in Hnsteps as [? [?%of_val_inj ?]]; subst.
iFrame.
iExists _,_. iFrame; iFrame "#".
iPureIntro. rewrite Hvveq. asimpl. econstructor.
Qed.
Lemma lam_beta_basic_type v τ τ' Γ e2' :
Γ ⊢ₜ e2' : τ ->
(∀ Γ, Γ ⊢ₜ of_val v : τ') ->
(∀ f, (of_val v).[f] = (of_val v)) →
(∀ `{ICG_ST Σ} `{LogRelG Σ} Δ, True ⊢ ⟦ τ' ⟧ Δ (v,v)) →
Γ ⊨ App (Rec (of_val v)) e2' ≤ctx≤ of_val v : τ'.
Proof.
iIntros (Ht' Htv Hsubst Hvv). apply binary_soundness; eauto.
* repeat econstructor; eauto.
* iIntros (Σ ST Hlr Δ vvs HΔ) "[#Hinv #Hvvs]"; asimpl.
iIntros (γh γh' σ1') "Hσ'". asimpl. rewrite ic_eq /ic_def.
iIntros (σ1 σ2 w n) "(Hrd & Hσ1)".
iDestruct "Hrd" as %Hrd; simpl in *.
iMod (allocate_full_state σ1') as (γh'') "(Hσ1' & Hfσ1')".
apply (nsteps_bind (fill_item (AppRCtx (RecV _)))) in Hrd
as [k [σ3 [w2 [Hkltn [He2 Hectx]]]]].
iPoseProof (binary_fundamental _ _ _ Ht' with "[]") as "Hx"; first iFrame "#".
iSpecialize ("Hx" $! γh γh'' σ1' with "Hσ1'"). rewrite ic_eq /ic_def.
iSpecialize ("Hx" $! _ _ _ _ with "[$Hσ1]"); eauto.
iMod "Hx" as "(Hx & Hσ2)".
iDestruct "Hx" as (σ2' v') "(Hrtc & Hσ2' & Hrel)".
iModIntro; iFrame; simpl in *.
apply rec_red_step in Hectx; asimpl in Hectx.
destruct Hectx as [m [HnkSm Hnsteps]]; rewrite Hsubst in Hnsteps.
apply (nsteps_val ST_lang) in Hnsteps as [? [?%of_val_inj ?]];
subst; iFrame.
iExists _,_. iFrame; iFrame "#". repeat iSplit; eauto; last iApply Hvv.
iPureIntro. rewrite Hsubst. econstructor.
Qed.
Lemma lam_beta_push_lam K Γ e2' τ e τ' τ'' :
(∀ f e, (fill_item K e).[f] = fill_item K (e.[f])) →
(∀ Γ e, Γ ⊢ₜ e : τ' → Γ ⊢ₜ fill_item K e : τ'') →
(∀ `{ICG_ST Σ} `{LogRelG Σ} Γ e e', Γ ⊨ e ≤log≤ e' : τ' →
Γ ⊨ fill_item K e ≤log≤ fill_item K e' : τ'') →
τ :: Γ ⊢ₜ e : τ' →
Γ ⊢ₜ e2' : τ →
Γ ⊨ App (Rec ((fill_item K e.[ren (+1)]))) e2' ≤ctx≤
fill_item K (App (Rec e.[ren (+1)]) e2') : τ''.
Proof.
intros Hsubst HKt HKc Ht1 Ht2.
apply binary_soundness; eauto.
- repeat econstructor; eauto.
eapply HKt. by apply (context_weakening [_]).
- eapply HKt. repeat econstructor; eauto.
by apply (context_weakening [_]).
- iIntros (Σ ST Hlr Δ vvs HΔ) "[#Hinv #Hvvs]"; asimpl.
rewrite !Hsubst. asimpl.
iDestruct (interp_env_length with "[]") as %Hvvsl; eauto.
iApply (logrel_bind [AppRCtx (RecV _)] [K; AppRCtx (RecV _)]). iSplitL.
{ iApply (binary_fundamental _ _ _ Ht2); iFrame "#". }
iIntros (vv) "#Hvv".
iIntros (γh γh' σ1') "Hσ1' /=".
iApply ic_rec; eauto.
iNext. rewrite Hsubst. asimpl.
erewrite typed_subst_head_simpl; eauto;
last by rewrite /= fmap_length Hvvsl.
rewrite -Hsubst.
pose proof (binary_fundamental _ _ _ Ht1) as He.
apply HKc in He.
iPoseProof (He _ (vv :: vvs)with "[]") as "He";
first rewrite interp_env_cons; iFrame "#".
iSpecialize ("He" $! _ _ _ with "[$Hσ1']").
iApply ic_wand_r; iSplitL; first iApply "He".
simpl. iIntros (v1 _). iDestruct 1 as (σ2' v1') "(Hrd & Hσ2' & Hvv1)".
iDestruct "Hrd" as %Hrd.
apply rtc_nsteps in Hrd; destruct Hrd as [n Hrd].
iExists _, _; iFrame.
iPureIntro.
etrans.
eapply nsteps_rtc.
eapply (nsteps_l _ _ (_, _) (_, _)); simpl.
{ eapply (Ectx_step _ _ _ _ _ [K]); eauto.
econstructor; eauto. }
asimpl.
erewrite typed_subst_head_simpl; eauto;
last by rewrite /= fmap_length Hvvsl.
rewrite -Hsubst. eauto.
econstructor.
Qed.
Lemma lam_beta_push_lam_RunST Γ e2' τ e τ'' :
subst (ren (+1)) <$> τ :: Γ ⊢ₜ e : TST (TVar 0) τ''.[ren (+1)] →
Γ ⊢ₜ e2' : τ →
Γ ⊨ App (Rec ((RunST e.[ren (+1)]))) e2' ≤ctx≤
RunST (App (Rec e.[ren (+1)]) e2') : τ''.
Proof.
intros Ht1 Ht2.
apply binary_soundness; eauto.
- repeat econstructor; eauto.
apply (context_weakening [_]); eauto.
- repeat econstructor; last eapply context_rename; eauto.
apply (context_weakening [_]); eauto.
- iIntros (Σ ST Hlr Δ vvs HΔ) "[#Hinv #Hvvs]"; asimpl.
iDestruct (interp_env_length with "[]") as %Hvvsl; eauto.
iApply (logrel_bind [AppRCtx (RecV _)] [RunSTCtx; AppRCtx (RecV _)]).
iSplitL. { iApply (binary_fundamental _ _ _ Ht2); iFrame "#". }
iIntros (vv) "#Hvv".
iIntros (γh γh' σ1') "Hσ1' /=".
iApply ic_rec; eauto.
iNext. asimpl.
erewrite typed_subst_head_simpl; eauto;
last by rewrite /= !fmap_length Hvvsl.
pose proof (binary_fundamental _ _ _ Ht1) as He.
eapply bin_log_related_runST in He.
iPoseProof (He _ (vv :: vvs)with "[]") as "He";
first rewrite interp_env_cons; iFrame "#".
iSpecialize ("He" $! _ _ _ with "[$Hσ1']").
iApply ic_wand_r; iSplitL; first iApply "He".
simpl. iIntros (v1 _). iDestruct 1 as (σ2' v1') "(Hrd & Hσ2' & Hvv1)".
iDestruct "Hrd" as %Hrd.
apply rtc_nsteps in Hrd; destruct Hrd as [n Hrd].
iExists _, _; iFrame.
iPureIntro.
etrans.
eapply nsteps_rtc.
eapply (nsteps_l _ _ (_, _) (_, _)); simpl.
{ eapply (Ectx_step _ _ _ _ _ [RunSTCtx]); eauto.
econstructor; eauto. }
asimpl.
erewrite typed_subst_head_simpl; eauto;
last by rewrite /= !fmap_length Hvvsl.
econstructor.
Qed.
future contextual_refinement saved_pred ST_Lang_reduction
logrel_shared soundness logrel fundamental.
From iris.proofmode Require Import tactics.
Hint Resolve to_of_val.
Lemma lam_beta_var ξ τ τ' Γ x e2' :
(ξ ++ τ :: Γ) !! x = Some τ' ->
ξ ++ Γ ⊢ₜ e2'.[ren (+length ξ)] : τ →
ξ ++ Γ
⊨ App (Rec (Var x).[ren (lookup_n (swap_list 1 (length ξ)))].[ren (+1)])
e2'.[ren (+length ξ)]
≤ctx≤ (Var x).[upn (length ξ) (e2' .: ids)] : τ'.
Proof.
iIntros (Hctxlu Ht'). asimpl.
destruct (decide (x = length ξ)) as [|Hx]; first subst.
- replace (upn (length ξ) (e2' .: ids) (length ξ)) with
e2'.[ren (+length ξ)]; last first.
{ rewrite iter_up; destruct lt_dec; first omega.
replace (length ξ - length ξ) with 0 by omega.
by asimpl. }
replace (ids (S (lookup_n (swap_list 1 (length ξ)) (length ξ)))) with
(Var 1); last first.
{ rewrite /lookup_n /swap_list lookup_app_r seq_length //=.
by replace (length ξ - length ξ) with 0 by omega. }
rewrite lookup_app_r //= in Hctxlu.
replace (length ξ - length ξ) with 0 in Hctxlu by omega.
inversion Hctxlu; subst.
apply binary_soundness; eauto.
+ repeat econstructor; eauto.
+ intros.
iIntros (Δ vvs HΔ) "[#Hinv #Hvvs]".
iApply (logrel_bind [AppRCtx (RecV _)] []). iSplitL.
iApply binary_fundamental; eauto.
iIntros ([v1 v2]) "#Hvv". asimpl.
iIntros (γh γh' σ1') "Hσ'". asimpl.
iApply ic_rec; eauto.
iNext; asimpl.
iApply ic_value. eauto. iExists _,_. iFrame; iFrame "#".
iPureIntro. constructor.
- unfold var in *.
assert (∃ y, (upn (length ξ) (e2' .: ids) x) = (Var y) ∧
(ids (S (lookup_n (swap_list 1 (length ξ)) x)))
= (Var (S (S y))) ∧
(ξ ++ Γ) !! y = Some τ') as Hvar.
{ rewrite iter_up; destruct lt_dec.
- rewrite lookup_app_l in Hctxlu; auto.
exists x. rewrite /lookup_n /swap_list !lookup_app_l ?seq_length; try omega.
rewrite lookup_seq; auto.
- rewrite lookup_app_r in Hctxlu; auto with omega.
rewrite /lookup_n /swap_list lookup_app_r ?seq_length; try omega.
destruct (x - length ξ) as [|z] eqn:Heqx; first omega.
asimpl.
exists (length ξ + z).
replace x with (length ξ + S z) by omega.
repeat split; auto.
rewrite lookup_app_r; auto with omega.
by replace (length ξ + z - length ξ) with z by omega. }
clear Hctxlu Hx. destruct Hvar as (y & Heq1 & Heq2 & Hctxlu).
rewrite Heq1 Heq2.
apply binary_soundness; eauto.
+ repeat econstructor; eauto.
+ by econstructor.
+ intros.
iIntros (Δ vvs HΔ) "[#Hinv #Hvvs]". asimpl.
iIntros (γh γh' σ1') "Hσ'". asimpl.
rewrite ic_eq /ic_def.
iIntros (σ1 σ2 v n) "(Hrd & Hσ1)".
iDestruct "Hrd" as %Hrd. simpl in *.
iMod (allocate_full_state σ1') as (γh'') "(Hσ1' & Hfσ1')".
apply (nsteps_bind (fill_item (AppRCtx (RecV _)))) in Hrd
as [k [σ3 [w [Hkltn [He2 Hectx]]]]].
iPoseProof (binary_fundamental _ _ _ Ht' with "[]") as "Hx";
first iFrame "#".
iSpecialize ("Hx" $! γh γh'' σ1' with "Hσ1'").
rewrite ic_eq /ic_def. asimpl.
iSpecialize ("Hx" $! _ _ _ _ with "[$Hσ1]"); eauto.
iMod "Hx" as "(Hx & Hσ2)".
iDestruct "Hx" as (σ2' v') "(Hrtc & Hσ2' & Hrel)".
iModIntro; iFrame.
simpl in Hectx.
apply rec_red_step in Hectx; asimpl in Hectx.
destruct Hectx as [m [HnkSm Hnsteps]].
iDestruct (interp_env_Some_l with "[]") as ([v1 v2]) "[Hvveq Hvv]"; eauto.
iDestruct "Hvveq" as %Hvveq.
rewrite /env_subst list_lookup_fmap in Hnsteps.
rewrite /env_subst list_lookup_fmap.
rewrite Hvveq in Hnsteps.
apply (nsteps_val ST_lang) in Hnsteps as [? [?%of_val_inj ?]]; subst.
iFrame.
iExists _,_. iFrame; iFrame "#".
iPureIntro. rewrite Hvveq. asimpl. econstructor.
Qed.
Lemma lam_beta_basic_type v τ τ' Γ e2' :
Γ ⊢ₜ e2' : τ ->
(∀ Γ, Γ ⊢ₜ of_val v : τ') ->
(∀ f, (of_val v).[f] = (of_val v)) →
(∀ `{ICG_ST Σ} `{LogRelG Σ} Δ, True ⊢ ⟦ τ' ⟧ Δ (v,v)) →
Γ ⊨ App (Rec (of_val v)) e2' ≤ctx≤ of_val v : τ'.
Proof.
iIntros (Ht' Htv Hsubst Hvv). apply binary_soundness; eauto.
* repeat econstructor; eauto.
* iIntros (Σ ST Hlr Δ vvs HΔ) "[#Hinv #Hvvs]"; asimpl.
iIntros (γh γh' σ1') "Hσ'". asimpl. rewrite ic_eq /ic_def.
iIntros (σ1 σ2 w n) "(Hrd & Hσ1)".
iDestruct "Hrd" as %Hrd; simpl in *.
iMod (allocate_full_state σ1') as (γh'') "(Hσ1' & Hfσ1')".
apply (nsteps_bind (fill_item (AppRCtx (RecV _)))) in Hrd
as [k [σ3 [w2 [Hkltn [He2 Hectx]]]]].
iPoseProof (binary_fundamental _ _ _ Ht' with "[]") as "Hx"; first iFrame "#".
iSpecialize ("Hx" $! γh γh'' σ1' with "Hσ1'"). rewrite ic_eq /ic_def.
iSpecialize ("Hx" $! _ _ _ _ with "[$Hσ1]"); eauto.
iMod "Hx" as "(Hx & Hσ2)".
iDestruct "Hx" as (σ2' v') "(Hrtc & Hσ2' & Hrel)".
iModIntro; iFrame; simpl in *.
apply rec_red_step in Hectx; asimpl in Hectx.
destruct Hectx as [m [HnkSm Hnsteps]]; rewrite Hsubst in Hnsteps.
apply (nsteps_val ST_lang) in Hnsteps as [? [?%of_val_inj ?]];
subst; iFrame.
iExists _,_. iFrame; iFrame "#". repeat iSplit; eauto; last iApply Hvv.
iPureIntro. rewrite Hsubst. econstructor.
Qed.
Lemma lam_beta_push_lam K Γ e2' τ e τ' τ'' :
(∀ f e, (fill_item K e).[f] = fill_item K (e.[f])) →
(∀ Γ e, Γ ⊢ₜ e : τ' → Γ ⊢ₜ fill_item K e : τ'') →
(∀ `{ICG_ST Σ} `{LogRelG Σ} Γ e e', Γ ⊨ e ≤log≤ e' : τ' →
Γ ⊨ fill_item K e ≤log≤ fill_item K e' : τ'') →
τ :: Γ ⊢ₜ e : τ' →
Γ ⊢ₜ e2' : τ →
Γ ⊨ App (Rec ((fill_item K e.[ren (+1)]))) e2' ≤ctx≤
fill_item K (App (Rec e.[ren (+1)]) e2') : τ''.
Proof.
intros Hsubst HKt HKc Ht1 Ht2.
apply binary_soundness; eauto.
- repeat econstructor; eauto.
eapply HKt. by apply (context_weakening [_]).
- eapply HKt. repeat econstructor; eauto.
by apply (context_weakening [_]).
- iIntros (Σ ST Hlr Δ vvs HΔ) "[#Hinv #Hvvs]"; asimpl.
rewrite !Hsubst. asimpl.
iDestruct (interp_env_length with "[]") as %Hvvsl; eauto.
iApply (logrel_bind [AppRCtx (RecV _)] [K; AppRCtx (RecV _)]). iSplitL.
{ iApply (binary_fundamental _ _ _ Ht2); iFrame "#". }
iIntros (vv) "#Hvv".
iIntros (γh γh' σ1') "Hσ1' /=".
iApply ic_rec; eauto.
iNext. rewrite Hsubst. asimpl.
erewrite typed_subst_head_simpl; eauto;
last by rewrite /= fmap_length Hvvsl.
rewrite -Hsubst.
pose proof (binary_fundamental _ _ _ Ht1) as He.
apply HKc in He.
iPoseProof (He _ (vv :: vvs)with "[]") as "He";
first rewrite interp_env_cons; iFrame "#".
iSpecialize ("He" $! _ _ _ with "[$Hσ1']").
iApply ic_wand_r; iSplitL; first iApply "He".
simpl. iIntros (v1 _). iDestruct 1 as (σ2' v1') "(Hrd & Hσ2' & Hvv1)".
iDestruct "Hrd" as %Hrd.
apply rtc_nsteps in Hrd; destruct Hrd as [n Hrd].
iExists _, _; iFrame.
iPureIntro.
etrans.
eapply nsteps_rtc.
eapply (nsteps_l _ _ (_, _) (_, _)); simpl.
{ eapply (Ectx_step _ _ _ _ _ [K]); eauto.
econstructor; eauto. }
asimpl.
erewrite typed_subst_head_simpl; eauto;
last by rewrite /= fmap_length Hvvsl.
rewrite -Hsubst. eauto.
econstructor.
Qed.
Lemma lam_beta_push_lam_RunST Γ e2' τ e τ'' :
subst (ren (+1)) <$> τ :: Γ ⊢ₜ e : TST (TVar 0) τ''.[ren (+1)] →
Γ ⊢ₜ e2' : τ →
Γ ⊨ App (Rec ((RunST e.[ren (+1)]))) e2' ≤ctx≤
RunST (App (Rec e.[ren (+1)]) e2') : τ''.
Proof.
intros Ht1 Ht2.
apply binary_soundness; eauto.
- repeat econstructor; eauto.
apply (context_weakening [_]); eauto.
- repeat econstructor; last eapply context_rename; eauto.
apply (context_weakening [_]); eauto.
- iIntros (Σ ST Hlr Δ vvs HΔ) "[#Hinv #Hvvs]"; asimpl.
iDestruct (interp_env_length with "[]") as %Hvvsl; eauto.
iApply (logrel_bind [AppRCtx (RecV _)] [RunSTCtx; AppRCtx (RecV _)]).
iSplitL. { iApply (binary_fundamental _ _ _ Ht2); iFrame "#". }
iIntros (vv) "#Hvv".
iIntros (γh γh' σ1') "Hσ1' /=".
iApply ic_rec; eauto.
iNext. asimpl.
erewrite typed_subst_head_simpl; eauto;
last by rewrite /= !fmap_length Hvvsl.
pose proof (binary_fundamental _ _ _ Ht1) as He.
eapply bin_log_related_runST in He.
iPoseProof (He _ (vv :: vvs)with "[]") as "He";
first rewrite interp_env_cons; iFrame "#".
iSpecialize ("He" $! _ _ _ with "[$Hσ1']").
iApply ic_wand_r; iSplitL; first iApply "He".
simpl. iIntros (v1 _). iDestruct 1 as (σ2' v1') "(Hrd & Hσ2' & Hvv1)".
iDestruct "Hrd" as %Hrd.
apply rtc_nsteps in Hrd; destruct Hrd as [n Hrd].
iExists _, _; iFrame.
iPureIntro.
etrans.
eapply nsteps_rtc.
eapply (nsteps_l _ _ (_, _) (_, _)); simpl.
{ eapply (Ectx_step _ _ _ _ _ [RunSTCtx]); eauto.
econstructor; eauto. }
asimpl.
erewrite typed_subst_head_simpl; eauto;
last by rewrite /= !fmap_length Hvvsl.
econstructor.
Qed.