]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/computation/fpbs_alt.ma
- nnAuto: we catch TypeCheckerFailure generated at the end of
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / fpbs_alt.ma
index f7419939036b416597cf5298c13f96380f3ba21f..cc6ae13627effee4aef12f62e67dd92d14bfc788 100644 (file)
 (**************************************************************************)
 
 include "basic_2/notation/relations/btpredstaralt_8.ma".
-include "basic_2/substitution/lleq_fqus.ma".
-include "basic_2/substitution/lleq_lleq.ma".
+include "basic_2/multiple/lleq_fqus.ma".
 include "basic_2/computation/cpxs_lleq.ma".
 include "basic_2/computation/lpxs_lleq.ma".
 include "basic_2/computation/fpbs.ma".
 
-(* "BIG TREE" PARALLEL COMPUTATION FOR CLOSURES *****************************)
+(* "QREST" PARALLEL COMPUTATION FOR CLOSURES ********************************)
 
 (* Note: alternative definition of fpbs *)
 definition fpbsa: ∀h. sd h → tri_relation genv lenv term ≝
                   λh,g,G1,L1,T1,G2,L2,T2.
                   ∃∃L0,L,T. ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] T &
-                         â¦\83G1, L1, Tâ¦\84 â\8a\83* ⦃G2, L0, T2⦄ &
-                         â¦\83G2, L0â¦\84 â\8a¢ â\9e¡*[h, g] L & L â\8b\95[T2, 0] L2.
+                         â¦\83G1, L1, Tâ¦\84 â\8a\90* ⦃G2, L0, T2⦄ &
+                         â¦\83G2, L0â¦\84 â\8a¢ â\9e¡*[h, g] L & L â\89¡[T2, 0] L2.
 
 interpretation "'big tree' parallel computation (closure) alternative"
    'BTPRedStarAlt h g G1 L1 T1 G2 L2 T2 = (fpbsa h g G1 L1 T1 G2 L2 T2).
@@ -64,20 +63,20 @@ qed.
 theorem fpbsa_inv_fpbs: ∀h,g,G1,G2,L1,L2,T1,T2.
                         ⦃G1, L1, T1⦄ ≥≥[h, g] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄.
 #h #g #G1 #G2 #L1 #L2 #T1 #T2 *
-/3 width=5 by cpxs_fqus_lpxs_fpbs, fpbs_strap1, fpb_lleq/
+/3 width=5 by cpxs_fqus_lpxs_fpbs, fpbs_strap1, fpbq_lleq/
 qed-.
 
 (* Advanced properties ******************************************************)
 
 lemma fpbs_intro_alt: ∀h,g,G1,G2,L1,L0,L,L2,T1,T,T2.
-                      â¦\83G1, L1â¦\84 â\8a¢ T1 â\9e¡*[h, g] T â\86\92 â¦\83G1, L1, Tâ¦\84 â\8a\83* ⦃G2, L0, T2⦄ →
-                      â¦\83G2, L0â¦\84 â\8a¢ â\9e¡*[h, g] L â\86\92 L â\8b\95[T2, 0] L2 →  ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄ .
+                      â¦\83G1, L1â¦\84 â\8a¢ T1 â\9e¡*[h, g] T â\86\92 â¦\83G1, L1, Tâ¦\84 â\8a\90* ⦃G2, L0, T2⦄ →
+                      â¦\83G2, L0â¦\84 â\8a¢ â\9e¡*[h, g] L â\86\92 L â\89¡[T2, 0] L2 →  ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄ .
 /3 width=7 by fpbsa_inv_fpbs, ex4_3_intro/ qed.
 
 (* Advanced inversion lemmas *************************************************)
 
 lemma fpbs_inv_alt: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄ →
                     ∃∃L0,L,T. ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] T &
-                              â¦\83G1, L1, Tâ¦\84 â\8a\83* ⦃G2, L0, T2⦄ &
-                              â¦\83G2, L0â¦\84 â\8a¢ â\9e¡*[h, g] L & L â\8b\95[T2, 0] L2.
+                              â¦\83G1, L1, Tâ¦\84 â\8a\90* ⦃G2, L0, T2⦄ &
+                              â¦\83G2, L0â¦\84 â\8a¢ â\9e¡*[h, g] L & L â\89¡[T2, 0] L2.
 /2 width=1 by  fpbs_fpbsa/ qed-.