X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2A%2Freduction%2Fcnr_crr.ma;fp=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2A%2Freduction%2Fcnr_crr.ma;h=e8b5dfb4b8e0dda9c29dbcf02aa0781a35271090;hb=d2545ffd201b1aa49887313791386add78fa8603;hp=0000000000000000000000000000000000000000;hpb=57ae1762497a5f3ea75740e2908e04adb8642cc2;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2A/reduction/cnr_crr.ma b/matita/matita/contribs/lambdadelta/basic_2A/reduction/cnr_crr.ma new file mode 100644 index 000000000..e8b5dfb4b --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2A/reduction/cnr_crr.ma @@ -0,0 +1,46 @@ +(**************************************************************************) +(* ___ *) +(* ||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/crr.ma". +include "basic_2A/reduction/cnr.ma". + +(* NORMAL TERMS FOR CONTEXT-SENSITIVE REDUCTION *****************************) + +(* Advanced inversion lemmas on reducibility ********************************) + +(* Note: this property is unusual *) +lemma cnr_inv_crr: ∀G,L,T. ⦃G, L⦄ ⊢ ➡ 𝐑⦃T⦄ → ⦃G, L⦄ ⊢ ➡ 𝐍⦃T⦄ → ⊥. +#G #L #T #H elim H -L -T +[ #L #K #V #i #HLK #H + elim (cnr_inv_delta … HLK H) +| #L #V #T #_ #IHV #H + elim (cnr_inv_appl … H) -H /2 width=1 by/ +| #L #V #T #_ #IHT #H + elim (cnr_inv_appl … H) -H /2 width=1 by/ +| #I #L #V #T * #H1 #H2 destruct + [ elim (cnr_inv_zeta … H2) + | elim (cnr_inv_eps … H2) + ] +|5,6: #a * [ elim a ] #L #V #T * #H1 #_ #IH #H2 destruct + [1,3: elim (cnr_inv_abbr … H2) -H2 /2 width=1 by/ + |*: elim (cnr_inv_abst … H2) -H2 /2 width=1 by/ + ] +| #a #L #V #W #T #H + elim (cnr_inv_appl … H) -H #_ #_ #H + elim (simple_inv_bind … H) +| #a #L #V #W #T #H + elim (cnr_inv_appl … H) -H #_ #_ #H + elim (simple_inv_bind … H) +] +qed-.