]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/computation/csx_fpbs.ma
commit of the "computation" component with lazy pointwise extensions
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / csx_fpbs.ma
index 7d88827934ccd97128876985bc1781c2fb4ec047..eabda3146522bfd76a69e7cf7e182a799f802df2 100644 (file)
@@ -13,7 +13,7 @@
 (**************************************************************************)
 
 include "basic_2/computation/fpbs.ma".
-include "basic_2/computation/csx_lpx.ma".
+include "basic_2/computation/csx_llpx.ma".
 
 (* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERMS ********************)
 
@@ -22,7 +22,7 @@ include "basic_2/computation/csx_lpx.ma".
 lemma csx_fpb_conf: ∀h,g,G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, g] T1 →
                     ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≽[h, g] ⦃G2, L2, T2⦄ → ⦃G2, L2⦄ ⊢ ⬊*[h, g] T2.
 #h #g #G1 #L1 #T1 #HT1 #G2 #L2 #T2 *
-/2 width=5 by csx_cpx_trans, csx_fquq_conf, csx_lpx_conf/
+/2 width=5 by csx_cpx_trans, csx_fquq_conf, csx_llpx_conf/
 qed-.
 
 lemma csx_fpbs_conf: ∀h,g,G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, g] T1 →
@@ -30,8 +30,3 @@ lemma csx_fpbs_conf: ∀h,g,G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, g] T1 →
 #h #g #G1 #L1 #T1 #HT1 #G2 #L2 #T2 #H @(fpbs_ind … H) -G2 -L2 -T2
 /2 width=5 by csx_fpb_conf/
 qed-.
-
-lemma csx_lpxs_conf: ∀h,g,G,L1. ∀T:term. ⦃G, L1⦄ ⊢ ⬊*[h, g] T →
-                     ∀L2. ⦃G, L1⦄ ⊢ ➡*[h, g] L2 → ⦃G, L2⦄ ⊢ ⬊*[h, g] T.
-/3 width=5 by csx_fpbs_conf, lpxs_fpbs/
-qed-.