X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground%2Frelocation%2Frtmap_basic.ma;h=f6671fa46e9d8ae78bd4f4e35e83550832ddc302;hp=60ea38e24bcdcf01636d393ee7f8b09dbfe55875;hb=8fdf1af656038d0245eba64ff2531bbe94ce0e9e;hpb=77c9255de3c5f7780aeacd745703a1cc76328a68 diff --git a/matita/matita/contribs/lambdadelta/ground/relocation/rtmap_basic.ma b/matita/matita/contribs/lambdadelta/ground/relocation/rtmap_basic.ma index 60ea38e24..f6671fa46 100644 --- a/matita/matita/contribs/lambdadelta/ground/relocation/rtmap_basic.ma +++ b/matita/matita/contribs/lambdadelta/ground/relocation/rtmap_basic.ma @@ -13,7 +13,7 @@ (**************************************************************************) include "ground/notation/functions/basic_2.ma". -include "ground/relocation/rtmap_at.ma". +include "ground/relocation/rtmap_uni.ma". (* RELOCATION MAP ***********************************************************) @@ -22,25 +22,12 @@ definition basic: nat → nat → rtmap ≝ λm,n. ⫯*[m] 𝐔❨n❩. interpretation "basic relocation (rtmap)" 'Basic m n = (basic m n). -(* Prioerties with application **********************************************) +(* Basic properties *********************************************************) -lemma at_basic_lt: ∀m,n,i. i < m → @❪i, 𝐁❨m,n❩❫ ≘ i. -#m elim m -m [ #n #i #H elim (lt_zero_false … H) ] -#m #IH #n * [ /2 width=2 by refl, at_refl/ ] -#i #H lapply (lt_S_S_to_lt … H) -H /3 width=7 by refl, at_push/ +lemma at_basic_succ_sn (m) (n): ⫯𝐁❨m,n❩ = 𝐁❨↑m,n❩. +#m #n >pushs_S // qed. -lemma at_basic_ge: ∀m,n,i. m ≤ i → @❪i, 𝐁❨m,n❩❫ ≘ n+i. -#m elim m -m // -#m #IH #n #j #H -elim (le_inv_S1 … H) -H #i #Hmi #H destruct -/3 width=7 by refl, at_push/ +lemma at_basic_zero_succ (n): ↑𝐁❨𝟎,n❩ = 𝐁❨𝟎,↑n❩. +#n >nexts_S // qed. - -(* Inversion lemmas with application ****************************************) - -lemma at_basic_inv_lt: ∀m,n,i,j. i < m → @❪i, 𝐁❨m,n❩❫ ≘ j → i = j. -/3 width=4 by at_basic_lt, at_mono/ qed-. - -lemma at_basic_inv_ge: ∀m,n,i,j. m ≤ i → @❪i, 𝐁❨m,n❩❫ ≘ j → n+i = j. -/3 width=4 by at_basic_ge, at_mono/ qed-.