Copyright (c) 1999 Association of Mizar Users
environ vocabulary ASYMPT_0, BOOLE, NEWTON, NAT_1, INT_1, POWER, SERIES_1, SQUARE_1, ABSVALUE, SEQ_2, SEQ_1, FUNCT_1, ARYTM_1, FUNCT_2, ARYTM_3, RELAT_1, ORDINAL2, RLVECT_1, FUNCOP_1, FRAENKEL, QC_LANG1, PROB_1, GROUP_1, FINSEQ_1, FINSEQ_2, ZF_LANG, ASYMPT_1, FINSEQ_4, ARYTM; notation TARSKI, XBOOLE_0, SUBSET_1, ORDINAL1, NUMBERS, XCMPLX_0, XREAL_0, REAL_1, FUNCT_1, FUNCT_2, FRAENKEL, INT_1, NAT_1, SEQ_1, SEQ_2, RELAT_1, NEWTON, POWER, SERIES_1, BHSP_4, PRE_CIRC, SQUARE_1, LIMFUNC1, ABSVALUE, SEQM_3, DOMAIN_1, FINSEQ_1, FINSEQ_2, FINSEQ_4, RVSUM_1, ASYMPT_0; constructors SERIES_1, BHSP_4, DOMAIN_1, FINSEQ_4, ASYMPT_0, SETWISEO, SEQ_2, PRE_CIRC, REAL_1, SFMASTR3, PARTFUN1, LIMFUNC1, PREPOWER, NAT_1, RVSUM_1; clusters XREAL_0, INT_1, RELSET_1, FINSEQ_2, ASYMPT_0, SEQ_1, SUBSET_1, MEMBERED, ORDINAL2; requirements REAL, NUMERALS, SUBSET, BOOLE, ARITHM; definitions ASYMPT_0; theorems TARSKI, FUNCT_2, INT_1, NAT_1, AXIOMS, REAL_1, SEQ_1, SQUARE_1, SEQ_2, ABSVALUE, REAL_2, FUNCOP_1, POWER, NEWTON, GROUP_4, EULER_2, GR_CY_1, BHSP_4, SERIES_1, SEQM_3, FINSEQ_1, UNIFORM1, FINSEQ_3, FINSEQ_4, WSIERP_1, PRE_FF, GROUP_5, SCMFSA9A, RVSUM_1, FINSEQ_2, ASYMPT_0, FRAENKEL, RELSET_1, XREAL_0, XBOOLE_0, XCMPLX_0, XCMPLX_1; schemes FUNCT_2, SEQ_1, INT_2, NAT_1, RECDEF_1; begin :: Examples from the text definition cluster -> non negative Element of NAT; coherence by NAT_1:18; end; reserve c, c1, c2, d, d1, d2, e for Real, i, k, n, m, N, n1, n2, N0, N1, N2, N3, M for Nat, x for set; Lm1: for n st n >= 2 holds 2 to_power n > n+1 proof defpred _P[Nat] means 2 to_power $1 > $1+1; 2 to_power 2 = 2^2 by POWER:53 .= 2*2 by SQUARE_1:def 3 .= 4; then A1: _P[2]; A2: for k st k >= 2 & _P[k] holds _P[k+1] proof let k such that k >= 2 and A3: 2 to_power k > k+1; 2 to_power (k+1) = (2 to_power k)*(2 to_power 1) by POWER:32 .= (2 to_power k)*2 by POWER:30 .= (2 to_power k) + (2 to_power k) by XCMPLX_1:11; then A4: 2 to_power (k+1) > (k+1) + (2 to_power k) by A3,REAL_1:53; 2 to_power k > 0 by POWER:39; then 2 to_power k >= 0 + 1 by INT_1:20; then (k+1) + (2 to_power k) >= (k+1) + 1 by AXIOMS:24; hence 2 to_power (k+1) > (k+1) + 1 by A4,AXIOMS:22; end; for n st n >= 2 holds _P[n] from Ind1(A1, A2); hence thesis; end; theorem for t,t1 being Real_Sequence st (t.0 = 0 & (for n st n > 0 holds t.n = 12*(n to_power 3)*log(2,n) - 5*n^2 + (log(2,n))^2 +36)) & (t1.0 = 0 & (for n st n > 0 holds t1.n = (n to_power 3)*log(2,n))) ex s,s1 being eventually-positive Real_Sequence st s = t & s1 = t1 & s in Big_Oh(s1) proof let f,g be Real_Sequence such that A1: f.0 = 0 & (for n st n > 0 holds f.n = 12*(n to_power 3)*log(2,n) - 5*n^2 + (log(2,n))^2 + 36) and A2: g.0 = 0 & (for n st n > 0 holds g.n = (n to_power 3)*log(2,n)); f is eventually-positive proof take 3; let n; assume A3: n >= 3; then n > 1 by AXIOMS:22; then A4: n to_power 3 > n to_power 2 by POWER:44; A5: n to_power 2 > 0 by A3,POWER:39; A6: log(2,n) >= log(2,3) by A3,PRE_FF:12; log(2,3) > log(2,2) by POWER:65; then log(2,3) > 1 by POWER:60; then log(2,n) > 1 by A6,AXIOMS:22; then (n to_power 3)*log(2,n) > (n to_power 2)*1 by A4,A5,REAL_2:199; then 12*((n to_power 3)*log(2,n)) > 5*(n to_power 2) by A5,REAL_2:199; then 12*(n to_power 3)*log(2,n) > 5*(n to_power 2) by XCMPLX_1:4; then 12*(n to_power 3)*log(2,n) > 5*n^2 + 0 by POWER:53; then A7: 12*(n to_power 3)*log(2,n) - 5*n^2 > 0 by REAL_1:86; (log(2,n))^2 >= 0 by SQUARE_1:72; then (12*(n to_power 3)*log(2,n)-5*n^2)+(log(2,n))^2>0+0 by A7,REAL_1: 67; then (12*(n to_power 3)*log(2,n)-5*n^2)+(log(2,n))^2+36 > 0+0 by REAL_1 :67; hence f.n > 0 by A1,A3; end; then reconsider f as eventually-positive Real_Sequence; g is eventually-positive proof take 2; let n; assume A8: n >= 2; then log(2,n) >= log(2,2) by PRE_FF:12; then A9: log(2,n) >= 1 by POWER:60; n to_power 3 > 0 by A8,POWER:39; then (n to_power 3)*log(2,n) > (n to_power 3)*0 by A9,REAL_1:70; hence g.n > 0 by A2,A8; end; then reconsider g as eventually-positive Real_Sequence; take f, g; ex s being Real_Sequence st (s.0 = 0 & (for n st n > 0 holds s.n = 12*(n to_power 3)*log(2,n) - 5*n^2 )) proof defpred P[Nat,Real] means ($1 = 0 implies $2 = 0) & ($1 > 0 implies $2 = 12*($1 to_power 3)*log(2,$1) - 5*($1)^2); A10: for x being Element of NAT ex y being Element of REAL st P[x,y] proof let n; n = 0 or n > 0; hence thesis; end; consider h being Function of NAT,REAL such that A11: for x being Element of NAT holds P[x,h.x] from FuncExD(A10); take h; thus h.0 = 0 by A11; let n; thus thesis by A11; end; then consider p being Real_Sequence such that A12: p.0 = 0 & (for n st n > 0 holds p.n = 12*(n to_power 3)*log(2,n) - 5*n^2); p is eventually-positive proof take 3; let n; assume A13: n >= 3; then n > 1 by AXIOMS:22; then A14: n to_power 3 > n to_power 2 by POWER:44; A15: n to_power 2 > 0 by A13,POWER:39; A16: log(2,n) >= log(2,3) by A13,PRE_FF:12; log(2,3) > log(2,2) by POWER:65; then log(2,3) > 1 by POWER:60; then log(2,n) > 1 by A16,AXIOMS:22; then (n to_power 3)*log(2,n) > (n to_power 2)*1 by A14,A15,REAL_2:199; then 12*((n to_power 3)*log(2,n)) > 5*(n to_power 2) by A15,REAL_2:199 ; then 12*(n to_power 3)*log(2,n) > 5*(n to_power 2) by XCMPLX_1:4; then 12*(n to_power 3)*log(2,n) > 5*n^2 + 0 by POWER:53; then 12*(n to_power 3)*log(2,n) - 5*n^2 > 0 by REAL_1:86; hence p.n > 0 by A12,A13; end; then reconsider p as eventually-positive Real_Sequence; ex s being Real_Sequence st s.0 = 0 & (for n st n > 0 holds s.n = (log(2,n))^2 + 36) proof defpred P[Nat,Real] means ($1 = 0 implies $2 = 0) & ($1 > 0 implies $2 = (log(2,$1))^2 + 36); A17: for x being Element of NAT ex y being Element of REAL st P[x,y] proof let n; per cases; suppose n = 0; hence thesis; suppose n > 0; hence thesis; end; consider h being Function of NAT,REAL such that A18: for x being Element of NAT holds P[x,h.x] from FuncExD(A17); take h; thus thesis by A18; end; then consider q being Real_Sequence such that A19: q.0 = 0 & (for n st n > 0 holds q.n = (log(2,n))^2 + 36); q is eventually-positive proof take 1; let n; assume A20: n >= 1; (log(2,n))^2 >= 0 by SQUARE_1:72; then (log(2,n))^2 + 36 > 0+0 by REAL_1:67; hence q.n > 0 by A19,A20; end; then reconsider q as eventually-positive Real_Sequence; set t = max(p,q); for n holds f.n = p.n + q.n proof let n; thus f.n = p.n + q.n proof per cases; suppose n = 0; hence thesis by A1,A12,A19; suppose A21: n > 0; then p.n = 12*(n to_power 3)*log(2,n) - 5*n^2 by A12; then p.n + q.n = (12*(n to_power 3)*log(2,n)-5*n^2)+((log(2,n))^2 +36) by A19,A21 .= 12*(n to_power 3)*log(2,n)-5*n^2+(log(2,n))^2 +36 by XCMPLX_1:1; hence thesis by A1,A21; end; end; then A22: Big_Oh(f) = Big_Oh(p+q) by SEQ_1:11 .= Big_Oh(t) by ASYMPT_0:9; 4 = 2*2 .= 2^2 by SQUARE_1:def 3 .= 2 to_power 2 by POWER:53; then A23: log(2,4) = 2*log(2,2) by POWER:63 .= 2*1 by POWER:60 .= 2; A24:7 = 12 - 5; A25: for n st n >= 4 holds 7*n^2 > q.n proof defpred _P[Nat] means 7*$1^2 > q.$1; A26: _P[4] proof A27: 7*4^2 = 7*(4*4) by SQUARE_1:def 3 .= 112; q.4 = 2^2 + 36 by A19,A23 .= 2*2 + 36 by SQUARE_1:def 3 .= 40; hence thesis by A27; end; A28: for k st k >= 4 & _P[k] holds _P[k+1] proof let k such that A29: k >= 4 and A30: 7*k^2 > q.k; A31: q.k = (log(2,k))^2 + 36 by A19,A29; 7*(k+1)^2 = 7*(k^2 + 2*k + 1) by SQUARE_1:65 .= 7*(k^2 + (2*k + 1)) by XCMPLX_1:1 .= 7*k^2 + (7*(2*k + 1)) by XCMPLX_1:8 .= 7*k^2 + (7*(2*k) + 7*1) by XCMPLX_1:8; then A32: 7*(k+1)^2 > q.k + (7*(2*k) + 7*1) by A30,REAL_1:53; A33: q.(k+1) = (log(2,k+1))^2 + 36 by A19; k >= 1 by A29,AXIOMS:22; then k+k >= k+1 by AXIOMS:24; then A34: log(2,k+k) >= log(2,k+1) by PRE_FF:12; k+1 >= 4+0 by A29,REAL_1:67; then log(2,k+1) >= 2 by A23,PRE_FF:12; then (log(2,k+k))^2 >= (log(2,k+1))^2 by A34,SQUARE_1:77; then A35: q.(k+1) <= (log(2,k+k))^2 + 36 by A33,AXIOMS:24; log(2,k+k) = log(2,2*k) by XCMPLX_1:11; then log(2,k+k) = log(2,k) + log(2,2) by A29,POWER:61; then A36: (log(2,k+k))^2 = (log(2,k) + 1)^2 by POWER:60 .= (log(2,k))^2 + 2*log(2,k) + 1 by SQUARE_1:65; k >= 2 by A29,AXIOMS:22; then A37: 2 to_power k > k + 1 by Lm1; A38: log(2,k) >= 2 by A23,A29,PRE_FF:12; k+1 > k+0 by REAL_1:67; then 2 to_power k > k by A37,AXIOMS:22; then log(2,2 to_power k) > log(2,k) by A29,POWER:65; then k*log(2,2) > log(2,k) by POWER:63; then k*1 > log(2,k) by POWER:60; then 14*k > 2*log(2,k) by A38,REAL_2:199; then (7*2)*k + 7 > 2*log(2,k) + 1 by REAL_1:67; then 7*(2*k) + 7 > 2*log(2,k) + 1 by XCMPLX_1:4; then (log(2,k))^2 + (2*log(2,k) + 1) < (log(2,k))^2 + (7*(2*k) + 7) by REAL_1:53; then (log(2,k+k))^2 < (log(2,k))^2 + (7*(2*k) + 7) by A36,XCMPLX_1:1; then (log(2,k+k))^2 + 36 < ((log(2,k))^2 + (7*(2*k) + 7)) + 36 by REAL_1:53; then (log(2,k+k))^2 + 36 < ((log(2,k))^2 + 36) + (7*(2*k) + 7*1) by XCMPLX_1:1; then q.k + (7*(2*k) + 7*1) > q.(k+1) by A31,A35,AXIOMS:22; hence thesis by A32,AXIOMS:22; end; for n st n >= 4 holds _P[n] from Ind1(A26, A28); hence thesis; end; A39: for n st n >= 4 holds p.n > 7*n^2 proof let n; assume A40: n >= 4; then A41: p.n = 12*(n to_power 3)*log(2,n) - 5*n^2 by A12; n > 1 by A40,AXIOMS:22; then A42: n to_power 3 > n to_power 2 by POWER:44; A43: n to_power 2 > 0 by A40,POWER:39; log(2,n) >= log(2,4) by A40,PRE_FF:12; then log(2,n) > 1 by A23,AXIOMS:22; then (n to_power 3)*log(2,n) > (n to_power 2)*1 by A42,A43,REAL_2:199; then 12*((n to_power 3)*log(2,n)) > 12*(n to_power 2) by REAL_1:70; then 12*(n to_power 3)*log(2,n) > 12*(n to_power 2) by XCMPLX_1:4; then 12*(n to_power 3)*log(2,n) > 12*n^2 by POWER:53; then p.n > 12*n^2 - 5*n^2 by A41,REAL_1:54; hence p.n > 7*n^2 by A24,XCMPLX_1:40; end; A44: for n st n >= 4 holds p.n > q.n proof let n; assume A45: n >= 4; then A46: p.n > 7*n^2 by A39; 7*n^2 > q.n by A25,A45; hence thesis by A46,AXIOMS:22; end; A47: for n st n >= 4 holds t.n = p.n proof let n; assume n >= 4; then A48: p.n > q.n by A44; thus t.n = max( p.n, q.n ) by ASYMPT_0:def 10 .= p.n by A48,SQUARE_1:def 2; end; A49: Big_Oh(g) = { s where s is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds s.n <= c*g.n & s.n >= 0 } by ASYMPT_0:def 12; A50: t is Element of Funcs(NAT, REAL) by FUNCT_2:11; consider N such that A51: for n st n >= N holds t.n > 0 by ASYMPT_0:def 6; now let n; assume A52: n >= max(4,N); A53: max(4,N) >= 4 & max(4,N) >= N by SQUARE_1:46; then A54: n >= 4 & n >= N by A52,AXIOMS:22; then A55: t.n = p.n by A47 .= 12*(n to_power 3)*log(2,n) - 5*n^2 by A12,A52,A53; n to_power 2 > 0 by A52,A53,POWER:39; then n^2 > 0 by POWER:53; then 5*n^2 > 5*0 by REAL_1:70; then t.n <= 12*(n to_power 3)*log(2,n) - 0 by A55,REAL_1:92; then t.n <= 12*((n to_power 3)*log(2,n)) by XCMPLX_1:4; hence t.n <= 12*g.n by A2,A52,A53; thus t.n >= 0 by A51,A54; end; then A56: t in Big_Oh(g) by A49,A50; f in Big_Oh(f) by ASYMPT_0:10; hence thesis by A22,A56,ASYMPT_0:12; end; Lm2: for a being logbase Real, f being Real_Sequence st a > 1 & (f.0 = 0 & (for n st n > 0 holds f.n = log(a,n))) holds f is eventually-positive proof let a be logbase Real, f be Real_Sequence such that A1: a > 1 and A2: (f.0 = 0 & (for n st n > 0 holds f.n = log(a,n))); A3: a > 0 & a <> 1 by ASYMPT_0:def 3; set N = [/a\]; A4: [/a\] >= a by INT_1:def 5; A5: N > 0 by A3,INT_1:def 5; then reconsider N as Nat by INT_1:16; now let n; assume A6: n >= N+1; N+1 > N+0 by REAL_1:67; then n > N by A6,AXIOMS:22; then A7: log(a,n) > log(a,N) by A1,A5,POWER:65; log(a,N) >= log(a,a) by A1,A4,PRE_FF:12; then log(a,n) > log(a,a) by A7,AXIOMS:22; then log(a,n) > 0 by A3,POWER:60; hence f.n > 0 by A2,A6; end; hence f is eventually-positive by ASYMPT_0:def 6; end; theorem for a,b being logbase Real, f,g being Real_Sequence st a > 1 & b > 1 & (f.0 = 0 & (for n st n > 0 holds f.n = log(a,n))) & (g.0 = 0 & (for n st n > 0 holds g.n = log(b,n))) holds ex s,s1 being eventually-positive Real_Sequence st s = f & s1 = g & Big_Oh(s) = Big_Oh(s1) proof let a,b be logbase Real, f,g be Real_Sequence such that A1: a > 1 and A2: b > 1 and A3: (f.0 = 0 & (for n st n > 0 holds f.n = log(a,n))) and A4: (g.0 = 0 & (for n st n > 0 holds g.n = log(b,n))); reconsider f as eventually-positive Real_Sequence by A1,A3,Lm2; reconsider g as eventually-positive Real_Sequence by A2,A4,Lm2; take f,g; A5: a > 0 & a <> 1 by ASYMPT_0:def 3; A6: b > 0 & b <> 1 by ASYMPT_0:def 3; A7: Big_Oh(f) = { t where t is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds t.n <= c*f.n & t.n >= 0 } by ASYMPT_0:def 12; A8: Big_Oh(g) = { t where t is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds t.n <= c*g.n & t.n >= 0 } by ASYMPT_0:def 12; now let x be set; hereby assume x in Big_Oh(f); then consider t being Element of Funcs(NAT, REAL) such that A9: x = t and A10: ex c,N st c > 0 & for n st n >= N holds t.n <= c*f.n & t.n >= 0 by A7; consider c,N such that A11: c > 0 and A12: for n st n >= N holds t.n <= c*f.n & t.n >= 0 by A10; log(a,b) > log(a,1) by A1,A2,POWER:65; then log(a,b) > 0 by A5,POWER:59; then A13: c*log(a,b) > c*0 by A11,REAL_1:70; now take N1 = N+1; let n; assume A14: n >= N1; N+1 > N+0 by REAL_1:67; then A15: n > N by A14,AXIOMS:22; then A16: t.n <= c*f.n by A12; f.n = log(a,n) by A3,A14 .= log(a,b)*log(b,n) by A5,A6,A14,POWER:64; then t.n <= (c*log(a,b))*log(b,n) by A16,XCMPLX_1:4; hence t.n <= (c*log(a,b))*g.n by A4,A14; thus t.n >= 0 by A12,A15; end; hence x in Big_Oh(g) by A8,A9,A13; end; assume x in Big_Oh(g); then consider t being Element of Funcs(NAT, REAL) such that A17: x = t and A18: ex c,N st c > 0 & for n st n >= N holds t.n <= c*g.n & t.n >= 0 by A8; consider c,N such that A19: c > 0 and A20: for n st n >= N holds t.n <= c*g.n & t.n >= 0 by A18; log(b,a) > log(b,1) by A1,A2,POWER:65; then log(b,a) > 0 by A6,POWER:59; then A21: c*log(b,a) > c*0 by A19,REAL_1:70; now take N1 = N+1; let n; assume A22: n >= N1; N+1 > N+0 by REAL_1:67; then A23: n > N by A22,AXIOMS:22; then A24: t.n <= c*g.n by A20; g.n = log(b,n) by A4,A22 .= log(b,a)*log(a,n) by A5,A6,A22,POWER:64; then t.n <= (c*log(b,a))*log(a,n) by A24,XCMPLX_1:4; hence t.n <= (c*log(b,a))*f.n by A3,A22; thus t.n >= 0 by A20,A23; end; hence x in Big_Oh(f) by A7,A17,A21; end; hence thesis by TARSKI:2; end; definition let a,b,c be Real; func seq_a^(a,b,c) -> Real_Sequence means :Def1: it.n = a to_power (b*n+c); existence proof deffunc _F(Nat) = a to_power (b*$1+c); consider h being Function of NAT, REAL such that A1: for n being Nat holds h.n = _F(n) from LambdaD; take h; let n; thus h.n = a to_power (b*n+c) by A1; end; uniqueness proof let j,k be Real_Sequence such that A2: for n holds j.n = a to_power (b*n+c) and A3: for n holds k.n = a to_power (b*n+c); now let n; thus j.n = a to_power (b*n+c) by A2 .= k.n by A3; end; hence j = k by FUNCT_2:113; end; end; definition let a be positive Real, b,c be Real; cluster seq_a^(a,b,c) -> eventually-positive; coherence proof set f = seq_a^(a,b,c); take 0; let n; assume n >= 0; f.n = a to_power (b*n+c) by Def1; hence f.n > 0 by POWER:39; end; end; Lm3: for a,b,c being Real st a > 0 & c > 0 & c <> 1 holds a to_power b = c to_power (b*log(c,a)) proof let a,b,c be Real; assume that A1: a > 0 and A2: c > 0 and A3: c <> 1; A4: a to_power b > 0 by A1,POWER:39; log(c, a to_power b) = b*log(c,a) by A1,A2,A3,POWER:63; hence thesis by A2,A3,A4,POWER:def 3; end; theorem for a,b being positive Real st a < b holds not seq_a^(b,1,0) in Big_Oh(seq_a^(a,1,0)) proof let a,b be positive Real such that A1: a < b; set f = seq_a^(b,1,0); set g = seq_a^(a,1,0); A2: Big_Oh(g) = { s where s is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds s.n <= c*g.n & s.n >= 0 } by ASYMPT_0:def 12; hereby assume f in Big_Oh(g); then consider s being Element of Funcs(NAT, REAL) such that A3: s = f and A4: ex c,N st c > 0 & for n st n >= N holds s.n <= c*g.n & s.n >= 0 by A2; consider c,N such that A5: c > 0 and A6: for n st n >= N holds s.n <= c*g.n & s.n >= 0 by A4; set d = (log(2,b) - log(2,a)); log(2,b) > log(2,a) + 0 by A1,POWER:65; then A7: d > 0 by REAL_1:86; set N0 = [/log(2,c) / d\]; set N1 = max( N, N0 ); A8: N1 >= N & N1 >= N0 by SQUARE_1:46; N1 = N or N1 = N0 by SQUARE_1:49; then reconsider N1 as Nat by A8,INT_1:16; set n = N1 + 1; set e = (2 to_power (n*log(2,a))); A9: e > 0 by POWER:39; N1 + 1 > N1 + 0 by REAL_1:67; then A10: n > N & n > N0 by A8,AXIOMS:22; N0 >= log(2,c) / d by INT_1:def 5; then n > log(2,c) / d by A10,AXIOMS:22; then n*d > (log(2,c) / d)*d by A7,REAL_1:70; then n*d > log(2,c) by A7,XCMPLX_1:88; then 2 to_power (n*d) > 2 to_power log(2,c) by POWER:44; then 2 to_power (n*d) > c by A5,POWER:def 3; then 2 to_power (n*log(2,b) - n*log(2,a)) > c by XCMPLX_1:40; then (2 to_power (n*log(2,b))) / e > c by POWER:34; then ((2 to_power (n*log(2,b)))/e)*e > c*e by A9,REAL_1:70; then 2 to_power (n*log(2,b)) > c*e by A9,XCMPLX_1:88; then b to_power n > c*(2 to_power (n*log(2,a))) by Lm3; then A11: b to_power n > c*(a to_power n) by Lm3; f.n <= c*g.n by A3,A6,A10; then b to_power (1*n + 0) <= c*g.n by Def1; hence contradiction by A11,Def1; end; end; :: Example p. 84 definition func seq_logn -> Real_Sequence means :Def2: it.0 = 0 & (for n st n > 0 holds it.n = log(2,n)); existence proof defpred P[Nat,Real] means ($1 = 0 implies $2 = 0) & ($1 > 0 implies $2 = log(2,$1)); A1: for x being Element of NAT ex y being Element of REAL st P[x,y] proof let n; per cases; suppose n = 0; hence thesis; suppose n > 0; hence thesis; end; consider h being Function of NAT,REAL such that A2: for x being Element of NAT holds P[x,h.x] from FuncExD(A1); take h; thus h.0 = 0 by A2; let n; thus thesis by A2; end; uniqueness proof let j,k be Real_Sequence such that A3: j.0 = 0 & (for n st n > 0 holds j.n = log(2,n)) and A4: k.0 = 0 & (for n st n > 0 holds k.n = log(2,n)); now let n; per cases; suppose n = 0; hence j.n = k.n by A3,A4; suppose A5: n > 0; then j.n = log(2,n) by A3; hence j.n = k.n by A4,A5; end; hence j = k by FUNCT_2:113; end; end; definition let a be Real; func seq_n^(a) -> Real_Sequence means :Def3: it.0 = 0 & (for n st n > 0 holds it.n = n to_power a); existence proof defpred P[Nat,Real] means ($1 = 0 implies $2 = 0) & ($1 > 0 implies $2 = $1 to_power a); A1: for x being Element of NAT ex y being Element of REAL st P[x,y] proof let n; per cases; suppose n = 0; hence thesis; suppose n > 0; hence thesis; end; consider h being Function of NAT,REAL such that A2: for x being Element of NAT holds P[x,h.x] from FuncExD(A1); take h; thus h.0 = 0 by A2; let n; thus thesis by A2; end; uniqueness proof let j,k be Real_Sequence such that A3: j.0 = 0 & (for n st n > 0 holds j.n = n to_power a) and A4: k.0 = 0 & (for n st n > 0 holds k.n = n to_power a); now let n; per cases; suppose n = 0; hence j.n = k.n by A3,A4; suppose A5: n > 0; then j.n = n to_power a by A3; hence j.n = k.n by A4,A5; end; hence j = k by FUNCT_2:113; end; end; definition cluster seq_logn -> eventually-positive; coherence proof set f = seq_logn; take 2; let n; assume A1: n >= 2; then A2: f.n = log(2,n) by Def2; log(2,n) >= log(2,2) by A1,PRE_FF:12; hence f.n > 0 by A2,POWER:60; end; end; definition let a be Real; cluster seq_n^(a) -> eventually-positive; coherence proof set f = seq_n^(a); take 1; let n; assume A1: n >= 1; then f.n = n to_power a by Def3; hence f.n > 0 by A1,POWER:39; end; end; Lm4: for f,g being Real_Sequence, n being Nat holds (f/"g).n = f.n/g.n proof let f,g be Real_Sequence, n be Nat; thus (f/"g).n = (f(#)g").n by SEQ_1:def 9 .= f.n*g".n by SEQ_1:12 .= f.n*(g.n)" by SEQ_1:def 8 .= f.n/g.n by XCMPLX_0:def 9; end; Lm5: for f,g being eventually-nonnegative Real_Sequence holds f in Big_Oh(g) & g in Big_Oh(f) iff Big_Oh(f) = Big_Oh(g) proof let f,g be eventually-nonnegative Real_Sequence; hereby assume f in Big_Oh(g) & g in Big_Oh(f); then Big_Oh(f) c= Big_Oh(g) & Big_Oh(g) c= Big_Oh(f) by ASYMPT_0:11; hence Big_Oh(f) = Big_Oh(g) by XBOOLE_0:def 10; end; thus thesis by ASYMPT_0:10; end; theorem Th4: for f,g being eventually-nonnegative Real_Sequence holds Big_Oh(f) c= Big_Oh(g) & not Big_Oh(f) = Big_Oh(g) iff f in Big_Oh(g) & not f in Big_Omega(g) proof let f,g be eventually-nonnegative Real_Sequence; A1: Big_Oh(f) = { t where t is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds t.n <= c*f.n & t.n >= 0 } by ASYMPT_0:def 12; hereby assume A2: Big_Oh(f) c= Big_Oh(g) & not Big_Oh(f) = Big_Oh(g); A3: f in Big_Oh(f) by ASYMPT_0:10; now assume f in Big_Omega(g); then g in Big_Oh(f) by ASYMPT_0:19; hence contradiction by A2,A3,Lm5; end; hence f in Big_Oh(g) & not f in Big_Omega(g) by A2,A3; end; assume A4: f in Big_Oh(g) & not f in Big_Omega(g); now let x be set; assume x in Big_Oh(f); then consider t being Element of Funcs(NAT, REAL) such that A5: x = t and A6: ex c,N st c > 0 & for n st n >= N holds t.n <= c*f.n & t.n >= 0 by A1; consider c,N such that c > 0 and A7: for n st n >= N holds t.n <= c*f.n & t.n >= 0 by A6; now take N; let n; assume n >= N; hence t.n >= 0 by A7; end; then A8: t is eventually-nonnegative by ASYMPT_0:def 4; t in Big_Oh(f) by A1,A6; hence x in Big_Oh(g) by A4,A5,A8,ASYMPT_0:12; end; hence Big_Oh(f) c= Big_Oh(g) by TARSKI:def 3; assume Big_Oh(f) = Big_Oh(g); then f in Big_Oh(g) & g in Big_Oh(f) by Lm5; hence contradiction by A4,ASYMPT_0:19; end; Lm6: for a, b, c being Real st 0 < a & a <= b & c >= 0 holds a to_power c <= b to_power c proof let a, b, c be Real; assume that A1: 0 < a and A2: a <= b and A3: c >= 0; per cases by A3; suppose c = 0; then a to_power c = 1 & b to_power c = 1 by POWER:29; hence thesis; suppose A4: c > 0; thus thesis proof per cases by A2,REAL_1:def 5; suppose a = b; hence thesis; suppose a < b; hence thesis by A1,A4,POWER:42; end; end; Lm7: 2 to_power 2 = 4 proof thus 2 to_power 2 = 2^2 by POWER:53 .= 2*2 by SQUARE_1:def 3 .= 4; end; Lm8: 2 to_power 3 = 8 proof thus 2 to_power 3 = 2 to_power (2+1) .= (2 to_power 2)*(2 to_power 1) by POWER:32 .= 4*2 by Lm7,POWER:30 .= 8; end; Lm9: 2 to_power 4 = 16 proof thus 2 to_power 4 = 2 to_power (2+2) .= (2 to_power 2)*(2 to_power 2) by POWER:32 .= 16 by Lm7; thus thesis; end; Lm10: 2 to_power 5 = 32 proof thus 2 to_power 5 = 2 to_power (3+2) .= (2 to_power 3)*(2 to_power 2) by POWER:32 .= 32 by Lm7,Lm8; end; Lm11: 2 to_power 6 = 64 proof thus 2 to_power 6 = 2 to_power (3+3) .= 2 to_power 3 * 2 to_power 3 by POWER:32 .= 64 by Lm8; end; Lm12: for n st n >= 4 holds 2*n + 3 < 2 to_power n proof defpred _P[Nat] means 2*$1 + 3 < 2 to_power $1; A1: _P[4] by Lm9; A2: for k st k >= 4 & _P[k] holds _P[k+1] proof let k such that A3: k >= 4 and A4: 2*k + 3 < 2 to_power k; 2*(k+1) + 3 = (2*k + 2*1) + 3 by XCMPLX_1:8 .= 2 + (2*k + 3) by XCMPLX_1:1; then A5: 2*(k+1) + 3 < 2 + (2 to_power k) by A4,REAL_1:53; k > 1 by A3,AXIOMS:22; then 2 to_power k > 2 to_power 1 by POWER:44; then 2 to_power k > 2 by POWER:30; then (2 to_power k) + (2 to_power k) > 2 + (2 to_power k) by REAL_1:53; then 2*(k+1) + 3 < (2 to_power k) + (2 to_power k) by A5,AXIOMS:22; then 2*(k+1) + 3 < 2*(2 to_power k) by XCMPLX_1:11; then 2*(k+1) + 3 < (2 to_power 1)*(2 to_power k) by POWER:30; hence thesis by POWER:32; end; for n st n >= 4 holds _P[n] from Ind1(A1, A2); hence thesis; end; Lm13: for n st n >= 6 holds (n+1)^2 < 2 to_power n proof defpred _P[Nat] means ($1+1)^2 < 2 to_power $1; (6+1)^2 = 7*7 by SQUARE_1:def 3; then A1: _P[6] by Lm11; A2: for k st k >= 6 & _P[k] holds _P[k+1] proof let k such that A3: k >= 6 and A4: (k+1)^2 < 2 to_power k; A5: ((k+1)+1)^2 = (k+(1+1))^2 by XCMPLX_1:1 .= k^2 + 2*k*2 + 2^2 by SQUARE_1:63 .= k^2 + (2*2)*k + 2^2 by XCMPLX_1:4 .= k^2 + (2+2)*k + 4 by SQUARE_1:def 3 .= k^2 + (2*k + 2*k) + 4 by XCMPLX_1:8 .= ((k^2 + 2*k) + 2*k) + 4 by XCMPLX_1:1 .= (k^2 + 2*k) + (2*k + (3+1)) by XCMPLX_1:1 .= (k^2 + 2*k) + ((2*k + 3) + 1) by XCMPLX_1:1 .= ((k^2 + 2*k) + 1) + (2*k + 3) by XCMPLX_1:1 .= (k+1)^2 + (2*k + 3) by SQUARE_1:65; k >= 4 by A3,AXIOMS:22; then 2*k + 3 < 2 to_power k by Lm12; then A6: (k+1)^2 + (2*k + 3) < (k+1)^2 + (2 to_power k) by REAL_1:53; (k+1)^2 + (2 to_power k) < (2 to_power k) + (2 to_power k) by A4,REAL_1:53; then (k+1)^2 + (2*k + 3) < (2 to_power k) + (2 to_power k) by A6,AXIOMS:22; then ((k+1)+1)^2 < 2*(2 to_power k) by A5,XCMPLX_1:11; then ((k+1)+1)^2 < (2 to_power 1)*(2 to_power k) by POWER:30; hence thesis by POWER:32; end; for n st n >= 6 holds _P[n] from Ind1(A1, A2); hence thesis; end; Lm14: for c being Real st c > 6 holds c^2 < 2 to_power c proof let c be Real such that A1: c > 6; A2: 5 = 6-1; A3: c >= 0 by A1,AXIOMS:22; set i0 = [\c/], i1 = [/c\]; per cases; suppose i0 = i1; then c is Integer by INT_1:59; then reconsider c as Nat by A3,INT_1:16; A4: (c+1)^2 < 2 to_power c by A1,Lm13; c+0 < c+1 by REAL_1:67; then c^2 < (c+1)^2 by SQUARE_1:78; hence thesis by A4,AXIOMS:22; suppose not i0 = i1; then A5: i0 + 1 = i1 by INT_1:66; then i0 + (1 + -1) = i1 + -1 by XCMPLX_1:1; then A6: i0 = i1 - 1 by XCMPLX_0:def 8; A7: i1 >= c by INT_1:def 5; then reconsider i1 as Nat by A3,INT_1:16; i1 > 6 by A1,A7,AXIOMS:22; then A8: i0 > 5 by A2,A6,REAL_1:54; then i0 >= 0 by AXIOMS:22; then reconsider i0 as Nat by INT_1:16; i0 >= 5 + 1 by A8,INT_1:20; then A9: i1^2 < 2 to_power i0 by A5,Lm13; i1 >= c by INT_1:def 5; then i1^2 >= c^2 by A3,SQUARE_1:77; then A10: c^2 < 2 to_power i0 by A9,AXIOMS:22; i0 <= c by INT_1:def 4; then 2 to_power i0 <= 2 to_power c by PRE_FF:10; hence thesis by A10,AXIOMS:22; end; Lm15: for e being positive Real, f being Real_Sequence st (f.0 = 0 & (for n st n > 0 holds f.n = log(2,n to_power e))) holds (f /" seq_n^(e)) is convergent & lim (f /" seq_n^(e)) = 0 proof let e be positive Real, f be Real_Sequence such that A1: (f.0 = 0 & (for n st n > 0 holds f.n = log(2,n to_power e))); set g = seq_n^(e); set h = f/"g; A2: now let p be real number; assume A3: p > 0; reconsider p1 = p as Real by XREAL_0:def 1; set i0 = [/(7/p1) to_power (1/e)\]; set i1 = [/(p1 to_power -(2/e)) + 1\]; set N = max( max(i0, i1), 2 ); A4: N >= max(i0, i1) & N >= 2 by SQUARE_1:46; max(i0, i1) >= i0 & max(i0, i1) >= i1 by SQUARE_1:46; then A5: N >= i0 & N >= i1 by A4,AXIOMS:22; N is Integer proof per cases by SQUARE_1:49; suppose N = max(i0, i1); hence thesis by SQUARE_1:49; suppose N = 2; hence thesis; end; then reconsider N as Nat by A4,INT_1:16; p" > 0 by A3,REAL_1:72; then 7*p" > 7*0 by REAL_1:70; then A6: 7/p > 0 by XCMPLX_0:def 9; take N; let n; assume A7: n >= N; then A8: n >= 2 & n >= i0 & n >= i1 by A4,A5,AXIOMS:22; i0 >= (7/p) to_power (1/e) by INT_1:def 5; then A9: n >= (7/p) to_power (1/e) by A8,AXIOMS:22; (7/p1) to_power (1/e) > 0 by A6,POWER:39; then n to_power e >= (7/p) to_power (1/e) to_power e by A9,Lm6; then n to_power e >= (7/p1) to_power (e*(1/e)) by A6,POWER:38; then n to_power e >= (7/p) to_power 1 by XCMPLX_1:107; then n to_power e >= 7/p1 by POWER:30; then p*(n to_power e) >= (7/p)*p by A3,AXIOMS:25; then p*(n to_power e) >= 7 by A3,XCMPLX_1:88; then p*(n to_power e) > 6 by AXIOMS:22; then A10: (p1*(n to_power e))^2 < 2 to_power (p1*(n to_power e)) by Lm14; set c = p1*(n to_power e); A11: n to_power e > 0 by A4,A7,POWER:39; then A12: (n to_power e)" > 0 by REAL_1:72; A13:p*(n to_power e) > p*0 by A3,A11,REAL_1:70; c*c < 2 to_power c by A10,SQUARE_1:def 3; then c < (2 to_power c)/c by A13,REAL_2:178; then n to_power e < ((2 to_power c)/c)/p by A3,REAL_2:178; then A14: n to_power e < (2 to_power c)/(c*p) by XCMPLX_1:79; A15:p1 to_power 2 > 0 by A3,POWER:39; A16:i1 >= (p to_power -(2/e)) + 1 by INT_1:def 5; (p to_power -(2/e)) + 1 > (p to_power -(2/e)) + 0 by REAL_1:67; then i1 > (p to_power -(2/e)) by A16,AXIOMS:22; then A17:n > p to_power -(2/e) by A8,AXIOMS:22; p1 to_power -(2/e) > 0 by A3,POWER:39; then n to_power e > p to_power (-(2/e)) to_power e by A17,POWER:42; then n to_power e > p1 to_power ((-(2/e))*e) by A3,POWER:38; then n to_power e > p to_power -((2/e)*e) by XCMPLX_1:175; then n to_power e > p to_power -2 by XCMPLX_1:88; then (p to_power 2)*(n to_power e) > (p to_power 2)*(p to_power -2) by A15,REAL_1:70; then (p1 to_power 2)*(n to_power e) > p1 to_power (2+-2) by A3,POWER:32; then (p1 to_power 2)*(n to_power e) > 1 by POWER:29; then p1^2*(n to_power e) > 1 by POWER:53; then (p*p)*(n to_power e) > 1 by SQUARE_1:def 3; then p*c > 1 by XCMPLX_1:4; then A18: 1/(p*c) < 1/1 by REAL_2:151; 2 to_power c > 0 by POWER:39; then (2 to_power c)*(1/(c*p)) < (2 to_power c)*(1/1) by A18,REAL_1:70; then (2 to_power c)/(c*p) < (2 to_power c)*1 by XCMPLX_1:100; then n to_power e < 2 to_power c by A14,AXIOMS:22; then log(2,n to_power e) < log(2,2 to_power c) by A11,POWER:65; then log(2,n to_power e) < c*log(2,2) by POWER:63; then log(2,n to_power e) < c*1 by POWER:60; then A19: log(2,n to_power e)/(n to_power e) < p by A11,REAL_2:178; n > 1 by A8,AXIOMS:22; then n to_power e > n to_power 0 by POWER:44; then n to_power e > 1 by POWER:29; then log(2,n to_power e) > log(2,1) by POWER:65; then log(2,n to_power e) > 0 by POWER:59; then log(2,n to_power e)*(n to_power e)" > 0*(n to_power e)" by A12,REAL_1:70; then A20: log(2,n to_power e)/(n to_power e) > 0 by XCMPLX_0:def 9; h.n = f.n/g.n by Lm4 .= log(2,n to_power e)/g.n by A1,A4,A7 .= log(2,n to_power e)/(n to_power e) by A4,A7,Def3; hence abs(h.n-0) < p by A19,A20,ABSVALUE:def 1; end; hence h is convergent by SEQ_2:def 6; hence lim h = 0 by A2,SEQ_2:def 7; end; Lm16: for e being Real st e > 0 holds (seq_logn /" seq_n^(e)) is convergent & lim(seq_logn /" seq_n^(e)) = 0 proof let e be Real; assume e > 0; then reconsider e as positive Real; A1: 1 = e/e by XCMPLX_1:60 .= e*(1/e) by XCMPLX_1:100; set f = seq_logn; set g = seq_n^(e); set h = f/"g; ex s being Real_Sequence st (s.0 = 0 & (for n st n > 0 holds s.n = log(2,n to_power e))) proof defpred P[Nat,Real] means ($1 = 0 implies $2 = 0) & ($1 > 0 implies $2 = log(2,$1 to_power e)); A2: for x being Element of NAT ex y being Element of REAL st P[x,y] proof let n; per cases; suppose n = 0; hence thesis; suppose n > 0; hence thesis; end; consider h being Function of NAT,REAL such that A3: for x being Element of NAT holds P[x,h.x] from FuncExD(A2); take h; thus h.0 = 0 by A3; let n; thus thesis by A3; end; then consider p being Real_Sequence such that A4: p.0 = 0 & (for n st n > 0 holds p.n = log(2,n to_power e)); set q = p/"g; A5: q is convergent & lim q = 0 by A4,Lm15; A6:for n holds h.n = (1/e)*q.n proof let n; A7: h.n = f.n / g.n by Lm4; A8: q.n = p.n / g.n by Lm4; thus thesis proof per cases; suppose A9: n = 0; then h.n = 0 / g.n by A7,Def2 .= 0*(1/e); hence thesis by A4,A8,A9; suppose A10: n > 0; then A11: n to_power e > 0 by POWER:39; h.n = log(2,n) / g.n by A7,A10,Def2 .= log(2,n to_power (e*(1/e))) / g.n by A1,POWER:30 .= log(2,n to_power e to_power (1/e)) / g.n by A10,POWER:38 .= ((1/e)*log(2,n to_power e)) / g.n by A11,POWER:63 .= ((1/e)*log(2,n to_power e)) * (g.n)" by XCMPLX_0:def 9 .= (1/e)*(log(2,n to_power e) * (g.n)") by XCMPLX_1:4 .= (1/e)*(log(2,n to_power e) / g.n) by XCMPLX_0:def 9; hence thesis by A4,A8,A10; end; end; then A12: h = (1/e)(#)q by SEQ_1:13; lim h = lim((1/e)(#)q) by A6,SEQ_1:13 .= (1/e)*0 by A5,SEQ_2:22; hence thesis by A5,A12,SEQ_2:21; end; theorem Th5: Big_Oh(seq_logn) c= Big_Oh(seq_n^(1/2)) & not Big_Oh(seq_logn) = Big_Oh(seq_n^(1/2)) proof set f = seq_logn; set g = seq_n^(1/2); f/"g is convergent & lim (f/"g) = 0 by Lm16; then f in Big_Oh(g) & not g in Big_Oh(f) by ASYMPT_0:16; then f in Big_Oh(g) & not f in Big_Omega(g) by ASYMPT_0:19; hence thesis by Th4; end; :: Example p. 86 theorem seq_n^(1/2) in Big_Omega(seq_logn) & not seq_logn in Big_Omega(seq_n^(1/2)) proof seq_logn in Big_Oh(seq_n^(1/2)) & not seq_logn in Big_Omega(seq_n^(1/2)) by Th4,Th5; hence thesis by ASYMPT_0:19; end; :: Example p. 88 Lm17: for f being Real_Sequence holds for N holds (for n st n <= N holds f.n >= 0) implies Sum(f,N) >= 0 proof let f be Real_Sequence; defpred _P[Nat] means (for n st n <= $1 holds f.n >= 0) implies Sum(f,$1) >= 0; A1: _P[0] proof assume for n st n <= 0 holds f.n >= 0; then f.0 >= 0; then (Partial_Sums(f)).0 >= 0 by SERIES_1:def 1; hence thesis by BHSP_4:def 6; end; A2: for N st _P[N] holds _P[N+1] proof let N; assume A3: (for n st n <= N holds f.n >= 0) implies Sum(f,N) >= 0; assume A4: (for n st n <= (N+1) holds f.n >= 0); A5: now let n; assume n <= N; then n+0 <= N+1 by REAL_1:55; hence f.n >= 0 by A4; end; f.(N+1) >= 0 by A4; then Sum(f,N) + f.(N+1) >= 0 + 0 by A3,A5,REAL_1:55; then (Partial_Sums(f)).N + f.(N+1) >= 0 by BHSP_4:def 6; then (Partial_Sums(f)).(N+1) >= 0 by SERIES_1:def 1; hence Sum(f,N+1) >= 0 by BHSP_4:def 6; end; for N holds _P[N] from Ind(A1, A2); hence thesis; end; Lm18: for f,g being Real_Sequence holds for N holds (for n st n <= N holds f.n <= g.n) implies Sum(f,N) <= Sum (g,N) proof let f,g be Real_Sequence; defpred _P[Nat] means (for n st n <= $1 holds f.n <= g.n) implies Sum(f,$1) <= Sum(g,$1); A1: _P[0] proof assume (for n st n <= 0 holds f.n <= g.n); then f.0 <= g.0; then Partial_Sums(f).0 <= g.0 by SERIES_1:def 1; then Partial_Sums(f).0 <= Partial_Sums(g).0 by SERIES_1:def 1; then Sum(f,0) <= Partial_Sums(g).0 by BHSP_4:def 6; hence Sum(f,0) <= Sum(g,0) by BHSP_4:def 6; end; A2: for N st _P[N] holds _P[N+1] proof let N; assume A3: (for n st n <= N holds f.n <= g.n) implies Sum(f,N) <= Sum(g,N); assume A4: (for n st n <= (N+1) holds f.n <= g.n); A5: now let n; assume n <= N; then n+0 <= N+1 by REAL_1:55; hence f.n <= g.n by A4; end; f.(N+1) <= g.(N+1) by A4; then Sum(f,N) + f.(N+1) <= Sum(g,N) + g.(N+1) by A3,A5,REAL_1:55; then (Partial_Sums(f)).N + f.(N+1) <= Sum(g,N) + g.(N+1) by BHSP_4:def 6; then (Partial_Sums(f)).(N+1) <= Sum(g,N) + g.(N+1) by SERIES_1:def 1 ; then Sum(f,N+1) <= Sum(g,N) + g.(N+1) by BHSP_4:def 6; then Sum(f,N+1) <= (Partial_Sums(g)).N + g.(N+1) by BHSP_4:def 6; then Sum(f,N+1) <= (Partial_Sums(g)).(N+1) by SERIES_1:def 1; hence Sum(f,N+1) <= Sum(g,N+1) by BHSP_4:def 6; end; for N holds _P[N] from Ind(A1, A2); hence thesis; end; Lm19: for f being Real_Sequence, b being Real st (f.0 = 0 & (for n st n > 0 holds f.n = b)) holds for N being Nat holds Sum(f,N) = b*N proof let f be Real_Sequence, b be Real; defpred _P[Nat] means Sum(f,$1) = b*$1; assume A1: f.0 = 0 & (for n st n > 0 holds f.n = b); then (Partial_Sums(f)).0 = 0 by SERIES_1:def 1; then A2: _P[0] by BHSP_4:def 6; A3: for N st _P[N] holds _P[N+1] proof let N; assume A4: Sum(f, N) = b*N; Sum(f, N+1) = (Partial_Sums(f)).(N+1) by BHSP_4:def 6 .= (Partial_Sums(f)).N + f.(N+1) by SERIES_1:def 1 .= b*N + f.(N+1) by A4,BHSP_4:def 6 .= b*N + b*1 by A1 .= b*(N+1) by XCMPLX_1:8; hence thesis; end; for N holds _P[N] from Ind( A2, A3 ); hence thesis; end; Lm20: for f,g being Real_Sequence, N,M being Nat holds Sum(f,N,M) + f.(N+1) = Sum(f,N+1,M) proof let f,g be Real_Sequence, N,M be Nat; Sum(f,N,M) + f.(N+1) = Sum(f,N) - Sum(f,M) + f.(N+1) by BHSP_4:def 7 .= Sum(f,N) + -Sum(f,M) + f.(N+1) by XCMPLX_0:def 8 .= Sum(f,N) + f.(N+1) + -Sum(f,M) by XCMPLX_1:1 .= (Partial_Sums(f)).N + f.(N+1) + -Sum (f,M) by BHSP_4:def 6 .= (Partial_Sums(f)).(N+1) + -Sum(f,M) by SERIES_1:def 1 .= Sum(f,N+1) + -Sum(f,M) by BHSP_4:def 6 .= Sum(f,N+1) - Sum(f,M) by XCMPLX_0:def 8 .= Sum(f,N+1,M) by BHSP_4:def 7; hence Sum(f,N,M) + f.(N+1) = Sum(f,N+1,M); end; Lm21: for f,g being Real_Sequence, M being Nat holds for N st N >= M+1 holds (for n st M+1 <= n & n <= N holds f.n <= g.n) implies Sum(f,N,M) <= Sum (g,N,M) proof let f,g be Real_Sequence, M be Nat; defpred _P[Nat] means (for n st M+1 <= n & n <= $1 holds f.n <= g.n) implies Sum(f,$1,M) <= Sum(g,$1,M); A1: _P[M+1] proof assume A2: (for n st M+1 <= n & n <= M+1 holds f.n <= g.n); A3: Sum(f,M+1,M) = Sum(f,M+1) - Sum(f,M) by BHSP_4:def 7 .= (Partial_Sums(f)).(M+1) - Sum(f,M) by BHSP_4:def 6 .= f.(M+1) + (Partial_Sums(f)).M - Sum(f,M) by SERIES_1:def 1 .= f.(M+1) + Sum(f,M) - Sum(f,M) by BHSP_4:def 6 .= f.(M+1) + Sum(f,M) + -Sum(f,M) by XCMPLX_0:def 8 .= f.(M+1) + (Sum(f,M) + -Sum(f,M)) by XCMPLX_1:1 .= f.(M+1) + (Sum(f,M) - Sum(f,M)) by XCMPLX_0:def 8 .= f.(M+1) + 0 by XCMPLX_1:14; Sum(g,M+1,M) = Sum(g,M+1) - Sum(g,M) by BHSP_4:def 7 .= (Partial_Sums(g)).(M+1) - Sum(g,M) by BHSP_4:def 6 .= g.(M+1) + (Partial_Sums(g)).M - Sum(g,M) by SERIES_1:def 1 .= g.(M+1) + Sum(g,M) - Sum(g,M) by BHSP_4:def 6 .= g.(M+1) + Sum(g,M) + -Sum(g,M) by XCMPLX_0:def 8 .= g.(M+1) + (Sum(g,M) + -Sum(g,M)) by XCMPLX_1:1 .= g.(M+1) + (Sum(g,M) - Sum(g,M)) by XCMPLX_0:def 8 .= g.(M+1) + 0 by XCMPLX_1:14; hence Sum(f,M+1,M) <= Sum(g,M+1,M) by A2,A3; end; A4: for N1 st N1 >= M+1 & _P[N1] holds _P[N1+1] proof let N1; assume that A5: N1 >= M+1 and A6: (for n st M+1 <= n & n <= N1 holds f.n <= g.n) implies Sum(f,N1,M) <= Sum(g,N1,M); assume A7: (for n st M+1 <= n & n <= N1+1 holds f.n <= g.n); A8: now let n; assume A9: M+1 <= n & n <= N1; then n + 0 <= N1 + 1 by REAL_1:55; hence f.n <= g.n by A7,A9; end; N1+1 >= M+1+0 by A5,REAL_1:55; then f.(N1+1) <= g.(N1+1) by A7; then Sum(f,N1,M) + f.(N1+1) <= g.(N1+1) + Sum(g,N1,M) by A6,A8,REAL_1: 55; then Sum(f,N1+1,M) <= g.(N1+1) + Sum(g,N1,M) by Lm20; hence Sum(f,N1+1,M) <= Sum(g,N1+1,M) by Lm20; end; for N st N >= M+1 holds _P[N] from Ind1(A1, A4); hence thesis; end; Lm22: for n holds [/n/2\] <= n proof let n; per cases; suppose n = 0; hence thesis by INT_1:54; suppose n > 0; then A1: n >= 0+1 by NAT_1:38; thus thesis proof per cases by A1,REAL_1:def 5; suppose A2: n = 1; now assume [/1/2\] > 1; then A3: [/1/2\] >= 1+1 by INT_1:20; [/1/2\] < 1/2 + 1 by INT_1:def 5; hence contradiction by A3,AXIOMS:22; end; hence thesis by A2; suppose n > 1; then A4: n >= 1+1 by NAT_1:38; A5: [/n/2\] < n/2 + 1 by INT_1:def 5; now assume n/2 + 1 > n; then 2*(n/2 + 1) > 2*n by REAL_1:70; then 2*(n/2) + 2*1 > 2*n by XCMPLX_1:8; then 2*(n*2") + 2 > 2*n by XCMPLX_0:def 9; then (2*2")*n + 2 > 2*n by XCMPLX_1:4; then 2 > 2*n - n by REAL_1:84; then 2 > 2*n + -n by XCMPLX_0:def 8; then 2 > 2*n + (-1)*n by XCMPLX_1:180; then 2 > (2+-1)*n by XCMPLX_1:8; hence contradiction by A4; end; hence thesis by A5,AXIOMS:22; end; end; Lm23: for f being Real_Sequence, b being Real, N being Nat st (f.0 = 0 & (for n st n > 0 holds f.n = b)) holds for M being Nat holds Sum(f, N, M) = b*(N-M) proof let f be Real_Sequence, b be Real, N be Nat such that A1: f.0 = 0 & (for n st n > 0 holds f.n = b); defpred _P[Nat] means Sum(f, N, $1) = b*(N-$1); A2: _P[0] proof Sum(f, 0) = (Partial_Sums(f)).0 by BHSP_4:def 6 .= 0 by A1,SERIES_1:def 1; then Sum(f, N, 0) = Sum(f, N) - 0 by BHSP_4:def 7 .= b*(N-0) by A1,Lm19; hence thesis; end; A3: for M st _P[M] holds _P[M+1] proof let M; assume A4: Sum(f, N, M) = b*(N-M); Sum(f, N, M+1) = Sum(f, N) - Sum(f, M+1) by BHSP_4:def 7 .= Sum(f, N) - (Partial_Sums(f)).(M+1) by BHSP_4:def 6 .= Sum (f, N) - ((Partial_Sums(f)).M + f.(M+1)) by SERIES_1:def 1 .= Sum (f, N) + -((Partial_Sums(f)).M + f.(M+1)) by XCMPLX_0:def 8 .= Sum(f, N) + (-(Partial_Sums(f)).M + -f.(M+1)) by XCMPLX_1: 140 .= Sum(f, N) + -(Partial_Sums(f)).M + -f.(M+1) by XCMPLX_1:1 .= (Sum (f, N) - (Partial_Sums(f)).M) + -f.(M+1) by XCMPLX_0:def 8 .= (Sum(f, N) - Sum(f, M)) + -f.(M+1) by BHSP_4:def 6 .= b*(N-M) + -f.(M+1) by A4,BHSP_4:def 7 .= b*(N-M) + -b by A1 .= b*(N-M) + (-1)*b by XCMPLX_1:180 .= b*(N-M+-1) by XCMPLX_1:8 .= b*(N+-M+-1) by XCMPLX_0:def 8 .= b*(N+(-M+-1)) by XCMPLX_1:1 .= b*(N+-(M+1)) by XCMPLX_1:140 .= b*(N-(M+1)) by XCMPLX_0:def 8; hence thesis; end; for M being Nat holds _P[M] from Ind(A2, A3); hence thesis; end; theorem for f being Real_Sequence, k being Nat st (for n holds f.n = Sum(seq_n^(k), n)) holds f in Big_Theta(seq_n^(k+1)) proof let f be Real_Sequence, k be Nat such that A1: for n holds f.n = Sum(seq_n^(k), n); set g = seq_n^(k+1); A2: Big_Oh(g) = { t where t is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds t.n <= c*g.n & t.n >= 0 } by ASYMPT_0:def 12; A3: Big_Omega(g) = { t where t is Element of Funcs(NAT, REAL) : ex d,N st d > 0 & for n st n >= N holds t.n>=d*g.n & t.n>=0 } by ASYMPT_0:def 13; A4: f is Element of Funcs(NAT, REAL) by FUNCT_2:11; 2 to_power (k+1) > 0 by POWER:39; then A5: (2 to_power (k+1))" > 0 by REAL_1:72; now let n; assume A6: n >= 1; set p = seq_n^(k); ex s being Real_Sequence st s.0 = 0 & (for m st m > 0 holds s.m = (n to_power k)) proof defpred P[Nat,Real] means ($1 = 0 implies $2 = 0) & ($1 > 0 implies $2 = (n to_power k)); A7: for x being Element of NAT ex y being Element of REAL st P[x,y] proof let n; per cases; suppose n = 0; hence thesis; suppose n > 0; hence thesis; end; consider h being Function of NAT,REAL such that A8: for x being Element of NAT holds P[x,h.x] from FuncExD(A7); take h; thus h.0 = 0 by A8; let n; thus thesis by A8; end; then consider q being Real_Sequence such that A9: q.0 = 0 & (for m st m > 0 holds q.m = n to_power k); A10: f.n = Sum(p, n) by A1; now let m; assume A11: m <= n; per cases; suppose m = 0; hence p.m <= q.m by A9,Def3; suppose A12: m > 0; then A13: p.m = m to_power k by Def3; q.m = n to_power k by A9,A12; hence p.m <= q.m by A11,A12,A13,Lm6; end; then A14: Sum(p, n) <= Sum(q, n) by Lm18; Sum(q, n) = (n to_power k)*n by A9,Lm19 .= (n to_power k)*(n to_power 1) by POWER:30 .= n to_power (k+1) by A6,POWER:32 .= g.n by A6,Def3; hence f.n <= 1*g.n by A1,A14; now let m; assume m <= n; per cases; suppose m = 0; hence p.m >= 0 by Def3; suppose m > 0; then p.m = m to_power k by Def3; hence p.m >= 0; end; hence f.n >= 0 by A10,Lm17; end; then A15: f in Big_Oh(g) by A2,A4; now let n; assume A16: n >= 1; set p = seq_n^(k); ex s being Real_Sequence st s.0 = 0 & (for m st m > 0 holds s.m = ((n/2) to_power k)) proof defpred P[Nat,Real] means ($1 = 0 implies $2 = 0) & ($1 > 0 implies $2 = ((n/2) to_power k)); A17: for x being Element of NAT ex y being Element of REAL st P[x,y] proof let n; per cases; suppose n = 0; hence thesis; suppose n > 0; hence thesis; end; consider h being Function of NAT,REAL such that A18: for x being Element of NAT holds P[x,h.x] from FuncExD(A17); take h; thus h.0 = 0 by A18; let n; thus thesis by A18; end; then consider q being Real_Sequence such that A19: q.0 = 0 & (for m st m > 0 holds q.m = (n/2) to_power k); A20: f.n = Sum(p, n) by A1; A21: g.n = n to_power (k+1) by A16,Def3; set n1 = [/n/2\]; A22: n*2" > 0*2" by A16,REAL_1:70; then A23: n/2 > 0 by XCMPLX_0:def 9; A24: [/n/2\] >= n/2 by INT_1:def 5; then reconsider n1 as Nat by INT_1:16; set n2 = n1-1; now assume n2 < 0; then n1-1 <= -1 by INT_1:21; then (n1-1)+1 <= -1+1 by AXIOMS:24; then (n1+-1)+1 <= 0 by XCMPLX_0:def 8; then n1+(-1+1) <= 0 by XCMPLX_1:1; hence contradiction by A22,A24,XCMPLX_0:def 9; end; then reconsider n2 as Nat by INT_1:16; A25: now let m; assume m <= n; per cases; suppose m = 0; hence p.m >= 0 by Def3; suppose m > 0; then p.m = m to_power k by Def3; hence p.m >= 0; end; then A26: Sum(p,n) >= 0 by Lm17; A27: Sum(p,n,n2) = Sum(p,n) - Sum(p,n2) by BHSP_4:def 7; now let m; assume A28: m <= n2; n1 <= n1+1 by NAT_1:29; then n2 <= n1 by REAL_1:86; then A29: m <= n1 by A28,AXIOMS:22; n1 <= n by Lm22; then m <= n by A29,AXIOMS:22; hence p.m >= 0 by A25; end; then Sum(p,n2) >= 0 by Lm17; then Sum(p,n) + Sum(p,n2) >= Sum(p,n) + 0 by REAL_1:55; then A30: Sum(p,n) >= Sum(p,n,n2) by A27,REAL_1:86; A31: n1 = n2+1 by XCMPLX_1:27; A32: for N0 st n1 <= N0 & N0 <= n holds q.N0 <= p.N0 proof let N0; assume A33: n1 <= N0 & N0 <= n; then A34: N0 > 0 by A22,A24,XCMPLX_0:def 9; then A35: q.N0 = (n/2) to_power k by A19; A36: p.N0 = N0 to_power k by A34,Def3; N0 >= n/2 by A24,A33,AXIOMS:22; hence q.N0 <= p.N0 by A23,A35,A36,Lm6; end; n >= n1 by Lm22; then n+-1 >= n1+-1 by AXIOMS:24; then n-1 >= n1+-1 by XCMPLX_0:def 8; then n-1 >= n2 by XCMPLX_0:def 8; then n >= n2 + 1 by REAL_1:84; then A37: Sum(p,n,n2) >= Sum(q,n,n2) by A31,A32,Lm21; A38: Sum(q,n,n2) = (n-n2)*((n/2) to_power k) by A19,Lm23; A39: now assume n-n2 < n/2; then A40: n < n/2 + n2 by REAL_1:84; [/n/2\] < n/2 + 1 by INT_1:def 5; then n2 < n/2 by REAL_1:84; then n/2 + n2 < n/2 + n/2 by REAL_1:53; then n < 1*(n/2) + 1*(n/2) by A40,AXIOMS:22; then n < (1+1)*(n/2) by XCMPLX_1:8; then n < 2*(n*2") by XCMPLX_0:def 9; then n < (2*2")*n by XCMPLX_1:4; hence contradiction; end; ((n/2) to_power k) > 0 by A23,POWER:39; then Sum(q,n,n2) >= (n/2) * ((n/2) to_power k) by A38,A39,AXIOMS:25; then Sum(q,n,n2) >= ((n/2) to_power 1) * ((n/2) to_power k) by POWER:30; then Sum(q,n,n2) >= ((n/2) to_power (k+1)) by A23,POWER:32; then Sum(q,n,n2) >= (n to_power (k+1))/(2 to_power (k+1)) by A16,POWER:36 ; then Sum (q,n,n2) >= (n to_power (k+1))*(2 to_power (k+1))" by XCMPLX_0:def 9; then Sum(p,n,n2) >= (n to_power (k+1))*(2 to_power (k+1))" by A37,AXIOMS: 22; hence (2 to_power (k+1))"*g.n <= f.n by A20,A21,A30,AXIOMS:22; thus f.n >= 0 by A1,A26; end; then f in Big_Omega(g) by A3,A4,A5; then f in Big_Oh(g) /\ Big_Omega(g) by A15,XBOOLE_0:def 3; hence thesis by ASYMPT_0:def 14; end; :: Example p. 89 theorem for f being Real_Sequence st (f.0 = 0 & (for n st n > 0 holds f.n = (n to_power log(2,n)))) holds ex s being eventually-positive Real_Sequence st s = f & not s is smooth proof let f be Real_Sequence such that A1: f.0 = 0 & (for n st n > 0 holds f.n = (n to_power log(2,n))); set g = f taken_every 2; f is eventually-positive proof take 1; let n; assume A2: n >= 1; then f.n = n to_power log(2,n) by A1; hence f.n > 0 by A2,POWER:39; end; then reconsider f as eventually-positive Real_Sequence; take f; A3: Big_Oh(f) = { t where t is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds t.n <= c*f.n & t.n >= 0 } by ASYMPT_0:def 12; now assume f is smooth; then f is_smooth_wrt 2 by ASYMPT_0:def 20; then g in Big_Oh(f) by ASYMPT_0:def 19; then consider t being Element of Funcs(NAT, REAL) such that A4: t = g and A5: ex c,N st c > 0 & for n st n >= N holds t.n <= c*f.n & t.n >= 0 by A3; consider c,N such that A6: c > 0 and A7: for n st n >= N holds t.n <= c*f.n & t.n >= 0 by A5; set N0 = [/sqrt c / sqrt 2\]; reconsider N2=max(N, N0) as Integer by SQUARE_1:49; set N1 = max( N2, 2 ); A8: N1 >= N2 & N1 >= 2 by SQUARE_1:46; N2 >= N & N2 >= N0 by SQUARE_1:46; then A9: N1 >= N & N1 >= N0 by A8,AXIOMS:22; N1 is Integer by SQUARE_1:49; then reconsider N1 as Nat by A8,INT_1:16; set n = N1 + 1; n > N1 + 0 by REAL_1:67; then A10: n > N0 & n > N & n > 2 by A8,A9,AXIOMS:22; A11: n to_power log(2,n) > 0 by POWER:39; A12: 2*n > 2*0 by REAL_1:70; A13: (2*n^2)*(n to_power log(2,n)) = (2*n) to_power log(2,2*n) proof thus (2*n^2)*(n to_power log(2,n)) = (2*(n*n))*(n to_power log(2,n)) by SQUARE_1:def 3 .= ((2*n)*n)*(n to_power log(2,n)) by XCMPLX_1:4 .= ((2*n)*(2 to_power (log(2,n))))*(n to_power log(2,n)) by POWER:def 3 .= (2*n)*((2 to_power (log(2,n)))*(n to_power log(2,n))) by XCMPLX_1:4 .= (2*n)*((2*n) to_power log(2,n)) by POWER:35 .= ((2*n) to_power 1)*((2*n) to_power log(2,n)) by POWER:30 .= (2*n) to_power (1 + log(2,n)) by A12,POWER:32 .= (2*n) to_power (log(2,2) + log(2,n)) by POWER:60 .= (2*n) to_power log(2,2*n) by POWER:61; end; A14: sqrt 2 > 0 by SQUARE_1:93; A15: sqrt 2 <> 0 by SQUARE_1:93; A16: sqrt c > 0 by A6,SQUARE_1:93; N0 >= sqrt c / sqrt 2 by INT_1:def 5; then n > sqrt c / sqrt 2 by A10,AXIOMS:22; then n*sqrt 2 > (sqrt c / sqrt 2) * sqrt 2 by A14,REAL_1:70; then n*sqrt 2 > sqrt c by A15,XCMPLX_1:88; then (n*sqrt 2)^2 > (sqrt c)^2 by A16,SQUARE_1:78; then (n*sqrt 2)^2 > c by A6,SQUARE_1:def 4; then n^2*(sqrt 2)^2 > c by SQUARE_1:68; then 2*n^2 > c by SQUARE_1:def 4; then (2*n) to_power log(2,2*n) > c*(n to_power log(2,n)) by A11,A13,REAL_1: 70; then f.(2*n) > c*(n to_power log(2,n)) by A1,A12; then t.n > c*(n to_power log(2,n)) by A4,ASYMPT_0:def 18; then t.n > c*f.n by A1; hence contradiction by A7,A10; end; hence thesis; end; :: Example p. 92 definition let b be Real; func seq_const(b) -> Real_Sequence equals :Def4: NAT --> b; coherence proof dom (NAT --> b) = NAT & rng (NAT --> b) = {b} by FUNCOP_1:14,19; hence NAT --> b is Real_Sequence by FUNCT_2:def 1,RELSET_1:11; end; end; definition cluster seq_const(1) -> eventually-nonnegative; coherence proof take 0; let n; assume n >= 0; (seq_const(1)).n = (NAT --> 1).n by Def4 .= 1 by FUNCOP_1:13; hence (seq_const(1)).n >= 0; end; end; Lm24: for a,b,c being Real holds a>1 & b>=a & c>=1 implies log(a,c) >= log(b,c) proof let a,b,c be Real; assume that A1: a > 1 and A2: b >= a and A3: c >= 1; A4: b > 1 by A1,A2,AXIOMS:22; A5: b > 0 & b <> 1 by A1,A2; log(b,c) >= log(b,1) by A3,A4,PRE_FF:12; then A6: log(b,c) >= 0 by A4,POWER:59; log(a,b) >= log(a,a) by A1,A2,PRE_FF:12; then log(a,b) >= 1 by A1,POWER:60; then log(a,b)*log(b,c) >= 1*log(b,c) by A6,AXIOMS:25; hence thesis by A1,A3,A5,POWER:64; end; theorem Th9: for f being eventually-nonnegative Real_Sequence holds ex F being FUNCTION_DOMAIN of NAT,REAL st F = { seq_n^(1) } & (f in F to_power Big_Oh(seq_const(1)) iff ex N,c,k st c>0 & for n st n >= N holds 1 <= f.n & f.n <= c*(seq_n^(k)).n) proof let h be eventually-nonnegative Real_Sequence; reconsider F = { seq_n^(1) } as FUNCTION_DOMAIN of NAT,REAL by FRAENKEL:10; take F; thus F = { seq_n^(1) }; set G = Big_Oh(seq_const(1)); set p = seq_const(1); A1: F to_power G = { t where t is Element of Funcs(NAT,REAL) : ex f,g being Element of Funcs(NAT,REAL), N being Element of NAT st f in F & g in G & for n being Element of NAT st n >= N holds t.n = (f.n) to_power (g.n) } by ASYMPT_0:def 23; A2:G = { t where t is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds t.n <= c*p.n & t.n >= 0 } by ASYMPT_0:def 12; now hereby assume h in F to_power Big_Oh(seq_const(1)); then consider t being Element of Funcs(NAT,REAL) such that A3: h = t and A4: ex f,g being Element of Funcs(NAT,REAL), N being Element of NAT st f in F & g in G & for n being Element of NAT st n>=N holds t.n=(f.n) to_power (g.n) by A1 ; consider f,g being Element of Funcs(NAT,REAL), N0 being Element of NAT such that A5: f in F and A6: g in G and A7: for n being Element of NAT st n>=N0 holds t.n = (f.n) to_power (g.n) by A4; A8:f = seq_n^(1) by A5,TARSKI:def 1; consider g' being Element of Funcs(NAT, REAL) such that A9: g = g' and A10: ex c,N st c > 0 & for n st n >= N holds g'.n <= c*p.n & g'.n >= 0 by A2, A6; consider c,N1 such that A11: c > 0 and A12: for n st n >= N1 holds g'.n <= c*p.n & g'.n >= 0 by A10; set N = max(2,max(N0,N1)); max(N0,N1) >= N0 & max(N0,N1) >= N1 & N >= max(N0,N1) by SQUARE_1:46; then A13: N >= N0 & N >= N1 & N >= 2 by AXIOMS:22,SQUARE_1:46; set k = [/c\]; A14: k >= c by INT_1:def 5; k > 0 by A11,INT_1:def 5; then reconsider k as Nat by INT_1:16; reconsider i = 1 as Nat; take N,i,k; thus i > 0; let n; assume A15: n >= N; then A16: n >= N0 & n >= N1 & n >= 2 by A13,AXIOMS:22; then A17: g'.n <= c*p.n & g'.n >= 0 by A12; A18: p.n = (NAT --> 1).n by Def4 .= 1 by FUNCOP_1:13; A19: n > 0 & n > 1 by A16,AXIOMS:22; f.n = n to_power 1 by A8,A13,A15,Def3 .= n by POWER:30; then A20: h.n = n to_power (g.n) by A3,A7,A16; n to_power (g.n) >= n to_power 0 by A9,A17,A19,PRE_FF:10; hence 1 <= h.n by A20,POWER:29; g.n <= c*p.n by A9,A12,A16; then A21: h.n <= n to_power (c*1) by A18,A19,A20,PRE_FF:10; n to_power c <= n to_power k by A14,A19,PRE_FF:10; then h.n <= n to_power k by A21,AXIOMS:22; hence h.n <= i*(seq_n^(k)).n by A13,A15,Def3; end; given N0,c,k such that c > 0 and A22: for n st n >= N0 holds 1 <= h.n & h.n <= c*(seq_n^(k)).n; reconsider t = h as Element of Funcs(NAT,REAL) by FUNCT_2:11; A23: G = { s where s is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds s.n <= c*(seq_const(1)).n & s.n>= 0 } by ASYMPT_0:def 12; set N = max(N0,2); A24: N >= N0 & N >= 2 by SQUARE_1:46; then A25:N > 1 by AXIOMS:22; A26:N > 0 & N <> 1 by SQUARE_1:46; set c1= max(c,2); A27: c1 >= c & c1 >= 2 by SQUARE_1:46; then A28:c1 > 0 & c1 > 1 by AXIOMS:22; set a = log(N,c1); set b = k+a; reconsider f = seq_n^(1) as Element of Funcs(NAT,REAL) by FUNCT_2:11; defpred Q[Nat,Real] means ($1 < N implies $2 = 1) & ($1 >= N implies $2 = log($1,t.$1)); A29:for x being Element of NAT ex y being Element of REAL st Q[x,y] proof let n; per cases; suppose n < N; hence thesis; suppose n >= N; hence thesis; end; consider g being Function of NAT,REAL such that A30:for x being Element of NAT holds Q[x,g.x] from FuncExD(A29); A31:g is Element of Funcs(NAT,REAL) by FUNCT_2:11; A32: now let n be Element of NAT; assume A33: n >= N; then n >= N0 by A24,AXIOMS:22; then A34: t.n >= 1 by A22; thus (f.n) to_power (g.n) = (n to_power 1) to_power (g.n) by A24,A33,Def3 .= n to_power (g.n) by POWER:30 .= n to_power (1*log(n,t.n)) by A30,A33 .= t.n by A24,A25,A33,A34,POWER:def 3; end; A35: f in F by TARSKI:def 1; now log(N,1) = 0 by A26,POWER:59; then a > 0 by A25,A28,POWER:65; then k+a > k+0 by REAL_1:53; hence b > 0; let n; assume A36:n >= N; then A37: n >= N0 & n >= 2 by A24,AXIOMS:22; then A38: n > 1 by AXIOMS:22; A39: n > 0 & n <> 1 by A24,A36,AXIOMS:22; A40: a >= log(n,c1) by A25,A28,A36,Lm24; A41: 1 <= t.n & t.n <= c*(seq_n^(k)).n by A22,A37; A42:(seq_n^(k)).n = n to_power k by A24,A36,Def3; then A43:(seq_n^(k)).n > 0 by A24,A36,POWER:39; c*(seq_n^(k)).n <= c1*(seq_n^(k)).n by A27,A42,AXIOMS:25; then t.n <= c1*(seq_n^(k)).n by A41,AXIOMS:22; then A44: log(n,t.n) <= log(n,c1*(seq_n^(k)).n) by A38,A41,PRE_FF:12; t.n = (f.n) to_power (g.n) by A32,A36 .= (n to_power 1) to_power (g.n) by A24,A36,Def3 .= n to_power (g.n) by POWER:30; then A45: log(n,t.n) = g.n*log(n,n) by A39,POWER:63 .= g.n*1 by A39,POWER: 60; A46: log(n,c1*(seq_n^(k)).n) = log(n,c1) + log(n,n to_power k) by A27,A39,A42,A43,POWER:61 .= log(n,c1) + k*log(n,n) by A39,POWER:63 .= log(n,c1) + k*1 by A39,POWER:60; A47: log(n,c1) + k <= a + k by A40,AXIOMS:24; (seq_const(1)).n = (NAT --> 1).n by Def4 .= 1 by FUNCOP_1:13; hence g.n <= b*(seq_const(1)).n by A44,A45,A46,A47,AXIOMS:22; g.n = log(n,t.n) by A30,A36; then g.n >= log(n,1) by A38,A41,PRE_FF:12; hence g.n >= 0 by A39,POWER:59; end; then g in G by A23,A31; hence h in F to_power Big_Oh(seq_const(1)) by A1,A31,A32,A35; end; hence thesis; end; begin :: Problem 3.1 theorem for f being Real_Sequence st (for n holds f.n = 3*(10 to_power 6) - 18*(10 to_power 3)*n + 27*n^2) holds f in Big_Oh(seq_n^(2)) proof let f be Real_Sequence such that A1: for n holds f.n = 3*(10 to_power 6) - 18*(10 to_power 3)*n + 27*n^2; set g = seq_n^(2); A2: Big_Oh(g) = { t where t is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds t.n <= c*g.n & t.n >= 0 } by ASYMPT_0:def 12; A3: f is Element of Funcs(NAT, REAL) by FUNCT_2:11; consider t1 being Nat such that A4: t1 = 10*10*10; consider t2 being Nat such that A5: t2 = t1*t1; t1 = 10*10^2 by A4,SQUARE_1:def 3; then t1 = 10*(10 to_power 2) by POWER:53; then t1 = (10 to_power 1)*(10 to_power 2) by POWER:30; then A6: t1 = 10 to_power (1+2) by POWER:32; then A7: t2 = 10 to_power (3+3) by A5,POWER:32 .= 10 to_power 6; A8: 10 to_power 3 = 10 to_power (2+1) .= (10 to_power 2)*(10 to_power 1) by POWER:32 .= (10 to_power 2)*10 by POWER:30 .= (10^2)*10 by POWER:53 .= (10*10)*10 by SQUARE_1:def 3 .= 1000; A9: t1 > 0 by A6,POWER:39; A10: for n st n >= 400 holds f.n <= 27*n^2 proof let n such that A11: n >= 400; now assume f.n > 27*n^2; then 3*t2 - 18*(10 to_power 3)*n + 27*n^2 > 27*n^2 by A1,A7; then 3*t2 + -18*t1*n + 27*n^2 > 27*n^2 by A6,XCMPLX_0:def 8; then 3*t2 + -18*t1*n > 27*n^2 - 27*n^2 by REAL_1:84; then 3*t2 + -18*t1*n > 0 by XCMPLX_1:14; then 3*t2 - 18*t1*n > 0 by XCMPLX_0:def 8; then A12: 3*t2 > 0 + 18*t1*n by REAL_1:86; (18*t1)*n >= (18*t1)*400 by A11,AXIOMS:25; then 3*t2 > (18*t1)*400 by A12,AXIOMS:22; then 3*t2 > t1*(18*400) by XCMPLX_1:4; then 3*(10 to_power (3+3)) > t1*7200 by A7; then 3*((10 to_power 3)*(10 to_power 3)) > t1*7200 by POWER:32; then (3*1000)*t1 > t1*7200 by A6,A8,XCMPLX_1:4; hence contradiction by A9,REAL_1:70; end; hence thesis; end; A13: for n st n >= 400 holds 18*t1*n - 3*t2 < 27*n^2 proof defpred _P[Nat] means 18*t1*$1 - 3*t2 < 27*$1^2; A14: _P[400] proof A15: 18*t1*400 - 3*t2 = 7200*t1 - (3*1000)*t1 by A5,A6,A8,XCMPLX_1:4 .= (7200 - 3000)*t1 by XCMPLX_1:40 .= 4200*t1; 27*400^2 = 27*(400*400) by SQUARE_1:def 3 .= 27*((4*40)*t1) by A6,A8 .= (27*160)*t1 by XCMPLX_1:4 .= 4320*t1; hence thesis by A9,A15,REAL_1:70; end; A16: for k st k >= 400 & _P[k] holds _P[k+1] proof let k such that A17: k >= 400 and A18: 18*t1*k - 3*t2 < 27*k^2; 18*t1*(k+1) - 3*t2 = (18*t1*k + 18*t1*1) - 3*t2 by XCMPLX_1:8 .= (18*t1*k + 18*t1*1) + -3*t2 by XCMPLX_0:def 8 .= (18*t1*k + -3*t2) + 18*t1 by XCMPLX_1:1 .= (18*t1*k - 3*t2) + 18*t1 by XCMPLX_0:def 8; then A19: 18*t1*(k+1) - 3*t2 < 27*k^2 + 18*t1 by A18,REAL_1:53; 54*400 <= 54*k by A17,AXIOMS:25; then A20: 18*t1 < 54*k by A6,A8,AXIOMS:22; 54*k + 0 <= 54*k + 27 by REAL_1:55; then 18*t1 < 54*k + 27 by A20,AXIOMS:22; then A21: 27*k^2 + 18*t1 < 27*k^2 + (54*k + 27) by REAL_1:53; 27*k^2 + (54*k + 27) = 27*k^2 + ((27*2)*k + 27) .= 27*k^2 + (27*(2*k) + 27*1) by XCMPLX_1:4 .= 27*k^2 + 27*(2*k + 1) by XCMPLX_1:8 .= 27*(k^2 + (2*k + 1)) by XCMPLX_1:8 .= 27*(k^2 + 2*k*1 + 1*1) by XCMPLX_1:1 .= 27*(k^2 + 2*k*1 + 1^2) by SQUARE_1:def 3 .= 27*(k+1)^2 by SQUARE_1:63; hence thesis by A19,A21,AXIOMS:22; end; for n st n >= 400 holds _P[n]from Ind1(A14, A16); hence thesis; end; now let n; assume A22: n >= 400; then f.n <= 27*n^2 by A10; then f.n <= 27*(n to_power 2) by POWER:53; hence f.n <= 27*g.n by A22,Def3; 0 + (18*t1*n - 3*t2) < 27*n^2 by A13,A22; then 0 < 27*n^2 - (18*t1*n - 3*t2) by REAL_1:86; then 0 < 27*n^2 - 18*t1*n + 3*t2 by XCMPLX_1:37; then 0 < 27*n^2 + -18*t1*n + 3*t2 by XCMPLX_0:def 8; then 0 < 3*t2 + -18*t1*n + 27*n^2 by XCMPLX_1:1; then 0 < 3*(10 to_power 6) - 18*t1*n + 27*n^2 by A7,XCMPLX_0:def 8; hence f.n >= 0 by A1,A6; end; hence thesis by A2,A3; end; :: Problem 3.2 -- omitted :: Problem 3.3 -- omitted :: Problem 3.4 -- omitted begin :: Problem 3.5 theorem :: Part 1 seq_n^(2) in Big_Oh(seq_n^(3)) proof set f = seq_n^(2); set g = seq_n^(3); A1: Big_Oh(g) = { t where t is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds t.n <= c*g.n & t.n >= 0 } by ASYMPT_0:def 12; A2: f is Element of Funcs(NAT, REAL) by FUNCT_2:11; now let n; assume A3: n >= 2; then A4: n > 1 by AXIOMS:22; A5: f.n = n to_power 2 by A3,Def3; g.n = n to_power 3 by A3,Def3; hence f.n <= 1*g.n by A4,A5,POWER:44; thus f.n >= 0 by A5; end; hence thesis by A1,A2; end; theorem :: Part 2 not seq_n^(2) in Big_Omega(seq_n^(3)) proof set f = seq_n^(2); set g = seq_n^(3); A1: Big_Omega(g) = { s where s is Element of Funcs(NAT, REAL) : ex d,N st d > 0 & for n st n >= N holds d*g.n <= s.n&s.n >= 0 } by ASYMPT_0:def 13; now assume seq_n^(2) in Big_Omega(seq_n^(3)); then consider s being Element of Funcs(NAT, REAL) such that A2: s = f and A3: ex d,N st d > 0 & for n st n >= N holds d*g.n <= s.n & s.n >= 0 by A1; consider d,N such that A4: d > 0 and A5: for n st n >= N holds d*g.n <= s.n & s.n >= 0 by A3; A6: N+2 > 1+0 by REAL_1:67; ex n st n >= N & d*g.n > s.n proof take n = max( N, [/(N+2)/d\] ); A7: n >= [/(N+2)/d\] & n >= N by SQUARE_1:46; d" > 0 by A4,REAL_1:72; then A8: (N+2)*d" > 0*d" by REAL_1:70; A9: [/(N+2)/d\] >= (N+2)/d by INT_1:def 5; then A10: n > 0 by A7,A8,XCMPLX_0:def 9; n is Integer by SQUARE_1:49; then reconsider n as Nat by A7,INT_1:16; A11: (d*g.n)*(n to_power -2) = d*(n to_power 3)*(n to_power -2) by A10,Def3 .= d*((n to_power 3)*(n to_power -2)) by XCMPLX_1:4 .= d*(n to_power (3+(-2))) by A10,POWER:32 .= d*n by POWER:30; A12: f.n*(n to_power -2) = (n to_power 2)*(n to_power -2) by A10,Def3 .= (n to_power (2+(-2))) by A10,POWER:32 .= 1 by POWER:29; A13: d*([/(N+2)/d\]) >= d*((N+2)/d) by A4,A9,AXIOMS:25; d*n >= d*([/(N+2)/d\]) by A4,A7,AXIOMS:25; then d*n >= ((N+2)/d)*d by A13,AXIOMS:22; then d*n >= N+2 by A4,XCMPLX_1:88; then A14: (d*g.n)*(n to_power -2) > f.n*(n to_power -2) by A6,A11,A12, AXIOMS:22; (n to_power -2) > 0 by A10,POWER:39; hence thesis by A2,A14,AXIOMS:25,SQUARE_1:46; end; hence contradiction by A5; end; hence thesis; end; Lm25: for a being positive Real, b,c being Real holds seq_a^(a,b,c) is eventually-positive; theorem :: Part 3 ex s being eventually-positive Real_Sequence st s = seq_a^(2,1,1) & seq_a^(2,1,0) in Big_Theta(s) proof set f = seq_a^(2,1,0); reconsider g = seq_a^(2,1,1) as eventually-positive Real_Sequence; take g; thus g = seq_a^(2,1,1); A1: Big_Theta(g) = { s where s is Element of Funcs(NAT, REAL) : ex c,d,N st c > 0 & d > 0 & for n st n >= N holds d*g.n <= s.n & s.n <= c*g.n } by ASYMPT_0:27; A2: f is Element of Funcs(NAT, REAL) by FUNCT_2:11; A3: (2 to_power -1) > 0 by POWER:39; now let n; assume n >= 2; A4: f.n = 2 to_power (1*n+0) by Def1; A5: g.n = 2 to_power (1*n+1) by Def1; then (2 to_power -1)*g.n = 2 to_power ((-1)+(n+1)) by POWER:32 .= 2 to_power (n+(1+-1)) by XCMPLX_1:1 .= f.n by A4; hence (2 to_power -1)*g.n <= f.n; n+0 <= n+1 by REAL_1:55; hence f.n <= 1*g.n by A4,A5,PRE_FF:10; end; hence thesis by A1,A2,A3; end; definition let a be Nat; func seq_n!(a) -> Real_Sequence means :Def5: it.n = (n+a)!; existence proof deffunc _F(Nat) = ($1+a)!; consider h being Function of NAT, REAL such that A1: for n being Nat holds h.n = _F(n) from LambdaD; take h; let n; thus h.n = (n+a)! by A1; end; uniqueness proof let j,k be Real_Sequence such that A2: for n holds j.n = (n+a)! and A3: for n holds k.n = (n+a)!; now let n; thus j.n = (n+a)! by A2 .= k.n by A3; end; hence j = k by FUNCT_2:113; end; end; definition let a be Nat; cluster seq_n!(a) -> eventually-positive; coherence proof set f = seq_n!(a); take 0; let n; assume n >= 0; f.n = (n+a)! by Def5; hence f.n > 0 by NEWTON:23; end; end; theorem :: Part 4 not seq_n!(0) in Big_Theta(seq_n!(1)) proof set f = seq_n!(0); set g = seq_n!(1); A1: Big_Theta(g) = { s where s is Element of Funcs(NAT, REAL) : ex c,d,N st c > 0 & d > 0 & for n st n >= N holds d*g.n <= s.n & s.n <= c*g.n } by ASYMPT_0:27; now assume f in Big_Theta(g); then consider s being Element of Funcs(NAT, REAL) such that A2: s = f and A3: ex c,d,N st c > 0 & d > 0 & for n st n >= N holds d*g.n <= s.n & s.n <= c*g.n by A1; consider c,d,N such that c > 0 and A4: d > 0 and A5: for n st n >= N holds d*g.n <= s.n & s.n <= c*g.n by A3; ex n st n >= N & d*g.n > f.n proof take n = max( N, [/(N+1)/d\] ); A6: n >= N & n >= [/(N+1)/d\] by SQUARE_1:46; n is Integer by SQUARE_1:49; then reconsider n as Nat by A6,INT_1:16; n! > 0 by NEWTON:23; then A7: (n!)" > 0 by REAL_1:72; A8: n! <> 0 by NEWTON:23; A9: (d*g.n)*(n!)" = (d*((n+1)!))*(n!)" by Def5 .= (d*((n+1)*(n!)))*(n!)" by NEWTON:21 .= ((d*(n+1))*(n!))*(n!)" by XCMPLX_1:4 .= (d*(n+1)*((n!)*(n!)")) by XCMPLX_1:4 .= (d*(n+1))*1 by A8,XCMPLX_0:def 7 .= d*(n+1); [/(N+1)/d\] >= (N+1)/d by INT_1:def 5; then [/(N+1)/d\] + 1 >= (N+1)/d + 1 by AXIOMS:24; then A10: d*([/(N+1)/d\] + 1) >= d*((N+1)/d + 1) by A4,AXIOMS:25; n+1 >= [/(N+1)/d\]+1 by A6,AXIOMS:24; then d*(n+1) >= d*([/(N+1)/d\]+1) by A4,AXIOMS:25; then A11: d*(n+1) >= d*(((N+1)/d) + 1) by A10,AXIOMS:22; A12: d*(((N+1)/d) + 1) = d*((N+1)/d) + d*1 by XCMPLX_1:8 .= (N+1) + d by A4,XCMPLX_1:88; N+1 >= 1+0 by AXIOMS:24; then d*(((N+1)/d) + 1) > 1 by A4,A12,REAL_1:67; then A13: (d*g.n)*(n!)" > 1 by A9,A11,AXIOMS:22; f.n*(n!)" = (n+0)!*(n!)" by Def5 .= 1 by A8,XCMPLX_0:def 7; hence thesis by A7,A13,AXIOMS:25,SQUARE_1:46; end; hence contradiction by A2,A5; end; hence thesis; end; begin :: Problem 3.6 Lm26: now let a,b,c,d be Real; assume 0 <= b & a <= b & 0 <= c & c <= d; then a*c <= b*c & b*c <= b*d by AXIOMS:25; hence a*c <= b*d by AXIOMS:22; end; theorem for f being Real_Sequence st f in Big_Oh(seq_n^(1)) holds f(#)f in Big_Oh(seq_n^(2)) proof set g = seq_n^(1); set h = seq_n^(2); let f be Real_Sequence; assume A1: f in Big_Oh(g); A2: Big_Oh(g) = { t where t is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds t.n <= c*g.n & t.n >= 0 } by ASYMPT_0:def 12; A3: Big_Oh(h) = { t where t is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds t.n <= c*h.n & t.n >= 0 } by ASYMPT_0:def 12; consider t being Element of Funcs(NAT, REAL) such that A4: t = f and A5: ex c,N st c > 0 & for n st n >= N holds t.n <= c*g.n & t.n >= 0 by A1,A2; consider c, N such that A6: c > 0 and A7: for n st n >= N holds t.n <= c*g.n & t.n >= 0 by A5; set d = max( c, c*c ); A8: d >= c & d >= c*c by SQUARE_1:46; A9: d > 0 by A6,SQUARE_1:46; A10: t(#)t is Element of Funcs(NAT, REAL) by FUNCT_2:11; A11: 0 to_power 1 = 0 by POWER:def 2; now take N; let n; assume n >= N; then A12: t.n <= c*g.n & t.n >= 0 by A7; A13: g.n >= 0 proof per cases; suppose n = 0; hence thesis by Def3; suppose n > 0; then g.n = n to_power 1 by Def3 .= n by POWER:30; hence thesis; end; A14: for n holds g.n <= h.n proof let n; per cases; suppose A15: n = 0; then g.n = 0 by Def3; hence g.n <= h.n by A15,Def3; suppose n > 0; then A16: n >= 0+1 by NAT_1:38; thus g.n <= h.n proof per cases by A16,REAL_1:def 5; suppose A17: n = 1; (1 to_power 1) = 1 & (1 to_power 2) = 1 by POWER:31; then g.n = (1 to_power 2) by A17,Def3; hence g.n <= h.n by A17,Def3; suppose A18: n > 1; then (n to_power 1) < (n to_power 2) by POWER:44; then g.n < (n to_power 2) by A18,Def3; hence g.n <= h.n by A18,Def3; end; end; t.n*t.n <= (c*g.n)*(c*g.n) by A12,Lm26; then t.n*t.n <= (c*(g.n*c*g.n)) by XCMPLX_1:4; then A19: t.n*t.n <= (c*(c*(g.n*g.n))) by XCMPLX_1:4; A20: (n to_power 1)*(n to_power 1) = (n to_power (1+1)) proof per cases; suppose n = 0; hence thesis by A11,POWER:def 2; suppose n > 0; hence thesis by POWER:32; end; g.n*g.n = h.n proof per cases; suppose A21: n = 0; hence g.n*g.n = 0*g.n by Def3 .= h.n by A21,Def3; suppose A22: n > 0; hence g.n*g.n = (n to_power 1)*g.n by Def3 .= (n to_power (1+1)) by A20,A22,Def3 .= h.n by A22,Def3; end; then A23: t.n*t.n <= (c*c)*h.n by A19,XCMPLX_1:4; h.n >= g.n by A14; then (c*c)*h.n <= d*h.n by A8,A13,AXIOMS:25; then t.n*t.n <= d*h.n by A23,AXIOMS:22; hence (t(#)t).n <= d*h.n by SEQ_1:12; t.n*t.n >= t.n*0 by A12,AXIOMS:25; hence (t(#)t).n >= 0 by SEQ_1:12; end; hence thesis by A3,A4,A9,A10; end; begin :: Problem 3.7 theorem ex s being eventually-positive Real_Sequence st s = seq_a^(2,1,0) & 2(#)seq_n^(1) in Big_Oh( seq_n^(1) ) & not seq_a^(2,2,0) in Big_Oh(s) proof set f = 2(#)seq_n^(1); set g = seq_n^(1); set p = seq_a^(2,2,0); reconsider q = seq_a^(2,1,0) as eventually-positive Real_Sequence; take q; thus q = seq_a^(2,1,0); A1: Big_Oh(g) = { t where t is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds t.n <= c*g.n & t.n >= 0 } by ASYMPT_0:def 12; A2: Big_Oh(q) = { t where t is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds t.n <= c*q.n & t.n >= 0 } by ASYMPT_0:def 12; A3: f is Element of Funcs(NAT, REAL) by FUNCT_2:11; now let n; assume n >= 0; thus f.n <= 2*g.n by SEQ_1:13; A4: g.n = n proof per cases; suppose n = 0; hence thesis by Def3; suppose n > 0; hence g.n = n to_power 1 by Def3 .= n by POWER:30; end; 2*n >= 2*0; hence f.n >= 0 by A4,SEQ_1:13; end; hence f in Big_Oh(g) by A1,A3; now assume p in Big_Oh(q); then consider t being Element of Funcs(NAT, REAL) such that A5: t = p and A6: ex c,N st c > 0 & for n st n >= N holds t.n <= c*q.n & t.n >= 0 by A2; consider c,N such that A7: c > 0 and A8: for n st n >= N holds t.n <= c*q.n & t.n >= 0 by A6; ex n st n >= N & t.n > c*q.n proof take n = max( N, [/log(2,c)+1\] ); A9: n >= N & n >= [/log(2,c)+1\] by SQUARE_1:46; n is Integer by SQUARE_1:49; then reconsider n as Nat by A9,INT_1:16; A10: (c*q.n)*(2 to_power -n) = (c*(2 to_power (1*n+0)))*(2 to_power -n) by Def1 .= c*((2 to_power n)*(2 to_power -n)) by XCMPLX_1:4 .= c*(2 to_power (n+-n)) by POWER:32 .= c*(2 to_power 0) by XCMPLX_0:def 6 .= c*1 by POWER:29; A11: p.n*(2 to_power -n) = (2 to_power (2*n+0))*(2 to_power -n) by Def1 .= (2 to_power (2*n+-(1*n))) by POWER:32 .= (2 to_power (2*n+(-1)*n)) by XCMPLX_1:175 .= (2 to_power ((2+-1)*n)) by XCMPLX_1:8 .= (2 to_power (1*n)); A12: 2 to_power n >= 2 to_power [/log(2,c)+1\] by A9,PRE_FF:10; [/log(2,c)+1\] >= log(2,c)+1 by INT_1:def 5; then A13: 2 to_power [/log(2,c)+1\] >= 2 to_power (log(2,c)+1) by PRE_FF:10 ; 2 to_power (log(2,c)+1) = (2 to_power log(2,c))*(2 to_power 1) by POWER:32 .= c*(2 to_power 1) by A7,POWER:def 3 .= c*2 by POWER:30; then 2 to_power (log(2,c)+1) > (c*q.n)*(2 to_power -n) by A7,A10,REAL_1:70; then 2 to_power [/log(2,c)+1\] > (c*q.n)*(2 to_power -n) by A13,AXIOMS: 22; then A14: p.n*(2 to_power -n) > (c*q.n)*(2 to_power -n) by A11,A12,AXIOMS: 22; 2 to_power -n > 0 by POWER:39; hence thesis by A5,A14,AXIOMS:25,SQUARE_1:46; end; hence contradiction by A8; end; hence thesis; end; begin :: Problem 3.8 theorem log(2,3) < 159/100 implies seq_n^(log(2,3)) in Big_Oh(seq_n^(159/100)) & not seq_n^(log(2,3)) in Big_Omega(seq_n^(159/100)) & not seq_n^(log(2,3)) in Big_Theta(seq_n^(159/100)) proof assume A1: log(2,3) < 159/100; set f = seq_n^(log(2,3)); set g = seq_n^(159/100); set h = f/"g; set c = 159/100 - log(2,3); log(2,3) - log(2,3) < 159/100 - log(2,3) by A1,REAL_1:54; then A2: 0 < c by XCMPLX_1:14; then A3: c*2" > 0*2" by REAL_1:70; then A4: c/2 > 0 by XCMPLX_0:def 9; A5: c/2 <> 0 by A3,XCMPLX_0:def 9; A6: now let p be real number such that A7: p > 0; reconsider p1 = p as Real by XREAL_0:def 1; set N1 = max([/((1/p1) to_power (1/(c/2)))\], 2); A8: N1 >= [/((1/p) to_power (1/(c/2)))\] & N1 >= 2 by SQUARE_1:46; then A9: N1 > 1 by AXIOMS:22; N1 is Integer by SQUARE_1:49; then reconsider N1 as Nat by A8,INT_1:16; take N1; let n; assume A10: n >= N1; p" > 0 by A7,REAL_1:72; then A11: 1/p > 0 by XCMPLX_1:217; [/((1/p) to_power (1/(c/2)))\]>= ((1/p) to_power (1/(c/2))) by INT_1:def 5; then A12: N1 >= ((1/p) to_power (1/(c/2))) by A8,AXIOMS:22; A13: ((1/p1) to_power (1/(c/2))) > 0 by A11,POWER:39; A14: n > 1 by A9,A10,AXIOMS:22; A15: h.n = f.n/g.n by Lm4; f.n = (n to_power log(2,3)) by A8,A10,Def3; then A16: h.n = (n to_power log(2,3)) / (n to_power (159/100)) by A8,A10,A15, Def3 .= (n to_power (log(2,3) - (159/100))) by A8,A10,POWER:34 .= (n to_power -c) by XCMPLX_1:143; n >= ((1/p) to_power (1/(c/2))) by A10,A12,AXIOMS:22; then n to_power (c/2) >= ((1/p) to_power (1/(c/2))) to_power (c/2) by A4,A13,Lm6; then n to_power (c/2) >= (1/p1) to_power ((1/(c/2))*(c/2)) by A11,POWER:38; then n to_power (c/2) >= (1/p) to_power 1 by A5,XCMPLX_1:88; then n to_power (c/2) >= 1/p1 by POWER:30; then 1 / (n to_power (c/2)) <= 1 / (1/p) by A11,REAL_2:152; then 1 / (n to_power (c/2)) <= 1 / (p") by XCMPLX_1:217; then 1 / (n to_power (c/2)) <= p by XCMPLX_1:218; then A17: n to_power -(c/2) <= p by A8,A10,POWER:33; A18: n to_power (c/2) > 0 by A8,A10,POWER:39; c*(1/2) < c*1 by A2,REAL_1:70; then c/2 < c by XCMPLX_1:100; then n to_power (c/2) < n to_power c by A14,POWER:44; then 1 / (n to_power (c/2)) > 1 / (n to_power c) by A18,REAL_2:151; then n to_power -(c/2) > 1 / (n to_power c) by A8,A10,POWER:33; then h.n < n to_power -(c/2) by A8,A10,A16,POWER:33; then A19: h.n < p by A17,AXIOMS:22; h.n > 0 by A8,A10,A16,POWER:39; hence abs(h.n-0) < p by A19,ABSVALUE:def 1; end; then A20: h is convergent by SEQ_2:def 6; then A21:lim h = 0 by A6,SEQ_2:def 7; then A22: f in Big_Oh(g) & not g in Big_Oh(f) by A20,ASYMPT_0:16; then A23: f in Big_Oh(g) & not f in Big_Omega(g) by ASYMPT_0:19; A24: now assume f in Big_Theta(g); then f in Big_Oh(g) /\ Big_Omega(g) by ASYMPT_0:def 14; hence contradiction by A23,XBOOLE_0:def 3; end; thus f in Big_Oh(g) by A20,A21,ASYMPT_0:16; thus not f in Big_Omega(g) by A22,ASYMPT_0:19; thus not f in Big_Theta(g) by A24; end; :: Problem 3.9 -- Proven in theorem ASYMPT_0:10 :: Problem 3.10 -- Proven in theorem ASYMPT_0:12 begin :: Problem 3.11 theorem for f,g being Real_Sequence st (for n holds f.n = n mod 2) & (for n holds g.n = n+1 mod 2) holds ex s,s1 being eventually-nonnegative Real_Sequence st s = f & s1 = g & not s in Big_Oh(s1) & not s1 in Big_Oh(s) proof let f,g be Real_Sequence such that A1: for n holds f.n = n mod 2 and A2: for n holds g.n = n+1 mod 2; f is eventually-nonnegative proof take 0; let n; assume n >= 0; A3: f.n = n mod 2 by A1; per cases by A3,GROUP_4:100; suppose f.n = 0; hence f.n >= 0; suppose f.n = 1; hence f.n >= 0; end; then reconsider f as eventually-nonnegative Real_Sequence; g is eventually-nonnegative proof take 0; let n; assume n >= 0; A4: g.n = n+1 mod 2 by A2; per cases by A4,GROUP_4:100; suppose g.n = 0; hence g.n >= 0; suppose g.n = 1; hence g.n >= 0; end; then reconsider g as eventually-nonnegative Real_Sequence; take f,g; A5: Big_Oh(f) = { t where t is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds t.n <= c*f.n & t.n >= 0 } by ASYMPT_0:def 12; A6: Big_Oh(g) = { t where t is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds t.n <= c*g.n & t.n >= 0 } by ASYMPT_0:def 12; A7: now assume f in Big_Oh(g); then consider t being Element of Funcs(NAT, REAL) such that A8: t = f and A9: ex c,N st c > 0 & for n st n >= N holds t.n <= c*g.n & t.n >= 0 by A6; consider c,N such that c > 0 and A10: for n st n >= N holds t.n <= c*g.n & t.n >= 0 by A9; ex n st n >= N & t.n > c*g.n proof per cases by GROUP_4:100; suppose A11: N mod 2 = 0; A12: N+1 >= N by NAT_1:38; A13: t.(N+1) = (N+1) mod 2 by A1,A8 .= (0+(1 mod 2)) mod 2 by A11,EULER_2:8 .= (0+1) mod 2 by GROUP_4:102 .= 1 by GROUP_4:102; g.(N+1) = ((N+1)+1) mod 2 by A2 .= (N+(1+1)) mod 2 by XCMPLX_1:1 .= (0+(2 mod 2)) mod 2 by A11,EULER_2:8 .= (0+0) mod 2 by GR_CY_1:5 .= 0 by GR_CY_1:6; then c*g.(N+1) = 0; hence thesis by A12,A13; suppose A14: N mod 2 = 1; then A15: t.N = 1 by A1,A8; g.N = (N+1) mod 2 by A2 .= (1+(1 mod 2)) mod 2 by A14,EULER_2:8 .= (1+1) mod 2 by GROUP_4:102 .= 0 by GR_CY_1:5; then c*g.N = 0; hence thesis by A15; end; hence contradiction by A10; end; now assume g in Big_Oh(f); then consider t being Element of Funcs(NAT, REAL) such that A16: t = g and A17: ex c,N st c > 0 & for n st n >= N holds t.n <= c*f.n & t.n >= 0 by A5; consider c,N such that c > 0 and A18: for n st n >= N holds t.n <= c*f.n & t.n >= 0 by A17; ex n st n >= N & t.n > c*f.n proof per cases by GROUP_4:100; suppose A19: N mod 2 = 0; A20: t.N = (N+1) mod 2 by A2,A16 .= (0+(1 mod 2)) mod 2 by A19,EULER_2:8 .= (0+1) mod 2 by GROUP_4:102 .= 1 by GROUP_4:102; f.N = 0 by A1,A19; then c*f.N = 0; hence thesis by A20; suppose A21: N mod 2 = 1; A22: N+1 >= N by NAT_1:38; A23: t.(N+1) = ((N+1)+1) mod 2 by A2,A16 .= (N+(1+1)) mod 2 by XCMPLX_1:1 .= (1+(2 mod 2)) mod 2 by A21,EULER_2:8 .= (1+0) mod 2 by GR_CY_1:5 .= 1 by GROUP_4:102; f.(N+1) = (N+1) mod 2 by A1 .= (1+(1 mod 2)) mod 2 by A21,EULER_2:8 .= (1+1) mod 2 by GROUP_4:102 .= 0 by GR_CY_1:5; then c*f.(N+1) = 0; hence thesis by A22,A23; end; hence contradiction by A18; end; hence thesis by A7; end; :: Problem 3.12 -- Proven in theorems ASYMPT_0:20, ASYMPT_0:21 :: Problem 3.13 -- omitted (uses a notation that we do not support) :: Problem 3.14 -- omitted (cannot Mizar an incorrect proof) :: Problem 3.15 -- omitted (cannot Mizar an incorrect proof) :: Problem 3.16 -- omitted (maximum rule over multiple functions can be :: reduced to repeated applications of the binary :: maximum rule) :: Problem 3.17 -- omitted (cannot Mizar an incorrect proof) :: Problem 3.18 -- Proven in theorems ASYMPT_0:28, ASYMPT_0:29, ASYMPT_0:30 begin :: Problem 3.19 theorem :: Part 1 for f,g being eventually-nonnegative Real_Sequence holds Big_Oh(f) = Big_Oh(g) iff f in Big_Theta(g) proof let f,g be eventually-nonnegative Real_Sequence; A1: Big_Oh(f) = { t where t is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds t.n <= c*f.n & t.n >= 0 } by ASYMPT_0:def 12; A2: Big_Oh(g) = { t where t is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds t.n <= c*g.n & t.n >= 0 } by ASYMPT_0:def 12; hereby assume A3: Big_Oh(f) = Big_Oh(g); then A4: f in Big_Oh(g) by ASYMPT_0:10; g in Big_Oh(f) by A3,ASYMPT_0:10; then f in Big_Omega(g) by ASYMPT_0:19; then f in Big_Oh(g) /\ Big_Omega(g) by A4,XBOOLE_0:def 3; hence f in Big_Theta(g) by ASYMPT_0:def 14; end; assume A5: f in Big_Theta(g); now let x be set; hereby assume x in Big_Oh(f); then consider t being Element of Funcs(NAT, REAL) such that A6: x = t and A7: ex c,N st c > 0 & for n st n >= N holds t.n <= c*f.n & t.n >= 0 by A1; consider c,N such that c > 0 and A8: for n st n >= N holds t.n <= c*f.n & t.n >= 0 by A7; now take N; let n; assume n >= N; hence t.n >= 0 by A8; end; then A9: t is eventually-nonnegative by ASYMPT_0:def 4; A10: t in Big_Oh(f) by A1,A7; f in Big_Oh(g) /\ Big_Omega(g) by A5,ASYMPT_0:def 14; then f in Big_Oh(g) by XBOOLE_0:def 3; hence x in Big_Oh(g) by A6,A9,A10,ASYMPT_0:12; end; assume x in Big_Oh(g); then consider t being Element of Funcs(NAT, REAL) such that A11: x = t and A12: ex c,N st c > 0 & for n st n >= N holds t.n <= c*g.n & t.n >= 0 by A2; consider c,N such that c > 0 and A13: for n st n >= N holds t.n <= c*g.n & t.n >= 0 by A12; now take N; let n; assume n >= N; hence t.n >= 0 by A13; end; then A14: t is eventually-nonnegative by ASYMPT_0:def 4; A15: t in Big_Oh(g) by A2,A12; f in Big_Oh(g) /\ Big_Omega(g) by A5,ASYMPT_0:def 14; then f in Big_Omega(g) by XBOOLE_0:def 3; then g in Big_Oh(f) by ASYMPT_0:19; hence x in Big_Oh(f) by A11,A14,A15,ASYMPT_0:12; end; hence Big_Oh(f) = Big_Oh(g) by TARSKI:2; end; theorem :: Part 2 for f,g being eventually-nonnegative Real_Sequence holds f in Big_Theta(g) iff Big_Theta(f) = Big_Theta(g) proof let f,g be eventually-nonnegative Real_Sequence; A1: Big_Theta(f) = { s where s is Element of Funcs(NAT, REAL) : ex c,d,N st c > 0 & d > 0 & for n st n >= N holds d*f.n <= s.n & s.n <= c*f.n } by ASYMPT_0:27; A2: Big_Theta(g) = { s where s is Element of Funcs(NAT, REAL) : ex c,d,N st c > 0 & d > 0 & for n st n >= N holds d*g.n <= s.n & s.n <= c*g.n } by ASYMPT_0:27; consider N1 such that A3: for n st n >= N1 holds f.n >= 0 by ASYMPT_0:def 4; consider N2 such that A4: for n st n >= N2 holds g.n >= 0 by ASYMPT_0:def 4; hereby assume A5: f in Big_Theta(g); now let x be set; hereby assume x in Big_Theta(f); then consider s being Element of Funcs(NAT, REAL) such that A6: s = x and A7: ex c,d,N st c > 0 & d > 0 & for n st n >= N holds d*f.n <= s.n & s.n <= c*f.n by A1; consider c,d,N3 such that c > 0 and A8: d > 0 and A9: for n st n >= N3 holds d*f.n <= s.n & s.n <= c*f.n by A7; set N = max( N1, N3 ); A10: N >= N1 & N >= N3 by SQUARE_1:46; now take N; let n; assume n >= N; then A11: n >= N1 & n >= N3 by A10,AXIOMS:22; then f.n >= 0 by A3; then d*f.n >= d*0 by A8,AXIOMS:25; hence s.n >= 0 by A9,A11; end; then A12: s is eventually-nonnegative by ASYMPT_0:def 4; s in Big_Theta(f) by A1,A7; hence x in Big_Theta(g) by A5,A6,A12,ASYMPT_0:30; end; assume x in Big_Theta(g); then consider s being Element of Funcs(NAT, REAL) such that A13: s = x and A14: ex c,d,N st c > 0 & d > 0 & for n st n >= N holds d*g.n <= s.n & s.n <= c*g.n by A2; consider c,d,N3 such that c > 0 and A15: d > 0 and A16: for n st n >= N3 holds d*g.n <= s.n & s.n <= c*g.n by A14; set N = max( N2, N3 ); A17: N >= N2 & N >= N3 by SQUARE_1:46; now take N; let n; assume n >= N; then A18: n >= N2 & n >= N3 by A17,AXIOMS:22; then g.n >= 0 by A4; then d*g.n >= d*0 by A15,AXIOMS:25; hence s.n >= 0 by A16,A18; end; then A19: s is eventually-nonnegative by ASYMPT_0:def 4; A20: s in Big_Theta(g) by A2,A14; g in Big_Theta(f) by A5,ASYMPT_0:29; hence x in Big_Theta(f) by A13,A19,A20,ASYMPT_0:30; end; hence Big_Theta(f) = Big_Theta(g) by TARSKI:2; end; assume Big_Theta(f) = Big_Theta(g); hence f in Big_Theta(g) by ASYMPT_0:28; end; :: Problem 3.20 -- Proven above in theorem ASYMPT_1:????? begin :: Problem 3.21 Lm27: for n holds (n^2 - n + 1) > 0 proof defpred _P[Nat] means ($1^2 - $1 + 1) > 0; A1: _P[0] proof 0^2 -0 + 1 = 0*0 - 0 + 1 by SQUARE_1:def 3 .= 1; hence thesis; end; A2: for k st _P[k] holds _P[k+1] proof let k such that A3: (k^2 - k + 1) > 0; A4: ((k+1)^2 - (k+1) + 1) = ((k^2 + 2*k*1 + 1^2) - (k+1) + 1) by SQUARE_1:63 .= ((k^2 + 2*k + 1*1) - (k+1) + 1) by SQUARE_1:def 3 .= ((k^2 + 2*k + 1) - k - 1 + 1) by XCMPLX_1:36 .= ((k^2 + 2*k + 1) - k + -1 + 1) by XCMPLX_0:def 8 .= ((k^2 + 2*k + 1) + -k + -1 + 1) by XCMPLX_0:def 8 .= ((k^2 + 2*k + 1) + -k + (-1 + 1)) by XCMPLX_1:1 .= ((k^2 + (2*k + 1)) + -k) by XCMPLX_1:1 .= (k^2 + ((1 + 2*k) + -k)) by XCMPLX_1:1 .= (k^2 + ((1 + -k) + 2*k)) by XCMPLX_1:1 .= ((k^2 + (1 + -k)) + 2*k) by XCMPLX_1:1 .= ((k^2 + -k + 1) + 2*k) by XCMPLX_1:1 .= ((k^2 -k + 1) + 2*k) by XCMPLX_0:def 8; (k^2 - k + 1) + 2*k > 0+0 by A3,REAL_1:67; hence thesis by A4; end; for n holds _P[n] from Ind(A1, A2); hence thesis; end; Lm28: for f,g being Real_Sequence, N being Nat, c being Real st f is convergent & lim f = c & for n st n >= N holds f.n = g.n holds g is convergent & lim g = c proof let f,g be Real_Sequence, N be Nat, c be Real such that A1: f is convergent and A2: lim f = c and A3: for n st n >= N holds f.n = g.n; A4: now let p be real number; assume p > 0; then consider M such that A5: for n st n >= M holds abs(f.n-c) < p by A1,A2,SEQ_2:def 7; set N1 = max(N, M); A6: N1 >= N & N1 >= M by SQUARE_1:46; take N1; let n; assume n >= N1; then A7: n >= N & n >= M by A6,AXIOMS:22; then abs(f.n-c) < p by A5; hence abs(g.n-c) < p by A3,A7; end; hence g is convergent by SEQ_2:def 6; hence lim g = c by A4,SEQ_2:def 7; end; Lm29: for n st n >= 1 holds (n^2 -n + 1) <= n^2 proof let n such that A1: n >= 1; now assume (n^2 -n + 1) > n^2; then (n^2 + -n + 1) > n^2 by XCMPLX_0:def 8; then n^2 + (-n + 1) > n^2 by XCMPLX_1:1; then -n^2 + (n^2 + (-n + 1)) > n^2 + -n^2 by REAL_1:53; then (-n^2 + n^2) + (-n + 1) > n^2 + -n^2 by XCMPLX_1:1; then 0 + (-n + 1) > n^2 + -n^2 by XCMPLX_0:def 6; then -n + 1 > 0 by XCMPLX_0:def 6; then 1 > 0 -(-n) by REAL_1:84; then 1 > 0 + -(-n) by XCMPLX_0:def 8; hence contradiction by A1; end; hence thesis; end; Lm30: for n st n >= 1 holds n^2 <= 2*(n^2 -n + 1) proof defpred _P[Nat] means $1^2 <= 2*($1^2 -$1 + 1); A1: _P[1] proof 1^2 = 1*1 by SQUARE_1:def 3 .= 1; hence thesis; end; A2: for k st k>=1 & _P[k] holds _P[k+1] proof let k such that A3: k >= 1 and A4: k^2 <= 2*(k^2 -k + 1); A5: (k+1)^2 = k^2 + 2*k*1 + 1^2 by SQUARE_1:63 .= k^2 + 2*k + 1*1 by SQUARE_1:def 3 .= k^2 + 2*k + 1; k^2 + (2*k + 1) <= 2*(k^2 -k + 1) + (2*k + 1) by A4,AXIOMS:24; then A6: (k+1)^2 <= 2*(k^2 -k + 1) + (2*k + 1) by A5,XCMPLX_1:1; A7: 2*(k^2 -k + 1) + (2*k + 1) = 2*(k^2 + -k + 1) + (2*k + 1) by XCMPLX_0:def 8 .= (2*(k^2 + -k) + 2*1) + (2*k + 1) by XCMPLX_1:8 .= ((2*k^2 + 2*(-k)) + 2) + (2*k + 1) by XCMPLX_1:8 .= (2*k^2 + (2*(-k) + 2)) + (2*k + 1) by XCMPLX_1:1 .= (2*k^2 + (-2*k + 2)) + (2*k + 1) by XCMPLX_1:175 .= ((2*k^2 + 2) + -2*k) + (2*k + 1) by XCMPLX_1:1 .= (2*k^2 + 2) + (-2*k + (2*k + 1)) by XCMPLX_1:1 .= (2*k^2 + 2) + ((-2*k + 2*k) + 1) by XCMPLX_1:1 .= (2*k^2 + 2) + (0 + 1) by XCMPLX_0:def 6 .= 2*k^2 + (2 + 1) by XCMPLX_1:1 .= 2*k^2 + 3; 2*k >= 2*1 by A3,AXIOMS:25; then A8: 2*k + 2 >= 2 + 2 by AXIOMS:24; A9: 2*k^2 + 3 <= 2*k^2 + 4 by AXIOMS:24; 2*k^2 + 4 <= 2*k^2 + (2*k + 2) by A8,AXIOMS:24; then A10: 2*(k^2 -k + 1) + (2*k + 1) <= 2*k^2 + (2*k + 2) by A7,A9,AXIOMS:22; 2*k^2 + (2*k + 2) = 2*k^2 + (2*k + 2*1) .= 2*k^2 + (2*(k + 1)) by XCMPLX_1:8 .= 2*(k^2 + (k + 1)) by XCMPLX_1:8 .= 2*(k^2 + (2+-1)*k + 1) by XCMPLX_1:1 .= 2*(k^2 + (2*k + (-1)*k) + 1) by XCMPLX_1:8 .= 2*(((k^2 + 2*k) + (-1)*k) + 1) by XCMPLX_1:1 .= 2*((k^2 + 2*k) + (1 + -1) + ((-1)*k + 1)) by XCMPLX_1:1 .= 2*((((k^2 + 2*k) + 1) + -1) + ((-1)*k + 1)) by XCMPLX_1:1 .= 2*(((k^2 + 2*k) + 1) + (-1 + ((-1)*k + 1))) by XCMPLX_1:1 .= 2*((k^2 + 2*k*1 + 1*1) + (((-1)*k + -1) + 1)) by XCMPLX_1:1 .= 2*((k^2 + 2*k*1 + 1^2) + (((-1)*k + -1) + 1)) by SQUARE_1: def 3 .= 2*((k+1)^2 + (((-1)*k + (-1)*1) + 1)) by SQUARE_1:63 .= 2*((k+1)^2 + ((-1)*(k + 1) + 1)) by XCMPLX_1:8 .= 2*((k+1)^2 + (-(k + 1) + 1)) by XCMPLX_1:180 .= 2*((k+1)^2 + -(k + 1) + 1) by XCMPLX_1:1 .= 2*((k+1)^2 -(k + 1) + 1) by XCMPLX_0:def 8; hence thesis by A6,A10,AXIOMS:22; end; for n st n >= 1 holds _P[n] from Ind1(A1, A2); hence thesis; end; Lm31: for e being Real st 0 < e & e < 1 holds ex N st for n st n >= N holds n*log(2,1+e) - 8*log(2,n) > 8*log(2,n) proof let e be Real such that A1: 0 < e and A2: e < 1; set f = seq_logn; set g = seq_n^(e); set h = f/"g; A3: h is convergent & lim h = 0 by A1,Lm16; set d = log(2,1+e); 0+1 < e+1 by A1,REAL_1:53; then log(2,1) < log(2,e+1) by POWER:65; then A4: d > 0 by POWER:59; then d*(1/16) > d*0 by REAL_1:70; then d/16 > 0 by XCMPLX_1:100; then consider N such that A5: for n st n >= N holds abs(h.n-0) < d/16 by A3,SEQ_2:def 7; ex N st for n st n >= N holds n*log(2,1+e) - 8*log(2,n) > 8*log(2,n) proof set N1 = max(2, N); A6: N1 >= 2 & N1 >= N by SQUARE_1:46; now take N1; let n; assume A7: n >= N1; then A8: n >= 2 & n >= N by A6,AXIOMS:22; then A9: abs(h.n-0) < d/16 by A5; log(2,2) <= log(2,n) by A8,PRE_FF:12; then A10: 1 <= log(2,n) by POWER:60; A11: h.n = f.n / g.n by Lm4 .= log(2,n) / g.n by A6,A7,Def2 .= log(2,n) / (n to_power e) by A6,A7,Def3; A12: n to_power e > 0 by A6,A7,POWER:39; then (n to_power e)" > 0 by REAL_1:72; then log(2,n)*(n to_power e)" > 0*(n to_power e)" by A10,REAL_1:70; then h.n > 0 by A11,XCMPLX_0:def 9; then A13: h.n < d/16 by A9,ABSVALUE:def 1; n > 1 by A8,AXIOMS:22; then n to_power 1 > n to_power e by A2,POWER:44; then 1 / (n to_power 1) < 1 / (n to_power e) by A12,REAL_2:151; then 1 / n < 1 / (n to_power e) by POWER:30; then log(2,n) * (1 / n) < log(2,n) * (1 / (n to_power e)) by A10,REAL_1:70; then log(2,n) / n < log(2,n) * (1/(n to_power e)) by XCMPLX_1:100; then log(2,n) / n < h.n by A11,XCMPLX_1:100; then A14: log(2,n) / n < d / 16 by A13,AXIOMS:22; n" > 0 by A6,A7,REAL_1:72; then log(2,n)*n" > 0*n" by A10,REAL_1:70; then log(2,n) / n > 0 by XCMPLX_0:def 9; then 1 / (log(2,n) / n) > 1 / (d / 16) by A14,REAL_2:151; then n / log(2,n) > 1 / (d / 16) by XCMPLX_1:57; then n / log(2,n) > 16 / d by XCMPLX_1:57; then d * (n / log(2,n)) > 16 / d * d by A4,REAL_1:70; then d * (n / log(2,n)) > 16 by A4,XCMPLX_1:88; then d * (n / log(2,n)) * log(2,n) > 16*log(2,n) by A10,REAL_1:70; then d * ((n / log(2,n)) * log(2,n)) > 16*log(2,n) by XCMPLX_1:4; then d*n > (8+8)*log(2,n) by A10,XCMPLX_1:88; then d*n > 8*log(2,n) + 8*log(2,n) by XCMPLX_1:8; then d*n - 8*log(2,n) > (8*log(2,n) + 8*log(2,n)) - 8*log(2,n) by REAL_1:54; then d*n - 8*log(2,n) > (8*log(2,n) + 8*log(2,n)) + -8*log(2,n) by XCMPLX_0:def 8; then d*n - 8*log(2,n) > 8*log(2,n) + (8*log(2,n) + -8*log(2,n)) by XCMPLX_1:1; then d*n - 8*log(2,n) > 8*log(2,n) + (8*log(2,n) - 8*log(2,n)) by XCMPLX_0:def 8; then d*n - 8*log(2,n) > 8*log(2,n) + 0 by XCMPLX_1:14; hence n*d - 8*log(2,n) > 8*log(2,n); end; hence thesis; end; hence thesis; end; theorem :: (Part 1, O(nlogn) c O(n^(1+e))), slightly generalized for e being Real, f being Real_Sequence st 0 < e & (f.0 = 0 & (for n st n > 0 holds f.n = n*log(2,n))) holds ex s being eventually-positive Real_Sequence st s = f & Big_Oh(s) c= Big_Oh(seq_n^(1+e)) & not Big_Oh(s) = Big_Oh(seq_n^(1+e)) proof let e be Real, f be Real_Sequence such that A1: 0 < e and A2: (f.0 = 0 & (for n st n > 0 holds f.n = n*log(2,n))); f is eventually-positive proof take 2; let n; assume A3: n >= 2; then log(2,n) >= log(2,2) by PRE_FF:12; then log(2,n) >= 1 by POWER:60; then n*log(2,n) > n*0 by A3,REAL_1:70; hence f.n > 0 by A2,A3; end; then reconsider f as eventually-positive Real_Sequence; take f; set g = seq_n^(1+e); set h = f /" g; set seq = seq_logn; set seq1 = seq_n^(e); set p = seq /" seq1; A4: p is convergent & lim p = 0 by A1,Lm16; for n st n >= 1 holds h.n = p.n proof let n; assume A5: n >= 1; h.n = f.n / g.n by Lm4 .= n*log(2,n) / g.n by A2,A5 .= n*log(2,n) / (n to_power (1+e)) by A5,Def3 .= (n to_power 1)*log(2,n) / (n to_power (1+e)) by POWER:30 .= (n to_power 1)*log(2,n) * (n to_power (1+e))" by XCMPLX_0:def 9 .= log(2,n)*((n to_power 1)*(n to_power (1+e))") by XCMPLX_1:4 .= log(2,n)*((n to_power 1)/(n to_power (1+e))) by XCMPLX_0:def 9 .= log(2,n)*(n to_power (1-(1+e))) by A5,POWER:34 .= log(2,n)*(n to_power (1+-(1+e))) by XCMPLX_0:def 8 .= log(2,n)*(n to_power (1+(-1+-e))) by XCMPLX_1:140 .= log(2,n)*(n to_power ((1+-1)+-e)) by XCMPLX_1:1 .= log(2,n)*(1/(n to_power e)) by A5,POWER:33 .= log(2,n)/(n to_power e) by XCMPLX_1:100 .= seq.n / (n to_power e) by A5,Def2 .= seq.n / seq1.n by A5,Def3 .= p.n by Lm4; hence h.n = p.n; end; then h is convergent & lim h = 0 by A4,Lm28; then f in Big_Oh(g) & not g in Big_Oh(f) by ASYMPT_0:16; then f in Big_Oh(g) & not f in Big_Omega(g) by ASYMPT_0:19; hence thesis by Th4; end; theorem :: (Part 2, O(n^(1+e)) c O(n^2/logn)) for e being Real, g being Real_Sequence st 0 < e & e < 1 & (g.0 = 0 & g.1 = 0 & (for n st n > 1 holds g.n = (n to_power 2)/log(2,n))) holds ex s being eventually-positive Real_Sequence st s = g & Big_Oh(seq_n^(1+e)) c= Big_Oh(s) & not Big_Oh(seq_n^(1+e)) = Big_Oh(s) proof let e be Real, g be Real_Sequence such that 0 < e and A1: e < 1 and A2: (g.0 = 0 & g.1 = 0 & (for n st n > 1 holds g.n = (n to_power 2)/log(2,n))); set f = seq_n^(1+e); set h = f /" g; set seq = seq_logn; set seq1 = seq_n^(1-e); set p = seq /" seq1; g is eventually-positive proof take 2; let n; assume A3: n >= 2; then n > 1 by AXIOMS:22; then A4: g.n = (n to_power 2)/log(2,n) by A2 .= (n to_power 2)*(log(2,n))" by XCMPLX_0:def 9; A5: n to_power 2 > 0 by A3,POWER:39; log(2,n) >= log(2,2) by A3,PRE_FF:12; then log(2,n) >= 1 by POWER:60; then (log(2,n))" > 0 by REAL_1:72; then (n to_power 2)*(log(2,n))" > (n to_power 2)*0 by A5,REAL_1:70; hence g.n > 0 by A4; end; then reconsider g as eventually-positive Real_Sequence; take g; 0+e < 1 by A1; then 0 < 1-e by REAL_1:86; then A6: p is convergent & lim p = 0 by Lm16; A7: (1+e)-2 = (e+1)+-2 by XCMPLX_0:def 8 .= e+(1+-2) by XCMPLX_1:1 .= e+-1 .= e-1 by XCMPLX_0:def 8; for n st n >= 2 holds h.n = p.n proof let n; assume A8: n >= 2; then A9: n > 1 by AXIOMS:22; h.n = f.n / g.n by Lm4 .= (n to_power (1+e)) / g.n by A8,Def3 .= (n to_power (1+e)) / ((n to_power 2) / log(2,n)) by A2,A9 .= (n to_power (1+e)) * ((n to_power 2) / log(2,n))" by XCMPLX_0:def 9 .= (n to_power (1+e)) * (log(2,n) / (n to_power 2)) by XCMPLX_1:215 .= (n to_power (1+e)) * (log(2,n) * (n to_power 2)") by XCMPLX_0:def 9 .= ((n to_power (1+e)) * (n to_power 2)") * log(2,n) by XCMPLX_1:4 .= ((n to_power (1+e)) / (n to_power 2)) * log(2,n) by XCMPLX_0:def 9 .= (n to_power (e-1)) * log(2,n) by A7,A8,POWER:34 .= (n to_power -(1-e)) * log(2,n) by XCMPLX_1:143 .= log(2,n) * (1 / (n to_power (1-e))) by A8,POWER:33 .= log(2,n) / (n to_power (1-e)) by XCMPLX_1:100 .= seq.n / (n to_power (1-e)) by A8,Def2 .= seq.n / seq1.n by A8,Def3 .= p.n by Lm4; hence thesis; end; then h is convergent & lim h = 0 by A6,Lm28; then f in Big_Oh(g) & not g in Big_Oh(f) by ASYMPT_0:16; then f in Big_Oh(g) & not f in Big_Omega(g) by ASYMPT_0:19; hence thesis by Th4; end; theorem :: (Part 3, O(n^2/logn) c O(n^8)) for f being Real_Sequence st (f.0 = 0 & f.1 = 0 & (for n st n > 1 holds f.n = (n to_power 2)/log(2,n))) holds ex s being eventually-positive Real_Sequence st s = f & Big_Oh(s) c= Big_Oh(seq_n^(8)) & not Big_Oh(s) = Big_Oh(seq_n^(8)) proof let f be Real_Sequence such that A1: (f.0 = 0 & f.1 = 0 & (for n st n > 1 holds f.n = (n to_power 2)/log(2,n))); set g = seq_n^(8); set h = f/"g; f is eventually-positive proof take 2; let n; assume A2: n >= 2; then n > 1 by AXIOMS:22; then A3: f.n = (n to_power 2)/log(2,n) by A1 .= (n to_power 2)*(log(2,n))" by XCMPLX_0:def 9; A4: n to_power 2 > 0 by A2,POWER:39; log(2,n) >= log(2,2) by A2,PRE_FF:12; then log(2,n) >= 1 by POWER:60; then (log(2,n))" > 0 by REAL_1:72; then (n to_power 2)*(log(2,n))" > (n to_power 2)*0 by A4,REAL_1:70; hence f.n > 0 by A3; end; then reconsider f as eventually-positive Real_Sequence; take f; A5: now let p be real number; assume A6: p > 0; reconsider p1 = p as Real by XREAL_0:def 1; set N = max(3,[/(p1 to_power -(1/6))\]); A7: N >= 3 & N >= [/(p to_power -(1/6))\] by SQUARE_1:46; N is Integer by SQUARE_1:49; then reconsider N as Nat by A7,INT_1:16; take N; let n; assume A8: n >= N; then A9: n >= 3 & n >= [/(p to_power -(1/6))\] by A7,AXIOMS:22; then A10: log(2,n) >= log(2,3) by PRE_FF:12; log(2,3) > log(2,2) by POWER:65; then log(2,n) > log(2,2) by A10,AXIOMS:22; then A11: log(2,n) > 1 by POWER:60; A12: n > 1 by A9,AXIOMS:22; A13: n to_power 6 > 0 by A7,A8,POWER:39; A14: h.n = f.n/g.n by Lm4 .= ((n to_power 2)/log(2,n)) / g.n by A1,A12 .= ((n to_power 2)/log(2,n)) / (n to_power 8) by A7,A8,Def3 .= ((n to_power 2)*(log(2,n))") / (n to_power 8) by XCMPLX_0:def 9 .= ((log(2,n))"*(n to_power 2)) * (n to_power 8)" by XCMPLX_0:def 9 .= (log(2,n))"*((n to_power 2)*(n to_power 8)") by XCMPLX_1:4 .= (log(2,n))"*((n to_power 2)/(n to_power 8)) by XCMPLX_0:def 9 .= (log(2,n))"*(n to_power (2-8)) by A7,A8,POWER:34 .= (log(2,n))"*(n to_power -6) .= (log(2,n))"*(1/(n to_power 6)) by A7,A8,POWER:33 .= (1/(n to_power 6))*(1/log(2,n)) by XCMPLX_1:217 .= 1/((n to_power 6)*log(2,n)) by XCMPLX_1:103; (n to_power 6)*1 < (n to_power 6)*log(2,n) by A11,A13,REAL_1:70; then A15: h.n < 1/(n to_power 6) by A13,A14,REAL_2:151; A16: (p1 to_power -(1/6)) > 0 by A6,POWER:39; [/(p to_power -(1/6))\] >= (p to_power -(1/6)) by INT_1:def 5; then n >= (p to_power -(1/6)) by A9,AXIOMS:22; then n to_power 6 >= (p to_power -(1/6)) to_power 6 by A16,Lm6; then A17: n to_power 6 >= p1 to_power ((-(1/6))*6) by A6,POWER:38; p1 to_power -1 > 0 by A6,POWER:39; then 1/(n to_power 6) <= 1/(p to_power -1) by A17,REAL_2:152; then 1/(n to_power 6) <= 1/(1/(p1 to_power 1)) by A6,POWER:33; then 1/(n to_power 6) <= 1/(1/p1) by POWER:30; then 1/(n to_power 6) <= p by XCMPLX_1:56; then A18: h.n < p by A15,AXIOMS:22; (n to_power 6)*log(2,n) > (n to_power 6)*0 by A11,A13,REAL_1:70; then ((n to_power 6)*log(2,n))" > 0 by REAL_1:72; then h.n > 0 by A14,XCMPLX_1:217; hence abs(h.n-0) < p by A18,ABSVALUE:def 1; end; then A19: h is convergent by SEQ_2:def 6; then lim h = 0 by A5,SEQ_2:def 7; then f in Big_Oh(g) & not g in Big_Oh(f) by A19,ASYMPT_0:16; then f in Big_Oh(g) & not f in Big_Omega(g) by ASYMPT_0:19; hence thesis by Th4; end; theorem :: (Part 4, O(n^8) = O((n^2-n+1)^4)) for g being Real_Sequence st (for n holds g.n = (n^2 - n + 1) to_power 4) holds ex s being eventually-positive Real_Sequence st s = g & Big_Oh(seq_n^(8)) = Big_Oh(s) proof let g be Real_Sequence such that A1: (for n holds g.n = (n^2 - n + 1) to_power 4); set f = seq_n^(8); g is eventually-positive proof take 0; let n; assume n >= 0; A2: g.n = (n^2 - n + 1) to_power 4 by A1; n^2 - n + 1 > 0 by Lm27; hence g.n > 0 by A2,POWER:39; end; then reconsider g as eventually-positive Real_Sequence; take g; A3: Big_Oh(f) = { t where t is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds t.n <= c*f.n & t.n >= 0 } by ASYMPT_0:def 12; A4: Big_Oh(g) = { t where t is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds t.n <= c*g.n & t.n >= 0 } by ASYMPT_0:def 12; A5: f is Element of Funcs(NAT, REAL) by FUNCT_2:11; A6: g is Element of Funcs(NAT, REAL) by FUNCT_2:11; now let n; assume A7: n >= 1; then A8: f.n = n to_power (2*4) by Def3 .= n to_power 2 to_power 4 by A7,POWER:38 .= n^2 to_power 4 by POWER:53; A9: g.n = (n^2 -n + 1) to_power 4 by A1; A10: n^2 -n + 1 <= n^2 by A7,Lm29; A11: n^2 -n + 1 > 0 by Lm27; hence g.n <= 1*f.n by A8,A9,A10,Lm6; thus g.n >= 0 by A9,A11,POWER:39; end; then A12: g in Big_Oh(f) by A3,A6; now let n; assume A13: n >= 1; then A14: f.n = n to_power (2*4) by Def3 .= n to_power 2 to_power 4 by A13,POWER:38 .= n^2 to_power 4 by POWER:53; A15: g.n = (n^2 -n + 1) to_power 4 by A1; A16: n^2 <= 2*(n^2 -n + 1) by A13,Lm30; A17: (n^2 -n + 1) > 0 by Lm27; n*n > n*0 by A13,REAL_1:70; then A18: n^2 > 0 by SQUARE_1:def 3; then f.n <= (2*(n^2 -n + 1)) to_power 4 by A14,A16,Lm6; hence f.n <= 16*g.n by A15,A17,Lm9,POWER:35; thus f.n >= 0 by A14,A18,POWER:39; end; then f in Big_Oh(g) by A4,A5; hence thesis by A12,Lm5; end; theorem :: (Part 5, O(n^8) c O((1+e)^n)) for e being Real st 0 < e & e < 1 holds ex s being eventually-positive Real_Sequence st s = seq_a^(1+e,1,0) & Big_Oh(seq_n^(8)) c= Big_Oh(s) & not Big_Oh(seq_n^(8)) = Big_Oh(s) proof let e be Real such that A1: 0 < e and A2: e < 1; set f = seq_n^(8); set g = seq_a^(1+e,1,0); set h = f/"g; A3: 1 + e > 0 + 0 by A1,REAL_1:67; then reconsider g as eventually-positive Real_Sequence by Lm25; take g; thus g = seq_a^(1+e,1,0); consider N such that A4: for n st n >= N holds n*log(2,1+e) - 8*log(2,n) > 8*log(2,n) by A1,A2,Lm31; A5: now let p be real number such that A6: p > 0; reconsider p1 = p as Real by XREAL_0:def 1; set N1 = max( N, max([/((1/p1) to_power (1/8))\], 2) ); A7: N1 >= N & N1 >= max([/((1/p) to_power (1/8))\], 2) by SQUARE_1:46; A8: max([/((1/p) to_power (1/8))\], 2) >= [/((1/p) to_power (1/8))\] & max([/((1/p) to_power (1/8))\], 2) >= 2 by SQUARE_1:46; then A9: N1 >= [/((1/p) to_power (1/8))\] & N1 >= 2 by A7,AXIOMS:22; N1 is Integer proof per cases by SQUARE_1:49; suppose N1 = N; hence thesis; suppose N1 = max([/((1/p) to_power (1/8))\], 2); hence thesis by SQUARE_1:49; end; then reconsider N1 as Nat by A7,INT_1:16; take N1; let n; assume A10: n >= N1; p" > 0 by A6,REAL_1:72; then A11: 1/p > 0 by XCMPLX_1:217; [/((1/p) to_power (1/8))\] >= ((1/p) to_power (1/8)) by INT_1:def 5; then A12: N1 >= ((1/p) to_power (1/8)) by A9,AXIOMS:22; A13: ((1/p1) to_power (1/8)) > 0 by A11,POWER:39; A14: n >= N by A7,A10,AXIOMS:22; A15: h.n = f.n/g.n by Lm4; g.n = ((1+e) to_power (1*n + 0)) by Def1; then A16: h.n = (n to_power 8) / ((1+e) to_power n) by A7,A8,A10,A15,Def3 .= (2 to_power (8*log(2,n))) / ((1+e) to_power n) by A7,A8,A10,Lm3 .= (2 to_power (8*log(2,n))) / (2 to_power (n*log(2,1+e))) by A3,Lm3 .= (2 to_power ((8*log(2,n)) - (n*log(2,1+e)))) by POWER:34 .= (2 to_power -((n*log(2,1+e)) - (8*log(2,n)))) by XCMPLX_1:143; n >= ((1/p) to_power (1/8)) by A10,A12,AXIOMS:22; then n to_power 8 >= ((1/p) to_power (1/8)) to_power 8 by A13,Lm6; then n to_power 8 >= (1/p1) to_power ((1/8)*8) by A11,POWER:38; then n to_power 8 >= 1/p1 by POWER:30; then 1 / (n to_power 8) <= 1 / (1/p) by A11,REAL_2:152; then 1 / (n to_power 8) <= 1 / (p") by XCMPLX_1:217; then 1 / (n to_power 8) <= p by XCMPLX_1:218; then 1 / (2 to_power (8*log(2,n))) <= p by A7,A8,A10,Lm3; then A17: 2 to_power -(8*log(2,n)) <= p by POWER:33; ((n*log(2,1+e)) - (8*log(2,n))) > (8*log(2,n)) by A4,A14; then A18: 2 to_power ((n*log(2,1+e)) - (8*log(2,n))) > 2 to_power (8*log(2,n)) by POWER:44; 2 to_power (8*log(2,n)) > 0 by POWER:39; then 1 / (2 to_power ((n*log(2,1+e)) - (8*log(2,n)))) < 1 / (2 to_power (8*log(2,n))) by A18,REAL_2:151; then 2 to_power -((n*log(2,1+e)) - (8*log(2,n))) < 1 / (2 to_power (8*log(2,n))) by POWER:33; then h.n < 2 to_power -(8*log(2,n)) by A16,POWER:33; then A19: h.n < p by A17,AXIOMS:22; h.n > 0 by A16,POWER:39; hence abs(h.n-0) < p by A19,ABSVALUE:def 1; end; then A20: h is convergent by SEQ_2:def 6; then lim h = 0 by A5,SEQ_2:def 7; then f in Big_Oh(g) & not g in Big_Oh(f) by A20,ASYMPT_0:16; then f in Big_Oh(g) & not f in Big_Omega(g) by ASYMPT_0:19; hence thesis by Th4; end; begin :: Problem 3.22 Lm32: 2 to_power 12 = 4096 proof thus 2 to_power 12 = 2 to_power (6+6) .= 64*64 by Lm11,POWER:32 .= 4096; end; Lm33: for n st n >= 3 holds n^2 > 2*n + 1 proof defpred _P[Nat] means $1^2 > 2*$1 + 1; 3^2 = 3*3 by SQUARE_1:def 3 .= 9; then A1: _P[3]; A2: for n st n >= 3 & _P[n] holds _P[n+1] proof let n; assume that A3: n >= 3 and A4: n^2 > 2*n + 1; A5: (n+1)^2 = (n+1)*(n+1) by SQUARE_1:def 3 .= (n+1)*n + (n+1)*1 by XCMPLX_1:8 .= n*n + 1*n + (n+1) by XCMPLX_1:8 .= n^2 + n + (n+1) by SQUARE_1:def 3; n^2 + (n + (n+1)) > 2*n + 1 + (n + (n+1)) by A4,REAL_1:53; then (n+1)^2 > 2*n + 1 + (n + (n+1)) by A5,XCMPLX_1:1; then (n+1)^2 > 2*n + (1 + (n + (n+1))) by XCMPLX_1:1; then (n+1)^2 > 2*n + ((n+1) + (n+1)) by XCMPLX_1:1; then A6: (n+1)^2 > 2*n + 2*(n+1) by XCMPLX_1:11; n > 0 & n > 1 by A3,AXIOMS:22; then n + n > 1 + 0 by REAL_1:67; then 2*n > 1 by XCMPLX_1:11; then 2*n + 2*(n+1) > 1 + 2*(n+1) by REAL_1:53; hence (n+1)^2 > 2*(n+1) + 1 by A6,AXIOMS:22; end; for n st n >= 3 holds _P[n] from Ind1(A1, A2); hence thesis; end; Lm34: for n st n >= 10 holds 2 to_power (n-1) > (2*n)^2 proof defpred _P[Nat] means 2 to_power ($1-1) > (2*$1)^2; A1: _P[10] proof A2: 2 to_power (10-1) = 2 to_power (6+3) .= 64*(2 to_power (2+1)) by Lm11,POWER:32 .= 64*((2 to_power 2)*(2 to_power 1)) by POWER:32 .= 64*((2 to_power (1+1))*2) by POWER:30 .= 64*(((2 to_power 1)*(2 to_power 1))*2) by POWER:32 .= 64*((2*(2 to_power 1))*2) by POWER:30 .= 64*((2*2)*2) by POWER:30 .= 512; (2*10)^2 = 20*20 by SQUARE_1:def 3 .= 400; hence thesis by A2; end; A3: for n st n >= 10 & _P[n] holds _P[n+1] proof let n; assume that A4: n >= 10 and A5: 2 to_power (n-1) > (2*n)^2; 2 to_power ((n+1)-1) = 2 to_power ((n+1)+-1) by XCMPLX_0:def 8 .= 2 to_power ((n+-1)+1) by XCMPLX_1:1 .= 2 to_power ((n-1)+1) by XCMPLX_0:def 8 .= (2 to_power (n-1))*(2 to_power 1) by POWER:32 .= (2 to_power (n-1))*2 by POWER:30; then A6: 2 to_power ((n+1)-1) > ((2*n)^2)*2 by A5,REAL_1:70; A7: (2*(n+1))^2 = (2*n + 2*1)^2 by XCMPLX_1:8 .= (2*n)^2 + 2*(2*n)*2 + 2^2 by SQUARE_1:63 .= (2*n)*(2*n) + 2*(2*n)*2 + 2^2 by SQUARE_1:def 3 .= (2*n)*(2*n) + (2*2)*(2*n) + 2^2 by XCMPLX_1:4 .= (2*n)*(2*n) + (2*2)*(2*n) + 2*2 by SQUARE_1:def 3 .= ((2*n)*2)*n + (2*2)*(2*n) + 2*2 by XCMPLX_1:4 .= (n*(2*2))*n + (2*2)*(2*n) + 2*2 by XCMPLX_1:4 .= (2*2)*(n*n) + (2*2)*(2*n) + 2*2 by XCMPLX_1:4 .= (2*2)*(n*n + 2*n) + 2*2*1 by XCMPLX_1:8 .= (2*2)*(n*n + 2*n + 1) by XCMPLX_1:8; now assume ((2*n)^2)*2 <= (2*2)*(n*n + 2*n + 1); then ((2*n)*(2*n))*2 <= (2*2)*(n*n + 2*n + 1) by SQUARE_1:def 3; then (2*(n*(2*n)))*2 <= (2*2)*(n*n + 2*n + 1) by XCMPLX_1:4; then (2*(2*(n*n)))*2 <= (2*2)*(n*n + 2*n + 1) by XCMPLX_1:4; then ((2*2)*(n*n))*2 <= (2*2)*(n*n + 2*n + 1) by XCMPLX_1:4; then ((2*2)*(n*n))*2 <= (2*2)*(n*n + (2*n + 1)) by XCMPLX_1:1; then ((2*2)*(n*n))*2 <= (2*2)*(n*n) + (2*2)*(2*n + 1) by XCMPLX_1:8; then ((2*2)*(n*n))*2 - (2*2)*(n*n) <= (2*2)*(2*n + 1) by REAL_1:86; then ((2*2)*(n*n))*2 + -((2*2)*(n*n)) <= (2*2)*(2*n + 1) by XCMPLX_0:def 8; then ((2*2)*(n*n))*2 + (-1)*((2*2)*(n*n)) <= (2*2)*(2*n + 1) by XCMPLX_1:180; then ((2*2)*(n*n))*(2+-1) <= (2*2)*(2*n + 1) by XCMPLX_1:8; then (2*2)"*((2*2)*(n*n)) <= (2*2)"*((2*2)*(2*n + 1)) by AXIOMS:25; then ((2*2)"*(2*2))*(n*n) <= (2*2)"*((2*2)*(2*n + 1)) by XCMPLX_1:4; then n*n <= ((2*2)"*(2*2))*(2*n + 1) by XCMPLX_1:4; then A8: n^2 <= 2*n + 1 by SQUARE_1:def 3; n >= 3 by A4,AXIOMS:22; hence contradiction by A8,Lm33; end; hence 2 to_power ((n+1)-1) > (2*(n+1))^2 by A6,A7,AXIOMS:22; end; for n st n >= 10 holds _P[n] from Ind1(A1, A3); hence thesis; end; Lm35: for n st n >= 9 holds (n+1) to_power 6 < 2*(n to_power 6) proof defpred _P[Nat] means ($1+1) to_power 6 < 2*($1 to_power 6); A1: _P[9] proof A2: 9 to_power 2 = 9 to_power (1+1) .= (9 to_power 1)*(9 to_power 1) by POWER:32 .= 9*(9 to_power 1) by POWER:30 .= 9*9 by POWER:30 .= 81; A3: 2*(9 to_power 4) = 2*(9 to_power (2+2)) .= 2*(81*81) by A2,POWER:32 .= 13122; consider t1 being Nat such that A4: t1 = 10; consider t2 being Nat such that A5: t2 = 10*10; consider t3 being Nat such that A6: t3 = 10*10*10; consider t4 being Nat such that A7: t4 = 10*10*10*10; consider t5 being Nat such that A8: t5 = 10*10*10*10*10; consider t6 being Nat such that A9: t6 = 10*10*10*10*10*10; A10: 10 to_power 3 = 10 to_power (2+1) .= (10 to_power 2)*(10 to_power 1) by POWER:32 .= (10 to_power (1+1))*10 by POWER:30 .= ((10 to_power 1)*(10 to_power 1))*10 by POWER:32 .= (10*(10 to_power 1))*10 by POWER:30 .= (10*10)*10 by POWER:30; A11: 10 to_power 6 = 10 to_power (3+3) .= (10*10*10)*(10*10*10) by A10,POWER:32 .= (10*10*10*10)*(10*10) .= t6 by A9,XCMPLX_1:4; A12: now thus 13122*9 = ((t4 + 3*t3 + t2) + (2*t1 + 2))*9 by A4,A5,A6,A7 .= (t4 + 3*t3 + t2)*9 + (2*t1 + 2)*9 by XCMPLX_1:8 .= (t4 + (3*t3 + t2))*9 + (2*t1 + 2)*9 by XCMPLX_1:1 .= 9*t4 + (3*t3 + t2)*9 + (2*t1 + 2)*9 by XCMPLX_1:8 .= 9*t4 + (9*(3*t3) + 9*t2) + (2*t1 + 2)*9 by XCMPLX_1:8 .= 9*t4 + ((9*3)*t3 + 9*t2) + (2*t1 + 2)*9 by XCMPLX_1:4 .= 9*t4 + (2*10 + 7)*t3 + 9*t2 + (2*t1 + 2)*9 by XCMPLX_1:1 .= 9*t4 + ((2*10)*t3 + 7*t3) + 9*t2 + (2*t1 + 2)*9 by XCMPLX_1:8 .= 9*t4 + 2*t4 + 7*t3 + 9*t2 + (2*t1 + 2)*9 by A6,A7,XCMPLX_1:1 .= (9+2)*t4 + 7*t3 + 9*t2 + (2*t1 + 2)*9 by XCMPLX_1:8 .= (10 + 1)*t4 + 7*t3 + 9*t2 + (2*t1 + 2)*9 .= 10*t4 + 1*t4 + 7*t3 + 9*t2 + (2*t1 + 2)*9 by XCMPLX_1:8 .= (t5 + t4 + 7*t3) + (9*t2 + (2*t1 + 2)*9) by A7,A8,XCMPLX_1:1 .= t5 + t4 + 7*t3 + (t3 + (8*t1 + 2*9)) by A4,A5,A6 .= (t5 + t4 + 7*t3 + t3) + (8*t1 + 2*9) by XCMPLX_1:1 .= (t5 + t4 + (7*t3 + 1*t3)) + (8*t1 + 2*9) by XCMPLX_1:1 .= (t5 + t4 + ((7+1)*t3)) + (8*t1 + 2*9) by XCMPLX_1:8 .= (t5 + t4 + 8*t3) + ((8+1)*t1 + 8) by A4 .= t5 + t4 + 8*t3 + 9*t1 + 8 by XCMPLX_1:1; end; now thus (13122*9)*9 = ((t5 + t4 + 8*t3) + (9*t1 + 8))*9 by A12,XCMPLX_1 :1 .= (t5 + t4 + 8*t3)*9 + (9*t1 +8 )*9 by XCMPLX_1:8 .= (t5 + (t4 + 8*t3))*9 + (9*t1 + 8)*9 by XCMPLX_1:1 .= (9*t5 + (t4 + 8*t3)*9) + (9*t1 + 8)*9 by XCMPLX_1:8 .= (9*t5 + (9*t4 + 9*(8*t3))) + (9*t1 + 8)*9 by XCMPLX_1:8 .= (9*t5 + (9*t4 + (9*8)*t3)) + (9*t1 + 8)*9 by XCMPLX_1:4 .= (9*t5 + (9*t4 + (70+2)*t3)) + (9*t1 + 8)*9 .= (9*t5 + (9*t4 + ((7*10)*t3 + 2*t3))) + (9*t1 + 8)*9 by XCMPLX_1:8 .= (9*t5 + ((9*t4 + 7*t4) + 2*t3)) + (9*t1 + 8)*9 by A6,A7,XCMPLX_1:1 .= (9*t5 + (((9+7)*t4) + 2*t3)) + (9*t1 + 8)*9 by XCMPLX_1:8 .= ((9*t5 + (10+6)*t4) + 2*t3) + (9*t1 + 8)*9 by XCMPLX_1:1 .= ((9*t5 + (t5 + 6*t4)) + 2*t3) + (9*t1 + 8)*9 by A7,A8,XCMPLX_1:8; end; then A13: (13122*9)*9 = (((9*t5 + 1*t5) + 6*t4) + 2*t3) + (9*t1 + 8)*9 by XCMPLX_1:1 .= ((((9+1)*t5) + 6*t4) + 2*t3) + (9*t1 + 8)*9 by XCMPLX_1:8 .= ((t6 + 6*t4) + 2*t3) + (8*t2 + (8*t1 + 2)) by A4,A5,A8,A9 .= (t6 + 6*t4 + 2*t3 + 8*t2) + (8*t1 + 2) by XCMPLX_1:1 .= t6 + 6*t4 + 2*t3 + 8*t2 + 8*t1 + 2 by XCMPLX_1:1; t6 < t6 + ((6*t4 + 2*t3 + 8*t2) + (8*t1 + 2)) by REAL_1:69; then t6 < (t6 + (6*t4 + 2*t3 + 8*t2)) + (8*t1 + 2) by XCMPLX_1:1; then t6 < (t6 + (6*t4 + (2*t3 + 8*t2))) + (8*t1 + 2) by XCMPLX_1:1; then t6 < ((t6 + 6*t4) + (2*t3 + 8*t2)) + (8*t1 + 2) by XCMPLX_1:1; then A14: t6 < (t6 + 6*t4 + 2*t3 + 8*t2) + (8*t1 + 2) by XCMPLX_1:1; (13122*9)*9 = (2*((9 to_power 4)*9))*9 by A3,XCMPLX_1:4 .= (2*((9 to_power 4)*(9 to_power 1)))*9 by POWER:30 .= (2*(9 to_power (4+1)))*9 by POWER:32 .= 2*((9 to_power 5)*9) by XCMPLX_1:4 .= 2*((9 to_power 5)*(9 to_power 1)) by POWER:30 .= 2*(9 to_power (5+1)) by POWER:32 .= 2*(9 to_power 6); hence thesis by A11,A13,A14,XCMPLX_1:1; end; A15: for n st n >= 9 & _P[n] holds _P[n+1] proof let n; assume that A16: n >= 9 and A17: (n+1) to_power 6 < 2*(n to_power 6); A18: n to_power 6 > 0 by A16,POWER:39; A19: (n+1) to_power 6 > 0 by POWER:39; (n+1)" > 0 by REAL_1:72; then (n+2)*((n+1)") > 0*((n+1)") by REAL_1:70; then A20: ((n+2) / (n+1)) > 0 by XCMPLX_0:def 9; now assume (n+2) / (n+1) >= (n+1) / n; then (n+2) / (n+1) * (n+1) >= (n+1) / n * (n+1) by AXIOMS:25; then (n+2) >= (n+1) / n * (n+1) by XCMPLX_1:88; then (n+2)*n >= (n+1) / n * (n+1) * n by AXIOMS:25; then (n+2)*n >= ((n+1) / n * n) * (n+1) by XCMPLX_1:4; then (n+2)*n >= (n+1)*(n+1) by A16,XCMPLX_1:88; then n*n + 2*n >= (n+1)*(n+1) by XCMPLX_1:8; then n^2 + 2*n >= (n+1)*(n+1) by SQUARE_1:def 3; then n^2 + 2*n >= (n+1)^2 by SQUARE_1:def 3; then n^2 + 2*n >= n^2 + 2*n*1 + 1^2 by SQUARE_1:63; then n^2 + 2*n >= n^2 + 2*n + 1*1 by SQUARE_1:def 3; then (n^2 + 2*n) - (n^2 + 2*n) >= 1 by REAL_1:84; hence contradiction by XCMPLX_1:14; end; then A21: ((n+2) / (n+1)) to_power 6 < ((n+1) / n) to_power 6 by A20,POWER :42; (n+1) to_power 6 / n to_power 6 < 2 by A17,A18,REAL_2:178; then ((n+1) / n) to_power 6 < 2 by A16,POWER:36; then ((n+2) / (n+1)) to_power 6 < 2 by A21,AXIOMS:22; then ((n+2) to_power 6) / ((n+1) to_power 6) < 2 by POWER:36; then ((n+2) to_power 6) / ((n+1) to_power 6) * ((n+1) to_power 6) < 2*((n+1) to_power 6) by A19,REAL_1:70; then (n+(1+1)) to_power 6 < 2*((n+1) to_power 6) by A19,XCMPLX_1:88; hence thesis by XCMPLX_1:1; end; for n st n >= 9 holds _P[n] from Ind1(A1, A15); hence thesis; end; Lm36: for n st n >= 30 holds 2 to_power n > n to_power 6 proof defpred _P[Nat] means 2 to_power $1 > $1 to_power 6; A1: _P[30] proof 2 to_power 30 = 2 to_power (5*6) .= 32 to_power 6 by Lm10,POWER:38; hence thesis by POWER:42; end; A2: for n st n >= 30 & _P[n] holds _P[n+1] proof let n; assume that A3: n >= 30 and A4: 2 to_power n > n to_power 6; A5: 2 to_power (n+1) = (2 to_power n)*(2 to_power 1) by POWER:32 .= (2 to_power n)*2 by POWER:30; A6: (2 to_power n)*2 > (n to_power 6)*2 by A4,REAL_1:70; n >= 9 by A3,AXIOMS:22; then (n+1) to_power 6 < 2*(n to_power 6) by Lm35; hence thesis by A5,A6,AXIOMS:22; end; for n st n >= 30 holds _P[n] from Ind1(A1, A2); hence thesis; end; Lm37: for x being Real st x > 9 holds 2 to_power x > (2*x)^2 proof let x be Real such that A1: x > 9; set n = [/x\]; A2: n >= x by INT_1:def 5; then n >= 9 by A1,AXIOMS:22; then reconsider n as Nat by INT_1:16; x >= [\x/] by INT_1:def 4; then A3: 2 to_power x >= 2 to_power [\x/] by PRE_FF:10; [/x\] - [\x/] <= 1 proof per cases; suppose x is Integer; then [\x/] = [/x\] by INT_1:59; hence thesis by XCMPLX_1:14; suppose not x is Integer; then not [\x/] = [/x\] by INT_1:59; then [\x/] + 1 = [/x\] by INT_1:66; hence thesis by XCMPLX_1:26; end; then [/x\] <= 1 + [\x/] by REAL_1:86; then [\x/] >= n-1 by REAL_1:86; then 2 to_power [\x/] >= 2 to_power (n-1) by PRE_FF:10; then A4: 2 to_power x >= 2 to_power (n-1) by A3,AXIOMS:22; n > 9 by A1,A2,AXIOMS:22; then n >= 9+1 by NAT_1:38; then 2 to_power (n-1) > (2*n)^2 by Lm34; then A5: 2 to_power x > (2*n)^2 by A4,AXIOMS:22; A6: 2*n >= 2*x by A2,AXIOMS:25; x >= 0 by A1,AXIOMS:22; then 2*x >= 0*x by AXIOMS:25; then (2*n)*(2*n) >= (2*x)*(2*x) by A6,Lm26; then (2*n)^2 >= (2*x)*(2*x) by SQUARE_1:def 3; then (2*n)^2 >= (2*x)^2 by SQUARE_1:def 3; hence 2 to_power x > (2*x)^2 by A5,AXIOMS:22; end; Lm38: ex N st for n st n >= N holds sqrt n - log(2, n) > 1 proof ex N st for n st n >= N holds n/6 > 1 proof now take N = 7; let n; assume n >= N; then n > 6 by AXIOMS:22; then n/6 > 6/6 by REAL_1:73; hence n/6 > 1; end; hence thesis; end; then consider N1 such that A1: for n st n >= N1 holds n/6 > 1; ex N st for n st n >= N holds n/3 > 2*log(2,n) proof now take N = 30; let n; assume A2: n >= N; then A3: 2 to_power n > n to_power 6 by Lm36; n to_power 6 > 0 by A2,POWER:39; then log(2,2 to_power n) > log(2,n to_power 6) by A3,POWER:65; then n*log(2,2) > log(2,n to_power 6) by POWER:63; then n*1 > log(2,n to_power 6) by POWER:60; then n > (3*2)*log(2,n) by A2,POWER:63; then n > 3*(2*log(2,n)) by XCMPLX_1:4; hence n/3 > 2*log(2,n) by REAL_2:178; end; hence thesis; end; then consider N2 such that A4: for n st n >= N2 holds n/3 > 2*log(2,n); ex N st for n st n >= N holds n/2 > log(2,n)*log(2,n) proof reconsider N = 2 to_power 10 as Nat; now take N; let n; assume A5: n >= N; set x = log(2,n); A6: n > 0 by A5,POWER:39; A7: 2 to_power 9 > 0 by POWER:39; 2 to_power 10 > 2 to_power 9 by POWER:44; then n > 2 to_power 9 by A5,AXIOMS:22; then log(2,n) > log(2,2 to_power 9) by A7,POWER:65; then x > 9*log(2,2) by POWER:63; then A8: x > 9*1 by POWER:60; then A9: 2*x > 0*x by REAL_1:70; then (2*x)*(2*x) > 0*(2*x) by REAL_1:70; then A10: (2*x)^2 > 0 by SQUARE_1:def 3; 2 to_power x > (2*x)^2 by A8,Lm37; then log(2, 2 to_power x) > log(2,(2*x)^2) by A10,POWER:65; then x*log(2,2) > log(2,(2*x)^2) by POWER:63; then x*1 > log(2,(2*x)^2) by POWER:60; then x > log(2,(2*x) to_power 2) by POWER:53; then x > 2*log(2,2*x) by A9,POWER:63; then log(2,n) > 2*log(2,log(2,n to_power 2)) by A6,POWER:63; then A11: log(2,n) > 2*log(2,log(2,n^2)) by POWER:53; 2 to_power 10 > 2 to_power 1 by POWER:44; then n > 2 to_power 1 by A5,AXIOMS:22; then A12: n > 2 by POWER:30; then A13: n*n > 2*n by REAL_1:70; 2*n > 2*2 by A12,REAL_1:70; then n*n > 2*2 by A13,AXIOMS:22; then n^2 > 2*2 by SQUARE_1:def 3; then log(2,n^2) > log(2,2*2) by POWER:65; then log(2,n^2) > log(2,2^2) by SQUARE_1:def 3; then log(2,n^2) > log(2,2 to_power 2) by POWER:53; then log(2,n^2) > 2*log(2,2) by POWER:63; then A14: log(2,n^2) > 2*1 by POWER:60; 2 to_power log(2,n) > 2 to_power (2*log(2,log(2,n^2))) by A11,POWER:44; then n > 2 to_power (2*log(2,log(2,n^2))) by A6,POWER:def 3; then n > 2 to_power (log(2,log(2,n^2) to_power 2)) by A14,POWER:63; then A15: n > 2 to_power (log(2,(log(2,n^2))^2)) by POWER:53; (log(2,n^2))^2 > 0 by A14,SQUARE_1:74; then n > (log(2,n^2))^2 by A15,POWER:def 3; then n > (log(2,n to_power 2))^2 by POWER:53; then n > (2*log(2,n))^2 by A6,POWER:63; then n > (2*log(2,n))*(2*log(2,n)) by SQUARE_1:def 3; then n > (2*(2*log(2,n)))*log(2,n) by XCMPLX_1:4; then n > ((2*2)*log(2,n))*log(2,n) by XCMPLX_1:4; then A16: n > 4*(log(2,n)*log(2,n)) by XCMPLX_1:4; 2 to_power 10 > 2 to_power 0 by POWER:44; then n > 2 to_power 0 by A5,AXIOMS:22; then n > 1 by POWER:29; then log(2,n) > log(2,1) by POWER:65; then log(2,n) > 0 by POWER:59; then log(2,n)*log(2,n) > 0*log(2,n) by REAL_1:70; then 4*(log(2,n)*log(2,n)) > 2*(log(2,n)*log(2,n)) by REAL_1:70; then n > 2*(log(2,n)*log(2,n)) by A16,AXIOMS:22; hence n/2 > log(2,n)*log(2,n) by REAL_2:178; end; hence thesis; end; then consider N3 such that A17: for n st n >= N3 holds n/2 > log(2,n)*log(2,n); set N = max( max(N1, 2), max( N2, N3 ) ); A18: N >= max(N1, 2) & N >= max( N2, N3 ) by SQUARE_1:46; A19: max( N2, N3 ) >= N2 & max( N2, N3 ) >= N3 by SQUARE_1:46; max(N1, 2) >= N1 & max(N1, 2) >= 2 by SQUARE_1:46; then A20: N >= N1 & N >= N2 & N >= N3 & N >= 2 by A18,A19,AXIOMS:22; now let n; assume n >= N; then A21: n >= N1 & n >= N2 & n >= N3 & n >= 2 by A20,AXIOMS:22; then A22: n/6 > 1 by A1; n/3 > 2*log(2,n) by A4,A21; then A23: n/6 + n/3 > 1 + 2*log(2,n) by A22,REAL_1:67; A24: n/2 > log(2,n)*log(2,n) by A17,A21; n/3 = (2*n) / (2*3) by XCMPLX_1:92 .= 2*n / 6; then n/6 + n/3 = (1*n + 2*n) / 6 by XCMPLX_1:63 .= ((1 + 2)*n) / 6 by XCMPLX_1:8 .= 3*n / 6; then A25: (n/6 + n/3) + n/2 = 3*n / 6 + (3*n) / (3*2) by XCMPLX_1:92 .= (3*n + 3*n) / 6 by XCMPLX_1:63 .= ((3 + 3)*n) / 6 by XCMPLX_1:8 .= (6*n) / (6*1) .= n / 1 by XCMPLX_1:92 .= n; (1 + 2*log(2,n)) + log(2,n)*log(2,n) = (1 + (1+1)*log(2,n)) + log(2,n)*log(2,n) .= (1 + (1*log(2,n) + 1*log(2,n))) + log(2,n)*log(2,n) by XCMPLX_1:8 .= (1 + 1*log(2,n) + 1*log(2,n)) + log(2,n)*log(2,n) by XCMPLX_1:1 .= (1 + 1*log(2,n)) + (1*log(2,n) + log(2,n)*log(2,n)) by XCMPLX_1:1 .= 1*(1 + log(2,n)) + log(2,n)*(1 + log(2,n)) by XCMPLX_1:8 .= (1 + log(2,n))*(1 + log(2,n)) by XCMPLX_1:8 .= (1 + log(2,n))^2 by SQUARE_1:def 3; then A26: n > (1 + log(2,n))^2 by A23,A24,A25,REAL_1:67; log(2,n) >= log(2,2) by A21,PRE_FF:12; then log(2,n) >= 1 by POWER:60; then A27: 1 + log(2,n) >= 1 + 0 by REAL_1:55; (1 + log(2,n))^2 >= 0 by SQUARE_1:72; then sqrt n > sqrt((1 + log(2,n))^2) by A26,SQUARE_1:95; then sqrt n > 1 + log(2,n) by A27,SQUARE_1:89; hence sqrt n - log(2,n) > 1 by REAL_1:86; end; hence thesis; end; Lm39: (4+1)! = 120 proof (4+1)! = (4+1)*(4!) by NEWTON:21 .= 5*((3+1)*(3!)) by NEWTON:21 .= 5*(4*((2+1)*(2!))) by NEWTON:21 .= 120 by NEWTON:20; hence thesis; end; Lm40: for n st n >= 10 holds 2 to_power (2*n) / (n!) < 1/(2 to_power (n-9)) proof defpred _P[Nat] means 2 to_power (2*$1) / ($1!) < 1/(2 to_power ($1-9)); A1: 7 = 8-1; A2: _P[10] proof A3: 2 to_power (2*10) / (10!) = 2 to_power 20 / ((9+1)*(9!)) by NEWTON:21 .= 2 to_power (1+19) / (10*(9!)) .= ((2 to_power 1)*(2 to_power 19)) / (10*(9!)) by POWER:32 .= (2*(2 to_power 19)) / ((2*5)*(9!)) by POWER:30 .= (2*(2 to_power 19)) / (2*(5*(9!))) by XCMPLX_1:4 .= (2 to_power 19) / (5*(9!)) by XCMPLX_1:92 .= (2 to_power 19) / (5*((8+1)*(8!))) by NEWTON:21 .= (2 to_power 19) / ((5*9)*(8!)) by XCMPLX_1:4 .= (2 to_power 19) / (45*((7+1)*(7!))) by NEWTON:21 .= (2 to_power (3+16)) / (8*(45*(7!))) by XCMPLX_1:4 .= (8*(2 to_power 16)) / (8*(45*(7!))) by Lm8,POWER:32 .= (2 to_power 16) / (45*(7!)) by XCMPLX_1:92 .= (2 to_power (4+12)) / (45*((6+1)*(6!))) by NEWTON:21 .= ((2 to_power (3+1))*4096) / (45*((6+1)*(6!))) by Lm32,POWER:32 .= ((8*(2 to_power 1))*4096) / (45*((6+1)*(6!))) by Lm8,POWER:32 .= ((8*2)*4096) / (45*((6+1)*(6!))) by POWER:30 .= (16*4096) / ((45*7)*(6!)) by XCMPLX_1:4 .= (16*4096) / (315*((5+1)*(5!))) by NEWTON:21 .= (16*4096) / ((315*45)*16) by Lm39 .= 4096 / 14175 by XCMPLX_1:92; not 4096 / 14175 >= 1 / 2; hence thesis by A3,POWER:30; end; A4: for k st k >= 10 & _P[k] holds _P[k+1] proof let k; assume that A5: k >= 10 and A6: 2 to_power (2*k) / (k!) < 1/(2 to_power (k-9)); A7: (2 to_power 2) > 0 by POWER:39; (k+1)" > 0 by REAL_1:72; then (2 to_power 2)*(k+1)" > 0*(k+1)" by A7,REAL_1:70; then A8: ((2 to_power 2)/(k+1)) > 0 by XCMPLX_0:def 9; 2 to_power -(k-9) > 0 by POWER:39; then A9: (1/(2 to_power (k-9))) > 0 by POWER:33; 2 to_power (2*(k+1)) / ((k+1)!) = (2 to_power (2*k + 2*1)) / ((k+1)!) by XCMPLX_1:8 .= ((2 to_power (2*k))*(2 to_power 2)) / ((k+1)!) by POWER:32 .= ((2 to_power (2*k))*(2 to_power 2)) / ((k+1)*(k!)) by NEWTON:21 .= ((2 to_power 2)/(k+1))* ((2 to_power (2*k))/(k!)) by XCMPLX_1:77; then A10: 2 to_power (2*(k+1)) / ((k+1)!) < ((2 to_power 2)/(k+1))*(1/(2 to_power (k-9))) by A6,A8,REAL_1:70; A11: 2 to_power 1 > 0 by POWER:39; now assume ((2 to_power 2)/(k+1)) >= (1/(2 to_power 1)); then (2 to_power 1)*((2 to_power 2)/(k+1)) >= (1/(2 to_power 1))*(2 to_power 1) by AXIOMS:25; then (2 to_power 1)*((2 to_power 2)*(k+1)") >= (1/(2 to_power 1))*(2 to_power 1) by XCMPLX_0:def 9; then ((2 to_power 1)*(2 to_power 2))*(k+1)" >= (1/(2 to_power 1))*(2 to_power 1) by XCMPLX_1:4; then (2 to_power (1+2))*(k+1)" >= (1/(2 to_power 1))*(2 to_power 1) by POWER:32; then 8/(k+1) >= (1/(2 to_power 1))*(2 to_power 1) by Lm8,XCMPLX_0:def 9; then 8/(k+1) >= 1 by A11,XCMPLX_1:107; then 8/(k+1)*(k+1) >= 1*(k+1) by AXIOMS:25; then 8 >= k+1 by XCMPLX_1:88; then 7 >= k by A1,REAL_1:84; hence contradiction by A5,AXIOMS:22; end; then ((2 to_power 2)/(k+1))*(1/(2 to_power (k-9))) < (1/(2 to_power 1))*(1/(2 to_power (k-9))) by A9,REAL_1:70; then 2 to_power (2*(k+1)) / ((k+1)!) < (1/(2 to_power 1))*(1/(2 to_power (k-9))) by A10,AXIOMS:22; then 2 to_power (2*(k+1)) / ((k+1)!) < (1/((2 to_power 1)*(2 to_power (k-9)))) by XCMPLX_1:103; then 2 to_power (2*(k+1)) / ((k+1)!) < (1/(2 to_power (1+(k-9)))) by POWER:32; then 2 to_power (2*(k+1)) / ((k+1)!) < (1/(2 to_power (1+(k+-9)))) by XCMPLX_0:def 8; then 2 to_power (2*(k+1)) / ((k+1)!) < (1/(2 to_power ((1+k)+-9))) by XCMPLX_1:1; hence 2 to_power (2*(k+1)) / ((k+1)!) < (1/(2 to_power ((k+1)-9))) by XCMPLX_0:def 8; end; for n st n >= 10 holds _P[n] from Ind1(A2, A4); hence thesis; end; Lm41: for n st n >= 3 holds 2*(n-2) >= n-1 proof defpred _P[Nat] means 2*($1-2) >= $1-1; A1: _P[3]; A2: for n st n >= 3 & _P[n] holds _P[n+1] proof let n; assume that n >= 3 and A3: 2*(n-2) >= n-1; A4: 2*((n+1)-2) = 2*((n+1)+-2) by XCMPLX_0:def 8 .= 2*(n+1) + 2*(-2) by XCMPLX_1:8 .= 2*n + 2*1 + 2*(-2) by XCMPLX_1:8 .= 2*n + 2*(-2) + 2 by XCMPLX_1:1 .= 2*(n + -2) + 2 by XCMPLX_1:8 .= 2*(n-2) + 2 by XCMPLX_0:def 8; 2*(n-2) + 2 >= (n-1) + 1 by A3,REAL_1:55; then 2*(n-2) + 2 >= (n+-1) +1 by XCMPLX_0:def 8; then 2*(n-2) + 2 >= (n + 1) + -1 by XCMPLX_1:1; hence 2*((n+1)-2) >= (n+1)-1 by A4,XCMPLX_0:def 8; end; for n st n >= 3 holds _P[n] from Ind1(A1, A2); hence thesis; end; Lm42: 5 to_power 5 = 3125 proof 5 to_power 5 = 5 to_power (4+1) .= (5 to_power 4)*(5 to_power 1) by POWER:32 .= (5 to_power (3+1))*5 by POWER:30 .= ((5 to_power 3)*(5 to_power 1))*5 by POWER:32 .= ((5 to_power (2+1))*5)*5 by POWER:30 .= (((5 to_power 2)*(5 to_power 1))*5)*5 by POWER:32 .= (((5 to_power (1+1))*5)*5)*5 by POWER:30 .= ((((5 to_power 1)*(5 to_power 1))*5)*5)*5 by POWER:32 .= ((((5 to_power 1)*5)*5)*5)*5 by POWER:30 .= (((5*5)*5)*5)*5 by POWER:30 .= 3125; hence thesis; end; Lm43: 4 to_power 4 = 256 proof 4 to_power 4 = 4 to_power (3+1) .= (4 to_power 3)*(4 to_power 1) by POWER:32 .= (4 to_power (2+1))*4 by POWER:30 .= ((4 to_power 2)*(4 to_power 1))*4 by POWER:32 .= ((4 to_power (1+1))*4)*4 by POWER:30 .= (((4 to_power 1)*(4 to_power 1))*4)*4 by POWER:32 .= (((4 to_power 1)*4)*4)*4 by POWER:30 .= ((4*4)*4)*4 by POWER:30 .= 256; hence thesis; end; Lm44: for a,b,d,e being Real holds (a/b) / (d/e) = (a/d) * (e/b) proof let a,b,d,e be Real; thus (a/b) / (d/e) = (a*e)/(b*d) by XCMPLX_1:85 .= (a/d) * (e/b) by XCMPLX_1:77; end; Lm45: for c being Real st c >= 0 holds (c to_power (1/2)) = sqrt c proof let c be Real such that A1: c >= 0; per cases by A1; suppose c = 0; hence thesis by POWER:def 2,SQUARE_1:82; suppose A2: c > 0; then A3: (c to_power (1/2)) > 0 by POWER:39; A4: sqrt c > 0 by A2,SQUARE_1:93; now assume A5: c to_power (1/2) <> sqrt c; thus contradiction proof per cases by A5,REAL_1:def 5; suppose (c to_power (1/2)) < sqrt c; then (c to_power (1/2))^2 < (sqrt c)^2 by A3,SQUARE_1:78; then (c to_power (1/2))^2 < c by A2,SQUARE_1:def 4; then (c to_power (1/2))*(c to_power (1/2)) < c by SQUARE_1:def 3; then (c to_power (1/2 + 1/2)) < c by A2,POWER:32; hence thesis by POWER:30; suppose (c to_power (1/2)) > sqrt c; then (c to_power (1/2))^2 > (sqrt c)^2 by A4,SQUARE_1:78; then (c to_power (1/2))^2 > c by A2,SQUARE_1:def 4; then (c to_power (1/2))*(c to_power (1/2)) > c by SQUARE_1:def 3; then (c to_power (1/2 + 1/2)) > c by A2,POWER:32; hence thesis by POWER:30; end; end; hence thesis; end; Lm46: ex N st for n st n >= N holds n - sqrt n*log(2,n) > n/2 proof set seq = seq_logn; set seq1 = seq_n^(1/2); set p = seq /" seq1; p is convergent & lim p = 0 by Lm16; then consider N such that A1: for n st n >= N holds abs(p.n-0) < 1/2 by SEQ_2:def 7; set N1 = max(2, N); A2: N1 >= 2 & N1 >= N by SQUARE_1:46; now let n; assume A3: n >= N1; then n >= 2 & n >= N by A2,AXIOMS:22; then A4: abs(p.n-0) < 1/2 by A1; A5: p.n = seq.n / seq1.n by Lm4 .= log(2,n) / seq1.n by A2,A3,Def2 .= log(2,n) / (n to_power (1/2)) by A2,A3,Def3 .= log(2,n) / sqrt n by Lm45; A6: sqrt n > 0 by A2,A3,SQUARE_1:93; A7: p.n < 1/2 by A4,ABSVALUE:def 1; A8: sqrt n <> 0 by A2,A3,SQUARE_1:93; log(2,n) / sqrt n * sqrt n < sqrt n * (1/2) by A5,A6,A7,REAL_1:70; then log(2,n) < sqrt n * (1/2) by A8,XCMPLX_1:88; then sqrt n*log(2,n) < sqrt n * (sqrt n * (1/2)) by A6,REAL_1:70; then sqrt n*log(2,n) < (sqrt n * sqrt n) * (1/2) by XCMPLX_1:4; then sqrt n*log(2,n) < ((sqrt n)^2) * (1/2) by SQUARE_1:def 3; then sqrt n*log(2,n) < n * (1/2) by SQUARE_1:def 4; then sqrt n*log(2,n) < n/2 by XCMPLX_1:100; then n/2 + sqrt n*log(2,n) < n/2 + n/2 by REAL_1:53; then n/2 + sqrt n*log(2,n) < (n+n)/2 by XCMPLX_1:63; then n/2 + sqrt n*log(2,n) < n by XCMPLX_1:65; hence n/2 < n - sqrt n*log(2,n) by REAL_1:86; end; hence thesis; end; Lm47: for s being Real_Sequence st for n holds s.n = (1 + 1/(n+1)) to_power (n+1) holds s is non-decreasing proof let s be Real_Sequence such that A1: for n holds s.n = (1 + 1/(n+1)) to_power (n+1); now let n; A2: (1+1/(n+1)) to_power (n+1) > 0 by POWER:39; A3: s.(n+1)/s.n =(1 + 1/(n+1+1)) to_power (n+1+1) / s.n by A1 .=(1 + 1/(n+1+1)) to_power (n+1+1)/(1+1/(n+1)) to_power (n+1)*1 by A1 .=(1 + 1/(n+1+1)) to_power (n+1+1)/(1 + 1/(n+1)) to_power (n+1) * ((1+1/(n+1))/(1+1/(n+1))) by XCMPLX_1:60 .=(1+1/(n+1)) * (1 + 1/(n+1+1)) to_power (n+1+1) / ((1 + 1/(n+1)) to_power (n+1) * (1+1/(n+1))) by XCMPLX_1:77 .=(1+1/(n+1)) * (1 + 1/(n+1+1)) to_power (n+1+1) / ((1 + 1/(n+1)) to_power (n+1) * (1+1/(n+1)) to_power 1) by POWER:30 .=(1+1/(n+1)) * (1 + 1/(n+1+1)) to_power (n+1+1) / (1 + 1/(n+1)) to_power (n+1+1) by POWER:32 .=(1+1/(n+1)) * ((1 + 1/(n+1+1)) to_power (n+1+1) / (1 + 1/(n+1)) to_power (n+1+1)) by XCMPLX_1:75 .=(1+1/(n+1)) * ((1 + 1/(n+1+1))/(1 + 1/(n+1))) to_power (n+1+1) by POWER:36; A4: (n+2)*(n+2)<>0 by XCMPLX_1:6; A5: (n+1)*(n+2)<>0 by XCMPLX_1:6; A6: (1 + 1/(n+1+1))/(1 + 1/(n+1)) = ((1*(n+1+1) + 1)/(n+1+1))/(1 + 1/(n+1)) by XCMPLX_1:114 .= ((n+1+1+1)/(n+1+1))/((1*(n+1) + 1)/(n+1)) by XCMPLX_1:114 .= ((n+1+1+1)*(n+1))/((n+1+1)*(n+1+1)) by XCMPLX_1:85 .= ((n+1+1+1)*(n+1))/((n+(1+1))*(n+1+1)) by XCMPLX_1:1 .= ((n+1+1+1)*(n+1))/((n+(1+1))*(n+(1+1))) by XCMPLX_1:1 .= ((n+(1+1)+1)*(n+1))/((n+2)*(n+2)) by XCMPLX_1:1 .= ((n+(2+1))*(n+1))/((n+2)*(n+2)) by XCMPLX_1:1 .= (n*n+n*1+3*n+3*1)/((n+2)*(n+2)) by XCMPLX_1:10 .= (n*n+(1*n+3*n)+3)/((n+2)*(n+2)) by XCMPLX_1:1 .= (n*n+(1+3)*n+3)/((n+2)*(n+2)) by XCMPLX_1:8 .= (n*n+(2+2)*n+3)/((n+2)*(n+2)) .= (n*n+(n*2+2*n)+3)/((n+2)*(n+2)) by XCMPLX_1:8 .= (n*n+n*2+2*n+3)/((n+2)*(n+2)) by XCMPLX_1:1 .= (n*n+n*2+2*n+3+1-1)/((n+2)*(n+2)) by XCMPLX_1:26 .= (n*n+n*2+2*n+(3+1)-1)/((n+2)*(n+2)) by XCMPLX_1:1 .= (n*n+n*2+2*n+2*2-1)/((n+2)*(n+2)) .= ((n+2)*(n+2)-1)/((n+2)*(n+2)) by XCMPLX_1:10 .= ((n+2)*(n+2))/((n+2)*(n+2)) - 1/((n+2)*(n+2)) by XCMPLX_1:121 .= 1 - 1/((n+2)*(n+2)) by A4,XCMPLX_1:60; n+1+1>0+1 by REAL_1:53; then n+(1+1)>0+1 by XCMPLX_1:1; then (n+2)*(n+2)>1 by REAL_2:137; then 1/((n+2)*(n+2))<1 by SQUARE_1:2; then - 1/((n+2)*(n+2)) > -1 by REAL_1:50; then (1 + -1/((n+2)*(n+2))) to_power (n+1+1) >= 1 + (n+1+1)*(-1/((n+2)*(n+2) )) by POWER:56; then (1 - 1/((n+2)*(n+2))) to_power (n+1+1) >= 1 + (n+1+1)*(-1/((n+2)*(n+2)) ) by XCMPLX_0:def 8; then (1 - 1/((n+2)*(n+2))) to_power (n+1+1) >= 1 + (n+(1+1))*(-1/((n+2)*(n+2 ))) by XCMPLX_1:1; then (1 - 1/((n+2)*(n+2))) to_power (n+1+1) >= 1 + -(n+2)*(1/((n+2)*(n+2))) by XCMPLX_1:175; then (1 - 1/((n+2)*(n+2))) to_power (n+1+1) >= 1 - (n+2)*(1/((n+2)*(n+2))) by XCMPLX_0:def 8; then (1 - 1/((n+2)*(n+2))) to_power (n+1+1) >= 1 - (n+2)*1/((n+2)*(n+2)) by XCMPLX_1:75; then (1 - 1/((n+2)*(n+2))) to_power (n+1+1) >= 1 - ((n+2)/(n+2)*1)/(n+2) by XCMPLX_1:84; then (1 - 1/((n+2)*(n+2))) to_power (n+1+1) >= 1 - (1*1)/(n+2) by XCMPLX_1:60; then s.(n+1)/s.n >= (1 + 1/(n+1)) * (1 - 1/(n+2)) by A3,A6,AXIOMS:25 ; then s.(n+1)/s.n >= ((1*(n+1) + 1)/(n+1)) * (1 - 1/(n+2)) by XCMPLX_1:114; then s.(n+1)/s.n >= ((n+(1+1))/(n+1)) * (1 - 1/(n+2)) by XCMPLX_1:1; then s.(n+1)/s.n >= ((n+2)/(n+1)) * ((1*(n+2) - 1)/(n+2)) by XCMPLX_1:128; then s.(n+1)/s.n >= ((n+2)/(n+1)) * ((n+(1+1-1))/(n+2)) by XCMPLX_1:29; then s.(n+1)/s.n >= ((n+1)*(n+2))/((n+1)*(n+2)) by XCMPLX_1:77; then A7: s.(n+1)/s.n >= 1 by A5,XCMPLX_1:60; s.n>0 by A1,A2; hence s.(n+1)>=s.n by A7,REAL_2:118; end; hence thesis by SEQM_3:def 13; end; Lm48: :: (1 + 1/n)^n is non-decreasing for n st n >= 1 holds ((n+1)/n) to_power n <= ((n+2)/(n+1)) to_power (n+1) proof let n; assume A1: n >= 1; A2: (n-1)+1 = (n+-1)+1 by XCMPLX_0:def 8 .= n+(-1+1) by XCMPLX_1:1 .= n; deffunc _F(Nat) = (1 + 1/($1+1)) to_power ($1+1); consider seq being Real_Sequence such that A3: for n holds seq.n = _F(n)from ExRealSeq; A4: seq is non-decreasing by A3,Lm47; n >= 0+1 by A1; then n-1 >= 0 by REAL_1:84; then reconsider m = n-1 as Nat by INT_1:16; seq.m <= seq.(m+1) by A4,SEQM_3:def 13; then (1 + 1/(m+1)) to_power (m+1) <= seq.(m+1) by A3; then (1 + 1/n) to_power n <= (1 + 1/(n+1)) to_power (n+1) by A2,A3; then (n/n + 1/n) to_power n <= (1 + 1/(n+1)) to_power (n+1) by A1,XCMPLX_1:60; then ((n+1)/n) to_power n <= (1 + 1/(n+1)) to_power (n+1) by XCMPLX_1:63 ; then ((n+1)/n) to_power n <= ((n+1)/(n+1) + 1/(n+1)) to_power (n+1) by XCMPLX_1:60; then ((n+1)/n) to_power n <= (((n+1)+1)/(n+1)) to_power (n+1) by XCMPLX_1:63; then ((n+1)/n) to_power n <= ((n+(1+1))/(n+1)) to_power (n+1) by XCMPLX_1:1; hence ((n+1)/n) to_power n <= ((n+2)/(n+1)) to_power (n+1); end; theorem :: (Part 1, O(n^logn) c O(n^sqrt n)) for f,g being Real_Sequence st (f.0 = 0 & (for n st n > 0 holds f.n = (n to_power log(2,n)))) & (g.0 = 0 & (for n st n > 0 holds g.n = (n to_power sqrt n))) holds ex s,s1 being eventually-positive Real_Sequence st s = f & s1 = g & Big_Oh(s) c= Big_Oh(s1) & not Big_Oh(s) = Big_Oh(s1) proof let f,g be Real_Sequence such that A1: (f.0 = 0 & (for n st n > 0 holds f.n = (n to_power log(2,n)))) and A2: (g.0 = 0 & (for n st n > 0 holds g.n = (n to_power sqrt n))); set h = f/"g; f is eventually-positive proof take 1; let n; assume A3: n >= 1; then f.n = n to_power log(2,n) by A1; hence f.n > 0 by A3,POWER:39; end; then reconsider f as eventually-positive Real_Sequence; g is eventually-positive proof take 1; let n; assume A4: n >= 1; then g.n = n to_power sqrt n by A2; hence g.n > 0 by A4,POWER:39; end; then reconsider g as eventually-positive Real_Sequence; take f,g; consider N such that A5: for n st n >= N holds sqrt n - log(2, n) > 1 by Lm38; A6: now let p be real number such that A7: p > 0; set N1 = max( N, max([/1/p\], 2) ); A8: N1 >= N & N1 >= max([/1/p\], 2) by SQUARE_1:46; A9: max([/1/p\], 2) >= [/1/p\] & max([/1/p\], 2) >= 2 by SQUARE_1:46; then A10: N1 >= [/1/p\] & N1 >= 2 by A8,AXIOMS:22; then A11: N1 > 1 by AXIOMS:22; N1 is Integer proof per cases by SQUARE_1:49; suppose N1 = N; hence thesis; suppose N1 = max([/1/p\], 2); hence thesis by SQUARE_1:49; end; then reconsider N1 as Nat by A8,INT_1:16; take N1; let n; assume A12: n >= N1; p" > 0 by A7,REAL_1:72; then A13: 1/p > 0 by XCMPLX_1:217; [/1/p\] >= 1/p by INT_1:def 5; then A14: N1 >= 1/p by A10,AXIOMS:22; A15: n > 1 by A11,A12,AXIOMS:22; A16: h.n = f.n/g.n by Lm4; f.n = n to_power log(2,n) by A1,A8,A9,A12; then A17: h.n = (n to_power log(2,n)) / (n to_power sqrt n) by A2,A8,A9,A12, A16 .= n to_power (log(2,n) - sqrt n) by A8,A9,A12,POWER:34 .= n to_power -(sqrt n - log(2,n)) by XCMPLX_1:143; n >= N by A8,A12,AXIOMS:22; then sqrt n - log(2, n) > 1 by A5; then (-1)*(sqrt n - log(2, n)) < (-1)*1 by REAL_1:71; then -(sqrt n - log(2, n)) < (-1)*1 by XCMPLX_1:180; then A18: n to_power -(sqrt n - log(2,n)) < n to_power -1 by A15,POWER:44; A19: n to_power -1 = 1/(n to_power 1) by A8,A9,A12,POWER:33 .= 1/n by POWER:30; n >= 1/p by A12,A14,AXIOMS:22; then 1/n <= 1/(1/p) by A13,REAL_2:152; then n to_power -1 <= p by A19,XCMPLX_1:56; then A20: h.n < p by A17,A18,AXIOMS:22; h.n > 0 by A8,A9,A12,A17,POWER:39; hence abs(h.n-0) < p by A20,ABSVALUE:def 1; end; then A21: h is convergent by SEQ_2:def 6; then lim h = 0 by A6,SEQ_2:def 7; then f in Big_Oh(g) & not g in Big_Oh(f) by A21,ASYMPT_0:16; then f in Big_Oh(g) & not f in Big_Omega(g) by ASYMPT_0:19; hence thesis by Th4; end; theorem :: (Part 2, O(n^sqrt n) c O(2^n)) for f being Real_Sequence st (f.0 = 0 & (for n st n > 0 holds f.n = (n to_power sqrt n))) holds ex s,s1 being eventually-positive Real_Sequence st s = f & s1 = seq_a^(2,1,0) & Big_Oh(s) c= Big_Oh(s1) & not Big_Oh(s) = Big_Oh(s1) proof let f be Real_Sequence such that A1: (f.0 = 0 & (for n st n > 0 holds f.n = (n to_power sqrt n))); set g = seq_a^(2,1,0); set h = f/"g; f is eventually-positive proof take 1; let n; assume A2: n >= 1; then f.n = n to_power sqrt n by A1; hence f.n > 0 by A2,POWER:39; end; then reconsider f as eventually-positive Real_Sequence; reconsider g as eventually-positive Real_Sequence; take f,g; consider N such that A3: for n st n >= N holds n - sqrt n*log(2,n) > n/2 by Lm46; A4: now let p be real number; assume p > 0; then p" > 0 by REAL_1:72; then A5: 1/p > 0 by XCMPLX_1:217; set N1 = max( N, max(2*[/log(2,1/p)\], 2) ); A6: N1 >= N & N1 >= max(2*[/log(2,1/p)\], 2) by SQUARE_1:46; A7: max(2*[/log(2,1/p)\], 2) >= 2*[/log(2,1/p)\] & max(2*[/log(2,1/p)\], 2) >= 2 by SQUARE_1:46; then A8: N1 >= 2*[/log(2,1/p)\] & N1 >= 2 by A6,AXIOMS:22; N1 is Integer proof per cases by SQUARE_1:49; suppose N1 = N; hence thesis; suppose N1 = max(2*[/log(2,1/p)\], 2); hence thesis by SQUARE_1:49; end; then reconsider N1 as Nat by A6,INT_1:16; take N1; let n; assume A9: n >= N1; A10: h.n = f.n/g.n by Lm4; A11: f.n = n to_power sqrt n by A1,A6,A7,A9 .= 2 to_power (sqrt n*log(2,n)) by A6,A7,A9,Lm3; g.n = 2 to_power (1*n+0) by Def1 .= 2 to_power n; then A12: h.n = 2 to_power ((sqrt n*log(2,n)) - n) by A10,A11,POWER:34 .= 2 to_power -(n - sqrt n*log(2,n)) by XCMPLX_1:143; n >= N by A6,A9,AXIOMS:22; then n - sqrt n*log(2,n) > n/2 by A3; then -(n - sqrt n*log(2,n)) < -(n/2) by REAL_1:50; then A13: 2 to_power -(n - sqrt n*log(2,n)) < 2 to_power -(n/2) by POWER:44; A14: [/log(2,1/p)\] >= log(2,1/p) by INT_1:def 5; n >= 2*[/log(2,1/p)\] by A8,A9,AXIOMS:22; then n/2 >= [/log(2,1/p)\] by REAL_2:177; then n/2 >= log(2,1/p) by A14,AXIOMS:22; then -(n/2) <= -log(2,1/p) by REAL_1:50; then 2 to_power -(n/2) <= 2 to_power -log(2,1/p) by PRE_FF:10; then 2 to_power -(n/2) <= 1/(2 to_power log(2,1/p)) by POWER:33; then 2 to_power -(n/2) <= 1/(1/p) by A5,POWER:def 3; then 2 to_power -(n/2) <= p by XCMPLX_1:56; then A15: h.n < p by A12,A13,AXIOMS:22; h.n > 0 by A12,POWER:39; hence abs(h.n-0) < p by A15,ABSVALUE:def 1; end; then A16: h is convergent by SEQ_2:def 6; then lim h = 0 by A4,SEQ_2:def 7; then f in Big_Oh(g) & not g in Big_Oh(f) by A16,ASYMPT_0:16; then f in Big_Oh(g) & not f in Big_Omega(g) by ASYMPT_0:19; hence thesis by Th4; end; theorem :: (Part 3, O(2^n+1) = O(2^n)) ex s, s1 being eventually-positive Real_Sequence st s = seq_a^(2,1,0) & s1 = seq_a^(2,1,1) & Big_Oh(s) = Big_Oh(s1) proof set f = seq_a^(2,1,0); set g = seq_a^(2,1,1); set h = f/"g; reconsider f as eventually-positive Real_Sequence; reconsider g as eventually-positive Real_Sequence; take f,g; thus f = seq_a^(2,1,0) & g = seq_a^(2,1,1); A1: 2" > 0; A2: for n holds h.n = 2" proof now let n; A3: f.n = 2 to_power (1*n + 0) by Def1; A4: g.n = 2 to_power (1*n + 1) by Def1; h.n = (2 to_power n)/g.n by A3,Lm4 .= 2 to_power (n-(n+1)) by A4,POWER:34 .= 2 to_power (n-n-1) by XCMPLX_1:36 .= 2 to_power (n+-n-1) by XCMPLX_0:def 8 .= 2 to_power (n+-n+-1) by XCMPLX_0:def 8 .= 2 to_power (0+-1) by XCMPLX_0:def 6 .= 1 / 2 to_power 1 by POWER:33 .= 1 / 2 by POWER:30 .= 2"; hence h.n = 2"; end; hence thesis; end; A5: now let p be real number such that A6: p > 0; take N = 0; let n; assume n >= N; abs(h.n-2") = abs(2"-2") by A2 .= 0 by ABSVALUE:7; hence abs(h.n-2") < p by A6; end; then A7: h is convergent by SEQ_2:def 6; then lim h > 0 by A1,A5,SEQ_2:def 7; hence thesis by A7,ASYMPT_0:15; end; theorem :: (Part 4, O(2^n) c O(2^2n)) ex s, s1 being eventually-positive Real_Sequence st s = seq_a^(2,1,0) & s1 = seq_a^(2,2,0) & Big_Oh(s) c= Big_Oh(s1) & not Big_Oh(s) = Big_Oh(s1) proof reconsider f = seq_a^(2,1,0) as eventually-positive Real_Sequence; reconsider g = seq_a^(2,2,0) as eventually-positive Real_Sequence; take f,g; thus f = seq_a^(2,1,0) & g = seq_a^(2,2,0); set h = f/"g; A1: for n holds h.n = (2 to_power -n) proof let n; h.n = f.n/g.n by Lm4 .= (2 to_power (1*n+0))/g.n by Def1 .= (2 to_power (1*n))/(2 to_power (2*n+0)) by Def1 .= 2 to_power (1*n-(2*n)) by POWER:34 .= 2 to_power (1*n+-(2*n)) by XCMPLX_0:def 8 .= 2 to_power (1*n+(-2)*n) by XCMPLX_1:175 .= 2 to_power ((1+-2)*n) by XCMPLX_1:8 .= 2 to_power ((-1)*n) .= 2 to_power (-(1*n)) by XCMPLX_1:175 .= 2 to_power -n; hence h.n = (2 to_power -n); end; A2: now let p be real number; assume A3: p > 0; then p" > 0 by REAL_1:72; then A4: 1/p > 0 by XCMPLX_1:217; set N = max(1,[/log(2,1/p)\] + 1); A5: N >= 1 & N >= ([/log(2,1/p)\] + 1) by SQUARE_1:46; N is Integer by SQUARE_1:49; then reconsider N as Nat by A5,INT_1:16; take N; let n; assume A6: n >= N; A7: abs(h.n-0) = abs((2 to_power -n)) by A1; (2 to_power -n) > 0 by POWER:39; then A8: abs(h.n-0) = (2 to_power -n) by A7,ABSVALUE:def 1; A9: 2 to_power n >= 2 to_power N by A6,PRE_FF:10; A10: 2 to_power N >= 2 to_power ([/log(2,1/p)\] + 1) by A5,PRE_FF:10; A11: [/log(2,1/p)\] + 1 > [/log(2,1/p)\] by REAL_1:69; [/log(2,1/p)\] >= log(2,1/p) by INT_1:def 5; then [/log(2,1/p)\] + 1 > log(2,1/p) by A11,AXIOMS:22; then 2 to_power ([/log(2,1/p)\]+1) > 2 to_power log(2,1/p) by POWER:44; then 2 to_power N > 2 to_power log(2,1/p) by A10,AXIOMS:22; then 2 to_power n > 2 to_power log(2,1/p) by A9,AXIOMS:22; then 2 to_power n > 1/p by A4,POWER:def 3; then (2 to_power n)*p > 1/p*p by A3,REAL_1:70; then A12: p*(2 to_power n) > 1 by A3,XCMPLX_1:88; 2 to_power n > 0 by POWER:39; then (2 to_power n)" > 0 by REAL_1:72; then (p*(2 to_power n))*(2 to_power n)" > 1*(2 to_power n)" by A12,REAL_1:70; then A13: p*((2 to_power n)*(2 to_power n)") > (2 to_power n)" by XCMPLX_1:4 ; 2 to_power n <> 0 by POWER:39; then p*1 > (2 to_power n)" by A13,XCMPLX_0:def 7; then p > 1/(2 to_power n) by XCMPLX_1:217; hence abs(h.n-0) < p by A8,POWER:33; end; then A14: h is convergent by SEQ_2:def 6; then lim h = 0 by A2,SEQ_2:def 7; then f in Big_Oh(g) & not g in Big_Oh(f) by A14,ASYMPT_0:16; then f in Big_Oh(g) & not f in Big_Omega(g) by ASYMPT_0:19; hence thesis by Th4; end; theorem :: (Part 5, O(2^2n) c O(n!)) ex s being eventually-positive Real_Sequence st s = seq_a^(2,2,0) & Big_Oh(s) c= Big_Oh(seq_n!(0)) & not Big_Oh(s) = Big_Oh(seq_n!(0)) proof set g = seq_n!(0); reconsider f = seq_a^(2,2,0)as eventually-positive Real_Sequence; take f; thus f = seq_a^(2,2,0); set h = f/"g; A1: now let p be real number; assume p > 0; then p" > 0 by REAL_1:72; then A2: 1/p > 0 by XCMPLX_1:217; set N = max(10,[/9+log(2,1/p)\]); A3: N >= 10 & N >= [/9+log(2,1/p)\] by SQUARE_1:46; N is Integer by SQUARE_1:49; then reconsider N as Nat by A3,INT_1:16; take N; let n; assume n >= N; then A4: n >= 10 & n >= [/9+log(2,1/p)\] by A3,AXIOMS:22; A5: h.n = f.n/g.n by Lm4 .= (2 to_power (2*n+0)) / g.n by Def1 .= (2 to_power (2*n+0)) / ((n+0)!) by Def5 .= (2 to_power (2*n)) / (n!); then A6: h.n < 1 / (2 to_power (n-9)) by A4,Lm40; [/9+log(2,1/p)\] >= 9+log(2,1/p) by INT_1:def 5; then n >= 9+log(2,1/p) by A4,AXIOMS:22; then n-9 >= log(2,1/p) by REAL_1:84; then 2 to_power (n-9) >= 2 to_power log(2,1/p) by PRE_FF:10; then 2 to_power (n-9) >= 1/p by A2,POWER:def 3; then 1 / (2 to_power (n-9)) <= 1/(1/p) by A2,REAL_2:152; then 1 / (2 to_power (n-9)) <= p by XCMPLX_1:56; then A7: h.n < p by A6,AXIOMS:22; A8: 2 to_power (2*n) > 0 by POWER:39; n! > 0 by NEWTON:23; then (n!)" > 0 by REAL_1:72; then (2 to_power (2*n))*(n!)" > 0*(n!)" by A8,REAL_1:70; then h.n > 0 by A5,XCMPLX_0:def 9; hence abs(h.n-0) < p by A7,ABSVALUE:def 1; end; then A9: h is convergent by SEQ_2:def 6; then lim h = 0 by A1,SEQ_2:def 7; then f in Big_Oh(g) & not g in Big_Oh(f) by A9,ASYMPT_0:16; then f in Big_Oh(g) & not f in Big_Omega(g) by ASYMPT_0:19; hence thesis by Th4; end; theorem :: (Part 6, O(n!) c O((n+1)!)) Big_Oh(seq_n!(0)) c= Big_Oh(seq_n!(1)) & not Big_Oh(seq_n!(0)) = Big_Oh(seq_n!(1)) proof set f = seq_n!(0); set g = seq_n!(1); set h = f/"g; A1: for n holds h.n = 1/(n+1) proof let n; A2: n! <> 0 by NEWTON:23; h.n = f.n/g.n by Lm4 .= (n+0)!/g.n by Def5 .= (n!)/((n+1)!) by Def5 .= ((n!)*1)/((n+1)*(n!)) by NEWTON:21 .= (1/(n+1))*((n!)/(n!)) by XCMPLX_1:77 .= (1/(n+1))*1 by A2,XCMPLX_1:60; hence h.n = 1/(n+1); end; A3: now let p be real number; assume A4: p > 0; then p" > 0 by REAL_1:72; then A5: 1/p > 0 by XCMPLX_1:217; set N = max( 1, [/1/p\] ); A6: N >= 1 & N >= ([/1/p\]) by SQUARE_1:46; N is Integer by SQUARE_1:49; then reconsider N as Nat by A6,INT_1:16; take N; let n; assume n >= N; then A7: n+1 > N by NAT_1:38; A8:-p < 0 by A4,REAL_1:26,50; 0 < (n+1)" by REAL_1:72; then 0 < 1/(n+1) by XCMPLX_1:217; then 0 < h.n by A1; then A9: -p < h.n by A8,AXIOMS:22; [/1/p\] >= 1/p by INT_1:def 5; then N >= 1/p by A6,AXIOMS:22; then (n+1) > 1/p by A7,AXIOMS:22; then 1/(n+1) < 1/(1/p) by A5,REAL_2:151; then 1/(n+1) < p by XCMPLX_1:56; then A10: h.n < p by A1; abs(h.n) < p proof per cases; suppose h.n >= 0; hence thesis by A10,ABSVALUE:def 1; suppose h.n < 0; then A11: abs(h.n) = -h.n by ABSVALUE:def 1; (-1)*(-p) > (-1)*h.n by A9,REAL_1:71; then p > (-1)*h.n by XCMPLX_1:181; hence thesis by A11,XCMPLX_1:180; end; hence abs(h.n-0) < p; end; then A12: h is convergent by SEQ_2:def 6; then lim h = 0 by A3,SEQ_2:def 7; then f in Big_Oh(g) & not g in Big_Oh(f) by A12,ASYMPT_0:16; then f in Big_Oh(g) & not f in Big_Omega(g) by ASYMPT_0:19; hence thesis by Th4; end; theorem :: (Part 7, O((n+1)!) c O(n^n)) for g being Real_Sequence st (g.0 = 0 & (for n st n > 0 holds g.n = (n to_power n))) holds ex s being eventually-positive Real_Sequence st s = g & Big_Oh(seq_n!(1)) c= Big_Oh(s) & not Big_Oh(seq_n!(1)) = Big_Oh(s) proof let g be Real_Sequence such that A1: (g.0 = 0 & (for n st n > 0 holds g.n = (n to_power n))); set f = seq_n!(1); set h = f/"g; g is eventually-positive proof take 1; let n; assume A2: n >= 1; then g.n = n to_power n by A1; hence g.n > 0 by A2,POWER:39; end; then reconsider g as eventually-positive Real_Sequence; take g; A3: 1-2 = -1; A4: 3 = 4-1; A5: for n st n >= 1 holds h.n > 0 proof let n; assume A6: n >= 1; A7: h.n = f.n/g.n by Lm4 .= (n+1)! / g.n by Def5 .= (n+1)! / (n to_power n) by A1,A6; n to_power n > 0 by A6,POWER:39; then (n to_power n)" > 0 by REAL_1:72; then A8: 1/(n to_power n) > 0 by XCMPLX_1:217; (n+1)! > 0 by NEWTON:23; then (n+1)!*(1/(n to_power n)) > (n+1)!*0 by A8,REAL_1:70; hence thesis by A7,XCMPLX_1:100; end; deffunc _F(Nat) = h.$1 / h.($1+1); consider p being Real_Sequence such that A9: for n holds p.n = _F(n) from ExRealSeq; A10: for n st n > 0 holds p.n = ((n+1)/(n+2)) * (((n+1)/n) to_power n) proof let n; assume A11: n > 0; A12: (n+1)! > 0 by NEWTON:23; p.n = h.n / h.(n+1) by A9 .= (f.n / g.n) / h.(n+1) by Lm4 .= ((n+1)! / g.n) / h.(n+1) by Def5 .= ((n+1)! / (n to_power n)) / h.(n+1) by A1,A11 .= ((n+1)! / (n to_power n)) / (f.(n+1) / g.(n+1)) by Lm4 .= ((n+1)! / (n to_power n)) / (((n+1)+1)! / g.(n+1)) by Def5 .= ((n+1)! / (n to_power n)) / (((n+1)+1)! / ((n+1) to_power (n+1))) by A1 .= ((n+1)! / (((n+1)+1)!)) * (((n+1) to_power (n+1)) / (n to_power n)) by Lm44 .= ((n+1)! / (((n+1)+1)*((n+1)!))) * (((n+1) to_power (n+1)) / (n to_power n)) by NEWTON:21 .= (1/((n+1)+1) * ((n+1)! / ((n+1)!))) * (((n+1) to_power (n+1)) / (n to_power n)) by XCMPLX_1:104 .= ((1/((n+1)+1)) * 1) * (((n+1) to_power (n+1)) / (n to_power n)) by A12,XCMPLX_1:60 .= (1/(n+(1+1))) * (((n+1) to_power (n+1)) / (n to_power n)) by XCMPLX_1:1 .= (1/(n+2)) * ((((n+1) to_power n)*((n+1) to_power 1)) / (n to_power n)) by POWER:32 .= (1/(n+2)) * ((((n+1) to_power n)*(n+1)) / (n to_power n)) by POWER:30 .= (1/(n+2)) * ((((n+1) to_power n)*(n+1)) * ((n to_power n)")) by XCMPLX_0: def 9 .= (1/(n+2)) * ((((n+1) to_power n)*((n to_power n)"))*(n+1)) by XCMPLX_1:4 .= (1/(n+2)) * ((((n+1) to_power n)/(n to_power n))*(n+1)) by XCMPLX_0:def 9 .= (1/(n+2)) * ((((n+1)/n) to_power n)*(n+1)) by A11,POWER:36 .= ((n+1)*(1/(n+2))) * (((n+1)/n) to_power n) by XCMPLX_1:4 .= ((n+1)/(n+2)) * (((n+1)/n) to_power n) by XCMPLX_1:100; hence p.n = ((n+1)/(n+2)) * (((n+1)/n) to_power n); end; defpred _P1[Nat] means p.$1 > 2; A13: _P1[4] proof p.4 = ((4+1) / (4+2)) * (((4+1)/4) to_power 4) by A10 .= (5 / 6) * ((5 to_power 4) / 256) by Lm43,POWER:36 .= (5 * (5 to_power 4)) / (6 * 256) by XCMPLX_1:77 .= ((5 to_power 1)*(5 to_power 4)) / 1536 by POWER:30 .= (5 to_power (4+1)) / 1536 by POWER:32 .= 3125 / 1536 by Lm42; hence thesis; end; A14: for k st k >= 4 & _P1[k] holds _P1[k+1] proof let k; assume that A15: k >= 4 and A16: p.k > 2; A17: now assume ((k+1)*(k+3)) >= ((k+2)*(k+2)); then ((k+1)*k + (k+1)*3) >= ((k+2)*(k+2)) by XCMPLX_1:8; then ((k*k + 1*k) + (k+1)*3) >= ((k+2)*(k+2)) by XCMPLX_1:8; then ((k*k + 1*k) + (3*k + 3*1)) >= ((k+2)*(k+2)) by XCMPLX_1:8; then (((k*k + 1*k) + 3*k) + 3*1) >= ((k+2)*(k+2)) by XCMPLX_1:1; then ((k*k + (1*k + 3*k)) + 3*1) >= ((k+2)*(k+2)) by XCMPLX_1:1; then ((k*k + ((1+3)*k)) + 3) >= ((k+2)*(k+2)) by XCMPLX_1:8; then (k*k + 4*k) + 3 >= (k+2)^2 by SQUARE_1:def 3; then (k*k + 4*k) + 3 >= k^2 + 2*k*2 + 2^2 by SQUARE_1:63; then (k*k + 4*k) + 3 >= k*k + 2*(2*k) + 2^2 by SQUARE_1:def 3; then (k*k + 4*k) + 3 >= k*k + (2*2)*k + 2^2 by XCMPLX_1:4; then (k*k + 4*k) + 3 >= k*k + 4*k + 2*2 by SQUARE_1:def 3; then -(k*k + 4*k) + ((k*k + 4*k) + 3) >= -(k*k + 4*k) + ((k*k + 4*k) + 4) by AXIOMS:24; then (-(k*k + 4*k) + (k*k + 4*k)) + 3 >= -(k*k + 4*k) + ((k*k + 4*k) + 4) by XCMPLX_1:1; then 0 + 3 >= -(k*k + 4*k) + ((k*k + 4*k) + 4) by XCMPLX_0:def 6; then 3 >= (-(k*k + 4*k) + (k*k + 4*k)) + 4 by XCMPLX_1:1; then 3 >= 0 + 4 by XCMPLX_0:def 6; hence contradiction; end; (k+1)" > 0 by REAL_1:72; then (k+2)*(k+1)" > 0*(k+1)" by REAL_1:70; then (k+2)/(k+1) > 0 by XCMPLX_0:def 9; then A18: (((k+2)/(k+1)) to_power (k+1)) > 0 by POWER:39; (k+3)" > 0 by REAL_1:72; then (k+2)*(k+3)" > 0*(k+3)" by REAL_1:70; then A19: ((k+2)/(k+3)) > 0 by XCMPLX_0:def 9; A20: p.(k+1) = ((((k+1)+1)/((k+1)+2))*((((k+1)+1)/(k+1)) to_power (k+1))) by A10 .= (((k+1)+1)/((k+1)+2)*(((k+(1+1))/(k+1)) to_power (k+1))) by XCMPLX_1: 1 .= ((k+(1+1))/((k+1)+2)*(((k+2)/(k+1)) to_power (k+1))) by XCMPLX_1:1 .= ((k+2)/(k+(1+2))*(((k+2)/(k+1)) to_power (k+1))) by XCMPLX_1:1 .= ((k+2)/(k+3)*(((k+2)/(k+1)) to_power (k+1))); A21: ((k+2)/(k+3))*(((k+2)/(k+1)) to_power (k+1)) > ((k+2)/(k+3))*0 by A18,A19,REAL_1:70; A22: k >= 1 by A15,AXIOMS:22; k" > 0 by A15,REAL_1:72; then (k+1)*k" > 0*k" by REAL_1:70; then (k+1)/k > 0 by XCMPLX_0:def 9; then A23: (((k+1)/k) to_power k) > 0 by POWER:39; (((k+2)/(k+1)) to_power (k+1))" > 0 by A18,REAL_1:72; then (((k+1)/k) to_power k)*(((k+2)/(k+1)) to_power (k+1))" > 0*(((k+2)/(k+1)) to_power (k+1))" by A23,REAL_1:70; then A24: ((((k+1)/k) to_power k) / (((k+2)/(k+1)) to_power (k+1))) > 0 by XCMPLX_0:def 9; A25: (k+1)*(k+3) > 0*(k+3) by REAL_1:70; ((k+2)*(k+2))" > 0 by A17,REAL_1:72; then (((k+1)*(k+3)) * ((k+2)*(k+2))") > 0*((k+2)*(k+2))" by A25,REAL_1:70; then A26: (((k+1)*(k+3)) / ((k+2)*(k+2))) > 0 by XCMPLX_0:def 9; A27: p.k / p.(k+1) = ((k+1)/(k+2)) * (((k+1)/k) to_power k) / ((k+2)/(k+3)*(((k+2)/(k+1)) to_power (k+1))) by A10,A15,A20 .= (((k+1)/(k+2)) / ((k+2)/(k+3))) * ((((k+1)/k) to_power k) / (((k+2)/(k+1)) to_power (k+1))) by XCMPLX_1:77 .= (((k+1)*(k+3)) / ((k+2)*(k+2))) * ((((k+1)/k) to_power k) / (((k+2)/(k+1)) to_power (k+1))) by XCMPLX_1:85; ((k+1)*(k+3)) < 1*((k+2)*(k+2)) by A17; then A28: (((k+1)*(k+3)) / ((k+2)*(k+2))) < 1 by REAL_2:178; (((k+1)/k) to_power k) <= 1*(((k+2)/(k+1)) to_power (k+1)) by A22,Lm48; then (((k+1)/k) to_power k) / (((k+2)/(k+1)) to_power (k+1)) <= 1 by A18,REAL_2:177; then (((k+1)*(k+3)) / ((k+2)*(k+2))) * ((((k+1)/k) to_power k) / (((k+2)/(k+1)) to_power (k+1))) < 1*1 by A24,A26,A28,REAL_2:199; then p.k / p.(k+1) * p.(k+1) < 1*p.(k+1) by A20,A21,A27,REAL_1:70; then p.(k+1) > p.k by A20,A21,XCMPLX_1:88; hence p.(k+1) > 2 by A16,AXIOMS:22; end; A29: for n st n >= 4 holds _P1[n] from Ind1(A13, A14); defpred _P[Nat] means h.$1 < 1/($1-2); A30: _P[4] proof h.4 = f.4/g.4 by Lm4 .= (4+1)!/g.4 by Def5 .= 120/256 by A1,Lm39,Lm43; hence thesis; end; A31: for k st k >= 4 & _P[k] holds _P[k+1] proof let k; assume that A32: k >= 4 and A33: h.k < 1/(k-2); p.k > 2 by A29,A32; then A34: h.k / h.(k+1) > 2 by A9; k+1 >= 1 by NAT_1:29; then A35: h.(k+1) > 0 by A5; then (h.k / h.(k+1)) * h.(k+1) > 2*h.(k+1) by A34,REAL_1:70; then h.k > 2*h.(k+1) by A35,XCMPLX_1:88; then A36: h.k/2 > h.(k+1) by REAL_2:178; h.k*(1/2) < (1/(k-2))*(1/2) by A33,REAL_1:70; then h.k/2 < (1/(k-2))*(1/2) by XCMPLX_1:100; then A37: h.k/2 < 1/(2*(k-2)) by XCMPLX_1:103; k >= 3 by A32,AXIOMS:22; then A38: 2*(k-2) >= k-1 by Lm41; k-1 >= 3 by A4,A32,REAL_1:49; then 1/(2*(k-2)) <= 1/(k-1) by A38,REAL_2:152; then h.k/2 < 1/(k-1) by A37,AXIOMS:22; then h.(k+1) < 1/(k-1) by A36,AXIOMS:22; then h.(k+1) < 1/(k+(1+-2)) by A3,XCMPLX_0:def 8; then h.(k+1) < 1/((k+1)+-2) by XCMPLX_1:1; hence thesis by XCMPLX_0:def 8; end; A39: for n st n >= 4 holds _P[n] from Ind1(A30, A31); A40: now let p be real number; assume p > 0; then p" > 0 by REAL_1:72; then A41: 1/p > 0 by XCMPLX_1:217; set N = [/1/p + 4\]; A42: [/1/p + 4\] >= 1/p + 4 by INT_1:def 5; 4 + 1/p > 4 by A41,REAL_1:69; then A43: N >= 4 by A42,AXIOMS:22; then reconsider N as Nat by INT_1:16; take N; let n; assume A44: n >= N; then A45: n >= 4 by A43,AXIOMS:22; then A46: h.n < 1/(n-2) by A39; n >= 1/p + (2 + 2) by A42,A44,AXIOMS:22; then n >= (1/p + 2) + 2 by XCMPLX_1:1; then A47: n - 2 >= 1/p + 2 by REAL_1:84; 1/p + 2 > 1/p by REAL_1:69; then n - 2 > 1/p by A47,AXIOMS:22; then 1/(n-2) < 1/(1/p) by A41,REAL_2:151; then 1/(n-2) < 1*p by XCMPLX_1:101; then A48: h.n < p by A46,AXIOMS:22; n >= 1 by A45,AXIOMS:22; then h.n > 0 by A5; hence abs(h.n-0) < p by A48,ABSVALUE:def 1; end; then A49: h is convergent by SEQ_2:def 6; then lim h = 0 by A40,SEQ_2:def 7; then f in Big_Oh(g) & not g in Big_Oh(f) by A49,ASYMPT_0:16; then f in Big_Oh(g) & not f in Big_Omega(g) by ASYMPT_0:19; hence thesis by Th4; end; begin :: Problem 3.23 Lm49: for k,n st k <= n holds n choose k >= ((n+1) choose k) / (n+1) proof let k,n; assume A1: k <= n; n + 0 <= n + 1 by AXIOMS:24; then A2: k <= n+1 by A1,AXIOMS:22; set n1 = n+1; reconsider l = n-k as Nat by A1,INT_1:18; set l1 = l+1; l1 = n1-k by XCMPLX_1:29; then A3: (n1 choose k) / n1 = ((n1!)/((k!) * (l1!))) / n1 by A2,NEWTON:def 3 .= (n1*(n!)/((k!) * (l1!))) / (n1*1) by NEWTON:21 .= (n1*(n!)*((k!) * (l1!))") / (n1*1) by XCMPLX_0:def 9 .= (n1*((n!)*((k!) * (l1!))")) / (n1*1) by XCMPLX_1:4 .= (n1*((n!)/((k!) * (l1!)))) / (n1*1) by XCMPLX_0:def 9 .= (n1/n1) * (((n!)/((k!) * (l1!)))/1) by XCMPLX_1:77 .= 1 * (((n!)/((k!) * (l1!)))/1) by XCMPLX_1:60 .= (n!)/((k!) * ((l!) * l1)) by NEWTON:21 .= (n!)*1/((k!) * (l!) * l1) by XCMPLX_1:4 .= (n!)/((k!) * (l!)) * (1/l1) by XCMPLX_1:77 .= (n choose k) * (1/l1) by A1,NEWTON:def 3; A4: n! > 0 & k! > 0 & l! > 0 by NEWTON:23; then (k!)*(l!) > (k!)*0 by REAL_1:70; then ((k!)*(l!))" > 0 by REAL_1:72; then (n!)*((k!)*(l!))" > (n!)*0 by A4,REAL_1:70; then (n!)/((k!)*(l!)) > 0 by XCMPLX_0:def 9; then A5: (n choose k) > 0 by A1,NEWTON:def 3; 0+1 <= l+1 by AXIOMS:24; then 1/1 >= 1/l1 by REAL_2:152; then (n choose k) * (1/1) >= (n choose k) * (1/l1) by A5,AXIOMS:25; hence thesis by A3; end; theorem for n st n >= 1 holds for f being Real_Sequence, k being Nat st (for n holds f.n = Sum(seq_n^(k), n)) holds f.n >= (n to_power (k+1)) / (k+1) proof defpred _P[Nat] means for f being Real_Sequence, k being Nat st (for n holds f.n = Sum(seq_n^(k), n)) holds f.$1 >= ($1 to_power (k+1)) / (k+1); A1: _P[1] proof let f be Real_Sequence, k be Nat such that A2: for n holds f.n = Sum(seq_n^(k), n); set g = seq_n^(k); A3: f.1 = Sum(g,1) by A2 .= (Partial_Sums(g)).(0+1) by BHSP_4:def 6 .= (Partial_Sums(g)).0 + g.1 by SERIES_1:def 1 .= g.1 + g.0 by SERIES_1:def 1 .= (1 to_power k) + g.0 by Def3 .= 1 + g.0 by POWER:31 .= 1 + 0 by Def3 .= 1/1; A4: (1 to_power (k+1)) / (k+1) = 1 / (k+1) by POWER:31; 0+1 <= k+1 by AXIOMS:24; hence thesis by A3,A4,REAL_2:152; end; A5: for n being Nat st n >= 1 & _P[n] holds _P[n+1] proof let n be Nat such that A6: n >= 1 and A7: for f being Real_Sequence, k being Nat st (for n holds f.n = Sum(seq_n^(k), n)) holds f.n >= (n to_power (k+1)) / (k+1); let f be Real_Sequence, k be Nat such that A8: for n holds f.n = Sum(seq_n^(k), n); set g = seq_n^(k); A9: f.(n+1) = Sum(g,n+1) by A8 .= (Partial_Sums(g)).(n+1) by BHSP_4:def 6 .= (Partial_Sums(g)).n + g.(n+1) by SERIES_1:def 1; f.n >= (n to_power (k+1)) / (k+1) by A7,A8; then Sum(g,n) >= (n to_power (k+1)) / (k+1) by A8; then (Partial_Sums(g)).n >= (n to_power (k+1)) / (k+1) by BHSP_4:def 6; then A10: f.(n+1) >= (n to_power (k+1)) / (k+1) + g.(n+1) by A9,AXIOMS:24; g.(n+1) = (n+1) to_power k by Def3 .= (n+1)|^k by POWER:48 .= Sum((n,1) In_Power k) by NEWTON:41; then A11: (n to_power (k+1)) / (k+1) + g.(n+1) = Sum (<*(n to_power (k+1)) / (k+1)*>^((n,1) In_Power k)) by RVSUM_1:106; A12: ((n+1) to_power (k+1)) / (k+1) = ((n+1) to_power (k+1)) * (k+1)" by XCMPLX_0:def 9 .= ((n+1)|^(k+1))*(k+1)" by POWER:48 .= Sum((n,1) In_Power (k+1))*(k+1)" by NEWTON:41 .= Sum(((k+1)")*((n,1) In_Power (k+1))) by RVSUM_1:117; set R1 = <*(n to_power (k+1)) / (k+1)*>^((n,1) In_Power k); set R2 = ((k+1)")*((n,1) In_Power (k+1)); A13: len ((n,1) In_Power k) = k+1 by NEWTON:def 4; A14: len <*(n to_power (k+1)) / (k+1)*> = 1 by FINSEQ_1:57; then A15: len R1 = (k+1)+1 by A13,FINSEQ_1:35 .= k+(1+1) by XCMPLX_1:1 .= k+2; then reconsider R1 as Element of (k+2)-tuples_on REAL by FINSEQ_2:110; len R2 = len ((n,1) In_Power (k+1)) by NEWTON:6 .= (k+1)+1 by NEWTON:def 4 .= k+(1+1) by XCMPLX_1:1 .= k+2; then reconsider R2 as Element of (k+2)-tuples_on REAL by FINSEQ_2:110; set R3 = ((n,1) In_Power (k+1)); len R3 = (k+1)+1 by NEWTON:def 4 .= k+(1+1) by XCMPLX_1:1 .= k+2; then reconsider R3 as Element of (k+2)-tuples_on REAL by FINSEQ_2:110; for i st i in Seg (k+2) holds R2.i<=R1.i proof let i such that A16: i in Seg (k+2); set r2 = R2.i, r1 = R1.i; A17: 1 <= i & i <= (k+2) by A16,FINSEQ_1:3; set k1 = (k+1)"; per cases by A17,REAL_1:def 5; suppose A18: i = 1; n|^(k+1) = R3.1 by NEWTON:39; then r2 = k1*(n|^(k+1)) by A18,RVSUM_1:67 .= (n to_power (k+1)) * k1 by A6,POWER:48 .= (n to_power (k+1)) / (k+1) by XCMPLX_0:def 9; hence thesis by A18,FINSEQ_1:58; suppose A19: i > 1; set i0 = i-1; set m = i0-1; set l = k-m; A20: i-1 > 1-1 by A19,REAL_1:54; then reconsider i0 as Nat by INT_1:16; A21: i0 >= 0+1 by A20,INT_1:20; then m >= 0 by REAL_1:84; then reconsider m as Nat by INT_1:16; m = i-(1+1) by XCMPLX_1:36 .= i-2; then A22: k >= m+0 by A17,REAL_1:86; then l >= 0 by REAL_1:84; then reconsider l as Nat by INT_1:16; i-1 <= (k+2)-1 by A17,REAL_1:49; then i0 <= (k+2)+-1 by XCMPLX_0:def 8; then A23: i0 <= k+(2+-1) by XCMPLX_1:1; len (((n,1) In_Power k)) = k+1 by NEWTON:def 4; then dom(((n,1) In_Power k)) = Seg(k+1) by FINSEQ_1:def 3; then A24: i0 in dom(((n,1) In_Power k)) by A21,A23,FINSEQ_1:3; A25: r1 = ((n,1) In_Power k).i0 by A14,A15,A17,A19,FINSEQ_1:37 .= (k choose m)*(n|^l)*(1|^m) by A24,NEWTON:def 4 .= (k choose l)*(n|^l)*(1|^m) by A22,NEWTON:30 .= (k choose l)*(n|^l)*1 by NEWTON:15 .= (k choose l)*(n to_power l) by A6,POWER:48; set i3 = (k+1)-i0; A26: i3 = (k+-(-1))-i0 .= (k-(-1))-i0 by XCMPLX_0:def 8 .= k-((-1)+i0) by XCMPLX_1:36 .= l by XCMPLX_0:def 8; then reconsider i3 as Nat; A27: i0 + 0 <= (k+1) by A26,REAL_1:84; len(((n,1) In_Power (k+1))) = (k+1)+1 by NEWTON:def 4; then dom(((n,1) In_Power (k+1))) = Seg((k+1)+1) by FINSEQ_1:def 3 .= Seg(k+(1+1)) by XCMPLX_1:1 .= Seg(k+2); then R3.i = ((k+1) choose i0)*(n|^i3)*(1|^i0) by A16,NEWTON:def 4; then A28: r2 = k1*(((k+1) choose i0)*(n|^i3)*(1|^i0)) by RVSUM_1:67 .= k1*(((k+1) choose l)*(n|^l)*(1|^i0)) by A26,A27,NEWTON:30 .= k1*(((k+1) choose l)*(n|^l)*1) by NEWTON:15 .= k1*(((k+1) choose l)*(n to_power l)) by A6,POWER:48 .= k1*((k+1) choose l)*(n to_power l) by XCMPLX_1:4; k-m <= k-0 by REAL_1:92; then ((k+1) choose l) / (k+1) <= (k choose l) by Lm49; then k1 * ((k+1) choose l) <= (k choose l) by XCMPLX_0:def 9; hence thesis by A25,A28,AXIOMS:25; end; then ((n+1) to_power (k+1)) / (k+1) <= Sum R1 by A12,RVSUM_1:112; hence f.(n+1) >= ((n+1) to_power (k+1)) / (k+1) by A10,A11,AXIOMS:22; end; for n st n >= 1 holds _P[n] from Ind1(A1, A5); hence thesis; end; begin :: Problem 3.24 Lm50: for f being Real_Sequence st (for n holds f.n = log(2,n!)) holds for n holds f.n = Sum(seq_logn, n) proof let f be Real_Sequence such that A1: for n holds f.n = log(2,n!); set g = seq_logn; defpred _P[Nat] means f.$1 = Sum(g, $1); A2: _P[0] proof A3: f.0 = log(2,1) by A1,NEWTON:18 .= 0 by POWER:59; Sum(g, 0) = Partial_Sums(g).0 by BHSP_4:def 6 .= g.0 by SERIES_1:def 1 .= 0 by Def2; hence thesis by A3; end; A4: for k st _P[k] holds _P[k+1] proof let k such that A5: f.k = Sum(g, k); A6: k! > 0 by NEWTON:23; f.(k+1) = log(2,(k+1)!) by A1 .= log(2,(k+1)*(k!)) by NEWTON:21 .= log(2,k+1) + log(2,k!) by A6,POWER:61 .= log(2,k+1) + Sum(g, k) by A1,A5 .= g.(k+1) + Sum(g, k) by Def2 .= g.(k+1) + Partial_Sums(g).k by BHSP_4:def 6 .= Partial_Sums(g).(k+1) by SERIES_1:def 1 .= Sum(g, k+1) by BHSP_4:def 6; hence thesis; end; for n holds _P[n] from Ind(A2, A4); hence thesis; end; Lm51: for n st n >= 4 holds n*log(2,n) >= 2*n proof let n; assume n >= 4; then log(2,n) >= log(2,2*2) by PRE_FF:12; then log(2,n) >= log(2,2^2) by SQUARE_1:def 3; then log(2,n) >= log(2,2 to_power 2) by POWER:53; then log(2,n) >= 2*log(2,2) by POWER:63; then log(2,n) >= 2*1 by POWER:60; hence n*log(2,n) >= 2*n by AXIOMS:25; end; theorem for f,g being Real_Sequence st (g.0 = 0 & (for n st n > 0 holds g.n = n*log(2,n))) & (for n holds f.n = log(2,n!)) holds ex s being eventually-nonnegative Real_Sequence st s = g & f in Big_Theta(s) proof let f,g be Real_Sequence such that A1: (g.0 = 0 & (for n st n > 0 holds g.n = n*log(2,n))) and A2: for n holds f.n = log(2,n!); g is eventually-positive proof take 2; let n; assume A3: n >= 2; then log(2,n) >= log(2,2) by PRE_FF:12; then log(2,n) >= 1 by POWER:60; then n*log(2,n) > n*0 by A3,REAL_1:70; hence g.n > 0 by A1,A3; end; then reconsider g as eventually-positive Real_Sequence; take g; set h = seq_logn; A4: Big_Theta(g) = { s where s is Element of Funcs(NAT, REAL) : ex c,d,N st c > 0 & d > 0 & for n st n >= N holds d*g.n <= s.n & s.n <= c*g.n } by ASYMPT_0:27; A5: f is Element of Funcs(NAT, REAL) by FUNCT_2:11; now let n; assume A6: n >= 4; ex s being Real_Sequence st s.0 = 0 & (for m st m > 0 holds s.m = (log(2,n/2))) proof defpred P[Nat,Real] means ($1 = 0 implies $2 = 0) & ($1 > 0 implies $2 = log(2,n/2)); A7: for x being Element of NAT ex y being Element of REAL st P[x,y] proof let n; per cases; suppose n = 0; hence thesis; suppose n > 0; hence thesis; end; consider h being Function of NAT,REAL such that A8: for x being Element of NAT holds P[x,h.x] from FuncExD(A7); take h; thus h.0 = 0 by A8; let n; thus thesis by A8; end; then consider p being Real_Sequence such that A9: p.0 = 0 & (for m st m > 0 holds p.m = log(2,n/2)); ex s being Real_Sequence st s.0 = 0 & (for m st m > 0 holds s.m = (log(2,n))) proof defpred P[Nat,Real] means ($1 = 0 implies $2 = 0) & ($1 > 0 implies $2 = log(2,n)); A10: for x being Element of NAT ex y being Element of REAL st P[x,y] proof let n; per cases; suppose n = 0; hence thesis; suppose n > 0; hence thesis; end; consider h being Function of NAT,REAL such that A11: for x being Element of NAT holds P[x,h.x] from FuncExD(A10); take h; thus h.0 = 0 by A11; let n; thus thesis by A11; end; then consider q being Real_Sequence such that A12: q.0 = 0 & (for m st m > 0 holds q.m = log(2,n)); set n1 = [/n/2\]; A13: n*2" > 0*2" by A6,REAL_1:70; then A14: n/2 > 0 by XCMPLX_0:def 9; A15: [/n/2\] >= n/2 by INT_1:def 5; then reconsider n1 as Nat by INT_1:16; set n2 = n1-1; now assume n2 < 0; then n1-1 <= -1 by INT_1:21; then (n1-1)+1 <= -1+1 by AXIOMS:24; then (n1+-1)+1 <= 0 by XCMPLX_0:def 8; then n1+(-1+1) <= 0 by XCMPLX_1:1; hence contradiction by A13,A15,XCMPLX_0:def 9; end; then reconsider n2 as Nat by INT_1:16; for k st k <= n2 holds h.k >= 0 proof let k such that k <= n2; per cases; suppose k = 0; hence thesis by Def2; suppose A16: k > 0; then k >= 0+1 by NAT_1:38; then log(2,k) >= log(2,1) by PRE_FF:12; then log(2,k) >= 0 by POWER:59; hence thesis by A16,Def2; end; then Sum(h,n2) >= 0 by Lm17; then Sum(h, n) + Sum(h,n2) >= Sum(h, n) + 0 by AXIOMS:24; then Sum(h, n) >= Sum(h, n) - Sum(h,n2) by REAL_1:86; then A17: Sum(h, n) >= Sum(h, n, n2) by BHSP_4:def 7; n >= n1 by Lm22; then n+-1 >= n1+-1 by AXIOMS:24; then n-1 >= n1+-1 by XCMPLX_0:def 8; then n-1 >= n2 by XCMPLX_0:def 8; then A18: n >= n2 + 1 by REAL_1:84; for k st n2+1 <= k & k <= n holds p.k <= h.k proof let k such that A19: n2+1 <= k and k <= n; n2+1 >= (n1+-1)+1 by XCMPLX_0:def 8; then n2+1 >= n1+(-1+1) by XCMPLX_1:1; then k >= n1 by A19,AXIOMS:22; then n/2 <= k by A15,AXIOMS:22; then log(2,n/2) <= log(2,k) by A14,PRE_FF:12; then p.k <= log(2,k) by A9,A19; hence thesis by A19,Def2; end; then A20: Sum(h, n, n2) >= Sum(p, n, n2) by A18,Lm21; A21: now assume n-n2 < n/2; then A22: n < n/2 + n2 by REAL_1:84; [/n/2\] < n/2 + 1 by INT_1:def 5; then n2 < n/2 by REAL_1:84; then n/2 + n2 < n/2 + n/2 by REAL_1:53; then n < 1*(n/2) + 1*(n/2) by A22,AXIOMS:22; then n < (1+1)*(n/2) by XCMPLX_1:8; then n < 2*(2"*n) by XCMPLX_0:def 9; then n < (2*2")*n by XCMPLX_1:4; hence contradiction; end; A23: Sum(p, n, n2) = (n - n2)*log(2,n/2) by A9,Lm23; 2*2 <= n by A6; then 2 <= n/2 by REAL_2:177; then log(2,2) <= log(2,n/2) by PRE_FF:12; then 1 <= log(2,n/2) by POWER:60; then A24: Sum(p, n, n2) >= (n/2)*log(2,n/2) by A21,A23,AXIOMS:25; A25: (n/2)*log(2,n/2) = (n/2)*(log(2,n) - log(2,2)) by A6,POWER:62 .= (n/2)*(log(2,n) - 1) by POWER:60 .= (n/2)*(log(2,n) + -1) by XCMPLX_0:def 8 .= log(2,n)*(n/2) + (-1)*(n/2) by XCMPLX_1:8 .= log(2,n)*(n/2) + -(n/2) by XCMPLX_1:180 .= log(2,n)*(n/2) - (n/2) by XCMPLX_0:def 8 .= log(2,n)*(n*2") - (n/2) by XCMPLX_0:def 9 .= (log(2,n)*n)*2" - (n/2) by XCMPLX_1:4 .= (n*log(2,n))/2 - (n/2) by XCMPLX_0:def 9; n*log(2,n) >= 2*n by A6,Lm51; then (n*log(2,n))*4" >= (2*n)*4" by AXIOMS:25; then (n*log(2,n))*4" >= (2*n) / (2*2) by XCMPLX_0:def 9; then ((2+-1)*(n*log(2,n))) * 4" >= n/2 by XCMPLX_1:92; then (2*(n*log(2,n)) + (-1)*(n*log(2,n))) * 4" >= n/2 by XCMPLX_1:8; then (2*(n*log(2,n)) + -(n*log(2,n))) * 4" >= n/2 by XCMPLX_1:180; then (2*(n*log(2,n)) - (n*log(2,n))) * 4" >= n/2 by XCMPLX_0:def 8; then (2*(n*log(2,n)))*4" - (n*log(2,n))*4" >= n/2 by XCMPLX_1:40; then (2*(n*log(2,n))) / 4 - (n*log(2,n))*4" >= n/2 by XCMPLX_0:def 9; then (2*(n*log(2,n))) / (2*2) - (n*log(2,n)) / 4 >= n/2 by XCMPLX_0:def 9 ; then (n*log(2,n)) / 2 - (n*log(2,n)) / 4 >= n/2 by XCMPLX_1:92; then (n*log(2,n)) / 2 >= n/2 + (n*log(2,n)) / 4 by REAL_1:84; then A26: (n*log(2,n))/2 - (n/2) >= n*log(2,n) / 4 by REAL_1:84; n*log(2,n) / 4 = g.n / 4 by A1,A6 .= (1/4)*g.n by XCMPLX_1:100; then Sum(p, n, n2) >= (1/4)*g.n by A24,A25,A26,AXIOMS:22; then Sum(h, n, n2) >= (1/4)*g.n by A20,AXIOMS:22; then Sum(h, n) >= (1/4)*g.n by A17,AXIOMS:22; hence (1/4)*g.n <= f.n by A2,Lm50; A27: log(2,n!) = f.n by A2 .= Sum(h, n) by A2,Lm50; for k st k <= n holds h.k <= q.k proof let k such that A28: k <= n; per cases; suppose k = 0; hence thesis by A12,Def2; suppose A29: k > 0; then log(2,k) <= log(2,n) by A28,PRE_FF:12; then h.k <= log(2,n) by A29,Def2; hence thesis by A12,A29; end; then A30: Sum(h, n) <= Sum(q, n) by Lm18; Sum(q, n) = n*log(2,n) by A12,Lm19; then log(2,n!) <= 1*g.n by A1,A6,A27,A30; hence f.n <= 1*g.n by A2; end; hence thesis by A4,A5; end; :: Problem 3.25 -- Proven in theorem ASYMPT_0:1 begin :: Problem 3.26 :: For a "natural" example of an algorithm that would take time in :: O(t(n)), consider a naive algorithm that tries to find at least one :: factor of a given natural number. For all even numbers, the task is trivial, :: and for all odd numbers, the algorithm might try dividing the number by all :: numbers smaller than it, giving a running time of n. theorem for f being eventually-nondecreasing eventually-nonnegative Real_Sequence, t being Real_Sequence st (for n holds (n mod 2 = 0 implies t.n = 1) & (n mod 2 = 1 implies t.n = n)) holds not t in Big_Theta(f) proof let f be eventually-nondecreasing eventually-nonnegative Real_Sequence, t be Real_Sequence such that A1: for n holds (n mod 2 = 0 implies t.n = 1) & (n mod 2 = 1 implies t.n = n); A2: Big_Theta(f) = { s where s is Element of Funcs(NAT, REAL) : ex c,d,N st c > 0 & d > 0 & for n st n >= N holds d*f.n <= s.n & s.n <= c*f.n } by ASYMPT_0:27; hereby assume t in Big_Theta(f); then consider s being Element of Funcs(NAT, REAL) such that A3: s = t and A4: ex c,d,N st c > 0 & d > 0 & for n st n >= N holds d*f.n <= s.n & s.n <= c*f.n by A2; consider c,d,N such that A5: c > 0 and A6: d > 0 and A7: for n st n >= N holds d*f.n <= s.n & s.n <= c*f.n by A4; consider N0 such that A8: for n st n >= N0 holds f.n <= f.(n+1) by ASYMPT_0:def 8; set N1 = max([/c/d\] + 1, max(N, N0)); A9: N1 >= [/c/d\] + 1 & N1 >= max(N, N0) by SQUARE_1:46; max(N, N0) >= N & max(N, N0) >= N0 by SQUARE_1:46; then A10: N1 >= N & N1 >= N0 by A9,AXIOMS:22; N1 is Integer by SQUARE_1:49; then reconsider N1 as Nat by A9,INT_1:16; thus contradiction proof per cases by GROUP_4:100; suppose A11: N1 mod 2 = 1; then s.N1 = N1 by A1,A3; then N1 <= c*f.N1 by A7,A10; then A12: N1/c <= f.N1 by A5,REAL_2:177; A13: c" > 0 by A5,REAL_1:72; A14: [/c/d\] + 1 > [/c/d\] + 0 by REAL_1:67; [/c/d\] >= c/d by INT_1:def 5; then [/c/d\] + 1 > c/d by A14,AXIOMS:22; then N1 > c/d by A9,AXIOMS:22; then N1*c" > c"*(c/d) by A13,REAL_1:70; then N1/c > c"*(c/d) by XCMPLX_0:def 9; then N1/c > c"*(c*(1/d)) by XCMPLX_1:100; then N1/c > (c"*c)*(1/d) by XCMPLX_1:4; then N1/c > 1*(1/d) by A5,XCMPLX_0:def 7; then A15: f.N1 > 1/d by A12,AXIOMS:22; f.(N1+1) >= f.(N1) by A8,A10; then A16: f.(N1+1) > 1/d by A15,AXIOMS:22; (N1+1) mod 2 = (1+(1 mod 2)) mod 2 by A11,EULER_2:8 .= (1+1) mod 2 by GROUP_4:102 .= 0 by GR_CY_1:5; then A17: t.(N1+1) = 1 by A1; N1+1 > N1+0 by REAL_1:67; then N1+1 > N by A10,AXIOMS:22; then A18: d*f.(N1+1) <= 1 by A3,A7,A17; d*(1/d) < d*f.(N1+1) by A6,A16,REAL_1:70; hence thesis by A6,A18,XCMPLX_1:107; suppose A19: N1 mod 2 = 0; then (N1+1) mod 2 = (0+(1 mod 2)) mod 2 by EULER_2:8 .= (0+1) mod 2 by GROUP_4:102 .= 1 by GROUP_4:102; then A20: s.(N1+1) = N1+1 by A1,A3; A21: N1+1 > N1+0 by REAL_1:67; then N1+1 > N by A10,AXIOMS:22; then N1+1 <= c*f.(N1+1) by A7,A20; then A22: (N1+1)/c <= f.(N1+1) by A5,REAL_2:177; A23: c" > 0 by A5,REAL_1:72; A24: [/c/d\] + 1 > [/c/d\] + 0 by REAL_1:67; [/c/d\] >= c/d by INT_1:def 5; then [/c/d\] + 1 > c/d by A24,AXIOMS:22; then N1 > c/d by A9,AXIOMS:22; then (N1+1) > c/d by A21,AXIOMS:22; then (N1+1)*c" > c"*(c/d) by A23,REAL_1:70; then (N1+1)/c > c"*(c/d) by XCMPLX_0:def 9; then (N1+1)/c > c"*(c*(1/d)) by XCMPLX_1:100; then (N1+1)/c > (c"*c)*(1/d) by XCMPLX_1:4; then (N1+1)/c > 1*(1/d) by A5,XCMPLX_0:def 7; then A25: f.(N1+1) > 1/d by A22,AXIOMS:22; N1+1 > N0 by A10,A21,AXIOMS:22; then f.((N1+1)+1) >= f.(N1+1) by A8; then f.(N1+(1+1)) >= f.(N1+1) by XCMPLX_1:1; then A26: f.(N1+2) > 1/d by A25,AXIOMS:22; (N1+2) mod 2 = (0+(2 mod 2)) mod 2 by A19,EULER_2:8 .= (0+0) mod 2 by GR_CY_1:5 .= 0 by GR_CY_1:6; then A27: t.(N1+2) = 1 by A1; N1+2 > N1+0 by REAL_1:67; then N1+2 > N by A10,AXIOMS:22; then A28: d*f.(N1+2) <= 1 by A3,A7,A27; d*(1/d) < d*f.(N1+2) by A6,A26,REAL_1:70; hence thesis by A6,A28,XCMPLX_1:107; end; end; end; :: Problem 3.27 -- Proven as part of theorem ASYMPT_0:37 begin :: Problem 3.28 definition let f be Function of NAT, REAL*, n be Nat; redefine func f.n -> FinSequence of REAL; coherence proof f.n is Element of REAL*; hence thesis; end; end; Lm52: now let a,b be positive Real; reconsider singlea = <*a*> as Element of REAL* by FINSEQ_1:def 11; defpred P[Nat,FinSequence of REAL,set] means ex n1 being Nat st n1 = [/(($1+1)+1)/2\] & $3 = $2^<*4*$2/.n1 + b*(($1+1)+1)*>; A1: now let n be Nat, x be Element of REAL*; reconsider n1 = [/((n+1)+1)/2\] as Nat by UNIFORM1:12; consider y being FinSequence of REAL such that A2: y = x^<*4*(x/.n1) + b*((n+1)+1)*>; reconsider y as Element of REAL* by FINSEQ_1:def 11; take y; thus P[n,x,y] by A2; end; consider prob28 being Function of NAT, REAL* such that A3: prob28.0 = singlea and A4: for n being Nat holds P[n,prob28.n,prob28.(n+1)] from RecExD(A1); take prob28; thus prob28.0 = <*a*> & for n being Nat holds ex n1 being Nat st n1 = [/((n+1)+1)/2\] & prob28.(n+1) = prob28.n^<*4*((prob28.n)/.n1) + b*((n+1)+1)*> by A3,A4; end; definition let n be Nat, a,b be positive Real; reconsider singlea = <*a*> as Element of REAL* by FINSEQ_1:def 11; defpred P[Nat,FinSequence of REAL,set] means ex n1 being Nat st n1 = [/(($1+1)+1)/2\] & $3 = $2^<*4*($2/.n1) + b*(($1+1)+1)*>; A1: for n being Nat, x,y1,y2 being Element of REAL* st P[n,x,y1] & P[n,x,y2] holds y1 = y2; consider prob28 being Function of NAT, REAL* such that A2: prob28.0 = singlea and A3: for n being Nat holds P[n,prob28.n,prob28.(n+1)] by Lm52; func Prob28 (n,a,b) -> Real means :Def6: it = 0 if n = 0 otherwise ex l being Nat, prob28 being Function of NAT, REAL* st l+1 = n & it = (prob28.l)/.n & prob28.0 = <*a*> & for n being Nat holds ex n1 being Nat st n1 = [/((n+1)+1)/2\] & prob28.(n+1) = prob28.n^<*4*((prob28.n)/.n1) + b*((n+1)+1)*>; consistency; existence proof thus n = 0 implies ex c being Real st c = 0; assume n <> 0; then consider l being Nat such that A4: n = l+1 by NAT_1:22; take (prob28.l)/.n, l, prob28; thus thesis by A2,A3,A4; end; uniqueness proof let c1, c2 be Real; thus n = 0 & c1 = 0 & c2 = 0 implies c1 = c2; assume n <> 0; assume A5: not thesis; then consider l1 being Nat, prob28_1 being Function of NAT, REAL* such that A6: l1+1 = n & c1 = (prob28_1.l1)/.n and A7: prob28_1.0 = singlea & for n being Nat holds P[n,prob28_1.n,prob28_1.(n+1)]; consider l2 being Nat, prob28_2 being Function of NAT, REAL* such that A8: l2+1 = n & c2 = (prob28_2.l2)/.n and A9: prob28_2.0 = singlea & for n being Nat holds P[n,prob28_2.n,prob28_2.(n+1)] by A5; prob28_1 = prob28_2 from RecUnD(A7,A9, A1); hence contradiction by A5,A6,A8,XCMPLX_1:2; end; end; definition let a,b be positive Real; func seq_prob28(a,b) -> Real_Sequence means :Def7: it.n = Prob28(n,a,b); existence proof deffunc _F(Nat) = Prob28($1,a,b); consider h being Function of NAT, REAL such that A1: for n being Nat holds h.n = _F(n) from LambdaD; take h; let n; thus h.n = Prob28(n,a,b) by A1; end; uniqueness proof let j,k be Real_Sequence such that A2: for n holds j.n = Prob28(n,a,b) and A3: for n holds k.n = Prob28(n,a,b); now let n; thus j.n = Prob28(n,a,b) by A2 .= k.n by A3; end; hence j = k by FUNCT_2:113; end; end; Lm53: for n st n >= 2 holds [/n/2\] < n proof let n such that A1: n >= 2; A2: [/n/2\] < n/2 + 1 by INT_1:def 5; now assume n/2 + 1 > n; then 2*(n/2 + 1) > 2*n by REAL_1:70; then 2*(n/2) + 2*1 > 2*n by XCMPLX_1:8; then 2*(2"*n) + 2 > 2*n by XCMPLX_0:def 9; then (2*2")*n + 2 > 2*n by XCMPLX_1:4; then 2 > 2*n - n by REAL_1:84; then 2 > 2*n + -n by XCMPLX_0:def 8; then 2 > 2*n + (-1)*n by XCMPLX_1:180; then 2 > (2+-1)*n by XCMPLX_1:8; hence contradiction by A1; end; hence thesis by A2,AXIOMS:22; end; Lm54: for a,b being positive Real holds Prob28(0,a,b) = 0 & Prob28(1,a,b) = a & for n st n >= 2 holds ex n1 st n1 = [/n/2\] & Prob28(n,a,b) = 4*Prob28(n1,a,b) + b*n proof let a,b be positive Real; defpred P[Nat,FinSequence of REAL,set] means ex n1 being Nat st n1 = [/(($1+1)+1)/2\] & $3 = $2^<*4*($2/.n1) + b*(($1+1)+1)*>; consider prob28 being Function of NAT, REAL* such that A1: prob28.0 = <*a*> and A2: for n being Nat holds P[n,prob28.n,prob28.(n+1)] by Lm52; thus Prob28(0,a,b) = 0 by Def6; 0 <> 1 & 0+1 = 1 & a = <*a*>/.1 by FINSEQ_4:25; hence Prob28(1,a,b) = a by A1,A2,Def6; let n such that A3: n >= 2; set m = n-2; n-2 >= 2-2 by A3,REAL_1:49; then reconsider m as Nat by INT_1:16; consider n1 such that A4: n1 = [/((m+1)+1)/2\] and A5: prob28.(m+1) = prob28.m^<*4*((prob28.m)/.n1) + b*((m+1)+1)*> by A2; A6: (m+1)+1 = m+(1+1) by XCMPLX_1:1 .= (n+-2)+2 by XCMPLX_0:def 8 .= n+(-2+2) by XCMPLX_1:1 .= n; defpred _P[Nat] means len(prob28.$1) = $1+1 & for k being Nat st k <= $1 holds (prob28.$1)/.(k+1) = Prob28(k+1,a,b); A7: _P[0] proof thus len(prob28.0) = 0+1 by A1,FINSEQ_1:57; let k be Nat such that A8: k <= 0; k = 0 by A8; hence (prob28.0)/.(k+1) = Prob28(k+1,a,b) by A1,A2,Def6; end; A9:for n being Nat st _P[n] holds _P[n+1] proof let n be Nat such that A10: len(prob28.n) = n+1 and A11: for k being Nat st k <= n holds (prob28.n)/.(k+1) = Prob28(k+1,a,b); consider n1 being Nat such that A12: n1 = [/((n+1)+1)/2\] & prob28.(n+1) = prob28.n^<*4*((prob28.n)/.n1) + b*((n+1)+1)*> by A2; len <*4*((prob28.n)/.n1) + b*((n+1)+1)*> = 1 by FINSEQ_1:57; hence len(prob28.(n+1)) = (n+1)+1 by A10,A12,FINSEQ_1:35; let k be Nat; assume k <= n+1; then A13: k = n+1 or k <= n by NAT_1:26; now assume A14: k <= n; then 0+1 <= k+1 & k+1 <= n+1 by REAL_1:55; then k+1 in dom (prob28.n) by A10,FINSEQ_3:27; hence (prob28.(n+1))/.(k+1) = (prob28.n)/.(k+1) by A12,GROUP_5:95 .= Prob28(k+1,a,b) by A11,A14; end; hence (prob28.(n+1))/.(k+1) = Prob28(k+1,a,b) by A1,A2,A13,Def6; end; A15: for n being Nat holds _P[n] from Ind(A7, A9); then A16:(prob28.(m+1))/.((m+1)+1) = Prob28(n,a,b) by A6; set n2 = n1-1; A17: (n1-1)+1 = (n1+-1)+1 by XCMPLX_0:def 8 .= n1+(-1+1) by XCMPLX_1:1 .= n1; n*2" > 0*2" by A3,REAL_1:70; then ((m+1)+1)/2 > 0 by A6,XCMPLX_0:def 9; then n1 > 0 by A4,INT_1:def 5; then n1 >= 0 + 1 by INT_1:20; then n2 >= 0 by REAL_1:84; then reconsider n2 as Nat by INT_1:16; A18: len(prob28.m) = m+1 by A15; len(<*4*((prob28.m)/.n1) + b*((m+1)+1)*>) = 1 by FINSEQ_1:57; then A19: 1 in dom <*4*((prob28.m)/.n1) + b*((m+1)+1)*> by FINSEQ_3:27; A20: n1-1 = [/(m+(1+1))/2\] - 1 by A4,XCMPLX_1:1 .= [/m/2 + 2/2\] - 1 by XCMPLX_1:63 .= [/m/2 + 1\] + -1 by XCMPLX_0:def 8 .= [/(m/2 + 1) + -1\] by INT_1:58 .= [/m/2 + (1 + -1)\] by XCMPLX_1:1 .= [/m/2\]; [/m/2\] <= m proof per cases; suppose m = 0; hence thesis by INT_1:54; suppose m > 0; then A21: m >= 0+1 by INT_1:20; thus thesis proof per cases by A21,REAL_1:def 5; suppose A22: m = 1; now assume A23: [/1/2\] <> 1; thus contradiction proof per cases by A23,REAL_1:def 5; suppose [/1/2\] > 1; then A24: [/1/2\] >= 1 + 1 by INT_1:20; [/1/2\] < 1/2 + 1 by INT_1:def 5; hence thesis by A24,AXIOMS:22; suppose [/1/2\] < 1; then [/1/2\] + 1 <= 1 by INT_1:20; then A25: [/1/2\] <= 1 - 1 by REAL_1:84; 1/2 <= [/1/2\] by INT_1:def 5; hence thesis by A25,AXIOMS:22; end; end; hence thesis by A22; suppose m > 1; then m >= 1+1 by INT_1:20; hence thesis by Lm53; end; end; then A26:(prob28.m)/.n1 = Prob28(n2+1,a,b) by A15,A17,A20; (prob28.(m+1))/.n = <*4* ((prob28.m)/.n1)+b*((m+1)+1)*>/.1 by A5,A6,A18,A19,GROUP_5:96 .= 4*Prob28(n1,a,b) + b*n by A6,A17,A26,FINSEQ_4:25; hence thesis by A4,A6,A16; end; theorem for a,b being positive Real holds seq_prob28(a,b) is eventually-nondecreasing proof let a,b be positive Real; set f = seq_prob28(a,b); defpred _R[Nat] means f.$1 <= f.($1+1); defpred _P[Nat] means (for n st n >= 1 & n < $1 holds _R[n]) implies _R[$1]; A1: for k st k >= 1 & (for n st n >= 1 & n < k holds _R[n]) holds _R[k] proof for k st k >= 1 holds _P[k] proof A2: _P[1] proof assume (for n st n >= 1 & n < 1 holds f.n <= f.(n+1)); A3: f.1 = Prob28(1,a,b) by Def7 .= a by Lm54; A4: [/2/2\] = 1 by INT_1:54; consider n1 such that A5: n1 = [/2/2\] and A6: Prob28(2,a,b) = 4*Prob28(n1,a,b) + b*2 by Lm54; A7: f.(1+1) = 4*Prob28(1,a,b) + b*2 by A4,A5,A6,Def7 .= 4*a + b*2 by Lm54; 4*a > 1*a by REAL_1:70; then 4*a + 2*b > a + 0 by REAL_1:67; hence thesis by A3,A7; end; A8: for m st m >= 1 & _P[m] holds _P[m+1] proof let m such that A9: m >= 1 and _P[m]; assume A10: (for n st n >= 1 & n < m+1 holds f.n <= f.(n+1)); A11: (m+1)+0 < (m+1)+1 by REAL_1:67; set t1 = [/(m+1)/2\]; set t2 = [/((m+1)+1)/2\]; (m+1)*2" > 0*2" by REAL_1:70; then A12: (m+1)/2 > 0 by XCMPLX_0:def 9; A13: [/(m+1)/2\] >= (m+1)/2 by INT_1:def 5; A14: t1 > 0 by A12,INT_1:def 5; reconsider t1 as Nat by A13,INT_1:16; [/((m+1)+1)/2\] >= ((m+1)+1)/2 by INT_1:def 5; then reconsider t2 as Nat by INT_1:16; f.t1 <= f.t2 proof per cases by GROUP_4:100; suppose (m+1) mod 2 = 0; then ((m qua Integer)+1) mod 2 = 0 by SCMFSA9A:5; then 0 = (m+1) - (((m qua Integer)+1) div 2)*2 by INT_1:def 8 .= (m+1) + -(((m qua Integer)+1) div 2)*2 by XCMPLX_0:def 8 ; then 0 -(-(((m qua Integer)+1) div 2) * 2) = m+1 by XCMPLX_1: 26 ; then 0+-(-(((m qua Integer)+1) div 2) * 2) = m+1 by XCMPLX_0: def 8 ; then 2 divides (m qua Integer)+1 by INT_1:def 9; then 2 divides ((m qua Integer)+1)+2 by WSIERP_1:9; then 2 divides (m qua Integer)+(2+1) by XCMPLX_1:1; then reconsider t4 = (m+3)/2 as Integer by WSIERP_1:22; A15: t4 - 1 is Integer; A16: (m+3)/2 - 1 = (m+3)/2 - 2/2 .= ((m+3)-2)/2 by XCMPLX_1:121 .= ((m+3)+-2)/2 by XCMPLX_0:def 8 .= (m+(3+-2))/2 by XCMPLX_1:1 .= (m+1)/2; t2 >= ((m+1)+1)/2 by INT_1:def 5; then A17: t2 >= (m+(1+1))/2 by XCMPLX_1:1; t2 < ((m+1)+1)/2 + 2/2 by INT_1:def 5; then t2 < (((m+1)+1)+2)/2 by XCMPLX_1:63; then t2 < ((m+1)+(1+2))/2 by XCMPLX_1:1; then A18: t2 < (m+(1+3))/2 by XCMPLX_1:1; A19: now assume A20: t2 <> t4; thus contradiction proof per cases by A20,REAL_1:def 5; suppose t2 > t4; then t2 >= (m+3)/2 + 2/2 by INT_1:20; then t2 >= ((m+3)+2)/2 by XCMPLX_1:63; then A21: t2 >= (m+(3+2))/2 by XCMPLX_1:1; m+5 > m+4 by REAL_1:67; then (m+5)*2" > (m+4)*2" by REAL_1:70; then (m+5)/2 > (m+4)*2" by XCMPLX_0:def 9; then (m+5)/2 > (m+4)/2 by XCMPLX_0:def 9; hence thesis by A18,A21,AXIOMS:22; suppose t2 < t4; then t2 + 1 <= t4 by INT_1:20; then A22: t2 <= (m+1)/2 by A16,REAL_1:84; m+1 < m+2 by REAL_1:67; then (m+1)*2" < (m+2)*2" by REAL_1:70; then (m+1)/2 < (m+2)*2" by XCMPLX_0:def 9; then (m+1)/2 < (m+2)/2 by XCMPLX_0:def 9; hence thesis by A17,A22,AXIOMS:22; end; end; (m+1)/2 + 1 = (m+1)/2 + 2/2 .= ((m+1)+2)/2 by XCMPLX_1:63 .= (m+(1+2))/2 by XCMPLX_1:1 .= t4; then A23: t1 + 1 = t2 by A15,A16,A19,INT_1:54; A24: t1 >= 1+0 by A14,NAT_1:38; m+1 >= 1+1 by A9,AXIOMS:24; then t1 < m+1 by Lm53; hence thesis by A10,A23,A24; suppose (m+1) mod 2 = 1; then A25: ((m qua Integer)+1) mod 2 = 1 by SCMFSA9A:5; t1 < (m+1)/2 + 2/2 by INT_1:def 5; then t1 < ((m+1)+2)/2 by XCMPLX_1:63; then A26: t1 < (m+(1+2))/2 by XCMPLX_1:1; A27: 2 divides ((m qua Integer)+2) iff (m+2)/2 is Integer by WSIERP_1:22; 1 = (m+1) - (((m qua Integer)+1) div 2) * 2 by A25,INT_1:def 8 .= (m+1) + -(((m qua Integer)+1) div 2) * 2 by XCMPLX_0:def 8 ; then 1 -(-(((m qua Integer)+1) div 2)*2) = (m+1) by XCMPLX_1: 26 ; then 1+-(-(((m qua Integer)+1) div 2)*2) = (m+1) by XCMPLX_0: def 8 ; then -(-(((m qua Integer)+1) div 2)*2) = (m+1)-1 by XCMPLX_1: 26 ; then (((m qua Integer)+1) div 2)*2 = (m+1) + -1 by XCMPLX_0: def 8 ; then (((m qua Integer)+1) div 2)*2 = m+(1 + -1) by XCMPLX_1:1 ; then 2 divides (m qua Integer) by INT_1:def 9; then reconsider t3 = (m+2)/2 as Integer by A27,WSIERP_1:9; A28: now assume A29: t1 <> t3; thus contradiction proof per cases by A29,REAL_1:def 5; suppose t1 > t3; then t1 >= t3 + 2/2 by INT_1:20; then t1 >= ((m+2)+2)/2 by XCMPLX_1:63; then A30: t1 >= (m+(2+2))/2 by XCMPLX_1:1; m+4 > m+3 by REAL_1:67; then (m+4)*2" > (m+3)*2" by REAL_1:70; then (m+4)/2 > (m+3)*2" by XCMPLX_0:def 9; then (m+4)/2 > (m+3)/2 by XCMPLX_0:def 9; hence thesis by A26,A30,AXIOMS:22; suppose t1 < t3; then t1 + 1 <= t3 by INT_1:20; then t1 <= t3 - 2/2 by REAL_1:84; then t1 <= ((m+2)-2)/2 by XCMPLX_1:121; then t1 <= ((m+2)+-2)/2 by XCMPLX_0:def 8; then A31: t1 <= (m+(2+-2))/2 by XCMPLX_1:1; m+0 < m+1 by REAL_1:67; then m*2" < (m+1)*2" by REAL_1:70; then m/2 < (m+1)*2" by XCMPLX_0:def 9; then m/2 < (m+1)/2 by XCMPLX_0:def 9; hence thesis by A13,A31,AXIOMS:22; end; end; t2 = [/(m+(1+1))/2\] by XCMPLX_1:1 .= [/t3\]; hence thesis by A28,INT_1:54; end; then A32: 4*(f.t1) <= 4*(f.t2) by AXIOMS:25; A33: b*(m+1) < b*((m+1)+1) by A11,REAL_1:70; A34: m+1 >= 1+1 by A9,AXIOMS:24; then consider n1 such that A35: n1 = [/(m+1)/2\] and A36: Prob28(m+1,a,b) = 4*Prob28(n1,a,b) + b*(m+1) by Lm54; A37: f.(m+1) = 4*Prob28(n1,a,b) + b*(m+1) by A36,Def7 .= 4*(f.t1) + b*(m+1) by A35,Def7; (m+1)+1 >= 2 by A11,A34,AXIOMS:22; then consider n2 such that A38: n2 = [/((m+1)+1)/2\] and A39: Prob28((m+1)+1,a,b) = 4*Prob28(n2,a,b) + b*((m+1)+1) by Lm54; f.((m+1)+1) = 4*Prob28(n2,a,b) + b*((m+1)+1) by A39,Def7 .= 4*(f.t2) + b*((m+1)+1) by A38,Def7; hence f.(m+1) <= f.((m+1)+1) by A32,A33,A37,REAL_1:67; end; thus thesis from Ind1(A2, A8); end; hence thesis; end; for n st n >= 1 holds _R[n] from Comp_Ind1(A1); hence thesis by ASYMPT_0:def 8; end; :: Problem 3.29 -- Proven in theorems ASYMPT_0:39 begin :: Problem 3.30 definition func POWEROF2SET -> non empty Subset of NAT equals :Def8: { 2 to_power n where n is Nat : not contradiction }; coherence proof set IT = { 2 to_power n where n is Nat : not contradiction }; A1: 2 to_power 1 in IT; now let x; assume x in IT; then consider n being Nat such that A2: 2 to_power n = x and not contradiction; thus x in NAT by A2; end; hence thesis by A1,TARSKI:def 3; end; end; Lm55: for n st n >= 2 holds n^2 > n+1 proof defpred _P[Nat] means $1^2 > $1+1; 2^2 = 2*2 by SQUARE_1:def 3 .= 4; then A1: _P[2]; A2: for k st k >= 2 & _P[k] holds _P[k+1] proof let k such that A3: k >= 2 and A4: k^2 > k+1; A5: (k+1)^2 = k^2 + 2*k + 1 by SQUARE_1:65; k^2 + (2*k + 1) > (k+1) + (2*k + 1) by A4,REAL_1:53; then A6: (k+1)^2 > (k+1) + (2*k + 1) by A5,XCMPLX_1:1; 2*k > 2*0 by A3,REAL_1:70; then 2*k + 1 > 0 + 1 by REAL_1:53; then (k+1) + (2*k + 1) > (k+1) + 1 by REAL_1:53; hence (k+1)^2 > (k+1) + 1 by A6,AXIOMS:22; end; for n st n >= 2 holds _P[n] from Ind1(A1, A2); hence thesis; end; Lm56: for n st n >= 1 holds (2 to_power (n+1)) - (2 to_power n) > 1 proof let n such that A1: n >= 1; 2 to_power n >= 2 to_power 1 by A1,PRE_FF:10; then (2 to_power n)*1 >= 2 by POWER:30; then (2 to_power n)*(2-1) > 1 by AXIOMS:22; then (2 to_power n)*2 - (2 to_power n)*1 > 1 by XCMPLX_1:40; then (2 to_power n)*(2 to_power 1) - (2 to_power n) > 1 by POWER:30; hence (2 to_power (n+1)) - (2 to_power n) > 1 by POWER:32; end; Lm57: for n st n >= 2 holds not ((2 to_power n) - 1) in POWEROF2SET proof let n such that A1: n >= 2; A2:1 = 2-1; assume (2 to_power n) - 1 in POWEROF2SET; then consider m such that A3: 2 to_power m = (2 to_power n) - 1 by Def8; A4: now assume m >= n; then A5: 2 to_power m >= 2 to_power n by PRE_FF:10; (2 to_power n) + 1 > (2 to_power n) + 0 by REAL_1:53; hence contradiction by A3,A5,REAL_1:84; end; A6: n-1 >= 1 by A1,A2,REAL_1:49; then n-1 is Nat by INT_1:16; then (2 to_power ((n-1)+1)) - (2 to_power (n-1)) > 1 by A6,Lm56; then (2 to_power ((n+-1)+1)) - (2 to_power (n-1)) > 1 by XCMPLX_0:def 8; then (2 to_power (n+(-1+1))) - (2 to_power (n-1)) > 1 by XCMPLX_1:1; then (2 to_power n) > 1 + (2 to_power (n-1)) by REAL_1:86; then A7: (2 to_power n) - 1 > (2 to_power (n-1)) by REAL_1:86; then A8: m >= n-1 by A3,POWER:44; m+1 <= n by A4,INT_1:20; then m <= n-1 by REAL_1:84; hence contradiction by A3,A7,A8,AXIOMS:21; end; theorem :: Part a) for f being Real_Sequence st (for n holds (n in POWEROF2SET implies f.n = n) & (not n in POWEROF2SET implies f.n = 2 to_power n)) holds f in Big_Theta(seq_n^(1), POWEROF2SET) & not f in Big_Theta(seq_n^(1)) & seq_n^(1) is smooth & not f is eventually-nondecreasing proof let f be Real_Sequence such that A1: for n holds (n in POWEROF2SET implies f.n = n) & (not n in POWEROF2SET implies f.n = 2 to_power n); set g = seq_n^(1); set h = (g taken_every 2); set p = seq_logn; set q = p/"g; set X = POWEROF2SET; A2: 3 = 4-1; A3: Big_Theta(g,X) = { t where t is Element of Funcs(NAT, REAL) : ex c,d,N st c > 0 & d > 0 & for n st n >= N & n in X holds d*g.n<=t.n & t.n<= c*g.n } by ASYMPT_0:def 17; A4: Big_Theta(g) = { t where t is Element of Funcs(NAT, REAL) : ex c,d,N st c > 0 & d > 0 & for n st n >= N holds d*g.n <= t.n & t.n <= c*g.n } by ASYMPT_0:27; A5: f is Element of Funcs(NAT, REAL) by FUNCT_2:11; now let n such that A6: n >= 1 and A7: n in X; A8: g.n = n to_power 1 by A6,Def3 .= n by POWER:30; hence 1*g.n <= f.n by A1,A7; thus f.n <= 1*g.n by A1,A7,A8; end; hence f in Big_Theta(g,X) by A3,A5; hereby assume f in Big_Theta(g); then consider t being Element of Funcs(NAT, REAL) such that A9: t = f and A10: ex c,d,N st c > 0 & d > 0 & for n st n >= N holds d*g.n <= t.n & t.n <= c*g.n by A4; consider c,d,N such that A11: c > 0 & d > 0 and A12: for n st n >= N holds d*g.n <= t.n & t.n <= c*g.n by A10; q is convergent & lim q = 0 by Lm16; then consider N0 such that A13: for m st m >= N0 holds abs(q.m-0) < 1/2 by SEQ_2:def 7; set N2 = max( max(N0, N), max([/c\], 2) ); A14: N2 >= max(N0, N) & N2 >= max([/c\],2) by SQUARE_1:46; max(N0, N) >= N0 & max(N0, N) >= N by SQUARE_1:46; then A15: N2 >= N0 & N2 >= N by A14,AXIOMS:22; max([/c\],2) >= [/c\] & max([/c\],2) >= 2 by SQUARE_1:46; then A16: N2 >= [/c\] & N2 >= 2 by A14,AXIOMS:22; N2 is Integer proof per cases by SQUARE_1:49; suppose N2 = max(N0, N); hence thesis; suppose N2 = max([/c\],2); hence thesis by SQUARE_1:49; end; then reconsider N2 as Nat by A14,INT_1:16; set N3 = (2 to_power N2) - 1; 2 to_power N2 > 0 by POWER:39; then 2 to_power N2 >= 0 + 1 by NAT_1:38; then N3 >= 0 by REAL_1:84; then reconsider N3 as Nat by INT_1:16; 2 to_power N2 > N2 + 1 by A16,Lm1; then A17: N3 > N2 by REAL_1:86; then A18: N3 >= N by A15,AXIOMS:22; not N3 in POWEROF2SET by A16,Lm57; then A19: t.N3 = 2 to_power N3 by A1,A9; g.N3 = N3 to_power 1 by A17,Def3 .= N3 by POWER:30; then A20: 2 to_power N3 <= c*N3 by A12,A18,A19; 2 to_power N3 > 0 by POWER:39; then log(2,2 to_power N3) <= log(2,c*N3) by A20,PRE_FF:12; then N3*log(2,2) <= log(2,c*N3) by POWER:63; then N3*1 <= log(2,c*N3) by POWER:60; then A21: N3 <= log(2,c) + log(2,N3) by A11,A17,POWER:61; A22: N3 >= [/c\] by A16,A17,AXIOMS:22; [/c\] >= c by INT_1:def 5; then N3 >= c by A22,AXIOMS:22; then log(2,N3) >= log(2,c) by A11,PRE_FF:12; then log(2,N3) + log(2,N3) >= log(2,c) + log(2,N3) by AXIOMS:24; then N3 <= log(2,N3) + log(2,N3) by A21,AXIOMS:22; then N3 <= 2*log(2,N3) by XCMPLX_1:11; then N3 / 2 <= log(2,N3) by REAL_2:177; then N3 * (1/2) <= log(2,N3) by XCMPLX_1:100; then N3" * (N3 * (1/2)) <= log(2,N3) * N3" by AXIOMS:25; then (N3" * N3) * (1/2) <= log(2,N3) * N3" by XCMPLX_1:4; then 1*(1/2) <= log(2,N3) * N3" by A17,XCMPLX_0:def 7; then A23: log(2,N3) / N3 >= 1/2 by XCMPLX_0:def 9; N3 >= N0 by A15,A17,AXIOMS:22; then A24: abs(q.N3-0) < 1/2 by A13; q.N3 = p.N3 / g.N3 by Lm4 .= log(2,N3) / g.N3 by A17,Def2 .= log(2,N3) / (N3 to_power 1) by A17,Def3 .= log(2,N3) / N3 by POWER:30; hence contradiction by A23,A24,ABSVALUE:def 1; end; now let n; assume n >= 0; A25: g.n = n proof per cases; suppose n = 0; hence thesis by Def3; suppose n > 0; hence g.n = n to_power 1 by Def3 .= n by POWER:30; end; A26: g.(n+1) = (n+1) to_power 1 by Def3 .= n+1 by POWER:30; n+0 <= n+1 by AXIOMS:24; hence g.n <= g.(n+1) by A25,A26; end; then A27: g is eventually-nondecreasing by ASYMPT_0:def 8; A28: Big_Oh(g) = { t where t is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds t.n <= c*g.n & t.n >= 0 } by ASYMPT_0:def 12; A29: h is Element of Funcs(NAT, REAL) by FUNCT_2:11; now let n; assume A30: n >= 1; then A31: 2*n > 2*0 by REAL_1:70; A32: h.n = g.(2*n) by ASYMPT_0:def 18 .= (2*n) to_power 1 by A31,Def3 .= 2*n by POWER:30; g.n = n to_power 1 by A30,Def3 .= n by POWER:30; hence h.n <= 2*g.n by A32; thus h.n >= 0 by A32; end; then h in Big_Oh(g) by A28,A29; then g is_smooth_wrt 2 by A27,ASYMPT_0:def 19; hence g is smooth by ASYMPT_0:37; hereby assume f is eventually-nondecreasing; then consider N such that A33: for n st n >= N holds f.n <= f.(n+1) by ASYMPT_0:def 8; set N1 = (2 to_power (N+2)) - 1; A34: N+2 >= 0+2 by AXIOMS:24; then A35: 2 to_power (N+2) >= 2 to_power 2 by PRE_FF:10; 2 to_power 2 = 2^2 by POWER:53 .= 2*2 by SQUARE_1:def 3 .= 4; then N1 >= 3 by A2,A35,REAL_1:49; then reconsider N1 as Nat by INT_1:16; not N1 in POWEROF2SET by A34,Lm57; then A36: f.N1 = 2 to_power N1 by A1; A37: N1+1 = (2 to_power (N+2)) +- 1 + 1 by XCMPLX_0:def 8 .= (2 to_power (N+2)) + (-1 + 1) by XCMPLX_1:1 .= 2 to_power (N+2); then N1+1 in POWEROF2SET by Def8; then A38: f.(N1+1) = N1+1 by A1; 2 to_power (N+2) > (N+2) + 1 by A34,Lm1; then A39: N1 > (N+2) by REAL_1:86; then A40: f.N1 > f.(N1+1) by A36,A37,A38,POWER:44; N+2 >= N+0 by AXIOMS:24; then N1 >= N by A39,AXIOMS:22; hence contradiction by A33,A40; end; end; theorem :: Part b) for f,g being Real_Sequence st (f.0 = 0 & (for n st n > 0 holds f.n = (n to_power (2 to_power [\log(2,n)/] )))) &(g.0 = 0 & (for n st n > 0 holds g.n = (n to_power n))) ex s being eventually-positive Real_Sequence st s = g & f in Big_Theta(s, POWEROF2SET) & not f in Big_Theta(s) & f is eventually-nondecreasing & s is eventually-nondecreasing & not s is_smooth_wrt 2 proof let f,g be Real_Sequence such that A1: f.0 = 0 & (for n st n > 0 holds f.n = (n to_power (2 to_power [\log(2,n)/] ))) and A2: (g.0 = 0 & (for n st n > 0 holds g.n = (n to_power n))); set h = (g taken_every 2); set X = POWEROF2SET; g is eventually-positive proof take 1; let n; assume A3: n >= 1; then g.n = n to_power n by A2; hence g.n > 0 by A3,POWER:39; end; then reconsider g as eventually-positive Real_Sequence; take g; A4: Big_Theta(g,X) = { t where t is Element of Funcs(NAT, REAL) : ex c,d,N st c > 0 & d > 0 & for n st n >= N & n in X holds d*g.n<=t.n & t.n<= c*g.n } by ASYMPT_0:def 17; A5: Big_Theta(g) = { t where t is Element of Funcs(NAT, REAL) : ex c,d,N st c > 0 & d > 0 & for n st n >= N holds d*g.n <= t.n & t.n <= c*g.n } by ASYMPT_0:27; A6: f is Element of Funcs(NAT, REAL) by FUNCT_2:11; A7: now let n such that A8: n >= 1 and A9: n in X; A10: f.n = n to_power (2 to_power [\log(2,n)/]) by A1,A8; consider n1 being Nat such that A11: n = 2 to_power n1 by A9,Def8; log(2,n) = n1*log(2,2) by A11,POWER:63 .= n1*1 by POWER:60; then A12: f.n = n to_power n by A10,A11,INT_1:47; hence 1*g.n <= f.n by A2,A8; thus f.n <= 1*g.n by A2,A8,A12; end; A13: now assume f in Big_Theta(g); then consider t being Element of Funcs(NAT, REAL) such that A14: t = f and A15: ex c,d,N st c > 0 & d > 0 & for n st n >= N holds d*g.n <= t.n & t.n <= c*g.n by A5; consider c,d,N such that A16: c > 0 & d > 0 and A17: for n st n >= N holds d*g.n <= t.n & t.n <= c*g.n by A15; d" > 0 by A16,REAL_1:72; then A18: 1/d > 0 by XCMPLX_1:217; set N1 = max([/1/d\], max(N, 2)); A19: N1 >= [/1/d\] & N1 >= max(N, 2) by SQUARE_1:46; max(N, 2) >= N & max(N, 2) >= 2 by SQUARE_1:46; then A20: N1 >= N & N1 >= 2 by A19,AXIOMS:22; N1 is Integer by SQUARE_1:49; then reconsider N1 as Nat by A19,INT_1:16; reconsider N2 = 2 to_power N1 as Nat; A21: N2 > N1 + 1 by A20,Lm1; N1+1 > N1+0 by REAL_1:67; then A22: N2 > N1 by A21,AXIOMS:22; A23: N2+1 > N2+0 by REAL_1:67; then N2+1 > N1 by A22,AXIOMS:22; then N2+1 > N by A20,AXIOMS:22; then A24: d*g.(N2+1) <= t.(N2+1) by A17; A25: g.(N2+1) = (N2+1) to_power (N2+1) by A2; A26: t.(N2+1) = (N2+1) to_power (2 to_power [\log(2,N2+1)/]) by A1,A14; log(2,N2) = N1*log(2,2) by POWER:63 .= N1*1 by POWER:60; then A27: log(2,N2+1) > N1 by A21,A23,POWER:65; N1 > 1 by A20,AXIOMS:22; then 2 to_power (N1+1) - 2 to_power N1 > 1 by Lm56; then 2 to_power (N1+1) > N2 + 1 by REAL_1:86; then log(2, 2 to_power (N1+1)) > log(2, N2+1) by POWER:65; then (N1+1)*log(2,2) > log(2, N2+1) by POWER:63; then A28: (N1+1)*1 > log(2, N2+1) by POWER:60; [\log(2,N2+1)/] >= [\N1/] by A27,PRE_FF:11; then A29: [\log(2,N2+1)/] >= N1 by INT_1:47; now assume [\log(2,N2+1)/] > N1; then A30: [\log(2,N2+1)/] >= N1 + 1 by INT_1:20; log(2,N2+1) >= [\log(2,N2+1)/] by INT_1:def 4; hence contradiction by A28,A30,AXIOMS:22; end; then A31: g.(N2+1) / t.(N2+1) = (N2+1) to_power (N2+1) / (N2+1) to_power N2 by A25,A26,A29,REAL_1:def 5 .= (N2+1) to_power ((N2+1) - N2) by POWER:34 .= (N2+1) to_power ((1+N2) + -N2) by XCMPLX_0:def 8 .= (N2+1) to_power (1+(N2 + -N2)) by XCMPLX_1:1 .= (N2+1) to_power (1+0) by XCMPLX_0:def 6 .= N2+1 by POWER:30; A32: g.(N2+1) > 0 by A25,POWER:39; [/1/d\] >= 1/d by INT_1:def 5; then N1 >= 1/d by A19,AXIOMS:22; then N2 >= 1/d by A22,AXIOMS:22; then N2+1 > 1/d + 0 by REAL_1:67; then 1 / (g.(N2+1) / t.(N2+1)) < 1/(1/d) by A18,A31,REAL_2:151; then 1 / (g.(N2+1) / t.(N2+1)) < d by XCMPLX_1:56; then t.(N2+1) / g.(N2+1) < d by XCMPLX_1:57; then t.(N2+1) / g.(N2+1) * g.(N2+1) < d*g.(N2+1) by A32,REAL_1:70; hence contradiction by A24,A32,XCMPLX_1:88; end; A33: now let n; assume A34: n >= 2; then A35: n + 1 > n + 0 & n+1 > 0+1 by REAL_1:67; A36: f.n = n to_power (2 to_power [\log(2,n)/]) by A1,A34; A37: f.(n+1) = (n+1) to_power (2 to_power [\log(2,(n+1))/]) by A1; log(2,n) >= log(2,2) by A34,PRE_FF:12; then log(2,n) >= 1 by POWER:60; then [\log(2,n)/] >= [\1/] by PRE_FF:11; then [\log(2,n)/] >= 1 by INT_1:47; then 2 to_power [\log(2,n)/] > 2 to_power 0 by POWER:44; then A38: n to_power (2 to_power [\log(2,n)/]) <= (n+1) to_power (2 to_power [\log(2,n)/]) by A34,A35,POWER:42; log(2,n) <= log(2,(n+1)) by A34,A35,POWER:65; then [\log(2,n)/] <= [\log(2,(n+1))/] by PRE_FF:11; then 2 to_power [\log(2,n)/] <= 2 to_power [\log(2,(n+1))/] by PRE_FF:10; then (n+1) to_power (2 to_power [\log(2,n)/]) <= (n+1) to_power (2 to_power [\log(2,(n+1))/]) by A35,PRE_FF:10; hence f.n <= f.(n+1) by A36,A37,A38,AXIOMS:22; end; A39: now let n; assume A40: n >= 1; A41: n + 1 > n + 0 by REAL_1:67; A42: g.n = n to_power n by A2,A40; A43: g.(n+1) = (n+1) to_power (n+1) by A2; n+1 >= 1+1 by A40,AXIOMS:24; then n+1 > 1 by AXIOMS:22; then A44: (n+1) to_power n < (n+1) to_power (n+1) by A41,POWER:44; n to_power n < (n+1) to_power n by A40,A41,POWER:42; hence g.n <= g.(n+1) by A42,A43,A44,AXIOMS:22; end; now assume g is_smooth_wrt 2; then A45: h in Big_Oh(g) by ASYMPT_0:def 19; Big_Oh(g) = { t where t is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds t.n <= c*g.n & t.n >= 0 } by ASYMPT_0:def 12; then consider t being Element of Funcs(NAT, REAL) such that A46: t = h and A47: ex c,N st c > 0 & for n st n >= N holds t.n <= c*g.n & t.n >= 0 by A45; consider c,N such that c > 0 and A48: for n st n >= N holds t.n <= c*g.n & t.n >= 0 by A47; set N0 = max( [/c\], max(N, 2) ); A49: N0 >= [/c\] & N0 >= max(N, 2) by SQUARE_1:46; A50: max(N, 2) >= N & max(N, 2) >= 2 by SQUARE_1:46; then A51: N0 >= N & N0 >= 2 by A49,AXIOMS:22; then A52: N0 > 0 & N0 > 1 by AXIOMS:22; A53: 2*N0 > 1*N0 by A49,A50,REAL_1:70; N0 is Integer by SQUARE_1:49; then reconsider N0 as Nat by A49,INT_1:16; h.N0 = g.(2*N0) by ASYMPT_0:def 18 .= (2*N0) to_power (2*N0) by A2,A53; then A54: h.N0 > N0 to_power (2*N0) by A49,A50,A53,POWER:42; N0 >= 1 by A51,AXIOMS:22; then N0 + N0 >= N0 + 1 by AXIOMS:24; then 2*N0 >= N0 + 1 by XCMPLX_1:11; then N0 to_power (2*N0) >= N0 to_power (N0+1) by A52,PRE_FF:10; then A55: h.N0 > N0 to_power (N0+1) by A54,AXIOMS:22; A56: N0 to_power (N0+1) = (N0 to_power N0)*(N0 to_power 1) by A49,A50,POWER:32 .= (N0 to_power N0)*N0 by POWER:30; [/c\] >= c by INT_1:def 5; then N0 >= c by A49,AXIOMS:22; then N0 to_power (N0+1) >= c*(N0 to_power N0) by A56,AXIOMS:25; then h.N0 > c*(N0 to_power N0) by A55,AXIOMS:22; then h.N0 > c*g.N0 by A2,A49,A50; hence contradiction by A46,A48,A51; end; hence thesis by A4,A6,A7,A13,A33,A39,ASYMPT_0:def 8; end; theorem :: Part c) for g being Real_Sequence st (for n holds (n in POWEROF2SET implies g.n = n) & (not n in POWEROF2SET implies g.n = n to_power 2)) holds ex s being eventually-positive Real_Sequence st s = g & seq_n^(1) in Big_Theta(s, POWEROF2SET) & not seq_n^(1) in Big_Theta(s) & s taken_every 2 in Big_Oh(s) & seq_n^(1) is eventually-nondecreasing & not s is eventually-nondecreasing proof let g be Real_Sequence such that A1: for n holds (n in POWEROF2SET implies g.n = n) & (not n in POWEROF2SET implies g.n = n to_power 2); set f = seq_n^(1); set h = (g taken_every 2); set X = POWEROF2SET; g is eventually-positive proof take 1; let n; assume A2: n >= 1; thus g.n > 0 proof per cases; suppose n in POWEROF2SET; hence thesis by A1,A2; suppose not n in POWEROF2SET; then g.n = n to_power 2 by A1; hence thesis by A2,POWER:39; end; end; then reconsider g as eventually-positive Real_Sequence; take g; A3: 3 = 4-1; A4: Big_Theta(g,X) = { t where t is Element of Funcs(NAT, REAL) : ex c,d,N st c > 0 & d > 0 & for n st n >= N & n in X holds d*g.n<=t.n & t.n<= c*g.n } by ASYMPT_0:def 17; A5: f is Element of Funcs(NAT, REAL) by FUNCT_2:11; A6: now let n such that A7: n >= 1 and A8: n in X; A9: f.n = n to_power 1 by A7,Def3 .= n by POWER:30; hence 1*g.n <= f.n by A1,A8; thus f.n <= 1*g.n by A1,A8,A9; end; A10: Big_Theta(g) = { t where t is Element of Funcs(NAT, REAL) : ex c,d,N st c > 0 & d > 0 & for n st n >= N holds d*g.n <= t.n & t.n <= c*g.n } by ASYMPT_0:27; A11: now assume f in Big_Theta(g); then consider t being Element of Funcs(NAT, REAL) such that A12: t = f and A13: ex c,d,N st c > 0 & d > 0 & for n st n >= N holds d*g.n <= t.n & t.n <= c*g.n by A10; consider c,d,N such that c > 0 and A14: d > 0 and A15: for n st n >= N holds d*g.n <= t.n & t.n <= c*g.n by A13; set N0 = max(max(N,2), [/1/d\]); A16: N0 >= max(N,2) & N0 >= [/1/d\] by SQUARE_1:46; max(N,2) >= N & max(N,2) >= 2 by SQUARE_1:46; then A17: N0 >= N & N0 >= 2 by A16,AXIOMS:22; N0 is Integer by SQUARE_1:49; then reconsider N0 as Nat by A16,INT_1:16; set N1 = (2 to_power N0) - 1; 2 to_power N0 > 0 by POWER:39; then 2 to_power N0 >= 0 + 1 by NAT_1:38; then N1 >= 0 by REAL_1:84; then reconsider N1 as Nat by INT_1:16; 2 to_power N0 > N0 + 1 by A17,Lm1; then A18: N1 > N0 by REAL_1:86; then A19: N1 >= N by A17,AXIOMS:22; A20: N1 > [/1/d\] by A16,A18,AXIOMS:22; [/1/d\] >= 1/d by INT_1:def 5; then N1 > 1/d by A20,AXIOMS:22; then N1*N1 > N1*(1/d) by A18,REAL_1:70; then N1^2 > N1*(1/d) by SQUARE_1:def 3; then d*N1^2 > (N1*(1/d))*d by A14,REAL_1:70; then d*N1^2 > N1*(1/d*d) by XCMPLX_1:4; then A21: d*N1^2 > N1*1 by A14,XCMPLX_1:88; not N1 in POWEROF2SET by A17,Lm57; then A22: g.N1 = N1 to_power 2 by A1 .= N1^2 by POWER:53; t.N1 = N1 to_power 1 by A12,A18,Def3 .= N1 by POWER:30; hence contradiction by A15,A19,A21,A22; end; A23: Big_Oh(g) = { t where t is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds t.n <= c*g.n & t.n >= 0 } by ASYMPT_0:def 12; A24: h is Element of Funcs(NAT, REAL) by FUNCT_2:11; A25: now let n; assume n >= 0; A26: h.n = g.(2*n) by ASYMPT_0:def 18; thus h.n <= 4*g.n proof per cases; suppose A27: n in POWEROF2SET; then consider m such that A28: n = 2 to_power m by Def8; 2*n = (2 to_power 1)*(2 to_power m) by A28,POWER:30 .= 2 to_power (m+1) by POWER:32; then 2*n in POWEROF2SET by Def8; then A29: g.(2*n) = 2*n by A1; g.n = n by A1,A27; hence thesis by A26,A29,AXIOMS:25; suppose A30: not n in POWEROF2SET; now assume 2*n in POWEROF2SET; then consider m such that A31: 2*n = 2 to_power m by Def8; thus contradiction proof per cases; suppose m = 0; then (n*2)*2" = 1*2" by A31,POWER:29; then A32: n*(2*2") = 1*2" by XCMPLX_1:4; now assume 1/2 is Nat; then 0+1 <= 1/2 by NAT_1:38; hence contradiction; end; hence thesis by A32; suppose m > 0; then m >= 0+1 by NAT_1:38; then A33: m-1 >= 0 by REAL_1:84; 2*n = 2 to_power (m+(-1+1)) by A31 .= 2 to_power ((m+-1)+1) by XCMPLX_1:1 .= 2 to_power ((m-1)+1) by XCMPLX_0:def 8 .= (2 to_power (m-1))*(2 to_power 1) by POWER:32 .= (2 to_power (m-1))*2 by POWER:30; then (2"*2)*n = ((2 to_power (m-1))*2)*2" by XCMPLX_1:4; then A34: n = (2 to_power (m-1))*(2*2") by XCMPLX_1:4; m-1 is Nat by A33,INT_1:16; hence thesis by A30,A34,Def8; end; end; then A35: g.(2*n) = (2*n) to_power 2 by A1 .= (2*n)^2 by POWER:53 .= 2^2*n^2 by SQUARE_1:68 .= (2*2)*n^2 by SQUARE_1:def 3 .= 4*n^2; g.n = n to_power 2 by A1,A30 .= n^2 by POWER:53; hence thesis by A35,ASYMPT_0:def 18; end; thus h.n >= 0 proof per cases; suppose (2*n) in POWEROF2SET; hence thesis by A1,A26; suppose not (2*n) in POWEROF2SET; then g.(2*n) = (2*n) to_power 2 by A1; hence thesis by ASYMPT_0:def 18; end; end; A36: now let n; assume n >= 0; A37: f.n = n proof per cases; suppose n = 0; hence thesis by Def3; suppose n > 0; hence f.n = n to_power 1 by Def3 .= n by POWER:30; end; A38: f.(n+1) = (n+1) to_power 1 by Def3 .= n+1 by POWER:30; n+0 <= n+1 by AXIOMS:24; hence f.n <= f.(n+1) by A37,A38; end; now assume g is eventually-nondecreasing; then consider N such that A39: for n st n >= N holds g.n <= g.(n+1) by ASYMPT_0:def 8; set N0 = max(N, 1); set N1 = (2 to_power (2*N0)) - 1; 2 to_power (2*N0) >= 2 to_power 0 by PRE_FF:10; then 2 to_power (2*N0) >= 1 by POWER:29; then A40: (2 to_power (2*N0)) - 1 >= 1 - 1 by REAL_1:49; A41: N0 >= 1 by SQUARE_1:46; then A42: 2*N0 >= 2*1 by AXIOMS:25; reconsider N1 as Nat by A40,INT_1:16; not N1 in POWEROF2SET by A42,Lm57; then A43: g.N1 = N1 to_power 2 by A1; N1+1 = (2 to_power (2*N0)) +- 1 + 1 by XCMPLX_0:def 8 .= (2 to_power (2*N0)) + (-1 + 1) by XCMPLX_1:1 .= 2 to_power (2*N0); then N1+1 in POWEROF2SET by Def8; then A44: g.(N1+1) = N1+1 by A1; A45: 2*N0 >= 2*1 by A41,AXIOMS:25; then 2 to_power (2*N0) >= 2 to_power 2 by PRE_FF:10; then 2 to_power (2*N0) >= 2^2 by POWER:53; then 2 to_power (2*N0) >= 2*2 by SQUARE_1:def 3; then N1 >= 3 by A3,REAL_1:49; then N1 >= 2 by AXIOMS:22; then N1^2 > N1+1 by Lm55; then A46: g.N1 > g.(N1+1) by A43,A44,POWER:53; A47: N0 >= N by SQUARE_1:46; 2 to_power (2*N0) > 2*N0 + 1 by A45,Lm1; then A48: N1 > 2*N0 by REAL_1:86; 2*N0 >= 1*N0 by AXIOMS:25; then N1 >= N0 by A48,AXIOMS:22; then N1 >= N by A47,AXIOMS:22; hence contradiction by A39,A46; end; hence thesis by A4,A5,A6,A11,A23,A24,A25,A36,ASYMPT_0:def 8; end; begin :: Problem 3.31 Lm58: for n st n >= 2 holds n! > 1 proof defpred _P[Nat] means $1! > 1; A1: _P[2] by NEWTON:20; A2: for k st k >= 2 & _P[k] holds _P[k+1] proof let k such that A3: k >= 2 and A4: k! > 1; A5: (k+1)*(k!) > (k+1)*1 by A4,REAL_1:70; k+1 > 0+1 by A3,REAL_1:53; then (k+1)*(k!) > 1 by A5,AXIOMS:22; hence thesis by NEWTON:21; end; for n st n >= 2 holds _P[n] from Ind1(A1, A2); hence thesis; end; Lm59: for n1, n st n <= n1 holds n! <= n1! proof defpred _P[Nat] means for n st n <= $1 holds n! <= ($1)!; A1: _P[0] by AXIOMS:21; A2: for k st _P[k] holds _P[k+1] proof let k such that A3: for n st n <= k holds n! <= k!; let n such that A4: n <= k+1; per cases by A4,NAT_1:26; suppose n <= k; then A5: n! <= k! by A3; A6: k+1 >= 0+1 by AXIOMS:24; k! >= 0 by NEWTON:23; then (k+1)*(k!) >= 1*(k!) by A6,AXIOMS:25; then (k+1)! >= k! by NEWTON:21; hence n! <= (k+1)! by A5,AXIOMS:22; suppose n = k+1; hence thesis; end; for n1 holds _P[n1] from Ind(A1, A2); hence thesis; end; Lm60: for k st k >= 1 holds ex n st (n! <= k & k < (n+1)! & for m st m! <= k & k < (m+1)! holds m = n) proof defpred _P[Nat] means ex n st (n! <= $1 & $1 < (n+1)! & for m st m! <= $1 & $1 < (m+1)! holds m = n); A1: _P[1] proof take 1; thus 1! <= 1 & 1 < (1+1)! by NEWTON:19,20; let m; assume A2: m! <= 1 & 1 < (m+1)!; A3: now assume m > 1; then m >= 1+1 by NAT_1:38; hence contradiction by A2,Lm58; end; m <> 0 by A2,NEWTON:19; then m > 0; then m >= 0+1 by NAT_1:38; hence m = 1 by A3,AXIOMS:21; end; A4:for k st k >= 1 & _P[k] holds _P[k+1] proof let k; assume that k >= 1 and A5: ex n st (n! <= k & k < (n+1)! & for m st m! <= k & k < (m+1)! holds m = n); consider n such that A6: n! <= k & k < (n+1)! and for m st m! <= k & k < (m+1)! holds m = n by A5; (n+1)! is Nat by NEWTON:22; then A7:k+1 <= (n+1)! by A6,INT_1:20; per cases by A7,REAL_1:def 5; suppose A8: k+1 < (n+1)!; take n; k+0 <= k+1 by AXIOMS:24; hence n! <= k+1 by A6,AXIOMS:22; thus k+1 < (n+1)! by A8; let m; assume that A9: m! <= k+1 and A10: k+1 < (m+1)!; now assume A11:m <> n; thus contradiction proof per cases by A11,REAL_1:def 5; suppose m > n; then m >= n+1 by NAT_1:38; then m! >= (n+1)! by Lm59; hence thesis by A8,A9,AXIOMS:22; suppose m < n; then m+1 <= n by NAT_1:38; then (m+1)! <= n! by Lm59; then A12: (m+1)! <= k by A6,AXIOMS:22; k <= k+1 by NAT_1:29; hence thesis by A10,A12,AXIOMS:22; end; end; hence m = n; suppose A13: k+1 = (n+1)!; take N = n+1; thus N! <= k+1 by A13; A14: N+1 > 0+1 by REAL_1:53; N! > 0 by NEWTON:23; then (N+1)*(N!) > 1*(N!) by A14,REAL_1:70; hence k+1 < (N+1)! by A13,NEWTON:21; let m; assume that A15: m! <= k+1 and A16: k+1 < (m+1)!; now assume A17:m <> N; thus contradiction proof per cases by A17,REAL_1:def 5; suppose m > N; then m >= N+1 by NAT_1:38; then m! >= (N+1)! by Lm59; then A18: k+1 >= (N+1)! by A15,AXIOMS:22; n + 2 >= 0 + 2 by AXIOMS:24; then n + (1+1) >= 2; then N+1 >= 2 by XCMPLX_1:1; then A19: N+1 > 1 by AXIOMS:22; N! > 0 by NEWTON:23; then (N+1)*(N!) > 1*(N!) by A19,REAL_1:70; hence thesis by A13,A18,NEWTON:21; suppose m < N; then m+1 <= N by NAT_1:38; hence thesis by A13,A16,Lm59; end; end; hence thesis; end; for k st k >= 1 holds _P[k] from Ind1(A1,A4); hence thesis; end; definition let x be Nat; func Step1(x) -> Nat means :Def9: ex n st n! <= x & x < (n+1)! & it = n! if x <> 0 otherwise it = 0; consistency; existence proof hereby assume x <> 0; then x > 0; then x >= 0+1 by NAT_1:38; then consider k being Nat such that A1: k! <= x & x < (k+1)! and for m st m! <= x & x < (m+1)! holds m = k by Lm60; consider k1 being Real such that A2: k1 = k!; reconsider k1 as Nat by A2,NEWTON:22; take k1; thus ex m st m! <= x & x < (m+1)! & k1 = m! by A1,A2; end; thus thesis; end; uniqueness proof let n1, n2 be Nat; now assume that A3: (ex n st n! <= x & x < (n+1)! & n1 = n!) and A4: (ex n st n! <= x & x < (n+1)! & n2 = n!); consider n such that A5: n! <= x & x < (n+1)! & n1 = n! by A3; consider m such that A6: m! <= x & x < (m+1)! & n2 = m! by A4; now assume A7:m <> n; thus contradiction proof per cases by A7,REAL_1:def 5; suppose m > n; then m >= n+1 by INT_1:20; then m! >= (n+1)! by Lm59; hence thesis by A5,A6,AXIOMS:22; suppose m < n; then m+1 <= n by INT_1:20; then (m+1)! <= n! by Lm59; hence thesis by A5,A6,AXIOMS:22; end; end; hence n1 = n2 by A5,A6; end; hence thesis; end; end; Lm61: for n st n >= 3 holds n! > n proof let n; assume A1: n >= 3; A2: 2 = 3-1; set n1 = n-1; A3: n1 >= 2 by A1,A2,REAL_1:49; then reconsider n1 as Nat by INT_1:16; A4: n1+1 = (n+-1)+1 by XCMPLX_0:def 8 .= n+(-1+1) by XCMPLX_1:1 .= n; n1! >= 2 by A3,Lm59,NEWTON:20; then n1! > 1 by AXIOMS:22; then n*(n1!) > n*1 by A1,REAL_1:70; hence thesis by A4,NEWTON:21; end; Lm62: for n,k st k >= 1 & n! <= k & k < (n+1)! holds Step1(k) = n! proof let n,k such that A1: k >= 1 and A2: n! <= k & k < (n+1)!; consider n1 such that A3: n1! <= k & k < (n1+1)! and A4: Step1(k) = n1! by A1,Def9; consider n2 such that n2! <= k & k < (n2+1)! and A5: for m st m! <= k & k < (m+1)! holds m = n2 by A1,Lm60; n = n2 by A2,A5; hence thesis by A3,A4,A5; end; theorem for f being Real_Sequence st (for n holds f.n = Step1(n)) holds ex s being eventually-positive Real_Sequence st s = f & f is eventually-nondecreasing & (for n holds f.n <= (seq_n^(1)).n) & not s is smooth proof let f be Real_Sequence such that A1: (for n holds f.n = Step1(n)); set g = seq_n^(1); f is eventually-positive proof take 1; let n; assume A2: n >= 1; A3: f.n = Step1(n) by A1; consider m such that A4: m! <= n & n < (m+1)! & Step1(n) = m! by A2,Def9; thus f.n > 0 by A3,A4,NEWTON:23; end; then reconsider f as eventually-positive Real_Sequence; take f; A5: 1 = 2-1; now let k; thus f.k <= f.(k+1) proof per cases; suppose A6: k = 0; A7: f.0 = Step1(0) by A1 .= 0 by Def9; f.(0+1) = Step1(1) by A1; hence thesis by A6,A7; suppose A8: k > 0; A9: k <= k+1 by NAT_1:29; A10: k+1 >= 1 by NAT_1:29; consider n1 such that A11: n1! <= k & k < (n1+1)! and A12: Step1(k) = n1! by A8,Def9; A13: f.k = n1! by A1,A12; (n1+1)! is Nat by NEWTON:22; then A14: k+1 <= (n1+1)! by A11,INT_1:20; thus thesis proof per cases by A14,REAL_1:def 5; suppose k+1 < (n1+1)!; then n1! <= k+1 & k+1 < (n1+1)! by A9,A11,AXIOMS:22; then Step1(k+1) = n1! by A10,Lm62; hence thesis by A1,A13; suppose A15: k+1 = (n1+1)!; A16: n1 + 2 > 0 + 1 by REAL_1:67; (n1+1)! > 0 by NEWTON:23; then 1*((n1+1)!) < (n1+2)*((n1+1)!) by A16,REAL_1:70; then (n1+1)! < (n1+(1+1))*((n1+1)!); then (n1+1)! < ((n1+1)+1)*((n1+1)!) by XCMPLX_1:1; then A17: k+1 < ((n1+1)+1)! by A15,NEWTON:21; A18: f.(k+1) = Step1(k+1) by A1 .= (n1+1)! by A10,A15,A17,Lm62; n1 <= n1+1 by NAT_1:29; hence thesis by A13,A18,Lm59; end; end; end; then A19: for k st k >= 0 holds f.k <= f.(k+1); A20: now let n; thus f.n <= g.n proof per cases; suppose A21: n = 0; f.0 = Step1(0) by A1 .= 0 by Def9; hence thesis by A21,Def3; suppose A22: n > 0; then consider n1 such that A23: n1! <= n & n < (n1+1)! and A24: Step1(n) = n1! by Def9; g.n = n to_power 1 by A22,Def3 .= n by POWER:30; hence thesis by A1,A23,A24; end; end; now assume A25: f is smooth; set h = (f taken_every 2); f is_smooth_wrt 2 by A25,ASYMPT_0:def 20; then A26: h in Big_Oh(f) by ASYMPT_0:def 19; Big_Oh(f) = { t where t is Element of Funcs(NAT, REAL) : ex c,N st c > 0 & for n st n >= N holds t.n <= c*f.n & t.n >= 0 } by ASYMPT_0:def 12; then consider t being Element of Funcs(NAT, REAL) such that A27: t = h and A28: ex c,N st c > 0 & for n st n >= N holds t.n <= c*f.n & t.n >= 0 by A26; consider c,N such that c > 0 and A29: for n st n >= N holds t.n <= c*f.n & t.n >= 0 by A28; set n2 = max(max(N,3), [/c\]+1); A30: n2 >= max(N,3) & n2 >= [/c\]+1 by SQUARE_1:46; A31: max(N,3) >= N & max(N,3) >= 3 by SQUARE_1:46; then A32:n2 >= N & n2 >= 3 by A30,AXIOMS:22; n2 is Integer by SQUARE_1:49; then reconsider n2 as Nat by A30,INT_1:16; set n1 = n2!-1; A33: n2 > 2 by A32,AXIOMS:22; then A34:n2! >= 2 by Lm59,NEWTON:20; then A35: n1 >= 1 by A5,REAL_1:49; A36: n2! is Nat by NEWTON:22; then n2! - 1 is Integer by INT_1:17; then reconsider n1 as Nat by A35,INT_1:16; set n3 = n2-1; A37:n3 >= 1 by A5,A33,REAL_1:49; then reconsider n3 as Nat by INT_1:16; n2! > n2 by A32,Lm61; then n2! >= n2 + 1 by A36,INT_1:20; then n1 >= n2 by REAL_1:84; then n1 >= N by A32,AXIOMS:22; then A38: t.n1 <= c*f.n1 by A29; A39:n2!-1 < n2!-0 by REAL_1:92; A40:n3+1 = (n2+-1)+1 by XCMPLX_0:def 8 .= n2+(-1+1) by XCMPLX_1:1 .= n2+0; then A41:n2*(n3!) = n2! by NEWTON:21; A42:n3! >= 1 by A37,Lm59,NEWTON:19; 1+1 <= n2 by A32,AXIOMS:22; then 1 <= n2-1 by REAL_1:84; then 1*1 <= (n2-1)*(n3!) by A42,Lm26; then n2*1 <= (n2-1)*(n3!)*n2 by AXIOMS:25; then n2 <= (n2-1)*((n3!)*(n3+1)) by A40,XCMPLX_1:4; then n2 <= (n2-1)*(n2!) by A40,NEWTON:21; then n2! + n2 <= (n2!)*1 + (n2-1)*(n2!) by AXIOMS:24; then n2! + n2 <= (n2!)*((n2-1)+1) by XCMPLX_1:8; then n2! + n2 <= (n2!)*((n2+-1)+1) by XCMPLX_0:def 8; then n2! + n2 <= (n2!)*(n2+(-1+1)) by XCMPLX_1:1; then n2*(n3!) <= n2*(n2!) - n2 by A41,REAL_1:84; then n3! <= (n2*(n2!) - n2) / n2 by A30,A31,REAL_2:177; then n3! <= (n2*(n2!-1))/(n2*1) by XCMPLX_1:40; then A43: n3! <= (n2! - 1)/1 by A30,A31,XCMPLX_1:92; A44: f.n1 = Step1(n1) by A1 .= n3! by A35,A39,A40,A43,Lm62; A45: t.n1 = f.(2*n1) by A27,ASYMPT_0:def 18; A46:2*(n2!) <= n2*(n2!) by A33,A34,AXIOMS:25; n2 < n2+1 by NAT_1:38; then n2*(n2!) < (n2+1)*(n2!) by A34,REAL_1:70; then n2*(n2!) < (n2+1)! by NEWTON:21; then A47:2*(n2!) < (n2+1)! by A46,AXIOMS:22; 2*n1 < 2*(n2!) by A39,REAL_1:70; then A48:2*n1 < (n2+1)! by A47,AXIOMS:22; n2! + 2 <= n2! + n2! by A34,AXIOMS:24; then n2! + 2 <= 2*(n2!) by XCMPLX_1:11; then n2! <= 2*(n2!) - 2*1 by REAL_1:84; then A49:n2! <= 2*n1 by XCMPLX_1:40; 2*n1 >= 1*n1 by AXIOMS:25; then A50:2*n1 >= 1 by A35,AXIOMS:22; A51:n2 = n2+(-1+1) .= (n2+-1)+1 by XCMPLX_1:1 .= n3+1 by XCMPLX_0:def 8; A52: f.(2*n1) = Step1(2*n1) by A1 .= (n3+1)! by A48,A49,A50,A51,Lm62 .= (n3+1)*(n3!) by NEWTON:21 .= ((n2+-1)+1)*(n3!) by XCMPLX_0:def 8 .= (n2+(-1+1))*(n3!) by XCMPLX_1:1 .= n2*(n3!); A53: [/c\] + 1 > [/c\] + 0 by REAL_1:67; [/c\] >= c by INT_1:def 5; then [/c\] + 1 > c by A53,AXIOMS:22; then A54: n2 > c by A30,AXIOMS:22; n3! > 0 by NEWTON:23; hence contradiction by A38,A44,A45,A52,A54,REAL_1:70; end; hence thesis by A19,A20,ASYMPT_0:def 8; end; :: Problem 3.32 -- omitted (corresponding theory section omitted) :: Problem 3.33 -- Proven in theorems ASYMPT_0:41, ASYMPT_0:42 begin :: Problem 3.34 definition cluster (seq_n^(1) - seq_const(1)) -> eventually-positive; coherence proof set f = seq_n^(1); set g = seq_const(1); take 2; let n; assume A1: n >= 2; then A2: n > 1+0 by AXIOMS:22; A3: f.n = n to_power 1 by A1,Def3 .= n by POWER:30; A4: g.n = (NAT --> 1).n by Def4 .= 1 by FUNCOP_1:13; (f-g).n = (f +- g).n by SEQ_1:15 .= f.n + (-g).n by SEQ_1:11 .= n + -1 by A3,A4,SEQ_1:14 .= n-1 by XCMPLX_0:def 8; hence (f-g).n > 0 by A2,REAL_1:86; end; end; :: This is to show that Big_Theta(n-1) + Big_Theta(n) = Big_Theta(n). :: Note that it is not true that Big_Theta(n) = Big_Theta(n) - Big_Theta(n-1). :: Consider n and n. The function n - n = 0, which is not in Big_Theta(n). theorem Big_Theta(seq_n^(1) - seq_const(1)) + Big_Theta(seq_n^(1)) = Big_Theta(seq_n^(1)) proof set p = seq_n^(1); set q = seq_const(1); set f = p - q; set g = p; A1: Big_Theta(f) + Big_Theta(g) = { t where t is Element of Funcs(NAT,REAL) : ex f',g' being Element of Funcs(NAT,REAL) st f' in Big_Theta(f) & g' in Big_Theta(g) & for n being Element of NAT holds t.n = f'.n + g'.n } by ASYMPT_0:def 21; A2: Big_Theta(f) = { t where t is Element of Funcs(NAT, REAL) : ex c,d,N st c > 0 & d > 0 & for n st n >= N holds d*f.n <= t.n & t.n <= c*f.n } by ASYMPT_0:27; A3: Big_Theta(g) = { t where t is Element of Funcs(NAT, REAL) : ex c,d,N st c > 0 & d > 0 & for n st n >= N holds d*g.n <= t.n & t.n <= c*g.n } by ASYMPT_0:27; now let x; hereby assume x in Big_Theta(f) + Big_Theta(g); then consider t being Element of Funcs(NAT, REAL) such that A4: t = x and A5: ex f',g' being Element of Funcs(NAT,REAL) st f' in Big_Theta(f) & g' in Big_Theta(g) & for n being Element of NAT holds t.n = f'.n + g'.n by A1; consider f',g' being Element of Funcs(NAT,REAL) such that A6: f' in Big_Theta(f) & g' in Big_Theta(g) and A7: for n being Element of NAT holds t.n = f'.n + g'.n by A5; consider r being Element of Funcs(NAT,REAL) such that A8: r = f' and A9: ex c,d,N st c > 0 & d > 0 & for n st n >= N holds d*f.n <= r.n & r.n <= c*f.n by A2,A6; consider s being Element of Funcs(NAT,REAL) such that A10: s = g' and A11: ex c,d,N st c > 0 & d > 0 & for n st n >= N holds d*g.n <= s.n & s.n <= c*g.n by A3,A6; consider c1,d1,N1 such that A12: c1 > 0 & d1 > 0 and A13: for n st n >= N1 holds d1*f.n <= f'.n & f'.n <= c1*f.n by A8,A9; consider c2,d2,N2 such that A14: c2 > 0 & d2 > 0 and A15: for n st n >= N2 holds d2*g.n <= g'.n & g'.n <= c2*g.n by A10,A11; set N = max(1,max(N1,N2)); A16: N >= 1 & N>=max(N1,N2) & max(N1,N2)>=N1 & max(N1,N2)>= N2 by SQUARE_1:46; then A17: N >= 1 & N >= N1 & N >= N2 by AXIOMS:22; set d = d2, c = c1+c2; c1 + c2 >= 0 + c2 by A12,AXIOMS:24; then A18: c > 0 & d > 0 by A14; now let n; assume A19: n >= N; then A20: n >= 1 & n >= N1 & n >= N2 by A17,AXIOMS:22; then d1*f.n <= f'.n by A13; then A21:d1*f.n + g'.n <= f'.n + g'.n by AXIOMS:24; d2*g.n <= g'.n by A15,A20; then d1*f.n + d2*g.n <= d1*f.n + g'.n by AXIOMS:24; then A22:d1*f.n + d2*g.n <= f'.n + g'.n by A21,AXIOMS:22; A23: g.n = (n to_power 1) by A16,A19,Def3 .= n by POWER:30; A24: (seq_const(1)).n = (NAT --> 1).n by Def4 .= 1 by FUNCOP_1:13; A25: f.n = (seq_n^(1) +- seq_const(1)).n by SEQ_1:15 .= (seq_n^(1)).n + (-seq_const(1)).n by SEQ_1:11 .= (seq_n^(1)).n + -((seq_const(1)).n) by SEQ_1:14 .= (n to_power 1) + -((seq_const(1)).n) by A16,A19,Def3 .= n + -1 by A24,POWER:30; then A26:d1*f.n + d2*g.n = d1*n + d1*(-1) + d2*n by A23,XCMPLX_1:8 .= d1*n + -(d1*1) + d2*n by XCMPLX_1:175 .= -d1 + (d1*n + d2*n) by XCMPLX_1:1 .= -d1 + (d1+d2)*n by XCMPLX_1:8; A27:n*(-d1) + (d1+d2)*n = (-d1+(d1+d2))*n by XCMPLX_1:8 .= ((-d1+d1)+d2)*n by XCMPLX_1:1 .= (0+d2)*n by XCMPLX_0:def 6; -n <= -1 by A20,REAL_1:50; then (-n)*d1 <= (-1)*d1 by A12,AXIOMS:25; then n*(-d1) <= (-1)*d1 by XCMPLX_1:176; then n*(-d1) <= -(1*d1) by XCMPLX_1:175; then n*(-d1) + (d1+d2)*n <= -d1 + (d1+d2)*n by AXIOMS:24; then d2*n <= f'.n + g'.n by A22,A26,A27,AXIOMS:22; hence d*g.n <= t.n by A7,A23; f'.n <= c1*f.n & g'.n <= c2*g.n by A13,A15,A20; then f'.n+g'.n<=c1*f.n+g'.n & c1*f.n+g'.n<=c1*f.n+c2*g.n by AXIOMS:24; then A28:f'.n + g'.n <= c1*f.n + c2*g.n by AXIOMS:22; A29:c1*f.n + c2*g.n = c1*(-1) + c1*n + c2*n by A23,A25,XCMPLX_1:8 .= c1*(-1) + (c1*n + c2*n) by XCMPLX_1:1 .= c1*(-1) + (c1+c2)*n by XCMPLX_1:8 .= -(c1*1) + (c1+c2)*n by XCMPLX_1:175; -c1 < -0 by A12,REAL_1:50; then -c1 + (c1+c2)*n <= 0 + (c1+c2)*n by AXIOMS:24; then f'.n + g'.n <= (c1+c2)*n by A28,A29,AXIOMS:22; hence t.n <= c*g.n by A7,A23; end; hence x in Big_Theta(g) by A3,A4,A18; end; assume x in Big_Theta(g); then consider t being Element of Funcs(NAT, REAL) such that A30: t = x and A31: ex c,d,N st c > 0 & d > 0 & for n st n >= N holds d*g.n <= t.n & t.n <= c*g.n by A3; consider c,d,N such that A32: c > 0 & d > 0 and A33: for n st n >= N holds d*g.n <= t.n & t.n <= c*g.n by A31; A34: 2"*c > 2"*0 & 2"*d > 2"*0 by A32,REAL_1:70; set f' = 2"(#)t, g' = 2"(#)t; A35: f' is Element of Funcs(NAT,REAL) & g' is Element of Funcs(NAT,REAL) by FUNCT_2:11; set N0 = max(N, 2); A36:N0 >= N & N0 >= 2 by SQUARE_1:46; reconsider N0 as Nat; now let n; assume A37: n >= N0; then A38:n >= N & n >= 2 by A36,AXIOMS:22; then d*g.n <= t.n & t.n <= c*g.n by A33; then A39:2"*(d*g.n) <= 2"*t.n & 2"*t.n <= 2"*(c*g.n) by AXIOMS:25; then A40:(2"*d)*g.n <= 2"*t.n by XCMPLX_1:4; A41:q.n = (NAT --> 1).n by Def4 .= 1 by FUNCOP_1:13; A42:f.n = (p+-q).n by SEQ_1:15 .= p.n + (-q).n by SEQ_1:11 .= p.n + -1 by A41,SEQ_1:14 .= p.n - 1 by XCMPLX_0:def 8 .= (n to_power 1) - 1 by A36,A37,Def3 .= n-1 by POWER:30; A43:g.n = (n to_power 1) - 0 by A36,A37,Def3 .= n-0 by POWER:30; then f.n <= g.n by A42,REAL_1:92; then 2"*d*f.n <= 2"*d*g.n by A34,AXIOMS:25; then 2"*d*f.n <= 2"*t.n by A40,AXIOMS:22; hence (2"*d)*f.n <= f'.n by SEQ_1:13; A44:2"*t.n <= (2"*c)*g.n by A39,XCMPLX_1:4; n + 2 <= n + n by A38,AXIOMS:24; then n+2 <= 2*n by XCMPLX_1:11; then n <= 2*n - 2*1 by REAL_1:84; then n <= 2*(n-1) by XCMPLX_1:40; then 2"*n <= 2"*(2*(n-1)) by AXIOMS:25; then 2"*n <= (2"*2)*(n-1) by XCMPLX_1:4; then c*(2"*n) <= c*(n-1) by A32,AXIOMS:25; then (2"*c)*n <= c*(n-1) by XCMPLX_1:4; then 2"*t.n <= c*f.n by A42,A43,A44,AXIOMS:22; hence f'.n <= c*f.n by SEQ_1:13; end; then A45: f' in Big_Theta(f) by A2,A32,A34,A35; now let n; assume n >= N0; then n >= N & n >= 2 by A36,AXIOMS:22; then d*g.n <= t.n & t.n <= c*g.n by A33; then 2"*(d*g.n) <= 2"*t.n & 2"*t.n <= 2"*(c*g.n) by AXIOMS:25; then 2"*d*g.n <= 2"*t.n & 2"*t.n <= 2"*c*g.n by XCMPLX_1:4; hence 2"*d*g.n <= g'.n & g'.n <= 2"*c*g.n by SEQ_1:13; end; then A46: g' in Big_Theta(g) by A3,A34,A35; for n being Element of NAT holds t.n = f'.n + g'.n proof let n be Element of NAT; f'.n = 2"*t.n by SEQ_1:13; then f'.n + g'.n = 2*(2"*t.n) by XCMPLX_1:11 .= (2*2")*t.n by XCMPLX_1:4 .= 1*t.n; hence thesis; end; hence x in Big_Theta(f) + Big_Theta(g) by A1,A30,A35,A45,A46; end; hence thesis by TARSKI:2; end; begin :: Problem 3.35 theorem ex F being FUNCTION_DOMAIN of NAT,REAL st F = { seq_n^(1) } & (for n holds (seq_n^(-1)).n <= (seq_n^(1)).n) & not seq_n^(-1) in F to_power Big_Oh(seq_const(1)) proof reconsider F = { seq_n^(1) } as FUNCTION_DOMAIN of NAT,REAL by FRAENKEL:10; take F; thus F = { seq_n^(1) }; A1:now let n; per cases; suppose n = 0; then (seq_n^(-1)).n = 0 & (seq_n^(1)).n = 0 by Def3; hence (seq_n^(-1)).n <= (seq_n^(1)).n; suppose A2: n > 0; then A3: (seq_n^(-1)).n=n to_power (-1) & (seq_n^(1)).n=n to_power 1 by Def3; A4: n >= 0+1 by A2,INT_1:20; n to_power (-1) <= n to_power 1 proof per cases by A4,REAL_1:def 5; suppose n = 1; then n to_power (-1) = 1 & n to_power 1 = 1 by POWER:31; hence thesis; suppose n > 1; hence thesis by PRE_FF:10; end; hence (seq_n^(-1)).n <= (seq_n^(1)).n by A3; end; set t = seq_n^(-1); now assume A5: t in F to_power Big_Oh(seq_const(1)); consider H being FUNCTION_DOMAIN of NAT,REAL such that A6: H = F and A7: t in H to_power Big_Oh(seq_const(1)) iff ex N,c,k st c>0 & for n st n >= N holds 1 <= t.n & t.n <= c*(seq_n^(k)).n by Th9; consider N0,c,k such that c>0 and A8: for n st n >= N0 holds 1 <= t.n & t.n <= c*(seq_n^(k)).n by A5,A6,A7; set N = max(N0,2); A9: N >= N0 & N >= 2 by SQUARE_1:46; now let n; assume A10: n >= N; then A11: n >= N0 & n >= 2 by A9,AXIOMS:22; then A12: t.n >= 1 by A8; A13: n > 1 by A11,AXIOMS:22; t.n = n to_power (-1) by A9,A10,Def3; hence contradiction by A12,A13,POWER:41; end; hence contradiction; end; hence thesis by A1; end; begin :: Addition :: In theorem ASYMPT_0:14, if we restrict our attentions to functions that :: do not have a subsequence that converges to 0, then the reverse implication :: is true. theorem for c being non negative Real, x,f being eventually-nonnegative Real_Sequence st ex e,N st e > 0 & for n st n >= N holds f.n >= e holds x in Big_Oh(c+f) implies x in Big_Oh(f) proof let c be non negative Real, x,f be eventually-nonnegative Real_Sequence; given e,N0 such that A1: e > 0 and A2: for n st n >= N0 holds f.n >= e; A3: Big_Oh(c+f) = { t where t is Element of Funcs(NAT, REAL) : ex d,N st d > 0 & for n st n >= N holds t.n <= d*(c+f).n & t.n >= 0 } by ASYMPT_0:def 12; A4: Big_Oh(f) = { t where t is Element of Funcs(NAT, REAL) : ex d,N st d > 0 & for n st n >= N holds t.n <= d*f.n & t.n >= 0 } by ASYMPT_0:def 12; assume x in Big_Oh(c+f); then consider t being Element of Funcs(NAT, REAL) such that A5: x = t and A6: ex d,N st d > 0 & for n st n >= N holds t.n <= d*(c+f).n & t.n >= 0 by A3; consider d,N1 such that A7: d > 0 and A8: for n st n >= N1 holds t.n <= d*(c+f).n & t.n >= 0 by A6; set b = max(2*d, (2*d*c)/e); A9: b >= 2*d & b >= (2*d*c)/e by SQUARE_1:46; 2*d > 2*0 by A7,REAL_1:70; then A10: b > 0 by SQUARE_1:46; set N = max(N0, N1); A11: N >= N0 & N >= N1 by SQUARE_1:46; now let n; assume n >= N; then A12: n >= N0 & n >= N1 by A11,AXIOMS:22; then t.n <= d*(c+f).n by A8; then t.n <= d*(c+f.n) by ASYMPT_0:def 9; then A13: t.n <= d*c + d*f.n by XCMPLX_1:8; thus t.n <= b*f.n proof per cases; suppose c >= f.n; then d*c >= d*f.n by A7,AXIOMS:25; then d*c + d*c >= d*c + d*f.n by AXIOMS:24; then t.n <= d*c + d*c by A13,AXIOMS:22; then t.n <= (2*(d*c))*1 by XCMPLX_1:11; then t.n <= (2*(d*c))*((1/e)*e) by A1,XCMPLX_1:107; then t.n <= ((2*(d*c))*(1/e))*e by XCMPLX_1:4; then t.n <= ((2*(d*c))/e)*e by XCMPLX_1:100; then A14: t.n <= ((2*d*c)/e)*e by XCMPLX_1:4; b*e >= ((2*d*c)/e)*e by A1,A9,AXIOMS:25; then A15: t.n <= b*e by A14,AXIOMS:22; f.n >= e by A2,A12; then b*f.n >= b*e by A10,AXIOMS:25; hence thesis by A15,AXIOMS:22; suppose c < f.n; then d*c < d*f.n by A7,REAL_1:70; then d*c + d*f.n < d*f.n + d*f.n by REAL_1:53; then t.n < d*f.n + d*f.n by A13,AXIOMS:22; then t.n < 2*(d*f.n) by XCMPLX_1:11; then A16: t.n < (2*d)*f.n by XCMPLX_1:4; f.n > 0 by A1,A2,A12; then b*f.n >= (2*d)*f.n by A9,AXIOMS:25; hence thesis by A16,AXIOMS:22; end; thus t.n >= 0 by A8,A12; end; hence x in Big_Oh(f) by A4,A5,A10; end; begin :: Potentially Useful Facts theorem 2 to_power 2 = 4 by Lm7; theorem 2 to_power 3 = 8 by Lm8; theorem 2 to_power 4 = 16 by Lm9; theorem 2 to_power 5 = 32 by Lm10; theorem 2 to_power 6 = 64 by Lm11; theorem 2 to_power 12 = 4096 by Lm32; theorem for n st n >= 3 holds n^2 > 2*n + 1 by Lm33; theorem for n st n >= 10 holds 2 to_power (n-1) > (2*n)^2 by Lm34; theorem for n st n >= 9 holds (n+1) to_power 6 < 2*(n to_power 6) by Lm35; theorem for n st n >= 30 holds 2 to_power n > n to_power 6 by Lm36; theorem for x being Real st x > 9 holds 2 to_power x > (2*x)^2 by Lm37; theorem ex N st for n st n >= N holds sqrt n - log(2, n) > 1 by Lm38; theorem for a,b,c being Real st a > 0 & c > 0 & c <> 1 holds a to_power b = c to_power (b*log(c,a)) by Lm3; theorem (4+1)! = 120 by Lm39; theorem 5 to_power 5 = 3125 by Lm42; theorem 4 to_power 4 = 256 by Lm43; theorem for n holds (n^2 - n + 1) > 0 by Lm27; theorem for n st n >= 2 holds n! > 1 by Lm58; theorem for n1, n st n <= n1 holds n! <= n1! by Lm59; theorem for k st k >= 1 holds ex n st (n! <= k & k < (n+1)! & for m st m! <= k & k < (m+1)! holds m = n) by Lm60; theorem for n st n >= 2 holds [/n/2\] < n by Lm53; theorem for n st n >= 3 holds n! > n by Lm61; canceled; theorem for n st n >= 2 holds 2 to_power n > n+1 by Lm1; theorem for a being logbase Real, f being Real_Sequence st a > 1 & (f.0 = 0 & (for n st n > 0 holds f.n = log(a,n))) holds f is eventually-positive by Lm2; theorem for f,g being eventually-nonnegative Real_Sequence holds f in Big_Oh(g) & g in Big_Oh(f) iff Big_Oh(f) = Big_Oh(g) by Lm5; theorem for a, b, c being Real st 0 < a & a <= b & c >= 0 holds a to_power c <= b to_power c by Lm6; theorem for n st n >= 4 holds 2*n + 3 < 2 to_power n by Lm12; theorem for n st n >= 6 holds (n+1)^2 < 2 to_power n by Lm13; theorem for c being Real st c > 6 holds c^2 < 2 to_power c by Lm14; theorem for e being positive Real, f being Real_Sequence st (f.0 = 0 & (for n st n > 0 holds f.n = log(2,n to_power e))) holds (f /" seq_n^(e)) is convergent & lim (f /" seq_n^(e)) = 0 by Lm15; theorem for e being Real st e > 0 holds (seq_logn/"seq_n^(e)) is convergent & lim(seq_logn/"seq_n^(e)) = 0 by Lm16; theorem for f being Real_Sequence holds for N holds (for n st n <= N holds f.n >= 0) implies Sum(f,N) >= 0 by Lm17; theorem for f,g being Real_Sequence holds for N holds (for n st n <= N holds f.n <= g.n) implies Sum(f,N) <= Sum (g,N) by Lm18; theorem for f being Real_Sequence, b being Real st (f.0 = 0 & (for n st n > 0 holds f.n = b)) holds for N being Nat holds Sum(f,N) = b*N by Lm19; theorem for f,g being Real_Sequence, N,M being Nat holds Sum(f,N,M) + f.(N+1) = Sum(f,N+1,M) by Lm20; theorem for f,g being Real_Sequence, M being Nat holds for N st N >= M+1 holds (for n st M+1 <= n & n <= N holds f.n <= g.n) implies Sum(f,N,M) <= Sum (g,N,M) by Lm21; theorem for n holds [/n/2\] <= n by Lm22; theorem for f being Real_Sequence, b being Real, N being Nat st (f.0 = 0 & (for n st n > 0 holds f.n = b)) holds for M being Nat holds Sum(f, N, M) = b*(N-M) by Lm23; theorem for f,g being Real_Sequence, N being Nat, c being Real st f is convergent & lim f = c & for n st n >= N holds f.n = g.n holds g is convergent & lim g = c by Lm28; theorem for n st n >= 1 holds (n^2 -n + 1) <= n^2 by Lm29; theorem for n st n >= 1 holds n^2 <= 2*(n^2 -n + 1) by Lm30; theorem for e being Real st 0 < e & e < 1 holds ex N st for n st n >= N holds n*log(2,1+e) - 8*log(2,n) > 8*log(2,n) by Lm31; theorem for n st n >= 10 holds 2 to_power (2*n) / (n!) < 1/(2 to_power (n-9)) by Lm40; theorem for n st n >= 3 holds 2*(n-2) >= n-1 by Lm41; theorem for c being Real st c >= 0 holds (c to_power (1/2)) = sqrt c by Lm45; theorem ex N st for n st n >= N holds n - sqrt n*log(2,n) > n/2 by Lm46; :: The proof of this theorem has been taken directly from the :: article POWER.miz (with very slight modifications to fit this new theorem). theorem for s being Real_Sequence st for n holds s.n = (1 + 1/(n+1)) to_power (n+1) holds s is non-decreasing by Lm47; theorem :: (1 + 1/n)^n is non-decreasing for n st n >= 1 holds ((n+1)/n) to_power n <= ((n+2)/(n+1)) to_power (n+1) by Lm48; theorem for k,n st k <= n holds n choose k >= ((n+1) choose k) / (n+1) by Lm49; theorem for f being Real_Sequence st (for n holds f.n = log(2,n!)) holds for n holds f.n = Sum(seq_logn, n) by Lm50; theorem for n st n >= 4 holds n*log(2,n) >= 2*n by Lm51; theorem for a,b being positive Real holds Prob28(0,a,b) = 0 & Prob28(1,a,b) = a & for n st n >= 2 holds ex n1 st n1 = [/n/2\] & Prob28(n,a,b) = 4*Prob28(n1,a,b) + b*n by Lm54; theorem for n st n >= 2 holds n^2 > n+1 by Lm55; theorem for n st n >= 1 holds (2 to_power (n+1)) - (2 to_power n) > 1 by Lm56; theorem for n st n >= 2 holds not ((2 to_power n) - 1) in POWEROF2SET by Lm57; theorem for n,k st k >= 1 & n! <= k & k < (n+1)! holds Step1(k) = n! by Lm62; theorem for a,b,c being Real holds a>1 & b>=a & c>=1 implies log(a,c) >= log(b,c) by Lm24;