X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Freducibility%2Fcpr.ma;h=072d951babf91b355a7005f522192e6f60afb2b2;hb=ea83c19f4cac864dd87eb059d8aeb2343eba480f;hp=caa36dfd58b65b0d292bb910c0c6eefa2a159611;hpb=78d4844bcccb3deb58a3179151c3045298782b18;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 caa36dfd5..072d951ba 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/reducibility/cpr.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/reducibility/cpr.ma @@ -49,16 +49,16 @@ lemma cpr_flat: ∀I,L,V1,V2,T1,T2. qed. lemma cpr_cast: ∀L,V,T1,T2. - L ⊢ T1 ➡ T2 → L ⊢ ⓣV. T1 ➡ T2. + 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_trans: ∀L1,T1,T2. L1 ⊢ T1 ➡ T2 → + ∀L2. L2 ≼ [0, |L1|] L1 → L2 ⊢ T1 ➡ T2. #L1 #T1 #T2 * #T #HT1 #HT2 #L2 #HL12 -lapply (tpss_lsubs_conf … HT2 … HL12) -HT2 -HL12 /3 width=4/ +lapply (tpss_lsubs_trans … HT2 … HL12) -HT2 -HL12 /3 width=4/ qed. (* Basic inversion lemmas ***************************************************) @@ -87,9 +87,9 @@ lemma cpr_inv_abbr1: ∀L,V1,T1,U2. L ⊢ ⓓV1. T1 ➡ U2 → 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_lsubs_trans … 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_lsubs_trans … 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/ @@ -97,9 +97,9 @@ elim (tpr_inv_abbr1 … H1) -H1 * qed-. (* Basic_1: was: pr2_gen_cast *) -lemma cpr_inv_cast1: ∀L,V1,T1,U2. L ⊢ ⓣV1. 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 + U2 = ⓝV2. T2 ) ∨ L ⊢ T1 ➡ U2. #L #V1 #T1 #U2 * #X #H #HU2 elim (tpr_inv_cast1 … H) -H /3 width=3/