]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/apps_2/functional/lift.ma
update in static_2 and app_2
[helm.git] / matita / matita / contribs / lambdadelta / apps_2 / functional / lift.ma
index 92c791912cacfdebf40ae9b94f61564c3b15d2b3..ea74090600b28b7fc081a3a9e582dc54a3e438ac 100644 (file)
@@ -12,7 +12,7 @@
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/relocation/lift.ma".
+include "basic_2/substitution/lift.ma".
 include "apps_2/functional/notation.ma".
 
 (* FUNCTIONAL RELOCATION ****************************************************)
@@ -33,7 +33,7 @@ interpretation "functional relocation" 'Lift d e T = (flift d e T).
 
 (* Main properties **********************************************************)
 
-theorem flift_lift: â\88\80T,d,e. â\87§[d, e] T ≡ ↑[d, e] T.
+theorem flift_lift: â\88\80T,d,e. â¬\86[d, e] T ≡ ↑[d, e] T.
 #T elim T -T
 [ * #i #d #e //
   elim (lt_or_eq_or_gt i d) #Hid normalize
@@ -47,7 +47,7 @@ qed.
 
 (* Main inversion properties ************************************************)
 
-theorem flift_inv_lift: â\88\80d,e,T1,T2. â\87§[d, e] T1 ≡ T2 → ↑[d, e] T1 = T2.
+theorem flift_inv_lift: â\88\80d,e,T1,T2. â¬\86[d, e] T1 ≡ T2 → ↑[d, e] T1 = T2.
 #d #e #T1 #T2 #H elim H -d -e -T1 -T2 normalize //
 [ #i #d #e #Hid >(tri_lt ?????? Hid) //
 | #i #d #e #Hid
@@ -60,9 +60,9 @@ qed-.
 
 (* Derived properties *******************************************************)
 
-lemma flift_join: â\88\80e1,e2,T. â\87§[e1, e2] ↑[0, e1] T ≡ ↑[0, e1 + e2] T.
+lemma flift_join: â\88\80e1,e2,T. â¬\86[e1, e2] ↑[0, e1] T ≡ ↑[0, e1 + e2] T.
 #e1 #e2 #T
 lapply (flift_lift T 0 (e1+e2)) #H
-elim (lift_split … H e1 e1 ? ? ?) -H // #U #H
+elim (lift_split … H e1 e1) -H // #U #H
 >(flift_inv_lift … H) -H //
 qed.