X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fdynamic%2Flsubsv_lsuba.ma;fp=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fdynamic%2Flsubsv_lsuba.ma;h=7928373cf8bd48e96c2d82ac6fb7870152d56bbf;hb=603c8b3cdab901c26f63b5fed2c65e49693cc9a3;hp=0000000000000000000000000000000000000000;hpb=b01109cef1d0cf392539cffa35751d9d716296ac;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_lsuba.ma b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_lsuba.ma new file mode 100644 index 000000000..7928373cf --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/dynamic/lsubsv_lsuba.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/dynamic/snv_aaa.ma". +include "basic_2/dynamic/lsubsv.ma". + +(* LOCAL ENVIRONMENT REFINEMENT FOR STRATIFIED NATIVE VALIDITY **************) + +(* Properties on local environment refinement for atomic arity assignment ***) + +lemma lsubsv_fwd_lsuba: ∀h,g,L1,L2. h ⊢ L1 ⊩:⊑[g] L2 → L1 ⁝⊑ L2. +#h #g #L1 #L2 #H elim H -L1 -L2 // /2 width=1/ +#L1 #L2 #V1 #V2 #W1 #W2 #l #HV1 #HVW1 #HW12 #HW2 #_ #_ #HL12 +elim (snv_fwd_aaa … HV1) -HV1 #A #HV1 +elim (snv_fwd_aaa … HW2) -HW2 #B #HW2 +lapply (ssta_aaa … HV1 … HVW1) -HVW1 #H1 +lapply (lsuba_aaa_trans … HW2 … HL12) #H2 +lapply (aaa_cpcs_mono … HW12 … H1 … H2) -W1 -H2 #H destruct /2 width=3/ +qed-.