X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground%2Frelocation%2Frtmap_after.ma;h=569ea3143372c1375106494f25bf596dceb80eeb;hp=ae6c2de4e502e0695ec001ed92fe5d33cf9311c4;hb=8fdf1af656038d0245eba64ff2531bbe94ce0e9e;hpb=77c9255de3c5f7780aeacd745703a1cc76328a68 diff --git a/matita/matita/contribs/lambdadelta/ground/relocation/rtmap_after.ma b/matita/matita/contribs/lambdadelta/ground/relocation/rtmap_after.ma index ae6c2de4e..569ea3143 100644 --- a/matita/matita/contribs/lambdadelta/ground/relocation/rtmap_after.ma +++ b/matita/matita/contribs/lambdadelta/ground/relocation/rtmap_after.ma @@ -13,6 +13,7 @@ (**************************************************************************) include "ground/notation/relations/rafter_3.ma". +include "ground/arith/nat_pred_succ.ma". include "ground/relocation/rtmap_istot.ma". (* RELOCATION MAP ***********************************************************) @@ -267,9 +268,10 @@ lemma after_mono_eq: ∀f1,f2,f. f1 ⊚ f2 ≘ f → ∀g1,g2,g. g1 ⊚ g2 ≘ g (* Properties on tls ********************************************************) -lemma after_tls: ∀n,f1,f2,f. @❪0, f1❫ ≘ n → +(* Note: this requires ↑ on first n *) +lemma after_tls: ∀n,f1,f2,f. @❪𝟏, f1❫ ≘ ↑n → f1 ⊚ f2 ≘ f → ⫱*[n]f1 ⊚ f2 ≘ ⫱*[n]f. -#n elim n -n // +#n @(nat_ind_succ … n) -n // #n #IH #f1 #f2 #f #Hf1 #Hf cases (at_inv_pxn … Hf1) -Hf1 [ |*: // ] #g1 #Hg1 #H1 cases (after_inv_nxx … Hf … H1) -Hf #g #Hg #H0 destruct @@ -315,31 +317,6 @@ qed-. lemma after_inv_isid3: ∀f1,f2,f. f1 ⊚ f2 ≘ f → 𝐈❪f❫ → 𝐈❪f1❫ ∧ 𝐈❪f2❫. /3 width=4 by after_fwd_isid2, after_fwd_isid1, conj/ qed-. -(* Properties on isuni ******************************************************) - -lemma after_isid_isuni: ∀f1,f2. 𝐈❪f2❫ → 𝐔❪f1❫ → f1 ⊚ ↑f2 ≘ ↑f1. -#f1 #f2 #Hf2 #H elim H -H -/5 width=7 by after_isid_dx, after_eq_repl_back2, after_next, after_push, eq_push_inv_isid/ -qed. - -lemma after_uni_next2: ∀f2. 𝐔❪f2❫ → ∀f1,f. ↑f2 ⊚ f1 ≘ f → f2 ⊚ ↑f1 ≘ f. -#f2 #H elim H -f2 -[ #f2 #Hf2 #f1 #f #Hf - elim (after_inv_nxx … Hf) -Hf [2,3: // ] #g #Hg #H0 destruct - /4 width=7 by after_isid_inv_sn, after_isid_sn, after_eq_repl_back0, eq_next/ -| #f2 #_ #g2 #H2 #IH #f1 #f #Hf - elim (after_inv_nxx … Hf) -Hf [2,3: // ] #g #Hg #H0 destruct - /3 width=5 by after_next/ -] -qed. - -(* Properties on uni ********************************************************) - -lemma after_uni: ∀n1,n2. 𝐔❨n1❩ ⊚ 𝐔❨n2❩ ≘ 𝐔❨n1+n2❩. -@nat_elim2 [3: #n #m