]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/computation/csx_fpbs.ma
- "big tree" theorem is now proved up to some conjectures involving
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / csx_fpbs.ma
index 99f39de1db7f17b06f03014556b62c2e66a21b00..7d88827934ccd97128876985bc1781c2fb4ec047 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/computation/lpxs.ma".
+include "basic_2/computation/fpbs.ma".
 include "basic_2/computation/csx_lpx.ma".
 
 (* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERMS ********************)
 
 (* Advanced properties ******************************************************)
 
+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/
+qed-.
+
+lemma csx_fpbs_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 #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.
-#h #g #G #L1 #T #HT #L2 #H @(lpxs_ind … H) -L2 /2 width=3 by csx_lpx_conf/
+/3 width=5 by csx_fpbs_conf, lpxs_fpbs/
 qed-.