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=6607013df43ccf24e2f81298b5e7a1ff3237c128;hp=cc2397d0f8dbf569165da561cac9bd99fc340cb1;hb=4173283e148199871d787c53c0301891deb90713;hpb=a67fc50ccfda64377e2c94c18c3a0d9265f651db 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 cc2397d0f..6607013df 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cnx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_transition/cnx.ma @@ -12,33 +12,24 @@ (* *) (**************************************************************************) -include "basic_2/notation/relations/predtynormal_5.ma". +include "basic_2/notation/relations/predtynormal_4.ma". include "static_2/syntax/tdeq.ma". include "basic_2/rt_transition/cpx.ma". (* NORMAL TERMS FOR UNBOUND CONTEXT-SENSITIVE PARALLEL RT-TRANSITION ********) -definition cnx: ∀h. sd h → relation3 genv lenv term ≝ - λh,o,G,L. NF … (cpx h G L) (tdeq h o …). +definition cnx: ∀h. relation3 genv lenv term ≝ + λh,G,L. NF … (cpx h G L) tdeq. interpretation "normality for unbound context-sensitive parallel rt-transition (term)" - 'PRedTyNormal h o G L T = (cnx h o G L T). + 'PRedTyNormal h G L T = (cnx h G L T). (* Basic inversion lemmas ***************************************************) -lemma cnx_inv_sort: ∀h,o,G,L,s. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐍⦃⋆s⦄ → deg h o s 0. -#h #o #G #L #s #H -lapply (H (⋆(next h s)) ?) -H /2 width=2 by cpx_ess/ -G -L #H -elim (tdeq_inv_sort1 … H) -H #s0 #d #H1 #H2 #H destruct -lapply (deg_next … H1) #H0 -lapply (deg_mono … H0 … H2) -H0 -H2 #H ->(pred_inv_fix_sn … H) -H // -qed-. - -lemma cnx_inv_abst: ∀h,o,p,G,L,V,T. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐍⦃ⓛ{p}V.T⦄ → - ⦃G, L⦄ ⊢ ⬈[h, o] 𝐍⦃V⦄ ∧ ⦃G, L.ⓛV⦄ ⊢ ⬈[h, o] 𝐍⦃T⦄. -#h #o #p #G #L #V1 #T1 #HVT1 @conj +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 ] @@ -46,9 +37,9 @@ lemma cnx_inv_abst: ∀h,o,p,G,L,V,T. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐍⦃ⓛ{p}V.T qed-. (* Basic_2A1: was: cnx_inv_abbr *) -lemma cnx_inv_abbr_neg: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐍⦃-ⓓV.T⦄ → - ⦃G, L⦄ ⊢ ⬈[h, o] 𝐍⦃V⦄ ∧ ⦃G, L.ⓓV⦄ ⊢ ⬈[h, o] 𝐍⦃T⦄. -#h #o #G #L #V1 #T1 #HVT1 @conj +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 ] @@ -56,26 +47,21 @@ lemma cnx_inv_abbr_neg: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐍⦃-ⓓV.T qed-. (* Basic_2A1: was: cnx_inv_eps *) -lemma cnx_inv_cast: ∀h,o,G,L,V,T. ⦃G, L⦄ ⊢ ⬈[h, o] 𝐍⦃ⓝV.T⦄ → ⊥. -#h #o #G #L #V #T #H lapply (H T ?) -H +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/ qed-. (* Basic properties *********************************************************) -lemma cnx_sort: ∀h,o,G,L,s. deg h o s 0 → ⦃G, L⦄ ⊢ ⬈[h, o] 𝐍⦃⋆s⦄. -#h #o #G #L #s #Hs #X #H elim (cpx_inv_sort1 … H) -H -/3 width=3 by tdeq_sort, deg_next/ -qed. - -lemma cnx_sort_iter: ∀h,o,G,L,s,d. deg h o s d → ⦃G, L⦄ ⊢ ⬈[h, o] 𝐍⦃⋆((next h)^d s)⦄. -#h #o #G #L #s #d #Hs lapply (deg_iter … d Hs) -Hs -