X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fcomputation%2Fcpxs_cpys.ma;h=7ab36e5be0595617575333d989766d0e75b1fb34;hb=62211b3b2d76ba387663c9e553fbe4d2dbd92c82;hp=0bbe9835cf4d98570f743bd2b3f9cc06ee3b15bf;hpb=2ba2dc23443ad764adab652e06d6f5ed10bd912d;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_cpys.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_cpys.ma index 0bbe9835c..7ab36e5be 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_cpys.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_cpys.ma @@ -13,23 +13,40 @@ (**************************************************************************) include "basic_2/substitution/cpys_cpys.ma". +include "basic_2/reduction/cpx_cpys.ma". include "basic_2/computation/cpxs_cpxs.ma". (* SN EXTENDED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS ********************) (* Main properties **********************************************************) -lemma cpx_fwd_cpys_tpxs: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → +axiom cpys_split_down: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶*×[d, e] T2 → + ∀i. d ≤ i → i ≤ d + e → + ∃∃T. ⦃G, L⦄ ⊢ T1 ▶*×[i, d+e-i] T & ⦃G, L⦄ ⊢ T ▶*×[d, i-d] T2. + +lemma cpys_tpxs_trans: ∀h,g,G,L,T1,T,d,e. ⦃G, L⦄ ⊢ T1 ▶*×[d, e] T → + ∀T2. ⦃G, ⋆⦄ ⊢ T ➡*[h, g] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2. +#h #g #G #L #T1 #T #d #e #HT1 #T2 #H @(cpxs_ind … H) -T2 +/3 width=5 by cpxs_strap1, cpys_cpx, lsubr_cpx_trans, cpx_cpxs/ +qed-. + +axiom cpx_fwd_cpys_tpxs: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → ∃∃T. ⦃G, L⦄ ⊢ T1 ▶*×[0, ∞] T & ⦃G, ⋆⦄ ⊢ T ➡*[h, g] T2. +(* #h #g #G #L #T1 #T2 #H elim H -G -L -T1 -T2 [ /2 width=3 by ex2_intro/ | /4 width=3 by cpx_cpxs, cpx_sort, ex2_intro/ | #I #G #L #K #V1 #V2 #W2 #i #HLK #_ #HVW2 * #V #HV1 #HV2 elim (lift_total V 0 (i+1)) #W #HVW - @(ex2_intro … W) /2 width=7 by cpys_subst_Y2/ + lapply (cpxs_lift … HV2 (⋆) (Ⓣ) … HVW … HVW2) + [ @ldrop_atom #H destruct | /3 width=7 by cpys_subst_Y2, ex2_intro/ ] | #a #I #G #L #V1 #V2 #T1 #T2 #_ #_ * #V #HV1 #HV2 * #T #HT1 #HT2 - elim (cpys_split_up … HT1 1) -HT1 // #T0 #HT10 #HT0 + elim (cpys_split_down … HT1 1) -HT1 // #T0 #HT10 #HT0 +(* @(ex2_intro … (ⓑ{a,I}V.T0)) - [ @(cpys_bind … HV1) -HV1 + [ @cpys_bind // | @(cpxs_bind … HV2) -HV2 - ] + + /2 width=5 by cpys_tpxs_trans/ + ] +*)*) \ No newline at end of file