X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Freduction%2Fcnx_lift.ma;h=91ea0091bc35784a63fd94dc7618e74b53d64714;hb=598a5c56535a8339f6533227ab580aff64e2d41c;hp=f52149d5b2c58270583cb89baa921fe067017cd2;hpb=2ba2dc23443ad764adab652e06d6f5ed10bd912d;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx_lift.ma index f52149d5b..91ea0091b 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx_lift.ma @@ -15,29 +15,20 @@ include "basic_2/reduction/cpx_lift.ma". include "basic_2/reduction/cnx.ma". -(* CONTEXT-SENSITIVE EXTENDED NORMAL TERMS **********************************) - -(* Advanced properties ******************************************************) - -lemma cnx_lref_atom: ∀h,g,G,L,i. ⇩[i] L ≡ ⋆ → ⦃G, L⦄ ⊢ 𝐍[h, g]⦃#i⦄. -#h #g #G #L #i #HL #X #H -elim (cpx_inv_lref1 … H) -H // * -#I #K #V1 #V2 #HLK #_ #_ -lapply (ldrop_mono … HL … HLK) -L #H destruct -qed. +(* NORMAL TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION ********************) (* Relocation properties ****************************************************) -lemma cnx_lift: ∀h,g,G,L0,L,T,T0,s,d,e. ⦃G, L⦄ ⊢ 𝐍[h, g]⦃T⦄ → ⇩[s, d, e] L0 ≡ L → - ⇧[d, e] T ≡ T0 → ⦃G, L0⦄ ⊢ 𝐍[h, g]⦃T0⦄. +lemma cnx_lift: ∀h,g,G,L0,L,T,T0,s,d,e. ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T⦄ → ⇩[s, d, e] L0 ≡ L → + ⇧[d, e] T ≡ T0 → ⦃G, L0⦄ ⊢ ➡[h, g] 𝐍⦃T0⦄. #h #g #G #L0 #L #T #T0 #s #d #e #HLT #HL0 #HT0 #X #H elim (cpx_inv_lift1 … H … HL0 … HT0) -L0 #T1 #HT10 #HT1 <(HLT … HT1) in HT0; -L #HT0 >(lift_mono … HT10 … HT0) -T1 -X // qed. -lemma cnx_inv_lift: ∀h,g,G,L0,L,T,T0,s,d,e. ⦃G, L0⦄ ⊢ 𝐍[h, g]⦃T0⦄ → ⇩[s, d, e] L0 ≡ L → - ⇧[d, e] T ≡ T0 → ⦃G, L⦄ ⊢ 𝐍[h, g]⦃T⦄. +lemma cnx_inv_lift: ∀h,g,G,L0,L,T,T0,s,d,e. ⦃G, L0⦄ ⊢ ➡[h, g] 𝐍⦃T0⦄ → ⇩[s, d, e] L0 ≡ L → + ⇧[d, e] T ≡ T0 → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T⦄. #h #g #G #L0 #L #T #T0 #s #d #e #HLT0 #HL0 #HT0 #X #H elim (lift_total X d e) #X0 #HX0 lapply (cpx_lift … H … HL0 … HT0 … HX0) -L #HTX0