begin
:: deftheorem Def1 defines middle_volume INTEGR15:def 1 :
for A being closed-interval Subset of REAL
for f being Function of A,REAL
for D being Division of A
for b4 being FinSequence of REAL holds
( b4 is middle_volume of f,D iff ( len b4 = len D & ( for i being Nat st i in dom D holds
ex r being Element of REAL st
( r in rng (f | (divset (D,i))) & b4 . i = r * (vol (divset (D,i))) ) ) ) );
Lm1:
for A being closed-interval Subset of REAL
for f being Function of A,REAL
for D being Division of A
for F being middle_volume of f,D
for i being Nat st f | A is bounded_below & i in dom D holds
(lower_volume (f,D)) . i <= F . i
Lm2:
for A being closed-interval Subset of REAL
for f being Function of A,REAL
for D being Division of A
for e being Real st f | A is bounded_below & 0 < e holds
ex F being middle_volume of f,D st
for i being Nat st i in dom D holds
( (lower_volume (f,D)) . i <= F . i & F . i < ((lower_volume (f,D)) . i) + e )
Lm3:
for A being closed-interval Subset of REAL
for f being Function of A,REAL
for D being Division of A
for e being Real st f | A is bounded_above & 0 < e holds
ex F being middle_volume of f,D st
for i being Nat st i in dom D holds
( F . i <= (upper_volume (f,D)) . i & ((upper_volume (f,D)) . i) - e < F . i )
Lm4:
for A being closed-interval Subset of REAL
for f being Function of A,REAL
for D being Division of A
for F being middle_volume of f,D
for i being Nat st f | A is bounded_above & i in dom D holds
F . i <= (upper_volume (f,D)) . i
:: deftheorem defines middle_sum INTEGR15:def 2 :
for A being closed-interval Subset of REAL
for f being Function of A,REAL
for D being Division of A
for F being middle_volume of f,D holds middle_sum (f,F) = Sum F;
theorem Th1:
theorem Th2:
theorem Th3:
theorem Th4:
:: deftheorem Def3 defines middle_volume_Sequence INTEGR15:def 3 :
for A being closed-interval Subset of REAL
for f being Function of A,REAL
for T being DivSequence of A
for b4 being Function of NAT,(REAL *) holds
( b4 is middle_volume_Sequence of f,T iff for k being Element of NAT holds b4 . k is middle_volume of f,T . k );
:: deftheorem Def4 defines middle_sum INTEGR15:def 4 :
for A being closed-interval Subset of REAL
for f being Function of A,REAL
for T being DivSequence of A
for S being middle_volume_Sequence of f,T
for b5 being Real_Sequence holds
( b5 = middle_sum (f,S) iff for i being Element of NAT holds b5 . i = middle_sum (f,(S . i)) );
theorem Th5:
theorem Th6:
theorem Th7:
theorem Th8:
Lm5:
for p, q, r being Real_Sequence st p is convergent & r is convergent & lim p = lim r & ( for i being Element of NAT holds p . i <= q . i ) & ( for i being Element of NAT holds q . i <= r . i ) holds
( q is convergent & lim p = lim q & lim r = lim q )
theorem Th9:
theorem Th10:
:: deftheorem Def5 defines middle_volume INTEGR15:def 5 :
for n being Element of NAT
for A being closed-interval Subset of REAL
for f being Function of A,(REAL n)
for D being Division of A
for b5 being FinSequence of REAL n holds
( b5 is middle_volume of f,D iff ( len b5 = len D & ( for i being Nat st i in dom D holds
ex r being Element of REAL n st
( r in rng (f | (divset (D,i))) & b5 . i = (vol (divset (D,i))) * r ) ) ) );
:: deftheorem Def6 defines middle_sum INTEGR15:def 6 :
for n being Element of NAT
for A being closed-interval Subset of REAL
for f being Function of A,(REAL n)
for D being Division of A
for F being middle_volume of f,D
for b6 being Element of REAL n holds
( b6 = middle_sum (f,F) iff for i being Element of NAT st i in Seg n holds
ex Fi being FinSequence of REAL st
( Fi = (proj (i,n)) * F & b6 . i = Sum Fi ) );
:: deftheorem Def7 defines middle_volume_Sequence INTEGR15:def 7 :
for n being Element of NAT
for A being closed-interval Subset of REAL
for f being Function of A,(REAL n)
for T being DivSequence of A
for b5 being Function of NAT,((REAL n) *) holds
( b5 is middle_volume_Sequence of f,T iff for k being Element of NAT holds b5 . k is middle_volume of f,T . k );
:: deftheorem Def8 defines middle_sum INTEGR15:def 8 :
for n being Element of NAT
for A being closed-interval Subset of REAL
for f being Function of A,(REAL n)
for T being DivSequence of A
for S being middle_volume_Sequence of f,T
for b6 being sequence of (REAL-NS n) holds
( b6 = middle_sum (f,S) iff for i being Element of NAT holds b6 . i = middle_sum (f,(S . i)) );
:: deftheorem defines + INTEGR15:def 9 :
for n being Nat
for Z being set
for f, g being PartFunc of Z,(REAL n) holds f + g = f <++> g;
:: deftheorem defines - INTEGR15:def 10 :
for n being Nat
for Z being set
for f, g being PartFunc of Z,(REAL n) holds f - g = f <--> g;
:: deftheorem defines (#) INTEGR15:def 11 :
for n being Nat
for r being real number
for Z being set
for f being PartFunc of Z,(REAL n) holds r (#) f = f [#] r;
begin
:: deftheorem Def12 defines bounded INTEGR15:def 12 :
for n being Element of NAT
for A being closed-interval Subset of REAL
for f being Function of A,(REAL n) holds
( f is bounded iff for i being Element of NAT st i in Seg n holds
(proj (i,n)) * f is bounded );
Lm6:
for n, i being Element of NAT
for f being PartFunc of REAL,(REAL n)
for i being Element of NAT
for A being Subset of REAL holds (proj (i,n)) * (f | A) = ((proj (i,n)) * f) | A
:: deftheorem Def13 defines integrable INTEGR15:def 13 :
for n being Element of NAT
for A being closed-interval Subset of REAL
for f being Function of A,(REAL n) holds
( f is integrable iff for i being Element of NAT st i in Seg n holds
(proj (i,n)) * f is integrable );
:: deftheorem Def14 defines integral INTEGR15:def 14 :
for n being Element of NAT
for A being closed-interval Subset of REAL
for f being Function of A,(REAL n)
for b4 being Element of REAL n holds
( b4 = integral f iff ( dom b4 = Seg n & ( for i being Element of NAT st i in Seg n holds
b4 . i = integral ((proj (i,n)) * f) ) ) );
theorem Th11:
theorem
:: deftheorem Def15 defines bounded INTEGR15:def 15 :
for n being Element of NAT
for f being PartFunc of REAL,(REAL n) holds
( f is bounded iff for i being Element of NAT st i in Seg n holds
(proj (i,n)) * f is bounded );
:: deftheorem Def16 defines is_integrable_on INTEGR15:def 16 :
for n being Element of NAT
for A being closed-interval Subset of REAL
for f being PartFunc of REAL,(REAL n) holds
( f is_integrable_on A iff for i being Element of NAT st i in Seg n holds
(proj (i,n)) * f is_integrable_on A );
definition
let n be
Element of
NAT ;
let A be
closed-interval Subset of
REAL;
let f be
PartFunc of
REAL,
(REAL n);
func integral (
f,
A)
-> Element of
REAL n means :
Def17:
(
dom it = Seg n & ( for
i being
Element of
NAT st
i in Seg n holds
it . i = integral (
((proj (i,n)) * f),
A) ) );
existence
ex b1 being Element of REAL n st
( dom b1 = Seg n & ( for i being Element of NAT st i in Seg n holds
b1 . i = integral (((proj (i,n)) * f),A) ) )
uniqueness
for b1, b2 being Element of REAL n st dom b1 = Seg n & ( for i being Element of NAT st i in Seg n holds
b1 . i = integral (((proj (i,n)) * f),A) ) & dom b2 = Seg n & ( for i being Element of NAT st i in Seg n holds
b2 . i = integral (((proj (i,n)) * f),A) ) holds
b1 = b2
end;
:: deftheorem Def17 defines integral INTEGR15:def 17 :
for n being Element of NAT
for A being closed-interval Subset of REAL
for f being PartFunc of REAL,(REAL n)
for b4 being Element of REAL n holds
( b4 = integral (f,A) iff ( dom b4 = Seg n & ( for i being Element of NAT st i in Seg n holds
b4 . i = integral (((proj (i,n)) * f),A) ) ) );
theorem
theorem
definition
let a,
b be
real number ;
let n be
Element of
NAT ;
let f be
PartFunc of
REAL,
(REAL n);
func integral (
f,
a,
b)
-> Element of
REAL n means :
Def18:
(
dom it = Seg n & ( for
i being
Element of
NAT st
i in Seg n holds
it . i = integral (
((proj (i,n)) * f),
a,
b) ) );
existence
ex b1 being Element of REAL n st
( dom b1 = Seg n & ( for i being Element of NAT st i in Seg n holds
b1 . i = integral (((proj (i,n)) * f),a,b) ) )
uniqueness
for b1, b2 being Element of REAL n st dom b1 = Seg n & ( for i being Element of NAT st i in Seg n holds
b1 . i = integral (((proj (i,n)) * f),a,b) ) & dom b2 = Seg n & ( for i being Element of NAT st i in Seg n holds
b2 . i = integral (((proj (i,n)) * f),a,b) ) holds
b1 = b2
end;
:: deftheorem Def18 defines integral INTEGR15:def 18 :
for a, b being real number
for n being Element of NAT
for f being PartFunc of REAL,(REAL n)
for b5 being Element of REAL n holds
( b5 = integral (f,a,b) iff ( dom b5 = Seg n & ( for i being Element of NAT st i in Seg n holds
b5 . i = integral (((proj (i,n)) * f),a,b) ) ) );
begin
theorem Th15:
theorem Th16:
theorem
theorem
theorem
theorem
theorem
theorem
theorem