X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Fcomputation%2Fltprs_alt.ma;fp=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Fcomputation%2Fltprs_alt.ma;h=7d532c973c69d4ff815b21260b77946c1c738284;hb=f79d97a42a84f94d37ad9589fcce46149ee23d12;hp=0000000000000000000000000000000000000000;hpb=99c8b28b92ec2c44774f664f9c9ec1a458593e1d;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/ltprs_alt.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/ltprs_alt.ma new file mode 100644 index 000000000..7d532c973 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/ltprs_alt.ma @@ -0,0 +1,34 @@ +(**************************************************************************) +(* ___ *) +(* ||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/computation/ltprs.ma". + +(* CONTEXT-FREE PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS ******************) + +(* alternative definition of ltprs *) +definition ltprsa: relation lenv ≝ lpx tprs. + +interpretation + "context-free parallel computation (environment) alternative" + 'PRedStarAlt L1 L2 = (ltprsa L1 L2). + +(* Basic properties *********************************************************) + +lemma ltprs_ltprsa: ∀L1,L2. L1 ➡* L2 → L1 ➡➡* L2. +/2 width=1/ qed. + +(* Basic inversion lemmas ***************************************************) + +lemma ltprsa_ltprs: ∀L1,L2. L1 ➡➡* L2 → L1 ➡* L2. +/2 width=1/ qed-.