X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda_delta%2Fbasic_2%2Fsubstitution%2Flsubs_sfr.ma;h=2b71621faf2ba690f8342347c570930034186d01;hb=ea83c19f4cac864dd87eb059d8aeb2343eba480f;hp=fb71a175c4b3d73d7c4af780e5f83b324f9039ad;hpb=2eef5f7f15de5fd3820075470c2937dba2012da6;p=helm.git diff --git a/matita/matita/contribs/lambda_delta/basic_2/substitution/lsubs_sfr.ma b/matita/matita/contribs/lambda_delta/basic_2/substitution/lsubs_sfr.ma index fb71a175c..2b71621fa 100644 --- a/matita/matita/contribs/lambda_delta/basic_2/substitution/lsubs_sfr.ma +++ b/matita/matita/contribs/lambda_delta/basic_2/substitution/lsubs_sfr.ma @@ -16,53 +16,53 @@ include "basic_2/substitution/lsubs.ma". (* LOCAL ENVIRONMENT REFINEMENT FOR SUBSTITUTION ****************************) -(* top element of the refinement *) +(* bottom element of the refinement *) definition sfr: nat → nat → predicate lenv ≝ - λd,e. NF … (lsubs d e) (lsubs d e …). + λd,e. NF_sn … (lsubs d e) (lsubs d e …). interpretation "local environment full refinement (substitution)" - 'SubEqTop d e L = (sfr d e L). + 'SubEqBottom d e L = (sfr d e L). (* Basic properties *********************************************************) -lemma sfr_atom: ∀d,e. ≼ [d, e] ⋆. +lemma sfr_atom: ∀d,e. ≽ [d, e] ⋆. #d #e #L #H -elim (lsubs_inv_atom1 … H) -H +elim (lsubs_inv_atom2 … H) -H [ #H destruct // | * #H1 #H2 destruct // ] qed. -lemma sfr_OO: ∀L. ≼ [0, 0] L. +lemma sfr_OO: ∀L. ≽ [0, 0] L. // qed. -lemma sfr_abbr: ∀L,V,e. ≼ [0, e] L → ≼ [0, e + 1] L.ⓓV. +lemma sfr_abbr: ∀L,V,e. ≽ [0, e] L → ≽ [0, e + 1] L.ⓓV. #L #V #e #HL #K #H -elim (lsubs_inv_abbr1 … H ?) -H //