X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Freducibility%2Fcpr.ma;h=caa36dfd58b65b0d292bb910c0c6eefa2a159611;hb=78d4844bcccb3deb58a3179151c3045298782b18;hp=220f8a648e19f02f1ee60555c497ccf57e137d15;hpb=a8c166f1e1baeeae04553058bd179420ada8bbe7;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/basic_2/reducibility/cpr.ma b/matita/matita/contribs/lambda_delta/basic_2/reducibility/cpr.ma index 220f8a648..caa36dfd5 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/reducibility/cpr.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/reducibility/cpr.ma @@ -20,7 +20,7 @@ include "basic_2/reducibility/tpr.ma". (* Basic_1: includes: pr2_delta1 *) definition cpr: lenv → relation term ≝ - λL,T1,T2. ∃∃T. T1 ➡ T & L ⊢ T [0, |L|] ▶* T2. + λL,T1,T2. ∃∃T. T1 ➡ T & L ⊢ T ▶* [0, |L|] T2. interpretation "context-sensitive parallel reduction (term)" @@ -28,14 +28,14 @@ interpretation (* Basic properties *********************************************************) -lemma cpr_intro: ∀L,T1,T,T2,d,e. T1 ➡ T → L ⊢ T [d, e] ▶* T2 → L ⊢ T1 ➡ T2. +lemma cpr_intro: ∀L,T1,T,T2,d,e. T1 ➡ T → L ⊢ T ▶* [d, e] T2 → L ⊢ T1 ➡ T2. /4 width=3/ qed-. (* Basic_1: was by definition: pr2_free *) -lemma cpr_pr: ∀T1,T2. T1 ➡ T2 → ∀L. L ⊢ T1 ➡ T2. +lemma cpr_tpr: ∀T1,T2. T1 ➡ T2 → ∀L. L ⊢ T1 ➡ T2. /2 width=3/ qed. -lemma cpr_tpss: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ▶* T2 → L ⊢ T1 ➡ T2. +lemma cpr_tpss: ∀L,T1,T2,d,e. L ⊢ T1 ▶* [d, e] T2 → L ⊢ T1 ➡ T2. /3 width=5/ qed. lemma cpr_refl: ∀L,T. L ⊢ T ➡ T. @@ -56,7 +56,7 @@ qed. (* Note: it does not hold replacing |L1| with |L2| *) (* Basic_1: was only: pr2_change *) lemma cpr_lsubs_conf: ∀L1,T1,T2. L1 ⊢ T1 ➡ T2 → - ∀L2. L1 [0, |L1|] ≼ L2 → L2 ⊢ T1 ➡ T2. + ∀L2. L1 ≼ [0, |L1|] L2 → L2 ⊢ T1 ➡ T2. #L1 #T1 #T2 * #T #HT1 #HT2 #L2 #HL12 lapply (tpss_lsubs_conf … HT2 … HL12) -HT2 -HL12 /3 width=4/ qed. @@ -78,7 +78,7 @@ qed-. (* Basic_1: was pr2_gen_abbr *) lemma cpr_inv_abbr1: ∀L,V1,T1,U2. L ⊢ ⓓV1. T1 ➡ U2 → - (∃∃V,V2,T2. V1 ➡ V & L ⊢ V [O, |L|] ▶* V2 & + (∃∃V,V2,T2. V1 ➡ V & L ⊢ V ▶* [O, |L|] V2 & L. ⓓV ⊢ T1 ➡ T2 & U2 = ⓓV2. T2 ) ∨ @@ -107,8 +107,9 @@ elim (tpr_inv_cast1 … H) -H /3 width=3/ elim (tpss_inv_flat1 … HU2) -HU2 #V2 #T2 #HV2 #HT2 #H destruct /4 width=5/ qed-. -(* Basic_1: removed theorems 4: +(* Basic_1: removed theorems 6: pr2_head_2 pr2_cflat pr2_gen_cflat clear_pr2_trans + pr2_gen_ctail pr2_ctail Basic_1: removed local theorems 3: pr2_free_free pr2_free_delta pr2_delta_delta *)