]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda-delta/Basic-2/reduction/cpr.ma
we now use non-telescopic substitution in parallel reduction, rather
[helm.git] / matita / matita / contribs / lambda-delta / Basic-2 / reduction / cpr.ma
index 21169192cb74c759affea24e2dddcd59e6ed79e6..f41999e45cf56f4e0772d48eb6e5653fdd110d0e 100644 (file)
@@ -29,7 +29,7 @@ lemma cpr_pr: ∀T1,T2. T1 ⇒ T2 → ∀L. L ⊢ T1 ⇒ T2.
 /2/ qed.
 
 lemma cpr_tps: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ≫ T2 → L ⊢ T1 ⇒ T2.
-/3 width=5/ qed. 
+/3 width=5/ qed.
 
 lemma cpr_refl: ∀L,T. L ⊢ T ⇒ T.
 /2/ qed.
@@ -40,11 +40,9 @@ lemma cpr_flat: ∀I,L,V1,V2,T1,T2.
 #I #L #V1 #V2 #T1 #T2 * #V #HV1 #HV2 * /3 width=5/
 qed.
 
-lemma cpr_delta: ∀L,K,V1,V2,V,i.
-                 ↓[0, i] L ≡ K. 𝕓{Abbr} V1 → K ⊢ V1 [0, |L| - i - 1] ≫ V2 →
-                 ↑[0, i + 1] V2 ≡ V → L ⊢ #i ⇒ V.
-#L #K #V1 #V2 #V #i #HLK #HV12 #HV2
-@ex2_1_intro [2: // | skip ] /3 width=8/ (**) (* /4/ is too slow *)
+lemma cpr_delta: ∀L,K,V,W,i.
+                 ↓[0, i] L ≡ K. 𝕓{Abbr} V → ↑[0, i + 1] V ≡ W → L ⊢ #i ⇒ W.
+/3/
 qed.
 
 lemma cpr_cast: ∀L,V,T1,T2.