X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground%2Frelocation%2Fgr_coafter_nat_tls.ma;h=0402e2929db2fca38b107edd0af2f52e77037bcd;hb=8bbe582d87984526f40182c4409cbfd43108cb79;hp=f297468f9bc80fb0c2be569210e3071e70d2d9a7;hpb=55c768d7e45babb300b5010463ba3196a68f1bbe;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/ground/relocation/gr_coafter_nat_tls.ma b/matita/matita/contribs/lambdadelta/ground/relocation/gr_coafter_nat_tls.ma index f297468f9..0402e2929 100644 --- a/matita/matita/contribs/lambdadelta/ground/relocation/gr_coafter_nat_tls.ma +++ b/matita/matita/contribs/lambdadelta/ground/relocation/gr_coafter_nat_tls.ma @@ -16,14 +16,14 @@ include "ground/relocation/gr_tls.ma". include "ground/relocation/gr_nat.ma". include "ground/relocation/gr_coafter.ma". -(* RELATIONAL CO-COMPOSITION FOR GENERIC RELOCATION MAPS ***********************************************************) +(* RELATIONAL CO-COMPOSITION FOR GENERIC RELOCATION MAPS ********************) -(* Properties with nat and iterated tail ********************************************) +(* Constructions with gr_nat and gr_tls *************************************) (*** coafter_tls *) lemma gr_coafter_tls_bi_tls (n2) (n1): ∀f1,f2,f. @↑❪n1, f1❫ ≘ n2 → - f1 ~⊚ f2 ≘ f → ⫱*[n2]f1 ~⊚ ⫱*[n1]f2 ≘ ⫱*[n2]f. + f1 ~⊚ f2 ≘ f → ⫰*[n2]f1 ~⊚ ⫰*[n1]f2 ≘ ⫰*[n2]f. #n2 @(nat_ind_succ … n2) -n2 [ #n1 | #n2 #IH * [| #n1 ] ] #f1 #f2 #f #Hf1 #Hf [ elim (gr_nat_inv_zero_dx … Hf1) -Hf1 [ |*: // ] #g1 #Hg1 #H1 destruct // | elim (gr_nat_inv_zero_succ … Hf1) -Hf1 [ |*: // ] #g1 #Hg1 #H1 @@ -41,5 +41,5 @@ qed. (*** coafter_tls_O *) lemma gr_coafter_tls_sn_tls: ∀n,f1,f2,f. @↑❪𝟎, f1❫ ≘ n → - f1 ~⊚ f2 ≘ f → ⫱*[n]f1 ~⊚ f2 ≘ ⫱*[n]f. + f1 ~⊚ f2 ≘ f → ⫰*[n]f1 ~⊚ f2 ≘ ⫰*[n]f. /2 width=1 by gr_coafter_tls_bi_tls/ qed.