X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Ffsb_fpbg.ma;h=cd2c4124815b517af7783d98d82d5843562fd9e7;hb=d71e53021b0c17e1a00c2d623e7139c6d18069d5;hp=b2093e068389d718efe3bc206056cd1459ea4453;hpb=5c186c72f508da0849058afeecc6877cd9ed6303;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fsb_fpbg.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fsb_fpbg.ma index b2093e068..cd2c41248 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fsb_fpbg.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fsb_fpbg.ma @@ -13,7 +13,7 @@ (**************************************************************************) include "basic_2/rt_computation/fpbg_fpbs.ma". -include "basic_2/rt_computation/fsb_ffdeq.ma". +include "basic_2/rt_computation/fsb_fdeq.ma". (* STRONGLY NORMALIZING CLOSURES FOR PARALLEL RST-TRANSITION ****************) @@ -24,7 +24,7 @@ lemma fsb_fpbs_trans: ∀h,o,G1,L1,T1. ≥[h, o] 𝐒⦃G1, L1, T1⦄ → #h #o #G1 #L1 #T1 #H @(fsb_ind_alt … H) -G1 -L1 -T1 #G1 #L1 #T1 #H1 #IH #G2 #L2 #T2 #H12 elim (fpbs_inv_fpbg … H12) -H12 -[ -IH /2 width=5 by fsb_ffdeq_trans/ +[ -IH /2 width=5 by fsb_fdeq_trans/ | -H1 * /2 width=5 by/ ] qed-. @@ -64,3 +64,12 @@ lemma fsb_ind_fpbg: ∀h,o. ∀Q:relation3 genv lenv term. #h #o #Q #IH #G1 #L1 #T1 #H @(fsb_ind_fpbg_fpbs … H) -H /3 width=1 by/ qed-. + +(* Inversion lemmas with proper parallel rst-computation for closures *******) + +lemma fsb_fpbg_refl_false (h) (o) (G) (L) (T): + ≥[h,o] 𝐒⦃G, L, T⦄ → ⦃G, L, T⦄ >[h,o] ⦃G, L, T⦄ → ⊥. +#h #o #G #L #T #H +@(fsb_ind_fpbg … H) -G -L -T #G1 #L1 #T1 #_ #IH #H +/2 width=5 by/ +qed-.