RunST.rec_hoisting.rec_hoisting_part1
From RunST Require Import IC base lang rules typing reduction regions
future contextual_refinement saved_pred ST_Lang_reduction
logrel_shared soundness soundnessNN logrelNN fundamentalNN
logrel fundamental.
From iris.proofmode Require Import tactics.
Hint Resolve to_of_val.
Definition bin_log_related' `{ICG_ST Σ} `{LogRelG Σ} (Γ : list type)
(e e' : expr) (τ : type) Δ vvs (HΔ : env_PersistentP Δ) : iProp Σ:=
(reg_inv ∧ ⟦ Γ ⟧* Δ vvs -★
⟦ τ ⟧ₑ Δ (e.[env_subst (vvs.*1)], e'.[env_subst (vvs.*2)]))%I.
Lemma weak_bin_log_related_rec `{ICG_ST Σ} `{LogRelG Σ} (Γ : list type)
(e e' : expr) (τ1 τ2 : type) Δ vvs (HΔ : env_PersistentP Δ) :
TArrow τ1 τ2 :: τ1 :: Γ ⊢ₜ e : τ2 →
TArrow τ1 τ2 :: τ1 :: Γ ⊢ₜ e' : τ2 →
(∀ vv1 vv2,
True ⊢ bin_log_related' (TArrow τ1 τ2 :: τ1 :: Γ)
e e' τ2 Δ (vv2 :: vv1 :: vvs) HΔ) →
True ⊢ bin_log_related' Γ (Rec e) (Rec e') (TArrow τ1 τ2) Δ (vvs) HΔ.
Proof.
iIntros (Ht1 Ht2 He) "#(Hinv & HΓ)". simpl.
iDestruct (interp_env_length with "HΓ") as %HΓ.
iIntros (? ? ?) "? /=". iApply ic_value; [cbn; rewrite ?to_of_val; trivial|].
iExists _, (RecV _); iFrame; iSplit; first by iPureIntro; econstructor.
iAlways. iLöb as "IH". iIntros ([v v']) "#Hvv".
iApply logrel_pure_det_head_step.
{ intros σ; econstructor; eauto. }
{ intros ? ? ? Hrd. inversion Hrd; subst; auto. }
{ intros σ. econstructor; eauto. }
iNext. asimpl. change (Rec ?e) with (of_val (RecV e)).
erewrite !n_closed_subst_head_simpl_2; (rewrite ?fmap_length; eauto).
- iApply (He (v, v') (_, _) with "[]"); first iFrame "#".
iApply interp_env_cons; iSplit; [|iApply interp_env_cons]; auto.
- rewrite -HΓ. apply (typed_n_closed _ _ _ Ht2).
- rewrite -HΓ. apply (typed_n_closed _ _ _ Ht1).
Qed.
Lemma Rec_Hoisting_LR1 `{ICG_ST Σ} `{LogRelG Σ} Γ e1 e1' τ e2 e2' τ' τ'' :
Γ ⊢ₜ e1 : τ → [TArrow τ' τ''; τ'; τ] ++ Γ ⊢ₜ e2 : τ'' →
Γ ⊢ₜ e1' : τ → [TArrow τ' τ''; τ'; τ] ++ Γ ⊢ₜ e2' : τ'' →
Γ ⊨ e1 ≤log≤ e1' : τ → [TArrow τ' τ''; τ'; τ] ++ Γ ⊨ e2 ≤log≤ e2' : τ'' →
Γ ⊨ App (Rec (Rec e2).[ren (+1)]) e1 ≤log≤
App (Rec (Rec (App (Rec e2'.[ren (+2)]) e1'.[ren (+3)])).[ren (+1)]) e1'
: TArrow τ' τ''.
Proof.
intros Ht1 Ht2 Ht1' Ht2' HLR1 HLR2.
iIntros (Δ 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.
apply (nsteps_bind (fill_item (AppRCtx (RecV _)))) in Hrd1.
destruct Hrd1 as (k & σ3 & w & Hk & Hrd11 & Hrd12). simpl in *.
set (Q := (|={⊤}=> Nat.iter k (λ Q, ▷ Q) (◇ (∀ σ1', ∃ (σ2' : state) (v' : val),
■ (rtc (pstep' det_lang)
(e1'.[env_subst (vvs.*2)], σ1') (of_val v', σ2') ∧
to_heap σ1' ≼ to_heap σ2' ∧ ✓ to_heap σ2'))))%I : iProp Σ).
iAssert Q as "HQ".
{ unfold Q. rewrite later_n_except_0_forall.
iApply (fupd_plain_forall_comm (E2 := ⊤)).
iIntros (σ3').
iMod (allocate_full_state σ3') as (γh3) "[Hnfσ3 Hnoσ3]".
iMod (allocate_full_state σ1) as (γh4) "[Hnfσ1 Hnoσ1]".
iPoseProof (HLR1 with "[]") as "He1"; iFrame "#".
iSpecialize ("He1" $! _ _ _ with "[$Hnfσ3]"); simpl.
rewrite ic_eq /ic_def /=.
iSpecialize ("He1" $! _ _ _ _ with "[Hnfσ1]"); iFrame; auto.
iApply (future_plain with "[-]").
iMod "He1" as "[He1 _]". iModIntro.
iDestruct "He1" as (σ4' ?) "(Hrd2&Hnfσ3&?)".
iDestruct "Hrd2" as %Hrd2.
iDestruct (heap_included with "[$Hnoσ3 $Hnfσ3]") as %[??].
iExists _, _; iPureIntro; repeat split; eauto. }
iMod "HQ" as "HQ". iModIntro. clear Q.
iPoseProof (HLR1 with "[]") as "He1"; iFrame "#".
iSpecialize ("He1" $! _ _ _ with "[$Hσ1']"); simpl.
rewrite ic_eq /ic_def /=.
iSpecialize ("He1" $! _ _ _ _ with "[How]"); iFrame; auto.
iDestruct (later_n_except_0_future (E := ⊤) with "[$HQ]") as "HQ".
iCombine "HQ" "He1" as "He1".
iMod "He1" as "(#HQ & [He1 HPh1])".
iDestruct "He1" as (σ2'' v'' ) "(Hrd21 & Ho1 & #Hrl1)".
iDestruct "Hrd21" as %Hrd21.
apply rtc_nsteps in Hrd21. destruct Hrd21 as [k1 Hrd21].
apply head_red_app in Hrd12. destruct Hrd12 as [? [? ?]]; subst.
iModIntro; iFrame. asimpl.
iExists σ2'', (RecV _); iSplit; iFrame; eauto.
{ 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. }
iAlways. iIntros (vv) "#Hvv". asimpl.
pose proof HLR2 as HLR2'.
eapply bin_log_related_rec in HLR2; eauto using (typed_n_closed _ _ _ Ht2);
eauto using (typed_n_closed _ _ _ Ht2').
iDestruct (interp_env_length with "Hvvs") as %Hvvsl.
assert (TArrow τ' τ'' :: τ' :: τ :: Γ ⊢ₜ
App (Rec e2'.[ren (+2)]) e1'.[ren (+3)] : τ'') as Ht3.
{ econstructor; last by eapply (context_weakening [_; _; _]).
econstructor. eapply (context_weakening [_; _]); eauto. }
replace (Rec e2.[ids 0 .: ids 1 .: (of_val w).[ren (+2)] .: env_subst (vvs.*1)
>> ren (+2)])
with (Rec e2).[env_subst (w :: (vvs).*1)]; last first.
{ assert (τ :: Γ ⊢ₜ (Rec e2) : TArrow τ' τ'') as Ht2''
by (econstructor; eauto).
rewrite -(n_closed_subst_head_simpl (S (length (Γ))));
first by asimpl.
{ apply (typed_n_closed _ _ _ Ht2''). }
{ by rewrite Hvvsl map_length. } }
replace (Rec (App (Rec e2'.[ids 2 .: ids 3 .: (of_val v'').[ren (+4)]
.: env_subst (vvs.*2) >> ren (+4)])
e1'.[env_subst (vvs.*2) >> ren (+2)]))
with (Rec (App (Rec e2'.[ren (+2)]) e1'.[ren (+3)]))
.[env_subst (v'' :: (vvs).*2)];
last first.
{ assert (τ :: Γ ⊢ₜ (Rec (App (Rec e2'.[ren (+2)]) e1'.[ren (+3)]))
: TArrow τ' τ'') as Ht3' by (econstructor; eauto).
rewrite -(n_closed_subst_head_simpl (S (length (Γ))));
first by asimpl.
{ apply (typed_n_closed _ _ _ Ht3'). }
{ by rewrite Hvvsl map_length. } }
iAssert (
■ ∀ σ1' : state,
∃ (σ2' : state) (v' : val),
rtc (pstep' det_lang) (e1'.[env_subst (vvs.*2)], σ1')
(of_val v', σ2') ∧ to_heap σ1' ≼ to_heap σ2' ∧ ✓ to_heap σ2')%I
as %HQ; last iClear "HQ".
{ rewrite uPred.pure_forall. iIntros (σq).
iDestruct ("HQ" $! σq) as (σq' vq) "%".
iPureIntro; eauto. }
iPoseProof (weak_bin_log_related_rec
(τ :: Γ) _ _ τ' τ'' Δ ((w, v'') :: vvs) HΔ) as "H";
first apply Ht2; first apply Ht3.
- iIntros (vv1 vv2) "[#Hinv #Hvvs']".
iIntros (γ1 γ2 σσ1) "Hσσ /=".
iPoseProof (HLR2' with "[]") as "He2"; iFrame "#".
destruct (HQ σσ1) as (σσ2 & vu & HQrd & HQinc & HQvl).
apply rtc_nsteps in HQrd. destruct HQrd as [kQ HQrd].
iMod (heap_catch_up with "Hσσ") as "Hσσ"; eauto.
iSpecialize ("He2" $! _ _ _ with "[$Hσσ]"); simpl.
iApply ic_wand_r; iSplitL; first iExact "He2". simpl.
iIntros (u _). iDestruct 1 as (σσ3 u') "(Hrd4 & Hσσ2 & #Huu)".
iDestruct "Hrd4" as %Hrd4.
apply rtc_nsteps in Hrd4. destruct Hrd4 as [k4 Hrd4].
iExists _, _; iFrame; iFrame "#". iPureIntro.
assert (TArrow τ' τ'' :: τ' :: τ :: Γ ⊢ₜ e1'.[ren (+3)] : τ) as H1''
by by eapply (context_weakening [_; _; _]).
erewrite <- typed_subst_head_simpl_3; eauto;
last by simpl; rewrite Hvvsl map_length.
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. by eapply (nsteps_rtc _).
- iIntros (γ1 γ2 σσ1) "Hσσ1".
rewrite /bin_log_related' /=.
iSpecialize ("H" with "[]"); first by rewrite interp_env_cons; iFrame "#".
iSpecialize ("H" $! γ1 _ _ with "[$Hσσ1]"). simpl.
iApply (ic_bind (fill_item (AppLCtx _))).
iApply ic_wand_r; iSplitL; first iExact "H". simpl.
iIntros (f _). iDestruct 1 as (σ4' f') "(Hrd3 & Hσ4' & #Hff)".
iDestruct "Hrd3" as %Hrd3.
apply rtc_nsteps in Hrd3. destruct Hrd3 as [k3 Hrd3].
iSpecialize ("Hff" $! vv with "[]"); iFrame "#".
iSpecialize ("Hff" $! _ _ _ with "[$Hσ4']"); iFrame "#".
iApply ic_wand_r; iSplitL; first iApply "Hff". simpl.
iIntros (u _); iDestruct 1 as (σ5' u5') "(Hrd5 & Hσ5' & #Huu5)".
iExists _, _; iFrame; iFrame "#".
iDestruct "Hrd5" as %Hrd5.
iPureIntro.
etrans.
+ eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (AppLCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
+ eauto.
Qed.
future contextual_refinement saved_pred ST_Lang_reduction
logrel_shared soundness soundnessNN logrelNN fundamentalNN
logrel fundamental.
From iris.proofmode Require Import tactics.
Hint Resolve to_of_val.
Definition bin_log_related' `{ICG_ST Σ} `{LogRelG Σ} (Γ : list type)
(e e' : expr) (τ : type) Δ vvs (HΔ : env_PersistentP Δ) : iProp Σ:=
(reg_inv ∧ ⟦ Γ ⟧* Δ vvs -★
⟦ τ ⟧ₑ Δ (e.[env_subst (vvs.*1)], e'.[env_subst (vvs.*2)]))%I.
Lemma weak_bin_log_related_rec `{ICG_ST Σ} `{LogRelG Σ} (Γ : list type)
(e e' : expr) (τ1 τ2 : type) Δ vvs (HΔ : env_PersistentP Δ) :
TArrow τ1 τ2 :: τ1 :: Γ ⊢ₜ e : τ2 →
TArrow τ1 τ2 :: τ1 :: Γ ⊢ₜ e' : τ2 →
(∀ vv1 vv2,
True ⊢ bin_log_related' (TArrow τ1 τ2 :: τ1 :: Γ)
e e' τ2 Δ (vv2 :: vv1 :: vvs) HΔ) →
True ⊢ bin_log_related' Γ (Rec e) (Rec e') (TArrow τ1 τ2) Δ (vvs) HΔ.
Proof.
iIntros (Ht1 Ht2 He) "#(Hinv & HΓ)". simpl.
iDestruct (interp_env_length with "HΓ") as %HΓ.
iIntros (? ? ?) "? /=". iApply ic_value; [cbn; rewrite ?to_of_val; trivial|].
iExists _, (RecV _); iFrame; iSplit; first by iPureIntro; econstructor.
iAlways. iLöb as "IH". iIntros ([v v']) "#Hvv".
iApply logrel_pure_det_head_step.
{ intros σ; econstructor; eauto. }
{ intros ? ? ? Hrd. inversion Hrd; subst; auto. }
{ intros σ. econstructor; eauto. }
iNext. asimpl. change (Rec ?e) with (of_val (RecV e)).
erewrite !n_closed_subst_head_simpl_2; (rewrite ?fmap_length; eauto).
- iApply (He (v, v') (_, _) with "[]"); first iFrame "#".
iApply interp_env_cons; iSplit; [|iApply interp_env_cons]; auto.
- rewrite -HΓ. apply (typed_n_closed _ _ _ Ht2).
- rewrite -HΓ. apply (typed_n_closed _ _ _ Ht1).
Qed.
Lemma Rec_Hoisting_LR1 `{ICG_ST Σ} `{LogRelG Σ} Γ e1 e1' τ e2 e2' τ' τ'' :
Γ ⊢ₜ e1 : τ → [TArrow τ' τ''; τ'; τ] ++ Γ ⊢ₜ e2 : τ'' →
Γ ⊢ₜ e1' : τ → [TArrow τ' τ''; τ'; τ] ++ Γ ⊢ₜ e2' : τ'' →
Γ ⊨ e1 ≤log≤ e1' : τ → [TArrow τ' τ''; τ'; τ] ++ Γ ⊨ e2 ≤log≤ e2' : τ'' →
Γ ⊨ App (Rec (Rec e2).[ren (+1)]) e1 ≤log≤
App (Rec (Rec (App (Rec e2'.[ren (+2)]) e1'.[ren (+3)])).[ren (+1)]) e1'
: TArrow τ' τ''.
Proof.
intros Ht1 Ht2 Ht1' Ht2' HLR1 HLR2.
iIntros (Δ 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.
apply (nsteps_bind (fill_item (AppRCtx (RecV _)))) in Hrd1.
destruct Hrd1 as (k & σ3 & w & Hk & Hrd11 & Hrd12). simpl in *.
set (Q := (|={⊤}=> Nat.iter k (λ Q, ▷ Q) (◇ (∀ σ1', ∃ (σ2' : state) (v' : val),
■ (rtc (pstep' det_lang)
(e1'.[env_subst (vvs.*2)], σ1') (of_val v', σ2') ∧
to_heap σ1' ≼ to_heap σ2' ∧ ✓ to_heap σ2'))))%I : iProp Σ).
iAssert Q as "HQ".
{ unfold Q. rewrite later_n_except_0_forall.
iApply (fupd_plain_forall_comm (E2 := ⊤)).
iIntros (σ3').
iMod (allocate_full_state σ3') as (γh3) "[Hnfσ3 Hnoσ3]".
iMod (allocate_full_state σ1) as (γh4) "[Hnfσ1 Hnoσ1]".
iPoseProof (HLR1 with "[]") as "He1"; iFrame "#".
iSpecialize ("He1" $! _ _ _ with "[$Hnfσ3]"); simpl.
rewrite ic_eq /ic_def /=.
iSpecialize ("He1" $! _ _ _ _ with "[Hnfσ1]"); iFrame; auto.
iApply (future_plain with "[-]").
iMod "He1" as "[He1 _]". iModIntro.
iDestruct "He1" as (σ4' ?) "(Hrd2&Hnfσ3&?)".
iDestruct "Hrd2" as %Hrd2.
iDestruct (heap_included with "[$Hnoσ3 $Hnfσ3]") as %[??].
iExists _, _; iPureIntro; repeat split; eauto. }
iMod "HQ" as "HQ". iModIntro. clear Q.
iPoseProof (HLR1 with "[]") as "He1"; iFrame "#".
iSpecialize ("He1" $! _ _ _ with "[$Hσ1']"); simpl.
rewrite ic_eq /ic_def /=.
iSpecialize ("He1" $! _ _ _ _ with "[How]"); iFrame; auto.
iDestruct (later_n_except_0_future (E := ⊤) with "[$HQ]") as "HQ".
iCombine "HQ" "He1" as "He1".
iMod "He1" as "(#HQ & [He1 HPh1])".
iDestruct "He1" as (σ2'' v'' ) "(Hrd21 & Ho1 & #Hrl1)".
iDestruct "Hrd21" as %Hrd21.
apply rtc_nsteps in Hrd21. destruct Hrd21 as [k1 Hrd21].
apply head_red_app in Hrd12. destruct Hrd12 as [? [? ?]]; subst.
iModIntro; iFrame. asimpl.
iExists σ2'', (RecV _); iSplit; iFrame; eauto.
{ 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. }
iAlways. iIntros (vv) "#Hvv". asimpl.
pose proof HLR2 as HLR2'.
eapply bin_log_related_rec in HLR2; eauto using (typed_n_closed _ _ _ Ht2);
eauto using (typed_n_closed _ _ _ Ht2').
iDestruct (interp_env_length with "Hvvs") as %Hvvsl.
assert (TArrow τ' τ'' :: τ' :: τ :: Γ ⊢ₜ
App (Rec e2'.[ren (+2)]) e1'.[ren (+3)] : τ'') as Ht3.
{ econstructor; last by eapply (context_weakening [_; _; _]).
econstructor. eapply (context_weakening [_; _]); eauto. }
replace (Rec e2.[ids 0 .: ids 1 .: (of_val w).[ren (+2)] .: env_subst (vvs.*1)
>> ren (+2)])
with (Rec e2).[env_subst (w :: (vvs).*1)]; last first.
{ assert (τ :: Γ ⊢ₜ (Rec e2) : TArrow τ' τ'') as Ht2''
by (econstructor; eauto).
rewrite -(n_closed_subst_head_simpl (S (length (Γ))));
first by asimpl.
{ apply (typed_n_closed _ _ _ Ht2''). }
{ by rewrite Hvvsl map_length. } }
replace (Rec (App (Rec e2'.[ids 2 .: ids 3 .: (of_val v'').[ren (+4)]
.: env_subst (vvs.*2) >> ren (+4)])
e1'.[env_subst (vvs.*2) >> ren (+2)]))
with (Rec (App (Rec e2'.[ren (+2)]) e1'.[ren (+3)]))
.[env_subst (v'' :: (vvs).*2)];
last first.
{ assert (τ :: Γ ⊢ₜ (Rec (App (Rec e2'.[ren (+2)]) e1'.[ren (+3)]))
: TArrow τ' τ'') as Ht3' by (econstructor; eauto).
rewrite -(n_closed_subst_head_simpl (S (length (Γ))));
first by asimpl.
{ apply (typed_n_closed _ _ _ Ht3'). }
{ by rewrite Hvvsl map_length. } }
iAssert (
■ ∀ σ1' : state,
∃ (σ2' : state) (v' : val),
rtc (pstep' det_lang) (e1'.[env_subst (vvs.*2)], σ1')
(of_val v', σ2') ∧ to_heap σ1' ≼ to_heap σ2' ∧ ✓ to_heap σ2')%I
as %HQ; last iClear "HQ".
{ rewrite uPred.pure_forall. iIntros (σq).
iDestruct ("HQ" $! σq) as (σq' vq) "%".
iPureIntro; eauto. }
iPoseProof (weak_bin_log_related_rec
(τ :: Γ) _ _ τ' τ'' Δ ((w, v'') :: vvs) HΔ) as "H";
first apply Ht2; first apply Ht3.
- iIntros (vv1 vv2) "[#Hinv #Hvvs']".
iIntros (γ1 γ2 σσ1) "Hσσ /=".
iPoseProof (HLR2' with "[]") as "He2"; iFrame "#".
destruct (HQ σσ1) as (σσ2 & vu & HQrd & HQinc & HQvl).
apply rtc_nsteps in HQrd. destruct HQrd as [kQ HQrd].
iMod (heap_catch_up with "Hσσ") as "Hσσ"; eauto.
iSpecialize ("He2" $! _ _ _ with "[$Hσσ]"); simpl.
iApply ic_wand_r; iSplitL; first iExact "He2". simpl.
iIntros (u _). iDestruct 1 as (σσ3 u') "(Hrd4 & Hσσ2 & #Huu)".
iDestruct "Hrd4" as %Hrd4.
apply rtc_nsteps in Hrd4. destruct Hrd4 as [k4 Hrd4].
iExists _, _; iFrame; iFrame "#". iPureIntro.
assert (TArrow τ' τ'' :: τ' :: τ :: Γ ⊢ₜ e1'.[ren (+3)] : τ) as H1''
by by eapply (context_weakening [_; _; _]).
erewrite <- typed_subst_head_simpl_3; eauto;
last by simpl; rewrite Hvvsl map_length.
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. by eapply (nsteps_rtc _).
- iIntros (γ1 γ2 σσ1) "Hσσ1".
rewrite /bin_log_related' /=.
iSpecialize ("H" with "[]"); first by rewrite interp_env_cons; iFrame "#".
iSpecialize ("H" $! γ1 _ _ with "[$Hσσ1]"). simpl.
iApply (ic_bind (fill_item (AppLCtx _))).
iApply ic_wand_r; iSplitL; first iExact "H". simpl.
iIntros (f _). iDestruct 1 as (σ4' f') "(Hrd3 & Hσ4' & #Hff)".
iDestruct "Hrd3" as %Hrd3.
apply rtc_nsteps in Hrd3. destruct Hrd3 as [k3 Hrd3].
iSpecialize ("Hff" $! vv with "[]"); iFrame "#".
iSpecialize ("Hff" $! _ _ _ with "[$Hσ4']"); iFrame "#".
iApply ic_wand_r; iSplitL; first iApply "Hff". simpl.
iIntros (u _); iDestruct 1 as (σ5' u5') "(Hrd5 & Hσ5' & #Huu5)".
iExists _, _; iFrame; iFrame "#".
iDestruct "Hrd5" as %Hrd5.
iPureIntro.
etrans.
+ eapply (nsteps_rtc _).
eapply (@nsteps_ctx det_lang (fill_item (AppLCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
+ eauto.
Qed.