]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/Basic_2/substitution/lift_lift.ma
- we proved that context-free reduction admits no one-step loops
[helm.git] / matita / matita / contribs / lambda_delta / Basic_2 / substitution / lift_lift.ma
index 2f1e3ac3c59c4f51828bffa8178a357acde1dce6..5bc36f943ededfbe6e9578acc3972cd0e034b4e7 100644 (file)
@@ -34,7 +34,7 @@ theorem lift_inj:  ∀d,e,T1,U. ↑[d,e] T1 ≡ U → ∀T2. ↑[d,e] T2 ≡ U 
 | #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
   elim (lift_inv_flat2 … HX) -HX #V #T #HV1 #HT1 #HX destruct -X /3/
 ]
-qed.
+qed-.
 
 (* Basic_1: was: lift_gen_lift *)
 theorem lift_div_le: ∀d1,e1,T1,T. ↑[d1, e1] T1 ≡ T →
@@ -69,7 +69,7 @@ theorem lift_div_le: ∀d1,e1,T1,T. ↑[d1, e1] T1 ≡ T →
 ]
 qed.
 
-theorem lift_mono:  ∀d,e,T,U1. ↑[d,e] T ≡ U1 → ∀U2. ↑[d,e] T ≡ U2 → U1 = U2.
+theorem lift_mono: ∀d,e,T,U1. ↑[d,e] T ≡ U1 → ∀U2. ↑[d,e] T ≡ U2 → U1 = U2.
 #d #e #T #U1 #H elim H -H d e T U1
 [ #k #d #e #X #HX
   lapply (lift_inv_sort1 … HX) -HX //
@@ -84,7 +84,7 @@ theorem lift_mono:  ∀d,e,T,U1. ↑[d,e] T ≡ U1 → ∀U2. ↑[d,e] T ≡ U2
 | #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX
   elim (lift_inv_flat1 … HX) -HX #V #T #HV1 #HT1 #HX destruct -X /3/
 ]
-qed.
+qed-.
 
 (* Basic_1: was: lift_free (left to right) *)
 theorem lift_trans_be: ∀d1,e1,T1,T. ↑[d1, e1] T1 ≡ T →
@@ -166,6 +166,6 @@ theorem lift_trans_ge: ∀d1,e1,T1,T. ↑[d1, e1] T1 ≡ T →
 | #I #V1 #V2 #T1 #T2 #d1 #e1 #_ #_ #IHV12 #IHT12 #d2 #e2 #X #HX #Hded
   elim (lift_inv_flat1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct -X;
   elim (IHV12 … HV20 ?) -IHV12 HV20 //
-  elim (IHT12 … HT20 ?) -IHT12 HT20 // /3 width=5/ (**) (* just /3 width=5/ crashes *)
+  elim (IHT12 … HT20 ?) -IHT12 HT20 // /3 width=5/
 ]
 qed.