X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fcomputation%2Flsubc_lsuba.ma;h=3a17f0e5b42c7eb1db6810352e88bc9f7e2dcf18;hb=fca909e9e53de73771e1b47e94434ae8f747d7fb;hp=b4f5be28952efa0d91e31ef5a83f45485e05bcd9;hpb=82500a9ceb53e1af0263c22afbd5954fa3a83190;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 b4f5be289..3a17f0e5b 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/lsubc_lsuba.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/lsubc_lsuba.ma @@ -13,14 +13,14 @@ (**************************************************************************) include "basic_2/static/lsuba.ma". -include "basic_2/computation/acp_aaa.ma". +include "basic_2/computation/gcp_aaa.ma". -(* LOCAL ENVIRONMENT REFINEMENT FOR ABSTRACT CANDIDATES OF REDUCIBILITY *****) +(* LOCAL ENVIRONMENT REFINEMENT FOR GENERIC REDUCIBILITY ********************) (* properties concerning lenv refinement for atomic arity assignment ********) -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/ +lemma lsuba_lsubc: ∀RR,RS,RP. gcp RR RS RP → gcr RR RS RP RP → + ∀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 by lsubc_pair/ +#L1 #L2 #W #V #A #H elim (aaa_inv_cast … H) -H /3 width=4 by acr_aaa, lsubc_beta/ qed.