]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/relocation/lifts_simple.ma
some renaming and reordering of variables
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / relocation / lifts_simple.ma
index ae73703f7c4ff3e909dd3ffa296fa059ed8370da..a8a9cc13350254ac6adfdf66038de4ecf789a42b 100644 (file)
@@ -17,16 +17,16 @@ include "basic_2/relocation/lifts.ma".
 
 (* GENERIC RELOCATION FOR TERMS *********************************************)
 
-(* Forward lemmas on simple terms *******************************************)
+(* Forward lemmas with simple terms *****************************************)
 
 (* Basic_2A1: includes: lift_simple_dx *)
-lemma lifts_simple_dx: ∀T1,T2,t. ⬆*[t] T1 ≡ T2 → 𝐒⦃T1⦄ → 𝐒⦃T2⦄.
-#T1 #T2 #t #H elim H -T1 -T2 -t //
-#a #I #V1 #V2 #T1 #T2 #t #_ #_ #_ #_ #H elim (simple_inv_bind … H)
+lemma lifts_simple_dx: ∀f,T1,T2. ⬆*[f] T1 ≡ T2 → 𝐒⦃T1⦄ → 𝐒⦃T2⦄.
+#f #T1 #T2 #H elim H -f -T1 -T2 //
+#f #p #I #V1 #V2 #T1 #T2 #_ #_ #_ #_ #H elim (simple_inv_bind … H)
 qed-.
 
 (* Basic_2A1: includes: lift_simple_sn *)
-lemma lifts_simple_sn: ∀T1,T2,t. ⬆*[t] T1 ≡ T2 → 𝐒⦃T2⦄ → 𝐒⦃T1⦄.
-#T1 #T2 #t #H elim H -T1 -T2 -t //
-#a #I #V1 #V2 #T1 #T2 #t #_ #_ #_ #_ #H elim (simple_inv_bind … H)
+lemma lifts_simple_sn: ∀f,T1,T2. ⬆*[f] T1 ≡ T2 → 𝐒⦃T2⦄ → 𝐒⦃T1⦄.
+#f #T1 #T2 #H elim H -f -T1 -T2 //
+#f #p #I #V1 #V2 #T1 #T2 #_ #_ #_ #_ #H elim (simple_inv_bind … H)
 qed-.