:: Integer and Rational Exponents
:: by Konrad Raczkowski
::
:: Received September 21, 1990
:: Copyright (c) 1990-2021 Association of Mizar Users


registration
let i be Integer;
cluster |.i.| -> natural ;
coherence
|.i.| is natural
;
end;

theorem Th1: :: PREPOWER:1
for a being Real
for s1 being Real_Sequence st s1 is convergent & ( for n being Nat holds s1 . n >= a ) holds
lim s1 >= a
proof end;

theorem Th2: :: PREPOWER:2
for a being Real
for s1 being Real_Sequence st s1 is convergent & ( for n being Nat holds s1 . n <= a ) holds
lim s1 <= a
proof end;

definition
let a be Real;
func a GeoSeq -> Real_Sequence means :Def1: :: PREPOWER:def 1
for m being Nat holds it . m = a |^ m;
existence
ex b1 being Real_Sequence st
for m being Nat holds b1 . m = a |^ m
proof end;
uniqueness
for b1, b2 being Real_Sequence st ( for m being Nat holds b1 . m = a |^ m ) & ( for m being Nat holds b2 . m = a |^ m ) holds
b1 = b2
proof end;
end;

:: deftheorem Def1 defines GeoSeq PREPOWER:def 1 :
for a being Real
for b2 being Real_Sequence holds
( b2 = a GeoSeq iff for m being Nat holds b2 . m = a |^ m );

theorem Th3: :: PREPOWER:3
for a being Real
for s1 being Real_Sequence holds
( s1 = a GeoSeq iff ( s1 . 0 = 1 & ( for m being Nat holds s1 . (m + 1) = (s1 . m) * a ) ) )
proof end;

theorem :: PREPOWER:4
for a being Real st a <> 0 holds
for m being Nat holds (a GeoSeq) . m <> 0
proof end;

theorem Th5: :: PREPOWER:5
for a being Complex
for n being natural Number st 0 <> a holds
0 <> a |^ n
proof end;

theorem Th6: :: PREPOWER:6
for a being Real
for n being natural Number st 0 < a holds
0 < a |^ n
proof end;

theorem Th7: :: PREPOWER:7
for a being Complex
for n being natural Number holds (1 / a) |^ n = 1 / (a |^ n)
proof end;

theorem :: PREPOWER:8
for a, b being Complex
for n being natural Number holds (b / a) |^ n = (b |^ n) / (a |^ n)
proof end;

theorem Th9: :: PREPOWER:9
for a, b being Real
for n being natural Number st 0 < a & a <= b holds
a |^ n <= b |^ n
proof end;

Lm1: for a, b being Real
for n being natural Number st 0 < a & a < b & 1 <= n holds
a |^ n < b |^ n

proof end;

theorem Th10: :: PREPOWER:10
for a, b being Real
for n being natural Number st 0 <= a & a < b & 1 <= n holds
a |^ n < b |^ n
proof end;

theorem Th11: :: PREPOWER:11
for a being Real
for n being natural Number st a >= 1 holds
a |^ n >= 1
proof end;

theorem Th12: :: PREPOWER:12
for a being Real
for n being natural Number st 1 <= a & 1 <= n holds
a <= a |^ n
proof end;

theorem :: PREPOWER:13
for a being Real
for n being Nat st 1 < a & 2 <= n holds
a < a |^ n
proof end;

theorem Th14: :: PREPOWER:14
for a being Real
for n being natural Number st 0 < a & a <= 1 & 1 <= n holds
a |^ n <= a
proof end;

theorem :: PREPOWER:15
for a being Real
for n being Nat st 0 < a & a < 1 & 2 <= n holds
a |^ n < a
proof end;

theorem Th16: :: PREPOWER:16
for a being Real
for n being natural Number st - 1 < a holds
(1 + a) |^ n >= 1 + (n * a)
proof end;

theorem Th17: :: PREPOWER:17
for a being Real
for n being natural Number st 0 < a & a < 1 holds
(1 + a) |^ n <= 1 + ((3 |^ n) * a)
proof end;

theorem Th18: :: PREPOWER:18
for m being Nat
for s1, s2 being Real_Sequence st s1 is convergent & ( for n being Nat holds s2 . n = (s1 . n) |^ m ) holds
( s2 is convergent & lim s2 = (lim s1) |^ m )
proof end;

definition
let n be natural Number ;
let a be Real;
assume A1: 1 <= n ;
func n -Root a -> Real means :Def2: :: PREPOWER:def 2
( it |^ n = a & it > 0 ) if a > 0
it = 0 if a = 0
;
consistency
for b1 being Real st a > 0 & a = 0 holds
( ( b1 |^ n = a & b1 > 0 ) iff b1 = 0 )
;
existence
( ( a > 0 implies ex b1 being Real st
( b1 |^ n = a & b1 > 0 ) ) & ( a = 0 implies ex b1 being Real st b1 = 0 ) )
proof end;
uniqueness
for b1, b2 being Real holds
( ( a > 0 & b1 |^ n = a & b1 > 0 & b2 |^ n = a & b2 > 0 implies b1 = b2 ) & ( a = 0 & b1 = 0 & b2 = 0 implies b1 = b2 ) )
proof end;
end;

:: deftheorem Def2 defines -Root PREPOWER:def 2 :
for n being natural Number
for a being Real st 1 <= n holds
for b3 being Real holds
( ( a > 0 implies ( b3 = n -Root a iff ( b3 |^ n = a & b3 > 0 ) ) ) & ( a = 0 implies ( b3 = n -Root a iff b3 = 0 ) ) );

registration
let n be Nat;
let a be Real;
cluster n -Root a -> ;
coherence
n -Root a is real
;
end;

Lm2: for a being Real
for n being Nat st a > 0 & n >= 1 holds
( (n -Root a) |^ n = a & n -Root (a |^ n) = a )

proof end;

theorem Th19: :: PREPOWER:19
for a being Real
for n being Nat st a >= 0 & n >= 1 holds
( (n -Root a) |^ n = a & n -Root (a |^ n) = a )
proof end;

theorem Th20: :: PREPOWER:20
for n being Nat st n >= 1 holds
n -Root 1 = 1
proof end;

theorem Th21: :: PREPOWER:21
for a being Real st a >= 0 holds
1 -Root a = a
proof end;

theorem Th22: :: PREPOWER:22
for a, b being Real
for n being Nat st a >= 0 & b >= 0 & n >= 1 holds
n -Root (a * b) = (n -Root a) * (n -Root b)
proof end;

theorem Th23: :: PREPOWER:23
for a being Real
for n being Nat st a > 0 & n >= 1 holds
n -Root (1 / a) = 1 / (n -Root a)
proof end;

theorem :: PREPOWER:24
for a, b being Real
for n being Nat st a >= 0 & b > 0 & n >= 1 holds
n -Root (a / b) = (n -Root a) / (n -Root b)
proof end;

theorem Th25: :: PREPOWER:25
for a being Real
for m, n being Nat st a >= 0 & n >= 1 & m >= 1 holds
n -Root (m -Root a) = (n * m) -Root a
proof end;

theorem Th26: :: PREPOWER:26
for a being Real
for m, n being Nat st a >= 0 & n >= 1 & m >= 1 holds
(n -Root a) * (m -Root a) = (n * m) -Root (a |^ (n + m))
proof end;

theorem Th27: :: PREPOWER:27
for a, b being Real
for n being Nat st 0 <= a & a <= b & n >= 1 holds
n -Root a <= n -Root b
proof end;

theorem Th28: :: PREPOWER:28
for a, b being Real
for n being Nat st a >= 0 & a < b & n >= 1 holds
n -Root a < n -Root b
proof end;

theorem Th29: :: PREPOWER:29
for a being Real
for n being Nat st a >= 1 & n >= 1 holds
( n -Root a >= 1 & a >= n -Root a )
proof end;

theorem Th30: :: PREPOWER:30
for a being Real
for n being Nat st 0 <= a & a < 1 & n >= 1 holds
( a <= n -Root a & n -Root a < 1 )
proof end;

theorem Th31: :: PREPOWER:31
for a being Real
for n being Nat st a > 0 & n >= 1 holds
(n -Root a) - 1 <= (a - 1) / n
proof end;

theorem :: PREPOWER:32
for a being Real st a >= 0 holds
2 -Root a = sqrt a
proof end;

Lm3: for s being Real_Sequence
for a being Real st a >= 1 & ( for n being Nat st n >= 1 holds
s . n = n -Root a ) holds
( s is convergent & lim s = 1 )

proof end;

theorem :: PREPOWER:33
for a being Real
for s being Real_Sequence st a > 0 & ( for n being Nat st n >= 1 holds
s . n = n -Root a ) holds
( s is convergent & lim s = 1 )
proof end;

definition
let a be Real;
let k be Integer;
func a #Z k -> number equals :Def3: :: PREPOWER:def 3
a |^ |.k.| if k >= 0
(a |^ |.k.|) " if k < 0
;
consistency
for b1 being number st k >= 0 & k < 0 holds
( b1 = a |^ |.k.| iff b1 = (a |^ |.k.|) " )
;
coherence
( ( k >= 0 implies a |^ |.k.| is number ) & ( k < 0 implies (a |^ |.k.|) " is number ) )
by TARSKI:1;
end;

:: deftheorem Def3 defines #Z PREPOWER:def 3 :
for a being Real
for k being Integer holds
( ( k >= 0 implies a #Z k = a |^ |.k.| ) & ( k < 0 implies a #Z k = (a |^ |.k.|) " ) );

registration
let a be Real;
let k be Integer;
cluster a #Z k -> real ;
coherence
a #Z k is real
proof end;
end;

theorem Th34: :: PREPOWER:34
for a being Real holds a #Z 0 = 1
proof end;

theorem Th35: :: PREPOWER:35
for a being Real holds a #Z 1 = a
proof end;

theorem Th36: :: PREPOWER:36
for a being Real
for n being Nat holds a #Z n = a |^ n
proof end;

Lm4: 1 " = 1
;

theorem Th37: :: PREPOWER:37
for k being Integer holds 1 #Z k = 1
proof end;

theorem :: PREPOWER:38
for a being Real
for k being Integer st a <> 0 holds
a #Z k <> 0
proof end;

theorem Th39: :: PREPOWER:39
for a being Real
for k being Integer st a > 0 holds
a #Z k > 0
proof end;

theorem Th40: :: PREPOWER:40
for a, b being Real
for k being Integer holds (a * b) #Z k = (a #Z k) * (b #Z k)
proof end;

theorem Th41: :: PREPOWER:41
for a being Real
for k being Integer holds a #Z (- k) = 1 / (a #Z k)
proof end;

theorem Th42: :: PREPOWER:42
for a being Real
for k being Integer holds (1 / a) #Z k = 1 / (a #Z k)
proof end;

theorem Th43: :: PREPOWER:43
for a being Real
for m, n being Nat st a <> 0 holds
a #Z (m - n) = (a |^ m) / (a |^ n)
proof end;

theorem Th44: :: PREPOWER:44
for a being Real
for k, l being Integer st a <> 0 holds
a #Z (k + l) = (a #Z k) * (a #Z l)
proof end;

theorem Th45: :: PREPOWER:45
for a being Real
for k, l being Integer holds (a #Z k) #Z l = a #Z (k * l)
proof end;

theorem Th46: :: PREPOWER:46
for a being Real
for n being Nat
for k being Integer st a > 0 & n >= 1 holds
(n -Root a) #Z k = n -Root (a #Z k)
proof end;

definition
let a be Real;
let p be Rational;
func a #Q p -> number equals :: PREPOWER:def 4
(denominator p) -Root (a #Z (numerator p));
coherence
(denominator p) -Root (a #Z (numerator p)) is number
by TARSKI:1;
end;

:: deftheorem defines #Q PREPOWER:def 4 :
for a being Real
for p being Rational holds a #Q p = (denominator p) -Root (a #Z (numerator p));

registration
let a be Real;
let p be Rational;
cluster a #Q p -> real ;
coherence
a #Q p is real
;
end;

theorem Th47: :: PREPOWER:47
for a being Real
for p being Rational st p = 0 holds
a #Q p = 1
proof end;

theorem Th48: :: PREPOWER:48
for a being Real
for p being Rational st a > 0 & p = 1 holds
a #Q p = a
proof end;

Lm5: for a being Real
for k being Integer st a >= 0 holds
a #Z k >= 0

proof end;

theorem Th49: :: PREPOWER:49
for a being Real
for n being Nat st 0 <= a holds
a #Q n = a |^ n
proof end;

theorem Th50: :: PREPOWER:50
for a being Real
for p being Rational
for n being Nat st n >= 1 & p = n " holds
a #Q p = n -Root a
proof end;

theorem Th51: :: PREPOWER:51
for p being Rational holds 1 #Q p = 1
proof end;

theorem Th52: :: PREPOWER:52
for a being Real
for p being Rational st a > 0 holds
a #Q p > 0
proof end;

theorem Th53: :: PREPOWER:53
for a being Real
for p, q being Rational st a > 0 holds
(a #Q p) * (a #Q q) = a #Q (p + q)
proof end;

theorem Th54: :: PREPOWER:54
for a being Real
for p being Rational st a > 0 holds
1 / (a #Q p) = a #Q (- p)
proof end;

theorem Th55: :: PREPOWER:55
for a being Real
for p, q being Rational st a > 0 holds
(a #Q p) / (a #Q q) = a #Q (p - q)
proof end;

theorem Th56: :: PREPOWER:56
for a, b being Real
for p being Rational st a > 0 & b > 0 holds
(a * b) #Q p = (a #Q p) * (b #Q p)
proof end;

theorem Th57: :: PREPOWER:57
for a being Real
for p being Rational st a > 0 holds
(1 / a) #Q p = 1 / (a #Q p)
proof end;

theorem Th58: :: PREPOWER:58
for a, b being Real
for p being Rational st a > 0 & b > 0 holds
(a / b) #Q p = (a #Q p) / (b #Q p)
proof end;

theorem Th59: :: PREPOWER:59
for a being Real
for p, q being Rational st a > 0 holds
(a #Q p) #Q q = a #Q (p * q)
proof end;

theorem Th60: :: PREPOWER:60
for a being Real
for p being Rational st a >= 1 & p >= 0 holds
a #Q p >= 1
proof end;

theorem Th61: :: PREPOWER:61
for a being Real
for p being Rational st a >= 1 & p <= 0 holds
a #Q p <= 1
proof end;

theorem Th62: :: PREPOWER:62
for a being Real
for p being Rational st a > 1 & p > 0 holds
a #Q p > 1
proof end;

theorem Th63: :: PREPOWER:63
for a being Real
for p, q being Rational st a >= 1 & p >= q holds
a #Q p >= a #Q q
proof end;

theorem Th64: :: PREPOWER:64
for a being Real
for p, q being Rational st a > 1 & p > q holds
a #Q p > a #Q q
proof end;

theorem Th65: :: PREPOWER:65
for a being Real
for p being Rational st a > 0 & a < 1 & p > 0 holds
a #Q p < 1
proof end;

theorem :: PREPOWER:66
for a being Real
for p being Rational st a > 0 & a <= 1 & p <= 0 holds
a #Q p >= 1
proof end;

registration
cluster Relation-like omega -defined REAL -valued RAT -valued Function-like V11() V14( omega ) V18( omega , REAL ) complex-valued ext-real-valued real-valued for Element of K19(K20(omega,REAL));
existence
ex b1 being Real_Sequence st b1 is RAT -valued
proof end;
end;

definition
mode Rational_Sequence is RAT -valued Real_Sequence;
end;

:: deftheorem PREPOWER:def 5 :
canceled;

theorem Th67: :: PREPOWER:67
for a being Real ex s being Rational_Sequence st
( s is convergent & lim s = a & ( for n being Nat holds s . n <= a ) )
proof end;

theorem Th68: :: PREPOWER:68
for a being Real ex s being Rational_Sequence st
( s is convergent & lim s = a & ( for n being Nat holds s . n >= a ) )
proof end;

definition
let a be Real;
let s be Rational_Sequence;
func a #Q s -> Real_Sequence means :Def5: :: PREPOWER:def 6
for n being Nat holds it . n = a #Q (s . n);
existence
ex b1 being Real_Sequence st
for n being Nat holds b1 . n = a #Q (s . n)
proof end;
uniqueness
for b1, b2 being Real_Sequence st ( for n being Nat holds b1 . n = a #Q (s . n) ) & ( for n being Nat holds b2 . n = a #Q (s . n) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def5 defines #Q PREPOWER:def 6 :
for a being Real
for s being Rational_Sequence
for b3 being Real_Sequence holds
( b3 = a #Q s iff for n being Nat holds b3 . n = a #Q (s . n) );

Lm6: for s being Rational_Sequence
for a being Real st s is convergent & a >= 1 holds
a #Q s is convergent

proof end;

theorem Th69: :: PREPOWER:69
for a being Real
for s being Rational_Sequence st s is convergent & a > 0 holds
a #Q s is convergent
proof end;

Lm7: for s1, s2 being Rational_Sequence
for a being Real st s1 is convergent & s2 is convergent & lim s1 = lim s2 & a >= 1 holds
lim (a #Q s1) = lim (a #Q s2)

proof end;

theorem Th70: :: PREPOWER:70
for s1, s2 being Rational_Sequence
for a being Real st s1 is convergent & s2 is convergent & lim s1 = lim s2 & a > 0 holds
( a #Q s1 is convergent & a #Q s2 is convergent & lim (a #Q s1) = lim (a #Q s2) )
proof end;

definition
let a, b be Real;
assume A1: a > 0 ;
func a #R b -> Real means :Def6: :: PREPOWER:def 7
ex s being Rational_Sequence st
( s is convergent & lim s = b & a #Q s is convergent & lim (a #Q s) = it );
existence
ex b1 being Real ex s being Rational_Sequence st
( s is convergent & lim s = b & a #Q s is convergent & lim (a #Q s) = b1 )
proof end;
uniqueness
for b1, b2 being Real st ex s being Rational_Sequence st
( s is convergent & lim s = b & a #Q s is convergent & lim (a #Q s) = b1 ) & ex s being Rational_Sequence st
( s is convergent & lim s = b & a #Q s is convergent & lim (a #Q s) = b2 ) holds
b1 = b2
by A1, Th70;
end;

:: deftheorem Def6 defines #R PREPOWER:def 7 :
for a, b being Real st a > 0 holds
for b3 being Real holds
( b3 = a #R b iff ex s being Rational_Sequence st
( s is convergent & lim s = b & a #Q s is convergent & lim (a #Q s) = b3 ) );

theorem Th71: :: PREPOWER:71
for a being Real st a > 0 holds
a #R 0 = 1
proof end;

theorem :: PREPOWER:72
for a being Real st a > 0 holds
a #R 1 = a
proof end;

theorem Th73: :: PREPOWER:73
for a being Real holds 1 #R a = 1
proof end;

theorem Th74: :: PREPOWER:74
for a being Real
for p being Rational st a > 0 holds
a #R p = a #Q p
proof end;

theorem Th75: :: PREPOWER:75
for a, b, c being Real st a > 0 holds
a #R (b + c) = (a #R b) * (a #R c)
proof end;

Lm8: for a, b being Real st a > 0 holds
a #R b >= 0

proof end;

Lm9: for a, b being Real st a > 0 holds
a #R b <> 0

proof end;

theorem Th76: :: PREPOWER:76
for a, c being Real st a > 0 holds
a #R (- c) = 1 / (a #R c)
proof end;

theorem Th77: :: PREPOWER:77
for a, b, c being Real st a > 0 holds
a #R (b - c) = (a #R b) / (a #R c)
proof end;

theorem Th78: :: PREPOWER:78
for a, b, c being Real st a > 0 & b > 0 holds
(a * b) #R c = (a #R c) * (b #R c)
proof end;

theorem Th79: :: PREPOWER:79
for a, c being Real st a > 0 holds
(1 / a) #R c = 1 / (a #R c)
proof end;

theorem :: PREPOWER:80
for a, b, c being Real st a > 0 & b > 0 holds
(a / b) #R c = (a #R c) / (b #R c)
proof end;

theorem Th81: :: PREPOWER:81
for a, b being Real st a > 0 holds
a #R b > 0
proof end;

theorem Th82: :: PREPOWER:82
for a, b, c being Real st a >= 1 & c >= b holds
a #R c >= a #R b
proof end;

theorem Th83: :: PREPOWER:83
for a, b, c being Real st a > 1 & c > b holds
a #R c > a #R b
proof end;

theorem Th84: :: PREPOWER:84
for a, b, c being Real st a > 0 & a <= 1 & c >= b holds
a #R c <= a #R b
proof end;

theorem Th85: :: PREPOWER:85
for a, b being Real st a >= 1 & b >= 0 holds
a #R b >= 1
proof end;

theorem Th86: :: PREPOWER:86
for a, b being Real st a > 1 & b > 0 holds
a #R b > 1
proof end;

theorem Th87: :: PREPOWER:87
for a, b being Real st a >= 1 & b <= 0 holds
a #R b <= 1
proof end;

theorem :: PREPOWER:88
for a, b being Real st a > 1 & b < 0 holds
a #R b < 1
proof end;

Lm10: for p being Rational
for s1, s2 being Real_Sequence st s1 is convergent & s2 is convergent & lim s1 > 0 & p >= 0 & ( for n being Nat holds
( s1 . n > 0 & s2 . n = (s1 . n) #Q p ) ) holds
lim s2 = (lim s1) #Q p

proof end;

theorem Th89: :: PREPOWER:89
for p being Rational
for s1, s2 being Real_Sequence st s1 is convergent & s2 is convergent & lim s1 > 0 & ( for n being Nat holds
( s1 . n > 0 & s2 . n = (s1 . n) #Q p ) ) holds
lim s2 = (lim s1) #Q p
proof end;

Lm11: for a, b being Real
for p being Rational st a > 0 holds
(a #R b) #Q p = a #R (b * p)

proof end;

Lm12: for a being Real
for s1, s2 being Real_Sequence st a >= 1 & s1 is convergent & s2 is convergent & ( for n being Nat holds s2 . n = a #R (s1 . n) ) holds
lim s2 = a #R (lim s1)

proof end;

theorem Th90: :: PREPOWER:90
for a being Real
for s1, s2 being Real_Sequence st a > 0 & s1 is convergent & s2 is convergent & ( for n being Nat holds s2 . n = a #R (s1 . n) ) holds
lim s2 = a #R (lim s1)
proof end;

theorem :: PREPOWER:91
for a, b, c being Real st a > 0 holds
(a #R b) #R c = a #R (b * c)
proof end;

theorem :: PREPOWER:92
for r, u being Real st r > 0 & u > 0 holds
ex k being Nat st u / (2 |^ k) <= r
proof end;

theorem :: PREPOWER:93
for n being Nat
for r being Real
for k being Nat st k >= n & r >= 1 holds
r |^ k >= r |^ n
proof end;

:: from SCPINVAR, 2006.03.14, A.T.
theorem Th94: :: PREPOWER:94
for n, m, l being Nat st n divides m & n divides l holds
n divides m - l
proof end;

theorem :: PREPOWER:95
for m, n being Nat holds
( m divides n iff m divides n ) ;

theorem Th96: :: PREPOWER:96
for m, n being Nat holds m gcd n = m gcd |.(n - m).|
proof end;

theorem :: PREPOWER:97
for a, b being Integer st a >= 0 & b >= 0 holds
a gcd b = a gcd (b - a)
proof end;

:: missing, 2008.03.05, A.T.
theorem :: PREPOWER:98
for a being Real
for l being Integer st a >= 0 holds
a #Z l >= 0 by Lm5;

theorem :: PREPOWER:99
for a being Real
for l being Integer st a > 0 holds
a #Q l = a #Z l
proof end;

:: missing, 2010.02.13, A.T.
theorem :: PREPOWER:100
for l being Integer st l <> 0 holds
0 #Z l = 0
proof end;