]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda-delta/Basic-2/reduction/tpr_tps.ma
- the shift function is now defined and cpr_shift_fwd is proved
[helm.git] / matita / matita / contribs / lambda-delta / Basic-2 / reduction / tpr_tps.ma
index 6fae790d80e23ed26bb728d13ea625d6696561e6..3a00c62c6b98a298ffb6ed1952f40bbdae09d769 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "Basic-2/reduction/lpr.ma".
+include "Basic-2/reduction/ltpr.ma".
 
 (* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************)
 
-axiom tpr_tps_lpr: ∀L1,L2. L1 ⇒ L2 → ∀T1,T2. T1 ⇒ T2 →
-                   ∀d,e,U1. L1 ⊢ T1 [d, e] ≫ U1 →
-                   ∃∃U2. U1 ⇒ U2 & L2 ⊢ T2 [d, e] ≫ U2.
+axiom tpr_tps_ltpr: ∀L1,L2. L1 ⇒ L2 → ∀T1,T2. T1 ⇒ T2 →
+                    ∀d,e,U1. L1 ⊢ T1 [d, e] ≫ U1 →
+                    ∃∃U2. U1 ⇒ U2 & L2 ⊢ T2 [d, e] ≫ U2.
 
 lemma tpr_tps_bind: ∀I,V1,V2,T1,T2,U1. V1 ⇒ V2 → T1 ⇒ T2 →
                     ⋆. 𝕓{I} V1 ⊢ T1 [0, 1] ≫ U1 →