RunST.rec_hoisting.rec_hoisting
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 RunST.rec_hoisting Require Import
rec_hoisting_part1 rec_hoisting_part2 rec_hoisting_part3.
From iris.proofmode Require Import tactics.
Hint Resolve to_of_val.
Lemma Rec_Hoisting_CR1 Γ e1 e1' τ e2 e2' τ' τ'' :
Γ ⊢ₜ e1 : τ → [TArrow τ' τ''; τ'; τ] ++ Γ ⊢ₜ e2 : τ'' →
Γ ⊢ₜ e1' : τ → [TArrow τ' τ''; τ'; τ] ++ Γ ⊢ₜ e2' : τ'' →
(∀ `{ICG_ST Σ} `{LogRelG Σ}, Γ ⊨ e1 ≤log≤ e1' : τ) →
(∀ `{ICG_ST Σ} `{LogRelG Σ},
[TArrow τ' τ''; τ'; τ] ++ Γ ⊨ e2 ≤log≤ e2' : τ'') →
Γ ⊨ App (Rec (Rec e2).[ren (+1)]) e1 ≤ctx≤
App (Rec (Rec (App (Rec e2'.[ren (+2)]) e1'.[ren (+3)])).[ren (+1)]) e1'
: TArrow τ' τ''.
Proof.
intros Ht1 Ht2 Ht1' Ht2' HLR1 HLR2.
apply binary_soundness; eauto.
- econstructor; last eauto. econstructor.
eapply (context_weakening [_]). econstructor; eauto.
- econstructor; eauto. econstructor.
eapply (context_weakening [_]).
econstructor. econstructor; last by eapply (context_weakening [_; _; _]).
econstructor. eapply (context_weakening [_; _]). eauto.
- intros. eapply Rec_Hoisting_LR1; eauto.
Qed.
Lemma Rec_Hoisting_CR2 Γ e1' τ e2' τ' τ'' :
Γ ⊢ₜ e1' : τ → [TArrow τ' τ''; τ'; τ] ++ Γ ⊢ₜ e2' : τ'' →
Γ ⊨ App (Rec (Rec (App (Rec e2'.[ren (+2)]) e1'.[ren (+3)])).[ren (+1)]) e1'
≤ctx≤
App (Rec (Rec (App (Rec e2'.[(ren (lookup_n (swap_list 1 2)))].[ren (+1)])
e1'.[ren (+2)])).[ren (+2)]) e1'
: TArrow τ' τ''.
Proof.
intros Ht1 Ht2.
apply binary_soundnessNN; eauto.
- econstructor; eauto. econstructor.
eapply (context_weakening [_]).
econstructor. econstructor; last by eapply (context_weakening [_; _; _]).
econstructor. eapply (context_weakening [_; _]). eauto.
- econstructor; eauto. econstructor.
eapply (context_weakening [_; _]).
econstructor; eauto.
econstructor; last by eapply (context_weakening [_; _]).
econstructor. eapply (context_weakening [_]).
eapply (context_swap_typed [] [_; _] [_]); auto.
- intros. eapply Rec_Hoisting_LR2; eauto.
Qed.
Lemma Rec_Hoisting_CR3 Γ e1' τ e2' τ' τ'' :
Γ ⊢ₜ e1' : τ → [TArrow τ' τ''; τ'; τ] ++ Γ ⊢ₜ e2' : τ'' →
Γ ⊨ App (Rec (Rec (App (Rec e2'.[(ren (lookup_n (swap_list 1 2)))].[ren (+1)])
e1'.[ren (+2)])).[ren (+2)]) e1'
≤ctx≤
Rec (App (Rec e2'.[(ren (lookup_n (swap_list 1 2)))].[ren (+1)])
e1'.[ren (+2)])
: TArrow τ' τ''.
Proof.
intros Ht1 Ht2.
apply binary_soundness; eauto.
- econstructor; eauto. econstructor.
eapply (context_weakening [_; _]).
econstructor; eauto.
econstructor; last by eapply (context_weakening [_; _]).
econstructor. eapply (context_weakening [_]).
eapply (context_swap_typed [] [_; _] [_]); auto.
- econstructor. econstructor; last by eapply (context_weakening [_; _]).
econstructor.
eapply (context_weakening [_]).
eapply (context_swap_typed [] [_; _] [_]); auto.
- intros. eapply Rec_Hoisting_LR3; eauto.
Qed.
Lemma Rec_Hoisting_basic Γ e1 e1' τ e2 e2' τ' τ'' :
Γ ⊢ₜ e1 : τ → [TArrow τ' τ''; τ'; τ] ++ Γ ⊢ₜ e2 : τ'' →
Γ ⊢ₜ e1' : τ → [TArrow τ' τ''; τ'; τ] ++ Γ ⊢ₜ e2' : τ'' →
(∀ `{ICG_ST Σ} `{LogRelG Σ}, Γ ⊨ e1 ≤log≤ e1' : τ) →
(∀ `{ICG_ST Σ} `{LogRelG Σ},
[TArrow τ' τ''; τ'; τ] ++ Γ ⊨ e2 ≤log≤ e2' : τ'') →
Γ ⊨ App (Rec (Rec e2).[ren (+1)]) e1 ≤ctx≤
Rec (App (Rec e2'.[(ren (lookup_n (swap_list 1 2)))].[ren (+1)])
e1'.[ren (+2)])
: TArrow τ' τ''.
Proof.
intros Ht1 Ht2 Ht1' Ht2' HLR1 HLR2.
eapply ctx_refines_trans.
eapply Rec_Hoisting_CR1; eauto.
eapply ctx_refines_trans.
eapply Rec_Hoisting_CR2; eauto.
eapply Rec_Hoisting_CR3; eauto.
Qed.
Lemma Rec_Hoisting Γ e1 τ e2 τ' τ'' :
Γ ⊢ₜ e1 : τ → [TArrow τ' τ''; τ'; τ] ++ Γ ⊢ₜ e2 : τ'' →
Γ ⊨ App (Rec (Rec e2).[ren (+1)]) e1 ≤ctx≤
Rec (App (Rec e2.[(ren (lookup_n (swap_list 1 2)))].[ren (+1)])
e1.[ren (+2)])
: TArrow τ' τ''.
Proof.
intros.
eapply Rec_Hoisting_basic; eauto.
- intros; apply binary_fundamental; auto.
- intros; apply binary_fundamental; auto.
Qed.
future contextual_refinement saved_pred ST_Lang_reduction
logrel_shared soundness soundnessNN logrelNN fundamentalNN
logrel fundamental.
From RunST.rec_hoisting Require Import
rec_hoisting_part1 rec_hoisting_part2 rec_hoisting_part3.
From iris.proofmode Require Import tactics.
Hint Resolve to_of_val.
Lemma Rec_Hoisting_CR1 Γ e1 e1' τ e2 e2' τ' τ'' :
Γ ⊢ₜ e1 : τ → [TArrow τ' τ''; τ'; τ] ++ Γ ⊢ₜ e2 : τ'' →
Γ ⊢ₜ e1' : τ → [TArrow τ' τ''; τ'; τ] ++ Γ ⊢ₜ e2' : τ'' →
(∀ `{ICG_ST Σ} `{LogRelG Σ}, Γ ⊨ e1 ≤log≤ e1' : τ) →
(∀ `{ICG_ST Σ} `{LogRelG Σ},
[TArrow τ' τ''; τ'; τ] ++ Γ ⊨ e2 ≤log≤ e2' : τ'') →
Γ ⊨ App (Rec (Rec e2).[ren (+1)]) e1 ≤ctx≤
App (Rec (Rec (App (Rec e2'.[ren (+2)]) e1'.[ren (+3)])).[ren (+1)]) e1'
: TArrow τ' τ''.
Proof.
intros Ht1 Ht2 Ht1' Ht2' HLR1 HLR2.
apply binary_soundness; eauto.
- econstructor; last eauto. econstructor.
eapply (context_weakening [_]). econstructor; eauto.
- econstructor; eauto. econstructor.
eapply (context_weakening [_]).
econstructor. econstructor; last by eapply (context_weakening [_; _; _]).
econstructor. eapply (context_weakening [_; _]). eauto.
- intros. eapply Rec_Hoisting_LR1; eauto.
Qed.
Lemma Rec_Hoisting_CR2 Γ e1' τ e2' τ' τ'' :
Γ ⊢ₜ e1' : τ → [TArrow τ' τ''; τ'; τ] ++ Γ ⊢ₜ e2' : τ'' →
Γ ⊨ App (Rec (Rec (App (Rec e2'.[ren (+2)]) e1'.[ren (+3)])).[ren (+1)]) e1'
≤ctx≤
App (Rec (Rec (App (Rec e2'.[(ren (lookup_n (swap_list 1 2)))].[ren (+1)])
e1'.[ren (+2)])).[ren (+2)]) e1'
: TArrow τ' τ''.
Proof.
intros Ht1 Ht2.
apply binary_soundnessNN; eauto.
- econstructor; eauto. econstructor.
eapply (context_weakening [_]).
econstructor. econstructor; last by eapply (context_weakening [_; _; _]).
econstructor. eapply (context_weakening [_; _]). eauto.
- econstructor; eauto. econstructor.
eapply (context_weakening [_; _]).
econstructor; eauto.
econstructor; last by eapply (context_weakening [_; _]).
econstructor. eapply (context_weakening [_]).
eapply (context_swap_typed [] [_; _] [_]); auto.
- intros. eapply Rec_Hoisting_LR2; eauto.
Qed.
Lemma Rec_Hoisting_CR3 Γ e1' τ e2' τ' τ'' :
Γ ⊢ₜ e1' : τ → [TArrow τ' τ''; τ'; τ] ++ Γ ⊢ₜ e2' : τ'' →
Γ ⊨ App (Rec (Rec (App (Rec e2'.[(ren (lookup_n (swap_list 1 2)))].[ren (+1)])
e1'.[ren (+2)])).[ren (+2)]) e1'
≤ctx≤
Rec (App (Rec e2'.[(ren (lookup_n (swap_list 1 2)))].[ren (+1)])
e1'.[ren (+2)])
: TArrow τ' τ''.
Proof.
intros Ht1 Ht2.
apply binary_soundness; eauto.
- econstructor; eauto. econstructor.
eapply (context_weakening [_; _]).
econstructor; eauto.
econstructor; last by eapply (context_weakening [_; _]).
econstructor. eapply (context_weakening [_]).
eapply (context_swap_typed [] [_; _] [_]); auto.
- econstructor. econstructor; last by eapply (context_weakening [_; _]).
econstructor.
eapply (context_weakening [_]).
eapply (context_swap_typed [] [_; _] [_]); auto.
- intros. eapply Rec_Hoisting_LR3; eauto.
Qed.
Lemma Rec_Hoisting_basic Γ e1 e1' τ e2 e2' τ' τ'' :
Γ ⊢ₜ e1 : τ → [TArrow τ' τ''; τ'; τ] ++ Γ ⊢ₜ e2 : τ'' →
Γ ⊢ₜ e1' : τ → [TArrow τ' τ''; τ'; τ] ++ Γ ⊢ₜ e2' : τ'' →
(∀ `{ICG_ST Σ} `{LogRelG Σ}, Γ ⊨ e1 ≤log≤ e1' : τ) →
(∀ `{ICG_ST Σ} `{LogRelG Σ},
[TArrow τ' τ''; τ'; τ] ++ Γ ⊨ e2 ≤log≤ e2' : τ'') →
Γ ⊨ App (Rec (Rec e2).[ren (+1)]) e1 ≤ctx≤
Rec (App (Rec e2'.[(ren (lookup_n (swap_list 1 2)))].[ren (+1)])
e1'.[ren (+2)])
: TArrow τ' τ''.
Proof.
intros Ht1 Ht2 Ht1' Ht2' HLR1 HLR2.
eapply ctx_refines_trans.
eapply Rec_Hoisting_CR1; eauto.
eapply ctx_refines_trans.
eapply Rec_Hoisting_CR2; eauto.
eapply Rec_Hoisting_CR3; eauto.
Qed.
Lemma Rec_Hoisting Γ e1 τ e2 τ' τ'' :
Γ ⊢ₜ e1 : τ → [TArrow τ' τ''; τ'; τ] ++ Γ ⊢ₜ e2 : τ'' →
Γ ⊨ App (Rec (Rec e2).[ren (+1)]) e1 ≤ctx≤
Rec (App (Rec e2.[(ren (lookup_n (swap_list 1 2)))].[ren (+1)])
e1.[ren (+2)])
: TArrow τ' τ''.
Proof.
intros.
eapply Rec_Hoisting_basic; eauto.
- intros; apply binary_fundamental; auto.
- intros; apply binary_fundamental; auto.
Qed.