begin
:: deftheorem SIN_COS:def 1 :
canceled;
:: deftheorem Def2 defines CHK SIN_COS:def 2 :
for m, k being Element of NAT holds
( ( m <= k implies CHK (m,k) = 1 ) & ( not m <= k implies CHK (m,k) = 0 ) );
:: deftheorem SIN_COS:def 3 :
canceled;
:: deftheorem SIN_COS:def 4 :
canceled;
:: deftheorem Def5 defines Prod_real_n SIN_COS:def 5 :
for b1 being Real_Sequence holds
( b1 = Prod_real_n iff ( b1 . 0 = 1 & ( for n being Element of NAT holds b1 . (n + 1) = (b1 . n) * (n + 1) ) ) );
:: deftheorem SIN_COS:def 6 :
canceled;
:: deftheorem defines ! SIN_COS:def 7 :
for n being Nat holds n ! = Prod_real_n . n;
:: deftheorem Def8 defines ExpSeq SIN_COS:def 8 :
for z being complex number
for b2 being Complex_Sequence holds
( b2 = z ExpSeq iff for n being Element of NAT holds b2 . n = (z #N n) / (n !) );
:: deftheorem Def9 defines rExpSeq SIN_COS:def 9 :
for a being real number
for b2 being Real_Sequence holds
( b2 = a rExpSeq iff for n being Element of NAT holds b2 . n = (a |^ n) / (n !) );
theorem Th1:
theorem
canceled;
theorem Th3:
:: deftheorem Def10 defines Coef SIN_COS:def 10 :
for n being Element of NAT
for b2 being Complex_Sequence holds
( b2 = Coef n iff for k being Element of NAT holds
( ( k <= n implies b2 . k = (n !) / ((k !) * ((n -' k) !)) ) & ( k > n implies b2 . k = 0 ) ) );
:: deftheorem Def11 defines Coef_e SIN_COS:def 11 :
for n being Element of NAT
for b2 being Complex_Sequence holds
( b2 = Coef_e n iff for k being Element of NAT holds
( ( k <= n implies b2 . k = 1r / ((k !) * ((n -' k) !)) ) & ( k > n implies b2 . k = 0 ) ) );
:: deftheorem Def12 defines Shift SIN_COS:def 12 :
for seq, b2 being Complex_Sequence holds
( b2 = Shift seq iff ( b2 . 0 = 0 & ( for k being Element of NAT holds b2 . (k + 1) = seq . k ) ) );
:: deftheorem Def13 defines Expan SIN_COS:def 13 :
for n being Element of NAT
for z, w being Element of COMPLEX
for b4 being Complex_Sequence holds
( b4 = Expan (n,z,w) iff for k being Element of NAT holds
( ( k <= n implies b4 . k = (((Coef n) . k) * (z |^ k)) * (w |^ (n -' k)) ) & ( n < k implies b4 . k = 0 ) ) );
:: deftheorem Def14 defines Expan_e SIN_COS:def 14 :
for n being Element of NAT
for z, w being Element of COMPLEX
for b4 being Complex_Sequence holds
( b4 = Expan_e (n,z,w) iff for k being Element of NAT holds
( ( k <= n implies b4 . k = (((Coef_e n) . k) * (z |^ k)) * (w |^ (n -' k)) ) & ( n < k implies b4 . k = 0 ) ) );
:: deftheorem Def15 defines Alfa SIN_COS:def 15 :
for n being Element of NAT
for z, w being Element of COMPLEX
for b4 being Complex_Sequence holds
( b4 = Alfa (n,z,w) iff for k being Element of NAT holds
( ( k <= n implies b4 . k = ((z ExpSeq) . k) * ((Partial_Sums (w ExpSeq)) . (n -' k)) ) & ( n < k implies b4 . k = 0 ) ) );
:: deftheorem defines Conj SIN_COS:def 16 :
for a, b being real number
for n being Element of NAT
for b4 being Real_Sequence holds
( b4 = Conj (n,a,b) iff for k being Element of 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 ) ) );
:: deftheorem Def17 defines Conj SIN_COS:def 17 :
for z, w being Element of COMPLEX
for n being Element of NAT
for b4 being Complex_Sequence holds
( b4 = Conj (n,z,w) iff for k being Element of 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 Element of REAL holds
( (p1 + (g1 * <i>)) * (p2 + (g2 * <i>)) = ((p1 * p2) - (g1 * g2)) + (((p1 * g2) + (p2 * g1)) * <i>) & (p2 + (g2 * <i>)) *' = p2 + ((- g2) * <i>) )
theorem Th4:
theorem Th5:
theorem Th6:
theorem Th7:
theorem Th8:
theorem Th9:
theorem Th10:
theorem Th11:
theorem Th12:
theorem Th13:
theorem Th14:
theorem Th15:
theorem Th16:
theorem Th17:
theorem Th18:
theorem Th19:
theorem Th20:
theorem Th21:
theorem Th22:
theorem Th23:
Lm2:
for z, w being complex number holds (Sum (z ExpSeq)) * (Sum (w ExpSeq)) = Sum ((z + w) ExpSeq)
begin
:: deftheorem Def18 defines exp SIN_COS:def 18 :
for b1 being Function of COMPLEX,COMPLEX holds
( b1 = exp iff for z being complex number holds b1 . z = Sum (z ExpSeq) );
:: deftheorem defines exp SIN_COS:def 19 :
for z being complex number holds exp z = exp . z;
theorem
begin
:: deftheorem Def20 defines sin SIN_COS:def 20 :
for b1 being Function of REAL,REAL holds
( b1 = sin iff for d being Element of REAL holds b1 . d = Im (Sum ((d * <i>) ExpSeq)) );
:: deftheorem defines sin SIN_COS:def 21 :
for th being real number holds sin th = sin . th;
:: deftheorem Def22 defines cos SIN_COS:def 22 :
for b1 being Function of REAL,REAL holds
( b1 = cos iff for d being Real holds b1 . d = Re (Sum ((d * <i>) ExpSeq)) );
:: deftheorem defines cos SIN_COS:def 23 :
for th being real number holds cos th = cos . th;
theorem
canceled;
theorem
canceled;
theorem Th27:
Lm3:
for th being Real holds Sum ((th * <i>) ExpSeq) = (cos . th) + ((sin . th) * <i>)
theorem
Lm4:
for th being Real holds (Sum ((th * <i>) ExpSeq)) *' = Sum ((- (th * <i>)) ExpSeq)
theorem
Lm5:
for th being Real
for th1 being real number st th = th1 holds
( |.(Sum ((th * <i>) ExpSeq)).| = 1 & abs (sin . th1) <= 1 & abs (cos . th1) <= 1 )
theorem
theorem Th31:
theorem
theorem Th33:
theorem
:: deftheorem Def24 defines P_sin SIN_COS:def 24 :
for th being real number
for b2 being Real_Sequence holds
( b2 = th P_sin iff for n being Element of NAT holds b2 . n = (((- 1) |^ n) * (th |^ ((2 * n) + 1))) / (((2 * n) + 1) !) );
:: deftheorem Def25 defines P_cos SIN_COS:def 25 :
for th being real number
for b2 being Real_Sequence holds
( b2 = th P_cos iff for n being Element of NAT holds b2 . n = (((- 1) |^ n) * (th |^ (2 * n))) / ((2 * n) !) );
theorem Th35:
theorem Th36:
theorem
theorem Th38:
theorem Th39:
theorem Th40:
theorem
deffunc H1( Real) -> Element of REAL = (2 * $1) + 1;
consider bq being Real_Sequence such that
Lm6:
for n being Element of NAT holds bq . n = H1(n)
from SEQ_1:sch 1();
bq is increasing sequence of NAT
then reconsider bq = bq as increasing sequence of NAT ;
Lm7:
for n being Element of NAT
for th, th1, th2, th3 being real number holds
( th |^ (2 * n) = (th |^ n) |^ 2 & (- 1) |^ (2 * n) = 1 & (- 1) |^ ((2 * n) + 1) = - 1 )
theorem Th42:
theorem Th43:
theorem
canceled;
theorem
theorem Th46:
theorem Th47:
theorem Th48:
theorem Th49:
Lm8:
for n being Element of NAT
for z being complex number holds
( (z ExpSeq) . 1 = z & (z ExpSeq) . 0 = 1r & |.(z |^ n).| = |.z.| |^ n )
Lm9:
for th being Real holds Sum (th ExpSeq) = Sum (th rExpSeq)
theorem Th50:
:: deftheorem Def26 defines exp_R SIN_COS:def 26 :
for b1 being Function of REAL,REAL holds
( b1 = exp_R iff for d being real number holds b1 . d = Sum (d rExpSeq) );
:: deftheorem defines exp_R SIN_COS:def 27 :
for th being real number holds exp_R th = exp_R . th;
theorem Th51:
theorem
canceled;
theorem Th53:
theorem
Lm10:
for p, q being real number holds exp_R . (p + q) = (exp_R . p) * (exp_R . q)
theorem
Lm11:
exp_R . 0 = 1
theorem
theorem Th57:
theorem Th58:
theorem Th59:
theorem
begin
:: deftheorem Def28 defines P_dt SIN_COS:def 28 :
for z being complex number
for b2 being Complex_Sequence holds
( b2 = z P_dt iff for n being Element of NAT holds b2 . n = (z |^ (n + 1)) / ((n + 2) !) );
:: deftheorem defines P_t SIN_COS:def 29 :
for z being complex number
for b2 being Complex_Sequence holds
( b2 = z P_t iff for n being Element of NAT holds b2 . n = (z #N n) / ((n + 2) !) );
Lm12:
for p being Real
for z being complex number 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) )
Lm13:
for p being real number
for z being complex number st p > 0 holds
( Re (z / (p * <i>)) = (Im z) / p & Im (z / (p * <i>)) = - ((Re z) / p) & |.(z / p).| = |.z.| / p )
theorem Th61:
theorem Th62:
theorem Th63:
theorem Th64:
theorem Th65:
theorem Th66:
theorem Th67:
theorem Th68:
theorem Th69:
theorem Th70:
theorem Th71:
theorem Th72:
theorem Th73:
theorem Th74:
:: deftheorem defines tan SIN_COS:def 30 :
tan = sin / cos;
:: deftheorem defines cot SIN_COS:def 31 :
cot = cos / sin;
theorem Th75:
Lm14:
( dom (tan | [.0,1.]) = [.0,1.] & ( for th being real number st th in [.0,1.] holds
(tan | [.0,1.]) . th = tan . th ) )
Lm15:
( tan is_differentiable_on ].0,1.[ & ( for th being real number st th in ].0,1.[ holds
diff (tan,th) > 0 ) )
theorem Th76:
theorem Th77:
Lm16:
( tan . 0 = 0 & tan . 1 > 1 )
begin
:: deftheorem Def32 defines PI SIN_COS:def 32 :
for b1 being real number holds
( b1 = PI iff ( tan . (b1 / 4) = 1 & b1 in ].0,4.[ ) );
theorem Th78:
begin
theorem Th79:
theorem
theorem Th81:
theorem
theorem Th83:
theorem
Lm17:
for th being real number st th in [.0,1.] holds
sin . th >= 0
theorem Th85:
theorem
begin
theorem
canceled;
theorem
theorem