X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Freducibility%2Flfpr_alt.ma;h=adff7ad07cc353a2a3c78854b8ed77dee52988d1;hb=ba575c0609015580c1419c17b350de19a158e8e3;hp=bfb483781044587c3af1e005256f3049cc514d83;hpb=cc21d0caa6229b7d1a905f9b62de2af4f40cc863;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/basic_2/reducibility/lfpr_alt.ma b/matita/matita/contribs/lambda_delta/basic_2/reducibility/lfpr_alt.ma index bfb483781..adff7ad07 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/reducibility/lfpr_alt.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/reducibility/lfpr_alt.ma @@ -13,7 +13,8 @@ (**************************************************************************) include "basic_2/grammar/lenv_px_bi.ma". -include "basic_2/reducibility/fpr.ma". +include "basic_2/reducibility/fpr_cpr.ma". +include "basic_2/reducibility/lfpr_fpr.ma". (* FOCALIZED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS **********************) @@ -29,6 +30,15 @@ interpretation lemma lfpra_refl: reflexive … lfpra. /2 width=1/ qed. +lemma fpr_lfpra: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ➡ ⦃L2, T2⦄ → ⦃L1⦄ ➡➡ ⦃L2⦄. +#L1 elim L1 -L1 +[ #L2 #T1 #T2 #H + elim (fpr_inv_atom1 … H) -H #_ #H destruct // +| #L1 #I #V1 #IH #L2 #T1 #T2 #H + elim (fpr_inv_pair1 … H) -H #L #V #HV1 #HL1 #H destruct /3 width=3/ +] +qed. + (* Basic inversion lemmas ***************************************************) lemma lfpra_inv_atom1: ∀L2. ⦃⋆⦄ ➡➡ ⦃L2⦄ → L2 = ⋆. @@ -47,48 +57,23 @@ lemma lfpra_inv_pair2: ∀L1,K2,I,V2. ⦃L1⦄ ➡➡ ⦃K2. ⓑ{I} V2⦄ → L1 = K1. ⓑ{I} V1. /2 width=1 by lpx_bi_inv_pair2/ qed-. +lemma lfpra_inv_fpr: ∀L1,L2. ⦃L1⦄ ➡➡ ⦃L2⦄ → ∀T.⦃L1, T⦄ ➡ ⦃L2, T⦄. +#L1 #L2 * -L1 -L2 // /3 width=1/ +qed-. + (* Basic forward lemmas *****************************************************) lemma lfpra_fwd_length: ∀L1,L2. ⦃L1⦄ ➡➡ ⦃L2⦄ → |L1| = |L2|. /2 width=2 by lpx_bi_fwd_length/ qed-. -(****************************************************************************) -(* -lemma fpr_lfpra: ∀L1,T1,L2,T2. ⦃L1, T1⦄ ➡ ⦃L2, T2⦄ → ⦃L1⦄ ➡➡ ⦃L2⦄. -#L1 #T1 @(cw_wf_ind … L1 T1) -L1 -T1 * -[ #T1 #_ #L2 #T2 #H - elim (fpr_inv_atom1 … H) -H #_ #H destruct // -| #L1 #I #V1 #T1 #IH #Y #X #H - elim (fpr_inv_pair1 … H) -H #L2 #V2 #HL12 #H destruct - @lpx_bi_pair - [ @(IH … HL12) - | @IH - - /3 width=4/ +(* Main properties **********************************************************) -lemma fpr_lfpra: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ➡ ⦃L2, T2⦄ → ⦃L1⦄ ➡➡ ⦃L2⦄. -#L1 elim L1 -L1 -[ #L2 #T1 #T2 #H - elim (fpr_inv_atom1 … H) -H #_ #H destruct // -| #L1 #I #V1 #IH #L2 #T1 #T2 #H - elim (fpr_inv_pair1 … H) -H #L #V #HL1 #H destruct - @lpx_bi_pair /2 width=3/ - - /4 width=3/ - -(* -include "basic_2/reducibility/lcpr.ma". - -lemma lcpr_pair2: ∀L1,L2. L1 ⊢ ➡ L2 → ∀V1,V2. ⦃L1, V1⦄ ➡ ⦃L2, V2⦄ → - ∀I. L1. ⓑ{I} V1 ⊢ ➡ L2. ⓑ{I} V2. -#L1 #L2 * #L #HL1 #HL2 #V1 #V2 * -#H1 #H2 #I -@(ex2_1_intro … (L.ⓑ{I}V1)) /2 width=1/ -@tpss_ - -(* -<(ltpss_fwd_length … HL2) /4 width=5/ +theorem lfpr_lfpra: ∀L1,L2. ⦃L1⦄ ➡ ⦃L2⦄ → ⦃L1⦄ ➡➡ ⦃L2⦄. +#L1 #L2 #H +lapply (lfpr_inv_fpr … H (⋆0)) -H /2 width=3/ qed. -*) -*) -*) \ No newline at end of file + +theorem lfpra_lfpr: ∀L1,L2. ⦃L1⦄ ➡➡ ⦃L2⦄ → ⦃L1⦄ ➡ ⦃L2⦄. +#L1 #L2 #H +lapply (lfpra_inv_fpr … H (⋆0)) -H /2 width=3/ +qed-.