X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fcomputation%2Ffsb_csx.ma;h=d40f30a850a60d0864cd9a633260d7a0e397ff7b;hb=1ca3d131ce61d857ebf691169e85ddb81250fd4e;hp=62eb01ba2dd63db9256fb7a25138900834a1a06d;hpb=c713c14cb3c69b1e9a4c693aed382eedc04512c1;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..d40f30a85 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/fsb_csx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/fsb_csx.ma @@ -15,9 +15,6 @@ include "basic_2/computation/lsx_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 & @@ -29,11 +26,11 @@ axiom fqup_lpxs_trans_nlleq: ∀h,g,G1,G2,K1,K2,T1,T2. ⦃G1, K1, T1⦄ ⊃+ ⦃ lemma csx_fsb: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → ⦃G, L⦄ ⊢ ⦥[h, g] T. #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 +#T1 #HT1 @(csx_ind_lsx … (yinj 0) … HT1) -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 ] +[ #G2 #L2 #T2 #H12 @IHu -IHu // /2 width=5 by csx_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