X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Freduction%2Fcix_lift.ma;h=7a3d5e68b24254e84e823599b635880153446ce2;hb=f21cc1fc7f776761926a7f017fda55735d63442e;hp=730a66fb16b730577806ef5c2f07f5a0a63d0f0c;hpb=29973426e0227ee48368d1c24dc0c17bf2baef77;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cix_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cix_lift.ma index 730a66fb1..7a3d5e68b 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cix_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cix_lift.ma @@ -15,21 +15,21 @@ include "basic_2/reduction/crx_lift.ma". include "basic_2/reduction/cix.ma". -(* CONTEXT-SENSITIVE EXTENDED IRREDUCIBLE TERMS *****************************) +(* IRREDUCIBLE TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION ***************) (* Advanced properties ******************************************************) -lemma cix_lref: ∀h,g,L,i. ⇩[0, i] L ≡ ⋆ → ⦃G, L⦄ ⊢ 𝐈[h, g]⦃#i⦄. -#h #g #L #i #HL #H elim (crx_inv_lref … H) -h #I #K #V #HLK +lemma cix_lref: ∀h,g,G,L,i. ⇩[i] L ≡ ⋆ → ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃#i⦄. +#h #g #G #L #i #HL #H elim (crx_inv_lref … H) -h #I #K #V #HLK lapply (ldrop_mono … HLK … HL) -L -i #H destruct qed. (* Properties on relocation *************************************************) -lemma cix_lift: ∀h,g,K,T. ⦃h, K⦄ ⊢ 𝐈[h, g]⦃T⦄ → ∀L,d,e. ⇩[d, e] L ≡ K → - ∀U. ⇧[d, e] T ≡ U → ⦃G, L⦄ ⊢ 𝐈[h, g]⦃U⦄. -/3 width=7 by crx_inv_lift/ qed. +lemma cix_lift: ∀h,g,G,K,T. ⦃G, K⦄ ⊢ ➡[h, g] 𝐈⦃T⦄ → ∀L,s,d,e. ⇩[s, d, e] L ≡ K → + ∀U. ⇧[d, e] T ≡ U → ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃U⦄. +/3 width=8 by crx_inv_lift/ qed. -lemma cix_inv_lift: ∀h,g,L,U. ⦃G, L⦄ ⊢ 𝐈[h, g]⦃U⦄ → ∀K,d,e. ⇩[d, e] L ≡ K → - ∀T. ⇧[d, e] T ≡ U → ⦃h, K⦄ ⊢ 𝐈[h, g]⦃T⦄. -/3 width=7/ qed-. +lemma cix_inv_lift: ∀h,g,G,L,U. ⦃G, L⦄ ⊢ ➡[h, g] 𝐈⦃U⦄ → ∀K,s,d,e. ⇩[s, d, e] L ≡ K → + ∀T. ⇧[d, e] T ≡ U → ⦃G, K⦄ ⊢ ➡[h, g] 𝐈⦃T⦄. +/3 width=8 by crx_lift/ qed-.