X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Freduction%2Fcix_lift.ma;h=730a66fb16b730577806ef5c2f07f5a0a63d0f0c;hb=29973426e0227ee48368d1c24dc0c17bf2baef77;hp=f81f17541e4fbe583124964018b22b48d109fd0a;hpb=f95f6cb21b86f3dad114b21f687aa5df36088064;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 f81f17541..730a66fb1 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cix_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cix_lift.ma @@ -19,17 +19,17 @@ include "basic_2/reduction/cix.ma". (* Advanced properties ******************************************************) -lemma cix_lref: ∀h,g,L,i. ⇩[0, i] L ≡ ⋆ → ⦃h, L⦄ ⊢ 𝐈[g]⦃#i⦄. +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 lapply (ldrop_mono … HLK … HL) -L -i #H destruct qed. (* Properties on relocation *************************************************) -lemma cix_lift: ∀h,g,K,T. ⦃h, K⦄ ⊢ 𝐈[g]⦃T⦄ → ∀L,d,e. ⇩[d, e] L ≡ K → - ∀U. ⇧[d, e] T ≡ U → ⦃h, L⦄ ⊢ 𝐈[g]⦃U⦄. +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_inv_lift: ∀h,g,L,U. ⦃h, L⦄ ⊢ 𝐈[g]⦃U⦄ → ∀K,d,e. ⇩[d, e] L ≡ K → - ∀T. ⇧[d, e] T ≡ U → ⦃h, K⦄ ⊢ 𝐈[g]⦃T⦄. +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-.