X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Freduction%2Fcrx.ma;h=5816b0028dae7fba4d361c3e555feb06acfd3fa1;hb=fca909e9e53de73771e1b47e94434ae8f747d7fb;hp=86796ae0bab0ba94d2d7d08003a25bba6e4c5778;hpb=606dab57f31b66eb3f30f603185124b88dfad4c1;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/crx.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/crx.ma index 86796ae0b..5816b0028 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/crx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/crx.ma @@ -12,11 +12,11 @@ (* *) (**************************************************************************) -include "basic_2/notation/relations/reducible_5.ma". +include "basic_2/notation/relations/predreducible_5.ma". include "basic_2/static/sd.ma". include "basic_2/reduction/crr.ma". -(* CONTEXT-SENSITIVE EXTENDED REDUCIBLE TERMS *******************************) +(* REDUCIBLE TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION *****************) (* activate genv *) (* extended reducible terms *) @@ -33,8 +33,8 @@ inductive crx (h) (g) (G:genv): relation2 lenv term ≝ . interpretation - "context-sensitive extended reducibility (term)" - 'Reducible h g G L T = (crx h g G L T). + "reducibility for context-sensitive extended reduction (term)" + 'PRedReducible h g G L T = (crx h g G L T). (* Basic properties *********************************************************) @@ -103,7 +103,7 @@ lemma trx_inv_atom: ∀h,g,I,G. ⦃G, ⋆⦄ ⊢ ➡[h, g] 𝐑⦃⓪{I}⦄ → #h #g * #i #G #H [ elim (crx_inv_sort … H) -H /2 width=4 by ex2_2_intro/ | elim (crx_inv_lref … H) -H #I #L #V #H - elim (ldrop_inv_atom1 … H) -H #H destruct + elim (drop_inv_atom1 … H) -H #H destruct | elim (crx_inv_gref … H) ] qed-.