X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Frelocation%2Fnstream_istot.ma;h=5ceef8a1be9738cc02d84ab3c2aa3073044ede81;hb=a77d0bd6a04e94f765d329d47b37d9e04d349b14;hp=c64c7011cd8caab1354a303d976ad67ff98c5e51;hpb=91ab6965be539b7ed0f3208e1c1fffd17aa151f7;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/ground_2/relocation/nstream_istot.ma b/matita/matita/contribs/lambdadelta/ground_2/relocation/nstream_istot.ma index c64c7011c..5ceef8a1b 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/relocation/nstream_istot.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/relocation/nstream_istot.ma @@ -18,11 +18,11 @@ include "ground_2/relocation/rtmap_istot.ma". (* RELOCATION N-STREAM ******************************************************) -let rec apply (i: nat) on i: rtmap → nat ≝ ?. +rec definition apply (i: nat) on i: rtmap → nat ≝ ?. * #n #f cases i -i [ @n | #i lapply (apply i f) -apply -i -f - #i @(⫯(n+i)) + #i @(↑(n+i)) ] defined. @@ -31,15 +31,15 @@ interpretation "functional application (nstream)" (* Specific properties on at ************************************************) -lemma at_O1: ∀i2,f. @⦃0, i2@f⦄ ≡ i2. +lemma at_O1: ∀i2,f. @⦃0, i2⨮f⦄ ≘ i2. #i2 elim i2 -i2 /2 width=5 by at_refl, at_next/ qed. -lemma at_S1: ∀n,f,i1,i2. @⦃i1, f⦄ ≡ i2 → @⦃⫯i1, n@f⦄ ≡ ⫯(n+i2). +lemma at_S1: ∀n,f,i1,i2. @⦃i1, f⦄ ≘ i2 → @⦃↑i1, n⨮f⦄ ≘ ↑(n+i2). #n elim n -n /3 width=7 by at_push, at_next/ qed. -lemma at_total: ∀i1,f. @⦃i1, f⦄ ≡ f@❴i1❵. +lemma at_total: ∀i1,f. @⦃i1, f⦄ ≘ f@❴i1❵. #i1 elim i1 -i1 [ * // | #i #IH * /3 width=1 by at_S1/ ] qed. @@ -47,32 +47,33 @@ qed. lemma at_istot: ∀f. 𝐓⦃f⦄. /2 width=2 by ex_intro/ qed. -lemma at_plus2: ∀f,i1,i,n,m. @⦃i1, n@f⦄ ≡ i → @⦃i1, (m+n)@f⦄ ≡ m+i. -#f #i1 #i #n #m #H elim m -m /2 width=5 by at_next/ +lemma at_plus2: ∀f,i1,i,n,m. @⦃i1, n⨮f⦄ ≘ i → @⦃i1, (m+n)⨮f⦄ ≘ m+i. +#f #i1 #i #n #m #H elim m -m // +#m