X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_transition%2Fcnx.ma;h=4b9059643cf1676c9f0b83e38db9581a3b684821;hp=0f45712626b3851ad7a6558089fd38632613cd75;hb=bd53c4e895203eb049e75434f638f26b5a161a2b;hpb=3b7b8afcb429a60d716d5226a5b6ab0d003228b1 diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cnx.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cnx.ma index 0f4571262..4b9059643 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cnx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cnx.ma @@ -27,18 +27,18 @@ interpretation (* Basic inversion lemmas ***************************************************) -lemma cnx_inv_abst: ∀h,p,G,L,V,T. ⦃G,L⦄ ⊢ ⬈[h] 𝐍⦃ⓛ{p}V.T⦄ → - ⦃G,L⦄ ⊢ ⬈[h] 𝐍⦃V⦄ ∧ ⦃G,L.ⓛV⦄ ⊢ ⬈[h] 𝐍⦃T⦄. +lemma cnx_inv_abst: ∀h,p,G,L,V,T. ❪G,L❫ ⊢ ⬈[h] 𝐍❪ⓛ[p]V.T❫ → + ❪G,L❫ ⊢ ⬈[h] 𝐍❪V❫ ∧ ❪G,L.ⓛV❫ ⊢ ⬈[h] 𝐍❪T❫. #h #p #G #L #V1 #T1 #HVT1 @conj -[ #V2 #HV2 lapply (HVT1 (ⓛ{p}V2.T1) ?) -HVT1 /2 width=2 by cpx_pair_sn/ -HV2 -| #T2 #HT2 lapply (HVT1 (ⓛ{p}V1.T2) ?) -HVT1 /2 width=2 by cpx_bind/ -HT2 +[ #V2 #HV2 lapply (HVT1 (ⓛ[p]V2.T1) ?) -HVT1 /2 width=2 by cpx_pair_sn/ -HV2 +| #T2 #HT2 lapply (HVT1 (ⓛ[p]V1.T2) ?) -HVT1 /2 width=2 by cpx_bind/ -HT2 ] #H elim (teqx_inv_pair … H) -H // qed-. (* Basic_2A1: was: cnx_inv_abbr *) -lemma cnx_inv_abbr_neg: ∀h,G,L,V,T. ⦃G,L⦄ ⊢ ⬈[h] 𝐍⦃-ⓓV.T⦄ → - ⦃G,L⦄ ⊢ ⬈[h] 𝐍⦃V⦄ ∧ ⦃G,L.ⓓV⦄ ⊢ ⬈[h] 𝐍⦃T⦄. +lemma cnx_inv_abbr_neg: ∀h,G,L,V,T. ❪G,L❫ ⊢ ⬈[h] 𝐍❪-ⓓV.T❫ → + ❪G,L❫ ⊢ ⬈[h] 𝐍❪V❫ ∧ ❪G,L.ⓓV❫ ⊢ ⬈[h] 𝐍❪T❫. #h #G #L #V1 #T1 #HVT1 @conj [ #V2 #HV2 lapply (HVT1 (-ⓓV2.T1) ?) -HVT1 /2 width=2 by cpx_pair_sn/ -HV2 | #T2 #HT2 lapply (HVT1 (-ⓓV1.T2) ?) -HVT1 /2 width=2 by cpx_bind/ -HT2 @@ -47,20 +47,20 @@ lemma cnx_inv_abbr_neg: ∀h,G,L,V,T. ⦃G,L⦄ ⊢ ⬈[h] 𝐍⦃-ⓓV.T⦄ → qed-. (* Basic_2A1: was: cnx_inv_eps *) -lemma cnx_inv_cast: ∀h,G,L,V,T. ⦃G,L⦄ ⊢ ⬈[h] 𝐍⦃ⓝV.T⦄ → ⊥. +lemma cnx_inv_cast: ∀h,G,L,V,T. ❪G,L❫ ⊢ ⬈[h] 𝐍❪ⓝV.T❫ → ⊥. #h #G #L #V #T #H lapply (H T ?) -H /2 width=6 by cpx_eps, teqx_inv_pair_xy_y/ qed-. (* Basic properties *********************************************************) -lemma cnx_sort: ∀h,G,L,s. ⦃G,L⦄ ⊢ ⬈[h] 𝐍⦃⋆s⦄. +lemma cnx_sort: ∀h,G,L,s. ❪G,L❫ ⊢ ⬈[h] 𝐍❪⋆s❫. #h #G #L #s #X #H elim (cpx_inv_sort1 … H) -H /2 width=1 by teqx_sort/ qed. -lemma cnx_abst: ∀h,p,G,L,W,T. ⦃G,L⦄ ⊢ ⬈[h] 𝐍⦃W⦄ → ⦃G,L.ⓛW⦄ ⊢ ⬈[h] 𝐍⦃T⦄ → - ⦃G,L⦄ ⊢ ⬈[h] 𝐍⦃ⓛ{p}W.T⦄. +lemma cnx_abst: ∀h,p,G,L,W,T. ❪G,L❫ ⊢ ⬈[h] 𝐍❪W❫ → ❪G,L.ⓛW❫ ⊢ ⬈[h] 𝐍❪T❫ → + ❪G,L❫ ⊢ ⬈[h] 𝐍❪ⓛ[p]W.T❫. #h #p #G #L #W #T #HW #HT #X #H elim (cpx_inv_abst1 … H) -H #W0 #T0 #HW0 #HT0 #H destruct @teqx_pair [ @HW | @HT ] // (**) (* auto fails because δ-expansion gets in the way *)