X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fcomputation%2Flsubc_lsuba.ma;h=b4f5be28952efa0d91e31ef5a83f45485e05bcd9;hb=ab0d181f9a89f461a9c280f42a949a2dc2abe44c;hp=b77900e2f9e95c20d292c336b84e967f651a1e0c;hpb=f62eeb3c7824564ccbe4fff6e75ddee46ca39cc0;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/lsubc_lsuba.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/lsubc_lsuba.ma index b77900e2f..b4f5be289 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/lsubc_lsuba.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/lsubc_lsuba.ma @@ -18,10 +18,9 @@ include "basic_2/computation/acp_aaa.ma". (* LOCAL ENVIRONMENT REFINEMENT FOR ABSTRACT CANDIDATES OF REDUCIBILITY *****) (* properties concerning lenv refinement for atomic arity assignment ********) -(* -lamma lsubc_lsuba: ∀RR,RS,RP. acp RR RS RP → acr RR RS RP (λL,T. RP L T) → - ∀L1,L2. L1 ⁝⊑ L2 → L1 ⊑[RP] L2. -#RR #RS #RP #H1RP #H2RP #L1 #L2 #H elim H -L1 -L2 -// /2 width=1/ /3 width=4/ + +lemma lsuba_lsubc: ∀RR,RS,RP. acp RR RS RP → acr RR RS RP (λG,L,T. RP G L T) → + ∀G,L1,L2. G ⊢ L1 ⁝⊑ L2 → G ⊢ L1 ⊑[RP] L2. +#RR #RS #RP #H1RP #H2RP #G #L1 #L2 #H elim H -L1 -L2 // /2 width=1/ +#L1 #L2 #W #V #A #H elim (aaa_inv_cast … H) -H /3 width=4/ qed. -*) \ No newline at end of file