X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Fcomputation%2Flsubc_ldrops.ma;fp=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Fcomputation%2Flsubc_ldrops.ma;h=95879a234de132f5f2dbf56396c3bf3754ebd50a;hb=eb918fc784eacd2094e3986ba321ef47690d9983;hp=0000000000000000000000000000000000000000;hpb=011cf6478141e69822a5b40933f2444d0522532f;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/basic_2/computation/lsubc_ldrops.ma b/matita/matita/contribs/lambda_delta/basic_2/computation/lsubc_ldrops.ma new file mode 100644 index 000000000..95879a234 --- /dev/null +++ b/matita/matita/contribs/lambda_delta/basic_2/computation/lsubc_ldrops.ma @@ -0,0 +1,32 @@ +(**************************************************************************) +(* ___ *) +(* ||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/lsubc_ldrop.ma". + +(* LOCAL ENVIRONMENT REFINEMENT FOR ABSTRACT CANDIDATES OF REDUCIBILITY *****) + +(* Properties concerning generic local environment slicing ******************) + +(* Basic_1: was: csubc_drop1_conf_rev *) +lemma ldrops_lsubc_trans: ∀RR,RS,RP. + acp RR RS RP → acr RR RS RP (λL,T. RP L T) → + ∀L1,K1,des. ⇩*[des] L1 ≡ K1 → ∀K2. K1 [RP] ⊑ K2 → + ∃∃L2. L1 [RP] ⊑ L2 & ⇩*[des] L2 ≡ K2. +#RR #RS #RP #Hacp #Hacr #L1 #K1 #des #H elim H -L1 -K1 -des +[ /2 width=3/ +| #L1 #L #K1 #des #d #e #_ #HLK1 #IHL #K2 #HK12 + elim (ldrop_lsubc_trans … Hacp Hacr … HLK1 … HK12) -Hacp -Hacr -K1 #K #HLK #HK2 + elim (IHL … HLK) -IHL -HLK /3 width=5/ +] +qed-.