X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Ffpbs_fpb.ma;h=632fe7dc6d656672c7727201a3ecd5f5cec83325;hp=b9c657735fb44ab5e58a20637fb189eb12354ad4;hb=dd74d1964ef07de249385a48f28302b427c0d287;hpb=54c9014b6657403c6e235c652176218e750d4b8a diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_fpb.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_fpb.ma index b9c657735..632fe7dc6 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_fpb.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_fpb.ma @@ -12,29 +12,12 @@ (* *) (**************************************************************************) -include "basic_2/reduction/fpbq_alt.ma". -include "basic_2/computation/fpbs_alt.ma". +include "basic_2/rt_transition/fpbq_fpb.ma". +include "basic_2/rt_computation/fpbs.ma". -(* "QRST" PARALLEL COMPUTATION FOR CLOSURES *********************************) +(* PARALLEL RST-COMPUTATION FOR CLOSURES ************************************) -(* Properties on extended context-sensitive parallel computation for terms **) - -lemma fpbs_cpx_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ → - ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h, o] U2 → (T2 = U2 → ⊥) → - ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ≥[h, o] ⦃G2, L2, U2⦄. -#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #U2 #HTU2 #HnTU2 elim (fpbs_inv_alt … H) -H -#L00 #L0 #T0 #HT10 #H10 #HL00 #HL02 lapply (lleq_cpx_trans … HTU2 … HL02) -HTU2 -#HTU2 lapply (cpx_lleq_conf_sn … HTU2 … HL02) -HL02 -#HL02 lapply (lpxs_cpx_trans … HTU2 … HL00) -HTU2 -#HTU2 elim (fqus_cpxs_trans_neq … H10 … HTU2 HnTU2) -H10 -HTU2 -HnTU2 -#U0 #HTU0 #HnTU0 #HU02 elim (eq_term_dec T1 T0) #HnT10 destruct -[ -HT10 elim (cpxs_neq_inv_step_sn … HTU0 HnTU0) -HTU0 -HnTU0 -| -HnTU0 elim (cpxs_neq_inv_step_sn … HT10 HnT10) -HT10 -HnT10 -] -/4 width=10 by fpbs_intro_alt, cpxs_trans, ex3_intro/ -qed-. - -(* Properties on "rst" proper parallel reduction on closures ****************) +(* Properties with proper parallel rst-reduction on closures ****************) lemma fpb_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.