]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_ldrop.ma
some improvements towards an alternative definition of fpbs ...
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / reduction / lpx_ldrop.ma
index 468419f6c95ada2f7039831112aa62abead0904e..176c160129949c4b8ac3272c77eda9284a57bee6 100644 (file)
@@ -28,3 +28,21 @@ lemma ldrop_lpx_trans: ∀h,g,G. dedropable_sn (lpx h g G).
 
 lemma lpx_ldrop_trans_O1: ∀h,g,G. dropable_dx (lpx h g G).
 /2 width=3 by lpx_sn_dropable/ qed-.
+
+(* Properties on supclosure *************************************************)
+
+lemma fsupq_lpx_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊃⸮ ⦃G2, L2, T2⦄ →
+                       ∀K2. ⦃G2, L2⦄ ⊢ ➡[h, g] K2 →
+                                  ∃∃K1,T. ⦃G1, L1⦄ ⊢ ➡[h, g] K1 & ⦃G1, L1⦄ ⊢ T1 ➡[h, g] T & ⦃G1, K1, T⦄ ⊃⸮ ⦃G2, K2, T2⦄.
+#h #g #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
+/3 width=5 by fsup_fsupq, fsupq_refl, lpx_pair, fsup_lref_O, fsup_pair_sn, fsup_flat_dx, ex3_2_intro/
+[ #a #I #G2 #L2 #V2 #T2 #X #H elim (lpx_inv_pair1 … H) -H
+  #K2 #W2 #HLK2 #HVW2 #H destruct
+  /3 width=5 by fsup_fsupq, cpx_pair_sn, fsup_bind_dx, ex3_2_intro/
+| #G1 #G2 #L1 #K1 #K2 #T1 #T2 #U1 #d #e #HLK1 #HTU1 #_ #IH12 #K0 #HK20
+       elim (IH12 … HK20) -K2 #K2 #T #HK12
+       elim (ldrop_lpx_trans … HLK1 … HK12) -HK12
+       elim (lift_total T d e)
+       /3 width=11 by cpx_lift, fsupq_ldrop, ex3_2_intro/
+]
+qed-.