X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2FBasic_2%2Freducibility%2Fcpr.ma;h=0ad09e4b4c22a1bd0dec07b9ed2645d2a143a073;hb=011cf6478141e69822a5b40933f2444d0522532f;hp=90bd96d4d9cce6b223e97a590a3d1ec680fa9e86;hpb=d38087520d6ce1d696b28da40f3811291fc8a311;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 90bd96d4d..0ad09e4b4 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,69 +28,87 @@ interpretation (* Basic properties *********************************************************) +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. -/2/ qed. +lemma cpr_pr: ∀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. -/2/ qed. +lemma cpr_refl: ∀L,T. L ⊢ T ➡ T. +/2 width=1/ qed. (* Note: new property *) (* Basic_1: was only: pr2_thin_dx *) lemma cpr_flat: ∀I,L,V1,V2,T1,T2. - L ⊢ V1 ⇒ V2 → L ⊢ T1 ⇒ T2 → L ⊢ 𝕗{I} V1. T1 ⇒ 𝕗{I} V2. T2. + L ⊢ V1 ➡ V2 → L ⊢ T1 ➡ T2 → L ⊢ ⓕ{I} V1. T1 ➡ ⓕ{I} V2. T2. #I #L #V1 #V2 #T1 #T2 * #V #HV1 #HV2 * /3 width=5/ qed. lemma cpr_cast: ∀L,V,T1,T2. - L ⊢ T1 ⇒ T2 → L ⊢ 𝕔{Cast} V. T1 ⇒ T2. -#L #V #T1 #T2 * /3/ + L ⊢ T1 ➡ T2 → L ⊢ ⓣV. T1 ➡ T2. +#L #V #T1 #T2 * /3 width=3/ 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. +lemma cpr_lsubs_conf: ∀L1,T1,T2. L1 ⊢ 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/ +lapply (tpss_lsubs_conf … HT2 … HL12) -HT2 -HL12 /3 width=4/ qed. (* Basic inversion lemmas ***************************************************) (* Basic_1: was: pr2_gen_csort *) -lemma cpr_inv_atom: ∀T1,T2. ⋆ ⊢ T1 ⇒ T2 → T1 ⇒ T2. +lemma cpr_inv_atom: ∀T1,T2. ⋆ ⊢ T1 ➡ T2 → T1 ➡ T2. #T1 #T2 * #T #HT normalize #HT2 <(tpss_inv_refl_O2 … HT2) -HT2 // qed-. (* Basic_1: was: pr2_gen_sort *) -lemma cpr_inv_sort1: ∀L,T2,k. L ⊢ ⋆k ⇒ T2 → T2 = ⋆k. +lemma cpr_inv_sort1: ∀L,T2,k. L ⊢ ⋆k ➡ T2 → T2 = ⋆k. #L #T2 #k * #X #H >(tpr_inv_atom1 … H) -H #H >(tpss_inv_sort1 … H) -H // 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 & + L. ⓓV ⊢ T1 ➡ T2 & + U2 = ⓓV2. T2 + ) ∨ + ∃∃T. ⇧[0,1] T ≡ T1 & L ⊢ T ➡ U2. +#L #V1 #T1 #Y * #X #H1 #H2 +elim (tpr_inv_abbr1 … H1) -H1 * +[ #V #T0 #T #HV1 #HT10 #HT0 #H destruct + elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV2 #HT2 #H destruct + lapply (tps_lsubs_conf … HT0 (L. ⓓV) ?) -HT0 /2 width=1/ #HT0 + lapply (tps_weak_all … HT0) -HT0 #HT0 + lapply (tpss_lsubs_conf … HT2 (L. ⓓV) ?) -HT2 /2 width=1/ #HT2 + lapply (tpss_weak_all … HT2) -HT2 #HT2 + lapply (tpss_strap … HT0 HT2) -T /4 width=7/ +| /4 width=5/ +] +qed-. + (* Basic_1: was: pr2_gen_cast *) -lemma cpr_inv_cast1: ∀L,V1,T1,U2. L ⊢ 𝕔{Cast} V1. T1 ⇒ U2 → ( - ∃∃V2,T2. L ⊢ V1 ⇒ V2 & L ⊢ T1 ⇒ T2 & - U2 = 𝕔{Cast} V2. T2 - ) ∨ L ⊢ T1 ⇒ U2. +lemma cpr_inv_cast1: ∀L,V1,T1,U2. L ⊢ ⓣV1. T1 ➡ U2 → ( + ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ T1 ➡ T2 & + U2 = ⓣV2. T2 + ) ∨ L ⊢ T1 ➡ U2. #L #V1 #T1 #U2 * #X #H #HU2 -elim (tpr_inv_cast1 … H) -H /3/ -* #V #T #HV1 #HT1 #H destruct -X; -elim (tpss_inv_flat1 … HU2) -HU2 #V2 #T2 #HV2 #HT2 #H destruct -U2 /4 width=5/ +elim (tpr_inv_cast1 … H) -H /3 width=3/ +* #V #T #HV1 #HT1 #H destruct +elim (tpss_inv_flat1 … HU2) -HU2 #V2 #T2 #HV2 #HT2 #H destruct /4 width=5/ qed-. -(* Basic_1: removed theorems 5: - pr2_head_1 pr2_head_2 pr2_cflat pr2_gen_cflat clear_pr2_trans +(* Basic_1: removed theorems 4: + pr2_head_2 pr2_cflat pr2_gen_cflat clear_pr2_trans Basic_1: removed local theorems 3: pr2_free_free pr2_free_delta pr2_delta_delta *) - -(* -pr2/fwd pr2_gen_appl -pr2/fwd pr2_gen_abbr -*)