X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2A%2Freduction%2Fcpr.ma;h=ea42ca983c605a4485c329b2379f7413a2615f05;hb=2f6f2b7c01d47d23f61dd48d767bcb37aecdcfea;hp=d79fba9fde9e61b078ebe2a245935d9786a50dda;hpb=d2545ffd201b1aa49887313791386add78fa8603;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2A/reduction/cpr.ma b/matita/matita/contribs/lambdadelta/basic_2A/reduction/cpr.ma index d79fba9fd..ea42ca983 100644 --- a/matita/matita/contribs/lambdadelta/basic_2A/reduction/cpr.ma +++ b/matita/matita/contribs/lambdadelta/basic_2A/reduction/cpr.ma @@ -12,6 +12,12 @@ (* *) (**************************************************************************) +include "ground/xoa/ex_4_1.ma". +include "ground/xoa/ex_5_6.ma". +include "ground/xoa/ex_6_6.ma". +include "ground/xoa/ex_6_7.ma". +include "ground/xoa/ex_7_7.ma". +include "ground/xoa/or_4.ma". include "basic_2A/notation/relations/pred_4.ma". include "basic_2A/static/lsubr.ma". include "basic_2A/unfold/lstas.ma". @@ -19,7 +25,6 @@ include "basic_2A/unfold/lstas.ma". (* CONTEXT-SENSITIVE PARALLEL REDUCTION FOR TERMS ***************************) (* activate genv *) -(* Basic_1: includes: pr0_delta1 pr2_delta1 pr2_thin_dx *) (* Note: cpr_flat: does not hold in basic_1 *) inductive cpr: relation4 genv lenv term term ≝ | cpr_atom : ∀I,G,L. cpr G L (⓪{I}) (⓪{I}) @@ -60,18 +65,15 @@ lemma lsubr_cpr_trans: ∀G. lsub_trans … (cpr G) lsubr. ] qed-. -(* Basic_1: was by definition: pr2_free *) lemma tpr_cpr: ∀G,T1,T2. ⦃G, ⋆⦄ ⊢ T1 ➡ T2 → ∀L. ⦃G, L⦄ ⊢ T1 ➡ T2. #G #T1 #T2 #HT12 #L lapply (lsubr_cpr_trans … HT12 L ?) // qed. -(* Basic_1: includes by definition: pr0_refl *) lemma cpr_refl: ∀G,T,L. ⦃G, L⦄ ⊢ T ➡ T. #G #T elim T -T // * /2 width=1 by cpr_bind, cpr_flat/ qed. -(* Basic_1: was: pr2_head_1 *) lemma cpr_pair_sn: ∀I,G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡ V2 → ∀T. ⦃G, L⦄ ⊢ ②{I}V1.T ➡ ②{I}V2.T. * /2 width=1 by cpr_bind, cpr_flat/ qed. @@ -97,8 +99,7 @@ fact lstas_cpr_aux: ∀h,G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 •*[h, d] T2 → d = 0 → ⦃G, L⦄ ⊢ T1 ➡ T2. #h #G #L #T1 #T2 #d #H elim H -G -L -T1 -T2 -d /3 width=1 by cpr_eps, cpr_flat, cpr_bind/ -[ #G #L #d #k #H0 destruct normalize // -| #G #L #K #V1 #V2 #W2 #i #d #HLK #_ #HVW2 #IHV12 #H destruct +[ #G #L #K #V1 #V2 #W2 #i #d #HLK #_ #HVW2 #IHV12 #H destruct /3 width=6 by cpr_delta/ | #G #L #K #V1 #V2 #W2 #i #d #_ #_ #_ #_