]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/reduction/cnx_lift.ma
- some renaming according to the written version of basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / reduction / cnx_lift.ma
index f313ee8bf0bea4aad8e6ae8a54d87e9206d60811..e66223cace592e819498ad35dab52c5fe6e5ca2e 100644 (file)
@@ -19,19 +19,19 @@ include "basic_2/reduction/cnx.ma".
 
 (* 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⦄.
-#h #g #G #L0 #L #T #T0 #s #d #e #HLT #HL0 #HT0 #X #H
+lemma cnx_lift: ∀h,g,G,L0,L,T,T0,s,l,m. ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T⦄ → ⬇[s, l, m] L0 ≡ L →
+                ⬆[l, m] T ≡ T0 → ⦃G, L0⦄ ⊢ ➡[h, g] 𝐍⦃T0⦄.
+#h #g #G #L0 #L #T #T0 #s #l #m #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⦄.
-#h #g #G #L0 #L #T #T0 #s #d #e #HLT0 #HL0 #HT0 #X #H
-elim (lift_total X d e) #X0 #HX0
+lemma cnx_inv_lift: ∀h,g,G,L0,L,T,T0,s,l,m. ⦃G, L0⦄ ⊢ ➡[h, g] 𝐍⦃T0⦄ → ⬇[s, l, m] L0 ≡ L →
+                    ⬆[l, m] T ≡ T0 → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T⦄.
+#h #g #G #L0 #L #T #T0 #s #l #m #HLT0 #HL0 #HT0 #X #H
+elim (lift_total X l m) #X0 #HX0
 lapply (cpx_lift … H … HL0 … HT0 … HX0) -L #HTX0
 >(HLT0 … HTX0) in HX0; -L0 -X0 #H
->(lift_inj … H … HT0) -T0 -X -d -e //
+>(lift_inj … H … HT0) -T0 -X -l -m //
 qed-.