:: Inner Products and Angles of Complex Numbers
:: by Wenpai Chang, Yatsuka Nakamura and Piotr Rudnicki
::
:: Received May 29, 2003
:: Copyright (c) 2003-2021 Association of Mizar Users


theorem Th1: :: COMPLEX2:1
for a, b being Real st b > 0 holds
ex r being Real st
( r = (b * (- [\(a / b)/])) + a & 0 <= r & r < b )
proof end;

theorem Th2: :: COMPLEX2:2
for a, b, c being Real st a > 0 & b >= 0 & c >= 0 & b < a & c < a holds
for i being Integer st b = c + (a * i) holds
b = c
proof end;

theorem Th3: :: COMPLEX2:3
for a, b being Real holds
( sin (a - b) = ((sin a) * (cos b)) - ((cos a) * (sin b)) & cos (a - b) = ((cos a) * (cos b)) + ((sin a) * (sin b)) )
proof end;

theorem :: COMPLEX2:4
for a being Real holds
( sin . (a - PI) = - (sin . a) & cos . (a - PI) = - (cos . a) )
proof end;

theorem Th5: :: COMPLEX2:5
for a being Real holds
( sin (a - PI) = - (sin a) & cos (a - PI) = - (cos a) )
proof end;

theorem Th6: :: COMPLEX2:6
for a, b being Real st a in ].0,(PI / 2).[ & b in ].0,(PI / 2).[ holds
( a < b iff sin a < sin b )
proof end;

theorem Th7: :: COMPLEX2:7
for a, b being Real st a in ].(PI / 2),PI.[ & b in ].(PI / 2),PI.[ holds
( a < b iff sin a > sin b )
proof end;

theorem Th8: :: COMPLEX2:8
for a being Real
for i being Integer holds sin a = sin (((2 * PI) * i) + a)
proof end;

theorem Th9: :: COMPLEX2:9
for a being Real
for i being Integer holds cos a = cos (((2 * PI) * i) + a)
proof end;

theorem Th10: :: COMPLEX2:10
for a being Real st sin a = 0 holds
cos a <> 0
proof end;

theorem Th11: :: COMPLEX2:11
for a, b being Real st 0 <= a & a < 2 * PI & 0 <= b & b < 2 * PI & sin a = sin b & cos a = cos b holds
a = b
proof end;

Lm1: Arg 0 = 0
by COMPTRIG:35;

:: ALL these to be changed (mainly deleted) after the change in COMPTRIG
theorem :: COMPLEX2:12
canceled;

::$CT
theorem Th12: :: COMPLEX2:13
for z being Complex st z <> 0 holds
( ( Arg z < PI implies Arg (- z) = (Arg z) + PI ) & ( Arg z >= PI implies Arg (- z) = (Arg z) - PI ) )
proof end;

theorem :: COMPLEX2:14
canceled;

::$CT
theorem Th13: :: COMPLEX2:15
for z being Complex holds
( Arg z = 0 iff z = |.z.| )
proof end;

theorem Th14: :: COMPLEX2:16
for z being Complex st z <> 0 holds
( Arg z < PI iff Arg (- z) >= PI )
proof end;

theorem Th15: :: COMPLEX2:17
for x, y being Complex st ( x <> y or x - y <> 0 ) holds
( Arg (x - y) < PI iff Arg (y - x) >= PI )
proof end;

theorem Th16: :: COMPLEX2:18
for z being Complex holds
( Arg z in ].0,PI.[ iff Im z > 0 )
proof end;

theorem :: COMPLEX2:19
for z being Complex st Arg z <> 0 holds
( Arg z < PI iff sin (Arg z) > 0 )
proof end;

theorem :: COMPLEX2:20
for x, y being Complex st Arg x < PI & Arg y < PI holds
Arg (x + y) < PI
proof end;

theorem Th19: :: COMPLEX2:21
for z being Complex holds
( Arg z = 0 iff ( Re z >= 0 & Im z = 0 ) )
proof end;

theorem Th20: :: COMPLEX2:22
for z being Complex holds
( Arg z = PI iff ( Re z < 0 & Im z = 0 ) )
proof end;

theorem Th21: :: COMPLEX2:23
for z being Complex holds
( Im z = 0 iff ( Arg z = 0 or Arg z = PI ) )
proof end;

theorem Th22: :: COMPLEX2:24
for z being Complex st Arg z <= PI holds
Im z >= 0
proof end;

theorem Th23: :: COMPLEX2:25
for z being Element of COMPLEX st z <> 0 holds
( cos (Arg (- z)) = - (cos (Arg z)) & sin (Arg (- z)) = - (sin (Arg z)) )
proof end;

theorem Th24: :: COMPLEX2:26
for a being Complex st a <> 0 holds
( cos (Arg a) = (Re a) / |.a.| & sin (Arg a) = (Im a) / |.a.| )
proof end;

theorem Th25: :: COMPLEX2:27
for a being Complex
for r being Real st r > 0 holds
Arg (a * r) = Arg a
proof end;

theorem Th26: :: COMPLEX2:28
for a being Complex
for r being Real st r < 0 holds
Arg (a * r) = Arg (- a)
proof end;

definition
let x, y be Complex;
func x .|. y -> Element of COMPLEX equals :: COMPLEX2:def 1
x * (y *');
correctness
coherence
x * (y *') is Element of COMPLEX
;
by XCMPLX_0:def 2;
end;

:: deftheorem defines .|. COMPLEX2:def 1 :
for x, y being Complex holds x .|. y = x * (y *');

theorem Th27: :: COMPLEX2:29
for x, y being Complex holds x .|. y = (((Re x) * (Re y)) + ((Im x) * (Im y))) + (((- ((Re x) * (Im y))) + ((Im x) * (Re y))) * <i>)
proof end;

theorem Th28: :: COMPLEX2:30
for z being Complex holds
( z .|. z = ((Re z) * (Re z)) + ((Im z) * (Im z)) & z .|. z = ((Re z) ^2) + ((Im z) ^2) )
proof end;

theorem Th29: :: COMPLEX2:31
for z being Complex holds
( z .|. z = |.z.| ^2 & |.z.| ^2 = Re (z .|. z) )
proof end;

theorem Th30: :: COMPLEX2:32
for x, y being Complex holds |.(x .|. y).| = |.x.| * |.y.|
proof end;

theorem :: COMPLEX2:33
for x being Complex st x .|. x = 0 holds
x = 0
proof end;

theorem Th32: :: COMPLEX2:34
for x, y being Complex holds y .|. x = (x .|. y) *'
proof end;

theorem :: COMPLEX2:35
for x, y, z being Complex holds (x + y) .|. z = (x .|. z) + (y .|. z) ;

theorem Th34: :: COMPLEX2:36
for x, y, z being Complex holds x .|. (y + z) = (x .|. y) + (x .|. z)
proof end;

theorem :: COMPLEX2:37
for a, x, y being Complex holds (a * x) .|. y = a * (x .|. y) ;

theorem Th36: :: COMPLEX2:38
for a, x, y being Complex holds x .|. (a * y) = (a *') * (x .|. y)
proof end;

theorem :: COMPLEX2:39
for a, x, y being Complex holds (a * x) .|. y = x .|. ((a *') * y)
proof end;

theorem :: COMPLEX2:40
for a, b, x, y, z being Complex holds ((a * x) + (b * y)) .|. z = (a * (x .|. z)) + (b * (y .|. z)) ;

theorem :: COMPLEX2:41
for a, b, x, y, z being Complex holds x .|. ((a * y) + (b * z)) = ((a *') * (x .|. y)) + ((b *') * (x .|. z))
proof end;

theorem Th40: :: COMPLEX2:42
for x, y being Complex holds (- x) .|. y = x .|. (- y)
proof end;

theorem :: COMPLEX2:43
for x, y being Complex holds (- x) .|. y = - (x .|. y) ;

theorem Th42: :: COMPLEX2:44
for x, y being Complex holds - (x .|. y) = x .|. (- y)
proof end;

theorem :: COMPLEX2:45
for x, y being Complex holds (- x) .|. (- y) = x .|. y
proof end;

theorem :: COMPLEX2:46
for x, y, z being Complex holds (x - y) .|. z = (x .|. z) - (y .|. z) ;

theorem Th45: :: COMPLEX2:47
for x, y, z being Complex holds x .|. (y - z) = (x .|. y) - (x .|. z)
proof end;

theorem :: COMPLEX2:48
for x, y being Complex holds (x + y) .|. (x + y) = (((x .|. x) + (x .|. y)) + (y .|. x)) + (y .|. y)
proof end;

theorem Th47: :: COMPLEX2:49
for x, y being Complex holds (x - y) .|. (x - y) = (((x .|. x) - (x .|. y)) - (y .|. x)) + (y .|. y)
proof end;

Lm2: for z being Complex holds |.z.| ^2 = ((Re z) ^2) + ((Im z) ^2)
proof end;

theorem Th48: :: COMPLEX2:50
for x, y being Complex holds
( Re (x .|. y) = 0 iff ( Im (x .|. y) = |.x.| * |.y.| or Im (x .|. y) = - (|.x.| * |.y.|) ) )
proof end;

definition
let a be Complex;
let r be Real;
func Rotate (a,r) -> Element of COMPLEX equals :: COMPLEX2:def 2
(|.a.| * (cos (r + (Arg a)))) + ((|.a.| * (sin (r + (Arg a)))) * <i>);
correctness
coherence
(|.a.| * (cos (r + (Arg a)))) + ((|.a.| * (sin (r + (Arg a)))) * <i>) is Element of COMPLEX
;
by XCMPLX_0:def 2;
end;

:: deftheorem defines Rotate COMPLEX2:def 2 :
for a being Complex
for r being Real holds Rotate (a,r) = (|.a.| * (cos (r + (Arg a)))) + ((|.a.| * (sin (r + (Arg a)))) * <i>);

theorem :: COMPLEX2:51
for a being Complex holds Rotate (a,0) = a by COMPTRIG:62;

theorem Th50: :: COMPLEX2:52
for r being Real
for a being Complex holds
( Rotate (a,r) = 0 iff a = 0 )
proof end;

theorem Th51: :: COMPLEX2:53
for r being Real
for a being Complex holds |.(Rotate (a,r)).| = |.a.|
proof end;

theorem Th52: :: COMPLEX2:54
for r being Real
for a being Complex st a <> 0 holds
ex i being Integer st Arg (Rotate (a,r)) = ((2 * PI) * i) + (r + (Arg a))
proof end;

theorem :: COMPLEX2:55
for a being Complex holds Rotate (a,(- (Arg a))) = |.a.| by SIN_COS:31;

theorem Th54: :: COMPLEX2:56
for a being Complex
for r being Real holds
( Re (Rotate (a,r)) = ((Re a) * (cos r)) - ((Im a) * (sin r)) & Im (Rotate (a,r)) = ((Re a) * (sin r)) + ((Im a) * (cos r)) )
proof end;

theorem Th55: :: COMPLEX2:57
for a, b being Complex
for r being Real holds Rotate ((a + b),r) = (Rotate (a,r)) + (Rotate (b,r))
proof end;

theorem Th56: :: COMPLEX2:58
for a being Complex
for r being Real holds Rotate ((- a),r) = - (Rotate (a,r))
proof end;

theorem Th57: :: COMPLEX2:59
for a, b being Complex
for r being Real holds Rotate ((a - b),r) = (Rotate (a,r)) - (Rotate (b,r))
proof end;

theorem Th58: :: COMPLEX2:60
for a being Complex holds Rotate (a,PI) = - a
proof end;

definition
let a, b be Complex;
func angle (a,b) -> Real equals :Def3: :: COMPLEX2:def 3
Arg (Rotate (b,(- (Arg a)))) if ( Arg a = 0 or b <> 0 )
otherwise (2 * PI) - (Arg a);
correctness
coherence
( ( ( Arg a = 0 or b <> 0 ) implies Arg (Rotate (b,(- (Arg a)))) is Real ) & ( Arg a = 0 or b <> 0 or (2 * PI) - (Arg a) is Real ) )
;
consistency
for b1 being Real holds verum
;
;
end;

:: deftheorem Def3 defines angle COMPLEX2:def 3 :
for a, b being Complex holds
( ( ( Arg a = 0 or b <> 0 ) implies angle (a,b) = Arg (Rotate (b,(- (Arg a)))) ) & ( Arg a = 0 or b <> 0 or angle (a,b) = (2 * PI) - (Arg a) ) );

theorem Th59: :: COMPLEX2:61
for r being Real
for a being Complex st r >= 0 holds
angle (r,a) = Arg a
proof end;

theorem Th60: :: COMPLEX2:62
for r being Real
for a, b being Complex st Arg a = Arg b & a <> 0 & b <> 0 holds
Arg (Rotate (a,r)) = Arg (Rotate (b,r))
proof end;

theorem Th61: :: COMPLEX2:63
for a, b being Complex
for r being Real st r > 0 holds
angle (a,b) = angle ((a * r),(b * r))
proof end;

theorem Th62: :: COMPLEX2:64
for a, b being Complex st a <> 0 & b <> 0 & Arg a = Arg b holds
Arg (- a) = Arg (- b)
proof end;

theorem Th63: :: COMPLEX2:65
for a, b being Complex
for r being Real st a <> 0 & b <> 0 holds
angle (a,b) = angle ((Rotate (a,r)),(Rotate (b,r)))
proof end;

theorem :: COMPLEX2:66
for a, b being Complex
for r being Real st r < 0 & a <> 0 & b <> 0 holds
angle (a,b) = angle ((a * r),(b * r))
proof end;

theorem :: COMPLEX2:67
for a, b being Complex st a <> 0 & b <> 0 holds
angle (a,b) = angle ((- a),(- b))
proof end;

theorem :: COMPLEX2:68
for a, b being Complex st b <> 0 & angle (a,b) = 0 holds
angle (a,(- b)) = PI
proof end;

theorem Th67: :: COMPLEX2:69
for a, b being Complex st a <> 0 & b <> 0 holds
( cos (angle (a,b)) = (Re (a .|. b)) / (|.a.| * |.b.|) & sin (angle (a,b)) = - ((Im (a .|. b)) / (|.a.| * |.b.|)) )
proof end;

definition
let x, y, z be Complex;
func angle (x,y,z) -> Real equals :Def4: :: COMPLEX2:def 4
(Arg (z - y)) - (Arg (x - y)) if (Arg (z - y)) - (Arg (x - y)) >= 0
otherwise (2 * PI) + ((Arg (z - y)) - (Arg (x - y)));
correctness
coherence
( ( (Arg (z - y)) - (Arg (x - y)) >= 0 implies (Arg (z - y)) - (Arg (x - y)) is Real ) & ( not (Arg (z - y)) - (Arg (x - y)) >= 0 implies (2 * PI) + ((Arg (z - y)) - (Arg (x - y))) is Real ) )
;
consistency
for b1 being Real holds verum
;
;
end;

:: deftheorem Def4 defines angle COMPLEX2:def 4 :
for x, y, z being Complex holds
( ( (Arg (z - y)) - (Arg (x - y)) >= 0 implies angle (x,y,z) = (Arg (z - y)) - (Arg (x - y)) ) & ( not (Arg (z - y)) - (Arg (x - y)) >= 0 implies angle (x,y,z) = (2 * PI) + ((Arg (z - y)) - (Arg (x - y))) ) );

theorem Th68: :: COMPLEX2:70
for x, y, z being Complex holds
( 0 <= angle (x,y,z) & angle (x,y,z) < 2 * PI )
proof end;

theorem Th69: :: COMPLEX2:71
for x, y, z being Complex holds angle (x,y,z) = angle ((x - y),0,(z - y))
proof end;

theorem Th70: :: COMPLEX2:72
for a, b, c, d being Complex holds angle (a,b,c) = angle ((a + d),(b + d),(c + d))
proof end;

theorem Th71: :: COMPLEX2:73
for a, b being Complex holds angle (a,b) = angle (a,0,b)
proof end;

theorem Th72: :: COMPLEX2:74
for x, y, z being Complex st angle (x,y,z) = 0 holds
( Arg (x - y) = Arg (z - y) & angle (z,y,x) = 0 )
proof end;

theorem Th73: :: COMPLEX2:75
for a, b being Complex st a <> 0 & b <> 0 holds
( Re (a .|. b) = 0 iff ( angle (a,0,b) = PI / 2 or angle (a,0,b) = (3 / 2) * PI ) )
proof end;

theorem :: COMPLEX2:76
for a, b being Complex st a <> 0 & b <> 0 holds
( ( ( not Im (a .|. b) = |.a.| * |.b.| & not Im (a .|. b) = - (|.a.| * |.b.|) ) or angle (a,0,b) = PI / 2 or angle (a,0,b) = (3 / 2) * PI ) & ( ( not angle (a,0,b) = PI / 2 & not angle (a,0,b) = (3 / 2) * PI ) or Im (a .|. b) = |.a.| * |.b.| or Im (a .|. b) = - (|.a.| * |.b.|) ) )
proof end;

Lm3: for a, b, c being Complex st a <> b & c <> b holds
( Re ((a - b) .|. (c - b)) = 0 iff ( angle (a,b,c) = PI / 2 or angle (a,b,c) = (3 / 2) * PI ) )

proof end;

theorem :: COMPLEX2:77
for x, y, z being Complex st x <> y & z <> y & ( angle (x,y,z) = PI / 2 or angle (x,y,z) = (3 / 2) * PI ) holds
(|.(x - y).| ^2) + (|.(z - y).| ^2) = |.(x - z).| ^2
proof end;

theorem Th76: :: COMPLEX2:78
for a, b, c being Complex
for r being Real st a <> b & b <> c holds
angle (a,b,c) = angle ((Rotate (a,r)),(Rotate (b,r)),(Rotate (c,r)))
proof end;

theorem Th77: :: COMPLEX2:79
for a, b being Complex holds angle (a,b,a) = 0
proof end;

Lm4: for x, y, z being Complex st angle (x,y,z) <> 0 holds
angle (z,y,x) = (2 * PI) - (angle (x,y,z))

proof end;

theorem Th78: :: COMPLEX2:80
for a, b, c being Complex holds
( angle (a,b,c) <> 0 iff (angle (a,b,c)) + (angle (c,b,a)) = 2 * PI )
proof end;

theorem :: COMPLEX2:81
for a, b, c being Complex st angle (a,b,c) <> 0 holds
angle (c,b,a) <> 0
proof end;

theorem :: COMPLEX2:82
for a, b, c being Complex st angle (a,b,c) = PI holds
angle (c,b,a) = PI
proof end;

theorem Th81: :: COMPLEX2:83
for a, b, c being Complex st a <> b & a <> c & b <> c & not angle (a,b,c) <> 0 & not angle (b,c,a) <> 0 holds
angle (c,a,b) <> 0
proof end;

Lm5: for a, b being Complex st Im a = 0 & Re a > 0 & 0 < Arg b & Arg b < PI holds
( ((angle (a,0c,b)) + (angle (0c,b,a))) + (angle (b,a,0c)) = PI & 0 < angle (0c,b,a) & 0 < angle (b,a,0c) )

proof end;

theorem Th82: :: COMPLEX2:84
for a, b, c being Complex st a <> b & b <> c & 0 < angle (a,b,c) & angle (a,b,c) < PI holds
( ((angle (a,b,c)) + (angle (b,c,a))) + (angle (c,a,b)) = PI & 0 < angle (b,c,a) & 0 < angle (c,a,b) )
proof end;

theorem Th83: :: COMPLEX2:85
for a, b, c being Complex st a <> b & b <> c & angle (a,b,c) > PI holds
( ((angle (a,b,c)) + (angle (b,c,a))) + (angle (c,a,b)) = 5 * PI & angle (b,c,a) > PI & angle (c,a,b) > PI )
proof end;

Lm6: for a, b being Complex st Im a = 0 & Re a > 0 & Arg b = PI holds
( ((angle (a,0,b)) + (angle (0,b,a))) + (angle (b,a,0)) = PI & 0 = angle (0,b,a) & 0 = angle (b,a,0) )

proof end;

theorem Th84: :: COMPLEX2:86
for a, b, c being Complex st a <> b & b <> c & angle (a,b,c) = PI holds
( angle (b,c,a) = 0 & angle (c,a,b) = 0 )
proof end;

theorem Th85: :: COMPLEX2:87
for a, b, c being Complex st a <> b & a <> c & b <> c & angle (a,b,c) = 0 & not ( angle (b,c,a) = 0 & angle (c,a,b) = PI ) holds
( angle (b,c,a) = PI & angle (c,a,b) = 0 )
proof end;

theorem :: COMPLEX2:88
for a, b, c being Complex holds
( ( ((angle (a,b,c)) + (angle (b,c,a))) + (angle (c,a,b)) = PI or ((angle (a,b,c)) + (angle (b,c,a))) + (angle (c,a,b)) = 5 * PI ) iff ( a <> b & a <> c & b <> c ) )
proof end;