X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Flfpxs_length.ma;h=70c9539c993b4d9ef1443e3fc7b2049b5df99bb5;hb=47a745462a714af9d65cea7b61af56524bd98fa1;hp=ac2f64dc599e92e5ad1856ff8bb8522016d1e230;hpb=10b733131aa2667d8ba4318d517f0ba3cf137359;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfpxs_length.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfpxs_length.ma index ac2f64dc5..70c9539c9 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfpxs_length.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfpxs_length.ma @@ -1,7 +1,23 @@ -(* Basic forward lemmas *****************************************************) +(**************************************************************************) +(* ___ *) +(* ||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/i_static/tc_lfxs_length.ma". +include "basic_2/rt_computation/lfpxs.ma". + +(* UNCOUNTED PARALLEL RT-COMPUTATION FOR LOCAL ENV.S ON REFERRED ENTRIES ****) (* Forward lemmas with length for local environments ************************) -(* Basic_2A1: was just: lpxs_fwd_length *) -lemma lfpxs_fwd_length: ∀h,o,G,L1,L2. ⦃G, L1⦄ ⊢ ⬈*[h, o] L2 → |L1| = |L2|. -/2 width=2 by TC_lpx_sn_fwd_length/ qed-. +lemma lfpxs_fwd_length: ∀h,G,L1,L2,T. ⦃G, L1⦄ ⊢ ⬈*[h, T] L2 → |L1| = |L2|. +/2 width=3 by tc_lfxs_fwd_length/ qed-.