:: Trigonometric Functions and Existence of Circle Ratio
:: by Yuguang Yang and Yasunari Shidama
::
:: Received October 22, 1998
:: Copyright (c) 1998-2021 Association of Mizar Users


:: Some definitions and properties of complex sequence
definition
let m, k be natural Number ;
func CHK (m,k) -> Element of COMPLEX equals :Def1: :: SIN_COS:def 1
1 if m <= k
otherwise 0 ;
correctness
coherence
( ( m <= k implies 1 is Element of COMPLEX ) & ( not m <= k implies 0 is Element of COMPLEX ) )
;
consistency
for b1 being Element of COMPLEX holds verum
;
proof end;
end;

:: deftheorem Def1 defines CHK SIN_COS:def 1 :
for m, k being natural Number holds
( ( m <= k implies CHK (m,k) = 1 ) & ( not m <= k implies CHK (m,k) = 0 ) );

registration
let m, k be natural Number ;
cluster CHK (m,k) -> real ;
coherence
CHK (m,k) is real
by Def1;
end;

scheme :: SIN_COS:sch 1
ExComplexCASE{ F1( Nat, Nat) -> Complex } :
for k being Nat ex seq being Complex_Sequence st
for n being Nat holds
( ( n <= k implies seq . n = F1(k,n) ) & ( n > k implies seq . n = 0 ) )
proof end;

scheme :: SIN_COS:sch 2
ExRealCASE{ F1( Nat, Nat) -> Real } :
for k being Nat ex rseq being Real_Sequence st
for n being Nat holds
( ( n <= k implies rseq . n = F1(k,n) ) & ( n > k implies rseq . n = 0 ) )
proof end;

1 in NAT
;

then reconsider jj = 1 as Element of REAL by NUMBERS:19;

definition
func Prod_real_n -> Real_Sequence means :Def2: :: SIN_COS:def 2
( it . 0 = 1 & ( for n being Nat holds it . (n + 1) = (it . n) * (n + 1) ) );
existence
ex b1 being Real_Sequence st
( b1 . 0 = 1 & ( for n being Nat holds b1 . (n + 1) = (b1 . n) * (n + 1) ) )
proof end;
uniqueness
for b1, b2 being Real_Sequence st b1 . 0 = 1 & ( for n being Nat holds b1 . (n + 1) = (b1 . n) * (n + 1) ) & b2 . 0 = 1 & ( for n being Nat holds b2 . (n + 1) = (b2 . n) * (n + 1) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def2 defines Prod_real_n SIN_COS:def 2 :
for b1 being Real_Sequence holds
( b1 = Prod_real_n iff ( b1 . 0 = 1 & ( for n being Nat holds b1 . (n + 1) = (b1 . n) * (n + 1) ) ) );

definition
let n be Nat;
redefine func n ! equals :: SIN_COS:def 3
Prod_real_n . n;
compatibility
for b1 being set holds
( b1 = n ! iff b1 = Prod_real_n . n )
proof end;
end;

:: deftheorem defines ! SIN_COS:def 3 :
for n being Nat holds n ! = Prod_real_n . n;

definition
let z be Complex;
deffunc H1( Nat) -> set = (z |^ $1) / ($1 !);
func z ExpSeq -> Complex_Sequence means :Def4: :: SIN_COS:def 4
for n being Nat holds it . n = (z |^ n) / (n !);
existence
ex b1 being Complex_Sequence st
for n being Nat holds b1 . n = (z |^ n) / (n !)
proof end;
uniqueness
for b1, b2 being Complex_Sequence st ( for n being Nat holds b1 . n = (z |^ n) / (n !) ) & ( for n being Nat holds b2 . n = (z |^ n) / (n !) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def4 defines ExpSeq SIN_COS:def 4 :
for z being Complex
for b2 being Complex_Sequence holds
( b2 = z ExpSeq iff for n being Nat holds b2 . n = (z |^ n) / (n !) );

definition
let a be Real;
deffunc H1( Nat) -> set = (a |^ $1) / ($1 !);
func a rExpSeq -> Real_Sequence means :Def5: :: SIN_COS:def 5
for n being Nat holds it . n = (a |^ n) / (n !);
existence
ex b1 being Real_Sequence st
for n being Nat holds b1 . n = (a |^ n) / (n !)
proof end;
uniqueness
for b1, b2 being Real_Sequence st ( for n being Nat holds b1 . n = (a |^ n) / (n !) ) & ( for n being Nat holds b2 . n = (a |^ n) / (n !) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def5 defines rExpSeq SIN_COS:def 5 :
for a being Real
for b2 being Real_Sequence holds
( b2 = a rExpSeq iff for n being Nat holds b2 . n = (a |^ n) / (n !) );

theorem Th1: :: SIN_COS:1
for n being Nat holds
( 0 ! = 1 & n ! <> 0 & (n + 1) ! = (n !) * (n + 1) ) by Def2;

theorem Th2: :: SIN_COS:2
for k, m being Nat holds
( ( 0 < k implies ((k -' 1) !) * k = k ! ) & ( k <= m implies ((m -' k) !) * ((m + 1) - k) = ((m + 1) -' k) ! ) )
proof end;

definition
let n be Nat;
func Coef n -> Complex_Sequence means :Def6: :: SIN_COS:def 6
for k being Nat holds
( ( k <= n implies it . k = (n !) / ((k !) * ((n -' k) !)) ) & ( k > n implies it . k = 0 ) );
existence
ex b1 being Complex_Sequence st
for k being Nat holds
( ( k <= n implies b1 . k = (n !) / ((k !) * ((n -' k) !)) ) & ( k > n implies b1 . k = 0 ) )
proof end;
uniqueness
for b1, b2 being Complex_Sequence st ( for k being Nat holds
( ( k <= n implies b1 . k = (n !) / ((k !) * ((n -' k) !)) ) & ( k > n implies b1 . k = 0 ) ) ) & ( for k being Nat holds
( ( k <= n implies b2 . k = (n !) / ((k !) * ((n -' k) !)) ) & ( k > n implies b2 . k = 0 ) ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def6 defines Coef SIN_COS:def 6 :
for n being Nat
for b2 being Complex_Sequence holds
( b2 = Coef n iff for k being Nat holds
( ( k <= n implies b2 . k = (n !) / ((k !) * ((n -' k) !)) ) & ( k > n implies b2 . k = 0 ) ) );

definition
let n be Nat;
func Coef_e n -> Complex_Sequence means :Def7: :: SIN_COS:def 7
for k being Nat holds
( ( k <= n implies it . k = 1r / ((k !) * ((n -' k) !)) ) & ( k > n implies it . k = 0 ) );
existence
ex b1 being Complex_Sequence st
for k being Nat holds
( ( k <= n implies b1 . k = 1r / ((k !) * ((n -' k) !)) ) & ( k > n implies b1 . k = 0 ) )
proof end;
uniqueness
for b1, b2 being Complex_Sequence st ( for k being Nat holds
( ( k <= n implies b1 . k = 1r / ((k !) * ((n -' k) !)) ) & ( k > n implies b1 . k = 0 ) ) ) & ( for k being Nat holds
( ( k <= n implies b2 . k = 1r / ((k !) * ((n -' k) !)) ) & ( k > n implies b2 . k = 0 ) ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def7 defines Coef_e SIN_COS:def 7 :
for n being Nat
for b2 being Complex_Sequence holds
( b2 = Coef_e n iff for k being Nat holds
( ( k <= n implies b2 . k = 1r / ((k !) * ((n -' k) !)) ) & ( k > n implies b2 . k = 0 ) ) );

definition
let seq be Complex_Sequence;
func Shift seq -> Complex_Sequence means :Def8: :: SIN_COS:def 8
( it . 0 = 0 & ( for k being Nat holds it . (k + 1) = seq . k ) );
existence
ex b1 being Complex_Sequence st
( b1 . 0 = 0 & ( for k being Nat holds b1 . (k + 1) = seq . k ) )
proof end;
uniqueness
for b1, b2 being Complex_Sequence st b1 . 0 = 0 & ( for k being Nat holds b1 . (k + 1) = seq . k ) & b2 . 0 = 0 & ( for k being Nat holds b2 . (k + 1) = seq . k ) holds
b1 = b2
proof end;
end;

:: deftheorem Def8 defines Shift SIN_COS:def 8 :
for seq, b2 being Complex_Sequence holds
( b2 = Shift seq iff ( b2 . 0 = 0 & ( for k being Nat holds b2 . (k + 1) = seq . k ) ) );

definition
let n be Nat;
let z, w be Complex;
func Expan (n,z,w) -> Complex_Sequence means :Def9: :: SIN_COS:def 9
for k being Nat holds
( ( k <= n implies it . k = (((Coef n) . k) * (z |^ k)) * (w |^ (n -' k)) ) & ( n < k implies it . k = 0 ) );
existence
ex b1 being Complex_Sequence st
for k being Nat holds
( ( k <= n implies b1 . k = (((Coef n) . k) * (z |^ k)) * (w |^ (n -' k)) ) & ( n < k implies b1 . k = 0 ) )
proof end;
uniqueness
for b1, b2 being Complex_Sequence st ( for k being Nat holds
( ( k <= n implies b1 . k = (((Coef n) . k) * (z |^ k)) * (w |^ (n -' k)) ) & ( n < k implies b1 . k = 0 ) ) ) & ( for k being Nat holds
( ( k <= n implies b2 . k = (((Coef n) . k) * (z |^ k)) * (w |^ (n -' k)) ) & ( n < k implies b2 . k = 0 ) ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def9 defines Expan SIN_COS:def 9 :
for n being Nat
for z, w being Complex
for b4 being Complex_Sequence holds
( b4 = Expan (n,z,w) iff for k being Nat holds
( ( k <= n implies b4 . k = (((Coef n) . k) * (z |^ k)) * (w |^ (n -' k)) ) & ( n < k implies b4 . k = 0 ) ) );

definition
let n be Nat;
let z, w be Complex;
func Expan_e (n,z,w) -> Complex_Sequence means :Def10: :: SIN_COS:def 10
for k being Nat holds
( ( k <= n implies it . k = (((Coef_e n) . k) * (z |^ k)) * (w |^ (n -' k)) ) & ( n < k implies it . k = 0 ) );
existence
ex b1 being Complex_Sequence st
for k being Nat holds
( ( k <= n implies b1 . k = (((Coef_e n) . k) * (z |^ k)) * (w |^ (n -' k)) ) & ( n < k implies b1 . k = 0 ) )
proof end;
uniqueness
for b1, b2 being Complex_Sequence st ( for k being Nat holds
( ( k <= n implies b1 . k = (((Coef_e n) . k) * (z |^ k)) * (w |^ (n -' k)) ) & ( n < k implies b1 . k = 0 ) ) ) & ( for k being Nat holds
( ( k <= n implies b2 . k = (((Coef_e n) . k) * (z |^ k)) * (w |^ (n -' k)) ) & ( n < k implies b2 . k = 0 ) ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def10 defines Expan_e SIN_COS:def 10 :
for n being Nat
for z, w being Complex
for b4 being Complex_Sequence holds
( b4 = Expan_e (n,z,w) iff for k being Nat holds
( ( k <= n implies b4 . k = (((Coef_e n) . k) * (z |^ k)) * (w |^ (n -' k)) ) & ( n < k implies b4 . k = 0 ) ) );

definition
let n be Nat;
let z, w be Complex;
func Alfa (n,z,w) -> Complex_Sequence means :Def11: :: SIN_COS:def 11
for k being Nat holds
( ( k <= n implies it . k = ((z ExpSeq) . k) * ((Partial_Sums (w ExpSeq)) . (n -' k)) ) & ( n < k implies it . k = 0 ) );
existence
ex b1 being Complex_Sequence st
for k being Nat holds
( ( k <= n implies b1 . k = ((z ExpSeq) . k) * ((Partial_Sums (w ExpSeq)) . (n -' k)) ) & ( n < k implies b1 . k = 0 ) )
proof end;
uniqueness
for b1, b2 being Complex_Sequence st ( for k being Nat holds
( ( k <= n implies b1 . k = ((z ExpSeq) . k) * ((Partial_Sums (w ExpSeq)) . (n -' k)) ) & ( n < k implies b1 . k = 0 ) ) ) & ( for k being Nat holds
( ( k <= n implies b2 . k = ((z ExpSeq) . k) * ((Partial_Sums (w ExpSeq)) . (n -' k)) ) & ( n < k implies b2 . k = 0 ) ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def11 defines Alfa SIN_COS:def 11 :
for n being Nat
for z, w being Complex
for b4 being Complex_Sequence holds
( b4 = Alfa (n,z,w) iff for k being Nat holds
( ( k <= n implies b4 . k = ((z ExpSeq) . k) * ((Partial_Sums (w ExpSeq)) . (n -' k)) ) & ( n < k implies b4 . k = 0 ) ) );

definition
let a, b be Real;
let n be Nat;
func Conj (n,a,b) -> Real_Sequence means :: SIN_COS:def 12
for k being Nat holds
( ( k <= n implies it . k = ((a rExpSeq) . k) * (((Partial_Sums (b rExpSeq)) . n) - ((Partial_Sums (b rExpSeq)) . (n -' k))) ) & ( n < k implies it . k = 0 ) );
existence
ex b1 being Real_Sequence st
for k being Nat holds
( ( k <= n implies b1 . k = ((a rExpSeq) . k) * (((Partial_Sums (b rExpSeq)) . n) - ((Partial_Sums (b rExpSeq)) . (n -' k))) ) & ( n < k implies b1 . k = 0 ) )
proof end;
uniqueness
for b1, b2 being Real_Sequence st ( for k being Nat holds
( ( k <= n implies b1 . k = ((a rExpSeq) . k) * (((Partial_Sums (b rExpSeq)) . n) - ((Partial_Sums (b rExpSeq)) . (n -' k))) ) & ( n < k implies b1 . k = 0 ) ) ) & ( for k being Nat holds
( ( k <= n implies b2 . k = ((a rExpSeq) . k) * (((Partial_Sums (b rExpSeq)) . n) - ((Partial_Sums (b rExpSeq)) . (n -' k))) ) & ( n < k implies b2 . k = 0 ) ) ) holds
b1 = b2
proof end;
end;

:: deftheorem defines Conj SIN_COS:def 12 :
for a, b being Real
for n being Nat
for b4 being Real_Sequence holds
( b4 = Conj (n,a,b) iff for k being Nat holds
( ( k <= n implies b4 . k = ((a rExpSeq) . k) * (((Partial_Sums (b rExpSeq)) . n) - ((Partial_Sums (b rExpSeq)) . (n -' k))) ) & ( n < k implies b4 . k = 0 ) ) );

definition
let z, w be Complex;
let n be Nat;
func Conj (n,z,w) -> Complex_Sequence means :Def13: :: SIN_COS:def 13
for k being Nat holds
( ( k <= n implies it . k = ((z ExpSeq) . k) * (((Partial_Sums (w ExpSeq)) . n) - ((Partial_Sums (w ExpSeq)) . (n -' k))) ) & ( n < k implies it . k = 0 ) );
existence
ex b1 being Complex_Sequence st
for k being Nat holds
( ( k <= n implies b1 . k = ((z ExpSeq) . k) * (((Partial_Sums (w ExpSeq)) . n) - ((Partial_Sums (w ExpSeq)) . (n -' k))) ) & ( n < k implies b1 . k = 0 ) )
proof end;
uniqueness
for b1, b2 being Complex_Sequence st ( for k being Nat holds
( ( k <= n implies b1 . k = ((z ExpSeq) . k) * (((Partial_Sums (w ExpSeq)) . n) - ((Partial_Sums (w ExpSeq)) . (n -' k))) ) & ( n < k implies b1 . k = 0 ) ) ) & ( for k being Nat holds
( ( k <= n implies b2 . k = ((z ExpSeq) . k) * (((Partial_Sums (w ExpSeq)) . n) - ((Partial_Sums (w ExpSeq)) . (n -' k))) ) & ( n < k implies b2 . k = 0 ) ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def13 defines Conj SIN_COS:def 13 :
for z, w being Complex
for n being Nat
for b4 being Complex_Sequence holds
( b4 = Conj (n,z,w) iff for k being Nat holds
( ( k <= n implies b4 . k = ((z ExpSeq) . k) * (((Partial_Sums (w ExpSeq)) . n) - ((Partial_Sums (w ExpSeq)) . (n -' k))) ) & ( n < k implies b4 . k = 0 ) ) );

Lm1: for p1, p2, g1, g2 being Real holds
( (p1 + (g1 * <i>)) * (p2 + (g2 * <i>)) = ((p1 * p2) - (g1 * g2)) + (((p1 * g2) + (p2 * g1)) * <i>) & (p2 + (g2 * <i>)) *' = p2 + ((- g2) * <i>) )

proof end;

theorem Th3: :: SIN_COS:3
for n being Nat
for z being Complex holds
( (z ExpSeq) . (n + 1) = (((z ExpSeq) . n) * z) / ((n + 1) + (0 * <i>)) & (z ExpSeq) . 0 = 1 & |.((z ExpSeq) . n).| = (|.z.| rExpSeq) . n )
proof end;

theorem Th4: :: SIN_COS:4
for k being Nat
for seq being Complex_Sequence st 0 < k holds
(Shift seq) . k = seq . (k -' 1)
proof end;

theorem Th5: :: SIN_COS:5
for k being Nat
for seq being Complex_Sequence holds (Partial_Sums seq) . k = ((Partial_Sums (Shift seq)) . k) + (seq . k)
proof end;

theorem Th6: :: SIN_COS:6
for w, z being Complex
for n being Nat holds (z + w) |^ n = (Partial_Sums (Expan (n,z,w))) . n
proof end;

theorem Th7: :: SIN_COS:7
for w, z being Complex
for n being Nat holds Expan_e (n,z,w) = (1r / (n !)) (#) (Expan (n,z,w))
proof end;

theorem Th8: :: SIN_COS:8
for w, z being Complex
for n being Nat holds ((z + w) |^ n) / (n !) = (Partial_Sums (Expan_e (n,z,w))) . n
proof end;

theorem Th9: :: SIN_COS:9
( 0c ExpSeq is absolutely_summable & Sum (0c ExpSeq) = 1r )
proof end;

registration
let z be Complex;
cluster z ExpSeq -> absolutely_summable ;
coherence
z ExpSeq is absolutely_summable
proof end;
end;

theorem Th10: :: SIN_COS:10
for w, z being Complex holds
( (z ExpSeq) . 0 = 1 & (Expan (0,z,w)) . 0 = 1 )
proof end;

theorem Th11: :: SIN_COS:11
for w, z being Complex
for k, l being Nat st l <= k holds
(Alfa ((k + 1),z,w)) . l = ((Alfa (k,z,w)) . l) + ((Expan_e ((k + 1),z,w)) . l)
proof end;

theorem Th12: :: SIN_COS:12
for w, z being Complex
for k being Nat holds (Partial_Sums (Alfa ((k + 1),z,w))) . k = ((Partial_Sums (Alfa (k,z,w))) . k) + ((Partial_Sums (Expan_e ((k + 1),z,w))) . k)
proof end;

theorem Th13: :: SIN_COS:13
for w, z being Complex
for k being Nat holds (z ExpSeq) . k = (Expan_e (k,z,w)) . k
proof end;

theorem Th14: :: SIN_COS:14
for w, z being Complex
for n being Nat holds (Partial_Sums ((z + w) ExpSeq)) . n = (Partial_Sums (Alfa (n,z,w))) . n
proof end;

theorem Th15: :: SIN_COS:15
for w, z being Complex
for k being Nat holds (((Partial_Sums (z ExpSeq)) . k) * ((Partial_Sums (w ExpSeq)) . k)) - ((Partial_Sums ((z + w) ExpSeq)) . k) = (Partial_Sums (Conj (k,z,w))) . k
proof end;

theorem Th16: :: SIN_COS:16
for z being Complex
for k being Nat holds
( |.((Partial_Sums (z ExpSeq)) . k).| <= (Partial_Sums (|.z.| rExpSeq)) . k & (Partial_Sums (|.z.| rExpSeq)) . k <= Sum (|.z.| rExpSeq) & |.((Partial_Sums (z ExpSeq)) . k).| <= Sum (|.z.| rExpSeq) )
proof end;

theorem Th17: :: SIN_COS:17
for z being Complex holds 1 <= Sum (|.z.| rExpSeq)
proof end;

theorem Th18: :: SIN_COS:18
for z being Complex
for n being Nat holds 0 <= (|.z.| rExpSeq) . n
proof end;

theorem Th19: :: SIN_COS:19
for z being Complex
for m, n being Nat holds
( |.((Partial_Sums (|.z.| rExpSeq)) . n).| = (Partial_Sums (|.z.| rExpSeq)) . n & ( n <= m implies |.(((Partial_Sums (|.z.| rExpSeq)) . m) - ((Partial_Sums (|.z.| rExpSeq)) . n)).| = ((Partial_Sums (|.z.| rExpSeq)) . m) - ((Partial_Sums (|.z.| rExpSeq)) . n) ) )
proof end;

theorem Th20: :: SIN_COS:20
for w, z being Complex
for k, n being Nat holds |.((Partial_Sums |.(Conj (k,z,w)).|) . n).| = (Partial_Sums |.(Conj (k,z,w)).|) . n
proof end;

theorem Th21: :: SIN_COS:21
for w, z being Complex
for p being Real st p > 0 holds
ex n being Nat st
for k being Nat st n <= k holds
|.((Partial_Sums |.(Conj (k,z,w)).|) . k).| < p
proof end;

theorem Th22: :: SIN_COS:22
for w, z being Complex
for seq being Complex_Sequence st ( for k being Nat holds seq . k = (Partial_Sums (Conj (k,z,w))) . k ) holds
( seq is convergent & lim seq = 0 )
proof end;

Lm2: for z, w being Complex holds (Sum (z ExpSeq)) * (Sum (w ExpSeq)) = Sum ((z + w) ExpSeq)
proof end;

definition
func exp -> Function of COMPLEX,COMPLEX means :Def14: :: SIN_COS:def 14
for z being Complex holds it . z = Sum (z ExpSeq);
existence
ex b1 being Function of COMPLEX,COMPLEX st
for z being Complex holds b1 . z = Sum (z ExpSeq)
proof end;
uniqueness
for b1, b2 being Function of COMPLEX,COMPLEX st ( for z being Complex holds b1 . z = Sum (z ExpSeq) ) & ( for z being Complex holds b2 . z = Sum (z ExpSeq) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def14 defines exp SIN_COS:def 14 :
for b1 being Function of COMPLEX,COMPLEX holds
( b1 = exp iff for z being Complex holds b1 . z = Sum (z ExpSeq) );

definition
let z be Complex;
func exp z -> Complex equals :: SIN_COS:def 15
exp . z;
coherence
exp . z is Complex
;
end;

:: deftheorem defines exp SIN_COS:def 15 :
for z being Complex holds exp z = exp . z;

definition
let z be Complex;
:: original: exp
redefine func exp z -> Element of COMPLEX ;
coherence
exp z is Element of COMPLEX
by XCMPLX_0:def 2;
end;

theorem :: SIN_COS:23
for z1, z2 being Complex holds exp (z1 + z2) = (exp z1) * (exp z2)
proof end;

definition
func sin -> Function of REAL,REAL means :Def16: :: SIN_COS:def 16
for d being Real holds it . d = Im (Sum ((d * <i>) ExpSeq));
existence
ex b1 being Function of REAL,REAL st
for d being Real holds b1 . d = Im (Sum ((d * <i>) ExpSeq))
proof end;
uniqueness
for b1, b2 being Function of REAL,REAL st ( for d being Real holds b1 . d = Im (Sum ((d * <i>) ExpSeq)) ) & ( for d being Real holds b2 . d = Im (Sum ((d * <i>) ExpSeq)) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def16 defines sin SIN_COS:def 16 :
for b1 being Function of REAL,REAL holds
( b1 = sin iff for d being Real holds b1 . d = Im (Sum ((d * <i>) ExpSeq)) );

definition
let th be Real;
func sin th -> number equals :: SIN_COS:def 17
sin . th;
coherence
sin . th is number
;
end;

:: deftheorem defines sin SIN_COS:def 17 :
for th being Real holds sin th = sin . th;

registration
let th be Real;
cluster sin th -> real ;
coherence
sin th is real
;
end;

definition
func cos -> Function of REAL,REAL means :Def18: :: SIN_COS:def 18
for d being Real holds it . d = Re (Sum ((d * <i>) ExpSeq));
existence
ex b1 being Function of REAL,REAL st
for d being Real holds b1 . d = Re (Sum ((d * <i>) ExpSeq))
proof end;
uniqueness
for b1, b2 being Function of REAL,REAL st ( for d being Real holds b1 . d = Re (Sum ((d * <i>) ExpSeq)) ) & ( for d being Real holds b2 . d = Re (Sum ((d * <i>) ExpSeq)) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def18 defines cos SIN_COS:def 18 :
for b1 being Function of REAL,REAL holds
( b1 = cos iff for d being Real holds b1 . d = Re (Sum ((d * <i>) ExpSeq)) );

definition
let th be Real;
func cos th -> number equals :: SIN_COS:def 19
cos . th;
coherence
cos . th is number
;
end;

:: deftheorem defines cos SIN_COS:def 19 :
for th being Real holds cos th = cos . th;

registration
let th be Real;
cluster cos th -> real ;
coherence
cos th is real
;
end;

theorem Th24: :: SIN_COS:24
( dom sin = REAL & dom cos = REAL ) by FUNCT_2:def 1;

Lm3: for th being Real holds Sum ((th * <i>) ExpSeq) = (cos . th) + ((sin . th) * <i>)
proof end;

theorem :: SIN_COS:25
for th being Real holds exp (th * <i>) = (cos th) + ((sin th) * <i>)
proof end;

Lm4: for th being Real holds (Sum ((th * <i>) ExpSeq)) *' = Sum ((- (th * <i>)) ExpSeq)
proof end;

theorem :: SIN_COS:26
for th being Real holds (exp (th * <i>)) *' = exp (- (th * <i>))
proof end;

Lm5: for th, th1 being Real st th = th1 holds
( |.(Sum ((th * <i>) ExpSeq)).| = 1 & |.(sin . th1).| <= 1 & |.(cos . th1).| <= 1 )

proof end;

theorem :: SIN_COS:27
for th being Real holds
( |.(exp (th * <i>)).| = 1 & ( for th being Real holds
( |.(sin th).| <= 1 & |.(cos th).| <= 1 ) ) )
proof end;

theorem Th28: :: SIN_COS:28
for th being Real holds
( ((cos . th) ^2) + ((sin . th) ^2) = 1 & ((cos . th) * (cos . th)) + ((sin . th) * (sin . th)) = 1 )
proof end;

theorem :: SIN_COS:29
for th being Real holds
( ((cos th) ^2) + ((sin th) ^2) = 1 & ((cos th) * (cos th)) + ((sin th) * (sin th)) = 1 ) by Th28;

theorem Th30: :: SIN_COS:30
for th being Real holds
( cos . 0 = 1 & sin . 0 = 0 & cos . (- th) = cos . th & sin . (- th) = - (sin . th) )
proof end;

theorem :: SIN_COS:31
for th being Real holds
( cos 0 = 1 & sin 0 = 0 & cos (- th) = cos th & sin (- th) = - (sin th) ) by Th30;

definition
let th be Real;
deffunc H1( Nat) -> set = (((- 1) |^ $1) * (th |^ ((2 * $1) + 1))) / (((2 * $1) + 1) !);
func th P_sin -> Real_Sequence means :Def20: :: SIN_COS:def 20
for n being Nat holds it . n = (((- 1) |^ n) * (th |^ ((2 * n) + 1))) / (((2 * n) + 1) !);
existence
ex b1 being Real_Sequence st
for n being Nat holds b1 . n = (((- 1) |^ n) * (th |^ ((2 * n) + 1))) / (((2 * n) + 1) !)
proof end;
uniqueness
for b1, b2 being Real_Sequence st ( for n being Nat holds b1 . n = (((- 1) |^ n) * (th |^ ((2 * n) + 1))) / (((2 * n) + 1) !) ) & ( for n being Nat holds b2 . n = (((- 1) |^ n) * (th |^ ((2 * n) + 1))) / (((2 * n) + 1) !) ) holds
b1 = b2
proof end;
deffunc H2( Nat) -> set = (((- 1) |^ $1) * (th |^ (2 * $1))) / ((2 * $1) !);
func th P_cos -> Real_Sequence means :Def21: :: SIN_COS:def 21
for n being Nat holds it . n = (((- 1) |^ n) * (th |^ (2 * n))) / ((2 * n) !);
existence
ex b1 being Real_Sequence st
for n being Nat holds b1 . n = (((- 1) |^ n) * (th |^ (2 * n))) / ((2 * n) !)
proof end;
uniqueness
for b1, b2 being Real_Sequence st ( for n being Nat holds b1 . n = (((- 1) |^ n) * (th |^ (2 * n))) / ((2 * n) !) ) & ( for n being Nat holds b2 . n = (((- 1) |^ n) * (th |^ (2 * n))) / ((2 * n) !) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def20 defines P_sin SIN_COS:def 20 :
for th being Real
for b2 being Real_Sequence holds
( b2 = th P_sin iff for n being Nat holds b2 . n = (((- 1) |^ n) * (th |^ ((2 * n) + 1))) / (((2 * n) + 1) !) );

:: deftheorem Def21 defines P_cos SIN_COS:def 21 :
for th being Real
for b2 being Real_Sequence holds
( b2 = th P_cos iff for n being Nat holds b2 . n = (((- 1) |^ n) * (th |^ (2 * n))) / ((2 * n) !) );

theorem Th32: :: SIN_COS:32
for z being Complex
for k being Nat holds
( z |^ (2 * k) = (z |^ k) |^ 2 & z |^ (2 * k) = (z |^ 2) |^ k )
proof end;

theorem Th33: :: SIN_COS:33
for k being Nat
for th being Real holds
( (th * <i>) |^ (2 * k) = ((- 1) |^ k) * (th |^ (2 * k)) & (th * <i>) |^ ((2 * k) + 1) = (((- 1) |^ k) * (th |^ ((2 * k) + 1))) * <i> )
proof end;

theorem :: SIN_COS:34
canceled;

::$CT
theorem Th34: :: SIN_COS:35
for n being Nat
for th being Real holds
( (Partial_Sums (th P_sin)) . n = (Partial_Sums (Im ((th * <i>) ExpSeq))) . ((2 * n) + 1) & (Partial_Sums (th P_cos)) . n = (Partial_Sums (Re ((th * <i>) ExpSeq))) . (2 * n) )
proof end;

theorem Th35: :: SIN_COS:36
for th being Real holds
( Partial_Sums (th P_sin) is convergent & Sum (th P_sin) = Im (Sum ((th * <i>) ExpSeq)) & Partial_Sums (th P_cos) is convergent & Sum (th P_cos) = Re (Sum ((th * <i>) ExpSeq)) )
proof end;

theorem Th36: :: SIN_COS:37
for th being Real holds
( cos . th = Sum (th P_cos) & sin . th = Sum (th P_sin) )
proof end;

theorem :: SIN_COS:38
for p, th being Real
for rseq being Real_Sequence st rseq is convergent & lim rseq = th & ( for n being Nat holds rseq . n >= p ) holds
th >= p by PREPOWER:1;

deffunc H1( Real) -> set = (2 * $1) + 1;

consider bq being Real_Sequence such that
Lm6: for n being Nat holds bq . n = H1(n) from SEQ_1:sch 1();

bq is increasing sequence of NAT
proof end;

then reconsider bq = bq as increasing sequence of NAT ;

Lm7: for n being Nat
for th, th1, th2, th3 being Real holds
( th |^ (2 * n) = (th |^ n) |^ 2 & (- 1) |^ (2 * n) = 1 & (- 1) |^ ((2 * n) + 1) = - 1 )

proof end;

theorem Th38: :: SIN_COS:39
for k, n being Nat st n <= k holds
n ! <= k !
proof end;

theorem Th39: :: SIN_COS:40
for k, n being Nat
for th being Real st 0 <= th & th <= 1 & n <= k holds
th |^ k <= th |^ n
proof end;

theorem :: SIN_COS:41
canceled;

::$CT
theorem Th41: :: SIN_COS:42
for p being Real holds Im (Sum (p ExpSeq)) = 0
proof end;

theorem Th42: :: SIN_COS:43
( cos . 1 > 0 & sin . 1 > 0 & cos . 1 < sin . 1 )
proof end;

theorem Th43: :: SIN_COS:44
for th being Real holds th rExpSeq = Re (th ExpSeq)
proof end;

theorem Th44: :: SIN_COS:45
for th being Real holds
( th rExpSeq is summable & Sum (th rExpSeq) = Re (Sum (th ExpSeq)) )
proof end;

Lm8: for n being Nat
for z being Complex holds
( (z ExpSeq) . 1 = z & (z ExpSeq) . 0 = 1r & |.(z |^ n).| = |.z.| |^ n )

proof end;

Lm9: for th being Real holds Sum (th ExpSeq) = Sum (th rExpSeq)
proof end;

theorem Th45: :: SIN_COS:46
for p, q being Real holds Sum ((p + q) rExpSeq) = (Sum (p rExpSeq)) * (Sum (q rExpSeq))
proof end;

definition
func exp_R -> Function of REAL,REAL means :Def22: :: SIN_COS:def 22
for d being Real holds it . d = Sum (d rExpSeq);
existence
ex b1 being Function of REAL,REAL st
for d being Real holds b1 . d = Sum (d rExpSeq)
proof end;
uniqueness
for b1, b2 being Function of REAL,REAL st ( for d being Real holds b1 . d = Sum (d rExpSeq) ) & ( for d being Real holds b2 . d = Sum (d rExpSeq) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def22 defines exp_R SIN_COS:def 22 :
for b1 being Function of REAL,REAL holds
( b1 = exp_R iff for d being Real holds b1 . d = Sum (d rExpSeq) );

definition
let th be Real;
func exp_R th -> number equals :: SIN_COS:def 23
exp_R . th;
coherence
exp_R . th is number
;
end;

:: deftheorem defines exp_R SIN_COS:def 23 :
for th being Real holds exp_R th = exp_R . th;

registration
let th be Real;
cluster exp_R th -> real ;
coherence
exp_R th is real
;
end;

theorem Th46: :: SIN_COS:47
dom exp_R = REAL by FUNCT_2:def 1;

theorem Th47: :: SIN_COS:48
for th being Real holds exp_R . th = Re (Sum (th ExpSeq))
proof end;

theorem :: SIN_COS:49
for th being Real holds exp th = exp_R th
proof end;

Lm10: for p, q being Real holds exp_R . (p + q) = (exp_R . p) * (exp_R . q)
proof end;

theorem :: SIN_COS:50
for p, q being Real holds exp_R (p + q) = (exp_R p) * (exp_R q) by Lm10;

Lm11: exp_R . 0 = 1
proof end;

theorem :: SIN_COS:51
exp_R 0 = 1 by Lm11;

theorem Th51: :: SIN_COS:52
for th being Real st th > 0 holds
exp_R . th >= 1
proof end;

theorem Th52: :: SIN_COS:53
for th being Real st th < 0 holds
( 0 < exp_R . th & exp_R . th <= 1 )
proof end;

theorem Th53: :: SIN_COS:54
for th being Real holds exp_R . th > 0
proof end;

theorem :: SIN_COS:55
for th being Real holds exp_R th > 0 by Th53;

definition
let z be Complex;
deffunc H2( Nat) -> set = (z |^ ($1 + 1)) / (($1 + 2) !);
func z P_dt -> Complex_Sequence means :Def24: :: SIN_COS:def 24
for n being Nat holds it . n = (z |^ (n + 1)) / ((n + 2) !);
existence
ex b1 being Complex_Sequence st
for n being Nat holds b1 . n = (z |^ (n + 1)) / ((n + 2) !)
proof end;
uniqueness
for b1, b2 being Complex_Sequence st ( for n being Nat holds b1 . n = (z |^ (n + 1)) / ((n + 2) !) ) & ( for n being Nat holds b2 . n = (z |^ (n + 1)) / ((n + 2) !) ) holds
b1 = b2
proof end;
deffunc H3( Nat) -> set = (z |^ $1) / (($1 + 2) !);
func z P_t -> Complex_Sequence means :: SIN_COS:def 25
for n being Nat holds it . n = (z |^ n) / ((n + 2) !);
existence
ex b1 being Complex_Sequence st
for n being Nat holds b1 . n = (z |^ n) / ((n + 2) !)
proof end;
uniqueness
for b1, b2 being Complex_Sequence st ( for n being Nat holds b1 . n = (z |^ n) / ((n + 2) !) ) & ( for n being Nat holds b2 . n = (z |^ n) / ((n + 2) !) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def24 defines P_dt SIN_COS:def 24 :
for z being Complex
for b2 being Complex_Sequence holds
( b2 = z P_dt iff for n being Nat holds b2 . n = (z |^ (n + 1)) / ((n + 2) !) );

:: deftheorem defines P_t SIN_COS:def 25 :
for z being Complex
for b2 being Complex_Sequence holds
( b2 = z P_t iff for n being Nat holds b2 . n = (z |^ n) / ((n + 2) !) );

Lm12: for p being Real
for z being Complex holds
( Re ((p * <i>) * z) = - (p * (Im z)) & Im ((p * <i>) * z) = p * (Re z) & Re (p * z) = p * (Re z) & Im (p * z) = p * (Im z) )

proof end;

Lm13: for p being Real
for z being Complex st p > 0 holds
( Re (z / (p * <i>)) = (Im z) / p & Im (z / (p * <i>)) = - ((Re z) / p) & |.(z / p).| = |.z.| / p )

proof end;

theorem Th55: :: SIN_COS:56
for z being Complex holds z P_dt is absolutely_summable
proof end;

theorem Th56: :: SIN_COS:57
for z being Complex holds z * (Sum (z P_dt)) = ((Sum (z ExpSeq)) - 1r) - z
proof end;

theorem Th57: :: SIN_COS:58
for p being Real st p > 0 holds
ex q being Real st
( q > 0 & ( for z being Complex st |.z.| < q holds
|.(Sum (z P_dt)).| < p ) )
proof end;

theorem Th58: :: SIN_COS:59
for z, z1 being Complex holds (Sum ((z1 + z) ExpSeq)) - (Sum (z1 ExpSeq)) = ((Sum (z1 ExpSeq)) * z) + ((z * (Sum (z P_dt))) * (Sum (z1 ExpSeq)))
proof end;

theorem Th59: :: SIN_COS:60
for p, q being Real holds (cos . (p + q)) - (cos . p) = (- (q * (sin . p))) - (q * (Im ((Sum ((q * <i>) P_dt)) * ((cos . p) + ((sin . p) * <i>)))))
proof end;

theorem Th60: :: SIN_COS:61
for p, q being Real holds (sin . (p + q)) - (sin . p) = (q * (cos . p)) + (q * (Re ((Sum ((q * <i>) P_dt)) * ((cos . p) + ((sin . p) * <i>)))))
proof end;

theorem Th61: :: SIN_COS:62
for p, q being Real holds (exp_R . (p + q)) - (exp_R . p) = (q * (exp_R . p)) + ((q * (exp_R . p)) * (Re (Sum (q P_dt))))
proof end;

theorem Th62: :: SIN_COS:63
for p being Real holds
( cos is_differentiable_in p & diff (cos,p) = - (sin . p) )
proof end;

theorem Th63: :: SIN_COS:64
for p being Real holds
( sin is_differentiable_in p & diff (sin,p) = cos . p )
proof end;

theorem Th64: :: SIN_COS:65
for p being Real holds
( exp_R is_differentiable_in p & diff (exp_R,p) = exp_R . p )
proof end;

theorem Th65: :: SIN_COS:66
for th being Real holds
( exp_R is_differentiable_on REAL & diff (exp_R,th) = exp_R . th )
proof end;

theorem Th66: :: SIN_COS:67
for th being Real holds
( cos is_differentiable_on REAL & diff (cos,th) = - (sin . th) )
proof end;

theorem Th67: :: SIN_COS:68
for th being Real holds
( sin is_differentiable_on REAL & diff (sin,th) = cos . th )
proof end;

theorem Th68: :: SIN_COS:69
for th being Real st th in [.0,1.] holds
( 0 < cos . th & cos . th >= 1 / 2 )
proof end;

definition
func tan -> PartFunc of REAL,REAL equals :: SIN_COS:def 26
sin / cos;
coherence
sin / cos is PartFunc of REAL,REAL
;
func cot -> PartFunc of REAL,REAL equals :: SIN_COS:def 27
cos / sin;
coherence
cos / sin is PartFunc of REAL,REAL
;
end;

:: deftheorem defines tan SIN_COS:def 26 :
tan = sin / cos;

:: deftheorem defines cot SIN_COS:def 27 :
cot = cos / sin;

theorem Th69: :: SIN_COS:70
( [.0,1.] c= dom tan & ].0,1.[ c= dom tan )
proof end;

Lm14: ( dom (tan | [.0,1.]) = [.0,1.] & ( for th being Real st th in [.0,1.] holds
(tan | [.0,1.]) . th = tan . th ) )

proof end;

Lm15: ( tan is_differentiable_on ].0,1.[ & ( for th being Real st th in ].0,1.[ holds
diff (tan,th) > 0 ) )

proof end;

theorem Th70: :: SIN_COS:71
tan | [.0,1.] is continuous
proof end;

theorem Th71: :: SIN_COS:72
for th1, th2 being Real st th1 in ].0,1.[ & th2 in ].0,1.[ & tan . th1 = tan . th2 holds
th1 = th2
proof end;

Lm16: ( tan . 0 = 0 & tan . 1 > 1 )
proof end;

definition
func PI -> Real means :Def28: :: SIN_COS:def 28
( tan . (it / 4) = 1 & it in ].0,4.[ );
existence
ex b1 being Real st
( tan . (b1 / 4) = 1 & b1 in ].0,4.[ )
proof end;
uniqueness
for b1, b2 being Real st tan . (b1 / 4) = 1 & b1 in ].0,4.[ & tan . (b2 / 4) = 1 & b2 in ].0,4.[ holds
b1 = b2
proof end;
end;

:: deftheorem Def28 defines PI SIN_COS:def 28 :
for b1 being Real holds
( b1 = PI iff ( tan . (b1 / 4) = 1 & b1 in ].0,4.[ ) );

theorem Th72: :: SIN_COS:73
sin . (PI / 4) = cos . (PI / 4)
proof end;

theorem Th73: :: SIN_COS:74
for th1, th2 being Real holds
( sin . (th1 + th2) = ((sin . th1) * (cos . th2)) + ((cos . th1) * (sin . th2)) & cos . (th1 + th2) = ((cos . th1) * (cos . th2)) - ((sin . th1) * (sin . th2)) )
proof end;

theorem :: SIN_COS:75
for th1, th2 being Real holds
( sin (th1 + th2) = ((sin th1) * (cos th2)) + ((cos th1) * (sin th2)) & cos (th1 + th2) = ((cos th1) * (cos th2)) - ((sin th1) * (sin th2)) ) by Th73;

theorem Th75: :: SIN_COS:76
( cos . (PI / 2) = 0 & sin . (PI / 2) = 1 & cos . PI = - 1 & sin . PI = 0 & cos . (PI + (PI / 2)) = 0 & sin . (PI + (PI / 2)) = - 1 & cos . (2 * PI) = 1 & sin . (2 * PI) = 0 )
proof end;

theorem :: SIN_COS:77
( cos (PI / 2) = 0 & sin (PI / 2) = 1 & cos PI = - 1 & sin PI = 0 & cos (PI + (PI / 2)) = 0 & sin (PI + (PI / 2)) = - 1 & cos (2 * PI) = 1 & sin (2 * PI) = 0 ) by Th75;

theorem Th77: :: SIN_COS:78
for th being Real holds
( sin . (th + (2 * PI)) = sin . th & cos . (th + (2 * PI)) = cos . th & sin . ((PI / 2) - th) = cos . th & cos . ((PI / 2) - th) = sin . th & sin . ((PI / 2) + th) = cos . th & cos . ((PI / 2) + th) = - (sin . th) & sin . (PI + th) = - (sin . th) & cos . (PI + th) = - (cos . th) )
proof end;

theorem :: SIN_COS:79
for th being Real holds
( sin (th + (2 * PI)) = sin th & cos (th + (2 * PI)) = cos th & sin ((PI / 2) - th) = cos th & cos ((PI / 2) - th) = sin th & sin ((PI / 2) + th) = cos th & cos ((PI / 2) + th) = - (sin th) & sin (PI + th) = - (sin th) & cos (PI + th) = - (cos th) ) by Th77;

Lm17: for th being Real st th in [.0,1.] holds
sin . th >= 0

proof end;

theorem Th79: :: SIN_COS:80
for th being Real st th in ].0,(PI / 2).[ holds
cos . th > 0
proof end;

theorem :: SIN_COS:81
for th being Real st th in ].0,(PI / 2).[ holds
cos th > 0 by Th79;

:: from COMPLEX2, 2006.03,06, A.T.
theorem :: SIN_COS:82
for a, b being Real holds sin (a - b) = ((sin a) * (cos b)) - ((cos a) * (sin b))
proof end;

theorem :: SIN_COS:83
for a, b being Real holds cos (a - b) = ((cos a) * (cos b)) + ((sin a) * (sin b))
proof end;

registration
cluster sin -> continuous ;
coherence
sin is continuous
proof end;
cluster cos -> continuous ;
coherence
cos is continuous
proof end;
cluster exp_R -> continuous ;
coherence
exp_R is continuous
proof end;
end;