X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fcomputation%2Ffsb_csx.ma;h=453383dc483bb300eb4414c40cd93742b824f71f;hb=a76f56fdad6348b167376093920650379c9936d4;hp=62eb01ba2dd63db9256fb7a25138900834a1a06d;hpb=376fd7774ef0fa2f30a4afb25aab6158e3cd04b7;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/fsb_csx.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/fsb_csx.ma index 62eb01ba2..453383dc4 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/fsb_csx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/fsb_csx.ma @@ -12,38 +12,42 @@ (* *) (**************************************************************************) -include "basic_2/computation/lsx_csx.ma". +include "basic_2/computation/fpbs_ext.ma". +include "basic_2/computation/csx_fpbs.ma". +include "basic_2/computation/llsx_csx.ma". include "basic_2/computation/fsb_alt.ma". -axiom lsx_fqup_conf: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊃+ ⦃G2, L2, T2⦄ → - G1 ⊢ ⋕⬊*[h, g, T1, 0] L1 → G2 ⊢ ⋕⬊*[h, g, T2, 0] L2. - -axiom fqup_lpxs_trans_nlleq: ∀h,g,G1,G2,K1,K2,T1,T2. ⦃G1, K1, T1⦄ ⊃+ ⦃G2, K2, T2⦄ → - ∀L2. ⦃G2, K2⦄ ⊢ ➡*[h, g] L2 → (K2 ⋕[T2, 0] L2 →⊥) → - ∃∃L1. ⦃G1, K1⦄ ⊢ ➡*[h, g] L1 & - K1 ⋕[T1, 0] L1 → ⊥ & ⦃G1, L1, T1⦄ ⊃+ ⦃G2, L2, T2⦄. - (* "BIG TREE" STRONGLY NORMALIZING TERMS ************************************) (* Advanced propreties on context-senstive extended bormalizing terms *******) -lemma csx_fsb: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → ⦃G, L⦄ ⊢ ⦥[h, g] T. +lemma csx_fsb_fpbs: ∀h,g,G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, g] T1 → + ∀G2,L2,T2. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄ → ⦃G2, L2⦄ ⊢ ⦥[h, g] T2. #h #g #G1 #L1 #T1 #H @(csx_ind_alt … H) -T1 -#T1 #HT1 @(lsx_ind h g G1 T1 0 … L1) /2 width=1 by csx_lsx/ -L1 -#L1 @(fqup_wf_ind … G1 L1 T1) -G1 -L1 -T1 -#G1 #L1 #T1 #IHu #H1 #IHl #IHc @fsb_intro -#G2 #L2 #T2 * -G2 -L2 -T2 -[ #G2 #L2 #T2 #H12 @IHu -IHu /2 width=5 by lsx_fqup_conf/ -H1 [| -IHl ] - [ #L0 #HL20 #HnL20 #_ elim (fqup_lpxs_trans_nlleq … H12 … HL20 HnL20) -L2 - /6 width=5 by fsb_fpbs_trans, lpxs_fpbs, fqup_fpbs, lpxs_cpxs_trans/ - | #T0 #HT20 #HnT20 elim (fqup_cpxs_trans_neq … H12 … HT20 HnT20) -T2 - /4 width=5 by fsb_fpbs_trans, fqup_fpbs/ +#T1 #HT1 #IHc #G2 #L2 #T2 @(fqup_wf_ind … G2 L2 T2) -G2 -L2 -T2 +#G0 #L0 #T0 #IHu #H10 lapply (csx_fpbs_conf … H10) // -HT1 +#HT0 generalize in match IHu; -IHu generalize in match H10; -H10 +@(llsx_ind_alt … (csx_llsx … HT0 0)) -L0 +#L0 #_ #IHl #H10 #IHu @fsb_intro +#G2 #L2 #T2 * -G2 -L2 -T2 [ -IHl -IHc | -IHu -IHl | ] +[ /3 width=5 by fpbs_fqup_trans/ +| #T2 #HT02 #HnT02 elim (fpbs_cpxs_trans_neq … H10 … HT02 HnT02) -T0 + /3 width=4 by/ +| #L2 #HL02 #HnL02 @(IHl … HL02 HnL02) -IHl -HnL02 [ -IHu -IHc | ] + [ /2 width=3 by fpbs_llpxs_trans/ + | #G3 #L3 #T3 #H03 #_ elim (llpxs_fqup_trans … H03 … HL02) -L2 + #L4 #T4 elim (eq_term_dec T0 T4) [ -IHc | -IHu ] + [ #H destruct /4 width=5 by fsb_fpbs_trans, llpxs_fpbs, fpbs_fqup_trans/ + | #HnT04 #HT04 #H04 elim (fpbs_cpxs_trans_neq … H10 … HT04 HnT04) -T0 + /4 width=8 by fpbs_fqup_trans, fpbs_llpxs_trans/ + ] ] -| -H1 -IHu -IHl /3 width=1 by/ -| -H1 -IHu /5 width=5 by fsb_fpbs_trans, lpxs_fpbs, lpxs_cpxs_trans/ ] qed. +lemma csx_fsb: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → ⦃G, L⦄ ⊢ ⦥[h, g] T. +/2 width=5 by csx_fsb_fpbs/ qed. + (* Advanced eliminators *****************************************************) lemma csx_ind_fpbu: ∀h,g. ∀R:relation3 genv lenv term.