]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambda_delta/basic_2/unfold/ltpss_ltpss.ma
- subject equivalence for atomic arity assignment completed!
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / unfold / ltpss_ltpss.ma
index ce2271b705f0b8f4b3622b0c593718a57b845021..d928491c2aff5b8da888331d107d9071b62ee291 100644 (file)
@@ -13,6 +13,7 @@
 (**************************************************************************)
 
 include "basic_2/unfold/tpss_tpss.ma".
+include "basic_2/unfold/tpss_alt.ma".
 include "basic_2/unfold/ltpss_tpss.ma".
 
 (* PARTIAL UNFOLD ON LOCAL ENVIRONMENTS *************************************)
@@ -42,6 +43,40 @@ lemma ltpss_tpss_trans_down: ∀L0,L1,T2,U2,d1,e1,d2,e2. d2 + e2 ≤ d1 →
 ]
 qed.
 
+fact ltpss_tpss_trans_eq_aux: ∀Y1,X2,L1,T2,U2,d,e.
+                              L1 ⊢ T2 [d, e] ▶* U2 → ∀L0. L0 [d, e] ▶* L1 →
+                              Y1 = L1 → X2 = T2 → L0 ⊢ T2 [d, e] ▶* U2.
+#Y1 #X2 @(cw_wf_ind … Y1 X2) -Y1 -X2 #Y1 #X2 #IH
+#L1 #T2 #U2 #d #e #H @(tpss_ind_alt … H) -L1 -T2 -U2 -d -e
+[ //
+| #L1 #K1 #V1 #V2 #W2 #i #d #e #Hdi #Hide #HLK1 #HV12 #HVW2 #_ #L0 #HL10 #H1 #H2 destruct
+  lapply (ldrop_fwd_lw … HLK1) #H1 normalize in H1;
+  elim (ltpss_ldrop_trans_be … HL10 … HLK1 ? ?) -HL10 -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 (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/ ]
+| #L #I #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #_ #_ #L0 #HL0 #H1 #H2 destruct
+  lapply (tpss_lsubs_conf … HT12 (L. ⓑ{I} V1) ?) -HT12 /2 width=1/ #HT12
+  lapply (IH … HV12 … HL0 ? ?) -HV12 [1,3: // |2,4: skip |5: /2 width=2/ ] #HV12
+  lapply (IH … HT12 (L0. ⓑ{I} V1) ? ? ?) -IH -HT12 [1,3,5: /2 width=2/ |2,4: skip | normalize // ] -HL0 #HT12
+  lapply (tpss_lsubs_conf … HT12 (L0. ⓑ{I} V2) ?) -HT12 /2 width=1/
+| #L #I #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #_ #_ #L0 #HL0 #H1 #H2 destruct
+  lapply (IH … HV12 … HL0 ? ?) -HV12 [1,3: // |2,4: skip |5: /2 width=3/ ]
+  lapply (IH … HT12 … HL0 ? ?) -IH -HT12 [1,3,5: normalize // |2,4: skip ] -HL0 /2 width=1/
+]
+qed.
+
+lemma ltpss_tpss_trans_eq: ∀L1,T2,U2,d,e. L1 ⊢ T2 [d, e] ▶* U2 →
+                           ∀L0. L0 [d, e] ▶* L1 → L0 ⊢ T2 [d, e] ▶* U2.
+/2 width=5/ qed.
+
+lemma ltpss_tps_trans_eq: ∀L0,L1,T2,U2,d,e. L0 [d, e] ▶* L1 →
+                          L1 ⊢ T2 [d, e] ▶ U2 → L0 ⊢ T2 [d, e] ▶* U2.
+/3 width=3/ qed.
+
 (* Main properties **********************************************************)
 
 fact ltpss_conf_aux: ∀K,K1,L1,d1,e1. K1 [d1, e1] ▶* L1 →