X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2A%2Freduction%2Fcnx_crx.ma;fp=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2A%2Freduction%2Fcnx_crx.ma;h=2be5d0d63fd81172c8167f344af393954920757b;hb=d2545ffd201b1aa49887313791386add78fa8603;hp=0000000000000000000000000000000000000000;hpb=57ae1762497a5f3ea75740e2908e04adb8642cc2;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2A/reduction/cnx_crx.ma b/matita/matita/contribs/lambdadelta/basic_2A/reduction/cnx_crx.ma new file mode 100644 index 000000000..2be5d0d63 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2A/reduction/cnx_crx.ma @@ -0,0 +1,49 @@ +(**************************************************************************) +(* ___ *) +(* ||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_2A/reduction/crx.ma". +include "basic_2A/reduction/cnx.ma". + +(* NORMAL TERMS FOR CONTEXT-SENSITIVE EXTENDED REDUCTION ********************) + +(* Advanced inversion lemmas on reducibility ********************************) + +(* Note: this property is unusual *) +lemma cnx_inv_crx: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ➡[h, g] 𝐑⦃T⦄ → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃T⦄ → ⊥. +#h #g #G #L #T #H elim H -L -T +[ #L #k #d #Hkd #H + lapply (cnx_inv_sort … H) -H #H + lapply (deg_mono … H Hkd) -h -L -k