X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Freduction%2Fcpr_cir.ma;h=99846d67ac4bafc17ee4a1f103b340d0d2ccd0f1;hb=598a5c56535a8339f6533227ab580aff64e2d41c;hp=5cee2e207a15f574b01d0f72b76a613a01a6b599;hpb=8ed01fd6a38bea715ceb449bb7b72a46bad87851;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr_cir.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr_cir.ma index 5cee2e207..99846d67a 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr_cir.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr_cir.ma @@ -17,9 +17,9 @@ include "basic_2/reduction/cpr.ma". (* CONTEXT-SENSITIVE PARALLEL REDUCTION FOR TERMS ***************************) -(* Advanced forward lemmas on context-sensitive irreducible terms ***********) +(* Advanced forward lemmas on irreducibility ********************************) -lemma cpr_fwd_cir: ∀G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L⦄ ⊢ 𝐈⦃T1⦄ → T2 = T1. +lemma cpr_fwd_cir: ∀G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L⦄ ⊢ ➡ 𝐈⦃T1⦄ → T2 = T1. #G #L #T1 #T2 #H elim H -G -L -T1 -T2 [ // | #G #L #K #V1 #V2 #W2 #i #HLK #_ #HVW2 #IHV12 #H @@ -28,15 +28,15 @@ lemma cpr_fwd_cir: ∀G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡ T2 → ⦃G, L⦄ ⊢ [ elim (cir_inv_bind … H) -H #HV1 #HT1 * #H destruct lapply (IHV1 … HV1) -IHV1 -HV1 #H destruct lapply (IHT1 … HT1) -IHT1 #H destruct // - | elim (cir_inv_ib2 … H) -H /2 width=1/ /3 width=2/ + | elim (cir_inv_ib2 … H) -H /3 width=2 by or_introl, eq_f2/ ] | * #G #L #V1 #V2 #T1 #T2 #_ #_ #IHV1 #IHT1 #H [ elim (cir_inv_appl … H) -H #HV1 #HT1 #_ >IHV1 -IHV1 // -HV1 >IHT1 -IHT1 // - | elim (cir_inv_ri2 … H) /2 width=1/ + | elim (cir_inv_ri2 … H) /2 width=1 by/ ] | #G #L #V1 #T1 #T #T2 #_ #_ #_ #H - elim (cir_inv_ri2 … H) /2 width=1/ + elim (cir_inv_ri2 … H) /2 width=1 by or_introl/ | #G #L #V1 #T1 #T2 #_ #_ #H elim (cir_inv_ri2 … H) /2 width=1/ | #a #G #L #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #_ #_ #H