]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/reducibility/lcpr_cpr.ma
Finalized copy sub-machine of the universal turing machine. Some new results
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / reducibility / lcpr_cpr.ma
index 3d833fc025d87776597f02dc4165fcf32598ef3e..69edad2aa15744087754bea50db6cd7528bb8988 100644 (file)
@@ -12,7 +12,7 @@
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/unfold/ltpsss_ltpsss.ma".
+include "basic_2/unfold/ltpss_ltpss.ma".
 include "basic_2/reducibility/cpr.ma".
 include "basic_2/reducibility/lcpr.ma".
 
@@ -23,5 +23,9 @@ include "basic_2/reducibility/lcpr.ma".
 lemma lcpr_pair: ∀L1,L2. L1 ⊢ ➡ L2 → ∀V1,V2. L2 ⊢ V1 ➡ V2 →
                  ∀I. L1. ⓑ{I} V1 ⊢ ➡ L2. ⓑ{I} V2.
 #L1 #L2 * #L #HL1 #HL2 #V1 #V2 *
-<(ltpsss_fwd_length … HL2) /4 width=5/
+<(ltpss_fwd_length … HL2) /4 width=5/
 qed.
+
+lemma ltpss_lcpr: ∀L1,L2,d,e. L1 ▶* [d, e] L2 → L1 ⊢ ➡ L2.
+#L1 #L2 #d #e #H elim H -L1 -L2 -d -e // /3 width=3/
+qed.