:: NAT_1 semantic presentation

definition
mode Nat is natural number ;
end;

theorem Th1: :: NAT_1:1
canceled;

theorem Th2: :: NAT_1:2
for X being Subset of REAL st 0 in X & ( for x being Real st x in X holds
x + 1 in X ) holds
for n being Nat holds n in X
proof end;

definition
let n be Element of NAT , k be Element of NAT ;
redefine func + as c1 + c2 -> Element of NAT ;
coherence
n + k is Element of NAT
proof end;
end;

registration
let n be Nat, k be Nat;
cluster a1 + a2 -> natural ;
coherence
n + k is natural
proof end;
end;

scheme :: NAT_1:sch 36
s36{ P1[ Nat] } :
for k being Element of NAT holds P1[k]
provided
E31: P1[0] and
E32: for k being Element of NAT st P1[k] holds
P1[k + 1]
proof end;

scheme :: NAT_1:sch 39
s39{ P1[ Nat] } :
for k being Nat holds P1[k]
provided
E31: P1[0] and
E32: for k being Nat st P1[k] holds
P1[k + 1]
proof end;

definition
let n be Element of NAT , k be Element of NAT ;
redefine func * as c1 * c2 -> Element of NAT ;
coherence
n * k is Element of NAT
proof end;
end;

registration
let n be Nat, k be Nat;
cluster a1 * a2 -> natural ;
coherence
n * k is natural
proof end;
end;

theorem Th3: :: NAT_1:3
canceled;

theorem Th4: :: NAT_1:4
canceled;

theorem Th5: :: NAT_1:5
canceled;

theorem Th6: :: NAT_1:6
canceled;

theorem Th7: :: NAT_1:7
canceled;

theorem Th8: :: NAT_1:8
canceled;

theorem Th9: :: NAT_1:9
canceled;

theorem Th10: :: NAT_1:10
canceled;

theorem Th11: :: NAT_1:11
canceled;

theorem Th12: :: NAT_1:12
canceled;

theorem Th13: :: NAT_1:13
canceled;

theorem Th14: :: NAT_1:14
canceled;

theorem Th15: :: NAT_1:15
canceled;

theorem Th16: :: NAT_1:16
canceled;

theorem Th17: :: NAT_1:17
canceled;

theorem Th18: :: NAT_1:18
for i being Nat holds 0 <= i
proof end;

theorem Th19: :: NAT_1:19
for i being Nat st 0 <> i holds
0 < i by ;

theorem Th20: :: NAT_1:20
for i, j, h being Nat st i <= j holds
i * h <= j * h
proof end;

theorem Th21: :: NAT_1:21
for i being Nat holds 0 < i + 1
proof end;

theorem Th22: :: NAT_1:22
for i being Nat holds
( i = 0 or ex k being Nat st i = k + 1 )
proof end;

theorem Th23: :: NAT_1:23
for i, j being Nat st i + j = 0 holds
( i = 0 & j = 0 )
proof end;

registration
cluster non zero set ;
existence
not for b1 being Nat holds b1 is empty
proof end;
end;

registration
let m be Nat;
let n be non zero Nat;
cluster a1 + a2 -> non zero natural ;
coherence
not m + n is empty
by ;
cluster a2 + a1 -> non zero natural ;
coherence
not n + m is empty
;
end;

scheme :: NAT_1:sch 44
s44{ F1() -> Nat, F2( Nat, Nat) -> Nat, P1[ Nat, Nat] } :
( ( for k being Element of NAT ex n being Element of NAT st P1[k,n] ) & ( for k, n, m being Element of NAT st P1[k,n] & P1[k,m] holds
n = m ) )
provided
E31: for k, n being Element of NAT holds
( P1[k,n] iff ( ( k = 0 & n = F1() ) or ex m, l being Element of NAT st
( k = m + 1 & P1[m,l] & n = F2(k,l) ) ) )
proof end;

theorem Th24: :: NAT_1:24
canceled;

theorem Th25: :: NAT_1:25
canceled;

theorem Th26: :: NAT_1:26
for i, j being Nat holds
( not i <= j + 1 or i <= j or i = j + 1 )
proof end;

theorem Th27: :: NAT_1:27
for i, j being Nat st i <= j & j <= i + 1 & not i = j holds
j = i + 1
proof end;

theorem Th28: :: NAT_1:28
for i, j being Nat st i <= j holds
ex k being Nat st j = i + k
proof end;

theorem Th29: :: NAT_1:29
for i, j being Nat holds i <= i + j
proof end;

scheme :: NAT_1:sch 58
s58{ P1[ Nat] } :
for k being Element of NAT holds P1[k]
provided
E31: for k being Element of NAT st ( for n being Element of NAT st n < k holds
P1[n] ) holds
P1[k]
proof end;

scheme :: NAT_1:sch 59
s59{ P1[ Nat] } :
ex k being Element of NAT st
( P1[k] & ( for n being Element of NAT st P1[n] holds
k <= n ) )
provided
E31: ex k being Element of NAT st P1[k]
proof end;

scheme :: NAT_1:sch 60
s60{ P1[ Nat], F1() -> Nat } :
ex k being Element of NAT st
( P1[k] & ( for n being Element of NAT st P1[n] holds
n <= k ) )
provided
E31: for k being Element of NAT st P1[k] holds
k <= F1() and
E32: ex k being Element of NAT st P1[k]
proof end;

theorem Th30: :: NAT_1:30
canceled;

theorem Th31: :: NAT_1:31
canceled;

theorem Th32: :: NAT_1:32
canceled;

theorem Th33: :: NAT_1:33
canceled;

theorem Th34: :: NAT_1:34
canceled;

theorem Th35: :: NAT_1:35
canceled;

theorem Th36: :: NAT_1:36
canceled;

theorem Th37: :: NAT_1:37
for i, j, h being Nat st i <= j holds
i <= j + h
proof end;

theorem Th38: :: NAT_1:38
for i, j being Nat holds
( i < j + 1 iff i <= j )
proof end;

theorem Th39: :: NAT_1:39
for j being Nat st j < 1 holds
j = 0
proof end;

theorem Th40: :: NAT_1:40
for i, j being Nat st i * j = 1 holds
( i = 1 & j = 1 )
proof end;

theorem Th41: :: NAT_1:41
for k, n being Nat st k <> 0 holds
n < n + k
proof end;

scheme :: NAT_1:sch 66
s66{ P1[ Nat] } :
P1[0]
provided
E31: ex k being Element of NAT st P1[k] and
E32: for k being Element of NAT st k <> 0 & P1[k] holds
ex n being Element of NAT st
( n < k & P1[n] )
proof end;

theorem Th42: :: NAT_1:42
for m being Nat st 0 < m holds
for n being Nat ex k, t being Nat st
( n = (m * k) + t & t < m )
proof end;

theorem Th43: :: NAT_1:43
for n, m, k, t, k1, t1 being Nat st n = (m * k) + t & t < m & n = (m * k1) + t1 & t1 < m holds
( k = k1 & t = t1 )
proof end;

definition
let k be Nat, l be Nat;
func c1 div c2 -> Nat means :Def1: :: NAT_1:def 1
( ex t being Nat st
( k = (l * it) + t & t < l ) or ( it = 0 & l = 0 ) );
existence
ex b1 being Nat st
( ex t being Nat st
( k = (l * b1) + t & t < l ) or ( b1 = 0 & l = 0 ) )
proof end;
uniqueness
for b1, b2 being Nat st ( ex t being Nat st
( k = (l * b1) + t & t < l ) or ( b1 = 0 & l = 0 ) ) & ( ex t being Nat st
( k = (l * b2) + t & t < l ) or ( b2 = 0 & l = 0 ) ) holds
b1 = b2
proof end;
func c1 mod c2 -> Nat means :Def2: :: NAT_1:def 2
( ex t being Nat st
( k = (l * t) + it & it < l ) or ( it = 0 & l = 0 ) );
existence
ex b1 being Nat st
( ex t being Nat st
( k = (l * t) + b1 & b1 < l ) or ( b1 = 0 & l = 0 ) )
proof end;
uniqueness
for b1, b2 being Nat st ( ex t being Nat st
( k = (l * t) + b1 & b1 < l ) or ( b1 = 0 & l = 0 ) ) & ( ex t being Nat st
( k = (l * t) + b2 & b2 < l ) or ( b2 = 0 & l = 0 ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def1 defines div NAT_1:def 1 :
for k, l, b3 being Nat holds
( b3 = k div l iff ( ex t being Nat st
( k = (l * b3) + t & t < l ) or ( b3 = 0 & l = 0 ) ) );

:: deftheorem Def2 defines mod NAT_1:def 2 :
for k, l, b3 being Nat holds
( b3 = k mod l iff ( ex t being Nat st
( k = (l * t) + b3 & b3 < l ) or ( b3 = 0 & l = 0 ) ) );

definition
let k be Nat, l be Nat;
redefine func div as c1 div c2 -> Element of NAT ;
coherence
k div l is Element of NAT
by ORDINAL1:def 13;
redefine func mod as c1 mod c2 -> Element of NAT ;
coherence
k mod l is Element of NAT
by ORDINAL1:def 13;
end;

theorem Th44: :: NAT_1:44
canceled;

theorem Th45: :: NAT_1:45
canceled;

theorem Th46: :: NAT_1:46
for i, j being Nat st 0 < i holds
j mod i < i
proof end;

theorem Th47: :: NAT_1:47
for i, j being Nat st 0 < i holds
j = (i * (j div i)) + (j mod i)
proof end;

definition
let k be Nat, l be Nat;
pred c1 divides c2 means :Def3: :: NAT_1:def 3
ex t being Nat st l = k * t;
reflexivity
for k being Nat ex t being Nat st k = k * t
proof end;
end;

:: deftheorem Def3 defines divides NAT_1:def 3 :
for k, l being Nat holds
( k divides l iff ex t being Nat st l = k * t );

theorem Th48: :: NAT_1:48
canceled;

theorem Th49: :: NAT_1:49
for j, i being Nat holds
( j divides i iff i = j * (i div j) )
proof end;

theorem Th50: :: NAT_1:50
canceled;

theorem Th51: :: NAT_1:51
for i, j, h being Nat st i divides j & j divides h holds
i divides h
proof end;

theorem Th52: :: NAT_1:52
for i, j being Nat st i divides j & j divides i holds
i = j
proof end;

theorem Th53: :: NAT_1:53
for i being Nat holds
( i divides 0 & 1 divides i )
proof end;

theorem Th54: :: NAT_1:54
for j, i being Nat st 0 < j & i divides j holds
i <= j
proof end;

theorem Th55: :: NAT_1:55
for i, j, h being Nat st i divides j & i divides h holds
i divides j + h
proof end;

theorem Th56: :: NAT_1:56
for i, j, h being Nat st i divides j holds
i divides j * h
proof end;

theorem Th57: :: NAT_1:57
for i, j, h being Nat st i divides j & i divides j + h holds
i divides h
proof end;

theorem Th58: :: NAT_1:58
for i, j, h being Nat st i divides j & i divides h holds
i divides j mod h
proof end;

definition
let k be Nat, n be Nat;
func c1 lcm c2 -> Nat means :Def4: :: NAT_1:def 4
( k divides it & n divides it & ( for m being Nat st k divides m & n divides m holds
it divides m ) );
existence
ex b1 being Nat st
( k divides b1 & n divides b1 & ( for m being Nat st k divides m & n divides m holds
b1 divides m ) )
proof end;
uniqueness
for b1, b2 being Nat st k divides b1 & n divides b1 & ( for m being Nat st k divides m & n divides m holds
b1 divides m ) & k divides b2 & n divides b2 & ( for m being Nat st k divides m & n divides m holds
b2 divides m ) holds
b1 = b2
proof end;
idempotence
for k being Nat holds
( k divides k & k divides k & ( for m being Nat st k divides m & k divides m holds
k divides m ) )
;
commutativity
for b1, k, n being Nat st k divides b1 & n divides b1 & ( for m being Nat st k divides m & n divides m holds
b1 divides m ) holds
( n divides b1 & k divides b1 & ( for m being Nat st n divides m & k divides m holds
b1 divides m ) )
;
end;

:: deftheorem Def4 defines lcm NAT_1:def 4 :
for k, n, b3 being Nat holds
( b3 = k lcm n iff ( k divides b3 & n divides b3 & ( for m being Nat st k divides m & n divides m holds
b3 divides m ) ) );

definition
let k be Nat, n be Nat;
redefine func lcm as c1 lcm c2 -> Element of NAT ;
coherence
k lcm n is Element of NAT
by ORDINAL1:def 13;
end;

definition
let k be Nat, n be Nat;
func c1 hcf c2 -> Nat means :Def5: :: NAT_1:def 5
( it divides k & it divides n & ( for m being Nat st m divides k & m divides n holds
m divides it ) );
existence
ex b1 being Nat st
( b1 divides k & b1 divides n & ( for m being Nat st m divides k & m divides n holds
m divides b1 ) )
proof end;
uniqueness
for b1, b2 being Nat st b1 divides k & b1 divides n & ( for m being Nat st m divides k & m divides n holds
m divides b1 ) & b2 divides k & b2 divides n & ( for m being Nat st m divides k & m divides n holds
m divides b2 ) holds
b1 = b2
proof end;
idempotence
for k being Nat holds
( k divides k & k divides k & ( for m being Nat st m divides k & m divides k holds
m divides k ) )
;
commutativity
for b1, k, n being Nat st b1 divides k & b1 divides n & ( for m being Nat st m divides k & m divides n holds
m divides b1 ) holds
( b1 divides n & b1 divides k & ( for m being Nat st m divides n & m divides k holds
m divides b1 ) )
;
end;

:: deftheorem Def5 defines hcf NAT_1:def 5 :
for k, n, b3 being Nat holds
( b3 = k hcf n iff ( b3 divides k & b3 divides n & ( for m being Nat st m divides k & m divides n holds
m divides b3 ) ) );

definition
let k be Nat, n be Nat;
redefine func hcf as c1 hcf c2 -> Element of NAT ;
coherence
k hcf n is Element of NAT
by ORDINAL1:def 13;
end;

scheme :: NAT_1:sch 91
s91{ F1( Nat) -> Nat, F2() -> Nat, F3() -> Nat } :
ex n being Element of NAT st
( F1(n) = F2() hcf F3() & F1((n + 1)) = 0 )
provided
E31: ( 0 < F3() & F3() < F2() ) and
E32: ( F1(0) = F2() & F1(1) = F3() ) and
E34: for n being Element of NAT holds F1((n + 2)) = F1(n) mod F1((n + 1))
proof end;

registration
cluster -> ordinal set ;
coherence
for b1 being Nat holds b1 is ordinal
;
end;

registration
cluster non empty ordinal Element of K18(REAL );
existence
ex b1 being Subset of REAL st
( not b1 is empty & b1 is ordinal )
proof end;
end;

theorem Th59: :: NAT_1:59
for k, n being Nat holds
( k < k + n iff 1 <= n )
proof end;

theorem Th60: :: NAT_1:60
for k, n being Nat st k < n holds
n - 1 is Element of NAT
proof end;

theorem Th61: :: NAT_1:61
for k, n being Nat st k <= n holds
n - k is Element of NAT
proof end;

theorem Th62: :: NAT_1:62
for n being Nat holds
( n mod 2 = 0 or n mod 2 = 1 )
proof end;

theorem Th63: :: NAT_1:63
for k, n being Nat holds (k * n) mod k = 0
proof end;

theorem Th64: :: NAT_1:64
for k being Nat st k > 1 holds
1 mod k = 1
proof end;

theorem Th65: :: NAT_1:65
for k, n, l, m being Nat st k mod n = 0 & l = k - (m * n) holds
l mod n = 0
proof end;

theorem Th66: :: NAT_1:66
for n, k, l being Nat st n <> 0 & k mod n = 0 & l < n holds
(k + l) mod n = l
proof end;

theorem Th67: :: NAT_1:67
for k, n, l being Nat st k mod n = 0 holds
(k + l) mod n = l mod n
proof end;

theorem Th68: :: NAT_1:68
for k, n being Nat st k <> 0 holds
(k * n) div k = n
proof end;

theorem Th69: :: NAT_1:69
for k, n, l being Nat st k mod n = 0 holds
(k + l) div n = (k div n) + (l div n)
proof end;

theorem Th70: :: NAT_1:70
for m, n being Nat holds
( not m < n + 1 or m < n or m = n )
proof end;

theorem Th71: :: NAT_1:71
for k being Nat holds
( not k < 2 or k = 0 or k = 1 )
proof end;

registration
cluster non zero Element of NAT ;
existence
not for b1 being Element of NAT holds b1 is empty
proof end;
end;

registration
cluster -> non negative Element of NAT ;
coherence
for b1 being Element of NAT holds not b1 is negative
by ;
end;

registration
cluster -> ordinal non negative set ;
coherence
for b1 being Nat holds not b1 is negative
by ;
end;

theorem Th72: :: NAT_1:72
for k, m being Nat st k <> 0 holds
(m * k) div k = m
proof end;

theorem Th73: :: NAT_1:73
for m, n, k being Nat holds m mod n = ((n * k) + m) mod n
proof end;

theorem Th74: :: NAT_1:74
for p, s, n being Nat holds (p + s) mod n = ((p mod n) + s) mod n
proof end;

theorem Th75: :: NAT_1:75
for p, s, n being Nat holds (p + s) mod n = (p + (s mod n)) mod n by ;

theorem Th76: :: NAT_1:76
for k, n being Nat st k < n holds
k mod n = k
proof end;

theorem Th77: :: NAT_1:77
for n being Nat holds n mod n = 0
proof end;

theorem Th78: :: NAT_1:78
for n being Nat holds 0 = 0 mod n
proof end;

theorem Th79: :: NAT_1:79
for i, h, j being Nat st i <> 0 & h = j * i holds
j <= h
proof end;

theorem Th80: :: NAT_1:80
for i, j being Nat st i < j holds
i div j = 0
proof end;

theorem Th81: :: NAT_1:81
for k, n being Nat holds
( k < n iff ( 1 <= k + 1 & k + 1 <= n ) )
proof end;

scheme :: NAT_1:sch 126
s126{ F1() -> Nat, P1[ Nat] } :
for i being Element of NAT st F1() <= i holds
P1[i]
provided
E31: P1[F1()] and
E32: for j being Element of NAT st F1() <= j & P1[j] holds
P1[j + 1]
proof end;

scheme :: NAT_1:sch 128
s128{ F1() -> Nat, P1[ Nat] } :
for k being Element of NAT st k >= F1() holds
P1[k]
provided
E31: for k being Element of NAT st k >= F1() & ( for n being Element of NAT st n >= F1() & n < k holds
P1[n] ) holds
P1[k]
proof end;

theorem Th82: :: NAT_1:82
for n being Nat holds
( not n <= 1 or n = 0 or n = 1 )
proof end;

theorem Th83: :: NAT_1:83
for n being Nat holds
( not n <= 2 or n = 0 or n = 1 or n = 2 )
proof end;

theorem Th84: :: NAT_1:84
for n being Nat holds
( not n <= 3 or n = 0 or n = 1 or n = 2 or n = 3 )
proof end;

theorem Th85: :: NAT_1:85
for n being Nat holds
( not n <= 4 or n = 0 or n = 1 or n = 2 or n = 3 or n = 4 )
proof end;

theorem Th86: :: NAT_1:86
for n being Nat holds
( not n <= 5 or n = 0 or n = 1 or n = 2 or n = 3 or n = 4 or n = 5 )
proof end;

theorem Th87: :: NAT_1:87
for n being Nat holds
( not n <= 6 or n = 0 or n = 1 or n = 2 or n = 3 or n = 4 or n = 5 or n = 6 )
proof end;

theorem Th88: :: NAT_1:88
for n being Nat holds
( not n <= 7 or n = 0 or n = 1 or n = 2 or n = 3 or n = 4 or n = 5 or n = 6 or n = 7 )
proof end;

theorem Th89: :: NAT_1:89
for n being Nat holds
( not n <= 8 or n = 0 or n = 1 or n = 2 or n = 3 or n = 4 or n = 5 or n = 6 or n = 7 or n = 8 )
proof end;

theorem Th90: :: NAT_1:90
for n being Nat holds
( not n <= 9 or n = 0 or n = 1 or n = 2 or n = 3 or n = 4 or n = 5 or n = 6 or n = 7 or n = 8 or n = 9 )
proof end;

scheme :: NAT_1:sch 137
s137{ P1[ Nat] } :
for k being non empty Nat holds P1[k]
provided
E31: P1[1] and
E32: for k being non empty Nat st P1[k] holds
P1[k + 1]
proof end;