X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_transition%2Fcnx.ma;h=f958591d49b12ee31cf5e980a6903a31c1a8d757;hb=3c7b4071a9ac096b02334c1d47468776b948e2de;hp=6607013df43ccf24e2f81298b5e7a1ff3237c128;hpb=4173283e148199871d787c53c0301891deb90713;p=helm.git 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 6607013df..f958591d4 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cnx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cnx.ma @@ -12,56 +12,60 @@ (* *) (**************************************************************************) -include "basic_2/notation/relations/predtynormal_4.ma". -include "static_2/syntax/tdeq.ma". +include "basic_2/notation/relations/predtynormal_3.ma". +include "static_2/syntax/teqx.ma". include "basic_2/rt_transition/cpx.ma". -(* NORMAL TERMS FOR UNBOUND CONTEXT-SENSITIVE PARALLEL RT-TRANSITION ********) +(* NORMAL TERMS FOR EXTENDED CONTEXT-SENSITIVE PARALLEL RT-TRANSITION *******) -definition cnx: ∀h. relation3 genv lenv term ≝ - λh,G,L. NF … (cpx h G L) tdeq. +definition cnx: relation3 genv lenv term ≝ + λG,L. NF … (cpx G L) teqx. interpretation - "normality for unbound context-sensitive parallel rt-transition (term)" - 'PRedTyNormal h G L T = (cnx h G L T). + "normality for extended context-sensitive parallel rt-transition (term)" + 'PRedTyNormal G L T = (cnx G L T). (* 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⦄. -#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 +lemma cnx_inv_abst (G) (L): + ∀p,V,T. ❪G,L❫ ⊢ ⬈𝐍 ⓛ[p]V.T → + ∧∧ ❪G,L❫ ⊢ ⬈𝐍 V & ❪G,L.ⓛV❫ ⊢ ⬈𝐍 T. +#G #L #p #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 ] -#H elim (tdeq_inv_pair … H) -H // +#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⦄. -#h #G #L #V1 #T1 #HVT1 @conj -[ #V2 #HV2 lapply (HVT1 (-ⓓV2.T1) ?) -HVT1 /2 width=2 by cpx_pair_sn/ -HV2 +lemma cnx_inv_abbr_neg (G) (L): + ∀V,T. ❪G,L❫ ⊢ ⬈𝐍 -ⓓV.T → + ∧∧ ❪G,L❫ ⊢ ⬈𝐍 V & ❪G,L.ⓓV❫ ⊢ ⬈𝐍 T. +#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 ] -#H elim (tdeq_inv_pair … H) -H // +#H elim (teqx_inv_pair … H) -H // qed-. (* Basic_2A1: was: cnx_inv_eps *) -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, tdeq_inv_pair_xy_y/ +lemma cnx_inv_cast (G) (L): + ∀V,T. ❪G,L❫ ⊢ ⬈𝐍 ⓝV.T → ⊥. +#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⦄. -#h #G #L #s #X #H elim (cpx_inv_sort1 … H) -H -/2 width=1 by tdeq_sort/ +lemma cnx_sort (G) (L): + ∀s. ❪G,L❫ ⊢ ⬈𝐍 ⋆s. +#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⦄. -#h #p #G #L #W #T #HW #HT #X #H +lemma cnx_abst (G) (L): + ∀p,W,T. ❪G,L❫ ⊢ ⬈𝐍 W → ❪G,L.ⓛW❫ ⊢ ⬈𝐍 T → ❪G,L❫ ⊢ ⬈𝐍 ⓛ[p]W.T. +#G #L #p #W #T #HW #HT #X #H elim (cpx_inv_abst1 … H) -H #W0 #T0 #HW0 #HT0 #H destruct -@tdeq_pair [ @HW | @HT ] // (**) (* auto fails because δ-expansion gets in the way *) +@teqx_pair [ @HW | @HT ] // (**) (* auto fails because δ-expansion gets in the way *) qed.