X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Freduction%2Ffpb.ma;h=4cf9d4bb0efbd5938eecdb69cff3840b960c3798;hb=78b27990925c54b2a34cff609fc9bcfbeb6b48f3;hp=947c4eda31bbf298095a30fde527d7325a861ff8;hpb=c28e3d93b588796bfbbfd6b2ec9dd86f405b2b00;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/fpb.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/fpb.ma index 947c4eda3..4cf9d4bb0 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/fpb.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/fpb.ma @@ -13,15 +13,15 @@ (**************************************************************************) include "basic_2/notation/relations/btpred_8.ma". -include "basic_2/relocation/fsup.ma". -include "basic_2/reduction/lpx.ma". +include "basic_2/relocation/fquq.ma". +include "basic_2/reduction/llpx.ma". (* "BIG TREE" PARALLEL REDUCTION FOR CLOSURES *******************************) inductive fpb (h) (g) (G1) (L1) (T1): relation3 genv lenv term ≝ -| fpb_fsup: ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊃ ⦃G2, L2, T2⦄ → fpb h g G1 L1 T1 G2 L2 T2 +| fpb_fquq: ∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊃⸮ ⦃G2, L2, T2⦄ → fpb h g G1 L1 T1 G2 L2 T2 | fpb_cpx : ∀T2. ⦃G1, L1⦄ ⊢ T1 ➡[h, g] T2 → fpb h g G1 L1 T1 G1 L1 T2 -| fpb_lpx : ∀L2. ⦃G1, L1⦄ ⊢ ➡[h, g] L2 → fpb h g G1 L1 T1 G1 L2 T1 +| fpb_llpx: ∀L2. ⦃G1, L1⦄ ⊢ ➡[h, g, T1, 0] L2 → fpb h g G1 L1 T1 G1 L2 T1 . interpretation @@ -35,6 +35,7 @@ lemma fpb_refl: ∀h,g. tri_reflexive … (fpb h g). lemma cpr_fpb: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L, T1⦄ ≽[h, g] ⦃G, L, T2⦄. /3 width=1 by fpb_cpx, cpr_cpx/ qed. - -lemma lpr_fpb: ∀h,g,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡ L2 → ⦃G, L1, T⦄ ≽[h, g] ⦃G, L2, T⦄. -/3 width=1 by fpb_lpx, lpr_lpx/ qed. +(* +lamma llpr_fpb: ∀h,g,G,L1,L2,T. ⦃G, L1⦄ ⊢ ➡[T, 0] L2 → ⦃G, L1, T⦄ ≽[h, g] ⦃G, L2, T⦄. +/3 width=1 by fpb_llpx, llpr_llpx/ qed. +*)