:: Several Integrability Formulas of Special Functions
:: by Cuiying Peng , Fuguo Ge and Xiquan Liang
::
:: Received August 28, 2007
:: Copyright (c) 2007 Association of Mizar Users


begin

Lm1: sin (PI / 4) > 0
proof end;

Lm2: sin (- (PI / 4)) < 0
proof end;

theorem Th1: :: INTEGRA8:1
for x being Real
for n being Element of NAT holds sin (x + ((2 * n) * PI )) = sin x
proof end;

theorem Th2: :: INTEGRA8:2
for x being Real
for n being Element of NAT holds sin (x + (((2 * n) + 1) * PI )) = - (sin x)
proof end;

theorem Th3: :: INTEGRA8:3
for x being Real
for n being Element of NAT holds cos (x + ((2 * n) * PI )) = cos x
proof end;

theorem Th4: :: INTEGRA8:4
for x being Real
for n being Element of NAT holds cos (x + (((2 * n) + 1) * PI )) = - (cos x)
proof end;

theorem Th5: :: INTEGRA8:5
for x being Real st sin (x / 2) >= 0 holds
sin (x / 2) = sqrt ((1 - (cos x)) / 2)
proof end;

theorem Th6: :: INTEGRA8:6
for x being Real st sin (x / 2) < 0 holds
sin (x / 2) = - (sqrt ((1 - (cos x)) / 2))
proof end;

theorem Th7: :: INTEGRA8:7
sin (PI / 4) = (sqrt 2) / 2
proof end;

theorem Th8: :: INTEGRA8:8
sin (- (PI / 4)) = - ((sqrt 2) / 2)
proof end;

theorem Th9: :: INTEGRA8:9
[.(- ((sqrt 2) / 2)),((sqrt 2) / 2).] c= ].(- 1),1.[
proof end;

theorem Th10: :: INTEGRA8:10
arcsin ((sqrt 2) / 2) = PI / 4
proof end;

theorem Th11: :: INTEGRA8:11
arcsin (- ((sqrt 2) / 2)) = - (PI / 4)
proof end;

theorem Th12: :: INTEGRA8:12
for x being Real st cos (x / 2) >= 0 holds
cos (x / 2) = sqrt ((1 + (cos x)) / 2)
proof end;

Lm3: cos (PI / 4) > 0
proof end;

theorem Th13: :: INTEGRA8:13
cos (PI / 4) = (sqrt 2) / 2
proof end;

theorem Th14: :: INTEGRA8:14
cos ((3 * PI ) / 4) = - ((sqrt 2) / 2)
proof end;

theorem Th15: :: INTEGRA8:15
arccos ((sqrt 2) / 2) = PI / 4
proof end;

theorem Th16: :: INTEGRA8:16
arccos (- ((sqrt 2) / 2)) = (3 * PI ) / 4
proof end;

theorem Th17: :: INTEGRA8:17
sinh . 1 = ((number_e ^2 ) - 1) / (2 * number_e )
proof end;

theorem Th18: :: INTEGRA8:18
cosh . 0 = 1
proof end;

theorem Th19: :: INTEGRA8:19
cosh . 1 = ((number_e ^2 ) + 1) / (2 * number_e )
proof end;

Lm4: for f1, f2 being PartFunc of REAL ,REAL holds (- f1) (#) (- f2) = f1 (#) f2
proof end;

theorem Th20: :: INTEGRA8:20
for L1 being LINEAR holds - L1 is LINEAR
proof end;

theorem Th21: :: INTEGRA8:21
for R1 being REST holds - R1 is REST
proof end;

theorem Th22: :: INTEGRA8:22
for f1 being PartFunc of REAL ,REAL
for x0 being Real st f1 is_differentiable_in x0 holds
( - f1 is_differentiable_in x0 & diff (- f1),x0 = - (diff f1,x0) )
proof end;

theorem Th23: :: INTEGRA8:23
for f1 being PartFunc of REAL ,REAL
for Z being open Subset of REAL st Z c= dom (- f1) & f1 is_differentiable_on Z holds
( - f1 is_differentiable_on Z & ( for x being Real st x in Z holds
((- f1) `| Z) . x = - (diff f1,x) ) )
proof end;

theorem :: INTEGRA8:24
- sin is_differentiable_on REAL
proof end;

theorem Th25: :: INTEGRA8:25
for x being Real holds
( - cos is_differentiable_in x & diff (- cos ),x = sin . x )
proof end;

theorem Th26: :: INTEGRA8:26
( - cos is_differentiable_on REAL & ( for x being Real st x in REAL holds
diff (- cos ),x = sin . x ) )
proof end;

theorem Th27: :: INTEGRA8:27
sin `| REAL = cos
proof end;

theorem Th28: :: INTEGRA8:28
cos `| REAL = - sin
proof end;

theorem Th29: :: INTEGRA8:29
(- cos ) `| REAL = sin by Th26, FDIFF_1:def 8, SIN_COS:27;

theorem Th30: :: INTEGRA8:30
sinh `| REAL = cosh
proof end;

theorem Th31: :: INTEGRA8:31
cosh `| REAL = sinh
proof end;

theorem Th32: :: INTEGRA8:32
exp_R `| REAL = exp_R
proof end;

theorem Th33: :: INTEGRA8:33
for f being PartFunc of REAL ,REAL
for Z being open Subset of REAL st Z c= dom tan & ( for x being Real st x in Z holds
( f . x = 1 / ((cos . x) ^2 ) & cos . x <> 0 ) ) holds
( tan is_differentiable_on Z & ( for x being Real st x in Z holds
(tan `| Z) . x = 1 / ((cos . x) ^2 ) ) )
proof end;

theorem Th34: :: INTEGRA8:34
for f being PartFunc of REAL ,REAL
for Z being open Subset of REAL st Z c= dom cot & ( for x being Real st x in Z holds
( f . x = - (1 / ((sin . x) ^2 )) & sin . x <> 0 ) ) holds
( cot is_differentiable_on Z & ( for x being Real st x in Z holds
(cot `| Z) . x = - (1 / ((sin . x) ^2 )) ) )
proof end;

theorem :: INTEGRA8:35
for r being Real holds rng (REAL --> r) c= REAL
proof end;

definition
let r be Real;
func Cst r -> Function of REAL ,REAL equals :: INTEGRA8:def 1
REAL --> r;
coherence
REAL --> r is Function of REAL ,REAL
;
end;

:: deftheorem defines Cst INTEGRA8:def 1 :
for r being Real holds Cst r = REAL --> r;

theorem Th36: :: INTEGRA8:36
for a, b being Real
for A being closed-interval Subset of REAL holds chi A,A = (Cst 1) | A
proof end;

theorem Th37: :: INTEGRA8:37
for a, b being Real
for A being closed-interval Subset of REAL st A = [.a,b.] holds
( upper_bound A = b & lower_bound A = a )
proof end;

begin

theorem :: INTEGRA8:38
for a, b being Real st a <= b holds
integral (Cst 1),a,b = b - a
proof end;

Lm5: dom sin = REAL
by FUNCT_2:def 1;

Lm6: dom cos = REAL
by FUNCT_2:def 1;

Lm7: dom (- sin ) = REAL
by FUNCT_2:def 1;

Lm8: dom exp_R = REAL
by FUNCT_2:def 1;

Lm9: dom sinh = REAL
by FUNCT_2:def 1;

Lm10: dom cosh = REAL
by FUNCT_2:def 1;

Lm11: for A being closed-interval Subset of REAL holds
( cos is_integrable_on A & cos | A is bounded )
proof end;

theorem Th39: :: INTEGRA8:39
for A being closed-interval Subset of REAL holds integral cos ,A = (sin . (upper_bound A)) - (sin . (lower_bound A))
proof end;

theorem :: INTEGRA8:40
for A being closed-interval Subset of REAL st A = [.0 ,(PI / 2).] holds
integral cos ,A = 1
proof end;

theorem :: INTEGRA8:41
for A being closed-interval Subset of REAL st A = [.0 ,PI .] holds
integral cos ,A = 0
proof end;

theorem :: INTEGRA8:42
for A being closed-interval Subset of REAL st A = [.0 ,((PI * 3) / 2).] holds
integral cos ,A = - 1
proof end;

theorem :: INTEGRA8:43
for A being closed-interval Subset of REAL st A = [.0 ,(PI * 2).] holds
integral cos ,A = 0
proof end;

theorem :: INTEGRA8:44
for A being closed-interval Subset of REAL
for n being Element of NAT st A = [.((2 * n) * PI ),(((2 * n) + 1) * PI ).] holds
integral cos ,A = 0
proof end;

theorem :: INTEGRA8:45
for A being closed-interval Subset of REAL
for x being Real
for n being Element of NAT st A = [.(x + ((2 * n) * PI )),(x + (((2 * n) + 1) * PI )).] holds
integral cos ,A = - (2 * (sin x))
proof end;

Lm12: for A being closed-interval Subset of REAL holds
( - sin is_integrable_on A & (- sin ) | A is bounded )
proof end;

theorem Th46: :: INTEGRA8:46
for A being closed-interval Subset of REAL holds integral (- sin ),A = (cos . (upper_bound A)) - (cos . (lower_bound A))
proof end;

theorem :: INTEGRA8:47
for A being closed-interval Subset of REAL st A = [.0 ,(PI / 2).] holds
integral (- sin ),A = - 1
proof end;

theorem :: INTEGRA8:48
for A being closed-interval Subset of REAL st A = [.0 ,PI .] holds
integral (- sin ),A = - 2
proof end;

theorem :: INTEGRA8:49
for A being closed-interval Subset of REAL st A = [.0 ,((PI * 3) / 2).] holds
integral (- sin ),A = - 1
proof end;

theorem :: INTEGRA8:50
for A being closed-interval Subset of REAL st A = [.0 ,(PI * 2).] holds
integral (- sin ),A = 0
proof end;

theorem :: INTEGRA8:51
for A being closed-interval Subset of REAL
for n being Element of NAT st A = [.((2 * n) * PI ),(((2 * n) + 1) * PI ).] holds
integral (- sin ),A = - 2
proof end;

theorem :: INTEGRA8:52
for A being closed-interval Subset of REAL
for x being Real
for n being Element of NAT st A = [.(x + ((2 * n) * PI )),(x + (((2 * n) + 1) * PI )).] holds
integral (- sin ),A = - (2 * (cos x))
proof end;

Lm13: for A being closed-interval Subset of REAL holds
( exp_R is_integrable_on A & exp_R | A is bounded )
proof end;

theorem Th53: :: INTEGRA8:53
for A being closed-interval Subset of REAL holds integral exp_R ,A = (exp_R . (upper_bound A)) - (exp_R . (lower_bound A))
proof end;

theorem :: INTEGRA8:54
for A being closed-interval Subset of REAL st A = [.0 ,1.] holds
integral exp_R ,A = number_e - 1
proof end;

Lm14: for A being closed-interval Subset of REAL holds sinh | A is continuous
proof end;

Lm15: for A being closed-interval Subset of REAL holds
( sinh is_integrable_on A & sinh | A is bounded )
proof end;

theorem Th55: :: INTEGRA8:55
for A being closed-interval Subset of REAL holds integral sinh ,A = (cosh . (upper_bound A)) - (cosh . (lower_bound A))
proof end;

theorem :: INTEGRA8:56
for A being closed-interval Subset of REAL st A = [.0 ,1.] holds
integral sinh ,A = ((number_e - 1) ^2 ) / (2 * number_e )
proof end;

Lm16: for A being closed-interval Subset of REAL holds cosh | A is continuous
proof end;

Lm17: for A being closed-interval Subset of REAL holds
( cosh is_integrable_on A & cosh | A is bounded )
proof end;

theorem Th57: :: INTEGRA8:57
for A being closed-interval Subset of REAL holds integral cosh ,A = (sinh . (upper_bound A)) - (sinh . (lower_bound A))
proof end;

theorem :: INTEGRA8:58
for A being closed-interval Subset of REAL st A = [.0 ,1.] holds
integral cosh ,A = ((number_e ^2 ) - 1) / (2 * number_e )
proof end;

theorem :: INTEGRA8:59
for f2 being PartFunc of REAL ,REAL
for A being closed-interval Subset of REAL
for Z being open Subset of REAL st A c= Z & dom tan = Z & dom tan = dom f2 & ( for x being Real st x in Z holds
( f2 . x = 1 / ((cos . x) ^2 ) & cos . x <> 0 ) ) & f2 | A is continuous holds
integral f2,A = (tan . (upper_bound A)) - (tan . (lower_bound A))
proof end;

theorem :: INTEGRA8:60
for f2 being PartFunc of REAL ,REAL
for A being closed-interval Subset of REAL
for Z being open Subset of REAL st A c= Z & dom cot = Z & dom cot = dom f2 & ( for x being Real st x in Z holds
( f2 . x = - (1 / ((sin . x) ^2 )) & sin . x <> 0 ) ) & f2 | A is continuous holds
integral f2,A = (cot . (upper_bound A)) - (cot . (lower_bound A))
proof end;

theorem :: INTEGRA8:61
for f2 being PartFunc of REAL ,REAL
for A being closed-interval Subset of REAL st dom tanh = dom f2 & ( for x being Real st x in REAL holds
f2 . x = 1 / ((cosh . x) ^2 ) ) & f2 | A is continuous holds
integral f2,A = (tanh . (upper_bound A)) - (tanh . (lower_bound A))
proof end;

Lm18: dom (arcsin `| ].(- 1),1.[) = ].(- 1),1.[
by FDIFF_1:def 8, SIN_COS6:84;

theorem Th62: :: INTEGRA8:62
for f2 being PartFunc of REAL ,REAL
for A being closed-interval Subset of REAL st A c= ].(- 1),1.[ & dom (arcsin `| ].(- 1),1.[) = dom f2 & ( for x being Real holds
( x in ].(- 1),1.[ & f2 . x = 1 / (sqrt (1 - (x ^2 ))) ) ) & f2 | A is continuous holds
integral f2,A = (arcsin . (upper_bound A)) - (arcsin . (lower_bound A))
proof end;

theorem Th63: :: INTEGRA8:63
for f2 being PartFunc of REAL ,REAL
for A being closed-interval Subset of REAL st A c= ].(- 1),1.[ & dom (arccos `| ].(- 1),1.[) = dom f2 & ( for x being Real holds
( x in ].(- 1),1.[ & f2 . x = - (1 / (sqrt (1 - (x ^2 )))) ) ) & f2 | A is continuous holds
integral f2,A = (arccos . (upper_bound A)) - (arccos . (lower_bound A))
proof end;

theorem :: INTEGRA8:64
for f2 being PartFunc of REAL ,REAL
for A being closed-interval Subset of REAL st A = [.(- ((sqrt 2) / 2)),((sqrt 2) / 2).] & dom (arcsin `| ].(- 1),1.[) = dom f2 & ( for x being Real holds
( x in ].(- 1),1.[ & f2 . x = 1 / (sqrt (1 - (x ^2 ))) ) ) & f2 | A is continuous holds
integral f2,A = PI / 2
proof end;

theorem :: INTEGRA8:65
for f2 being PartFunc of REAL ,REAL
for A being closed-interval Subset of REAL st A = [.(- ((sqrt 2) / 2)),((sqrt 2) / 2).] & dom (arccos `| ].(- 1),1.[) = dom f2 & ( for x being Real holds
( x in ].(- 1),1.[ & f2 . x = - (1 / (sqrt (1 - (x ^2 )))) ) ) & f2 | A is continuous holds
integral f2,A = - (PI / 2)
proof end;

theorem Th66: :: INTEGRA8:66
for f, g being PartFunc of REAL ,REAL
for A being closed-interval Subset of REAL
for Z being open Subset of REAL st f is_differentiable_on Z & g is_differentiable_on Z & A c= Z & f `| Z is_integrable_on A & (f `| Z) | A is bounded & g `| Z is_integrable_on A & (g `| Z) | A is bounded holds
integral ((f `| Z) + (g `| Z)),A = (((f . (upper_bound A)) - (f . (lower_bound A))) + (g . (upper_bound A))) - (g . (lower_bound A))
proof end;

theorem Th67: :: INTEGRA8:67
for f, g being PartFunc of REAL ,REAL
for A being closed-interval Subset of REAL
for Z being open Subset of REAL st f is_differentiable_on Z & g is_differentiable_on Z & A c= Z & f `| Z is_integrable_on A & (f `| Z) | A is bounded & g `| Z is_integrable_on A & (g `| Z) | A is bounded holds
integral ((f `| Z) - (g `| Z)),A = ((f . (upper_bound A)) - (f . (lower_bound A))) - ((g . (upper_bound A)) - (g . (lower_bound A)))
proof end;

theorem Th68: :: INTEGRA8:68
for f being PartFunc of REAL ,REAL
for A being closed-interval Subset of REAL
for r being Real
for Z being open Subset of REAL st f is_differentiable_on Z & A c= Z & f `| Z is_integrable_on A & (f `| Z) | A is bounded holds
integral (r (#) (f `| Z)),A = (r * (f . (upper_bound A))) - (r * (f . (lower_bound A)))
proof end;

Lm19: for A being closed-interval Subset of REAL holds
( sin is_integrable_on A & sin | A is bounded )
proof end;

theorem Th69: :: INTEGRA8:69
for A being closed-interval Subset of REAL holds integral (sin + cos ),A = ((((- cos ) . (upper_bound A)) - ((- cos ) . (lower_bound A))) + (sin . (upper_bound A))) - (sin . (lower_bound A))
proof end;

theorem :: INTEGRA8:70
for A being closed-interval Subset of REAL st A = [.0 ,(PI / 2).] holds
integral (sin + cos ),A = 2
proof end;

theorem :: INTEGRA8:71
for A being closed-interval Subset of REAL st A = [.0 ,PI .] holds
integral (sin + cos ),A = 2
proof end;

theorem :: INTEGRA8:72
for A being closed-interval Subset of REAL st A = [.0 ,((PI * 3) / 2).] holds
integral (sin + cos ),A = 0
proof end;

theorem :: INTEGRA8:73
for A being closed-interval Subset of REAL st A = [.0 ,(PI * 2).] holds
integral (sin + cos ),A = 0
proof end;

theorem :: INTEGRA8:74
for A being closed-interval Subset of REAL
for n being Element of NAT st A = [.((2 * n) * PI ),(((2 * n) + 1) * PI ).] holds
integral (sin + cos ),A = 2
proof end;

theorem :: INTEGRA8:75
for A being closed-interval Subset of REAL
for x being Real
for n being Element of NAT st A = [.(x + ((2 * n) * PI )),(x + (((2 * n) + 1) * PI )).] holds
integral (sin + cos ),A = (2 * (cos x)) - (2 * (sin x))
proof end;

theorem Th76: :: INTEGRA8:76
for A being closed-interval Subset of REAL holds integral (sinh + cosh ),A = (((cosh . (upper_bound A)) - (cosh . (lower_bound A))) + (sinh . (upper_bound A))) - (sinh . (lower_bound A))
proof end;

theorem :: INTEGRA8:77
for A being closed-interval Subset of REAL st A = [.0 ,1.] holds
integral (sinh + cosh ),A = number_e - 1
proof end;

theorem Th78: :: INTEGRA8:78
for A being closed-interval Subset of REAL holds integral (sin - cos ),A = (((- cos ) . (upper_bound A)) - ((- cos ) . (lower_bound A))) - ((sin . (upper_bound A)) - (sin . (lower_bound A)))
proof end;

theorem :: INTEGRA8:79
for A being closed-interval Subset of REAL st A = [.0 ,(PI / 2).] holds
integral (sin - cos ),A = 0
proof end;

theorem :: INTEGRA8:80
for A being closed-interval Subset of REAL st A = [.0 ,PI .] holds
integral (sin - cos ),A = 2
proof end;

theorem :: INTEGRA8:81
for A being closed-interval Subset of REAL st A = [.0 ,((PI * 3) / 2).] holds
integral (sin - cos ),A = 2
proof end;

theorem :: INTEGRA8:82
for A being closed-interval Subset of REAL st A = [.0 ,(PI * 2).] holds
integral (sin - cos ),A = 0
proof end;

theorem :: INTEGRA8:83
for A being closed-interval Subset of REAL
for n being Element of NAT st A = [.((2 * n) * PI ),(((2 * n) + 1) * PI ).] holds
integral (sin - cos ),A = 2
proof end;

theorem :: INTEGRA8:84
for A being closed-interval Subset of REAL
for x being Real
for n being Element of NAT st A = [.(x + ((2 * n) * PI )),(x + (((2 * n) + 1) * PI )).] holds
integral (sin - cos ),A = (2 * (cos x)) + (2 * (sin x))
proof end;

theorem :: INTEGRA8:85
for A being closed-interval Subset of REAL
for r being Real holds integral (r (#) sin ),A = (r * ((- cos ) . (upper_bound A))) - (r * ((- cos ) . (lower_bound A)))
proof end;

theorem :: INTEGRA8:86
for A being closed-interval Subset of REAL
for r being Real holds integral (r (#) cos ),A = (r * (sin . (upper_bound A))) - (r * (sin . (lower_bound A)))
proof end;

theorem :: INTEGRA8:87
for A being closed-interval Subset of REAL
for r being Real holds integral (r (#) sinh ),A = (r * (cosh . (upper_bound A))) - (r * (cosh . (lower_bound A)))
proof end;

theorem :: INTEGRA8:88
for A being closed-interval Subset of REAL
for r being Real holds integral (r (#) cosh ),A = (r * (sinh . (upper_bound A))) - (r * (sinh . (lower_bound A)))
proof end;

theorem :: INTEGRA8:89
for A being closed-interval Subset of REAL
for r being Real holds integral (r (#) exp_R ),A = (r * (exp_R . (upper_bound A))) - (r * (exp_R . (lower_bound A)))
proof end;

theorem Th90: :: INTEGRA8:90
for A being closed-interval Subset of REAL holds integral (sin (#) cos ),A = (1 / 2) * (((cos . (lower_bound A)) * (cos . (lower_bound A))) - ((cos . (upper_bound A)) * (cos . (upper_bound A))))
proof end;

theorem :: INTEGRA8:91
for A being closed-interval Subset of REAL st A = [.0 ,(PI / 2).] holds
integral (sin (#) cos ),A = 1 / 2
proof end;

theorem :: INTEGRA8:92
for A being closed-interval Subset of REAL st A = [.0 ,PI .] holds
integral (sin (#) cos ),A = 0
proof end;

theorem :: INTEGRA8:93
for A being closed-interval Subset of REAL st A = [.0 ,(PI * (3 / 2)).] holds
integral (sin (#) cos ),A = 1 / 2
proof end;

theorem :: INTEGRA8:94
for A being closed-interval Subset of REAL st A = [.0 ,(PI * 2).] holds
integral (sin (#) cos ),A = 0
proof end;

theorem :: INTEGRA8:95
for A being closed-interval Subset of REAL
for n being Element of NAT st A = [.((2 * n) * PI ),(((2 * n) + 1) * PI ).] holds
integral (sin (#) cos ),A = 0
proof end;

theorem :: INTEGRA8:96
for A being closed-interval Subset of REAL
for x being Real
for n being Element of NAT st A = [.(x + ((2 * n) * PI )),(x + (((2 * n) + 1) * PI )).] holds
integral (sin (#) cos ),A = 0
proof end;

Lm20: for A being closed-interval Subset of REAL holds
( cos (#) cos is_integrable_on A & (cos (#) cos ) | A is bounded )
proof end;

theorem :: INTEGRA8:97
for A being closed-interval Subset of REAL holds integral (sin (#) sin ),A = (((cos . (lower_bound A)) * (sin . (lower_bound A))) - ((cos . (upper_bound A)) * (sin . (upper_bound A)))) + (integral (cos (#) cos ),A)
proof end;

theorem :: INTEGRA8:98
for A being closed-interval Subset of REAL holds integral (sinh (#) sinh ),A = (((cosh . (upper_bound A)) * (sinh . (upper_bound A))) - ((cosh . (lower_bound A)) * (sinh . (lower_bound A)))) - (integral (cosh (#) cosh ),A)
proof end;

theorem :: INTEGRA8:99
for A being closed-interval Subset of REAL holds integral (sinh (#) cosh ),A = (1 / 2) * (((cosh . (upper_bound A)) * (cosh . (upper_bound A))) - ((cosh . (lower_bound A)) * (cosh . (lower_bound A))))
proof end;

theorem :: INTEGRA8:100
for A being closed-interval Subset of REAL holds integral (exp_R (#) exp_R ),A = (1 / 2) * (((exp_R . (upper_bound A)) ^2 ) - ((exp_R . (lower_bound A)) ^2 ))
proof end;

Lm21: for A being closed-interval Subset of REAL holds
( exp_R (#) (sin + cos ) is_integrable_on A & (exp_R (#) (sin + cos )) | A is bounded )
proof end;

Lm22: for A being closed-interval Subset of REAL holds
( exp_R (#) (cos - sin ) is_integrable_on A & (exp_R (#) (cos - sin )) | A is bounded )
proof end;

theorem :: INTEGRA8:101
for A being closed-interval Subset of REAL holds integral (exp_R (#) (sin + cos )),A = ((exp_R (#) sin ) . (upper_bound A)) - ((exp_R (#) sin ) . (lower_bound A))
proof end;

theorem :: INTEGRA8:102
for A being closed-interval Subset of REAL holds integral (exp_R (#) (cos - sin )),A = ((exp_R (#) cos ) . (upper_bound A)) - ((exp_R (#) cos ) . (lower_bound A))
proof end;