:: The Complex Numbers
:: by Czes{\l}aw Byli\'nski
::
:: Received March 1, 1990
:: Copyright (c) 1990-2021 Association of Mizar Users


:: Auxiliary theorems
theorem Th1: :: COMPLEX1:1
for a, b being Real st (a ^2) + (b ^2) = 0 holds
a = 0
proof end;

Lm1: for a, b being Real holds 0 <= (a ^2) + (b ^2)
proof end;

:: Complex Numbers
definition
let z be Complex;
func Re z -> number means :Def1: :: COMPLEX1:def 1
it = z if z is real
otherwise ex f being Function of 2,REAL st
( z = f & it = f . 0 );
existence
( ( z is real implies ex b1 being number st b1 = z ) & ( not z is real implies ex b1 being number ex f being Function of 2,REAL st
( z = f & b1 = f . 0 ) ) )
proof end;
uniqueness
for b1, b2 being number holds
( ( z is real & b1 = z & b2 = z implies b1 = b2 ) & ( not z is real & ex f being Function of 2,REAL st
( z = f & b1 = f . 0 ) & ex f being Function of 2,REAL st
( z = f & b2 = f . 0 ) implies b1 = b2 ) )
;
consistency
for b1 being number holds verum
;
func Im z -> number means :Def2: :: COMPLEX1:def 2
it = 0 if z is real
otherwise ex f being Function of 2,REAL st
( z = f & it = f . 1 );
existence
( ( z is real implies ex b1 being number st b1 = 0 ) & ( not z is real implies ex b1 being number ex f being Function of 2,REAL st
( z = f & b1 = f . 1 ) ) )
proof end;
uniqueness
for b1, b2 being number holds
( ( z is real & b1 = 0 & b2 = 0 implies b1 = b2 ) & ( not z is real & ex f being Function of 2,REAL st
( z = f & b1 = f . 1 ) & ex f being Function of 2,REAL st
( z = f & b2 = f . 1 ) implies b1 = b2 ) )
;
consistency
for b1 being number holds verum
;
end;

:: deftheorem Def1 defines Re COMPLEX1:def 1 :
for z being Complex
for b2 being number holds
( ( z is real implies ( b2 = Re z iff b2 = z ) ) & ( not z is real implies ( b2 = Re z iff ex f being Function of 2,REAL st
( z = f & b2 = f . 0 ) ) ) );

:: deftheorem Def2 defines Im COMPLEX1:def 2 :
for z being Complex
for b2 being number holds
( ( z is real implies ( b2 = Im z iff b2 = 0 ) ) & ( not z is real implies ( b2 = Im z iff ex f being Function of 2,REAL st
( z = f & b2 = f . 1 ) ) ) );

registration
let z be Complex;
cluster Re z -> real ;
coherence
Re z is real
proof end;
cluster Im z -> real ;
coherence
Im z is real
proof end;
end;

definition
let z be Complex;
:: original: Re
redefine func Re z -> Element of REAL ;
coherence
Re z is Element of REAL
by XREAL_0:def 1;
:: original: Im
redefine func Im z -> Element of REAL ;
coherence
Im z is Element of REAL
by XREAL_0:def 1;
end;

registration
let r be Real;
cluster Im r -> zero ;
coherence
Im r is zero
by Def2;
end;

theorem Th2: :: COMPLEX1:2
for f being Function of 2,REAL ex a, b being Element of REAL st f = (0,1) --> (a,b)
proof end;

Lm2: for a, b being Element of REAL holds
( Re [*a,b*] = a & Im [*a,b*] = b )

proof end;

Lm3: for z being Complex holds [*(Re z),(Im z)*] = z
proof end;

theorem Th3: :: COMPLEX1:3
for z1, z2 being Complex st Re z1 = Re z2 & Im z1 = Im z2 holds
z1 = z2
proof end;

definition
let z1, z2 be Complex;
redefine pred z1 = z2 means :: COMPLEX1:def 3
( Re z1 = Re z2 & Im z1 = Im z2 );
compatibility
( z1 = z2 iff ( Re z1 = Re z2 & Im z1 = Im z2 ) )
by Th3;
end;

:: deftheorem defines = COMPLEX1:def 3 :
for z1, z2 being Complex holds
( z1 = z2 iff ( Re z1 = Re z2 & Im z1 = Im z2 ) );

notation
synonym 0c for 0 ;
end;

definition
:: original: 0c
redefine func 0c -> Element of COMPLEX ;
correctness
coherence
0c is Element of COMPLEX
;
by XCMPLX_0:def 2;
end;

definition
func 1r -> Element of COMPLEX equals :: COMPLEX1:def 4
1;
correctness
coherence
1 is Element of COMPLEX
;
by XCMPLX_0:def 2;
:: original: <i>
redefine func <i> -> Element of COMPLEX ;
coherence
<i> is Element of COMPLEX
by XCMPLX_0:def 2;
end;

:: deftheorem defines 1r COMPLEX1:def 4 :
1r = 1;

theorem Th4: :: COMPLEX1:4
( Re 0 = 0 & Im 0 = 0 )
proof end;

theorem Th5: :: COMPLEX1:5
for z being Complex holds
( z = 0 iff ((Re z) ^2) + ((Im z) ^2) = 0 ) by Th4, Th1;

theorem Th6: :: COMPLEX1:6
( Re 1r = 1 & Im 1r = 0 )
proof end;

theorem Th7: :: COMPLEX1:7
( Re <i> = 0 & Im <i> = 1 )
proof end;

Lm7: for x being Real
for x1, x2 being Element of REAL st x = [*x1,x2*] holds
( x2 = 0 & x = x1 )

proof end;

Lm8: for x9, y9 being Element of REAL
for x, y being Real st x9 = x & y9 = y holds
+ (x9,y9) = x + y

proof end;

Lm9: for x being Element of REAL holds opp x = - x
proof end;

Lm10: for x9, y9 being Element of REAL
for x, y being Real st x9 = x & y9 = y holds
* (x9,y9) = x * y

proof end;

Lm11: for x, y, z being Complex st z = x + y holds
Re z = (Re x) + (Re y)

proof end;

Lm12: for x, y, z being Complex st z = x + y holds
Im z = (Im x) + (Im y)

proof end;

Lm13: for x, y, z being Complex st z = x * y holds
Re z = ((Re x) * (Re y)) - ((Im x) * (Im y))

proof end;

Lm14: for x, y, z being Complex st z = x * y holds
Im z = ((Re x) * (Im y)) + ((Im x) * (Re y))

proof end;

Lm15: for z1, z2 being Complex holds z1 + z2 = [*((Re z1) + (Re z2)),((Im z1) + (Im z2))*]
proof end;

Lm16: for z1, z2 being Complex holds z1 * z2 = [*(((Re z1) * (Re z2)) - ((Im z1) * (Im z2))),(((Re z1) * (Im z2)) + ((Re z2) * (Im z1)))*]
proof end;

Lm17: for z1, z2 being Complex holds
( Re (z1 * z2) = ((Re z1) * (Re z2)) - ((Im z1) * (Im z2)) & Im (z1 * z2) = ((Re z1) * (Im z2)) + ((Re z2) * (Im z1)) )

proof end;

Lm18: for z1, z2 being Complex holds
( Re (z1 + z2) = (Re z1) + (Re z2) & Im (z1 + z2) = (Im z1) + (Im z2) )

proof end;

Lm19: for x being Element of REAL holds Re (x * <i>) = 0
proof end;

Lm20: for x being Element of REAL holds Im (x * <i>) = x
proof end;

Lm21: for x, y being Element of REAL holds [*x,y*] = x + (y * <i>)
proof end;

definition
end;

:: deftheorem COMPLEX1:def 5 :
canceled;

theorem Th8: :: COMPLEX1:8
for z1, z2 being Complex holds
( Re (z1 + z2) = (Re z1) + (Re z2) & Im (z1 + z2) = (Im z1) + (Im z2) )
proof end;

definition
end;

:: deftheorem COMPLEX1:def 6 :
canceled;

theorem Th9: :: COMPLEX1:9
for z1, z2 being Complex holds
( Re (z1 * z2) = ((Re z1) * (Re z2)) - ((Im z1) * (Im z2)) & Im (z1 * z2) = ((Re z1) * (Im z2)) + ((Re z2) * (Im z1)) )
proof end;

theorem :: COMPLEX1:10
for a being Real holds Re (a * <i>) = 0
proof end;

theorem :: COMPLEX1:11
for a being Real holds Im (a * <i>) = a
proof end;

theorem Th12: :: COMPLEX1:12
for a, b being Real holds
( Re (a + (b * <i>)) = a & Im (a + (b * <i>)) = b )
proof end;

theorem Th13: :: COMPLEX1:13
for z being Complex holds (Re z) + ((Im z) * <i>) = z
proof end;

theorem Th14: :: COMPLEX1:14
for z1, z2 being Complex st Im z1 = 0 & Im z2 = 0 holds
( Re (z1 * z2) = (Re z1) * (Re z2) & Im (z1 * z2) = 0 )
proof end;

theorem Th15: :: COMPLEX1:15
for z1, z2 being Complex st Re z1 = 0 & Re z2 = 0 holds
( Re (z1 * z2) = - ((Im z1) * (Im z2)) & Im (z1 * z2) = 0 )
proof end;

theorem :: COMPLEX1:16
for z being Complex holds
( Re (z * z) = ((Re z) ^2) - ((Im z) ^2) & Im (z * z) = 2 * ((Re z) * (Im z)) )
proof end;

definition
end;

:: deftheorem COMPLEX1:def 7 :
canceled;

Lm22: for z being Complex holds - z = (- (Re z)) + ((- (Im z)) * <i>)
proof end;

theorem Th17: :: COMPLEX1:17
for z being Complex holds
( Re (- z) = - (Re z) & Im (- z) = - (Im z) )
proof end;

theorem :: COMPLEX1:18
<i> * <i> = - 1r ;

definition
end;

:: deftheorem COMPLEX1:def 8 :
canceled;

Lm23: for z1, z2 being Complex holds z1 - z2 = ((Re z1) - (Re z2)) + (((Im z1) - (Im z2)) * <i>)
proof end;

theorem Th19: :: COMPLEX1:19
for z1, z2 being Complex holds
( Re (z1 - z2) = (Re z1) - (Re z2) & Im (z1 - z2) = (Im z1) - (Im z2) )
proof end;

definition
end;

:: deftheorem COMPLEX1:def 9 :
canceled;

Lm24: for z being Complex holds z " = ((Re z) / (((Re z) ^2) + ((Im z) ^2))) + (((- (Im z)) / (((Re z) ^2) + ((Im z) ^2))) * <i>)
proof end;

theorem Th20: :: COMPLEX1:20
for z being Complex holds
( Re (z ") = (Re z) / (((Re z) ^2) + ((Im z) ^2)) & Im (z ") = (- (Im z)) / (((Re z) ^2) + ((Im z) ^2)) )
proof end;

theorem :: COMPLEX1:21
<i> " = - <i> ;

theorem Th22: :: COMPLEX1:22
for z being Complex st Re z <> 0 & Im z = 0 holds
( Re (z ") = (Re z) " & Im (z ") = 0 )
proof end;

theorem Th23: :: COMPLEX1:23
for z being Complex st Re z = 0 & Im z <> 0 holds
( Re (z ") = 0 & Im (z ") = - ((Im z) ") )
proof end;

definition
end;

:: deftheorem COMPLEX1:def 10 :
canceled;

Lm25: for z1, z2 being Complex holds z1 / z2 = ((((Re z1) * (Re z2)) + ((Im z1) * (Im z2))) / (((Re z2) ^2) + ((Im z2) ^2))) + (((((Re z2) * (Im z1)) - ((Re z1) * (Im z2))) / (((Re z2) ^2) + ((Im z2) ^2))) * <i>)
proof end;

theorem :: COMPLEX1:24
for z1, z2 being Complex holds
( Re (z1 / z2) = (((Re z1) * (Re z2)) + ((Im z1) * (Im z2))) / (((Re z2) ^2) + ((Im z2) ^2)) & Im (z1 / z2) = (((Re z2) * (Im z1)) - ((Re z1) * (Im z2))) / (((Re z2) ^2) + ((Im z2) ^2)) )
proof end;

theorem :: COMPLEX1:25
for z1, z2 being Complex st Im z1 = 0 & Im z2 = 0 & Re z2 <> 0 holds
( Re (z1 / z2) = (Re z1) / (Re z2) & Im (z1 / z2) = 0 )
proof end;

theorem :: COMPLEX1:26
for z1, z2 being Complex st Re z1 = 0 & Re z2 = 0 & Im z2 <> 0 holds
( Re (z1 / z2) = (Im z1) / (Im z2) & Im (z1 / z2) = 0 )
proof end;

definition
let z be Complex;
func z *' -> Complex equals :: COMPLEX1:def 11
(Re z) - ((Im z) * <i>);
correctness
coherence
(Re z) - ((Im z) * <i>) is Complex
;
;
involutiveness
for b1, b2 being Complex st b1 = (Re b2) - ((Im b2) * <i>) holds
b2 = (Re b1) - ((Im b1) * <i>)
proof end;
end;

:: deftheorem defines *' COMPLEX1:def 11 :
for z being Complex holds z *' = (Re z) - ((Im z) * <i>);

theorem Th27: :: COMPLEX1:27
for z being Complex holds
( Re (z *') = Re z & Im (z *') = - (Im z) )
proof end;

theorem :: COMPLEX1:28
0 *' = 0 by Th4;

theorem :: COMPLEX1:29
for z being Complex st z *' = 0 holds
z = 0
proof end;

theorem :: COMPLEX1:30
1r *' = 1r by Th6;

theorem :: COMPLEX1:31
<i> *' = - <i> by Th7;

theorem Th32: :: COMPLEX1:32
for z1, z2 being Complex holds (z1 + z2) *' = (z1 *') + (z2 *')
proof end;

theorem Th33: :: COMPLEX1:33
for z being Complex holds (- z) *' = - (z *')
proof end;

theorem :: COMPLEX1:34
for z1, z2 being Complex holds (z1 - z2) *' = (z1 *') - (z2 *')
proof end;

theorem Th35: :: COMPLEX1:35
for z1, z2 being Complex holds (z1 * z2) *' = (z1 *') * (z2 *')
proof end;

theorem Th36: :: COMPLEX1:36
for z being Complex holds (z ") *' = (z *') "
proof end;

theorem :: COMPLEX1:37
for z1, z2 being Complex holds (z1 / z2) *' = (z1 *') / (z2 *')
proof end;

theorem Th38: :: COMPLEX1:38
for z being Complex st Im z = 0 holds
z *' = z by Th27;

registration
let r be Real;
reduce r *' to r;
reducibility
r *' = r
by Th38;
end;

theorem :: COMPLEX1:39
for z being Complex st Re z = 0 holds
z *' = - z
proof end;

theorem :: COMPLEX1:40
for z being Complex holds
( Re (z * (z *')) = ((Re z) ^2) + ((Im z) ^2) & Im (z * (z *')) = 0 )
proof end;

theorem :: COMPLEX1:41
for z being Complex holds
( Re (z + (z *')) = 2 * (Re z) & Im (z + (z *')) = 0 )
proof end;

theorem :: COMPLEX1:42
for z being Complex holds
( Re (z - (z *')) = 0 & Im (z - (z *')) = 2 * (Im z) )
proof end;

definition
let z be Complex;
func |.z.| -> Real equals :: COMPLEX1:def 12
sqrt (((Re z) ^2) + ((Im z) ^2));
coherence
sqrt (((Re z) ^2) + ((Im z) ^2)) is Real
;
projectivity
for b1 being Real
for b2 being Complex st b1 = sqrt (((Re b2) ^2) + ((Im b2) ^2)) holds
b1 = sqrt (((Re b1) ^2) + ((Im b1) ^2))
proof end;
end;

:: deftheorem defines |. COMPLEX1:def 12 :
for z being Complex holds |.z.| = sqrt (((Re z) ^2) + ((Im z) ^2));

theorem Th43: :: COMPLEX1:43
for a being Real st a >= 0 holds
|.a.| = a
proof end;

registration
let z be zero Complex;
cluster |.z.| -> zero ;
coherence
|.z.| is zero
by Th4, SQUARE_1:17;
end;

theorem :: COMPLEX1:44
|.0.| = 0 ;

registration
let z be non zero Complex;
cluster |.z.| -> non zero ;
coherence
not |.z.| is zero
proof end;
end;

theorem :: COMPLEX1:45
for z being Complex st |.z.| = 0 holds
z = 0 ;

registration
let z be Complex;
cluster |.z.| -> non negative ;
coherence
not |.z.| is negative
proof end;
end;

theorem :: COMPLEX1:46
for z being Complex holds 0 <= |.z.| ;

theorem :: COMPLEX1:47
for z being Complex holds
( z <> 0 iff 0 < |.z.| ) ;

theorem Th48: :: COMPLEX1:48
|.1r.| = 1 by Th6, SQUARE_1:18;

theorem :: COMPLEX1:49
|.<i>.| = 1 by Th7, SQUARE_1:18;

Lm26: for z being Complex holds |.(- z).| = |.z.|
proof end;

Lm27: for a being Real st a <= 0 holds
|.a.| = - a

proof end;

Lm28: for a being Real holds sqrt (a ^2) = |.a.|
proof end;

theorem :: COMPLEX1:50
for z being Complex st Im z = 0 holds
|.z.| = |.(Re z).| by Lm28;

theorem :: COMPLEX1:51
for z being Complex st Re z = 0 holds
|.z.| = |.(Im z).| by Lm28;

theorem :: COMPLEX1:52
for z being Complex holds |.(- z).| = |.z.| by Lm26;

theorem Th53: :: COMPLEX1:53
for z being Complex holds |.(z *').| = |.z.|
proof end;

Lm29: for a being Real holds
( - |.a.| <= a & a <= |.a.| )

proof end;

theorem :: COMPLEX1:54
for z being Complex holds Re z <= |.z.|
proof end;

theorem :: COMPLEX1:55
for z being Complex holds Im z <= |.z.|
proof end;

theorem Th56: :: COMPLEX1:56
for z1, z2 being Complex holds |.(z1 + z2).| <= |.z1.| + |.z2.|
proof end;

theorem Th57: :: COMPLEX1:57
for z1, z2 being Complex holds |.(z1 - z2).| <= |.z1.| + |.z2.|
proof end;

theorem :: COMPLEX1:58
for z1, z2 being Complex holds |.z1.| - |.z2.| <= |.(z1 + z2).|
proof end;

theorem Th59: :: COMPLEX1:59
for z1, z2 being Complex holds |.z1.| - |.z2.| <= |.(z1 - z2).|
proof end;

theorem Th60: :: COMPLEX1:60
for z1, z2 being Complex holds |.(z1 - z2).| = |.(z2 - z1).|
proof end;

theorem Th61: :: COMPLEX1:61
for z1, z2 being Complex holds
( |.(z1 - z2).| = 0 iff z1 = z2 )
proof end;

theorem :: COMPLEX1:62
for z1, z2 being Complex holds
( z1 <> z2 iff 0 < |.(z1 - z2).| )
proof end;

theorem :: COMPLEX1:63
for z, z1, z2 being Complex holds |.(z1 - z2).| <= |.(z1 - z).| + |.(z - z2).|
proof end;

Lm30: for a, b being Real holds
( ( - b <= a & a <= b ) iff |.a.| <= b )

proof end;

theorem :: COMPLEX1:64
for z1, z2 being Complex holds |.(|.z1.| - |.z2.|).| <= |.(z1 - z2).|
proof end;

theorem Th65: :: COMPLEX1:65
for z1, z2 being Complex holds |.(z1 * z2).| = |.z1.| * |.z2.|
proof end;

theorem Th66: :: COMPLEX1:66
for z being Complex holds |.(z ").| = |.z.| "
proof end;

theorem Th67: :: COMPLEX1:67
for z1, z2 being Complex holds |.z1.| / |.z2.| = |.(z1 / z2).|
proof end;

theorem :: COMPLEX1:68
for z being Complex holds |.(z * z).| = ((Re z) ^2) + ((Im z) ^2)
proof end;

theorem :: COMPLEX1:69
for z being Complex holds |.(z * z).| = |.(z * (z *')).|
proof end;

:: Originally from SQUARE_1
theorem :: COMPLEX1:70
for a being Real st a <= 0 holds
|.a.| = - a by Lm27;

theorem Th71: :: COMPLEX1:71
for a being Real holds
( |.a.| = a or |.a.| = - a )
proof end;

theorem :: COMPLEX1:72
for a being Real holds sqrt (a ^2) = |.a.| by Lm28;

theorem :: COMPLEX1:73
for a, b being Real holds min (a,b) = ((a + b) - |.(a - b).|) / 2
proof end;

theorem :: COMPLEX1:74
for a, b being Real holds max (a,b) = ((a + b) + |.(a - b).|) / 2
proof end;

theorem Th75: :: COMPLEX1:75
for a being Real holds |.a.| ^2 = a ^2
proof end;

theorem :: COMPLEX1:76
for a being Real holds
( - |.a.| <= a & a <= |.a.| ) by Lm29;

theorem :: COMPLEX1:77
for a, b, c, d being Real st a + (b * <i>) = c + (d * <i>) holds
( a = c & b = d )
proof end;

:: from JGRAPH_1, 29.12.2006, AK
theorem :: COMPLEX1:78
for a, b being Real holds sqrt ((a ^2) + (b ^2)) <= |.a.| + |.b.|
proof end;

theorem :: COMPLEX1:79
for a, b being Real holds |.a.| <= sqrt ((a ^2) + (b ^2))
proof end;

theorem :: COMPLEX1:80
for z1 being Complex holds |.(1 / z1).| = 1 / |.z1.| by Th48, Th67;

theorem :: COMPLEX1:81
for z1, z2 being Complex holds z1 + z2 = ((Re z1) + (Re z2)) + (((Im z1) + (Im z2)) * <i>)
proof end;

theorem :: COMPLEX1:82
for z1, z2 being Complex holds z1 * z2 = (((Re z1) * (Re z2)) - ((Im z1) * (Im z2))) + ((((Re z1) * (Im z2)) + ((Re z2) * (Im z1))) * <i>)
proof end;

theorem :: COMPLEX1:83
for z being Complex holds - z = (- (Re z)) + ((- (Im z)) * <i>) by Lm22;

theorem :: COMPLEX1:84
for z1, z2 being Complex holds z1 - z2 = ((Re z1) - (Re z2)) + (((Im z1) - (Im z2)) * <i>) by Lm23;

theorem :: COMPLEX1:85
for z being Complex holds z " = ((Re z) / (((Re z) ^2) + ((Im z) ^2))) + (((- (Im z)) / (((Re z) ^2) + ((Im z) ^2))) * <i>) by Lm24;

theorem :: COMPLEX1:86
for z1, z2 being Complex holds z1 / z2 = ((((Re z1) * (Re z2)) + ((Im z1) * (Im z2))) / (((Re z2) ^2) + ((Im z2) ^2))) + (((((Re z2) * (Im z1)) - ((Re z1) * (Im z2))) / (((Re z2) ^2) + ((Im z2) ^2))) * <i>) by Lm25;