]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/Basic_2/reduction/tpr_lift.ma
1. we compare the expected branching with the actual one and
[helm.git] / matita / matita / contribs / lambda_delta / Basic_2 / reduction / tpr_lift.ma
index 311a1433a9af933928bd329e13b05bab019251ca..7ebd4af8c68a19c19e618834e525103bb98bfd59 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "Basic-2/substitution/tps_lift.ma".
-include "Basic-2/reduction/tpr.ma".
+include "Basic_2/substitution/tps_lift.ma".
+include "Basic_2/reduction/tpr.ma".
 
 (* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************)
 
 (* Relocation properties ****************************************************)
 
-(* Basic-1: was: pr0_lift *)
+(* Basic_1: was: pr0_lift *)
 lemma tpr_lift: ∀T1,T2. T1 ⇒ T2 →
                 ∀d,e,U1. ↑[d, e] T1 ≡ U1 → ∀U2. ↑[d, e] T2 ≡ U2 → U1 ⇒ U2.
 #T1 #T2 #H elim H -H T1 T2
@@ -27,6 +27,7 @@ lemma tpr_lift: ∀T1,T2. T1 ⇒ T2 →
   lapply (lift_mono … HU1 … HU2) -HU1 #H destruct -U1
   [ lapply (lift_inv_sort1 … HU2) -HU2 #H destruct -U2 //
   | lapply (lift_inv_lref1 … HU2) * * #Hid #H destruct -U2 //
+  | lapply (lift_inv_gref1 … HU2) -HU2 #H destruct -U2 //
   ]
 | #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #X1 #HX1 #X2 #HX2
   elim (lift_inv_flat1 … HX1) -HX1 #W1 #U1 #HVW1 #HTU1 #HX1 destruct -X1;
@@ -55,7 +56,7 @@ lemma tpr_lift: ∀T1,T2. T1 ⇒ T2 →
 ]
 qed.
 
-(* Basic-1: was: pr0_gen_lift *)
+(* Basic_1: was: pr0_gen_lift *)
 lemma tpr_inv_lift: ∀T1,T2. T1 ⇒ T2 →
                     ∀d,e,U1. ↑[d, e] U1 ≡ T1 →
                     ∃∃U2. ↑[d, e] U2 ≡ T2 & U1 ⇒ U2.
@@ -63,6 +64,7 @@ lemma tpr_inv_lift: ∀T1,T2. T1 ⇒ T2 →
 [ * #i #d #e #U1 #HU1
   [ lapply (lift_inv_sort2 … HU1) -HU1 #H destruct -U1 /2/
   | lapply (lift_inv_lref2 … HU1) -HU1 * * #Hid #H destruct -U1 /3/
+  | lapply (lift_inv_gref2 … HU1) -HU1 #H destruct -U1 /2/
   ]
 | #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #e #X #HX
   elim (lift_inv_flat2 … HX) -HX #V0 #T0 #HV01 #HT01 #HX destruct -X;
@@ -113,7 +115,7 @@ fact tpr_inv_abst1_aux: ∀U1,U2. U1 ⇒ U2 → ∀V1,T1. U1 = 𝕔{Abst} V1. T1
 ]
 qed.
 
-(* Basic-1: was pr0_gen_abst *)
+(* Basic_1: was pr0_gen_abst *)
 lemma tpr_inv_abst1: ∀V1,T1,U2. 𝕔{Abst} V1. T1 ⇒ U2 →
                      ∃∃V2,T2. V1 ⇒ V2 & T1 ⇒ T2 & U2 = 𝕔{Abst} V2. T2.
 /2/ qed.