X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Ffpbs.ma;h=f7437d26db043bb4702af1512cf983e01018ae78;hb=e23331eef5817eaa6c5e1c442d1d6bbb18650573;hp=9d6f5c2801c0cd12f1ec64ad7e491630bb4739c1;hpb=db020b4218272e2e35641ce3bc3b0a9b3afda899;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs.ma index 9d6f5c280..f7437d26d 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs.ma @@ -12,65 +12,38 @@ (* *) (**************************************************************************) -include "ground_2/lib/star.ma". -include "basic_2/notation/relations/predsubtystar_7.ma". -include "basic_2/rt_transition/fpbq.ma". +include "ground/lib/star.ma". +include "basic_2/notation/relations/predsubtystar_6.ma". +include "basic_2/rt_transition/fpb.ma". (* PARALLEL RST-COMPUTATION FOR CLOSURES ************************************) -definition fpbs: ∀h. tri_relation genv lenv term ≝ - λh. tri_TC … (fpbq h). +definition fpbs: tri_relation genv lenv term ≝ + tri_TC … fpb. -interpretation "parallel rst-computation (closure)" - 'PRedSubTyStar h G1 L1 T1 G2 L2 T2 = (fpbs h G1 L1 T1 G2 L2 T2). +interpretation + "parallel rst-computation (closure)" + 'PRedSubTyStar G1 L1 T1 G2 L2 T2 = (fpbs G1 L1 T1 G2 L2 T2). -(* Basic eliminators ********************************************************) - -lemma fpbs_ind: ∀h,G1,L1,T1. ∀Q:relation3 genv lenv term. Q G1 L1 T1 → - (∀G,G2,L,L2,T,T2. ⦃G1,L1,T1⦄ ≥[h] ⦃G,L,T⦄ → ⦃G,L,T⦄ ≽[h] ⦃G2,L2,T2⦄ → Q G L T → Q G2 L2 T2) → - ∀G2,L2,T2. ⦃G1,L1,T1⦄ ≥[h] ⦃G2,L2,T2⦄ → Q G2 L2 T2. -/3 width=8 by tri_TC_star_ind/ qed-. - -lemma fpbs_ind_dx: ∀h,G2,L2,T2. ∀Q:relation3 genv lenv term. Q G2 L2 T2 → - (∀G1,G,L1,L,T1,T. ⦃G1,L1,T1⦄ ≽[h] ⦃G,L,T⦄ → ⦃G,L,T⦄ ≥[h] ⦃G2,L2,T2⦄ → Q G L T → Q G1 L1 T1) → - ∀G1,L1,T1. ⦃G1,L1,T1⦄ ≥[h] ⦃G2,L2,T2⦄ → Q G1 L1 T1. -/3 width=8 by tri_TC_star_ind_dx/ qed-. (* Basic properties *********************************************************) -lemma fpbs_refl: ∀h. tri_reflexive … (fpbs h). -/2 width=1 by tri_inj/ qed. - -lemma fpbq_fpbs: ∀h,G1,G2,L1,L2,T1,T2. ⦃G1,L1,T1⦄ ≽[h] ⦃G2,L2,T2⦄ → - ⦃G1,L1,T1⦄ ≥[h] ⦃G2,L2,T2⦄. +(* Basic_2A1: uses: fpbq_fpbs *) +lemma fpb_fpbs: + ∀G1,G2,L1,L2,T1,T2. ❪G1,L1,T1❫ ≽ ❪G2,L2,T2❫ → + ❪G1,L1,T1❫ ≥ ❪G2,L2,T2❫. /2 width=1 by tri_inj/ qed. -lemma fpbs_strap1: ∀h,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1,L1,T1⦄ ≥[h] ⦃G,L,T⦄ → - ⦃G,L,T⦄ ≽[h] ⦃G2,L2,T2⦄ → ⦃G1,L1,T1⦄ ≥[h] ⦃G2,L2,T2⦄. +lemma fpbs_strap1: + ∀G1,G,G2,L1,L,L2,T1,T,T2. ❪G1,L1,T1❫ ≥ ❪G,L,T❫ → + ❪G,L,T❫ ≽ ❪G2,L2,T2❫ → ❪G1,L1,T1❫ ≥ ❪G2,L2,T2❫. /2 width=5 by tri_step/ qed-. -lemma fpbs_strap2: ∀h,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1,L1,T1⦄ ≽[h] ⦃G,L,T⦄ → - ⦃G,L,T⦄ ≥[h] ⦃G2,L2,T2⦄ → ⦃G1,L1,T1⦄ ≥[h] ⦃G2,L2,T2⦄. +lemma fpbs_strap2: + ∀G1,G,G2,L1,L,L2,T1,T,T2. ❪G1,L1,T1❫ ≽ ❪G,L,T❫ → + ❪G,L,T❫ ≥ ❪G2,L2,T2❫ → ❪G1,L1,T1❫ ≥ ❪G2,L2,T2❫. /2 width=5 by tri_TC_strap/ qed-. -(* Basic_2A1: uses: lleq_fpbs fleq_fpbs *) -lemma fdeq_fpbs: ∀h,G1,G2,L1,L2,T1,T2. ⦃G1,L1,T1⦄ ≛ ⦃G2,L2,T2⦄ → ⦃G1,L1,T1⦄ ≥[h] ⦃G2,L2,T2⦄. -/3 width=1 by fpbq_fpbs, fpbq_fdeq/ qed. - -(* Basic_2A1: uses: fpbs_lleq_trans *) -lemma fpbs_fdeq_trans: ∀h,G1,G,L1,L,T1,T. ⦃G1,L1,T1⦄ ≥[h] ⦃G,L,T⦄ → - ∀G2,L2,T2. ⦃G,L,T⦄ ≛ ⦃G2,L2,T2⦄ → ⦃G1,L1,T1⦄ ≥[h] ⦃G2,L2,T2⦄. -/3 width=9 by fpbs_strap1, fpbq_fdeq/ qed-. - -(* Basic_2A1: uses: lleq_fpbs_trans *) -lemma fdeq_fpbs_trans: ∀h,G,G2,L,L2,T,T2. ⦃G,L,T⦄ ≥[h] ⦃G2,L2,T2⦄ → - ∀G1,L1,T1. ⦃G1,L1,T1⦄ ≛ ⦃G,L,T⦄ → ⦃G1,L1,T1⦄ ≥[h] ⦃G2,L2,T2⦄. -/3 width=5 by fpbs_strap2, fpbq_fdeq/ qed-. - -lemma tdeq_rdeq_lpx_fpbs: ∀h,T1,T2. T1 ≛ T2 → ∀L1,L0. L1 ≛[T2] L0 → - ∀G,L2. ⦃G,L0⦄ ⊢ ⬈[h] L2 → ⦃G,L1,T1⦄ ≥[h] ⦃G,L2,T2⦄. -/4 width=5 by fdeq_fpbs, fpbs_strap1, fpbq_lpx, fdeq_intro_dx/ qed. - (* Basic_2A1: removed theorems 3: fpb_fpbsa_trans fpbs_fpbsa fpbsa_inv_fpbs *)