]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/unfold/ltpss_ltpss.ma
the partial commit continues ...
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / unfold / ltpss_ltpss.ma
index 8885b233916d43b4d5979f6e14ba1e37f9d65ac4..f866adc126cd41eb357b208098568ee8a483a81c 100644 (file)
@@ -54,7 +54,7 @@ fact ltpss_tpss_trans_eq_aux: ∀Y1,X2,L1,T2,U2,d,e.
   elim (ltpss_ldrop_trans_be … HL01 … HLK1 ? ?) -HL01 -HLK1 // /2 width=2/ #X #H #HLK0
   elim (ltpss_inv_tpss22 … H ?) -H /2 width=1/ #K0 #V0 #HK01 #HV01 #H destruct
   lapply (tpss_fwd_tw … HV01) #H2
-  lapply (transitive_le (#[K1] + #[V0]) … H1) -H1 /2 width=1/ -H2 #H
+  lapply (transitive_le (#{K1} + #{V0}) … H1) -H1 /2 width=1/ -H2 #H
   lapply (tpss_trans_eq … HV01 HV12) -V1 #HV02
   lapply (IH … HV02 … HK01 ? ?) -IH -HV02 -HK01
   [1,3: // |2,4: skip | normalize /2 width=1/ | /2 width=6/ ]
@@ -133,3 +133,14 @@ lemma ltpss_conf: ∀L0,L1,d1,e1. L0 ▶* [d1, e1] L1 →
                   ∀L2,d2,e2. L0 ▶* [d2, e2] L2 →
                   ∃∃L. L1 ▶* [d2, e2] L & L2 ▶* [d1, e1] L.
 /2 width=7/ qed.
+
+theorem ltpss_trans_eq: ∀L1,L,d,e. L1 ▶* [d, e] L →
+                        ∀L2. L ▶* [d, e] L2 → L1 ▶* [d, e] L2.
+#L1 #L #d #e #H elim H -L1 -L -d -e //
+[ #L1 #L #I #V1 #V #e #_ #HV1 #IHL1 #X #H
+  elim (ltpss_inv_tpss21 … H ?) -H // <minus_plus_m_m #L2 #V2 #HL2 #HV2 #H destruct
+  elim (ltpss_tpss_conf … HV1 … HL2) -HV1 #V0 #HV10 #HV0
+  elim (tpss_conf_eq … HV2 … HV0) -V #V #HV2 #HV0
+  lapply (tpss_trans_eq … HV10 … HV0) -V0 #HV1
+  @ltpss_tpss2 /2 width=1/
+        
\ No newline at end of file