(**************************************************************************) (* ___ *) (* ||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 //