From cfc43911db215e21036317b26bd1dcf9c3e5d435 Mon Sep 17 00:00:00 2001 From: Ferruccio Guidi Date: Wed, 15 Jan 2014 16:37:48 +0000 Subject: [PATCH] we start a criterion on extended computation --- .../basic_2/computation/cpxs_cpys.ma | 35 +++++++++++++++++++ .../lambdadelta/basic_2/web/basic_2_src.tbl | 2 +- 2 files changed, 36 insertions(+), 1 deletion(-) create mode 100644 matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_cpys.ma diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_cpys.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_cpys.ma new file mode 100644 index 000000000..0bbe9835c --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_cpys.ma @@ -0,0 +1,35 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/substitution/cpys_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 → + ∃∃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/ +| #a #I #G #L #V1 #V2 #T1 #T2 #_ #_ * #V #HV1 #HV2 * #T #HT1 #HT2 + elim (cpys_split_up … HT1 1) -HT1 // #T0 #HT10 #HT0 + @(ex2_intro … (ⓑ{a,I}V.T0)) + [ @(cpys_bind … HV1) -HV1 + | @(cpxs_bind … HV2) -HV2 + ] diff --git a/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl b/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl index 920657dd2..ceacb2e53 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl +++ b/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl @@ -106,7 +106,7 @@ table { ] [ { "context-sensitive extended computation" * } { [ "lpxs ( ⦃?,?⦄ ⊢ ➡*[?,?] ? )" "lpxs_alt ( ⦃?,?⦄ ⊢ ➡➡*[?,?] ? )" "lpxs_ldrop" + "lpxs_lleq" + "lpxs_aaa" + "lpxs_cpxs" + "lpxs_lpxs" * ] - [ "cpxs ( ⦃?,?⦄ ⊢ ? ➡*[?,?] ? )" "cpxs_tstc" + "cpxs_tstc_vector" + "cpxs_lift" + "cpxs_lleq" + "cpxs_aaa" + "cpxs_cpxs" * ] + [ "cpxs ( ⦃?,?⦄ ⊢ ? ➡*[?,?] ? )" "cpxs_tstc" + "cpxs_tstc_vector" + "cpxs_lift" + "cpxs_cpys" + "cpxs_lleq" + "cpxs_aaa" + "cpxs_cpxs" * ] } ] [ { "context-sensitive computation" * } { -- 2.39.2