diff --git a/_CoqProject b/_CoqProject index 47e3af4..a043dab 100644 --- a/_CoqProject +++ b/_CoqProject @@ -31,6 +31,8 @@ theories/LogicalRelation/Weakening.v theories/LogicalRelation/Universe.v theories/LogicalRelation/Neutral.v theories/LogicalRelation/Transitivity.v +theories/LogicalRelation/InstKripke.v +theories/LogicalRelation/EqRedRight.v theories/LogicalRelation/Reduction.v theories/LogicalRelation/NormalRed.v theories/LogicalRelation/Application.v diff --git a/theories/Consequences.v b/theories/Consequences.v index 3a9a45a..4707700 100644 --- a/theories/Consequences.v +++ b/theories/Consequences.v @@ -56,20 +56,20 @@ Section NatCanonicityInduction. Qed. Lemma nat_red_empty_ind : - (forall t, [ε ||-Nat t : tNat | red_nat_empty] -> + (forall t u, [ε ||-Nat t ≅ u : tNat | red_nat_empty] -> ∑ n : nat, [ε |- t ≅ n : tNat]) × - (forall t, NatProp red_nat_empty t -> ∑ n : nat, [ε |- t ≅ n : tNat]). + (forall t u, NatPropEq red_nat_empty t u -> ∑ n : nat, [ε |- t ≅ n : tNat]). Proof. - apply NatRedInduction. - - intros * [? []] ? _ [n] ; refold. + apply NatRedEqInduction. + - intros * [? []] ? ? _ [n] ; refold. exists n. now etransitivity. - exists 0 ; cbn. now repeat constructor. - - intros ? _ [n]. + - intros ? ? _ [n]. exists (S n) ; simpl. now econstructor. - - intros ? [? [? _ _]]. + - intros ? ? [? ? []]. exfalso. now eapply no_neutral_empty_ctx. Qed. @@ -80,7 +80,7 @@ Section NatCanonicityInduction. intros Ht. assert [LRNat_ one red_nat_empty | ε ||- t : tNat] as ?%nat_red_empty_ind. { - apply Fundamental in Ht as [?? Vt%reducibleTm]. + apply Fundamental in Ht as [?? Vt%reducibleTmEq]. irrelevance. } now assumption. diff --git a/theories/DeclarativeSubst.v b/theories/DeclarativeSubst.v index 801c5aa..51df620 100644 --- a/theories/DeclarativeSubst.v +++ b/theories/DeclarativeSubst.v @@ -7,7 +7,7 @@ From LogRel Require Import LogicalRelation Validity Fundamental. From LogRel.LogicalRelation Require Import Induction Escape Irrelevance Transitivity. From LogRel.Substitution Require Import Properties Irrelevance. -(** Currently, this is obtained as a consequence of the fundamental lemma. +(** Currently, this is obtained as a consequence of the fundamental lemma. However, it could be alternatively proven much earlier, by a direct induction. *) Set Printing Primitive Projection Parameters. @@ -44,20 +44,14 @@ Proof. unshelve eapply transEq. 4: now apply Vconv. 2-3: unshelve eapply VB ; tea. - all: irrValid. + eapply urefl; now eapply irrelevanceSubstEq. - intros * Ht * HΔ Hσ. unshelve eapply Fundamental_subst_conv in Hσ as []. 1,3: boundary. - apply Fundamental in Ht as [VΓ VA Vt Vu Vtu] ; cbn in *. + apply Fundamental in Ht as [VΓ VA Vtu] ; cbn in *. unshelve eapply escapeEqTerm. 2: now unshelve eapply VA ; tea ; irrValid. - cbn. - eapply transEqTerm. - + cbn. - unshelve eapply Vtu. - + cbn. - eapply Vu. - all: irrValid. + cbn. unshelve eapply Vtu. Qed. @@ -149,16 +143,16 @@ Section MoreSubst. Theorem typing_substmap1 Γ T : (forall (t : term), [Γ ,, T |- t : T⟨↑⟩] -> - forall (A : term), [Γ,, T |- A] -> + forall (A : term), [Γ,, T |- A] -> [Γ,, T |- A[t]⇑]) × (forall (t : term), [Γ ,, T |- t : T⟨↑⟩] -> - forall (A u : term), [Γ,, T |- u : A] -> + forall (A u : term), [Γ,, T |- u : A] -> [Γ,, T |- u[t]⇑ : A[t]⇑]) × (forall (t t' : term), [Γ ,, T |- t ≅ t' : T⟨↑⟩] -> forall (A B : term), [Γ,, T |- A ≅ B] -> [Γ,, T |- A[t]⇑ ≅ B[t']⇑]) × (forall (t t' : term), [Γ ,, T |- t ≅ t' : T⟨↑⟩] -> - forall (A u v : term), [Γ,, T |- u ≅ v : A] -> + forall (A u v : term), [Γ,, T |- u ≅ v : A] -> [Γ,, T |- u[t]⇑ ≅ v[t']⇑ : A[t]⇑]). Proof. repeat match goal with |- _ × _ => split end. diff --git a/theories/Fundamental.v b/theories/Fundamental.v index 73719d1..e165736 100644 --- a/theories/Fundamental.v +++ b/theories/Fundamental.v @@ -68,8 +68,6 @@ Module FundTmEq. : Type := { VΓ : [||-v Γ ]; VA : [ Γ ||-v< one > A | VΓ ]; - Vt : [ Γ ||-v< one > t : A | VΓ | VA ]; - Vu : [ Γ ||-v< one > u : A | VΓ | VA ]; Vtu : [ Γ ||-v< one > t ≅ u : A | VΓ | VA ]; }. Arguments FundTmEq {_ _ _ _ _ _ _ _ _ _ _ _}. @@ -94,9 +92,7 @@ Module FundSubstConv. {Γ Δ : context} {wfΓ : [|- Γ]} {σ σ' : nat -> term} : Type := { VΔ : [||-v Δ ] ; - Vσ : [VΔ | Γ ||-v σ : Δ | wfΓ] ; - Vσ' : [VΔ | Γ ||-v σ' : Δ | wfΓ ] ; - Veq : [VΔ | Γ ||-v σ ≅ σ' : Δ | wfΓ | Vσ] ; + Veq : [VΔ | Γ ||-v σ ≅ σ' : Δ | wfΓ ] ; }. Arguments FundSubstConv {_ _ _ _ _ _ _ _ _ _ _ _}. End FundSubstConv. @@ -117,9 +113,7 @@ Section Fundamental. Lemma FundConNil : FundCon ε. Proof. unshelve econstructor. - + unshelve econstructor. - - intros; exact unit. - - intros; exact unit. + + unshelve econstructor; intros; exact unit. + constructor. Qed. @@ -132,14 +126,7 @@ Section Fundamental. Lemma FundTyU (Γ : context) (fΓ : FundCon Γ) : FundTy Γ U. Proof. - unshelve econstructor. - - assumption. - - unshelve econstructor. - + intros * _. apply LRU_. - econstructor; tea; [constructor|]. - cbn; eapply redtywf_refl; gen_typing. - + intros * _ _. simpl. constructor. - cbn; eapply redtywf_refl; gen_typing. + now unshelve (econstructor; eapply UValid). Qed. Lemma FundTyPi (Γ : context) (F G : term) @@ -157,21 +144,17 @@ Section Fundamental. (fA : FundTm Γ U A) : FundTy Γ A. Proof. - destruct fA as [ VΓ VU [ RA RAext ] ]. econstructor. - unshelve econstructor. - - intros * vσ. - eapply UnivEq. exact (RA _ _ wfΔ vσ). - - intros * vσ' vσσ'. - eapply UnivEqEq. exact (RAext _ _ _ wfΔ vσ vσ' vσσ'). + destruct fA as [ VΓ VU RA]. econstructor. + now eapply univValid. Qed. Lemma FundTmVar : forall (Γ : context) (n : nat) decl, FundCon Γ -> in_ctx Γ n decl -> FundTm Γ decl (tRel n). Proof. - intros Γ n d FΓ hin. + intros Γ n d FΓ hin. unshelve econstructor; tea. - + eapply in_ctx_valid in hin as []; now eapply embValidTyOne. + + eapply in_ctx_valid in hin as []; now eapply embValidTyOne. + now eapply varnValid. Qed. @@ -181,9 +164,9 @@ Section Fundamental. Proof. intros * [] []; econstructor. eapply PiValidU; irrValid. - Unshelve. - 3: eapply UValid. - 2: eapply univValid. + Unshelve. + 3: eapply UValid. + 2: eapply univValid. all:tea. Qed. @@ -209,7 +192,7 @@ Section Fundamental. FundTm Γ A t -> FundTyEq Γ A B -> FundTm Γ B t. Proof. - intros * [] []; econstructor. + intros * [] []; econstructor. eapply conv; irrValid. Unshelve. all: tea. Qed. @@ -246,13 +229,16 @@ Section Fundamental. Proof. intros * [] []; unshelve econstructor; tea. 1:irrValid. eapply transValidTyEq; irrValid. - Unshelve. tea. + Unshelve. 2:tea. Qed. Lemma FundTyEqUniv : forall (Γ : context) (A B : term), FundTmEq Γ U A B -> FundTyEq Γ A B. Proof. - intros * []; unshelve econstructor; tea. + intros * [?? VAB]. + pose proof (lreflValidTm _ VAB). + pose proof (ureflValidTm VAB). + unshelve econstructor; tea. 1,2: now eapply univValid. now eapply univEqValid. Qed. @@ -263,12 +249,8 @@ Section Fundamental. FundTm Γ A a -> FundTmEq Γ B[a..] (tApp (tLambda A t) a) t[a..]. Proof. intros * [] [] []; econstructor. - - eapply appValid. eapply lamValid. irrValid. - - unshelve epose (substSTm _ _). - 8-12: irrValid. - tea. - unshelve epose (betaValid VA _ _ _). 2,5-7:irrValid. - Unshelve. all: tea; try irrValid. + Unshelve. 1:tea. eapply substS; tea; irrValid. Qed. Lemma FundTmEqPiCong : forall (Γ : context) (A B C D : term), @@ -279,25 +261,17 @@ Section Fundamental. Proof. intros * [] [] []. assert (VA' : [Γ ||-v A | VΓ]) by now eapply univValid. - assert [Γ ||-v A ≅ B | VΓ | VA'] by (eapply univEqValid; irrValid). + assert (VAB : [Γ ||-v A ≅ B | VΓ | VA']) by (eapply univEqValid; irrValid). + pose proof (ureflValidTy VAB). opector; tea. - eapply UValid. - - edestruct FundTmProd. 3: irrValid. - all: unshelve econstructor; irrValid. - - edestruct FundTmProd. 3: irrValid. - 1: unshelve econstructor; irrValid. - unshelve econstructor. - + eapply validSnoc; now eapply univValid. - + eapply UValid. - + eapply irrelevanceTmLift; irrValid. - unshelve epose (PiCongTm _ _ _ _ _ _ _ _ _ _ _). - 16: irrValid. + 17: irrValid. 2: tea. - 2,3,8: irrValid. - all: try irrValid. - + eapply univValid; irrValid. - + eapply irrelevanceLift; irrValid. - + eapply irrelevanceTmLift; irrValid. + 2,3,8: eapply irrelevanceTmEq; tea; first [now eapply lreflValidTm| now eapply ureflValidTm]. + all: try (irrValid + eapply irrelevanceTmEq ;now eapply ureflValidTm). + + eapply UValid. + + eapply irrelevanceTmLift; tea; now eapply irrelevanceTmEq, ureflValidTm. Unshelve. all: try irrValid. Qed. @@ -308,13 +282,8 @@ Section Fundamental. FundTmEq Γ B[a..] (tApp f a) (tApp g b). Proof. intros * [] []; econstructor. - - eapply appValid; irrValid. - - eapply conv. 2: eapply appValid; irrValid. - eapply substSΠeq; try irrValid. - 1: eapply reflValidTy. - now eapply symValidTmEq. - - eapply appcongValid; irrValid. - Unshelve. all: irrValid. + eapply appcongValid; first [irrValid | now eapply ureflValidTm]. + Unshelve. 1: irrValid. now eapply lreflValidTm. Qed. Lemma FundTmEqLambdaCong : forall (Γ : context) (t u A A' A'' B : term), @@ -324,158 +293,34 @@ Section Fundamental. FundTmEq (Γ,, A) B t u -> FundTmEq Γ (tProd A B) (tLambda A' t) (tLambda A'' u). Proof. intros * [VΓ] [? ? VA'] [? ? VA''] []. - assert (VΠAB : [Γ ||-v< one > tProd A B | VΓ]). - { unshelve eapply PiValid; irrValid. } - assert (VB' : [Γ,, A' ||-v< one > B | validSnoc _ VA']). - { eapply irrelevanceLift; [tea|]; irrValid. } - assert (VB'' : [Γ,, A'' ||-v< one > B | validSnoc _ VA'']). - { eapply irrelevanceLift; [tea|]; irrValid. } - assert (Vλt : [Γ ||-v< one > tLambda A' t : tProd A B | VΓ | VΠAB ]). - { eapply conv; [|eapply lamValid]. - + eapply symValidTyEq, PiCong; tea; try irrValid. - eapply reflValidTy. - + eapply irrelevanceTmLift; irrValid. - } - assert (Vλu : [Γ ||-v< one > tLambda A'' u : tProd A B | VΓ | VΠAB ]). - { eapply conv; [|eapply lamValid]. - + eapply symValidTyEq, PiCong; tea; try irrValid. - eapply reflValidTy. - + eapply irrelevanceTmLift; irrValid. + econstructor. + eapply conv. + 2:{ eapply lamCongValid; tea. + + unshelve (eapply transValidTyEq; tea; eapply symValidTyEq); try irrValid. + + now eapply convCtx1. + + eapply convEqCtx1; tea; now eapply reflValidTy. + + now eapply convTmEqCtx1. } - Unshelve. all: try irrValid. - assert [Γ,, A ||-v< one > A⟨@wk1 Γ A⟩ | validSnoc VΓ VA]. - { apply wk1ValidTy; irrValid. } - assert (VΓAA' : [Γ,, A ||-v< one > A'⟨@wk1 Γ A⟩ | validSnoc VΓ VA]). - { apply wk1ValidTy; irrValid. } - assert (VΓAA'' : [Γ,, A ||-v< one > A''⟨@wk1 Γ A⟩ | validSnoc VΓ VA]). - { apply wk1ValidTy; irrValid. } - assert (VAdequate (ta := ta) (Γ,, A) VR). - { now eapply validSnoc. } - assert (VAdequate (ta := ta) (Γ,, A,, A'⟨@wk1 Γ A⟩) VR). - { unshelve eapply validSnoc; tea; try irrValid. } - assert (VAdequate (ta := ta) (Γ,, A,, A''⟨@wk1 Γ A⟩) VR). - { unshelve eapply validSnoc; tea; try irrValid. } - assert (eq0 : forall t, t⟨upRen_term_term ↑⟩[(tRel 0)..] = t). - { clear; intros; bsimpl; apply idSubst_term; intros []; reflexivity. } - econstructor; tea. - eapply irrelevanceTmEq, etaeqValid; try irrValid. - eapply transValidTmEq; [|eapply transValidTmEq]; [|eapply irrelevanceTmEq; tea|]. - + eapply irrelevanceTmEq'; [eapply eq0|]. - eapply transValidTmEq; [eapply betaValid|]; refold. - - eapply irrelevanceTm'. - 2: eapply (wkValidTm (A := B)) with (ρ := wk_up A' (@wk1 Γ A)). - * now bsimpl. - * eapply irrelevanceTmLift; irrValid. - - eapply irrelevanceTmEq'; [symmetry; eapply eq0|]. - match goal with |- [_ ||-v< _ > ?t ≅ ?u : _ | _ | _ ] => assert (Hrw : t = u) end. - { bsimpl; apply idSubst_term; intros []; reflexivity. } - set (t' := t) at 2; rewrite Hrw. - apply reflValidTm; tea. - + apply symValidTmEq; eapply irrelevanceTmEq'; [eapply eq0|]. - eapply transValidTmEq; [eapply betaValid|]; refold. - - eapply irrelevanceTm'. - 2: eapply (wkValidTm (A := B)) with (ρ := wk_up A'' (@wk1 Γ A)). - * now bsimpl. - * eapply irrelevanceTmLift; irrValid. - - eapply irrelevanceTmEq'; [symmetry; eapply eq0|]. - match goal with |- [_ ||-v< _ > ?t ≅ ?u : _ | _ | _ ] => assert (Hrw : t = u) end. - { bsimpl; apply idSubst_term; intros []; reflexivity. } - set (u' := u) at 2; rewrite Hrw. - apply reflValidTm; tea. + unshelve (eapply PiCong; [|eapply symValidTyEq|eapply reflValidTy]; irrValid); try irrValid. Unshelve. - all: refold; try irrValid. - * unshelve eapply irrelevanceTy; tea. - rewrite <- (wk_up_wk1_ren_on Γ A' A). - eapply wkValidTy, irrelevanceLift; irrValid. - * eapply conv; [now eapply irrelevanceTyEq, wk1ValidTyEq|]. - eapply irrelevanceTm'; [symmetry; eapply wk1_ren_on|]. - eapply var0Valid'. - * eapply irrelevanceTy. - rewrite <- (wk_up_wk1_ren_on Γ A'' A). - eapply wkValidTy, irrelevanceLift; irrValid. - * eapply conv; [now eapply irrelevanceTyEq, wk1ValidTyEq|]. - eapply irrelevanceTm'; [symmetry; eapply wk1_ren_on|]. - eapply var0Valid'. - Unshelve. - all: try irrValid. - 2,4: rewrite <- (wk1_ren_on Γ A); irrValid. + 5: tea. + 3: now eapply convCtx1. + 2: irrValid. + unshelve eapply PiValid; irrValid. Qed. Lemma FundTmEqFunEta : forall (Γ : context) (f A B : term), FundTm Γ (tProd A B) f -> FundTmEq Γ (tProd A B) (tLambda A (tApp f⟨↑⟩ (tRel 0))) f. Proof. - intros * [VΓ VΠ Vf]. - assert (eq0 : forall t, t⟨upRen_term_term ↑⟩[(tRel 0)..] = t). - { clear; intros; bsimpl; apply idSubst_term; intros []; reflexivity. } - assert (VA : [Γ ||-v< one > A | VΓ]). - { now eapply PiValidDom. } - assert (VΓA := validSnoc VΓ VA). - assert (VΓAA : [Γ,, A ||-v< one > A⟨@wk1 Γ A⟩ | VΓA]). - { now eapply irrelevanceTy, wk1ValidTy. } - assert (VΓAA0 : VAdequate (ta := ta) (Γ,, A,, A⟨@wk1 Γ A⟩) VR). - { now eapply validSnoc. } - assert (VΓAA' : [Γ,, A ||-v< one > A⟨↑⟩ | VΓA]). - { now rewrite wk1_ren_on in VΓAA. } - assert [Γ,, A ||-v< one > B | VΓA]. - { eapply irrelevanceTy, PiValidCod. } - assert ([Γ,, A ||-v< one > tProd A⟨@wk1 Γ A⟩ B⟨upRen_term_term ↑⟩ | VΓA]). - { rewrite <- (wk_up_wk1_ren_on Γ A A). - now eapply irrelevanceTy, (wk1ValidTy (A := tProd A B)). } - assert ([Γ,, A ||-v< one > tRel 0 : A⟨wk1 A⟩ | VΓA | VΓAA]). - { eapply irrelevanceTm'; [rewrite wk1_ren_on; reflexivity|]. - unshelve apply var0Valid'. } - assert ([(Γ,, A),, A⟨wk1 A⟩ ||-v< one > tProd A⟨↑⟩⟨↑⟩ B⟨upRen_term_term ↑⟩⟨upRen_term_term ↑⟩ | VΓAA0]). - { assert (eq1 : (tProd A B)⟨@wk1 Γ A⟩⟨@wk1 (Γ,, A) A⟨@wk1 Γ A⟩⟩ = tProd (A⟨↑⟩⟨↑⟩) (B⟨upRen_term_term ↑⟩⟨upRen_term_term ↑⟩)). - { rewrite wk1_ren_on, wk1_ren_on; reflexivity. } - eapply irrelevanceTy'; [|eapply eq1|reflexivity]. - now eapply wkValidTy, wkValidTy. } - assert ([Γ ||-v< one > tLambda A (tApp f⟨↑⟩ (tRel 0)) : tProd A B | VΓ | VΠ]). - { eapply irrelevanceTm, lamValid. - eapply irrelevanceTm'; [apply eq0|eapply (appValid (F := A⟨@wk1 Γ A⟩))]. - rewrite <- (wk1_ren_on Γ A). - eapply irrelevanceTm'; [|now eapply wkValidTm]. - rewrite <- (wk_up_wk1_ren_on Γ A A). - reflexivity. } - Unshelve. all: try irrValid. - econstructor; tea. - eapply irrelevanceTmEq, etaeqValid; try irrValid. - eapply transValidTmEq; refold. - + eapply irrelevanceTmEq'; [eapply eq0|]. - eapply betaValid; refold. - eapply irrelevanceTm'; [apply eq0|]. - eapply appValid. - rewrite <- shift_upRen. - assert (eq1 : (tProd A B)⟨@wk1 Γ A⟩⟨@wk1 (Γ,, A) A⟨@wk1 Γ A⟩⟩ = tProd (A⟨↑⟩⟨↑⟩) (B⟨upRen_term_term ↑⟩⟨upRen_term_term ↑⟩)). - { rewrite wk1_ren_on, wk1_ren_on; reflexivity. } - eapply irrelevanceTm'; [eapply eq1|]. - rewrite <- (wk1_ren_on (Γ,, A) A⟨@wk1 Γ A⟩). - now eapply wkValidTm, wkValidTm. - + refold. - match goal with |- [_ ||-v< _ > ?t ≅ ?u : _ | _ | _ ] => assert (Hrw : t = u) end. - { rewrite <- eq0; now bsimpl. } - rewrite Hrw. - apply reflValidTm; tea. - eapply irrelevanceTm'; [apply eq0|]. - eapply (appValid (F := A⟨@wk1 Γ A⟩)). - eapply irrelevanceTm'; [|now eapply wk1ValidTm]. - now rewrite <- (wk_up_wk1_ren_on Γ A A). - Unshelve. - all: refold; try irrValid. - { unshelve eapply irrelevanceTy; [tea|]. - rewrite <- (wk_up_wk1_ren_on Γ A A). - now eapply wkValidTy. } - { eapply (irrelevanceTy' (A := A⟨↑⟩⟨@wk1 (Γ,, A) A⟨@wk1 Γ A⟩⟩)); [|now rewrite wk1_ren_on|reflexivity]. - now eapply wkValidTy. } - { eapply (irrelevanceTm' (A := A⟨@wk1 Γ A⟩⟨↑⟩)); [now rewrite wk1_ren_on|]. - apply var0Valid'. } + intros * [VΓ VΠ Vf]; econstructor. + eapply etaValid; irrValid. Unshelve. - all: try irrValid. - { shelve. } - { rewrite <- (wk1_ren_on (Γ,, A) A⟨@wk1 Γ A⟩). - now eapply wkValidTy. } + + tea. + + now eapply PiValidDom. + + now eapply PiValidCod. Qed. - Lemma FundTmEqFunExt : forall (Γ : context) (f g A B : term), + (* Lemma FundTmEqFunExt : forall (Γ : context) (f g A B : term), FundTy Γ A -> FundTm Γ (tProd A B) f -> FundTm Γ (tProd A B) g -> @@ -485,18 +330,18 @@ Section Fundamental. intros * [] [VΓ0 VA0] [] []. assert [Γ ||-v< one > g : tProd A B | VΓ0 | VA0]. 1:{ - eapply conv. + eapply conv. 2: irrValid. eapply symValidTyEq. eapply PiCong. - eapply irrelevanceLift. + eapply irrelevanceLift. 1,3,4: eapply reflValidTy. irrValid. } - econstructor. - 3: eapply etaeqValid. + econstructor. + 3: eapply etaeqValid. 5: do 2 rewrite wk1_ren_on. Unshelve. all: irrValid. - Qed. + Qed. *) Lemma FundTmEqRefl : forall (Γ : context) (t A : term), FundTm Γ A t -> @@ -518,8 +363,8 @@ Section Fundamental. FundTmEq Γ A t t''. Proof. intros * [] []; econstructor; tea. - 1: irrValid. eapply transValidTmEq; irrValid. + Unshelve. all: tea. Qed. Lemma FundTmEqConv : forall (Γ : context) (t t' A B : term), @@ -528,8 +373,7 @@ Section Fundamental. FundTmEq Γ B t t'. Proof. intros * [] []; econstructor. - 1,2: eapply conv; irrValid. - eapply convEq; irrValid. + eapply conv; irrValid. Unshelve. all: irrValid. Qed. @@ -546,7 +390,7 @@ Section Fundamental. Lemma FundTmZero : forall Γ : context, FundCon Γ -> FundTm Γ tNat tZero. Proof. - intros; unshelve econstructor; tea. + intros; unshelve econstructor; tea. 2:eapply zeroValid. Qed. @@ -566,7 +410,8 @@ Section Fundamental. Proof. intros * [] [] [] []; unshelve econstructor; tea. 2: eapply natElimValid; irrValid. - Unshelve. all: irrValid. + Unshelve. + 2,3: irrValid. Qed. Lemma FundTyEmpty : forall Γ : context, FundCon Γ -> FundTy Γ tEmpty. @@ -586,14 +431,13 @@ Section Fundamental. Proof. intros * [] []; unshelve econstructor; tea. 2: eapply emptyElimValid; irrValid. - Unshelve. 1,2: irrValid. + Unshelve. 2,3: irrValid. Qed. Lemma FundTmEqSuccCong : forall (Γ : context) (n n' : term), FundTmEq Γ tNat n n' -> FundTmEq Γ tNat (tSucc n) (tSucc n'). Proof. intros * []; unshelve econstructor; tea. - 1,2: eapply irrelevanceTm; eapply succValid; irrValid. eapply irrelevanceTmEq; eapply succcongValid; irrValid. Unshelve. all: tea. Qed. @@ -606,32 +450,10 @@ Section Fundamental. FundTmEq Γ tNat n n' -> FundTmEq Γ P[n..] (tNatElim P hz hs n) (tNatElim P' hz' hs' n'). Proof. - intros * [? VP0 VP0'] [VΓ0] [] []. - pose (VN := natValid (l:=one) VΓ0). - assert (VP' : [ _ ||-v P' | validSnoc VΓ0 VN]) by irrValid. - assert [Γ ||-v< one > hz' : P'[tZero..] | VΓ0 | substS VP' (zeroValid VΓ0)]. 1:{ - eapply conv. 2: irrValid. - eapply substSEq. 2,3: irrValid. - 1: eapply reflValidTy. - 2: eapply reflValidTm. - all: eapply zeroValid. - } - assert [Γ ||-v< one > hs' : elimSuccHypTy P' | VΓ0 | elimSuccHypTyValid VΓ0 VP']. 1:{ - eapply conv. 2: irrValid. - eapply elimSuccHypTyCongValid; irrValid. - } - unshelve econstructor; tea. - 2: eapply natElimValid; irrValid. - + eapply conv. - 2: eapply irrelevanceTm; now eapply natElimValid. - eapply symValidTyEq. - eapply substSEq; tea. - 2,3: irrValid. - eapply reflValidTy. - + eapply natElimCongValid; tea; try irrValid. - Unshelve. all: try irrValid. - 1: eapply zeroValid. - 1: unshelve eapply substS; try irrValid. + intros * [? VP0 VP0'] [VΓ0] [] []; unshelve econstructor. + 3:eapply natElimCongValid; try irrValid. + tea. + Unshelve. all: irrValid. Qed. Lemma FundTmEqNatElimZero : forall (Γ : context) (P hz hs : term), @@ -641,9 +463,8 @@ Section Fundamental. FundTmEq Γ P[tZero..] (tNatElim P hz hs tZero) hz. Proof. intros * [] [] []; unshelve econstructor; tea. - 3: irrValid. - 3: eapply natElimZeroValid; irrValid. - eapply natElimValid; irrValid. + 2: eapply natElimZeroValid; try irrValid. + now eapply reflValidTy. Unshelve. irrValid. Qed. @@ -656,17 +477,9 @@ Section Fundamental. (tApp (tApp hs n) (tNatElim P hz hs n)). Proof. intros * [] [] [] []; unshelve econstructor; tea. - 4: eapply natElimSuccValid; irrValid. - 1: eapply natElimValid; irrValid. - eapply simple_appValid. - 2: eapply natElimValid; irrValid. - eapply irrelevanceTm'. - 2: now eapply appValid. - now bsimpl. - Unshelve. all: try irrValid. - eapply simpleArrValid; eapply substS; tea. - 1,2: irrValid. - eapply succValid; irrValid. + 2: eapply natElimSuccValid; try irrValid. + now eapply reflValidTy. + Unshelve. all: irrValid. Qed. Lemma FundTmEqEmptyElimCong : forall (Γ : context) @@ -675,20 +488,9 @@ Section Fundamental. FundTmEq Γ tEmpty n n' -> FundTmEq Γ P[n..] (tEmptyElim P n) (tEmptyElim P' n'). Proof. - intros * [? VP0 VP0'] [VΓ0]. - pose (VN := emptyValid (l:=one) VΓ0). - assert (VP' : [ _ ||-v P' | validSnoc VΓ0 VN]) by irrValid. - unshelve econstructor; tea. - 2: eapply emptyElimValid; irrValid. - + eapply conv. - 2: eapply irrelevanceTm; now eapply emptyElimValid. - eapply symValidTyEq. - eapply substSEq; tea. - 2,3: irrValid. - eapply reflValidTy. - + eapply emptyElimCongValid; tea; try irrValid. - Unshelve. all: try irrValid. - 1: unshelve eapply substS; try irrValid. + intros * [? VP0 VP0'] [VΓ0]; unshelve econstructor; tea. + 2: eapply emptyElimCongValid; irrValid. + Unshelve. all: irrValid. Qed. Lemma FundTySig : forall (Γ : context) (A B : term), @@ -702,11 +504,12 @@ Section Fundamental. FundTm Γ U A -> FundTm (Γ,, A) U B -> FundTm Γ U (tSig A B). Proof. - intros * [] []; unshelve econstructor. - 3: unshelve eapply SigValidU. - 3-5: irrValid. - 1: tea. - unshelve eapply univValid; tea. + intros * [] []; unshelve econstructor; tea. + unshelve epose (SigCongValidU (F:=A) (G:=B) _ _ ). + 6,7,8: irrValid. + + tea. + + now eapply univValid. + + irrValid. Qed. Lemma FundTmPair : forall (Γ : context) (A B a b : term), @@ -759,13 +562,11 @@ Section Fundamental. FundTmEq (Γ,, A) U B B' -> FundTmEq Γ U (tSig A B) (tSig A' B'). Proof. intros * [] [] []; unshelve econstructor. - 5: eapply SigCongTm; tea; try irrValid. - 4: eapply irrelevanceTmLift ; try irrValid; now eapply univEqValid. - 1,2: unshelve eapply SigValidU; try irrValid. - 1,2: now eapply univValid. - 1: eapply UValid. - eapply irrelevanceTmLift ; try irrValid; now eapply univEqValid. - Unshelve. all: solve [ eapply UValid | now eapply univValid | irrValid]. + 3: eapply SigCongValidU ; tea; try irrValid. + 1: tea. + Unshelve. + + unshelve (eapply univValid; irrValid); irrValid. + + irrValid. Qed. Lemma FundTmEqPairCong : forall (Γ : context) (A A' A'' B B' B'' a a' b b' : term), @@ -777,166 +578,41 @@ Section Fundamental. FundTmEq Γ A a a' -> FundTmEq Γ B[a..] b b' -> FundTmEq Γ (tSig A B) (tPair A' B' a b) (tPair A'' B'' a' b'). Proof. - intros * [VΓ VA] [] [] [] [] [] []. - assert (VΣ : [Γ ||-v< one > tSig A B | VΓ]). - { unshelve eapply SigValid; irrValid. } - assert (VA' : [Γ ||-v< one > A' | VΓ]) by irrValid. - assert (VA'' : [Γ ||-v< one > A'' | VΓ]) by irrValid. - assert ([Γ ||-v< one > a : A' | VΓ | VA']). - { eapply conv; [|irrValid]; irrValid. Unshelve. tea. } - assert ([Γ ||-v< one > a : A'' | VΓ | VA'']). - { eapply conv; [irrValid|]; irrValid. Unshelve. tea. } - assert ([Γ ||-v< one > a' : A'' | VΓ | VA'']). - { eapply conv; [irrValid|]; irrValid. Unshelve. tea. } - assert (VΓA' : VAdequate (ta := ta) (Γ,, A') VR) by now eapply validSnoc. - assert (VΓA'' : VAdequate (ta := ta) (Γ,, A'') VR) by now eapply validSnoc. - assert (VA'B : [Γ,, A' ||-v< one > B | VΓA']). - { eapply irrelevanceTy, irrelevanceLift; irrValid. - Unshelve. all: irrValid. } - assert (VA''B : [Γ,, A'' ||-v< one > B | VΓA'']). - { eapply irrelevanceTy, irrelevanceLift; irrValid. - Unshelve. all: irrValid. } - assert (VA'B' : [Γ,, A' ||-v< one > B' | VΓA']). - { eapply irrelevanceTy, irrelevanceLift; try irrValid. - Unshelve. all: irrValid. } - assert (VA''B' : [Γ,, A'' ||-v< one > B' | VΓA'']). - { eapply irrelevanceTy, irrelevanceLift; try irrValid. - Unshelve. all: irrValid. } - assert (VA''B'' : [Γ,, A'' ||-v< one > B'' | VΓA'']). - { eapply irrelevanceTy, irrelevanceLift; try irrValid. - Unshelve. all: irrValid. } - assert (VBa : [Γ ||-v< one > B[a..] | VΓ]). - { irrValid. } - assert (VB'a : [Γ ||-v< one > B'[a..] | VΓ]). - { eapply substS; irrValid. Unshelve. all: irrValid. } - assert (VB''a : [Γ ||-v< one > B''[a..] | VΓ]). - { eapply substS; irrValid. Unshelve. all: irrValid. } - assert (VB''a' : [Γ ||-v< one > B''[a'..] | VΓ]). - { eapply substS; irrValid. Unshelve. all: irrValid. } - assert [Γ ||-v< one > B[a..] ≅ B'[a..] | VΓ | VBa]. - { eapply irrelevanceTyEq, substSEq; try irrValid. - apply reflValidTm. irrValid. Unshelve. all: irrValid. } - assert [Γ ||-v< one > B[a..] ≅ B''[a'..] | VΓ | VBa]. - { eapply irrelevanceTyEq, substSEq; [..|irrValid|irrValid]. - all: irrValid. Unshelve. all: irrValid. } - assert (VΣA'B' : [Γ ||-v< one > tSig A' B' | VΓ]). - { unshelve eapply SigValid; try irrValid. } - assert (VΣA''B'' : [Γ ||-v< one > tSig A'' B'' | VΓ]). - { unshelve eapply SigValid; try irrValid. } - assert ([Γ ||-v< one > tSig A B ≅ tSig A' B' | VΓ | VΣ]). - { unshelve eapply irrelevanceTyEq, SigCong; try irrValid. } - assert ([Γ ||-v< one > tSig A B ≅ tSig A'' B'' | VΓ | VΣ]). - { unshelve eapply irrelevanceTyEq, SigCong; try irrValid. } - assert [Γ ||-v< one > tPair A' B' a b : tSig A B | _ | VΣ ]. - { eapply conv; [|unshelve eapply pairValid]; try irrValid. - - unshelve eapply symValidTyEq; irrValid. - - eapply conv; irrValid. Unshelve. all: irrValid. } - assert [Γ ||-v< one > tPair A'' B'' a' b' : tSig A B | _ | VΣ ]. - { eapply conv; [|unshelve eapply pairValid]; try irrValid. - - unshelve eapply symValidTyEq; irrValid. - - eapply conv; irrValid. - Unshelve. all: irrValid. } - assert [Γ ||-v< one > b : B'[a..] | VΓ | VB'a]. - { eapply conv; irrValid. Unshelve. all: irrValid. } - assert [Γ ||-v< one > b' : B''[a'..] | VΓ | VB''a']. - { eapply conv; irrValid. Unshelve. all: irrValid. } - assert ([Γ ||-v< one > tFst (tPair A' B' a b) ≅ a : A | VΓ | VA]). - { eapply (convEq (A := A')); [eapply symValidTyEq; irrValid|]. - eapply pairFstValid. Unshelve. all: try irrValid. } - assert ([Γ ||-v< one > tFst (tPair A'' B'' a' b') ≅ a' : A | VΓ | VA]). - { eapply (convEq (A := A'')); [eapply symValidTyEq; irrValid|]. - eapply pairFstValid. Unshelve. all: irrValid. } - assert ([Γ ||-v< one > tFst (tPair A' B' a b) ≅ tFst (tPair A'' B'' a' b') : A | VΓ | VA]). - { eapply transValidTmEq; [irrValid|]. - eapply transValidTmEq; [irrValid|]. - eapply symValidTmEq; irrValid. } - assert ([Γ ||-v< one > tFst (tPair A' B' a b) : A | VΓ | VA]). - { eapply fstValid. Unshelve. all: try irrValid. } - assert ([Γ ||-v< one > tFst (tPair A'' B'' a' b') : A | VΓ | VA]). - { eapply fstValid. Unshelve. all: try irrValid. } - assert (VBfst : [Γ ||-v< one > B[(tFst (tPair A' B' a b))..] | VΓ]). - { eapply (substS (F := A)). Unshelve. all: try irrValid. } - assert (VB'fst : [Γ ||-v< one > B'[(tFst (tPair A' B' a b))..] | VΓ]). - { eapply (substS (F := A)). Unshelve. all: try irrValid. } - assert (VB''fst' : [Γ ||-v< one > B''[(tFst (tPair A'' B'' a' b'))..] | VΓ]). - { eapply (substS (F := A)). Unshelve. all: try irrValid. } - assert ([Γ ||-v< one > B[(tFst (tPair A' B' a b))..] ≅ B[a..] | VΓ | VBfst]). - { eapply irrelevanceTyEq, substSEq; [..|irrValid|irrValid]. - Unshelve. all: try irrValid. eapply reflValidTy. } - assert ([Γ ||-v< one > B'[(tFst (tPair A' B' a b))..] ≅ B[a..] | VΓ | VB'fst]). - { eapply irrelevanceTyEq, substSEq; [..|irrValid|irrValid]. - Unshelve. all: try irrValid. - eapply symValidTyEq; irrValid. Unshelve. all: try irrValid. } - assert ([Γ ||-v< one > B''[(tFst (tPair A'' B'' a' b'))..] ≅ B[a'..] | VΓ | VB''fst']). - { eapply irrelevanceTyEq, substSEq; [..|irrValid|]. - Unshelve. all: try irrValid. - eapply symValidTyEq; irrValid. Unshelve. all: try irrValid. } - assert ([Γ ||-v< one > B''[(tFst (tPair A'' B'' a' b'))..] ≅ B[a..] | VΓ | VB''fst']). - { eapply irrelevanceTyEq, substSEq; [..|irrValid|]. - Unshelve. all: try irrValid. - eapply symValidTyEq; irrValid. Unshelve. all: try irrValid. - eapply transValidTmEq; [irrValid|]. - eapply symValidTmEq; irrValid. - } - assert ([Γ ||-v< one > tSnd (tPair A' B' a b) ≅ b : B[(tFst (tPair A' B' a b))..] | VΓ | VBfst]). - { eapply (convEq (A := B'[(tFst (tPair A' B' a b))..])). - shelve. eapply pairSndValid. - Unshelve. all: try irrValid. - eapply irrelevanceTyEq, substSEq; [..|irrValid|apply reflValidTm; irrValid]. + intros * [VΓ VA] [?? VA'] [] [] [] [] []. + assert [Γ ,, A' ||-v B' | validSnoc _ VA']. + 1:{ eapply convCtx1; tea. } + econstructor. + eapply conv; cycle 1. + - eapply pairCongValid. + + eapply transValidTyEq; tea; eapply symValidTyEq; irrValid. + + eapply convEqCtx1; tea; eapply transValidTyEq; tea. + eapply symValidTyEq; irrValid. + + eapply conv; [eapply substSEq; tea|]; irrValid. + - eapply symValidTyEq; eapply SigCong; try irrValid. + Unshelve. all: try irrValid. - - unshelve apply reflValidTy. - - unshelve eapply symValidTyEq; irrValid. } - assert ([Γ ||-v< one > tSnd (tPair A'' B'' a' b') ≅ b' : B[a..] | VΓ | VBa]). - { eapply (convEq (A := B''[(tFst (tPair A'' B'' a' b'))..])). - shelve. eapply pairSndValid. - Unshelve. all: try irrValid. } - unshelve econstructor. 1-4: irrValid. - eapply irrelevanceTmEq, sigEtaValid; try irrValid. - eapply transValidTmEq; [irrValid|]. - eapply transValidTmEq. - - eapply convEq; [|irrValid]. - eapply symValidTyEq; try irrValid. - - eapply symValidTmEq; try irrValid. - eapply convEq; [|irrValid]. - eapply symValidTyEq; try irrValid. - Unshelve. all: try irrValid. + + unshelve eapply SigValid; irrValid. + + unshelve (eapply conv; tea; irrValid); irrValid. + + eapply lreflValidTm; irrValid. Qed. Lemma FundTmEqSigEta : forall (Γ : context) (A B p : term), FundTm Γ (tSig A B) p -> FundTmEq Γ (tSig A B) (tPair A B (tFst p) (tSnd p)) p. Proof. - intros * [VΓ VΣ Vp]. - assert (VA := domSigValid _ VΣ). - assert (VB := codSigValid _ VΣ). - assert (Vfst : [Γ ||-v< one > tFst p : A | _ | VA]). - { unshelve eapply irrelevanceTm, fstValid; try irrValid. } - assert (VBfst : [Γ ||-v< one > B[(tFst p)..] | VΓ]). - { unshelve eapply substS; try irrValid. } - assert (Vsnd : [Γ ||-v< one > tSnd p : B[(tFst p)..] | _ | VBfst]). - { unshelve eapply irrelevanceTm, sndValid. - 5: irrValid. all: irrValid. } - assert (Vηp : [Γ ||-v< one > tPair A B (tFst p) (tSnd p) : tSig A B | VΓ | VΣ]). - { unshelve eapply irrelevanceTm, pairValid; try irrValid. } - unshelve econstructor; try irrValid. - eapply irrelevanceTmEq, sigEtaValid; try irrValid. - - eapply transValidTmEq; [eapply pairFstValid|eapply reflValidTm]. - Unshelve. all: try irrValid. - - eapply transValidTmEq; [unshelve eapply irrelevanceTmEq, pairSndValid|unshelve eapply reflValidTm]; try irrValid. - unshelve eapply conv; try irrValid. - eapply irrelevanceTyEq, substSEq; try irrValid. - 1,2: unshelve apply reflValidTy. - { unshelve eapply fstValid; try irrValid. } - { unshelve eapply symValidTmEq, pairFstValid; try irrValid. } - Unshelve. all: try irrValid. + intros * [VΓ VΣ Vp]; econstructor. + eapply sigEtaValid; irrValid. + Unshelve. + + tea. + + now eapply domSigValid. + + now eapply codSigValid. Qed. Lemma FundTmEqFstCong : forall (Γ : context) (A B p p' : term), FundTmEq Γ (tSig A B) p p' -> FundTmEq Γ A (tFst p) (tFst p'). Proof. intros * []; unshelve econstructor. - 5: eapply fstCongValid; irrValid. + 3: eapply fstValid; irrValid. 2: now eapply domSigValid. - 1,2: eapply fstValid; irrValid. Unshelve. all: now eapply codSigValid. Qed. @@ -948,8 +624,8 @@ Section Fundamental. Proof. intros * [] [] [] []. unshelve econstructor. - 3,5: eapply pairFstValid; irrValid. - 2,3: irrValid. + 3: eapply pairFstValid; irrValid. + 2: irrValid. tea. Unshelve. all: irrValid. Qed. @@ -958,19 +634,12 @@ Section Fundamental. FundTmEq Γ (tSig A B) p p' -> FundTmEq Γ B[(tFst p)..] (tSnd p) (tSnd p'). Proof. intros * []; unshelve econstructor. - 5: eapply sndCongValid; irrValid. + 3: eapply sndValid; irrValid. 1: tea. - 1: eapply sndValid. - eapply conv; [| eapply sndValid]. - eapply substSEq. - 5: eapply fstCongValid. - 4: eapply fstValid. - 4-6: irrValid. - + eapply reflValidTy. - + eapply codSigValid. - + eapply reflValidTy. - + eapply symValidTmEq; irrValid. - Unshelve. all: try solve [now eapply domSigValid| now eapply codSigValid| irrValid]. + Unshelve. + 2: now eapply domSigValid. + 1: now eapply codSigValid. + irrValid. Qed. Lemma FundTmEqSndBeta : forall (Γ : context) (A B a b : term), @@ -981,15 +650,8 @@ Section Fundamental. FundTmEq Γ B[(tFst (tPair A B a b))..] (tSnd (tPair A B a b)) b. Proof. intros * [] [] [] []; unshelve econstructor. - 3,5: unshelve eapply pairSndValid; irrValid. - + tea. - + eapply conv; tea. - eapply irrelevanceTyEq. - eapply substSEq. - 1,3: eapply reflValidTy. - 1: irrValid. - 2: eapply symValidTmEq. - 1,2: eapply pairFstValid; irrValid. + 3: eapply pairSndValid. + 1: tea. Unshelve. all: irrValid. Qed. @@ -998,7 +660,7 @@ Section Fundamental. Proof. intros * [] [] []. unshelve econstructor; tea. - unshelve eapply IdValid; irrValid. + unshelve eapply IdValid; try irrValid. Qed. @@ -1028,6 +690,7 @@ Section Fundamental. Proof. intros * [] [] [] [] [] []; unshelve econstructor. 3: unshelve eapply IdElimValid; try irrValid. + 2,3: eapply reflValidTy. tea. Qed. @@ -1039,7 +702,7 @@ Section Fundamental. unshelve econstructor; tea. 3: eapply IdCongValid; irrValid. 1,2: eapply IdValid; try irrValid. - 1,2: eapply conv; irrValid. + 1,2: eapply ureflValidTm, conv; tea; try irrValid. Unshelve. all: irrValid. Qed. @@ -1048,10 +711,8 @@ Section Fundamental. FundTmEq Γ A x x' -> FundTmEq Γ A y y' -> FundTmEq Γ U (tId A x y) (tId A' x' y'). Proof. intros * [] [] []; unshelve econstructor. - 5: eapply IdTmCongValid; try irrValid; tea. - 2,3: eapply IdTmValid; try irrValid; tea. + 3: eapply IdTmCongValid; try irrValid; tea. 1: tea. - 1,2: eapply conv; try irrValid; eapply univEqValid; irrValid. Unshelve. all: first [eapply UValid| irrValid | tea]. Qed. @@ -1059,14 +720,9 @@ Section Fundamental. FundTyEq Γ A A' -> FundTmEq Γ A x x' -> FundTmEq Γ (tId A x x) (tRefl A x) (tRefl A' x'). Proof. intros * [] []; unshelve econstructor. - 5: eapply reflCongValid; tea; irrValid. - 3: eapply conv. - 2,4: eapply reflValid; try irrValid. - 2: eapply conv; irrValid. - 2: eapply symValidTyEq; eapply IdCongValid; tea; try irrValid. + 3: eapply reflCongValid; tea; irrValid. eapply IdValid; irrValid. - Unshelve. all: try eapply IdValid; try irrValid; eapply conv; irrValid. - Unshelve. all: irrValid. + Unshelve. tea. Qed. Lemma FundTmEqIdElimCong : forall (Γ : context) (A A' x x' P P' hr hr' y y' e e' : term), @@ -1080,34 +736,11 @@ Section Fundamental. FundTmEq Γ (tId A x y) e e' -> FundTmEq Γ P[e .: y..] (tIdElim A x P hr y e) (tIdElim A' x' P' hr' y' e'). Proof. intros * [] [] [] [] [] [] [] []. - assert [_ ||-v x' : _ | _ | VB] by (eapply conv; irrValid). - assert [_ ||-v y' : _ | _ | VB] by (eapply conv; irrValid). - assert (VId' : [_ ||-v tId A' x' y' | VΓ]) by (eapply IdValid; irrValid). - assert [_ ||-v _ ≅ tId A' x' y' | _ | VA6] by (eapply IdCongValid; irrValid). - assert [_ ||-v e' : _ | _ | VId'] by (eapply conv; irrValid). - unshelve econstructor. - 5: eapply IdElimCongValid; tea; irrValid. - 1: eapply IdElimValid; irrValid. - eapply convsym. - 2: eapply IdElimValid; eapply conv; [|irrValid]. - 1:{ - eapply substEqIdElimMotive. 7: tea. 2,3,5-12: tea; try irrValid. 1,2: irrValid. - } - eapply substEqIdElimMotive. 7: tea. 2,3,5-9: tea; try irrValid. - 1,2: irrValid. - 2: eapply convsym. - 1,3: eapply reflValid; first [irrValid| eapply conv; irrValid]. - 1:eapply IdCongValid; irrValid. - eapply reflCongValid; irrValid. - Unshelve. all: tea; try irrValid. - 3,4: eapply IdValid; irrValid. - 1: eapply validSnoc; now eapply idElimMotiveCtxIdValid. - eapply convCtx2'; tea. - 1: eapply convCtx1; tea; [eapply symValidTyEq; irrValid| ]. - 1,3: now eapply idElimMotiveCtxIdValid. - eapply idElimMotiveCtxIdCongValid; tea; irrValid. + econstructor. + eapply IdElimValid; tea; irrValid. Unshelve. - 3: eapply idElimMotiveCtxIdValid. all: irrValid. + 3: unshelve (eapply IdValid; irrValid); irrValid. + all: irrValid. Qed. Lemma FundTmEqIdElimRefl : forall (Γ : context) (A x P hr y A' z : term), @@ -1123,33 +756,23 @@ Section Fundamental. FundTmEq Γ A x z -> FundTmEq Γ P[tRefl A' z .: y..] (tIdElim A x P hr y (tRefl A' z)) hr. Proof. intros * [] [] [] [] [] [] [] [] [] []. - assert (VId : [_ ||-v tId A x y | VΓ]) by (unshelve eapply IdValid; irrValid). - assert [Γ ||-v< one > tRefl A' z : tId A x y | _ | VId]. - 1:{ - eapply convsym ; [| eapply reflValid; try irrValid]. - 2: eapply conv; irrValid. - eapply IdCongValid; try irrValid. - eapply transValidTmEq; [eapply symValidTmEq|]; irrValid. - Unshelve. - 1: unshelve eapply IdValid; try eapply conv. - all: try irrValid. - Unshelve. all: irrValid. - } - unshelve econstructor. - 5: eapply IdElimReflValid; tea; try irrValid. - 3: eapply conv; irrValid. - 2:{ - eapply conv; [|irrValid]. - eapply substExtIdElimMotive; cycle 1; tea. - + eapply reflValid; irrValid. - + irrValid. - + eapply reflCongValid; tea; irrValid. - } - eapply IdElimValid; irrValid. - Unshelve. all: tea; try irrValid. - unshelve eapply IdValid; irrValid. + econstructor. + eapply IdElimReflValid; tea; try irrValid. + Unshelve. all: try irrValid. + + unshelve (eapply IdValid; irrValid); irrValid. + + assert (VId : [Γ ||-v tId A x y | VΓ]) by (unshelve (eapply IdValid; irrValid); irrValid). + assert [_ ||-v y ≅ z : _ | _ | VA]. + 1: eapply transValidTmEq; [eapply symValidTmEq|]; irrValid. + assert [Γ ||-v tId A x y ≅ tId A' z z | _ | VId ]. + 1: eapply IdCongValid; irrValid. + eapply conv; [|eapply reflValid, ureflValidTm, conv; irrValid]. + eapply symValidTyEq; irrValid. + Unshelve. + 2: unshelve eapply ureflValidTy. + all: try irrValid. Qed. + Lemma Fundamental : (forall Γ : context, [ |-[ de ] Γ ] -> FundCon (ta := ta) Γ) × (forall (Γ : context) (A : term), [Γ |-[ de ] A] -> FundTy (ta := ta) Γ A) × (forall (Γ : context) (A t : term), [Γ |-[ de ] t : A] -> FundTm (ta := ta) Γ A t) @@ -1233,7 +856,7 @@ Lemma Fundamental : (forall Γ : context, [ |-[ de ] Γ ] -> FundCon (ta := ta) destruct IH ; tea. apply Fundamental in Hσ0 as [?? Hσ0]. cbn in *. - eapply reducibleTm in Hσ0. + eapply reducibleTmEq in Hσ0. eapply Fundamental in HA as []. unshelve econstructor. 1: now eapply validSnoc. @@ -1256,7 +879,7 @@ Lemma Fundamental : (forall Γ : context, [ |-[ de ] Γ ] -> FundCon (ta := ta) all: now econstructor. - inversion HΔ as [|?? HΔ' HA] ; subst ; clear HΔ ; refold. destruct IH ; tea. - apply Fundamental in Hσ0 as [?? Hσ0 Hτ0 Ηστ] ; cbn in *. + apply Fundamental in Hσ0 as [?? Ηστ] ; cbn in *. eapply Fundamental in HA as [? HA]. unshelve econstructor. + now eapply validSnoc. @@ -1264,24 +887,7 @@ Lemma Fundamental : (forall Γ : context, [ |-[ de ] Γ ] -> FundCon (ta := ta) 1: now irrValid. cbn. irrelevanceRefl. - now eapply reducibleTm. - + unshelve econstructor. - 1: now eapply irrelevanceSubst. - cbn. - unshelve irrelevanceRefl. - 2: now unshelve eapply HA ; tea ; irrValid. - cbn. - unshelve eapply LRTmRedConv. - 3: now eapply reducibleTy. - * irrelevanceRefl. - unshelve eapply HA ; tea. - all: now irrValid. - * irrelevanceRefl. - now eapply reducibleTm. - + unshelve econstructor ; cbn in *. - * now irrValid. - * irrelevanceRefl. - now eapply reducibleTmEq. + now eapply reducibleTmEq. Qed. End Fundamental. diff --git a/theories/GenericTyping.v b/theories/GenericTyping.v index c33bf32..1d94f46 100644 --- a/theories/GenericTyping.v +++ b/theories/GenericTyping.v @@ -11,7 +11,7 @@ times with different instances of this abstract notion of typing, gathering more and more properties. *) (** -More precisely, an instance consists of giving notions of +More precisely, an instance consists of giving notions of - context well-formation [|- Γ] - type well-formation [Γ |- A] - term well-formation [Γ |- t : A] @@ -48,7 +48,7 @@ Section RedDefinitions. }. Record TypeConvWf (Γ : context) (A B : term) : Type := - { + { tyc_wf_l : [Γ |- A] ; tyc_wf_r : [Γ |- B] ; tyc_wf_conv :> [Γ |- A ≅ B] @@ -151,7 +151,7 @@ Notation "[ Γ |- A ↘ B ]" := (TypeRedWhnf Γ A B) (only parsing) : typing_sco Notation "[ Γ |-[ ta ] A ↘ B ]" := (TypeRedWhnf (ta := ta) Γ A B) : typing_scope. Notation "[ Γ |- t ↘ u : A ]" := (TermRedWhnf Γ A t u) (only parsing ): typing_scope. Notation "[ Γ |-[ ta ] t ↘ u : A ]" := (TermRedWhnf (ta := ta) Γ A t u) : typing_scope. -Notation "[ Γ |- A :≅: B ]" := (TypeConvWf Γ A B) (only parsing) : typing_scope. +Notation "[ Γ |- A :≅: B ]" := (TypeConvWf Γ A B) (only parsing) : typing_scope. Notation "[ Γ |-[ ta ] A :≅: B ]" := (TypeConvWf (ta := ta) Γ A B) : typing_scope. Notation "[ Γ |- t :≅: u : A ]" := (TermConvWf Γ A t u) (only parsing) : typing_scope. Notation "[ Γ |-[ ta ] t :≅: u : A ]" := (TermConvWf (ta := ta) Γ A t u) : typing_scope. @@ -209,16 +209,16 @@ Section GenericTyping. { wft_wk {Γ Δ A} (ρ : Δ ≤ Γ) : [|- Δ ] -> [Γ |- A] -> [Δ |- A⟨ρ⟩] ; - wft_U {Γ} : + wft_U {Γ} : [ |- Γ ] -> [ Γ |- U ] ; - wft_prod {Γ} {A B} : - [ Γ |- A ] -> - [Γ ,, A |- B ] -> + wft_prod {Γ} {A B} : + [ Γ |- A ] -> + [Γ ,, A |- B ] -> [ Γ |- tProd A B ] ; - wft_sig {Γ} {A B} : - [ Γ |- A ] -> - [Γ ,, A |- B ] -> + wft_sig {Γ} {A B} : + [ Γ |- A ] -> + [Γ ,, A |- B ] -> [ Γ |- tSig A B ] ; wft_Id {Γ} {A x y} : [Γ |- A] -> @@ -226,7 +226,7 @@ Section GenericTyping. [Γ |- y : A] -> [Γ |- tId A x y] ; wft_term {Γ} {A} : - [ Γ |- A : U ] -> + [ Γ |- A : U ] -> [ Γ |- A ] ; }. @@ -239,16 +239,16 @@ Section GenericTyping. in_ctx Γ n decl -> [ Γ |- tRel n : decl ] ; ty_prod {Γ} {A B} : - [ Γ |- A : U] -> + [ Γ |- A : U] -> [Γ ,, A |- B : U ] -> [ Γ |- tProd A B : U ] ; ty_lam {Γ} {A B t} : [ Γ |- A ] -> - [ Γ ,, A |- t : B ] -> + [ Γ ,, A |- t : B ] -> [ Γ |- tLambda A t : tProd A B] ; ty_app {Γ} {f a A B} : - [ Γ |- f : tProd A B ] -> - [ Γ |- a : A ] -> + [ Γ |- f : tProd A B ] -> + [ Γ |- a : A ] -> [ Γ |- tApp f a : B[a ..] ] ; ty_nat {Γ} : [|-Γ] -> @@ -273,11 +273,11 @@ Section GenericTyping. [Γ |- e : tEmpty] -> [Γ |- tEmptyElim P e : P[e..]] ; ty_sig {Γ} {A B} : - [ Γ |- A : U] -> + [ Γ |- A : U] -> [Γ ,, A |- B : U ] -> [ Γ |- tSig A B : U ] ; ty_pair {Γ} {A B a b} : - [ Γ |- A ] -> + [ Γ |- A ] -> [Γ ,, A |- B ] -> [Γ |- a : A] -> [Γ |- b : B[a..]] -> @@ -528,7 +528,7 @@ Section GenericTyping. [Γ |- y : A] -> [Γ |- e ⤳* e' : tId A x y] -> [Γ |- tIdElim A x P hr y e ⤳* tIdElim A x P hr y e' : P[e .: y..]]; - redtm_conv {Γ t u A A'} : + redtm_conv {Γ t u A A'} : [Γ |- t ⤳* u : A] -> [Γ |- A ≅ A'] -> [Γ |- t ⤳* u : A'] ; @@ -662,8 +662,8 @@ Ltac renToWk := fold ren_term; repeat change (ren_term ?x ?y) with y⟨x⟩; repeat change S with ↑; - repeat lazymatch goal with - | [ _ : _ |- ?G] => renToWk0 G + repeat lazymatch goal with + | [ _ : _ |- ?G] => renToWk0 G end. @@ -723,7 +723,7 @@ Section GenericConsequences. A' = A -> [Γ |- t :⤳*: u : A']. Proof. - now intros ? ->. + now intros ? ->. Qed. (** *** Properties of well-typed reduction *) @@ -732,7 +732,7 @@ Section GenericConsequences. Proof. intros []; now eapply redty_ty_src. Qed. - + Lemma tmr_wf_l {Γ t u A} : [Γ |- t :⤳*: u : A] -> [Γ |- t : A]. Proof. intros []; now eapply redtm_ty_src. @@ -746,11 +746,11 @@ Section GenericConsequences. Lemma redty_red {Γ A B} : [Γ |- A ⤳* B] -> [ A ⤳* B ]. Proof. - intros ?%redty_sound. + intros ?%redty_sound. assumption. Qed. - Lemma redtm_red {Γ t u A} : + Lemma redtm_red {Γ t u A} : [Γ |- t ⤳* u : A] -> [t ⤳* u]. Proof. @@ -770,7 +770,7 @@ Section GenericConsequences. Proof. intros []; now eapply redty_red. Qed. - + Lemma redtywf_term {Γ A B} : [ Γ |- A :⤳*: B : U] -> [Γ |- A :⤳*: B ]. Proof. @@ -786,10 +786,10 @@ Section GenericConsequences. intros ??? [] []; unshelve econstructor; try etransitivity; tea. Qed. - (** Almost all of the RedTermProperties can be derived + (** Almost all of the RedTermProperties can be derived for the well-formed reduction [Γ |- t :⤳*: u : A] but for application (which requires stability of typing under substitution). *) - + Definition redtmwf_wk {Γ Δ t u A} (ρ : Δ ≤ Γ) : [|- Δ ] -> [Γ |- t :⤳*: u : A] -> [Δ |- t⟨ρ⟩ :⤳*: u⟨ρ⟩ : A⟨ρ⟩]. Proof. intros ? []; constructor; gen_typing. Qed. @@ -827,7 +827,7 @@ Section GenericConsequences. Proof. intros [] ?; constructor; gen_typing. Qed. - + Lemma redtmwf_appwk {Γ Δ A B B' t u a} (ρ : Δ ≤ Γ) : [Γ |- t :⤳*: u : tProd A B] -> [Δ |- a : A⟨ρ⟩] -> @@ -874,10 +874,10 @@ Section GenericConsequences. (** *** Derived typing, reduction and conversion judgements *) - Lemma ty_var0 {Γ A} : + Lemma ty_var0 {Γ A} : [Γ |- A] -> [Γ ,, A |- tRel 0 : A⟨↑⟩]. - Proof. + Proof. intros; refine (ty_var _ (in_here _ _)); gen_typing. Qed. @@ -923,8 +923,8 @@ Section GenericConsequences. #[local] Hint Resolve ty_simple_app : gen_typing. - - Lemma ty_id {Γ A B C} : + + Lemma ty_id {Γ A B C} : [Γ |- A] -> [Γ |- A ≅ B] -> [Γ |- A ≅ C] -> @@ -937,7 +937,7 @@ Section GenericConsequences. now eapply ty_var0. Qed. - Lemma ty_id' {Γ A} : + Lemma ty_id' {Γ A} : [Γ |- A] -> [Γ |- idterm A : arr A A]. Proof. @@ -947,7 +947,7 @@ Section GenericConsequences. eapply ty_lam; tea. now eapply ty_var0. Qed. - + Lemma redtm_id_beta {Γ a A} : [Γ |- A] -> [Γ |- A ≅ A] -> @@ -962,7 +962,7 @@ Section GenericConsequences. + now asimpl. Qed. - Lemma convtm_id {Γ A A' B C} : + Lemma convtm_id {Γ A A' B C} : [|- Γ] -> [Γ |- A] -> [Γ |- A'] -> @@ -985,14 +985,14 @@ Section GenericConsequences. now apply wfc_cons. } 1,2: eapply ty_id; tea; now symmetry. assert [|- Γ,, A] by gen_typing. - assert [Γ,, A |-[ ta ] A⟨@wk1 Γ A⟩] by now eapply wft_wk. + assert [Γ,, A |-[ ta ] A⟨@wk1 Γ A⟩] by now eapply wft_wk. eapply convtm_exp. - cbn. eapply redtm_id_beta. 3: now eapply ty_var0. 1,2: renToWk; tea; now eapply convty_wk. - - cbn. + - cbn. assert [Γ,, A |- A'⟨↑⟩ ≅ A⟨↑⟩] - by (renToWk; symmetry; now eapply convty_wk). + by (renToWk; symmetry; now eapply convty_wk). eapply redtm_conv; tea. eapply redtm_id_beta. 1: renToWk; now eapply wft_wk. @@ -1014,28 +1014,28 @@ Section GenericConsequences. [Γ |- f : arr B C] -> [Γ |- comp A f g : arr A C]. Proof. - intros tyA tyB **. + intros tyA tyB **. eapply ty_lam; tea. assert [|- Γ,, A] by gen_typing. pose (r := @wk1 Γ A). eapply ty_simple_app; renToWk. - - unshelve eapply (wft_wk _ _ tyB) ; tea. + - unshelve eapply (wft_wk _ _ tyB) ; tea. - now eapply wft_wk. - replace (arr _ _) with (arr B C)⟨r⟩ by (unfold r; now bsimpl). now eapply ty_wk. - eapply ty_simple_app; renToWk. - + unshelve eapply (wft_wk _ _ tyA) ; tea. + + unshelve eapply (wft_wk _ _ tyA) ; tea. + now eapply wft_wk. + replace (arr _ _) with (arr A B)⟨r⟩ by (unfold r; now bsimpl). now eapply ty_wk. + unfold r; rewrite wk1_ren_on; now refine (ty_var _ (in_here _ _)). Qed. - + Lemma wft_wk1 {Γ A B} : [Γ |- A] -> [Γ |- B] -> [Γ ,, A |- B⟨↑⟩]. Proof. intros; renToWk; eapply wft_wk; gen_typing. Qed. - + Lemma redtm_comp_beta {Γ A B C f g a} : [Γ |- A] -> [Γ |- B] -> @@ -1218,8 +1218,8 @@ Section GenericConsequences. rewrite scons_eta'. now bsimpl. Qed. - - + + (** *** Lifting determinism properties from untyped reduction to typed reduction. *) Lemma redtm_whnf {Γ t u A} : [Γ |- t ⤳* u : A] -> whnf t -> t = u. diff --git a/theories/LogicalRelation.v b/theories/LogicalRelation.v index f7a0447..6f01364 100644 --- a/theories/LogicalRelation.v +++ b/theories/LogicalRelation.v @@ -29,7 +29,6 @@ Ltac logrel := eauto with logrel. context -> term -> (term -> Type@{i}) -> - (term -> Type@{i}) -> (term -> term -> Type@{i}) -> Type@{j}. @@ -40,7 +39,6 @@ Module LRPack. Record LRPack@{i} {Γ : context} {A : term} := { eqTy : term -> Type@{i}; - redTm : term -> Type@{i}; eqTm : term -> term -> Type@{i}; }. @@ -51,13 +49,13 @@ End LRPack. Export LRPack(LRPack,Build_LRPack). Notation "[ P | Γ ||- A ≅ B ]" := (@LRPack.eqTy Γ A P B). -Notation "[ P | Γ ||- t : A ]" := (@LRPack.redTm Γ A P t). +Notation "[ P | Γ ||- t : A ]" := (@LRPack.eqTm Γ A P t t). Notation "[ P | Γ ||- t ≅ u : A ]" := (@LRPack.eqTm Γ A P t u). (** An LRPack is adequate wrt. a RedRel when its three unpacked components are. *) Definition LRPackAdequate@{i j} {Γ : context} {A : term} (R : RedRel@{i j}) (P : LRPack@{i} Γ A) : Type@{j} := - R Γ A P.(LRPack.eqTy) P.(LRPack.redTm) P.(LRPack.eqTm). + R Γ A P.(LRPack.eqTy) P.(LRPack.eqTm). Arguments LRPackAdequate _ _ /. @@ -84,7 +82,6 @@ Coercion LRAd.adequate : LRAdequate >-> LRPackAdequate. Notation "[ R | Γ ||- A ]" := (@LRAdequate Γ A R). Notation "[ R | Γ ||- A ≅ B | RA ]" := (RA.(@LRAd.pack Γ A R).(LRPack.eqTy) B). -Notation "[ R | Γ ||- t : A | RA ]" := (RA.(@LRAd.pack Γ A R).(LRPack.redTm) t). Notation "[ R | Γ ||- t ≅ u : A | RA ]" := (RA.(@LRAd.pack Γ A R).(LRPack.eqTm) t u). (** ** Universe levels *) @@ -147,26 +144,6 @@ End neRedTyEq. Export neRedTyEq(neRedTyEq,Build_neRedTyEq). Notation "[ Γ ||-ne A ≅ B | R ]" := (neRedTyEq Γ A B R). -Module neRedTm. - - Record neRedTm `{ta : tag} - `{WfType ta} `{RedType ta} - `{Typing ta} `{ConvNeuConv ta} `{ConvType ta} `{RedTerm ta} - {Γ : context} {t A : term} {R : [ Γ ||-ne A ]} - : Set := { - te : term; - red : [ Γ |- t :⤳*: te : R.(neRedTy.ty)]; - eq : [Γ |- te ~ te : R.(neRedTy.ty)] ; - }. - - Arguments neRedTm {_ _ _ _ _ _ _}. - -End neRedTm. - -Export neRedTm(neRedTm, Build_neRedTm). - -Notation "[ Γ ||-ne t : A | B ]" := (neRedTm Γ t A B). - Module neRedTmEq. Record neRedTmEq `{ta : tag} @@ -236,8 +213,6 @@ Module URedTm. te : term; red : [ Γ |- t :⤳*: te : U ]; type : isType te; - eqr : [Γ |- te ≅ te : U]; - rel : [rec R.(URedTy.lt) | Γ ||- t ] ; }. Arguments URedTm {_ _ _ _ _ _ _ _} rec. @@ -260,7 +235,6 @@ Module URedTm. End URedTm. Export URedTm(URedTm,Build_URedTm,URedTmEq,Build_URedTmEq). -Notation "[ R | Γ ||-U t : A | l ]" := (URedTm R Γ t A l) (at level 0, R, Γ, t, A, l at level 50). Notation "[ R | Γ ||-U t ≅ u : A | l ]" := (URedTmEq R Γ t u A l) (at level 0, R, Γ, t, u, A, l at level 50). (** ** Reducibility of a polynomial A,, B *) @@ -278,17 +252,15 @@ Module PolyRedPack. shpTy : [Γ |- shp]; posTy : [Γ ,, shp |- pos]; shpRed {Δ} (ρ : Δ ≤ Γ) : [ |- Δ ] -> LRPack@{i} Δ shp⟨ρ⟩ ; - posRed {Δ} {a} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) : - [ (shpRed ρ h) | Δ ||- a : shp⟨ρ⟩] -> + posRed {Δ} {a b} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) : + [ (shpRed ρ h) | Δ ||- a ≅ b : shp⟨ρ⟩] -> LRPack@{i} Δ (pos[a .: (ρ >> tRel)]) ; posExt {Δ a b} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) - (ha : [ (shpRed ρ h) | Δ ||- a : shp⟨ρ⟩ ]) : - [ (shpRed ρ h) | Δ ||- b : shp⟨ρ⟩] -> - [ (shpRed ρ h) | Δ ||- a ≅ b : shp⟨ρ⟩] -> - [ (posRed ρ h ha) | Δ ||- (pos[a .: (ρ >> tRel)]) ≅ (pos[b .: (ρ >> tRel)]) ] + (hab : [ (shpRed ρ h) | Δ ||- a ≅ b : shp⟨ρ⟩]) : + [ (posRed ρ h hab) | Δ ||- (pos[a .: (ρ >> tRel)]) ≅ (pos[b .: (ρ >> tRel)]) ] }. Arguments PolyRedPack {_ _ _ _}. @@ -303,7 +275,7 @@ Module PolyRedPack. {Γ : context} {R : RedRel@{i j}} {PA : PolyRedPack@{i} Γ shp pos} : Type@{j} := { shpAd {Δ} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) : LRPackAdequate@{i j} R (PA.(shpRed) ρ h); - posAd {Δ a} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) (ha : [ PA.(shpRed) ρ h | Δ ||- a : shp⟨ρ⟩ ]) + posAd {Δ a b} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) (ha : [ PA.(shpRed) ρ h | Δ ||- a ≅ b : shp⟨ρ⟩ ]) : LRPackAdequate@{i j} R (PA.(posRed) ρ h ha); }. @@ -319,10 +291,10 @@ Module PolyRedEq. `{WfType ta} `{ConvType ta} {Γ : context} {shp pos: term} {PA : PolyRedPack Γ shp pos} {shp' pos' : term} : Type := { - shpRed {Δ} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) : + shpRed [Δ] (ρ : Δ ≤ Γ) (h : [ |- Δ ]) : [ PA.(PolyRedPack.shpRed) ρ h | Δ ||- shp⟨ρ⟩ ≅ shp'⟨ρ⟩ ]; - posRed {Δ a} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) - (ha : [ PA.(PolyRedPack.shpRed) ρ h | Δ ||- a : shp⟨ρ⟩]) : + posRed [Δ a b] (ρ : Δ ≤ Γ) (h : [ |- Δ ]) + (ha : [ PA.(PolyRedPack.shpRed) ρ h | Δ ||- a ≅ b : shp⟨ρ⟩]) : [ PA.(PolyRedPack.posRed) ρ h ha | Δ ||- pos[a .: (ρ >> tRel)] ≅ pos'[a .: (ρ >> tRel)] ]; }. @@ -406,61 +378,60 @@ Inductive isLRFun `{ta : tag} `{WfContext ta} | LamLRFun : forall A' t : term, (forall {Δ} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) (domRed:= ΠA.(PolyRedPack.shpRed) ρ h), [domRed | Δ ||- (PiRedTy.dom ΠA)⟨ρ⟩ ≅ A'⟨ρ⟩]) -> - (forall {Δ a} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) - (ha : [ ΠA.(PolyRedPack.shpRed) ρ h | Δ ||- a : ΠA.(PiRedTy.dom)⟨ρ⟩ ]), - [ΠA.(PolyRedPack.posRed) ρ h ha | Δ ||- t[a .: (ρ >> tRel)] : ΠA.(PiRedTy.cod)[a .: (ρ >> tRel)]]) -> + (forall {Δ a b} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) + (ha : [ ΠA.(PolyRedPack.shpRed) ρ h | Δ ||- a ≅ b : ΠA.(PiRedTy.dom)⟨ρ⟩ ]), + [ΠA.(PolyRedPack.posRed) ρ h ha | Δ ||- t[a .: (ρ >> tRel)] ≅ t[b .: (ρ >> tRel)] : ΠA.(PiRedTy.cod)[a .: (ρ >> tRel)]]) -> isLRFun ΠA (tLambda A' t) | NeLRFun : forall f : term, [Γ |- f ~ f : tProd (PiRedTy.dom ΠA) (PiRedTy.cod ΠA)] -> isLRFun ΠA f. -Module PiRedTm. +Module PiRedTmEq. + + Import PiRedTy. + + Definition appRed `{ta : tag} `{WfContext ta} `{WfType ta} `{ConvType ta} `{RedType ta} + {Γ A} (ΠA : PiRedTy Γ A) (nfL nfR : term) Δ a b := + forall (ρ : Δ ≤ Γ) (h : [ |- Δ ]) + (hab : [ΠA.(PolyRedPack.shpRed) ρ h | Δ ||- a ≅ b : ΠA.(PiRedTy.dom)⟨ρ⟩ ] ), + [ ΠA.(PolyRedPack.posRed) ρ h hab | Δ ||- tApp nfL⟨ρ⟩ a ≅ tApp nfR⟨ρ⟩ b : _ ]. + + Arguments appRed /. Record PiRedTm `{ta : tag} `{WfContext ta} `{WfType ta} `{ConvType ta} `{RedType ta} `{Typing ta} `{ConvTerm ta} `{ConvNeuConv ta} `{RedTerm ta} - {Γ : context} {t A : term} {ΠA : PiRedTy Γ A} + {Γ : context} {A} {ΠA : PiRedTy Γ A} {t : term} : Type := { nf : term; - red : [ Γ |- t :⤳*: nf : tProd ΠA.(PiRedTy.dom) ΠA.(PiRedTy.cod) ]; + red : [ Γ |- t :⤳*: nf : tProd ΠA.(dom) ΠA.(cod) ]; isfun : isLRFun ΠA nf; - refl : [ Γ |- nf ≅ nf : tProd ΠA.(PiRedTy.dom) ΠA.(PiRedTy.cod) ]; - app {Δ a} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) - (ha : [ ΠA.(PolyRedPack.shpRed) ρ h | Δ ||- a : ΠA.(PiRedTy.dom)⟨ρ⟩ ]) - : [ΠA.(PolyRedPack.posRed) ρ h ha | Δ ||- tApp nf⟨ρ⟩ a : ΠA.(PiRedTy.cod)[a .: (ρ >> tRel)]] ; - eq {Δ a b} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) (domRed:= ΠA.(PolyRedPack.shpRed) ρ h) - (ha : [ domRed | Δ ||- a : ΠA.(PiRedTy.dom)⟨ρ⟩ ]) - (hb : [ domRed | Δ ||- b : ΠA.(PiRedTy.dom)⟨ρ⟩ ]) - (eq : [ domRed | Δ ||- a ≅ b : ΠA.(PiRedTy.dom)⟨ρ⟩ ]) - : [ ΠA.(PolyRedPack.posRed) ρ h ha | Δ ||- tApp nf⟨ρ⟩ a ≅ tApp nf⟨ρ⟩ b : ΠA.(PiRedTy.cod)[a .: (ρ >> tRel)] ] + (* app {Δ a b} : appRed ΠA nf nf Δ a b; *) }. - Arguments PiRedTm {_ _ _ _ _ _ _ _ _}. - -End PiRedTm. - -Export PiRedTm(PiRedTm,Build_PiRedTm). -Notation "[ Γ ||-Π t : A | ΠA ]" := (PiRedTm Γ t A ΠA). + Arguments PiRedTm {_ _ _ _ _ _ _ _ _ _ _}. -Module PiRedTmEq. + Lemma whnf `{GenericTypingProperties} {Γ A} {ΠA : PiRedTy Γ A} {t} : forall (red : PiRedTm ΠA t), + whnf (nf red). + Proof. + intros [? ? isfun]; simpl; destruct isfun; constructor; tea. + now eapply convneu_whne. + Qed. Record PiRedTmEq `{ta : tag} `{WfContext ta} `{WfType ta} `{ConvType ta} `{RedType ta} `{Typing ta} `{ConvTerm ta} `{ConvNeuConv ta} `{RedTerm ta} {Γ : context} {t u A : term} {ΠA : PiRedTy Γ A} : Type := { - redL : [ Γ ||-Π t : A | ΠA ] ; - redR : [ Γ ||-Π u : A | ΠA ] ; - eq : [ Γ |- redL.(PiRedTm.nf) ≅ redR.(PiRedTm.nf) : tProd ΠA.(PiRedTy.dom) ΠA.(PiRedTy.cod) ]; - eqApp {Δ a} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) - (ha : [ΠA.(PolyRedPack.shpRed) ρ h | Δ ||- a : ΠA.(PiRedTy.dom)⟨ρ⟩ ] ) - : [ ΠA.(PolyRedPack.posRed) ρ h ha | Δ ||- - tApp redL.(PiRedTm.nf)⟨ρ⟩ a ≅ tApp redR.(PiRedTm.nf)⟨ρ⟩ a : ΠA.(PiRedTy.cod)[a .: (ρ >> tRel)]] + redL : PiRedTm ΠA t ; + redR : PiRedTm ΠA u ; + eq : [ Γ |- redL.(nf) ≅ redR.(nf) : tProd ΠA.(dom) ΠA.(cod) ]; + eqApp {Δ a b} : appRed ΠA redL.(nf) redR.(nf) Δ a b; }. Arguments PiRedTmEq {_ _ _ _ _ _ _ _ _}. End PiRedTmEq. -Export PiRedTmEq(PiRedTmEq,Build_PiRedTmEq). +Export PiRedTmEq(PiRedTm,Build_PiRedTm,PiRedTmEq,Build_PiRedTmEq). Notation "[ Γ ||-Π t ≅ u : A | ΠA ]" := (PiRedTmEq Γ t u A ΠA). @@ -486,19 +457,19 @@ Inductive isLRPair `{ta : tag} `{WfContext ta} | PairLRpair : forall (A' B' a b : term) (rdom : forall {Δ} (ρ : Δ ≤ Γ) (h : [ |- Δ ]), [ΣA.(PolyRedPack.shpRed) ρ h | Δ ||- (SigRedTy.dom ΣA)⟨ρ⟩ ≅ A'⟨ρ⟩]) - (rcod : forall {Δ a} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) - (ha : [ ΣA.(PolyRedPack.shpRed) ρ h | Δ ||- a : ΣA.(PiRedTy.dom)⟨ρ⟩ ]), - [ΣA.(PolyRedPack.posRed) ρ h ha | Δ ||- (SigRedTy.cod ΣA)[a .: (ρ >> tRel)] ≅ B'[a .: (ρ >> tRel)]]) + (rcod : forall {Δ a b} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) + (ha : [ ΣA.(PolyRedPack.shpRed) ρ h | Δ ||- a ≅ b: ΣA.(PiRedTy.dom)⟨ρ⟩ ]), + [ΣA.(PolyRedPack.posRed) ρ h ha | Δ ||- (SigRedTy.cod ΣA)[a .: (ρ >> tRel)] ≅ B'[b .: (ρ >> tRel)]]) (rfst : forall {Δ} (ρ : Δ ≤ Γ) (h : [ |- Δ ]), - [ΣA.(PolyRedPack.shpRed) ρ h | Δ ||- a⟨ρ⟩ : (SigRedTy.dom ΣA)⟨ρ⟩]) + [ΣA.(PolyRedPack.shpRed) ρ h | Δ ||- a⟨ρ⟩ ≅ a⟨ρ⟩ : (SigRedTy.dom ΣA)⟨ρ⟩]) (rsnd : forall {Δ} (ρ : Δ ≤ Γ) (h : [ |- Δ ]), - [ΣA.(PolyRedPack.posRed) ρ h (rfst ρ h) | Δ ||- b⟨ρ⟩ : (SigRedTy.cod ΣA)[a⟨ρ⟩ .: (ρ >> tRel)] ]), + [ΣA.(PolyRedPack.posRed) ρ h (rfst ρ h) | Δ ||- b⟨ρ⟩ ≅ b⟨ρ⟩ : (SigRedTy.cod ΣA)[a⟨ρ⟩ .: (ρ >> tRel)] ]), isLRPair ΣA (tPair A' B' a b) | NeLRPair : forall p : term, [Γ |- p ~ p : tSig (SigRedTy.dom ΣA) (SigRedTy.cod ΣA)] -> isLRPair ΣA p. -Module SigRedTm. +Module SigRedTmEq. Record SigRedTm `{ta : tag} `{WfContext ta} `{WfType ta} `{ConvType ta} `{RedType ta} @@ -508,41 +479,36 @@ Module SigRedTm. nf : term; red : [ Γ |- t :⤳*: nf : ΣA.(outTy) ]; ispair : isLRPair ΣA nf; - refl : [ Γ |- nf ≅ nf : ΣA.(outTy) ]; - fstRed {Δ} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) : - [ΣA.(PolyRedPack.shpRed) ρ h | Δ ||- tFst nf⟨ρ⟩ : ΣA.(ParamRedTyPack.dom)⟨ρ⟩] ; - sndRed {Δ} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) : - [ΣA.(PolyRedPack.posRed) ρ h (fstRed ρ h) | Δ ||- tSnd nf⟨ρ⟩ : _] ; }. Arguments SigRedTm {_ _ _ _ _ _ _ _ _ _ _}. -End SigRedTm. - -Export SigRedTm(SigRedTm,Build_SigRedTm). -Notation "[ Γ ||-Σ t : A | ΣA ]" := (SigRedTm (Γ:=Γ) (A:=A) ΣA t) (at level 0, Γ, t, A, ΣA at level 50). - -Module SigRedTmEq. + Lemma whnf `{GenericTypingProperties} {Γ A} {ΣA : SigRedTy Γ A} {t} : + forall (red : SigRedTm ΣA t), whnf (nf red). + Proof. + intros [? ? ispair]; simpl; destruct ispair; constructor; tea. + now eapply convneu_whne. + Qed. Record SigRedTmEq `{ta : tag} `{WfContext ta} `{WfType ta} `{ConvType ta} `{RedType ta} `{Typing ta} `{ConvTerm ta} `{ConvNeuConv ta} `{RedTerm ta} {Γ : context} {A : term} {ΣA : SigRedTy Γ A} {t u : term} : Type := { - redL : [ Γ ||-Σ t : A | ΣA ] ; - redR : [ Γ ||-Σ u : A | ΣA ] ; - eq : [ Γ |- redL.(SigRedTm.nf) ≅ redR.(SigRedTm.nf) : ΣA.(outTy) ]; - eqFst {Δ} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) : - [ΣA.(PolyRedPack.shpRed) ρ h | Δ ||- tFst redL.(SigRedTm.nf)⟨ρ⟩ ≅ tFst redR.(SigRedTm.nf)⟨ρ⟩ : ΣA.(ParamRedTyPack.dom)⟨ρ⟩] ; - eqSnd {Δ} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) (redfstL := redL.(SigRedTm.fstRed) ρ h) : - [ΣA.(PolyRedPack.posRed) ρ h redfstL | Δ ||- tSnd redL.(SigRedTm.nf)⟨ρ⟩ ≅ tSnd redR.(SigRedTm.nf)⟨ρ⟩ : _] ; + redL : SigRedTm ΣA t ; + redR : SigRedTm ΣA u ; + eq : [ Γ |- redL.(nf) ≅ redR.(nf) : ΣA.(outTy) ]; + eqFst [Δ] (ρ : Δ ≤ Γ) (h : [ |- Δ ]) : + [ΣA.(PolyRedPack.shpRed) ρ h | Δ ||- tFst redL.(nf)⟨ρ⟩ ≅ tFst redR.(nf)⟨ρ⟩ : ΣA.(ParamRedTyPack.dom)⟨ρ⟩] ; + eqSnd [Δ] (ρ : Δ ≤ Γ) (h : [ |- Δ ]) : + [ΣA.(PolyRedPack.posRed) ρ h (eqFst ρ h) | Δ ||- tSnd redL.(nf)⟨ρ⟩ ≅ tSnd redR.(nf)⟨ρ⟩ : _] ; }. Arguments SigRedTmEq {_ _ _ _ _ _ _ _ _ _ _}. End SigRedTmEq. -Export SigRedTmEq(SigRedTmEq,Build_SigRedTmEq). +Export SigRedTmEq(SigRedTm,Build_SigRedTm, SigRedTmEq,Build_SigRedTmEq). Notation "[ Γ ||-Σ t ≅ u : A | ΣA ]" := (SigRedTmEq (Γ:=Γ) (A:=A) ΣA t u) (at level 0, Γ, t, u, A, ΣA at level 50). @@ -551,22 +517,17 @@ Notation "[ Γ ||-Σ t ≅ u : A | ΣA ]" := (SigRedTmEq (Γ:=Γ) (A:=A) ΣA t u (** ** Reducibility of neutrals at an arbitrary type *) Module NeNf. - Record RedTm `{ta : tag} `{Typing ta} `{ConvNeuConv ta} {Γ k A} : Set := - { - ty : [Γ |- k : A] ; - refl : [Γ |- k ~ k : A] - }. - Arguments RedTm {_ _ _}. - Record RedTmEq `{ta : tag} `{ConvNeuConv ta} {Γ k l A} : Set := + Record RedTmEq `{ta : tag} `{Typing ta} `{ConvNeuConv ta} {Γ k l A} : Set := { + tyL : [Γ |- k : A] ; + tyR : [Γ |- l : A] ; conv : [Γ |- k ~ l : A] }. - Arguments RedTmEq {_ _}. + Arguments RedTmEq {_ _ _}. End NeNf. -Notation "[ Γ ||-NeNf k : A ]" := (NeNf.RedTm Γ k A) (at level 0, Γ, k, A at level 50). Notation "[ Γ ||-NeNf k ≅ l : A ]" := (NeNf.RedTmEq Γ k l A) (at level 0, Γ, k, l, A at level 50). (** ** Reducibility of natural number type *) @@ -601,72 +562,6 @@ Export NatRedTyEq(NatRedTyEq,Build_NatRedTyEq). Notation "[ Γ ||-Nat A ≅ B | RA ]" := (@NatRedTyEq _ _ _ Γ A RA B) (at level 0, Γ, A, B, RA at level 50). -Module NatRedTm. -Section NatRedTm. - Context `{ta : tag} `{WfType ta} - `{RedType ta} `{Typing ta} `{ConvNeuConv ta} `{ConvTerm ta} - `{RedTerm ta}. - - Inductive NatRedTm {Γ : context} {A: term} {NA : NatRedTy Γ A} : term -> Set := - | Build_NatRedTm {t} - (nf : term) - (red : [Γ |- t :⤳*: nf : tNat ]) - (eq : [Γ |- nf ≅ nf : tNat]) - (prop : NatProp nf) : NatRedTm t - - with NatProp {Γ : context} {A: term} {NA : NatRedTy Γ A} : term -> Set := - | zeroR : NatProp tZero - | succR {n} : - NatRedTm n -> - NatProp (tSucc n) - | neR {ne} : [Γ ||-NeNf ne : tNat] -> NatProp ne. - - -Scheme NatRedTm_mut_rect := Induction for NatRedTm Sort Type with - NatProp_mut_rect := Induction for NatProp Sort Type. - -Combined Scheme _NatRedInduction from - NatRedTm_mut_rect, - NatProp_mut_rect. - -Let _NatRedInductionType := - ltac:(let ind := fresh "ind" in - pose (ind := _NatRedInduction); - let ind_ty := type of ind in - exact ind_ty). - -Let NatRedInductionType := - ltac: (let ind := eval cbv delta [_NatRedInductionType] zeta - in _NatRedInductionType in - let ind' := polymorphise ind in - exact ind'). - -(* KM: looks like there is a bunch of polymorphic universes appearing there... *) -Lemma NatRedInduction : NatRedInductionType. -Proof. - intros ??? PRed PProp **; split; now apply (_NatRedInduction _ _ _ PRed PProp). -Defined. - -Definition nf {Γ A n} {NA : [Γ ||-Nat A]} : @NatRedTm _ _ NA n -> term. -Proof. - intros [? nf]. exact nf. -Defined. - -Definition red {Γ A n} {NA : [Γ ||-Nat A]} (Rn : @NatRedTm _ _ NA n) : [Γ |- n :⤳*: nf Rn : tNat]. -Proof. - dependent inversion Rn; subst; cbn; tea. -Defined. - -End NatRedTm. -Arguments NatRedTm {_ _ _ _ _ _ _ _ _}. -Arguments NatProp {_ _ _ _ _ _ _ _ _}. - -End NatRedTm. - -Export NatRedTm(NatRedTm,Build_NatRedTm, NatProp, NatRedInduction). - -Notation "[ Γ ||-Nat t : A | RA ]" := (@NatRedTm _ _ _ _ _ _ _ Γ A RA t) (at level 0, Γ, t, A, RA at level 50). - Module NatRedTmEq. Section NatRedTmEq. @@ -690,6 +585,26 @@ Section NatRedTmEq. NatPropEq (tSucc n) (tSucc n') | neReq {ne ne'} : [Γ ||-NeNf ne ≅ ne' : tNat] -> NatPropEq ne ne'. +Definition nfL {Γ A n m} {NA : [Γ ||-Nat A]} : NatRedTmEq (NA:=NA) n m -> term. +Proof. + intros [?? nf]. exact nf. +Defined. + +Definition nfR {Γ A n m} {NA : [Γ ||-Nat A]} : NatRedTmEq (NA:=NA) n m -> term. +Proof. + intros [?? ? nf]. exact nf. +Defined. + +Definition redL {Γ A n m} {NA : [Γ ||-Nat A]} (Rnm : NatRedTmEq (NA:=NA) n m) : [Γ |- n :⤳*: nfL Rnm : tNat]. +Proof. + dependent inversion Rnm; subst; cbn; tea. +Defined. + +Definition redR {Γ A n m} {NA : [Γ ||-Nat A]} (Rnm : NatRedTmEq (NA:=NA) n m) : [Γ |- m :⤳*: nfR Rnm : tNat]. +Proof. + dependent inversion Rnm; subst; cbn; tea. +Defined. + Scheme NatRedTmEq_mut_rect := Induction for NatRedTmEq Sort Type with NatPropEq_mut_rect := Induction for NatPropEq Sort Type. @@ -760,43 +675,6 @@ Export EmptyRedTyEq(EmptyRedTyEq,Build_EmptyRedTyEq). Notation "[ Γ ||-Empty A ≅ B | RA ]" := (@EmptyRedTyEq _ _ _ Γ A RA B) (at level 0, Γ, A, B, RA at level 50). -Module EmptyRedTm. -Section EmptyRedTm. - Context `{ta : tag} `{WfType ta} - `{RedType ta} `{Typing ta} `{ConvNeuConv ta} `{ConvTerm ta} - `{RedTerm ta}. - - Inductive EmptyProp {Γ : context} : term -> Set := - | neR {ne} : [Γ ||-NeNf ne : tEmpty] -> EmptyProp ne. - - Inductive EmptyRedTm {Γ : context} {A: term} {NA : EmptyRedTy Γ A} : term -> Set := - | Build_EmptyRedTm {t} - (nf : term) - (red : [Γ |- t :⤳*: nf : tEmpty ]) - (eq : [Γ |- nf ≅ nf : tEmpty]) - (prop : @EmptyProp Γ nf) : EmptyRedTm t. - -Definition nf {Γ A n} {NA : [Γ ||-Empty A]} : @EmptyRedTm _ _ NA n -> term. -Proof. - intros [? nf]. exact nf. -Defined. - -Definition red {Γ A n} {NA : [Γ ||-Empty A]} (Rn : @EmptyRedTm _ _ NA n) : [Γ |- n :⤳*: nf Rn : tEmpty]. -Proof. - dependent inversion Rn; subst; cbn; tea. -Defined. - -End EmptyRedTm. -Arguments EmptyRedTm {_ _ _ _ _ _ _ _ _}. -Arguments EmptyProp {_ _ _}. - -End EmptyRedTm. - -Export EmptyRedTm(EmptyRedTm,Build_EmptyRedTm, EmptyProp). - -Notation "[ Γ ||-Empty t : A | RA ]" := (@EmptyRedTm _ _ _ _ _ _ _ Γ A RA t) (at level 0, Γ, t, A, RA at level 50). - - Module EmptyRedTmEq. Section EmptyRedTmEq. Context `{ta : tag} `{WfContext ta} `{WfType ta} `{ConvType ta} @@ -814,9 +692,29 @@ Section EmptyRedTmEq. (eq : [Γ |- nfL ≅ nfR : tEmpty]) (prop : @EmptyPropEq Γ nfL nfR) : EmptyRedTmEq t u. + Definition nfL {Γ A n m} {NA : [Γ ||-Empty A]} : EmptyRedTmEq (NA:=NA) n m -> term. + Proof. + intros [?? nf]. exact nf. + Defined. + + Definition nfR {Γ A n m} {NA : [Γ ||-Empty A]} : EmptyRedTmEq (NA:=NA) n m -> term. + Proof. + intros [?? ? nf]. exact nf. + Defined. + + Definition redL {Γ A n m} {NA : [Γ ||-Empty A]} (Rn : EmptyRedTmEq (NA:=NA) n m) : [Γ |- n :⤳*: nfL Rn : tEmpty]. + Proof. + dependent inversion Rn; subst; cbn; tea. + Defined. + + Definition redR {Γ A n m} {NA : [Γ ||-Empty A]} (Rn : EmptyRedTmEq (NA:=NA) n m) : [Γ |- n :⤳*: nfL Rn : tEmpty]. + Proof. + dependent inversion Rn; subst; cbn; tea. + Defined. + End EmptyRedTmEq. Arguments EmptyRedTmEq {_ _ _ _ _ _ _ _ _}. -Arguments EmptyPropEq {_ _}. +Arguments EmptyPropEq {_ _ _}. End EmptyRedTmEq. Export EmptyRedTmEq(EmptyRedTmEq,Build_EmptyRedTmEq, EmptyPropEq). @@ -839,17 +737,13 @@ Module IdRedTyPack. red : [Γ |- A :⤳*: outTy] ; eq : [Γ |- outTy ≅ outTy] ; tyRed : LRPack@{i} Γ ty ; - lhsRed : [ tyRed | Γ ||- lhs : _ ] ; - rhsRed : [ tyRed | Γ ||- rhs : _ ] ; + lhsRed : [ tyRed | Γ ||- lhs ≅ lhs : _ ] ; + rhsRed : [ tyRed | Γ ||- rhs ≅ rhs : _ ] ; (* Bake in PER property for reducible conversion at ty to cut dependency cycles *) - lhsRedRefl : [ tyRed | Γ ||- lhs ≅ lhs : _ ] ; - rhsRedRefl : [ tyRed | Γ ||- rhs ≅ rhs : _ ] ; - tyPER : PER (fun t u => [tyRed | _ ||- t ≅ u : _]) ; + tyPER : PER tyRed.(LRPack.eqTm) ; tyKripke : forall {Δ} (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), LRPack@{i} Δ ty⟨ρ⟩ ; tyKripkeEq : forall {Δ Ξ} (ρ : Δ ≤ Γ) (ρ' : Ξ ≤ Γ) (ρ'' : Ξ ≤ Δ) (wfΔ : [|-Δ]) (wfΞ : [|-Ξ]) B, ρ' =1 ρ'' ∘w ρ -> [tyKripke ρ wfΔ | _ ||- _ ≅ B] -> [tyKripke ρ' wfΞ | _ ||- _ ≅ B⟨ρ''⟩]; - tyKripkeTm : forall {Δ Ξ} (ρ : Δ ≤ Γ) (ρ' : Ξ ≤ Γ) (ρ'' : Ξ ≤ Δ) (wfΔ : [|-Δ]) (wfΞ : [|-Ξ]) t, - ρ' =1 ρ'' ∘w ρ -> [tyKripke ρ wfΔ | _ ||- t : _] -> [tyKripke ρ' wfΞ | _ ||- t⟨ρ''⟩ : _]; tyKripkeTmEq : forall {Δ Ξ} (ρ : Δ ≤ Γ) (ρ' : Ξ ≤ Γ) (ρ'' : Ξ ≤ Δ) (wfΔ : [|-Δ]) (wfΞ : [|-Ξ]) t u, ρ' =1 ρ'' ∘w ρ -> [tyKripke ρ wfΔ | _ ||- t ≅ u : _] -> [tyKripke ρ' wfΞ | _ ||- t⟨ρ''⟩ ≅ u⟨ρ''⟩ : _]; }. @@ -894,40 +788,6 @@ End IdRedTyEq. Export IdRedTyEq(IdRedTyEq,Build_IdRedTyEq). -Module IdRedTm. -Section IdRedTm. - Universe i. - Context `{ta : tag} `{WfContext ta} `{WfType ta} - `{RedType ta} `{Typing ta} `{ConvType ta} `{ConvNeuConv ta} `{ConvTerm ta} - `{RedTerm ta} {Γ : context} {A: term} {IA : IdRedTyPack@{i} Γ A}. - - Inductive IdProp : term -> Type:= - | reflR {A x} : - [Γ |- A] -> - [Γ |- x : A] -> - [IA.(IdRedTyPack.tyRed) | _ ||- _ ≅ A] -> - [IA.(IdRedTyPack.tyRed) | _ ||- IA.(IdRedTyPack.lhs) ≅ x : _ ] -> - (* Should the index only be conversion ? *) - [IA.(IdRedTyPack.tyRed) | _ ||- IA.(IdRedTyPack.rhs) ≅ x : _ ] -> - IdProp (tRefl A x) - | neR {ne} : [Γ ||-NeNf ne : IA.(IdRedTyPack.outTy)] -> IdProp ne. - - Record IdRedTm {t : term} : Type := - Build_IdRedTm { - nf : term ; - red : [Γ |- t :⤳*: nf : IA.(IdRedTyPack.outTy) ] ; - eq : [Γ |- nf ≅ nf : IA.(IdRedTyPack.outTy)] ; - prop : IdProp nf ; - }. - -End IdRedTm. -Arguments IdRedTm {_ _ _ _ _ _ _ _ _ _ _}. -Arguments IdProp {_ _ _ _ _ _ _ _ _}. - -End IdRedTm. - -Export IdRedTm(IdRedTm,Build_IdRedTm, IdProp). - Module IdRedTmEq. Section IdRedTmEq. @@ -985,35 +845,25 @@ Inductive LR@{i j k} `{ta : tag} | LRU {Γ A} (H : [Γ ||-U A]) : LR rec Γ A (fun B => [Γ ||-U≅ B ]) - (fun t => [ rec | Γ ||-U t : A | H ]) (fun t u => [ rec | Γ ||-U t ≅ u : A | H ]) | LRne {Γ A} (neA : [ Γ ||-ne A ]) : LR rec Γ A (fun B => [ Γ ||-ne A ≅ B | neA]) - (fun t => [ Γ ||-ne t : A | neA]) (fun t u => [ Γ ||-ne t ≅ u : A | neA]) | LRPi {Γ : context} {A : term} (ΠA : PiRedTy@{j} Γ A) (ΠAad : PiRedTyAdequate@{j k} (LR rec) ΠA) : LR rec Γ A (fun B => [ Γ ||-Π A ≅ B | ΠA ]) - (fun t => [ Γ ||-Π t : A | ΠA ]) (fun t u => [ Γ ||-Π t ≅ u : A | ΠA ]) | LRNat {Γ A} (NA : [Γ ||-Nat A]) : - LR rec Γ A (NatRedTyEq NA) (NatRedTm NA) (NatRedTmEq NA) + LR rec Γ A (NatRedTyEq NA) (NatRedTmEq NA) | LREmpty {Γ A} (NA : [Γ ||-Empty A]) : - LR rec Γ A (EmptyRedTyEq NA) (EmptyRedTm NA) (EmptyRedTmEq NA) + LR rec Γ A (EmptyRedTyEq NA) (EmptyRedTmEq NA) | LRSig {Γ : context} {A : term} (ΣA : SigRedTy@{j} Γ A) (ΣAad : SigRedTyAdequate@{j k} (LR rec) ΣA) : - LR rec Γ A (SigRedTyEq ΣA) (SigRedTm ΣA) (SigRedTmEq ΣA) + LR rec Γ A (SigRedTyEq ΣA) (SigRedTmEq ΣA) | LRId {Γ A} (IA : IdRedTyPack@{j} Γ A) (IAad : IdRedTyAdequate@{j k} (LR rec) IA) : - LR rec Γ A (IdRedTyEq IA) (IdRedTm IA) (IdRedTmEq IA) + LR rec Γ A (IdRedTyEq IA) (IdRedTmEq IA) . - (** Removed, as it is provable (!), cf LR_embedding in LRInduction. *) - (* | LREmb {Γ A l'} (l_ : l' << l) (H : [ rec l' l_ | Γ ||- A]) : - LR rec Γ A - (fun B => [ rec l' l_ | Γ ||- A ≅ B | H ]) - (fun t => [ rec l' l_ | Γ ||- t : A | H ]) - (fun t u => [ rec l' l_ | Γ ||- t ≅ u : A | H ]) *) - Set Elimination Schemes. (** ** Bundling of the logical relation *) @@ -1033,10 +883,10 @@ Section MoreDefs. Definition LogRel0@{i j k} := LR@{i j k} rec0@{i j}. - Definition LRbuild0@{i j k} {Γ A rEq rTe rTeEq} : - LogRel0@{i j k} Γ A rEq rTe rTeEq -> [ LogRel0@{i j k} | Γ ||- A ] := + Definition LRbuild0@{i j k} {Γ A rEq rTeEq} : + LogRel0@{i j k} Γ A rEq rTeEq -> [ LogRel0@{i j k} | Γ ||- A ] := fun H => {| - LRAd.pack := {| LRPack.eqTy := rEq ; LRPack.redTm := rTe ; LRPack.eqTm := rTeEq |} ; + LRAd.pack := {| LRPack.eqTy := rEq ; LRPack.eqTm := rTeEq |} ; LRAd.adequate := H |}. Definition LogRelRec@{i j k} (l : TypeLevel) : forall l', l' << l -> RedRel@{j k} := @@ -1056,16 +906,16 @@ Section MoreDefs. Definition LogRel@{i j k l} (l : TypeLevel) := LR@{j k l} (LogRelRec@{i j k} l). - Definition LRbuild@{i j k l} {Γ l A rEq rTe rTeEq} : - LR@{j k l} (LogRelRec@{i j k} l) Γ A rEq rTe rTeEq -> [ LogRel l | Γ ||- A ] := + Definition LRbuild@{i j k l} {Γ l A rEq rTeEq} : + LR@{j k l} (LogRelRec@{i j k} l) Γ A rEq rTeEq -> [ LogRel l | Γ ||- A ] := fun H => {| - LRAd.pack := {| LRPack.eqTy := rEq ; LRPack.redTm := rTe ; LRPack.eqTm := rTeEq |} ; + LRAd.pack := {| LRPack.eqTy := rEq ; LRPack.eqTm := rTeEq |} ; LRAd.adequate := H |}. Definition LRunbuild {Γ l A} : [ LogRel l | Γ ||- A ] -> - ∑ rEq rTe rTeEq , LR (LogRelRec l) Γ A rEq rTe rTeEq := - fun H => (LRPack.eqTy H; LRPack.redTm H; LRPack.eqTm H; H.(LRAd.adequate)). + ∑ rEq rTeEq , LR (LogRelRec l) Γ A rEq rTeEq := + fun H => (LRPack.eqTy H; LRPack.eqTm H; H.(LRAd.adequate)). Definition LRU_@{i j k l} {l Γ A} (H : [Γ ||-U A]) : [ LogRel@{i j k l} l | Γ ||- A ] := @@ -1100,7 +950,7 @@ End MoreDefs. *) Notation "[ Γ ||-< l > A ]" := [ LogRel l | Γ ||- A ]. Notation "[ Γ ||-< l > A ≅ B | RA ]" := [ LogRel l | Γ ||- A ≅ B | RA ]. -Notation "[ Γ ||-< l > t : A | RA ]" := [ LogRel l | Γ ||- t : A | RA ]. +Notation "[ Γ ||-< l > t : A | RA ]" := [ LogRel l | Γ ||- t ≅ t : A | RA ]. Notation "[ Γ ||-< l > t ≅ u : A | RA ]" := [ LogRel l | Γ ||- t ≅ u : A | RA ]. Lemma instKripke `{GenericTypingProperties} {Γ A l} (wfΓ : [|-Γ]) (h : forall Δ (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), [Δ ||- A⟨ρ⟩]) : [Γ ||- A]. @@ -1127,18 +977,16 @@ Section PolyRed. { shpTy : [Γ |- shp] ; posTy : [Γ,, shp |- pos] ; - shpRed {Δ} (ρ : Δ ≤ Γ) : [ |- Δ ] -> [ LogRel@{i j k l} l | Δ ||- shp⟨ρ⟩ ] ; - posRed {Δ} {a} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) : - [ (shpRed ρ h) | Δ ||- a : shp⟨ρ⟩] -> + shpRed [Δ] (ρ : Δ ≤ Γ) : [ |- Δ ] -> [ LogRel@{i j k l} l | Δ ||- shp⟨ρ⟩ ] ; + posRed [Δ a b] (ρ : Δ ≤ Γ) (h : [ |- Δ ]) : + [ (shpRed ρ h) | Δ ||- a ≅ b : shp⟨ρ⟩] -> [ LogRel@{i j k l} l | Δ ||- pos[a .: (ρ >> tRel)]] ; posExt - {Δ a b} + [Δ a b] (ρ : Δ ≤ Γ) (h : [ |- Δ ]) - (ha : [ (shpRed ρ h) | Δ ||- a : shp⟨ρ⟩ ]) : - [ (shpRed ρ h) | Δ ||- b : shp⟨ρ⟩] -> - [ (shpRed ρ h) | Δ ||- a ≅ b : shp⟨ρ⟩] -> - [ (posRed ρ h ha) | Δ ||- (pos[a .: (ρ >> tRel)]) ≅ (pos[b .: (ρ >> tRel)]) ] + (hab : [ (shpRed ρ h) | Δ ||- a ≅ b : shp⟨ρ⟩]) : + [ (posRed ρ h hab) | Δ ||- (pos[a .: (ρ >> tRel)]) ≅ (pos[b .: (ρ >> tRel)]) ] }. Definition from@{i j k l} {PA : PolyRedPack@{k} Γ shp pos} @@ -1312,15 +1160,14 @@ Section LogRelRecFoldLemmas. Defined. (* This is a duplicate of the above, no ? *) - Lemma LogRelRec_unfold {Γ l A t eqTy redTm eqTm} (h: [Γ ||-U A]) : - LogRelRec l (URedTy.level h) (URedTy.lt h) Γ t eqTy redTm eqTm <~> - LogRel (URedTy.level h) Γ t eqTy redTm eqTm. + Lemma LogRelRec_unfold {Γ l A t eqTy eqTm} (h: [Γ ||-U A]) : + LogRelRec l (URedTy.level h) (URedTy.lt h) Γ t eqTy eqTm <~> + LogRel (URedTy.level h) Γ t eqTy eqTm. Proof. destruct l; [destruct (elim (URedTy.lt h))|]. destruct h; inversion lt; subst; cbn; now split. Qed. - Lemma TyEqRecFwd {l Γ A t u} (h : [Γ ||-U A]) (lr : [LogRelRec l (URedTy.level h) (URedTy.lt h) | Γ ||- t]) : [lr | Γ ||- t ≅ u] <~> [RedTyRecFwd h lr | Γ ||- t ≅ u]. @@ -1344,48 +1191,14 @@ End LogRelRecFoldLemmas. (** ** Properties of reducibility at Nat and Empty *) -Section NatPropProperties. - Context `{GenericTypingProperties}. - Lemma NatProp_whnf {Γ A t} {NA : [Γ ||-Nat A]} : NatProp NA t -> whnf t. - Proof. intros [ | | ? []]; now (econstructor; eapply convneu_whne). Qed. - - Lemma NatPropEq_whnf {Γ A t u} {NA : [Γ ||-Nat A]} : NatPropEq NA t u -> whnf t × whnf u. - Proof. intros [ | | ? ? []]; split; econstructor; eapply convneu_whne; first [eassumption|symmetry; eassumption]. Qed. - -End NatPropProperties. - -Section EmptyPropProperties. - Context `{GenericTypingProperties}. - Lemma EmptyProp_whnf {Γ A t} {NA : [Γ ||-Empty A]} : @EmptyProp _ _ _ Γ t -> whnf t. - Proof. intros [ ? []]; now (econstructor; eapply convneu_whne). Qed. +Lemma NatPropEq_whnf `{GenericTypingProperties} {Γ A t u} {NA : [Γ ||-Nat A]} : NatPropEq NA t u -> whnf t × whnf u. +Proof. intros [ | | ? ? []]; split; econstructor; eapply convneu_whne; first [eassumption|symmetry; eassumption]. Qed. - Lemma EmptyPropEq_whnf {Γ A t u} {NA : [Γ ||-Empty A]} : @EmptyPropEq _ _ Γ t u -> whnf t × whnf u. - Proof. intros [ ? ? []]; split; econstructor; eapply convneu_whne; first [eassumption|symmetry; eassumption]. Qed. - -End EmptyPropProperties. +Lemma EmptyPropEq_whnf `{GenericTypingProperties} {Γ A t u} {NA : [Γ ||-Empty A]} : EmptyPropEq Γ t u -> whnf t × whnf u. +Proof. intros [ ? ? []]; split; econstructor; eapply convneu_whne; first [eassumption|symmetry; eassumption]. Qed. (* A&Y: We prove the hand-crafted induction principles here: *) -Lemma EmptyRedInduction : - forall {ta : tag} {H : WfType ta} {H0 : RedType ta} {H1 : Typing ta} - {H2 : ConvNeuConv ta} {H3 : ConvTerm ta} {H4 : RedTerm ta} - (Γ : context) (A : term) (NA : [Γ ||-Empty A]) - (P : forall t : term, [Γ ||-Empty t : A | NA] -> Type) - (P0 : forall t : term, EmptyProp Γ t -> Type), - (forall (t nf : term) (red : [Γ |-[ ta ] t :⤳*: nf : tEmpty]) - (eq : [Γ |-[ ta ] nf ≅ nf : tEmpty]) (prop : EmptyProp Γ nf), - P0 nf prop -> P t (Build_EmptyRedTm nf red eq prop)) -> - (forall (ne : term) (r : [Γ ||-NeNf ne : tEmpty]), P0 ne (EmptyRedTm.neR r)) -> - (forall (t : term) (n : [Γ ||-Empty t : A | NA]), P t n) - × (forall (t : term) (n : EmptyProp Γ t), P0 t n). -Proof. - intros. split. - - intros. induction n. - eapply X. destruct prop. eapply X0. - - intros. induction n. eapply X0. -Qed. - - Lemma EmptyRedEqInduction : forall {ta : tag} {H0 : WfType ta} {H2 : RedType ta} {H3 : Typing ta} {H4 : ConvNeuConv ta} {H5 : ConvTerm ta} {H6 : RedTerm ta} @@ -1427,14 +1240,10 @@ Section IdRedTy. tyRed : [ LogRel@{i j k l} l | Γ ||- ty ] ; lhsRed : [ tyRed | Γ ||- lhs : _ ] ; rhsRed : [ tyRed | Γ ||- rhs : _ ] ; - lhsRedRefl : [ tyRed | Γ ||- lhs ≅ lhs : _ ] ; - rhsRedRefl : [ tyRed | Γ ||- rhs ≅ rhs : _ ] ; tyPER : PER (fun t u => [tyRed | _ ||- t ≅ u : _]) ; tyKripke : forall {Δ} (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), [ LogRel@{i j k l} l | Δ ||- ty⟨ρ⟩ ] ; tyKripkeEq : forall {Δ Ξ} (ρ : Δ ≤ Γ) (ρ' : Ξ ≤ Γ) (ρ'' : Ξ ≤ Δ) (wfΔ : [|-Δ]) (wfΞ : [|-Ξ]) B, ρ' =1 ρ'' ∘w ρ -> [tyKripke ρ wfΔ | _ ||- _ ≅ B] -> [tyKripke ρ' wfΞ | _ ||- _ ≅ B⟨ρ''⟩]; - tyKripkeTm : forall {Δ Ξ} (ρ : Δ ≤ Γ) (ρ' : Ξ ≤ Γ) (ρ'' : Ξ ≤ Δ) (wfΔ : [|-Δ]) (wfΞ : [|-Ξ]) t, - ρ' =1 ρ'' ∘w ρ -> [tyKripke ρ wfΔ | _ ||- t : _] -> [tyKripke ρ' wfΞ | _ ||- t⟨ρ''⟩ : _]; tyKripkeTmEq : forall {Δ Ξ} (ρ : Δ ≤ Γ) (ρ' : Ξ ≤ Γ) (ρ'' : Ξ ≤ Δ) (wfΔ : [|-Δ]) (wfΞ : [|-Ξ]) t u, ρ' =1 ρ'' ∘w ρ -> [tyKripke ρ wfΔ | _ ||- t ≅ u : _] -> [tyKripke ρ' wfΞ | _ ||- t⟨ρ''⟩ ≅ u⟨ρ''⟩ : _]; }. @@ -1449,11 +1258,8 @@ Section IdRedTy. - exact IA.(IdRedTyPack.eq). - exact IA.(IdRedTyPack.lhsRed). - exact IA.(IdRedTyPack.rhsRed). - - exact IA.(IdRedTyPack.lhsRedRefl). - - exact IA.(IdRedTyPack.rhsRedRefl). - exact IA.(IdRedTyPack.tyPER). - intros; now eapply IA.(IdRedTyPack.tyKripkeEq). - - intros; now eapply IA.(IdRedTyPack.tyKripkeTm). - intros; now eapply IA.(IdRedTyPack.tyKripkeTmEq). Defined. @@ -1465,11 +1271,8 @@ Section IdRedTy. - exact IA.(eq). - exact IA.(lhsRed). - exact IA.(rhsRed). - - exact IA.(IdRedTy.lhsRedRefl). - - exact IA.(IdRedTy.rhsRedRefl). - exact IA.(IdRedTy.tyPER). - intros; now eapply IA.(IdRedTy.tyKripkeEq). - - intros; now eapply IA.(IdRedTy.tyKripkeTm). - intros; now eapply IA.(IdRedTy.tyKripkeTmEq). Defined. @@ -1490,8 +1293,6 @@ Section IdRedTy. Proof. reflexivity. Qed. Definition IdRedTyEq {Γ l A} (IA : @IdRedTy Γ l A) := IdRedTyEq (toPack IA). - Definition IdRedTm {Γ l A} (IA : @IdRedTy Γ l A) := IdRedTm (toPack IA). - Definition IdProp {Γ l A} (IA : @IdRedTy Γ l A) := IdProp (toPack IA). Definition IdRedTmEq {Γ l A} (IA : @IdRedTy Γ l A) := IdRedTmEq (toPack IA). Definition IdPropEq {Γ l A} (IA : @IdRedTy Γ l A) := IdPropEq (toPack IA). @@ -1503,14 +1304,14 @@ End IdRedTy. Arguments IdRedTy {_ _ _ _ _ _ _ _ _}. End IdRedTy. -Export IdRedTy(IdRedTy, Build_IdRedTy,IdRedTyEq,IdRedTm,IdProp,IdRedTmEq,IdPropEq,LRId'). +Export IdRedTy(IdRedTy, Build_IdRedTy,IdRedTyEq,IdRedTmEq,IdPropEq,LRId'). Ltac unfold_id_outTy := change (IdRedTyPack.outTy (IdRedTy.toPack ?IA)) with (tId IA.(IdRedTy.ty) IA.(IdRedTy.lhs) IA.(IdRedTy.rhs)) in *. Notation "[ Γ ||-Id< l > A ]" := (IdRedTy Γ l A) (at level 0, Γ, l, A at level 50). Notation "[ Γ ||-Id< l > A ≅ B | RA ]" := (IdRedTyEq (Γ:=Γ) (l:=l) (A:=A) RA B) (at level 0, Γ, l, A, B, RA at level 50). -Notation "[ Γ ||-Id< l > t : A | RA ]" := (IdRedTm (Γ:=Γ) (l:=l) (A:=A) RA t) (at level 0, Γ, l, t, A, RA at level 50). +Notation "[ Γ ||-Id< l > t : A | RA ]" := (IdRedTmEq (Γ:=Γ) (l:=l) (A:=A) RA t t) (at level 0, Γ, l, t, A, RA at level 50). Notation "[ Γ ||-Id< l > t ≅ u : A | RA ]" := (IdRedTmEq (Γ:=Γ) (l:=l) (A:=A) RA t u) (at level 0, Γ, l, t, u, A, RA at level 50). diff --git a/theories/LogicalRelation/Application.v b/theories/LogicalRelation/Application.v index 916862a..2cf3d22 100644 --- a/theories/LogicalRelation/Application.v +++ b/theories/LogicalRelation/Application.v @@ -1,7 +1,7 @@ From LogRel.AutoSubst Require Import core unscoped Ast Extra. From LogRel Require Import Utils BasicAst Notations Context NormalForms Weakening GenericTyping LogicalRelation. -From LogRel.LogicalRelation Require Import Induction Irrelevance Escape Reflexivity Weakening Neutral Transitivity Reduction NormalRed. +From LogRel.LogicalRelation Require Import Induction Escape Irrelevance Reflexivity Weakening Neutral Transitivity Reduction NormalRed. Set Universe Polymorphism. @@ -14,7 +14,7 @@ Context `{GenericTypingProperties}. Set Printing Primitive Projection Parameters. -Section AppTerm. +(* Section AppTerm. Context {Γ t u F G l l' l''} (hΠ : [Γ ||-Π tProd F G]) {RF : [Γ ||- F]} @@ -22,7 +22,7 @@ Section AppTerm. (Ru : [Γ ||- u : F | RF]) (RGu : [Γ ||- G[u..]]). - Lemma app_id : [Γ ||- tApp (PiRedTm.nf Rt) u : G[u..] | RGu]. + Lemma app_id : [Γ ||- tApp (PiRedTmEq.nf Rt) u : G[u..] | RGu]. Proof. assert (wfΓ := wfc_wft (escape RF)). replace (PiRedTm.nf _) with (PiRedTm.nf Rt)⟨@wk_id Γ⟩ by now bsimpl. @@ -44,16 +44,16 @@ Section AppTerm. easy. Qed. -End AppTerm. +End AppTerm. *) -Lemma appTerm {Γ t u F G l} +(* Lemma appTerm {Γ t u F G l} (RΠ : [Γ ||- tProd F G]) {RF : [Γ ||- F]} (Rt : [Γ ||- t : tProd F G | RΠ]) (Ru : [Γ ||- u : F | RF]) - (RGu : [Γ ||- G[u..]]) : + (RGu : [Γ ||- G[u..]]) : [Γ ||- tApp t u : G[u..]| RGu]. -Proof. +Proof. unshelve eapply appTerm0. 7:irrelevance. 3: exact (invLRΠ RΠ). @@ -67,65 +67,55 @@ Lemma appTerm' {Γ t u F G l X} (Rt : [Γ ||- t : tProd F G | RΠ]) (Ru : [Γ ||- u : F | RF]) (eq : X = G[u..]) - (RX : [Γ ||- X]) : + (RX : [Γ ||- X]) : [Γ ||- tApp t u : X | RX]. -Proof. +Proof. irrelevance0; [symmetry; tea|]. unshelve eapply appTerm; cycle 1; tea. Unshelve. now rewrite <- eq. -Qed. +Qed. *) +Lemma scons_wk_id {Γ t u} : t[u .: wk_id (Γ:=Γ) >> tRel] = t[u..]. +Proof. now bsimpl. Qed. + +Lemma codSubst {Γ u u' F G l l'} + (RΠ : [Γ ||- tProd F G]) + {RF : [Γ ||- F]} + (Ruu' : [Γ ||- u ≅ u' : F | RF ]) + (RGu : [Γ ||- G[u..]]) : + [RGu | Γ ||- G[u..] ≅ G[u'..]]. +Proof. + pose (RΠ' := invLRcan RΠ ProdType); cbn in RΠ'. + set (h := normRedΠ0 _) in RΠ'. + assert (wfΓ : [|-Γ]) by (escape; gen_typing). + unshelve epose proof (PolyRed.posExt h wk_id wfΓ _). + 3:cbn; irrelevance. + cbn -[wk_id] in *. rewrite scons_wk_id in X. + irrelevance. +Qed. Lemma appcongTerm {Γ t t' u u' F G l l'} - (RΠ : [Γ ||- tProd F G]) + (RΠ : [Γ ||- tProd F G]) {RF : [Γ ||- F]} (Rtt' : [Γ ||- t ≅ t' : tProd F G | RΠ]) - (Ru : [Γ ||- u : F | RF]) - (Ru' : [Γ ||- u' : F | RF]) (Ruu' : [Γ ||- u ≅ u' : F | RF ]) - (RGu : [Γ ||- G[u..]]) - : + (RGu : [Γ ||- G[u..]]) : [Γ ||- tApp t u ≅ tApp t' u' : G[u..] | RGu]. Proof. - set (hΠ := invLRΠ RΠ); pose (RΠ' := LRPi' (normRedΠ0 hΠ)). - assert [Γ ||- t ≅ t' : tProd F G | RΠ'] as [Rt Rt' ? eqApp] by irrelevance. - set (h := invLRΠ _) in hΠ. - epose proof (e := redtywf_whnf (PiRedTyPack.red h) whnf_tProd); - symmetry in e; injection e; clear e; - destruct h as [?????? [?? domRed codRed codExt]] ; clear RΠ Rtt'; - intros; cbn in *; subst. + pose proof (codSubst RΠ Ruu' RGu). + normRedΠin Rtt'; destruct Rtt' as [Rt Rt' ? app]. assert (wfΓ : [|-Γ]) by gen_typing. - assert [Γ ||- u' : F⟨@wk_id Γ⟩ | domRed _ (@wk_id Γ) wfΓ] by irrelevance. - assert [Γ ||- u : F⟨@wk_id Γ⟩ | domRed _ (@wk_id Γ) wfΓ] by irrelevance. - assert (RGu' : [Γ ||- G[u'..]]). - 1:{ - replace G[u'..] with G[u' .: @wk_id Γ >> tRel] by now bsimpl. - now eapply (codRed _ u' (@wk_id Γ)). - } - assert (RGuu' : [Γ ||- G [u'..] ≅ G[u..] | RGu']). - 1:{ - replace G[u..] with G[u .: @wk_id Γ >> tRel] by now bsimpl. - irrelevance0. - 2: unshelve eapply codExt. - 6: eapply LRTmEqSym; irrelevance. - 2-4: tea. - now bsimpl. - } - eapply transEqTerm; eapply transEqTerm. - - eapply (snd (appTerm0 hΠ Rt Ru RGu)). - - unshelve epose proof (eqApp _ u (@wk_id Γ) wfΓ _). 1: irrelevance. - replace (PiRedTm.nf Rt) with (PiRedTm.nf Rt)⟨@wk_id Γ⟩ by now bsimpl. - irrelevance. - - unshelve epose proof (PiRedTm.eq Rt' (a:= u) (b:=u') (@wk_id Γ) wfΓ _ _ _). - all: irrelevance. - - replace (_)⟨_⟩ with (PiRedTm.nf Rt') by now bsimpl. - eapply LRTmEqRedConv; tea. - eapply LRTmEqSym. - eapply (snd (appTerm0 hΠ Rt' Ru' RGu')). + unshelve epose proof (hX := app Γ u u' (@wk_id Γ) wfΓ _). + 1: abstract (irrelevance0; [| exact Ruu']; cbn; now bsimpl). + pose proof (PiRedTmEq.red Rt); pose proof (PiRedTmEq.red Rt'). + unshelve epose proof (redSubstTmEq (A':=G[u'..]) (tl:=tApp t u) (tr:=tApp t' u') _ hX _ _ _). + all: rewrite ?wk_id_ren_on, ?scons_wk_id; escape; cbn in *; tea. + 1,2: gen_typing. + irrelevance. Qed. Lemma appcongTerm' {Γ t t' u u' F F' G l l' X} - (RΠ : [Γ ||- tProd F G]) + (RΠ : [Γ ||- tProd F G]) {RF : [Γ ||- F]} {RF' : [Γ ||- F']} (RFF' : [Γ ||- F ≅ F' | RF]) @@ -139,9 +129,7 @@ Lemma appcongTerm' {Γ t t' u u' F F' G l l' X} Proof. intros eq. irrelevance0 ; [symmetry; apply eq|]. - eapply appcongTerm; tea. - eapply LRTmRedConv; tea. - now eapply LRTyEqSym. + unshelve eapply appcongTerm; cycle 2; tea. Unshelve. now rewrite <- eq. Qed. diff --git a/theories/LogicalRelation/EqRedRight.v b/theories/LogicalRelation/EqRedRight.v new file mode 100644 index 0000000..c4a7abd --- /dev/null +++ b/theories/LogicalRelation/EqRedRight.v @@ -0,0 +1,128 @@ +(** * LogRel.EqRedRight: Reducibility of the rhs of a reducible conversion between types. *) +From Coq Require Import CRelationClasses. +From LogRel.AutoSubst Require Import core unscoped Ast Extra. +From LogRel Require Import Utils BasicAst Notations Context NormalForms UntypedReduction Weakening GenericTyping LogicalRelation. +From LogRel.LogicalRelation Require Import Induction Reflexivity Escape Irrelevance Weakening Transitivity Neutral. + +Set Universe Polymorphism. + + +Section Consequence. +Context `{GenericTypingProperties}. + +Lemma eq_id_subst_scons {Γ A} B : B = B[tRel 0 .: @wk1 Γ A >> tRel]. +Proof. + clear; bsimpl; rewrite scons_eta'; now bsimpl. +Qed. + +Set Printing Primitive Projection Parameters. + +Lemma posRedExt {Γ l A B A' B'} {PA : PolyRed Γ l A B} (PA' : PolyRedEq PA A' B') [Δ a b] + (ρ : Δ ≤ Γ) (h : [|- Δ]) (ha : [PolyRed.shpRed PA ρ h | Δ ||- a ≅ b : A⟨ρ⟩]) : + [PolyRed.posRed PA ρ h ha | Δ ||- B[a .: ρ >> tRel] ≅ B'[b .: ρ >> tRel]]. +Proof. + eapply LRTransEq. + 1: eapply (PolyRed.posExt PA). + unshelve eapply (PolyRedEq.posRed PA'). + 2: tea. + 2: now symmetry. +Qed. + + +Lemma PolyRedEqRedRight {Γ l A B A' B'} (PA : PolyRed Γ l A B) + (ihA : forall (Δ : context) (ρ : Δ ≤ Γ) (h : [|- Δ]) X, + [PolyRed.shpRed PA ρ h | Δ ||- A⟨ρ⟩ ≅ X] -> [Δ ||-< l > X]) + (ihB : forall (Δ : context) (a b : term) (ρ : Δ ≤ Γ) (h : [|- Δ]) + (ha : [PolyRed.shpRed PA ρ h | Δ ||- a ≅ b : A⟨ρ⟩]) X, + [PolyRed.posRed PA ρ h ha | Δ ||- B[a .: ρ >> tRel] ≅ X] -> + [Δ ||-< l > X]) + (PA' : PolyRedEq PA A' B') : PolyRed Γ l A' B'. +Proof. + destruct PA; pose proof PA' as []; cbn in *. + assert [|-Γ] by gen_typing. + assert (hdom: forall (Δ : context) (ρ : Δ ≤ Γ), [ |-[ ta ] Δ] -> [Δ ||-< l > A'⟨ρ⟩]). + 1:{ intros; eapply ihA; eauto. Unshelve. tea. } + assert (hcod: forall (Δ : context) (a b : term) (ρ : Δ ≤ Γ) + (h : [ |-[ ta ] Δ]), [hdom Δ ρ h | Δ ||- a ≅ b : A'⟨ρ⟩] -> + [Δ ||-< l > B'[a .: ρ >> tRel]]). + 1:{ + intros; eapply ihB; eauto. Unshelve. + 2: tea. + 2: eapply LRTmEqConv; tea. + unshelve eapply LRTyEqSym; cycle 1 ; eauto. + } + assert [Γ |- A'] by now eapply escape, instKripke. + assert [|-Γ,,A'] by gen_typing. + unshelve econstructor; tea. + - unshelve epose proof (X := hcod (Γ,,A') (tRel 0) (tRel 0) (wk1 _) _ _); tea. + 1: eapply var0; tea; now bsimpl. + escape; now rewrite <- eq_id_subst_scons in EscX. + - intros. + assert [hdom Δ ρ h | _ ||- _ ≅ A⟨ρ⟩]. + 1: eapply LRTyEqSym ; unshelve eauto; tea. + assert [shpRed Δ ρ h | _ ||- a ≅ b : _]. + 1: eapply LRTmEqConv; tea. + eapply LRTransEq. + + unshelve eapply LRTyEqSym, (posRedExt PA'); [|tea|now symmetry]. + + unshelve eapply (PolyRedEq.posRed PA'); [|tea|now symmetry]. +Qed. + +#[program] +Definition mkIdRedTy {Γ l A} ty lhs rhs (outTy := tId ty lhs rhs) + (red : [Γ |- A :⤳*: outTy]) + (tyRed : [ LogRel l | Γ ||- ty ]) + (lhsRed : [ tyRed | Γ ||- lhs : _ ]) + (rhsRed : [ tyRed | Γ ||- rhs : _ ]) : [Γ ||-Id A] := + {| IdRedTy.ty := ty ; + IdRedTy.lhs := lhs ; + IdRedTy.rhs := rhs |}. +Next Obligation. + pose proof (reflLRTyEq tyRed). + escape. timeout 1 gen_typing. +Qed. +Next Obligation. apply LRTmEqPER. Qed. +Next Obligation. now apply wk. Qed. +Next Obligation. + irrelevance0. 2: now eapply wkEq. + match goal with [ H : _ =1 _ |- _] => + clear -H; bsimpl; rewrite H; now bsimpl + end. + Unshelve. tea. +Qed. +Next Obligation. + irrelevance0. 2: now eapply wkTermEq. + match goal with [ H : _ =1 _ |- _] => + clear -H; bsimpl; rewrite H; now bsimpl + end. + Unshelve. tea. +Qed. + +Lemma LRTyEqRedRight {Γ l A B} (RA : [Γ ||- A]) : + [RA | Γ ||- A ≅ B] -> [Γ ||- B]. +Proof. + revert B; pattern l, Γ, A, RA. + apply LR_rect_TyUr; clear l Γ A RA. + - intros ??? [l0] ? []; eapply LRU_; exists l0; tea. + - intros ??? [] ? [nf']; eapply LRne_; exists nf'; tea. + cbn in *; now eapply urefl. + - intros ??? [] ??? [A' B']; cbn in *; eapply LRPi'. + exists A' B'; tea. + 1,2: now eapply urefl. + now eapply PolyRedEqRedRight. + - intros ????? []; eapply LRNat_; now constructor. + - intros ????? []; eapply LREmpty_; now constructor. + - intros ??? [] ??? [A' B']; cbn in *; eapply LRSig'. + exists A' B'; tea. + 1,2: now eapply urefl. + now eapply PolyRedEqRedRight. + - intros ??? [] ??? [tynf l r]; cbn in *; eapply LRId', (mkIdRedTy tynf l r); tea. + all: eapply LRTmEqConv;[| now eapply urefl]; tea. + Unshelve. eauto. + Qed. + +End Consequence. + +Ltac eqty_escape_right RA H ::= + let X := fresh "EscR" H in + pose proof (X := escape (LRTyEqRedRight RA H)). + diff --git a/theories/LogicalRelation/Escape.v b/theories/LogicalRelation/Escape.v index de7e880..af5cc20 100644 --- a/theories/LogicalRelation/Escape.v +++ b/theories/LogicalRelation/Escape.v @@ -33,7 +33,7 @@ Section Escapes. [Γ |- A ≅ B]. Proof. pattern l, Γ, A, lr ; eapply LR_rect_TyUr. - + intros ??? [] []. + + intros ??? [] []. gen_typing. + intros ??? [] []. cbn in *. @@ -45,26 +45,26 @@ Section Escapes. + intros ??? [] []; gen_typing. + intros ??? [] * ? ? []; cbn in *. eapply convty_exp. all: gen_typing. - + intros ??? [???? red] ?? [???? red']; cbn in *. + + intros ??? [???? red] ?? [???? red']; cbn in *. eapply convty_exp; tea;[eapply red | eapply red']. Qed. - Definition escapeTerm {l Γ t A} (lr : [Γ ||-< l > A ]) : - [Γ ||-< l > t : A | lr ] -> + Definition escapeTerm {l Γ t u A} (lr : [Γ ||-< l > A ]) : + [Γ ||-< l > t ≅ u : A | lr ] -> [Γ |- t : A]. Proof. pattern l, Γ, A, lr ; eapply LR_rect_TyUr. - - intros ??? [] [] ; cbn in *. + - intros ??? [] [[]] ; cbn in *. gen_typing. - intros ??? [] [] ; cbn in *. gen_typing. - - intros ??? [] * ?? [] ; cbn in *. + - intros ??? [] * ?? [[]] ; cbn in *. gen_typing. - intros ??? [] []; gen_typing. - intros ??? [] []; gen_typing. - - intros ??? [] * ?? [] ; cbn in *. + - intros ??? [] * ?? [[]] ; cbn in *. gen_typing. - - intros ??? IA _ _ []. + - intros ??? IA _ _ []. unfold_id_outTy; destruct IA; cbn in *; gen_typing. Qed. @@ -81,7 +81,7 @@ Section Escapes. assert (isPosType ty). { constructor. - now eapply convneu_whne. + now eapply convneu_whne. } eapply (convtm_conv (A := ty)). eapply convtm_exp ; tea. @@ -119,25 +119,34 @@ Section Escapes. - intros * ??? []; gen_typing. - intros * _ _ ? []; gen_typing. Qed. - + End Escapes. + +(* hack to redefine the symmetric version for term later *) +Ltac sym_escape RA H := idtac. +Ltac eqty_escape_right RA H := idtac. + Ltac escape := repeat lazymatch goal with - | [H : [_ ||-< _ > _] |- _] => - let X := fresh "Esc" H in - try pose proof (X := escape H) ; + | [H : [_ ||-< _ > _] |- _] => + try + (let X := fresh "Esc" H in + pose proof (X := escape H)); block H | [H : [_ ||-<_> _ ≅ _ | ?RA ] |- _] => - let X := fresh "Esc" H in - try pose proof (X := escapeEq RA H) ; + try + (let X := fresh "Esc" H in + pose proof (X := escapeEq RA H)) ; + try (eqty_escape_right RA H) ; block H - | [H : [_ ||-<_> _ : _ | ?RA] |- _] => + (* | [H : [_ ||-<_> _ : _ | ?RA] |- _] => let X := fresh "R" H in try pose proof (X := escapeTerm RA H) ; - block H + block H *) | [H : [_ ||-<_> _ ≅ _ : _ | ?RA] |- _] => - let X := fresh "R" H in - try pose proof (X := escapeEqTerm RA H) ; + try (let X := fresh "R" H in pose proof (X := escapeEqTerm RA H)) ; + try (let X := fresh "Rl" H in pose proof (X := escapeTerm RA H)) ; + try (sym_escape RA H) ; block H end; unblock. diff --git a/theories/LogicalRelation/Id.v b/theories/LogicalRelation/Id.v index d110840..375b3d6 100644 --- a/theories/LogicalRelation/Id.v +++ b/theories/LogicalRelation/Id.v @@ -1,6 +1,6 @@ From LogRel.AutoSubst Require Import core unscoped Ast Extra. From LogRel Require Import Notations Utils BasicAst Context NormalForms Weakening GenericTyping LogicalRelation. -From LogRel.LogicalRelation Require Import Induction Irrelevance Weakening Neutral Escape Reflexivity NormalRed Reduction Transitivity Universe. +From LogRel.LogicalRelation Require Import Induction Escape Irrelevance Weakening Neutral Reflexivity NormalRed Reduction Transitivity Universe EqRedRight. Set Universe Polymorphism. Set Printing Primitive Projection Parameters. @@ -8,32 +8,6 @@ Set Printing Primitive Projection Parameters. Section IdRed. Context `{GenericTypingProperties}. - Lemma mkIdRedTy {Γ l A} : - forall (ty lhs rhs : term) - (tyRed : [Γ ||- ty]), - [Γ |- A :⤳*: tId ty lhs rhs] -> - [tyRed | Γ ||- lhs : _] -> - [tyRed | Γ ||- rhs : _] -> - [Γ ||-Id A]. - Proof. - intros; unshelve econstructor; cycle 3; tea. - 1: intros; now eapply wk. - 2,3: now eapply reflLRTmEq. - 2: apply perLRTmEq. - - eapply convty_Id. - 1: eapply escapeEq; now eapply reflLRTyEq. - 1,2: eapply escapeEqTerm; now eapply reflLRTmEq. - - cbn; intros; irrelevance0; [|now eapply wkEq]. - bsimpl; rewrite H11; now bsimpl. - Unshelve. all:tea. - - cbn; intros; irrelevance0; [| now eapply wkTerm]. - bsimpl; rewrite H11; now bsimpl. - Unshelve. all:tea. - - cbn; intros; irrelevance0; [| now eapply wkTermEq]. - bsimpl; rewrite H11; now bsimpl. - Unshelve. all:tea. - Defined. - Lemma IdRed0 {Γ l A t u} (RA : [Γ ||- A]) : [RA | Γ ||- t : _] -> [RA | Γ ||- u : _] -> @@ -43,20 +17,20 @@ Section IdRed. 1: eapply redtywf_refl; escape; gen_typing. all: tea. Defined. - + Lemma IdRed {Γ l A t u} (RA : [Γ ||- A]) : [RA | Γ ||- t : _] -> [RA | Γ ||- u : _] -> [Γ ||- tId A t u]. Proof. intros; apply LRId'; now eapply IdRed0. Defined. - + Lemma IdRedTy_inv {Γ l A t u} (RIA : [Γ ||-Id tId A t u]) : [× A = RIA.(IdRedTy.ty), t = RIA.(IdRedTy.lhs) & u = RIA.(IdRedTy.rhs)]. Proof. pose proof (redtywf_whnf RIA.(IdRedTy.red) whnf_tId) as e; injection e; now split. Qed. - Lemma IdCongRed {Γ l A A' t t' u u'} + Lemma IdCongRed {Γ l A A' t t' u u'} (RA : [Γ ||- A]) (RIA : [Γ ||- tId A t u]) : [Γ |- tId A' t' u'] -> @@ -74,70 +48,60 @@ Section IdRed. + irrelevance. + cbn; rewrite <- elhs; irrelevance. + cbn; rewrite <- erhs; irrelevance. - Qed. - - Lemma IdRedU@{i j k l} {Γ l A t u} + Qed. + + (* Lemma IdRedU@{i j k l} {Γ l A t u} (RU : [LogRel@{i j k l} l | Γ ||- U]) (RU' := invLRU RU) (RA : [LogRel@{i j k l} (URedTy.level RU') | Γ ||- A]) : - [RU | Γ ||- A : U] -> - [RA | Γ ||- t : _] -> - [RA | Γ ||- u : _] -> + [RU | Γ ||- A ≅ A' : U] -> + [RA | Γ ||- t ≅ t': _] -> + [RA | Γ ||- u ≅: _] -> [RU | Γ ||- tId A t u : U]. Proof. intros RAU Rt Ru. enough [LRU_ RU' | _ ||- tId A t u : U] by irrelevance. econstructor. - - eapply redtmwf_refl; escape; now eapply ty_Id. + (* - eapply redtmwf_refl; escape; now eapply ty_Id. *) - constructor. - eapply convtm_Id; eapply escapeEq + eapply escapeEqTerm; now eapply reflLRTmEq + eapply reflLRTyEq. - eapply RedTyRecBwd. eapply IdRed; irrelevanceCum. Unshelve. irrelevanceCum. - Qed. - - Lemma IdCongRedU@{i j k l} {Γ l A A' t t' u u'} + Qed. *) + + Lemma IdCongRedU@{i j k l} {Γ l A A' t t' u u'} (RU : [LogRel@{i j k l} l | Γ ||- U]) (RU' := invLRU RU) - (RA : [LogRel@{i j k l} (URedTy.level RU') | Γ ||- A]) - (RA' : [LogRel@{i j k l} (URedTy.level RU') | Γ ||- A']) : + (RA : [LogRel@{i j k l} (URedTy.level RU') | Γ ||- A]) : [RU | Γ ||- A : U] -> - [RU | Γ ||- A' : U] -> [RU | _ ||- A ≅ A' : U] -> - [RA | _ ||- t : _] -> - [RA' | _ ||- t' : _] -> [RA | _ ||- t ≅ t' : _] -> - [RA | _ ||- u : _] -> - [RA' | _ ||- u' : _] -> [RA | _ ||- u ≅ u' : _] -> [RU | _ ||- tId A t u ≅ tId A' t' u' : U]. Proof. - intros RAU RAU' RAAU' Rt Rt' Rtt' Ru Ru' Ruu'. + intros RAU RAAU' Rtt' Ruu'. enough [LRU_ RU' | _ ||- tId A t u ≅ tId A' t' u': U] by irrelevance. - opector. - - change [LRU_ RU' | _ ||- tId A t u : _]. - enough [RU | _ ||- tId A t u : _] by irrelevance. - now unshelve eapply IdRedU. - - change [LRU_ RU' | _ ||- tId A' t' u' : _]. - enough [RU | _ ||- tId A' t' u' : _] by irrelevance. - now unshelve eapply IdRedU. - - eapply RedTyRecBwd. eapply IdRed; irrelevanceCum. - Unshelve. clear dependent RA'; irrelevanceCum. + assert (hAA' : [RA | _ ||- A ≅ A']). + 1: unshelve eapply UnivEqEq; try irrelevance + irrelevanceCum0. + escape. opector. + 1,2: eexists (tId _ _ _); [eapply redtmwf_refl; gen_typing| constructor]. + - eapply RedTyRecBwd; unshelve eapply IdRed. + all: try eapply lrefl; irrelevanceCum. - pose proof (redtmwf_whnf (URedTm.red u0) whnf_tId) as <-. pose proof (redtmwf_whnf (URedTm.red u1) whnf_tId) as <-. - eapply convtm_Id; now escape. - - eapply RedTyRecBwd. eapply IdRed; irrelevanceCum. - Unshelve. irrelevanceCum. + now eapply convtm_Id. + - eapply RedTyRecBwd; unshelve eapply IdRed. + 1: unshelve eapply UnivEq; [| |irrelevanceCum0| symmetry; irrelevanceCum]. + all: eapply urefl; eapply LRTmEqConv; tea. - eapply TyEqRecFwd. - eapply LRTyEqIrrelevantCum. - unshelve eapply IdCongRed; tea. - + escape; gen_typing. - + now eapply UnivEqEq. - Unshelve. now eapply IdRed. + unshelve eapply LRTyEqIrrelevantCum. + 2: eapply RedTyRecFwd in l0 ; irrelevanceCum. + unshelve eapply IdCongRed; tea; gen_typing. Qed. -Lemma reflRed {Γ l A x} (RA : [Γ ||- A]) (Rx : [RA | _ ||- x : _]) (RIA : [Γ ||- tId A x x]) : +(* Lemma reflRed {Γ l A x} (RA : [Γ ||- A]) (Rx : [RA | _ ||- x : _]) (RIA : [Γ ||- tId A x x]) : [RIA | _ ||- tRefl A x : _]. Proof. set (RIA' := invLRId RIA). @@ -157,374 +121,162 @@ Proof. Unshelve. all: tea. Qed. -Lemma reflRed' {Γ l A x} (RA : [Γ ||- A]) (Rx : [RA | _ ||- x : _]) +Lemma reflRed' {Γ l A x} (RA : [Γ ||- A]) (Rx : [RA | _ ||- x : _]) (RIA := IdRed RA Rx Rx): [RIA | _ ||- tRefl A x : _]. -Proof. now eapply reflRed. Qed. +Proof. now eapply reflRed. Qed. *) - -Lemma reflCongRed {Γ l A A' x x'} +Lemma reflCongRed {Γ l A A' x x'} (RA : [Γ ||- A]) - (TA' : [Γ |- A']) - (RAA' : [RA | _ ||- _ ≅ A']) - (Rx : [RA | _ ||- x : _]) - (Tx' : [Γ |- x' : A']) - (Rxx' : [RA | _ ||- x ≅ x' : _]) + (RAA : [RA | _ ||- _ ≅ A']) + (Rxx : [RA | _ ||- x ≅ x' : _]) (RIA : [Γ ||- tId A x x]) : [RIA | _ ||- tRefl A x ≅ tRefl A' x' : _]. Proof. - set (RIA' := invLRId RIA). - enough [LRId' RIA' | _ ||- tRefl A x ≅ tRefl A' x' : _] by irrelevance. - pose proof (IdRedTy_inv RIA') as [eA ex ex']. - assert (e : tId (IdRedTy.ty RIA') (IdRedTy.lhs RIA') (IdRedTy.rhs RIA') = tId A x x). - 1: f_equal; now symmetry. - assert [Γ |- tId A x x ≅ tId A' x' x'] by (escape ; gen_typing). - econstructor; unfold_id_outTy; cbn; rewrite ?e. - 1,2: eapply redtmwf_refl. - 1: escape; gen_typing. - - eapply ty_conv; [| now symmetry]; now eapply ty_refl. - - eapply convtm_refl; now escape. - - constructor; cbn. - 1-4: now escape. - all: irrelevance0; tea. - 1: apply reflLRTyEq. - 1,2: rewrite <- ex; tea; now eapply reflLRTmEq. - 1,2: rewrite <- ex'; tea; now eapply reflLRTmEq. - Unshelve. all: tea. -Qed. - -Definition idElimProp {Γ l} (A x P hr y e : term) {IA : [Γ ||-Id tId A x y]} (Pe : IdProp IA e) : term := - match Pe with - | IdRedTm.reflR _ _ _ _ _ => hr - | IdRedTm.neR _ => tIdElim A x P hr y e - end. - -Lemma idElimPropIrr {Γ l} {A x P hr y e : term} {IA : [Γ ||-Id tId A x y]} (Pe Pe' : IdProp IA e) : - idElimProp A x P hr y e Pe = idElimProp A x P hr y e Pe'. -Proof. - destruct Pe; cbn. - 2: dependent inversion Pe'; try reflexivity. - 2: subst; match goal with H : [_ ||-NeNf _ : _] |- _ => destruct H as [_ ?%convneu_whne]; inv_whne end; tea. - refine (match Pe' as Pe0 in IdRedTm.IdProp _ e return match e as e return IdProp _ e -> Type with | tRefl A0 x0 => fun Pe0 => hr = idElimProp A x P hr y (tRefl A0 x0) Pe0 | _ => fun _ => unit end Pe0 with | IdRedTm.reflR _ _ _ _ _ => _ | IdRedTm.neR r => _ end). - 1: reflexivity. - 1: match type of r with [_ ||-NeNf ?t : _] => destruct t ; try easy end. - exfalso; destruct r as [_ ?%convneu_whne]; inv_whne. + set (RIA' := normRedId RIA). + enough [RIA' | _ ||- tRefl A x ≅ tRefl A' x' : _] by irrelevance. + escape. + assert [Γ |- tId A' x' x' ≅ tId A x x] by (symmetry; timeout 1 gen_typing). + (* assert [Γ |- tId Ar xr xr ≅ tId A x x] by (symmetry; timeout 1 gen_typing). *) + exists (tRefl A x) (tRefl A' x'). + 1,2: eapply redtmwf_refl; unfold_id_outTy; cbn. + 1: gen_typing. + 1: eapply ty_conv; [|tea]; gen_typing. + 1: unfold_id_outTy; cbn; timeout 1 gen_typing. + constructor; tea. + 1: eapply ty_conv; tea. + all: try irrelevance. + 2,3: eapply lrefl; cbn; irrelevance. + eapply reflLRTyEq. Qed. -Lemma IdProp_refl_inv {Γ l A x y A' x'} {IA : [Γ ||-Id tId A x y]} (Pe : IdProp IA (tRefl A' x')) : - IdProp IA (tRefl A' x'). -Proof. - econstructor; inversion Pe. - all: try match goal with H : [_ ||-NeNf _ : _] |- _ => destruct H as [_ ?%convneu_whne]; inv_whne end; tea. -Defined. -Lemma IdRedTm_whnf_prop {Γ l A x y e} {IA : [Γ ||-Id tId A x y]} (Re : [_ ||-Id e : _ | IA]) : - whnf e -> IdProp IA e. -Proof. - intros wh; rewrite (redtmwf_whnf (IdRedTm.red Re) wh). - exact (IdRedTm.prop Re). -Qed. - -Lemma idElimPropRed {Γ l A x P hr y e} - (RA : [Γ ||- A]) - (Rx : [RA | _ ||- x : _]) - (RP0 : [Γ ,, A ,, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0) ||- P]) - (RP : forall y e (Ry : [RA | Γ ||- y : _]) (RIAxy : [Γ ||- tId A x y]), - [ RIAxy | _ ||- e : _] -> [Γ ||- P[e .: y..]]) - (RPeq : forall A' x' y y' e e' - (RA' : [Γ ||- A']) - (RAA' : [RA | _ ||- _ ≅ A']) - (Rx' : [RA | _ ||- x' : _]) - (Rxx' : [RA | _ ||- x ≅ x' : _]) - (Ry : [RA | Γ ||- y : _]) - (Ry' : [RA' | _ ||- y' : _]) - (Ryy' : [RA | Γ ||- y ≅ y' : _]) - (RIAxy : [Γ ||- tId A x y]) - (RIAxy' : [Γ ||- tId A x y']) - (Re : [ RIAxy | _ ||- e : _]) - (Re' : [ RIAxy' | _ ||- e' : _]) - (Ree' : [ RIAxy | _ ||- e ≅ e' : _]), - [RP y e Ry RIAxy Re | Γ ||- P[e .: y..] ≅ P[e' .: y' ..]]) - (Rhr : [RP x (tRefl A x) Rx (IdRed RA Rx Rx) (reflRed' RA Rx) | _ ||- hr : _]) - (Ry : [RA | _ ||- y : _]) - (RIAxy : [Γ ||- tId A x y]) - (Re : [RIAxy | _ ||- e : _]) - (RIAxy0 : [Γ ||-Id tId A x y]) - (Pe : IdProp RIAxy0 e) : - [RP y e Ry _ Re | _ ||- tIdElim A x P hr y e : _] × - [RP y e Ry _ Re | _ ||- tIdElim A x P hr y e ≅ idElimProp A x P hr y e Pe : _]. -Proof. - pose proof (IdRedTy_inv RIAxy0) as [eA ex ey]. - eapply redSubstTerm. - - destruct Pe; cbn in *. - + eapply LRTmRedConv; tea. - unshelve eapply RPeq; cycle 3; tea. - 2,3: eapply transEqTerm; [|eapply LRTmEqSym]; rewrite ?ex, ?ey; irrelevance. - * eapply reflLRTyEq. - * eapply reflCongRed; tea. - 1: irrelevance0; [symmetry; tea|]; tea. - rewrite ex; irrelevance. - + eapply neuTerm. - * escape; eapply ty_IdElim; tea. - * destruct r. - pose proof (reflLRTyEq RA). - pose proof (reflLRTyEq RP0). - pose proof Rx as ?%reflLRTmEq. - pose proof Ry as ?%reflLRTmEq. - pose proof Rhr as ?%reflLRTmEq. - escape; eapply convneu_IdElim; tea. - eapply convneu_conv; tea. unfold_id_outTy. - rewrite <- eA, <- ex, <- ey; eapply convty_Id; tea. - - destruct Pe; cbn in *. - + escape; eapply redtm_idElimRefl; tea. - * eapply ty_conv; tea; rewrite eA; now symmetry. - * now rewrite eA. - * rewrite eA, ex, ey; etransitivity; tea; now symmetry. - * now rewrite eA, ex. - + eapply redtm_refl; escape; now eapply ty_IdElim. -Qed. - - -Lemma idElimRed {Γ l A x P hr y e} +Lemma reflCongRed' {Γ l A Al Ar x xl xr} (RA : [Γ ||- A]) - (Rx : [RA | _ ||- x : _]) - (RP0 : [Γ ,, A ,, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0) ||- P]) - (RP : forall y e (Ry : [RA | Γ ||- y : _]) (RIAxy : [Γ ||- tId A x y]), - [ RIAxy | _ ||- e : _] -> [Γ ||- P[e .: y..]]) - (RPeq : forall A' x' y y' e e' - (RA' : [Γ ||- A']) - (RAA' : [RA | _ ||- _ ≅ A']) - (Rx' : [RA | _ ||- x' : _]) - (Rxx' : [RA | _ ||- x ≅ x' : _]) - (Ry : [RA | Γ ||- y : _]) - (Ry' : [RA' | _ ||- y' : _]) - (Ryy' : [RA | Γ ||- y ≅ y' : _]) - (RIAxy : [Γ ||- tId A x y]) - (RIAxy' : [Γ ||- tId A x y']) - (Re : [ RIAxy | _ ||- e : _]) - (Re' : [ RIAxy' | _ ||- e' : _]) - (Ree' : [ RIAxy | _ ||- e ≅ e' : _]), - [RP y e Ry RIAxy Re | Γ ||- P[e .: y..] ≅ P[e' .: y' ..]]) - (Rhr : [RP x (tRefl A x) Rx (IdRed RA Rx Rx) (reflRed' RA Rx) | _ ||- hr : _]) - (Ry : [RA | _ ||- y : _]) - (RIAxy : [Γ ||- tId A x y]) - (RIAxy' := LRId' (invLRId RIAxy)) - (Re : [RIAxy' | _ ||- e : _]) : - [RP y e Ry _ Re | _ ||- tIdElim A x P hr y e : _] × - [RP y e Ry _ Re | _ ||- tIdElim A x P hr y e ≅ tIdElim A x P hr y (IdRedTm.nf Re) : _]. -Proof. - pose proof (IdRedTy_inv (invLRId RIAxy)) as [eA ex ey]. - pose proof (hred := Re.(IdRedTm.red)); unfold_id_outTy; rewrite <-eA,<-ex,<-ey in hred. - eapply redSubstTerm. - - pose proof (redTmFwdConv Re hred (IdProp_whnf _ _ (IdRedTm.prop Re))) as [Rnf Rnfeq]. - eapply LRTmRedConv. - 2: eapply idElimPropRed; tea; exact (IdRedTm.prop Re). - unshelve eapply LRTyEqSym. - 2: now eapply RP. - eapply RPeq; cycle 2; first [eapply reflLRTyEq | now eapply reflLRTmEq | tea]. - Unshelve. all: tea. - - escape; eapply redtm_idElim; tea; apply hred. -Qed. + (RAAl : [RA | _ ||- _ ≅ Al]) + (RAAr : [RA | _ ||- _ ≅ Ar]) + (Rxxl : [RA | _ ||- x ≅ xl : _]) + (Rxxr : [RA | _ ||- x ≅ xr : _]) + (RIA : [Γ ||- tId A x x]) : + [RIA | _ ||- tRefl Al xl ≅ tRefl Ar xr : _]. +Proof. etransitivity; [symmetry|]; now eapply reflCongRed. Qed. Lemma idElimPropCongRed {Γ l A A' x x' P P' hr hr' y y' e e'} (RA : [Γ ||- A]) (RA' : [Γ ||- A']) (RAA' : [RA | Γ ||- A ≅ A']) - (Rx : [RA | _ ||- x : _]) - (Rx' : [RA' | _ ||- x' : _]) (Rxx' : [RA | _ ||- x ≅ x' : _]) - (RP0 : [Γ ,, A ,, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0) ||- P]) - (RP0' : [Γ ,, A' ,, tId A'⟨@wk1 Γ A'⟩ x'⟨@wk1 Γ A'⟩ (tRel 0) ||- P']) - (RPP0 : [RP0 | _ ||- _ ≅ P']) - (RP : forall y e (Ry : [RA | Γ ||- y : _]) (RIAxy : [Γ ||- tId A x y]), - [ RIAxy | _ ||- e : _] -> [Γ ||- P[e .: y..]]) - (RP' : forall y' e' (Ry' : [RA' | Γ ||- y' : _]) (RIAxy' : [Γ ||- tId A' x' y']), - [ RIAxy' | _ ||- e' : _] -> [Γ ||- P'[e' .: y'..]]) - (RPP' : forall y y' e e' - (Ry : [RA | Γ ||- y : _]) - (Ry' : [RA' | Γ ||- y' : _]) + (RIAxx : [Γ ||- tId A x x]) + (RP0 : [Γ ,, A ,, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0) |- P]) + (RP0' : [Γ ,, A' ,, tId A'⟨@wk1 Γ A'⟩ x'⟨@wk1 Γ A'⟩ (tRel 0) |- P']) + (RPP0 : [Γ ,, A ,, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0) |- P ≅ P']) + (RP : forall {y y' e e'} (Ryy' : [RA | Γ ||- y ≅ y' : _]) {RIAxy : [Γ ||- tId A x y]}, + [ RIAxy | _ ||- e ≅ e' : _] -> [Γ ||- P[e .: y..]]) + (RPP' : forall y y' e e' (Ryy' : [RA | _ ||- y ≅ y' : _]) (RIAxy : [Γ ||- tId A x y]) - (RIAxy' : [Γ ||- tId A' x' y']) - (Re : [ RIAxy | _ ||- e : _]) - (Re' : [ RIAxy' | _ ||- e' : _]) (Ree' : [ RIAxy | _ ||- e ≅ e' : _]), - [RP y e Ry _ Re | _ ||- P[e .: y..] ≅ P'[e' .: y'..]]) - (RPeq : forall A' x' y y' e e' - (RA' : [Γ ||- A']) + [RP Ryy' Ree' | _ ||- P[e .: y..] ≅ P'[e' .: y'..]]) + (RPeq : forall A' x' y y' e e' (RAA' : [RA | _ ||- _ ≅ A']) - (Rx' : [RA | _ ||- x' : _]) (Rxx' : [RA | _ ||- x ≅ x' : _]) - (Ry : [RA | Γ ||- y : _]) - (Ry' : [RA' | _ ||- y' : _]) (Ryy' : [RA | Γ ||- y ≅ y' : _]) (RIAxy : [Γ ||- tId A x y]) - (RIAxy' : [Γ ||- tId A x y']) - (Re : [ RIAxy | _ ||- e : _]) - (Re' : [ RIAxy' | _ ||- e' : _]) (Ree' : [ RIAxy | _ ||- e ≅ e' : _]), - [RP y e Ry RIAxy Re | Γ ||- P[e .: y..] ≅ P[e' .: y' ..]]) - (RPeq' : forall A1 x1 y' y1 e' e1 - (RA1 : [Γ ||- A1]) - (RAA1 : [RA' | _ ||- _ ≅ A1]) - (Rx1 : [RA' | _ ||- x1 : _]) - (Rxx1 : [RA' | _ ||- x' ≅ x1 : _]) - (Ry' : [RA' | Γ ||- y' : _]) - (Ry1 : [RA1 | _ ||- y1 : _]) - (Ryy1 : [RA' | Γ ||- y' ≅ y1 : _]) - (RIAxy' : [Γ ||- tId A' x' y']) - (RIAxy1 : [Γ ||- tId A' x' y1]) - (Re' : [ RIAxy' | _ ||- e' : _]) - (Re1 : [ RIAxy1 | _ ||- e1 : _]) - (Ree1 : [ RIAxy' | _ ||- e' ≅ e1 : _]), - [RP' y' e' Ry' RIAxy' Re' | Γ ||- P'[e' .: y'..] ≅ P'[e1 .: y1 ..]]) - (Rhr : [RP x (tRefl A x) Rx (IdRed RA Rx Rx) (reflRed' RA Rx) | _ ||- hr : _]) - (Rhr' : [RP' x' (tRefl A' x') Rx' (IdRed RA' Rx' Rx') (reflRed' RA' Rx') | _ ||- hr' : _]) - (Rhrhr' : [RP x (tRefl A x) Rx (IdRed RA Rx Rx) (reflRed' RA Rx) | _ ||- hr ≅ hr' : _]) - (Ry : [RA | _ ||- y : _]) - (Ry' : [RA' | _ ||- y' : _]) + [RP Ryy' Ree' | Γ ||- P[e .: y..] ≅ P[e' .: y' ..]]) + (Rhrhr' : [RP Rxx' (reflCongRed RA RAA' Rxx' RIAxx) | _ ||- hr ≅ hr' : _]) (Ryy' : [RA | _ ||- y ≅ y' : _]) (RIAxy : [Γ ||- tId A x y]) - (RIAxy' : [Γ ||- tId A' x' y']) - (Re : [RIAxy | _ ||- e : _]) - (Re' : [RIAxy' | _ ||- e' : _]) (Ree' : [RIAxy | _ ||- e ≅ e' : _]) - (RIAxy0 : [Γ ||-Id tId A x y]) - (Pee' : IdPropEq RIAxy0 e e') : - [RP y e Ry _ Re | _ ||- tIdElim A x P hr y e ≅ tIdElim A' x' P' hr' y' e' : _]. + (Pee' : IdPropEq (normRedId0 (invLRId RIAxy)) e e') : + [RP Ryy' Ree' | _ ||- tIdElim A x P hr y e ≅ tIdElim A' x' P' hr' y' e' : _]. Proof. - pose proof (IdRedTy_inv RIAxy0) as [eA ex ey]. - (* pose proof (IdRedTy_inv (invLRId RIAxy))) as [eA ex ey]. - pose proof (IdRedTy_inv (invLRId RIAxy')) as [eA' ex' ey']. *) - pose proof (IdPropEq_whnf _ _ _ Pee') as [whe whe']. - assert (Rei : [LRId' RIAxy0 | _ ||- e : _]) by irrelevance. - assert (Rei' : [LRId' (invLRId RIAxy') | _ ||- e' : _]) by irrelevance. - pose proof (IdRedTm_whnf_prop Rei whe). - pose proof (IdRedTm_whnf_prop Rei' whe'). - eapply LREqTermHelper. - 1,2: unshelve eapply idElimPropRed; tea. - 1: now eapply RPP'. - destruct Pee'; cbn in *. - + unshelve erewrite (idElimPropIrr X), (idElimPropIrr X0). - 1,2: now eapply IdProp_refl_inv. - cbn; eapply LRTmEqRedConv; tea. - eapply RPeq; cycle 3; tea. - 1,4: rewrite ex, ey; eapply transEqTerm; [|eapply LRTmEqSym]; irrelevance. - 2: eapply reflLRTyEq. - eapply reflCongRed; tea. - 1: irrelevance. - rewrite ex; irrelevance. - + unshelve erewrite (idElimPropIrr X), (idElimPropIrr X0); destruct r; unfold_id_outTy. - * econstructor; escape ; constructor; unfold_id_outTy. - 1: rewrite <- ex, <- ey, <-eA; tea. - now eapply lrefl. - * econstructor; pose proof (IdRedTy_inv (invLRId RIAxy')) as [eA' ex' ey']. - escape; constructor; unfold_id_outTy. - all: rewrite <- ex', <- ey', <-eA'; tea. - eapply convneu_conv; [now eapply urefl|]. - rewrite <- ex, <- ey, <-eA; now eapply convty_Id. - * cbn. escape; eapply neuTermEq. - - now eapply ty_IdElim. - - eapply ty_conv; [now eapply ty_IdElim|]. - symmetry; eapply escapeEq; now eapply RPP'. - Unshelve. all: tea. - - eapply convneu_IdElim; tea. - now rewrite ex, ey, eA. + pose proof (RPP' _ _ _ _ Rxx' _ (reflCongRed RA RAA' Rxx' RIAxx)). + pose proof (RPP' _ _ _ _ Ryy' _ Ree'). + destruct Pee'; cbn in *; escape. + - eapply redSubstTmEq; cycle 1. + + eapply redtm_idElimRefl; tea. + transitivity x'0; [|symmetry]; tea. + + eapply redtm_idElimRefl; tea. + 1-4: eapply ty_conv; tea. + 1: transitivity A; tea; now symmetry. + 2: eapply convtm_conv; tea; transitivity x; tea; now symmetry. + eapply convtm_conv; tea. + transitivity y; tea. + transitivity x0; symmetry; tea. + transitivity x; tea. now symmetry. + + unshelve eapply escapeEq, RPP'; tea. + + eapply LRTmEqConv; tea. + irrelevanceRefl. + eapply RPeq; tea. + Unshelve. + * transitivity x0; [|symmetry]; irrelevance. + * tea. + * now eapply reflCongRed. + - eapply neuTermEq. + + now eapply ty_IdElim. + + eapply ty_conv; [|symmetry; tea]. + eapply ty_IdElim; tea. + all: eapply ty_conv; tea. + now eapply convty_Id. + + eapply convneu_IdElim; tea. + destruct r; unfold_id_outTy; now cbn in *. Qed. + + + + Lemma idElimCongRed {Γ l A A' x x' P P' hr hr' y y' e e'} (RA : [Γ ||- A]) (RA' : [Γ ||- A']) (RAA' : [RA | Γ ||- A ≅ A']) - (Rx : [RA | _ ||- x : _]) - (Rx' : [RA' | _ ||- x' : _]) (Rxx' : [RA | _ ||- x ≅ x' : _]) - (RP0 : [Γ ,, A ,, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0) ||- P]) - (RP0' : [Γ ,, A' ,, tId A'⟨@wk1 Γ A'⟩ x'⟨@wk1 Γ A'⟩ (tRel 0) ||- P']) - (RPP0 : [RP0 | _ ||- _ ≅ P']) - (RP : forall y e (Ry : [RA | Γ ||- y : _]) (RIAxy : [Γ ||- tId A x y]), - [ RIAxy | _ ||- e : _] -> [Γ ||- P[e .: y..]]) - (RP' : forall y' e' (Ry' : [RA' | Γ ||- y' : _]) (RIAxy' : [Γ ||- tId A' x' y']), - [ RIAxy' | _ ||- e' : _] -> [Γ ||- P'[e' .: y'..]]) - (RPP' : forall y y' e e' - (Ry : [RA | Γ ||- y : _]) - (Ry' : [RA' | Γ ||- y' : _]) + (RIAxx : [Γ ||- tId A x x]) + (RP0 : [Γ ,, A ,, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0) |- P]) + (RP0' : [Γ ,, A' ,, tId A'⟨@wk1 Γ A'⟩ x'⟨@wk1 Γ A'⟩ (tRel 0) |- P']) + (RPP0 : [Γ ,, A ,, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0) |- P ≅ P']) + (RP : forall {y y' e e'} (Ryy' : [RA | Γ ||- y ≅ y' : _]) {RIAxy : [Γ ||- tId A x y]}, + [ RIAxy | _ ||- e ≅ e' : _] -> [Γ ||- P[e .: y..]]) + (RPP' : forall y y' e e' (Ryy' : [RA | _ ||- y ≅ y' : _]) (RIAxy : [Γ ||- tId A x y]) - (RIAxy' : [Γ ||- tId A' x' y']) - (Re : [ RIAxy | _ ||- e : _]) - (Re' : [ RIAxy' | _ ||- e' : _]) (Ree' : [ RIAxy | _ ||- e ≅ e' : _]), - [RP y e Ry _ Re | _ ||- P[e .: y..] ≅ P'[e' .: y'..]]) - (RPeq : forall A' x' y y' e e' - (RA' : [Γ ||- A']) + [RP Ryy' Ree' | _ ||- P[e .: y..] ≅ P'[e' .: y'..]]) + (RPeq : forall A' x' y y' e e' (RAA' : [RA | _ ||- _ ≅ A']) - (Rx' : [RA | _ ||- x' : _]) (Rxx' : [RA | _ ||- x ≅ x' : _]) - (Ry : [RA | Γ ||- y : _]) - (Ry' : [RA' | _ ||- y' : _]) (Ryy' : [RA | Γ ||- y ≅ y' : _]) (RIAxy : [Γ ||- tId A x y]) - (RIAxy' : [Γ ||- tId A x y']) - (Re : [ RIAxy | _ ||- e : _]) - (Re' : [ RIAxy' | _ ||- e' : _]) (Ree' : [ RIAxy | _ ||- e ≅ e' : _]), - [RP y e Ry RIAxy Re | Γ ||- P[e .: y..] ≅ P[e' .: y' ..]]) - (RPeq' : forall A1 x1 y' y1 e' e1 - (RA1 : [Γ ||- A1]) - (RAA1 : [RA' | _ ||- _ ≅ A1]) - (Rx1 : [RA' | _ ||- x1 : _]) - (Rxx1 : [RA' | _ ||- x' ≅ x1 : _]) - (Ry' : [RA' | Γ ||- y' : _]) - (Ry1 : [RA1 | _ ||- y1 : _]) - (Ryy1 : [RA' | Γ ||- y' ≅ y1 : _]) - (RIAxy' : [Γ ||- tId A' x' y']) - (RIAxy1 : [Γ ||- tId A' x' y1]) - (Re' : [ RIAxy' | _ ||- e' : _]) - (Re1 : [ RIAxy1 | _ ||- e1 : _]) - (Ree1 : [ RIAxy' | _ ||- e' ≅ e1 : _]), - [RP' y' e' Ry' RIAxy' Re' | Γ ||- P'[e' .: y'..] ≅ P'[e1 .: y1 ..]]) - (Rhr : [RP x (tRefl A x) Rx (IdRed RA Rx Rx) (reflRed' RA Rx) | _ ||- hr : _]) - (Rhr' : [RP' x' (tRefl A' x') Rx' (IdRed RA' Rx' Rx') (reflRed' RA' Rx') | _ ||- hr' : _]) - (Rhrhr' : [RP x (tRefl A x) Rx (IdRed RA Rx Rx) (reflRed' RA Rx) | _ ||- hr ≅ hr' : _]) - (Ry : [RA | _ ||- y : _]) - (Ry' : [RA' | _ ||- y' : _]) + [RP Ryy' Ree' | Γ ||- P[e .: y..] ≅ P[e' .: y' ..]]) + (Rhrhr' : [RP Rxx' (reflCongRed RA RAA' Rxx' RIAxx) | _ ||- hr ≅ hr' : _]) (Ryy' : [RA | _ ||- y ≅ y' : _]) (RIAxy : [Γ ||- tId A x y]) - (RIAxy' : [Γ ||- tId A' x' y']) - (Re : [RIAxy | _ ||- e : _]) - (Re' : [RIAxy' | _ ||- e' : _]) - (Ree' : [RIAxy | _ ||- e ≅ e' : _]) : - [RP y e Ry _ Re | _ ||- tIdElim A x P hr y e ≅ tIdElim A' x' P' hr' y' e' : _]. + (Ree' : [RIAxy | _ ||- e ≅ e' : _]) : + [RP Ryy' Ree' | _ ||- tIdElim A x P hr y e ≅ tIdElim A' x' P' hr' y' e' : _]. Proof. - pose proof (IdRedTy_inv (invLRId RIAxy)) as [eA ex ey]. - assert (RIAxyeq : [RIAxy | _ ||- _ ≅ tId A' x' y']) by (escape; now eapply IdCongRed). - assert (Req : [LRId' (invLRId RIAxy) | _ ||- e ≅ e' : _]) by irrelevance. - cbn in Req; inversion Req; unfold_id_outTy. - pose proof redL as [? _]; pose proof redR as [? _]. - rewrite <-eA,<-ex,<-ey in redL, redR. - pose proof (IdPropEq_whnf _ _ _ prop) as [whL whR]. - pose proof (redTmFwdConv Re redL whL) as [RnfL RnfLeq]. - unshelve epose proof (redTmFwdConv Re' _ whR) as [RnfR RnfReq]. - 1: eapply redtmwf_conv; tea; now escape. - eapply LREqTermHelper; cycle -1. - - eapply LRTmEqRedConv. - 2: eapply idElimPropCongRed. - 16: exact prop. - all: tea. - 1: eapply RPeq; cycle 2; first [now eapply reflLRTmEq | now eapply LRTmEqSym| eapply reflLRTyEq| tea]. - enough [LRId' (invLRId RIAxy) | _ ||- nfL ≅ nfR : _] by irrelevance. - exists nfL nfR; tea; eapply redtmwf_refl; unfold_id_outTy; tea. - - assert (Re1 : [LRId' (invLRId RIAxy) | _ ||- e : _]) by irrelevance. - rewrite (redtmwf_det whL (IdProp_whnf _ _ Re1.(IdRedTm.prop)) redL Re1.(IdRedTm.red)). - irrelevanceRefl; eapply idElimRed; tea. - - assert (Re1' : [LRId' (invLRId RIAxy') | _ ||- e' : _]) by irrelevance. - rewrite (redtmwf_det whR (IdProp_whnf _ _ Re1'.(IdRedTm.prop)) redR Re1'.(IdRedTm.red)). - irrelevanceRefl; eapply idElimRed; tea. - Unshelve. - all: tea. - now eapply RP'. - - now eapply RPP'. + assert (nRee' : [normRedId RIAxy | Γ ||- e ≅ e' : _]) by irrelevance. + destruct nRee' as [nfL nfR ??? prop]; unfold_id_outTy; cbn in *. + pose proof (RPP' _ _ _ _ Rxx' _ (reflCongRed RA RAA' Rxx' RIAxx)). + pose proof (RPP' _ _ _ _ Ryy' _ Ree'). + pose proof (IdPropEq_whnf _ _ _ prop) as []. + assert (RenfL : [RIAxy | _ ||- e ≅ nfL : _]). + 1: symmetry; eapply redTmEqFwd; tea; now eapply lrefl. + assert (RnfLR : [RIAxy | _ ||- nfL ≅ nfR : _]) + by (eapply redTmEqFwdBoth; cycle 1; tea). + assert (Ryy : [RA | _ ||- y ≅ y : _]) by now eapply lrefl. + pose proof (RAA := reflLRTyEq RA). + pose proof (RPeq _ _ _ _ _ _ RAA Rxx' Ryy _ RenfL). + escape. + assert [Γ |- tId A x y ≅ tId A' x' y'] by gen_typing. + eapply redSubstTmEq; cycle 1; tea. + 1,2: eapply redtm_idElim; tea. + 1: gen_typing. + 1-3: now eapply ty_conv. + 1: eapply redtm_conv; tea; gen_typing. + eapply LRTmEqConv. 1: now eapply LRTyEqSym. + now unshelve eapply idElimPropCongRed. Qed. End IdRed. \ No newline at end of file diff --git a/theories/LogicalRelation/Induction.v b/theories/LogicalRelation/Induction.v index dce530c..4083151 100644 --- a/theories/LogicalRelation/Induction.v +++ b/theories/LogicalRelation/Induction.v @@ -27,16 +27,16 @@ Section Inductions. (** Reducibility at a lower level implies reducibility at a higher level, and their decoding are the same. Both need to be proven simultaneously, because of contravariance in the product case. *) - + Fixpoint LR_embedding@{i j k l} {l l'} (l_ : l << l') - {Γ A rEq rTe rTeEq} (lr : LogRel@{i j k l} l Γ A rEq rTe rTeEq) {struct lr} - : (LogRel@{i j k l} l' Γ A rEq rTe rTeEq) := + {Γ A rEq rTeEq} (lr : LogRel@{i j k l} l Γ A rEq rTeEq) {struct lr} + : (LogRel@{i j k l} l' Γ A rEq rTeEq) := let embedPolyAd {Γ A B} {PA : PolyRedPack Γ A B} (PAad : PolyRedPackAdequate _ PA) := {| PolyRedPack.shpAd (Δ : context) (ρ : Δ ≤ _) (h : [ |- Δ]) := LR_embedding l_ (PAad.(PolyRedPack.shpAd) ρ h) ; - PolyRedPack.posAd (Δ : context) (a : term) (ρ : Δ ≤ _) (h : [ |- Δ]) - (ha : [PolyRedPack.shpRed PA ρ h | Δ ||- a : _]) := + PolyRedPack.posAd (Δ : context) (a b : term) (ρ : Δ ≤ _) (h : [ |- Δ]) + (ha : [PolyRedPack.shpRed PA ρ h | Δ ||- a ≅ b : _]) := LR_embedding l_ (PAad.(PolyRedPack.posAd) ρ h ha) |} in @@ -50,11 +50,11 @@ same. Both need to be proven simultaneously, because of contravariance in the pr | LRNat _ NA => LRNat _ NA | LREmpty _ NA => LREmpty _ NA | LRSig _ PA PAad => LRSig _ PA (embedPolyAd PAad) - | LRId _ IA IAad => - let embedIdAd := + | LRId _ IA IAad => + let embedIdAd := {| IdRedTyPack.tyAd := LR_embedding l_ IAad.(IdRedTyPack.tyAd) ; IdRedTyPack.tyKripkeAd (Δ : context) (ρ : Δ ≤ _) (wfΔ : [ |- Δ]) := - LR_embedding l_ (IAad.(IdRedTyPack.tyKripkeAd) ρ wfΔ) |} + LR_embedding l_ (IAad.(IdRedTyPack.tyKripkeAd) ρ wfΔ) |} in LRId _ IA embedIdAd end. @@ -62,22 +62,22 @@ same. Both need to be proven simultaneously, because of contravariance in the pr Notation PolyHyp P Γ ΠA HAad G := ((forall {Δ} (ρ : Δ ≤ Γ) (h : [ |- Δ]), P (HAad.(PolyRedPack.shpAd) ρ h)) -> - (forall {Δ a} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) - (ha : [ ΠA.(PolyRedPack.shpRed) ρ h | Δ ||- a : _ ]), + (forall {Δ a b} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) + (ha : [ ΠA.(PolyRedPack.shpRed) ρ h | Δ ||- a ≅ b: _ ]), P (HAad.(PolyRedPack.posAd) ρ h ha)) -> G). Theorem LR_rect@{i j k o} (l : TypeLevel) (rec : forall l', l' << l -> RedRel@{i j}) - (P : forall {c t rEq rTe rTeEq}, - LR@{i j k} rec c t rEq rTe rTeEq -> Type@{o}) : + (P : forall {c t rEq rTeEq}, + LR@{i j k} rec c t rEq rTeEq -> Type@{o}) : (forall (Γ : context) A (h : [Γ ||-U A]), P (LRU rec h)) -> (forall (Γ : context) (A : term) (neA : [Γ ||-ne A]), - P (LRne rec neA)) -> + P (LRne rec neA)) -> (forall (Γ : context) (A : term) (ΠA : PiRedTy@{j} Γ A) (HAad : PiRedTyAdequate (LR rec) ΠA), PolyHyp P Γ ΠA HAad (P (LRPi rec ΠA HAad))) -> @@ -89,18 +89,18 @@ same. Both need to be proven simultaneously, because of contravariance in the pr (forall (Γ : context) (A : term) (ΠA : SigRedTy@{j} Γ A) (HAad : SigRedTyAdequate (LR rec) ΠA), PolyHyp P Γ ΠA HAad (P (LRSig rec ΠA HAad))) -> - (forall Γ A (IA : IdRedTyPack@{j} Γ A) (IAad : IdRedTyAdequate (LR rec) IA), + (forall Γ A (IA : IdRedTyPack@{j} Γ A) (IAad : IdRedTyAdequate (LR rec) IA), P IAad.(IdRedTyPack.tyAd) -> (forall Δ (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), P (IAad.(IdRedTyPack.tyKripkeAd) ρ wfΔ)) -> P (LRId rec IA IAad)) -> - forall (Γ : context) (t : term) (rEq rTe : term -> Type@{j}) - (rTeEq : term -> term -> Type@{j}) (lr : LR@{i j k} rec Γ t rEq rTe rTeEq), + forall (Γ : context) (t : term) (rEq : term -> Type@{j}) + (rTeEq : term -> term -> Type@{j}) (lr : LR@{i j k} rec Γ t rEq rTeEq), P lr. Proof. cbn. intros HU Hne HPi HNat HEmpty HSig HId. - fix HRec 6. + fix HRec 5. destruct lr. - eapply HU. - eapply Hne. @@ -114,17 +114,17 @@ same. Both need to be proven simultaneously, because of contravariance in the pr Defined. Definition LR_rec@{i j k} := LR_rect@{i j k Set}. - + Notation PolyHypLogRel P Γ ΠA G := ((forall {Δ} (ρ : Δ ≤ Γ) (h : [ |- Δ]), P (ΠA.(PolyRed.shpRed) ρ h).(LRAd.adequate)) -> - (forall {Δ a} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) - (ha : [ Δ ||-< _ > a : ΠA.(ParamRedTy.dom)⟨ρ⟩ | ΠA.(PolyRed.shpRed) ρ h ]), + (forall {Δ a b} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) + (ha : [ Δ ||-< _ > a ≅ b : ΠA.(ParamRedTy.dom)⟨ρ⟩ | ΠA.(PolyRed.shpRed) ρ h ]), P (ΠA.(PolyRed.posRed) ρ h ha).(LRAd.adequate)) -> G). (** Induction principle specialized to LogRel as the reducibility relation on lower levels *) Theorem LR_rect_LogRelRec@{i j k l o} - (P : forall {l Γ t rEq rTe rTeEq}, - LogRel@{i j k l} l Γ t rEq rTe rTeEq -> Type@{o}) : + (P : forall {l Γ t rEq rTeEq}, + LogRel@{i j k l} l Γ t rEq rTeEq -> Type@{o}) : (forall l (Γ : context) A (h : [Γ ||-U A]), P (LRU (LogRelRec l) h)) -> @@ -134,22 +134,22 @@ same. Both need to be proven simultaneously, because of contravariance in the pr (forall (l : TypeLevel) (Γ : context) (A : term) (ΠA : ParamRedTy@{i j k l} tProd Γ l A), PolyHypLogRel P Γ ΠA (P (LRPi' ΠA).(LRAd.adequate ))) -> - + (forall l Γ A (NA : [Γ ||-Nat A]), P (LRNat (LogRelRec l) NA)) -> (forall l Γ A (NA : [Γ ||-Empty A]), P (LREmpty (LogRelRec l) NA)) -> - + (forall (l : TypeLevel) (Γ : context) (A : term) (ΠA : ParamRedTy@{i j k l} tSig Γ l A), PolyHypLogRel P Γ ΠA (P (LRSig' ΠA).(LRAd.adequate ))) -> - - (forall l Γ A (IA : [Γ ||-Id A]), - P (IA.(IdRedTy.tyRed).(LRAd.adequate)) -> + + (forall l Γ A (IA : [Γ ||-Id A]), + P (IA.(IdRedTy.tyRed).(LRAd.adequate)) -> (forall Δ (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), P (IA.(IdRedTy.tyKripke) ρ wfΔ).(LRAd.adequate)) -> - + P (LRId' IA).(LRAd.adequate)) -> - forall (l : TypeLevel) (Γ : context) (t : term) (rEq rTe : term -> Type@{k}) - (rTeEq : term -> term -> Type@{k}) (lr : LR@{j k l} (LogRelRec@{i j k} l) Γ t rEq rTe rTeEq), + forall (l : TypeLevel) (Γ : context) (t : term) (rEq : term -> Type@{k}) + (rTeEq : term -> term -> Type@{k}) (lr : LR@{j k l} (LogRelRec@{i j k} l) Γ t rEq rTeEq), P lr. Proof. intros ?? HPi ?? HSig HId **; eapply LR_rect@{j k l o}. @@ -161,8 +161,8 @@ same. Both need to be proven simultaneously, because of contravariance in the pr Notation PolyHypTyUr P Γ ΠA G := ((forall {Δ} (ρ : Δ ≤ Γ) (h : [ |- Δ]), P (ΠA.(PolyRed.shpRed) ρ h)) -> - (forall {Δ a} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) - (ha : [ ΠA.(PolyRed.shpRed) ρ h | Δ ||- a : ΠA.(ParamRedTy.dom)⟨ρ⟩ ]), + (forall {Δ a b} (ρ : Δ ≤ Γ) (h : [ |- Δ ]) + (ha : [ ΠA.(PolyRed.shpRed) ρ h | Δ ||- a ≅ b : ΠA.(ParamRedTy.dom)⟨ρ⟩ ]), P (ΠA.(PolyRed.posRed) ρ h ha)) -> G). Theorem LR_rect_TyUr@{i j k l o} @@ -180,21 +180,21 @@ same. Both need to be proven simultaneously, because of contravariance in the pr (forall l Γ A (NA : [Γ ||-Nat A]), P (LRNat_ l NA)) -> (forall l Γ A (NA : [Γ ||-Empty A]), P (LREmpty_ l NA)) -> - + (forall (l : TypeLevel) (Γ : context) (A : term) (ΠA : ParamRedTy@{i j k l} tSig Γ l A), PolyHypTyUr P Γ ΠA (P (LRSig' ΠA))) -> - - (forall l Γ A (IA : [Γ ||-Id A]), - P (IA.(IdRedTy.tyRed)) -> + + (forall l Γ A (IA : [Γ ||-Id A]), + P (IA.(IdRedTy.tyRed)) -> (forall Δ (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), P (IA.(IdRedTy.tyKripke) ρ wfΔ)) -> - + P (LRId' IA)) -> forall (l : TypeLevel) (Γ : context) (A : term) (lr : [LogRel@{i j k l} l | Γ ||- A]), P lr. Proof. intros HU Hne HPi HNat HEmpty HSig HId l Γ A lr. - apply (LR_rect_LogRelRec@{i j k l o} (fun l Γ A _ _ _ lr => P l Γ A (LRbuild lr))). + apply (LR_rect_LogRelRec@{i j k l o} (fun l Γ A _ _ lr => P l Γ A (LRbuild lr))). all: auto. Defined. @@ -211,11 +211,11 @@ same. Both need to be proven simultaneously, because of contravariance in the pr (forall Γ A (NA : [Γ ||-Nat A]), P (LRNat_ l NA)) -> (forall Γ A (NA : [Γ ||-Empty A]), P (LREmpty_ l NA)) -> - + (forall (Γ : context) (A : term) (ΠA : ParamRedTy@{i j k l} tSig Γ l A), PolyHypTyUr P Γ ΠA (P (LRSig' ΠA))) -> - - (forall Γ A (IA : [Γ ||-Id A]), + + (forall Γ A (IA : [Γ ||-Id A]), P (IA.(IdRedTy.tyRed)) -> (forall Δ (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), P (IA.(IdRedTy.tyKripke) ρ wfΔ)) -> P (LRId' IA)) -> @@ -228,10 +228,10 @@ same. Both need to be proven simultaneously, because of contravariance in the pr generalize l Γ A lr. eapply LR_rect_TyUr; intros [] *; constructor + auto. exfalso. pose (x := URedTy.lt h). inversion x. - Defined. + Defined. (* Induction principle with inductive hypothesis for the universe at lower levels *) - + Theorem LR_rect_TyUrGen@{i j k l o} (P : forall {l Γ A}, [LogRel@{i j k l} l | Γ ||- A] -> Type@{o}) : @@ -247,12 +247,12 @@ same. Both need to be proven simultaneously, because of contravariance in the pr (forall l Γ A (NA : [Γ ||-Nat A]), P (LRNat_ l NA)) -> (forall l Γ A (NA : [Γ ||-Empty A]), P (LREmpty_ l NA)) -> - + (forall (l : TypeLevel) (Γ : context) (A : term) (ΠA : ParamRedTy@{i j k l} tSig Γ l A), PolyHypTyUr P Γ ΠA (P (LRSig' ΠA))) -> - + (forall l Γ A (IA : [Γ ||-Id A]), - P (IA.(IdRedTy.tyRed)) -> + P (IA.(IdRedTy.tyRed)) -> (forall Δ (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), P (IA.(IdRedTy.tyKripke) ρ wfΔ)) -> P (LRId' IA)) -> @@ -269,12 +269,8 @@ End Inductions. (** ** Inversion principles *) Section Inversions. - Context `{ta : tag} - `{!WfContext ta} `{!WfType ta} `{!Typing ta} - `{!ConvType ta} `{!ConvTerm ta} `{!ConvNeuConv ta} - `{!RedType ta} `{!RedTerm ta} `{!RedTypeProperties} - `{!ConvNeuProperties}. - + Context `{GenericTypingProperties}. + Definition invLRTy {Γ l A A'} (lr : [Γ ||- A]) (r : [A ⤳* A']) (w : isType A') := match w return Type with | UnivType => [Γ ||-U A] @@ -325,14 +321,14 @@ Section Inversions. eapply redty_red, redA. - intros ??? [redA] ???. enough (A' = tNat) as ->. - + dependent inversion w. + + dependent inversion w. 1: now econstructor. inv_whne. + eapply whred_det; tea. all: gen_typing. - intros ??? [redA] ???. enough (A' = tEmpty) as ->. - + dependent inversion w. + + dependent inversion w. 1: now econstructor. inv_whne. + eapply whred_det; tea. @@ -372,7 +368,7 @@ Section Inversions. intros. now unshelve eapply (invLR _ redIdAlg ProdType). Qed. - + Lemma invLRΣ {Γ l dom cod} : [Γ ||- tSig dom cod] -> [Γ ||-Σ tSig dom cod]. Proof. intros. @@ -387,4 +383,40 @@ Section Inversions. (* invLRNat is useless *) + Lemma invLRConvU {Γ l A} : [Γ ||-U A] -> [Γ |- A ≅ U]. + Proof. intros []; gen_typing. Qed. + + Lemma invLRConvNe {Γ A} (RA : [Γ ||-ne A]) : [Γ |- A ≅ RA.(neRedTy.ty)]. + Proof. + destruct RA; cbn in *. + eapply convty_exp. + 2: apply redtywf_refl; gen_typing. + 1: gen_typing. + apply convty_term; apply convtm_convneu. + all: gen_typing. + Qed. + + Lemma invLRConvPi {Γ l A} (RA : [Γ ||-Π A]) : [Γ |- A ≅ RA.(PiRedTy.outTy)]. + Proof. + destruct RA as [?? []]. cbn in *. + eapply convty_exp; tea. + now apply redtywf_refl. + Qed. + + Lemma invLRConvSig {Γ l A} (RA : [Γ ||-Σ A]) : [Γ |- A ≅ RA.(SigRedTy.outTy)]. + Proof. + destruct RA as [?? []]. cbn in *. + eapply convty_exp; tea. + now apply redtywf_refl. + Qed. + + Lemma invLRConvNat {Γ A} (RA : [Γ ||-Nat A]) : [Γ |- A ≅ tNat]. + Proof. destruct RA; gen_typing. Qed. + + Lemma invLRConvEmpty {Γ A} (RA : [Γ ||-Empty A]) : [Γ |- A ≅ tEmpty]. + Proof. destruct RA; gen_typing. Qed. + + Lemma invLRConvId {Γ l A} (RA: [Γ ||-Id A]) : [Γ |- A ≅ RA.(IdRedTy.outTy)]. + Proof. destruct RA; unfold IdRedTy.outTy; cbn; gen_typing. Qed. + End Inversions. diff --git a/theories/LogicalRelation/InstKripke.v b/theories/LogicalRelation/InstKripke.v new file mode 100644 index 0000000..d823af1 --- /dev/null +++ b/theories/LogicalRelation/InstKripke.v @@ -0,0 +1,130 @@ +(** * LogRel.LogicalRelation.InstKripke: combinators to instantiate Kripke-style quantifications *) +From LogRel.AutoSubst Require Import core unscoped Ast Extra. +From LogRel Require Import Utils BasicAst Notations Context NormalForms Weakening + GenericTyping LogicalRelation. +From LogRel.LogicalRelation Require Import Induction Escape Reflexivity Neutral Weakening Irrelevance Application Reduction Transitivity NormalRed EqRedRight. + +Set Universe Polymorphism. +Set Printing Primitive Projection Parameters. + + +Lemma instKripkeEq `{GenericTypingProperties} {Γ A B l} (wfΓ : [|-Γ]) + {h : forall Δ (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), [Δ ||- A⟨ρ⟩]} + (eq : forall Δ (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), [h Δ ρ wfΔ | Δ ||- A⟨ρ⟩ ≅ B⟨ρ⟩]) + : [instKripke wfΓ h | Γ ||- A ≅ B]. +Proof. + specialize (eq Γ wk_id wfΓ); rewrite wk_id_ren_on in eq. + irrelevance. +Qed. + +Lemma instKripkeTmEq `{GenericTypingProperties} {Γ A t u l} (wfΓ : [|-Γ]) + {h : forall Δ (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), [Δ ||- A⟨ρ⟩]} + (eq : forall Δ (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), [h Δ ρ wfΔ | Δ ||- t⟨ρ⟩ ≅ u⟨ρ⟩ : _]) + : [instKripke wfΓ h | Γ ||- t ≅ u : _]. +Proof. + specialize (eq Γ wk_id wfΓ); rewrite !wk_id_ren_on in eq. + irrelevance. +Qed. + +Lemma instKripkeSubst `{GenericTypingProperties} {Γ A B l} (wfΓ : [|-Γ]) + {hA : forall Δ (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), [Δ ||- A⟨ρ⟩]} + (hB : forall Δ a b (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]) + (hab : [hA Δ ρ wfΔ | Δ ||- a ≅ b : _]), + [Δ ||- B[a .: ρ >> tRel]]) + : [ Γ ,, A ||- B]. +Proof. + pose proof (instKripke wfΓ hA). + escape. assert (wfΓA : [|- Γ ,, A]) by gen_typing. + unshelve epose proof (hinst := hB (Γ ,, A) (tRel 0) (tRel 0) (@wk1 Γ A) wfΓA _). + 1: eapply var0; tea; now bsimpl. + now rewrite <- eq_id_subst_scons in hinst. +Qed. + +Lemma instKripkeSubstEq `{GenericTypingProperties} {Γ A B B' l} (wfΓ : [|-Γ]) + {hA : forall Δ (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), [Δ ||- A⟨ρ⟩]} + {hB : forall Δ a b (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]) + (hab : [hA Δ ρ wfΔ | Δ ||- a ≅ b : _]), + [Δ ||- B[a .: ρ >> tRel]]} + (eq : forall Δ a b (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]) + (hab : [hA Δ ρ wfΔ | Δ ||- a ≅ b : _]), + [hB Δ a b ρ wfΔ hab | Δ ||- B[a .: ρ >> tRel] ≅ B'[b .: ρ >> tRel]]) + : [ instKripkeSubst wfΓ hB | Γ ,, A ||- _ ≅ B']. +Proof. + pose proof (instKripke wfΓ hA). + escape. assert (wfΓA : [|- Γ ,, A]) by gen_typing. + unshelve epose proof (hinst := eq (Γ ,, A) (tRel 0) (tRel 0) (@wk1 Γ A) wfΓA _). + 1: eapply var0; tea; now bsimpl. + rewrite <- eq_id_subst_scons in hinst. + irrelevance; now rewrite <- eq_id_subst_scons. +Qed. + +Lemma instKripkeSubstTmEq `{GenericTypingProperties} {Γ A B t u l} (wfΓ : [|-Γ]) + {hA : forall Δ (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), [Δ ||- A⟨ρ⟩]} + {hB : forall Δ a b (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]) + (hab : [hA Δ ρ wfΔ | Δ ||- a ≅ b : _]), + [Δ ||- B[a .: ρ >> tRel]]} + (eq : forall Δ a b (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]) + (hab : [hA Δ ρ wfΔ | Δ ||- a ≅ b : _]), + [hB Δ a b ρ wfΔ hab | Δ ||- t[a .: ρ >> tRel] ≅ u[b .: ρ >> tRel] : _]) + : [ instKripkeSubst wfΓ hB | Γ ,, A ||- t ≅ u : _]. +Proof. + pose proof (instKripke wfΓ hA). + escape. assert (wfΓA : [|- Γ ,, A]) by gen_typing. + unshelve epose proof (hinst := eq (Γ ,, A) (tRel 0) (tRel 0) (@wk1 Γ A) wfΓA _). + 1: eapply var0; tea; now bsimpl. + rewrite <- ! eq_id_subst_scons in hinst. + irrelevance; now rewrite <- eq_id_subst_scons. +Qed. + +Lemma instKripkeSubstConv `{GenericTypingProperties} {Γ A A' B l} (wfΓ : [|-Γ]) + {hA : forall Δ (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), [Δ ||- A⟨ρ⟩]} + (eq : forall Δ (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), [hA Δ ρ wfΔ | Δ ||- A⟨ρ⟩ ≅ A'⟨ρ⟩]) + (hB : forall Δ a b (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]) + (hab : [hA Δ ρ wfΔ | Δ ||- a ≅ b : _]), + [Δ ||- B[a .: ρ >> tRel]]) + : [ Γ ,, A' ||- B]. +Proof. + pose proof (instKripkeEq wfΓ eq). + escape. assert (wfΓA' : [|- Γ ,, A']) by gen_typing. + unshelve epose proof (hinst := hB (Γ ,, A') (tRel 0) (tRel 0) (@wk1 Γ A') wfΓA' _). + 1: now unshelve (eapply var0conv; tea; symmetry; erewrite <- wk1_ren_on; now eapply escapeEq, wkEq). + now rewrite <- eq_id_subst_scons in hinst. +Qed. + +Lemma instKripkeSubstConvEq `{GenericTypingProperties} {Γ A A' B B' l} (wfΓ : [|-Γ]) + {hA : forall Δ (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), [Δ ||- A⟨ρ⟩]} + (eqA : forall Δ (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), [hA Δ ρ wfΔ | Δ ||- A⟨ρ⟩ ≅ A'⟨ρ⟩]) + {hB : forall Δ a b (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]) + (hab : [hA Δ ρ wfΔ | Δ ||- a ≅ b : _]), + [Δ ||- B[a .: ρ >> tRel]]} + (eq : forall Δ a b (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]) + (hab : [hA Δ ρ wfΔ | Δ ||- a ≅ b : _]), + [hB Δ a b ρ wfΔ hab | Δ ||- B[a .: ρ >> tRel] ≅ B'[b .: ρ >> tRel]]) + : [ instKripkeSubstConv wfΓ eqA hB | Γ ,, A' ||- _ ≅ B']. +Proof. + pose proof (instKripkeEq wfΓ eqA). + escape. assert (wfΓA' : [|- Γ ,, A']) by gen_typing. + unshelve epose proof (hinst := eq (Γ ,, A') (tRel 0) (tRel 0) (@wk1 Γ A') wfΓA' _). + 1: now unshelve (eapply var0conv; tea; symmetry; erewrite <- wk1_ren_on; now eapply escapeEq, wkEq). + rewrite <- eq_id_subst_scons in hinst. + irrelevance; now rewrite <- eq_id_subst_scons. +Qed. + +Lemma instKripkeSubstConvTmEq `{GenericTypingProperties} {Γ A A' B t u l} (wfΓ : [|-Γ]) + {hA : forall Δ (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), [Δ ||- A⟨ρ⟩]} + (eqA : forall Δ (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), [hA Δ ρ wfΔ | Δ ||- A⟨ρ⟩ ≅ A'⟨ρ⟩]) + {hB : forall Δ a b (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]) + (hab : [hA Δ ρ wfΔ | Δ ||- a ≅ b : _]), + [Δ ||- B[a .: ρ >> tRel]]} + (eq : forall Δ a b (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]) + (hab : [hA Δ ρ wfΔ | Δ ||- a ≅ b : _]), + [hB Δ a b ρ wfΔ hab | Δ ||- t[a .: ρ >> tRel] ≅ u[b .: ρ >> tRel] : _]) + : [ instKripkeSubstConv wfΓ eqA hB | Γ ,, A' ||- t ≅ u : _]. +Proof. + pose proof (instKripkeEq wfΓ eqA). + escape. assert (wfΓA' : [|- Γ ,, A']) by gen_typing. + unshelve epose proof (hinst := eq (Γ ,, A') (tRel 0) (tRel 0) (@wk1 Γ A') wfΓA' _). + 1: now unshelve (eapply var0conv; tea; symmetry; erewrite <- wk1_ren_on; now eapply escapeEq, wkEq). + rewrite <- ! eq_id_subst_scons in hinst. + irrelevance; now rewrite <- eq_id_subst_scons. +Qed. diff --git a/theories/LogicalRelation/Irrelevance.v b/theories/LogicalRelation/Irrelevance.v index c4afbd2..4e6a578 100644 --- a/theories/LogicalRelation/Irrelevance.v +++ b/theories/LogicalRelation/Irrelevance.v @@ -29,19 +29,18 @@ Section EquivLRPack. Notation "A <≈> B" := (prod@{v v} (A -> B) (B -> A)) (at level 90). Definition equivLRPack {Γ Γ' A A'} (P: LRPack@{i} Γ A) (P': LRPack@{i'} Γ' A'):= - and3@{v v v} + prod@{v v} (forall B, [P | Γ ||- A ≅ B] <≈> [P' | Γ' ||- A' ≅ B]) - (forall t, [P | Γ ||- t : A] <≈> [P' | Γ' ||- t : A']) + (* (forall t, [P | Γ ||- t : A] <≈> [P' | Γ' ||- t : A']) *) (forall t u, [P | Γ ||- t ≅ u : A] <≈> [P' | Γ' ||- t ≅ u : A']). End EquivLRPack. Lemma symLRPack@{i i' v} {Γ Γ' A A'} {P: LRPack@{i} Γ A} {P': LRPack@{i'} Γ' A'} : equivLRPack@{i i' v} P P' -> equivLRPack@{i' i v} P' P. Proof. - intros [eqT rTm eqTm]; constructor;split ; - apply eqT + apply rTm + apply eqTm. + intros [eqT eqTm]; constructor;split ; apply eqT + apply eqTm. Qed. - + Record equivPolyRed@{i j k l i' j' k' l' v} {Γ l l' shp shp' pos pos'} @@ -50,10 +49,10 @@ Record equivPolyRed@{i j k l i' j' k' l' v} { eqvShp : forall {Δ} (ρ : Δ ≤ Γ) (wfΔ : [ |- Δ]), equivLRPack@{k k' v} (PolyRed.shpRed PA ρ wfΔ) (PolyRed.shpRed PA' ρ wfΔ) ; - eqvPos : forall {Δ a} (ρ : Δ ≤ Γ) (wfΔ : [ |- Δ]) - (ha : [PolyRed.shpRed PA ρ wfΔ| Δ ||- a : _]) - (ha' : [PolyRed.shpRed PA' ρ wfΔ | Δ ||- a : _]), - equivLRPack@{k k' v} + eqvPos : forall {Δ a b} (ρ : Δ ≤ Γ) (wfΔ : [ |- Δ]) + (ha : [PolyRed.shpRed PA ρ wfΔ| Δ ||- a ≅ b : _]) + (ha' : [PolyRed.shpRed PA' ρ wfΔ | Δ ||- a ≅ b : _]), + equivLRPack@{k k' v} (PolyRed.posRed PA ρ wfΔ ha) (PolyRed.posRed PA' ρ wfΔ ha') }. @@ -81,8 +80,8 @@ avoid having to basically duplicate their proofs. *) Section ΠIrrelevanceLemmas. Universe i j k l i' j' k' l' v. -Context {Γ lA A lA' A'} - (ΠA : ParamRedTy@{i j k l} tProd Γ lA A) +Context {Γ lA A lA' A'} + (ΠA : ParamRedTy@{i j k l} tProd Γ lA A) (ΠA' : ParamRedTy@{i' j' k' l'} tProd Γ lA' A') (RA := LRPi' ΠA) (RA' := LRPi' ΠA') @@ -102,7 +101,7 @@ Proof. now eapply eqv.(eqvShp). Qed. -Lemma ΠIrrelevanceTm t : [Γ ||- t : A | RA] -> [Γ ||- t : A' | RA']. +Lemma ΠIrrelevanceTm t : PiRedTm ΠA t -> PiRedTm ΠA' t. Proof. intros []; cbn in *; econstructor; tea. - now eapply redtmwf_conv. @@ -112,13 +111,9 @@ Proof. * intros; unshelve eapply eqv.(eqvPos); [|eauto]. now apply eqv.(eqvShp). + constructor; now eapply convneu_conv. - - eapply (convtm_conv refl). - apply eqPi. - - intros; unshelve eapply eqv.(eqvPos). + (* - unfold PiRedTmEq.appRed in *; intros; unshelve eapply eqv.(eqvPos). 2: now auto. - now apply eqv.(eqvShp). - - intros; unshelve eapply eqv.(eqvPos), eq. - all: now eapply eqv.(eqvShp). + now apply eqv.(eqvShp). *) Defined. Lemma ΠIrrelevanceTmEq t u : [Γ ||- t ≅ u : A | RA] -> [Γ ||- t ≅ u : A' | RA']. @@ -126,7 +121,7 @@ Proof. intros [] ; cbn in *; unshelve econstructor. 1,2: now eapply ΠIrrelevanceTm. - now eapply convtm_conv. - - intros; unshelve eapply eqv.(eqvPos). + - unfold PiRedTmEq.appRed in *; intros; unshelve eapply eqv.(eqvPos). 2: now auto. now apply eqv.(eqvShp). Qed. @@ -134,8 +129,8 @@ Qed. End ΠIrrelevanceLemmas. Lemma ΠIrrelevanceLRPack@{i j k l i' j' k' l' v} - {Γ lA A lA' A'} - (ΠA : ParamRedTy@{i j k l} tProd Γ lA A) + {Γ lA A lA' A'} + (ΠA : ParamRedTy@{i j k l} tProd Γ lA A) (ΠA' : ParamRedTy@{i' j' k' l'} tProd Γ lA' A') (RA := LRPi' ΠA) (RA' := LRPi' ΠA') @@ -147,7 +142,6 @@ Proof. pose proof (equivPolyRedSym eqv). constructor. - split; now apply ΠIrrelevanceTyEq. - - split; now apply ΠIrrelevanceTm. - split; now apply ΠIrrelevanceTmEq. Qed. @@ -156,8 +150,8 @@ Qed. Section ΣIrrelevanceLemmas. Universe i j k l i' j' k' l' v. -Context {Γ lA A lA' A'} - (ΣA : ParamRedTy@{i j k l} tSig Γ lA A) +Context {Γ lA A lA' A'} + (ΣA : ParamRedTy@{i j k l} tSig Γ lA A) (ΣA' : ParamRedTy@{i' j' k' l'} tSig Γ lA' A') (RA := LRSig' ΣA) (RA' := LRSig' ΣA') @@ -177,10 +171,9 @@ Proof. now eapply eqv.(eqvShp). Qed. -Lemma ΣIrrelevanceTm t : [Γ ||- t : A | RA] -> [Γ ||- t : A' | RA']. +Lemma ΣIrrelevanceTm t : SigRedTm ΣA t -> SigRedTm ΣA' t. Proof. intros []; cbn in *; unshelve econstructor; tea. - - intros; unshelve eapply eqv.(eqvShp); now auto. - now eapply redtmwf_conv. - destruct ispair as [A₀ B₀ a b|n Hn]. + unshelve econstructor. @@ -190,24 +183,22 @@ Proof. now unshelve eapply eqv.(eqvShp). * intros; now eapply eqv.(eqvPos). + constructor; now eapply convneu_conv. - - now eapply convtm_conv. - - intros; unshelve eapply eqv.(eqvPos); now auto. Defined. Lemma ΣIrrelevanceTmEq t u : [Γ ||- t ≅ u : A | RA] -> [Γ ||- t ≅ u : A' | RA']. Proof. intros [] ; cbn in *; unshelve econstructor. 1,2: now eapply ΣIrrelevanceTm. - - now eapply convtm_conv. - intros; unshelve eapply eqv.(eqvShp); auto. + - now eapply convtm_conv. - intros; unshelve eapply eqv.(eqvPos); auto. Qed. End ΣIrrelevanceLemmas. Lemma ΣIrrelevanceLRPack@{i j k l i' j' k' l' v} - {Γ lA A lA' A'} - (ΣA : ParamRedTy@{i j k l} tSig Γ lA A) + {Γ lA A lA' A'} + (ΣA : ParamRedTy@{i j k l} tSig Γ lA A) (ΣA' : ParamRedTy@{i' j' k' l'} tSig Γ lA' A') (RA := LRSig' ΣA) (RA' := LRSig' ΣA') @@ -219,28 +210,22 @@ Proof. pose proof (equivPolyRedSym eqv). constructor. - split; now apply ΣIrrelevanceTyEq. - - split; now apply ΣIrrelevanceTm. - split; now apply ΣIrrelevanceTmEq. Qed. (** *** Lemmas for conversion of reducible neutral terms at arbitrary types *) -Lemma NeNfconv {Γ k A A'} : [Γ |- A'] -> [Γ |- A ≅ A'] -> [Γ ||-NeNf k : A] -> [Γ ||-NeNf k : A']. -Proof. - intros ?? []; econstructor; tea. all: gen_typing. -Qed. - Lemma NeNfEqconv {Γ k k' A A'} : [Γ |- A'] -> [Γ |- A ≅ A'] -> [Γ ||-NeNf k ≅ k' : A] -> [Γ ||-NeNf k ≅ k' : A']. Proof. - intros ?? []; econstructor; tea. gen_typing. + intros ?? []; econstructor; tea; gen_typing. Qed. (** *** Irrelevance for Identity types *) Section IdIrrelevance. Universe i j k l i' j' k' l' v. - Context {Γ lA A lA' A'} - (IA : IdRedTy@{i j k l} Γ lA A) + Context {Γ lA A lA' A'} + (IA : IdRedTy@{i j k l} Γ lA A) (IA' : IdRedTy@{i' j' k' l'} Γ lA' A') (RA := LRId' IA) (RA' := LRId' IA') @@ -261,22 +246,6 @@ Section IdIrrelevance. - apply eqv; etransitivity; tea; now symmetry. - apply eqv; etransitivity; tea; now symmetry. Qed. - - Lemma IdIrrelevanceProp t : IdProp IA t -> IdProp IA' t. - Proof. - intros []; constructor; tea; cycle -1. - 1: eapply NeNfconv; tea; unfold_id_outTy ; destruct IA'; escape; cbn in *; gen_typing. - all: apply eqv; tea. - all: etransitivity; [now symmetry|]; tea. - Qed. - - Lemma IdIrrelevanceTm t : [Γ ||- t : A | RA] -> [Γ ||- t : A' | RA']. - Proof. - intros []; cbn in *; unshelve econstructor; unfold_id_outTy; tea. - - now eapply redtmwf_conv. - - now eapply convtm_conv. - - now eapply IdIrrelevanceProp. - Qed. Lemma IdIrrelevancePropEq t u : IdPropEq IA t u -> IdPropEq IA' t u. Proof. @@ -285,7 +254,7 @@ Section IdIrrelevance. all: apply eqv; tea. all: etransitivity; [now symmetry|]; tea. Qed. - + Lemma IdIrrelevanceTmEq t u : [Γ ||- t ≅ u : A | RA] -> [Γ ||- t ≅ u : A' | RA']. Proof. intros []; cbn in *; unshelve econstructor; unfold_id_outTy. @@ -293,12 +262,12 @@ Section IdIrrelevance. - now eapply convtm_conv. - now eapply IdIrrelevancePropEq. Qed. - + End IdIrrelevance. Lemma IdIrrelevanceLRPack@{i j k l i' j' k' l' v} - {Γ lA A lA' A'} - (IA : IdRedTy@{i j k l} Γ lA A) + {Γ lA A lA' A'} + (IA : IdRedTy@{i j k l} Γ lA A) (IA' : IdRedTy@{i' j' k' l'} Γ lA' A') (RA := LRId' IA) (RA' := LRId' IA') @@ -315,7 +284,6 @@ Proof. assert [IA'.(IdRedTy.tyRed) | Γ ||- IA'.(IdRedTy.rhs) ≅ IA.(IdRedTy.rhs) : _ ] by (apply eqv; now symmetry). constructor. - split; now apply IdIrrelevanceTyEq. - - split; now apply IdIrrelevanceTm. - split; now apply IdIrrelevanceTmEq. Qed. @@ -338,10 +306,7 @@ Proof. + intros ?; split; intros []; econstructor; cbn in *; tea. all: etransitivity ; [| tea]; tea; now symmetry. + intros ?; split; intros []; econstructor; cbn in *; tea. - 1,3: now eapply redtmwf_conv. - all: now eapply convneu_conv; first [eassumption|symmetry; eassumption|gen_typing]. - + intros ??; split; intros []; econstructor; cbn in *. - 1-2,4-5: now eapply redtmwf_conv. + 1,2,4,5: now eapply redtmwf_conv. all: now eapply convneu_conv; first [eassumption|symmetry; eassumption|gen_typing]. Qed. @@ -352,19 +317,12 @@ Section NatIrrelevant. Context {Γ lA lA' A A'} (NA : [Γ ||-Nat A]) (NA' : [Γ ||-Nat A']) (RA := LRNat_@{i j k l} lA NA) (RA' := LRNat_@{i' j' k' l'} lA' NA'). - + Lemma NatIrrelevanceTyEq B : [Γ ||- A ≅ B | RA] -> [Γ ||- A' ≅ B | RA']. Proof. intros []; now econstructor. Qed. - Lemma NatIrrelevanceTm : - (forall t, [Γ ||- t : A | RA] -> [Γ ||- t : A' | RA']) - × (forall t, NatProp NA t -> NatProp NA' t). - Proof. - apply NatRedInduction; now econstructor. - Qed. - Lemma NatIrrelevanceTmEq : (forall t u, [Γ ||- t ≅ u : A | RA] -> [Γ ||- t ≅ u : A' | RA']) × (forall t u, NatPropEq NA t u -> NatPropEq NA' t u). @@ -380,7 +338,6 @@ Lemma NatIrrelevanceLRPack@{i j k l i' j' k' l' v} Proof. constructor. - split; apply NatIrrelevanceTyEq. - - split; apply NatIrrelevanceTm. - split; apply NatIrrelevanceTmEq. Qed. @@ -389,18 +346,12 @@ Section EmptyIrrelevant. Context {Γ lA lA' A A'} (NA : [Γ ||-Empty A]) (NA' : [Γ ||-Empty A']) (RA := LREmpty_@{i j k l} lA NA) (RA' := LREmpty_@{i' j' k' l'} lA' NA'). - + Lemma EmptyIrrelevanceTyEq B : [Γ ||- A ≅ B | RA] -> [Γ ||- A' ≅ B | RA']. Proof. intros []; now econstructor. Qed. - Lemma EmptyIrrelevanceTm : - (forall t, [Γ ||- t : A | RA] -> [Γ ||- t : A' | RA']). - Proof. - intros t Ht. induction Ht; now econstructor. - Qed. - Lemma EmptyIrrelevanceTmEq : (forall t u, [Γ ||- t ≅ u : A | RA] -> [Γ ||- t ≅ u : A' | RA']). Proof. @@ -415,7 +366,6 @@ Lemma EmptyIrrelevanceLRPack@{i j k l i' j' k' l' v} Proof. constructor. - split; apply EmptyIrrelevanceTyEq. - - split; apply EmptyIrrelevanceTm. - split; apply EmptyIrrelevanceTmEq. Qed. @@ -457,11 +407,6 @@ Proof. constructor. + intros; cbn; split; intros []; now constructor. - + intros ?; destruct (IHty Γ t) as [tfwd tbwd]; split; intros []; - unshelve econstructor. - 6: apply tfwd; assumption. - 9: apply tbwd; assumption. - all : tea. + cbn ; intros ? ?; destruct (IHty Γ t) as [tfwd tbwd]; destruct (IHty Γ u) as [ufwd ubwd]. @@ -470,16 +415,11 @@ Proof. 5: apply tbwd; assumption. 6: apply ufwd; assumption. 8: apply ubwd; assumption. - (* all: apply todo. *) all: cbn. 6: refine (fst (IHeq _ _ _ _ _) _); eassumption. 7: refine (snd (IHeq _ _ _ _ _) _); eassumption. (* Regression here: now/eassumption adds universe constraints that we do not want to accept but can't prevent *) - 1-4:econstructor; cycle -1; [|tea..]. - 1: eapply tfwd; eassumption. - 1: eapply ufwd; eassumption. - 1: eapply tbwd; eassumption. - 1: eapply ubwd; eassumption. + 1-4:econstructor; cycle -1; [|tea..]; tea. all: cbn; tea. Qed. @@ -490,21 +430,21 @@ Qed. Lemma LRIrrelevantPreds {lA lA'} (IH : IHStatement lA lA') (Γ : context) (A A' : term) - {eqTyA redTmA : term -> Type@{k}} - {eqTyA' redTmA' : term -> Type@{k'}} + {eqTyA : term -> Type@{k}} + {eqTyA' : term -> Type@{k'}} {eqTmA : term -> term -> Type@{k}} {eqTmA' : term -> term -> Type@{k'}} - (lrA : LogRel@{i j k l} lA Γ A eqTyA redTmA eqTmA) - (lrA' : LogRel@{i' j' k' l'} lA' Γ A' eqTyA' redTmA' eqTmA') - (RA := Build_LRPack Γ A eqTyA redTmA eqTmA) - (RA' := Build_LRPack Γ A' eqTyA' redTmA' eqTmA') : + (lrA : LogRel@{i j k l} lA Γ A eqTyA eqTmA) + (lrA' : LogRel@{i' j' k' l'} lA' Γ A' eqTyA' eqTmA') + (RA := Build_LRPack Γ A eqTyA eqTmA) + (RA' := Build_LRPack Γ A' eqTyA' eqTmA') : eqTyA A' -> equivLRPack@{k k' v} RA RA'. Proof. intros he. set (s := ShapeViewConv lrA lrA' he). induction lrA as [? ? h1 | ? ? neA | ? A ΠA HAad IHdom IHcod | ?? NA | ?? NA|? A ΠA HAad IHdom IHcod | ?? IAP IAad IHPar] - in RA, A', RA', eqTyA', eqTmA', redTmA', lrA', he, s |- *. + in RA, A', RA', eqTyA', eqTmA', lrA', he, s |- *. - destruct lrA' ; try solve [destruct s] ; clear s. now apply UnivIrrelevanceLRPack. - destruct lrA' ; try solve [destruct s] ; clear s. @@ -522,7 +462,7 @@ Proof. + intros; unshelve eapply IHdom. 2: eapply (LRAd.adequate (PolyRed.shpRed PA' _ _)). eapply domRed. - + intros; unshelve eapply IHcod. + + intros. unshelve eapply IHcod. 2: eapply (LRAd.adequate (PolyRed.posRed PA' _ _ _)). eapply codRed. - destruct lrA' ; try solve [destruct s] ; clear s. @@ -563,35 +503,33 @@ Lemma LRIrrelevantCumPolyRed {lA} (Γ : context) (shp pos : term) (PA : PolyRed@{i j k l} Γ lA shp pos) (IHshp : forall (Δ : context) (ρ : Δ ≤ Γ), [ |-[ ta ] Δ] -> [Δ ||-< lA > shp⟨ρ⟩]) - (IHpos : forall (Δ : context) (a : term) (ρ : Δ ≤ Γ) (h : [ |-[ ta ] Δ]), - [PolyRed.shpRed PA ρ h | _ ||- a : _] -> + (IHpos : forall (Δ : context) (a b : term) (ρ : Δ ≤ Γ) (h : [ |-[ ta ] Δ]), + [PolyRed.shpRed PA ρ h | _ ||- a ≅ b : _] -> [Δ ||-< lA > pos[a .: ρ >> tRel]]) : PolyRed@{i' j' k' l'} Γ lA shp pos. Proof. unshelve econstructor. + exact IHshp. - + intros Δ a ρ tΔ ra. eapply IHpos. + + intros Δ a b ρ tΔ ra. eapply IHpos. pose (shpRed := PA.(PolyRed.shpRed) ρ tΔ). destruct (LRIrrelevantPreds IH _ _ _ (LRAd.adequate shpRed) (LRAd.adequate (IHshp Δ ρ tΔ)) - (reflLRTyEq shpRed)) as [_ irrTmRed _]. - now eapply (snd (irrTmRed a)). + (reflLRTyEq shpRed)) as [_ irrTmEq]. + now eapply (snd (irrTmEq a b)). + now destruct PA. + now destruct PA. - + cbn. intros Δ a b ρ tΔ ra rb rab. + + cbn. intros Δ a b ρ tΔ rab. set (p := LRIrrelevantPreds _ _ _ _ _ _ _). - destruct p as [_ irrTmRed irrTmEq]. - pose (ra' := snd (irrTmRed a) ra). + destruct p as [_ irrTmEq]. + pose (ra' := snd (irrTmEq a b) rab). pose (posRed := PA.(PolyRed.posRed) ρ tΔ ra'). destruct (LRIrrelevantPreds IH _ _ _ (LRAd.adequate posRed) - (LRAd.adequate (IHpos Δ a ρ tΔ ra')) - (reflLRTyEq posRed)) as [irrTyEq _ _]. + (LRAd.adequate (IHpos Δ a b ρ tΔ ra')) + (reflLRTyEq posRed)) as [irrTyEq _]. eapply (fst (irrTyEq (pos[b .: ρ >> tRel]))). eapply PolyRed.posExt. - 1: exact (snd (irrTmRed b) rb). - exact (snd (irrTmEq a b) rab). Qed. @@ -620,7 +558,7 @@ Proof. unshelve eapply LRIrrelevantCumPolyRed; tea. + intros; now eapply IHdom. + intros; now eapply IHcod. - - intros [] IHPar IHKripke IH. + - intros [] IHPar IHKripke IH. specialize (IHPar IH). pose (IHK Δ ρ wfΔ := IHKripke Δ ρ wfΔ IH). cbn in *; eapply LRId'. assert (eqv: equivLRPack tyRed IHPar). @@ -629,8 +567,8 @@ Proof. 1: intros; eapply LRIrrelevantPreds; tea; try eapply reflLRTyEq; now eapply LRAd.adequate. unshelve econstructor. 4-7: tea. - 1-4: now apply eqv. - 2-4: intros * ? ?%eqvK; apply eqvK; eauto. + 1-2: now apply eqv. + 2-3: intros * ? ?%eqvK; apply eqvK; eauto. econstructor. + intros ?? ?%eqv; apply eqv; now symmetry. + intros ??? ?%eqv ?%eqv; apply eqv; now etransitivity. @@ -669,47 +607,53 @@ Proof. destruct (LRIrrelevantPreds@{u i j k u i' j' k'} IrrRec0@{u i j k u i' j' k'} Γ t t (lr1 : LRPackAdequate (LogRel@{u i j k} zero) lr1) (lr2 : LRPackAdequate (LogRel@{u i' j' k'} zero) lr2) - (reflLRTyEq lr1)) as [tyEq _ _]. + (reflLRTyEq lr1)) as [tyEq _]. exact (tyEq u). Qed. #[local] -Theorem LRIrrelevantCum@{i j k l i' j' k' l'} +Theorem IrrelevantCum@{i j k l i' j' k' l'} (Γ : context) (A A' : term) {lA lA'} - {eqTyA redTmA : term -> Type@{k}} - {eqTyA' redTmA' : term -> Type@{k'}} + {eqTyA : term -> Type@{k}} + {eqTyA' : term -> Type@{k'}} {eqTmA : term -> term -> Type@{k}} {eqTmA' : term -> term -> Type@{k'}} - (lrA : LogRel@{i j k l} lA Γ A eqTyA redTmA eqTmA) - (lrA' : LogRel@{i' j' k' l'} lA' Γ A' eqTyA' redTmA' eqTmA') : + (lrA : LogRel@{i j k l} lA Γ A eqTyA eqTmA) + (lrA' : LogRel@{i' j' k' l'} lA' Γ A' eqTyA' eqTmA') : eqTyA A' -> - @and3@{v v v} (forall B, eqTyA B <≈> eqTyA' B) - (forall t, redTmA t <≈> redTmA' t) + @prod@{v v} (forall B, eqTyA B <≈> eqTyA' B) (forall t u, eqTmA t u <≈> eqTmA' t u). Proof. exact (LRIrrelevantPreds@{i j k l i' j' k' l'} IrrRec Γ A A' lrA lrA'). Qed. -Theorem LRIrrelevantPack@{i j k l} - (Γ : context) (A A' : term) {lA lA'} +Theorem LRIrrelevantCum@{i j k l i' j' k' l'} + {Γ : context} {A A' : term} {lA lA'} (RA : [ LogRel@{i j k l} lA | Γ ||- A ]) - (RA' : [ LogRel@{i j k l} lA' | Γ ||- A' ]) + (RA' : [ LogRel@{i' j' k' l'} lA' | Γ ||- A' ]) (RAA' : [Γ ||- A ≅ A' | RA]) : equivLRPack@{v v v} RA RA'. Proof. - pose proof (LRIrrelevantCum@{i j k l i j k l} Γ A A' (LRAd.adequate RA) (LRAd.adequate RA') RAA') as []. - constructor; eauto. -Defined. + exact (IrrelevantCum _ _ _ RA.(LRAd.adequate) RA'.(LRAd.adequate) RAA'). +Qed. + +Theorem LRIrrelevantPack@{i j k l} + (Γ : context) (A A' : term) {lA lA'} + (RA : [ LogRel@{i j k l} lA | Γ ||- A ]) + (RA' : [ LogRel@{i j k l} lA' | Γ ||- A' ]) + (RAA' : [Γ ||- A ≅ A' | RA]) : + equivLRPack@{v v v} RA RA'. +Proof. now apply LRIrrelevantCum. Qed. -Theorem LRTransEq@{i j k l} - (Γ : context) (A B C : term) {lA lB} +Theorem LRTransEq@{i j k l} + (Γ : context) (A B C : term) {lA lB} (RA : [ LogRel@{i j k l} lA | Γ ||- A ]) (RB : [ LogRel@{i j k l} lB | Γ ||- B ]) (RAB : [Γ ||- A ≅ B | RA]) (RBC : [Γ ||- B ≅ C | RB]) : [Γ ||- A ≅ C | RA]. Proof. - pose proof (LRIrrelevantPack Γ A B RA RB RAB) as [h _ _]. + pose proof (LRIrrelevantPack Γ A B RA RB RAB) as [h _]. now apply h. Defined. @@ -730,22 +674,11 @@ Qed. End LRIrrelevant. -#[local] -Corollary TyEqIrrelevantCum Γ A {lA eqTyA redTmA eqTmA lA' eqTyA' redTmA' eqTmA'} - (lrA : LogRel lA Γ A eqTyA redTmA eqTmA) (lrA' : LogRel lA' Γ A eqTyA' redTmA' eqTmA') : - forall B, eqTyA B -> eqTyA' B. -Proof. - apply (LRIrrelevantCum _ _ _ lrA lrA'). - now eapply LRTyEqRefl. -Qed. - Corollary LRTyEqIrrelevantCum@{i j k l i' j' k' l'} lA lA' Γ A (lrA : [LogRel@{i j k l} lA | Γ ||- A]) (lrA' : [LogRel@{i' j' k' l'} lA' | Γ ||- A]) : forall B, [Γ ||-< lA > A ≅ B | lrA] -> [Γ ||-< lA' > A ≅ B | lrA']. Proof. - destruct lrA, lrA'. - cbn in *. - now eapply TyEqIrrelevantCum. + apply (LRIrrelevantCum lrA lrA'), reflLRTyEq. Qed. Corollary LRTyEqIrrelevantCum'@{i j k l i' j' k' l'} lA lA' Γ A A' (e : A = A') @@ -762,55 +695,11 @@ Proof. revert lrA'; rewrite <- e; now apply LRTyEqIrrelevantCum. Qed. -#[local] -Corollary RedTmIrrelevantCum Γ A {lA eqTyA redTmA eqTmA lA' eqTyA' redTmA' eqTmA'} - (lrA : LogRel lA Γ A eqTyA redTmA eqTmA) (lrA' : LogRel lA' Γ A eqTyA' redTmA' eqTmA') : - forall t, redTmA t -> redTmA' t. -Proof. - apply (LRIrrelevantCum _ _ _ lrA lrA'). - now eapply LRTyEqRefl. -Qed. - -Corollary LRTmRedIrrelevantCum@{i j k l i' j' k' l'} lA lA' Γ A - (lrA : [LogRel@{i j k l} lA | Γ ||- A]) (lrA' : [LogRel@{i' j' k' l'} lA' | Γ ||- A]) : - forall t, [Γ ||-< lA > t : A | lrA] -> [Γ ||-< lA' > t : A | lrA']. -Proof. - destruct lrA, lrA'. - cbn in *. - now eapply RedTmIrrelevantCum. -Qed. - -Corollary LRTmRedIrrelevantCum'@{i j k l i' j' k' l'} lA lA' Γ A A' (e : A = A') - (lrA : [LogRel@{i j k l} lA | Γ ||- A]) (lrA' : [LogRel@{i' j' k' l'} lA' | Γ ||- A']) : - forall t, [Γ ||-< lA > t : A | lrA] -> [Γ ||-< lA' > t : A' | lrA']. -Proof. - revert lrA'; rewrite <- e; now apply LRTmRedIrrelevantCum. -Qed. - -Corollary LRTmRedIrrelevant'@{i j k l} lA lA' Γ A A' (e : A = A') - (lrA : [LogRel@{i j k l} lA | Γ ||- A]) (lrA' : [LogRel@{i j k l} lA' | Γ ||- A']) : - forall t, [Γ ||-< lA > t : A | lrA] -> [Γ ||-< lA' > t : A' | lrA']. -Proof. - revert lrA'; rewrite <- e; now apply LRTmRedIrrelevantCum. -Qed. - - -#[local] -Corollary TmEqIrrelevantCum Γ A {lA eqTyA redTmA eqTmA lA' eqTyA' redTmA' eqTmA'} - (lrA : LogRel lA Γ A eqTyA redTmA eqTmA) (lrA' : LogRel lA' Γ A eqTyA' redTmA' eqTmA') : - forall t u, eqTmA t u -> eqTmA' t u. -Proof. - apply (LRIrrelevantCum _ _ _ lrA lrA'). - now eapply LRTyEqRefl. -Qed. - Corollary LRTmEqIrrelevantCum@{i j k l i' j' k' l'} lA lA' Γ A (lrA : [LogRel@{i j k l} lA | Γ ||- A]) (lrA' : [LogRel@{i' j' k' l'} lA' | Γ ||- A]) : forall t u, [Γ ||-< lA > t ≅ u : A | lrA] -> [Γ ||-< lA' > t ≅ u : A | lrA']. Proof. - destruct lrA, lrA'. - cbn in *. - now eapply TmEqIrrelevantCum. + apply (LRIrrelevantCum lrA lrA'), reflLRTyEq. Qed. Corollary LRTmEqIrrelevantCum'@{i j k l i' j' k' l'} lA lA' Γ A A' (e : A = A') @@ -827,42 +716,18 @@ Proof. revert lrA'; rewrite <- e; now apply LRTmEqIrrelevantCum. Qed. - - -Corollary TyEqSym Γ A A' {lA eqTyA redTmA eqTmA lA' eqTyA' redTmA' eqTmA'} - (lrA : LogRel lA Γ A eqTyA redTmA eqTmA) (lrA' : LogRel lA' Γ A' eqTyA' redTmA' eqTmA') : - eqTyA A' -> eqTyA' A. -Proof. - intros. - apply (LRIrrelevantCum _ _ _ lrA lrA'). - 1: eauto. - now eapply LRTyEqRefl. -Qed. - Corollary LRTyEqSym lA lA' Γ A A' (lrA : [Γ ||-< lA > A]) (lrA' : [Γ ||-< lA'> A']) : [Γ ||-< lA > A ≅ A' | lrA] -> [Γ ||-< lA' > A' ≅ A | lrA']. Proof. - destruct lrA, lrA'. - cbn in *. - now eapply TyEqSym. + intros; eapply (LRIrrelevantCum lrA lrA'); tea. + now eapply reflLRTyEq. Qed. -#[local] -Corollary RedTmConv Γ A A' {lA eqTyA redTmA eqTmA lA' eqTyA' redTmA' eqTmA'} - (lrA : LogRel lA Γ A eqTyA redTmA eqTmA) (lrA' : LogRel lA' Γ A' eqTyA' redTmA' eqTmA') : - eqTyA A' -> - forall t, redTmA t -> redTmA' t. -Proof. - apply (LRIrrelevantCum _ _ _ lrA lrA'). -Qed. - -Corollary LRTmRedConv lA lA' Γ A A' (lrA : [Γ ||-< lA > A]) (lrA' : [Γ ||-< lA'> A' ]) : +Corollary LRTmEqConv lA lA' Γ A A' (lrA : [Γ ||-< lA > A]) (lrA' : [Γ ||-< lA'> A' ]) : [Γ ||-< lA > A ≅ A' | lrA ] -> - forall t, [Γ ||-< lA > t : A | lrA] -> [Γ ||-< lA' > t : A' | lrA']. + forall t u, [Γ ||-< lA > t ≅ u : A | lrA] -> [Γ ||-< lA' > t ≅ u: A' | lrA']. Proof. - destruct lrA, lrA'. - cbn in *. - now eapply RedTmConv. + intros; now apply (LRIrrelevantCum lrA lrA'). Qed. Corollary PolyRedEqSym {Γ l l' shp shp' pos pos'} @@ -873,37 +738,16 @@ Proof. intros []; unshelve econstructor. - intros; eapply LRTyEqSym; eauto. - intros. eapply LRTyEqSym. unshelve eapply posRed; tea. - eapply LRTmRedConv; tea. + eapply LRTmEqConv; tea. now eapply LRTyEqSym. Unshelve. all: tea. Qed. -#[local] -Corollary TmEqRedConv Γ A A' {lA eqTyA redTmA eqTmA lA' eqTyA' redTmA' eqTmA'} - (lrA : LogRel lA Γ A eqTyA redTmA eqTmA) (lrA' : LogRel lA' Γ A' eqTyA' redTmA' eqTmA') : - eqTyA A' -> - forall t u, eqTmA t u -> eqTmA' t u. -Proof. - apply (LRIrrelevantCum _ _ _ lrA lrA'). -Qed. - +#[deprecated(note="Use LRTmEqConv")] Corollary LRTmEqRedConv lA lA' Γ A A' (lrA : [Γ ||-< lA > A]) (lrA' : [Γ ||-< lA'> A']) : [Γ ||-< lA > A ≅ A' | lrA ] -> forall t u, [Γ ||-< lA > t ≅ u : A | lrA] -> [Γ ||-< lA' > t ≅ u : A' | lrA']. -Proof. - destruct lrA, lrA'. - cbn in *. - now eapply TmEqRedConv. -Qed. - -Corollary LRTmTmEqIrrelevant' lA lA' Γ A A' (e : A = A') - (lrA : [Γ ||-< lA > A]) (lrA' : [Γ ||-< lA'> A']) : - forall t u, - [Γ ||- t : A | lrA] × [Γ ||-< lA > t ≅ u : A | lrA] -> - [Γ ||- t : A' | lrA'] × [Γ ||-< lA' > t ≅ u : A' | lrA']. -Proof. - intros ?? []; split; [eapply LRTmRedIrrelevant'| eapply LRTmEqIrrelevant']; tea. -Qed. +Proof. apply LRTmEqConv. Qed. Set Printing Primitive Projection Parameters. @@ -918,16 +762,19 @@ Proof. pattern lA, Γ, A, lrA. apply LR_rect_TyUr; clear lA Γ A lrA. - intros * []. unshelve econstructor; try eassumption. 1: symmetry; eassumption. - (* Need an additional universe level h < i *) - eapply TyEqSym@{h i j k h i j k}. 3:exact relEq. - all: eapply LogRelRec_unfold; eapply LRAd.adequate; eassumption. + now eapply TyEqRecFwd, LRTyEqSym@{h i j k h i j k}, TyEqRecFwd. - intros * []. unshelve econstructor. 3,4: eassumption. symmetry; eassumption. - intros * ihdom ihcod * []. unshelve econstructor. 1,2: eassumption. 1: symmetry; eassumption. - intros. apply ihcod. eapply eqApp. + unfold PiRedTmEq.appRed in *. intros. + assert (hba := ihdom _ _ _ _ _ hab). + apply ihcod. + eapply LRTmEqConv. + 2: eapply eqApp with (hab:=hba). + eapply PolyRed.posExt. - intros ??? NA. set (G := _); enough (h : G × (forall t u, NatPropEq NA t u -> NatPropEq NA u t)) by apply h. subst G; apply NatRedEqInduction. @@ -938,12 +785,13 @@ Proof. 2: econstructor; now eapply NeNfEqSym. symmetry; eassumption. - intros * ihshp ihpos * []; unshelve econstructor; tea. - 1: now symmetry. + intros; now eapply ihshp. + + now symmetry. + intros; eapply ihpos. - eapply LRTmEqRedConv. + eapply LRTmEqConv. 2: eapply eqSnd. now eapply PolyRed.posExt. + Unshelve. eassumption. - intros ??? [] ???? [????? hprop]; unshelve econstructor; unfold_id_outTy; cbn in *. 3,4: tea. 1: now symmetry. @@ -951,9 +799,23 @@ Proof. now eapply NeNfEqSym. Qed. + End Irrelevances. +(* Could it be useful to redefine reducible convertibility independently from + a witness of type reducibility ? *) +Definition LRTyConv `{GenericTypingProperties} Γ l (A B : term) := ∑ (C : term) (RC : [Γ ||- C]), [Γ ||- _ ≅ A | RC] × [Γ ||- _ ≅ B | RC]. + +Notation "[ Γ ||-< l > A ≅ B ]" := (LRTyConv Γ l A B). + +#[global] +Instance LRTmEqSymmetric `{GenericTypingProperties} {Γ A l} (RA : [Γ ||- A]): Symmetric (RA.(LRPack.eqTm)). +Proof. intros x y; apply LRTmEqSym. Defined. + +Ltac sym_escape RA H ::= + let X := fresh "Rr" H in pose proof (X := escapeTerm RA (LRTmEqSym _ _ _ RA _ _ H)). + (** ** Tactics for irrelevance, with and without universe cumulativity *) Ltac irrelevanceCum0 := @@ -961,8 +823,8 @@ Ltac irrelevanceCum0 := | [|- [_ ||-<_> _]] => (now eapply LRCumulative) + eapply LRCumulative' | [|- [_ | _ ||- _ ≅ _ ] ] => eapply LRTyEqIrrelevantCum' | [|- [_ ||-<_> _ ≅ _ | _ ] ] => eapply LRTyEqIrrelevantCum' - | [|- [_ | _ ||- _ : _ ] ] => eapply LRTmRedIrrelevantCum' - | [|- [_ ||-<_> _ : _ | _ ] ] => eapply LRTmRedIrrelevantCum' + (* | [|- [_ | _ ||- _ : _ ] ] => eapply LRTmRedIrrelevantCum' *) + (* | [|- [_ ||-<_> _ : _ | _ ] ] => eapply LRTmRedIrrelevantCum' *) | [|- [_ | _ ||- _ ≅ _ : _ ] ] => eapply LRTmEqIrrelevantCum' | [|- [_ ||-<_> _ ≅ _ : _ | _ ] ] => eapply LRTmEqIrrelevantCum' end. @@ -975,11 +837,11 @@ Ltac irrelevance0 := lazymatch goal with | [|- [_ | _ ||- _ ≅ _ ] ] => eapply LRTyEqIrrelevant' | [|- [_ ||-<_> _ ≅ _ | _ ] ] => eapply LRTyEqIrrelevant' - | [|- [_ | _ ||- _ : _ ] ] => eapply LRTmRedIrrelevant' - | [|- [_ ||-<_> _ : _ | _ ] ] => eapply LRTmRedIrrelevant' + (* | [|- [_ | _ ||- _ : _ ] ] => eapply LRTmRedIrrelevant' *) + (* | [|- [_ ||-<_> _ : _ | _ ] ] => eapply LRTmRedIrrelevant' *) | [|- [_ | _ ||- _ ≅ _ : _ ] ] => eapply LRTmEqIrrelevant' | [|- [_ ||-<_> _ ≅ _ : _ | _ ] ] => eapply LRTmEqIrrelevant' - | [|- [_ ||-<_> _ : _ | _] × [_ ||-<_> _≅ _ : _ | _]] => eapply LRTmTmEqIrrelevant' + (* | [|- [_ ||-<_> _ : _ | _] × [_ ||-<_> _≅ _ : _ | _]] => eapply LRTmTmEqIrrelevant' *) end. Ltac irrelevance := irrelevance0 ; [|eassumption] ; try first [reflexivity| now bsimpl]. diff --git a/theories/LogicalRelation/Neutral.v b/theories/LogicalRelation/Neutral.v index d80036d..97995d9 100644 --- a/theories/LogicalRelation/Neutral.v +++ b/theories/LogicalRelation/Neutral.v @@ -1,6 +1,8 @@ +(** * LogRel.Neutral: Reducibility of neutral types and terms. *) +From Coq Require Import CRelationClasses. From LogRel.AutoSubst Require Import core unscoped Ast Extra. From LogRel Require Import Utils BasicAst Notations Context NormalForms UntypedReduction Weakening GenericTyping LogicalRelation. -From LogRel.LogicalRelation Require Import Induction Reflexivity Irrelevance Escape. +From LogRel.LogicalRelation Require Import Induction Reflexivity Escape Irrelevance Transitivity. Set Universe Polymorphism. @@ -15,17 +17,12 @@ Defined. Lemma neU {l Γ A n} (h : [Γ ||-U A]) : [Γ |- n : A] -> - [Γ |- n ~ n : A] -> - [LogRelRec l | Γ ||-U n : A | h]. + [Γ |- n ~ n : A] -> URedTm (LogRelRec l) Γ n A h. Proof. assert [Γ |- A ≅ U] by (destruct h; gen_typing). - intros; exists n. + intros; unshelve eexists n. * eapply redtmwf_conv; tea; now eapply redtmwf_refl. * now eapply NeType, convneu_whne. - * apply convtm_convneu ; tea. - 1: now constructor. - now eapply convneu_conv. - * eapply RedTyRecBwd, neu. 1,2: gen_typing. Defined. @@ -45,6 +42,7 @@ Proof. all: cbn; assumption. Qed. +(* TODO MOVE ME *) Lemma ty_app_ren {Γ Δ A f a dom cod} (ρ : Δ ≤ Γ) : [Γ |- f : A] -> [Γ |- A ≅ tProd dom cod] -> [Δ |- a : dom⟨ρ⟩] -> [Δ |- tApp f⟨ρ⟩ a : cod[a .: ρ >> tRel]]. Proof. @@ -55,6 +53,7 @@ Proof. gen_typing. Qed. +(* TODO MOVE ME *) Lemma convneu_app_ren {Γ Δ A f g a b dom cod} (ρ : Δ ≤ Γ) : [Γ |- f ~ g : A] -> [Γ |- A ≅ tProd dom cod] -> @@ -68,23 +67,22 @@ Proof. gen_typing. Qed. -Record complete {l Γ A} (RA : [Γ ||- A]) := { - reflect : forall n n', +Definition reflect {l Γ A} (RA : [Γ ||- A]) := + forall n n', [Γ |- n : A] -> [Γ |- n' : A] -> [Γ |- n ~ n' : A] -> - [Γ ||- n : A | RA] × [Γ ||- n ≅ n' : A| RA]; -}. + [Γ ||- n ≅ n' : A| RA]. -Lemma complete_reflect_simpl {l Γ A} (RA : [Γ ||- A]) (c : complete RA) : +Lemma reflect_diag {l Γ A} (RA : [Γ ||- A]) (c : reflect RA) : forall n, [Γ |- n : A] -> [Γ |- n ~ n : A] -> [Γ ||- n : A | RA]. Proof. intros; eapply c. all: eassumption. Qed. -Lemma complete_var0 {l Γ A A'} (RA : [Γ ,, A ||- A']) : - complete RA -> +Lemma reflect_var0 {l Γ A A'} (RA : [Γ ,, A ||- A']) : + reflect RA -> [Γ ,, A |- A⟨↑⟩ ≅ A'] -> [Γ |- A] -> [Γ ,, A ||- tRel 0 : A' | RA]. @@ -92,18 +90,17 @@ Proof. intros cRA conv HA. assert [Γ ,, A |- tRel 0 : A'] by (eapply ty_conv; tea; escape; eapply (ty_var (wfc_wft EscRA) (in_here _ _))). - eapply complete_reflect_simpl; tea. - - eapply convneu_var; tea. + eapply reflect_diag; tea. + eapply convneu_var; tea. Qed. -Lemma complete_U : forall l Γ A (RA : [Γ ||-U< l > A]), complete (LRU_ RA). +Lemma reflect_U : forall l Γ A (RA : [Γ ||-U< l > A]), reflect (LRU_ RA). Proof. -intros l Γ A h0; split. -- intros ???? h; pose proof (lrefl h); pose proof (urefl h). - assert [Γ |- A ≅ U] by (destruct h0; gen_typing); split. - 2: unshelve econstructor. - 1-3: now apply neU. + intros l Γ A h0 ???? h; pose proof (lrefl h); pose proof (urefl h). + assert [Γ |- A ≅ U] by (destruct h0; gen_typing). + unshelve econstructor. + 1-2: now apply neU. + eapply RedTyRecBwd, neu. 1,2: try gen_typing. + cbn. gen_typing. + eapply RedTyRecBwd; apply neu. 1,2: gen_typing. @@ -111,168 +108,105 @@ intros l Γ A h0; split. all: eapply ty_ne_term, tm_ne_conv; tea; gen_typing. Qed. -Lemma complete_ne : forall l Γ A (RA : [Γ ||-ne A]), complete (LRne_ l RA). +Lemma reflect_ne : forall l Γ A (RA : [Γ ||-ne A]), reflect (LRne_ l RA). Proof. -intros l Γ A h0; split. -- destruct h0 as [B []]; intros ** ; assert ([Γ |- A ≅ B]) by gen_typing ; split. - + exists n; cbn. - * eapply redtmwf_refl ; gen_typing. - * eapply lrefl; eapply convneu_conv; eassumption. - + exists n n'; cbn. - 1,2: eapply redtmwf_refl ; eapply ty_conv; gen_typing. - gen_typing. + intros l Γ A h0 ?????. + destruct h0 as [B []]; intros ** ; assert ([Γ |- A ≅ B]) by gen_typing. + exists n n'; cbn. + 1,2: eapply redtmwf_refl ; eapply ty_conv; gen_typing. + gen_typing. Qed. -Lemma complete_Pi : forall l Γ A (RA : [Γ ||-Π< l > A]), +Lemma reflect_Pi : forall l Γ A (RA : [Γ ||-Π< l > A]), (forall (Δ : context) (ρ : Δ ≤ Γ) (h : [ |-[ ta ] Δ]), - complete (PolyRed.shpRed RA ρ h)) -> - (forall (Δ : context) (a : term) (ρ : Δ ≤ Γ) (h : [ |-[ ta ] Δ]) - (ha : [PolyRed.shpRed RA ρ h | Δ ||- a : _]), - complete (PolyRed.posRed RA ρ h ha)) -> - complete (LRPi' RA). + reflect (PolyRed.shpRed RA ρ h)) -> + (forall (Δ : context) (a b : term) (ρ : Δ ≤ Γ) (h : [ |-[ ta ] Δ]) + (ha : [PolyRed.shpRed RA ρ h | Δ ||- a ≅ b: _]), + reflect (PolyRed.posRed RA ρ h ha)) -> + reflect (LRPi' RA). Proof. -intros l Γ A ΠA0 ihdom ihcod; split. -- set (ΠA := ΠA0); destruct ΠA0 as [dom cod]. + intros l Γ A ΠA0 ihdom ihcod. + set (ΠA := ΠA0); destruct ΠA0 as [dom cod]. simpl in ihdom, ihcod. assert [Γ |- A ≅ tProd dom cod] by gen_typing. - assert [Γ |- A ≅ tProd dom cod] by gen_typing. - assert [Γ |- dom]. - { - erewrite <- wk_id_ren_on. - eapply escape, polyRed. - gen_typing. - } + assert [Γ |- dom] by eapply PolyRed.shpTy, polyRed. assert [|- Γ ,, dom] as Hext by gen_typing. - assert [Γ,, dom |-[ ta ] tRel 0 : dom⟨@wk1 Γ dom⟩]. - { - eapply ty_var ; tea. - rewrite wk1_ren_on. - econstructor. - } + assert [Γ,, dom |-[ ta ] tRel 0 : dom⟨@wk1 Γ dom⟩] + by (rewrite wk1_ren_on; now eapply ty_var0). assert [Γ,, dom |-[ ta ] tRel 0 ~ tRel 0 : dom⟨@wk1 Γ dom⟩] by now apply convneu_var. assert [PolyRed.shpRed polyRed (wk1 dom) Hext | Γ,, dom ||- tRel 0 : dom⟨wk1 dom⟩] by now eapply ihdom. - assert [Γ ,, dom |- cod]. - { - replace cod with cod[tRel 0 .: @wk1 Γ dom >> tRel]. - 2: bsimpl; rewrite scons_eta'; now asimpl. - now eapply escape, polyRed. - } + assert [Γ ,, dom |- cod] by eapply PolyRed.posTy, polyRed. + assert (forall n Δ a (ρ : Δ ≤ Γ), [|- Δ] -> [Γ |- n : A] -> [Δ |- a : dom⟨ρ⟩] -> [Δ |-[ ta ] tApp n⟨ρ⟩ a : cod[a .: ρ >> tRel]]) as Happ. - { - intros. - eapply typing_meta_conv. - 1: eapply ty_app ; tea. - 1: eapply typing_meta_conv. - 1: eapply ty_wk. - - eassumption. - - eapply ty_conv ; tea. - - cbn ; reflexivity. - - now bsimpl. - } - assert (forall n, [Γ |- n : A] -> [Γ,, dom |-[ ta ] tApp n⟨@wk1 Γ dom⟩ (tRel 0) : cod]). - { - intros. - eapply typing_meta_conv. - 1: apply Happ ; tea. - bsimpl. rewrite scons_eta'. now bsimpl. - } + { + intros. + eapply typing_meta_conv. + 1: eapply ty_app ; tea. + 1: eapply typing_meta_conv. + 1: eapply ty_wk. + - eassumption. + - eapply ty_conv ; tea. + - cbn ; reflexivity. + - now bsimpl. + } + assert (forall n n', [Γ |- n : A] -> [Γ |- n' : A] -> - [Γ |- n ~ n : A] -> [Γ |- n' ~ n' : A] -> + [Γ |- n ~ n : A] -> [Γ |- n' ~ n' : A] -> [Γ |- n ~ n' : A] -> [Γ |-[ ta ] n ≅ n' : tProd dom cod]). { - intros. + intros ?????? hnn'. eapply convtm_eta ; tea. - - now eapply ty_conv. - - constructor. - now eapply convneu_conv. - - now eapply ty_conv. - - econstructor. - now eapply convneu_conv. - - eapply convneu_app_ren in H21 ; tea ; cycle -1. - 2: eapply ihcod in H21 as [_ hred]. - + now eapply escapeEqTerm, reflLRTmEq. - + erewrite <- wk1_ren_on. - eapply convtm_meta_conv. - 1: now escape. - 1: bsimpl; rewrite scons_eta' ; now bsimpl. - now bsimpl. - + eapply typing_meta_conv ; eauto. - bsimpl. rewrite scons_eta'. now bsimpl. - + eapply typing_meta_conv ; eauto. - bsimpl. rewrite scons_eta'. now bsimpl. + 1-4: first [now eapply ty_conv| constructor; now eapply convneu_conv]. + eapply convneu_app_ren in hnn' ; tea ; cycle -1. + 1: now eapply escapeEqTerm. + eapply ihcod in hnn' as hred. + + erewrite <- wk1_ren_on. + eapply convtm_meta_conv. + 1: now escape. + 1: bsimpl; rewrite scons_eta' ; now bsimpl. + now bsimpl. + + eapply typing_meta_conv ; eauto. + + eapply typing_meta_conv ; eauto. + Unshelve. 3: eassumption. } - unshelve refine ( let funred : forall n, [Γ |- n : A] -> [Γ |- n ~ n : A] -> [Γ ||-Π n : A | ΠA] := _ in _). + unshelve refine ( let funred : forall n, [Γ |- n : A] -> [Γ |- n ~ n : A] -> PiRedTm ΠA n := _ in _). { intros. exists n; cbn. - eapply redtmwf_refl ; gen_typing. - constructor; now eapply convneu_conv. - - eauto. - - intros. - eapply ihcod ; last first. - + eapply convne_meta_conv. - 1: eapply convneu_app. - * eapply convne_meta_conv. - 1: eapply convneu_wk. - 2: eapply convneu_conv ; tea. - all: cbn ; easy. - * now eapply escapeEqTerm, reflLRTmEq. - * now bsimpl. - * reflexivity. - + eapply Happ ; tea. - now escape. - + eapply Happ ; tea. - now escape. - - intros. - eapply ihcod ; last first. - + eapply convne_meta_conv. - 1: eapply convneu_app. - * eapply convne_meta_conv. - 1: eapply convneu_wk. - 2: eapply convneu_conv ; tea. - all: cbn ; easy. - * now escape. - * now bsimpl. - * reflexivity. - + eapply ty_conv. - 1: eapply Happ ; tea ; now escape. - symmetry. - eapply escapeEq, PolyRed.posExt ; tea. - + eapply Happ ; tea. - now escape. } intros ???? h. pose proof (lrefl h); pose proof (urefl h). - split. 1: now apply funred. unshelve econstructor. 1,2: now apply funred. all: cbn; clear funred. * eauto. * intros. apply ihcod; cbn. - + apply escapeTerm in ha; now eapply ty_app_ren. - + apply escapeTerm in ha; now eapply ty_app_ren. + + apply escapeTerm in hab; now eapply ty_app_ren. + + pose proof (escapeEq _ (ΠA.(PolyRed.posExt) _ _ hab)). + apply LRTmEqSym, escapeTerm in hab. + eapply ty_conv; [| now symmetry]. + now eapply ty_app_ren. + eapply convneu_app_ren. 1,2: eassumption. - eapply escapeEqTerm; eapply reflLRTmEq; eassumption. - - Unshelve. - all: eauto. + now eapply escapeEqTerm. Qed. Arguments ParamRedTy.outTy /. -Lemma complete_Sig : forall l Γ A (RA : [Γ ||-Σ< l > A]), +Lemma reflect_Sig : forall l Γ A (RA : [Γ ||-Σ< l > A]), (forall (Δ : context) (ρ : Δ ≤ Γ) (h : [ |-[ ta ] Δ]), - complete (PolyRed.shpRed RA ρ h)) -> - (forall (Δ : context) (a : term) (ρ : Δ ≤ Γ) (h : [ |-[ ta ] Δ]) - (ha : [PolyRed.shpRed RA ρ h | Δ ||- a : _]), - complete (PolyRed.posRed RA ρ h ha)) -> - complete (LRSig' RA). + reflect (PolyRed.shpRed RA ρ h)) -> + (forall (Δ : context) (a b : term) (ρ : Δ ≤ Γ) (h : [ |-[ ta ] Δ]) + (ha : [PolyRed.shpRed RA ρ h | Δ ||- a ≅ b : _]), + reflect (PolyRed.posRed RA ρ h ha)) -> + reflect (LRSig' RA). Proof. intros l Γ A ΣA0 ihdom ihcod. set (ΣA := ΣA0); destruct ΣA0 as [dom cod] ; cbn in *. @@ -281,39 +215,17 @@ Proof. by (destruct ΣA; cbn in *; gen_typing). assert [Γ |- ΣA.(outTy)] by (destruct ΣA; cbn in *; gen_typing). - assert [Γ |- dom]. - { - erewrite <- wk_id_ren_on. - eapply escape, polyRed. - gen_typing. - } + assert [Γ |- dom] by now eapply PolyRed.shpTy. assert [|- Γ ,, dom] as Hext by gen_typing. - assert [Γ,, dom |-[ ta ] tRel 0 : dom⟨@wk1 Γ dom⟩]. - { - eapply ty_var ; tea. - rewrite wk1_ren_on. - econstructor. - } + assert [Γ,, dom |-[ ta ] tRel 0 : dom⟨@wk1 Γ dom⟩] + by (rewrite wk1_ren_on; now eapply ty_var0). assert [Γ,, dom |-[ ta ] tRel 0 ~ tRel 0 : dom⟨@wk1 Γ dom⟩] by now apply convneu_var. assert [PolyRed.shpRed polyRed (wk1 dom) Hext | Γ,, dom ||- tRel 0 : dom⟨wk1 dom⟩] by now eapply ihdom. - assert [Γ ,, dom |- cod]. - { - replace cod with cod[tRel 0 .: @wk1 Γ dom >> tRel]. - 2: bsimpl; rewrite scons_eta'; now asimpl. - now eapply escape, polyRed. - } - assert (hfst : forall n Δ (ρ : Δ ≤ Γ) (h : [ |- Δ]), [Γ |- n : A] -> [Γ |- n ~ n : A] -> - [PolyRedPack.shpRed ΣA ρ h | Δ ||- tFst n⟨ρ⟩ : _]). - 1:{ - intros; eapply complete_reflect_simpl. - * eapply ihdom. - * rewrite wk_fst; eapply ty_wk; tea. - eapply ty_fst; now eapply ty_conv. - * rewrite wk_fst; eapply convneu_wk; tea. - eapply convneu_fst; now eapply convneu_conv. - } + assert [Γ ,, dom |- cod] + by now eapply PolyRed.posTy. + assert (hconv_fst : forall n n' Δ (ρ : Δ ≤ Γ) (h : [ |- Δ]), [Γ |- n : A] -> [Γ |- n' : A] -> [Γ |- n ~ n' : A] -> [PolyRedPack.shpRed ΣA ρ h | Δ ||- tFst n⟨ρ⟩ ≅ tFst n'⟨ρ⟩ : _]). 1:{ @@ -322,165 +234,128 @@ Proof. * rewrite wk_fst; eapply ty_wk; tea. eapply ty_fst; now eapply ty_conv. * rewrite wk_fst ; eapply ty_wk ; tea. - eapply ty_fst ; now eapply ty_conv. + eapply ty_fst ; now eapply ty_conv. * repeat rewrite wk_fst; eapply convneu_wk; tea. eapply convneu_fst; now eapply convneu_conv. } + assert (hconv : forall n n', [Γ |- n : A] -> [Γ |- n' : A] -> [Γ |- n ~ n : A] -> [Γ |- n' ~ n' : A] -> [Γ |- n ~ n' : A] -> [Γ |-[ ta ] n ≅ n' : tSig dom cod]). { - intros. + intros n n' hn hn' hnn hn'n' hnn'. + assert (wfΓ : [|- Γ]) by gen_typing. + pose proof (hfst := hconv_fst n n' Γ wk_id wfΓ hn hn' hnn'). eapply convtm_eta_sig ; cbn in * ; tea. - - now eapply ty_conv. - - econstructor. - now eapply convneu_conv. - - now eapply ty_conv. - - econstructor. - now eapply convneu_conv. - - eapply convtm_meta_conv. - 1: eapply escapeEqTerm, ihdom. - 4: now rewrite wk_id_ren_on. - 4: reflexivity. - all: rewrite wk_id_ren_on. - + now eapply ty_fst, ty_conv. - + now eapply ty_fst, ty_conv. - + eapply convneu_fst, convneu_conv ; tea. - Unshelve. - gen_typing. - - eapply convtm_meta_conv. - 1: eapply escapeEqTerm, (ihcod _ (tFst n) wk_id). - 5: reflexivity. - Unshelve. - + eapply typing_meta_conv. - 1: gen_typing. - now bsimpl. - + eapply ty_conv. - 1: gen_typing. - symmetry. - replace (cod[(tFst n')..]) with (cod[(tFst n') .: (@wk_id Γ) >> tRel]) by (now bsimpl). - eapply escapeEq, polyRed.(PolyRed.posExt) ; tea. - Unshelve. - * now erewrite <- wk_id_ren_on. - * now erewrite <- (wk_id_ren_on _ n), <- (wk_id_ren_on _ n'). - * gen_typing. - * now erewrite <- wk_id_ren_on. - + eapply convne_meta_conv. - 1:now eapply convneu_snd, convneu_conv. - 1: now bsimpl. - easy. - + now bsimpl. - + gen_typing. - + now erewrite <- wk_id_ren_on. - } - split. + 1-4: first [now eapply ty_conv| constructor; now eapply convneu_conv]. + 1: generalize (escapeEqTerm _ hfst); now rewrite !wk_id_ren_on. + + eapply convtm_meta_conv. + 1: unshelve eapply escapeEqTerm, (ihcod _ (tFst n)⟨wk_id⟩ (tFst n)⟨wk_id⟩ wk_id); tea. + 6: reflexivity. + + now eapply lrefl. + + eapply typing_meta_conv. + 1: gen_typing. + now bsimpl. + + eapply ty_conv. + 1: gen_typing. + symmetry. + replace (cod[(tFst n')..]) with (cod[(tFst n') .: (@wk_id Γ) >> tRel]) by (now bsimpl). + unshelve eapply escapeEq, polyRed.(PolyRed.posExt) ; tea. + now erewrite <- (wk_id_ren_on _ n'). + + eapply convne_meta_conv. + 1:now eapply convneu_snd, convneu_conv. + 1: now bsimpl. + easy. + + now bsimpl. + } + unshelve refine ( let funred : forall n, [Γ |- n : A] -> - [Γ |- n ~ n : A] -> [Γ ||-Σ n : A | ΣA] := _ in _). + [Γ |- n ~ n : A] -> SigRedTm ΣA n := _ in _). { intros n **. cbn in *. - unshelve eexists n _. - - intros. now eapply hfst. + unshelve eexists n. - eapply redtmwf_refl; now eapply ty_conv. - constructor ; cbn ; now eapply convneu_conv. - - eauto. - - intros. - cbn. - eapply complete_reflect_simpl. - * eapply ihcod. - * rewrite wk_snd. - eapply typing_meta_conv. - 1: eapply ty_wk ; tea. - 1: now eapply ty_snd, ty_conv. - now bsimpl. - * eapply convne_meta_conv. - 3: reflexivity. - 1: rewrite wk_snd. - 1: eapply convneu_wk ; tea. - 1: now eapply convneu_snd, convneu_conv. - now bsimpl. } intros ???? h. - pose proof (lrefl h); pose proof (urefl h). - unshelve refine (let Rn :[Γ ||-Σ n : A | ΣA] := _ in _). - 1: eapply funred; tea; now eapply lrefl. - unshelve refine (let Rn' :[Γ ||-Σ n' : A | ΣA] := _ in _). - 1: eapply funred; tea; now eapply urefl. - assert (forall (Δ : context) (ρ : Δ ≤ Γ) (h : [ |-[ ta ] Δ]), - [Δ |- cod[tFst n⟨ρ⟩ .: ρ >> tRel] ≅ cod[tFst n'⟨ρ⟩ .: ρ >> tRel]]). - { - intros; eapply escapeEq; unshelve eapply (PolyRed.posExt ΣA); tea. - + eapply Rn. - + eapply Rn'. - + now eapply hconv_fst. - } - split; tea; eexists Rn Rn'. - + cbn. - eapply hconv ; tea. - + cbn. intros. now eapply hconv_fst. - + intros; cbn; eapply ihcod. - all: rewrite wk_fst; rewrite !wk_snd. - 2: eapply ty_conv; [|now symmetry]; rewrite wk_fst. - all: rewrite <- subst_ren_subst_mixed. - * eapply ty_wk; tea; eapply ty_snd; now eapply ty_conv. - * eapply ty_wk; tea; eapply ty_snd; now eapply ty_conv. - * eapply convneu_wk; tea; eapply convneu_snd; now eapply convneu_conv. + pose proof (lrefl h); pose proof (urefl h). + unshelve refine (let Rn : SigRedTm ΣA n := _ in _). + 1: now eapply funred. + unshelve refine (let Rn' : SigRedTm ΣA n' := _ in _). + 1: now eapply funred. + assert (forall (Δ : context) (ρ : Δ ≤ Γ) (h : [ |-[ ta ] Δ]), + [Δ |- cod[tFst n⟨ρ⟩ .: ρ >> tRel] ≅ cod[tFst n'⟨ρ⟩ .: ρ >> tRel]]). + { + intros; eapply escapeEq; unshelve eapply (PolyRed.posExt ΣA); tea. + now eapply hconv_fst. + } + unshelve eexists Rn Rn' _. + + cbn. intros. now eapply hconv_fst. + + cbn. eapply hconv ; tea. + + intros; cbn. eapply ihcod. + all: rewrite wk_fst; rewrite !wk_snd. + 2: eapply ty_conv; [|now symmetry]; rewrite wk_fst. + all: rewrite <- subst_ren_subst_mixed. + * eapply ty_wk; tea; eapply ty_snd; now eapply ty_conv. + * eapply ty_wk; tea; eapply ty_snd; now eapply ty_conv. + * eapply convneu_wk; tea; eapply convneu_snd; now eapply convneu_conv. Qed. -Lemma complete_Nat {l Γ A} (NA : [Γ ||-Nat A]) : complete (LRNat_ l NA). +Lemma reflect_Nat {l Γ A} (NA : [Γ ||-Nat A]) : reflect (LRNat_ l NA). Proof. - split. - - intros. - assert [Γ |- A ≅ tNat] by (destruct NA; gen_typing). - assert [Γ |- n : tNat] by now eapply ty_conv. - split; econstructor. - 1,4,5: eapply redtmwf_refl; tea; now eapply ty_conv. - 2,4: do 2 constructor; tea. - 1,4: eapply convtm_convneu ; [now constructor|..]. - all: eapply convneu_conv; [|eassumption]. - all: first [assumption|now eapply lrefl]. + red; intros. + assert [Γ |- A ≅ tNat] by (destruct NA; gen_typing). + assert [Γ |- n : tNat] by now eapply ty_conv. + econstructor. + 1,2: eapply redtmwf_refl; tea; now eapply ty_conv. + 2: do 2 constructor; tea. + 1: eapply convtm_convneu ; [now constructor|..]. + 1,3: eapply convneu_conv; [|eassumption]; tea. + eapply ty_conv; eassumption. Qed. -Lemma complete_Empty {l Γ A} (NA : [Γ ||-Empty A]) : complete (LREmpty_ l NA). +Lemma reflect_Empty {l Γ A} (NA : [Γ ||-Empty A]) : reflect (LREmpty_ l NA). Proof. - split. - - intros. - assert [Γ |- A ≅ tEmpty] by (destruct NA; gen_typing). - assert [Γ |- n : tEmpty] by now eapply ty_conv. - split; econstructor. - 1,4,5: eapply redtmwf_refl; tea; now eapply ty_conv. - 2,4: do 2 constructor; tea. - 1,4: eapply convtm_convneu ; [now constructor|..]. - all: eapply convneu_conv; [|eassumption]. - all: try first [assumption|now eapply lrefl]. + red; intros. + assert [Γ |- A ≅ tEmpty] by (destruct NA; gen_typing). + assert [Γ |- n : tEmpty] by now eapply ty_conv. + assert [Γ |- n' : tEmpty] by now eapply ty_conv. + econstructor. + 1,2: eapply redtmwf_refl; tea; now eapply ty_conv. + 2: do 2 constructor; tea. + 1: eapply convtm_convneu ; [now constructor|..]. + all: now eapply convneu_conv. Qed. -Lemma complete_Id {l Γ A} (IA : [Γ ||-Id A]) : - complete (LRId' IA). +Lemma reflect_Id {l Γ A} (IA : [Γ ||-Id A]) : + reflect (LRId' IA). Proof. - split; intros. + red; intros. assert [Γ |- A ≅ IA.(IdRedTy.outTy)] by (destruct IA; unfold IdRedTy.outTy; cbn; gen_typing). assert [Γ |- n : IA.(IdRedTy.outTy)] by now eapply ty_conv. - split; econstructor. - 1,4,5: eapply redtmwf_refl; tea; now eapply ty_conv. - 2,4: do 2 constructor; tea. - 1,4: eapply convtm_convneu ; [now constructor|..]. - all: eapply convneu_conv; tea; now eapply lrefl. + assert [Γ |- n' : IA.(IdRedTy.outTy)] by now eapply ty_conv. + econstructor. + 1,2: eapply redtmwf_refl; tea; now eapply ty_conv. + 2: do 2 constructor; tea. + 1: eapply convtm_convneu ; [now constructor|..]. + all: now eapply convneu_conv. Qed. -Lemma completeness {l Γ A} (RA : [Γ ||- A]) : complete RA. +Lemma reflectness {l Γ A} (RA : [Γ ||- A]) : reflect RA. Proof. revert l Γ A RA; eapply LR_rect_TyUr; cbn; intros. -- now apply complete_U. -- now apply complete_ne. -- now apply complete_Pi. -- now apply complete_Nat. -- now apply complete_Empty. -- now apply complete_Sig. -- now apply complete_Id. +- now apply reflect_U. +- now apply reflect_ne. +- now apply reflect_Pi. +- now apply reflect_Nat. +- now apply reflect_Empty. +- now apply reflect_Sig. +- now apply reflect_Id. Qed. Lemma neuTerm {l Γ A} (RA : [Γ ||- A]) {n} : @@ -488,7 +363,7 @@ Lemma neuTerm {l Γ A} (RA : [Γ ||- A]) {n} : [Γ |- n ~ n : A] -> [Γ ||- n : A | RA]. Proof. - intros. now eapply completeness. + intros. now eapply reflectness. Qed. Lemma neuTermEq {l Γ A} (RA : [Γ ||- A]) {n n'} : @@ -497,15 +372,19 @@ Lemma neuTermEq {l Γ A} (RA : [Γ ||- A]) {n n'} : [Γ |- n ~ n' : A] -> [Γ ||- n ≅ n' : A| RA]. Proof. - intros; now eapply completeness. + intros; now eapply reflectness. Qed. +Lemma neNfTermEq {Γ l A n n'} (RA : [Γ ||- A]) : [Γ ||-NeNf n ≅ n' : A] -> [RA | Γ ||- n ≅ n' : A]. +Proof. intros []; now eapply neuTermEq. Qed. + + Lemma var0conv {l Γ A A'} (RA : [Γ ,, A ||- A']) : [Γ,, A |- A⟨↑⟩ ≅ A'] -> [Γ |- A] -> [Γ ,, A ||- tRel 0 : A' | RA]. Proof. - apply complete_var0 ; now eapply completeness. + apply reflect_var0 ; now eapply reflectness. Qed. Lemma var0 {l Γ A A'} (RA : [Γ ,, A ||- A']) : @@ -520,4 +399,6 @@ Proof. eapply reflLRTyEq. Qed. -End Neutral. \ No newline at end of file +End Neutral. + + diff --git a/theories/LogicalRelation/NormalRed.v b/theories/LogicalRelation/NormalRed.v index 94e6846..782a419 100644 --- a/theories/LogicalRelation/NormalRed.v +++ b/theories/LogicalRelation/NormalRed.v @@ -2,13 +2,13 @@ From Coq Require Import CRelationClasses. From LogRel.AutoSubst Require Import core unscoped Ast Extra. From LogRel Require Import Utils BasicAst Notations Context NormalForms UntypedReduction Weakening GenericTyping LogicalRelation. -From LogRel.LogicalRelation Require Import Induction Irrelevance Escape Reflexivity Weakening Neutral Transitivity Reduction. +From LogRel.LogicalRelation Require Import Induction Irrelevance Escape Reflexivity Weakening Neutral Transitivity Reduction EqRedRight. Set Universe Polymorphism. Ltac redSubst := match goal with - | [H : [_ |- ?X :⤳*: ?Y] |- _] => + | [H : [_ |- ?X :⤳*: ?Y] |- _] => assert (X = Y)by (eapply redtywf_whnf ; gen_typing); subst; clear H end. @@ -31,28 +31,41 @@ Program Definition normRedΠ0 {Γ F G l} (h : [Γ ||-Π tProd F G]) : [Γ ||-Π tProd F G] := {| PiRedTyPack.dom := F ; PiRedTyPack.cod := G |}. -Solve All Obligations with +Solve All Obligations with intros; pose proof (e := redtywf_whnf (PiRedTyPack.red h) whnf_tProd); - symmetry in e; injection e; clear e; + symmetry in e; injection e; clear e; destruct h ; intros; cbn in *; subst; eassumption. Program Definition normRedΣ0 {Γ F G l} (h : [Γ ||-Σ tSig F G]) : [Γ ||-Σ tSig F G] := {| PiRedTyPack.dom := F ; PiRedTyPack.cod := G |}. -Solve All Obligations with +Solve All Obligations with intros; pose proof (e := redtywf_whnf (PiRedTyPack.red h) whnf_tSig); - symmetry in e; injection e; clear e; + symmetry in e; injection e; clear e; destruct h ; intros; cbn in *; subst; eassumption. Definition normRedΠ {Γ F G l} (h : [Γ ||- tProd F G]) : [Γ ||- tProd F G] := LRPi' (normRedΠ0 (invLRΠ h)). +#[program] +Definition normEqRedΠ {Γ F F' G G' l} (h : [Γ ||- tProd F G]) + (heq : [Γ ||- _ ≅ tProd F' G' | h]) : [Γ ||- _ ≅ tProd F' G' | normRedΠ h] := + {| + PiRedTyEq.dom := F'; + PiRedTyEq.cod := G'; + |}. +Solve All Obligations with + intros; assert[Γ ||- _ ≅ tProd F' G' | normRedΠ h] as [?? red] by irrelevance; + pose proof (e := redtywf_whnf red whnf_tProd); + symmetry in e; injection e; clear e; + destruct h ; intros; cbn in *; subst; eassumption. + Definition normRedΣ {Γ F G l} (h : [Γ ||- tSig F G]) : [Γ ||- tSig F G] := LRSig' (normRedΣ0 (invLRΣ h)). #[program] -Definition normEqRedΣ {Γ F F' G G' l} (h : [Γ ||- tSig F G]) +Definition normEqRedΣ {Γ F F' G G' l} (h : [Γ ||- tSig F G]) (heq : [Γ ||- _ ≅ tSig F' G' | h]) : [Γ ||- _ ≅ tSig F' G' | normRedΣ h] := {| PiRedTyEq.dom := F'; @@ -61,29 +74,49 @@ Definition normEqRedΣ {Γ F F' G G' l} (h : [Γ ||- tSig F G]) Solve All Obligations with intros; assert[Γ ||- _ ≅ tSig F' G' | normRedΣ h] as [?? red] by irrelevance; pose proof (e := redtywf_whnf red whnf_tSig); - symmetry in e; injection e; clear e; + symmetry in e; injection e; clear e; destruct h ; intros; cbn in *; subst; eassumption. #[program] -Definition normLambda {Γ F F' G t l RΠ} - (Rlam : [Γ ||- tLambda F' t : tProd F G | normRedΠ RΠ ]) : - [Γ ||- tLambda F' t : tProd F G | normRedΠ RΠ ] := - {| PiRedTm.nf := tLambda F' t |}. +Definition normRedId0 {Γ A x y l} (h : [Γ ||-Id tId A x y]) + : [Γ ||-Id tId A x y] := + mkIdRedTy A x y _ _ _ _. +Solve All Obligations with + intros; pose proof (e := redtywf_whnf (IdRedTy.red h) whnf_tId); + symmetry in e; injection e; clear e; + destruct h ; intros; cbn in *; subst; + first [eassumption | irrelevance]. + +Definition normRedId {Γ A x y l} (h : [Γ ||- tId A x y]) : [Γ ||- tId A x y] := + LRId' (normRedId0 (invLRId h)). + +#[program] +Definition normLambda {Γ F Fl Fr G tl tr l RΠ} + (Rlam : [Γ ||- tLambda Fl tl ≅ tLambda Fr tr : tProd F G | normRedΠ RΠ ]) : + [Γ ||- tLambda Fl tl ≅ tLambda Fr tr : tProd F G | normRedΠ RΠ ] := + {| + PiRedTmEq.redL := {| PiRedTmEq.nf := tLambda Fl tl |} ; + PiRedTmEq.redR := {| PiRedTmEq.nf := tLambda Fr tr |} + |}. Solve All Obligations with intros; - pose proof (e := redtmwf_whnf (PiRedTm.red Rlam) whnf_tLambda); - destruct Rlam as [???? app eqq]; cbn in *; subst; + pose proof (eL := redtmwf_whnf (PiRedTmEq.red Rlam.(PiRedTmEq.redL)) whnf_tLambda); + pose proof (eR := redtmwf_whnf (PiRedTmEq.red Rlam.(PiRedTmEq.redR)) whnf_tLambda); + destruct Rlam as [[] [] app eqq]; cbn in *; subst; first [eapply app | now eapply eqq| eassumption]. + #[program] -Definition normPair {Γ F F' G G' f g l RΣ} - (Rp : [Γ ||- tPair F' G' f g : tSig F G | normRedΣ RΣ ]) : - [Γ ||- tPair F' G' f g : tSig F G | normRedΣ RΣ ] := - {| SigRedTm.nf := tPair F' G' f g |}. +Definition normPair {Γ F Fl Fr G Gl Gr fl fr gl gr l RΣ} + (Rp : [Γ ||- tPair Fl Gl fl gl ≅ tPair Fr Gr fr gr : tSig F G | normRedΣ RΣ ]) : + [Γ ||- tPair Fl Gl fl gl ≅ tPair Fr Gr fr gr : tSig F G | normRedΣ RΣ ] := + {| SigRedTmEq.redL := {| SigRedTmEq.nf := tPair Fl Gl fl gl |} ; + SigRedTmEq.redR := {| SigRedTmEq.nf := tPair Fr Gr fr gr |} |}. Solve All Obligations with intros; - pose proof (e := redtmwf_whnf (SigRedTm.red Rp) whnf_tPair); - destruct Rp as [???? fstRed sndRed]; cbn in *; subst; + pose proof (eL := redtmwf_whnf (SigRedTmEq.red (SigRedTmEq.redL Rp)) whnf_tPair); + pose proof (eR := redtmwf_whnf (SigRedTmEq.red (SigRedTmEq.redR Rp)) whnf_tPair); + destruct Rp as [[] [] ? fstRed sndRed]; cbn in *; subst; first [eapply fstRed | irrelevanceRefl; now unshelve eapply sndRed| eassumption]. Definition invLRcan {Γ l A} (lr : [Γ ||- A]) (w : isType A) : [Γ ||- A] := @@ -96,7 +129,6 @@ Definition invLRcan {Γ l A} (lr : [Γ ||- A]) (w : isType A) : [Γ ||- A] | IdType => fun _ x => LRId' x | NeType wh => fun _ x => LRne_ _ (normRedNe0 x wh) end lr (invLR lr (reflexivity A) w). - End Normalization. @@ -107,10 +139,10 @@ End Normalization. Ltac normRedΠin X := let g := type of X in match g with - | [ LRAd.pack ?R | _ ||- ?t : _] => + (* | [ LRAd.pack ?R | _ ||- ?t : _] => let X' := fresh X in rename X into X' ; - assert (X : [_ ||-<_> t : _ | LRPi' (normRedΠ0 (invLRΠ R))]) by irrelevance; clear X' + assert (X : [_ ||-<_> t : _ | LRPi' (normRedΠ0 (invLRΠ R))]) by irrelevance; clear X' *) | [ LRAd.pack ?R | _ ||- _ ≅ ?B] => let X' := fresh X in rename X into X' ; @@ -123,12 +155,12 @@ Ltac normRedΠin X := Goal forall `{GenericTypingProperties} Γ A B l R f X (RABX : [Γ ||- arr A B ≅ X | R]) - (Rf : [Γ ||- f : arr A B | R]) + (* (Rf : [Γ ||- f : arr A B | R]) *) (Rff : [Γ ||- f ≅ f : arr A B | R]) , True. Proof. intros. - normRedΠin Rf. + (* normRedΠin Rf. *) normRedΠin Rff. normRedΠin RABX. constructor. @@ -141,9 +173,9 @@ Ltac normRedΠ id := set (G := _); let g := eval unfold G in G in match g with - | [ LRAd.pack ?R | _ ||- ?t : _] => + | [ LRAd.pack ?R | _ ||- ?t ≅ ?u : _] => pose (id := normRedΠ0 (invLRΠ R)); subst G; - enough [_ ||-<_> t : _ | LRPi' id] by irrelevance + enough [_ ||-<_> t ≅ u : _ | LRPi' id] by irrelevance end. (* Normalizes a term reducible at a Π type *) @@ -151,10 +183,10 @@ Ltac normRedΠ id := Ltac normRedΣin X := let g := type of X in match g with - | [ LRAd.pack ?R | _ ||- ?t : _] => + (* | [ LRAd.pack ?R | _ ||- ?t : _] => let X' := fresh X in rename X into X' ; - assert (X : [_ ||-<_> t : _ | LRSig' (normRedΣ0 (invLRΣ R))]) by irrelevance; clear X' + assert (X : [_ ||-<_> t : _ | LRSig' (normRedΣ0 (invLRΣ R))]) by irrelevance; clear X' *) | [ LRAd.pack ?R | _ ||- _ ≅ ?B] => let X' := fresh X in rename X into X' ; diff --git a/theories/LogicalRelation/Poly.v b/theories/LogicalRelation/Poly.v new file mode 100644 index 0000000..2ea342a --- /dev/null +++ b/theories/LogicalRelation/Poly.v @@ -0,0 +1,2 @@ + +(** * LogRel.LogicalRelation.Poly: derived properties of reducible polynomials.*) \ No newline at end of file diff --git a/theories/LogicalRelation/Reduction.v b/theories/LogicalRelation/Reduction.v index b1c1eba..c1294c1 100644 --- a/theories/LogicalRelation/Reduction.v +++ b/theories/LogicalRelation/Reduction.v @@ -1,7 +1,7 @@ -(* From Coq.Classes Require Import CRelationClasses. *) +From Coq.Classes Require Import CRelationClasses. From LogRel.AutoSubst Require Import core unscoped Ast Extra. From LogRel Require Import Notations Utils BasicAst Context NormalForms UntypedReduction Weakening GenericTyping LogicalRelation. -From LogRel.LogicalRelation Require Import Induction Reflexivity Universe Escape Irrelevance. +From LogRel.LogicalRelation Require Import Induction Reflexivity Universe Escape Irrelevance Transitivity. Set Universe Polymorphism. @@ -57,116 +57,135 @@ Proof. intros ? []; now eapply redSubst. Qed. -Lemma redSubstTerm {Γ A t u l} (RA : [Γ ||- A]) : - [Γ ||- u : A | RA] -> + +Lemma redURedTm {Γ l A t u} {h : [Γ ||-U A]} : + [Γ |- t ⤳* u : A] -> + URedTm (LogRelRec l) Γ u A h -> + URedTm (LogRelRec l) Γ t A h. +Proof. + intros redtu ru; pose proof (invLRConvU h). + assert [Γ |- t : U] by (eapply ty_conv; [gen_typing| now escape]). + destruct ru as [nf]; exists nf; tea. + etransitivity; tea; gen_typing. +Defined. + +Lemma redPiRedTm {Γ l A t u} {h : [Γ ||-Π A]} : + [Γ |- t ⤳* u : A] -> + PiRedTm h u -> + PiRedTm h t. +Proof. + intros redtu ru; pose proof (invLRConvPi h). + assert [Γ |-[ ta ] t ⤳* u : h.(outTy)] by now eapply redtm_conv. + destruct ru as [nf]; exists nf; tea. + constructor; destruct red; tea; now etransitivity. +Defined. + +Lemma redSigRedTm {Γ l A t u} {h : [Γ ||-Σ A]} : + [Γ |- t ⤳* u : A] -> + SigRedTm h u -> + SigRedTm h t. +Proof. + intros redtu ru; pose proof (invLRConvSig h). + assert [Γ |-[ ta ] t ⤳* u : h.(outTy)] by now eapply redtm_conv. + destruct ru as [nf]; exists nf; tea. + constructor; destruct red; tea; now etransitivity. +Defined. + +Lemma eqAppLRefl {Γ l A u v} (ΠA : [Γ ||-Π A]): + (forall Δ a b, PiRedTmEq.appRed ΠA u v Δ a b) -> + (forall Δ a b, PiRedTmEq.appRed ΠA u u Δ a b). +Proof. + cbn; intros eqApp; etransitivity. + 1: eapply eqApp. + symmetry; eapply LRTmEqConv. + 2: unshelve eapply eqApp; tea; now eapply urefl. + irrelevanceRefl; unshelve eapply (ΠA.(PolyRed.posExt) ρ h). + now symmetry. +Qed. + +Lemma redSubstLeftTmEq {Γ A t u v l} (RA : [Γ ||- A]) : + [Γ ||- u ≅ v : A | RA] -> [Γ |- t ⤳* u : A ] -> - [Γ ||- t : A | RA] × [Γ ||- t ≅ u : A | RA]. + [Γ ||- t ≅ v : A | RA]. Proof. - revert t u; pattern l, Γ, A, RA; apply LR_rect_TyUr; clear l Γ A RA; intros l Γ. - - intros ? h ?? ru' redtu; pose (ru := ru'); destruct ru' as [T]. - assert [Γ |- A ≅ U] by (destruct h; gen_typing). - assert [Γ |- t : U] by (eapply ty_conv; [gen_typing| now escape]). - assert (redtu' : [Γ |-[ ta ] t ⤳* u]) by gen_typing. - destruct (redSubst (A:=t) (RedTyRecFwd h rel) redtu') as [rTyt0]. - pose proof (rTyt := RedTyRecBwd h rTyt0). - unshelve refine (let rt : [LRU_ h | Γ ||- t : A]:= _ in _). - + exists T; tea; constructor; destruct red; tea. - etransitivity; tea; eapply redtm_conv; tea; now escape. - + split; tea; exists rt ru rTyt; tea. - apply TyEqRecFwd; irrelevance. - - intros ???? ru' ?; pose (ru := ru'); destruct ru' as [n]. - assert [Γ |- A ≅ neRedTy.ty neA]. 1:{ - destruct neA; cbn in *. - eapply convty_exp. - 2: apply redtywf_refl; gen_typing. - 1: gen_typing. - apply convty_term; apply convtm_convneu. - all: gen_typing. - } - assert [Γ |-[ ta ] t ⤳* u : neRedTy.ty neA] by (eapply redtm_conv; tea). - assert [Γ |-[ ta ] t : neRedTy.ty neA] by (eapply ty_conv; tea; gen_typing). - unshelve refine (let rt : [LRne_ l neA | Γ ||- t : A] := _ in _). - + exists n; tea; destruct red; constructor; tea; now etransitivity. - + split; tea; exists n n; tea; destruct red; constructor; tea; now etransitivity. - - intros ? ΠA ihdom ihcod ?? ru' ?; pose (ru := ru'); destruct ru' as [f]. - assert [Γ |- A ≅ ΠA.(outTy)]. 1:{ - destruct ΠA as [?? []]. cbn in *. - eapply convty_exp; tea. - now apply redtywf_refl. - } - assert [Γ |-[ ta ] t ⤳* u : ΠA.(outTy)] by now eapply redtm_conv. - assert [Γ |-[ ta ] t : ΠA.(outTy)] by (eapply ty_conv; gen_typing). - unshelve refine (let rt : [LRPi' ΠA | Γ ||- t : A] := _ in _). - 1: exists f; tea; constructor; destruct red; tea; etransitivity; tea. - split; tea; exists rt ru; tea. - intros; cbn. apply eq; tea. - now apply reflLRTmEq. - - intros ? NA t ? Ru red; inversion Ru; subst. - assert [Γ |- A ≅ tNat] by (destruct NA; gen_typing). - assert [Γ |- t :⤳*: nf : tNat]. 1:{ + intros uv tu; transitivity u; [|assumption]. + revert t u v uv tu; pattern l, Γ, A, RA; apply LR_rect_TyUr; clear l Γ A RA; intros l Γ. + - intros ? h ??? [ru] redtu ; pose proof (invLRConvU h). + assert (redtu' : [Γ |-[ ta ] t ⤳* u]) by (eapply redty_term; gen_typing). + destruct (redSubst (A:=t) (RedTyRecFwd h relL) redtu') as [rTyt0 ?]. + unshelve eexists (redURedTm redtu ru) ru (RedTyRecBwd h rTyt0); tea. + 1: now eapply lrefl. + apply TyEqRecFwd; irrelevance. + - intros ????? ru' ?; pose (ru := ru'); destruct ru' as [n]. + assert [Γ |- A ≅ neRedTy.ty neA] by now eapply invLRConvNe. + assert [Γ |-[ ta ] t ⤳* u : neRedTy.ty neA] by (eapply redtm_conv; tea). + assert [Γ |-[ ta ] t : neRedTy.ty neA] by (eapply ty_conv; tea; gen_typing). + exists n n; tea. + 2: now eapply lrefl. + etransitivity; tea; gen_typing. + - intros ? ΠA ihdom ihcod ??? ru redtu. + pose proof (lrefl ru); destruct ru as [ru]. + exists (redPiRedTm redtu ru) ru ; tea. + 1: cbn; now eapply lrefl. + now eapply eqAppLRefl. + - intros ? NA t u ? Ru red. + inversion Ru as [?? nfL ???? [? _]%reflNatRedTmEq]; subst. + pose proof (invLRConvNat NA). + assert [Γ |- t :⤳*: nfL : tNat]. 1:{ constructor. 1: gen_typing. etransitivity. 2: gen_typing. now eapply redtm_conv. } - split; econstructor; tea. - now eapply reflNatRedTmEq. - - intros ? NA t ? Ru red; inversion Ru; subst. - assert [Γ |- A ≅ tEmpty] by (destruct NA; gen_typing). - assert [Γ |- t :⤳*: nf : tEmpty]. 1:{ - constructor. - 1: eapply ty_conv; gen_typing. + econstructor; tea; now eapply lrefl. + - intros ? NA t u ? Ru red. + inversion Ru as [?? nfL ???? [? _]%(reflEmptyRedTmEq (NA:=NA))]; subst. + pose proof (invLRConvEmpty NA). + assert [Γ |- t :⤳*: nfL : tEmpty]. 1:{ + constructor. 1: gen_typing. etransitivity. 2: gen_typing. now eapply redtm_conv. } - split; econstructor; tea. - now eapply reflEmptyRedTmEq. - Unshelve. 2: tea. - - intros ? PA ihdom ihcod ?? ru' ?; pose (ru := ru'); destruct ru' as [p]. - assert [Γ |- A ≅ PA.(outTy)]. 1:{ - destruct PA as [?? []]. cbn in *. - eapply convty_exp; tea. - now apply redtywf_refl. - } - assert [Γ |-[ ta ] t ⤳* u : PA.(outTy)] by now eapply redtm_conv. - assert [Γ |-[ ta ] t : PA.(outTy)] by (eapply ty_conv; gen_typing). - unshelve refine (let rt : [LRSig' PA | Γ ||- t : A] := _ in _). - 1: unshelve eexists p _; tea; constructor; destruct red; tea; etransitivity; tea. - split; tea; exists rt ru; tea; intros; cbn; now apply reflLRTmEq. - - intros ? IA ih _ t u [nf ?? prop] redt. - assert [Γ |- A ≅ IA.(IdRedTy.outTy)] by (destruct IA; unfold IdRedTy.outTy; cbn; gen_typing). - assert [Γ |-[ ta ] t ⤳* u : IA.(IdRedTy.outTy)] by now eapply redtm_conv. + econstructor; tea; now eapply lrefl. + - intros ? PA ihdom ihcod ??? [ru] redtu. + unshelve eexists (redSigRedTm redtu ru) ru _. + all:cbn; intros; eapply lrefl; eauto. + irrelevanceRefl; unshelve eauto; tea. + - intros ? IA ih _ t u ? [nf ? redL ?? prop] redt. + assert [Γ |- A ≅ IA.(IdRedTy.outTy)] by now eapply invLRConvId. + assert [Γ |-[ ta ] t ⤳* u : IA.(IdRedTy.outTy)] by now eapply redtm_conv. assert [Γ |-[ ta ] t : IA.(IdRedTy.outTy)] by (eapply ty_conv; gen_typing). assert [Γ |-[ ta ] t :⤳*: nf : IA.(IdRedTy.outTy)]. - 1: constructor; [apply red|etransitivity; tea; apply red]. - split; eexists; tea. - now eapply reflIdPropEq. + 1: destruct redL; constructor; tea; now etransitivity. + eexists; tea; [now eapply lrefl|]. + now eapply reflIdPropEq in prop as []. Qed. -Lemma redSubstTerm' {Γ A t u l} (RA : [Γ ||- A]) : - [Γ ||- u : A | RA] -> - [Γ |- t ⤳* u : A ] -> - [Γ ||- t : A | RA] × - [Γ ||- u : A | RA] × - [Γ ||- t ≅ u : A | RA]. +Lemma redSubstTmEq {Γ A A' tl tr ul ur l} (RA : [Γ ||- A]) : + [Γ ||- ul ≅ ur : A | RA] -> + [Γ |- tl ⤳* ul : A ] -> + [Γ |- tr ⤳* ur : A' ] -> + [Γ |- A ≅ A'] -> + [Γ ||- tl ≅ tr : A | RA]. Proof. - intros. assert ([Γ ||- t : A | RA] × [Γ ||- t ≅ u : A | RA]) by now eapply redSubstTerm. - now repeat split. + intros. + assert [Γ |- tr ⤳* ur : A ]. + 1: eapply redtm_conv; tea; now symmetry. + eapply redSubstLeftTmEq; tea; symmetry. + eapply redSubstLeftTmEq; tea; now symmetry. Qed. - -Lemma redwfSubstTerm {Γ A t u l} (RA : [Γ ||- A]) : - [Γ ||- u : A | RA] -> +Lemma redwfSubstTmEq {Γ A t u v l} (RA : [Γ ||- A]) : + [Γ ||- u ≅ v : A | RA] -> [Γ |- t :⤳*: u : A ] -> - [Γ ||- t : A | RA] × [Γ ||- t ≅ u : A | RA]. + [Γ ||- t ≅ v : A | RA]. Proof. - intros ? []; now eapply redSubstTerm. + intros ? []; now eapply redSubstLeftTmEq. Qed. - Lemma redFwd {Γ l A B} : [Γ ||- A] -> [Γ |- A :⤳*: B] -> whnf B -> [Γ ||- B]. Proof. intros RA; revert B; pattern l, Γ, A, RA; apply LR_rect_TyUr; clear l Γ A RA. - - intros ??? [??? red] ? red' ?. + - intros ??? [??? red] ? red' ?. eapply LRU_. unshelve erewrite (redtywf_det _ _ red' red); tea. 1: constructor. econstructor; tea. eapply redtywf_refl; gen_typing. - intros ??? [? red] ? red' ?. @@ -176,7 +195,7 @@ Proof. econstructor; tea. eapply redtywf_refl; gen_typing. - intros ??? [?? red] ?? ? red' ?. - eapply LRPi'. + eapply LRPi'. unshelve erewrite (redtywf_det _ _ red' red); tea. 1: constructor. econstructor. @@ -197,7 +216,7 @@ Proof. econstructor; tea. eapply redtywf_refl; gen_typing. - intros ??? [?? red] ?? ? red' ?. - eapply LRSig'. + eapply LRSig'. unshelve erewrite (redtywf_det _ _ red' red); tea. 1: constructor. econstructor. @@ -217,59 +236,140 @@ Proof. split; tea; irrelevance. Qed. -Lemma IdProp_whnf {Γ l A} (IA : [Γ ||-Id A]) t : IdProp IA t -> whnf t. -Proof. intros []; constructor; destruct r; now eapply convneu_whne. Qed. +Lemma redFwdURedTm {Γ l A t u} {h : [Γ ||-U A]} : + [Γ |- t :⤳*: u : A] -> + whnf u -> + URedTm (LogRelRec l) Γ t A h -> + URedTm (LogRelRec l) Γ u A h. +Proof. + intros redtu whu rt. + destruct rt as [nf rednf]; exists nf; tea. + unshelve erewrite (redtmwf_det _ _ redtu rednf); tea. + 1: now eapply isType_whnf. + eapply redtmwf_refl; gen_typing. +Defined. -Lemma redTmFwd {Γ l A t u} {RA : [Γ ||- A]} : - [Γ ||- t : A | RA] -> [Γ |- t :⤳*: u : A] -> whnf u -> [Γ ||- u : A | RA]. +Lemma redFwdPiRedTm {Γ l A t u} {h : [Γ ||-Π A]} : + [Γ |- t :⤳*: u : A] -> + whnf u -> + PiRedTm h t -> + PiRedTm h u. Proof. - revert t u; pattern l,Γ,A,RA; apply LR_rect_TyUr; clear l Γ A RA. - - intros ?????? [? red] red' ?. - unshelve erewrite (redtmwf_det _ _ red' red); tea. - 1: now eapply isType_whnf. - econstructor; tea. - 1: eapply redtmwf_refl; gen_typing. - eapply RedTyRecBwd. eapply redFwd. - 2,3: gen_typing. - now eapply RedTyRecFwd. - - intros ??? ??? [? red] red' ?. + intros redtu whu rt; pose proof (PiRedTmEq.whnf rt). + destruct rt as [nf rednf]; exists nf; tea. + unshelve erewrite (redtmwf_det _ _ redtu rednf); tea. + eapply redtmwf_refl; gen_typing. +Defined. + +Lemma redFwdSigRedTm {Γ l A t u} {h : [Γ ||-Σ A]} : + [Γ |- t :⤳*: u : A] -> + whnf u -> + SigRedTm h t -> + SigRedTm h u. +Proof. + intros redtu whu rt; pose proof (SigRedTmEq.whnf rt). + destruct rt as [nf rednf]; exists nf; tea. + unshelve erewrite (redtmwf_det _ _ redtu rednf); tea. + eapply redtmwf_refl; gen_typing. +Defined. + + +Lemma IdProp_whnf {Γ l A} (IA : [Γ ||-Id A]) t u : IdPropEq IA t u -> whnf t × whnf u. +Proof. + intros []; split; constructor. + 1: destruct r; now eapply convneu_whne. + destruct r as[?? conv]; symmetry in conv; now eapply convneu_whne. +Qed. + +Lemma redTmEqFwd {Γ l A t u v} {RA : [Γ ||- A]} : + [Γ ||- t ≅ v : A | RA] -> [Γ |- t :⤳*: u : A] -> whnf u -> [Γ ||- u ≅ v : A | RA]. +Proof. + intros tv tu wu; transitivity t; tea. + revert t u tv tu wu; pattern l,Γ,A,RA; apply LR_rect_TyUr; clear l Γ A RA. + - intros ?????? [red] red' ?. + pose proof (invLRConvU h). + assert [Γ |- t :⤳*: u] by (eapply redtywf_term; gen_typing). + assert (ru : [Γ ||- u]). + 1:eapply redFwd; tea; now eapply RedTyRecFwd. + exists (redFwdURedTm red' wu red) red (RedTyRecBwd _ ru); cbn; tea. + 1: now eapply lrefl. + eapply TyEqRecBwd. + eapply LRTyEqSym. + eapply redSubst in ru as []; [| now eapply tyr_wf_red ]. + irrelevance. + Unshelve. 2: now eapply RedTyRecFwd. + - intros ??? ??? [?? red] red' ?. unshelve erewrite (redtmwf_det _ _ red' red); tea. 1: constructor; now eapply convneu_whne. econstructor; tea. eapply redtmwf_refl; gen_typing. - - intros ???????? [? red] red' ?. + now eapply lrefl. + - intros ???????? [red] red' wu. + pose proof (PiRedTmEq.whnf red). + unshelve eexists (redFwdPiRedTm red' wu red) red. + 1: now eapply lrefl. + now eapply eqAppLRefl. + - intros ?????? Rt red' ?; inversion Rt as [???? red ?? prop] ; subst. unshelve erewrite (redtmwf_det _ _ red' red); tea. - 1: eapply isFun_whnf; destruct isfun; constructor; tea; now eapply convneu_whne. + 1: now eapply fst, NatPropEq_whnf. econstructor; tea. eapply redtmwf_refl; gen_typing. - - intros ?????? Rt red' ?; inversion Rt; subst. + 1: now eapply lrefl. + now eapply reflNatRedTmEq in prop as []. + - intros ?????? Rt red' ?; inversion Rt as [???? red ?? prop]; subst. unshelve erewrite (redtmwf_det _ _ red' red); tea. - 1: now eapply NatProp_whnf. + 1:now eapply fst, (EmptyPropEq_whnf (NA:=NA)). econstructor; tea. eapply redtmwf_refl; gen_typing. - - intros ?????? Rt red' ?; inversion Rt; subst. + 1: now eapply lrefl. + now eapply (reflEmptyRedTmEq (NA:=NA)) in prop as []. + - intros ???????? [red] red' ?. + pose proof (SigRedTmEq.whnf red). + unshelve eexists (redFwdSigRedTm red' wu red) red _. + all:cbn; intros; eapply lrefl; eauto. + irrelevanceRefl; unshelve eauto; tea. + - intros ???????? [?? red ?? prop] red' ?. unshelve erewrite (redtmwf_det _ _ red' red); tea. - 1: now eapply EmptyProp_whnf. - econstructor; tea. - eapply redtmwf_refl; gen_typing. - Unshelve. 2: tea. - - intros ???????? [? red] red' ?. - unshelve erewrite (redtmwf_det _ _ red' red); tea. - 1: eapply isPair_whnf; destruct ispair; constructor; tea; now eapply convneu_whne. - econstructor; tea. - eapply redtmwf_refl; gen_typing. - - intros ???????? [? red] red' ?. - unshelve erewrite (redtmwf_det _ _ red' red); tea. - 1: now eapply IdProp_whnf. + 1: now eapply fst, IdPropEq_whnf. econstructor; tea. eapply redtmwf_refl; gen_typing. + 1: now eapply lrefl. + now eapply fst, reflIdPropEq. Qed. -Lemma redTmFwdConv {Γ l A t u} {RA : [Γ ||- A]} : - [Γ ||- t : A | RA] -> [Γ |- t :⤳*: u : A] -> whnf u -> [Γ ||- u : A | RA] × [Γ ||- t ≅ u : A | RA]. + +Lemma redTmEqNf {Γ l A t u} {RA : [Γ ||- A]} : + [Γ ||- t ≅ u : A | RA] -> ∑ (v : term), [Γ |- t :⤳*: v : A] × whnf v. +Proof. + revert t u; pattern l,Γ,A,RA; apply LR_rect_TyUr; clear l Γ A RA. + - intros ??? h ?? [[]]; eexists; split; tea. + 2: now eapply isType_whnf. + eapply redtmwf_conv; tea. + destruct h; cbn in *; gen_typing. + - intros ??? h ?? []; eexists; split; tea. + 1: eapply redtmwf_conv; tea; destruct h; timeout 2 gen_typing. + constructor; now eapply convneu_whne. + - intros ??? h ???? [Rt]; pose proof (PiRedTmEq.whnf Rt); destruct Rt; cbn in *. + eexists; split; tea. + eapply redtmwf_conv; tea; destruct h; cbn in *; timeout 2 gen_typing. + - intros ??? h ?? [??????? []%NatPropEq_whnf]; eexists; split. + 1: eapply redtmwf_conv; tea; destruct h; cbn in *; eapply convty_exp; gen_typing. + assumption. + - intros ??? h ?? [??????? []%(EmptyPropEq_whnf (NA:=h))]; eexists; split. + 1: eapply redtmwf_conv; tea; destruct h; cbn in *; eapply convty_exp; gen_typing. + assumption. + - intros ??? h ???? [Rt]; pose proof (SigRedTmEq.whnf Rt); destruct Rt. + eexists; split; tea; cbn in *. + eapply redtmwf_conv; tea; destruct h; cbn in *; eapply convty_exp; gen_typing. + - intros ??? h ???? [????? [? _]%IdPropEq_whnf]; eexists; split; tea; cbn in *. + eapply redtmwf_conv; tea; destruct h; cbn in *; tea; unfold_id_outTy; cbn. + eapply convty_exp; gen_typing. + Qed. + +Lemma redTmEqFwdBoth {Γ l A t t' w w'} {RA : [Γ ||- A]} : + [Γ ||- t ≅ t' : A | RA] -> [Γ |- t :⤳*: w : A] -> [Γ |- t' :⤳*: w' : A] -> whnf w -> whnf w' -> [Γ ||- w ≅ w' : A | RA]. Proof. - intros Rt red wh. pose (Ru := redTmFwd Rt red wh). - destruct (redwfSubstTerm RA Ru red); now split. + intros; eapply redTmEqFwd; tea; symmetry; eapply redTmEqFwd; tea; now symmetry. Qed. End Reduction. diff --git a/theories/LogicalRelation/Reflexivity.v b/theories/LogicalRelation/Reflexivity.v index ae6c6ae..45ac10b 100644 --- a/theories/LogicalRelation/Reflexivity.v +++ b/theories/LogicalRelation/Reflexivity.v @@ -1,6 +1,6 @@ (** * LogRel.LogicalRelation.Reflexivity: reflexivity of the logical relation. *) From LogRel.AutoSubst Require Import core unscoped Ast Extra. -From LogRel Require Import Utils BasicAst Context NormalForms Weakening GenericTyping LogicalRelation. +From LogRel Require Import Utils BasicAst Notations Context NormalForms Weakening GenericTyping LogicalRelation. From LogRel.LogicalRelation Require Import Induction Escape. Set Universe Polymorphism. @@ -19,91 +19,93 @@ Section Reflexivities. (* Deprecated *) - Corollary LRTyEqRefl {l Γ A eqTy redTm eqTm} - (lr : LogRel l Γ A eqTy redTm eqTm) : eqTy A. + #[deprecated(note="Use reflLRTyEq")] + Corollary LRTyEqRefl {l Γ A eqTy eqTm} + (lr : LogRel l Γ A eqTy eqTm) : eqTy A. Proof. - pose (R := Build_LRPack Γ A eqTy redTm eqTm). + pose (R := Build_LRPack Γ A eqTy eqTm). pose (Rad := Build_LRAdequate (pack:=R) lr). change [Rad | _ ||- _ ≅ A ]; now eapply reflLRTyEq. Qed. - Lemma NeNfEqRefl {Γ k A} : [Γ ||-NeNf k : A] -> [Γ ||-NeNf k ≅ k : A]. - Proof. - intros []; now econstructor. - Qed. + Ltac eqrefl := etransitivity; tea; now symmetry. + + Lemma NeNfEqRefl {Γ k l A} : [Γ ||-NeNf k ≅ l : A] -> [Γ ||-NeNf k ≅ k : A] × [Γ ||-NeNf l ≅ l : A]. + Proof. intros []; split; econstructor; tea; eqrefl. Qed. + + Lemma NeNfEqLRefl {Γ k l A} : [Γ ||-NeNf k ≅ l : A] -> [Γ ||-NeNf k ≅ k : A]. + Proof. now intros []%NeNfEqRefl. Qed. + + Lemma NeNfEqRRefl {Γ k l A} : [Γ ||-NeNf k ≅ l : A] -> [Γ ||-NeNf l ≅ l : A]. + Proof. now intros []%NeNfEqRefl. Qed. Lemma reflNatRedTmEq {Γ A} {NA : [Γ ||-Nat A]} : - (forall t : term, [Γ ||-Nat t : A | NA] -> [Γ ||-Nat t ≅ t : A | NA]) - × (forall t : term, NatProp NA t -> NatPropEq NA t t). + (forall t u : term, [Γ ||-Nat t ≅ u : A | NA] -> [Γ ||-Nat t ≅ t : A | NA] × [Γ ||-Nat u ≅ u : A | NA]) + × (forall t u : term, NatPropEq NA t u -> NatPropEq NA t t × NatPropEq NA u u). Proof. - eapply NatRedInduction. - 1-3: now econstructor. - intros; econstructor. - now eapply NeNfEqRefl. + eapply NatRedEqInduction. + - intros * ???? [] ; split; econstructor; tea; eqrefl. + - split; now econstructor. + - intros * ? []; split; now econstructor. + - intros * []%NeNfEqRefl; split; now econstructor. Qed. Lemma reflEmptyRedTmEq {Γ A} {NA : [Γ ||-Empty A]} : - (forall t : term, [Γ ||-Empty t : A | NA] -> [Γ ||-Empty t ≅ t : A | NA]) - × (forall t : term, @EmptyProp _ _ _ Γ t -> @EmptyPropEq _ _ Γ t t). + (forall t u : term, [Γ ||-Empty t ≅ u : A | NA] -> [Γ ||-Empty t ≅ t : A | NA] × [Γ ||-Empty u ≅ u : A | NA]) + × (forall t u : term, EmptyPropEq Γ t u -> EmptyPropEq Γ t t × EmptyPropEq Γ u u). Proof. - split. - - intros t Ht. induction Ht. - econstructor; eauto. - destruct prop; econstructor. - now eapply NeNfEqRefl. - - intros ? []. econstructor. - now eapply NeNfEqRefl. + apply EmptyRedEqInduction. + - intros * ???? []; split; econstructor; tea; eqrefl. + - intros ?? []%NeNfEqRefl; split; now econstructor. Qed. - Lemma reflIdPropEq {Γ l A} (IA : [Γ ||-Id A]) t (Pt : IdProp IA t) : IdPropEq IA t t. + Lemma reflIdPropEq {Γ l A} (IA : [Γ ||-Id A]) t u + (Pt : IdPropEq IA t u) : IdPropEq IA t t × IdPropEq IA u u. Proof. - destruct Pt; constructor; tea; now eapply NeNfEqRefl. + destruct Pt as[|??[]%NeNfEqRefl]; split; now constructor. Qed. - Lemma reflIdRedTmEq {Γ l A} (IA : [Γ ||-Id A]) t (Rt : [Γ ||-Id t : _ | IA]) : [Γ ||-Id t ≅ t : _ | IA]. - Proof. destruct Rt; econstructor; tea; now eapply reflIdPropEq. Qed. + Lemma reflIdRedTmEq {Γ l A} (IA : [Γ ||-Id A]) t u (Rt : [Γ ||-Id t ≅ u : _ | IA]) : [Γ ||-Id t ≅ t : _ | IA] × [Γ ||-Id u ≅ u : _ | IA]. + Proof. destruct Rt as [????? []%reflIdPropEq] ; split; econstructor; tea; eqrefl. Qed. +(* Definition reflLRTmEq@{h i j k l} {l Γ A} (lr : [ LogRel@{i j k l} l | Γ ||- A ] ) : - forall t, - [ Γ ||- t : A | lr ] -> - [ Γ ||- t ≅ t : A | lr ]. + forall t u, + [ Γ ||- t ≅ u : A | lr ] -> + [ Γ ||- t ≅ t : A | lr ] × [ Γ ||- u ≅ u : A | lr ]. Proof. pattern l, Γ, A, lr; eapply LR_rect_TyUr; clear l Γ A lr; intros l Γ A. - - intros h t [? ? ? ? Rt%RedTyRecFwd@{h i j k}] ; cbn in *. - (* Need an additional universe level h < i *) - Unset Printing Notations. - pose proof (reflLRTyEq@{h i k j} Rt). - unshelve econstructor. - all : cbn. - 1-2: econstructor ; tea ; cbn. - 1-3,5: eapply RedTyRecBwd; tea. - 1: cbn; easy. - now eapply TyEqRecBwd. - - intros [] t []. - econstructor ; cbn in *. - all: eassumption. - - intros ??? t []. - unshelve econstructor ; cbn in *. - 1-2: now econstructor. - all: cbn; now eauto. + - intros h t u [? ? ? Rt Ru Rtu%TyEqRecFwd]; cbn in *. + pose proof (Rt' := reflLRTyEq@{h i k j} (RedTyRecFwd@{h i j k} _ Rt)). + pose proof (Ru' := reflLRTyEq@{h i k j} (RedTyRecFwd@{h i j k} _ Ru)). + split; unshelve econstructor; tea. + 1,3: eqrefl. + all: now eapply TyEqRecFwd. + - intros [] t u []. + split; econstructor ; cbn in *; tea; eqrefl. + - intros ??? t u []. + split; unshelve econstructor ; cbn in *; tea; try eqrefl. + all: apply PiRedTmEq.app. - intros; now apply reflNatRedTmEq. - intros; now apply reflEmptyRedTmEq. - - intros ??? t []. - unshelve econstructor ; cbn in *. - 1-2: now econstructor. - all: cbn; now eauto. + - intros ? ihdom ihcod t u [??? eqfst eqsnd]. + split; unshelve econstructor ; cbn in *; tea; try eqrefl. + all: intros Δ ρ h; pose proof (ihdom Δ ρ h _ _ (eqfst _ _ _)) as [hL hR]; tea. + (* Problem of irrelevance of the relation *) + 1,2: admit. - intros; now eapply reflIdRedTmEq. - Qed. + Admitted. + (* Qed. *) (* Deprecated *) - Corollary LRTmEqRefl@{h i j k l} {l Γ A eqTy redTm eqTm} (lr : LogRel@{i j k l} l Γ A eqTy redTm eqTm) : - forall t, redTm t -> eqTm t t. + Corollary LRTmEqRefl@{h i j k l} {l Γ A eqTy eqTm} (lr : LogRel@{i j k l} l Γ A eqTy eqTm) : + forall t u, eqTm t u -> eqTm t t × eqTm u u. Proof. - pose (R := Build_LRPack Γ A eqTy redTm eqTm). + pose (R := Build_LRPack Γ A eqTy eqTm). pose (Rad := Build_LRAdequate (pack:=R) lr). - intros t ?; change [Rad | _ ||- t ≅ t : _ ]; now eapply reflLRTmEq. + intros t u ?; change ([Rad | _ ||- t ≅ t : _ ] × [Rad | _ ||- u ≅ u : _ ]); now eapply reflLRTmEq. Qed. - +*) End Reflexivities. diff --git a/theories/LogicalRelation/ShapeView.v b/theories/LogicalRelation/ShapeView.v index 0e3f645..3b97008 100644 --- a/theories/LogicalRelation/ShapeView.v +++ b/theories/LogicalRelation/ShapeView.v @@ -14,8 +14,8 @@ Section ShapeViews. in the same way. *) Definition ShapeView@{i j k l i' j' k' l'} Γ - A {lA eqTyA redTmA redTyA} B {lB eqTyB redTmB redTyB} - (lrA : LogRel@{i j k l} lA Γ A eqTyA redTmA redTyA) (lrB : LogRel@{i' j' k' l'} lB Γ B eqTyB redTmB redTyB) : Set := + A {lA eqTyA redTyA} B {lB eqTyB redTyB} + (lrA : LogRel@{i j k l} lA Γ A eqTyA redTyA) (lrB : LogRel@{i' j' k' l'} lB Γ B eqTyB redTyB) : Set := match lrA, lrB with | LRU _ _, LRU _ _ => True | LRne _ _, LRne _ _ => True @@ -27,7 +27,7 @@ Section ShapeViews. | _, _ => False end. - Arguments ShapeView Γ A {lA eqTyA redTmA redTyA} B {lB eqTyB redTmB redTyB} + Arguments ShapeView Γ A {lA eqTyA redTyA} B {lB eqTyB redTyB} !lrA !lrB. (** ** The main property *) @@ -37,32 +37,32 @@ if two reducible types are reducibly convertible, then they must be reducible in This lets us relate different reducibility proofs when we have multiple such proofs, typically when showing symmetry or transitivity of the logical relation. *) - Arguments ShapeView Γ A {lA eqTyA redTmA redTyA} B {lB eqTyB redTmB redTyB} + Arguments ShapeView Γ A {lA eqTyA redTyA} B {lB eqTyB redTyB} !lrA !lrB. - Lemma red_whnf@{i j k l} {Γ A lA eqTyA redTmA eqTmA} - (lrA : LogRel@{i j k l} lA Γ A eqTyA redTmA eqTmA) : + Lemma red_whnf@{i j k l} {Γ A lA eqTyA eqTmA} + (lrA : LogRel@{i j k l} lA Γ A eqTyA eqTmA) : ∑ nf, [Γ |- A :⤳*: nf] × whnf nf. Proof. - pattern lA, Γ, A, eqTyA, redTmA, eqTmA, lrA; eapply LR_rect; intros ??[]. + pattern lA, Γ, A, eqTyA, eqTmA, lrA; eapply LR_rect; intros ??[]. all: eexists; split; tea; constructor; tea. now eapply convneu_whne. Defined. - Lemma eqTy_red_whnf@{i j k l} {Γ A lA eqTyA redTmA eqTmA B} - (lrA : LogRel@{i j k l} lA Γ A eqTyA redTmA eqTmA) : + Lemma eqTy_red_whnf@{i j k l} {Γ A lA eqTyA eqTmA B} + (lrA : LogRel@{i j k l} lA Γ A eqTyA eqTmA) : eqTyA B -> ∑ nf, [Γ |- B :⤳*: nf] × whnf nf. Proof. - pattern lA, Γ, A, eqTyA, redTmA, eqTmA, lrA. + pattern lA, Γ, A, eqTyA, eqTmA, lrA. eapply LR_rect_LogRelRec@{i j k l k}; intros ??? []. 3,6,7: intros ??. all: intros []; eexists; split; tea; constructor; tea. eapply convneu_whne; now symmetry. Defined. - Lemma ShapeViewConv@{i j k l i' j' k' l'} {Γ A lA eqTyA redTmA eqTmA B lB eqTyB redTmB eqTmB} - (lrA : LogRel@{i j k l} lA Γ A eqTyA redTmA eqTmA) (lrB : LogRel@{i' j' k' l'} lB Γ B eqTyB redTmB eqTmB) : + Lemma ShapeViewConv@{i j k l i' j' k' l'} {Γ A lA eqTyA eqTmA B lB eqTyB eqTmB} + (lrA : LogRel@{i j k l} lA Γ A eqTyA eqTmA) (lrB : LogRel@{i' j' k' l'} lB Γ B eqTyB eqTmB) : eqTyA B -> ShapeView@{i j k l i' j' k' l'} Γ A B lrA lrB. Proof. @@ -70,7 +70,7 @@ when showing symmetry or transitivity of the logical relation. *) pose (x := eqTy_red_whnf lrA eqAB). pose (y:= red_whnf lrB). pose proof (h := redtywf_det (snd x.π2) (snd y.π2) (fst x.π2) (fst y.π2)). - revert eqAB x y h. + revert eqAB x y h. destruct lrA; destruct lrB; intros []; cbn; try easy; try discriminate. all: try now (intros e; destruct neA as [? ? ne]; subst; apply convneu_whne in ne; inversion ne). all: try now (intros e; subst; symmetry in eq; apply convneu_whne in eq; inversion eq). diff --git a/theories/LogicalRelation/SimpleArr.v b/theories/LogicalRelation/SimpleArr.v index f59ed0f..2873b34 100644 --- a/theories/LogicalRelation/SimpleArr.v +++ b/theories/LogicalRelation/SimpleArr.v @@ -1,13 +1,13 @@ From LogRel.AutoSubst Require Import core unscoped Ast Extra. From LogRel Require Import Notations Utils BasicAst Context NormalForms Weakening GenericTyping LogicalRelation. -From LogRel.LogicalRelation Require Import Induction Irrelevance Weakening Neutral Escape Reflexivity NormalRed Reduction Transitivity Application. +From LogRel.LogicalRelation Require Import Induction Escape Irrelevance Weakening Neutral Escape Reflexivity NormalRed Reduction Transitivity Application InstKripke. Set Universe Polymorphism. Set Printing Primitive Projection Parameters. Section SimpleArrow. Context `{GenericTypingProperties}. - + Lemma shiftPolyRed {Γ l A B} : [Γ ||- A] -> [Γ ||- B] -> PolyRed Γ l A B⟨↑⟩. Proof. intros; escape; unshelve econstructor. @@ -64,7 +64,7 @@ Section SimpleArrow. + now eapply shiftPolyRedEq. Qed. - + Lemma arrRedCong {Γ l A A' B B'} (RA : [Γ ||- A]) (RB : [Γ ||- B]) (tyA' : [Γ |- A']) (tyB' : [Γ |- B']) : [Γ ||- A ≅ A' | RA] -> @@ -95,14 +95,14 @@ Section SimpleArrow. + intros. irrelevance0; rewrite subst_arr; [refold; reflexivity|]; refold. eapply arrRedCong. - 1,2: eapply escape; first [eapply RB| eapply RC]; now irrelevanceRefl. + 1,2: eapply escape; first [eapply RB| eapply RC]; symmetry; now irrelevanceRefl. 1,2: first [eapply RBext| eapply RCext]; now irrelevanceRefl. Unshelve. all: now irrelevanceRefl + tea. Qed. Lemma polyRedEqArrExt {Γ l A A' B B' C C'} - (PAB : PolyRed Γ l A B) (PAC : PolyRed Γ l A C) - (PAB' : PolyRed Γ l A' B') (PAC' : PolyRed Γ l A' C') + (PAB : PolyRed Γ l A B) (PAC : PolyRed Γ l A C) + (PAB' : PolyRed Γ l A' B') (PAC' : PolyRed Γ l A' C') (PABC : PolyRed Γ l A (arr B C)) (PABeq : PolyRedEq PAB A' B') (PACeq : PolyRedEq PAC A' C') @@ -113,75 +113,23 @@ Section SimpleArrow. + intros; irrelevance0; rewrite subst_arr; refold; [reflexivity|]. apply arrRedCong. * eapply escape; unshelve eapply (PolyRed.posRed); cycle 1; tea. - eapply LRTmRedConv; tea; irrelevanceRefl; now unshelve eapply (PolyRedEq.shpRed PABeq). + eapply LRTmEqConv ; tea; irrelevanceRefl; now unshelve eapply (PolyRedEq.shpRed PABeq). * eapply escape; unshelve eapply (PolyRed.posRed); cycle 1; tea. - eapply LRTmRedConv; tea. irrelevanceRefl; now unshelve eapply (PolyRedEq.shpRed PABeq). + eapply LRTmEqConv; tea; irrelevanceRefl; now unshelve eapply (PolyRedEq.shpRed PABeq). * unshelve eapply (PolyRedEq.posRed PABeq); tea; now irrelevanceRefl. * unshelve eapply (PolyRedEq.posRed PACeq); tea; now irrelevanceRefl. Qed. - Lemma idred {Γ l A} (RA : [Γ ||- A]) : - [Γ ||- idterm A : arr A A | ArrRedTy RA RA]. - Proof. - eassert [_ | Γ,, A ||- tRel 0 : A⟨_⟩] as hrel. - { - eapply var0. - 1: reflexivity. - now escape. - } - Unshelve. - all: cycle -1. - { - erewrite <- wk1_ren_on. - eapply wk ; tea. - escape. - gen_typing. - } - eapply reflLRTmEq in hrel. - normRedΠ ΠAA. - pose proof (reflLRTyEq RA). - escape. - assert (h : forall Δ a (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]) RA (ha : [Δ ||- a : A⟨ρ⟩ | RA]), - [Δ ||- tApp (idterm A)⟨ρ⟩ a : A⟨ρ⟩| RA] × - [Δ ||- tApp (idterm A)⟨ρ⟩ a ≅ a : A⟨ρ⟩| RA] - ). - 1:{ - intros; cbn; escape. - assert [Δ |- A⟨ρ⟩] by now eapply wft_wk. - assert [Δ |- A⟨ρ⟩ ≅ A⟨ρ⟩] by now eapply convty_wk. - eapply redSubstTerm; tea. - now eapply redtm_id_beta. - } - econstructor; cbn. - - eapply redtmwf_refl. - now eapply ty_id. - - constructor. - + intros; cbn; apply reflLRTyEq. - + intros; cbn. - irrelevance0; [|apply ha]. - cbn; bsimpl; now rewrite rinstInst'_term. - - eapply convtm_id; tea. - eapply wfc_wft; now escape. - - intros; cbn; irrelevance0. - 2: now eapply h. - bsimpl; now rewrite rinstInst'_term. - - intros ; cbn; irrelevance0; cycle 1. - 1: eapply transEqTerm;[|eapply transEqTerm]. - + now eapply h. - + eassumption. - + eapply LRTmEqSym. now eapply h. - + bsimpl; now rewrite rinstInst'_term. - Qed. - Section SimpleAppTerm. + (* Section SimpleAppTerm. Context {Γ t u F G l} {RF : [Γ ||- F]} (RG : [Γ ||- G]) (hΠ := normRedΠ0 (ArrRedTy0 RF RG)) (Rt : [Γ ||- t : arr F G | LRPi' hΠ]) - (Ru : [Γ ||- u : F | RF]). + (Ru : [Γ ||- u : F | RF]). *) - Lemma simple_app_id : [Γ ||- tApp (PiRedTm.nf Rt) u : G | RG ]. + (* Lemma simple_app_id : [Γ ||- tApp (PiRedTm.nf Rt) u : G | RG ]. Proof. irrelevance0. 2: now eapply app_id. @@ -205,39 +153,35 @@ Section SimpleArrow. Lemma simple_appTerm {Γ t u F G l} {RF : [Γ ||- F]} - (RG : [Γ ||- G]) + (RG : [Γ ||- G]) (RΠ : [Γ ||- arr F G]) (Rt : [Γ ||- t : arr F G | RΠ]) (Ru : [Γ ||- u : F | RF]) : [Γ ||- tApp t u : G| RG]. - Proof. + Proof. unshelve eapply simple_appTerm0. 3: irrelevance. all: tea. - Qed. + Qed. *) Lemma simple_appcongTerm {Γ t t' u u' F G l} {RF : [Γ ||- F]} (RG : [Γ ||- G]) - (RΠ : [Γ ||- arr F G]) + (RΠ : [Γ ||- arr F G]) (Rtt' : [Γ ||- t ≅ t' : _ | RΠ]) - (Ru : [Γ ||- u : F | RF]) - (Ru' : [Γ ||- u' : F | RF]) (Ruu' : [Γ ||- u ≅ u' : F | RF ]) : [Γ ||- tApp t u ≅ tApp t' u' : G | RG]. Proof. irrelevance0. - 2: eapply appcongTerm. - 2-5: tea. - now bsimpl. - Unshelve. tea. + 2: eapply appcongTerm; tea. now bsimpl. + Unshelve. now bsimpl. Qed. - Lemma wkRedArr {Γ l A B f} - (RA : [Γ ||- A]) - (RB : [Γ ||- B]) + (* Lemma wkRedArr {Γ l A B f} + (RA : [Γ ||- A]) + (RB : [Γ ||- B]) {Δ} (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]) : [Γ ||- f : arr A B | ArrRedTy RA RB] -> [Δ ||- f⟨ρ⟩ : arr A⟨ρ⟩ B⟨ρ⟩ | ArrRedTy (wk ρ wfΔ RA) (wk ρ wfΔ RB)]. @@ -247,16 +191,16 @@ Section SimpleArrow. symmetry; apply wk_arr. Qed. - Lemma compred {Γ l A B C f g} - (RA : [Γ ||- A]) - (RB : [Γ ||- B]) + Lemma compred {Γ l A B C f g} + (RA : [Γ ||- A]) + (RB : [Γ ||- B]) (RC : [Γ ||- C]) : [Γ ||- f : arr A B | ArrRedTy RA RB] -> [Γ ||- g : arr B C | ArrRedTy RB RC] -> [Γ ||- comp A g f : arr A C | ArrRedTy RA RC]. Proof. intros Rf Rg. - normRedΠin Rf; normRedΠin Rg; normRedΠ ΠAC. + normRedΠin Rf; normRedΠin Rg; normRedΠ ΠAC. assert (h : forall Δ a (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]) RA (ha : [Δ ||- a : A⟨ρ⟩ | RA]), [Δ ||- tApp g⟨ρ⟩ (tApp f⟨ρ⟩ a) : _ | wk ρ wfΔ RC]). 1:{ @@ -280,7 +224,7 @@ Section SimpleArrow. [Δ ||- tApp (comp A g f)⟨ρ⟩ a : _ | wk ρ wfΔ RC] × [Δ ||- tApp (comp A g f)⟨ρ⟩ a ≅ tApp g⟨ρ⟩ (tApp f⟨ρ⟩ a) : _ | wk ρ wfΔ RC]). 1:{ - intros; cbn. + intros; cbn. assert (eq : forall t : term, t⟨↑⟩⟨upRen_term_term ρ⟩ = t⟨ρ⟩⟨↑⟩) by (intros; now asimpl). do 2 rewrite eq. eapply redSubstTerm. @@ -325,6 +269,6 @@ Section SimpleArrow. + eapply LRTmEqSym. unshelve eapply beta; tea. + cbn; bsimpl; now rewrite rinstInst'_term. - Qed. + Qed. *) End SimpleArrow. diff --git a/theories/LogicalRelation/Transitivity.v b/theories/LogicalRelation/Transitivity.v index 369380c..788f0a8 100644 --- a/theories/LogicalRelation/Transitivity.v +++ b/theories/LogicalRelation/Transitivity.v @@ -12,7 +12,7 @@ Set Printing Primitive Projection Parameters. Set Printing Universes. - Lemma transEq@{i j k l} {Γ A B C lA lB} + Lemma transEq@{i j k l} {Γ A B C lA lB} {RA : [LogRel@{i j k l} lA | Γ ||- A]} {RB : [LogRel@{i j k l} lB | Γ ||- B]} (RAB : [Γ ||- A ≅ B | RA]) @@ -30,7 +30,7 @@ Proof. unshelve erewrite (redtmwf_det _ _ (URedTm.red redR) (URedTm.red redL0)) ; tea. all: apply isType_whnf; apply URedTm.type. + apply TyEqRecFwd; unshelve eapply transEq@{h i j k}. - 4,5: now apply (TyEqRecFwd h). + 4,5: now apply (TyEqRecFwd h). Qed. Lemma transEqTermNeu {Γ A t u v} {RA : [Γ ||-ne A]} : @@ -50,8 +50,8 @@ Lemma transEqTermΠ {Γ lA A t u v} {ΠA : [Γ ||-Π A]} [PolyRed.shpRed ΠA ρ h | Δ ||- t ≅ u : _] -> [PolyRed.shpRed ΠA ρ h | Δ ||- u ≅ v : _] -> [PolyRed.shpRed ΠA ρ h | Δ ||- t ≅ v : _]) - (ihcod : forall (Δ : context) (a : term) (ρ : Δ ≤ Γ) (h : [ |-[ ta ] Δ]) - (ha : [PolyRed.shpRed ΠA ρ h | Δ ||- a : _]) (t u v : term), + (ihcod : forall (Δ : context) (a b: term) (ρ : Δ ≤ Γ) (h : [ |-[ ta ] Δ]) + (ha : [PolyRed.shpRed ΠA ρ h | Δ ||- a ≅ b : _]) (t u v : term), [PolyRed.posRed ΠA ρ h ha | Δ ||- t ≅ u : _] -> [PolyRed.posRed ΠA ρ h ha | Δ ||- u ≅ v : _] -> [PolyRed.posRed ΠA ρ h ha | Δ ||- t ≅ v : _]) : @@ -60,16 +60,16 @@ Lemma transEqTermΠ {Γ lA A t u v} {ΠA : [Γ ||-Π A]} [Γ ||-Π t ≅ v : A | ΠA ]. Proof. intros [tL] [? tR]. - assert (forall t (red : [Γ ||-Π t : _ | ΠA]), whnf (PiRedTm.nf red)). - { intros ? [? ? isfun]; simpl; destruct isfun; constructor; tea. - now eapply convneu_whne. } - unshelve epose proof (e := redtmwf_det _ _ (PiRedTm.red redR) (PiRedTm.red redL)); tea. - 1,2: now eauto. + unshelve epose proof (e := redtmwf_det _ _ (PiRedTmEq.red redR) (PiRedTmEq.red redL)); tea. + 1,2: now apply PiRedTmEq.whnf. exists tL tR. + etransitivity; tea. now rewrite e. - + intros. eapply ihcod. - 1: eapply eqApp. - rewrite e; apply eqApp0. + + cbn in *; intros. eapply ihcod. + 2: eapply eqApp0. + irrelevanceRefl. + rewrite <- e. + unshelve apply eqApp; [assumption|]. + eapply ihdom; [| eapply LRTmEqSym]; eassumption. Qed. Lemma transEqTermΣ {Γ lA A t u v} {ΣA : [Γ ||-Σ A]} @@ -77,8 +77,8 @@ Lemma transEqTermΣ {Γ lA A t u v} {ΣA : [Γ ||-Σ A]} [PolyRed.shpRed ΣA ρ h | Δ ||- t ≅ u : _] -> [PolyRed.shpRed ΣA ρ h | Δ ||- u ≅ v : _] -> [PolyRed.shpRed ΣA ρ h | Δ ||- t ≅ v : _]) - (ihcod : forall (Δ : context) (a : term) (ρ : Δ ≤ Γ) (h : [ |-[ ta ] Δ]) - (ha : [PolyRed.shpRed ΣA ρ h | Δ ||- a : _]) (t u v : term), + (ihcod : forall (Δ : context) (a b : term) (ρ : Δ ≤ Γ) (h : [ |-[ ta ] Δ]) + (ha : [PolyRed.shpRed ΣA ρ h | Δ ||- a ≅ b : _]) (t u v : term), [PolyRed.posRed ΣA ρ h ha | Δ ||- t ≅ u : _] -> [PolyRed.posRed ΣA ρ h ha | Δ ||- u ≅ v : _] -> [PolyRed.posRed ΣA ρ h ha | Δ ||- t ≅ v : _]) : @@ -87,23 +87,20 @@ Lemma transEqTermΣ {Γ lA A t u v} {ΣA : [Γ ||-Σ A]} [Γ ||-Σ t ≅ v : A | ΣA ]. Proof. intros [tL ?? eqfst eqsnd] [? tR ? eqfst' eqsnd']. - assert (forall t (red : [Γ ||-Σ t : _ | ΣA]), whnf (SigRedTm.nf red)). - { intros ? [? ? ispair]; simpl; destruct ispair; constructor; tea. - now eapply convneu_whne. } - unshelve epose proof (e := redtmwf_det _ _ (SigRedTm.red redR) (SigRedTm.red redL)); tea. - 1,2: now eauto. - exists tL tR. - + etransitivity; tea. now rewrite e. - + intros; eapply ihdom ; [eapply eqfst| rewrite e; eapply eqfst']. - + intros; eapply ihcod; [eapply eqsnd|]. - rewrite e. - eapply LRTmEqRedConv. + unshelve epose proof (e := redtmwf_det _ _ (SigRedTmEq.red redR) (SigRedTmEq.red redL)); tea. + 1,2: now apply SigRedTmEq.whnf. + induction e. + unshelve eexists tL tR _. + + intros; eapply ihdom ; [eapply eqfst| eapply eqfst']. + + etransitivity; tea. + + intros; eapply ihcod. + 1: irrelevanceRefl; eapply eqsnd. + eapply LRTmEqConv. 2: eapply eqsnd'. - eapply PolyRed.posExt. - 1: eapply (SigRedTm.fstRed tL). - eapply LRTmEqSym. rewrite <- e. - eapply eqfst. - Unshelve. tea. + irrelevanceRefl; unshelve eapply PolyRed.posExt. + 4: eapply LRTmEqSym, eqfst. + assumption. + Unshelve. all:tea. Qed. @@ -116,9 +113,9 @@ Proof. Qed. Lemma transEqTermNat {Γ A} (NA : [Γ ||-Nat A]) : - (forall t u, + (forall t u, [Γ ||-Nat t ≅ u : A | NA] -> forall v, - [Γ ||-Nat u ≅ v : A | NA] -> + [Γ ||-Nat u ≅ v : A | NA] -> [Γ ||-Nat t ≅ v : A | NA]) × (forall t u, NatPropEq NA t u -> forall v, @@ -127,7 +124,7 @@ Lemma transEqTermNat {Γ A} (NA : [Γ ||-Nat A]) : Proof. apply NatRedEqInduction. - intros * ???? ih ? uv; inversion uv; subst. - destruct (NatPropEq_whnf prop), (NatPropEq_whnf prop0). + destruct (NatPropEq_whnf prop), (NatPropEq_whnf prop0). unshelve epose proof (redtmwf_det _ _ redR redL0); tea; subst. econstructor; tea. 1: now etransitivity. @@ -148,9 +145,9 @@ Proof. Qed. Lemma transEqTermEmpty {Γ A} (NA : [Γ ||-Empty A]) : - (forall t u, + (forall t u, [Γ ||-Empty t ≅ u : A | NA] -> forall v, - [Γ ||-Empty u ≅ v : A | NA] -> + [Γ ||-Empty u ≅ v : A | NA] -> [Γ ||-Empty t ≅ v : A | NA]) × (forall t u, EmptyPropEq Γ t u -> forall v, @@ -200,12 +197,12 @@ Proof. Qed. -Lemma transEqTerm@{h i j k l} {Γ lA A t u v} +Lemma transEqTerm@{h i j k l} {Γ lA A t u v} {RA : [LogRel@{i j k l} lA | Γ ||- A]} : [Γ ||- t ≅ u : A | RA] -> [Γ ||- u ≅ v : A | RA] -> [Γ ||- t ≅ v : A | RA]. -Proof. +Proof. revert t u v; pattern lA, Γ, A, RA; apply LR_rect_TyUr; clear lA Γ A RA; intros l Γ. - intros *; apply transEqTermU@{h i j k}. - intros *; apply transEqTermNeu. @@ -217,35 +214,35 @@ Proof. Qed. -#[global] -Instance perLRTmEq@{i j k l} {Γ l A} (RA : [LogRel@{i j k l} l | Γ ||- A]): - PER (fun t u => [RA | _ ||- t ≅ u : _]). -Proof. - econstructor. - - intros ???; now eapply LRTmEqSym. - - intros ???; now eapply transEqTerm. -Qed. - Lemma LREqTermSymConv {Γ t u G G' l RG RG'} : - [Γ ||- t ≅ u : G | RG] -> + [Γ ||- t ≅ u : G | RG] -> [Γ ||- G' ≅ G | RG'] -> [Γ ||- u ≅ t : G' | RG']. Proof. intros Rtu RGG'. - eapply LRTmEqSym; eapply LRTmEqRedConv; tea. + eapply LRTmEqSym; eapply LRTmEqConv; tea. now eapply LRTyEqSym. -Qed. +Qed. Lemma LREqTermHelper {Γ t t' u u' G G' l RG RG'} : - [Γ ||- t ≅ u : G | RG] -> - [Γ ||- t' ≅ u' : G' | RG'] -> + [Γ ||- t ≅ u : G | RG] -> + [Γ ||- t' ≅ u' : G' | RG'] -> [Γ ||- G ≅ G' | RG] -> - [Γ ||- u ≅ u' : G | RG] -> + [Γ ||- u ≅ u' : G | RG] -> [Γ ||- t ≅ t' : G | RG]. Proof. intros Rtu Rtu' RGG' Ruu'. do 2 (eapply transEqTerm; tea). now eapply LREqTermSymConv. -Qed. +Qed. End Transitivity. + +#[global] +Instance LRTmEqTransitive `{GenericTypingProperties} {Γ A l} (RA : [Γ ||- A]): Transitive (RA.(LRPack.eqTm)). +Proof. intros x y z ; apply transEqTerm. Defined. + +#[global] +Instance LRTmEqPER@{i j k l} `{GenericTypingProperties} {Γ l A} (RA : [LogRel@{i j k l} l | Γ ||- A]): + PER (RA.(LRPack.eqTm)). +Proof. econstructor; typeclasses eauto. Qed. diff --git a/theories/LogicalRelation/Universe.v b/theories/LogicalRelation/Universe.v index 53f0ec7..e6db959 100644 --- a/theories/LogicalRelation/Universe.v +++ b/theories/LogicalRelation/Universe.v @@ -8,20 +8,25 @@ Set Printing Universes. Section UniverseReducibility. Context `{GenericTypingProperties}. + Lemma redUOneCtx {Γ} : [|- Γ] -> [Γ ||-U U]. + Proof. + intros ; econstructor; [easy| gen_typing|eapply redtywf_refl; gen_typing]. + Defined. + Lemma redUOne {Γ l A} : [Γ ||- A] -> [Γ ||-U U]. Proof. - intros ?%escape; econstructor; [easy| gen_typing|eapply redtywf_refl; gen_typing]. + intros ?%escape; eapply redUOneCtx; gen_typing. Qed. - Lemma UnivEq'@{i j k l} {Γ A l} (rU : [ LogRel@{i j k l} l | Γ ||- U ]) (rA : [ LogRel@{i j k l} l | Γ ||- A : U | rU]) + Lemma UnivEq'@{i j k l} {Γ A B l} (rU : [ LogRel@{i j k l} l | Γ ||- U ]) (rA : [ LogRel@{i j k l} l | Γ ||- A ≅ B : U | rU]) : [ LogRel@{i j k l} zero | Γ ||- A]. Proof. - assert [ LogRel@{i j k l} one | Γ ||- A : U | LRU_@{i j k l} (redUOne rU)] - as [ _ _ _ _ hA ] by irrelevance. + assert [ LogRel@{i j k l} one | Γ ||- A ≅ B : U | LRU_@{i j k l} (redUOne rU)] + as [ ? ?? hA ] by irrelevance. exact (LRCumulative hA). Qed. - Lemma UnivEq@{i j k l} {Γ A l} l' (rU : [ LogRel@{i j k l} l | Γ ||- U]) (rA : [ LogRel@{i j k l} l | Γ ||- A : U | rU]) + Lemma UnivEq@{i j k l} {Γ A B l} l' (rU : [ LogRel@{i j k l} l | Γ ||- U]) (rA : [ LogRel@{i j k l} l | Γ ||- A ≅ B : U | rU]) : [ LogRel@{i j k l} l' | Γ ||- A]. Proof. destruct l'. @@ -38,30 +43,4 @@ Section UniverseReducibility. eapply LRTyEqIrrelevantCum. exact hAB. Qed. - - (* The lemmas below does not seem to be - at the right levels for fundamental lemma ! *) - - Set Printing Universes. - Lemma univTmTy'@{h i j k l} {Γ l UU A} (h : [Γ ||-U UU ]) : - [LogRel@{i j k l} l | Γ ||- A : UU | LRU_ h] -> [LogRel@{h i j k} (URedTy.level h) | Γ ||- A]. - Proof. intros []; now eapply RedTyRecFwd. Qed. - - Lemma univTmTy@{h i j k l} {Γ l A} (RU : [Γ ||- U]) : - [LogRel@{i j k l} l | Γ ||- A : U | RU] -> [LogRel@{h i j k} (URedTy.level (invLRU RU)) | Γ ||- A]. - Proof. - intros h; apply univTmTy'. - irrelevance. - Qed. - - Lemma univEqTmEqTy'@{h i j k l} {Γ l l' UU A B} (h : [Γ ||-U UU]) (RA : [Γ ||- A]) : - [LogRel@{i j k l} l | Γ ||- A ≅ B : UU | LRU_ h] -> - [LogRel@{h i j k} l' | Γ ||- A ≅ B | RA]. - Proof. intros [????? RA']. apply TyEqRecFwd in RA'. irrelevance. Qed. - - Lemma univEqTmEqTy@{h i j k l} {Γ l l' A B} (RU : [Γ ||- U]) (RA : [Γ ||- A]) : - [LogRel@{i j k l} l | Γ ||- A ≅ B : U | RU] -> - [LogRel@{h i j k} l' | Γ ||- A ≅ B | RA]. - Proof. intros h; eapply (univEqTmEqTy' (invLRU RU)); irrelevance. Qed. - End UniverseReducibility. diff --git a/theories/LogicalRelation/Weakening.v b/theories/LogicalRelation/Weakening.v index 2bb3a25..45f5c09 100644 --- a/theories/LogicalRelation/Weakening.v +++ b/theories/LogicalRelation/Weakening.v @@ -11,21 +11,21 @@ Section Weakenings. Lemma wkU {Γ Δ l A} (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]) (h : [Γ ||-U A]) : [Δ ||-U A⟨ρ⟩]. Proof. destruct h; econstructor; tea; change U with U⟨ρ⟩; gen_typing. Defined. - Lemma wkPoly {Γ l shp pos Δ} (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]) : - PolyRed Γ l shp pos -> + Lemma wkPoly {Γ l shp pos Δ} (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]) : + PolyRed Γ l shp pos -> PolyRed Δ l shp⟨ρ⟩ pos⟨wk_up shp ρ⟩. Proof. intros []; opector. - intros ? ρ' ?; replace (_⟨_⟩) with (shp⟨ρ' ∘w ρ⟩) by now bsimpl. now eapply shpRed. - - intros ? a ρ' **. + - intros ? a b ρ' **. replace (pos⟨_⟩[a .: ρ' >> tRel]) with (pos[ a .: (ρ' ∘w ρ) >> tRel]) by now bsimpl. econstructor; unshelve eapply posRed; tea; irrelevance. - now eapply wft_wk. - eapply wft_wk; tea; eapply wfc_cons; tea; now eapply wft_wk. - intros ? a b ρ' wfΔ' **. replace (_[b .: ρ' >> tRel]) with (pos[ b .: (ρ' ∘w ρ) >> tRel]) by (now bsimpl). - unshelve epose (posExt _ a b (ρ' ∘w ρ) wfΔ' _ _ _); irrelevance. + unshelve epose (posExt _ a b (ρ' ∘w ρ) wfΔ' _); irrelevance. Qed. Lemma wkΠ {Γ Δ A l} (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]) (ΠA : [Γ ||-Π< l > A]) : @@ -47,45 +47,38 @@ Section Weakenings. Defined. Lemma wkNat {Γ A Δ} (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]) : [Γ ||-Nat A] -> [Δ ||-Nat A⟨ρ⟩]. - Proof. + Proof. intros []; constructor. change tNat with tNat⟨ρ⟩. - gen_typing. + gen_typing. Qed. Lemma wkEmpty {Γ A Δ} (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]) : [Γ ||-Empty A] -> [Δ ||-Empty A⟨ρ⟩]. - Proof. + Proof. intros []; constructor. change tEmpty with tEmpty⟨ρ⟩. - gen_typing. + gen_typing. Qed. Lemma wkId@{i j k l} {Γ l A Δ} (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]) : IdRedTy@{i j k l} Γ l A -> IdRedTy@{i j k l} Δ l A⟨ρ⟩. (* [Γ ||-Id A] -> [Δ ||-Id A⟨ρ⟩]. *) - Proof. + Proof. intros []; unshelve econstructor. 6: erewrite wk_Id; now eapply redtywf_wk. 3: rewrite wk_Id; gen_typing. - now apply tyKripke. - intros; rewrite wk_comp_ren_on; now apply tyKripke. - - unshelve eapply tyKripkeTm; [eapply wk_id| gen_typing| now rewrite wk_comp_runit| irrelevance]. - - unshelve eapply tyKripkeTm; [eapply wk_id| gen_typing| now rewrite wk_comp_runit| irrelevance]. (* could also employ reflexivity instead *) - unshelve eapply tyKripkeTmEq; [eapply wk_id| gen_typing| now rewrite wk_comp_runit|irrelevance]. - unshelve eapply tyKripkeTmEq; [eapply wk_id| gen_typing| now rewrite wk_comp_runit|irrelevance]. - - apply perLRTmEq. - - intros; irrelevance0. + - apply LRTmEqPER. + - intros; irrelevance0. 1: now rewrite wk_comp_ren_on. unshelve eapply tyKripkeEq; tea. 3: irrelevance; now rewrite wk_comp_ren_on. bsimpl; setoid_rewrite H10; now bsimpl. - - intros; irrelevance0. - 1: now rewrite wk_comp_ren_on. - unshelve eapply tyKripkeTm; tea. - 3: irrelevance; now rewrite wk_comp_ren_on. - bsimpl; setoid_rewrite H10; now bsimpl. - - intros; irrelevance0. + - intros; irrelevance0. 1: now rewrite wk_comp_ren_on. unshelve eapply tyKripkeTmEq; tea. 3: irrelevance; now rewrite wk_comp_ren_on. @@ -115,7 +108,7 @@ Section Weakenings. Lemma wkΠ_eq {Γ Δ A l} (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]) (ΠA : [Γ ||-Π< l > A]) : wk ρ wfΔ (LRPi' ΠA) = LRPi' (wkΠ ρ wfΔ ΠA). Proof. reflexivity. Qed. - + #[local] Lemma wkΣ_eq {Γ Δ A l} (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]) (ΠA : [Γ ||-Σ< l > A]) : wk ρ wfΔ (LRSig' ΠA) = LRSig' (wkΣ ρ wfΔ ΠA). @@ -123,13 +116,13 @@ Section Weakenings. Set Printing Primitive Projection Parameters. - Lemma wkPolyEq {Γ l shp shp' pos pos' Δ} (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]) (PA : PolyRed Γ l shp pos) : + Lemma wkPolyEq {Γ l shp shp' pos pos' Δ} (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]) (PA : PolyRed Γ l shp pos) : PolyRedEq PA shp' pos' -> PolyRedEq (wkPoly ρ wfΔ PA) shp'⟨ρ⟩ pos'⟨wk_up shp' ρ⟩. Proof. intros []; opector. - intros ? ρ' wfΔ'; replace (_⟨_⟩) with (shp'⟨ρ' ∘w ρ⟩) by now bsimpl. pose (shpRed _ (ρ' ∘w ρ) wfΔ'); irrelevance. - - intros ?? ρ' wfΔ' ha. + - intros ??? ρ' wfΔ' ha. replace (_[_ .: ρ' >> tRel]) with (pos'[ a .: (ρ' ∘w ρ) >> tRel]) by now bsimpl. irrelevance0. 2: unshelve eapply posRed; tea; irrelevance. @@ -137,7 +130,7 @@ Section Weakenings. Qed. - Lemma wkEq@{i j k l} {Γ Δ A B l} (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]) (lrA : [Γ ||- A]) : + Lemma wkEq@{i j k l} {Γ Δ A B l} (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]) (lrA : [Γ ||- A]) : [LogRel@{i j k l} l | Γ ||- A ≅ B | lrA] -> [LogRel@{i j k l} l | Δ ||- A⟨ρ⟩ ≅ B⟨ρ⟩ | wk ρ wfΔ lrA]. Proof. @@ -183,11 +176,12 @@ Section Weakenings. irrelevance0; [apply eq|]. rewrite <- eq. now unshelve apply Hdom. - + intros Ξ a ρ' wfΞ *; cbn. + + intros Ξ a b ρ' wfΞ *; cbn. assert (eq : forall t, t⟨ρ' ∘w ρ⟩ = t⟨ρ⟩⟨ρ'⟩) by now bsimpl. - unshelve eassert (Ht0 := Ht Ξ a (ρ' ∘w ρ) wfΞ _). + unshelve eassert (Ht0 := Ht Ξ a b (ρ' ∘w ρ) wfΞ _). { cbn in ha; irrelevance0; [symmetry; apply eq|tea]. } replace (t'⟨upRen_term_term ρ⟩[a .: ρ' >> tRel]) with (t'[a .: (ρ' ∘w ρ) >> tRel]) by now bsimpl. + replace (t'⟨upRen_term_term ρ⟩[b .: ρ' >> tRel]) with (t'[b .: (ρ' ∘w ρ) >> tRel]) by now bsimpl. irrelevance0; [|apply Ht0]. now bsimpl. + change [Δ |- f⟨ρ⟩ ~ f⟨ρ⟩ : (tProd (PiRedTy.dom ΠA) (PiRedTy.cod ΠA))⟨ρ⟩]. @@ -195,34 +189,16 @@ Section Weakenings. Qed. (* TODO: use program or equivalent to have only the first field non-opaque *) - Lemma wkΠTerm {Γ Δ u A l} (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]) (ΠA : [Γ ||-Π< l > A]) - (ΠA' := wkΠ ρ wfΔ ΠA) : - [Γ||-Π u : A | ΠA] -> - [Δ ||-Π u⟨ρ⟩ : A⟨ρ⟩ | ΠA' ]. + Lemma wkΠTerm {Γ Δ u A l} (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]) (ΠA : [Γ ||-Π< l > A]) + (ΠA' := wkΠ ρ wfΔ ΠA) : PiRedTm ΠA u -> PiRedTm ΠA' u⟨ρ⟩. Proof. intros [t]. exists (t⟨ρ⟩); try change (tProd _ _) with (ΠA.(outTy)⟨ρ⟩). + now eapply redtmwf_wk. + now apply isLRFun_ren. - + now apply convtm_wk. - + intros ? a ρ' ??. - replace ((t ⟨ρ⟩)⟨ ρ' ⟩) with (t⟨ρ' ∘w ρ⟩) by now bsimpl. - irrelevance0. - 2: unshelve apply app; [eassumption|]; subst ΠA'; irrelevance. - subst ΠA'; bsimpl; try rewrite scons_comp'; reflexivity. - + intros ??? ρ' ?????. - replace ((t ⟨ρ⟩)⟨ ρ' ⟩) with (t⟨ρ' ∘w ρ⟩) by now bsimpl. - irrelevance0. - 2: unshelve apply eq; [eassumption|..]; subst ΠA'; irrelevance. - subst ΠA'; bsimpl; try rewrite scons_comp'; reflexivity. Defined. - Lemma wkNeNf {Γ Δ k A} (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]) : - [Γ ||-NeNf k : A] -> [Δ ||-NeNf k⟨ρ⟩ : A⟨ρ⟩]. - Proof. - intros []; constructor. all: gen_typing. - Qed. - + Lemma isLRPair_ren : forall Γ Δ t A l (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]) (ΣA : [Γ ||-Σ< l > A]), isLRPair ΣA t -> isLRPair (wkΣ ρ wfΔ ΣA) t⟨ρ⟩. Proof. @@ -236,46 +212,37 @@ Section Weakenings. irrelevance0; [apply eq|]. rewrite <- eq. now unshelve apply Hdom. - + intros Ξ a' ρ' wfΞ ha'; cbn. + + intros Ξ a' b' ρ' wfΞ ha'; cbn. assert (eq : forall t, t⟨ρ' ∘w ρ⟩ = t⟨ρ⟩⟨ρ'⟩) by now bsimpl. - unshelve eassert (Hcod0 := Hcod Ξ a' (ρ' ∘w ρ) wfΞ _). + unshelve eassert (Hcod0 := Hcod Ξ a' b' (ρ' ∘w ρ) wfΞ _). { cbn in ha'; irrelevance0; [symmetry; apply eq|tea]. } - replace (B'⟨upRen_term_term ρ⟩[a' .: ρ' >> tRel]) with B'[a' .: (ρ' ∘w ρ) >> tRel] by now bsimpl. + replace (B'⟨upRen_term_term ρ⟩[b' .: ρ' >> tRel]) with B'[b' .: (ρ' ∘w ρ) >> tRel] by now bsimpl. irrelevance0; [|apply Hcod0]. now bsimpl. + refold; intros Ξ ρ' wfΞ. assert (eq : forall t, t⟨ρ' ∘w ρ⟩ = t⟨ρ⟩⟨ρ'⟩) by now bsimpl. - rewrite <- eq. - irrelevance0; [|now unshelve apply Hsnd]. + irrelevance0. + 2:rewrite <- eq; now unshelve apply Hsnd. now bsimpl. + change [Δ |- p⟨ρ⟩ ~ p⟨ρ⟩ : (tSig (SigRedTy.dom ΣA) (SigRedTy.cod ΣA))⟨ρ⟩]. now eapply convneu_wk. Qed. - Lemma wkΣTerm {Γ Δ u A l} (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]) (ΠA : [Γ ||-Σ< l > A]) - (ΠA' := wkΣ ρ wfΔ ΠA) : - [Γ||-Σ u : A | ΠA] -> - [Δ ||-Σ u⟨ρ⟩ : A⟨ρ⟩ | ΠA' ]. + Lemma wkΣTerm {Γ Δ u A l} (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]) (ΠA : [Γ ||-Σ< l > A]) + (ΠA' := wkΣ ρ wfΔ ΠA) : SigRedTm ΠA u -> SigRedTm ΠA' u⟨ρ⟩. Proof. - intros [t]. - unshelve eexists (t⟨ρ⟩) _; try (cbn; rewrite wk_sig). - + intros ? ρ' wfΔ'; rewrite wk_comp_ren_on; irrelevance0. - 2: now unshelve eapply fstRed. - cbn; symmetry; apply wk_comp_ren_on. + intros [t]. + unshelve eexists (t⟨ρ⟩); try (cbn; rewrite wk_sig). + now eapply redtmwf_wk. + apply isLRPair_ren; assumption. - + eapply convtm_wk; eassumption. - + intros ? ρ' ?; irrelevance0. - 2: rewrite wk_comp_ren_on; now unshelve eapply sndRed. - rewrite <- wk_comp_ren_on; cbn; now rewrite <- wk_up_ren_subst. Defined. - Lemma wkTerm {Γ Δ t A l} (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]) (lrA : [Γ ||- A]) : + (* Lemma wkTerm {Γ Δ t A l} (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]) (lrA : [Γ ||- A]) : [Γ ||- t : A | lrA] -> [Δ ||- t⟨ρ⟩ : A⟨ρ⟩ | wk ρ wfΔ lrA]. Proof. revert t Δ ρ wfΔ. pattern l, Γ, A, lrA. eapply LR_rect_TyUr; clear l Γ A lrA. - - intros ?????? ρ ? [te]; exists te⟨ρ⟩; try change U with U⟨ρ⟩. + - intros ?????? ρ ? [te]; unshelve eexists te⟨ρ⟩ _ _; try change U with U⟨ρ⟩. + gen_typing. + apply isType_ren; assumption. + now apply convtm_wk. @@ -283,14 +250,14 @@ Section Weakenings. - intros ?????? ρ ? [te]. exists te⟨ρ⟩; cbn. + now eapply redtmwf_wk. + apply convneu_wk; assumption. - - intros; now apply wkΠTerm. + - intros; now apply wkΠTerm. - intros??? NA t ? ρ wfΔ; revert t; pose (NA' := wkNat ρ wfΔ NA). set (G := _); enough (h : G × (forall t, NatProp NA t -> NatProp NA' t⟨ρ⟩)) by apply h. subst G; apply NatRedInduction. + intros; econstructor; tea; change tNat with tNat⟨ρ⟩; gen_typing. + constructor. + now constructor. - + intros; constructor. + + intros; constructor. change tNat with tNat⟨ρ⟩. now eapply wkNeNf. - intros??? NA t ? ρ wfΔ; revert t; pose (NA' := wkEmpty ρ wfΔ NA). @@ -306,7 +273,7 @@ Section Weakenings. + destruct prop. econstructor. change tEmpty with tEmpty⟨ρ⟩. now eapply wkNeNf. - - intros; now apply wkΣTerm. + - intros; now apply wkΣTerm. - intros * _ _ * [??? prop]; econstructor; unfold_id_outTy; cbn; rewrite ?wk_Id. 1: now eapply redtmwf_wk. 1: now eapply convtm_wk. @@ -318,26 +285,23 @@ Section Weakenings. 2,3: eapply IA.(IdRedTy.tyKripkeTmEq); [now rewrite wk_comp_runit| irrelevance]. eapply IA.(IdRedTy.tyKripkeEq); [now rewrite wk_comp_runit| irrelevance]. Unshelve. all: tea. - Qed. + Qed. *) Lemma wkUTerm {Γ Δ l A t} (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]) (h : [Γ ||-U A]) : - [LogRelRec l| Γ ||-U t : A | h ] -> [LogRelRec l | Δ||-U t⟨ρ⟩ : A⟨ρ⟩ | wkU ρ wfΔ h]. + URedTm (LogRelRec l) Γ t A h -> URedTm (LogRelRec l) Δ t⟨ρ⟩ A⟨ρ⟩ (wkU ρ wfΔ h). Proof. intros [te]. exists te⟨ρ⟩; change U with U⟨ρ⟩. - gen_typing. - apply isType_ren; assumption. - - now apply convtm_wk. - - destruct l; [destruct (elim (URedTy.lt h))|cbn]. - eapply (wk (l:=zero)); eassumption. Defined. - Lemma wkNeNfEq {Γ Δ k k' A} (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]) : + Lemma wkNeNfEq {Γ Δ k k' A} (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]) : [Γ ||-NeNf k ≅ k' : A] -> [Δ ||-NeNf k⟨ρ⟩ ≅ k'⟨ρ⟩ : A⟨ρ⟩]. Proof. - intros []; constructor. gen_typing. - Qed. + intros []; constructor; gen_typing. + Qed. - Lemma wkTermEq {Γ Δ t u A l} (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]) (lrA : [Γ ||- A]) : + Lemma wkTermEq {Γ Δ t u A l} (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]) (lrA : [Γ ||- A]) : [Γ ||- t ≅ u : A | lrA] -> [Δ ||- t⟨ρ⟩ ≅ u⟨ρ⟩: A⟨ρ⟩ | wk ρ wfΔ lrA]. Proof. revert t u Δ ρ wfΔ. pattern l, Γ, A, lrA. @@ -347,7 +311,7 @@ Section Weakenings. + exact (wkUTerm ρ wfΔ h rL). + exact (wkUTerm ρ wfΔ h rR). + apply RedTyRecBwd; apply wk; [assumption|]; now apply (RedTyRecFwd h). - + cbn. change U with U⟨ρ⟩. + + cbn. change U with U⟨ρ⟩. now eapply convtm_wk. + apply RedTyRecBwd; apply wk; [assumption|]; now apply (RedTyRecFwd h). + apply TyEqRecBwd. eapply wkEq. now apply TyEqRecFwd. @@ -355,7 +319,7 @@ Section Weakenings. exists (tL⟨ρ⟩) (tR⟨ρ⟩); cbn. 1,2: now eapply redtmwf_wk. now eapply convneu_wk. - - intros * ?? * []; rewrite wkΠ_eq. + - intros * ?? * []; rewrite wkΠ_eq. unshelve econstructor; cbn; try rewrite wk_prod. 1,2: now eapply wkΠTerm. + now eapply convtm_wk. @@ -369,7 +333,7 @@ Section Weakenings. + intros; econstructor; tea; change tNat with tNat⟨ρ⟩; gen_typing. + constructor. + now constructor. - + intros; constructor. + + intros; constructor. change tNat with tNat⟨ρ⟩. now eapply wkNeNfEq. - intros??? NA t u ? ρ wfΔ; revert t u; pose (NA' := wkEmpty ρ wfΔ NA). @@ -385,13 +349,13 @@ Section Weakenings. + destruct prop. econstructor. change tEmpty with tEmpty⟨ρ⟩. now eapply wkNeNfEq. - - intros * ?? * []; rewrite wkΣ_eq. + - intros * ?? * []; rewrite wkΣ_eq. unshelve econstructor; cbn; try rewrite wk_sig. 1,2: now eapply wkΣTerm. - + now eapply convtm_wk. + intros; cbn; do 2 rewrite wk_comp_ren_on. irrelevance0. 2: now unshelve eapply eqFst. now rewrite wk_comp_ren_on. + + now eapply convtm_wk. + intros; cbn; irrelevance0. 2: do 2 rewrite wk_comp_ren_on; now unshelve eapply eqSnd. rewrite wk_comp_ren_on; now rewrite <- wk_up_ren_subst. diff --git a/theories/Normalisation.v b/theories/Normalisation.v index e6eb221..34667ba 100644 --- a/theories/Normalisation.v +++ b/theories/Normalisation.v @@ -168,7 +168,7 @@ Section Normalisation. - constructor. - destruct H as [? ? ? H]. apply escapeEq in H as []; now split. - - destruct H as [? ? ? ? H]. + - destruct H as [? ? H]. apply escapeTmEq in H as []; now split. Qed. diff --git a/theories/Substitution/Conversion.v b/theories/Substitution/Conversion.v index 2094821..5babfbe 100644 --- a/theories/Substitution/Conversion.v +++ b/theories/Substitution/Conversion.v @@ -1,7 +1,8 @@ +From Coq Require Import CRelationClasses. From LogRel.AutoSubst Require Import core unscoped Ast Extra. From LogRel Require Import Utils BasicAst Notations Context NormalForms Weakening GenericTyping LogicalRelation Validity. From LogRel.LogicalRelation Require Import Irrelevance Escape Reflexivity Weakening Neutral. -From LogRel.Substitution Require Import Properties Irrelevance. +From LogRel.Substitution Require Import Irrelevance Properties Reflexivity. Set Universe Polymorphism. @@ -10,135 +11,86 @@ Context `{GenericTypingProperties}. Set Printing Primitive Projection Parameters. -Lemma conv {Γ t A B l} (VΓ : [||-v Γ]) + +Lemma conv {Γ t u A B l} (VΓ : [||-v Γ]) (VA : [Γ ||-v A | VΓ]) (VB : [Γ ||-v B | VΓ]) (VeqAB : [Γ ||-v A ≅ B | VΓ | VA]) - (Vt : [Γ ||-v t : A | VΓ | VA]) : - [Γ ||-v t : B | VΓ | VB]. + (Vt : [Γ ||-v t ≅ u : A | VΓ | VA]) : + [Γ ||-v t ≅ u : B | VΓ | VB]. Proof. - constructor; intros; [eapply LRTmRedConv| eapply LRTmEqRedConv]. - 1,3: now unshelve now eapply validTyEq. - 1: now eapply validTm. - now eapply validTmExt. + constructor; intros; eapply LRTmEqConv. + 2: now unshelve now eapply validTmEq. + now eapply validTyEqLeft. Qed. -Lemma convsym {Γ t A B l} (VΓ : [||-v Γ]) +Lemma convsym {Γ t u A B l} (VΓ : [||-v Γ]) (VA : [Γ ||-v A | VΓ]) (VB : [Γ ||-v B | VΓ]) (VeqAB : [Γ ||-v A ≅ B | VΓ | VA]) - (Vt : [Γ ||-v t : B | VΓ | VB]) : - [Γ ||-v t : A | VΓ | VA]. + (Vt : [Γ ||-v t ≅ u : B | VΓ | VB]) : + [Γ ||-v t ≅ u : A | VΓ | VA]. Proof. - constructor; intros; [eapply LRTmRedConv| eapply LRTmEqRedConv]. - 1,3: unshelve eapply LRTyEqSym; tea; [|now unshelve now eapply validTyEq]. - 1: now unshelve now eapply validTm. - now eapply validTmExt. + eapply conv; tea; now eapply symValidTyEq. Qed. +#[deprecated(note="Use conv")] Lemma convEq {Γ t u A B l} (VΓ : [||-v Γ]) (VA : [Γ ||-v A | VΓ]) (VB : [Γ ||-v B | VΓ]) (VeqAB : [Γ ||-v A ≅ B | VΓ | VA]) (Vt : [Γ ||-v t ≅ u : A | VΓ | VA]) : [Γ ||-v t ≅ u : B | VΓ | VB]. -Proof. - constructor; intros; eapply LRTmEqRedConv. - 1: now unshelve now eapply validTyEq. - now eapply validTmEq. -Qed. +Proof. now eapply conv. Qed. -Lemma convSubstCtx1 {Γ Δ A B l σ} +Lemma convSubstEqCtx1 {Γ Δ A B l σ σ'} (VΓ : [||-v Γ]) (wfΔ : [|- Δ]) - (VΓA : [||-v Γ ,, A]) - (VΓB : [||-v Γ ,, B]) + (VΓA : [||-v Γ ,, A]) + (VΓB : [||-v Γ ,, B]) (VA : [_ ||-v A | VΓ]) - (VB : [_ ||-v B | VΓ]) (VAB : [_ ||-v A ≅ B | VΓ | VA]) - (Vσ : [_ ||-v σ : _ | VΓA | wfΔ]) : - [_ ||-v σ : _ | VΓB | wfΔ]. + (Vσσ': [_ ||-v σ ≅ σ' : _ | VΓA | wfΔ]) : + [_ ||-v σ ≅ σ' : _ | VΓB | wfΔ]. Proof. pose proof (invValiditySnoc VΓA) as [? [? [?]]]; subst. pose proof (invValiditySnoc VΓB) as [? [? [?]]]; subst. - eapply irrelevanceSubstExt. + eapply irrelevanceSubstEqExt. + 1: rewrite <- scons_eta'; reflexivity. 1: rewrite <- scons_eta'; reflexivity. - unshelve eapply consSubstS. - 1: epose (validTail Vσ); irrValid. - eapply LRTmRedConv. - 2: irrelevanceRefl; eapply validHead. - now eapply validTyEq. + unshelve eapply consSubstEq. + 1: epose (eqTail Vσσ'); irrValid. + eapply LRTmEqConv. + 2: irrelevanceRefl; eapply eqHead. + eapply validTyEqLeft; [eapply ureflValidTy |]; exact VAB. Unshelve. 6: tea. - 2: eapply irrelevanceSubst; now eapply validTail. + 3: eapply irrelevanceSubstEq; now eapply eqTail. tea. Qed. -Lemma convSubstEqCtx1 {Γ Δ A B l σ σ'} +Lemma convCtx1 {Γ A B P l} (VΓ : [||-v Γ]) - (wfΔ : [|- Δ]) - (VΓA : [||-v Γ ,, A]) - (VΓB : [||-v Γ ,, B]) + (VΓA : [||-v Γ ,, A]) + (VΓB : [||-v Γ ,, B]) (VA : [_ ||-v A | VΓ]) - (VB : [_ ||-v B | VΓ]) - (VAB : [_ ||-v A ≅ B | VΓ | VA]) - (Vσ : [_ ||-v σ : _ | VΓA | wfΔ]) - (Vσ' : [_ ||-v σ' : _ | VΓA | wfΔ]) - (Vσσ' : [_ ||-v σ ≅ σ' : _ | VΓA | wfΔ | Vσ]) - (VσB : [_ ||-v σ : _ | VΓB | wfΔ]) : - [_ ||-v σ ≅ σ' : _ | VΓB | wfΔ | VσB]. -Proof. - pose proof (invValiditySnoc VΓA) as [? [? [?]]]; subst. - pose proof (invValiditySnoc VΓB) as [? [? [?]]]; subst. - eapply irrelevanceSubstEq. - eapply irrelevanceSubstEqExt. - 1: rewrite <- scons_eta'; reflexivity. - unshelve eapply consSubstSEq'. - 8: now eapply eqTail. - 3: irrValid. - 3: eapply LRTmEqRedConv. - 4: now eapply eqHead. - 2: irrelevanceRefl; eapply validTyEq; irrValid. - eapply LRTmRedConv. - 2: irrelevanceRefl; eapply validHead. - eapply validTyExt. - 1: now eapply validTail. - eapply reflSubst. - Unshelve. - 1: now rewrite scons_eta'. - 9: tea. - 7: tea. - 2,6: irrValid. - 1,2: tea. - 1,2: eapply irrelevanceSubst; now eapply validTail. -Qed. - - -Lemma convCtx1 {Γ A B P l} - (VΓ : [||-v Γ]) - (VΓA : [||-v Γ ,, A]) - (VΓB : [||-v Γ ,, B]) - (VA : [_ ||-v A | VΓ]) - (VB : [_ ||-v B | VΓ]) (VAB : [_ ||-v A ≅ B | VΓ | VA]) (VP : [_ ||-v P | VΓA]) : [_ ||-v P | VΓB]. Proof. opector; intros. - - eapply validTy; tea; eapply convSubstCtx1; tea; now eapply symValidTyEq. + - eapply validTy; tea; eapply convSubstEqCtx1; tea; now eapply symValidTyEq. - irrelevanceRefl; unshelve eapply validTyExt. - 3,4: tea. - 1,2: eapply convSubstCtx1; tea; now eapply symValidTyEq. - eapply convSubstEqCtx1; cycle 2; tea; now eapply symValidTyEq. - Unshelve. all: tea. + 3,4: tea. + eapply convSubstEqCtx1; tea; now eapply symValidTyEq. + Unshelve. all: tea; now eapply ureflValidTy. Qed. -Lemma convEqCtx1 {Γ A B P Q l} +Lemma convEqCtx1 {Γ A B P Q l} (VΓ : [||-v Γ]) - (VΓA : [||-v Γ ,, A]) - (VΓB : [||-v Γ ,, B]) + (VΓA : [||-v Γ ,, A]) + (VΓB : [||-v Γ ,, B]) (VA : [_ ||-v A | VΓ]) - (VB : [_ ||-v B | VΓ]) (VAB : [_ ||-v A ≅ B | VΓ | VA]) (VP : [_ ||-v P | VΓA]) (VPB : [_ ||-v P | VΓB]) @@ -147,143 +99,91 @@ Lemma convEqCtx1 {Γ A B P Q l} Proof. constructor; intros; irrelevanceRefl. eapply validTyEq; tea. - Unshelve. 1: tea. - unshelve eapply convSubstCtx1; cycle 5; tea; now eapply symValidTyEq. + Unshelve. 1: tea. + unshelve eapply convSubstEqCtx1; cycle 5; tea. + 1: now eapply symValidTyEq. + now eapply ureflValidTy. Qed. -Lemma convSubstCtx2 {Γ Δ A1 B1 A2 B2 l σ} +Lemma convTmEqCtx1 {Γ A B C t u l} (VΓ : [||-v Γ]) - (wfΔ : [|- Δ]) - (VA1 : [_ ||-v A1 | VΓ]) - (VB1 : [_ ||-v B1 | VΓ]) - (VAB1 : [_ ||-v A1 ≅ B1 | VΓ | VA1]) - (VΓA1 := validSnoc VΓ VA1) - (VΓB1 := validSnoc VΓ VB1) - (VA2 : [_ ||-v A2 | VΓA1]) - (VB2 : [_ ||-v B2 | VΓA1]) - (VAB : [_ ||-v A2 ≅ B2 | VΓA1 | VA2]) - (VB2' := convCtx1 VΓ VΓA1 VΓB1 VA1 VB1 VAB1 VB2) - (VΓA12 := validSnoc VΓA1 VA2) - (VΓB12 := validSnoc VΓB1 VB2') - (Vσ : [_ ||-v σ : _ | VΓA12 | wfΔ]) : - [_ ||-v σ : _ | VΓB12 | wfΔ]. + (VΓA : [||-v Γ ,, A]) + (VΓB : [||-v Γ ,, B]) + (VA : [_ ||-v A | VΓ]) + (VC : [_ ||-v C | VΓA]) + (VC' : [_ ||-v C | VΓB]) + (VAB : [_ ||-v A ≅ B | VΓ | VA]) + (VPtu : [_ ||-v t ≅ u : _ | VΓA | VC]) : + [_ ||-v t ≅ u : _ | VΓB | VC']. Proof. - eapply irrelevanceSubstExt. - 1: rewrite <- scons_eta'; reflexivity. - unshelve eapply consSubstS. - + eapply convSubstCtx1; tea. - now eapply validTail. - + eapply LRTmRedConv. - 2: now eapply validHead. - eapply validTyEq; tea. - Unshelve. all: tea. + constructor; intros; irrelevanceRefl. + (unshelve now eapply validTmEq); tea. + unshelve (eapply convSubstEqCtx1; tea; now eapply symValidTyEq). + 2: now eapply ureflValidTy. Qed. -Lemma convSubstCtx2' {Γ Δ A1 B1 A2 B2 l σ} - (VΓ : [||-v Γ]) - (wfΔ : [|- Δ]) - (VΓA1 : [||-v Γ ,, A1]) - (VΓA12 : [||-v Γ ,, A1 ,, A2]) - (VΓB12 : [||-v Γ ,, B1 ,, B2]) - (VA1 : [_ ||-v A1 | VΓ]) - (VB1 : [_ ||-v B1 | VΓ]) - (VAB1 : [_ ||-v A1 ≅ B1 | VΓ | VA1]) - (VA2 : [_ ||-v A2 | VΓA1]) - (VB2 : [_ ||-v B2 | VΓA1]) - (VAB : [_ ||-v A2 ≅ B2 | VΓA1 | VA2]) - (Vσ : [_ ||-v σ : _ | VΓA12 | wfΔ]) : - [_ ||-v σ : _ | VΓB12 | wfΔ]. -Proof. - eapply irrelevanceSubst. - eapply convSubstCtx2; irrValid. - Unshelve. all: tea; irrValid. -Qed. -Lemma convSubstEqCtx2 {Γ Δ A1 B1 A2 B2 l σ σ'} +Lemma convSubstEqCtx2 {Γ Δ A1 B1 A2 B2 l σ σ'} (VΓ : [||-v Γ]) (wfΔ : [|- Δ]) (VA1 : [_ ||-v A1 | VΓ]) (VB1 : [_ ||-v B1 | VΓ]) (VAB1 : [_ ||-v A1 ≅ B1 | VΓ | VA1]) - (VΓA1 := validSnoc VΓ VA1) - (VΓB1 := validSnoc VΓ VB1) + (VΓA1 := validSnoc VΓ VA1) + (VΓB1 := validSnoc VΓ VB1) (VA2 : [_ ||-v A2 | VΓA1]) (VB2 : [_ ||-v B2 | VΓA1]) (VAB : [_ ||-v A2 ≅ B2 | VΓA1 | VA2]) - (VB2' := convCtx1 VΓ VΓA1 VΓB1 VA1 VB1 VAB1 VB2) - (VΓA12 := validSnoc VΓA1 VA2) - (VΓB12 := validSnoc VΓB1 VB2') - (Vσ : [_ ||-v σ : _ | VΓA12 | wfΔ]) - (Vσ' : [_ ||-v σ' : _ | VΓA12 | wfΔ]) - (Vσσ' : [_ ||-v σ ≅ σ' : _ | VΓA12 | wfΔ | Vσ]) - (VσB : [_ ||-v σ : _ | VΓB12 | wfΔ]) : - [_ ||-v σ ≅ σ' : _ | VΓB12 | wfΔ | VσB]. + (VB2' := convCtx1 VΓ VΓA1 VΓB1 VA1 VAB1 VB2) + (VΓA12 := validSnoc VΓA1 VA2) + (VΓB12 := validSnoc VΓB1 VB2') + (Vσσ' : [_ ||-v σ ≅ σ' : _ | VΓA12 | wfΔ]) : + [_ ||-v σ ≅ σ' : _ | VΓB12 | wfΔ]. Proof. - eapply irrelevanceSubstEq. eapply irrelevanceSubstEqExt. - 1: rewrite <- scons_eta'; reflexivity. - unshelve eapply consSubstSEq'. - 8:{ eapply convSubstEqCtx1; tea. - 1: now eapply validTail. + 1,2: rewrite <- scons_eta'; reflexivity. + unshelve eapply consSubstEq. + + unshelve (eapply convSubstEqCtx1; tea); tea. now eapply eqTail. - } - 3,4: irrValid. - 2: eapply convSubstCtx1; tea; now eapply validTail. - 3: eapply LRTmEqRedConv. - 4: now eapply eqHead. - 2: irrelevanceRefl; eapply validTyEq; irrValid. - eapply LRTmRedConv. - 2: irrelevanceRefl; eapply validHead. - eapply validTyExt. - 1: now eapply validTail. - eapply reflSubst. - Unshelve. - 1: now rewrite scons_eta'. - 11: tea. - 1,2,6: irrValid. - 1: tea. - 2: eapply convSubstCtx1; tea. - 1,2: now eapply validTail. -Qed. - -Lemma convSubstEqCtx2' {Γ Δ A1 B1 A2 B2 l σ σ'} + + unshelve eapply LRTmEqConv; cycle 3. + 2: now eapply eqHead. + 1: eapply validTyEqLeft; [|tea]; tea. + Unshelve. all: tea. +Qed. + +Lemma convSubstEqCtx2' {Γ Δ A1 B1 A2 B2 l σ σ'} (VΓ : [||-v Γ]) (wfΔ : [|- Δ]) + (VΓA1 : [||-v Γ ,, A1]) + (VΓA12 : [||-v Γ ,, A1 ,, A2]) + (VΓB12 : [||-v Γ ,, B1 ,, B2]) (VA1 : [_ ||-v A1 | VΓ]) - (VB1 : [_ ||-v B1 | VΓ]) (VAB1 : [_ ||-v A1 ≅ B1 | VΓ | VA1]) - (VΓA1 : [||-v Γ ,, A1]) - (VΓB1 : [||-v Γ ,, B1]) (VA2 : [_ ||-v A2 | VΓA1]) - (VB2 : [_ ||-v B2 | VΓA1]) (VAB : [_ ||-v A2 ≅ B2 | VΓA1 | VA2]) - (VB2' : [_ ||-v B2 | VΓB1]) - (VΓA12 : [||-v Γ ,, A1 ,, A2]) - (VΓB12 : [||-v Γ ,, B1,, B2]) - (Vσ : [_ ||-v σ : _ | VΓA12 | wfΔ]) - (Vσ' : [_ ||-v σ' : _ | VΓA12 | wfΔ]) - (Vσσ' : [_ ||-v σ ≅ σ' : _ | VΓA12 | wfΔ | Vσ]) - (VσB : [_ ||-v σ : _ | VΓB12 | wfΔ]) : - [_ ||-v σ ≅ σ' : _ | VΓB12 | wfΔ | VσB]. + (Vσσ' : [_ ||-v σ ≅ σ' : _ | VΓA12 | wfΔ]) : + [_ ||-v σ ≅ σ' : _ | VΓB12 | wfΔ]. Proof. + pose proof (ureflValidTy VAB1). + pose proof (ureflValidTy VAB). eapply irrelevanceSubstEq. eapply convSubstEqCtx2; irrValid. - Unshelve. all: tea; irrValid. + Unshelve. all: tea; irrValid. Qed. -Lemma convCtx2 {Γ A1 B1 A2 B2 P l} +Lemma convCtx2 {Γ A1 B1 A2 B2 P l} (VΓ : [||-v Γ]) (VA1 : [_ ||-v A1 | VΓ]) (VB1 : [_ ||-v B1 | VΓ]) (VAB1 : [_ ||-v A1 ≅ B1 | VΓ | VA1]) - (VΓA1 := validSnoc VΓ VA1) - (VΓB1 := validSnoc VΓ VB1) + (VΓA1 := validSnoc VΓ VA1) + (VΓB1 := validSnoc VΓ VB1) (VA2 : [_ ||-v A2 | VΓA1]) (VB2 : [_ ||-v B2 | VΓA1]) (VAB : [_ ||-v A2 ≅ B2 | VΓA1 | VA2]) - (VB2' := convCtx1 VΓ VΓA1 VΓB1 VA1 VB1 VAB1 VB2) - (VΓA12 := validSnoc VΓA1 VA2) - (VΓB12 := validSnoc VΓB1 VB2') + (VB2' := convCtx1 VΓ VΓA1 VΓB1 VA1 VAB1 VB2) + (VΓA12 := validSnoc VΓA1 VA2) + (VΓB12 := validSnoc VΓB1 VB2') (VP : [_ ||-v P | VΓA12]) : [_ ||-v P | VΓB12]. Proof. @@ -291,33 +191,216 @@ Proof. assert [_ ||-v B1 ≅ A1 | _ | VB1] by now eapply symValidTyEq. assert [_ ||-v B2 ≅ A2 | _ | VB2'] by (eapply convEqCtx1; tea; now eapply symValidTyEq). opector; intros. - - eapply validTy; tea; now eapply convSubstCtx2'. + - eapply validTy; tea; now eapply convSubstEqCtx2'. - irrelevanceRefl; unshelve eapply validTyExt. - 3,4: tea. - 1,2: now eapply convSubstCtx2'. + 3,4: tea. eapply convSubstEqCtx2'; tea. Unshelve. tea. Qed. -Lemma convCtx2' {Γ A1 A2 B1 B2 P l} +Lemma convCtx2' {Γ A1 A2 B1 B2 P l} (VΓ : [||-v Γ]) (VA1 : [_ ||-v A1 | VΓ]) - (VB1 : [_ ||-v B1 | VΓ]) (VAB1 : [_ ||-v A1 ≅ B1 | VΓ | VA1]) - (VΓA1 : [||-v Γ ,, A1]) - (VΓB1 : [||-v Γ ,, B1]) + (VΓA1 : [||-v Γ ,, A1]) (VA2 : [_ ||-v A2 | VΓA1]) - (VB2 : [_ ||-v B2 | VΓA1]) (VAB : [_ ||-v A2 ≅ B2 | VΓA1 | VA2]) - (VB2' : [_ ||-v B2 | VΓB1]) - (VΓA12 : [||-v Γ ,, A1 ,, A2]) + (VΓA12 : [||-v Γ ,, A1 ,, A2]) (VΓB12 : [||-v Γ ,, B1,, B2]) (VP : [_ ||-v P | VΓA12]) : [_ ||-v P | VΓB12]. Proof. + pose proof (ureflValidTy VAB1). + pose proof (ureflValidTy VAB). eapply irrelevanceTy; eapply convCtx2; irrValid. Unshelve. all: tea; irrValid. Qed. + + +Inductive validCtxEq : forall {Γ Γ'}, [||-v Γ] -> [||-v Γ'] -> Type := + | validCtxNil : validCtxEq validEmpty validEmpty + | validCtxSnoc {l l' Γ Γ' A A'} (VΓ :[||-v Γ]) (VΓ' : [||-v Γ']) + (VA : [_ ||-v A| VΓ]) + (VA' : [_ ||-v A'| VΓ']) : + validCtxEq VΓ VΓ' -> [_ ||-v A ≅ A'| VΓ| VA] -> validCtxEq (validSnoc VΓ VA) (validSnoc VΓ' VA'). + +Import EqNotations. + +Definition transpValidCtxEq {Γ1 Γ1' Γ2 Γ2'} + {VΓ1 : [||-v Γ1]} + {VΓ1' : [||-v Γ1']} + {VΓ2 : [||-v Γ2]} + {VΓ2' : [||-v Γ2']} + (h1 : validCtxEq VΓ1 VΓ1') + (h2 : validCtxEq VΓ2 VΓ2') + (e : Γ1 = Γ2) (e' : Γ1' = Γ2') + (eV : rew [fun Δ => [||-v Δ]] e in VΓ1 = VΓ2) + (eV' : rew [fun Δ => [||-v Δ]] e' in VΓ1' = VΓ2') : Type. +Proof. subst; cbn in *; exact (h1 = h2). Defined. + + +Lemma invValidCtxEq {Γ Γ'} {VΓ : [||-v Γ]} {VΓ' : [||-v Γ']} (h : validCtxEq VΓ VΓ'): + match Γ as Γ return forall Γ' (VΓ : [||-v Γ]) (VΓ' : [||-v Γ']) (h : validCtxEq VΓ VΓ'), Type with + | nil => fun Γ' VΓ VΓ' h => + ∑ (eΓ' : Γ' = ε), + transpValidCtxEq h validCtxNil eq_refl eΓ' (invValidityEmpty _) (invValidityEmpty _) + | (A :: Γ)%list => fun ΓA' VΓA VΓA' h => + let '(l;VΓ; VA;eVΓA) := invValiditySnoc VΓA in + ∑ l' A' Γ' (eΓ' : ΓA' = (A' :: Γ')%list) + (VΓ' : [||-v Γ']) + (VA' : [_ ||-v A'| VΓ']) + (VΓΓ' : validCtxEq VΓ VΓ') + (VAA' : [_ ||-v A ≅ A'| VΓ| VA]) + (eVΓA' : rew [fun Δ => [||-v Δ]] eΓ' in VΓA' = validSnoc VΓ' VA'), + transpValidCtxEq h (validCtxSnoc VΓ VΓ' VA VA' VΓΓ' VAA') eq_refl eΓ' eVΓA eVΓA' + end Γ' VΓ VΓ' h. +Proof. + induction h. + + exists eq_refl; reflexivity. + + unshelve eexists _,_, _, eq_refl, _, _, _, _, eq_refl; tea. + reflexivity. +Defined. + + +Notation "[||-v Γ ≅ Γ' ]" := (∑ VΓ VΓ', @validCtxEq Γ Γ' VΓ VΓ'). + +Lemma reflValidCtxEq {Γ} (VΓ : [||-v Γ]) : validCtxEq VΓ VΓ. +Proof. + induction Γ, VΓ using validity_rect; constructor; tea. + now eapply reflValidTy. +Qed. + +Lemma convSubsS [Γ Γ'] [VΓ : [||-v Γ]] [VΓ' : [||-v Γ']] : validCtxEq VΓ VΓ' -> + forall {Δ} (wfΔ : [|-Δ]) {σ σ'}, + [_ ||-v σ ≅ σ' : _ | VΓ | wfΔ ] -> [_ ||-v σ ≅ σ' : _ | VΓ' | wfΔ ]. +Proof. + intros h; induction h. + 1: constructor. + intros ???? Vσ. + exists (IHh _ _ _ _ (eqTail Vσ)). + eapply LRTmEqConv. 2: exact (eqHead Vσ). + instValid (lrefl (eqTail Vσ)); irrelevance. +Qed. + +(* If needed there is a much easier proof using convSubstS and idSubstS *) +(* Lemma convValidCtxEqIdSubst {Γ Γ'} (VΓ : [||-v Γ]) (VΓ' : [||-v Γ']) : validCtxEq VΓ VΓ' -> + [_ ||-v tRel : _ | VΓ | validWf VΓ']. +Proof. + intros h; induction h; unshelve econstructor. + - eapply irrelevanceSubstEqExt. + 3: eapply (wkSubstSEq _ _ _ (wk1 A') IHh). + all: intros a; cbn -[wk1]; now rewrite (wk1_ren a). + - eapply var0conv. 2: now eapply validTyWf. + replace A[↑ >> tRel] with A⟨↑⟩ by (bsimpl; now substify). + erewrite <-2!wk1_ren_on; symmetry; unshelve eapply escapeEq, wkEq; tea. + 1: now eapply validWf, validSnoc. + 1: generalize (validTy VA _ IHh); now rewrite subst_rel. + rewrite <-(subst_rel A'); irrelevance0; [apply subst_rel|]. + eapply (validTyEq t _ IHh). +Qed. *) + +Lemma convValidTy [Γ Γ'] [VΓ : [||-v Γ]] [VΓ' : [||-v Γ']] : validCtxEq VΓ VΓ' -> + forall [l A], [_ ||-v A | VΓ'] -> [_ ||-v A | VΓ]. +Proof. + intros h l A VA; unshelve econstructor; intros ???? Vσ. + all: instValid (convSubsS h wfΔ Vσ); tea ; irrelevance. +Qed. + +Lemma convValidTyEq [Γ Γ'] [VΓ : [||-v Γ]] [VΓ' : [||-v Γ']] : validCtxEq VΓ VΓ' -> + forall [l A A'] (VAΓ : [_ ||-v A | VΓ]) (VAΓ' : [_ ||-v A | VΓ']), + [_ ||-v A ≅ A' | VΓ' | VAΓ'] -> [_ ||-v A ≅ A' | VΓ | VAΓ]. +Proof. + intros * h ** ; unshelve econstructor; intros ???? Vσ. + instValid (convSubsS h wfΔ Vσ); irrelevance. +Qed. + +Lemma convValidTyEq' [Γ Γ'] [VΓ : [||-v Γ]] [VΓ' : [||-v Γ']] (h : validCtxEq VΓ VΓ') : + forall [l A A'] (VAΓ' : [_ ||-v A | VΓ']), + [_ ||-v A ≅ A' | VΓ' | VAΓ'] -> [_ ||-v A ≅ A' | VΓ | convValidTy h VAΓ']. +Proof. intros; now eapply convValidTyEq. Qed. + +Lemma convValidTmEq {Γ Γ'} (VΓ : [||-v Γ]) (VΓ' : [||-v Γ']) : validCtxEq VΓ VΓ' -> + forall l A t u (VAΓ : [_ ||-v A | VΓ]) (VAΓ' : [_ ||-v A | VΓ']), + [_ ||-v t ≅ u : _ | VΓ' | VAΓ'] -> [_ ||-v t ≅ u : _ | VΓ | VAΓ]. +Proof. + intros * h ** ; unshelve econstructor; intros ???? Vσ. + instValid (convSubsS h wfΔ Vσ); irrelevance. +Qed. + +Lemma convValidTmEq' [Γ Γ'] [VΓ : [||-v Γ]] [VΓ' : [||-v Γ']] (h : validCtxEq VΓ VΓ') : + forall [l A t u] (VAΓ' : [_ ||-v A | VΓ']), + [_ ||-v t ≅ u : _ | VΓ' | VAΓ'] -> [_ ||-v t ≅ u : _ | VΓ | convValidTy h VAΓ']. +Proof. intros; now eapply convValidTmEq. Qed. + +Lemma symValidCtxEq [Γ Δ] [VΓ : [||-v Γ]] [VΔ : [||-v Δ]] : validCtxEq VΓ VΔ -> validCtxEq VΔ VΓ. +Proof. + intros h; induction h; constructor; tea. + eapply symValidTyEq. + now unshelve now eapply convValidTyEq'. +Qed. + +Lemma transValidCtxEq [Γ Δ Ξ] [VΓ : [||-v Γ]] [VΔ : [||-v Δ]] [VΞ : [||-v Ξ]] : + validCtxEq VΓ VΔ -> validCtxEq VΔ VΞ -> validCtxEq VΓ VΞ. +Proof. + intros h; induction h in Ξ, VΞ |- *; intros h'; pose proof (h'' := invValidCtxEq h'); cbn in h''. + - destruct h'' as []; subst; rewrite (invValidityEmpty VΞ); constructor. + - destruct h'' as (l''&A''&Γ''&?&?&?&?&?&?&?); do 3 (subst; cbn in *); econstructor. + + now eapply IHh. + + eapply transValidTyEq; tea. + now eapply convValidTyEq. + Unshelve. now eapply convValidTy. +Qed. + +Lemma irrValidCtxEq [Γ Δ] [VΓ1 VΓ2 : [||-v Γ]] [VΔ1 VΔ2 : [||-v Δ]] : validCtxEq VΓ1 VΔ1 -> validCtxEq VΓ2 VΔ2. +Proof. + intros h; induction h. + - rewrite (invValidityEmpty VΓ2), (invValidityEmpty VΔ2); constructor. + - pose proof (invValiditySnoc VΓ2) as (l1&VΓ1&?&->). + pose proof (invValiditySnoc VΔ2) as (l2&VΓ1'&?&->). + econstructor. 1: now eapply IHh. irrValid. +Qed. + +Lemma reflValidCtx {Γ} : [||-v Γ] -> [||-v Γ ≅ Γ]. +Proof. + intros VΓ; exists VΓ, VΓ; apply reflValidCtxEq. +Defined. + +Lemma symValidCtx [Γ Δ] : [||-v Γ ≅ Δ] -> [||-v Δ ≅ Γ]. +Proof. + intros (?&?&?); eexists _, _; now eapply symValidCtxEq. +Defined. + +Lemma transValidCtx [Γ Δ Ξ] : [||-v Γ ≅ Δ] -> [||-v Δ ≅ Ξ] -> [||-v Γ ≅ Ξ]. +Proof. + intros (VΓ&VΔ&?) (?&?& h); eexists VΓ, _; eapply transValidCtxEq; tea. + eapply irrValidCtxEq, h. + Unshelve. tea. +Qed. + +Lemma validCtxSnoc' [l Γ Γ' A A'] (VΓΓ' :[||-v Γ ≅ Γ']) (VA : [_ ||-v A| VΓΓ'.π1]) : + [_ ||-v A ≅ A'| _ | VA] -> [||-v Γ,,A ≅ Γ',,A']. +Proof. + intros VAA'; destruct VΓΓ' as (VΓ& VΓ'& VΓΓ'). + unshelve eexists (validSnoc _ VA), (validSnoc (l:=l) VΓ' _). + 1: eapply convValidTy; [now eapply symValidCtxEq| exact (ureflValidTy VAA')]. + econstructor; cbn; tea. +Defined. + +Lemma convVTy [Γ Γ' l A] (Veq : [||-v Γ ≅ Γ']) : + [_ ||-v A | Veq.π1] -> [_ ||-v A | Veq.π2.π1]. +Proof. apply convValidTy, symValidCtxEq, Veq.π2.π2. Qed. + +Lemma convVTyEq [Γ Γ' l A B] (Veq : [||-v Γ ≅ Γ']) (VA : [_ ||-v A | Veq.π1]) : + [_ ||-v A ≅ B | Veq.π1 | VA ] -> [_ ||-v A ≅ B | Veq.π2.π1 | convVTy Veq VA]. +Proof. apply convValidTyEq, symValidCtxEq, Veq.π2.π2. Qed. + +Lemma convVTmEq [Γ Γ' l A t u] (Veq : [||-v Γ ≅ Γ']) (VA : [_ ||-v A | Veq.π1]) : + [_ ||-v t ≅ u : _ | Veq.π1 | VA ] -> [_ ||-v t ≅ u : _ | Veq.π2.π1 | convVTy Veq VA]. +Proof. apply convValidTmEq, symValidCtxEq, Veq.π2.π2. Qed. + + End Conversion. + +Notation "[||-v Γ ≅ Γ' ]" := (∑ VΓ VΓ', validCtxEq (Γ:=Γ) (Γ':=Γ') VΓ VΓ'). diff --git a/theories/Substitution/Escape.v b/theories/Substitution/Escape.v index 0ccfb7e..b6a6559 100644 --- a/theories/Substitution/Escape.v +++ b/theories/Substitution/Escape.v @@ -26,17 +26,7 @@ Proof. irrelevance. Qed. -Lemma reducibleTm {Γ l A t} (VΓ : [||-v Γ]) (VA : [Γ ||-v A | VΓ]) : - [Γ ||-v t : A | VΓ | VA] -> [Γ ||- t : A | reducibleTy VΓ VA]. -Proof. - intros. - replace A with A[tRel] by now asimpl. - replace t with t[tRel] by now asimpl. - unshelve epose proof (validTm X _ (idSubstS VΓ)). - irrelevance. -Qed. - -Lemma reducibleTmEq {Γ l A t u} (VΓ : [||-v Γ]) (VA : [Γ ||-v A | VΓ]) : +Lemma reducibleTmEq {Γ l A t u} (VΓ : [||-v Γ]) (VA : [Γ ||-v A | VΓ]) : [Γ ||-v t ≅ u : A | VΓ | VA] -> [Γ ||- t ≅ u : A | reducibleTy VΓ VA]. Proof. intros. @@ -51,7 +41,7 @@ Lemma escapeTy {Γ l A} (VΓ : [||-v Γ]) (VA : [Γ ||-v A | VΓ]) : [Γ |- A Proof. eapply escape; now eapply reducibleTy. Qed. -Lemma escapeEq {Γ l A B} (VΓ : [||-v Γ]) (VA : [Γ ||-v A | VΓ]) : +Lemma escapeEq {Γ l A B} (VΓ : [||-v Γ]) (VA : [Γ ||-v A | VΓ]) : [Γ ||-v A ≅ B | VΓ | VA] -> [Γ |- A ≅ B]. Proof. intros; unshelve eapply escapeEq; tea. @@ -59,18 +49,18 @@ Proof. now eapply reducibleTyEq. Qed. -Lemma escapeTm {Γ l A t} (VΓ : [||-v Γ]) (VA : [Γ ||-v A | VΓ]) : - [Γ ||-v t : A | VΓ | VA] -> [Γ |- t : A]. +Lemma escapeTmEq {Γ l A t u} (VΓ : [||-v Γ]) (VA : [Γ ||-v A | VΓ]) : + [Γ ||-v t ≅ u : A | VΓ | VA] -> [Γ |- t ≅ u : A]. Proof. - intros; unshelve eapply escapeTerm; tea. + intros; unshelve eapply escapeEqTerm; tea. 1: now eapply reducibleTy. - now eapply reducibleTm. + now eapply reducibleTmEq. Qed. -Lemma escapeTmEq {Γ l A t u} (VΓ : [||-v Γ]) (VA : [Γ ||-v A | VΓ]) : - [Γ ||-v t ≅ u : A | VΓ | VA] -> [Γ |- t ≅ u : A]. +Lemma escapeTm {Γ l A t u} (VΓ : [||-v Γ]) (VA : [Γ ||-v A | VΓ]) : + [Γ ||-v t ≅ u : A | VΓ | VA] -> [Γ |- t : A]. Proof. - intros; unshelve eapply escapeEqTerm; tea. + intros; unshelve eapply escapeTerm; tea. 1: now eapply reducibleTy. now eapply reducibleTmEq. Qed. diff --git a/theories/Substitution/Introductions/Application.v b/theories/Substitution/Introductions/Application.v index 08fb78c..e89aa87 100644 --- a/theories/Substitution/Introductions/Application.v +++ b/theories/Substitution/Introductions/Application.v @@ -12,26 +12,6 @@ Set Printing Primitive Projection Parameters. Set Printing Universes. -Lemma appValid {Γ F G t u l} - {VΓ : [||-v Γ]} - {VF : [Γ ||-v F | VΓ]} - {VΠFG : [Γ ||-v tProd F G | VΓ]} - (Vt : [Γ ||-v t : tProd F G | VΓ | VΠFG]) - (Vu : [Γ ||-v u : F | VΓ | VF]) - (VGu := substSΠ VΠFG Vu) : - [Γ ||-v tApp t u : G[u..] | VΓ | VGu]. -Proof. - opector; intros. - - instValid Vσ. - epose (appTerm RVΠFG RVt RVu (substSΠaux VΠFG Vu _ _ wfΔ Vσ)). - irrelevance. - - instAllValid Vσ Vσ' Vσσ'. - unshelve epose (appcongTerm _ REVt RVu _ REVu (substSΠaux VΠFG Vu _ _ wfΔ Vσ)). - 2: irrelevance. - eapply LRTmRedConv; tea. - unshelve eapply LRTyEqSym. 2,3: tea. -Qed. - Lemma appcongValid {Γ F G t u a b l} {VΓ : [||-v Γ]} {VF : [Γ ||-v F | VΓ]} @@ -43,10 +23,25 @@ Lemma appcongValid {Γ F G t u a b l} (VGa := substSΠ VΠFG Va) : [Γ ||-v tApp t a ≅ tApp u b : G[a..] | VΓ | VGa]. Proof. - constructor; intros; instValid Vσ. - unshelve epose proof (appcongTerm _ RVtu _ _ _ (substSΠaux VΠFG Va _ _ wfΔ Vσ)); fold subst_term; cycle 5. - all: try irrelevance. - now eapply LRCumulative. + constructor; intros; instValid Vσσ'. + pose proof (h := substSΠaux VΠFG Va _ _ _ wfΔ Vσσ'). + pose proof (appcongTerm _ RVtu RVab h). + irrelevance. Qed. +Lemma appValid {Γ F G t u l} + {VΓ : [||-v Γ]} + {VF : [Γ ||-v F | VΓ]} + {VΠFG : [Γ ||-v tProd F G | VΓ]} + (Vt : [Γ ||-v t : tProd F G | VΓ | VΠFG]) + (Vu : [Γ ||-v u : F | VΓ | VF]) + (VGu := substSΠ VΠFG Vu) : + [Γ ||-v tApp t u : G[u..] | VΓ | VGu]. +Proof. + eapply lreflValidTm, appcongValid. + 1,3: now eapply reflValidTm. + tea. +Qed. + + End Application. diff --git a/theories/Substitution/Introductions/Empty.v b/theories/Substitution/Introductions/Empty.v index b356cd4..dd0d354 100644 --- a/theories/Substitution/Introductions/Empty.v +++ b/theories/Substitution/Introductions/Empty.v @@ -1,8 +1,9 @@ +From Coq Require Import CRelationClasses. From LogRel.AutoSubst Require Import core unscoped Ast Extra. From LogRel Require Import Utils BasicAst Notations Context NormalForms Weakening GenericTyping LogicalRelation Validity. -From LogRel.LogicalRelation Require Import Induction Irrelevance Escape Reflexivity Weakening Neutral Transitivity Reduction Application Universe SimpleArr. -From LogRel.Substitution Require Import Irrelevance Properties Conversion SingleSubst Reflexivity. -From LogRel.Substitution.Introductions Require Import Universe Pi SimpleArr Var. +From LogRel.LogicalRelation Require Import Induction Escape Irrelevance Reflexivity Irrelevance Weakening Neutral Transitivity Reduction Application Universe EqRedRight SimpleArr. +From LogRel.Substitution Require Import Irrelevance Properties Conversion SingleSubst Reflexivity Reduction. +From LogRel.Substitution.Introductions Require Import Universe Pi SimpleArr Var Application. Set Universe Polymorphism. @@ -13,7 +14,7 @@ Set Printing Primitive Projection Parameters. Lemma emptyRed {Γ l} : [|- Γ] -> [Γ ||- tEmpty]. -Proof. +Proof. intros; eapply LREmpty_; constructor; eapply redtywf_refl; gen_typing. Defined. @@ -24,7 +25,7 @@ Proof. + cbn; constructor; eapply redtywf_refl; gen_typing. Defined. -Lemma emptyconvValid {Γ l} (VΓ : [||-v Γ]) (VEmpty : [Γ ||-v tEmpty | VΓ]) : +Lemma emptyconvValid {Γ l} (VΓ : [||-v Γ]) (VEmpty : [Γ ||-v tEmpty | VΓ]) : [Γ ||-v tEmpty ≅ tEmpty | VΓ | VEmpty]. Proof. constructor; intros. @@ -32,105 +33,31 @@ Proof. constructor; cbn; eapply redtywf_refl; gen_typing. Qed. -(* TODO: also appears in Nat.v, should probably be more central *) -Lemma redUOne {Γ} : [|- Γ] -> [Γ ||-U U]. +Lemma emptyURedTm {Δ} (wfΔ : [|-Δ]) : URedTm (LogRelRec one) Δ tEmpty U (redUOneCtx wfΔ). Proof. - intros ; econstructor; [easy| gen_typing|eapply redtywf_refl; gen_typing]. + exists tEmpty; [| constructor]. + eapply redtmwf_refl; gen_typing. Defined. -Lemma emptyTermRed {Δ} (wfΔ : [|-Δ]) : [Δ ||- tEmpty : U | LRU_ (redUOne wfΔ)]. +Lemma emptyTermRed {Δ} (wfΔ : [|-Δ]) : [Δ ||- tEmpty : U | LRU_ (redUOneCtx wfΔ)]. Proof. - econstructor. - + eapply redtmwf_refl; gen_typing. - + constructor. - + gen_typing. - + now eapply (emptyRed (l:= zero)). + unshelve eexists (emptyURedTm wfΔ) (emptyURedTm wfΔ) _; cbn. + 1,3: now eapply (emptyRed (l:=zero)). + 1: gen_typing. + apply reflLRTyEq. Defined. Lemma emptyTermValid {Γ} (VΓ : [||-v Γ]): [Γ ||-v tEmpty : U | VΓ | UValid VΓ]. Proof. - constructor; intros. - - eapply emptyTermRed. - - eexists (emptyTermRed wfΔ) (emptyTermRed wfΔ) (emptyRed wfΔ); cbn. - + gen_typing. - + now eapply (emptyRed (l:=zero)). - + constructor; eapply redtywf_refl; gen_typing. -Qed. - - -(* TODO: move *) -Lemma up_single_subst {t σ u} : t[up_term_term σ][u..] = t[u .: σ]. -Proof. now bsimpl. Qed. - -(* TODO: move *) -Lemma up_liftSubst_eq {σ t u} : t[up_term_term σ][u]⇑ = t[u .: ↑ >> up_term_term σ]. -Proof. - bsimpl. eapply ext_term; intros [|n]; cbn. - 1: reflexivity. - unfold funcomp; now rewrite rinstInst'_term. + constructor; intros; eapply emptyTermRed. Qed. -(* TODO: move *) Lemma liftSubst_singleSubst_eq {t u v: term} : t[u]⇑[v..] = t[u[v..]..]. Proof. now bsimpl. Qed. -Definition emptyElim {Γ A} (P : term) {n} (NA : [Γ ||-Empty A]) (p : EmptyProp Γ n) : term. -Proof. - destruct p. - - exact (tEmptyElim P ne). -Defined. - -Section EmptyElimRed. - Context {Γ l P} - (wfΓ : [|- Γ]) - (NN : [Γ ||-Empty tEmpty]) - (RN := LREmpty_ _ NN) - (RP : [Γ,, tEmpty ||- P]) - (RPpt : forall n, [Γ ||- n : _ | RN] -> [Γ ||- P[n..]]) - (RPext : forall n n' (Rn : [Γ ||- n : _ | RN]), - [Γ ||- n' : _ | RN] -> - [Γ ||- n ≅ n' : _ | RN] -> - [Γ ||- P[n..] ≅ P[n'..] | RPpt _ Rn]). - - Definition emptyRedElimStmt := - (forall n (Rn : [Γ ||- n : _ | RN]), - [Γ ||- tEmptyElim P n : _ | RPpt _ Rn ] × - [Γ ||- tEmptyElim P n ≅ tEmptyElim P (EmptyRedTm.nf Rn) : _ | RPpt _ Rn]) × - (forall n (Nn : EmptyProp Γ n) (Rn : [Γ ||- n : _ | RN]), - [Γ ||- tEmptyElim P n : P[n..] | RPpt _ Rn ] × - [Γ ||- tEmptyElim P n ≅ emptyElim P NN Nn : _ | RPpt _ Rn]). - - Lemma emptyElimRedAux : emptyRedElimStmt. - Proof. - escape. - apply EmptyRedInduction. - - intros ? nf red ? nfprop ih. - assert [Γ ||- nf : _ | RN]. 1:{ - econstructor; tea. - eapply redtmwf_refl; gen_typing. - } - eapply redSubstTerm. - + eapply LRTmRedConv. - 2: unshelve eapply ih; tea. - eapply RPext. - 2: eapply LRTmEqSym. - 1,2: eapply redwfSubstTerm; tea. - + eapply redtm_emptyelim; tea. - cbn; gen_typing. - - intros ? [] ?. - apply reflect. - + apply completeness. - + now eapply ty_emptyElim. - + now eapply ty_emptyElim. - + eapply convneu_emptyElim; tea. - { eapply escapeEq, reflLRTyEq. } - Unshelve. all: tea. - Qed. - - Lemma emptyElimRed : forall n (Rn : [Γ ||- n : _ | RN]), [Γ ||- tEmptyElim P n : _ | RPpt _ Rn ]. - Proof. intros. apply (fst emptyElimRedAux). Qed. -End EmptyElimRed. - +(* TODO: move in Neutral *) +Lemma NeNfRed {Γ l A n n'} (RA : [Γ ||- A]) : [Γ ||-NeNf n ≅ n' : A] -> [RA | Γ ||- n ≅ n' : A]. +Proof. intros []; now eapply neuTermEq. Qed. Section EmptyElimRedEq. Context {Γ l P Q} @@ -140,182 +67,87 @@ Section EmptyElimRedEq. (RP : [Γ,, tEmpty ||- P]) (RQ : [Γ,, tEmpty ||- Q]) (eqPQ : [Γ,, tEmpty |- P ≅ Q]) - (RPpt : forall n, [Γ ||- n : _ | RN] -> [Γ ||- P[n..]]) - (RQpt : forall n, [Γ ||- n : _ | RN] -> [Γ ||- Q[n..]]) - (RPQext : forall n n' (Rn : [Γ ||- n : _ | RN]), - [Γ ||- n' : _ | RN] -> - [Γ ||- n ≅ n' : _ | RN] -> - [Γ ||- P[n..] ≅ Q[n'..] | RPpt _ Rn]). - - Lemma RPext : forall n n' (Rn : [Γ ||- n : _ | RN]), - [Γ ||- n' : _ | RN] -> - [Γ ||- n ≅ n' : _ | RN] -> - [Γ ||- P[n..] ≅ P[n'..] | RPpt _ Rn]. - Proof. - intros. eapply transEq; [| eapply LRTyEqSym ]; eapply RPQext; cycle 1; tea. - now eapply reflLRTmEq. - Unshelve. 2,3: eauto. - Qed. + (RPpt : forall n n', [Γ ||- n ≅ n' : _ | RN] -> [Γ ||- P[n..]]) + (RPQext : forall n n' (Rn : [Γ ||- n ≅ n' : _ | RN]), + [Γ ||- P[n..] ≅ Q[n'..] | RPpt _ _ Rn]). - Lemma emptyElimRedAuxLeft : @emptyRedElimStmt _ _ P NN RPpt. - Proof. - eapply emptyElimRedAux; tea. - + eapply RPext. - Qed. + #[local] + Lemma RQpt : forall n n', [Γ ||- n ≅ n': _ | RN] -> [Γ ||- Q[n..]]. + Proof. intros; now unshelve eapply LRTyEqRedRight, RPQext. Qed. - Lemma emptyElimRedAuxRight : @emptyRedElimStmt _ _ Q NN RQpt. + #[local] + Lemma RPext : forall n n' (Rn : [Γ ||- n ≅ n' : _ | RN]), + [Γ ||- P[n..] ≅ P[n'..] | RPpt _ _ Rn]. Proof. - eapply emptyElimRedAux; tea. - + intros. eapply transEq; [eapply LRTyEqSym |]; eapply RPQext; cycle 1; tea. - now eapply reflLRTmEq. - Unshelve. all:tea. + intros; eapply transEq; [| eapply LRTyEqSym ]. + all: unshelve (eapply RPQext; cycle 1; tea); now eapply urefl. + Unshelve. 2: eapply RQpt; now symmetry. Qed. - Lemma emptyElimRedEqAux : - (forall n n' (Rnn' : [Γ ||- n ≅ n' : _ | RN]) (Rn : [Γ ||- n : _ | RN]) (Rn' : [Γ ||- n' : _ | RN]), - [Γ ||- tEmptyElim P n ≅ tEmptyElim Q n' : _ | RPpt _ Rn ]) × - (forall n n' (Rnn' : EmptyPropEq Γ n n') (Rn : [Γ ||- n : _ | RN]) (Rn' : [Γ ||- n' : _ | RN]), - [Γ ||- tEmptyElim P n ≅ tEmptyElim Q n' : _ | RPpt _ Rn ]). + Lemma emptyElimRedEq n n' (Rnn' : [Γ ||- n ≅ n' : _ | RN]) : + [Γ ||- tEmptyElim P n ≅ tEmptyElim Q n' : _ | RPpt _ _ Rnn' ]. Proof. - apply EmptyRedEqInduction. - - intros ?? nfL nfR redL redR ? prop ih RL RR; edestruct @EmptyPropEq_whnf; eauto. - assert [Γ ||- nfL : _ | RN] by (eapply redTmFwd; cycle 1; tea). - assert [Γ ||- nfR : _ | RN] by (eapply redTmFwd; cycle 1; tea). - eapply LREqTermHelper. - * eapply (fst emptyElimRedAuxLeft). - * eapply (fst emptyElimRedAuxRight). - * eapply RPQext. 1: tea. - now econstructor. - * eapply LRTmEqRedConv. - + eapply RPext; tea. - eapply LRTmEqSym; eapply redwfSubstTerm; cycle 1; tea. - + unshelve erewrite (redtmwf_det _ _ (EmptyRedTm.red RL) redL); tea. - 1: dependent inversion RL; subst; cbn; now eapply EmptyProp_whnf. - unshelve erewrite (redtmwf_det _ _ (EmptyRedTm.red RR) redR); tea. - 1: dependent inversion RR; subst; cbn; now eapply EmptyProp_whnf. - now eapply ih. - Unshelve. tea. 2, 4: tea. - - intros ?? neueq ??. escape. inversion neueq. - assert [Γ |- P[ne..] ≅ Q[ne'..]]. 1:{ - eapply escapeEq. eapply RPQext; tea. - econstructor. - 1,2: now eapply redtmwf_refl. - 2: now constructor. - gen_typing. - } + inversion Rnn' as [??????? prop]; subst. + inversion prop as [?? h]; pose proof h as [?? conv]; subst. + pose proof (convneu_whne conv); pose proof (convneu_whne (symmetry conv)). + pose proof (Rne := neNfTermEq RN h); pose proof (Rnne := redwfSubstTmEq _ (lrefl Rne) redL). + pose proof (RPQext _ _ Rne). + escape. + eapply redSubstTmEq; cycle 1. + + eapply redtm_emptyelim; tea; gen_typing. + + eapply redtm_emptyelim; tea; gen_typing. + + now unshelve eapply escapeEq, RPQext. + + eapply LRTmEqConv; [unshelve eapply RPext;[|now symmetry]|]. eapply neuTermEq. - + eapply ty_emptyElim; tea. - + eapply ty_conv. - 1: eapply ty_emptyElim; tea. - now symmetry. - + eapply convneu_emptyElim ; tea. - Unshelve. tea. + * eapply ty_emptyElim; tea. + * eapply ty_conv; [now eapply ty_emptyElim|now symmetry]. + * now eapply convneu_emptyElim. Qed. - - Lemma emptyElimRedEq : - (forall n n' (Rnn' : [Γ ||- n ≅ n' : _ | RN]) (Rn : [Γ ||- n : _ | RN]) (Rn' : [Γ ||- n' : _ | RN]), - [Γ ||- tEmptyElim P n ≅ tEmptyElim Q n' : _ | RPpt _ Rn ]). - Proof. apply emptyElimRedEqAux. Qed. End EmptyElimRedEq. - Section EmptyElimValid. - Context {Γ l P} + Context {Γ l} (VΓ : [||-v Γ]) (VN := emptyValid (l:=l) VΓ) (VΓN := validSnoc VΓ VN) - (VP : [Γ ,, tEmpty ||-v P | VΓN]). - + { P P' } + (VP : [Γ ,, tEmpty ||-v P | VΓN]) + (VPP' : [Γ ,, tEmpty ||-v P ≅ P' | VΓN | VP ]) + (VP' := ureflValidTy VPP'). - Lemma emptyElimValid {n} - (Vn : [Γ ||-v n : tEmpty | VΓ | VN]) + Lemma emptyElimCongValid {n n'} + (Vn : [Γ ||-v n ≅ n' : tEmpty | VΓ | VN]) (VPn := substS VP Vn) - : [Γ ||-v tEmptyElim P n : _ | VΓ | VPn]. + : [Γ ||-v tEmptyElim P n ≅ tEmptyElim P' n' : _ | VΓ | VPn]. Proof. - constructor; intros. - - instValid Vσ; cbn. - irrelevance0. 1: now rewrite singleSubstComm'. - epose proof (Vuσ := liftSubstS' VN Vσ). - instValid Vuσ; escape. - unshelve eapply emptyElimRed; tea. - + intros m Rm. - rewrite up_single_subst. - unshelve eapply validTy; cycle 1; tea. - unshelve econstructor; [now bsimpl| now cbn]. - + intros m m' Rm Rm' Rmm'; cbn. - irrelevance0. 1: now rewrite up_single_subst. - rewrite up_single_subst. - unshelve eapply validTyExt; cycle 1 ; tea. - 1,2: unshelve econstructor; [now bsimpl| now cbn]. - unshelve econstructor; [|now cbn]. - bsimpl; cbn; eapply reflSubst. - - instAllValid Vσ Vσ' Vσσ'. - irrelevance0. 1: now rewrite singleSubstComm'. - pose (Vuσ := liftSubstS' VN Vσ). - pose proof (Vuσ' := liftSubstS' VN Vσ'). - pose proof (Vuσrea := liftSubstSrealign' VN Vσσ' Vσ'). - pose proof (Vuσσ' := liftSubstSEq' VN Vσσ'). - instValid Vuσ'; instAllValid Vuσ Vuσrea Vuσσ'; escape. - unshelve eapply emptyElimRedEq; tea; fold subst_term. - all: try now (irrelevance; now rewrite elimSuccHypTy_subst). - + intros m Rm. - rewrite up_single_subst. - unshelve eapply validTy; cycle 1; tea. - unshelve econstructor; [now bsimpl| now cbn]. - + intros m Rm. - rewrite up_single_subst. - unshelve eapply validTy; cycle 1; tea. - unshelve econstructor; [now bsimpl| now cbn]. - + intros m m' Rm Rm' Rmm'; cbn. - irrelevance0. 1: now rewrite up_single_subst. - rewrite up_single_subst. - unshelve eapply validTyExt; cycle 1 ; tea. - 1,2: unshelve econstructor; [now bsimpl| now cbn]. - unshelve econstructor; [|now cbn]. - now bsimpl. + constructor; intros; instValid Vσσ'; epose proof (Vuσ := liftSubstEq' VN Vσσ'). + instValid Vuσ; cbn in *; escape. + irrelevance0. 1: now rewrite singleSubstComm'. + eapply emptyElimRedEq; rewrite ?elimSuccHypTy_subst; tea. + + intros; irrelevance0; rewrite up_single_subst; [reflexivity|]. + unshelve eapply validTyEq; cycle 1 ; tea. + now unshelve eapply consSubstEq. + Unshelve. + 2: tea. + intros ?? Rn; rewrite up_single_subst. + unshelve eapply validTy. + 5: tea. + 3: unshelve eapply consSubstEq; tea. Qed. - End EmptyElimValid. -Lemma emptyElimCongValid {Γ l P Q n n'} +Lemma emptyElimValid {Γ l P n n'} (VΓ : [||-v Γ]) (VN := emptyValid (l:=l) VΓ) (VΓN := validSnoc VΓ VN) (VP : [Γ ,, tEmpty ||-v P | VΓN]) - (VQ : [Γ ,, tEmpty ||-v Q | VΓN]) - (VPQ : [Γ ,, tEmpty ||-v P ≅ Q | VΓN | VP]) - (Vn : [Γ ||-v n : _ | VΓ | VN]) - (Vn' : [Γ ||-v n' : _ | VΓ | VN]) - (Veqn : [Γ ||-v n ≅ n' : _ | VΓ | VN]) - (VPn := substS VP Vn) : - [Γ ||-v tEmptyElim P n ≅ tEmptyElim Q n' : _ | VΓ | VPn]. + (Vn : [Γ ||-v n ≅ n' : _ | VΓ | VN]) + (VPn := substS VP Vn): + [Γ ||-v tEmptyElim P n : _ | VΓ | VPn]. Proof. - constructor; intros. - pose (Vuσ := liftSubstS' VN Vσ). - instValid Vσ; instValid Vuσ; escape. - irrelevance0. 1: now rewrite singleSubstComm'. - unshelve eapply emptyElimRedEq; tea; fold subst_term. - all: try (irrelevance; now rewrite elimSuccHypTy_subst). - + intros m Rm. - rewrite up_single_subst. - unshelve eapply validTy; cycle 1; tea. - unshelve econstructor; [now bsimpl| now cbn]. - + intros m Rm. - rewrite up_single_subst. - unshelve eapply validTy; cycle 1; tea. - unshelve econstructor; [now bsimpl| now cbn]. - + intros m m' Rm Rm' Rmm'; cbn. - irrelevance0. 1: now rewrite up_single_subst. - rewrite up_single_subst. - eapply transEq; cycle 1. - - unshelve eapply validTyEq. - 7: tea. 1: tea. - unshelve econstructor; [now bsimpl| now cbn]. - - unshelve eapply validTyExt; cycle 1 ; tea. - 1,2: unshelve econstructor; [now bsimpl| now cbn]. - unshelve econstructor; [|now cbn]. - bsimpl. eapply reflSubst. + eapply lreflValidTm , emptyElimCongValid; tea. + now eapply reflValidTy. Qed. End Empty. diff --git a/theories/Substitution/Introductions/Id.v b/theories/Substitution/Introductions/Id.v index 1c17324..d7d574f 100644 --- a/theories/Substitution/Introductions/Id.v +++ b/theories/Substitution/Introductions/Id.v @@ -1,129 +1,105 @@ +From Coq Require Import ssrbool CRelationClasses. From LogRel.AutoSubst Require Import core unscoped Ast Extra. From LogRel Require Import Utils BasicAst Notations Context NormalForms Weakening GenericTyping LogicalRelation Validity. -From LogRel.LogicalRelation Require Import Induction Irrelevance Escape Reflexivity Weakening Neutral Transitivity Reduction Application Universe Id. +From LogRel.LogicalRelation Require Import Induction Escape Reflexivity Irrelevance Weakening Neutral Transitivity Reduction Application Universe Id EqRedRight NormalRed InstKripke. From LogRel.Substitution Require Import Irrelevance Properties Conversion SingleSubst Reflexivity Reduction. -From LogRel.Substitution.Introductions Require Import Universe Var. +From LogRel.Substitution.Introductions Require Import Universe Var Poly. + +Set Universe Polymorphism. +Set Printing Primitive Projection Parameters. Set Universe Polymorphism. Section Id. Context `{GenericTypingProperties}. - Lemma IdValid {Γ l A x y} - (VΓ : [||-v Γ]) - (VA : [_ ||-v A | VΓ]) - (Vx : [_ ||-v x : _ | _ | VA]) - (Vy : [_ ||-v y : _ | _ | VA]) : + Lemma IdValid {Γ l A x x' y y'} + (VΓ : [||-v Γ]) + (VA : [_ ||-v A | VΓ]) + (Vx : [_ ||-v x ≅ x' : _ | _ | VA]) + (Vy : [_ ||-v y ≅ y' : _ | _ | VA]) : [_ ||-v tId A x y | VΓ]. Proof. + pose proof (lreflValidTm _ Vx). + pose proof (lreflValidTm _ Vy). unshelve econstructor; intros. - - instValid vσ; cbn; now eapply IdRed. - - instAllValid vσ vσ' vσσ'; eapply IdCongRed; refold; tea. - eapply wft_Id; now escape. + - instValid (lrefl vσ); cbn; now eapply IdRed. + - instValid vσσ'; eapply IdCongRed; refold; tea. + eapply wft_Id; escape; tea; now eapply ty_conv. Qed. Lemma IdCongValid {Γ l A A' x x' y y'} - (VΓ : [||-v Γ]) - (VA : [_ ||-v A | VΓ]) - (VA' : [_ ||-v A' | VΓ]) - (VAA' : [_ ||-v A ≅ A' | VΓ | VA]) - (Vx : [_ ||-v x : _ | _ | VA]) - (Vx' : [_ ||-v x' : _ | _ | VA]) + (VΓ : [||-v Γ]) + (VA : [_ ||-v A | VΓ]) + (VAA' : [_ ||-v A ≅ A' | VΓ | VA]) (Vxx' : [_ ||-v x ≅ x' : _ | _ | VA]) - (Vy : [_ ||-v y : _ | _ | VA]) - (Vy' : [_ ||-v y' : _ | _ | VA]) - (Vyy' : [_ ||-v y ≅ y' : _ | _ | VA]) + (Vyy' : [_ ||-v y ≅ y' : _ | _ | VA]) (VId : [_ ||-v tId A x y | VΓ]) : [_ ||-v tId A x y ≅ tId A' x' y' | VΓ | VId]. Proof. constructor; intros. - instValid Vσ; eapply IdCongRed; refold; tea. - eapply wft_Id. - 2,3: eapply ty_conv. - all: now escape. + instValid Vσσ'; eapply IdCongRed; refold; tea. + escape; eapply wft_Id; tea; now eapply ty_conv. Qed. - - Lemma IdTmValid {Γ l A x y} - (VΓ : [||-v Γ]) - (VU := UValid VΓ) - (VAU : [_ ||-v A : U | VΓ | VU]) - (VA := univValid VΓ VU VAU) - (Vx : [_ ||-v x : _ | _ | VA]) - (Vy : [_ ||-v y : _ | _ | VA]) : - [_ ||-v tId A x y : _ | VΓ | VU]. - Proof. - unshelve econstructor; intros. - - instValid Vσ. - unshelve eapply IdRedU; refold; tea. - 1: now eapply univValid. - all: irrelevance. - - instAllValid Vσ Vσ' Vσσ'. - unshelve eapply IdCongRedU; refold; tea. - 1,2: now eapply univValid. - all: irrelevance. - Qed. - Lemma IdTmCongValid {Γ l A A' x x' y y'} - (VΓ : [||-v Γ]) + (VΓ : [||-v Γ]) (VU := UValid VΓ) - (VAU : [_ ||-v A : _ | VΓ | VU]) + (VAUeq : [_ ||-v A ≅ A' : U | VΓ | VU]) + (VAU := (lreflValidTm _ VAUeq)) (VA := univValid VΓ VU VAU) - (VAU' : [_ ||-v A' : _ | VΓ | VU]) - (VAAU' : [_ ||-v A ≅ A' : _ | VΓ | VU]) - (Vx : [_ ||-v x : _ | _ | VA]) - (Vx' : [_ ||-v x' : _ | _ | VA]) - (Vxx' : [_ ||-v x ≅ x' : _ | _ | VA]) - (Vy : [_ ||-v y : _ | _ | VA]) - (Vy' : [_ ||-v y' : _ | _ | VA]) - (Vyy' : [_ ||-v y ≅ y' : _ | _ | VA]) : + (Vx : [_ ||-v x ≅ x' : _ | _ | VA]) + (Vy : [_ ||-v y ≅ y' : _ | _ | VA]) : [_ ||-v tId A x y ≅ tId A' x' y' : _ | VΓ | VU]. Proof. - constructor; intros; instValid Vσ. + constructor; intros; instValid Vσσ'. unshelve eapply IdCongRedU; refold; tea. - 1,2: now eapply univValid. - 2,5: eapply LRTmRedConv; [eapply univEqValid; irrValid|]. + 1: now eapply univValid. + 1: now eapply lrefl. all: irrelevance. - Unshelve. 3,9: now eapply univValid. - all: tea. Qed. - - Lemma reflValid {Γ l A x} - (VΓ : [||-v Γ]) - (VA : [_ ||-v A | VΓ]) + Lemma IdTmValid {Γ l A x y} + (VΓ : [||-v Γ]) + (VU := UValid VΓ) + (VAU : [_ ||-v A : U | VΓ | VU]) + (VA := univValid VΓ VU VAU) (Vx : [_ ||-v x : _ | _ | VA]) - (VId : [_ ||-v tId A x x | VΓ]) : - [_ ||-v tRefl A x : _ | _ | VId]. + (Vy : [_ ||-v y : _ | _ | VA]) : + [_ ||-v tId A x y : _ | VΓ | VU]. Proof. - unshelve econstructor; intros. - - instValid Vσ; now unshelve eapply reflRed. - - instAllValid Vσ Vσ' Vσσ'; escape; now unshelve eapply reflCongRed. + unshelve eapply IdTmCongValid; tea; irrValid. Qed. Lemma reflCongValid {Γ l A A' x x'} - (VΓ : [||-v Γ]) - (VA : [_ ||-v A | VΓ]) - (VA' : [_ ||-v A' | VΓ]) - (VAA' : [_ ||-v A ≅ A' | VΓ | VA]) - (Vx : [_ ||-v x : _ | _ | VA]) - (Vx' : [_ ||-v x' : _ | _ | VA]) - (Vxx' : [_ ||-v x ≅ x' : _ | _ | VA]) + (VΓ : [||-v Γ]) + (VA : [_ ||-v A | VΓ]) + (VAA' : [_ ||-v A ≅ A' | VΓ | VA]) + (Vx : [_ ||-v x ≅ x' : _ | _ | VA]) (VId : [_ ||-v tId A x x | VΓ]) : [_ ||-v tRefl A x ≅ tRefl A' x' : _ | _ | VId]. Proof. - constructor; intros; instValid Vσ. - escape; unshelve eapply reflCongRed; refold; tea. - now eapply ty_conv. + constructor; intros; instValid Vσσ'. + escape; now unshelve eapply reflCongRed. + Qed. + + Lemma reflValid {Γ l A x x'} + (VΓ : [||-v Γ]) + (VA : [_ ||-v A | VΓ]) + (Vx : [_ ||-v x ≅ x': _ | _ | VA]) + (VId : [_ ||-v tId A x x | VΓ]) : + [_ ||-v tRefl A x : _ | _ | VId]. + Proof. + eapply reflCongValid;[eapply reflValidTy| now eapply lreflValidTm]. Qed. Lemma subst_scons2 (P e y : term) (σ : nat -> term) : P[e .: y..][σ] = P [e[σ] .: (y[σ] .: σ)]. Proof. now asimpl. Qed. - + Lemma subst_upup_scons2 (P e y : term) (σ : nat -> term) : P[up_term_term (up_term_term σ)][e .: y..] = P [e .: (y .: σ)]. Proof. now asimpl. Qed. - Lemma consSubstS' {Γ σ t l A Δ} + (* Lemma consSubstS' {Γ σ t l A Δ} (VΓ : [||-v Γ]) (VΓA : [||-v Γ ,, A]) (wfΔ : [|- Δ]) @@ -131,12 +107,12 @@ Context `{GenericTypingProperties}. (VA : [Γ ||-v A | VΓ]) (Vt : [ Δ ||- t : A[σ] | validTy VA wfΔ Vσ]) : [Δ ||-v (t .: σ) : Γ ,, A | VΓA | wfΔ]. - Proof. + Proof. pose proof (invValiditySnoc VΓA) as [? [? [? ->]]]. unshelve eapply consSubstS; [ irrValid| irrelevance]. Qed. - Lemma consSubstSEq'' {Γ σ σ' t u l A Δ} + Lemma consSubstSEq'' {Γ σ σ' t u l A Δ} (VΓ : [||-v Γ]) (VΓA : [||-v Γ ,, A]) (wfΔ : [|- Δ]) @@ -151,604 +127,208 @@ Context `{GenericTypingProperties}. pose proof (invValiditySnoc VΓA) as [? [? [? ->]]]. pose proof (consSubstSEq' VΓ wfΔ Vσ Vσσ' VA Vt Vtu). irrValid. - Qed. + Qed. + consSubstEq *) - Lemma idElimMotiveCtxIdValid {Γ l A x} - (VΓ : [||-v Γ]) - (VA : [_ ||-v A | VΓ]) - (Vx : [_ ||-v x : _ | _ | VA]) : + Lemma idElimMotiveCtxIdValid {Γ l A x x'} + (VΓ : [||-v Γ]) + (VA : [_ ||-v A | VΓ]) + (Vx : [_ ||-v x ≅ x' : _ | _ | VA]) : [Γ,, A ||-v< l > tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0) | validSnoc VΓ VA]. Proof. unshelve eapply IdValid. - - now eapply wk1ValidTy. - - now eapply wk1ValidTm. - - exact (var0Valid VΓ VA). + 3: now eapply wk1ValidTy. + 3: now eapply wk1ValidTmEq. + 2: exact (var0Valid VΓ VA). Qed. - + Lemma idElimMotiveCtxIdCongValid {Γ l A A' x x'} - (VΓ : [||-v Γ]) - (VA : [_ ||-v A | VΓ]) - (VA' : [_ ||-v A' | VΓ]) + (VΓ : [||-v Γ]) + (VA : [_ ||-v A | VΓ]) (VAA' : [_ ||-v A ≅ A' | VΓ | VA]) - (Vx : [_ ||-v x : _ | _ | VA]) - (Vx' : [_ ||-v x' : _ | _ | VA]) - (Vxx' : [_ ||-v x ≅ x' : _ | _ | VA]) + (Vxx' : [_ ||-v x ≅ x' : _ | _ | VA]) (VId : [Γ,, A ||-v< l > tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0) | validSnoc VΓ VA]) : [_ ||-v _ ≅ tId A'⟨@wk1 Γ A'⟩ x'⟨@wk1 Γ A'⟩ (tRel 0) | validSnoc VΓ VA | VId]. Proof. assert (h : forall t, t⟨@wk1 Γ A'⟩ = t⟨@wk1 Γ A⟩) by reflexivity. - unshelve eapply IdCongValid. + unshelve eapply IdCongValid; rewrite ?h. - now eapply wk1ValidTy. - - rewrite h; now eapply wk1ValidTy. - - rewrite h; now eapply wk1ValidTyEq. - - now eapply wk1ValidTm. - - rewrite h; now eapply wk1ValidTm. - - rewrite h; now eapply wk1ValidTmEq. + - now eapply wk1ValidTyEq. + - now eapply wk1ValidTmEq. - eapply var0Valid. - - eapply var0Valid. - - eapply reflValidTm; eapply var0Valid. Qed. - - - Lemma idElimMotiveScons2Valid {Γ l A x y e} - (VΓ : [||-v Γ]) - (VA : [_ ||-v A | VΓ]) - (Vx : [_ ||-v x : _ | _ | VA]) - (Vy : [Γ ||-v y : _ | _ | VA]) + + Lemma idElimMotiveCtxEq {Γ l A A' x x'} + (VΓ : [||-v Γ]) + (VA : [_ ||-v A | VΓ]) + (VAA' : [_ ||-v A ≅ A' | VΓ | VA]) + (Vxx' : [_ ||-v x ≅ x' : _ | _ | VA]) : + [||-v (Γ,, A ,, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0)) ≅ (Γ,, A' ,, tId A'⟨@wk1 Γ A'⟩ x'⟨@wk1 Γ A'⟩ (tRel 0))]. + Proof. + unshelve eapply validCtxSnoc'. + 2: unshelve eapply validCtxSnoc'; [| now eapply reflValidCtx|..]; cbn; tea. + 3: cbn; now eapply idElimMotiveCtxIdCongValid. + cbn; now eapply idElimMotiveCtxIdValid. + Defined. + + + Lemma idElimMotiveScons2Valid {Γ l A x x' y y' e e'} + (VΓ : [||-v Γ]) + (VA : [_ ||-v A | VΓ]) + (Vx : [_ ||-v x ≅ x' : _ | _ | VA]) + (Vy : [Γ ||-v y ≅ y' : _ | _ | VA]) (VId : [Γ ||-v tId A x y | VΓ]) - (Ve : [_ ||-v e : _ | _ | VId]) + (Ve : [_ ||-v e ≅ e' : _ | _ | VId]) (VΓext : [||-v (Γ ,, A) ,, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0)]) - Δ σ (wfΔ: [ |-[ ta ] Δ]) (vσ: [VΓ | Δ ||-v σ : _ | wfΔ]) : - [VΓext | Δ ||-v (e[σ] .: (y[σ] .: σ)) : _ | wfΔ]. + Δ (wfΔ: [ |-[ ta ] Δ]) {σ σ'} (Vσσ': [VΓ | Δ ||-v σ ≅ σ' : _ | wfΔ]) : + [VΓext | Δ ||-v (e[σ] .: (y[σ] .: σ)) ≅ (e'[σ'] .: (y'[σ'] .: σ')) : _ | wfΔ]. Proof. - intros; unshelve eapply consSubstS'; tea. - 2: now eapply consSubstSvalid. - 1: now eapply idElimMotiveCtxIdValid. - instValid vσ; irrelevance. + epose (Vσy := consSubstEqvalid Vσσ' Vy). + unshelve epose (consSubstEq _ _ Vσy (idElimMotiveCtxIdValid VΓ VA Vx) _). + 4: irrValid. + instValid Vσσ'; irrelevance. Qed. - - Lemma substIdElimMotive {Γ l A x P y e} - (VΓ : [||-v Γ]) - (VA : [_ ||-v A | VΓ]) - (Vx : [_ ||-v x : _ | _ | VA]) + + Lemma substIdElimMotive {Γ l A x x' P y y' e e'} + (VΓ : [||-v Γ]) + (VA : [_ ||-v A | VΓ]) + (Vx : [_ ||-v x ≅ x' : _ | _ | VA]) (VΓext : [||-v (Γ ,, A) ,, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0)]) (VP : [_ ||-v P | VΓext]) - (Vy : [Γ ||-v y : _ | _ | VA]) + (Vy : [Γ ||-v y ≅ y' : _ | _ | VA]) (VId : [Γ ||-v tId A x y | VΓ]) - (Ve : [_ ||-v e : _ | _ | VId]) : + (Ve : [_ ||-v e ≅ e' : _ | _ | VId]) : [_ ||-v P[e .: y ..] | VΓ]. Proof. opector; intros. - rewrite subst_scons2; eapply validTy; tea; now eapply idElimMotiveScons2Valid. - irrelevance0 ; rewrite subst_scons2;[reflexivity|]. unshelve eapply validTyExt. - 5,6: eapply idElimMotiveScons2Valid; cycle 1; tea. - 1: tea. - eapply consSubstSEq''. - + now unshelve eapply consSubstSEqvalid. - + instValid vσ; irrelevance. - + instAllValid vσ vσ' vσσ'; irrelevance. - Unshelve. 2: now eapply idElimMotiveCtxIdValid. + 5: eapply idElimMotiveScons2Valid; cycle 1; first [now eapply lreflValidTm| tea]. + tea. Qed. - - Lemma subst_idElimMotive_substupup {Γ Δ l A x P y e σ} - (VΓ : [||-v Γ]) - (wfΔ : [|- Δ]) - (Vσ : [_ ||-v σ : _ | VΓ | wfΔ]) - (VA : [_ ||-v A | VΓ]) - (RVA := validTy VA wfΔ Vσ) - (Vx : [_ ||-v x : _ | _ | VA]) - (VΓext : [||-v (Γ ,, A) ,, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0)]) - (VP : [_ ||-v P | VΓext]) - (Ry : [ RVA | _ ||- y : _]) - (RId : [Δ ||- tId A[σ] x[σ] y]) - (Re : [RId | _ ||- e : _]) : - [Δ ||- P[up_term_term (up_term_term σ)][e .: y ..]]. - Proof. - pose (VΓA := validSnoc VΓ VA). - rewrite subst_upup_scons2. - unshelve eapply validTy; cycle 2; tea. - unshelve eapply consSubstS'; tea. - + now eapply consSubstS. - + now eapply idElimMotiveCtxIdValid. - + irrelevance. - Qed. - - Lemma irrelevanceSubst' {Γ} (VΓ VΓ' : [||-v Γ]) {σ Δ Δ'} (wfΔ : [|- Δ]) (wfΔ' : [|- Δ']) : Δ = Δ' -> - [Δ ||-v σ : Γ | VΓ | wfΔ] -> [Δ' ||-v σ : Γ | VΓ' | wfΔ']. - Proof. - intros ->; eapply irrelevanceSubst. - Qed. + Lemma up_twice_subst t a b σ : + t[up_term_term (up_term_term σ)][a[σ] .: b[σ]..] = + t[a .: b..][σ]. + Proof. now bsimpl. Qed. - Lemma idElimMotive_Idsubst_eq {Γ Δ A x σ} : + Lemma idElimMotive_Idsubst_eq {Γ Δ A x σ} : tId A[σ]⟨@wk1 Δ A[σ]⟩ x[σ]⟨@wk1 Δ A[σ]⟩ (tRel 0) = (tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0))[up_term_term σ]. Proof. now bsimpl. Qed. - - Lemma red_idElimMotive_substupup {Γ Δ l A x P σ} - (VΓ : [||-v Γ]) - (wfΔ : [|- Δ]) - (Vσ : [_ ||-v σ : _ | VΓ | wfΔ]) - (VA : [_ ||-v A | VΓ]) - (Vx : [_ ||-v x : _ | _ | VA]) - (VΓext : [||-v (Γ ,, A) ,, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0)]) - (VP : [_ ||-v P | VΓext]) : - [(Δ ,, A[σ]),, tId A[σ]⟨@wk1 Δ A[σ]⟩ x[σ]⟨@wk1 Δ A[σ]⟩ (tRel 0) ||- P[up_term_term (up_term_term σ)]]. - Proof. - pose (VΓA := validSnoc VΓ VA). - instValid Vσ. - unshelve eapply validTy; cycle 2; tea. - * escape. - assert [|- Δ,, A[σ]] by now eapply wfc_cons. - eapply wfc_cons; tea. - eapply wft_Id. - 1: now eapply wft_wk. - 1: now eapply ty_wk. - rewrite wk1_ren_on; now eapply ty_var0. - * epose (v1:= liftSubstS' VA Vσ). - epose (v2:= liftSubstS' _ v1). - eapply irrelevanceSubst'. - 1: now erewrite idElimMotive_Idsubst_eq. - eapply v2. - Unshelve. 2: now eapply idElimMotiveCtxIdValid. - Qed. - - Lemma irrelevanceSubstEq' {Γ} (VΓ VΓ' : [||-v Γ]) {σ σ' Δ Δ'} (wfΔ : [|- Δ]) (wfΔ' : [|- Δ']) - (Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ]) (Vσ' : [Δ' ||-v σ : Γ | VΓ' | wfΔ']) : - Δ = Δ' -> - [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ | Vσ] -> [Δ' ||-v σ ≅ σ' : Γ | VΓ' | wfΔ' | Vσ']. - Proof. - intros ->; eapply irrelevanceSubstEq. - Qed. - - Lemma red_idElimMotive_substupup_cong {Γ Δ l A x P σ σ'} - (VΓ : [||-v Γ]) - (wfΔ : [|- Δ]) - (Vσ : [_ ||-v σ : _ | VΓ | wfΔ]) - (Vσ' : [_ ||-v σ' : _ | VΓ | wfΔ]) - (Vσσ' : [_ ||-v σ ≅ σ' : _ | VΓ | wfΔ | Vσ]) - (VA : [_ ||-v A | VΓ]) - (Vx : [_ ||-v x : _ | _ | VA]) - (VΓext : [||-v (Γ ,, A) ,, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0)]) - (VP : [_ ||-v P | VΓext]) - (RPσ : [(Δ ,, A[σ]),, tId A[σ]⟨@wk1 Δ A[σ]⟩ x[σ]⟨@wk1 Δ A[σ]⟩ (tRel 0) ||- P[up_term_term (up_term_term σ)]]) : - [RPσ | _ ||- _ ≅ P[up_term_term (up_term_term σ')]]. - Proof. - pose (VΓA := validSnoc VΓ VA). - instAllValid Vσ Vσ' Vσσ'. - irrelevanceRefl; unshelve eapply validTyExt; cycle 2; tea. - * escape. - assert [|- Δ,, A[σ]] by now eapply wfc_cons. - eapply wfc_cons; tea. - eapply wft_Id. - 1: now eapply wft_wk. - 1: now eapply ty_wk. - rewrite wk1_ren_on; now eapply ty_var0. - * epose (v1:= liftSubstS' VA Vσ). - epose (v2:= liftSubstS' _ v1). - eapply irrelevanceSubst'. - 1: now erewrite idElimMotive_Idsubst_eq. - eapply v2. - Unshelve. 2: now eapply idElimMotiveCtxIdValid. - * eapply irrelevanceSubst'. - 1: now erewrite idElimMotive_Idsubst_eq. - eapply liftSubstSrealign'. - + now eapply liftSubstSEq'. - + now eapply liftSubstSrealign'. - Unshelve. - 2: now eapply idElimMotiveCtxIdValid. - * eapply irrelevanceSubstEq'. - 1: now erewrite idElimMotive_Idsubst_eq. - eapply liftSubstSEq'. - now eapply liftSubstSEq'. - Unshelve. - 2: now eapply idElimMotiveCtxIdValid. - Qed. - - Lemma redEq_idElimMotive_substupup {Γ Δ l A A' x x' P P' σ} - (VΓ : [||-v Γ]) - (wfΔ : [|- Δ]) - (Vσ : [_ ||-v σ : _ | VΓ | wfΔ]) - (VA : [_ ||-v A | VΓ]) - (VA' : [_ ||-v A' | VΓ]) - (VAA' : [_ ||-v A ≅ A' | _ | VA]) - (Vx : [_ ||-v x : _ | _ | VA]) - (Vx' : [_ ||-v x' : _ | _ | VA]) - (Vxx' : [_ ||-v x ≅ x' : _ | _ | VA]) - (VΓext : [||-v (Γ ,, A) ,, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0)]) - (VP : [_ ||-v P | VΓext]) - (VP' : [_ ||-v P' | VΓext]) - (VPP' : [_ ||-v P ≅ P' | _ | VP]) - (VPsubstupup : [(Δ ,, A[σ]),, tId A[σ]⟨@wk1 Δ A[σ]⟩ x[σ]⟨@wk1 Δ A[σ]⟩ (tRel 0) ||- P[up_term_term (up_term_term σ)]]) : - [_ ||- _ ≅ P'[up_term_term (up_term_term σ)] | VPsubstupup]. - Proof. - pose (VΓA := validSnoc VΓ VA). - instValid Vσ. - irrelevanceRefl; unshelve eapply validTyEq; cycle 2; tea. - * escape. - assert [|- Δ,, A[σ]] by now eapply wfc_cons. - eapply wfc_cons; tea. - eapply wft_Id. - 1: now eapply wft_wk. - 1: now eapply ty_wk. - rewrite wk1_ren_on; now eapply ty_var0. - * epose (v1:= liftSubstS' VA Vσ). - epose (v2:= liftSubstS' _ v1). - eapply irrelevanceSubst'. - 1: now erewrite idElimMotive_Idsubst_eq. - eapply v2. - Unshelve. 2: now eapply idElimMotiveCtxIdValid. - Qed. - - - Lemma substEq_idElimMotive_substupupEq {Γ Δ l A x P y y' e e' σ σ'} - (VΓ : [||-v Γ]) - (wfΔ : [|- Δ]) - (Vσ : [_ ||-v σ : _ | VΓ | wfΔ]) - (Vσ' : [_ ||-v σ' : _ | VΓ | wfΔ]) - (Vσσ' : [_ ||-v σ ≅ σ' : _ | VΓ | wfΔ | Vσ]) - (VA : [_ ||-v A | VΓ]) - (RVA := validTy VA wfΔ Vσ) - (RVA' := validTy VA wfΔ Vσ') - (Vx : [_ ||-v x : _ | _ | VA]) - (VΓext : [||-v (Γ ,, A) ,, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0)]) - (VP : [_ ||-v P | VΓext]) - (Ry : [ RVA | _ ||- y : _]) - (Ry' : [ RVA' | _ ||- y' : _]) - (Ryy' : [ RVA | _ ||- y ≅ y' : _]) - (RId : [Δ ||- tId A[σ] x[σ] y]) - (RId' : [Δ ||- tId A[σ'] x[σ'] y']) - (Re : [RId | _ ||- e : _]) - (Re' : [RId' | _ ||- e' : _]) - (Ree' : [RId | _ ||- e ≅ e' : _]) - (RPsubst : [Δ ||- P[up_term_term (up_term_term σ)][e .: y ..]]) : - [RPsubst | Δ ||- P[up_term_term (up_term_term σ)][e .: y ..] ≅ P[up_term_term (up_term_term σ')][e' .: y' ..]]. - Proof. - pose (VΓA := validSnoc VΓ VA). - irrelevance0; rewrite subst_upup_scons2; [reflexivity|]. - unshelve eapply validTyExt; cycle 2; tea. - - unshelve eapply consSubstS'; tea. - + now eapply consSubstS. - + now eapply idElimMotiveCtxIdValid. - + irrelevance. - - unshelve eapply consSubstS'; tea. - + now eapply consSubstS. - + now eapply idElimMotiveCtxIdValid. - + irrelevance. - - unshelve eapply consSubstSEq''; tea. - 4,5: irrelevance. - 2: now eapply idElimMotiveCtxIdValid. - 2: unshelve eapply consSubstSEq'; tea. - Qed. - - Lemma substEq_idElimMotive_substupup {Γ Δ l A x P y y' e e' σ} - (VΓ : [||-v Γ]) - (wfΔ : [|- Δ]) - (Vσ : [_ ||-v σ : _ | VΓ | wfΔ]) - (VA : [_ ||-v A | VΓ]) - (RVA := validTy VA wfΔ Vσ) - (Vx : [_ ||-v x : _ | _ | VA]) - (VΓext : [||-v (Γ ,, A) ,, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0)]) - (VP : [_ ||-v P | VΓext]) - (Ry : [ RVA | _ ||- y : _]) - (Ry' : [ RVA | _ ||- y' : _]) - (Ryy' : [ RVA | _ ||- y ≅ y' : _]) - (RId : [Δ ||- tId A[σ] x[σ] y]) - (RId' : [Δ ||- tId A[σ] x[σ] y']) - (Re : [RId | _ ||- e : _]) - (Re' : [RId' | _ ||- e' : _]) - (Ree' : [RId | _ ||- e ≅ e' : _]) - (RPsubst : [Δ ||- P[up_term_term (up_term_term σ)][e .: y ..]]) : - [RPsubst | Δ ||- P[up_term_term (up_term_term σ)][e .: y ..] ≅ P[up_term_term (up_term_term σ)][e' .: y' ..]]. - Proof. - eapply substEq_idElimMotive_substupupEq; tea; eapply reflSubst. - Qed. - - Set Printing Primitive Projection Parameters. - - Lemma substExt_idElimMotive_substupup {Γ Δ l A A' x x' P P' y y' e e' σ} - (VΓ : [||-v Γ]) - (wfΔ : [|- Δ]) - (Vσ : [_ ||-v σ : _ | VΓ | wfΔ]) - (VA : [_ ||-v A | VΓ]) - (VA' : [_ ||-v A' | VΓ]) - (VAA' : [_ ||-v A ≅ A' | _ | VA]) - (RVA := validTy VA wfΔ Vσ) - (Vx : [_ ||-v x : _ | _ | VA]) - (Vx' : [_ ||-v x' : _ | _ | VA]) - (Vxx' : [_ ||-v x ≅ x' : _ | _ | VA]) - (VΓext : [||-v (Γ ,, A) ,, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0)]) - (VP : [_ ||-v P | VΓext]) - (VP' : [_ ||-v P' | VΓext]) - (VPP' : [_ ||-v P ≅ P' | _ | VP]) - (Ry : [ RVA | _ ||- y : _]) - (Ry' : [ RVA | _ ||- y' : _]) - (Ryy' : [ RVA | _ ||- y ≅ y' : _]) - (RId : [Δ ||- tId A[σ] x[σ] y]) - (RId' : [Δ ||- tId A[σ] x[σ] y']) - (Re : [RId | _ ||- e : _]) - (Re' : [RId' | _ ||- e' : _]) - (Ree' : [RId | _ ||- e ≅ e' : _]) - (RPsubst : [Δ ||- P[up_term_term (up_term_term σ)][e .: y ..]]) : - [RPsubst | Δ ||- P[up_term_term (up_term_term σ)][e .: y ..] ≅ P'[up_term_term (up_term_term σ)][e' .: y' ..]]. - Proof. - pose (VΓA := validSnoc VΓ VA). - eapply LRTransEq. - eapply substEq_idElimMotive_substupup. - 2,4,8: tea. all:tea. - irrelevance0; rewrite subst_upup_scons2; [reflexivity|]. - unshelve eapply validTyEq; cycle 2; tea. - eapply irrelevanceSubst. - unshelve eapply consSubstS. - 3: now eapply idElimMotiveCtxIdValid. - 1: now eapply consSubstS. - irrelevance. - Unshelve. 2: eapply subst_idElimMotive_substupup; cycle 1; tea. - Qed. - - - Lemma substExtIdElimMotive {Γ l A x P y y' e e'} - (VΓ : [||-v Γ]) - (VA : [_ ||-v A | VΓ]) - (Vx : [_ ||-v x : _ | _ | VA]) - (VΓext : [||-v (Γ ,, A) ,, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0)]) - (VP : [_ ||-v P | VΓext]) - (Vy : [Γ ||-v y : _ | _ | VA]) - (Vy' : [Γ ||-v y' : _ | _ | VA]) - (Vyy' : [Γ ||-v y ≅ y' : _ | _ | VA]) - (VId : [Γ ||-v tId A x y | VΓ]) - (VId' : [Γ ||-v tId A x y' | VΓ]) - (Ve : [_ ||-v e : _ | _ | VId]) - (Ve' : [_ ||-v e' : _ | _ | VId']) - (Vee' : [_ ||-v e ≅ e' : _ | _ | VId]) - (VPey : [_ ||-v P[e .: y ..] | VΓ]) : - [_ ||-v P[e .: y ..] ≅ P[e' .: y' ..] | VΓ | VPey]. - Proof. - constructor; intros. - irrelevance0; rewrite subst_scons2; [reflexivity|]. - unshelve eapply validTyExt. - 5,6: eapply idElimMotiveScons2Valid; cycle 1; tea. - 1: tea. - instValid Vσ; unshelve eapply consSubstSEq''. - 4: now eapply idElimMotiveCtxIdValid. - 2: eapply consSubstSEq'; [now eapply reflSubst|]; irrelevance. - all: irrelevance. - Unshelve. 1:tea. irrelevance. - Qed. - Lemma substEqIdElimMotive {Γ l A A' x x' P P' y y' e e'} - (VΓ : [||-v Γ]) - (VA : [_ ||-v A | VΓ]) - (VA' : [_ ||-v A' | VΓ]) - (VAA' : [_ ||-v A ≅ A' | VΓ | VA]) - (Vx : [_ ||-v x : _ | _ | VA]) - (Vx' : [_ ||-v x' : _ | _ | VA]) - (Vxx' : [_ ||-v x ≅ x' : _ | _ | VA]) - (VΓext : [||-v (Γ ,, A) ,, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0)]) - (VP : [_ ||-v P | VΓext]) - (VP' : [_ ||-v P | VΓext]) - (VPP' : [_ ||-v P ≅ P' | _ | VP]) - (Vy : [Γ ||-v y : _ | _ | VA]) - (Vy' : [Γ ||-v y' : _ | _ | VA]) - (Vyy' : [Γ ||-v y ≅ y' : _ | _ | VA]) - (VId : [Γ ||-v tId A x y | VΓ]) - (Ve : [_ ||-v e : _ | _ | VId]) - (Ve' : [_ ||-v e' : _ | _ | VId]) - (Vee' : [_ ||-v e ≅ e' : _ | _ | VId]) - (VPey : [_ ||-v P[e .: y ..] | VΓ]) : - [_ ||-v P[e .: y ..] ≅ P'[e' .: y' ..] | VΓ | VPey]. - Proof. - assert (VId' : [Γ ||-v tId A x y' | VΓ]) by now eapply IdValid. - assert [Γ ||-v< l > e' : tId A x y' | VΓ | VId']. - 1:{ - eapply conv; tea. - eapply IdCongValid; tea. - 1: eapply reflValidTy. - now eapply reflValidTm. - } - eapply transValidTyEq. - - eapply substExtIdElimMotive. - 2: tea. all: tea. - Unshelve. eapply substIdElimMotive; cycle 1; tea. - - constructor; intros; irrelevance0; rewrite subst_scons2; [reflexivity|]. - unshelve eapply validTyEq. - 6: tea. 1: tea. - eapply idElimMotiveScons2Valid; tea. - Qed. - - (* Lemma liftSubstS' {Γ σ Δ lF F} - {VΓ : [||-v Γ]} - {wfΔ : [|- Δ]} - (VF : [Γ ||-v F | VΓ]) - (VΓF : [||-v Γ,, F]) - {wfΔF : [|- Δ ,, F[σ]]} - (Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ ]) : - [Δ ,, F[σ] ||-v up_term_term σ : Γ ,, F | VΓF | wfΔF ]. - Proof. - eapply irrelevanceSubst. - now unshelve eapply liftSubstS'. - Qed. *) - - - Lemma IdElimValid {Γ l A x P hr y e} - (VΓ : [||-v Γ]) - (VA : [_ ||-v A | VΓ]) - (Vx : [_ ||-v x : _ | _ | VA]) + Lemma idElimMotiveScons2Red {Γ l A x x' y y' e e'} + {VΓ : [||-v Γ]} + {VA : [_ ||-v A | VΓ]} + (Vx : [_ ||-v x ≅ x' : _ | _ | VA]) (VΓext : [||-v (Γ ,, A) ,, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0)]) - (VP : [_ ||-v P | VΓext]) - (VPhr := substIdElimMotive VΓ VA Vx VΓext VP Vx (IdValid VΓ VA Vx Vx) (reflValid VΓ VA Vx _)) - (Vhr : [_ ||-v hr : _ | _ | VPhr ]) - (Vy : [_ ||-v y : _ | _ | VA]) - (VId : [Γ ||-v tId A x y | VΓ]) - (Ve : [_ ||-v e : _ | _ | VId]) - (VPye := substIdElimMotive VΓ VA Vx VΓext VP Vy VId Ve) : - [_ ||-v tIdElim A x P hr y e : _ | _ | VPye]. + {Δ} {wfΔ : [|-Δ]} + {σ σ'} (Vσσ' : [_ ||-v σ ≅ σ' : _ | VΓ | wfΔ]) + {RVA : [Δ ||- A[σ]]} + (Ry : [ RVA | _ ||- y ≅ y' : _]) + {RId : [Δ ||- tId A[σ] x[σ] y]} + (Re : [RId | _ ||- e ≅ e' : _]) : + [VΓext | Δ ||-v (e .: (y .: σ)) ≅ (e' .: (y' .: σ')) : _ | wfΔ]. Proof. - unshelve epose (h := _ : [||-v Γ,, A]). 1: now eapply validSnoc. - constructor; intros. - - instValid Vσ. - irrelevance0. - 2: unshelve eapply idElimRed; refold; tea. - 1: refold; now rewrite subst_upup_scons2, subst_scons2. - 2,5: irrelevance. - + intros; eapply subst_idElimMotive_substupup; cycle 1; tea. - + now eapply red_idElimMotive_substupup. - + intros; eapply substEq_idElimMotive_substupup; cycle 1; tea. - eapply LRTmRedConv; tea; now eapply LRTyEqSym. - - instAllValid Vσ Vσ' Vσσ'. - irrelevance0. - 2: unshelve eapply idElimCongRed; refold. - 1: refold; now rewrite subst_upup_scons2, subst_scons2. - all: try now eapply LRCumulative. - all: tea; try irrelevanceCum. - 2,4: now eapply red_idElimMotive_substupup. - + intros; eapply subst_idElimMotive_substupup; cycle 1; tea; irrelevanceCum. - Unshelve. all: tea. irrelevanceCum. - + intros; eapply subst_idElimMotive_substupup; cycle 1; tea; irrelevanceCum. - Unshelve. all: tea. - + now eapply red_idElimMotive_substupup_cong. - + intros; eapply substEq_idElimMotive_substupupEq; cycle 2; tea; irrelevanceCum. - Unshelve. all: tea. now eapply LRCumulative. - + intros; eapply substEq_idElimMotive_substupup; cycle 1; tea. - 1,3-6: irrelevanceCum. - eapply LRTmRedConv;[|irrelevanceCum]; eapply LRTyEqSym; irrelevanceCum. - Unshelve. all: tea; now eapply LRCumulative. - + Set Printing Primitive Projection Parameters. - intros; eapply substEq_idElimMotive_substupup; cycle 1; tea. - 1,3: irrelevanceCum. - eapply LRTmRedConv; [|irrelevanceCum]; eapply LRTyEqSym; irrelevanceCum. - Unshelve. all: tea. + pose proof (invValiditySnoc VΓext) as (?&VΓA& VIdA &->). + pose proof (invValiditySnoc VΓA) as (?&?& ? &->). + do 2 (unshelve eapply consSubstEq; [|irrelevance]); irrValid. Qed. - Lemma IdElimCongValid {Γ l A A' x x' P P' hr hr' y y' e e'} - (VΓ : [||-v Γ]) - (VA : [_ ||-v A | VΓ]) - (VA' : [_ ||-v A' | VΓ]) - (VAA' : [_ ||-v A ≅ A' | _ | VA]) - (Vx : [_ ||-v x : _ | _ | VA]) - (Vx' : [_ ||-v x' : _ | _ | VA]) - (Vxx' : [_ ||-v x ≅ x' : _ | _ | VA]) + Lemma IdElimValid {Γ l A A' x x' P P' hr hr' y y' e e'} + (VΓ : [||-v Γ]) + (VA : [_ ||-v A | VΓ]) + (VAA' : [_ ||-v A ≅ A' | VΓ | VA]) + (Vx : [_ ||-v x ≅ x' : _ | _ | VA]) (VΓext : [||-v (Γ ,, A) ,, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0)]) (VP : [_ ||-v P | VΓext]) - (VP' : [_ ||-v P' | VΓext]) - (VPP' : [_ ||-v P ≅ P' | _ | VP]) - (VPhr := substIdElimMotive VΓ VA Vx VΓext VP Vx (IdValid VΓ VA Vx Vx) (reflValid VΓ VA Vx _)) - (Vhr : [_ ||-v hr : _ | _ | VPhr ]) - (Vhr' : [_ ||-v hr' : _ | _ | VPhr ]) - (Vhrhr' : [_ ||-v hr ≅ hr' : _ | _ | VPhr]) - (Vy : [_ ||-v y : _ | _ | VA]) - (Vy' : [_ ||-v y' : _ | _ | VA]) - (Vyy' : [_ ||-v y ≅ y' : _ | _ | VA]) + (VPP' : [_ ||-v P ≅ P' | VΓext | VP]) + (VIdxx := (IdValid VΓ VA Vx Vx)) + (VPhr := substIdElimMotive VΓ VA Vx VΓext VP Vx VIdxx (reflValid VΓ VA Vx _)) + (Vhr : [_ ||-v hr ≅ hr' : _ | _ | VPhr ]) + (Vy : [_ ||-v y ≅ y' : _ | _ | VA]) (VId : [Γ ||-v tId A x y | VΓ]) - (Ve : [_ ||-v e : _ | _ | VId]) - (Ve' : [_ ||-v e' : _ | _ | VId]) - (Vee' : [_ ||-v e ≅ e' : _ | _ | VId]) + (Ve : [_ ||-v e ≅ e' : _ | _ | VId]) (VPye := substIdElimMotive VΓ VA Vx VΓext VP Vy VId Ve) : [_ ||-v tIdElim A x P hr y e ≅ tIdElim A' x' P' hr' y' e' : _ | _ | VPye]. Proof. - assert [_ ||-v x' : _ | _ | VA'] by now eapply conv. - assert [_ ||-v y' : _ | _ | VA'] by now eapply conv. - assert (VΓext' : [||-v (Γ ,, A'),, tId A'⟨@wk1 Γ A'⟩ x'⟨@wk1 Γ A'⟩ (tRel 0)]). - 1: eapply validSnoc; now eapply idElimMotiveCtxIdValid. - assert (h : forall t, t⟨@wk1 Γ A'⟩ = t⟨@wk1 Γ A⟩) by reflexivity. - assert (VPalt' : [_ ||-v P' | VΓext']). - 1:{ - eapply convCtx2'; tea. - 1: eapply convCtx1; tea; [now eapply symValidTyEq| ]. - 1,3: now eapply idElimMotiveCtxIdValid. - eapply idElimMotiveCtxIdCongValid; tea. - Unshelve. 1: now eapply idElimMotiveCtxIdValid. tea. - } - constructor; intros. - instValid Vσ. - irrelevance0. - 2: unshelve eapply idElimCongRed; refold. - 1: refold; now rewrite subst_upup_scons2, subst_scons2. - all: first [now eapply LRCumulative | solve [irrelevance | irrelevanceCum] | now eapply LRTmRedConv | tea]. - + intros ; eapply subst_idElimMotive_substupup; cycle 1; tea; irrelevanceCum. - Unshelve. all: tea. irrelevanceCum. - + intros; eapply red_idElimMotive_substupup; tea. - + intros; eapply subst_idElimMotive_substupup; cycle 1; tea. - irrelevance. - Unshelve. all: tea. - + unshelve eapply IdRed; tea; eapply LRTmRedConv; tea. - + eapply red_idElimMotive_substupup; tea. - + eapply redEq_idElimMotive_substupup. - 6-8: tea. 3-5: tea. all: tea. - + intros. - assert [_ ||- y'0 : _ | RVA] by (eapply LRTmRedConv; tea; now eapply LRTyEqSym). - eapply substExt_idElimMotive_substupup. - 7: tea. 2,3,5,6: tea. all: tea; try solve [irrelevanceCum]. - 1: eapply LRTmRedConv; tea; eapply LRTyEqSym; tea. - eapply IdCongRed; tea; [now escape| now eapply reflLRTmEq]. - Unshelve. 1: now eapply LRCumulative. 1,2: now eapply IdRed. - + intros; eapply substEq_idElimMotive_substupup; cycle 1; tea; try solve [irrelevanceCum]. - eapply LRTmRedConv. 2:irrelevanceCum. eapply LRTyEqSym; irrelevanceCum. - Unshelve. all: tea; now eapply LRCumulative. - + intros; eapply substEq_idElimMotive_substupup; cycle 1; tea; try solve [irrelevanceCum]. - eapply LRTmRedConv. 2:irrelevanceCum. eapply LRTyEqSym; irrelevanceCum. - Unshelve. all: tea; now eapply LRCumulative. - + eapply LRTmRedConv; tea. - rewrite subst_upup_scons2; change (tRefl A'[σ] x'[σ]) with (tRefl A' x')[σ]. - rewrite <- subst_scons2. - eapply validTyEq. eapply substEqIdElimMotive. - 6,7: tea. 5-8: tea. 2-4: tea. 1: tea. - 2: eapply conv. - 1,3: now eapply reflValid. - 1: eapply symValidTyEq; now eapply IdCongValid. - now eapply reflCongValid. - Unshelve. all: now eapply IdValid. - + eapply LRTmRedConv; tea. - now eapply IdCongValid. + constructor; intros; cbn. + instValid Vσσ'. + pose proof (Vuu0 := liftSubstEq' (idElimMotiveCtxIdValid VΓ VA Vx) (liftSubstEq' VA Vσσ')). + set (wfΔ' := wfc_cons _ _) in Vuu0. + epose proof (Vuu := irrelevanceSubstEq _ VΓext _ wfΔ' Vuu0). + instValid Vuu. + irrelevance0; [now rewrite <-up_twice_subst|]. + unshelve (eapply idElimCongRed; tea); tea. + - intros * Ry ? Re. + epose proof (Vext := idElimMotiveScons2Red Vx VΓext Vσσ' Ry Re). + instValid Vext; now rewrite subst_upup_scons2. + - erewrite idElimMotive_Idsubst_eq; now eapply escape. + - erewrite idElimMotive_Idsubst_eq. + epose proof (VA' := ureflValidTy VAA'). + epose proof (X := liftSubstEq' VA' (urefl Vσσ')). + epose proof (X0 :=idElimMotiveCtxIdValid _ VA' (conv _ _ VA' VAA' (ureflValidTm Vx))). + pose proof (Vuu1 := liftSubstEq' X0 X). + set (wfΔ'' := wfc_cons _ _) in Vuu1. + pose proof (eqvCtx := idElimMotiveCtxEq VΓ VA VAA' Vx). + pose proof (Vuu2 := irrelevanceSubstEq _ eqvCtx.π2.π1 _ wfΔ'' Vuu1). + pose proof (VP' := convVTy eqvCtx (irrelevanceTy _ _ (ureflValidTy VPP'))). + instValid Vuu2; now eapply escape. + - erewrite idElimMotive_Idsubst_eq; now eapply escapeEq. + - intros ???? Ry ? Re. + epose proof (Vext := idElimMotiveScons2Red Vx VΓext Vσσ' Ry Re). + instValid Vext; irrelevance0; now rewrite subst_upup_scons2. + - intros ???????? Ry ? Re. + pose proof (reflValidTy VP). + epose proof (Vext := idElimMotiveScons2Red Vx VΓext (lrefl Vσσ') Ry Re). + instValid Vext; irrelevance0; now rewrite subst_upup_scons2. + - irrelevance0; tea; clear; now bsimpl. Qed. + Lemma subst_subst_twice t a b σ : + t[a .: b..][σ] = t[a[σ] .: (b[σ] .: σ)]. + Proof. now bsimpl. Qed. + + Lemma subst_refl A x σ : (tRefl A x)[σ] = tRefl A[σ] x[σ]. + Proof. reflexivity. Qed. Lemma IdElimReflValid {Γ l A x P hr y B z} - (VΓ : [||-v Γ]) - (VA : [_ ||-v A | VΓ]) - (Vx : [_ ||-v x : _ | _ | VA]) + (VΓ : [||-v Γ]) + (VA : [_ ||-v A | VΓ]) + (Vxy : [_ ||-v x ≅ y : _ | _ | VA]) (VΓext : [||-v (Γ ,, A) ,, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0)]) (VP : [_ ||-v P | VΓext]) - (VPhr := substIdElimMotive VΓ VA Vx VΓext VP Vx (IdValid VΓ VA Vx Vx) (reflValid VΓ VA Vx _)) + (VIdxx := (IdValid VΓ VA Vxy Vxy)) + (Vrflx := reflValid VΓ VA Vxy _) + (VPhr := substIdElimMotive VΓ VA Vxy VΓext VP Vxy VIdxx Vrflx) (Vhr : [_ ||-v hr : _ | _ | VPhr ]) - (Vy : [_ ||-v y : _ | _ | VA]) - (Vxy : [_ ||-v x ≅ y : _ | _ | VA]) - (VB : [_ ||-v B | VΓ]) (VAB : [_ ||-v _ ≅ B | VΓ | VA]) - (Vz : [_ ||-v z : _ | _ | VB]) - (Vxz : [_ ||-v x ≅ z : _ | _ | VA]) - (VId : [Γ ||-v tId A x y | VΓ]) + (Vxz : [_ ||-v x ≅ z : _ | _ | VA]) + (VId : [Γ ||-v tId A x y | VΓ]) (VRefl : [_ ||-v tRefl B z : _ | _ | VId]) - (VPye := substIdElimMotive VΓ VA Vx VΓext VP Vy VId VRefl) : + (VPye := substIdElimMotive VΓ VA Vxy VΓext VP (ureflValidTm Vxy) VId VRefl) : [_ ||-v tIdElim A x P hr y (tRefl B z) ≅ hr : _ | _ | VPye]. Proof. - constructor; intros. - assert [Γ ||-v< l > P[tRefl A x .: x..] ≅ P[tRefl B z .: y..] | VΓ | VPhr]. - 1:{ - eapply substExtIdElimMotive; cycle 1; tea. - 1: now eapply reflValid. - eapply reflCongValid; tea. - eapply conv; tea. - now eapply symValidTyEq. - Unshelve. now eapply IdValid. - } - eapply redwfSubstValid; cycle 1. + eapply redSubstValid. + + constructor; intros; cbn; rewrite <-up_twice_subst. + pose proof (Vuu0 := liftSubstEq' (idElimMotiveCtxIdValid VΓ VA Vxy) (liftSubstEq' VA Vσσ')). + set (wfΔ' := wfc_cons _ _) in Vuu0. + epose proof (Vuu := irrelevanceSubstEq _ VΓext _ wfΔ' Vuu0). + instValid (lrefl Vσσ') ; instValid Vuu ; escape. + eapply redtm_idElimRefl; tea. + - now erewrite idElimMotive_Idsubst_eq. + - now rewrite <-subst_refl, up_twice_subst. + eapply conv; tea. - + constructor; intros. - instValid Vσ0; escape. - constructor. - - eapply ty_conv; tea. - - rewrite subst_scons2, <- subst_upup_scons2. - eapply redtm_idElimRefl; refold; tea. - * eapply escape. now eapply red_idElimMotive_substupup. - * rewrite subst_upup_scons2; change (tRefl ?A[?σ] ?x[_]) with (tRefl A x)[σ]. - now rewrite <- subst_scons2. - * now eapply ty_conv. + constructor; intros. + epose (Vr := reflCongValid _ _ VAB Vxz VIdxx). + epose (Vext := idElimMotiveScons2Valid _ VA Vxy Vxy _ Vr VΓext _ _ Vσσ'). + instValid Vext. + irrelevance0; now rewrite subst_subst_twice. Qed. End Id. - + diff --git a/theories/Substitution/Introductions/Lambda.v b/theories/Substitution/Introductions/Lambda.v index 314d7e0..ee052c0 100644 --- a/theories/Substitution/Introductions/Lambda.v +++ b/theories/Substitution/Introductions/Lambda.v @@ -1,22 +1,36 @@ From LogRel.AutoSubst Require Import core unscoped Ast Extra. From LogRel Require Import Utils BasicAst Notations Context NormalForms Weakening GenericTyping LogicalRelation Validity. -From LogRel.LogicalRelation Require Import Escape Reflexivity Neutral Weakening Irrelevance Application Reduction Transitivity NormalRed. -From LogRel.Substitution Require Import Irrelevance Properties SingleSubst. -From LogRel.Substitution.Introductions Require Import Pi. +From LogRel.LogicalRelation Require Import Induction Escape Reflexivity Neutral Weakening Irrelevance Application Reduction Transitivity NormalRed EqRedRight InstKripke. +From LogRel.Substitution Require Import Irrelevance Properties SingleSubst Reflexivity Conversion Reduction. +From LogRel.Substitution.Introductions Require Import Pi Application Var. Set Universe Polymorphism. Set Printing Primitive Projection Parameters. +Lemma isLRFun_isWfFun `{GenericTypingProperties} + {l Γ F G t} (RΠFG : [Γ ||-< l > tProd F G]) + (Rt : isLRFun (normRedΠ0 (invLRΠ RΠFG)) t) + : isWfFun Γ F G t. +Proof. + assert (wfΓ: [|- Γ]) by (escape ; gen_typing). + destruct Rt as [?? eqA eqt|]; cbn in *. + 2: now constructor. + pose proof (RFA := instKripkeEq wfΓ eqA). + pose proof (LRTyEqRedRight _ RFA). + pose proof (instKripkeSubstTmEq wfΓ eqt). + pose proof (instKripkeSubstConvTmEq wfΓ eqA eqt). + escape. + now constructor. +Qed. + Section LambdaValid. Context `{GenericTypingProperties}. - - - - -Context {Γ F G l} {VΓ : [||-v Γ]} (VF : [Γ ||-v F | VΓ]) +Context {Γ F G l} + {VΓ : [||-v Γ]} + (VF : [Γ ||-v F | VΓ]) (VΓF := validSnoc VΓ VF) (VG : [Γ ,, F ||-v G | VΓF]) (VΠFG := PiValid VΓ VF VG). @@ -28,285 +42,204 @@ Proof. bsimpl; cbn; now rewrite rinstInst'_term_pointwise. Qed. Lemma consWkEq' {Δ Ξ} σ (ρ : Δ ≤ Ξ) a Z : Z[up_term_term σ][a .: ρ >> tRel] = Z[a .: σ⟨ρ⟩]. Proof. bsimpl; cbn; now rewrite rinstInst'_term_pointwise. Qed. -Lemma lamBetaRed {t} (Vt : [Γ ,, F ||-v t : G | VΓF | VG]) - {Δ σ} (wfΔ : [ |-[ ta ] Δ]) (Vσ : [VΓ | Δ ||-v σ : Γ | wfΔ]) - {Δ0 a} (ρ : Δ0 ≤ Δ) (wfΔ0 : [|-Δ0]) - (RFσ : [Δ0 ||- F[σ]⟨ρ⟩]) - (ha : [RFσ | _ ||- a : _]) - (RGσ : [Δ0 ||- G[up_term_term σ][a .: ρ >> tRel]]) : - [_ ||- tApp (tLambda F t)[σ]⟨ρ⟩ a : _ | RGσ] × - [_ ||- tApp (tLambda F t)[σ]⟨ρ⟩ a ≅ t[a .: σ⟨ρ⟩] : _ | RGσ]. +Lemma eq_subst_3 t a ρ σ : t[up_term_term σ][a .: ρ >> tRel] = t[up_term_term σ⟨ρ⟩][a..]. Proof. - pose proof (Vσa := consWkSubstS VF ρ wfΔ0 Vσ ha); instValid Vσa. - pose proof (VσUp := liftSubstS' VF Vσ). - instValid Vσ. instValid VσUp. escape. - irrelevance0. 1: now rewrite consWkEq'. - eapply redwfSubstTerm; tea. - constructor; tea. - do 2 rewrite <- consWkEq. - eapply redtm_beta; tea. - fold subst_term; renToWk; eapply ty_wk; tea. - eapply wfc_cons; tea; eapply wft_wk; tea. + bsimpl ; now substify. Qed. -Lemma betaValid {t a} (Vt : [Γ ,, F ||-v t : G | VΓF | VG]) - (Va : [Γ ||-v a : F | VΓ | VF]) : - [Γ ||-v tApp (tLambda F t) a ≅ t[a..] : G[a..] | VΓ | substS VG Va]. +Lemma eq_subst_4 t a σ : t[up_term_term σ][a..] = t[a .: σ]. Proof. - constructor; intros. instValid Vσ. - pose (Vσa := consSubstS _ _ Vσ _ RVa). instValid Vσa. - pose proof (VσUp := liftSubstS' VF Vσ). - instValid Vσ. instValid VσUp. escape. - assert (eq : forall t, t[a..][σ] = t[up_term_term σ][a[σ]..]) by (intros; now bsimpl). - assert (eq' : forall t, t[a..][σ] = t[a[σ].: σ]) by (intros; now bsimpl). - irrelevance0. 1: symmetry; apply eq. - rewrite eq; eapply redwfSubstTerm; tea. - 1: rewrite <- eq; rewrite eq'; irrelevance. - constructor; tea. - * do 2 (rewrite <- eq; rewrite eq'); tea. - * eapply redtm_beta; tea. - Unshelve. 2:rewrite <- eq; now rewrite eq'. + bsimpl ; now substify. Qed. -Lemma lamValid0 {t} (Vt : [Γ ,, F ||-v t : G | VΓF | VG]) - {Δ σ} (wfΔ : [ |-[ ta ] Δ]) (Vσ : [VΓ | Δ ||-v σ : Γ | wfΔ]) : - [validTy VΠFG wfΔ Vσ | Δ ||- (tLambda F t)[σ] : (tProd F G)[σ]]. + +Lemma eq_upren t σ ρ : t[up_term_term σ]⟨upRen_term_term ρ⟩ = t[up_term_term σ⟨ρ⟩]. +Proof. asimpl; unfold Ren1_subst; asimpl; substify; now asimpl. Qed. + +Lemma eq_substren t σ ρ : t[σ]⟨ρ⟩ = t[σ⟨ρ⟩]. +Proof. now asimpl. Qed. + + +Lemma lamPiRedTm {F' G'} + (VF' : [Γ ||-v F' | VΓ]) + (VFF' : [Γ ||-v F ≅ F' | VΓ | VF]) + (VG' : [Γ ,, F ||-v G' | VΓF]) + (VGG' : [Γ ,, F ||-v G ≅ G' | VΓF | VG]) + {t t'} (Vtt' : [Γ ,, F ||-v t ≅ t' : G | VΓF | VG]) + {Δ} (wfΔ : [|-Δ]) {σ σ'} (Vσσ' : [VΓ | Δ ||-v σ ≅ σ' : Γ | wfΔ]) + (R0 : [Δ ||-Π (tProd F G)[σ]]) + (R := normRedΠ0 R0) + : PiRedTm R (tLambda F' t')[σ']. Proof. - pose proof (VσUp := liftSubstS' VF Vσ). - instValid Vσ. instValid VσUp. escape. - pose (RΠ := normRedΠ RVΠFG); refold. - enough [RΠ | Δ ||- (tLambda F t)[σ] : (tProd F G)[σ]] by irrelevance. - assert [Δ |- (tLambda F t)[σ] : (tProd F G)[σ]] by (escape; cbn; gen_typing). - exists (tLambda F t)[σ]; intros; cbn in *. - + now eapply redtmwf_refl. - + constructor; cbn; intros. - - apply reflLRTyEq. - - rewrite Poly.eq_subst_2. - irrelevance0; [symmetry; apply Poly.eq_subst_2|]. - unshelve apply Vt; tea. - eapply consSubstS. - irrelevance0; [|tea]. - now bsimpl. - Unshelve. - now unshelve eapply wkSubstS. - + eapply convtm_eta; tea. - 1,2: constructor; tea; eapply escapeEq, reflLRTyEq. - assert (eqσ : forall Z, Z[up_term_term σ] = Z[up_term_term σ]⟨upRen_term_term S⟩[(tRel 0) ..]) - by (intro; bsimpl; cbn; now rewrite rinstInst'_term_pointwise). - assert [Δ,, F[σ] |-[ ta ] tApp (tLambda F[σ] t[up_term_term σ])⟨S⟩ (tRel 0) ⤳* t[up_term_term σ]⟨upRen_term_term S⟩[(tRel 0)..] : G[up_term_term σ]]. - { - rewrite (eqσ G). eapply redtm_beta. - * renToWk; eapply wft_wk; tea; gen_typing. - * renToWk; eapply ty_wk; tea. - eapply wfc_cons; [gen_typing | eapply wft_wk; gen_typing]. - * fold ren_term; refine (ty_var _ (in_here _ _)); gen_typing. - } - eapply convtm_exp; tea. - - rewrite <- eqσ; tea. - - rewrite <- eqσ; tea. - - unshelve eapply escapeEq, reflLRTyEq; [|tea]. - - rewrite <- (eqσ t); eapply escapeEqTerm; now eapply reflLRTmEq. - + eapply lamBetaRed; tea. - + pose proof (Vσa := consWkSubstS VF ρ h Vσ ha). - pose proof (Vσb := consWkSubstS VF ρ h Vσ hb). - assert (Vσab : [_ ||-v _ ≅ (b .: σ⟨ρ⟩) : _ | _ | _ | Vσa]). - 1:{ - unshelve eapply consSubstSEq'. - 2: irrelevance. - eapply irrelevanceSubstEq. - eapply wkSubstSEq ; eapply reflSubst. - Unshelve. all: tea. - } - eapply LREqTermHelper. - 1,2: eapply lamBetaRed; tea. - - irrelevance0; rewrite consWkEq';[reflexivity|]. - unshelve eapply validTyExt; cycle 3; tea. - Unshelve. rewrite consWkEq'; eapply validTy; tea. - - epose proof (validTmExt Vt _ _ Vσb Vσab). - irrelevance. now rewrite consWkEq'. -Qed. + refold. + pose proof (VG'':=convCtx1 VΓ VΓF (validSnoc VΓ VF') VF VFF' VG'). + epose proof (PiCong VΓ VF VG VF' VG'' VFF' VGG'). + epose proof (VtG' := conv VΓF VG VG' VGG' Vtt'). + epose proof (convTmEqCtx1 VΓ VΓF (validSnoc VΓ VF') VF VG' VG'' VFF' VtG'). + instValid (liftSubstEq' VF' (urefl Vσσ')); + instValid Vσσ'; instValid (urefl Vσσ'). escape. -Lemma lamValid {t} (Vt : [Γ ,, F ||-v t : G | VΓF | VG]) - : - [Γ ||-v tLambda F t : tProd F G | VΓ | VΠFG ]. + econstructor. + 1: eapply redtmwf_refl, ty_conv; [now eapply ty_lam| now symmetry]. + + constructor; intros; cbn. + 2: epose (Vaσ := consWkSubstEq VF Vσσ' ρ h (lrefl ha)). + all: irrelevanceRefl. + 1: now unshelve eapply wkEq. + + rewrite 2! consWkEq'. + pose proof (ureflValidTm Vtt'). + unshelve epose (Vaσ' := consWkSubstEq VF (urefl Vσσ') ρ h _). + 4: eapply LRTmEqConv; [|tea]; irrelevanceRefl; eapply wkEq; tea. + 1: now eapply wk. + instValid Vaσ'. + eapply LRTmEqConv;[|tea]. + rewrite consWkEq'; eapply LRTyEqSym. + now instValid Vaσ. + Unshelve. 2: rewrite consWkEq'; now instValid Vaσ. tea. +Defined. + +Lemma lamCongValid {F' G'} + (VF' : [Γ ||-v F' | VΓ]) + (VFF' : [Γ ||-v F ≅ F' | VΓ | VF]) + (VG' : [Γ ,, F ||-v G' | VΓF]) + (VGG' : [Γ ,, F ||-v G ≅ G' | VΓF | VG]) + {t t'} (Vtt' : [Γ ,, F ||-v t ≅ t' : G | VΓF | VG]) : + [Γ ||-v tLambda F t ≅ tLambda F' t' : tProd F G | VΓ | VΠFG ]. Proof. - opector. 1: now apply lamValid0. - intros. - pose (VσUp := liftSubstS' VF Vσ). - pose proof (VσUp' := liftSubstS' VF Vσ'). - pose proof (VσUprea := liftSubstSrealign' VF Vσσ' Vσ'). - pose proof (VσσUp := liftSubstSEq' VF Vσσ'). - instAllValid Vσ Vσ' Vσσ'; instValid VσUp'; instAllValid VσUp VσUprea VσσUp. - pose (RΠ := normRedΠ RVΠFG); refold. - enough [RΠ | Δ ||- (tLambda F t)[σ] ≅ (tLambda F t)[σ'] : (tProd F G)[σ]] by irrelevance. - unshelve econstructor. - - change [RΠ | Δ ||- (tLambda F t)[σ] : (tProd F G)[σ]]. - eapply normLambda. - epose (lamValid0 Vt wfΔ Vσ). - irrelevance. - - change [RΠ | Δ ||- (tLambda F t)[σ'] : (tProd F G)[σ]]. - eapply normLambda. - eapply LRTmRedConv. - 2: epose (lamValid0 Vt wfΔ Vσ'); irrelevance. - eapply LRTyEqSym. eapply (validTyExt VΠFG); tea. - Unshelve. 2: now eapply validTy. - - refold; cbn; escape. - eapply convtm_eta; refold; tea. - 2,4: constructor; first [assumption|now eapply lrefl|idtac]. + econstructor; intros. cbn -[PiValid]. + normRedΠ R; refold. + pose proof (VG'':=convCtx1 VΓ VΓF (validSnoc VΓ VF') VF VFF' VG'). + epose proof (PiCong VΓ VF VG VF' VG'' VFF' VGG'). + epose proof (VtG' := conv VΓF VG VG' VGG' Vtt'). + epose proof (convTmEqCtx1 VΓ VΓF (validSnoc VΓ VF') VF VG' VG'' VFF' VtG'). + simple refine (let ht : PiRedTm R (tLambda F t)[σ] := _ in _). + 1:{ + pose proof (VFF := reflValidTy VF). + pose proof (VGG := reflValidTy VG). + pose proof (Vtt := lreflValidTm _ Vtt'). + eapply (lamPiRedTm VF VFF VG VGG Vtt wfΔ (lrefl Vσσ') _). + } + simple refine (let ht' : PiRedTm R (tLambda F' t')[σ'] := _ in _). + 1: eapply (lamPiRedTm VF' VFF' VG' VGG' Vtt' wfΔ Vσσ' _). + assert (h : forall (Δ0 : context) (a b : term), + PiRedTmEq.appRed R (PiRedTmEq.nf ht) (PiRedTmEq.nf ht') Δ0 a b). + 1:{ + cbn; intros. + set (RF := PolyRed.shpRed _ _ _) in hab. + pose proof (Vσρ := wkSubstSEq VΓ wfΔ h ρ Vσσ'). + pose proof (Vσρ' := wkSubstSEq VΓ wfΔ h ρ (urefl Vσσ')). + instValid Vσρ'; instValid (liftSubstEq' VF Vσρ). + instValid Vσρ; instValid (liftSubstEq' VF' Vσρ'). + instValid (consWkSubstEq VF Vσσ' ρ h hab). + escape. irrelevance0. + 1: now rewrite eq_subst_3. + eapply redSubstTmEq; cycle 1. + * eapply redtm_beta; tea. + now rewrite eq_upren, eq_substren. + * eapply redtm_beta. + 1,3: rewrite eq_substren; tea. + 1: eapply ty_conv; tea; cbn. + 1: replace F[_]⟨_⟩ with F[σ⟨ρ⟩]; tea; symmetry; apply eq_substren. + now rewrite eq_upren, eq_substren. + Unshelve. 2: now rewrite eq_subst_4. + * now rewrite !eq_subst_4. + * now rewrite !eq_upren, !eq_subst_4. + } + unshelve econstructor; tea. + cbn; instValid (liftSubstEq' VF Vσσ'); instValid Vσσ'; escape. + eapply convtm_eta; tea. + * destruct ht; gen_typing. + * eapply isLRFun_isWfFun; exact (PiRedTmEq.isfun ht). + * destruct ht' as [? red ?]; cbn in red; gen_typing. + * eapply isLRFun_isWfFun; exact (PiRedTmEq.isfun ht'). + * unshelve epose proof (r := h (Δ ,, F[σ]) (tRel 0) (tRel 0) (wk1 F[σ]) _ _). + gen_typing. - + eapply ty_conv; [tea|now symmetry]. - + eapply ty_conv; [tea|]. - assert (Vσ'σ := symmetrySubstEq _ _ _ _ _ Vσ' Vσσ'). - assert (Vupσ'σ := liftSubstSEq' VF Vσ'σ). - unshelve eassert (VG' := validTyExt VG _ _ _ Vupσ'σ). - { eapply liftSubstSrealign'; tea. } - eapply escapeEq, VG'. - + eapply ty_conv; [now eapply ty_lam|]. - symmetry; now eapply convty_prod. - + assert (eqσ : forall σ Z, Z[up_term_term σ] = Z[up_term_term σ]⟨upRen_term_term S⟩[(tRel 0) ..]) - by (intros; bsimpl; cbn; now rewrite rinstInst'_term_pointwise). - eapply convtm_exp. - * rewrite (eqσ σ G). eapply redtm_beta. - -- renToWk; eapply wft_wk; tea; gen_typing. - -- renToWk; eapply ty_wk; tea. - eapply wfc_cons; [gen_typing | eapply wft_wk; gen_typing]. - -- fold ren_term; refine (ty_var _ (in_here _ _)); gen_typing. - * eapply redtm_conv. 2: now symmetry. - rewrite (eqσ σ' G). eapply redtm_beta. - -- renToWk; eapply wft_wk; tea; gen_typing. - -- renToWk; eapply ty_wk; tea. - eapply wfc_cons; [gen_typing | eapply wft_wk; gen_typing]. - -- fold ren_term. eapply ty_conv. - refine (ty_var _ (in_here _ _)). 1: gen_typing. - cbn; renToWk; eapply convty_wk; tea; gen_typing. - * tea. - * fold ren_term; rewrite <- eqσ; tea. - * fold ren_term; rewrite <- eqσ. - eapply ty_conv; [tea|symmetry; tea]. - * now eapply lrefl. - * fold ren_term. - set (x := ren_term _ _); change x with (t[up_term_term σ]⟨upRen_term_term S⟩); clear x. - set (x := ren_term _ _); change x with (t[up_term_term σ']⟨upRen_term_term S⟩); clear x. - do 2 rewrite <- eqσ ; eapply escapeEqTerm; now eapply validTmExt. - - refold; cbn; intros. - pose proof (Vσa := consWkSubstS VF ρ h Vσ ha). - assert (ha' : [ _||- a : _| wk ρ h RVF0]). - 1: { - eapply LRTmRedConv; tea. - irrelevance0; [reflexivity|]. - eapply wkEq; eapply (validTyExt VF); tea. - } - pose proof (Vσa' := consWkSubstS VF ρ h Vσ' ha'). - assert (Vσσa : [_ ||-v _ ≅ (a .: σ'⟨ρ⟩) : _ | _ | _ | Vσa]). - { - unshelve eapply consSubstSEq'. - 2: eapply reflLRTmEq; irrelevance0; [|exact ha]; now bsimpl. - eapply irrelevanceSubstEq; now eapply wkSubstSEq. - Unshelve. all: tea. - } - eapply LREqTermHelper. - 1,2: eapply lamBetaRed; tea. - + irrelevance0; rewrite consWkEq';[reflexivity|]. - unshelve eapply validTyExt; cycle 3; tea. - Unshelve. rewrite consWkEq'; eapply validTy; tea. - + epose proof (validTmExt Vt _ _ Vσa' Vσσa). - irrelevance. now rewrite consWkEq'. + + eapply var0; [now rewrite wk1_ren_on| assumption]. + + erewrite <- 2!wk1_ren_on; eapply convtm_meta_conv; [| | reflexivity]. + 1: eapply escapeEqTerm, r. + now rewrite <- eq_id_subst_scons. +Qed. + +Lemma lamValid + {t t'} (Vtt' : [Γ ,, F ||-v t ≅ t' : G | VΓF | VG]) : + [Γ ||-v tLambda F t : tProd F G | VΓ | VΠFG ]. +Proof. + eapply lamCongValid; tea; first [eapply reflValidTy | now eapply lreflValidTm]. Qed. +Lemma singleSubst_subst_eq t a σ : t[a..][σ] = t[up_term_term σ][a[σ]..]. +Proof. now bsimpl. Qed. -Lemma ηeqEqTermConvNf {σ Δ f} (ρ := @wk1 Γ F) +Lemma singleSubst_subst_eq' t a σ : t[a..][σ] = t[a[σ] .: σ]. +Proof. now bsimpl. Qed. + +Lemma betaValid {t a} (Vt : [Γ ,, F ||-v t : G | VΓF | VG]) + (Va : [Γ ||-v a : F | VΓ | VF]) : + [Γ ||-v tApp (tLambda F t) a ≅ t[a..] : G[a..] | VΓ | substS VG Va]. +Proof. + eapply redSubstValid. + 2: eapply substSTm ; irrValid. + constructor; intros; cbn. + rewrite 2!singleSubst_subst_eq. + instValid (lrefl Vσσ'). + instValid (liftSubstEq' VF (lrefl Vσσ')). + escape; now eapply redtm_beta. +Qed. + + +Lemma redtm_app_helper {Δ Δ' f nf a σ} (ρ : Δ' ≤ Δ) : + [|- Δ'] -> + [Δ |- f[σ] ⤳* nf : (tProd F G)[σ]] -> + [Δ' |- a : F[σ]⟨ρ⟩] -> + [Δ' |- tApp f[σ]⟨ρ⟩ a ⤳* tApp nf⟨ρ⟩ a : G[up_term_term σ][a .: ρ >> tRel]]. +Proof. + intros red tya. + rewrite eq_subst_3; eapply redtm_app; tea. + eapply redtm_meta_conv; [now eapply redtm_wk| | reflexivity]. + cbn; now rewrite eq_upren. +Qed. + + +Lemma ηeqEqTermNf {σ Δ f} (ρ := @wk1 Γ F) (wfΔ : [|- Δ]) (Vσ : [Δ ||-v σ : Γ | VΓ| wfΔ]) - (RΠFG := normRedΠ (validTy VΠFG wfΔ Vσ)) - (Rf : [Δ ||- f[σ] : (tProd F G)[σ] | RΠFG ]) : - [Δ ,, F[σ] |- tApp f⟨ρ⟩[up_term_term σ] (tRel 0) ≅ tApp (PiRedTm.nf Rf)⟨↑⟩ (tRel 0) : G[up_term_term σ]]. + (RΠFG := normRedΠ0 (invLRΠ (validTy VΠFG wfΔ Vσ))) + (RGσ : [Δ ,, F[σ] ||- G[up_term_term σ]]) + (Rf : [Δ ||- f[σ] : (tProd F G)[σ] | LRPi' RΠFG ]) : + [RGσ | Δ ,, F[σ] ||- tApp f⟨ρ⟩[up_term_term σ] (tRel 0) ≅ tApp Rf.(PiRedTmEq.redR).(PiRedTmEq.nf)⟨↑⟩ (tRel 0) : G[up_term_term σ]]. Proof. refold. - pose (VσUp := liftSubstS' VF Vσ). + pose (VσUp := liftSubstEq' VF Vσ). instValid Vσ; instValid VσUp; escape. - destruct (PiRedTm.red Rf); cbn in *. assert (wfΔF : [|- Δ,, F[σ]]) by gen_typing. - unshelve epose proof (r := PiRedTm.app Rf (@wk1 Δ F[σ]) wfΔF (var0 _ _ _)). + unshelve epose proof (r := PiRedTmEq.eqApp Rf (@wk1 Δ F[σ]) wfΔF (var0 _ _ _)); tea. 1: now rewrite wk1_ren_on. - assumption. - assert (eqσ : forall σ Z, Z[up_term_term σ] = Z[up_term_term σ][(tRel 0) .: @wk1 Δ F[σ] >> tRel]) - by (intros; bsimpl; cbn; now rewrite rinstInst'_term_pointwise). - eapply convtm_exp. - 7: rewrite eqσ; eapply escapeEqTerm; eapply reflLRTmEq; irrelevance. - * eapply redtm_meta_conv. 3: reflexivity. - 1: eapply redtm_app. - 2: eapply (ty_var wfΔF (in_here _ _)). - 1:{ - replace (f⟨_⟩[_]) with f[σ]⟨@wk1 Δ F[σ]⟩ by (unfold ρ; now bsimpl). - eapply redtm_meta_conv. 3: reflexivity. - 1: eapply redtm_wk; tea. - now rewrite wk1_ren_on. - } - fold ren_term. rewrite eqσ; now bsimpl. - * rewrite <- (wk1_ren_on Δ F[σ]); unshelve eapply redtmwf_refl. - rewrite eqσ; eapply escapeTerm ; irrelevance. - Unshelve. 2,4: rewrite <- eqσ; tea. - * tea. - * rewrite eqσ; eapply escapeTerm ; irrelevance. - Unshelve. 2: rewrite <- eqσ; tea. - * rewrite eqσ; eapply escapeTerm ; irrelevance. - Unshelve. 2: rewrite <- eqσ; tea. - * unshelve eapply escapeEq, reflLRTyEq; [|tea]. + irrelevance0; [ erewrite (eq_id_subst_scons (G[_])); reflexivity|]. + eapply redSubstLeftTmEq. + + erewrite <- wk1_ren_on; irrelevance. + + clear r; destruct Rf.(PiRedTmEq.redL) as [? [? red] ?]; cbn -[wk1] in *. + replace f⟨ρ⟩[up_term_term σ] with f[σ]⟨@wk1 Δ F[σ]⟩. + eapply redtm_app_helper; tea. + 1: rewrite wk1_ren_on; now eapply ty_var0. + unfold ρ; now bsimpl. + Unshelve. 2: now rewrite <- eq_id_subst_scons. Qed. -Lemma isLRFun_isWfFun : forall {σ Δ f} (wfΔ : [|- Δ]) (vσ : [Δ ||-v σ : Γ | VΓ | wfΔ]) (RΠFG : [Δ ||-< l > (tProd F G)[σ]]) - (p : [Δ ||-Π f[σ] : tProd F[σ] G[up_term_term σ] | normRedΠ0 (Induction.invLRΠ RΠFG)]), - isWfFun Δ F[σ] G[up_term_term σ] (PiRedTm.nf p). + +Lemma ηeqEqTermConvNf {σ Δ f} (ρ := @wk1 Γ F) + (wfΔ : [|- Δ]) (Vσ : [Δ ||-v σ : Γ | VΓ| wfΔ]) + (RΠFG := normRedΠ0 (invLRΠ (validTy VΠFG wfΔ Vσ))) + (Rf : [Δ ||- f[σ] : (tProd F G)[σ] | LRPi' RΠFG ]) : + (* (Rf0 : PiRedTm RΠFG (f[σ])) : *) + [Δ ,, F[σ] |- tApp f⟨ρ⟩[up_term_term σ] (tRel 0) ≅ tApp Rf.(PiRedTmEq.redR).(PiRedTmEq.nf)⟨↑⟩ (tRel 0) : G[up_term_term σ]]. Proof. -intros. -instValid vσ. -destruct RΠFG; simpl in *. -destruct p as [nf ? isfun]; simpl in *. -destruct isfun as [A t vA vt|]; simpl in *; constructor; tea. -+ rewrite <- (@wk_id_ren_on Δ). - unshelve eapply escapeConv, vA; tea. -+ rewrite <- (wk_id_ren_on Δ F[σ]), <- (wk_id_ren_on Δ A). - now unshelve eapply escapeEq, vA. -+ assert (Hrw : forall t, t[tRel 0 .: @wk1 Δ A >> tRel] = t). - { intros; bsimpl; apply idSubst_term; intros []; reflexivity. } - assert [Δ |- F[σ]] by now eapply escape. - assert (wfΔF : [|- Δ,, F[σ]]) by now apply wfc_cons. - unshelve eassert (vt0 := vt _ (tRel 0) (@wk1 Δ F[σ]) _ _); tea. - { eapply var0; [now bsimpl|tea]. } - eapply escapeTerm in vt0. - now rewrite !Hrw in vt0. -+ assert (Hrw : forall t, t[tRel 0 .: @wk1 Δ A >> tRel] = t). - { intros; bsimpl; apply idSubst_term; intros []; reflexivity. } - assert [Δ |- A]. - { rewrite <- (@wk_id_ren_on Δ). - unshelve eapply escapeConv, vA; tea. } - assert (wfΔA : [|- Δ,, A]) by now apply wfc_cons. - assert [Δ |-[ ta ] A ≅ F[σ]]. - { rewrite <- (wk_id_ren_on Δ F[σ]), <- (wk_id_ren_on Δ A). - symmetry; now unshelve eapply escapeEq, vA. } - assert [Δ,, A ||-< l > G[up_term_term σ]]. - { eapply validTy with (wfΔ := wfΔA); tea. - unshelve econstructor; [shelve|]; cbn. - apply var0conv; [|tea]. - replace F[↑ >> up_term_term σ] with F[σ]⟨@wk1 Δ A⟩ by now bsimpl. - rewrite <- (wk1_ren_on Δ A); eapply convty_wk; tea. - Unshelve. - eapply irrelevanceSubstExt with (σ := σ⟨@wk1 Δ A⟩). - { now bsimpl. } - now eapply wkSubstS. - } - rewrite <- (Hrw t); eapply escapeTerm. - irrelevance0; [apply Hrw|unshelve apply vt]. - - apply wfc_cons; tea. - - apply var0conv; tea. - rewrite <- (wk1_ren_on Δ A A). - apply convty_wk; [now apply wfc_cons|]. - rewrite <- (wk_id_ren_on Δ F[σ]), <- (wk_id_ren_on Δ A). - symmetry; now unshelve eapply escapeEq, vA. - Unshelve. - all: tea. + refold. + pose (VσUp := liftSubstEq' VF Vσ); instValid VσUp. + now unshelve eapply escapeEqTerm, ηeqEqTermNf. Qed. + Lemma ηeqEqTerm {σ Δ f g} (ρ := @wk1 Γ F) (Vfg : [Γ ,, F ||-v tApp f⟨ρ⟩ (tRel 0) ≅ tApp g⟨ρ⟩ (tRel 0) : G | VΓF | VG ]) (wfΔ : [|- Δ]) (Vσ : [Δ ||-v σ : Γ | VΓ| wfΔ]) @@ -315,50 +248,100 @@ Lemma ηeqEqTerm {σ Δ f g} (ρ := @wk1 Γ F) (Rg : [Δ ||- g[σ] : (tProd F G)[σ] | RΠFG ]) : [Δ ||- f[σ] ≅ g[σ] : (tProd F G)[σ] | RΠFG ]. Proof. - set (RΠ := normRedΠ RΠFG); refold. - enough [Δ ||- f[σ] ≅ g[σ] : (tProd F G)[σ] | RΠ ] by irrelevance. - opector. - - change [Δ ||- f[σ] : (tProd F G)[σ] | RΠ ]; irrelevance. - - change [Δ ||- g[σ] : (tProd F G)[σ] | RΠ ]; irrelevance. - - cbn; pose (VσUp := liftSubstS' VF Vσ). + normRedΠin Rf; normRedΠin Rg; set (RΠ := normRedΠ0 _) in Rf, Rg; refold. + enough [Δ ||- f[σ] ≅ g[σ] : (tProd F G)[σ] | LRPi' RΠ ] by irrelevance. + pose (Rf0 := Rf.(PiRedTmEq.redR)); pose (Rg0 := Rg.(PiRedTmEq.redR)). + exists Rf0 Rg0. + - cbn; pose (VσUp := liftSubstEq' VF Vσ). instValid Vσ; instValid VσUp; escape. eapply convtm_eta; tea. - + destruct (PiRedTm.red p0); cbn in *; tea. - + now eapply isLRFun_isWfFun. - + destruct (PiRedTm.red p); cbn in *; tea. - + now eapply isLRFun_isWfFun. + 2,4: eapply isLRFun_isWfFun, PiRedTmEq.isfun. + + destruct Rf0; cbn in *; gen_typing. + + destruct Rg0; cbn in *; gen_typing. + etransitivity ; [symmetry| etransitivity]; tea; eapply ηeqEqTermConvNf. - - match goal with H : [_ ||-Π f[σ] : _ | _] |- _ => rename H into Rfσ end. - match goal with H : [_ ||-Π g[σ] : _ | _] |- _ => rename H into Rgσ end. - cbn; intros ?? ρ' h ha. - set (ν := (a .: σ⟨ρ'⟩)). - pose (Vν := consWkSubstS VF ρ' h Vσ ha). - instValid Vσ; instValid Vν; escape. - assert (wfΔF : [|- Δ,, F[σ]]) by gen_typing. - cbn in *. - assert (eq : forall t, t⟨ρ⟩[a .: σ⟨ρ'⟩] = t[σ]⟨ρ'⟩) by (intros; unfold ρ; now bsimpl). - do 2 rewrite eq in RVfg. - eapply transEqTerm; [|eapply transEqTerm]. - 2: irrelevance; symmetry; apply consWkEq'. - + eapply LRTmEqSym; eapply redwfSubstTerm. - 1: unshelve epose proof (r := PiRedTm.app Rfσ ρ' h ha); irrelevance. - eapply redtmwf_appwk; tea. - 1: exact (PiRedTm.red Rfσ). - now bsimpl. - + eapply redwfSubstTerm. - 1: unshelve epose proof (r := PiRedTm.app Rgσ ρ' h ha); irrelevance. - eapply redtmwf_appwk; tea. - 1: exact (PiRedTm.red Rgσ). - now bsimpl. + - cbn; intros ??? ρ' h ha. + epose (Vν := consWkSubstEq VF Vσ ρ' h (lrefl ha)); instValid Vν. + assert (eq : forall t a, t⟨ρ⟩[a .: σ⟨ρ'⟩] = t[σ]⟨ρ'⟩) by (intros; unfold ρ; now bsimpl). + cbn in RVfg; rewrite 2!eq in RVfg. + etransitivity; [| etransitivity]. + + symmetry; eapply redSubstLeftTmEq. + 1: pose proof (urefl (PiRedTmEq.eqApp Rf ρ' h (lrefl ha))); irrelevance. + eapply redtm_app_helper; tea; [| now escape]. + now destruct (PiRedTmEq.redR Rf) as [? []]. + + irrelevance0; tea; now rewrite Poly.eq_subst_2. + + eapply redSubstLeftTmEq. + 1: pose proof (rg := PiRedTmEq.eqApp Rg ρ' h ha); irrelevance. + eapply redtm_app_helper; tea; [| now escape]. + now destruct (PiRedTmEq.redL Rg) as [? []]. Qed. Lemma etaeqValid {f g} (ρ := @wk1 Γ F) (Vf : [Γ ||-v f : tProd F G | VΓ | VΠFG ]) - (Vg : [Γ ||-v g : tProd F G | VΓ | VΠFG ]) + (Vg : [Γ ||-v g : tProd F G | VΓ | VΠFG ]) (Vfg : [Γ ,, F ||-v tApp f⟨ρ⟩ (tRel 0) ≅ tApp g⟨ρ⟩ (tRel 0) : G | VΓF | VG ]) : [Γ ||-v f ≅ g : tProd F G | VΓ | VΠFG]. Proof. - constructor; intros ??? Vσ; instValid Vσ; now eapply ηeqEqTerm. + constructor; intros ???? Vσ; instValid Vσ; instValid (lrefl Vσ). + etransitivity. + + irrelevanceRefl;eapply ηeqEqTerm; tea. + + irrelevance. +Qed. + +Lemma upren_subst_rel0 t : t[(tRel 0)]⇑ = t. +Proof. bsimpl; rewrite scons_eta'; now bsimpl. Qed. + +Lemma etaExpandValid {f} (ρ := @wk1 Γ F) + (Vf : [Γ ||-v f : tProd F G | VΓ | VΠFG ]) : + [Γ ,, F ||-v eta_expand f : G | VΓF | VG]. +Proof. + unshelve epose (VF' := wkValidTy ρ _ _ VF); tea. + unshelve epose (wkValidTy ρ _ _ VΠFG); tea. + unshelve epose (wkValidTm ρ _ _ _ Vf); tea. + eapply irrelevanceTm'. + 2: eapply appValid; erewrite <-wk1_ren_on; irrValid. + refold; unfold ρ; bsimpl; apply upren_subst_rel0. + Unshelve. all: refold; tea. + pose (var0Valid _ VF); irrValid. Qed. End LambdaValid. + +Section EtaValid. +Context `{GenericTypingProperties}. + +Context {Γ F G l} + {VΓ : [||-v Γ]} + (VF : [Γ ||-v F | VΓ]) + (VΓF := validSnoc VΓ VF) + (VG : [Γ ,, F ||-v G | VΓF]) + (VΠFG := PiValid VΓ VF VG). + + +Lemma subst_rel0 t : t⟨↑⟩[(tRel 0)..] = t. +Proof. now bsimpl. Qed. + +Lemma etaValid {f} (ρ := @wk1 Γ F) + (Vf : [Γ ||-v f : tProd F G | VΓ | VΠFG ]) : + [Γ ||-v (tLambda F (eta_expand f)) ≅ f : tProd F G | VΓ | VΠFG]. +Proof. + assert [||-v Γ,, F] by now eapply validSnoc. + unshelve epose (wkValidTm ρ _ _ _ Vf); tea. + unshelve epose (VF' := wkValidTy ρ _ _ VF); tea. + unshelve epose (VG' := wkValidTy (wk_up F ρ) _ _ VG). + 1:now eapply validSnoc. + unshelve epose (wkValidTy ρ _ _ VΠFG); tea. + eapply etaeqValid; tea. + 1: now eapply lamValid, etaExpandValid. + unshelve epose (x := betaValid VF' VG' (t:=(eta_expand f⟨ρ⟩)) (a:=(tRel 0)) _ _). + 3: eapply irrelevanceTmEq'; cycle 1. + + eapply etaExpandValid; irrValid. + + pose (var0Valid _ VF); irrValid. + + cbn -[wk1] in x |- *. + rewrite subst_rel0 in x. + replace f⟨↑⟩⟨upRen_term_term (wk1 F)⟩ with f⟨ρ⟩⟨↑⟩. + 2: clear; unfold ρ; now bsimpl. + apply x. + + rewrite wk_up_wk1_ren_on; bsimpl; apply upren_subst_rel0. +Qed. + +End EtaValid. diff --git a/theories/Substitution/Introductions/Nat.v b/theories/Substitution/Introductions/Nat.v index 9b78283..7c11676 100644 --- a/theories/Substitution/Introductions/Nat.v +++ b/theories/Substitution/Introductions/Nat.v @@ -1,8 +1,8 @@ From LogRel.AutoSubst Require Import core unscoped Ast Extra. From LogRel Require Import Utils BasicAst Notations Context NormalForms Weakening GenericTyping LogicalRelation Validity. -From LogRel.LogicalRelation Require Import Induction Irrelevance Escape Reflexivity Weakening Neutral Transitivity Reduction Application Universe SimpleArr. -From LogRel.Substitution Require Import Irrelevance Properties Conversion SingleSubst Reflexivity. -From LogRel.Substitution.Introductions Require Import Universe Pi SimpleArr Var. +From LogRel.LogicalRelation Require Import Induction Escape Irrelevance Reflexivity Irrelevance Weakening Neutral Transitivity Reduction Application Universe EqRedRight SimpleArr. +From LogRel.Substitution Require Import Irrelevance Properties Conversion SingleSubst Reflexivity Reduction. +From LogRel.Substitution.Introductions Require Import Universe Pi SimpleArr Var Application. Set Universe Polymorphism. @@ -13,7 +13,7 @@ Set Printing Primitive Projection Parameters. Lemma natRed {Γ l} : [|- Γ] -> [Γ ||- tNat]. -Proof. +Proof. intros; eapply LRNat_; constructor; eapply redtywf_refl; gen_typing. Defined. @@ -24,7 +24,7 @@ Proof. + cbn; constructor; eapply redtywf_refl; gen_typing. Defined. -Lemma natconvValid {Γ l} (VΓ : [||-v Γ]) (VNat : [Γ ||-v tNat | VΓ]) : +Lemma natconvValid {Γ l} (VΓ : [||-v Γ]) (VNat : [Γ ||-v tNat | VΓ]) : [Γ ||-v tNat ≅ tNat | VΓ | VNat]. Proof. constructor; intros. @@ -32,198 +32,95 @@ Proof. constructor; cbn; eapply redtywf_refl; gen_typing. Qed. -Lemma redUOne {Γ} : [|- Γ] -> [Γ ||-U U]. +Lemma natURedTm {Δ} (wfΔ : [|-Δ]) : URedTm (LogRelRec one) Δ tNat U (redUOneCtx wfΔ). Proof. - intros ; econstructor; [easy| gen_typing|eapply redtywf_refl; gen_typing]. + exists tNat; [| constructor]. + eapply redtmwf_refl; gen_typing. Defined. -Lemma natTermRed {Δ} (wfΔ : [|-Δ]) : [Δ ||- tNat : U | LRU_ (redUOne wfΔ)]. +Lemma natTermRed {Δ} (wfΔ : [|-Δ]) : [Δ ||- tNat : U | LRU_ (redUOneCtx wfΔ)]. Proof. - econstructor. - + eapply redtmwf_refl; gen_typing. - + constructor. - + gen_typing. - + now eapply (natRed (l:= zero)). + unshelve eexists (natURedTm wfΔ) (natURedTm wfΔ) _; cbn. + 1,3: now eapply (natRed (l:=zero)). + 1: gen_typing. + apply reflLRTyEq. Defined. Lemma natTermValid {Γ} (VΓ : [||-v Γ]): [Γ ||-v tNat : U | VΓ | UValid VΓ]. Proof. - constructor; intros. - - eapply natTermRed. - - eexists (natTermRed wfΔ) (natTermRed wfΔ) (natRed wfΔ); cbn. - + gen_typing. - + now eapply (natRed (l:=zero)). - + constructor; eapply redtywf_refl; gen_typing. + constructor; intros; eapply natTermRed. Qed. Lemma zeroRed {Γ l} {NN : [Γ ||-Nat tNat]} (wfΓ : [|- Γ]): [Γ ||- tZero : _ | LRNat_ l NN]. -Proof. - exists tZero. - 1,2: gen_typing. - constructor. -Defined. - -Lemma zeroRedEq {Γ l} {NN : [Γ ||-Nat tNat]} (wfΓ : [|- Γ]): [Γ ||- tZero ≅ tZero : _ | LRNat_ l NN]. Proof. exists tZero tZero. 1-3: gen_typing. constructor. Defined. -Lemma zeroValid {Γ l} (VΓ : [||-v Γ]): +Lemma zeroRedEq {Γ l} {NN : [Γ ||-Nat tNat]} (wfΓ : [|- Γ]): [Γ ||- tZero ≅ tZero : _ | LRNat_ l NN]. +Proof. now apply zeroRed. Defined. + +Lemma zeroValid {Γ l} (VΓ : [||-v Γ]): [Γ ||-v tZero : tNat | VΓ | natValid VΓ]. Proof. - constructor; intros; cbn; [unshelve eapply zeroRed| unshelve eapply zeroRedEq]; tea. + constructor; intros; cbn; unshelve eapply zeroRed; tea. Qed. -Lemma succRed {Γ l n} {NN : [Γ ||-Nat tNat]} : - [Γ ||- n : _ | LRNat_ l NN] -> - [Γ ||- tSucc n : _ | LRNat_ l NN]. -Proof. - intros Rn; exists (tSucc n). - + eapply redtmwf_refl; eapply ty_succ; now escape. - + eapply convtm_succ; eapply escapeEqTerm; now eapply reflLRTmEq. - + now constructor. -Defined. - Lemma succRedEq {Γ l n n'} {NN : [Γ ||-Nat tNat]} : - [Γ ||- n : _ | LRNat_ l NN] -> - [Γ ||- n' : _ | LRNat_ l NN] -> [Γ ||- n ≅ n' : _ | LRNat_ l NN] -> [Γ ||- tSucc n ≅ tSucc n' : _ | LRNat_ l NN]. Proof. - intros ???; escape; exists (tSucc n) (tSucc n'). - 1,2: eapply redtmwf_refl; now eapply ty_succ. - + now eapply convtm_succ. + intros h; pose proof (h' :=h); inversion h'; subst. + escape; exists (tSucc n) (tSucc n'). + 1,2: eapply redtmwf_refl; eapply ty_succ; gen_typing. + + eapply convtm_succ; eapply convtm_exp; gen_typing. + now constructor. Defined. -Lemma succRedInv {Γ l n} {NN : [Γ ||-Nat tNat]} : - [Γ ||- tSucc n : _ | LRNat_ l NN] -> - [Γ ||- n : _ | LRNat_ l NN]. +Lemma succRed {Γ l n} {NN : [Γ ||-Nat tNat]} : + [Γ ||- n : _ | LRNat_ l NN] -> + [Γ ||- tSucc n : _ | LRNat_ l NN]. +Proof. eapply succRedEq. Defined. + + +Lemma succRedInv {Γ l n m } {NN : [Γ ||-Nat tNat]} : + [Γ ||- tSucc n ≅ tSucc m : _ | LRNat_ l NN] -> + [Γ ||- n ≅ m : _ | LRNat_ l NN]. Proof. - intros RSn; inversion RSn; subst. - unshelve epose proof (redtmwf_whnf red _). 1: constructor. + intros RSn; inversion RSn; subst. + unshelve epose proof (redtmwf_whnf redL _); try constructor. + unshelve epose proof (redtmwf_whnf redR _); try constructor. subst. inversion prop; subst; tea. - match goal with H : [ _ ||-NeNf _ : _] |- _ => destruct H; apply convneu_whne in refl; inv_whne end. + match goal with H : [ _ ||-NeNf _ ≅ _ : _] |- _ => destruct H as [?? refl] end. + apply convneu_whne in refl; inv_whne. Qed. -Lemma succValid {Γ l n} (VΓ : [||-v Γ]) - (Vn : [Γ ||-v n : tNat | VΓ | natValid VΓ]) : - [Γ ||-v tSucc n : tNat | VΓ | natValid VΓ]. -Proof. - constructor; intros; cbn. - - instValid Vσ ; now unshelve eapply succRed. - - instAllValid Vσ Vσ' Vσσ'; now unshelve eapply succRedEq. -Qed. -Lemma succcongValid {Γ l n n'} (VΓ : [||-v Γ]) - (Vn : [Γ ||-v n : tNat | VΓ | natValid VΓ]) - (Vn' : [Γ ||-v n' : tNat | VΓ | natValid VΓ]) +Lemma succcongValid {Γ l n n'} (VΓ : [||-v Γ]) (Veqn : [Γ ||-v n ≅ n' : tNat | VΓ | natValid VΓ]) : [Γ ||-v tSucc n ≅ tSucc n' : tNat | VΓ | natValid VΓ]. Proof. - constructor; intros; cbn; instValid Vσ; now unshelve eapply succRedEq. + constructor; intros; cbn; instValid Vσσ'; now unshelve eapply succRedEq. Qed. +Lemma succValid {Γ l n} (VΓ : [||-v Γ]) + (Vn : [Γ ||-v n : tNat | VΓ | natValid VΓ]) : + [Γ ||-v tSucc n : tNat | VΓ | natValid VΓ]. +Proof. now eapply lreflValidTm, succcongValid, reflValidTm. Qed. Lemma elimSuccHypTy_subst {P} σ : elimSuccHypTy P[up_term_term σ] = (elimSuccHypTy P)[σ]. Proof. unfold elimSuccHypTy. cbn. rewrite shift_up_eq. - rewrite liftSubstComm. cbn. + rewrite liftSubstComm. cbn. now rewrite up_liftSubst_eq. Qed. Lemma liftSubst_singleSubst_eq {t u v: term} : t[u]⇑[v..] = t[u[v..]..]. Proof. now bsimpl. Qed. -Definition natElim {Γ A} (P hz hs : term) {n} {NA : [Γ ||-Nat A]} (p : NatProp NA n) : term. -Proof. - destruct p. - - exact hz. - - exact (tApp (tApp hs n) (tNatElim P hz hs n)). - - exact (tNatElim P hz hs ne). -Defined. - -Section NatElimRed. - Context {Γ l P hs hz} - (wfΓ : [|- Γ]) - (NN : [Γ ||-Nat tNat]) - (RN := LRNat_ _ NN) - (RP : [Γ,, tNat ||- P]) - (RPpt : forall n, [Γ ||- n : _ | RN] -> [Γ ||- P[n..]]) - (RPext : forall n n' (Rn : [Γ ||- n : _ | RN]), - [Γ ||- n' : _ | RN] -> - [Γ ||- n ≅ n' : _ | RN] -> - [Γ ||- P[n..] ≅ P[n'..] | RPpt _ Rn]) - (RPz := RPpt _ (zeroRed wfΓ)) - (Rhz : [Γ ||- hz : P[tZero..] | RPz]) - (RPs : [Γ ||- elimSuccHypTy P]) - (Rhs : [Γ ||- hs : _ | RPs]). - - Definition natRedElimStmt := - (forall n (Rn : [Γ ||- n : _ | RN]), - [Γ ||- tNatElim P hz hs n : _ | RPpt _ Rn ] × - [Γ ||- tNatElim P hz hs n ≅ tNatElim P hz hs (NatRedTm.nf Rn) : _ | RPpt _ Rn]) × - (forall n (Nn : NatProp NN n) (Rn : [Γ ||- n : _ | RN]), - [Γ ||- tNatElim P hz hs n : _ | RPpt _ Rn ] × - [Γ ||- tNatElim P hz hs n ≅ natElim P hz hs Nn : _ | RPpt _ Rn]). - - Lemma natElimRedAux : natRedElimStmt. - Proof. - escape. - apply NatRedInduction. - - intros ? nf red ? nfprop ih. - assert [Γ ||- nf : _ | RN]. 1:{ - econstructor; tea. - eapply redtmwf_refl; gen_typing. - } - eapply redSubstTerm. - + eapply LRTmRedConv. - 2: unshelve eapply ih; tea. - eapply RPext. - 2: eapply LRTmEqSym. - 1,2: eapply redwfSubstTerm; tea. - + eapply redtm_natelim; tea. - cbn; gen_typing. - - intros. - eapply redSubstTerm. - 2: eapply redtm_natElimZero; tea. - irrelevance. - - intros n Rn ih RSucc; change [Γ ||- n : tNat | RN] in Rn. - eapply redSubstTerm. - 2: eapply redtm_natElimSucc; tea. - 2: now escape. - eapply simple_appTerm. - 2: eapply ih. - irrelevance0. - 2: unshelve eapply (appTerm RPs Rhs Rn). - now bsimpl. - - intros ? [] ?. - apply reflect. - + apply completeness. - + now eapply ty_natElim. - + now eapply ty_natElim. - + eapply convneu_natElim; tea. - { eapply escapeEq, reflLRTyEq. } - { eapply escapeEqTerm; now eapply reflLRTmEq. } - { eapply escapeEqTerm; now eapply reflLRTmEq. } - Unshelve. - * eapply ArrRedTy; now eapply RPpt. - * rewrite subst_arr. eapply ArrRedTy. - 2: rewrite liftSubst_singleSubst_eq; cbn. - all: now eapply RPpt. - * exact l. - * assumption. - Qed. - - Lemma natElimRed : forall n (Rn : [Γ ||- n : _ | RN]), [Γ ||- tNatElim P hz hs n : _ | RPpt _ Rn ]. - Proof. intros. apply (fst natElimRedAux). Qed. -End NatElimRed. - - Section NatElimRedEq. Context {Γ l P Q hs hs' hz hz'} (wfΓ : [|- Γ]) @@ -232,128 +129,94 @@ Section NatElimRedEq. (RP : [Γ,, tNat ||- P]) (RQ : [Γ,, tNat ||- Q]) (eqPQ : [Γ,, tNat |- P ≅ Q]) - (RPpt : forall n, [Γ ||- n : _ | RN] -> [Γ ||- P[n..]]) - (RQpt : forall n, [Γ ||- n : _ | RN] -> [Γ ||- Q[n..]]) - (RPQext : forall n n' (Rn : [Γ ||- n : _ | RN]), - [Γ ||- n' : _ | RN] -> - [Γ ||- n ≅ n' : _ | RN] -> - [Γ ||- P[n..] ≅ Q[n'..] | RPpt _ Rn]) - (RPz := RPpt _ (zeroRed wfΓ)) - (RQz := RQpt _ (zeroRed wfΓ)) - (Rhz : [Γ ||- hz : P[tZero..] | RPz]) - (RQhz : [Γ ||- hz' : Q[tZero..] | RQz]) - (RPQhz : [Γ ||- hz ≅ hz' : _ | RPz]) + (RPpt : forall n n', [Γ ||- n ≅ n' : _ | RN] -> [Γ ||- P[n..]]) + (* (RQpt : forall n n', [Γ ||- n ≅ n': _ | RN] -> [Γ ||- Q[n..]]) *) + (RPQext : forall n n' (Rn : [Γ ||- n ≅ n' : _ | RN]), + [Γ ||- P[n..] ≅ Q[n'..] | RPpt _ _ Rn]) + (RPz := RPpt _ _ (zeroRed wfΓ)) + (RPQz := RPQext _ _ (zeroRedEq wfΓ)) + (Rhz : [Γ ||- hz ≅ hz' : _ | RPz]) (RPs : [Γ ||- elimSuccHypTy P]) (RQs : [Γ ||- elimSuccHypTy Q]) - (Rhs : [Γ ||- hs : _ | RPs]) - (RQhs : [Γ ||- hs' : _ | RQs]) - (RPQhs : [Γ ||- hs ≅ hs' : _ | RPs ]) + (RPQs : [Γ ||- elimSuccHypTy P ≅ elimSuccHypTy Q | RPs]) + (Rhs : [Γ ||- hs ≅ hs' : _ | RPs ]) . - Lemma RPext : forall n n' (Rn : [Γ ||- n : _ | RN]), - [Γ ||- n' : _ | RN] -> - [Γ ||- n ≅ n' : _ | RN] -> - [Γ ||- P[n..] ≅ P[n'..] | RPpt _ Rn]. - Proof. - intros. eapply transEq; [| eapply LRTyEqSym ]; eapply RPQext; cycle 1; tea. - now eapply reflLRTmEq. - Unshelve. 2,3: eauto. - Qed. + #[local] + Lemma RQpt : forall n n', [Γ ||- n ≅ n': _ | RN] -> [Γ ||- Q[n..]]. + Proof. intros; now unshelve eapply LRTyEqRedRight, RPQext. Qed. - Lemma natElimRedAuxLeft : @natRedElimStmt _ _ P hs hz NN RPpt. - Proof. - eapply natElimRedAux; tea. - + eapply RPext. - Qed. + Let RQz := RQpt _ _ (zeroRed wfΓ). - Lemma natElimRedAuxRight : @natRedElimStmt _ _ Q hs' hz' NN RQpt. + #[local] + Lemma RPext : forall n n' (Rn : [Γ ||- n ≅ n' : _ | RN]), + [Γ ||- P[n..] ≅ P[n'..] | RPpt _ _ Rn]. Proof. - eapply natElimRedAux; tea. - + intros. eapply transEq; [eapply LRTyEqSym |]; eapply RPQext; cycle 1; tea. - now eapply reflLRTmEq. - Unshelve. all:tea. + intros; eapply transEq; [| eapply LRTyEqSym ]. + all: unshelve (eapply RPQext; cycle 1; tea); now eapply urefl. + Unshelve. 2: eapply RQpt; now symmetry. Qed. Lemma natElimRedEqAux : - (forall n n' (Rnn' : [Γ ||- n ≅ n' : _ | RN]) (Rn : [Γ ||- n : _ | RN]) (Rn' : [Γ ||- n' : _ | RN]), - [Γ ||- tNatElim P hz hs n ≅ tNatElim Q hz' hs' n' : _ | RPpt _ Rn ]) × - (forall n n' (Rnn' : NatPropEq NN n n') (Rn : [Γ ||- n : _ | RN]) (Rn' : [Γ ||- n' : _ | RN]), - [Γ ||- tNatElim P hz hs n ≅ tNatElim Q hz' hs' n' : _ | RPpt _ Rn ]). + (forall n n' (Rnn' : [Γ ||- n ≅ n' : _ | RN]), [Γ ||- n ≅ n' : _ | RN] -> + [Γ ||- tNatElim P hz hs n ≅ tNatElim Q hz' hs' n' : _ | RPpt _ _ Rnn' ]) × + (forall n n' (Rnn' : NatPropEq NN n n') (RP : [Γ ||- P[n..]]), + (* (Rn : [Γ ||- n ≅ n' : _ | RN]), *) + [Γ ||- tNatElim P hz hs n ≅ tNatElim Q hz' hs' n' : _ | RP (*RPpt _ _ Rn*) ]). Proof. apply NatRedEqInduction. - - intros ?? nfL nfR redL redR ? prop ih RL RR; destruct (NatPropEq_whnf prop). - assert [Γ ||- nfL : _ | RN] by (eapply redTmFwd; cycle 1; tea). - assert [Γ ||- nfR : _ | RN] by (eapply redTmFwd; cycle 1; tea). - eapply LREqTermHelper. - * eapply (fst natElimRedAuxLeft). - * eapply (fst natElimRedAuxRight). - * eapply RPQext. 1: tea. - now econstructor. - * eapply LRTmEqRedConv. - + eapply RPext; tea. - eapply LRTmEqSym; eapply redwfSubstTerm; cycle 1; tea. - + unshelve erewrite (redtmwf_det _ _ (NatRedTm.red RL) redL); tea. - 1: dependent inversion RL; subst; cbn; now eapply NatProp_whnf. - unshelve erewrite (redtmwf_det _ _ (NatRedTm.red RR) redR); tea. - 1: dependent inversion RR; subst; cbn; now eapply NatProp_whnf. - now eapply ih. - Unshelve. tea. - - intros. eapply LREqTermHelper. - * unshelve eapply (snd natElimRedAuxLeft); constructor. - * unshelve eapply (snd natElimRedAuxRight); tea; constructor. - * eapply RPQext; [unshelve eapply zeroRed| unshelve eapply zeroRedEq]; tea. - * cbn; irrelevance. - - intros ??? ih Rn Rn'. pose proof (succRedInv Rn); pose proof (succRedInv Rn'). - eapply LREqTermHelper. - * unshelve eapply (snd natElimRedAuxLeft); now constructor. - * unshelve eapply (snd natElimRedAuxRight); tea; now constructor. - * eapply RPQext; [unshelve eapply succRed| unshelve eapply succRedEq]; tea. - * cbn. - eapply simple_appcongTerm. - 4: now eapply ih. - + irrelevance0. 2: eapply appcongTerm; tea. - now bsimpl. - + eapply (fst natElimRedAuxLeft). - + eapply LRTmRedConv. 2: eapply (fst natElimRedAuxRight). - eapply LRTyEqSym; now eapply RPQext. - Unshelve. 2,4,5: tea. - 1: eapply ArrRedTy; now eapply RPpt. - rewrite subst_arr. eapply ArrRedTy. - 2: rewrite liftSubst_singleSubst_eq; cbn. - all: now eapply RPpt. - - intros ?? neueq ??. escape. inversion neueq. - assert [Γ |- P[ne..] ≅ Q[ne'..]]. 1:{ - eapply escapeEq. eapply RPQext; tea. - econstructor. - 1,2: now eapply redtmwf_refl. - 2: now constructor. - gen_typing. - } - eapply neuTermEq. - + eapply ty_natElim; tea. - + eapply ty_conv. - 1: eapply ty_natElim; tea. - now symmetry. - + eapply convneu_natElim ; tea. - Unshelve. tea. + - intros ?? nfL nfR redL redR ? prop ih Rtu ; destruct (NatPropEq_whnf prop). + eapply redSubstTmEq. + + eapply LRTmEqConv, ih; unshelve eapply RPext. + eapply redTmEqFwd; tea; now eapply lrefl. + + escape; eapply redtm_natelim; tea; gen_typing. + + escape; eapply redtm_natelim; tea; [..| gen_typing]. + 1,2: now eapply ty_conv. + + now unshelve eapply escapeEq, RPQext. + - intros. eapply redSubstTmEq. + + irrelevanceRefl; eapply Rhz. + + escape; eapply redtm_natElimZero; tea. + + escape; eapply redtm_natElimZero; tea. + 1,2: now eapply ty_conv. + + now escape. + - intros n n' Rn ih ?; change [Γ ||- n ≅ n' : _ | RN] in Rn. + eapply redSubstTmEq; cycle 1. + + escape; eapply redtm_natElimSucc; tea. + + escape; eapply redtm_natElimSucc; tea. + 1,2: now eapply ty_conv. + + unshelve eapply escapeEq, RPQext. + now eapply succRedEq. + + assert [Γ ||- arr P[n..] P[(tSucc n)..]]. + 1: now (eapply ArrRedTy; eapply RPpt; [| eapply succRedEq]). + unshelve eapply simple_appcongTerm; [..| eauto]; tea. + unshelve (irrelevance0; [| eapply appcongTerm; tea]). + all: now rewrite subst_arr, liftSubst_singleSubst_eq. + - intros n n' Rn RP0. + pose proof (neNfTermEq RN Rn). + eapply neuTermEq. 2: eapply ty_conv ; [| eapply escapeEq; eapply LRTyEqSym]. + + escape; now eapply ty_natElim. + + escape; eapply ty_natElim; tea. + all: now eapply ty_conv. + + now unshelve eapply RPQext. + + escape ; destruct Rn; now eapply convneu_natElim. + Unshelve. 2: eapply RQpt; now eapply urefl. Qed. Lemma natElimRedEq : - (forall n n' (Rnn' : [Γ ||- n ≅ n' : _ | RN]) (Rn : [Γ ||- n : _ | RN]) (Rn' : [Γ ||- n' : _ | RN]), - [Γ ||- tNatElim P hz hs n ≅ tNatElim Q hz' hs' n' : _ | RPpt _ Rn ]). - Proof. apply natElimRedEqAux. Qed. + (forall n n' (Rnn' : [Γ ||- n ≅ n' : _ | RN]), + [Γ ||- tNatElim P hz hs n ≅ tNatElim Q hz' hs' n' : _ | RPpt _ _ Rnn' ]). + Proof. intros; now apply (fst natElimRedEqAux). Qed. End NatElimRedEq. - Section NatElimValid. - Context {Γ l P} + Context {Γ l} (VΓ : [||-v Γ]) (VN := natValid (l:=l) VΓ) - (VΓN := validSnoc VΓ VN) - (VP : [Γ ,, tNat ||-v P | VΓN]). - - Lemma elimSuccHypTyValid : + (VΓN := validSnoc VΓ VN). + + Lemma elimSuccHypTyValid {P} + (VP : [Γ ,, tNat ||-v P | VΓN]) : [Γ ||-v elimSuccHypTy P | VΓ]. Proof. unfold elimSuccHypTy. @@ -367,183 +230,14 @@ Section NatElimValid. change tNat with tNat⟨@wk1 Γ tNat⟩ at 2. eapply var0Valid. Unshelve. all: tea. - Qed. - - Context {hz hs} - (VPz := substS VP (zeroValid VΓ)) - (Vhz : [Γ ||-v hz : P[tZero..] | VΓ | VPz]) - (Vhs : [Γ ||-v hs : _ | VΓ | elimSuccHypTyValid]). - - - Lemma natElimValid {n} - (Vn : [Γ ||-v n : tNat | VΓ | VN]) - (VPn := substS VP Vn) - : [Γ ||-v tNatElim P hz hs n : _ | VΓ | VPn]. - Proof. - constructor; intros. - - instValid Vσ; cbn. - irrelevance0. 1: now rewrite singleSubstComm'. - epose proof (Vuσ := liftSubstS' VN Vσ). - instValid Vuσ; escape. - unshelve eapply natElimRed; tea. - + intros m Rm. - rewrite up_single_subst. - unshelve eapply validTy; cycle 1; tea. - unshelve econstructor; [now bsimpl| now cbn]. - + rewrite elimSuccHypTy_subst. eapply validTy; tea; eapply elimSuccHypTyValid. - + intros m m' Rm Rm' Rmm'; cbn. - irrelevance0. 1: now rewrite up_single_subst. - rewrite up_single_subst. - unshelve eapply validTyExt; cycle 1 ; tea. - 1,2: unshelve econstructor; [now bsimpl| now cbn]. - unshelve econstructor; [|now cbn]. - bsimpl; cbn; eapply reflSubst. - + irrelevance. - + irrelevance. now rewrite elimSuccHypTy_subst. - - instAllValid Vσ Vσ' Vσσ'. - irrelevance0. 1: now rewrite singleSubstComm'. - pose (Vuσ := liftSubstS' VN Vσ). - pose proof (Vuσ' := liftSubstS' VN Vσ'). - pose proof (Vuσrea := liftSubstSrealign' VN Vσσ' Vσ'). - pose proof (Vuσσ' := liftSubstSEq' VN Vσσ'). - instValid Vuσ'; instAllValid Vuσ Vuσrea Vuσσ'; escape. - unshelve eapply natElimRedEq; tea; fold subst_term. - 6-11: irrelevance; now rewrite elimSuccHypTy_subst. - 3,4: rewrite elimSuccHypTy_subst; eapply validTy; tea; eapply elimSuccHypTyValid. - + intros m Rm. - rewrite up_single_subst. - unshelve eapply validTy; cycle 1; tea. - unshelve econstructor; [now bsimpl| now cbn]. - + intros m Rm. - rewrite up_single_subst. - unshelve eapply validTy; cycle 1; tea. - unshelve econstructor; [now bsimpl| now cbn]. - + intros m m' Rm Rm' Rmm'; cbn. - irrelevance0. 1: now rewrite up_single_subst. - rewrite up_single_subst. - unshelve eapply validTyExt; cycle 1 ; tea. - 1,2: unshelve econstructor; [now bsimpl| now cbn]. - unshelve econstructor; [|now cbn]. - now bsimpl. - Qed. - - Lemma natElimZeroValid : - [Γ ||-v tNatElim P hz hs tZero ≅ hz : _ | VΓ | VPz]. - Proof. - constructor; intros. - epose proof (Vuσ := liftSubstS' VN Vσ). - instValid Vσ; instValid Vuσ; escape. - irrelevance0. 1: now rewrite singleSubstComm'. - unshelve eapply (snd (natElimRedAux _ _ _ _ _ _ _ _) _ NatRedTm.zeroR); tea; fold subst_term. - 7: eapply zeroValid. - + intros m Rm. - rewrite up_single_subst. - unshelve eapply validTy; cycle 1; tea. - unshelve econstructor; [now bsimpl| now cbn]. - + intros m m' Rm Rm' Rmm'; cbn. - irrelevance0. 1: now rewrite up_single_subst. - rewrite up_single_subst. - unshelve eapply validTyExt; cycle 1 ; tea. - 1,2: unshelve econstructor; [now bsimpl| now cbn]. - unshelve econstructor; [|now cbn]. - bsimpl; cbn; eapply reflSubst. - + irrelevance. - + rewrite elimSuccHypTy_subst. eapply validTy; tea; eapply elimSuccHypTyValid. - + irrelevance. now rewrite elimSuccHypTy_subst. - Unshelve. 4: tea. - Qed. + Defined. - Lemma natElimSuccValid {n} - (Vn : [Γ ||-v n : tNat | VΓ | VN]) - (VPSn := substS VP (succValid _ Vn)) : - [Γ ||-v tNatElim P hz hs (tSucc n) ≅ tApp (tApp hs n) (tNatElim P hz hs n) : _ | VΓ | VPSn]. + Lemma elimSuccHypTyCongValid {P P'} + (VP : [Γ ,, tNat ||-v P | VΓN]) + (VeqP : [Γ ,, tNat ||-v P ≅ P' | VΓN | VP ]) : + [Γ ||-v elimSuccHypTy P ≅ elimSuccHypTy P' | VΓ | elimSuccHypTyValid VP]. Proof. - constructor; intros. - epose proof (Vuσ := liftSubstS' VN Vσ). - instValid Vσ; instValid Vuσ; escape. - irrelevance0. 1: now rewrite singleSubstComm'. - pose (NSn := NatRedTm.succR RVn). - unshelve eapply (snd (natElimRedAux _ _ _ _ _ _ _ _) _ NSn); tea; fold subst_term. - + intros m Rm. - rewrite up_single_subst. - unshelve eapply validTy; cycle 1; tea. - unshelve econstructor; [now bsimpl| now cbn]. - + intros m m' Rm Rm' Rmm'; cbn. - irrelevance0. 1: now rewrite up_single_subst. - rewrite up_single_subst. - unshelve eapply validTyExt; cycle 1 ; tea. - 1,2: unshelve econstructor; [now bsimpl| now cbn]. - unshelve econstructor; [|now cbn]. - bsimpl; cbn; eapply reflSubst. - + irrelevance. - + rewrite elimSuccHypTy_subst. eapply validTy; tea; eapply elimSuccHypTyValid. - + irrelevance. now rewrite elimSuccHypTy_subst. - + now eapply succRed. - Qed. - -End NatElimValid. - -Lemma natElimCongValid {Γ l P Q hz hz' hs hs' n n'} - (VΓ : [||-v Γ]) - (VN := natValid (l:=l) VΓ) - (VΓN := validSnoc VΓ VN) - (VP : [Γ ,, tNat ||-v P | VΓN]) - (VQ : [Γ ,, tNat ||-v Q | VΓN]) - (VPQ : [Γ ,, tNat ||-v P ≅ Q | VΓN | VP]) - (VPz := substS VP (zeroValid VΓ)) - (VQz := substS VQ (zeroValid VΓ)) - (Vhz : [Γ ||-v hz : P[tZero..] | VΓ | VPz]) - (Vhz' : [Γ ||-v hz' : Q[tZero..] | VΓ | VQz]) - (Vheqz : [Γ ||-v hz ≅ hz' : P[tZero..] | VΓ | VPz]) - (VPs := elimSuccHypTyValid VΓ VP) - (VQs := elimSuccHypTyValid VΓ VQ) - (Vhs : [Γ ||-v hs : _ | VΓ | VPs]) - (Vhs' : [Γ ||-v hs' : _ | VΓ | VQs]) - (Vheqs : [Γ ||-v hs ≅ hs' : _ | VΓ | VPs]) - (Vn : [Γ ||-v n : _ | VΓ | VN]) - (Vn' : [Γ ||-v n' : _ | VΓ | VN]) - (Veqn : [Γ ||-v n ≅ n' : _ | VΓ | VN]) - (VPn := substS VP Vn) : - [Γ ||-v tNatElim P hz hs n ≅ tNatElim Q hz' hs' n' : _ | VΓ | VPn]. -Proof. - constructor; intros. - pose (Vuσ := liftSubstS' VN Vσ). - instValid Vσ; instValid Vuσ; escape. - irrelevance0. 1: now rewrite singleSubstComm'. - unshelve eapply natElimRedEq; tea; fold subst_term. - 6-11: irrelevance; now rewrite elimSuccHypTy_subst. - 3,4: rewrite elimSuccHypTy_subst; eapply validTy; tea; eapply elimSuccHypTyValid. - + intros m Rm. - rewrite up_single_subst. - unshelve eapply validTy; cycle 1; tea. - unshelve econstructor; [now bsimpl| now cbn]. - + intros m Rm. - rewrite up_single_subst. - unshelve eapply validTy; cycle 1; tea. - unshelve econstructor; [now bsimpl| now cbn]. - + intros m m' Rm Rm' Rmm'; cbn. - irrelevance0. 1: now rewrite up_single_subst. - rewrite up_single_subst. - eapply transEq; cycle 1. - - unshelve eapply validTyEq. - 7: tea. 1: tea. - unshelve econstructor; [now bsimpl| now cbn]. - - unshelve eapply validTyExt; cycle 1 ; tea. - 1,2: unshelve econstructor; [now bsimpl| now cbn]. - unshelve econstructor; [|now cbn]. - bsimpl. eapply reflSubst. -Qed. - -Lemma elimSuccHypTyCongValid {Γ l P P'} - (VΓ : [||-v Γ]) - (VN := natValid (l:=l) VΓ) - (VΓN := validSnoc VΓ VN) - (VP : [Γ ,, tNat ||-v P | VΓN]) - (VP' : [Γ ,, tNat ||-v P' | VΓN]) - (VeqP : [Γ ,, tNat ||-v P ≅ P' | VΓN | VP]) : - [Γ ||-v elimSuccHypTy P ≅ elimSuccHypTy P' | VΓ | elimSuccHypTyValid VΓ VP]. - Proof. - unfold elimSuccHypTy. + unfold elimSuccHypTy. pose proof (ureflValidTy VeqP). eapply irrelevanceTyEq. assert [Γ,, tNat ||-v< l > P'[tSucc (tRel 0)]⇑ | validSnoc VΓ VN]. 1:{ eapply substLiftS; tea. @@ -567,4 +261,106 @@ Lemma elimSuccHypTyCongValid {Γ l P P'} now bsimpl. Qed. + + + Context { P P' hz hz' hs hs'} + (VP : [Γ ,, tNat ||-v P | VΓN]) + (VPP' : [Γ ,, tNat ||-v P ≅ P' | VΓN | VP ]) + (VP' := ureflValidTy VPP') + (VPz := substS VP (zeroValid VΓ)) + (Vhz : [Γ ||-v hz ≅ hz' : P[tZero..] | VΓ | VPz]) + (Vhs : [Γ ||-v hs ≅ hs' : _ | VΓ | elimSuccHypTyValid VP]). + + (* Lemma elimSuccHypTyCongValid ? *) + + Lemma natElimCongValid {n n'} + (Vn : [Γ ||-v n ≅ n' : tNat | VΓ | VN]) + (VPn := substS VP Vn) + : [Γ ||-v tNatElim P hz hs n ≅ tNatElim P' hz' hs' n' : _ | VΓ | VPn]. + Proof. + pose proof (elimSuccHypTyValid VP). + pose proof (elimSuccHypTyValid VP'). + pose proof (elimSuccHypTyCongValid VP VPP'). + constructor; intros; instValid Vσσ'; epose proof (Vuσ := liftSubstEq' VN Vσσ'). + instValid Vuσ; cbn -[elimSuccHypTy elimSuccHypTyValid] in *; escape. + irrelevance0. 1: now rewrite singleSubstComm'. + eapply natElimRedEq; rewrite ?elimSuccHypTy_subst; tea. + + intros; irrelevance0; rewrite up_single_subst; [reflexivity|]. + unshelve eapply validTyEq; cycle 1 ; tea. + now unshelve eapply consSubstEq. + + irrelevance0; [exact (singleSubstComm' _ tZero σ)|]; tea. + + irrelevance0; [now rewrite elimSuccHypTy_subst|]; tea. + + irrelevance0; [now rewrite elimSuccHypTy_subst|]; tea. + Unshelve. + 2,3: tea. + 2: now rewrite elimSuccHypTy_subst. + intros ?? Rn; rewrite up_single_subst. + unshelve eapply validTy. + 5: tea. + 3: unshelve eapply consSubstEq; tea. + Qed. +End NatElimValid. + +Lemma natElimValid {Γ l P hz hz' hs hs' n n'} + (VΓ : [||-v Γ]) + (VN := natValid (l:=l) VΓ) + (VΓN := validSnoc VΓ VN) + (VP : [Γ ,, tNat ||-v P | VΓN]) + (VPz := substS VP (zeroValid VΓ)) + (Vhz : [Γ ||-v hz ≅ hz' : P[tZero..] | VΓ | VPz]) + (Vhs : [Γ ||-v hs ≅ hs' : _ | VΓ | elimSuccHypTyValid VΓ VP]) + (Vn : [Γ ||-v n ≅ n' : _ | VΓ | VN]) + (VPn := substS VP Vn): + [Γ ||-v tNatElim P hz hs n : _ | VΓ | VPn]. +Proof. + eapply lreflValidTm , natElimCongValid; tea. + now eapply reflValidTy. +Qed. + +Section NatElimRedValid. + Context {Γ l} + (VΓ : [||-v Γ]) + (VN := natValid (l:=l) VΓ) + (VΓN := validSnoc VΓ VN) + { P P' hz hz' hs hs'} + (VP : [Γ ,, tNat ||-v P | VΓN]) + (VPP' : [Γ ,, tNat ||-v P ≅ P' | VΓN | VP ]) + (VP' := ureflValidTy VPP') + (VPz := substS VP (zeroValid VΓ)) + (Vhz : [Γ ||-v hz ≅ hz' : P[tZero..] | VΓ | VPz]) + (Vhs : [Γ ||-v hs ≅ hs' : _ | VΓ | elimSuccHypTyValid VΓ VP]). + + Lemma natElimZeroValid : + [Γ ||-v tNatElim P hz hs tZero ≅ hz : _ | VΓ | VPz]. + Proof. + eapply redSubstValid. 2: now eapply lreflValidTm. + constructor; intros; cbn; rewrite singleSubstComm'. + instValid Vσσ'; instValid (liftSubstEq' VN Vσσ'); escape. + eapply redtm_natElimZero; tea. + + now rewrite <- (singleSubstComm' _ tZero σ). + + now rewrite elimSuccHypTy_subst. + Qed. + + Lemma natElimSuccValid {n} + (Vn : [Γ ||-v n : tNat | VΓ | VN]) + (VPSn := substS VP (succValid _ Vn)) : + [Γ ||-v tNatElim P hz hs (tSucc n) ≅ tApp (tApp hs n) (tNatElim P hz hs n) : _ | VΓ | VPSn]. + Proof. + eapply redSubstValid. + * constructor; intros; rewrite singleSubstComm'. + instValid Vσσ'; instValid (liftSubstEq' VN Vσσ'); escape. + eapply redtm_natElimSucc; tea; refold. + + now rewrite <- (singleSubstComm' _ tZero σ). + + now rewrite elimSuccHypTy_subst. + * eapply simple_appValid. + 2: now unshelve (eapply natElimValid; tea). + eapply irrelevanceTm'; [| exact (appValid (lreflValidTm _ Vhs) Vn)]. + clear; now bsimpl. + Unshelve. + eapply simpleArrValid; eapply substS; tea. + now eapply succValid. + Qed. + +End NatElimRedValid. + End Nat. diff --git a/theories/Substitution/Introductions/Pi.v b/theories/Substitution/Introductions/Pi.v index db0bd22..6ca9b45 100644 --- a/theories/Substitution/Introductions/Pi.v +++ b/theories/Substitution/Introductions/Pi.v @@ -2,21 +2,12 @@ From Coq Require Import ssrbool. From LogRel.AutoSubst Require Import core unscoped Ast Extra. From LogRel Require Import Utils BasicAst Notations Context NormalForms Weakening GenericTyping LogicalRelation Validity. -From LogRel.LogicalRelation Require Import Escape Reflexivity Neutral Weakening Irrelevance. -From LogRel.Substitution Require Import Irrelevance Properties. +From LogRel.LogicalRelation Require Import Escape Reflexivity Neutral Weakening Irrelevance Induction NormalRed EqRedRight. +From LogRel.Substitution Require Import Irrelevance Properties SingleSubst Reflexivity. From LogRel.Substitution.Introductions Require Import Universe Poly. Set Universe Polymorphism. -(* Lemma eq_subst_1 F G Δ σ : G[up_term_term σ] = G[tRel 0 .: σ⟨ @wk1 Δ F[σ] ⟩]. -Proof. - now bsimpl. -Qed. - -Lemma eq_subst_2 G a ρ σ : G[up_term_term σ][a .: ρ >> tRel] = G[a .: σ⟨ρ⟩]. -Proof. - bsimpl ; now substify. -Qed. *) Section PolyRedPi. Context `{GenericTypingProperties}. @@ -56,7 +47,7 @@ Section PolyRedPi. Proof. intros; irrelevanceRefl; now eapply LRPiPolyEq. Qed. - + End PolyRedPi. @@ -67,40 +58,39 @@ Section PiTyValidity. (vF : [Γ ||-v< l > F | vΓ ]) (vG : [Γ ,, F ||-v< l > G | validSnoc vΓ vF]). - Lemma PiRed {Δ σ} (tΔ : [ |-[ ta ] Δ]) - (vσ : [vΓ | Δ ||-v σ : _ | tΔ]) + Lemma PiRed {Δ σ σ'} (tΔ : [ |-[ ta ] Δ]) + (vσ : [vΓ | Δ ||-v σ ≅ σ' : _ | tΔ]) : [ Δ ||-< l > (tProd F G)[σ] ]. Proof. eapply LRPi'. - pose (p := substPolyRed vΓ vF vG _ vσ). + pose (p := substPolyRed vΓ vF vG _ (lrefl vσ)). escape; cbn; econstructor; tea; destruct (polyRedId p); - destruct (polyRedEqId p (substPolyRedEq vΓ vF vG _ vσ vσ (reflSubst _ _ vσ))); escape. + destruct (polyRedEqId p (substPolyRedEq vΓ vF vG _ (lrefl vσ) p)); escape. - apply redtywf_refl; gen_typing. - gen_typing. - gen_typing. Defined. Lemma PiEqRed1 {Δ σ σ'} (tΔ : [ |-[ ta ] Δ]) - (vσ : [vΓ | Δ ||-v σ : _ | tΔ]) - (vσ' : [vΓ | Δ ||-v σ' : _ | tΔ]) - (vσσ' : [vΓ | Δ ||-v σ ≅ σ' : _ | tΔ | vσ]) - : [ Δ ||-< l > (tProd F G)[σ] ≅ (tProd F G)[σ'] | PiRed tΔ vσ ]. + (vσσ' : [vΓ | Δ ||-v σ ≅ σ' : _ | tΔ ]) + : [ Δ ||-< l > (tProd F G)[σ] ≅ (tProd F G)[σ'] | PiRed tΔ vσσ' ]. Proof. - pose (p := substPolyRed vΓ vF vG _ vσ). - pose (p' := substPolyRed vΓ vF vG _ vσ'). - pose (peq := substPolyRedEq vΓ vF vG _ vσ vσ' vσσ'). - destruct (polyRedId p); destruct (polyRedId p'); destruct (polyRedEqId p peq). + pose (p := substPolyRed vΓ vF vG _ vσσ'). + pose (p' := substPolyRed vΓ vF vG _ (urefl vσσ')). + pose (peq := substPolyRedEq vΓ vF vG _ vσσ' p). + destruct (polyRedId p), (polyRedId p'), (polyRedEqId p peq). escape; econstructor; cbn; tea. - apply redtywf_refl; gen_typing. - gen_typing. + - now eapply substPolyRedEq. Defined. Lemma PiValid : [Γ ||-v< l > tProd F G | vΓ]. Proof. unshelve econstructor. - - intros Δ σ. eapply PiRed. - - intros Δ σ σ'. eapply PiEqRed1. + - intros; now eapply PiRed. + - intros; eapply PiEqRed1. Defined. End PiTyValidity. @@ -114,91 +104,26 @@ Section PiTyDomValidity. Lemma PiValidDom : [Γ ||-v< l > F | vΓ]. Proof. unshelve econstructor. - - intros Δ σ vΔ vσ; instValid vσ. - cbn in RvΠFG; apply Induction.invLRΠ in RvΠFG. - destruct RvΠFG as [dom ? red ? ? ? [? ? Hdom]]. - assert (Hrw : F[σ] = dom); [|subst dom]. - { eapply redtywf_whnf in red as [=]; [tea|constructor]. } - rewrite <- (wk_id_ren_on Δ F[σ]). - now eapply Hdom. - - cbn; refold. - intros Δ σ σ' vΔ vσ vσ' vσσ'. - match goal with |- [ LRAd.pack ?P | _ ||- _ ≅ _] => set (vF := P); clearbody vF end. - instValid vσ. - cbn in RvΠFG; apply Induction.invLRΠ in RvΠFG. - assert (vΠ := validTyExt vΠFG _ vσ vσ' vσσ'). - eapply LRTyEqIrrelevant' with (lrA' := LRPi' RvΠFG) in vΠ; [|reflexivity]. - destruct RvΠFG as [dom ? red ? ? ? []]. - destruct vΠ as [dom' ? red' ? ? [Heq _]]; simpl in *. - specialize (Heq Δ wk_id vΔ). - assert (Hrw : F[σ] = dom). - { eapply redtywf_whnf in red as [=]; [tea|constructor]. } - assert (Hrw' : F[σ'] = dom'). - { eapply redtywf_whnf in red' as [=]; [tea|constructor]. } - rewrite wk_id_ren_on, <- Hrw' in Heq. - eapply Transitivity.transEq; [|eapply Heq]. - rewrite <- (wk_id_ren_on Δ dom) in Hrw; rewrite <- Hrw. - eapply LRTyEqIrrelevant' with (lrA := vF); [reflexivity|]. - eapply reflLRTyEq. + - intros Δ vΔ σ σ' vσ; instValid vσ. + now pose proof (polyRedId (normRedΠ0 (invLRΠ RvΠFG))) as []. + - intros Δ vΔ σ σ' vσσ'; instValid vσσ'. + pose proof (polyRedEqId _ (normEqRedΠ _ REvΠFG)) as [? _]. + irrelevance. Qed. + Lemma eq_subst_eta A σ : A[σ] = A[up_term_term (↑ >> σ)][(σ var_zero)..]. + Proof. bsimpl; now rewrite scons_eta'. Qed. + Lemma PiValidCod : [Γ,, F ||-v< l > G | validSnoc vΓ PiValidDom]. Proof. - unshelve econstructor. - - intros Δ σ vΔ [vσ v0]. - instValid vσ; cbn in *. - apply Induction.invLRΠ in RvΠFG. - destruct RvΠFG as [dom cod red ? ? ? [? ? Hdom Hcod _]]. - specialize (Hcod Δ (σ 0) wk_id vΔ). - assert (HF : F[↑ >> σ] = dom). - { eapply redtywf_whnf in red as [=]; [tea|constructor]. } - assert (HG0 : G[up_term_term (↑ >> σ)] = cod). - { eapply redtywf_whnf in red as [=]; [tea|constructor]. } - assert (HG : G[σ] = cod[σ 0 .: @wk_id Δ >> tRel]). - { rewrite <- HG0; bsimpl; apply ext_term; intros []; reflexivity. } - rewrite HG; apply Hcod. - irrelevance0; [|eapply v0]. - now rewrite wk_id_ren_on. - - cbn; refold. - intros Δ σ σ' vΔ [vσ v0] [vσ' v0'] [vσσ' v00']. - match goal with |- [ LRAd.pack ?P | _ ||- _ ≅ _] => set (vG := P); clearbody vG end. - instValid vσ. - cbn in RvΠFG; apply Induction.invLRΠ in RvΠFG. - assert (vΠ := validTyExt vΠFG _ vσ vσ' vσσ'). - eapply LRTyEqIrrelevant' with (lrA' := LRPi' RvΠFG) in vΠ; [|reflexivity]. - destruct RvΠFG as [dom cod red ? ? ? [? ? ? ? Hcod]]. - destruct vΠ as [dom' cod' red' ? ? [Hdom' Hcod']]; simpl in *. - specialize (Hcod' Δ (σ' 0) wk_id vΔ). - assert (HF : F[↑ >> σ] = dom). - { eapply redtywf_whnf in red as [=]; [tea|constructor]. } - assert (HF' : F[↑ >> σ'] = dom'). - { eapply redtywf_whnf in red' as [=]; [tea|constructor]. } - assert (HG0 : G[up_term_term (↑ >> σ)] = cod). - { eapply redtywf_whnf in red as [=]; [tea|constructor]. } - assert (HG : G[σ] = cod[σ 0 .: @wk_id Δ >> tRel]). - { rewrite <- HG0; bsimpl; apply ext_term; intros []; reflexivity. } - assert (HG0' : G[up_term_term (↑ >> σ')] = cod'). - { eapply redtywf_whnf in red' as [=]; [tea|constructor]. } - assert (HG' : G[σ'] = cod'[σ' 0 .: @wk_id Δ >> tRel]). - { rewrite <- HG0'; bsimpl; apply ext_term; intros []; reflexivity. } - rewrite HG'. - assert (Hσ0 : [shpRed Δ wk_id vΔ | _ ||- σ 0 : _]). - { irrelevance0; [|eapply v0]. - now rewrite wk_id_ren_on. } - assert (Hσ0' : [shpRed Δ wk_id vΔ | _ ||- σ' 0 : _]). - { eapply LRTmRedConv; [|eapply v0']. - eapply LRTyEqSym; rewrite HF'. - rewrite <- (wk_id_ren_on Δ dom'). - apply (Hdom' _ _ vΔ). } - eassert (Hcod0 := Hcod Δ (σ 0) (σ' 0) wk_id vΔ Hσ0 Hσ0'). - eapply Transitivity.transEq; [|now unshelve eapply Hcod']. - eapply Transitivity.transEq; [|unshelve eapply Hcod0]. - + unshelve (irrelevance0; [|eapply reflLRTyEq]); try now symmetry. - shelve. - now eauto. - + irrelevance0; [|eapply v00']. - now rewrite wk_id_ren_on. - Qed. + unshelve econstructor. + - intros Δ vΔ σ σ' [vσ v0]; instValid vσ; cbn in *. + rewrite eq_subst_eta; eapply singleSubstΠ1; tea. + - cbn; refold. + intros Δ vΔ σ σ' [vσσ' v00']; instValid vσσ'; instValid (urefl vσσ'); cbn in *. + pose proof (polyRedEqId _ (normEqRedΠ _ REvΠFG)) as [? _]. + rewrite 2!eq_subst_eta; eapply singleSubstΠ2; tea. + Qed. End PiTyDomValidity. @@ -214,12 +139,12 @@ Section PiTyCongruence. (vFF' : [ Γ ||-v< l > F ≅ F' | vΓ | vF ]) (vGG' : [ Γ ,, F ||-v< l > G ≅ G' | validSnoc vΓ vF | vG ]). - Lemma PiEqRed2 {Δ σ} (tΔ : [ |-[ ta ] Δ]) (vσ : [vΓ | Δ ||-v σ : _ | tΔ]) - : [ Δ ||-< l > (tProd F G)[σ] ≅ (tProd F' G')[σ] | validTy (PiValid vΓ vF vG) tΔ vσ ]. + Lemma PiEqRed2 {Δ σ σ'} (tΔ : [ |-[ ta ] Δ]) (vσ : [vΓ | Δ ||-v σ ≅ σ' : _ | tΔ]) + : [ Δ ||-< l > (tProd F G)[σ] ≅ (tProd F' G')[σ'] | validTy (PiValid vΓ vF vG) tΔ vσ ]. Proof. - pose (p := substPolyRed vΓ vF vG _ vσ). - pose (p' := substPolyRed vΓ vF' vG' _ vσ). - pose (peq := substEqPolyRedEq vΓ vF vG _ vσ vF' vG' vFF' vGG'). + pose (p := substPolyRed vΓ vF vG _ (lrefl vσ)). + pose (p' := substPolyRed vΓ vF' vG' _ (urefl vσ)). + pose (peq := substEqPolyRedEq vΓ vF vG _ vσ vF' vG' vFF' vGG' p). destruct (polyRedId p); destruct (polyRedId p'); destruct (polyRedEqId p peq). escape; econstructor; cbn; tea. - apply redtywf_refl; gen_typing. @@ -228,73 +153,12 @@ Section PiTyCongruence. Lemma PiCong : [ Γ ||-v< l > tProd F G ≅ tProd F' G' | vΓ | PiValid vΓ vF vG ]. Proof. - econstructor. intros Δ σ. eapply PiEqRed2. + econstructor. intros *; eapply PiEqRed2. Qed. End PiTyCongruence. -Section PiTmValidity. - - Context `{GenericTypingProperties}. - Context {Γ F G} (VΓ : [||-v Γ]) - (VF : [ Γ ||-v< one > F | VΓ ]) - (VU : [ Γ ,, F ||-v< one > U | validSnoc VΓ VF ]) - (VFU : [ Γ ||-v< one > F : U | VΓ | UValid VΓ ]) - (VGU : [ Γ ,, F ||-v< one > G : U | validSnoc VΓ VF | VU ]) . - (* (VF := univValid (l':=zero) _ _ vFU) - (VG := univValid (l':=zero) _ _ vGU). *) - - Lemma prodTyEqU {Δ σ σ'} (tΔ : [ |-[ ta ] Δ]) - (Vσ : [VΓ | Δ ||-v σ : _ | tΔ]) - (Vσ' : [VΓ | Δ ||-v σ' : _ | tΔ]) - (Vσσ' : [VΓ | Δ ||-v σ ≅ σ' : _ | tΔ | Vσ ]) - : [Δ |-[ ta ] tProd F[σ] G[up_term_term σ] ≅ tProd F[σ'] G[up_term_term σ'] : U]. - Proof. - pose proof (Vuσ := liftSubstS' VF Vσ). - pose proof (Vureaσ' := liftSubstSrealign' VF Vσσ' Vσ'). - pose proof (Vuσσ' := liftSubstSEq' VF Vσσ'). - instAllValid Vσ Vσ' Vσσ'; instAllValid Vuσ Vureaσ' Vuσσ'; escape. - eapply convtm_prod; tea. - Qed. - - Lemma PiRedU {Δ σ} (tΔ : [ |-[ ta ] Δ]) (Vσ : [VΓ | Δ ||-v σ : _ | tΔ]) - : [ Δ ||-< one > (tProd F G)[σ] : U | validTy (UValid VΓ) tΔ Vσ ]. - Proof. - pose proof (Vσσ := reflSubst _ _ Vσ). - pose proof (Vuσ := liftSubstS' VF Vσ). - pose proof (Vuσσ := liftSubstSEq' VF Vσσ). - instAllValid Vσ Vσ Vσσ; instAllValid Vuσ Vuσ Vuσσ; escape. - econstructor; cbn. - - apply redtmwf_refl; cbn in *; now eapply ty_prod. - - constructor. - - now eapply convtm_prod. - - cbn. unshelve refine (LRCumulative (PiRed _ _ _ tΔ Vσ)); - unshelve eapply univValid; tea; try irrValid. - Defined. - - Lemma PiValidU : [ Γ ||-v< one > tProd F G : U | VΓ | UValid VΓ ]. - Proof. - econstructor. - - intros Δ σ tΔ Vσ. - exact (PiRedU tΔ Vσ). - - intros Δ σ σ' tΔ Vσ Vσ' Vσσ'. - pose proof (univValid (l' := zero) _ _ VFU) as VF0. - pose proof (irrelevanceTy (validSnoc VΓ VF) - (validSnoc (l := zero) VΓ VF0) - (univValid (l' := zero) _ _ VGU)) as VG0. - unshelve econstructor ; cbn. - + exact (PiRedU tΔ Vσ). - + exact (PiRedU tΔ Vσ'). - + exact (LRCumulative (PiRed VΓ VF0 VG0 tΔ Vσ)). - + exact (prodTyEqU tΔ Vσ Vσ' Vσσ'). - + exact (LRCumulative (PiRed VΓ VF0 VG0 tΔ Vσ')). - + pose proof (PiEqRed1 VΓ VF0 VG0 tΔ Vσ Vσ' Vσσ'). - irrelevanceCum. - Qed. - -End PiTmValidity. - Section PiTmCongruence. @@ -315,13 +179,13 @@ Section PiTmCongruence. Lemma PiCongTm : [ Γ ||-v< one > tProd F G ≅ tProd F' G' : U | vΓ | UValid vΓ ]. Proof. econstructor. - intros Δ σ tΔ Vσ. + intros Δ tΔ σ σ' Vσσ'. pose proof (univValid (l' := zero) _ _ vFU) as vF0. pose proof (univValid (l' := zero) _ _ vF'U) as vF'0. - pose proof (Vσσ := reflSubst _ _ Vσ). - pose proof (Vuσ := liftSubstS' vF Vσ). - pose proof (Vuσσ := liftSubstSEq' vF Vσσ). - instAllValid Vσ Vσ Vσσ; instAllValid Vuσ Vuσ Vuσσ; escape. + pose proof (Vuσ := liftSubstEq' vF Vσσ'). + (* pose proof (Vuσ' := liftSubstEq' vF (urefl Vσσ')). *) + pose proof (Vuσ' := liftSubstEq' vF' (urefl Vσσ')). + instValid Vσσ'; instValid (urefl Vσσ'); instValid Vuσ ; instValid Vuσ'; escape. pose proof (irrelevanceTy (validSnoc vΓ vF) (validSnoc (l := zero) vΓ vF0) (univValid (l' := zero) _ _ vGU)) as vG0. @@ -329,13 +193,14 @@ Section PiTmCongruence. (validSnoc (l := zero) vΓ vF'0) (univValid (l' := zero) _ _ vG'U)) as vG'0. unshelve econstructor ; cbn. - - exact (PiRedU vΓ vF vU vFU vGU tΔ Vσ). - - exact (PiRedU vΓ vF' vU' vF'U vG'U tΔ Vσ). - - exact (LRCumulative (PiRed vΓ vF0 vG0 tΔ Vσ)). + 1,2: econstructor; [apply redtmwf_refl; cbn; now eapply ty_prod| constructor]. + - unshelve refine (LRCumulative (PiRed _ _ _ tΔ Vσσ')). + all: unshelve eapply univValid; tea; try irrValid. - now eapply convtm_prod. - - exact (LRCumulative (PiRed vΓ vF'0 vG'0 tΔ Vσ)). - - enough ([ Δ ||-< zero > (tProd F G)[σ] ≅ (tProd F' G')[σ] | PiRed vΓ vF0 vG0 tΔ Vσ]) by irrelevanceCum. - refine (PiEqRed2 vΓ vF0 vG0 vF'0 vG'0 _ _ tΔ Vσ). + - unshelve refine (LRCumulative (PiRed _ _ _ tΔ (urefl Vσσ'))). + all: unshelve eapply univValid; tea; try irrValid. + - enough ([ Δ ||-< zero > (tProd F G)[σ] ≅ (tProd F' G')[σ'] | PiRed vΓ vF0 vG0 tΔ Vσσ']) by irrelevanceCum. + refine (PiEqRed2 vΓ vF0 vG0 vF'0 vG'0 _ _ tΔ Vσσ'). + exact (univEqValid vΓ (UValid vΓ) vF0 vFF'). + pose proof (univEqValid (validSnoc vΓ vF) vU (univValid (l' := zero) _ _ vGU) vGG') as vGG'0. refine (irrelevanceTyEq _ _ _ _ vGG'0). @@ -343,3 +208,19 @@ Section PiTmCongruence. End PiTmCongruence. + +Section PiTmValidity. + + Context `{GenericTypingProperties}. + Context {Γ F G} (VΓ : [||-v Γ]) + (VF : [ Γ ||-v< one > F | VΓ ]) + (VU : [ Γ ,, F ||-v< one > U | validSnoc VΓ VF ]) + (VFU : [ Γ ||-v< one > F : U | VΓ | UValid VΓ ]) + (VGU : [ Γ ,, F ||-v< one > G : U | validSnoc VΓ VF | VU ]) . + + Lemma PiValidU : [ Γ ||-v< one > tProd F G : U | VΓ | UValid VΓ ]. + Proof. + now eapply lreflValidTm, PiCongTm. + Qed. + +End PiTmValidity. diff --git a/theories/Substitution/Introductions/Poly.v b/theories/Substitution/Introductions/Poly.v index c7133fe..8592cc0 100644 --- a/theories/Substitution/Introductions/Poly.v +++ b/theories/Substitution/Introductions/Poly.v @@ -2,8 +2,8 @@ From Coq Require Import ssrbool. From LogRel.AutoSubst Require Import core unscoped Ast Extra. From LogRel Require Import Utils BasicAst Notations Context NormalForms Weakening GenericTyping LogicalRelation Validity. -From LogRel.LogicalRelation Require Import Escape Reflexivity Neutral Weakening Irrelevance. -From LogRel.Substitution Require Import Irrelevance Properties. +From LogRel.LogicalRelation Require Import Escape Reflexivity Neutral Weakening Irrelevance Transitivity EqRedRight InstKripke. +From LogRel.Substitution Require Import Irrelevance Properties Reflexivity. From LogRel.Substitution.Introductions Require Import Universe. Set Universe Polymorphism. @@ -21,7 +21,6 @@ Qed. Lemma subst_wk_id_tail Γ P t : P[t .: @wk_id Γ >> tRel] = P[t..]. Proof. setoid_rewrite id_ren; now bsimpl. Qed. - Set Printing Primitive Projection Parameters. Section PolyValidity. @@ -29,27 +28,24 @@ Section PolyValidity. Context `{GenericTypingProperties}. Lemma mkPolyRed {Γ l A B} (wfΓ : [|-Γ]) - (RA : forall Δ (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]), [Δ ||- A⟨ρ⟩]) - (RB : forall Δ a (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]), [Δ ||- a : _ | RA Δ ρ wfΔ] -> [Δ ||- B[a .: ρ >> tRel]]) - (RBext : forall Δ a b (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]) (Ra : [Δ ||- a : _ | RA Δ ρ wfΔ]), - [Δ ||- b : _ | RA Δ ρ wfΔ] -> - [Δ ||- a ≅ b : _ | RA Δ ρ wfΔ] -> [Δ ||- B[a .: ρ >> tRel] ≅ B[b .: ρ >> tRel] | RB Δ a ρ wfΔ Ra]) : + (RA : forall Δ (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]), [Δ ||- A⟨ρ⟩]) + (RB : forall Δ a b (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]), [Δ ||- a ≅ b : _ | RA Δ ρ wfΔ] -> [Δ ||- B[a .: ρ >> tRel]]) + (RBext : forall Δ a b (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]) (Rab : [Δ ||- a ≅ b : _ | RA Δ ρ wfΔ]), + [Δ ||- B[a .: ρ >> tRel] ≅ B[b .: ρ >> tRel] | RB Δ a b ρ wfΔ Rab]) : PolyRed Γ l A B. Proof. assert [Γ |- A] by (eapply escape; now eapply instKripke). unshelve econstructor; tea. - unshelve epose proof (RB _ (tRel 0) (@wk1 Γ A) _ _). + unshelve epose proof (h := RB _ (tRel 0) (tRel 0) (@wk1 Γ A) _ _). + gen_typing. + eapply var0; tea; now rewrite wk1_ren_on. - + enough (B = B[tRel 0 .: @wk1 Γ A >> tRel]) as -> by now escape. - setoid_rewrite wk1_ren; rewrite scons_eta'; now asimpl. + + escape. now rewrite <- eq_id_subst_scons in Esch. Qed. - Lemma mkPolyRed' {Γ l A B} (RA : [Γ ||- A]) - (RB : forall Δ a (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]), [Δ ||- a : _ | wk ρ wfΔ RA] -> [Δ ||- B[a .: ρ >> tRel]]) - (RBext : forall Δ a b (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]) (Ra : [Δ ||- a : _ | wk ρ wfΔ RA]), - [Δ ||- b : _ | wk ρ wfΔ RA] -> - [Δ ||- a ≅ b : _ | wk ρ wfΔ RA] -> [Δ ||- B[a .: ρ >> tRel] ≅ B[b .: ρ >> tRel] | RB Δ a ρ wfΔ Ra]) : + Lemma mkPolyRed' {Γ l A B} (RA : [Γ ||- A]) + (RB : forall Δ a b (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]), [Δ ||- a ≅ b : _ | wk ρ wfΔ RA] -> [Δ ||- B[a .: ρ >> tRel]]) + (RBext : forall Δ a b (ρ : Δ ≤ Γ) (wfΔ : [|- Δ]) (Rab : [Δ ||- a ≅ b : _ | wk ρ wfΔ RA]), + [Δ ||- B[a .: ρ >> tRel] ≅ B[b .: ρ >> tRel] | RB Δ a b ρ wfΔ Rab]) : PolyRed Γ l A B. Proof. unshelve eapply mkPolyRed; tea. @@ -57,57 +53,59 @@ Section PolyValidity. Qed. - Lemma polyCodSubstRed {Γ l F G} (RF : [Γ ||- F]) : - PolyRed Γ l F G -> forall t, [Γ ||- t : _ | RF] -> [Γ ||- G[t..]]. + Lemma polyCodSubstRed {Γ l F G} (RF : [Γ ||- F]) : + PolyRed Γ l F G -> forall t u, [Γ ||- t ≅ u : _ | RF] -> [Γ ||- G[t..]]. Proof. - intros PFG ??. - escape. assert (wfΓ : [|- Γ]) by gen_typing. + intros PFG ???. + assert (wfΓ : [|- Γ]) by (escape ; gen_typing). erewrite <- subst_wk_id_tail. - eapply (PolyRed.posRed PFG wk_id wfΓ). + eapply (PolyRed.posRed PFG wk_id wfΓ ). irrelevance. Qed. Lemma polyCodSubstExtRed {Γ l F G} (RF : [Γ ||- F]) (PFG : PolyRed Γ l F G) (RG := polyCodSubstRed RF PFG) : - forall t u (Rt : [Γ ||- t : _ | RF]), [Γ ||- u : _ | RF] -> [Γ ||- t ≅ u : _ | RF] -> - [Γ ||- G[t..] ≅ G[u..]| RG t Rt]. + forall t u (Rt : [Γ ||- t ≅ u : _ | RF]), + [Γ ||- G[t..] ≅ G[u..]| RG t u Rt]. Proof. - intros. escape. assert (wfΓ : [|- Γ]) by gen_typing. + intros. assert (wfΓ : [|- Γ]) by (escape; gen_typing). irrelevance0; erewrite <- subst_wk_id_tail; [reflexivity|]. unshelve eapply (PolyRed.posExt PFG wk_id wfΓ); irrelevance. - Qed. - + Qed. Lemma polyRedId {Γ l F G} : PolyRed Γ l F G -> [Γ ||- F] × [Γ ,, F ||- G]. Proof. - intros [?? RF RG]; split. - - rewrite <- (wk_id_ren_on Γ F). eapply RF; gen_typing. - - replace G with G[tRel 0 .: @wk1 Γ F >> tRel]. - 2: bsimpl; rewrite scons_eta'; now asimpl. - eapply RG. eapply var0; tea; now bsimpl. - Unshelve. gen_typing. + intros PFG; assert (wfΓ : [|- Γ]) by (destruct PFG; gen_typing). + split; [exact (instKripke wfΓ (PolyRed.shpRed PFG))| exact (instKripkeSubst wfΓ (PolyRed.posRed PFG))]. Qed. Lemma polyRedEqId {Γ l F F' G G'} (PFG : PolyRed Γ l F G) (RFG := polyRedId PFG) : PolyRedEq PFG F' G' -> [Γ ||- F ≅ F' | fst RFG] × [Γ ,, F ||- G ≅ G' | snd RFG]. Proof. - intros [RFF' RGG']; destruct RFG; escape; split. - - rewrite <- (wk_id_ren_on Γ F'); irrelevance0. - 2: unshelve eapply RFF'; gen_typing. - apply wk_id_ren_on. - - replace G' with G'[tRel 0 .: @wk1 Γ F >> tRel]. - 2: bsimpl; rewrite scons_eta'; now asimpl. - irrelevance0. - 2: eapply RGG'. - bsimpl; rewrite scons_eta'; now asimpl. - Unshelve. 1: gen_typing. - eapply var0; tea; now bsimpl. + intros PFGeq; assert (wfΓ : [|- Γ]) by (destruct PFG; gen_typing). + split. + - pose (instKripkeEq wfΓ (PolyRedEq.shpRed PFGeq)); irrelevance. + - unshelve epose proof (instKripkeSubstEq wfΓ _). + 7: intros; eapply LRTransEq; [|unshelve eapply (PolyRedEq.posRed PFGeq _ _ (urefl hab))]; + try (unshelve eapply (PolyRed.posExt PFG); tea). + irrelevance. Qed. - Lemma polyRedSubst {Γ l A B t} (PAB : PolyRed Γ l A B) - (Rt : forall Δ a (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), - [Δ ||- a : _ | PolyRed.shpRed PAB ρ wfΔ] -> + Lemma polyRedEqCodSubstExt {Γ l F F' G G'} (RF : [Γ ||- F]) (PFG : PolyRed Γ l F G) (RG := polyCodSubstRed RF PFG) : + PolyRedEq PFG F' G' -> + forall t u (Rt : [Γ ||- t ≅ u : _ | RF]), + [Γ ||- G[t..] ≅ G'[u..]| RG t u Rt]. + Proof. + intros Peq *; assert (wfΓ : [|- Γ]) by (destruct PFG; gen_typing). + unshelve epose proof (h := posRedExt Peq wk_id wfΓ _). + 3: irrelevance. + rewrite subst_wk_id_tail in h; irrelevance. + Qed. + + Lemma polyRedSubst {Γ l A B t} (PAB : PolyRed Γ l A B) + (Rt : forall Δ a b (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), + [Δ ||- a ≅ b : _ | PolyRed.shpRed PAB ρ wfΔ] -> [Δ ||- t[a .: ρ >> tRel] : _ | PolyRed.shpRed PAB ρ wfΔ ]) - (Rtext : forall Δ a b (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), + (Rtext : forall Δ a b (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), [Δ ||- a : _ | PolyRed.shpRed PAB ρ wfΔ] -> [Δ ||- b : _ | PolyRed.shpRed PAB ρ wfΔ] -> [Δ ||- a ≅ b : _ | PolyRed.shpRed PAB ρ wfΔ] -> @@ -116,30 +114,25 @@ Section PolyValidity. Proof. destruct PAB; opector; tea. + intros; rewrite liftSubst_scons_eq. - unshelve eapply posRed; tea; eapply Rt; now irrelevanceRefl. - + unshelve epose proof (posRed _ t (@wk1 Γ A) _ _). + unshelve eapply posRed; [..| eapply Rt]; tea ; now irrelevanceRefl. + + unshelve epose proof (posRed _ t t (@wk1 Γ A) _ _). - escape; gen_typing. - replace t with t[tRel 0 .: @wk1 Γ A >> tRel]. 2:{ bsimpl; rewrite scons_eta'; now asimpl. } eapply Rt. eapply var0; tea; now rewrite wk1_ren_on. - - escape. + - escape. replace (B[_]) with B[t .: @wk1 Γ A >> tRel]; tea. now setoid_rewrite wk1_ren. + intros; irrelevance0; rewrite liftSubst_scons_eq;[reflexivity|]. - unshelve eapply posExt; tea; eapply Rt + eapply Rtext; now irrelevanceRefl. + unshelve eapply posExt; tea; eapply Rt + eapply Rtext; cbn; irrelevanceRefl. + 1: now eapply lrefl. + 1: now eapply urefl. + eassumption. Qed. - Lemma polyRedEqSubst {Γ l A B t u} (PAB : PolyRed Γ l A B) - (Rt : forall Δ a (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), - [Δ ||- a : _ | PolyRed.shpRed PAB ρ wfΔ] -> - [Δ ||- t[a .: ρ >> tRel] : _ | PolyRed.shpRed PAB ρ wfΔ ]) - (Ru : forall Δ a (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), - [Δ ||- a : _ | PolyRed.shpRed PAB ρ wfΔ] -> - [Δ ||- u[a .: ρ >> tRel] : _ | PolyRed.shpRed PAB ρ wfΔ ]) - (Rtu : forall Δ a b (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), - [Δ ||- a : _ | PolyRed.shpRed PAB ρ wfΔ] -> - [Δ ||- b : _ | PolyRed.shpRed PAB ρ wfΔ] -> + Lemma polyRedEqSubst {Γ l A B t u} (PAB : PolyRed Γ l A B) + (Rtu : forall Δ a b (ρ : Δ ≤ Γ) (wfΔ : [|-Δ]), [Δ ||- a ≅ b : _ | PolyRed.shpRed PAB ρ wfΔ] -> [Δ ||- t[a .: ρ >> tRel] ≅ u[b .: ρ >> tRel] : _ | PolyRed.shpRed PAB ρ wfΔ ]) (PABt : PolyRed Γ l A B[t]⇑) @@ -149,69 +142,52 @@ Section PolyValidity. - intros; eapply reflLRTyEq. - intros; irrelevance0; rewrite liftSubst_scons_eq; [reflexivity|]. unshelve eapply PolyRed.posExt; cycle 1; tea. - + eapply Rt; now irrelevanceRefl. - + eapply Ru; now irrelevanceRefl. - + eapply Rtu; try eapply reflLRTmEq; now irrelevanceRefl. + eapply Rtu; irrelevanceRefl; now eapply lrefl. Qed. - + Context {l Γ F G} (VΓ : [||-v Γ]) (VF : [Γ ||-v< l > F | VΓ ]) (VG : [Γ ,, F ||-v< l > G | validSnoc VΓ VF]). - Context {Δ σ} (tΔ : [ |-[ ta ] Δ]) (Vσ : [VΓ | Δ ||-v σ : _ | tΔ]). - + Context {Δ σ σ'} (tΔ : [ |-[ ta ] Δ]) (Vσσ' : [VΓ | Δ ||-v σ ≅ σ' : _ | tΔ]). + Lemma substPolyRed : PolyRed Δ l F[σ] G[up_term_term σ]. Proof. - pose proof (Vuσ := liftSubstS' VF Vσ). - instValid Vσ; instValid Vuσ; escape. + pose proof (Vuσ := liftSubstEq' VF Vσσ'). + instValid Vσσ'; instValid Vuσ; escape. unshelve econstructor; tea. - intros; now eapply wk. - - cbn; intros ?? ρ wfΔ' ha. + - cbn; intros ??? ρ wfΔ' hab. rewrite eq_subst_2. - pose proof (Vawkσ := consWkSubstS VF ρ wfΔ' Vσ ha). + pose proof (Vawkσ := consWkSubstEq VF Vσσ' ρ wfΔ' hab). now instValid Vawkσ. - - cbn; intros ??? ρ wfΔ' ha hb hab. - irrelevance0; rewrite eq_subst_2. - 1: reflexivity. - pose proof (Vawkσ := consWkSubstS VF ρ wfΔ' Vσ ha). - pose proof (Vbwkσ := consWkSubstS VF ρ wfΔ' Vσ hb). - epose proof (Vabwkσ := consWkSubstSEq' VF Vσ (reflSubst _ _ Vσ) ρ wfΔ' ha hab). - now instAllValid Vawkσ Vbwkσ Vabwkσ. + - cbn; intros ??? ρ wfΔ' hab. + epose proof (Vabwkσ := consWkSubstEq VF (lrefl Vσσ') ρ wfΔ' hab). + instValid Vabwkσ. + rewrite !eq_subst_2; irrelevance. Qed. Lemma substEqPolyRedEq {F' G'} (VF' : [Γ ||-v< l > F' | VΓ ]) (VG' : [Γ ,, F' ||-v< l > G' | validSnoc VΓ VF']) (VFF' : [Γ ||-v< l > F ≅ F' | VΓ | VF]) (VGG' : [Γ ,, F ||-v< l > G ≅ G' | validSnoc VΓ VF | VG]) - : PolyRedEq substPolyRed F'[σ] G'[up_term_term σ]. + (PFGσ : PolyRed Δ l F[σ] G[up_term_term σ]) + : PolyRedEq PFGσ F'[σ'] G'[up_term_term σ']. Proof. - instValid Vσ. + instValid Vσσ'. unshelve econstructor. - - intros; irrelevanceRefl; now unshelve eapply wkEq. - - intros ?? ρ wfΔ' ha; irrelevance0; rewrite eq_subst_2. - 1: reflexivity. - unshelve epose proof (Vabwkσ := consWkSubstSEq' VF Vσ (reflSubst _ _ Vσ) ρ wfΔ' ha _). - 2: now eapply reflLRTmEq. - unshelve eapply validTyEq; cycle 2; tea. - now eapply consWkSubstS. + - intros; irrelevanceRefl; now unshelve now eapply wkEq. + - intros ??? ρ wfΔ' hab; rewrite eq_subst_2. + unshelve epose proof (Vabwkσ := consWkSubstEq VF Vσσ' ρ wfΔ' (lrefl hab)). + instValid Vabwkσ; irrelevance0; [|tea]. + now rewrite eq_subst_2. Qed. - Context {σ'} (Vσ' : [VΓ | Δ ||-v σ' : _ | tΔ]) (Vσσ' : [VΓ | Δ ||-v σ ≅ σ' : _ | tΔ | Vσ]). - - Lemma substPolyRedEq : PolyRedEq substPolyRed F[σ'] G[up_term_term σ']. + Lemma substPolyRedEq (PFGσ : PolyRed Δ l F[σ] G[up_term_term σ]) + : PolyRedEq PFGσ F[σ'] G[up_term_term σ']. Proof. - instAllValid Vσ Vσ' Vσσ'. - unshelve econstructor. - - intros; irrelevanceRefl; now eapply wkEq. - - intros ?? ρ wfΔ' ha; irrelevance0; rewrite eq_subst_2. - 1: reflexivity. - unshelve epose proof (Vabwkσ := consWkSubstSEq' VF Vσ Vσσ' ρ wfΔ' ha _). - 2: now eapply reflLRTmEq. - eapply validTyExt; tea. - eapply consWkSubstS; tea. - eapply LRTmRedConv; tea. - irrelevanceRefl; now eapply wkEq. - Unshelve. 1,3,5: tea. now eapply wk. + eapply substEqPolyRedEq; tea. + all: eapply reflValidTy. Qed. End PolyValidity. diff --git a/theories/Substitution/Introductions/Sigma.v b/theories/Substitution/Introductions/Sigma.v index cf5c624..cefc9c9 100644 --- a/theories/Substitution/Introductions/Sigma.v +++ b/theories/Substitution/Introductions/Sigma.v @@ -1,52 +1,54 @@ -From Coq Require Import ssrbool. +From Coq Require Import ssrbool CRelationClasses. From LogRel.AutoSubst Require Import core unscoped Ast Extra. -From LogRel Require Import Utils BasicAst Notations Context NormalForms Weakening - GenericTyping LogicalRelation Validity. -From LogRel.LogicalRelation Require Import Escape Reflexivity Neutral Weakening Irrelevance Reduction NormalRed Induction Transitivity. -From LogRel.Substitution Require Import Irrelevance Properties SingleSubst Reflexivity. +From LogRel Require Import Utils BasicAst Notations Context NormalForms Weakening GenericTyping LogicalRelation Validity. +From LogRel.LogicalRelation Require Import Induction Escape Reflexivity Irrelevance Weakening Neutral Transitivity Reduction Application Universe EqRedRight SimpleArr NormalRed InstKripke. +From LogRel.Substitution Require Import Irrelevance Properties Conversion SingleSubst Reflexivity Reduction. From LogRel.Substitution.Introductions Require Import Universe Poly. + Set Universe Polymorphism. Set Printing Primitive Projection Parameters. -Section SigmaValidity. +Section SigmaCongRed. + Context `{GenericTypingProperties}. - Context {l Γ F G} (VΓ : [||-v Γ]) - (VF : [Γ ||-v< l > F | VΓ ]) - (VG : [Γ ,, F ||-v< l > G | validSnoc VΓ VF]). + Context {Γ F G F' G' l} + (VΓ : [||-v Γ]) + (VF : [ Γ ||-v< l > F | VΓ ]) + (VG : [ Γ ,, F ||-v< l > G | validSnoc VΓ VF ]) + (VF' : [ Γ ||-v< l > F' | VΓ ]) + (VG' : [ Γ ,, F' ||-v< l > G' | validSnoc VΓ VF' ]) + (VFF' : [ Γ ||-v< l > F ≅ F' | VΓ | VF ]) + (VGG' : [ Γ ,, F ||-v< l > G ≅ G' | validSnoc VΓ VF | VG ]). - Lemma SigRed {Δ σ} (wfΔ : [ |-[ ta ] Δ]) (Vσ : [VΓ | Δ ||-v σ : _ | wfΔ]) + Lemma SigRed {Δ σ σ'} (wfΔ : [ |-[ ta ] Δ]) (Vσ : [VΓ | Δ ||-v σ ≅ σ' : _ | wfΔ]) : [ Δ ||-< l > (tSig F G)[σ] ]. Proof. eapply LRSig'. pose (p := substPolyRed VΓ VF VG _ Vσ). - escape; cbn; econstructor; tea; - destruct (polyRedId p); - destruct (polyRedEqId p (substPolyRedEq VΓ VF VG _ Vσ Vσ (reflSubst _ _ Vσ))); escape. - - apply redtywf_refl; gen_typing. - - gen_typing. - - gen_typing. + set (pid := polyRedId p). + set (peq := polyRedEqId p (substPolyRedEq VΓ VF VG _ Vσ p)). + econstructor; tea. + - eapply redtywf_refl, wft_sig; refold; destruct pid; now escape. + - eapply lrefl; destruct peq; now escape. + - destruct pid, peq; escape; eapply convty_sig; tea; now eapply lrefl. Defined. - Lemma SigEqRed {Δ σ σ'} (tΔ : [ |-[ ta ] Δ]) - (Vσ : [VΓ | Δ ||-v σ : _ | tΔ]) - (Vσ' : [VΓ | Δ ||-v σ' : _ | tΔ]) - (Vσσ' : [VΓ | Δ ||-v σ ≅ σ' : _ | tΔ | Vσ]) - : [ Δ ||-< l > (tSig F G)[σ] ≅ (tSig F G)[σ'] | SigRed tΔ Vσ ]. + + Lemma SigCongRed {Δ σ σ'} (wfΔ : [|- Δ]) (Vσ : [VΓ | Δ ||-v σ ≅ σ' : _ | wfΔ]) + : [ Δ ||-< l > (tSig F G)[σ] ≅ (tSig F' G')[σ'] | SigRed wfΔ Vσ ]. Proof. pose (p := substPolyRed VΓ VF VG _ Vσ). - pose (p' := substPolyRed VΓ VF VG _ Vσ'). - pose (peq := substPolyRedEq VΓ VF VG _ Vσ Vσ' Vσσ'). + pose (p' := substPolyRed VΓ VF' VG' _ Vσ). + pose (peq := substEqPolyRedEq VΓ VF VG _ Vσ VF' VG' VFF' VGG' p). destruct (polyRedId p); destruct (polyRedId p'); destruct (polyRedEqId p peq). escape; econstructor; cbn; tea. - - apply redtywf_refl; gen_typing. - - gen_typing. - Defined. + - apply redtywf_refl; eapply wft_sig; tea. + eapply escape; now instValid (liftSubstEq' VF' (urefl Vσ)). + - now eapply convty_sig. + Qed. - Lemma SigValid : [Γ ||-v< l > tSig F G | VΓ]. - Proof. unshelve econstructor; intros; [now eapply SigRed| now eapply SigEqRed]. Defined. - -End SigmaValidity. +End SigmaCongRed. Section SigmaCongValidity. @@ -60,24 +62,17 @@ Section SigmaCongValidity. (VFF' : [ Γ ||-v< l > F ≅ F' | VΓ | VF ]) (VGG' : [ Γ ,, F ||-v< l > G ≅ G' | validSnoc VΓ VF | VG ]). - Lemma SigCongRed {Δ σ} (wfΔ : [|- Δ]) (Vσ : [VΓ | Δ ||-v σ : _ | wfΔ]) - : [ Δ ||-< l > (tSig F G)[σ] ≅ (tSig F' G')[σ] | validTy (SigValid VΓ VF VG) wfΔ Vσ ]. + Lemma SigValid : [Γ ||-v< l > tSig F G | VΓ]. Proof. - pose (p := substPolyRed VΓ VF VG _ Vσ). - pose (p' := substPolyRed VΓ VF' VG' _ Vσ). - pose (peq := substEqPolyRedEq VΓ VF VG _ Vσ VF' VG' VFF' VGG'). - destruct (polyRedId p); destruct (polyRedId p'); destruct (polyRedEqId p peq). - escape; econstructor; cbn; tea. - - apply redtywf_refl; gen_typing. - - gen_typing. + unshelve econstructor; intros; [now eapply SigRed|]. + eapply SigCongRed; tea; now eapply reflValidTy. Qed. - Lemma SigCong : [ Γ ||-v< l > tSig F G ≅ tSig F' G' | VΓ | SigValid VΓ VF VG ]. - Proof. econstructor; intros; eapply SigCongRed. Qed. + Lemma SigCong : [ Γ ||-v< l > tSig F G ≅ tSig F' G' | VΓ | SigValid ]. + Proof. econstructor; intros; irrelevanceRefl; now unshelve now eapply SigCongRed. Qed. End SigmaCongValidity. - Lemma up_subst_single {t σ a} : t[up_term_term σ][a..] = t[a .: σ]. Proof. now asimpl. Qed. @@ -91,13 +86,13 @@ Proof. now rewrite wk_id_shift, up_subst_single, scons_eta'. Qed. Section SigInjValid. Context `{GenericTypingProperties}. Context {l Γ F G} (VΓ : [||-v Γ]) (VΣ : [Γ ||-v tSig F G | VΓ]). - + Lemma domSigValid : [Γ ||-v F | VΓ]. Proof. unshelve econstructor. - - intros ??? Vσ; instValid Vσ. + - intros ???? Vσ; instValid Vσ. apply (polyRedId (normRedΣ0 (invLRΣ RVΣ))). - - intros ???? Vσ Vσ' Vσσ' ; instAllValid Vσ Vσ' Vσσ'. + - intros ???? Vσσ' ; instValid Vσσ'. set (P := (polyRedId _)); destruct P. pose (X := normEqRedΣ _ REVΣ); fold subst_term in *. irrelevanceRefl; apply (polyRedEqId _ X). @@ -106,354 +101,309 @@ Section SigInjValid. Lemma codSigValid : [Γ,, F ||-v G | validSnoc VΓ domSigValid ]. Proof. pose (domSigValid). - assert (h : forall (Δ : context) (σ : nat -> term) (wfΔ : [ |-[ ta ] Δ]), - [validSnoc VΓ domSigValid | Δ ||-v σ : Γ,, F | wfΔ] -> [Δ ||-< l > G[σ]]). + assert (h : forall (Δ : context) (wfΔ : [ |-[ ta ] Δ]) (σ σ' : nat -> term), + [validSnoc VΓ domSigValid | Δ ||-v σ ≅ σ' : Γ,, F | wfΔ] -> [Δ ||-< l > G[σ]]). 1:{ - intros ?? wfΔ Vσ; instValid (validTail Vσ). + intros ? wfΔ ?? Vσ; instValid (eqTail Vσ). pose (p := normRedΣ0 (invLRΣ RVΣ)); fold subst_term in *. erewrite split_valid_subst_wk_id. unshelve eapply (PolyRed.posRed p (wk_id) wfΔ). - irrelevance0; [|exact (validHead Vσ)]; now rewrite wk_id_ren_on. + 2: irrelevance0; [|exact (eqHead Vσ)]; now rewrite wk_id_ren_on. } unshelve econstructor; tea. - intros ??? wfΔ Vσ Vσ' Vσσ' ; instAllValid (validTail Vσ) (validTail Vσ') (eqTail Vσσ'). + intros ? wfΔ ?? Vσσ' ; instValid (eqTail Vσσ'). pose (p := normRedΣ0 (invLRΣ RVΣ)); pose (q := normEqRedΣ _ REVΣ); fold subst_term in *. irrelevance0. 1: now erewrite split_valid_subst_wk_id. assert [PolyRed.shpRed p wk_id wfΔ | Δ ||- σ' var_zero : (ParamRedTy.dom p)⟨wk_id⟩]. 1:{ - eapply LRTmRedConv. - 2: exact (validHead Vσ'). + eapply LRTmEqConv. + 2: exact (urefl (eqHead Vσσ')). rewrite wk_id_ren_on; cbn. - now eapply LRTyEqSym. + now eapply reflLRTyEq. } eapply transEq; cycle 1. + erewrite split_valid_subst_wk_id. unshelve eapply (PolyRedEq.posRed q wk_id wfΔ). - irrelevance. + 2: irrelevance. + unshelve eapply (PolyRed.posExt p wk_id wfΔ); tea. - 1: irrelevance0; [|exact (validHead Vσ)]; now rewrite wk_id_ren_on. irrelevance0; [|exact (eqHead Vσσ')]; now rewrite wk_id_ren_on. Qed. - + End SigInjValid. Section SigTmValidity. - Context `{GenericTypingProperties}. - Context {Γ F G} (VΓ : [||-v Γ]) - (VF : [ Γ ||-v< one > F | VΓ ]) + + Section Lemmas. + Context {Γ F F' G G'} {VΓ : [||-v Γ]} + {VF : [ Γ ||-v< one > F | VΓ ]} (VU : [ Γ ,, F ||-v< one > U | validSnoc VΓ VF ]) - (VFU : [ Γ ||-v< one > F : U | VΓ | UValid VΓ ]) - (VGU : [ Γ ,, F ||-v< one > G : U | validSnoc VΓ VF | VU ]) . + (VFeqU : [ Γ ||-v< one > F ≅ F' : U | VΓ | UValid VΓ ]) + (VGeqU : [ Γ ,, F ||-v< one > G ≅ G' : U | validSnoc VΓ VF | VU ]). + Lemma sigTmEq {Δ σ σ'} (tΔ : [ |-[ ta ] Δ]) - (Vσ : [VΓ | Δ ||-v σ : _ | tΔ]) - (Vσ' : [VΓ | Δ ||-v σ' : _ | tΔ]) - (Vσσ' : [VΓ | Δ ||-v σ ≅ σ' : _ | tΔ | Vσ ]) - : [Δ |-[ ta ] tSig F[σ] G[up_term_term σ] ≅ tSig F[σ'] G[up_term_term σ'] : U]. - Proof. - pose proof (Vuσ := liftSubstS' VF Vσ). - pose proof (Vureaσ' := liftSubstSrealign' VF Vσσ' Vσ'). - pose proof (Vuσσ' := liftSubstSEq' VF Vσσ'). - instAllValid Vσ Vσ' Vσσ'; instAllValid Vuσ Vureaσ' Vuσσ'; escape. + (Vσσ' : [VΓ | Δ ||-v σ ≅ σ' : _ | tΔ ]) + : [Δ |-[ ta ] tSig F[σ] G[up_term_term σ] ≅ tSig F'[σ'] G'[up_term_term σ'] : U]. + Proof. + pose proof (Vuσ := liftSubstEq' VF Vσσ'). + instValid Vσσ'; instValid Vuσ; escape. eapply convtm_sig; tea. Qed. - Lemma SigRedU {Δ σ} (tΔ : [ |-[ ta ] Δ]) (Vσ : [VΓ | Δ ||-v σ : _ | tΔ]) - : [ Δ ||-< one > (tSig F G)[σ] : U | validTy (UValid VΓ) tΔ Vσ ]. + + Lemma SigURedTm {Δ σ σ'} (tΔ : [ |-[ ta ] Δ]) (Vσ : [VΓ | Δ ||-v σ ≅ σ' : _ | tΔ]) + : URedTm (LogRelRec _) Δ (tSig F G)[σ] U (redUOneCtx tΔ). Proof. - pose proof (Vσσ := reflSubst _ _ Vσ). - pose proof (Vuσ := liftSubstS' VF Vσ). - pose proof (Vuσσ := liftSubstSEq' VF Vσσ). - instAllValid Vσ Vσ Vσσ; instAllValid Vuσ Vuσ Vuσσ; escape. - econstructor; cbn. - - apply redtmwf_refl; cbn in *; now eapply ty_sig. - - constructor. - - now eapply convtm_sig. - - cbn. unshelve refine (LRCumulative (SigRed _ _ _ tΔ Vσ)); - unshelve eapply univValid; tea; try irrValid. + pose proof (Vuσσ := liftSubstEq' VF Vσ); instValid Vσ ; instValid Vuσσ; escape. + econstructor; [eapply redtmwf_refl; cbn; now eapply ty_sig|constructor]. Defined. - - Lemma SigValidU : [ Γ ||-v< one > tSig F G : U | VΓ | UValid VΓ ]. - Proof. - econstructor. - - intros Δ σ tΔ Vσ. - exact (SigRedU tΔ Vσ). - - intros Δ σ σ' tΔ Vσ Vσ' Vσσ'. - pose proof (univValid (l' := zero) _ _ VFU) as VF0. - pose proof (irrelevanceTy (validSnoc VΓ VF) - (validSnoc (l := zero) VΓ VF0) - (univValid (l' := zero) _ _ VGU)) as VG0. - unshelve econstructor ; cbn. - + exact (SigRedU tΔ Vσ). - + exact (SigRedU tΔ Vσ'). - + exact (LRCumulative (SigRed VΓ VF0 VG0 tΔ Vσ)). - + exact (sigTmEq tΔ Vσ Vσ' Vσσ'). - + exact (LRCumulative (SigRed VΓ VF0 VG0 tΔ Vσ')). - + pose proof (SigEqRed VΓ VF0 VG0 tΔ Vσ Vσ' Vσσ'). - irrelevanceCum. + End Lemmas. + + Context {Γ F F' G G'} {VΓ : [||-v Γ]} + {VF : [ Γ ||-v< one > F | VΓ ]} + {VU : [ Γ ,, F ||-v< one > U | validSnoc VΓ VF ]} + (VFeqU : [ Γ ||-v< one > F ≅ F' : U | VΓ | UValid VΓ ]) + (VFU := lreflValidTm _ VFeqU) + (VGeqU : [ Γ ,, F ||-v< one > G ≅ G' : U | validSnoc VΓ VF | VU ]) + (VGU := lreflValidTm _ VGeqU). + + Lemma SigCongRedU {Δ σ σ'} (tΔ : [ |-[ ta ] Δ]) (Vσ : [VΓ | Δ ||-v σ ≅ σ' : _ | tΔ]) + : [ Δ ||-< one > (tSig F G)[σ] ≅ (tSig F' G')[σ'] : U | validTy (UValid VΓ) tΔ Vσ ]. + Proof. + pose proof (Vuσσ := liftSubstEq' VF Vσ); instValid Vσ ; instValid Vuσσ; escape. + set (RΣ := SigURedTm VU VFeqU VGeqU tΔ Vσ). + pose proof (symValidTmEq VFeqU). + epose proof (VF' := univValid _ _ (ureflValidTm VFeqU)). + epose proof (VFF' := univEqValid _ _ VF VFeqU). + epose proof (VGU' := convTmEqCtx1 _ _ (validSnoc VΓ VF') VF _ (UValid _) VFF' (ureflValidTm VGeqU)). + set (RΣ' := SigURedTm (UValid _) (ureflValidTm VFeqU) VGU' tΔ (urefl Vσ)). + unshelve eexists RΣ RΣ' _. + - unshelve (eapply RedTyRecBwd, LRCumulative, SigRed; [| tea]). + all: unshelve (eapply univValid, lreflValidTm; irrValid); eapply UValid. + - eapply convtm_sig; refold; tea; now eapply lrefl. + - unshelve (eapply RedTyRecBwd, LRCumulative, SigRed; [| now eapply urefl]). + all: unshelve (eapply univValid, lreflValidTm; irrValid); eapply UValid. + - irrelevanceCum0; [reflexivity|]. + unshelve (eapply SigCongRed; tea; [now eapply univValid| now eapply univEqValid]); tea. + unshelve (eapply univValid, lreflValidTm; irrValid); eapply UValid. Qed. + Lemma SigCongValidU : [ Γ ||-v< one > tSig F G ≅ tSig F' G' : U | VΓ | UValid VΓ ]. + Proof. econstructor; intros Δ tΔ σ σ' Vσσ'; eapply SigCongRedU. Qed. + End SigTmValidity. -Section SigTmCongruence. +Section SigRedTmEqHelper. + Import SigRedTmEq. + Context `{GenericTypingProperties}. + + Lemma isLRPair_isWfPair {Γ A B l p} (ΣA : [Γ ||-Σ tSig A B]) + (RA : [Γ ||- A]) + (RΣ := (normRedΣ0 ΣA)) + (Rp : isLRPair RΣ p) : + isWfPair Γ A B p. + Proof. + assert (wfΓ: [|- Γ]) by (escape ; gen_typing). + destruct Rp as [???? eqA eqB rfst rsnd|]. + 2: now econstructor. + pose proof (RFA := instKripkeEq wfΓ eqA). + pose proof (LRTyEqRedRight _ RFA). + pose proof (instKripkeSubstConv wfΓ eqA (PolyRed.posRed RΣ)). + pose proof (instKripkeSubstConvEq wfΓ eqA eqB). + pose proof (Ra := instKripkeTmEq wfΓ rfst). + pose proof (instKripkeSubstEq wfΓ eqB). + pose proof (polyCodSubstRed _ RΣ _ _ Ra). + unshelve epose proof (hB:=eqB _ _ _ (@wk_id Γ) wfΓ _). + 3: irrelevance0; [| exact Ra]; now rewrite wk_id_ren_on. + pose proof (polyCodSubstExtRed _ RΣ _ _ Ra). + epose proof (hb := rsnd _ wk_id wfΓ). + cbn -[wk_id] in *. + escape. + rewrite 2!subst_wk_id_tail in EschB. + rewrite subst_wk_id_tail in EscRhB, Rlhb. + rewrite 2!wk_id_ren_on in Rlhb. + now econstructor. + Qed. + + + Context {Γ l A} (RA : [Γ ||-Σ A]) + {t u} (Rt : SigRedTm RA t) (Ru : SigRedTm RA u). + + (* Let Rout : [Γ ||-Σ RA.(ParamRedTy.outTy)]. + Proof. + apply normRedΣ0. + econstructor; [eapply redtywf_refl|..]; destruct RA; tea. + cbn in *; gen_typing. + Defined. *) + + Lemma build_sigRedTmEq + (eqnf : [Γ |- nf Rt ≅ nf Ru : ParamRedTy.outTy RA]) + (Rdom := fst (polyRedId RA)) + (Rfst : [ Rdom | _ ||- tFst (nf Rt) ≅ tFst (nf Ru) : _ ]) + (Rcod := polyCodSubstRed Rdom RA _ _ Rfst) + (Rsnd : [ Rcod | _ ||- tSnd (nf Rt) ≅ tSnd (nf Ru) : _ ]) : + [LRSig' RA | _ ||- t ≅ u : _]. + Proof. + unshelve eexists Rt Ru _; tea. + - intros; unshelve (irrelevanceRefl; rewrite 2!wk_fst; now eapply wkTermEq); tea. + - intros; irrelevance0. + 2: unshelve (rewrite 2!wk_snd; now eapply wkTermEq); tea. + rewrite wk_fst; clear; now bsimpl. + Qed. + + Lemma redtmwf_fst {F G f f'} : + [ Γ |- f :⤳*: f' : tSig F G ] -> + [ Γ |- tFst f :⤳*: tFst f' : F ]. + Proof. + intros [] ; constructor; [| now eapply redtm_fst]. + timeout 1 gen_typing. + Qed. + + Lemma redtmwf_snd {F G f f'} : + [ Γ |- f :⤳*: f' : tSig F G ] -> + [ Γ |- G[(tFst f)..] ≅ G[(tFst f')..]] -> + [ Γ |- tSnd f :⤳*: tSnd f' : G[(tFst f)..] ]. + Proof. + intros [] ? ; constructor; [| now eapply redtm_snd]. + eapply ty_conv; [| now symmetry]; timeout 1 gen_typing. + Qed. + + (* Lemma build_sigRedTmEq' + (eqnf : [Γ |- nf Rt ≅ nf Ru : ParamRedTy.outTy RA]) + (Rdom := fst (polyRedId RA)) + (Rfst : [ Rdom | _ ||- tFst t ≅ tFst u : _ ]) + (Rcod := polyCodSubstRed Rdom RA _ _ Rfst) + (Rsnd : [ Rcod | _ ||- tSnd t ≅ tSnd u : _ ]) : + [LRSig' RA | _ ||- t ≅ u : _]. + Proof. + unshelve eapply build_sigRedTmEq; tea. + - dependent inversion Rt; dependent inversion Ru; eapply redTmEqFwdBoth. + 2,3: cbn in *; now eapply redtmwf_fst. + 1: tea. + cbn; constructor. + 3:{ cbn in *. tea. } + try solve [constructor]. *) + +End SigRedTmEqHelper. +Section SigRedTmEqHelper. Context `{GenericTypingProperties}. - Context {Γ F G F' G'} - (VΓ : [||-v Γ]) - (VF : [ Γ ||-v< one > F | VΓ ]) - (VU : [ Γ ,, F ||-v< one > U | validSnoc VΓ VF ]) - (VFU : [ Γ ||-v< one > F : U | VΓ | UValid VΓ ]) - (VGU : [ Γ ,, F ||-v< one > G : U | validSnoc VΓ VF | VU ]) - (VF' : [ Γ ||-v< one > F' | VΓ ]) - (VU' : [ Γ ,, F' ||-v< one > U | validSnoc VΓ VF' ]) - (VF'U : [ Γ ||-v< one > F' : U | VΓ | UValid VΓ ]) - (VG'U : [ Γ ,, F' ||-v< one > G' : U | validSnoc VΓ VF' | VU' ]) - (VFF' : [ Γ ||-v< one > F ≅ F' : U | VΓ | UValid VΓ ]) - (VGG' : [ Γ ,, F ||-v< one > G ≅ G' : U | validSnoc VΓ VF | VU ]). - - Lemma SigCongTm : [ Γ ||-v< one > tSig F G ≅ tSig F' G' : U | VΓ | UValid VΓ ]. - Proof. - econstructor. - intros Δ σ tΔ Vσ. - pose proof (univValid (l' := zero) _ _ VFU) as VF0. - pose proof (univValid (l' := zero) _ _ VF'U) as VF'0. - pose proof (Vσσ := reflSubst _ _ Vσ). - pose proof (Vuσ := liftSubstS' VF Vσ). - pose proof (Vuσσ := liftSubstSEq' VF Vσσ). - instAllValid Vσ Vσ Vσσ; instAllValid Vuσ Vuσ Vuσσ; escape. - pose proof (irrelevanceTy (validSnoc VΓ VF) - (validSnoc (l := zero) VΓ VF0) - (univValid (l' := zero) _ _ VGU)) as VG0. - pose proof (irrelevanceTy (validSnoc VΓ VF') - (validSnoc (l := zero) VΓ VF'0) - (univValid (l' := zero) _ _ VG'U)) as VG'0. - unshelve econstructor ; cbn. - - exact (SigRedU VΓ VF VU VFU VGU tΔ Vσ). - - exact (SigRedU VΓ VF' VU' VF'U VG'U tΔ Vσ). - - exact (LRCumulative (SigRed VΓ VF0 VG0 tΔ Vσ)). - - now eapply convtm_sig. - - exact (LRCumulative (SigRed VΓ VF'0 VG'0 tΔ Vσ)). - - enough ([ Δ ||-< zero > (tSig F G)[σ] ≅ (tSig F' G')[σ] | SigRed VΓ VF0 VG0 tΔ Vσ]) by irrelevanceCum. - refine (SigCongRed VΓ VF0 VG0 VF'0 VG'0 _ _ tΔ Vσ). - + exact (univEqValid VΓ (UValid VΓ) VF0 VFF'). - + pose proof (univEqValid (validSnoc VΓ VF) VU (univValid (l' := zero) _ _ VGU) VGG') as VGG'0. - refine (irrelevanceTyEq _ _ _ _ VGG'0). - Qed. - -End SigTmCongruence. + Import SigRedTmEq. + + Lemma build_sigRedTmEq' {Γ l F G} + (RΣ0 : [Γ ||-Σ tSig F G]) + (RΣ := normRedΣ0 RΣ0) + {t u} (Rt : SigRedTm RΣ t) (Ru : SigRedTm RΣ u) + (Rdom := fst (polyRedId RΣ)) + (Rfst : [ Rdom | _ ||- tFst (nf Rt) ≅ tFst (nf Ru) : _ ]) + (Rcod := polyCodSubstRed Rdom RΣ _ _ Rfst) + (Rsnd : [ Rcod | _ ||- tSnd (nf Rt) ≅ tSnd (nf Ru) : _ ]) : + [LRSig' RΣ | _ ||- t ≅ u : _]. + Proof. + eapply build_sigRedTmEq; tea. + cbn; destruct (polyRedId RΣ); escape; eapply convtm_eta_sig; tea. + 2,4: eapply isLRPair_isWfPair; [| eapply ispair]; tea. + + destruct Rt; cbn in *; gen_typing. + + destruct Ru; cbn in *; gen_typing. + Qed. + +End SigRedTmEqHelper. Section ProjRed. + Import SigRedTmEq. Context `{GenericTypingProperties}. - Lemma fstRed {l Δ F G} {p} + (* Lemma redSigRedTm {l Δ F G} {p p'} (RΣ : [Δ ||-Σ tSig F G]) - (RF : [Δ ||- F]) - (Rp : [Δ ||- p : _ | LRSig' (normRedΣ0 RΣ)]) : - ([Δ ||- tFst p : _ | RF] × [Δ ||- tFst p ≅ tFst Rp.(SigRedTm.nf) : _ | RF]) × [Δ ||- tFst Rp.(SigRedTm.nf) : _ | RF]. + (Rp : SigRedTm ΣA p) : + [Δ ||- p ≅ p' : _ | LRSig' (normRedΣ0 RΣ)] -> + [Δ ||- p ≅ : _ | LRSig' (normRedΣ0 RΣ)] -> + +SigRedTmEq.whnf *) + + Lemma sigRedTmEqNf {l Δ F G p p'} + (RΣ : [Δ ||-Σ tSig F G]) + (Rp : [Δ ||- p ≅ p' : _ | LRSig' (normRedΣ0 RΣ)]) : + [Δ ||- p ≅ Rp.(redL).(nf) : _ | LRSig' (normRedΣ0 RΣ)]. Proof. - assert [Δ ||- tFst Rp.(SigRedTm.nf) : _ | RF]. - 1:{ - assert (wfΔ : [|-Δ]) by (escape; gen_typing). - pose (r := SigRedTm.fstRed Rp (@wk_id Δ) wfΔ). - rewrite wk_id_ren_on in r. - irrelevance. - } - split; tea. - irrelevanceRefl. - eapply redSubstTerm; tea. - eapply redtm_fst; now destruct (SigRedTm.red Rp). + symmetry; eapply redTmEqFwd. + + eapply lrefl; irrelevance. + + eapply Rp.(redL).(red). + + eapply whnf. Qed. - Lemma fstRedEq {l Δ F G} {p p'} - (RΣ : [Δ ||-Σ tSig F G]) - (RF : [Δ ||- F]) - (RΣn := LRSig' (normRedΣ0 RΣ)) - (Rpp' : [Δ ||- p ≅ p' : _ | RΣn]) - (Rp := SigRedTmEq.redL Rpp') - (Rp' := SigRedTmEq.redR Rpp') : - [× [Δ ||- tFst p ≅ tFst Rp.(SigRedTm.nf) : _ | RF], - [Δ ||- tFst p' ≅ tFst Rp'.(SigRedTm.nf) : _ | RF] & - [Δ ||- tFst Rp.(SigRedTm.nf) ≅ tFst Rp'.(SigRedTm.nf) : _ | RF]]. - Proof. - split. - - now eapply fstRed. - - now eapply fstRed. - - assert (wfΔ : [|-Δ]) by (escape; gen_typing). - epose (e := SigRedTmEq.eqFst Rpp' wk_id wfΔ). - do 2 rewrite wk_id_ren_on in e. - irrelevance. - Qed. - - - Lemma sndRed {l Δ F G} {p} + + Lemma fstRed {l Δ F G p p'} (RΣ : [Δ ||-Σ tSig F G]) (RF : [Δ ||- F]) - (RΣn := LRSig' (normRedΣ0 RΣ)) - (Rp : [Δ ||- p : _ | LRSig' (normRedΣ0 RΣ)]) - (RGfstp : [Δ ||- G[(tFst p)..]]) - (RGfstpEq : [Δ ||- G[(tFst p)..] ≅ G[(tFst Rp.(SigRedTm.nf))..] | RGfstp]) : - [Δ ||- tSnd p : _ | RGfstp] × [Δ ||- tSnd p ≅ tSnd Rp.(SigRedTm.nf) : _ | RGfstp]. - Proof. - eapply redSubstTerm. - 2: eapply redtm_snd; now destruct (SigRedTm.red Rp). - assert (wfΔ : [|-Δ]) by (escape; gen_typing). - eapply LRTmRedConv; cycle 1. - + pose proof (r := SigRedTm.sndRed Rp (@wk_id Δ) wfΔ). - rewrite <- (wk_id_ren_on Δ (SigRedTm.nf Rp)). - eassumption. - + unshelve eapply LRTyEqSym. 2: tea. - rewrite wk_id_shift; rewrite wk_id_ren_on; tea. - Qed. - - Lemma sndRedEq {l Δ F G} {p p'} + (Rp : [Δ ||- p ≅ p' : _ | LRSig' (normRedΣ0 RΣ)]) : + [Δ ||- tFst p ≅ tFst p' : _ | RF]. + Proof. + eapply redSubstTmEq; cycle 1. + + eapply redtm_fst, tmr_wf_red; exact (SigRedTmEq.red (SigRedTmEq.redL Rp)). + + eapply redtm_fst, tmr_wf_red; exact (SigRedTmEq.red (SigRedTmEq.redR Rp)). + + cbn; now unshelve eapply escapeEq, reflLRTyEq. + + irrelevanceRefl; eapply instKripkeTmEq. + intros ? ρ h; rewrite <- 2!wk_fst; eapply (SigRedTmEq.eqFst Rp). + Unshelve. all: escape; gen_typing. + Qed. + + Lemma sndRed {l Δ F G} {p p'} (RΣ : [Δ ||-Σ tSig F G]) (RF : [Δ ||- F]) (RΣn := LRSig' (normRedΣ0 RΣ)) - (Rpp' : [Δ ||- p ≅ p' : _ | RΣn]) - (Rp := SigRedTmEq.redL Rpp') - (Rp' := SigRedTmEq.redR Rpp') - (RGfstp : [Δ ||- G[(tFst p)..]]) - (RGfstpEq : [Δ ||- G[(tFst p)..] ≅ G[(tFst Rp.(SigRedTm.nf))..] | RGfstp]) - (RGfstp' : [Δ ||- G[(tFst p')..]]) - (RGfstpEq' : [Δ ||- G[(tFst p')..] ≅ G[(tFst Rp'.(SigRedTm.nf))..] | RGfstp']) : - [× [Δ ||- tSnd p ≅ tSnd Rp.(SigRedTm.nf) : _ | RGfstp], - [Δ ||- tSnd p' ≅ tSnd Rp'.(SigRedTm.nf) : _ | RGfstp'] & - [Δ ||- tSnd Rp.(SigRedTm.nf) ≅ tSnd Rp'.(SigRedTm.nf) : _ | RGfstp]]. - Proof. - split. - - now eapply sndRed. - - now eapply sndRed. - - assert (wfΔ : [|-Δ]) by (escape; gen_typing). - eapply LRTmEqRedConv; cycle 1. - + epose proof (e := SigRedTmEq.eqSnd Rpp' wk_id wfΔ). - rewrite <- (wk_id_ren_on Δ (SigRedTm.nf Rp)). - rewrite <- (wk_id_ren_on Δ (SigRedTm.nf Rp')). - tea. - + unshelve eapply LRTyEqSym. 2: tea. - rewrite wk_id_shift; rewrite wk_id_ren_on; tea. + (Rp : [Δ ||- p ≅ p' : _ | LRSig' (normRedΣ0 RΣ)]) + (RGfstp : [Δ ||- G[(tFst p)..]]) : + [Δ ||- tSnd p ≅ tSnd p' : _ | RGfstp]. + Proof. + eapply redSubstTmEq; cycle 1. + + eapply redtm_snd, tmr_wf_red; exact (SigRedTmEq.red (SigRedTmEq.redL Rp)). + + eapply redtm_snd, tmr_wf_red; exact (SigRedTmEq.red (SigRedTmEq.redR Rp)). + + epose proof (singleSubstEqΣ (fstRed RΣ RF Rp) RGfstp); cbn; now escape. + + irrelevanceRefl. + assert (wfΔ : [|- Δ]) by (escape; gen_typing). + pose proof (fstRed RΣ RF (sigRedTmEqNf RΣ Rp)). + erewrite <- wk_id_ren_on, <- (wk_id_ren_on _ (tSnd (SigRedTmEq.nf (SigRedTmEq.redL _)))), <-2!wk_snd. + eapply LRTmEqConv; [| eapply (SigRedTmEq.eqSnd Rp wk_id wfΔ)]. + symmetry in X. + cbn -[wk_id]; irrelevance0. + 2: exact (singleSubstEqΣ (RFG:=LRSig' RΣ) X (singleSubstΣ1 (LRSig' RΣ) X)). + now rewrite subst_wk_id_tail, wk_id_ren_on. + Unshelve. 2,4: tea. Qed. - Context {l Γ F G } (VΓ : [||-v Γ]) (VF : [Γ ||-v< l > F | VΓ ]) (VG : [Γ ,, F ||-v< l > G | validSnoc VΓ VF]) (VΣ := SigValid VΓ VF VG). - Lemma fstValid {p} (Vp : [Γ ||-v p : _ | VΓ | VΣ]): [Γ ||-v tFst p : _ | VΓ | VF]. - Proof. - unshelve econstructor. - - intros ?? wfΔ Vσ. - instValid Vσ. - pose proof (invLRΣ RVΣ); refold. - unshelve eapply fstRed. 2: tea. irrelevance. - - intros ??? wfΔ Vσ Vσ' Vσσ'. - instAllValid Vσ Vσ' Vσσ'. - pose (RΣinv := invLRΣ RVΣ); normRedΣin REVp; fold subst_term in *. - destruct (fstRedEq RΣinv RVF REVp). - eapply LREqTermHelper; tea; eapply reflLRTyEq. - Qed. - - Lemma fstCongValid {p p'} - (Vp : [Γ ||-v p : _ | VΓ | VΣ]) - (Vp' : [Γ ||-v p' : _ | VΓ | VΣ]) - (Vpp' : [Γ ||-v p ≅ p' : _ | VΓ | VΣ]) - : [Γ ||-v tFst p ≅ tFst p' : _ | VΓ | VF]. - Proof. - constructor; intros ?? wfΔ Vσ. - instValid Vσ; instValidExt Vσ (reflSubst _ _ Vσ). - pose (RΣinv := invLRΣ RVΣ); normRedΣin RVpp'; fold subst_term in *. - destruct (fstRedEq RΣinv RVF RVpp'). - eapply LREqTermHelper; tea. - Qed. - - Lemma sndValid {p} (Vp : [Γ ||-v p : _ | VΓ | VΣ]) - (VGfst := substS VG (fstValid Vp)) : - [Γ ||-v tSnd p : _ | VΓ | VGfst]. + Lemma fstValid {p p'} (Vp : [Γ ||-v p ≅ p' : _ | VΓ | VΣ]): + [Γ ||-v tFst p ≅ tFst p' : _ | VΓ | VF]. Proof. - unshelve econstructor. - - intros ?? wfΔ Vσ. - instValid Vσ. - pose (RΣinv := invLRΣ RVΣ); normRedΣin RVp; fold subst_term in *. - destruct (fstRed RΣinv RVF RVp) as [[Rfstp Rfstpeq] Rfstnf]. - pose (Vpσ := consSubstS _ _ Vσ VF Rfstp). - pose (Vnfσ := consSubstS _ _ Vσ VF Rfstnf). - pose (Veqσ := consSubstSEq' _ _ Vσ (reflSubst _ _ Vσ) VF Rfstp Rfstpeq). - instAllValid Vpσ Vnfσ Veqσ. - unshelve epose (fst (sndRed RΣinv RVF RVp _ _)). - + now rewrite up_subst_single. - + irrelevance0; rewrite up_subst_single; tea; reflexivity. - + irrelevance. - - intros ??? wfΔ Vσ Vσ' Vσσ'. - pose proof (Vfstp := fstValid Vp). - instAllValid Vσ Vσ' Vσσ'. - pose (RΣinv := invLRΣ RVΣ); pose (RΣinv' := invLRΣ RVΣ0). - normRedΣin RVp; normRedΣin RVp0; normRedΣin REVp; fold subst_term in *. - destruct (fstRed RΣinv RVF REVp.(SigRedTmEq.redL)) as [[Rfstp Rfstpeq] Rfstnf]. - pose (Vpσ := consSubstS _ _ Vσ VF Rfstp). - pose (Vnfσ := consSubstS _ _ Vσ VF Rfstnf). - pose (Veqσ := consSubstSEq' _ _ Vσ (reflSubst _ _ Vσ) VF Rfstp Rfstpeq). - instAllValid Vpσ Vnfσ Veqσ. - destruct (fstRed RΣinv RVF REVp.(SigRedTmEq.redR)) as [[Rfstp' Rfstpeq'] Rfstnf']. - pose (Vpσ' := consSubstS _ _ Vσ VF Rfstp'). - pose (Vnfσ' := consSubstS _ _ Vσ VF Rfstnf'). - pose (Veqσ' := consSubstSEq' _ _ Vσ (reflSubst _ _ Vσ) VF Rfstp' Rfstpeq'). - instAllValid Vpσ' Vnfσ' Veqσ'. - unshelve epose proof(r := sndRedEq RΣinv RVF REVp _ _ _ _). - + now rewrite up_subst_single. - + irrelevance0; rewrite up_subst_single; tea; reflexivity. - + now rewrite up_subst_single. - + irrelevance0; rewrite up_subst_single; tea; reflexivity. - + destruct r; irrelevance0; cycle 1. - 1: eapply LREqTermHelper. - 1,2,4: tea. - 2: now rewrite singleSubstComm'. - epose (Veqσ0 := consSubstSEq' _ _ Vσ (reflSubst _ _ Vσ) VF Rfstp REVfstp). - instAllValid Vpσ Vpσ' Veqσ0. - irrelevance0; rewrite up_subst_single; tea; reflexivity. - Qed. - - Lemma sndCongValid {p p'} - (Vp : [Γ ||-v p : _ | VΓ | VΣ]) - (Vp' : [Γ ||-v p' : _ | VΓ | VΣ]) - (Vpp' : [Γ ||-v p ≅ p' : _ | VΓ | VΣ]) - (VGfst := substS VG (fstValid Vp)) : + constructor; intros; cbn; instValid Vσσ'. + unshelve (eapply fstRed; irrelevance); now apply invLRΣ. + Qed. + + Lemma subst_fst {t σ} : tFst t[σ] = (tFst t)[σ]. + Proof. reflexivity. Qed. + + Lemma sndValid {p p'} (Vp : [Γ ||-v p ≅ p' : _ | VΓ | VΣ]) + (VGfst := substS VG (fstValid Vp)) : [Γ ||-v tSnd p ≅ tSnd p' : _ | VΓ | VGfst]. Proof. - constructor; intros ?? wfΔ Vσ. - pose proof (VGfsteq:= substSEq (reflValidTy _ VF) VG (reflValidTy _ VG) (fstValid Vp) (fstValid Vp') (fstCongValid Vp Vp' Vpp')). - cbn in VGfsteq. - instValid Vσ ; instValidExt Vσ (reflSubst _ _ Vσ). - pose (RΣinv := invLRΣ RVΣ); normRedΣin RVpp'; fold subst_term in *. - destruct (fstRed RΣinv RVF RVpp'.(SigRedTmEq.redL)) as [[Rfstp Rfstpeq] Rfstnf]. - pose (Vpσ := consSubstS _ _ Vσ VF Rfstp). - pose (Vnfσ := consSubstS _ _ Vσ VF Rfstnf). - pose (Veqσ := consSubstSEq' _ _ Vσ (reflSubst _ _ Vσ) VF Rfstp Rfstpeq). - instAllValid Vpσ Vnfσ Veqσ. - destruct (fstRed RΣinv RVF RVpp'.(SigRedTmEq.redR)) as [[Rfstp' Rfstpeq'] Rfstnf']. - pose (Vpσ' := consSubstS _ _ Vσ VF Rfstp'). - pose (Vnfσ' := consSubstS _ _ Vσ VF Rfstnf'). - pose (Veqσ' := consSubstSEq' _ _ Vσ (reflSubst _ _ Vσ) VF Rfstp' Rfstpeq'). - instAllValid Vpσ' Vnfσ' Veqσ'. - unshelve epose proof(r := sndRedEq RΣinv RVF RVpp' _ _ _ _). - + now rewrite up_subst_single. - + irrelevance0; rewrite up_subst_single; tea; reflexivity. - + now rewrite up_subst_single. - + irrelevance0; rewrite up_subst_single; tea; reflexivity. - + destruct r; irrelevance0; cycle 1. - 1: eapply LREqTermHelper. - 1,2,4: tea. - 2: now rewrite singleSubstComm'. - irrelevance0; rewrite up_subst_single; change (tFst ?p[?σ]) with (tFst p)[σ]; rewrite <- singleSubstComm; tea. - reflexivity. - Qed. - - + constructor; intros; cbn; instValid Vσσ'. + unshelve (irrelevance0; [|eapply sndRed; [| irrelevance]; tea]). + all:refold. 3: now rewrite singleSubstComm'. + 2: now apply invLRΣ. + now rewrite subst_fst, <- singleSubstComm'. + Qed. + End ProjRed. @@ -461,300 +411,267 @@ End ProjRed. Section PairRed. Context `{GenericTypingProperties}. - Lemma pairFstRed {Γ A B a b l} + Lemma pairFstRed {Γ A A' B B' a a' b b' l} (RA : [Γ ||- A]) + (RAA : [Γ ||- A ≅ A' | RA]) (RB : [Γ ,, A ||- B]) + (RB' : [Γ ,, A' ||- B']) (RBa : [Γ ||- B[a..]]) - (Ra : [Γ ||- a : A | RA]) - (Rb : [Γ ||- b : _ | RBa ]) : - [Γ ||- tFst (tPair A B a b) : _ | RA] × [Γ ||- tFst (tPair A B a b) ≅ a : _ | RA]. + (RBB : [Γ ||- B[a..] ≅ B'[a'..] | RBa ]) + (Ra : [Γ ||- a ≅ a' : A | RA]) + (Rb : [Γ ||- b ≅ b' : _ | RBa ]) : + [Γ ||- tFst (tPair A B a b) ≅ tFst (tPair A' B' a' b') : _ | RA]. Proof. + pose proof (LRTyEqRedRight RA RAA). escape. - eapply redSubstTerm; tea. - eapply redtm_fst_beta; tea. + eapply redSubstTmEq; tea. + 1,2: eapply redtm_fst_beta; tea; now eapply ty_conv. Qed. - Lemma pairSndRed {Γ A B a b l} + Lemma pairSndRed {Γ A A' B B' a a' b b' l} (RA : [Γ ||- A]) + (RAA : [Γ ||- A ≅ A' | RA]) (RB : [Γ ,, A ||- B]) - (RBa : [Γ ||- B[a..]]) + (RB' : [Γ ,, A' ||- B']) (RBfst : [Γ ||- B[(tFst (tPair A B a b))..] ]) - (RBconv : [Γ ||- B[a..] ≅ B[(tFst (tPair A B a b))..] | RBa ]) + (RBBfst : [Γ ||- B[(tFst (tPair A B a b))..] ≅ B'[(tFst (tPair A' B' a' b'))..] | RBfst]) + (RBa : [Γ ||- B[a..]]) + (RBfsta : [Γ ||- B[a..] ≅ B[(tFst (tPair A B a b))..] | RBa]) + (RBB : [Γ ||- B[a..] ≅ B'[a'..] | RBa ]) + (Ra : [Γ ||- a ≅ a' : A | RA]) + (Rb : [Γ ||- b ≅ b' : _ | RBa ]) : + [Γ ||- tSnd (tPair A B a b) ≅ tSnd (tPair A' B' a' b') : _ | RBfst ]. + Proof. + pose proof (LRTyEqRedRight RA RAA). + (* pose proof (LRTyEqRedRight _ RBBfst). *) + escape. + eapply redSubstTmEq; tea. + 1: now eapply LRTmEqConv. + 1,2: eapply redtm_snd_beta; tea; now eapply ty_conv. + Qed. + + Import SigRedTmEq. + + (* #[program] + Definition pairSigRedTm {Γ A B a b l} + (RΣ0 : [Γ ||-Σ tSig A B]) + (RΣ := normRedΣ0 RΣ0) + (RA : [Γ ||- A]) + (RBa : [Γ ||- B[a..]]) + (Ra : [Γ ||- a : A | RA]) + (Rb : [Γ ||- b : _ | RBa ]) : + SigRedTm RΣ (tPair A B a b) := {| nf := tPair A B a b |}. + Next Obligation. + destruct (polyRedId (normRedΣ0 RΣ0)); escape. + eapply redtmwf_refl; cbn; now eapply ty_pair. + Qed. + Next Obligation. + unshelve econstructor; intros; cbn. + * irrelevanceCumRefl; now eapply wkTermEq. + * eapply reflLRTyEq. + * irrelevanceRefl. + unshelve eapply (PolyRed.posExt (normRedΣ0 RΣ0)); tea. + * irrelevanceCum0. + 2: now eapply wkTermEq. + clear; now bsimpl. + Unshelve. all: tea. + Qed. *) + + Set Printing Universes. + + Obligation Tactic := idtac. + #[program] + Definition pairSigRedTm@{i j k l} {Γ F G A B a b l} + (RΣ0 : ParamRedTy@{i j k l} tSig Γ l (tSig F G)) + (RΣ := normRedΣ0@{i j k l} RΣ0) + (RΣ' := LRSig'@{i j k l} RΣ) + (RAB : [_ ||- _ ≅ tSig A B | RΣ']) + (Rdom : [LogRel@{i j k l} l | Γ ||- F]) + (Rcoda : [LogRel@{i j k l} l | Γ ||- G[a..]]) + (Rcod := snd@{l l} (polyRedId RΣ)) + (Ra : [Γ ||- a : _ | Rdom]) + (Rb : [Γ ||- b : _ | Rcoda ]) + (RA : [_ ||- _ ≅ A | Rdom]) + (RBa : [_ ||- _ ≅ B[a..] | Rcoda]) : + SigRedTm RΣ (tPair A B a b) := + {| nf := tPair A B a b |}. + Next Obligation. + intros. + destruct (polyRedId (normRedΣ0 (invLRΣ (LRTyEqRedRight _ RAB)))). + eapply redtmwf_refl; cbn; eapply ty_conv. + 2: now (symmetry; eapply escapeEq). + eapply ty_pair; first [now eapply escape| eapply ty_conv; [now eapply escapeTerm| now eapply escapeEq]]. + Qed. + Next Obligation. + unshelve econstructor; intros; cbn. + 1-3:irrelevanceRefl. + * now eapply wkTermEq. + * now eapply wkEq. + * unshelve eapply (posRedExt (normEqRedΣ _ RAB)); tea; irrelevance. + * irrelevance0. + 2: now eapply wkTermEq. + clear; now bsimpl. + Unshelve. all: tea. + Qed. + + Definition pairSigRedTm0@{i j k l} {Γ A B a b l} + (RΣ0 : ParamRedTy@{i j k l} tSig Γ l (tSig A B)) + (RΣ := normRedΣ0@{i j k l} RΣ0) + (RΣ' := LRSig'@{i j k l} RΣ) + (RA : [LogRel@{i j k l} l | Γ ||- A]) + (RBa : [LogRel@{i j k l} l | Γ ||- B[a..]]) + (* (RBfst : [Γ ||- B[(tFst (tPair A B a b))..] ]) *) (Ra : [Γ ||- a : A | RA]) (Rb : [Γ ||- b : _ | RBa ]) : - [Γ ||- tSnd (tPair A B a b) : _ | RBfst ] × [Γ ||- tSnd (tPair A B a b) ≅ b : _ | RBfst]. + SigRedTm RΣ (tPair A B a b). + Proof. + unshelve eapply pairSigRedTm; tea; eapply reflLRTyEq. + Defined. + + Lemma pairCongRed0 {Γ A A' B B' a a' b b' l} + (RΣ0 : [Γ ||-Σ tSig A B]) + (RΣ := normRedΣ0 RΣ0) + (RΣ' := LRSig' RΣ) + (RΣeq : [_ ||- _ ≅ tSig A' B' | RΣ']) + (RA : [Γ ||- A]) + (RA' : [Γ ||- A']) + (RBa : [Γ ||- B[a..]]) + (RBconv : [Γ ||- B[a..] ≅ B[(tFst (tPair A B a b))..] | RBa ]) + (Ra : [Γ ||- a ≅ a' : A | RA]) + (Rb : [Γ ||- b ≅ b' : _ | RBa ]) : + [Γ ||- tPair A B a b ≅ tPair A' B' a' b' : _ | RΣ']. Proof. + pose proof (RBa' := polyCodSubstRed RA RΣ _ _ (urefl Ra)). + destruct (polyRedId RΣ) as [_ ?]. + destruct (polyRedId (normRedΣ0 (invLRΣ (LRTyEqRedRight _ RΣeq)))). + destruct (polyRedEqId _ (normEqRedΣ _ RΣeq)) as [? ?]. + assert (RAA' : [RA | _ ||- _ ≅ A']) by irrelevance. + epose proof (RBaeq := polyCodSubstExtRed _ RΣ _ _ Ra). + epose proof (polyRedEqCodSubstExt _ _ (normEqRedΣ _ RΣeq) _ _ (urefl Ra)). + assert (RBBa' : [RBa' | _ ||- _ ≅ B'[a'..]]) by irrelevance. + epose proof (polyRedEqCodSubstExt _ _ (normEqRedΣ _ RΣeq) _ _ Ra). + cbn in RBaeq. + assert (Rb' : [_ ||- b' : _ | RBa']) by (eapply LRTmEqConv; tea; eapply urefl; irrelevance). escape. - eapply redSubstTerm; tea. - 2: eapply redtm_snd_beta; tea. - now eapply LRTmRedConv. + set (Rpair := pairSigRedTm0 RΣ0 RA RBa (lrefl Ra) (lrefl Rb)). + set (Rpair' := pairSigRedTm RΣ0 RΣeq RA RBa' (urefl Ra) Rb' RAA' RBBa'). + assert [RA | _ ||- tFst (tPair A B a b) ≅ tFst (tPair A' B' a' b') : _]. + 1:{ + eapply redSubstTmEq; tea. + all: eapply redtm_fst_beta; tea; cbn; eapply ty_conv; tea. + } + unshelve eapply (build_sigRedTmEq' _ Rpair Rpair'); subst Rpair Rpair'; cbn. + 1: irrelevance. + eapply redSubstTmEq. + 2,3: eapply redtm_snd_beta; tea; now eapply ty_conv. + 1: now eapply LRTmEqConv. + cbn; unshelve eapply escapeEq, (polyRedEqCodSubstExt _ _ (normEqRedΣ _ RΣeq)); tea. Qed. - - Lemma pairRed {Γ A B a b l} + (* Lemma pairRed0 {Γ A B a b l} (RΣ0 : [Γ ||-Σ tSig A B]) (RΣ := normRedΣ0 RΣ0) + (RΣ' := LRSig' RΣ) (RA : [Γ ||- A]) (RBa : [Γ ||- B[a..]]) (RBconv : [Γ ||- B[a..] ≅ B[(tFst (tPair A B a b))..] | RBa ]) (RBfst : [Γ ||- B[(tFst (tPair A B a b))..] ]) (Ra : [Γ ||- a : A | RA]) (Rb : [Γ ||- b : _ | RBa ]) : - [Γ ||- tPair A B a b : _ | LRSig' RΣ]. - Proof. - destruct (polyRedId RΣ); escape. - unshelve eexists (tPair A B a b) _. - + intros ? ρ wfΔ. - rewrite wk_fst; irrelevanceRefl; unshelve eapply wkTerm. - 1:tea. - 2: unshelve eapply pairFstRed; tea. - + eapply redtmwf_refl; cbn. - eapply ty_pair; tea. - + unshelve econstructor; cbn; intros. - - irrelevance0; [reflexivity|eapply wkTerm]. - apply Ra. - Unshelve. tea. - - apply reflLRTyEq. - - apply reflLRTyEq. - - assert (Hrw : B[a⟨ρ⟩ .: ρ >> tRel] = B[a..]⟨ρ⟩) by now bsimpl. - irrelevance0; [symmetry; apply Hrw|eapply wkTerm]. - apply Rb. - Unshelve. tea. - + eapply convtm_eta_sig; tea. - * now eapply ty_pair. - * constructor; tea; (unshelve eapply escapeEq, reflLRTyEq; [|tea]). - * now eapply ty_pair. - * constructor; tea; (unshelve eapply escapeEq, reflLRTyEq; [|tea]). - * enough [Γ |- tFst (tPair A B a b) ≅ a : A]. - 1: transitivity a; tea; now symmetry. - eapply convtm_exp. - - now eapply redtm_fst_beta. - - now eapply redtmwf_refl. - - tea. - - tea. - - tea. - - unshelve eapply escapeEq, reflLRTyEq; [|tea]. - - eapply escapeEqTerm; now eapply reflLRTmEq. - * enough [Γ |- tSnd (tPair A B a b) ≅ b : B[(tFst (tPair A B a b))..]]. - 1: transitivity b; tea; now symmetry. - eapply convtm_conv; tea. - eapply convtm_exp. - - eapply redtm_conv; [| now symmetry]; now eapply redtm_snd_beta. - - now eapply redtm_refl. - - tea. - - tea. - - tea. - - unshelve eapply escapeEq, reflLRTyEq; [|tea]. - - eapply escapeEqTerm; now eapply reflLRTmEq. - + intros ? ρ wfΔ. - irrelevance0. - 2: rewrite wk_snd; eapply wkTerm; now eapply pairSndRed. - now rewrite subst_ren_subst_mixed. - Unshelve. all: tea. - Qed. - - Lemma isLRPair_isWfPair {Γ A B l p} (wfΓ : [|- Γ]) (ΣA : [Γ ||-Σ tSig A B]) - (RA : [Γ ||- A]) - (Rp : [Γ ||-Σ p : _ | normRedΣ0 ΣA]) : - isWfPair Γ A B (SigRedTm.nf Rp). - Proof. - destruct Rp; simpl; intros. - destruct ispair as [A' B' a b HA HB Ha Hb|]; [|constructor; tea]. - assert (Hrw : forall A t, t[tRel 0 .: @wk1 Γ A >> tRel] = t). - { intros; bsimpl; apply idSubst_term; intros []; reflexivity. } - assert (Hrw' : forall a t, t[a .: @wk_id Γ >> tRel] = t[a..]). - { intros; now bsimpl. } - assert [Γ |-[ ta ] A]. - { now unshelve eapply escape. } - assert [Γ |-[ ta ] A']. - { rewrite <- (wk_id_ren_on Γ A'); now unshelve eapply escapeConv, HA. } - assert [|-[ ta ] Γ,, A']. - { now eapply wfc_cons. } - assert [Γ |-[ ta ] A ≅ A']. - { rewrite <- (wk_id_ren_on Γ A), <- (wk_id_ren_on Γ A'). - eapply escapeEq; now unshelve apply HA. } - assert (RB : - forall (Δ : context) (a : term) (ρ : Δ ≤ Γ) (h : [ |-[ ta ] Δ]) - (ha : [PolyRedPack.shpRed (normRedΣ0 ΣA) ρ h | Δ ||- a - : (ParamRedTyPack.dom (normRedΣ0 ΣA))⟨ρ⟩]), - [Δ ||- B[a .: ρ >> tRel]] - ). - { intros Δ a0 ? ? ha0. - unshelve econstructor. - + eapply PolyRedPack.posRed, ha0. - + eapply PolyRedPack.posAd, PolyRed.toAd. } - constructor; tea. - - rewrite <- (Hrw A' B); unshelve eapply escape, RB; tea. - apply var0conv; cbn; [|tea]. - rewrite <- (@wk1_ren_on Γ A' A'); apply convty_wk; [tea|now symmetry]. - - rewrite <- (Hrw A' B'); unshelve eapply escapeConv, HB; tea. - apply var0conv; cbn; [|tea]. - rewrite <- (@wk1_ren_on Γ A' A'); apply convty_wk; [tea|now symmetry]. - - rewrite <- (Hrw A B'); unshelve eapply escapeConv, HB. - + apply wfc_cons; tea. - + apply var0; cbn; [now bsimpl|tea]. - - rewrite <- (Hrw A B), <- (Hrw A B'); unshelve eapply escapeEq, HB. - + apply wfc_cons; tea. - + apply var0; cbn; [now bsimpl|tea]. - - rewrite <- (Hrw A' B), <- (Hrw A' B'); unshelve eapply escapeEq, HB. - + apply wfc_cons; tea. - + apply var0conv; cbn; tea. - rewrite <- (@wk1_ren_on Γ A' A'); apply convty_wk; [tea|now symmetry]. - - rewrite <- (wk_id_ren_on Γ A), <- (wk_id_ren_on Γ a). - now unshelve eapply escapeTerm, Ha. - - rewrite <- Hrw'. - unshelve eapply escape, RB; tea. - rewrite <- (wk_id_ren_on Γ a); now unshelve apply Ha. - - rewrite <- Hrw'; unshelve eapply escapeConv, HB; tea. - rewrite <- (wk_id_ren_on Γ a); now unshelve apply Ha. - - rewrite <- !Hrw'; unshelve eapply escapeEq, HB; tea. - rewrite <- (wk_id_ren_on Γ a); now unshelve apply Ha. - - rewrite <- Hrw', <- (wk_id_ren_on Γ b), <- (wk_id_ren_on Γ a). - now unshelve eapply escapeTerm, Hb. - Qed. - - Lemma isLRPair_isPair {Γ A B l p} (ΣA : [Γ ||-Σ tSig A B]) (Rp : [Γ ||-Σ p : _ | normRedΣ0 ΣA]) : - isPair (SigRedTm.nf Rp). - Proof. - destruct Rp; simpl; intros. - destruct ispair; constructor; tea. - now eapply convneu_whne. + [Γ ||- tPair A B a b : _ | RΣ']. + Proof. + destruct (polyRedId RΣ) as [_ ?]; escape. + set (Rpair := pairSigRedTm RΣ0 RA RBa Ra Rb). + unshelve eapply (build_sigRedTmEq' _ Rpair Rpair); subst Rpair; cbn. + + eapply redSubstTmEq. + 2,3: now eapply redtm_fst_beta. + 1: irrelevanceCum. + now unshelve eapply escapeEq, reflLRTyEq. + + eapply redSubstTmEq. + 2,3: now eapply redtm_snd_beta. + 1: now eapply LRTmEqConv. + now unshelve eapply escapeEq, reflLRTyEq. Qed. + Lemma pairRed {Γ A B a b l} + (RΣ0 : [Γ ||-Σ tSig A B]) + (RΣ := normRedΣ0 RΣ0) + (RΣ' := LRSig' RΣ) + (RA : [Γ ||- A]) + (RBa : [Γ ||- B[a..]]) + (Ra : [Γ ||- a : A | RA]) + (Rb : [Γ ||- b : _ | RBa ]) : + [Γ ||- tPair A B a b : _ | RΣ']. + Proof. + assert [_ ||- a ≅ tFst (tPair A B a b) : _ | RA]. + 1:{ symmetry; eapply redSubstLeftTmEq; tea . + destruct (polyRedId RΣ) as [_ ?]; escape. + now eapply redtm_fst_beta. } + eapply pairRed0; tea. + + now unshelve now eapply singleSubstEqΣ. + + eapply singleSubstΣ1; tea; now symmetry. + Qed. *) + Lemma sigEtaRed {Γ A B l p p'} (RΣ0 : [Γ ||-Σ tSig A B]) - (RΣ := LRSig' (normRedΣ0 RΣ0)) + (RΣn := normRedΣ0 RΣ0) + (RΣ := LRSig' RΣn) (RA : [Γ ||- A]) (RBfst : [Γ ||- B[(tFst p)..]]) - (RBfst' : [Γ ||- B[(tFst p')..]]) (Rp : [Γ ||- p : _ | RΣ]) - (Rp' : [Γ ||- p' : _ | RΣ]) - (RBfstEq0 : [Γ ||- B[(tFst p)..] ≅ B[(tFst p')..] | RBfst]) - (RBfstnf : [Γ ||- B[(tFst Rp.(SigRedTm.nf))..]]) - (RBfstEq : [Γ ||- B[(tFst p)..] ≅ B[(tFst Rp.(SigRedTm.nf))..] | RBfst]) - (RBfstEq' : [Γ ||- B[(tFst p')..] ≅ B[(tFst Rp'.(SigRedTm.nf))..] | RBfst']) + (Rp' : [Γ ||- p' : _ | RΣ]) (Rfstpp' : [Γ ||- tFst p ≅ tFst p' : _ | RA]) (Rsndpp' : [Γ ||- tSnd p ≅ tSnd p' : _ | RBfst]) : [Γ ||- p ≅ p' : _ | RΣ]. Proof. - exists Rp Rp'. - - destruct (polyRedId (normRedΣ0 RΣ0)) as [_ RB]. - assert ([Γ ||- SigRedTm.nf Rp : _ | RΣ] × [Γ ||- p ≅ SigRedTm.nf Rp : _ | RΣ]) as [Rnf Rpnf]. - 1: eapply (redTmFwdConv Rp (SigRedTm.red Rp)), isPair_whnf, isLRPair_isPair. - assert ([Γ ||- SigRedTm.nf Rp' : _ | RΣ]× [Γ ||- p' ≅ SigRedTm.nf Rp' : _ | RΣ]) as [Rnf' Rpnf']. - 1: eapply (redTmFwdConv Rp' (SigRedTm.red Rp')), isPair_whnf, isLRPair_isPair. - destruct (fstRed RΣ0 RA Rp) as [[Rfstp Rfsteq] Rfstnf]. - destruct (fstRed RΣ0 RA Rp') as [[Rfstp' Rfsteq'] Rfstnf']. - destruct (sndRed RΣ0 RA Rp RBfst RBfstEq). - destruct (sndRed RΣ0 RA Rp' RBfst' RBfstEq'). - escape. - assert [|- Γ] by now eapply wfc_wft. - eapply convtm_eta_sig; tea. - 1, 2: now eapply isLRPair_isWfPair. - + transitivity (tFst p). - 1: now symmetry. - transitivity (tFst p'); tea. - + eapply convtm_conv; tea. - transitivity (tSnd p). - 1: now symmetry. - transitivity (tSnd p'); tea. - eapply convtm_conv; tea. - now symmetry. - - intros; do 2 rewrite wk_fst. - irrelevanceRefl. eapply wkTermEq. - destruct (fstRed RΣ0 RA Rp) as [[Rfstp Rfsteq] Rfstnf]. - destruct (fstRed RΣ0 RA Rp') as [[Rfstp' Rfsteq'] Rfstnf']. - eapply transEqTerm; tea. - eapply transEqTerm; tea. - now eapply LRTmEqSym. - Unshelve. tea. - - intros. do 2 rewrite wk_snd. - eapply LRTmEqRedConv. - 2:{ - eapply wkTermEq. - destruct (sndRed RΣ0 RA Rp RBfst RBfstEq). - destruct (sndRed RΣ0 RA Rp' RBfst' RBfstEq'). - eapply transEqTerm; tea. - eapply LRTmEqRedConv; tea. - eapply transEqTerm; tea. - now eapply LRTmEqSym. - Unshelve. tea. - } - rewrite wk_fst, <- subst_ren_subst_mixed. - eapply wkEq. cbn. - eapply LRTyEqSym. - eapply transEq; tea. - now eapply LRTyEqSym. - Unshelve. tea. + assert (RBfst' : [Γ ||- B[(tFst p')..]]). + 1: eapply singleSubstΣ1; tea; now symmetry. + pose proof (Rpnf := sigRedTmEqNf _ Rp). + pose proof (Rpnf' := sigRedTmEqNf _ Rp'). + pose proof (fstRed _ RA Rpnf). + pose proof (fstRed _ RA Rpnf'). + unshelve eapply (build_sigRedTmEq' _ Rp.(redL) Rp'.(redL)). + - transitivity (tFst p); [symmetry|transitivity (tFst p')]; irrelevance. + - pose proof (sndRed _ RA Rpnf RBfst). + pose proof (sndRed _ RA Rpnf' RBfst'). + unshelve (eapply LRTmEqConv; [cbn; now unshelve now eapply singleSubstEqΣ|]); tea. + transitivity (tSnd p); [symmetry |transitivity (tSnd p')]. + 1,2: irrelevance. + eapply LRTmEqConv; tea. + now unshelve (eapply singleSubstEqΣ; now symmetry). Qed. + + Lemma subst_sig {A B σ} : (tSig A B)[σ] = (tSig A[σ] B[up_term_term σ]). Proof. reflexivity. Qed. Lemma subst_pair {A B a b σ} : (tPair A B a b)[σ] = (tPair A[σ] B[up_term_term σ] a[σ] b[σ]). Proof. reflexivity. Qed. - - Lemma subst_fst {p σ} : (tFst p)[σ] = tFst p[σ]. - Proof. reflexivity. Qed. - + Lemma subst_snd {p σ} : (tSnd p)[σ] = tSnd p[σ]. Proof. reflexivity. Qed. - - Lemma pairFstRedEq {Γ A A' B B' a a' b b' l} - (RA : [Γ ||- A]) - (RA' : [Γ ||- A']) - (RAA' :[Γ ||- A ≅ A' | RA]) - (RB : [Γ ,, A ||- B]) - (RB' : [Γ ,, A' ||- B']) - (RBa : [Γ ||- B[a..]]) - (RBa' : [Γ ||- B'[a'..]]) - (Ra : [Γ ||- a : A | RA]) - (Ra' : [Γ ||- a' : A' | RA']) - (Raa' : [Γ ||- a ≅ a' : A | RA]) - (Rb : [Γ ||- b : _ | RBa ]) - (Rb' : [Γ ||- b' : _ | RBa' ]) : - [Γ ||- tFst (tPair A B a b) ≅ tFst (tPair A' B' a' b') : _ | RA]. - Proof. - destruct (pairFstRed RA RB RBa Ra Rb). - destruct (pairFstRed RA' RB' RBa' Ra' Rb'). - eapply transEqTerm; tea. - eapply transEqTerm; tea. - eapply LRTmEqRedConv. - + now eapply LRTyEqSym. - + now eapply LRTmEqSym. - Qed. - - Lemma pairSndRedEq {Γ A A' B B' a a' b b' l} - (RA : [Γ ||- A]) - (RA' : [Γ ||- A']) - (RAA' :[Γ ||- A ≅ A' | RA]) - (RB : [Γ ,, A ||- B]) - (RB' : [Γ ,, A' ||- B']) - (RBa : [Γ ||- B[a..]]) - (RBa' : [Γ ||- B'[a'..]]) - (RBaBa' : [Γ ||- B[a..] ≅ B'[a'..] | RBa ]) - (RBfst : [Γ ||- B[(tFst (tPair A B a b))..] ]) - (RBconv : [Γ ||- B[a..] ≅ B[(tFst (tPair A B a b))..] | RBa ]) - (RBfst' : [Γ ||- B'[(tFst (tPair A' B' a' b'))..] ]) - (RBconv' : [Γ ||- B'[a'..] ≅ B'[(tFst (tPair A' B' a' b'))..] | RBa' ]) - (Ra : [Γ ||- a : A | RA]) - (Ra' : [Γ ||- a' : A' | RA']) - (Rb : [Γ ||- b : _ | RBa ]) - (Rb' : [Γ ||- b' : _ | RBa' ]) - (Rbb' : [Γ ||- b ≅ b' : _ | RBa]) : - [Γ ||- tSnd (tPair A B a b) ≅ tSnd (tPair A' B' a' b') : _ | RBfst]. + + Lemma up_subst_single' t a σ : t[up_term_term σ][(a[σ])..] = t[a..][σ]. + Proof. now bsimpl. Qed. + + Lemma pairFstValid0 {Γ A B a b l} + (VΓ : [||-v Γ]) + (VA : [ Γ ||-v A | VΓ ]) + (VB : [ Γ ,, A ||-v B | validSnoc VΓ VA]) + (VΣ := SigValid VΓ VA VB) + (Va : [Γ ||-v a : A | VΓ | VA]) + (VBa := substS VB Va) + (Vb : [Γ ||-v b : B[a..] | VΓ | VBa]) : + [Γ ||-v tFst (tPair A B a b) ≅ a : _ | VΓ | VA]. Proof. - destruct (pairSndRed RA RB RBa RBfst RBconv Ra Rb). - destruct (pairSndRed RA' RB' RBa' RBfst' RBconv' Ra' Rb'). - eapply transEqTerm; tea. - eapply LRTmEqRedConv; tea. - eapply transEqTerm; tea. - eapply LRTmEqRedConv; cycle 1. - + now eapply LRTmEqSym. - + eapply LRTyEqSym; eapply transEq; cycle 1; tea. + eapply redSubstValid; tea. + constructor; intros; rewrite <-subst_fst, subst_pair. + instValid Vσσ'; instValid (liftSubstEq' VA Vσσ'); escape. + eapply redtm_fst_beta; tea. + now rewrite up_subst_single'. Qed. - Lemma pairFstValid {Γ A B a b l} (VΓ : [||-v Γ]) (VA : [ Γ ||-v A | VΓ ]) @@ -766,31 +683,34 @@ Section PairRed. [Γ ||-v tFst (tPair A B a b) : _ | VΓ | VA] × [Γ ||-v tFst (tPair A B a b) ≅ a : _ | VΓ | VA]. Proof. - assert (h : forall {Δ σ} (wfΔ : [|- Δ]) (Vσ : [VΓ | Δ ||-v σ : Γ | wfΔ]), - [validTy VA wfΔ Vσ | Δ ||- (tFst (tPair A B a b))[σ] : A[σ]] × - [validTy VA wfΔ Vσ | Δ ||- (tFst (tPair A B a b))[σ] ≅ a[σ] : A[σ]]). - 1:{ - intros ?? wfΔ Vσ; instValid Vσ. - pose (RVΣ0 := normRedΣ0 (invLRΣ RVΣ)). - pose (RVB := snd (polyRedId RVΣ0)). - unshelve eapply pairFstRed; tea; fold subst_term. - + now rewrite <- singleSubstComm'. - + irrelevance0; tea; now rewrite singleSubstComm'. - } - split; unshelve econstructor. - - apply h. - - intros ??? wfΔ Vσ Vσ' Vσσ'. - instAllValid Vσ Vσ' Vσσ'. - pose (RΣ := normRedΣ0 (invLRΣ RVΣ)); - pose (RΣ' := normRedΣ0 (invLRΣ RVΣ0)); fold subst_term in *. - pose (RVB := snd (polyRedId RΣ)). - pose (RVB' := snd (polyRedId RΣ')). - unshelve eapply pairFstRedEq; tea; fold subst_term. - 1,2: now rewrite <- singleSubstComm'. - all: irrelevance0; tea; now rewrite singleSubstComm'. - - apply h. - Qed. - + split; [eapply lreflValidTm|]; now eapply pairFstValid0. + Qed. + + Lemma pairSndValid0 {Γ A B a b l} + (VΓ : [||-v Γ]) + (VA : [ Γ ||-v A | VΓ ]) + (VB : [ Γ ,, A ||-v B | validSnoc VΓ VA]) + (VΣ := SigValid VΓ VA VB) + (Va : [Γ ||-v a : A | VΓ | VA]) + (VBa := substS VB Va) + (Vb : [Γ ||-v b : B[a..] | VΓ | VBa]) + (Vfstall := pairFstValid VΓ VA VB Va Vb) + (VBfst := substS VB (fst Vfstall)) : + [Γ ||-v tSnd (tPair A B a b) ≅ b : _ | VΓ | VBfst]. + Proof. + eapply redSubstValid; cycle 1. + + eapply conv; tea. + pose proof (h := symValidTmEq (pairFstValid0 VΓ VA VB Va Vb)). + epose proof (substSEq (reflValidTy VA) (reflValidTy VB) h). + irrValid. + + constructor; intros. + rewrite <-up_subst_single', subst_snd, <-subst_fst, subst_pair. + instValid Vσσ'; instValid (liftSubstEq' VA Vσσ'); escape. + eapply redtm_snd_beta; tea. + now rewrite up_subst_single'. + Qed. + + Lemma pairSndValid {Γ A B a b l} (VΓ : [||-v Γ]) (VA : [ Γ ||-v A | VΓ ]) @@ -804,42 +724,45 @@ Section PairRed. [Γ ||-v tSnd (tPair A B a b) : _ | VΓ | VBfst] × [Γ ||-v tSnd (tPair A B a b) ≅ b : _ | VΓ | VBfst]. Proof. - pose proof (VBfsteq := substSEq (reflValidTy _ VA) VB (reflValidTy _ VB) Va (fst Vfstall) (symValidTmEq (snd (Vfstall)))). - cbn in VBfsteq. - assert (h : forall {Δ σ} (wfΔ : [|- Δ]) (Vσ : [VΓ | Δ ||-v σ : Γ | wfΔ]), - [validTy VBfst wfΔ Vσ | Δ ||- (tSnd (tPair A B a b))[σ] : _ ] × - [validTy VBfst wfΔ Vσ | Δ ||- (tSnd (tPair A B a b))[σ] ≅ b[σ] : _ ]). + split; [eapply lreflValidTm|]; now eapply pairSndValid0. + Qed. + + Lemma pairCongValid {Γ A A' B B' a a' b b' l} + (VΓ : [||-v Γ]) + (VA : [ Γ ||-v A | VΓ ]) + (VA' : [ Γ ||-v A ≅ A' | VΓ | VA]) + (VB : [ Γ ,, A ||-v B | validSnoc VΓ VA ]) + (VBB' : [ Γ ,, A ||-v B ≅ B' | validSnoc VΓ VA | VB]) + (VΣ := SigValid VΓ VA VB) + (Va : [Γ ||-v a ≅ a' : A | VΓ | VA]) + (VBa := substS VB Va) + (Vb : [Γ ||-v b ≅ b' : B[a..] | VΓ | VBa]) : + [Γ ||-v tPair A B a b ≅ tPair A' B' a' b' : _ | VΓ | VΣ]. + Proof. + constructor; intros; rewrite 2!subst_pair. + assert [_ ||-v tSig A B ≅ tSig A' B' | VΓ | VΣ]. 1:{ - intros ?? wfΔ Vσ ; instValid Vσ. - pose (RVΣ0 := normRedΣ0 (invLRΣ RVΣ)). - pose (RVB := snd (polyRedId RVΣ0)); fold subst_term in *. - irrelevance0. 1: now rewrite singleSubstComm'. - unshelve eapply pairSndRed; tea; fold subst_term. - + now rewrite <- singleSubstComm'. - + rewrite <- subst_pair, <- subst_fst; irrelevance0; - rewrite <- singleSubstComm'; tea; reflexivity. - Unshelve. now rewrite <- singleSubstComm'. - + irrelevance0; tea; now rewrite <- singleSubstComm'. + eapply SigCong; tea. + eapply convCtx1; tea; now eapply ureflValidTy. } - split; unshelve econstructor. - - apply h. - - intros ??? wfΔ Vσ Vσ' Vσσ' ; instAllValid Vσ Vσ' Vσσ'. - pose (RΣ := normRedΣ0 (invLRΣ RVΣ)). - pose (RΣ' := normRedΣ0 (invLRΣ RVΣ0)). - pose (RB := snd (polyRedId RΣ)); - pose (RB' := snd (polyRedId RΣ')); fold subst_term in *. - irrelevance0. 1: now rewrite singleSubstComm'. - unshelve eapply pairSndRedEq; tea; fold subst_term. - 1,2: now rewrite <- singleSubstComm'. - all: try (irrelevance0; tea; now rewrite <- singleSubstComm'). - all: rewrite <- subst_pair, <- subst_fst. - 2: rewrite <- singleSubstComm'; tea. - all: irrelevance0; rewrite <- singleSubstComm'; try reflexivity; tea. - Unshelve. now rewrite <- singleSubstComm'. - - apply h. + instValid Vσσ'; instValid (liftSubstEq' VA Vσσ'). + irrelevance0; [now rewrite subst_sig|]. + eapply pairCongRed0; tea. + + irrelevance. + + + pose proof (lreflValidTm _ Vb). + unshelve epose proof (Rafst := symValidTmEq (pairFstValid0 VΓ VA VB (lreflValidTm _ Va) _)). + 2: irrValid. + pose proof (RBaBfst := substSEq (reflValidTy VA) (reflValidTy VB) Rafst). + pose proof (validTyEq RBaBfst _ (lrefl Vσσ')). + rewrite <-subst_pair, subst_fst, up_subst_single'; irrelevance. + + irrelevance. + Unshelve. + 1: now eapply ureflValidTy. + 1: rewrite <-subst_sig; now apply invLRΣ. + now rewrite up_subst_single'. Qed. - Lemma pairValid {Γ A B a b l} (VΓ : [||-v Γ]) (VA : [ Γ ||-v A | VΓ ]) @@ -850,90 +773,10 @@ Section PairRed. (Vb : [Γ ||-v b : B[a..] | VΓ | VBa]) : [Γ ||-v tPair A B a b : _ | VΓ | VΣ]. Proof. - destruct (pairFstValid VΓ VA VB Va Vb) as [Vfst Vfsteq]. - destruct (pairSndValid VΓ VA VB Va Vb) as [Vsnd Vsndeq]. - pose proof (VBfst := substS VB Vfst). - pose proof (VBfsteq := substSEq (reflValidTy _ VA) VB (reflValidTy _ VB) Va Vfst (symValidTmEq Vfsteq)). - cbn in VBfsteq. - assert (pairSubstRed : forall (Δ : context) (σ : nat -> term) (wfΔ : [ |-[ ta ] Δ]) - (Vσ : [VΓ | Δ ||-v σ : Γ | wfΔ]), - [Δ ||- (tPair A B a b)[σ] : (tSig A B)[σ] | validTy VΣ wfΔ Vσ]). - 1:{ - intros ?? wfΔ Vσ. - instValid Vσ; instValidExt Vσ (reflSubst _ _ Vσ). - pose (RVΣ0 := normRedΣ0 (invLRΣ RVΣ)). - pose (RVB := snd (polyRedId RVΣ0)); fold subst_term in *. - pose (Vaσ := consSubstSvalid Vσ Va); instValid Vaσ. - rewrite <- up_subst_single in RVB0. - assert (RVb' : [RVB0 | Δ ||- b[σ] : B[up_term_term σ][(a[σ])..]]) - by (irrelevance0; tea; apply singleSubstComm'). - unshelve epose (p := pairFstRed RVA RVB RVB0 RVa RVb'). - destruct p as [Rfst Rfsteq%LRTmEqSym]. - pose proof (Vfstσ := consSubstS _ _ Vσ VA Rfst). - epose proof (Veqσ := consSubstSEq' _ _ Vσ (reflSubst _ _ Vσ) VA RVa Rfsteq). - instValid Vfstσ; instValidExt Vfstσ Veqσ. - unshelve epose (pairRed RVΣ0 RVA RVB0 _ _ RVa RVb'); fold subst_term in *. - 3: irrelevance. - + rewrite <- up_subst_single in REVB. - irrelevance0; tea; now rewrite up_subst_single. - + now rewrite <- up_subst_single in RVB1. - } - unshelve econstructor. - 1: intros; now eapply pairSubstRed. - intros ??? wfΔ Vσ Vσ' Vσσ'. - instAllValid Vσ Vσ' Vσσ'. - set (p := _[_]); set (p' := _[_]). - pose (RΣ := normRedΣ RVΣ); pose (RΣ' := normRedΣ RVΣ0); fold subst_term in *. - pose proof (Rp0 := pairSubstRed _ _ wfΔ Vσ). - pose proof (Rp0' := pairSubstRed _ _ wfΔ Vσ'). - assert (Rp : [Δ ||- p : _ | RΣ]) by irrelevance. - assert (Rp' : [Δ ||- p' : _ | RΣ]). - 1: eapply LRTmRedConv; [|exact (pairSubstRed _ _ wfΔ Vσ')]; now eapply LRTyEqSym. - irrelevance0. - 1: symmetry; apply subst_sig. - destruct (fstRed (invLRΣ RVΣ) RVA Rp) as [[Rfstp Rfsteq] Rfstnf]. - destruct (fstRed (invLRΣ RVΣ) RVA Rp') as [[Rfstp' Rfsteq'] Rfstnf']. - pose (Vfstpσ := consSubstS _ _ Vσ VA Rfstp). - pose (Vfstnfσ := consSubstS _ _ Vσ VA Rfstnf). - pose proof (Vfsteqσ := consSubstSEq' _ _ Vσ (reflSubst _ _ Vσ) VA Rfstp Rfsteq). - epose (Vfstpσ' := consSubstS _ _ Vσ VA Rfstp'). - pose (Vfstnfσ' := consSubstS _ _ Vσ VA Rfstnf'). - pose proof (Vfsteqσ' := consSubstSEq' _ _ Vσ (reflSubst _ _ Vσ) VA Rfstp' Rfsteq'). - pose (Vuσ := liftSubstS' VA Vσ). - pose (Vuσ' := liftSubstS' VA Vσ'). - pose (Vurσ' := liftSubstSrealign' VA Vσσ' Vσ'). - instValid Vuσ; instValid Vuσ'; instValid Vurσ'. - assert(Rfstpp' : [RVA | Δ ||- tFst p ≅ tFst p' : A[σ]]). - 1:{ - eapply pairFstRedEq; tea; irrelevance0; tea; fold subst_term. - 1,2: now rewrite singleSubstComm'. - Unshelve. all: fold subst_term. - 2: tea. - 1,2: now rewrite <- singleSubstComm'. - } - pose proof (Vfstppσ := consSubstSEq' _ _ Vσ (reflSubst _ _ Vσ) VA Rfstp Rfstpp'). - instAllValid Vfstpσ Vfstnfσ Vfsteqσ. - instAllValid Vfstpσ' Vfstnfσ' Vfsteqσ'. - instValidExt Vfstpσ' Vfstppσ. - rewrite <- up_subst_single in RVB2. - rewrite <- up_subst_single in RVB3, REVB. - rewrite <- up_subst_single in RVB4. - rewrite <- up_subst_single in RVB5, REVB0. - rewrite <- up_subst_single in REVB1. - eapply (sigEtaRed _ RVA RVB2 RVB4 Rp Rp');tea. - + irrelevance0; tea; now rewrite up_subst_single. - + irrelevance0; tea; now rewrite up_subst_single. - + irrelevance0; tea; now rewrite up_subst_single. - + unshelve eapply pairSndRedEq; tea; fold subst_term. - 7-9: irrelevance0; tea; now rewrite singleSubstComm'. - 1,2: now rewrite <- singleSubstComm'. - 1: irrelevance0; fold subst_term; now rewrite <- singleSubstComm'. - 2: rewrite <- subst_pair, <- subst_fst, <- singleSubstComm'; tea. - all: rewrite <- subst_pair, <- subst_fst; irrelevance0; - rewrite <- singleSubstComm'; try reflexivity; tea. - Qed. - - Lemma sigEtaValid {Γ A B p p' l} + eapply pairCongValid; tea; now eapply reflValidTy. + Qed. + + Lemma sigEtaEqValid {Γ A B p p' l} (VΓ : [||-v Γ]) (VA : [ Γ ||-v A | VΓ ]) (VB : [ Γ ,, A ||-v B | validSnoc VΓ VA]) @@ -946,36 +789,34 @@ Section PairRed. (Vsndpp' : [Γ ||-v tSnd p ≅ tSnd p' : _| VΓ | VBfst]) : [Γ ||-v p ≅ p' : _ | VΓ | VΣ]. Proof. - pose proof (Vfst' := fstValid VΓ VA VB Vp'). - pose proof (VBfst' := substS VB Vfst'). - pose proof (VBfsteq := substSEq (reflValidTy _ VA) VB (reflValidTy _ VB) Vfst Vfst' Vfstpp'). - cbn in VBfsteq. - constructor; intros ?? wfΔ Vσ; instValid Vσ. - pose proof (RΣ0 := invLRΣ RVΣ). - pose (RΣ := LRSig' (normRedΣ0 RΣ0)). - fold subst_term in *. - assert (Rp : [Δ ||- p[σ] : _ | RΣ]) by irrelevance. - assert (Rp' : [Δ ||- p'[σ] : _ | RΣ]) by irrelevance. - destruct (fstRed RΣ0 RVA Rp) as [[Rfstp Rfsteq] Rfstnf]. - destruct (fstRed RΣ0 RVA Rp') as [[Rfstp' Rfsteq'] Rfstnf']. - pose (Vfstpσ := consSubstS _ _ Vσ VA Rfstp). - pose (Vfstnfσ := consSubstS _ _ Vσ VA Rfstnf). - pose proof (Vfsteqσ := consSubstSEq' _ _ Vσ (reflSubst _ _ Vσ) VA Rfstp Rfsteq). - epose (Vfstpσ' := consSubstS _ _ Vσ VA Rfstp'). - pose (Vfstnfσ' := consSubstS _ _ Vσ VA Rfstnf'). - pose proof (Vfsteqσ' := consSubstSEq' _ _ Vσ (reflSubst _ _ Vσ) VA Rfstp' Rfsteq'). - instAllValid Vfstpσ Vfstnfσ Vfsteqσ. - instAllValid Vfstpσ' Vfstnfσ' Vfsteqσ'. - irrelevance0. - 1: now rewrite subst_sig. - unshelve eapply (sigEtaRed RΣ0 RVA _ _ Rp Rp') ; tea. - 1,2: rewrite <- subst_fst,<- singleSubstComm'; tea. - + irrelevance0; rewrite <- subst_fst, <- singleSubstComm'; try reflexivity; tea. - + now rewrite up_subst_single. - + irrelevance0; rewrite up_subst_single; try reflexivity; tea. - + irrelevance0; rewrite up_subst_single; try reflexivity; tea. - + irrelevance0. 1: now rewrite <- subst_fst, <- singleSubstComm'. tea. - Qed. + constructor; intros. + instValid Vσσ'. + irrelevance0; [now rewrite subst_sig|]. + eapply sigEtaRed; try irrelevance. + + eapply lrefl; irrelevance. + + eapply urefl; irrelevance. + Unshelve. + 2: rewrite <- subst_sig; now apply invLRΣ. + 1: tea. + now rewrite subst_fst, up_subst_single'. + Qed. + + + Lemma sigEtaValid {Γ A B p l} + (VΓ : [||-v Γ]) + (VA : [ Γ ||-v A | VΓ ]) + (VB : [ Γ ,, A ||-v B | validSnoc VΓ VA]) + (VΣ := SigValid VΓ VA VB) + (Vp : [Γ ||-v p : _ | VΓ | VΣ]) : + [Γ ||-v tPair A B (tFst p) (tSnd p) ≅ p : _ | VΓ | VΣ]. + Proof. + pose (Vfst := fstValid _ _ _ Vp). + pose (Vsnd := sndValid _ _ _ Vp). + pose proof (pairFstValid0 _ _ _ Vfst Vsnd). + pose proof (pairSndValid0 _ _ _ Vfst Vsnd). + pose proof (pairValid _ _ _ Vfst Vsnd). + unshelve eapply sigEtaEqValid; tea; try irrValid. + Qed. End PairRed. diff --git a/theories/Substitution/Introductions/SimpleArr.v b/theories/Substitution/Introductions/SimpleArr.v index 8240dee..7d15ded 100644 --- a/theories/Substitution/Introductions/SimpleArr.v +++ b/theories/Substitution/Introductions/SimpleArr.v @@ -2,7 +2,7 @@ From LogRel.AutoSubst Require Import core unscoped Ast Extra. From LogRel Require Import Utils BasicAst Notations Context NormalForms Weakening GenericTyping LogicalRelation Validity. From LogRel.LogicalRelation Require Import Escape Reflexivity Neutral Weakening Irrelevance. -From LogRel.Substitution Require Import Irrelevance Properties Conversion. +From LogRel.Substitution Require Import Irrelevance Properties Conversion Reflexivity. From LogRel.Substitution.Introductions Require Import Universe Pi Application Lambda Var. Set Universe Polymorphism. @@ -58,54 +58,4 @@ Section SimpleArrValidity. Unshelve. all: tea. Qed. - Lemma simple_idValid {Γ A l} - (VΓ : [||-v Γ]) - {VF : [Γ ||-v A | VΓ]} - (VΠ : [Γ ||-v arr A A | VΓ]) : - [Γ ||-v idterm A : arr A A | _ | VΠ]. - Proof. - eapply irrelevanceTm'. - 2: unshelve eapply lamValid. - 5: unshelve eapply var0Valid. - - now rewrite wk1_ren_on. - - tea. - Qed. - - Lemma simple_compValid {Γ A B C f g l} - (VΓ : [||-v Γ]) - {VA : [Γ ||-v A | VΓ]} - (VB : [Γ ||-v B | VΓ]) - (VC : [Γ ||-v C | VΓ]) - (VAB : [Γ ||-v arr A B | VΓ]) - (VBC : [Γ ||-v arr B C | VΓ]) - (VAC : [Γ ||-v arr A C | VΓ]) - (Vf : [Γ ||-v f : arr A B | _ | VAB]) - (Vg : [Γ ||-v g : arr B C | _ | VBC]) : - [Γ ||-v comp A g f : arr A C | _ | VAC]. - Proof. - eapply irrelevanceTm'. - 2: unshelve eapply lamValid. - 5: unshelve eapply simple_appValid. - 9: unshelve eapply simple_appValid. - 13: eapply var0Valid. - 4: eapply wk1ValidTy; exact VC. - 4: eapply wk1ValidTy; exact VB. - - now rewrite wk1_ren_on. - - eassumption. - - do 2 erewrite wk1_ren_on. rewrite<- arr_ren1. - erewrite<- wk1_ren_on. - now eapply wk1ValidTy. - - eapply irrelevanceTm'. 2: erewrite<- wk1_ren_on;now eapply wk1ValidTm. - rewrite wk1_ren_on. rewrite arr_ren1. - do 2 rewrite wk1_ren_on. reflexivity. - - do 2 erewrite wk1_ren_on. rewrite<- arr_ren1. - erewrite<- wk1_ren_on. - now eapply wk1ValidTy. - - eapply irrelevanceTm'. 2: erewrite<- wk1_ren_on;now eapply wk1ValidTm. - rewrite wk1_ren_on. rewrite arr_ren1. - do 2 rewrite wk1_ren_on. reflexivity. - - Unshelve. all: tea. - Qed. - End SimpleArrValidity. diff --git a/theories/Substitution/Introductions/Universe.v b/theories/Substitution/Introductions/Universe.v index 5da1986..5d65d07 100644 --- a/theories/Substitution/Introductions/Universe.v +++ b/theories/Substitution/Introductions/Universe.v @@ -11,8 +11,7 @@ Context `{GenericTypingProperties} {Γ : context}. Lemma UValid (VΓ : [||-v Γ]) : [Γ ||-v U | VΓ]. Proof. unshelve econstructor; intros. - - eapply LRU_; econstructor; tea; [constructor|]. - cbn; eapply redtywf_refl; gen_typing. + - now eapply LRU_, redUOneCtx. - cbn; constructor; eapply redtywf_refl; gen_typing. Defined. @@ -23,7 +22,7 @@ Lemma univValid {A l l'} (VΓ : [||-v Γ]) Proof. unshelve econstructor; intros. - instValid vσ. now eapply UnivEq. - - instAllValid vσ vσ' vσσ'; now eapply UnivEqEq. + - instValid vσσ'; now eapply UnivEqEq. Qed. Lemma univEqValid {A B l l'} (VΓ : [||-v Γ]) @@ -32,7 +31,7 @@ Lemma univEqValid {A B l l'} (VΓ : [||-v Γ]) (VAB : [Γ ||-v A ≅ B : U | VΓ | VU]) : [Γ ||-v A ≅ B | VΓ | VA]. Proof. - constructor; intros; instValid Vσ. + constructor; intros; instValid Vσσ'. now eapply UnivEqEq. Qed. diff --git a/theories/Substitution/Introductions/Var.v b/theories/Substitution/Introductions/Var.v index 0e59fb8..34d9784 100644 --- a/theories/Substitution/Introductions/Var.v +++ b/theories/Substitution/Introductions/Var.v @@ -10,22 +10,18 @@ Set Printing Primitive Projection Parameters. Section Var. Context `{GenericTypingProperties}. - + Lemma var0Valid {Γ l A} (VΓ : [||-v Γ]) (VA : [Γ ||-v A | VΓ]) : [Γ,, A ||-v tRel 0 : _ | validSnoc VΓ VA | wk1ValidTy _ VA ]. Proof. - constructor; intros; cbn. - + epose (validHead Vσ); irrelevance. - + epose (eqHead Vσσ'); irrelevance. + constructor; intros; cbn; epose (eqHead Vσσ'); irrelevance. Qed. - + Lemma var0Valid' {Γ l A} (VΓ : [||-v Γ,,A]) (VA : [Γ,,A ||-v A⟨↑⟩ | VΓ]) : [Γ,, A ||-v tRel 0 : _ | VΓ | VA ]. Proof. pose proof (invValiditySnoc VΓ) as [? [? [? ]]]; subst. - constructor; intros; cbn. - + epose (validHead Vσ); irrelevance. - + epose (eqHead Vσσ'); irrelevance. + constructor; intros; cbn; epose (eqHead Vσσ'); irrelevance. Qed. Lemma in_ctx_valid {Γ A n} (hin : in_ctx Γ n A) : forall (VΓ : [||-v Γ]), ∑ l, [Γ ||-v A | VΓ]. @@ -48,11 +44,11 @@ Section Var. eapply irrelevanceTm'; tea. now rewrite wk1_ren_on. Qed. - + Lemma var1Valid {Γ l A B} (VΓ : [||-v (Γ,, A) ,, B]) (VA : [_ ||-v A⟨↑⟩⟨↑⟩ | VΓ]) : [(Γ,, A) ,, B ||-v tRel 1 : _ | VΓ | VA ]. Proof. eapply varnValid; do 2 constructor. Qed. - + End Var. \ No newline at end of file diff --git a/theories/Substitution/Irrelevance.v b/theories/Substitution/Irrelevance.v index 5515837..2203700 100644 --- a/theories/Substitution/Irrelevance.v +++ b/theories/Substitution/Irrelevance.v @@ -1,4 +1,5 @@ +From Coq Require Import CRelationClasses. From LogRel.AutoSubst Require Import core unscoped Ast Extra. From LogRel Require Import Utils BasicAst Notations Context NormalForms Weakening GenericTyping LogicalRelation Validity. From LogRel.LogicalRelation Require Import Irrelevance Reflexivity Transitivity. @@ -8,42 +9,49 @@ Set Universe Polymorphism. Section Irrelevances. Context `{GenericTypingProperties}. +Section VRIrrelevant. +Universes u1 u2 u3 u4 v1 v2 v3 v4. -Lemma VRirrelevant Γ {vsubst vsubst' veqsubst veqsubst'} - (vr : VR Γ vsubst veqsubst) (vr' : VR Γ vsubst' veqsubst') : - (forall Δ σ wfΔ wfΔ', vsubst Δ σ wfΔ <~> vsubst' Δ σ wfΔ') × - (forall Δ σ σ' wfΔ wfΔ' vs vs', veqsubst Δ σ σ' wfΔ vs <~> veqsubst' Δ σ σ' wfΔ' vs'). +Lemma VRirrelevant Γ {veqsubst veqsubst'} + (vr : VR@{u1 u2 u3 u4} Γ veqsubst) (vr' : VR@{v1 v2 v3 v4} Γ veqsubst') : + (forall Δ σ σ' wfΔ wfΔ', veqsubst Δ wfΔ σ σ' <~> veqsubst' Δ wfΔ' σ σ'). Proof. - revert vsubst' veqsubst' vr'. pattern Γ, vsubst, veqsubst, vr. - apply VR_rect; clear Γ vsubst veqsubst vr. - - intros ?? h. inversion h. split; reflexivity. - - intros ?????? ih ?? h. inversion h. - specialize (ih _ _ VΓad0); destruct ih as [ih1 ih2]. - split. - + intros. split; intros []; unshelve econstructor. - 1,2: eapply ih1; eassumption. - 1,2: irrelevance. - + intros; split; intros []; unshelve econstructor. - 1,3: eapply ih2; eassumption. - 1,2: irrelevance. + revert veqsubst' vr'. pattern Γ, veqsubst, vr. + apply VR_rect; clear Γ veqsubst vr. + - intros ? h; inversion h; intros; split; intros; constructor. + - intros ?????? ih ? h; inversion h. + specialize (ih _ VΓad0). + intros; split; intros []; unshelve econstructor. + 1,2: now eapply ih. + all: irrelevanceCum. Qed. +Fail Fail Constraint u1 < v1. +Fail Fail Constraint v1 < u1. +Fail Fail Constraint u2 < v2. +Fail Fail Constraint v2 < u2. +Fail Fail Constraint u3 < v3. +Fail Fail Constraint v3 < u3. +Fail Fail Constraint u4 < v4. +Fail Fail Constraint v4 < u4. + +End VRIrrelevant. + Lemma irrelevanceSubst {Γ} (VΓ VΓ' : [||-v Γ]) {σ Δ} (wfΔ wfΔ' : [|- Δ]) : [Δ ||-v σ : Γ | VΓ | wfΔ] -> [Δ ||-v σ : Γ | VΓ' | wfΔ']. Proof. - apply (fst (VRirrelevant Γ VΓ.(VAd.adequate) VΓ'.(VAd.adequate))). + eapply VRirrelevant; eapply VAd.adequate. Qed. -Lemma irrelevanceSubstEq {Γ} (VΓ VΓ' : [||-v Γ]) {σ σ' Δ} (wfΔ wfΔ' : [|- Δ]) - (Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ]) (Vσ' : [Δ ||-v σ : Γ | VΓ' | wfΔ']) : - [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ | Vσ] -> [Δ ||-v σ ≅ σ' : Γ | VΓ' | wfΔ' | Vσ']. +Lemma irrelevanceSubstEq {Γ} (VΓ VΓ' : [||-v Γ]) {σ σ' Δ} (wfΔ wfΔ' : [|- Δ]) : + [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ ] -> [Δ ||-v σ ≅ σ' : Γ | VΓ' | wfΔ']. Proof. - apply (snd (VRirrelevant Γ VΓ.(VAd.adequate) VΓ'.(VAd.adequate))). + eapply VRirrelevant; eapply VAd.adequate. Qed. Set Printing Primitive Projection Parameters. -Lemma reflSubst {Γ} (VΓ : [||-v Γ]) : forall {σ Δ} (wfΔ : [|- Δ]) +(* Lemma reflSubst {Γ} (VΓ : [||-v Γ]) : forall {σ Δ} (wfΔ : [|- Δ]) (Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ]), [Δ ||-v σ ≅ σ : Γ | VΓ | wfΔ | Vσ]. Proof. @@ -51,56 +59,70 @@ Proof. - constructor. - intros * ih. unshelve econstructor. 1: apply ih. - apply reflLRTmEq. exact (validHead Vσ). -Qed. - -Lemma symmetrySubstEq {Γ} (VΓ VΓ' : [||-v Γ]) : forall {σ σ' Δ} (wfΔ wfΔ' : [|- Δ]) - (Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ]) (Vσ' : [Δ ||-v σ' : Γ | VΓ' | wfΔ']), - [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ | Vσ] -> [Δ ||-v σ' ≅ σ : Γ | VΓ' | wfΔ' | Vσ']. + exact (validHead Vσ). +Qed. *) + +Unset Printing Notations. +Lemma symmetrySubstEq@{u1 u2 u3 u4} {Γ} + (VΓ VΓ' : VAdequate@{u3 u4} Γ VR@{u1 u2 u3 u4}) : + (* (VΓ VΓ' : [||-v Γ]) : *) + forall {σ σ' Δ} (wfΔ wfΔ' : [|- Δ]), + [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ ] -> [Δ ||-v σ' ≅ σ : Γ | VΓ' | wfΔ']. Proof. revert VΓ'; pattern Γ, VΓ; apply validity_rect; clear Γ VΓ. - intros VΓ'. rewrite (invValidityEmpty VΓ'). constructor. - intros * ih VΓ'. pose proof (x := invValiditySnoc VΓ'). destruct x as [lA'[ VΓ'' [VA' ->]]]. - intros ????? [tl hd] [tl' hd'] [tleq hdeq]. + intros ????? [tleq hdeq]. unshelve econstructor. 1: now eapply ih. - eapply LRTmEqSym. cbn in *. - revert hdeq. apply LRTmEqRedConv. - eapply validTyExt. 2:eassumption. - eapply irrelevanceSubst; eassumption. + symmetry; revert hdeq; apply LRTmEqConv. + eapply validTyExt. Qed. +#[global] +Instance substEqSymmetric {Γ Δ} (VΓ : [||-v Γ]) (wfΔ : [|- Δ]) : + Symmetric (VΓ.(VPack.eqSubst) Δ wfΔ). +Proof. intros x y. eapply (symmetrySubstEq VΓ VΓ wfΔ wfΔ). Qed. + Lemma transSubstEq {Γ} (VΓ : [||-v Γ]) : - forall {σ σ' σ'' Δ} (wfΔ : [|- Δ]) - (Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ]) - (Vσ' : [Δ ||-v σ' : Γ | VΓ | wfΔ]), - [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ | Vσ] -> - [Δ ||-v σ' ≅ σ'' : Γ | VΓ | wfΔ | Vσ'] -> - [Δ ||-v σ ≅ σ'' : Γ | VΓ | wfΔ | Vσ]. + forall {σ σ' σ'' Δ} (wfΔ : [|- Δ]), + [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ ] -> + [Δ ||-v σ' ≅ σ'' : Γ | VΓ | wfΔ ] -> + [Δ ||-v σ ≅ σ'' : Γ | VΓ | wfΔ ]. Proof. pattern Γ, VΓ; apply validity_rect; clear Γ VΓ. - constructor. - intros * ih * [] []; unshelve econstructor. - 1: now eapply ih. - eapply transEqTerm; tea. - eapply LRTmEqRedConv; tea. + 1: now eapply ih. + etransitivity; [irrelevance|]. + eapply LRTmEqConv; tea. unshelve eapply LRTyEqSym; tea. 2: unshelve eapply validTyExt. - 7: eassumption. + 5: eassumption. 1: tea. - now eapply validTail. Qed. +#[global] +Instance substEqTransitive {Γ Δ} (VΓ : [||-v Γ]) (wfΔ : [|- Δ]) : + Transitive (VΓ.(VPack.eqSubst) Δ wfΔ). +Proof. intros x y z. eapply (transSubstEq VΓ wfΔ). Qed. + +#[global] +Instance substEqPER {Γ Δ} (VΓ : [||-v Γ]) (wfΔ : [|- Δ]) : + PER (VΓ.(VPack.eqSubst) Δ wfΔ). +Proof. constructor; typeclasses eauto. Qed. + + Lemma irrelevanceTy {Γ} : forall (VΓ VΓ' : [||-v Γ]) {l A}, [Γ ||-v A | VΓ] -> [Γ ||-v A | VΓ']. Proof. intros VΓ VΓ' l A [VA VAext]; unshelve econstructor; intros. - - unshelve eapply VA. 2: eapply irrelevanceSubst. all:eassumption. - - eapply VAext; [eapply irrelevanceSubst| eapply irrelevanceSubstEq]; eassumption. + - unshelve (eapply VA; now eapply irrelevanceSubstEq); tea. + - eapply VAext; eapply irrelevanceSubstEq; eassumption. Qed. -Lemma irrelevanceTy' {Γ Γ' A A' l} (VΓ : [||-v Γ]) (VΓ' : [||-v Γ']) (VA : [Γ ||-v A | VΓ]) : +Lemma irrelevanceTy' {Γ Γ' A A' l} (VΓ : [||-v Γ]) (VΓ' : [||-v Γ']) (VA : [Γ ||-v A | VΓ]) : A = A' -> Γ = Γ' -> [Γ' ||-v A' | VΓ']. Proof. intros eqA eqΓ; subst; now eapply irrelevanceTy. @@ -113,26 +135,20 @@ Lemma irrelevanceLift {l A F G Γ} (VΓ : [||-v Γ]) [Γ ,, G ||-v A | validSnoc VΓ VG]. Proof. intros [VA VAext]; unshelve econstructor. - - intros ??? [hd tl]. eapply VA. + - intros ???? [hd tl]. eapply VA. unshelve econstructor. 1: eassumption. - eapply LRTmRedConv. 2: eassumption. - eapply LRTyEqSym; unshelve eapply VFeqG; eassumption. - - intros ???? [??] [??] [??]. eapply VAext. - + unshelve econstructor. 1: eassumption. - eapply LRTmRedConv. 2: eassumption. - eapply LRTyEqSym; unshelve eapply VFeqG; eassumption. - + unshelve econstructor. 1: eassumption. - eapply LRTmEqRedConv. 2: eassumption. - eapply LRTyEqSym; unshelve eapply VFeqG; eassumption. + eapply LRTmEqConv. 2: eassumption. + eapply LRTyEqSym; unshelve eapply VFeqG; tea; now eapply lrefl. + - intros ???? [??]. eapply VAext. Qed. -Lemma irrelevanceTyEq {Γ l A B} (VΓ VΓ' : [||-v Γ]) (VA : [Γ ||-v A | VΓ]) (VA' : [Γ||-v A | VΓ']) : - [Γ ||-v< l > A ≅ B | VΓ | VA] -> [Γ ||-v< l > A ≅ B | VΓ' | VA']. +Lemma irrelevanceTyEq {Γ l l' A B} (VΓ VΓ' : [||-v Γ]) (VA : [Γ ||-v A | VΓ]) (VA' : [Γ||-v A | VΓ']) : + [Γ ||-v< l > A ≅ B | VΓ | VA] -> [Γ ||-v< l' > A ≅ B | VΓ' | VA']. Proof. intros [h]; constructor; intros. irrelevanceRefl. unshelve apply h. 1:eassumption. - eapply irrelevanceSubst; eassumption. + eapply irrelevanceSubstEq; eassumption. Qed. Lemma irrelevanceTyEq' {Γ l A A' B} (VΓ VΓ' : [||-v Γ]) (VA : [Γ ||-v A | VΓ]) (VA' : [Γ||-v A' | VΓ']) : A = A' -> @@ -141,33 +157,30 @@ Proof. intros ->; now eapply irrelevanceTyEq. Qed. -Lemma symValidTyEq {Γ l A B} {VΓ : [||-v Γ]} {VA : [Γ ||-v A | VΓ]} (VB : [Γ ||-v B | VΓ]) : - [Γ ||-v A ≅ B | VΓ | VA] -> [Γ ||-v B ≅ A | VΓ | VB]. +Lemma symValidTyEq {Γ l l' A B} {VΓ : [||-v Γ]} {VA : [Γ ||-v A | VΓ]} (VB : [Γ ||-v B | VΓ]) : + [Γ ||-v A ≅ B | VΓ | VA] -> [Γ ||-v B ≅ A | VΓ | VB]. Proof. intros; constructor; intros. eapply LRTyEqSym; now eapply validTyEq. - Unshelve. all: tea. + Unshelve. 1:tea. now symmetry. Qed. -Lemma transValidTyEq {Γ l A B C} {VΓ : [||-v Γ]} - {VA : [Γ ||-v A | VΓ]} {VB : [Γ ||-v B | VΓ]} : - [Γ ||-v A ≅ B | VΓ | VA] -> [Γ ||-v B ≅ C | VΓ | VB] -> [Γ ||-v A ≅ C | VΓ | VA]. +Lemma transValidTyEq {Γ l l' A B C} {VΓ : [||-v Γ]} + {VA : [Γ ||-v A | VΓ]} {VB : [Γ ||-v B | VΓ]} : + [Γ ||-v A ≅ B | VΓ | VA] -> [Γ ||-v B ≅ C | VΓ | VB] -> [Γ ||-v A ≅ C | VΓ | VA]. Proof. constructor; intros; eapply LRTransEq; now eapply validTyEq. - Unshelve. all: tea. + Unshelve. 1: tea. now eapply urefl. Qed. + Lemma irrelevanceTm'' {Γ l l' t A} (VΓ VΓ' : [||-v Γ]) (VA : [Γ ||-v A | VΓ]) (VA' : [Γ||-v A | VΓ']) : [Γ ||-v t : A | VΓ | VA] -> [Γ ||-v t : A | VΓ' | VA']. Proof. - intros [h1 h2]; unshelve econstructor. - - intros. irrelevanceRefl. - unshelve apply h1. 1:eassumption. - eapply irrelevanceSubst; eassumption. - - intros. irrelevanceRefl. - unshelve eapply h2. 1: eassumption. - 1,2: eapply irrelevanceSubst; eassumption. - eapply irrelevanceSubstEq; eassumption. + intros [h2]; unshelve econstructor. + intros. irrelevanceRefl. + unshelve apply h2. 1:eassumption. + now eapply irrelevanceSubstEq. Qed. Lemma irrelevanceTm' {Γ l l' t A A'} (VΓ VΓ' : [||-v Γ]) (VA : [Γ ||-v A | VΓ]) (VA' : [Γ||-v A' | VΓ']) : @@ -192,23 +205,13 @@ Lemma irrelevanceTmLift {l t A F G Γ} (VΓ : [||-v Γ]) [Γ ,, F ||-v t : A | validSnoc VΓ VF | VA] -> [Γ ,, G ||-v t : A | validSnoc VΓ VG | VA']. Proof. - intros [Vt Vtext]; unshelve econstructor. - - intros ??? [hd tl]. irrelevanceRefl. - unshelve eapply Vt; tea. - unshelve econstructor; tea. - eapply LRTmRedConv; tea. - eapply LRTyEqSym; unshelve eapply VFeqG; eassumption. - - intros ???? [??] [??] [??]. irrelevanceRefl. - unshelve eapply Vtext; tea. - + unshelve econstructor; tea. - eapply LRTmRedConv; tea. - eapply LRTyEqSym; unshelve eapply VFeqG; eassumption. - + unshelve econstructor; tea. - eapply LRTmRedConv; tea. - eapply LRTyEqSym; unshelve eapply VFeqG; eassumption. - + unshelve econstructor; tea. - eapply LRTmEqRedConv; tea. - eapply LRTyEqSym; unshelve eapply VFeqG; eassumption. + intros [Vtext]; unshelve econstructor. + intros ???? [hd tl]. irrelevanceRefl. + unshelve eapply Vtext; tea. + unshelve econstructor; tea. + eapply LRTmEqConv; tea. + eapply LRTyEqSym; unshelve eapply VFeqG; tea. + now eapply lrefl. Qed. Lemma irrelevanceTmEq {Γ l t u A} (VΓ VΓ' : [||-v Γ]) (VA : [Γ ||-v A | VΓ]) (VA' : [Γ||-v A | VΓ']) : @@ -216,7 +219,7 @@ Lemma irrelevanceTmEq {Γ l t u A} (VΓ VΓ' : [||-v Γ]) (VA : [Γ ||-v A | Proof. intros [h]; constructor; intros; irrelevanceRefl. unshelve apply h; tea. - eapply irrelevanceSubst; eassumption. + eapply irrelevanceSubstEq; eassumption. Qed. Lemma irrelevanceTmEq' {Γ l t u A A'} (VΓ VΓ' : [||-v Γ]) (VA : [Γ ||-v A | VΓ]) (VA' : [Γ||-v A' | VΓ']) : @@ -229,44 +232,61 @@ Lemma symValidTmEq {Γ l t u A} {VΓ : [||-v Γ]} {VA : [Γ ||-v A | VΓ]} : [Γ ||-v t ≅ u : A| VΓ | VA] -> [Γ ||-v u ≅ t : A | VΓ | VA]. Proof. intros; constructor; intros. - eapply LRTmEqSym; now eapply validTmEq. + eapply LRTmEqSym; eapply LRTmEqConv. + 2: now eapply validTmEq. + eapply validTyExt. + Unshelve. 1: tea. now symmetry. Qed. Lemma transValidTmEq {Γ l t u v A} {VΓ : [||-v Γ]} {VA : [Γ ||-v A | VΓ]} : - [Γ ||-v t ≅ u : A | VΓ | VA] -> - [Γ ||-v u ≅ v : A | VΓ | VA] -> + [Γ ||-v t ≅ u : A | VΓ | VA] -> + [Γ ||-v u ≅ v : A | VΓ | VA] -> [Γ ||-v t ≅ v : A | VΓ | VA]. Proof. - constructor; intros; eapply transEqTerm; now eapply validTmEq. + constructor; intros; etransitivity. + 1: now eapply validTmEq. + eapply LRTmEqConv. + 2: now eapply validTmEq. + eapply LRTyEqSym; eapply validTyExt. + Unshelve. 1,6-8:tea. now eapply urefl. Qed. -Lemma irrelevanceSubstExt {Γ} (VΓ : [||-v Γ]) {σ σ' Δ} (wfΔ : [|- Δ]) : +#[global] +Instance validTmEqSymmetric {Γ l A} (VΓ : [||-v Γ]) (VA : [Γ ||-v A | VΓ]) : + Symmetric (fun t u => [Γ ||-v t ≅ u : _ | VΓ | VA]). +Proof. intros x y; eapply symValidTmEq. Qed. + +#[global] +Instance validTmEqTransitive {Γ l A} (VΓ : [||-v Γ]) (VA : [Γ ||-v A | VΓ]) : + Transitive (fun t u => [Γ ||-v t ≅ u : _ | VΓ | VA]). +Proof. intros x y z. eapply transValidTmEq. Qed. + +(* Lemma irrelevanceSubstExt {Γ} (VΓ : [||-v Γ]) {σ σ' Δ} (wfΔ : [|- Δ]) : σ =1 σ' -> [Δ ||-v σ : Γ | VΓ | wfΔ] -> [Δ ||-v σ' : Γ | VΓ | wfΔ]. Proof. revert σ σ'; pattern Γ, VΓ; apply validity_rect; clear Γ VΓ. - constructor. - intros ????? ih ?? eq. unshelve econstructor. - + eapply ih. 2: now eapply validTail. + + eapply ih. 2: now eapply eqTail. now rewrite eq. + rewrite <- (eq var_zero). - pose proof (validHead X). + pose proof (eqHead X). irrelevance. now rewrite eq. -Qed. +Qed. *) Lemma irrelevanceSubstEqExt {Γ} (VΓ : [||-v Γ]) {σ1 σ1' σ2 σ2' Δ} - (wfΔ : [|- Δ]) (eq1 : σ1 =1 σ1') (eq2 : σ2 =1 σ2') - (Vσ1 : [Δ ||-v σ1 : Γ | VΓ | wfΔ]) : - [Δ ||-v σ1 ≅ σ2 : Γ | VΓ | wfΔ | Vσ1] -> - [Δ ||-v σ1' ≅ σ2' : Γ | VΓ | wfΔ | irrelevanceSubstExt VΓ wfΔ eq1 Vσ1]. + (wfΔ : [|- Δ]) (eq1 : σ1 =1 σ1') (eq2 : σ2 =1 σ2') : + [Δ ||-v σ1 ≅ σ2 : Γ | VΓ | wfΔ ] -> + [Δ ||-v σ1' ≅ σ2' : Γ | VΓ | wfΔ ]. Proof. - revert σ1 σ1' σ2 σ2' eq1 eq2 Vσ1; pattern Γ, VΓ; apply validity_rect; clear Γ VΓ. + revert σ1 σ1' σ2 σ2' eq1 eq2; pattern Γ, VΓ; apply validity_rect; clear Γ VΓ. - constructor. - - intros ????? ih ???? eq1 eq2 ? X. unshelve econstructor. + - intros ????? ih ???? eq1 eq2 X. unshelve econstructor. + eapply irrelevanceSubstEq. unshelve eapply ih. - 6: now eapply eqTail. - all: now (rewrite eq1 + rewrite eq2). + 5: now eapply eqTail. + all: now rewrite ?eq1, ?eq2. + rewrite <- (eq1 var_zero); rewrite <- (eq2 var_zero). pose proof (eqHead X). irrelevance. @@ -279,7 +299,7 @@ Ltac irrValid := match goal with | [_ : _ |- [||-v _]] => idtac | [_ : _ |- [ _ ||-v _ : _ | _ | _]] => eapply irrelevanceSubst - | [_ : _ |- [ _ ||-v _ ≅ _ : _ | _ | _ | _]] => eapply irrelevanceSubstEq + | [_ : _ |- [ _ ||-v _ ≅ _ : _ | _ | _ ]] => eapply irrelevanceSubstEq | [_ : _ |- [_ ||-v<_> _ | _]] => eapply irrelevanceTy | [_ : _ |- [_ ||-v<_> _ ≅ _ | _ | _]] => eapply irrelevanceTyEq | [_ : _ |- [_ ||-v<_> _ : _ | _ | _]] => eapply irrelevanceTm diff --git a/theories/Substitution/Properties.v b/theories/Substitution/Properties.v index 89ff10a..686691a 100644 --- a/theories/Substitution/Properties.v +++ b/theories/Substitution/Properties.v @@ -8,98 +8,58 @@ Set Universe Polymorphism. Section Properties. Context `{GenericTypingProperties}. - -Lemma wellformedSubst {Γ σ Δ} (VΓ : [||-v Γ]) (wfΔ : [|- Δ]) : - [Δ ||-v σ : Γ | VΓ| wfΔ] -> [Δ |-s σ : Γ ]. +Definition validTyEqLeft {Γ A B l} (VΓ : [||-v Γ]) + (VA : [Γ ||-v A | VΓ]) + (VB : [Γ ||-v B | VΓ]) + (VeqAB : [Γ ||-v A ≅ B | VΓ | VA]) + {Δ} (wfΔ : [|- Δ]) {σ σ'} + (Vσσ' : [VΓ | Δ ||-v σ ≅ σ' : Γ | wfΔ]) : + [validTy VA wfΔ Vσσ' | _ ||- _ ≅ B[σ]]. Proof. - revert σ; pattern Γ, VΓ; apply validity_rect; clear Γ VΓ. - - intros. apply well_sempty. - - intros * ih σ [tl hd]. - eapply well_scons. - + now apply ih. - + now escape. + eapply LRTransEq; [eapply validTyEq| eapply validTyExt]; tea. + Unshelve. + (* 6: now symmetry. *) + (* why is the instance not found here and found later in this file ? *) + 6: now eapply symmetrySubstEq. all: tea. Qed. -Lemma wellformedSubstEq {Γ σ σ' Δ} (VΓ : [||-v Γ]) (wfΔ : [|- Δ]) (Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ]) : - [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ | Vσ] -> [Δ |-s σ ≅ σ' : Γ]. + +Lemma wellformedSubstEq {Γ σ σ' Δ} (VΓ : [||-v Γ]) (wfΔ : [|- Δ]) : + [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ] -> [Δ |-s σ ≅ σ' : Γ]. Proof. - revert σ σ' Vσ; pattern Γ, VΓ; apply validity_rect; clear Γ VΓ. + revert σ σ'; pattern Γ, VΓ; apply validity_rect; clear Γ VΓ. - intros. apply conv_sempty. - - intros * ih ??? [tl hd]. apply conv_scons. + - intros * ih ?? []. apply conv_scons. + now eapply ih. + now escape. Qed. -Lemma consSubstS {Γ σ t l A Δ} (VΓ : [||-v Γ]) (wfΔ : [|- Δ]) - (Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ]) (VA : [Γ ||-v A | VΓ]) - (Vt : [ Δ ||- t : A[σ] | validTy VA wfΔ Vσ]) : - [Δ ||-v (t .: σ) : Γ ,, A | validSnoc VΓ VA | wfΔ]. -Proof. unshelve econstructor; eassumption. Defined. - - -Lemma consSubstSEq {Γ σ σ' t l A Δ} (VΓ : [||-v Γ]) (wfΔ : [|- Δ]) - (Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ]) - (Vσσ' : [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ | Vσ]) - (VA : [Γ ||-v A | VΓ]) - (Vt : [Δ ||- t : A[σ] | validTy VA wfΔ Vσ]) : - [Δ ||-v (t .: σ) ≅ (t .: σ') : Γ ,, A | validSnoc VΓ VA | wfΔ | consSubstS VΓ wfΔ Vσ VA Vt]. -Proof. - unshelve econstructor. - 1: eassumption. - eapply LRTmEqRefl; [apply (validTy VA wfΔ)| exact Vt]. -Qed. -Lemma consSubstSEq' {Γ σ σ' t u l A Δ} (VΓ : [||-v Γ]) (wfΔ : [|- Δ]) - (Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ]) - (Vσσ' : [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ | Vσ]) +Lemma consSubstEq {Γ σ σ' t u l A Δ} (VΓ : [||-v Γ]) (wfΔ : [|- Δ]) + (Vσσ' : [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ ]) (VA : [Γ ||-v A | VΓ]) - (Vt : [Δ ||- t : A[σ] | validTy VA wfΔ Vσ]) - (Vtu : [Δ ||- t ≅ u : A[σ] | validTy VA wfΔ Vσ]) : - [Δ ||-v (t .: σ) ≅ (u .: σ') : Γ ,, A | validSnoc VΓ VA | wfΔ | consSubstS VΓ wfΔ Vσ VA Vt]. + (Vtu : [Δ ||- t ≅ u : A[σ] | validTy VA wfΔ Vσσ']) : + [Δ ||-v (t .: σ) ≅ (u .: σ') : Γ ,, A | validSnoc VΓ VA | wfΔ ]. Proof. unshelve econstructor; tea. -Qed. - - -Lemma consSubstSvalid {Γ σ t l A Δ} {VΓ : [||-v Γ]} {wfΔ : [|- Δ]} - (Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ]) {VA : [Γ ||-v A | VΓ]} - (Vt : [ Γ ||-v t : A | VΓ | VA]) : - [Δ ||-v (t[σ] .: σ) : Γ ,, A | validSnoc VΓ VA | wfΔ]. -Proof. unshelve eapply consSubstS; tea; now eapply validTm. Defined. +Qed. Set Printing Primitive Projection Parameters. -Lemma consSubstSEqvalid {Γ σ σ' t l A Δ} {VΓ : [||-v Γ]} {wfΔ : [|- Δ]} - (Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ]) - (Vσ' : [Δ ||-v σ' : Γ | VΓ | wfΔ]) - (Vσσ' : [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ | Vσ]) +Lemma consSubstEqvalid {Γ σ σ' t u l A Δ} {VΓ : [||-v Γ]} {wfΔ : [|- Δ]} + (Vσσ' : [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ ]) {VA : [Γ ||-v A | VΓ]} - (Vt : [Γ ||-v t : A | VΓ | VA]) : - [Δ ||-v (t[σ] .: σ) ≅ (t[σ'] .: σ') : Γ ,, A | validSnoc VΓ VA | wfΔ | consSubstSvalid Vσ Vt]. + (Vt : [Γ ||-v t ≅ u : A | VΓ | VA]) : + [Δ ||-v (t[σ] .: σ) ≅ (u[σ'] .: σ') : Γ ,, A | validSnoc VΓ VA | wfΔ ]. Proof. unshelve econstructor; intros; tea. - now apply validTmExt. + now apply validTmEq. Qed. -Lemma wkSubstS {Γ} (VΓ : [||-v Γ]) : - forall {σ Δ Δ'} (wfΔ : [|- Δ]) (wfΔ' : [|- Δ']) (ρ : Δ' ≤ Δ), - [Δ ||-v σ : Γ | VΓ | wfΔ] -> [Δ' ||-v σ ⟨ ρ ⟩ : Γ | VΓ | wfΔ']. -Proof. - pattern Γ, VΓ ; apply validity_rect; clear Γ VΓ. - - constructor. - - intros * ih * [tl hd]. unshelve econstructor. - + eapply ih; eassumption. - + eapply LRTmRedIrrelevant'. - 2: eapply (wkTerm _ wfΔ') ; exact hd. - now asimpl. -Defined. - - Lemma wkSubstSEq {Γ} (VΓ : [||-v Γ]) : - forall {σ σ' Δ Δ'} (wfΔ : [|- Δ]) (wfΔ' : [|- Δ']) (ρ : Δ' ≤ Δ) - (Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ]), - [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ | Vσ] -> - [Δ' ||-v σ ⟨ ρ ⟩ ≅ σ' ⟨ ρ ⟩ : Γ | VΓ | wfΔ' | wkSubstS VΓ wfΔ wfΔ' ρ Vσ]. + forall {σ σ' Δ Δ'} (wfΔ : [|- Δ]) (wfΔ' : [|- Δ']) (ρ : Δ' ≤ Δ), + [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ ] -> + [Δ' ||-v σ ⟨ ρ ⟩ ≅ σ' ⟨ ρ ⟩ : Γ | VΓ | wfΔ' ]. Proof. pattern Γ, VΓ; apply validity_rect; clear Γ VΓ. - constructor. @@ -110,171 +70,85 @@ Proof. now asimpl. Qed. -Lemma wk1SubstS {Γ σ Δ F} (VΓ : [||-v Γ]) (wfΔ : [|- Δ]) (wfF : [Δ |- F]) : - [Δ ||-v σ : Γ | VΓ | wfΔ ] -> - [Δ ,, F ||-v σ ⟨ @wk1 Δ F ⟩ : Γ | VΓ | wfc_cons wfΔ wfF]. -Proof. eapply wkSubstS. Defined. Lemma wk1SubstSEq {Γ σ σ' Δ F} (VΓ : [||-v Γ]) - (wfΔ : [|- Δ]) (wfF : [Δ |- F]) - (Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ ]) : - [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ | Vσ] -> + (wfΔ : [|- Δ]) (wfF : [Δ |- F]) : + [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ ] -> let ρ := @wk1 Δ F in - [Δ ,, F ||-v σ ⟨ ρ ⟩ ≅ σ' ⟨ ρ ⟩ : Γ | VΓ | wfc_cons wfΔ wfF | wk1SubstS VΓ wfΔ wfF Vσ]. + [Δ ,, F ||-v σ ⟨ ρ ⟩ ≅ σ' ⟨ ρ ⟩ : Γ | VΓ | wfc_cons wfΔ wfF ]. Proof. intro vσσ'. eapply wkSubstSEq ; eassumption. Qed. -Lemma consWkSubstS {Γ F Δ Ξ σ a l VΓ wfΔ } VF - (ρ : Ξ ≤ Δ) wfΞ {RF}: - [Δ ||-v σ : Γ | VΓ | wfΔ] -> - [Ξ ||- a : F[σ]⟨ρ⟩ | RF] -> - [Ξ ||-v (a .: σ⟨ρ⟩) : Γ,, F | validSnoc (l:=l) VΓ VF | wfΞ]. -Proof. - intros. unshelve eapply consSubstS. 2: irrelevance. - now eapply wkSubstS. -Qed. - -Lemma consWkSubstSEq' {Γ Δ Ξ A σ σ' a b l} {VΓ : [||-v Γ]} {wfΔ : [|- Δ]} +Lemma consWkSubstEq {Γ Δ Ξ A σ σ' a b l} {VΓ : [||-v Γ]} {wfΔ : [|- Δ]} (VA : [Γ ||-v A | VΓ]) - (Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ]) - (Vσσ' : [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ | Vσ]) + (Vσσ' : [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ ]) (ρ : Ξ ≤ Δ) wfΞ {RA} - (Ra : [Ξ ||- a : A[σ]⟨ρ⟩ | RA]) - (Rab : [Ξ ||- a ≅ b : A[σ]⟨ρ⟩ | RA]) - (Vawkσ := consWkSubstS VA ρ wfΞ Vσ Ra) : - [Ξ ||-v (a .: σ⟨ρ⟩) ≅ (b .: σ'⟨ρ⟩) : Γ ,, A | validSnoc VΓ VA | wfΞ | Vawkσ]. + (Rab : [Ξ ||- a ≅ b : A[σ]⟨ρ⟩ | RA]) : + [Ξ ||-v (a .: σ⟨ρ⟩) ≅ (b .: σ'⟨ρ⟩) : Γ ,, A | validSnoc VΓ VA | wfΞ ]. Proof. - unshelve eapply consSubstSEq'. + unshelve eapply consSubstEq. - unshelve eapply irrelevanceSubstEq. - 4: now eapply wkSubstSEq. + 3: now eapply wkSubstSEq. tea. - irrelevance0; tea. now bsimpl. -Qed. - - -Lemma liftSubstS {Γ σ Δ lF F} (VΓ : [||-v Γ]) (wfΔ : [|- Δ]) - (VF : [Γ ||-v F | VΓ]) - (Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ ]) : - let VΓF := validSnoc VΓ VF in - let ρ := @wk1 Δ F[σ] in - let wfΔF := wfc_cons wfΔ (escape (validTy VF wfΔ Vσ)) in - [Δ ,, F[σ] ||-v (tRel 0 .: σ ⟨ ρ ⟩) : Γ ,, F | VΓF | wfΔF ]. -Proof. - intros; unshelve econstructor. - - now eapply wk1SubstS. - - eapply var0; unfold ρ; [now bsimpl|]. - now eapply escape, VF. -Defined. - -Lemma liftSubstSrealign {Γ σ σ' Δ lF F} {VΓ : [||-v Γ]} {wfΔ : [|- Δ]} - (VF : [Γ ||-v F | VΓ]) - {Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ ]} : - let VΓF := validSnoc VΓ VF in - let ρ := @wk1 Δ F[σ] in - let wfΔF := wfc_cons wfΔ (escape (validTy VF wfΔ Vσ)) in - [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ | Vσ] -> - [Δ ||-v σ' : Γ | VΓ | wfΔ ] -> - [Δ ,, F[σ] ||-v (tRel 0 .: σ'⟨ρ⟩) : Γ ,, F | VΓF | wfΔF]. -Proof. - intros; unshelve econstructor. - + now eapply wk1SubstS. - + cbn. - assert [Δ,, F[σ] |-[ ta ] tRel 0 : F[S >> (tRel 0 .: σ'⟨ρ⟩)]]. - { replace F[_ >> _] with F[σ']⟨S⟩ by (unfold ρ; now bsimpl). - eapply ty_conv. 1: apply (ty_var wfΔF (in_here _ _)). - cbn; renToWk. eapply convty_wk; tea. - eapply escapeEq; unshelve eapply validTyExt; cycle 3; tea. } - apply neuTerm; tea. - - apply convneu_var; tea. Qed. -Lemma liftSubstS' {Γ σ Δ lF F} {VΓ : [||-v Γ]} {wfΔ : [|- Δ]} - (VF : [Γ ||-v F | VΓ]) - (Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ ]) : - let VΓF := validSnoc VΓ VF in - let wfΔF := wfc_cons wfΔ (escape (validTy VF wfΔ Vσ)) in - [Δ ,, F[σ] ||-v up_term_term σ : Γ ,, F | VΓF | wfΔF ]. -Proof. - eapply irrelevanceSubstExt. - 2: eapply liftSubstS. - intros ?; now bsimpl. -Qed. -Lemma liftSubstSEq {Γ σ σ' Δ lF F} (VΓ : [||-v Γ]) (wfΔ : [|- Δ]) +Lemma liftSubstEq {Γ σ σ' Δ lF F} (VΓ : [||-v Γ]) (wfΔ : [|- Δ]) (VF : [Γ ||-v F | VΓ]) - (Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ ]) : + (Vσσ' : [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ ]) : let VΓF := validSnoc VΓ VF in let ρ := @wk1 Δ F[σ] in - let wfΔF := wfc_cons wfΔ (escape (validTy VF wfΔ Vσ)) in - let Vliftσ := liftSubstS VΓ wfΔ VF Vσ in - [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ | Vσ] -> - [Δ ,, F[σ] ||-v (tRel 0 .: σ ⟨ ρ ⟩) ≅ (tRel 0 .: σ' ⟨ ρ ⟩) : Γ ,, F | VΓF | wfΔF | Vliftσ]. + let wfΔF := wfc_cons wfΔ (escape (validTy VF wfΔ Vσσ')) in + [Δ ,, F[σ] ||-v (tRel 0 .: σ ⟨ ρ ⟩) ≅ (tRel 0 .: σ' ⟨ ρ ⟩) : Γ ,, F | VΓF | wfΔF ]. Proof. intros; unshelve econstructor. + now apply wk1SubstSEq. - + apply reflLRTmEq; exact (validHead Vliftσ). + + cbn. eapply var0; unfold ρ; [now bsimpl|]. + now eapply escape, VF. Qed. -Lemma liftSubstSEq' {Γ σ σ' Δ lF F} {VΓ : [||-v Γ]} {wfΔ : [|- Δ]} +Lemma liftSubstEq' {Γ σ σ' Δ lF F} {VΓ : [||-v Γ]} {wfΔ : [|- Δ]} (VF : [Γ ||-v F | VΓ]) - {Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ ]} : + (Vσ : [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ ]) : let VΓF := validSnoc VΓ VF in let ρ := wk_up F (@wk_id Γ) in let wfΔF := wfc_cons wfΔ (escape (validTy VF wfΔ Vσ)) in - let Vliftσ := liftSubstS' VF Vσ in - [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ | Vσ] -> - [Δ ,, F[σ] ||-v up_term_term σ ≅ up_term_term σ' : Γ ,, F | VΓF | wfΔF | Vliftσ]. + [Δ ,, F[σ] ||-v up_term_term σ ≅ up_term_term σ' : Γ ,, F | VΓF | wfΔF ]. Proof. intros. eapply irrelevanceSubstEq. unshelve eapply irrelevanceSubstEqExt. - 6: now eapply liftSubstSEq. - all: intros ?; now bsimpl. - Unshelve. all: tea. + 5: unshelve eapply liftSubstEq. + 6,8: tea. + 1-2: intros ?; now bsimpl. Qed. -Lemma liftSubstSrealign' {Γ σ σ' Δ lF F} {VΓ : [||-v Γ]} {wfΔ : [|- Δ]} - (VF : [Γ ||-v F | VΓ]) - {Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ ]} : - let VΓF := validSnoc VΓ VF in - let ρ := wk_up F (@wk_id Γ) in - let wfΔF := wfc_cons wfΔ (escape (validTy VF wfΔ Vσ)) in - [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ | Vσ] -> - [Δ ||-v σ' : Γ | VΓ | wfΔ ] -> - [Δ ,, F[σ] ||-v up_term_term σ' : Γ ,, F | VΓF | wfΔF]. -Proof. - intros. - eapply irrelevanceSubstExt. - 2: eapply liftSubstSrealign; tea. - intros ?; now bsimpl. -Qed. Lemma wk1ValidTy {Γ lA A lF F} {VΓ : [||-v Γ]} (VF : [Γ ||-v F | VΓ]) : - [Γ ||-v A | VΓ] -> + [Γ ||-v A | VΓ] -> [Γ ,, F ||-v A ⟨ @wk1 Γ F ⟩ | validSnoc VΓ VF ]. Proof. assert (forall σ, (A ⟨@wk1 Γ F⟩)[σ] = A[↑ >> σ]) as h by (intros; asimpl; now rewrite wk1_ren) ; intros [VA VAext]; unshelve econstructor. - - abstract (intros * [tl _]; rewrite h; exact (VA _ _ wfΔ tl)). - - intros * [tl _] [tleq _]. + - abstract (intros * [tl _]; rewrite h; exact (VA _ wfΔ _ _ tl)). + - intros ???? [tleq ?]. rewrite (h σ'); unshelve eapply LRTyEqSym. - 2: eapply VA; eassumption. - rewrite (h σ). - eapply VAext. 1: exact (validTail vσ). - eapply symmetrySubstEq. eassumption. + 2: eapply VA; now symmetry. + rewrite (h σ); eapply VAext. Qed. -Lemma wk1ValidTyEq {Γ lA A B lF F} {VΓ : [||-v Γ]} (VF : [Γ ||-v F | VΓ]) +Lemma wk1ValidTyEq {Γ lA A B lF F} {VΓ : [||-v Γ]} (VF : [Γ ||-v F | VΓ]) {VA : [Γ ||-v A | VΓ]} : - [Γ ||-v A ≅ B | VΓ | VA] -> + [Γ ||-v A ≅ B | VΓ | VA] -> [Γ ,, F ||-v A ⟨ @wk1 Γ F ⟩ ≅ B ⟨ @wk1 Γ F ⟩ | validSnoc VΓ VF | wk1ValidTy VF VA]. Proof. assert (forall A σ, (A ⟨@wk1 Γ F⟩)[σ] = A[↑ >> σ]) as h by (intros; asimpl; now rewrite wk1_ren). intros []; constructor; intros. rewrite h. irrelevance0. 1: symmetry; apply h. - unshelve intuition; tea; now eapply validTail. + unshelve intuition; tea; now eapply eqTail. Qed. Lemma wk1ValidTm {Γ lA t A lF F} {VΓ : [||-v Γ]} @@ -285,8 +159,7 @@ Lemma wk1ValidTm {Γ lA t A lF F} {VΓ : [||-v Γ]} Proof. assert (forall A σ, (A ⟨@wk1 Γ F⟩)[σ] = A[↑ >> σ]) as h by (intros; asimpl; now rewrite wk1_ren). constructor; intros; repeat rewrite h. - - instValid (validTail Vσ); irrelevance. - - instValidExt (validTail Vσ') (eqTail Vσσ'); irrelevance. + instValid (eqTail Vσσ'); irrelevance. Qed. Lemma wk1ValidTmEq {Γ lA t u A lF F} {VΓ : [||-v Γ]} @@ -297,7 +170,7 @@ Lemma wk1ValidTmEq {Γ lA t u A lF F} {VΓ : [||-v Γ]} Proof. assert (forall A σ, (A ⟨@wk1 Γ F⟩)[σ] = A[↑ >> σ]) as h by (intros; asimpl; now rewrite wk1_ren). constructor; intros; repeat rewrite h. - instValid (validTail Vσ); irrelevance. + instValid (eqTail Vσσ'); irrelevance. Qed. @@ -307,7 +180,7 @@ Lemma embValidTy@{u i j k l} {Γ l l' A} typeValidity@{u i j k l} Γ VΓ l' A (*[Γ ||-v A |VΓ]*). Proof. unshelve econstructor. - - intros ??? Vσ; destruct (validTy VA _ Vσ) as [pack]; exists pack. + - intros ???? Vσ; destruct (validTy VA _ Vσ) as [pack]; exists pack. eapply LR_embedding; tea. - intros; now eapply validTyExt. Defined. @@ -328,95 +201,50 @@ Proof. pattern Γ, VΓ; apply validity_rect; clear Γ VΓ. - exists ε, eq_refl, wfc_nil; constructor. - intros * [Δ [e [wfΔ Vid]]]. - exists (Δ,, A[tRel]); unshelve eexists. + exists (Δ,, A[tRel]); unshelve eexists. 1: asimpl; now rewrite e. unshelve eexists. + apply wfc_cons; tea. eapply escape. apply (validTy VA wfΔ Vid). - + eapply irrelevanceSubstExt. - 2: eapply irrelevanceSubst; now unshelve eapply liftSubstS. - intros []; [| bsimpl]; reflexivity. + + eapply irrelevanceSubstEqExt. + 3: eapply irrelevanceSubstEq. + 3: unshelve eapply liftSubstEq; cycle 2; tea. + all:intros []; [| bsimpl]; reflexivity. Qed. Definition soundCtx {Γ} (VΓ : [||-v Γ]) : [|-Γ] := (soundCtxId VΓ).π1. Definition idSubstS {Γ} (VΓ : [||-v Γ]) : [Γ ||-v tRel : Γ | VΓ | _] := (soundCtxId VΓ).π2. -Lemma reflIdSubstS {Γ} (VΓ : [||-v Γ]) : [Γ ||-v tRel ≅ tRel : Γ | VΓ | _ | idSubstS VΓ]. -Proof. apply reflSubst. Qed. - -Lemma substS_wk {Γ Δ} (ρ : Δ ≤ Γ) : +Lemma substEq_wk {Γ Δ} (ρ : Δ ≤ Γ) : forall (VΓ : [||-v Γ]) - (VΔ : [||-v Δ]) - {Ξ σ} (wfΞ : [|- Ξ]), [VΔ | Ξ ||-v σ : _ | wfΞ] -> [VΓ | Ξ ||-v ρ >> σ : _ | wfΞ]. + (VΔ : [||-v Δ]) + {Ξ σ σ'} (wfΞ : [|- Ξ]), [VΔ | Ξ ||-v σ ≅ σ' : _ | wfΞ] -> [VΓ | Ξ ||-v ρ >> σ ≅ ρ >> σ' : _ | wfΞ]. Proof. destruct ρ as [? wwk]; induction wwk. + intros; rewrite (invValidityEmpty VΓ); constructor. + intros. pose proof (invValiditySnoc VΔ) as [? [? [? eq]]]. rewrite eq in X; cbn in X; inversion X. - eapply irrelevanceSubstExt. + eapply irrelevanceSubstEqExt. 1: rewrite <- (scons_eta' σ); reflexivity. + 1: rewrite <- (scons_eta' σ'); reflexivity. cbn. asimpl. now eapply IHwwk. + intros. pose proof (invValiditySnoc VΔ) as [? [? [? eq]]]. rewrite eq in X; cbn in X; inversion X. - eapply irrelevanceSubstExt. + eapply irrelevanceSubstEqExt. 1:{ rewrite <- (scons_eta' σ); cbn; unfold up_ren; rewrite scons_comp'; cbn. reflexivity. } + 1:{ rewrite <- (scons_eta' σ'); cbn; unfold up_ren; rewrite scons_comp'; cbn. reflexivity. } asimpl. pose proof (invValiditySnoc VΓ) as [? [? [? eq']]]. - rewrite eq'. - unshelve eapply consSubstS. + rewrite eq'; unshelve econstructor. * now eapply IHwwk. * irrelevance. Defined. -Lemma substSEq_wk {Γ Δ} (ρ : Δ ≤ Γ) : - forall (VΓ : [||-v Γ]) - (VΔ : [||-v Δ]) - Ξ σ σ' (wfΞ : [|- Ξ]) - (Vσ : [VΔ | Ξ ||-v σ : _ | wfΞ]), - [VΔ | Ξ ||-v σ' : _ | wfΞ] -> - [VΔ | Ξ ||-v σ ≅ σ' : _ | wfΞ | Vσ] -> - [VΓ | Ξ ||-v ρ >> σ ≅ ρ >> σ' : _ | wfΞ | substS_wk ρ VΓ VΔ wfΞ Vσ]. -Proof. - destruct ρ as [? wwk]; induction wwk. - + intros; rewrite (invValidityEmpty VΓ); constructor. - + intros. - pose proof (invValiditySnoc VΔ) as [? [? [? eq]]]. - revert Vσ X X0; rewrite eq; intros Vσ Vσ' Vσσ'. - cbn; asimpl; eapply irrelevanceSubstEq. - unshelve eapply IHwwk; tea. - 1,2: now eapply validTail. - now eapply eqTail. - Unshelve. tea. -+ intros ??????. - set (ρ0 := {| well_wk := _ |}); unfold ρ0. - pose proof (invValiditySnoc VΔ) as [? [VΓ0 [? eq]]]. - pose proof (invValiditySnoc VΓ) as [? [VΔ0 [? eqΓ]]]. - rewrite eq; intros Vσ Vσ' Vσσ'. - assert (subst_eq : forall τ : nat -> term, τ var_zero .: (ρ >> ↑ >> τ) =1 (0 .: ρ >> S) >> τ). - 1:{ intros τ; asimpl; reflexivity. } - pose proof (v := substS_wk ρ0 VΓ _ wfΞ Vσ). - cbn; asimpl ; eapply irrelevanceSubstEq; unshelve eapply irrelevanceSubstEqExt. - 2,5: apply subst_eq. - - eapply irrelevanceSubstExt. - 1: symmetry; apply subst_eq. - exact v. - - eapply irrelevanceSubstEq. - eapply consSubstSEq'. - * exact (IHwwk VΔ0 VΓ0 Ξ (↑ >> σ) (↑ >> σ') wfΞ (validTail Vσ) (validTail Vσ') (eqTail Vσσ')). - * irrelevance0. - 2: now eapply eqHead. - now asimpl. - Unshelve. 2: tea. - rewrite eqΓ in v. - irrelevanceRefl. - eapply (validHead v). -Qed. - Lemma wkValidTy {l Γ Δ A} (ρ : Δ ≤ Γ) (VΓ : [||-v Γ]) (VΔ : [||-v Δ]) @@ -427,12 +255,11 @@ Proof. unshelve econstructor. - intros; rewrite h. eapply validTy; tea. - now eapply substS_wk. + now eapply substEq_wk. - intros; irrelevance0; rewrite h; [reflexivity|]. - eapply validTyExt. - 1: now eapply substS_wk. - now eapply substSEq_wk. - Unshelve. 2,3: tea. + unshelve eapply validTyExt. + 5: now eapply substEq_wk. + tea. Qed. Lemma wkValidTm {l Γ Δ A t} (ρ : Δ ≤ Γ) @@ -444,16 +271,10 @@ Lemma wkValidTm {l Γ Δ A t} (ρ : Δ ≤ Γ) Proof. assert (hA : forall σ, A⟨ρ⟩[σ] = A[ρ >> σ]) by (intros; now asimpl). assert (ht : forall σ, t⟨ρ⟩[σ] = t[ρ >> σ]) by (intros; now asimpl). - unshelve econstructor. - - intros; rewrite ht. - irrelevance0; [symmetry; apply hA|]. - eapply validTm, Vt. - - intros; do 2 rewrite ht. - irrelevance0; [symmetry; apply hA|]. - eapply validTmExt; [apply Vt|now eapply substS_wk|]. - now eapply substSEq_wk. - Unshelve. all: tea. - now eapply substS_wk. + econstructor; intros; rewrite 2!ht. + irrelevance0; [symmetry; apply hA|]. + eapply validTmExt, Vt. + Unshelve. 1: tea. now eapply substEq_wk. Qed. Lemma wkValidTyEq {l Γ Δ A B} (ρ : Δ ≤ Γ) @@ -467,9 +288,43 @@ Proof. unshelve econstructor; intros; irrelevance0; rewrite h; [reflexivity|]. now eapply validTyEq. Unshelve. 1: tea. - now eapply substS_wk. + now eapply substEq_wk. Qed. +Lemma subst_rel t : t[tRel] = t. +Proof. now bsimpl. Qed. + +Lemma validTyWf {Γ A l} (VΓ : [||-v Γ]) : [_ ||-v A | VΓ ] -> [Γ |- A]. +Proof. + intros h; generalize (validTy h _ (idSubstS VΓ)); rewrite subst_rel. + eapply escape. +Qed. + +Lemma validWf {Γ} (VΓ: [||-v Γ]) : [|- Γ]. +Proof. + induction Γ, VΓ using validity_rect. + gen_typing. + eapply wfc_cons; tea; now eapply validTyWf. +Qed. + +(* Lemma compSubstS {Γ} (VΓ : [||-v Γ]) : + forall Δ (VΔ : [||-v Δ]) + {σ σ'} (Vσσ' : [_ ||-v σ ≅ σ' : _ | VΓ | validWf VΔ]) + Ξ (wfΞ : [|- Ξ]) + {τ τ'} (Vττ' : [_ ||-v τ ≅ τ' : _ | VΔ | wfΞ]), + [_ ||-v σ >> subst_term τ ≅ σ' >> subst_term τ' : _ | VΓ | wfΞ]. +Proof. + induction Γ, VΓ using validity_rect. + 1: intros; constructor. + intros ; unshelve econstructor. + - eapply irrelevanceSubstEqExt. + 3:{ eapply IHVΓ; [eapply eqTail, Vσσ'|]; tea. } + 1,2: intros ?; reflexivity. + - cbn. unfold funcomp. cbv. bsimpl. substify. asimpl. + + intros Δ VΔ; induction Δ, VΔ using validity_rect. unshelve econstructor. *) + + diff --git a/theories/Substitution/Reduction.v b/theories/Substitution/Reduction.v index 192acdf..ad76c83 100644 --- a/theories/Substitution/Reduction.v +++ b/theories/Substitution/Reduction.v @@ -9,30 +9,16 @@ Context `{GenericTypingProperties}. Set Printing Primitive Projection Parameters. -Lemma redwfSubstValid {Γ A t u l} +Lemma redSubstValid {Γ A t u l} (VΓ : [||-v Γ]) - (red : [Γ ||-v t :⤳*: u : A | VΓ]) + (red : [Γ ||-v t ⤳* u : A | VΓ]) (VA : [Γ ||-v A | VΓ]) (Vu : [Γ ||-v u : A | VΓ | VA]) : - [Γ ||-v t : A | VΓ | VA] × [Γ ||-v t ≅ u : A | VΓ | VA]. + [Γ ||-v t ≅ u : A | VΓ | VA]. Proof. - assert (Veq : [Γ ||-v t ≅ u : A | VΓ | VA]). - { - constructor; intros; eapply redwfSubstTerm. - 1: now eapply validTm. - now eapply validRed. - } - split; tea; constructor; intros. - - eapply redwfSubstTerm. - 1: now eapply validTm. - now eapply validRed. - - eapply transEqTerm. 2: eapply transEqTerm. - + now eapply validTmEq. - + now eapply validTmExt. - + eapply LRTmEqSym. eapply LRTmEqRedConv. - 1: eapply LRTyEqSym; now eapply validTyExt. - now eapply validTmEq. - Unshelve. all: tea. + constructor; intros. eapply redSubstLeftTmEq. + 1: now eapply validTmExt. + now eapply validRed. Qed. - + End Reduction. \ No newline at end of file diff --git a/theories/Substitution/Reflexivity.v b/theories/Substitution/Reflexivity.v index 93c6a65..2a26403 100644 --- a/theories/Substitution/Reflexivity.v +++ b/theories/Substitution/Reflexivity.v @@ -1,25 +1,62 @@ From LogRel.AutoSubst Require Import core unscoped Ast Extra. From LogRel Require Import Utils BasicAst Notations Context NormalForms Weakening GenericTyping LogicalRelation Validity. -From LogRel.LogicalRelation Require Import Irrelevance Escape Reflexivity Weakening Neutral. +From LogRel.LogicalRelation Require Import Irrelevance EqRedRight. +From LogRel.Substitution Require Import Irrelevance Properties. Set Universe Polymorphism. Section Reflexivity. Context `{GenericTypingProperties}. -Lemma reflValidTy {Γ A l} (VΓ : [||-v Γ]) +Lemma reflValidTy {Γ A l} {VΓ : [||-v Γ]} (VA : [Γ ||-v A | VΓ]) : [Γ ||-v A ≅ A | VΓ | VA]. Proof. - constructor; intros; apply reflLRTyEq. + constructor; intros; eapply validTyExt. Qed. +Set Printing Primitive Projection Parameters. -Lemma reflValidTm {Γ t A l} (VΓ : [||-v Γ]) - (VA : [Γ ||-v A | VΓ]) +Lemma ureflValidTy {Γ A B l} {VΓ : [||-v Γ]} + {VA : [Γ ||-v A | VΓ]} + (VAB : [Γ ||-v A ≅ B | VΓ | VA]) : + [Γ ||-v B | VΓ ]. +Proof. + unshelve econstructor; intros. + 1: (unshelve now eapply LRTyEqRedRight, validTyEq); [tea| |now eapply lrefl]. + eapply LRTransEq. + 2: (unshelve now eapply validTyEq); cycle 2; tea. + eapply LRTyEqSym. + (unshelve now eapply validTyEq); tea; now eapply lrefl. +Qed. + +#[deprecated(note="Unneeded")] +Lemma reflValidTm {Γ t A l} {VΓ : [||-v Γ]} + {VA : [Γ ||-v A | VΓ]} (Vt : [Γ ||-v t : A | VΓ | VA]) : [Γ ||-v t ≅ t : A | VΓ | VA]. +Proof. exact Vt. Qed. + +Lemma lreflValidTm {Γ t u A l} (VΓ : [||-v Γ]) + {VA : [Γ ||-v A | VΓ]} + (Vtu : [Γ ||-v t ≅ u : A | VΓ | VA]) : + [Γ ||-v t : A | VΓ | VA]. +Proof. + constructor; intros. + etransitivity; [now eapply validTmEq|symmetry]. + eapply LRTmEqConv. + 2: now eapply validTmEq. + eapply LRTyEqSym. + eapply validTyExt. + Unshelve. 1,6-8:tea. now eapply urefl. +Qed. + +Lemma ureflValidTm {Γ t u A l} + {VΓ : [||-v Γ]} + {VA : [Γ ||-v A | VΓ]} + (Vtu : [Γ ||-v t ≅ u : A | VΓ | VA]) : + [Γ ||-v u : A | VΓ | VA]. Proof. - constructor; intros; apply reflLRTmEq; now eapply validTm. + eapply lreflValidTm; now eapply symValidTmEq. Qed. End Reflexivity. \ No newline at end of file diff --git a/theories/Substitution/SingleSubst.v b/theories/Substitution/SingleSubst.v index 69cda66..732b6ad 100644 --- a/theories/Substitution/SingleSubst.v +++ b/theories/Substitution/SingleSubst.v @@ -1,7 +1,7 @@ From LogRel.AutoSubst Require Import core unscoped Ast Extra. From LogRel Require Import Utils BasicAst Notations Context NormalForms Weakening GenericTyping LogicalRelation Validity. From LogRel.LogicalRelation Require Import Induction Irrelevance Escape Reflexivity Weakening Neutral Transitivity NormalRed. -From LogRel.Substitution Require Import Irrelevance Properties Conversion. +From LogRel.Substitution Require Import Irrelevance Properties Conversion Reflexivity. Set Universe Polymorphism. @@ -14,121 +14,70 @@ Lemma singleSubstComm G t σ : G[t..][σ] = G[t[σ] .: σ]. Proof. now asimpl. Qed. -Lemma substS {Γ F G t l} {VΓ : [||-v Γ]} +Lemma substS {Γ F G t t' l} {VΓ : [||-v Γ]} {VF : [Γ ||-v F | VΓ]} (VG : [Γ,, F ||-v G | validSnoc VΓ VF]) - (Vt : [Γ ||-v t : F | VΓ | VF]) : + (Vt : [Γ ||-v t ≅ t' : F | VΓ | VF]) : [Γ ||-v G[t..] | VΓ]. Proof. opector; intros; rewrite singleSubstComm. - - unshelve eapply validTy. 3,4: tea. - now eapply consSubstSvalid. + - unshelve eapply validTy. + 6: now eapply consSubstEqvalid. + tea. - irrelevance0. 1: symmetry; apply singleSubstComm. - eapply validTyExt. - 1: eapply consSubstS; now eapply validTm. - now eapply consSubstSEqvalid. - Unshelve. all: eassumption. + unshelve eapply validTyExt. + 5: now eapply consSubstEqvalid, lreflValidTm. + tea. Qed. -Lemma substSEq {Γ F F' G G' t t' l} {VΓ : [||-v Γ]} +Lemma substSEq {Γ F F' G G' t t' l} {VΓ : [||-v Γ]} {VF : [Γ ||-v F | VΓ]} - {VF' : [Γ ||-v F' | VΓ]} (VFF' : [Γ ||-v F ≅ F' | VΓ | VF]) (VΓF := validSnoc VΓ VF) - (VΓF' := validSnoc VΓ VF') {VG : [Γ,, F ||-v G | VΓF]} - (VG' : [Γ,, F' ||-v G' | VΓF']) (VGG' : [Γ ,, F ||-v G ≅ G' | VΓF | VG]) - (Vt : [Γ ||-v t : F | VΓ | VF]) - (Vt' : [Γ ||-v t' : F' | VΓ | VF']) (Vtt' : [Γ ||-v t ≅ t' : F | VΓ | VF]) - (VGt := substS VG Vt) : + (VGt := substS VG Vtt') : [Γ ||-v G[t..] ≅ G'[t'..] | VΓ | VGt]. Proof. constructor; intros. - assert (VtF' : [Γ ||-v t : F' | VΓ | VF']) by now eapply conv. - pose proof (consSubstSvalid Vσ Vt'). - pose proof (consSubstSvalid Vσ VtF'). rewrite singleSubstComm; irrelevance0. 1: symmetry; apply singleSubstComm. - eapply transEq. - - unshelve now eapply validTyEq. - 2: now eapply consSubstSvalid. - - eapply validTyExt; tea. - unshelve econstructor. - 1: eapply reflSubst. - eapply validTmEq. - now eapply convEq. - Unshelve. all: tea. + unshelve now eapply validTyEq. + 2: now eapply consSubstEqvalid. Qed. - - - -Lemma substSTm {Γ F G t f l} {VΓ : [||-v Γ]} - {VF : [Γ ||-v F | VΓ]} - (VΓF := validSnoc VΓ VF) - {VG : [Γ,, F ||-v G | VΓF]} - (Vt : [Γ ||-v t : F | VΓ | VF]) - (Vf : [Γ ,, F ||-v f : G | VΓF | VG]) - (VGt := substS VG Vt) : - [Γ ||-v f[t..] : G[t..] | VΓ | VGt]. -Proof. - constructor; intros; rewrite !singleSubstComm; irrelevance0. - 1,3: symmetry; apply singleSubstComm. - - now eapply validTm. - - eapply validTmExt; tea. - 1: now apply consSubstSvalid. - now apply consSubstSEqvalid. - Unshelve. 1,3: eassumption. - now apply consSubstSvalid. -Qed. - -Lemma substSTmEq {Γ F F' G G' t t' f f' l} (VΓ : [||-v Γ]) +Lemma substSTmEq {Γ F F' G G' t t' f f' l} (VΓ : [||-v Γ]) (VF : [Γ ||-v F | VΓ]) - (VF' : [Γ ||-v F' | VΓ]) (VFF' : [Γ ||-v F ≅ F' | VΓ | VF]) (VΓF := validSnoc VΓ VF) - (VΓF' := validSnoc VΓ VF') (VG : [Γ,, F ||-v G | VΓF]) - (VG' : [Γ,, F' ||-v G' | VΓF']) (VGG' : [Γ ,, F ||-v G ≅ G' | VΓF | VG]) - (Vt : [Γ ||-v t : F | VΓ | VF]) - (Vt' : [Γ ||-v t' : F' | VΓ | VF']) - (Vtt' : [Γ ||-v t ≅ t' : F | VΓ | VF]) - (Vf : [Γ ,, F ||-v f : G | VΓF | VG]) - (Vf' : [Γ ,, F' ||-v f' : G' | VΓF' | VG']) + (Vtt' : [Γ ||-v t ≅ t' : F | VΓ | VF]) (Vff' : [Γ ,, F ||-v f ≅ f' : G | VΓF | VG]) : - [Γ ||-v f[t..] ≅ f'[t'..] : G[t..] | VΓ | substS VG Vt]. + [Γ ||-v f[t..] ≅ f'[t'..] : G[t..] | VΓ | substS VG Vtt']. Proof. - constructor; intros; rewrite !singleSubstComm; irrelevance0. + constructor; intros; rewrite !singleSubstComm; irrelevance0. 1: symmetry; apply singleSubstComm. - eapply transEqTerm. - + unshelve now eapply validTmEq. - 2: now eapply consSubstSvalid. - + assert (Vσt : [Δ ||-v (t[σ] .: σ) : _ | VΓF' | wfΔ]) - by (eapply consSubstSvalid; tea; now eapply conv). - assert (Vσt' : [Δ ||-v (t'[σ] .: σ) : _ | VΓF' | wfΔ]) - by (eapply consSubstSvalid; tea; now eapply conv). - assert (Vσtσt' : [Δ ||-v (t[σ] .: σ) ≅ (t'[σ] .: σ) : _ | VΓF' | wfΔ | Vσt]). - 1:{ - constructor. - - bsimpl; epose (reflSubst _ _ Vσ); now eapply irrelevanceSubstEq. - - bsimpl; eapply validTmEq. now eapply convEq. - } - eapply LRTmEqRedConv. - 2: eapply (validTmExt Vf' _ Vσt Vσt' Vσtσt'). - eapply LRTyEqSym. now eapply validTyEq. - Unshelve. 2: now eapply consSubstSvalid. + unshelve now eapply validTmEq. + 1: tea. + now unshelve now eapply consSubstEq, validTmEq. Qed. -(* Skipping a series of lemmas on single substitution of a weakened term *) +Lemma substSTm {Γ F G t f l} {VΓ : [||-v Γ]} + {VF : [Γ ||-v F | VΓ]} + (VΓF := validSnoc VΓ VF) + {VG : [Γ,, F ||-v G | VΓF]} + (Vt : [Γ ||-v t : F | VΓ | VF]) + (Vf : [Γ ,, F ||-v f : G | VΓF | VG]) + (VGt := substS VG Vt) : + [Γ ||-v f[t..] : G[t..] | VΓ | VGt]. +Proof. eapply substSTmEq; tea; now eapply reflValidTy. Qed. Lemma liftSubstComm G t σ : G[t]⇑[σ] = G[t[σ] .: ↑ >> σ]. Proof. now bsimpl. Qed. - Lemma substLiftS {Γ F G t l} (VΓ : [||-v Γ]) (VF : [Γ ||-v F | VΓ]) (VΓF := validSnoc VΓ VF) @@ -137,30 +86,26 @@ Lemma substLiftS {Γ F G t l} (VΓ : [||-v Γ]) (Vt : [Γ,, F ||-v t : F⟨@wk1 Γ F⟩ | VΓF | VF']) : [Γ ,, F ||-v G[t]⇑ | VΓF]. Proof. - assert (h : forall Δ σ (wfΔ: [|- Δ]) - (vσ: [VΓF | Δ ||-v σ : Γ,, F | wfΔ]), - [VΓF | Δ ||-v (t[σ] .: ↑ >> σ) : _ | wfΔ ]). + assert (h : forall Δ σ σ' (wfΔ: [|- Δ]) + (vσσ': [VΓF | Δ ||-v σ ≅ σ' : Γ,, F | wfΔ]), + [VΓF | Δ ||-v (t[σ] .: ↑ >> σ) ≅ (t[σ'] .: ↑ >> σ') : _ | wfΔ ]). 1:{ unshelve econstructor. - + asimpl; now eapply validTail. + + asimpl; now eapply eqTail. + cbn. irrelevance0. - 2: now eapply validTm. + 2: now eapply validTmExt. now bsimpl. } opector; intros; rewrite liftSubstComm. - - unshelve eapply validTy; cycle 2; tea; now eapply h. + - unshelve eapply validTy. + 6: now eapply h. + tea. - irrelevance0. 2: unshelve eapply validTyExt. - 8: now eapply h. - 4: now eapply (h _ _ _ vσ). - 1: now bsimpl. - 1: tea. - constructor. - + asimpl; eapply irrelevanceSubstEq; now eapply eqTail. - + cbn. irrelevance0. - 2: now eapply validTmExt. - now bsimpl. - Unshelve. all:tea. + 7: now eapply h. + 2: tea. + now bsimpl. + Unshelve. all:tea. Qed. Lemma substLiftSEq {Γ F G G' t l} (VΓ : [||-v Γ]) @@ -174,10 +119,10 @@ Lemma substLiftSEq {Γ F G G' t l} (VΓ : [||-v Γ]) [Γ ,, F ||-v G[t]⇑ ≅ G'[t]⇑ | VΓF | substLiftS _ VF VG Vt]. Proof. constructor; intros; rewrite liftSubstComm. - assert (Vσt : [Δ ||-v (t[σ] .: ↑ >> σ) : _ | VΓF | wfΔ ]). 1:{ + assert (Vσt : [Δ ||-v (t[σ] .: ↑ >> σ) ≅ (t[σ'] .: ↑ >> σ') : _ | VΓF | wfΔ ]). 1:{ unshelve econstructor. - + bsimpl. now eapply validTail. - + bsimpl. instValid Vσ. irrelevance. + + bsimpl. now eapply eqTail. + + cbn; instValid Vσσ'; irrelevance. } instValid Vσt. irrelevance. Qed. @@ -189,102 +134,105 @@ Lemma substLiftSEq' {Γ F G G' t t' l} (VΓ : [||-v Γ]) (VG' : [Γ,, F ||-v G' | VΓF]) (VGeq : [Γ,, F ||-v G ≅ G' | VΓF | VG]) (VF' := wk1ValidTy VF VF) - (Vt : [Γ,, F ||-v t : F⟨@wk1 Γ F⟩ | VΓF | VF']) + (Vt : [Γ,, F ||-v t : F⟨@wk1 Γ F⟩ | VΓF | VF']) (Vt' : [Γ,, F ||-v t' : F⟨@wk1 Γ F⟩ | VΓF | VF']) (Vtt' : [Γ,, F ||-v t ≅ t' : F⟨@wk1 Γ F⟩ | VΓF | VF']) : [Γ ,, F ||-v G[t]⇑ ≅ G'[t']⇑ | VΓF | substLiftS _ VF VG Vt]. Proof. - eapply transValidTyEq. - 1: eapply substLiftSEq; [| exact VGeq]; tea. - constructor; intros; irrelevance0; rewrite liftSubstComm ; [reflexivity|]. - instValid Vσ. - eapply validTyExt. - + unshelve eapply consSubstS. - 6: now eapply validTail. - 3: exact VF. - irrelevance. - + unshelve eapply consSubstSEq'. - 1: now eapply validTail. - 1,3: irrelevance. - eapply reflSubst. - Unshelve. 3: tea. now eapply substLiftS. + constructor; intros; rewrite liftSubstComm. + assert (Vσt : [Δ ||-v (t[σ] .: ↑ >> σ) ≅ (t'[σ'] .: ↑ >> σ') : _ | VΓF | wfΔ ]). 1:{ + unshelve econstructor. + + bsimpl. now eapply eqTail. + + cbn; instValid Vσσ'; irrelevance. + } + instValid Vσt. irrelevance. Qed. -Lemma singleSubstPoly {Γ F G t l lF} +Lemma singleSubstPoly {Γ F G t u l lF} (RFG : PolyRed Γ l F G) {RF : [Γ ||- F]} - (Rt : [Γ ||- t : F | RF]) : + (Rt : [Γ ||- t ≅ u : F | RF]) : [Γ ||- G[t..]]. Proof. replace G[t..] with G[t .: wk_id (Γ:=Γ) >> tRel] by now bsimpl. unshelve eapply (PolyRed.posRed RFG). - 1: escape; gen_typing. - irrelevance0; tea. + 2: escape; gen_typing. + 2: irrelevance0; tea. now bsimpl. Qed. -Lemma singleSubstΠ1 {Γ F G t l lF} +Lemma singleSubstΠ1 {Γ F G t u l lF} (ΠFG : [Γ ||- tProd F G]) {RF : [Γ ||- F]} - (Rt : [Γ ||- t : F | RF]) : + (Rt : [Γ ||- t ≅ u : F | RF]) : [Γ ||- G[t..]]. Proof. eapply singleSubstPoly; tea. eapply (ParamRedTy.polyRed (normRedΠ0 (invLRΠ ΠFG))). Qed. -Lemma singleSubstΣ1 {Γ F G t l lF} +Lemma singleSubstΣ1 {Γ F G t u l lF} (ΠFG : [Γ ||- tSig F G]) {RF : [Γ ||- F]} - (Rt : [Γ ||- t : F | RF]) : + (Rt : [Γ ||- t ≅ u : F | RF]) : [Γ ||- G[t..]]. Proof. eapply singleSubstPoly; tea. eapply (ParamRedTy.polyRed (normRedΣ0 (invLRΣ ΠFG))). Qed. -Lemma singleSubstPoly2 {Γ F F' G G' t t' l lF lF'} +Lemma substId Γ t u : t[u..] = t[u .: wk_id (Γ:=Γ) >> tRel ]. +Proof. now bsimpl. Qed. + +Lemma singleSubstPoly2 {Γ F F' G G' t t' l lF} {RFG : PolyRed Γ l F G} (RFGeq : PolyRedEq RFG F' G') {RF : [Γ ||- F]} - {RF' : [Γ ||- F']} - (Rt : [Γ ||- t : F | RF]) - (Rt' : [Γ ||- t' : F' | RF']) (Rteq : [Γ ||- t ≅ t' : F | RF]) - (RGt : [Γ ||- G[t..]]) - (RGt' : [Γ ||- G'[t'..]]) : + (RGt : [Γ ||- G[t..]]) : [Γ ||- G[t..] ≅ G'[t'..] | RGt ]. Proof. assert (wfΓ : [|-Γ]) by (escape ; gen_typing). - assert [Γ ||- t' : F⟨wk_id (Γ:=Γ)⟩ | PolyRed.shpRed RFG wk_id wfΓ]. - { - eapply LRTmRedConv; tea. - eapply LRTyEqSym. - replace F' with F'⟨wk_id (Γ := Γ)⟩ by now bsimpl. - eapply (PolyRedEq.shpRed RFGeq). - } + rewrite (substId Γ). + irrelevance0; [now rewrite (substId Γ)|]. eapply transEq. - 2: (replace G'[t'..] with G'[t' .: wk_id (Γ:=Γ) >> tRel] by now bsimpl); eapply (PolyRedEq.posRed RFGeq). - irrelevance0. - 2: eapply (PolyRed.posExt RFG). - 3: irrelevance0; tea; now bsimpl. - 1: now bsimpl. - eassumption. - Unshelve. all: tea. - irrelevance0; tea; now bsimpl. + + unshelve eapply PolyRed.posExt. + 2,4: tea. + 2: irrelevance. + + unshelve eapply (PolyRedEq.posRed RFGeq). + 2: tea. + 2: irrelevance0; [| now eapply urefl]; now bsimpl. Qed. -Lemma singleSubstΠ2 {Γ F F' G G' t t' l lF lF'} +Lemma singleSubstEqPoly {Γ F G t t' l lF} + {RFG : PolyRed Γ l F G} + {RF : [Γ ||- F]} + (Rteq : [Γ ||- t ≅ t' : F | RF]) + (RGt : [Γ ||- G[t..]]) : + [Γ ||- G[t..] ≅ G[t'..] | RGt ]. +Proof. + eapply (singleSubstPoly2 (F':=F) (RFG:=RFG)); tea. + destruct RFG; econstructor; cbn; tea; intros; eapply reflLRTyEq. +Qed. + +Lemma singleSubstEqΣ {Γ F G t t' l lF} + {RFG : [Γ ||- tSig F G]} + {RF : [Γ ||- F]} + (Rteq : [Γ ||- t ≅ t' : F | RF]) + (RGt : [Γ ||- G[t..]]) : + [Γ ||- G[t..] ≅ G[t'..] | RGt ]. +Proof. + unshelve (eapply singleSubstEqPoly; tea). + 2: exact (normRedΣ0 (invLRΣ RFG)). +Qed. + +Lemma singleSubstΠ2 {Γ F F' G G' t t' l lF} {ΠFG : [Γ ||- tProd F G]} (ΠFGeq : [Γ ||- tProd F G ≅ tProd F' G' | ΠFG]) {RF : [Γ ||- F]} - {RF' : [Γ ||- F']} - (Rt : [Γ ||- t : F | RF]) - (Rt' : [Γ ||- t' : F' | RF']) (Rteq : [Γ ||- t ≅ t' : F | RF]) - (RGt : [Γ ||- G[t..]]) - (RGt' : [Γ ||- G'[t'..]]) : + (RGt : [Γ ||- G[t..]]) : [Γ ||- G[t..] ≅ G'[t'..] | RGt ]. Proof. eapply singleSubstPoly2; tea. @@ -296,25 +244,25 @@ Proof. exact polyRed. Qed. -Lemma substSΠaux {Γ F G t l} +Lemma substSΠaux {Γ F G t u l} {VΓ : [||-v Γ]} {VF : [Γ ||-v F | VΓ]} (VΠFG : [Γ ||-v tProd F G | VΓ]) - (Vt : [Γ ||-v t : F | VΓ | VF]) - (Δ : context) (σ : nat -> term) - (wfΔ : [ |-[ ta ] Δ]) (vσ : [VΓ | Δ ||-v σ : Γ | wfΔ]) : + (Vt : [Γ ||-v t ≅ u : F | VΓ | VF]) + (Δ : context) (σ σ' : nat -> term) + (wfΔ : [ |-[ ta ] Δ]) (vσ : [VΓ | Δ ||-v σ ≅ σ' : Γ | wfΔ]) : [Δ ||-< l > G[up_term_term σ][t[σ]..]]. Proof. eapply singleSubstΠ1. eapply (validTy VΠFG); tea. - now eapply validTm. - Unshelve. all: eassumption. + unshelve now eapply validTmEq. + 1,3: tea. Qed. Lemma singleSubstComm' G t σ : G[t..][σ] = G[up_term_term σ][t[σ]..]. Proof. now asimpl. Qed. -Lemma substSΠ {Γ F G t l} +Lemma substSΠ {Γ F G t l} {VΓ : [||-v Γ]} {VF : [Γ ||-v F | VΓ]} (VΠFG : [Γ ||-v tProd F G | VΓ]) @@ -322,44 +270,39 @@ Lemma substSΠ {Γ F G t l} [Γ ||-v G[t..] | VΓ]. Proof. opector; intros. - - rewrite singleSubstComm'; now eapply substSΠaux. + - rewrite singleSubstComm'. now eapply substSΠaux. - rewrite singleSubstComm'. irrelevance0. 1: symmetry; apply singleSubstComm'. eapply singleSubstΠ2. 1: eapply (validTyExt VΠFG). - 1, 2: tea. - 1, 2: now eapply validTm. 1: now eapply validTmExt. - now eapply substSΠaux. Unshelve. all: tea. now eapply substSΠaux. Qed. -Lemma substSΠeq {Γ F F' G G' t u l} + +Lemma substSΠeq {Γ F F' G G' t u l} {VΓ : [||-v Γ]} {VF : [Γ ||-v F | VΓ]} {VF' : [Γ ||-v F' | VΓ]} {VΠFG : [Γ ||-v tProd F G | VΓ]} - (VΠFG' : [Γ ||-v tProd F' G' | VΓ]) (VΠFGeq : [Γ ||-v tProd F G ≅ tProd F' G' | VΓ | VΠFG]) - (Vt : [Γ ||-v t : F | VΓ | VF]) - (Vu : [Γ ||-v u : F' | VΓ | VF']) - (Vtu : [Γ ||-v t ≅ u : F | VΓ | VF]) + (Vt : [Γ ||-v t : F | VΓ | VF]) + (Vu : [Γ ||-v u : F' | VΓ | VF']) + (Vtu : [Γ ||-v t ≅ u : F | VΓ | VF]) (VGt := substSΠ VΠFG Vt) : [Γ ||-v G[t..] ≅ G'[u..] | VΓ | VGt]. Proof. + pose proof (ureflValidTy VΠFGeq). constructor; intros. rewrite singleSubstComm'. irrelevance0. 1: symmetry; apply singleSubstComm'. eapply singleSubstΠ2. 1: now eapply (validTyEq VΠFGeq). - 3: now eapply validTmEq. - 1,2: now eapply validTm. - now eapply substSΠaux. + 1: now eapply validTmEq. Unshelve. all: tea. now eapply substSΠaux. Qed. - End SingleSubst. \ No newline at end of file diff --git a/theories/TypeConstructorsInj.v b/theories/TypeConstructorsInj.v index d5f9e1e..d4a8fef 100644 --- a/theories/TypeConstructorsInj.v +++ b/theories/TypeConstructorsInj.v @@ -5,7 +5,7 @@ From LogRel Require Import Utils BasicAst Notations Context NormalForms Weakenin GenericTyping DeclarativeTyping DeclarativeInstance. From LogRel Require Import LogicalRelation Validity Fundamental DeclarativeSubst. From LogRel.LogicalRelation Require Import Escape Irrelevance Neutral Induction NormalRed. -From LogRel.Substitution Require Import Escape Poly. +From LogRel.Substitution Require Import Escape Poly Reflexivity. Set Printing Primitive Projection Parameters. @@ -53,15 +53,15 @@ Section TypeConstructors. clearbody HTred. clear HT. eapply reducibleTy in HT'. - revert nfT T' nfT' HΓ HT' Hconv. - revert HTred. + revert nfT T' nfT' HΓ HT' Hconv. + revert HTred. generalize (eq_refl : one = one). generalize one at 1 3; intros l eql HTred; revert eql. pattern l, Γ, T, HTred; apply LR_rect_TyUr; clear l Γ T HTred. all: intros ? Γ T. - intros [] -> nfT T' nfT' HΓ HT' []. assert (T' = U) as HeqT' by (eapply redtywf_whnf ; gen_typing); subst. - assert (T = U) as HeqU by (eapply redtywf_whnf ; gen_typing). + assert (T = U) as HeqU by (eapply redtywf_whnf ; gen_typing). destruct nfT ; inversion HeqU ; subst. 2: now exfalso ; gen_typing. clear HeqU. @@ -81,7 +81,7 @@ Section TypeConstructors. 1-6: symmetry in ne'; apply convneu_whne in ne'; inversion ne'. cbn. gen_typing. - intros [dom cod red] _ _ -> nfT T' nfT' HΓ HT'[dom' cod' red']; cbn in *. - assert (T = tProd dom cod) as HeqT by (apply red_whnf ; gen_typing). + assert (T = tProd dom cod) as HeqT by (apply red_whnf ; gen_typing). assert (T' = tProd dom' cod') as HeqT' by (apply red_whnf ; gen_typing). destruct nfT; cycle -1. 1: subst ; exfalso ; gen_typing. @@ -109,7 +109,7 @@ Section TypeConstructors. * exfalso; subst; inversion w. + exfalso; subst; inversion w. - intros [dom cod red] _ _ -> nfT T' nfT' HΓ HT' [dom' cod' red']. - assert (T = tSig dom cod) as HeqT by (apply red_whnf ; gen_typing). + assert (T = tSig dom cod) as HeqT by (apply red_whnf ; gen_typing). assert (T' = tSig dom' cod') as HeqT' by (apply red_whnf ; gen_typing). destruct nfT; cycle -1. 1: subst; inv_whne. @@ -214,7 +214,7 @@ Section TypeConstructors. pose proof HT as HT'. unshelve eapply red_ty_complete in HT as (T''&[? nfT]). 2: econstructor. - assert [Γ |- tProd A B ≅ T''] as Hconv by + assert [Γ |- tProd A B ≅ T''] as Hconv by (etransitivity ; [eassumption|now eapply RedConvTyC]). unshelve eapply ty_conv_inj in Hconv. 1: constructor. @@ -242,7 +242,7 @@ Section TypeConstructors. pose proof HT as HT'. unshelve eapply red_ty_complete in HT as (T''&[? nfT]). 2: econstructor. - assert [Γ |- tSig A B ≅ T''] as Hconv by + assert [Γ |- tSig A B ≅ T''] as Hconv by (etransitivity ; [eassumption|now eapply RedConvTyC]). unshelve eapply ty_conv_inj in Hconv. 1: constructor. @@ -251,7 +251,7 @@ Section TypeConstructors. do 2 eexists ; split. all: eassumption. Qed. - + Corollary red_ty_compl_sig_r Γ A B T : [Γ |- T ≅ tSig A B] -> ∑ A' B', [× [Γ |- T ⤳* tSig A' B'], [Γ |- A' ≅ A] & [Γ,, A' |- B ≅ B']]. @@ -277,7 +277,7 @@ Section TypeConstructors. pose proof HT as HT'. unshelve eapply red_ty_complete in HT as (T''&[? nfT]). 2: econstructor. - assert [Γ |- tId A x y ≅ T''] as Hconv by + assert [Γ |- tId A x y ≅ T''] as Hconv by (etransitivity ; [eassumption|now eapply RedConvTyC]). unshelve eapply ty_conv_inj in Hconv. 1: constructor. @@ -285,7 +285,7 @@ Section TypeConstructors. destruct nfT, Hconv. do 3 eexists ; split; eassumption. Qed. - + Corollary red_ty_compl_id_r Γ A x y T : [Γ |- T ≅ tId A x y] -> ∑ A' x' y', [× [Γ |- T ⤳* tId A' x' y'], [Γ |- A' ≅ A], [Γ |- x ≅ x' : A] & [Γ |- y ≅ y' : A]]. @@ -322,7 +322,7 @@ Section Boundary. now eapply typing_wk. Qed. - Lemma scons2_well_subst {Γ A B} : + Lemma scons2_well_subst {Γ A B} : (forall a b, [Γ |- a : A] -> [Γ |- b : B[a..]] -> [Γ |-s (b .: a ..) : (Γ ,, A),, B]) × (forall a a' b b', [Γ |- a ≅ a' : A] -> [Γ |- b ≅ b' : B[a..]] -> [Γ |-s (b .: a..) ≅ (b' .: a'..) : (Γ ,, A),, B]). Proof. @@ -350,7 +350,7 @@ Section Boundary. Lemma shift_subst_eq t a : t⟨↑⟩[a..] = t. Proof. now asimpl. Qed. - Lemma idElimMotiveCtx {Γ A x} : + Lemma idElimMotiveCtx {Γ A x} : [Γ |- A] -> [Γ |- x : A] -> [|- (Γ,, A),, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0)]. @@ -359,7 +359,7 @@ Section Boundary. assert [|- Γ,, A] by now econstructor. econstructor; tea. econstructor. - 1: now eapply wft_wk. + 1: now eapply wft_wk. 1: eapply ty_wk; tea; econstructor; tea. rewrite wk1_ren_on; now eapply ty_var0. Qed. @@ -398,7 +398,9 @@ Section Boundary. - intros Γ A B H; apply Fundamental in H as [? VA VB _]; split. + now eapply escapeTy. + now eapply escapeTy. - - intros Γ A t u H; apply Fundamental in H as [? VA Vt Vu]; prod_splitter. + - intros Γ A t u H; apply Fundamental in H as [? VA Vtu]. + pose proof (ureflValidTm Vtu). + prod_splitter. + now eapply escapeTy. + now eapply escapeTm. + now eapply escapeTm. @@ -548,7 +550,7 @@ End Stability. (** ** Generation *) -(** The generation lemma (the name comes from the PTS literature), gives a +(** The generation lemma (the name comes from the PTS literature), gives a stronger inversion principle on typing derivations, that give direct access to the last non-conversion rule, and bundle together all conversions. @@ -580,7 +582,7 @@ Definition termGenData (Γ : context) (t T : term) : Type := | tSnd p => ∑ A B, T = B[(tFst p)..] × [Γ |- p : tSig A B] | tId A x y => [× T = U, [Γ |- A : U], [Γ |- x : A] & [Γ |- y : A]] | tRefl A x => [× T = tId A x x, [Γ |- A] & [Γ |- x : A]] - | tIdElim A x P hr y e => + | tIdElim A x P hr y e => [× T = P[e .: y..], [Γ |- A], [Γ |- x : A], [Γ,, A,, tId A⟨@wk1 Γ A⟩ x⟨@wk1 Γ A⟩ (tRel 0) |- P], [Γ |- hr : P[tRefl A x .: x..]], [Γ |- y : A] & [Γ |- e : tId A x y]] end. @@ -801,7 +803,7 @@ Proof. all: try now econstructor. all: try now cbn in Hconv. Qed. - + Lemma type_isType Γ A : [Γ |-[de] A] -> whnf A -> diff --git a/theories/Validity.v b/theories/Validity.v index 15d21eb..0e1b6e9 100644 --- a/theories/Validity.v +++ b/theories/Validity.v @@ -1,5 +1,5 @@ From LogRel.AutoSubst Require Import core unscoped Ast Extra. -From LogRel Require Import Utils BasicAst Notations Context NormalForms Weakening GenericTyping LogicalRelation. +From LogRel Require Import Utils BasicAst Notations Context NormalForms Weakening GenericTyping LogicalRelation EqRedRight. Set Primitive Projections. Set Universe Polymorphism. @@ -21,8 +21,8 @@ Ltac substitution := eauto with substitution. Definition VRel@{i j | i < j +} `{ta : tag} `{!WfContext ta} := forall (Γ : context) - (validSubst : forall (Δ : context) (σ : nat -> term) (wfΔ : [|- Δ]), Type@{i}) - (eqSubst : forall (Δ : context) (σ σ' : nat -> term) (wfΔ : [|- Δ ]), validSubst Δ σ wfΔ -> Type@{i}) + (* (validSubst : forall (Δ : context) (σ : nat -> term) (wfΔ : [|- Δ]), Type@{i}) *) + (eqSubst : forall (Δ : context) (wfΔ : [|- Δ ]) (σ σ' : nat -> term) , Type@{i}) , Type@{j}. (* A VPack contains the data corresponding to the codomain of VRel seen as a functional relation *) @@ -31,8 +31,8 @@ Module VPack. Record VPack@{i} `{ta : tag} `{!WfContext ta} {Γ : context} := { - validSubst : forall (Δ : context) (σ : nat -> term) (wfΔ : [|- Δ]), Type@{i} ; - eqSubst : forall (Δ : context) (σ σ' : nat -> term) (wfΔ : [|- Δ ]), validSubst Δ σ wfΔ -> Type@{i} ; + (* validSubst : forall (Δ : context) (σ : nat -> term) (wfΔ : [|- Δ]), Type@{i} ; *) + eqSubst : forall (Δ : context) (wfΔ : [|- Δ ]) (σ σ' : nat -> term) , Type@{i} ; }. Arguments VPack : clear implicits. @@ -42,13 +42,13 @@ End VPack. Export VPack(VPack,Build_VPack). -Notation "[ P | Δ ||-v σ : Γ | wfΔ ]" := (@VPack.validSubst _ _ Γ P Δ σ wfΔ) (at level 0, P, Δ, σ, Γ, wfΔ at level 50). -Notation "[ P | Δ ||-v σ ≅ σ' : Γ | wfΔ | vσ ]" := (@VPack.eqSubst _ _ Γ P Δ σ σ' wfΔ vσ) (at level 0, P, Δ, σ, σ', Γ, wfΔ, vσ at level 50). +Notation "[ P | Δ ||-v σ : Γ | wfΔ ]" := (@VPack.eqSubst _ _ Γ P Δ wfΔ σ σ) (at level 0, P, Δ, σ, Γ, wfΔ at level 50). +Notation "[ P | Δ ||-v σ ≅ σ' : Γ | wfΔ ]" := (@VPack.eqSubst _ _ Γ P Δ wfΔ σ σ') (at level 0, P, Δ, σ, σ', Γ, wfΔ at level 50). (* An VPack it adequate wrt. a VRel when its three unpacked components are *) #[universes(polymorphic)] Definition VPackAdequate@{i j} `{ta : tag} `{!WfContext ta} {Γ : context} (R : VRel@{i j}) (P : VPack@{i} Γ) : Type@{j} := - R Γ P.(VPack.validSubst) P.(VPack.eqSubst). + R Γ P.(VPack.eqSubst). Arguments VPackAdequate _ _ /. @@ -73,8 +73,8 @@ Coercion VAd.pack : VAdequate >-> VPack. Coercion VAd.adequate : VAdequate >-> VPackAdequate. Notation "[ R | ||-v Γ ]" := (VAdequate Γ R) (at level 0, R, Γ at level 50). -Notation "[ R | Δ ||-v σ : Γ | RΓ | wfΔ ]" := (RΓ.(@VAd.pack _ _ Γ R).(VPack.validSubst) Δ σ wfΔ) (at level 0, R, Δ, σ, Γ, RΓ, wfΔ at level 50). -Notation "[ R | Δ ||-v σ ≅ σ' : Γ | RΓ | wfΔ | vσ ]" := (RΓ.(@VAd.pack _ _ Γ R).(VPack.eqSubst) Δ σ σ' wfΔ vσ) (at level 0, R, Δ, σ, σ', Γ, RΓ, wfΔ, vσ at level 50). +Notation "[ R | Δ ||-v σ : Γ | RΓ | wfΔ ]" := (RΓ.(@VAd.pack _ _ Γ R).(VPack.eqSubst) Δ wfΔ σ σ) (at level 0, R, Δ, σ, Γ, RΓ, wfΔ at level 50). +Notation "[ R | Δ ||-v σ ≅ σ' : Γ | RΓ | wfΔ ]" := (RΓ.(@VAd.pack _ _ Γ R).(VPack.eqSubst) Δ wfΔ σ σ') (at level 0, R, Δ, σ, σ', Γ, RΓ, wfΔ at level 50). Record typeValidity@{u i j k l} `{ta : tag} `{!WfContext ta} `{!WfType ta} `{!Typing ta} `{!ConvType ta} @@ -82,16 +82,13 @@ Record typeValidity@{u i j k l} `{ta : tag} `{!WfContext ta} {Γ : context} {VΓ : VPack@{u} Γ} {l : TypeLevel} {A : term} := { - validTy : forall {Δ : context} {σ : nat -> term} - (wfΔ : [|- Δ ]) - (vσ : [ VΓ | Δ ||-v σ : Γ | wfΔ ]) + validTy : forall {Δ : context} (wfΔ : [|- Δ ]) {σ σ' : nat -> term} + (vσ : [ VΓ | Δ ||-v σ ≅ σ' : Γ | wfΔ ]) , [LogRel@{i j k l} l | Δ ||- A [ σ ] ] ; - validTyExt : forall {Δ : context} {σ σ' : nat -> term} - (wfΔ : [|- Δ ]) - (vσ : [ VΓ | Δ ||-v σ : Γ | wfΔ ]) - (vσ' : [ VΓ | Δ ||-v σ' : Γ | wfΔ ]) - (vσσ' : [ VΓ | Δ ||-v σ ≅ σ' : Γ | wfΔ | vσ ]) - , [LogRel@{i j k l} l | Δ ||- A [ σ ] ≅ A [ σ' ] | validTy wfΔ vσ ] + validTyExt : forall {Δ : context}(wfΔ : [|- Δ ]) + {σ σ' : nat -> term} + (vσσ' : [ VΓ | Δ ||-v σ ≅ σ' : Γ | wfΔ ]) + , [LogRel@{i j k l} l | Δ ||- A [ σ ] ≅ A [ σ' ] | validTy wfΔ vσσ' ] }. Arguments typeValidity : clear implicits. @@ -99,10 +96,9 @@ Arguments typeValidity {_ _ _ _ _ _ _ _ _}. Notation "[ P | Γ ||-v< l > A ]" := (typeValidity Γ P l A) (at level 0, P, Γ, l, A at level 50). -Definition emptyValidSubst `{ta : tag} `{!WfContext ta} : forall (Δ : context) (σ : nat -> term) (wfΔ : [|- Δ]), Type := fun _ _ _ => unit. -Definition emptyEqSubst@{u} `{ta : tag} `{!WfContext ta} : forall (Δ : context) (σ σ' : nat -> term) (wfΔ : [|- Δ]), emptyValidSubst@{u} Δ σ wfΔ -> Type@{u} := fun _ _ _ _ _ => unit. +Definition emptyEqSubst@{u} `{ta : tag} `{!WfContext ta} : forall (Δ : context) (wfΔ : [|- Δ]) (σ σ' : nat -> term), Type@{u} := fun _ _ _ _ => unit. Definition emptyVPack `{ta : tag} `{!WfContext ta} : VPack ε := - Build_VPack _ emptyValidSubst emptyEqSubst. + Build_VPack _ emptyEqSubst. Section snocValid. Universe u i j k l. @@ -112,26 +108,16 @@ Section snocValid. {Γ : context} {VΓ : VPack@{u} Γ} {A : term} {l : TypeLevel} {vA : typeValidity@{u i j k l} Γ VΓ l A (* [ VΓ | Γ ||-v< l > A ] *)}. - Record snocValidSubst {Δ : context} {σ : nat -> term} {wfΔ : [|- Δ]} : Type := - { - validTail : [ VΓ | Δ ||-v ↑ >> σ : Γ | wfΔ ] ; - validHead : [ Δ ||-< l > σ var_zero : A[↑ >> σ] | validTy vA wfΔ validTail ] - }. - Arguments snocValidSubst : clear implicits. - - Record snocEqSubst {Δ : context} {σ σ' : nat -> term} {wfΔ : [|- Δ]} {vσ : snocValidSubst Δ σ wfΔ} : Type := + Record snocEqSubst {Δ : context} {wfΔ : [|- Δ]} {σ σ' : nat -> term} : Type := { - eqTail : [ VΓ | Δ ||-v ↑ >> σ ≅ ↑ >> σ' : Γ | wfΔ | validTail vσ ] ; - eqHead : [ Δ ||-< l > σ var_zero ≅ σ' var_zero : A[↑ >> σ] | validTy vA wfΔ (validTail vσ) ] + eqTail : [ VΓ | Δ ||-v ↑ >> σ ≅ ↑ >> σ' : Γ | wfΔ ] ; + eqHead : [ Δ ||-< l > σ var_zero ≅ σ' var_zero : A[↑ >> σ] | validTy vA wfΔ eqTail ] }. - Definition snocVPack := Build_VPack@{u (* max(u,k) *)} (Γ ,, A) snocValidSubst (@snocEqSubst). + Definition snocVPack := Build_VPack@{u (* max(u,k) *)} (Γ ,, A) (@snocEqSubst). End snocValid. -Arguments snocValidSubst : clear implicits. -Arguments snocValidSubst {_ _ _ _ _ _ _ _ _}. - Arguments snocEqSubst : clear implicits. Arguments snocEqSubst {_ _ _ _ _ _ _ _ _}. @@ -144,19 +130,19 @@ Inductive VR@{i j k l} `{ta : tag} `{WfContext ta} `{WfType ta} `{Typing ta} `{ConvType ta} `{ConvTerm ta} `{ConvNeuConv ta} `{RedType ta} `{RedTerm ta} : VRel@{k l} := - | VREmpty : VR ε emptyValidSubst@{k} emptyEqSubst@{k} + | VREmpty : VR ε emptyEqSubst@{k} | VRSnoc : forall {Γ A l} (VΓ : VPack@{k} Γ) (VΓad : VPackAdequate@{k l} VR VΓ) (VA : typeValidity@{k i j k l} Γ VΓ l A (*[ VΓ | Γ ||-v< l > A ]*)), - VR (Γ ,, A) (snocValidSubst Γ VΓ A l VA) (snocEqSubst Γ VΓ A l VA). + VR (Γ ,, A) (snocEqSubst Γ VΓ A l VA). Set Elimination Schemes. Notation "[||-v Γ ]" := [ VR | ||-v Γ ] (at level 0, Γ at level 50). Notation "[ Δ ||-v σ : Γ | VΓ | wfΔ ]" := [ VR | Δ ||-v σ : Γ | VΓ | wfΔ ] (at level 0, Δ, σ, Γ, VΓ, wfΔ at level 50). -Notation "[ Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ | vσ ]" := [ VR | Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ | vσ ] (at level 0, Δ, σ, σ', Γ, VΓ, wfΔ, vσ at level 50). +Notation "[ Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ ]" := [ VR | Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ ] (at level 0, Δ, σ, σ', Γ, VΓ, wfΔ at level 50). Notation "[ Γ ||-v< l > A | VΓ ]" := [ VΓ | Γ ||-v< l > A ] (at level 0, Γ, l , A, VΓ at level 50). @@ -171,58 +157,57 @@ Section MoreDefs. : [||-v Γ ,, A ] := Build_VAdequate (snocVPack Γ VΓ A l VA) (VRSnoc VΓ VΓ VA). - Record termValidity@{i j k l} {Γ l} {t A : term} - {VΓ : [VR@{i j k l}| ||-v Γ]} - {VA : typeValidity@{k i j k l} Γ VΓ l A (*[Γ ||-v A |VΓ]*)} : Type := - { - validTm : forall {Δ σ} - (wfΔ : [|- Δ]) (Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ]), - [Δ ||- t[σ] : A[σ] | validTy VA wfΔ Vσ ] ; - validTmExt : forall {Δ : context} {σ σ' : nat -> term} - (wfΔ : [|- Δ ]) - (Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ ]) - (Vσ' : [ Δ ||-v σ' : Γ | VΓ | wfΔ ]) - (Vσσ' : [ Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ | Vσ ]), - [Δ ||- t[σ] ≅ t[σ'] : A[σ] | validTy VA wfΔ Vσ ] - }. - Record typeEqValidity@{i j k l} {Γ l} {A B : term} {VΓ : [VR@{i j k l}| ||-v Γ]} {VA : typeValidity@{k i j k l} Γ VΓ l A (*[Γ ||-v A |VΓ]*)} : Type := { - validTyEq : forall {Δ σ} - (wfΔ : [|- Δ]) (Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ]), - [Δ ||- A[σ] ≅ B[σ] | validTy VA wfΔ Vσ] + validTyEq : forall {Δ} (wfΔ : [|- Δ]) + {σ σ'} (Vσσ' : [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ]), + [Δ ||- A[σ] ≅ B[σ'] | validTy VA wfΔ Vσσ'] }. Record termEqValidity@{i j k l} {Γ l} {t u A : term} {VΓ : [VR@{i j k l}| ||-v Γ]} {VA : typeValidity@{k i j k l} Γ VΓ l A (*[Γ ||-v A |VΓ]*)} : Type := { - validTmEq : forall {Δ σ} - (wfΔ : [|- Δ]) (Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ]), - [Δ ||- t[σ] ≅ u[σ] : A[σ] | validTy VA wfΔ Vσ] + validTmEq : forall {Δ}(wfΔ : [|- Δ]) {σ σ'} + (Vσσ' : [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ]), + [Δ ||- t[σ] ≅ u[σ'] : A[σ] | validTy VA wfΔ Vσσ'] }. + Definition termValidity {Γ l} {t A : term} + {VΓ : [VR| ||-v Γ]} + {VA : typeValidity Γ VΓ l A (*[Γ ||-v A |VΓ]*)} : Type := + @termEqValidity Γ l t t A VΓ VA. + + Definition validTmExt {Γ l} {t A : term} + {VΓ : [VR| ||-v Γ]} + {VA : typeValidity Γ VΓ l A (*[Γ ||-v A |VΓ]*)} + (Vt : @termValidity Γ l t A VΓ VA) + : forall {Δ : context} (wfΔ : [|- Δ ]) {σ σ' : nat -> term} + (Vσσ' : [ Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ ]), + [Δ ||- t[σ] ≅ t[σ'] : A[σ] | validTy VA wfΔ Vσσ' ] := + fun _ => Vt.(validTmEq). + + Record tmEqValidity {Γ l} {t u A : term} {VΓ : [||-v Γ]} : Type := { Vty : [Γ ||-v< l > A | VΓ] ; - Vlhs : @termValidity Γ l t A VΓ Vty ; - Vrhs : @termValidity Γ l u A VΓ Vty ; + (* Vlhs : @termValidity Γ l t A VΓ Vty ; + Vrhs : @termValidity Γ l u A VΓ Vty ; *) Veq : @termEqValidity Γ l t u A VΓ Vty }. Record redValidity {Γ} {t u A : term} {VΓ : [||-v Γ]} : Type := { - validRed : forall {Δ σ} - (wfΔ : [|- Δ]) (Vσ : [Δ ||-v σ : Γ | VΓ | wfΔ]), - [Δ |- t[σ] :⤳*: u[σ] : A[σ]] + validRed : forall {Δ} (wfΔ : [|- Δ]) {σ σ'} (Vσσ' : [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ]), + [Δ |- t[σ] ⤳* u[σ] : A[σ]] }. End MoreDefs. Arguments termValidity : clear implicits. Arguments termValidity {_ _ _ _ _ _ _ _ _}. -Arguments Build_termValidity {_ _ _ _ _ _ _ _ _}. +(* Arguments Build_termValidity {_ _ _ _ _ _ _ _ _}. *) Arguments typeEqValidity : clear implicits. Arguments typeEqValidity {_ _ _ _ _ _ _ _ _}. @@ -244,20 +229,21 @@ Notation "[ Γ ||-v< l > t : A | VΓ | VA ]" := (termValidity Γ l t A VΓ V Notation "[ Γ ||-v< l > A ≅ B | VΓ | VA ]" := (typeEqValidity Γ l A B VΓ VA) (at level 0, Γ, l, A, B, VΓ, VA at level 50). Notation "[ Γ ||-v< l > t ≅ u : A | VΓ | VA ]" := (termEqValidity Γ l t u A VΓ VA) (at level 0, Γ, l, t, u, A, VΓ, VA at level 50). Notation "[ Γ ||-v< l > t ≅ u : A | VΓ ]" := (tmEqValidity Γ l t u A VΓ) (at level 0, Γ, l, t, u, A, VΓ at level 50). -Notation "[ Γ ||-v t :⤳*: u : A | VΓ ]" := (redValidity Γ t u A VΓ) (at level 0, Γ, t, u, A, VΓ at level 50). +(* Notation "[ Γ ||-v t :⤳*: u : A | VΓ ]" := (redValidity Γ t u A VΓ) (at level 0, Γ, t, u, A, VΓ at level 50). *) +Notation "[ Γ ||-v t ⤳* u : A | VΓ ]" := (redValidity Γ t u A VΓ) (at level 0, Γ, t, u, A, VΓ at level 50). Section Inductions. Context `{ta : tag} `{WfContext ta} `{WfType ta} `{Typing ta} `{ConvType ta} `{ConvTerm ta} `{ConvNeuConv ta} `{RedType ta} `{RedTerm ta}. Theorem VR_rect - (P : forall {Γ vSubst vSubstExt}, VR Γ vSubst vSubstExt -> Type) + (P : forall {Γ vSubstExt}, VR Γ vSubstExt -> Type) (hε : P VREmpty) (hsnoc : forall {Γ A l VΓ VΓad VA}, P VΓad -> P (VRSnoc (Γ := Γ) (A := A) (l := l) VΓ VΓad VA)) : - forall {Γ vSubst vSubstExt} (VΓ : VR Γ vSubst vSubstExt), P VΓ. + forall {Γ vSubstExt} (VΓ : VR Γ vSubstExt), P VΓ. Proof. - fix ih 4; destruct VΓ; [exact hε | apply hsnoc; apply ih]. + fix ih 3; destruct VΓ; [exact hε | apply hsnoc; apply ih]. Defined. Theorem validity_rect @@ -266,7 +252,7 @@ Section Inductions. (hsnoc : forall {Γ A l} (VΓ : [||-v Γ]) (VA : [Γ ||-v< l > A | VΓ]), P VΓ -> P (validSnoc VΓ VA)) : forall {Γ : context} (VΓ : [||-v Γ]), P VΓ. Proof. - intros Γ [[s eq] VΓad]; revert Γ s eq VΓad. + intros Γ [[eq] VΓad]; revert Γ eq VΓad. apply VR_rect. - apply hε. - intros *; apply hsnoc. @@ -282,14 +268,14 @@ Section Inductions. pattern Γ, VΓ. apply validity_rect. - reflexivity. - intros; do 3 eexists; reflexivity. - Qed. + Defined. Lemma invValidityEmpty (VΓ : [||-v ε]) : VΓ = validEmpty. - Proof. apply (invValidity VΓ). Qed. + Proof. apply (invValidity VΓ). Defined. Lemma invValiditySnoc {Γ A} (VΓ₀ : [||-v Γ ,, A ]) : ∑ l (VΓ : [||-v Γ]) (VA : [Γ ||-v< l > A | VΓ]), VΓ₀ = validSnoc VΓ VA. - Proof. apply (invValidity VΓ₀). Qed. + Proof. apply (invValidity VΓ₀). Defined. End Inductions. @@ -297,23 +283,24 @@ End Inductions. valid substitions *) Definition wfCtxOfsubstS `{GenericTypingProperties} - {Γ Δ σ} {VΓ : [||-v Γ]} {wfΔ : [|- Δ]} : - [Δ ||-v σ : Γ | VΓ | wfΔ] -> [|- Δ] := fun _ => wfΔ. + {Γ Δ σ σ'} {VΓ : [||-v Γ]} {wfΔ : [|- Δ]} : + [Δ ||-v σ ≅ σ' : Γ | VΓ | wfΔ] -> [|- Δ] := fun _ => wfΔ. Ltac instValid vσ := let wfΔ := (eval unfold wfCtxOfsubstS in (wfCtxOfsubstS vσ)) in repeat lazymatch goal with | [H : typeValidity _ _ _ _ |- _] => - let X := fresh "R" H in - try pose (X := validTy H wfΔ vσ) ; + try (let X := fresh "R" H in pose (X := validTy H wfΔ vσ)); + try (let X := fresh "RE" H in pose (X := validTyExt H wfΔ vσ)) ; + try (let X := fresh "RSym" H in pose (X := validTy H wfΔ (urefl vσ))) ; (* should only do that if vσ : [.. |- σ ≅ σ' : ...] with σ != σ' *) block H | [H : termValidity _ _ _ _ _ _ |- _] => let X := fresh "R" H in - try pose (X := validTm H wfΔ vσ) ; + try pose (X := validTmExt H wfΔ vσ) ; block H | [H : typeEqValidity _ _ _ _ _ _ |- _] => - let X := fresh "R" H in - try pose (X := validTyEq H wfΔ vσ) ; + try (let X := fresh "R" H in pose (X := validTyEq H wfΔ vσ)) ; + try (let X := fresh "RSym" H in pose (X := LRTyEqRedRight _ (validTyEq H wfΔ vσ))) ; block H | [H : termEqValidity _ _ _ _ _ _ _ |- _] => let X := fresh "R" H in @@ -321,7 +308,7 @@ Ltac instValid vσ := block H end; unblock. -Ltac instValidExt vσ' vσσ' := +(* Ltac instValidExt vσ' vσσ' := repeat lazymatch goal with | [H : typeValidity _ _ _ _ |- _] => let X := fresh "RE" H in @@ -331,9 +318,9 @@ Ltac instValidExt vσ' vσσ' := let X := fresh "RE" H in try pose (X := validTmExt H _ _ vσ' vσσ') ; block H - end; unblock. + end; unblock. *) -Ltac instAllValid vσ vσ' vσσ' := +(* Ltac instAllValid vσ vσ' vσσ' := instValid vσ ; instValid vσ' ; - instValidExt vσ' vσσ'. + instValidExt vσ' vσσ'. *)