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=3b05c3ea67fc63945a79c43fadb64a421065c9f6;hb=9aa9a54946719d3fdb4cadb7c7d33fd13956c083;hp=13a5ca15009bf6a84957a48329062c43dd9cee30;hpb=f75be90562ddd964ef7ed43b956eb908f3133e3a;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 13a5ca150..3b05c3ea6 100644 --- a/matita/matita/contribs/lambda_delta/Basic_2/reducibility/cpr.ma +++ b/matita/matita/contribs/lambda_delta/Basic_2/reducibility/cpr.ma @@ -30,13 +30,13 @@ interpretation (* Basic_1: was by definition: pr2_free *) lemma cpr_pr: ∀T1,T2. T1 ⇒ T2 → ∀L. L ⊢ T1 ⇒ T2. -/2/ qed. +/2 width=3/ qed. 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. +/2 width=1/ qed. (* Note: new property *) (* Basic_1: was only: pr2_thin_dx *) @@ -47,7 +47,7 @@ qed. lemma cpr_cast: ∀L,V,T1,T2. L ⊢ T1 ⇒ T2 → L ⊢ 𝕔{Cast} V. T1 ⇒ T2. -#L #V #T1 #T2 * /3/ +#L #V #T1 #T2 * /3 width=3/ qed. (* Note: it does not hold replacing |L1| with |L2| *) @@ -55,7 +55,7 @@ qed. 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 ***************************************************) @@ -64,14 +64,14 @@ qed. lemma cpr_inv_atom: ∀T1,T2. ⋆ ⊢ T1 ⇒ T2 → T1 ⇒ T2. #T1 #T2 * #T #HT normalize #HT2 <(tpss_inv_refl_O2 … HT2) -HT2 // -qed. +qed-. (* Basic_1: was: pr2_gen_sort *) 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. +qed-. (* Basic_1: was: pr2_gen_cast *) lemma cpr_inv_cast1: ∀L,V1,T1,U2. L ⊢ 𝕔{Cast} V1. T1 ⇒ U2 → ( @@ -79,10 +79,10 @@ lemma cpr_inv_cast1: ∀L,V1,T1,U2. L ⊢ 𝕔{Cast} V1. T1 ⇒ U2 → ( U2 = 𝕔{Cast} 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/ -qed. +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