]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/reducibility/lcpr_cpr.ma
- full commit for the transtive closure of ltpss!
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / reducibility / lcpr_cpr.ma
index 69edad2aa15744087754bea50db6cd7528bb8988..22f3a398c7bdc17941bb7fdfe94d22ba81c3da85 100644 (file)
@@ -12,7 +12,7 @@
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/unfold/ltpss_ltpss.ma".
+include "basic_2/unfold/ltpss_sn_ltpss_sn.ma".
 include "basic_2/reducibility/cpr.ma".
 include "basic_2/reducibility/lcpr.ma".
 
@@ -23,9 +23,7 @@ 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 *
-<(ltpss_fwd_length … HL2) /4 width=5/
+<(ltpss_sn_fwd_length … HL2) #V #HV1 #HV2 #I
+lapply (ltpss_sn_tpss_trans_eq … HV2 … HL2) -HV2 #V2
+@(ex2_1_intro … (L.ⓑ{I}V)) /2 width=1/ (**) (* explicit constructor *)
 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.