]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/computation/cprs_ltpr.ma
- some additions and corrections
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / computation / cprs_ltpr.ma
index 8859a46d6b19fa4f7babd1cd004e7d24f10d3abf..2682a7609311bfb11884ac526f5328860603fd8d 100644 (file)
@@ -20,9 +20,9 @@ include "basic_2/computation/cprs.ma".
 (* Properties concerning parallel unfold on terms ***************************)
 
 (* Basic_1: was only: pr3_subst1 *)
-lemma cprs_tpss_ltpr: ∀L1,T1,U1,d,e. L1 ⊢ T1 [d, e] ▶* U1 →
+lemma cprs_tpss_ltpr: ∀L1,T1,U1,d,e. L1 ⊢ T1 ▶* [d, e] U1 →
                       ∀L2. L1 ➡ L2 → ∀T2. L2 ⊢ T1 ➡* T2 →
-                      ∃∃U2. L2 ⊢ U1 ➡* U2 & L2 ⊢ T2 [d, e] ▶* U2.
+                      ∃∃U2. L2 ⊢ U1 ➡* U2 & L2 ⊢ T2 ▶* [d, e] U2.
 #L1 #T1 #U1 #d #e #HTU1 #L2 #HL12 #T2 #HT12 elim HT12 -T2
 [ #T2 #HT12
   elim (cpr_tpss_ltpr … HL12 … HT12 … HTU1) -L1 -T1 /3 width=3/