]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/Basic_2/functional/lift.ma
- main lemmas about abstract reducibility candidates closed
[helm.git] / matita / matita / contribs / lambda_delta / Basic_2 / functional / lift.ma
index 4debfaabf6cfe34a31eb55eeeeaa72a8955b2f8b..78046720905bc2a961738df7879d27fca81a8352 100644 (file)
@@ -24,8 +24,8 @@ let rec flift d e U on U ≝ match U with
   | GRef _ ⇒ U
   ]
 | TPair I V T ⇒ match I with
-  [ Bind I ⇒ 𝕓{I} (flift d e V). (flift (d+1) e T)
-  | Flat I ⇒ 𝕗{I} (flift d e V). (flift d e T)
+  [ Bind2 I ⇒ ⓑ{I} (flift d e V). (flift (d+1) e T)
+  | Flat2 I ⇒ ⓕ{I} (flift d e V). (flift d e T)
   ]
 ].
 
@@ -33,7 +33,7 @@ interpretation "functional relocation" 'Lift d e T = (flift d e T).
 
 (* Main properties **********************************************************)
 
-theorem flift_lift: â\88\80T,d,e. â\86\91[d, e] T â\89¡ â\86\9f[d, e] T.
+theorem flift_lift: â\88\80T,d,e. â\87§[d, e] T â\89¡ â\86\91[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. â\86\91[d, e] T1 â\89¡ T2 â\86\92 â\86\9f[d, e] T1 = T2.
+theorem flift_inv_lift: â\88\80d,e,T1,T2. â\87§[d, e] T1 â\89¡ T2 â\86\92 â\86\91[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,7 +60,7 @@ qed-.
 
 (* Derived properties *******************************************************)
 
-lemma flift_join: â\88\80e1,e2,T. â\86\91[e1, e2] â\86\9f[0, e1] T â\89¡ â\86\9f[0, e1 + e2] T.
+lemma flift_join: â\88\80e1,e2,T. â\87§[e1, e2] â\86\91[0, e1] T â\89¡ â\86\91[0, e1 + e2] T.
 #e1 #e2 #T
 lapply (flift_lift T 0 (e1+e2)) #H
 elim (lift_split … H e1 e1 ? ? ?) -H // #U #H