X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Ffpbg.ma;h=f975e1b2f4cfa1161899a504370c4c21e0019441;hb=d8d00d6f6694155be5be486a8239f5953efe28b7;hp=b14847da03a338bfb3bda7ebad945b551e82f4fd;hpb=a454837a256907d2f83d42ced7be847e10361ea9;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbg.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbg.ma index b14847da0..f975e1b2f 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbg.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbg.ma @@ -12,6 +12,7 @@ (* *) (**************************************************************************) +include "ground_2/xoa/ex_2_3.ma". include "basic_2/notation/relations/predsubtystarproper_7.ma". include "basic_2/rt_transition/fpb.ma". include "basic_2/rt_computation/fpbs.ma". @@ -52,13 +53,13 @@ lemma fpbg_fpbs_trans: ∀h,G,G2,L,L2,T,T2. ⦃G,L,T⦄ ≥[h] ⦃G2,L2,T2⦄ qed-. (* Basic_2A1: uses: fpbg_fleq_trans *) -lemma fpbg_fdeq_trans: ∀h,G1,G,L1,L,T1,T. ⦃G1,L1,T1⦄ >[h] ⦃G,L,T⦄ → +lemma fpbg_feqx_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=5 by fpbg_fpbq_trans, fpbq_fdeq/ qed-. +/3 width=5 by fpbg_fpbq_trans, fpbq_feqx/ qed-. (* Properties with t-bound rt-transition for terms **************************) -lemma cpm_tdneq_cpm_fpbg (h) (G) (L): +lemma cpm_tneqx_cpm_fpbg (h) (G) (L): ∀n1,T1,T. ⦃G,L⦄ ⊢ T1 ➡[n1,h] T → (T1 ≛ T → ⊥) → ∀n2,T2. ⦃G,L⦄ ⊢ T ➡[n2,h] T2 → ⦃G,L,T1⦄ >[h] ⦃G,L,T2⦄. -/4 width=5 by fpbq_fpbs, cpm_fpbq, cpm_fpb, ex2_3_intro/ qed. +/4 width=5 by fpbq_fpbs, cpm_fpbq, cpm_fpb, ex2_3_intro/ qed.