X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fetc_2A1%2Fcpr%2Flsubr_lbotr.etc;fp=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fetc_2A1%2Fcpr%2Flsubr_lbotr.etc;h=b95a19b66b603affe4af28b206fbe78211c336ff;hb=09b4420070d6a71990e16211e499b51dbb0742cb;hp=0000000000000000000000000000000000000000;hpb=bba53a83579540bc3925d47d679e2aad22e85755;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_2A1/cpr/lsubr_lbotr.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_2A1/cpr/lsubr_lbotr.etc new file mode 100644 index 000000000..b95a19b66 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_2A1/cpr/lsubr_lbotr.etc @@ -0,0 +1,77 @@ +(**************************************************************************) +(* ___ *) +(* ||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 *) +(* *) +(**************************************************************************) + +notation "hvbox( ⊒ [ term 46 d , break term 46 e ] break term 46 L2 )" + non associative with precedence 45 + for @{ 'SubEqBottom $d $e $L2 }. + +include "basic_2/relocation/lsubr.ma". + +(* LOCAL ENVIRONMENT REFINEMENT FOR SUBSTITUTION ****************************) + +(* bottom element of the refinement *) +definition lbotr: nat → nat → predicate lenv ≝ + λd,e. NF_sn … (lsubr d e) (lsubr d e …). + +interpretation + "local environment full refinement (substitution)" + 'SubEqBottom d e L = (lbotr d e L). + +(* Basic properties *********************************************************) + +lemma lbotr_atom: ∀d,e. ⊒[d, e] ⋆. +#d #e #L #H +elim (lsubr_inv_atom2 … H) -H +[ #H destruct // +| * #H1 #H2 destruct // +] +qed. + +lemma lbotr_OO: ∀L. ⊒[0, 0] L. +// qed. + +lemma lbotr_abbr: ∀L,V,e. ⊒[0, e] L → ⊒[0, e + 1] L.ⓓV. +#L #V #e #HL #K #H +elim (lsubr_inv_abbr2 … H ?) -H //