X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Ffpbg_fpbs.ma;h=3641326f8177364d577bedc2f4d0e964177c08cf;hp=b188e7880b23f30c40e3a56c134fd00384bab86d;hb=ff612dc35167ec0c145864c9aa8ae5e1ebe20a48;hpb=afd4afa9489fa65019ad7cdc274e261f6993b871 diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbg_fpbs.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbg_fpbs.ma index b188e7880..3641326f8 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbg_fpbs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbg_fpbs.ma @@ -12,8 +12,8 @@ (* *) (**************************************************************************) -include "basic_2/static/ffdeq_fqup.ma". -include "basic_2/static/ffdeq_ffdeq.ma". +include "static_2/static/fdeq_fqup.ma". +include "static_2/static/fdeq_fdeq.ma". include "basic_2/rt_transition/fpbq_fpb.ma". include "basic_2/rt_computation/fpbg.ma". @@ -30,11 +30,11 @@ qed-. (* Advanced properties with degree-based equivalence on closures ************) (* Basic_2A1: uses: fleq_fpbg_trans *) -lemma ffdeq_fpbg_trans: ∀h,o,G,G2,L,L2,T,T2. ⦃G, L, T⦄ >[h, o] ⦃G2, L2, T2⦄ → - ∀G1,L1,T1. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ >[h, o] ⦃G2, L2, T2⦄. +lemma fdeq_fpbg_trans: ∀h,o,G,G2,L,L2,T,T2. ⦃G, L, T⦄ >[h, o] ⦃G2, L2, T2⦄ → + ∀G1,L1,T1. ⦃G1, L1, T1⦄ ≛[h, o] ⦃G, L, T⦄ → ⦃G1, L1, T1⦄ >[h, o] ⦃G2, L2, T2⦄. #h #o #G #G2 #L #L2 #T #T2 * #G0 #L0 #T0 #H0 #H02 #G1 #L1 #T1 #H1 -elim (ffdeq_fpb_trans … H1 … H0) -G -L -T -/4 width=9 by fpbs_strap2, fpbq_ffdeq, ex2_3_intro/ +elim (fdeq_fpb_trans … H1 … H0) -G -L -T +/4 width=9 by fpbs_strap2, fpbq_fdeq, ex2_3_intro/ qed-. (* Properties with parallel proper rst-reduction on closures ****************) @@ -51,7 +51,7 @@ lemma fpbq_fpbg_trans: ∀h,o,G1,G,G2,L1,L,L2,T1,T,T2. ⦃G1, L1, T1⦄ >[h, o] ⦃G2, L2, T2⦄. #h #o #G1 #G #G2 #L1 #L #L2 #T1 #T #T2 #H1 #H2 elim (fpbq_inv_fpb … H1) -H1 -/2 width=5 by ffdeq_fpbg_trans, fpb_fpbg_trans/ +/2 width=5 by fdeq_fpbg_trans, fpb_fpbg_trans/ qed-. (* Properties with parallel rst-compuutation on closures ********************) @@ -71,10 +71,10 @@ lemma fpbs_inv_fpbg: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, [ /2 width=1 by or_introl/ | #G #G2 #L #L2 #T #T2 #_ #H2 * #H1 elim (fpbq_inv_fpb … H2) -H2 #H2 - [ /3 width=5 by ffdeq_trans, or_introl/ - | elim (ffdeq_fpb_trans … H1 … H2) -G -L -T - /4 width=5 by ex2_3_intro, or_intror, ffdeq_fpbs/ - | /3 width=5 by fpbg_ffdeq_trans, or_intror/ + [ /3 width=5 by fdeq_trans, or_introl/ + | elim (fdeq_fpb_trans … H1 … H2) -G -L -T + /4 width=5 by ex2_3_intro, or_intror, fdeq_fpbs/ + | /3 width=5 by fpbg_fdeq_trans, or_intror/ | /4 width=5 by fpbg_fpbq_trans, fpb_fpbq, or_intror/ ] ] @@ -86,8 +86,8 @@ lemma fpbs_fpb_trans: ∀h,o,F1,F2,K1,K2,T1,T2. ⦃F1, K1, T1⦄ ≥[h, o] ⦃F2 ∀G2,L2,U2. ⦃F2, K2, T2⦄ ≻[h, o] ⦃G2, L2, U2⦄ → ∃∃G1,L1,U1. ⦃F1, K1, T1⦄ ≻[h, o] ⦃G1, L1, U1⦄ & ⦃G1, L1, U1⦄ ≥[h, o] ⦃G2, L2, U2⦄. #h #o #F1 #F2 #K1 #K2 #T1 #T2 #H elim (fpbs_inv_fpbg … H) -H -[ #H12 #G2 #L2 #U2 #H2 elim (ffdeq_fpb_trans … H12 … H2) -F2 -K2 -T2 - /3 width=5 by ffdeq_fpbs, ex2_3_intro/ +[ #H12 #G2 #L2 #U2 #H2 elim (fdeq_fpb_trans … H12 … H2) -F2 -K2 -T2 + /3 width=5 by fdeq_fpbs, ex2_3_intro/ | * #H1 #H2 #H3 #H4 #H5 #H6 #H7 #H8 #H9 @(ex2_3_intro … H4) -H4 /3 width=5 by fpbs_strap1, fpb_fpbq/ ]