X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frt_computation%2Flfpxs_drops.ma;h=caeb2bf778a110238aa4c0c2145da9c526556f28;hb=05b047be6817f430c8c72fd9b0902df8bb9f579e;hp=7f4060200e08d2b434f44ff12d637097681d5d19;hpb=d4ab51c20dbccd1e88cd2c4dcdaf3b4e56301155;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfpxs_drops.ma b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfpxs_drops.ma index 7f4060200..caeb2bf77 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfpxs_drops.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/rt_computation/lfpxs_drops.ma @@ -15,20 +15,20 @@ include "basic_2/i_static/tc_lfxs_drops.ma". include "basic_2/rt_transition/lfpx_drops.ma". -(* UNCOUNTED PARALLEL RT-COMPUTATION FOR LOCAL ENV.S ON REFERRED ENTRIES ****) +(* UNBOUND PARALLEL RT-COMPUTATION FOR LOCAL ENV.S ON REFERRED ENTRIES ******) (* Properties with generic slicing for local environments *******************) (* Basic_2A1: uses: drop_lpxs_trans *) lemma drops_lfpxs_trans: ∀h,G. tc_dedropable_sn (cpx h G). -/3 width=5 by drops_lfpx_trans, dedropable_sn_LTC/ qed-. +/3 width=5 by drops_lfpx_trans, dedropable_sn_CTC/ qed-. (* Inversion lemmas with generic slicing for local environments *************) (* Basic_2A1: uses: lpxs_drop_conf *) lemma lfpxs_drops_conf: ∀h,G. tc_dropable_sn (cpx h G). -/3 width=5 by lfpx_drops_conf, dropable_sn_LTC/ qed-. +/3 width=5 by lfpx_drops_conf, dropable_sn_CTC/ qed-. (* Basic_2A1: uses: lpxs_drop_trans_O1 *) lemma lfpxs_drops_trans: ∀h,G. tc_dropable_dx (cpx h G). -/3 width=5 by lfpx_drops_trans, dropable_dx_LTC/ qed-. +/3 width=5 by lfpx_drops_trans, dropable_dx_CTC/ qed-.