]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_ldrop.ma
- the relation for pointwise extensions now takes a binder as argument
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / reduction / lpx_ldrop.ma
index 188e8d11b6b343c1799332bb62745f283071e4d2..5097cc1143e1d49a8c69af92af3492671cd9ba43 100644 (file)
@@ -12,7 +12,7 @@
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/relocation/ldrop_lpx_sn.ma".
+include "basic_2/relocation/lpx_sn_ldrop.ma".
 include "basic_2/reduction/cpx_lift.ma".
 include "basic_2/reduction/lpx.ma".
 
@@ -31,9 +31,9 @@ lemma lpx_ldrop_trans_O1: ∀h,g,G. dropable_dx (lpx h g G).
 
 (* Properties on supclosure *************************************************)
 
-lemma fqu_lpx_trans: â\88\80h,g,G1,G2,L1,L2,T1,T2. â¦\83G1, L1, T1â¦\84 â\8a\83 ⦃G2, L2, T2⦄ →
+lemma fqu_lpx_trans: â\88\80h,g,G1,G2,L1,L2,T1,T2. â¦\83G1, L1, T1â¦\84 â\8a\90 ⦃G2, L2, T2⦄ →
                      ∀K2. ⦃G2, L2⦄ ⊢ ➡[h, g] K2 →
-                     â\88\83â\88\83K1,T. â¦\83G1, L1â¦\84 â\8a¢ â\9e¡[h, g] K1 & â¦\83G1, L1â¦\84 â\8a¢ T1 â\9e¡[h, g] T & â¦\83G1, K1, Tâ¦\84 â\8a\83 ⦃G2, K2, T2⦄.
+                     â\88\83â\88\83K1,T. â¦\83G1, L1â¦\84 â\8a¢ â\9e¡[h, g] K1 & â¦\83G1, L1â¦\84 â\8a¢ T1 â\9e¡[h, g] T & â¦\83G1, K1, Tâ¦\84 â\8a\90 ⦃G2, K2, T2⦄.
 #h #g #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
 /3 width=5 by fqu_lref_O, fqu_pair_sn, fqu_flat_dx, lpx_pair, ex3_2_intro/
 [ #a #I #G2 #L2 #V2 #T2 #X #H elim (lpx_inv_pair1 … H) -H
@@ -45,18 +45,18 @@ lemma fqu_lpx_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊃ ⦃G2, L2, T
 ]
 qed-.
 
-lemma fquq_lpx_trans: â\88\80h,g,G1,G2,L1,L2,T1,T2. â¦\83G1, L1, T1â¦\84 â\8a\83⸮ ⦃G2, L2, T2⦄ →
+lemma fquq_lpx_trans: â\88\80h,g,G1,G2,L1,L2,T1,T2. â¦\83G1, L1, T1â¦\84 â\8a\90⸮ ⦃G2, L2, T2⦄ →
                       ∀K2. ⦃G2, L2⦄ ⊢ ➡[h, g] K2 →
-                      â\88\83â\88\83K1,T. â¦\83G1, L1â¦\84 â\8a¢ â\9e¡[h, g] K1 & â¦\83G1, L1â¦\84 â\8a¢ T1 â\9e¡[h, g] T & â¦\83G1, K1, Tâ¦\84 â\8a\83⸮ ⦃G2, K2, T2⦄.
+                      â\88\83â\88\83K1,T. â¦\83G1, L1â¦\84 â\8a¢ â\9e¡[h, g] K1 & â¦\83G1, L1â¦\84 â\8a¢ T1 â\9e¡[h, g] T & â¦\83G1, K1, Tâ¦\84 â\8a\90⸮ ⦃G2, K2, T2⦄.
 #h #g #G1 #G2 #L1 #L2 #T1 #T2 #H #K2 #HLK2 elim (fquq_inv_gen … H) -H
 [ #HT12 elim (fqu_lpx_trans … HT12 … HLK2) /3 width=5 by fqu_fquq, ex3_2_intro/
 | * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/
 ]
 qed-.
 
-lemma lpx_fqu_trans: â\88\80h,g,G1,G2,L1,L2,T1,T2. â¦\83G1, L1, T1â¦\84 â\8a\83 ⦃G2, L2, T2⦄ →
+lemma lpx_fqu_trans: â\88\80h,g,G1,G2,L1,L2,T1,T2. â¦\83G1, L1, T1â¦\84 â\8a\90 ⦃G2, L2, T2⦄ →
                      ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, g] L1 →
-                     â\88\83â\88\83K2,T. â¦\83G1, K1â¦\84 â\8a¢ T1 â\9e¡[h, g] T & â¦\83G1, K1, Tâ¦\84 â\8a\83 ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, g] L2.
+                     â\88\83â\88\83K2,T. â¦\83G1, K1â¦\84 â\8a¢ T1 â\9e¡[h, g] T & â¦\83G1, K1, Tâ¦\84 â\8a\90 ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, g] L2.
 #h #g #G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2
 /3 width=7 by fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, lpx_pair, ex3_2_intro/
 [ #I #G1 #L1 #V1 #X #H elim (lpx_inv_pair2 … H) -H
@@ -68,9 +68,9 @@ lemma lpx_fqu_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊃ ⦃G2, L2, T
 ]
 qed-.
 
-lemma lpx_fquq_trans: â\88\80h,g,G1,G2,L1,L2,T1,T2. â¦\83G1, L1, T1â¦\84 â\8a\83⸮ ⦃G2, L2, T2⦄ →
+lemma lpx_fquq_trans: â\88\80h,g,G1,G2,L1,L2,T1,T2. â¦\83G1, L1, T1â¦\84 â\8a\90⸮ ⦃G2, L2, T2⦄ →
                       ∀K1. ⦃G1, K1⦄ ⊢ ➡[h, g] L1 →
-                      â\88\83â\88\83K2,T. â¦\83G1, K1â¦\84 â\8a¢ T1 â\9e¡[h, g] T & â¦\83G1, K1, Tâ¦\84 â\8a\83⸮ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, g] L2.
+                      â\88\83â\88\83K2,T. â¦\83G1, K1â¦\84 â\8a¢ T1 â\9e¡[h, g] T & â¦\83G1, K1, Tâ¦\84 â\8a\90⸮ ⦃G2, K2, T2⦄ & ⦃G2, K2⦄ ⊢ ➡[h, g] L2.
 #h #g #G1 #G2 #L1 #L2 #T1 #T2 #H #K1 #HKL1 elim (fquq_inv_gen … H) -H
 [ #HT12 elim (lpx_fqu_trans … HT12 … HKL1) /3 width=5 by fqu_fquq, ex3_2_intro/
 | * #H1 #H2 #H3 destruct /2 width=5 by ex3_2_intro/