begin
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem
canceled;
Th3:
for F, G being FinSequence st len F = len G & ( for i being Nat st i in dom F holds
F . i = G . i ) holds
F = G
by FINSEQ_2:10;
theorem Th5:
theorem Th6:
theorem Th7:
:: deftheorem defines |^ NEWTON:def 1 :
for x being complex number
for n being Nat holds x |^ n = Product (n |-> x);
theorem
canceled;
theorem
theorem Th10:
theorem Th11:
theorem
theorem Th13:
theorem
theorem Th15:
for
s being
Nat holds 1
|^ s = 1
theorem Th16:
:: deftheorem defines ! NEWTON:def 2 :
for n being Nat holds n ! = Product (idseq n);
theorem
canceled;
theorem Th18:
theorem
theorem
theorem Th21:
for
s being
Nat holds
(s + 1) ! = (s !) * (s + 1)
theorem Th22:
theorem Th23:
theorem
canceled;
theorem
for
s,
t being
Nat holds
(s !) * (t !) <> 0 ;
definition
let k,
n be
Nat;
func n choose k -> set means :
Def3:
for
l being
Nat st
l = n - k holds
it = (n !) / ((k !) * (l !)) if n >= k otherwise it = 0 ;
consistency
for b1 being set holds verum
;
existence
( ( n >= k implies ex b1 being set st
for l being Nat st l = n - k holds
b1 = (n !) / ((k !) * (l !)) ) & ( not n >= k implies ex b1 being set st b1 = 0 ) )
uniqueness
for b1, b2 being set holds
( ( n >= k & ( for l being Nat st l = n - k holds
b1 = (n !) / ((k !) * (l !)) ) & ( for l being Nat st l = n - k holds
b2 = (n !) / ((k !) * (l !)) ) implies b1 = b2 ) & ( not n >= k & b1 = 0 & b2 = 0 implies b1 = b2 ) )
end;
:: deftheorem Def3 defines choose NEWTON:def 3 :
for k, n being Nat
for b3 being set holds
( ( n >= k implies ( b3 = n choose k iff for l being Nat st l = n - k holds
b3 = (n !) / ((k !) * (l !)) ) ) & ( not n >= k implies ( b3 = n choose k iff b3 = 0 ) ) );
theorem
canceled;
theorem
canceled;
theorem
canceled;
theorem Th29:
theorem Th30:
theorem Th31:
theorem Th32:
theorem Th33:
theorem
theorem Th35:
theorem
definition
let a,
b be
real number ;
let n be
Nat;
func (
a,
b)
In_Power n -> FinSequence of
REAL means :
Def4:
(
len it = n + 1 & ( for
i,
l,
m being
Nat st
i in dom it &
m = i - 1 &
l = n - m holds
it . i = ((n choose m) * (a |^ l)) * (b |^ m) ) );
existence
ex b1 being FinSequence of REAL st
( len b1 = n + 1 & ( for i, l, m being Nat st i in dom b1 & m = i - 1 & l = n - m holds
b1 . i = ((n choose m) * (a |^ l)) * (b |^ m) ) )
uniqueness
for b1, b2 being FinSequence of REAL st len b1 = n + 1 & ( for i, l, m being Nat st i in dom b1 & m = i - 1 & l = n - m holds
b1 . i = ((n choose m) * (a |^ l)) * (b |^ m) ) & len b2 = n + 1 & ( for i, l, m being Nat st i in dom b2 & m = i - 1 & l = n - m holds
b2 . i = ((n choose m) * (a |^ l)) * (b |^ m) ) holds
b1 = b2
end;
:: deftheorem Def4 defines In_Power NEWTON:def 4 :
for a, b being real number
for n being Nat
for b4 being FinSequence of REAL holds
( b4 = (a,b) In_Power n iff ( len b4 = n + 1 & ( for i, l, m being Nat st i in dom b4 & m = i - 1 & l = n - m holds
b4 . i = ((n choose m) * (a |^ l)) * (b |^ m) ) ) );
theorem
canceled;
theorem Th38:
theorem Th39:
theorem Th40:
theorem Th41:
:: deftheorem Def5 defines Newton_Coeff NEWTON:def 5 :
for n being Nat
for b2 being FinSequence of REAL holds
( b2 = Newton_Coeff n iff ( len b2 = n + 1 & ( for i, k being Nat st i in dom b2 & k = i - 1 holds
b2 . i = n choose k ) ) );
theorem
canceled;
theorem Th43:
theorem
begin
theorem Th45:
for
l,
k being
Nat st
l >= 1 holds
k * l >= k
theorem Th46:
for
l,
n,
k being
Nat st
l >= 1 &
n >= k * l holds
n >= k
theorem Th47:
theorem
for
n being
Nat st
n <> 0 holds
(n + 1) / n > 1
theorem Th49:
for
k being
Nat holds
k / (k + 1) < 1
theorem Th50:
for
l being
Nat holds
l ! >= l
theorem Th51:
theorem
canceled;
theorem Th53:
theorem Th54:
theorem
theorem
theorem Th57:
theorem
theorem
theorem Th60:
theorem
theorem Th62:
theorem
theorem
theorem
theorem Th66:
theorem Th67:
theorem
theorem
theorem
theorem Th71:
theorem
canceled;
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
theorem
:: deftheorem Def6 defines SetPrimes NEWTON:def 6 :
for b1 being Subset of NAT holds
( b1 = SetPrimes iff for n being Nat holds
( n in b1 iff n is prime ) );
:: deftheorem Def7 defines SetPrimenumber NEWTON:def 7 :
for p being Nat
for b2 being Subset of NAT holds
( b2 = SetPrimenumber p iff for q being Nat holds
( q in b2 iff ( q < p & q is prime ) ) );
theorem Th82:
theorem
theorem Th84:
theorem Th85:
theorem Th86:
theorem
canceled;
theorem
canceled;
theorem
canceled;
Th88:
SetPrimenumber 2 = {}
theorem
theorem
canceled;
theorem
theorem
theorem Th94:
theorem
Lm1:
for n being Nat holds SetPrimenumber n = { k where k is Element of NAT : ( k < n & k is prime ) }
:: deftheorem Def8 defines primenumber NEWTON:def 8 :
for n being Nat
for b2 being prime Element of NAT holds
( b2 = primenumber n iff n = card (SetPrimenumber b2) );
theorem
theorem Th97:
theorem
theorem
canceled;
theorem
theorem
theorem
theorem
defpred S1[ Nat] means 2 |^ $1 >= $1 + 1;
Lm2:
S1[ 0 ]
by RVSUM_1:124;
Lm3:
for n being Nat st S1[n] holds
S1[n + 1]
theorem Th104:
theorem
for
n being
Nat holds 2
|^ n > n
theorem Th106:
Lm4:
for n1, m1, n2, m2 being natural number st (2 |^ n1) * ((2 * m1) + 1) = (2 |^ n2) * ((2 * m2) + 1) holds
n1 <= n2
theorem