]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/computation/cprs_lcpr.ma
- lambda_delta: we updated some notation
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / computation / cprs_lcpr.ma
index 14bcfd5c47d028b1519456bac8a70f97b1fd583d..86f2a13ea67212a5e094b274881b29ddae0441db 100644 (file)
@@ -21,8 +21,8 @@ include "basic_2/computation/cprs.ma".
 
 (* Properties concerning context-sensitive parallel reduction on lenv's *****)
 
-lemma ltpr_tpss_trans: ∀L1,L2. L1 ➡ L2 → ∀T1,T2,d,e. L2 ⊢ T1 [d, e] ▶* T2 →
-                       ∃∃T. L1 ⊢ T1 [d, e] ▶* T & L1 ⊢ T ➡* T2.
+lemma ltpr_tpss_trans: ∀L1,L2. L1 ➡ L2 → ∀T1,T2,d,e. L2 ⊢ T1 ▶* [d, e] T2 →
+                       ∃∃T. L1 ⊢ T1 ▶* [d, e] T & L1 ⊢ T ➡* T2.
 #L1 #L2 #HL12 #T1 #T2 #d #e #H @(tpss_ind … H) -T2
 [ /2 width=3/
 | #T #T2 #_ #HT2 * #T0 #HT10 #HT0