X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Fequivalence%2Ffpcs_aaa.ma;fp=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Fequivalence%2Ffpcs_aaa.ma;h=9f4327bff22edf311fbb9dfe71e4b5d92bc97315;hb=4bea40e6589ce21c15ecf99bdd5bd2a1c62f6809;hp=0000000000000000000000000000000000000000;hpb=943bfd67310c34c05cbba627272864eb10800143;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/basic_2/equivalence/fpcs_aaa.ma b/matita/matita/contribs/lambda_delta/basic_2/equivalence/fpcs_aaa.ma new file mode 100644 index 000000000..9f4327bff --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/equivalence/fpcs_aaa.ma @@ -0,0 +1,30 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/computation/fprs_aaa.ma". +include "basic_2/equivalence/fpcs_fpcs.ma". + +(* CONTEXT-FREE PARALLEL EQUIVALENCE ON CLOSURES ****************************) + +(* Main properties about atomic arity assignment on terms *******************) + +theorem aaa_fpcs_mono: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⬌* ⦃L2, T2⦄ → + ∀A1. L1 ⊢ T1 ⁝ A1 → ∀A2. L2 ⊢ T2 ⁝ A2 → + A1 = A2. +#L1 #L2 #T1 #T2 #H12 #A1 #HT1 #A2 #HT2 +elim (fpcs_inv_fprs … H12) -H12 #L #T #H1 #H2 +lapply (aaa_fprs_conf … HT1 … H1) -L1 -T1 #HT1 +lapply (aaa_fprs_conf … HT2 … H2) -L2 -T2 #HT2 +lapply (aaa_mono … HT1 … HT2) -L -T // +qed-.