RunST.rec_hoisting.rec_hoisting_part2
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_relatedNN' `{ICG_ST Σ} `{LogRelG Σ} (Γ : list type)
(e e' : expr) (τ : type) Δ vvs (HΔ : env_PersistentP Δ) : iProp Σ:=
(reg_inv ∧ ⟦ Γ ⟧NN* Δ vvs -★
⟦ τ ⟧NNₑ Δ (e.[env_subst (vvs.*1)], e'.[env_subst (vvs.*2)]))%I.
Lemma weak_bin_log_related_recNN `{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_relatedNN' (TArrow τ1 τ2 :: τ1 :: Γ)
e e' τ2 Δ (vv2 :: vv1 :: vvs) HΔ) →
True ⊢ bin_log_relatedNN' Γ (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 fundamentalNN.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.
Require Import Coq.Logic.FunctionalExtensionality.
Lemma lookup_swap_3 : (lookup_n (swap_list 1 2)) = (1 .: 2 .: 0 .: (+3)).
Proof.
extensionality i; destruct i as [|[|[|?]]]; by cbv.
Qed.
Lemma env_subst_irrel l1 l1' l2 :
length l1 = length l1' →
(+(length l1)) >>> env_subst (l1 ++ l2) =
(+(length l1')) >>> env_subst (l1' ++ l2).
Proof.
intros Hi. unfold env_subst. extensionality i. asimpl.
by rewrite -!lookup_drop !drop_app Hi.
Qed.
Lemma f_unfold_3 (f : var → expr) : f 0 .: f 1 .: f 2 .: (+3) >>> f = f.
Proof. by asimpl. Qed.
Lemma Rec_Hoisting_LR2 `{ICG_ST Σ} `{LogRelG Σ} Γ e1' τ e2' τ' τ'' :
Γ ⊢ₜ e1' : τ → [TArrow τ' τ''; τ'; τ] ++ Γ ⊢ₜ e2' : τ'' →
Γ ⊨ App (Rec (Rec (App (Rec e2'.[ren (+2)]) e1'.[ren (+3)])).[ren (+1)]) e1'
≤logNN≤
App (Rec (Rec (App (Rec e2'.[(ren (lookup_n (swap_list 1 2)))].[ren (+1)])
e1'.[ren (+2)])).[ren (+2)]) e1'
: TArrow τ' τ''.
Proof.
rewrite lookup_swap_3. asimpl.
iIntros (Ht1 Ht2 Δ 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 *.
iPoseProof (fundamentalNN.binary_fundamental _ _ _ Ht1 with "[]") as "He1";
iFrame "#".
iSpecialize ("He1" $! _ _ _ with "[$Hσ1']"); simpl.
rewrite ic_eq /ic_def /=.
iSpecialize ("He1" $! _ _ _ _ with "[How]"); iFrame; auto.
iMod "He1" as "[Hrd How]". iDestruct "Hrd" as (σ2' w') "(Hrd2 & Hσ2' & #Hww)".
iDestruct "Hrd2" as %Hrd2.
apply head_red_app in Hrd12; destruct Hrd12 as [? [? ?]]; subst.
iModIntro; iFrame. iExists _, (RecV _); iFrame; iSplit.
{ iPureIntro. asimpl.
replace n with (k + 1) by omega; eapply nsteps_trans.
- eapply (@nsteps_ctx det_lang (fill_item (AppRCtx (RecV _))));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
- eapply (nsteps_l _ _ _ (_, _) (_, _)); simpl.
{ apply head_prim_step. econstructor; eauto. }
asimpl. econstructor. }
asimpl.
assert (TArrow τ' τ'' :: τ' :: τ :: Γ ⊢ₜ
App (Rec e2'.[ren (+2)]) e1'.[ren (+3)] : τ'') as Ht3.
{ econstructor; last by eapply (context_weakening [_; _; _]).
econstructor. eapply (context_weakening [_; _]); eauto. }
iDestruct (interp_env_length with "Hvvs") as %Hvvsl.
replace (Rec
(App (Rec e2'.[Var 2 .: Var 3 .: (of_val w).[ren (+4)]
.: env_subst (vvs.*1) >> ren (+4)])
e1'.[env_subst (vvs.*1) >> ren (+2)]))
with (Rec (App (Rec e2'.[ren (+2)]) e1'.[ren (+3)]))
.[env_subst (w :: (vvs).*1)]; 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. } }
assert (TArrow τ' τ'' :: τ' :: τ :: Γ ⊢ₜ
(App (Rec e2'.[ren (1 .: 2 .: 0 .: (+3))]
.[ren (+1)]) .[upn 2 (ren (+1))]
e1'.[ren (+2)].[upn 2 (ren (+1))])
: τ'') as Ht3'.
{ rewrite -lookup_swap_3. econstructor; last first.
- asimpl. eapply (context_weakening [_; _; _]); eauto.
- econstructor.
replace (e2'.[ren (lookup_n (swap_list 1 2))].[ren (+1)].[
upn 2 (upn 2 (ren (+1)))]) with
(e2'.[ren (lookup_n (swap_list 1 2))].[ren (+1)].[(upn 4 (ren (+1)))])
by by asimpl.
eapply (context_gen_weakening [_] [_; _; _; _] ).
eapply (context_weakening [_]); eauto.
eapply (context_swap_typed [] [_; _] [_]); auto. }
replace (Rec
(App
(Rec
e2'.[ids 2 .: ids 3 .: ids 1 .: env_subst (vvs.*2) >> ren (+4)])
e1'.[env_subst (vvs.*2) >> ren (+2)]))
with (Rec (App (Rec e2'.[ren (1 .: 2 .: 0 .: (+3))].[ren (+1)])
.[upn 2 (ren (+1))]
e1'.[ren (+2)].[upn 2 (ren (+1))])).[env_subst (w' :: (vvs).*2)];
last first.
{ assert (τ :: Γ ⊢ₜ (Rec (App (Rec e2'.[ren (1 .: 2 .: 0 .: (+3))]
.[ren (+1)]) .[upn 2 (ren (+1))]
e1'.[ren (+2)].[upn 2 (ren (+1))]))
: TArrow τ' τ'') as Ht3'' by by econstructor; eauto.
rewrite -(n_closed_subst_head_simpl (S (length (Γ))));
first by asimpl.
{ apply (typed_n_closed _ _ _ Ht3''). }
{ by rewrite Hvvsl map_length. } }
iPoseProof (weak_bin_log_related_recNN
(τ :: Γ) _ _ τ' τ'' Δ ((w, w') :: vvs) HΔ) as "H";
first apply Ht3; first apply Ht3'; last first.
- iAlways. iIntros (vv) "#Hvv".
iIntros (γ1 γ2 σσ1) "Hσσ1".
rewrite /bin_log_relatedNN' /=.
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 k'). iDestruct 1 as (σ4' f') "(Hrd3 & Hσ4' & #Hff)".
iDestruct "Hrd3" as %Hrd3.
iSpecialize ("Hff" $! vv with "[]"); iFrame "#".
iSpecialize ("Hff" $! _ _ _ with "[$Hσ4']"); iFrame "#".
iApply ic_wand_r; iSplitL; first iApply "Hff". simpl.
iIntros (u k''); iDestruct 1 as (σ5' u5') "(Hrd5 & Hσ5' & #Huu5)".
iExists _, _; iFrame; iFrame "#".
iDestruct "Hrd5" as %Hrd5.
iPureIntro.
eapply nsteps_trans.
+ eapply (@nsteps_ctx det_lang (fill_item (AppLCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
+ eauto.
- unfold bin_log_relatedNN'. asimpl.
iIntros (vv1 vv2) "[#Hinv #Hvvs']".
iIntros (γ1 γ2 σσ1) "Hσσ /=".
rewrite ?interp_env_cons.
iDestruct "Hvvs'" as "(Hvv2 & Hvv1 & Hww & Hvvs)".
rewrite ic_eq /ic_def /=.
iIntros (σ4 σ5 u k') "(Hrd5 & Hσ4)".
iDestruct "Hrd5" as %Hrd5.
assert (TArrow τ' τ'' :: τ' :: τ :: Γ ⊢ₜ e1'.[ren (+3)] : τ) as H1''
by by eapply (context_weakening [_; _; _]).
erewrite <- typed_subst_head_simpl_3 in Hrd5; eauto;
last by simpl; rewrite Hvvsl map_length.
asimpl in Hrd5.
apply (nsteps_bind (fill_item (AppRCtx (RecV _)))) in Hrd5.
destruct Hrd5 as (m & σ6 & w5 & Hm & Hrd51 & Hrd52). simpl in *.
iMod (allocate_full_state σ4) as (γ3) "[H341 _]".
iMod (allocate_full_state σ1') as (γ4) "[H411 _]".
iPoseProof (fundamentalNN.binary_fundamental _ _ _ Ht1 with "[]") as "He1";
iFrame "#".
iSpecialize ("He1" $! _ _ _ with "[$H411]"); simpl.
rewrite ic_eq /ic_def /=.
iSpecialize ("He1" $! _ _ _ _ with "[$H341]"); auto.
iMod (allocate_full_state σ1) as (γ5) "[H511 _]".
iMod (allocate_full_state σσ1) as (γ6) "[H611 _]".
iPoseProof (fundamentalNN.binary_fundamental _ _ _ Ht1 with "[]") as "He1'";
iFrame "#".
iSpecialize ("He1'" $! _ _ _ with "[$H611]"); simpl.
rewrite ic_eq /ic_def /=.
iSpecialize ("He1'" $! _ _ _ _ with "[$H511]"); auto.
iPoseProof (fundamentalNN.binary_fundamental _ _ _ Ht1 with "[]") as "He1''";
iFrame "#".
iSpecialize ("He1''" $! _ _ _ with "[$Hσσ]"); simpl.
rewrite ic_eq /ic_def /=.
iSpecialize ("He1''" $! _ _ _ _ with "[$Hσ4]"); auto.
iCombine "He1" "He1''" as "He1".
iPoseProof (future_cancel_2 with "[He1 He1']") as "He"; first trivial.
{ iSplitL "He1'"; iFrame. }
iMod "He" as "[He1 [[He2 Hσ26] [He3 Hσ36]]]".
iDestruct "He2" as (σ7' v7') "(Hrd7 & Hσ7' & #Hvv7)".
iDestruct "Hrd7" as %Hrd7.
destruct (nsteps_deterministic' Hrd7 Hrd2) as [?[??]]; subst.
iDestruct "He3" as (σ8' v8') "(Hrd8 & Hσ8' & #Hvv8)".
iDestruct "Hrd8" as %Hrd8.
replace (k- k) with 0 by omega. rewrite future_unfold_O.
iMod "He1" as "[He1 Hσ5]".
iDestruct "He1" as (σ9' v9') "(Hrd9 & Hσ9' & #Hvv9)".
iDestruct "Hrd9" as %Hrd9.
destruct (nsteps_deterministic' Hrd8 Hrd9) as [?[??]]; subst.
apply rec_red_step in Hrd52. destruct Hrd52 as [m [Hkm Hrd52]].
assert (τ :: TArrow τ' τ'' :: τ' :: Γ ⊢ₜ
e2'.[ren (1 .: 2 .: 0 .: (+3))] : τ'') as Ht2''.
{ rewrite -lookup_swap_3; eapply (context_swap_typed [] [_; _] [_]); auto. }
match type of Hrd52 with
nsteps _ _ (?A, _) _ =>
replace A with
e2'.[ren (1 .: 2 .: 0 .: (+3))]
.[env_subst (((w, v9') :: vv2 :: vv1 :: vvs).*1)]
in Hrd52; last first
end.
{ simpl.
erewrite <- typed_subst_head_simpl_3; eauto;
last rewrite map_length -Hvvsl //.
asimpl.
erewrite <- typed_subst_head_simpl_3; eauto;
last rewrite map_length -Hvvsl //. }
iPoseProof (fundamentalNN.binary_fundamental _ _ _
Ht2'' _ (((w, v9') :: vv2 :: vv1 :: vvs)) with "[]") as "He2".
{ rewrite ?interp_env_cons; iFrame "#". }
simpl.
repeat (erewrite <- typed_subst_head_simpl_3; eauto;
last rewrite map_length -Hvvsl //).
simpl in Hrd52.
erewrite <- typed_subst_head_simpl_3 in Hrd52; eauto;
last rewrite map_length -Hvvsl //.
iSpecialize ("He2" $! _ _ _ with "[$Hσ8']"); simpl.
rewrite ic_eq /ic_def /=.
iSpecialize ("He2" $! _ _ _ _ with "[$Hσ36]"); auto.
iModIntro. iMod "He2" as "[He2 Hσ8']".
iDestruct "He2" as (σ10 v10) "(Hrd10 & Hσ10 & #Hvv10)".
iDestruct "Hrd10" as %Hrd10.
iModIntro; iFrame. iExists _, _; iFrame; iFrame "#".
iPureIntro.
replace k' with (k + S m) by omega.
eapply nsteps_trans.
+ eapply (@nsteps_ctx det_lang (fill_item (AppRCtx (RecV _))));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
asimpl. eauto.
+ eapply (nsteps_l _ _ (_, _) (_, _)); simpl.
{ apply head_prim_step. econstructor; eauto. }
asimpl in Hrd10.
match goal with
|- nsteps _ _ (?A, _) _ =>
replace A with
e2'.[env_subst (vv2.2 :: vv1.2 :: v9' :: (vvs.*2))]
end.
{ erewrite <- typed_subst_head_simpl_3; eauto;
last rewrite map_length -Hvvsl //. }
{ match goal with
|- _ = e2'.[_].[?A].[?B,_/] => remember B; remember A as z; asimpl
end. rewrite Heqz /env_subst. asimpl.
replace (((+3) >>>
(λ x : var,
from_option id (Var x)
(of_val <$> (vv2.2 :: vv1.2 :: w' :: vvs.*2) !! x))))
with ((+3) >>> env_subst (vv2.2 :: vv1.2 :: w' :: vvs.*2))
by trivial.
rewrite (env_subst_irrel [vv2.2;vv1.2;w'] [vv2.2;vv1.2;v9']); auto.
unfold env_subst; asimpl.
replace (((+3) >>>
(λ x : var,
from_option id (Var x)
(of_val <$>
(vv2.2 :: vv1.2 :: v9' :: list_fmap (val * val)%type val snd vvs)
!! x))))
with ((+3) >>> env_subst (vv2.2 :: vv1.2 :: v9' :: vvs.*2))
by trivial.
by rewrite f_unfold_3. }
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_relatedNN' `{ICG_ST Σ} `{LogRelG Σ} (Γ : list type)
(e e' : expr) (τ : type) Δ vvs (HΔ : env_PersistentP Δ) : iProp Σ:=
(reg_inv ∧ ⟦ Γ ⟧NN* Δ vvs -★
⟦ τ ⟧NNₑ Δ (e.[env_subst (vvs.*1)], e'.[env_subst (vvs.*2)]))%I.
Lemma weak_bin_log_related_recNN `{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_relatedNN' (TArrow τ1 τ2 :: τ1 :: Γ)
e e' τ2 Δ (vv2 :: vv1 :: vvs) HΔ) →
True ⊢ bin_log_relatedNN' Γ (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 fundamentalNN.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.
Require Import Coq.Logic.FunctionalExtensionality.
Lemma lookup_swap_3 : (lookup_n (swap_list 1 2)) = (1 .: 2 .: 0 .: (+3)).
Proof.
extensionality i; destruct i as [|[|[|?]]]; by cbv.
Qed.
Lemma env_subst_irrel l1 l1' l2 :
length l1 = length l1' →
(+(length l1)) >>> env_subst (l1 ++ l2) =
(+(length l1')) >>> env_subst (l1' ++ l2).
Proof.
intros Hi. unfold env_subst. extensionality i. asimpl.
by rewrite -!lookup_drop !drop_app Hi.
Qed.
Lemma f_unfold_3 (f : var → expr) : f 0 .: f 1 .: f 2 .: (+3) >>> f = f.
Proof. by asimpl. Qed.
Lemma Rec_Hoisting_LR2 `{ICG_ST Σ} `{LogRelG Σ} Γ e1' τ e2' τ' τ'' :
Γ ⊢ₜ e1' : τ → [TArrow τ' τ''; τ'; τ] ++ Γ ⊢ₜ e2' : τ'' →
Γ ⊨ App (Rec (Rec (App (Rec e2'.[ren (+2)]) e1'.[ren (+3)])).[ren (+1)]) e1'
≤logNN≤
App (Rec (Rec (App (Rec e2'.[(ren (lookup_n (swap_list 1 2)))].[ren (+1)])
e1'.[ren (+2)])).[ren (+2)]) e1'
: TArrow τ' τ''.
Proof.
rewrite lookup_swap_3. asimpl.
iIntros (Ht1 Ht2 Δ 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 *.
iPoseProof (fundamentalNN.binary_fundamental _ _ _ Ht1 with "[]") as "He1";
iFrame "#".
iSpecialize ("He1" $! _ _ _ with "[$Hσ1']"); simpl.
rewrite ic_eq /ic_def /=.
iSpecialize ("He1" $! _ _ _ _ with "[How]"); iFrame; auto.
iMod "He1" as "[Hrd How]". iDestruct "Hrd" as (σ2' w') "(Hrd2 & Hσ2' & #Hww)".
iDestruct "Hrd2" as %Hrd2.
apply head_red_app in Hrd12; destruct Hrd12 as [? [? ?]]; subst.
iModIntro; iFrame. iExists _, (RecV _); iFrame; iSplit.
{ iPureIntro. asimpl.
replace n with (k + 1) by omega; eapply nsteps_trans.
- eapply (@nsteps_ctx det_lang (fill_item (AppRCtx (RecV _))));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
- eapply (nsteps_l _ _ _ (_, _) (_, _)); simpl.
{ apply head_prim_step. econstructor; eauto. }
asimpl. econstructor. }
asimpl.
assert (TArrow τ' τ'' :: τ' :: τ :: Γ ⊢ₜ
App (Rec e2'.[ren (+2)]) e1'.[ren (+3)] : τ'') as Ht3.
{ econstructor; last by eapply (context_weakening [_; _; _]).
econstructor. eapply (context_weakening [_; _]); eauto. }
iDestruct (interp_env_length with "Hvvs") as %Hvvsl.
replace (Rec
(App (Rec e2'.[Var 2 .: Var 3 .: (of_val w).[ren (+4)]
.: env_subst (vvs.*1) >> ren (+4)])
e1'.[env_subst (vvs.*1) >> ren (+2)]))
with (Rec (App (Rec e2'.[ren (+2)]) e1'.[ren (+3)]))
.[env_subst (w :: (vvs).*1)]; 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. } }
assert (TArrow τ' τ'' :: τ' :: τ :: Γ ⊢ₜ
(App (Rec e2'.[ren (1 .: 2 .: 0 .: (+3))]
.[ren (+1)]) .[upn 2 (ren (+1))]
e1'.[ren (+2)].[upn 2 (ren (+1))])
: τ'') as Ht3'.
{ rewrite -lookup_swap_3. econstructor; last first.
- asimpl. eapply (context_weakening [_; _; _]); eauto.
- econstructor.
replace (e2'.[ren (lookup_n (swap_list 1 2))].[ren (+1)].[
upn 2 (upn 2 (ren (+1)))]) with
(e2'.[ren (lookup_n (swap_list 1 2))].[ren (+1)].[(upn 4 (ren (+1)))])
by by asimpl.
eapply (context_gen_weakening [_] [_; _; _; _] ).
eapply (context_weakening [_]); eauto.
eapply (context_swap_typed [] [_; _] [_]); auto. }
replace (Rec
(App
(Rec
e2'.[ids 2 .: ids 3 .: ids 1 .: env_subst (vvs.*2) >> ren (+4)])
e1'.[env_subst (vvs.*2) >> ren (+2)]))
with (Rec (App (Rec e2'.[ren (1 .: 2 .: 0 .: (+3))].[ren (+1)])
.[upn 2 (ren (+1))]
e1'.[ren (+2)].[upn 2 (ren (+1))])).[env_subst (w' :: (vvs).*2)];
last first.
{ assert (τ :: Γ ⊢ₜ (Rec (App (Rec e2'.[ren (1 .: 2 .: 0 .: (+3))]
.[ren (+1)]) .[upn 2 (ren (+1))]
e1'.[ren (+2)].[upn 2 (ren (+1))]))
: TArrow τ' τ'') as Ht3'' by by econstructor; eauto.
rewrite -(n_closed_subst_head_simpl (S (length (Γ))));
first by asimpl.
{ apply (typed_n_closed _ _ _ Ht3''). }
{ by rewrite Hvvsl map_length. } }
iPoseProof (weak_bin_log_related_recNN
(τ :: Γ) _ _ τ' τ'' Δ ((w, w') :: vvs) HΔ) as "H";
first apply Ht3; first apply Ht3'; last first.
- iAlways. iIntros (vv) "#Hvv".
iIntros (γ1 γ2 σσ1) "Hσσ1".
rewrite /bin_log_relatedNN' /=.
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 k'). iDestruct 1 as (σ4' f') "(Hrd3 & Hσ4' & #Hff)".
iDestruct "Hrd3" as %Hrd3.
iSpecialize ("Hff" $! vv with "[]"); iFrame "#".
iSpecialize ("Hff" $! _ _ _ with "[$Hσ4']"); iFrame "#".
iApply ic_wand_r; iSplitL; first iApply "Hff". simpl.
iIntros (u k''); iDestruct 1 as (σ5' u5') "(Hrd5 & Hσ5' & #Huu5)".
iExists _, _; iFrame; iFrame "#".
iDestruct "Hrd5" as %Hrd5.
iPureIntro.
eapply nsteps_trans.
+ eapply (@nsteps_ctx det_lang (fill_item (AppLCtx _)));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
+ eauto.
- unfold bin_log_relatedNN'. asimpl.
iIntros (vv1 vv2) "[#Hinv #Hvvs']".
iIntros (γ1 γ2 σσ1) "Hσσ /=".
rewrite ?interp_env_cons.
iDestruct "Hvvs'" as "(Hvv2 & Hvv1 & Hww & Hvvs)".
rewrite ic_eq /ic_def /=.
iIntros (σ4 σ5 u k') "(Hrd5 & Hσ4)".
iDestruct "Hrd5" as %Hrd5.
assert (TArrow τ' τ'' :: τ' :: τ :: Γ ⊢ₜ e1'.[ren (+3)] : τ) as H1''
by by eapply (context_weakening [_; _; _]).
erewrite <- typed_subst_head_simpl_3 in Hrd5; eauto;
last by simpl; rewrite Hvvsl map_length.
asimpl in Hrd5.
apply (nsteps_bind (fill_item (AppRCtx (RecV _)))) in Hrd5.
destruct Hrd5 as (m & σ6 & w5 & Hm & Hrd51 & Hrd52). simpl in *.
iMod (allocate_full_state σ4) as (γ3) "[H341 _]".
iMod (allocate_full_state σ1') as (γ4) "[H411 _]".
iPoseProof (fundamentalNN.binary_fundamental _ _ _ Ht1 with "[]") as "He1";
iFrame "#".
iSpecialize ("He1" $! _ _ _ with "[$H411]"); simpl.
rewrite ic_eq /ic_def /=.
iSpecialize ("He1" $! _ _ _ _ with "[$H341]"); auto.
iMod (allocate_full_state σ1) as (γ5) "[H511 _]".
iMod (allocate_full_state σσ1) as (γ6) "[H611 _]".
iPoseProof (fundamentalNN.binary_fundamental _ _ _ Ht1 with "[]") as "He1'";
iFrame "#".
iSpecialize ("He1'" $! _ _ _ with "[$H611]"); simpl.
rewrite ic_eq /ic_def /=.
iSpecialize ("He1'" $! _ _ _ _ with "[$H511]"); auto.
iPoseProof (fundamentalNN.binary_fundamental _ _ _ Ht1 with "[]") as "He1''";
iFrame "#".
iSpecialize ("He1''" $! _ _ _ with "[$Hσσ]"); simpl.
rewrite ic_eq /ic_def /=.
iSpecialize ("He1''" $! _ _ _ _ with "[$Hσ4]"); auto.
iCombine "He1" "He1''" as "He1".
iPoseProof (future_cancel_2 with "[He1 He1']") as "He"; first trivial.
{ iSplitL "He1'"; iFrame. }
iMod "He" as "[He1 [[He2 Hσ26] [He3 Hσ36]]]".
iDestruct "He2" as (σ7' v7') "(Hrd7 & Hσ7' & #Hvv7)".
iDestruct "Hrd7" as %Hrd7.
destruct (nsteps_deterministic' Hrd7 Hrd2) as [?[??]]; subst.
iDestruct "He3" as (σ8' v8') "(Hrd8 & Hσ8' & #Hvv8)".
iDestruct "Hrd8" as %Hrd8.
replace (k- k) with 0 by omega. rewrite future_unfold_O.
iMod "He1" as "[He1 Hσ5]".
iDestruct "He1" as (σ9' v9') "(Hrd9 & Hσ9' & #Hvv9)".
iDestruct "Hrd9" as %Hrd9.
destruct (nsteps_deterministic' Hrd8 Hrd9) as [?[??]]; subst.
apply rec_red_step in Hrd52. destruct Hrd52 as [m [Hkm Hrd52]].
assert (τ :: TArrow τ' τ'' :: τ' :: Γ ⊢ₜ
e2'.[ren (1 .: 2 .: 0 .: (+3))] : τ'') as Ht2''.
{ rewrite -lookup_swap_3; eapply (context_swap_typed [] [_; _] [_]); auto. }
match type of Hrd52 with
nsteps _ _ (?A, _) _ =>
replace A with
e2'.[ren (1 .: 2 .: 0 .: (+3))]
.[env_subst (((w, v9') :: vv2 :: vv1 :: vvs).*1)]
in Hrd52; last first
end.
{ simpl.
erewrite <- typed_subst_head_simpl_3; eauto;
last rewrite map_length -Hvvsl //.
asimpl.
erewrite <- typed_subst_head_simpl_3; eauto;
last rewrite map_length -Hvvsl //. }
iPoseProof (fundamentalNN.binary_fundamental _ _ _
Ht2'' _ (((w, v9') :: vv2 :: vv1 :: vvs)) with "[]") as "He2".
{ rewrite ?interp_env_cons; iFrame "#". }
simpl.
repeat (erewrite <- typed_subst_head_simpl_3; eauto;
last rewrite map_length -Hvvsl //).
simpl in Hrd52.
erewrite <- typed_subst_head_simpl_3 in Hrd52; eauto;
last rewrite map_length -Hvvsl //.
iSpecialize ("He2" $! _ _ _ with "[$Hσ8']"); simpl.
rewrite ic_eq /ic_def /=.
iSpecialize ("He2" $! _ _ _ _ with "[$Hσ36]"); auto.
iModIntro. iMod "He2" as "[He2 Hσ8']".
iDestruct "He2" as (σ10 v10) "(Hrd10 & Hσ10 & #Hvv10)".
iDestruct "Hrd10" as %Hrd10.
iModIntro; iFrame. iExists _, _; iFrame; iFrame "#".
iPureIntro.
replace k' with (k + S m) by omega.
eapply nsteps_trans.
+ eapply (@nsteps_ctx det_lang (fill_item (AppRCtx (RecV _))));
eauto using (@ectxi_lang_ctx_item _ _ _ _ det_ectxi_lang).
asimpl. eauto.
+ eapply (nsteps_l _ _ (_, _) (_, _)); simpl.
{ apply head_prim_step. econstructor; eauto. }
asimpl in Hrd10.
match goal with
|- nsteps _ _ (?A, _) _ =>
replace A with
e2'.[env_subst (vv2.2 :: vv1.2 :: v9' :: (vvs.*2))]
end.
{ erewrite <- typed_subst_head_simpl_3; eauto;
last rewrite map_length -Hvvsl //. }
{ match goal with
|- _ = e2'.[_].[?A].[?B,_/] => remember B; remember A as z; asimpl
end. rewrite Heqz /env_subst. asimpl.
replace (((+3) >>>
(λ x : var,
from_option id (Var x)
(of_val <$> (vv2.2 :: vv1.2 :: w' :: vvs.*2) !! x))))
with ((+3) >>> env_subst (vv2.2 :: vv1.2 :: w' :: vvs.*2))
by trivial.
rewrite (env_subst_irrel [vv2.2;vv1.2;w'] [vv2.2;vv1.2;v9']); auto.
unfold env_subst; asimpl.
replace (((+3) >>>
(λ x : var,
from_option id (Var x)
(of_val <$>
(vv2.2 :: vv1.2 :: v9' :: list_fmap (val * val)%type val snd vvs)
!! x))))
with ((+3) >>> env_subst (vv2.2 :: vv1.2 :: v9' :: vvs.*2))
by trivial.
by rewrite f_unfold_3. }
Qed.