:: The Definition of Riemann Definite Integral and some Related Lemmas
:: by Noboru Endou and Artur Korni{\l}owicz
::
:: Received March 13, 1999
:: Copyright (c) 1999-2021 Association of Mizar Users


registration
cluster closed_interval -> compact for Element of bool REAL;
coherence
for b1 being Subset of REAL st b1 is closed_interval holds
b1 is compact
proof end;
end;

theorem :: INTEGRA1:1
canceled;

theorem :: INTEGRA1:2
canceled;

::$CT 2
theorem Th1: :: INTEGRA1:3
for A being non empty closed_interval Subset of REAL holds
( A is bounded_below & A is bounded_above )
proof end;

registration
cluster non empty closed_interval -> real-bounded for Element of bool REAL;
coherence
for b1 being Subset of REAL st not b1 is empty & b1 is closed_interval holds
b1 is real-bounded
by Th1;
end;

theorem Th2: :: INTEGRA1:4
for A being non empty closed_interval Subset of REAL holds A = [.(lower_bound A),(upper_bound A).]
proof end;

theorem Th3: :: INTEGRA1:5
for A being non empty closed_interval Subset of REAL
for a1, a2, b1, b2 being Real st A = [.a1,b1.] & A = [.a2,b2.] holds
( a1 = a2 & b1 = b2 )
proof end;

definition
let A be non empty compact Subset of REAL;
mode Division of A -> non empty increasing FinSequence of REAL means :Def1: :: INTEGRA1:def 2
( rng it c= A & it . (len it) = upper_bound A );
existence
ex b1 being non empty increasing FinSequence of REAL st
( rng b1 c= A & b1 . (len b1) = upper_bound A )
proof end;
end;

:: deftheorem INTEGRA1:def 1 :
canceled;

:: deftheorem Def1 defines Division INTEGRA1:def 2 :
for A being non empty compact Subset of REAL
for b2 being non empty increasing FinSequence of REAL holds
( b2 is Division of A iff ( rng b2 c= A & b2 . (len b2) = upper_bound A ) );

definition
let A be non empty compact Subset of REAL;
func divs A -> set means :Def2: :: INTEGRA1:def 3
for x1 being set holds
( x1 in it iff x1 is Division of A );
existence
ex b1 being set st
for x1 being set holds
( x1 in b1 iff x1 is Division of A )
proof end;
uniqueness
for b1, b2 being set st ( for x1 being set holds
( x1 in b1 iff x1 is Division of A ) ) & ( for x1 being set holds
( x1 in b2 iff x1 is Division of A ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def2 defines divs INTEGRA1:def 3 :
for A being non empty compact Subset of REAL
for b2 being set holds
( b2 = divs A iff for x1 being set holds
( x1 in b2 iff x1 is Division of A ) );

registration
let A be non empty compact Subset of REAL;
cluster divs A -> non empty ;
coherence
not divs A is empty
proof end;
end;

registration
let A be non empty compact Subset of REAL;
cluster -> Relation-like Function-like for Element of divs A;
coherence
for b1 being Element of divs A holds
( b1 is Function-like & b1 is Relation-like )
by Def2;
end;

registration
let A be non empty compact Subset of REAL;
cluster -> FinSequence-like real-valued for Element of divs A;
coherence
for b1 being Element of divs A holds
( b1 is real-valued & b1 is FinSequence-like )
by Def2;
end;

theorem Th4: :: INTEGRA1:6
for i being Element of NAT
for A being non empty closed_interval Subset of REAL
for D being Division of A st i in dom D holds
D . i in A
proof end;

theorem Th5: :: INTEGRA1:7
for i being Element of NAT
for A being non empty closed_interval Subset of REAL
for D being Division of A st i in dom D & i <> 1 holds
( i - 1 in dom D & D . (i - 1) in A & i - 1 in NAT )
proof end;

definition
let A be non empty closed_interval Subset of REAL;
let D be Division of A;
let i be Nat;
assume A1: i in dom D ;
func divset (D,i) -> non empty closed_interval Subset of REAL means :Def3: :: INTEGRA1:def 4
( lower_bound it = lower_bound A & upper_bound it = D . i ) if i = 1
otherwise ( lower_bound it = D . (i - 1) & upper_bound it = D . i );
existence
( ( i = 1 implies ex b1 being non empty closed_interval Subset of REAL st
( lower_bound b1 = lower_bound A & upper_bound b1 = D . i ) ) & ( not i = 1 implies ex b1 being non empty closed_interval Subset of REAL st
( lower_bound b1 = D . (i - 1) & upper_bound b1 = D . i ) ) )
proof end;
uniqueness
for b1, b2 being non empty closed_interval Subset of REAL holds
( ( i = 1 & lower_bound b1 = lower_bound A & upper_bound b1 = D . i & lower_bound b2 = lower_bound A & upper_bound b2 = D . i implies b1 = b2 ) & ( not i = 1 & lower_bound b1 = D . (i - 1) & upper_bound b1 = D . i & lower_bound b2 = D . (i - 1) & upper_bound b2 = D . i implies b1 = b2 ) )
proof end;
correctness
consistency
for b1 being non empty closed_interval Subset of REAL holds verum
;
;
end;

:: deftheorem Def3 defines divset INTEGRA1:def 4 :
for A being non empty closed_interval Subset of REAL
for D being Division of A
for i being Nat st i in dom D holds
for b4 being non empty closed_interval Subset of REAL holds
( ( i = 1 implies ( b4 = divset (D,i) iff ( lower_bound b4 = lower_bound A & upper_bound b4 = D . i ) ) ) & ( not i = 1 implies ( b4 = divset (D,i) iff ( lower_bound b4 = D . (i - 1) & upper_bound b4 = D . i ) ) ) );

theorem Th6: :: INTEGRA1:8
for i being Element of NAT
for A being non empty closed_interval Subset of REAL
for D being Division of A st i in dom D holds
divset (D,i) c= A
proof end;

definition
let A be Subset of REAL;
func vol A -> Real equals :: INTEGRA1:def 5
(upper_bound A) - (lower_bound A);
correctness
coherence
(upper_bound A) - (lower_bound A) is Real
;
;
end;

:: deftheorem defines vol INTEGRA1:def 5 :
for A being Subset of REAL holds vol A = (upper_bound A) - (lower_bound A);

theorem :: INTEGRA1:9
for A being non empty real-bounded Subset of REAL holds 0 <= vol A by SEQ_4:11, XREAL_1:48;

definition
let A be non empty closed_interval Subset of REAL;
let f be PartFunc of A,REAL;
let D be Division of A;
func upper_volume (f,D) -> FinSequence of REAL means :Def5: :: INTEGRA1:def 6
( len it = len D & ( for i being Nat st i in dom D holds
it . i = (upper_bound (rng (f | (divset (D,i))))) * (vol (divset (D,i))) ) );
existence
ex b1 being FinSequence of REAL st
( len b1 = len D & ( for i being Nat st i in dom D holds
b1 . i = (upper_bound (rng (f | (divset (D,i))))) * (vol (divset (D,i))) ) )
proof end;
uniqueness
for b1, b2 being FinSequence of REAL st len b1 = len D & ( for i being Nat st i in dom D holds
b1 . i = (upper_bound (rng (f | (divset (D,i))))) * (vol (divset (D,i))) ) & len b2 = len D & ( for i being Nat st i in dom D holds
b2 . i = (upper_bound (rng (f | (divset (D,i))))) * (vol (divset (D,i))) ) holds
b1 = b2
proof end;
func lower_volume (f,D) -> FinSequence of REAL means :Def6: :: INTEGRA1:def 7
( len it = len D & ( for i being Nat st i in dom D holds
it . i = (lower_bound (rng (f | (divset (D,i))))) * (vol (divset (D,i))) ) );
existence
ex b1 being FinSequence of REAL st
( len b1 = len D & ( for i being Nat st i in dom D holds
b1 . i = (lower_bound (rng (f | (divset (D,i))))) * (vol (divset (D,i))) ) )
proof end;
uniqueness
for b1, b2 being FinSequence of REAL st len b1 = len D & ( for i being Nat st i in dom D holds
b1 . i = (lower_bound (rng (f | (divset (D,i))))) * (vol (divset (D,i))) ) & len b2 = len D & ( for i being Nat st i in dom D holds
b2 . i = (lower_bound (rng (f | (divset (D,i))))) * (vol (divset (D,i))) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def5 defines upper_volume INTEGRA1:def 6 :
for A being non empty closed_interval Subset of REAL
for f being PartFunc of A,REAL
for D being Division of A
for b4 being FinSequence of REAL holds
( b4 = upper_volume (f,D) iff ( len b4 = len D & ( for i being Nat st i in dom D holds
b4 . i = (upper_bound (rng (f | (divset (D,i))))) * (vol (divset (D,i))) ) ) );

:: deftheorem Def6 defines lower_volume INTEGRA1:def 7 :
for A being non empty closed_interval Subset of REAL
for f being PartFunc of A,REAL
for D being Division of A
for b4 being FinSequence of REAL holds
( b4 = lower_volume (f,D) iff ( len b4 = len D & ( for i being Nat st i in dom D holds
b4 . i = (lower_bound (rng (f | (divset (D,i))))) * (vol (divset (D,i))) ) ) );

definition
let A be non empty closed_interval Subset of REAL;
let f be PartFunc of A,REAL;
let D be Division of A;
func upper_sum (f,D) -> Real equals :: INTEGRA1:def 8
Sum (upper_volume (f,D));
correctness
coherence
Sum (upper_volume (f,D)) is Real
;
;
func lower_sum (f,D) -> Real equals :: INTEGRA1:def 9
Sum (lower_volume (f,D));
correctness
coherence
Sum (lower_volume (f,D)) is Real
;
;
end;

:: deftheorem defines upper_sum INTEGRA1:def 8 :
for A being non empty closed_interval Subset of REAL
for f being PartFunc of A,REAL
for D being Division of A holds upper_sum (f,D) = Sum (upper_volume (f,D));

:: deftheorem defines lower_sum INTEGRA1:def 9 :
for A being non empty closed_interval Subset of REAL
for f being PartFunc of A,REAL
for D being Division of A holds lower_sum (f,D) = Sum (lower_volume (f,D));

definition
let A be non empty closed_interval Subset of REAL;
let f be PartFunc of A,REAL;
set S = divs A;
func upper_sum_set f -> Function of (divs A),REAL means :Def9: :: INTEGRA1:def 10
for D being Division of A holds it . D = upper_sum (f,D);
existence
ex b1 being Function of (divs A),REAL st
for D being Division of A holds b1 . D = upper_sum (f,D)
proof end;
uniqueness
for b1, b2 being Function of (divs A),REAL st ( for D being Division of A holds b1 . D = upper_sum (f,D) ) & ( for D being Division of A holds b2 . D = upper_sum (f,D) ) holds
b1 = b2
proof end;
func lower_sum_set f -> Function of (divs A),REAL means :Def10: :: INTEGRA1:def 11
for D being Division of A holds it . D = lower_sum (f,D);
existence
ex b1 being Function of (divs A),REAL st
for D being Division of A holds b1 . D = lower_sum (f,D)
proof end;
uniqueness
for b1, b2 being Function of (divs A),REAL st ( for D being Division of A holds b1 . D = lower_sum (f,D) ) & ( for D being Division of A holds b2 . D = lower_sum (f,D) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def9 defines upper_sum_set INTEGRA1:def 10 :
for A being non empty closed_interval Subset of REAL
for f being PartFunc of A,REAL
for b3 being Function of (divs A),REAL holds
( b3 = upper_sum_set f iff for D being Division of A holds b3 . D = upper_sum (f,D) );

:: deftheorem Def10 defines lower_sum_set INTEGRA1:def 11 :
for A being non empty closed_interval Subset of REAL
for f being PartFunc of A,REAL
for b3 being Function of (divs A),REAL holds
( b3 = lower_sum_set f iff for D being Division of A holds b3 . D = lower_sum (f,D) );

definition
let A be non empty closed_interval Subset of REAL;
let f be PartFunc of A,REAL;
attr f is upper_integrable means :: INTEGRA1:def 12
rng (upper_sum_set f) is bounded_below ;
attr f is lower_integrable means :: INTEGRA1:def 13
rng (lower_sum_set f) is bounded_above ;
end;

:: deftheorem defines upper_integrable INTEGRA1:def 12 :
for A being non empty closed_interval Subset of REAL
for f being PartFunc of A,REAL holds
( f is upper_integrable iff rng (upper_sum_set f) is bounded_below );

:: deftheorem defines lower_integrable INTEGRA1:def 13 :
for A being non empty closed_interval Subset of REAL
for f being PartFunc of A,REAL holds
( f is lower_integrable iff rng (lower_sum_set f) is bounded_above );

definition
let A be non empty closed_interval Subset of REAL;
let f be PartFunc of A,REAL;
func upper_integral f -> Real equals :: INTEGRA1:def 14
lower_bound (rng (upper_sum_set f));
correctness
coherence
lower_bound (rng (upper_sum_set f)) is Real
;
;
end;

:: deftheorem defines upper_integral INTEGRA1:def 14 :
for A being non empty closed_interval Subset of REAL
for f being PartFunc of A,REAL holds upper_integral f = lower_bound (rng (upper_sum_set f));

definition
let A be non empty closed_interval Subset of REAL;
let f be PartFunc of A,REAL;
func lower_integral f -> Real equals :: INTEGRA1:def 15
upper_bound (rng (lower_sum_set f));
coherence
upper_bound (rng (lower_sum_set f)) is Real
;
end;

:: deftheorem defines lower_integral INTEGRA1:def 15 :
for A being non empty closed_interval Subset of REAL
for f being PartFunc of A,REAL holds lower_integral f = upper_bound (rng (lower_sum_set f));

definition
let A be non empty closed_interval Subset of REAL;
let f be PartFunc of A,REAL;
attr f is integrable means :: INTEGRA1:def 16
( f is upper_integrable & f is lower_integrable & upper_integral f = lower_integral f );
end;

:: deftheorem defines integrable INTEGRA1:def 16 :
for A being non empty closed_interval Subset of REAL
for f being PartFunc of A,REAL holds
( f is integrable iff ( f is upper_integrable & f is lower_integrable & upper_integral f = lower_integral f ) );

definition
let A be non empty closed_interval Subset of REAL;
let f be PartFunc of A,REAL;
func integral f -> Real equals :: INTEGRA1:def 17
upper_integral f;
coherence
upper_integral f is Real
;
end;

:: deftheorem defines integral INTEGRA1:def 17 :
for A being non empty closed_interval Subset of REAL
for f being PartFunc of A,REAL holds integral f = upper_integral f;

theorem Th8: :: INTEGRA1:10
for X being non empty set
for f, g being PartFunc of X,REAL holds rng (f + g) c= (rng f) ++ (rng g)
proof end;

theorem Th9: :: INTEGRA1:11
for f being real-valued Function st f is bounded_below holds
rng f is bounded_below
proof end;

theorem :: INTEGRA1:12
for f being real-valued Function st rng f is bounded_below holds
f is bounded_below
proof end;

theorem Th11: :: INTEGRA1:13
for f being real-valued Function st f is bounded_above holds
rng f is bounded_above
proof end;

theorem :: INTEGRA1:14
for f being real-valued Function st rng f is bounded_above holds
f is bounded_above
proof end;

theorem :: INTEGRA1:15
for f being real-valued Function st f is bounded holds
rng f is real-bounded by Th11, Th9;

theorem Th14: :: INTEGRA1:16
for A being non empty set holds (chi (A,A)) | A is V8()
proof end;

theorem Th15: :: INTEGRA1:17
for X being non empty set
for A being non empty Subset of X holds rng (chi (A,A)) = {1}
proof end;

theorem Th16: :: INTEGRA1:18
for X being non empty set
for A being non empty Subset of X
for B being set st B meets dom (chi (A,A)) holds
rng ((chi (A,A)) | B) = {1}
proof end;

theorem Th17: :: INTEGRA1:19
for i being Element of NAT
for A being non empty closed_interval Subset of REAL
for D being Division of A st i in dom D holds
vol (divset (D,i)) = (lower_volume ((chi (A,A)),D)) . i
proof end;

theorem Th18: :: INTEGRA1:20
for i being Element of NAT
for A being non empty closed_interval Subset of REAL
for D being Division of A st i in dom D holds
vol (divset (D,i)) = (upper_volume ((chi (A,A)),D)) . i
proof end;

theorem :: INTEGRA1:21
for F, G, H being FinSequence of REAL st len F = len G & len F = len H & ( for k being Element of NAT st k in dom F holds
H . k = (F /. k) + (G /. k) ) holds
Sum H = (Sum F) + (Sum G)
proof end;

theorem Th20: :: INTEGRA1:22
for F, G, H being FinSequence of REAL st len F = len G & len F = len H & ( for k being Element of NAT st k in dom F holds
H . k = (F /. k) - (G /. k) ) holds
Sum H = (Sum F) - (Sum G)
proof end;

theorem Th21: :: INTEGRA1:23
for A being non empty closed_interval Subset of REAL
for D being Division of A holds Sum (lower_volume ((chi (A,A)),D)) = vol A
proof end;

theorem Th22: :: INTEGRA1:24
for A being non empty closed_interval Subset of REAL
for D being Division of A holds Sum (upper_volume ((chi (A,A)),D)) = vol A
proof end;

registration
let A be non empty closed_interval Subset of REAL;
let f be PartFunc of A,REAL;
let D be Division of A;
cluster upper_volume (f,D) -> non empty ;
coherence
not upper_volume (f,D) is empty
proof end;
cluster lower_volume (f,D) -> non empty ;
coherence
not lower_volume (f,D) is empty
proof end;
end;

theorem Th23: :: INTEGRA1:25
for A being non empty closed_interval Subset of REAL
for D being Division of A
for f being Function of A,REAL st f | A is bounded_below holds
(lower_bound (rng f)) * (vol A) <= lower_sum (f,D)
proof end;

theorem :: INTEGRA1:26
for i being Element of NAT
for A being non empty closed_interval Subset of REAL
for D being Division of A
for f being Function of A,REAL st f | A is bounded_above & i in dom D holds
(upper_bound (rng f)) * (vol (divset (D,i))) >= (upper_bound (rng (f | (divset (D,i))))) * (vol (divset (D,i)))
proof end;

theorem Th25: :: INTEGRA1:27
for A being non empty closed_interval Subset of REAL
for D being Division of A
for f being Function of A,REAL st f | A is bounded_above holds
upper_sum (f,D) <= (upper_bound (rng f)) * (vol A)
proof end;

theorem Th26: :: INTEGRA1:28
for A being non empty closed_interval Subset of REAL
for D being Division of A
for f being Function of A,REAL st f | A is bounded holds
lower_sum (f,D) <= upper_sum (f,D)
proof end;

definition
let D1, D2 be FinSequence;
pred D1 <= D2 means :: INTEGRA1:def 18
( len D1 <= len D2 & rng D1 c= rng D2 );
reflexivity
for D1 being FinSequence holds
( len D1 <= len D1 & rng D1 c= rng D1 )
;
end;

:: deftheorem defines <= INTEGRA1:def 18 :
for D1, D2 being FinSequence holds
( D1 <= D2 iff ( len D1 <= len D2 & rng D1 c= rng D2 ) );

notation
let D1, D2 be FinSequence;
synonym D2 >= D1 for D1 <= D2;
end;

theorem :: INTEGRA1:29
for A being non empty closed_interval Subset of REAL
for D1, D2 being Division of A st len D1 = 1 holds
D1 <= D2
proof end;

theorem Th28: :: INTEGRA1:30
for A being non empty closed_interval Subset of REAL
for D1, D2 being Division of A
for f being Function of A,REAL st f | A is bounded_above & len D1 = 1 holds
upper_sum (f,D1) >= upper_sum (f,D2)
proof end;

theorem Th29: :: INTEGRA1:31
for A being non empty closed_interval Subset of REAL
for D1, D2 being Division of A
for f being Function of A,REAL st f | A is bounded_below & len D1 = 1 holds
lower_sum (f,D1) <= lower_sum (f,D2)
proof end;

theorem :: INTEGRA1:32
for i being Element of NAT
for A being non empty closed_interval Subset of REAL
for D being Division of A st i in dom D holds
ex A1, A2 being non empty closed_interval Subset of REAL st
( A1 = [.(lower_bound A),(D . i).] & A2 = [.(D . i),(upper_bound A).] & A = A1 \/ A2 )
proof end;

theorem Th31: :: INTEGRA1:33
for i being Element of NAT
for A being non empty closed_interval Subset of REAL
for D1, D2 being Division of A st i in dom D1 & D1 <= D2 holds
ex j being Element of NAT st
( j in dom D2 & D1 . i = D2 . j )
proof end;

definition
let A be non empty closed_interval Subset of REAL;
let D1, D2 be Division of A;
let i be Nat;
assume A1: D1 <= D2 ;
func indx (D2,D1,i) -> Element of NAT means :Def18: :: INTEGRA1:def 19
( it in dom D2 & D1 . i = D2 . it ) if i in dom D1
otherwise it = 0 ;
existence
( ( i in dom D1 implies ex b1 being Element of NAT st
( b1 in dom D2 & D1 . i = D2 . b1 ) ) & ( not i in dom D1 implies ex b1 being Element of NAT st b1 = 0 ) )
by A1, Th31;
uniqueness
for b1, b2 being Element of NAT holds
( ( i in dom D1 & b1 in dom D2 & D1 . i = D2 . b1 & b2 in dom D2 & D1 . i = D2 . b2 implies b1 = b2 ) & ( not i in dom D1 & b1 = 0 & b2 = 0 implies b1 = b2 ) )
proof end;
correctness
consistency
for b1 being Element of NAT holds verum
;
;
end;

:: deftheorem Def18 defines indx INTEGRA1:def 19 :
for A being non empty closed_interval Subset of REAL
for D1, D2 being Division of A
for i being Nat st D1 <= D2 holds
for b5 being Element of NAT holds
( ( i in dom D1 implies ( b5 = indx (D2,D1,i) iff ( b5 in dom D2 & D1 . i = D2 . b5 ) ) ) & ( not i in dom D1 implies ( b5 = indx (D2,D1,i) iff b5 = 0 ) ) );

theorem Th32: :: INTEGRA1:34
for p being increasing FinSequence of REAL
for n being Element of NAT st n <= len p holds
p /^ n is increasing FinSequence of REAL
proof end;

theorem Th33: :: INTEGRA1:35
for p being increasing FinSequence of REAL
for i, j being Element of NAT st j in dom p & i <= j holds
mid (p,i,j) is increasing FinSequence of REAL
proof end;

Lm1: for i, j being Element of NAT
for f being FinSequence st i in dom f & j in dom f & i <= j holds
len (mid (f,i,j)) = (j - i) + 1

proof end;

theorem Th34: :: INTEGRA1:36
for i, j being Element of NAT
for A being non empty closed_interval Subset of REAL
for D being Division of A st i in dom D & j in dom D & i <= j holds
ex B being non empty closed_interval Subset of REAL st
( lower_bound B = (mid (D,i,j)) . 1 & upper_bound B = (mid (D,i,j)) . (len (mid (D,i,j))) & mid (D,i,j) is Division of B )
proof end;

theorem Th35: :: INTEGRA1:37
for i, j being Element of NAT
for A, B being non empty closed_interval Subset of REAL
for D being Division of A st i in dom D & j in dom D & i <= j & D . i >= lower_bound B & D . j = upper_bound B holds
mid (D,i,j) is Division of B
proof end;

definition
let p be FinSequence of REAL ;
func PartSums p -> FinSequence of REAL means :Def19: :: INTEGRA1:def 20
( len it = len p & ( for i being Nat st i in dom p holds
it . i = Sum (p | i) ) );
existence
ex b1 being FinSequence of REAL st
( len b1 = len p & ( for i being Nat st i in dom p holds
b1 . i = Sum (p | i) ) )
proof end;
uniqueness
for b1, b2 being FinSequence of REAL st len b1 = len p & ( for i being Nat st i in dom p holds
b1 . i = Sum (p | i) ) & len b2 = len p & ( for i being Nat st i in dom p holds
b2 . i = Sum (p | i) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def19 defines PartSums INTEGRA1:def 20 :
for p, b2 being FinSequence of REAL holds
( b2 = PartSums p iff ( len b2 = len p & ( for i being Nat st i in dom p holds
b2 . i = Sum (p | i) ) ) );

theorem Th36: :: INTEGRA1:38
for A being non empty closed_interval Subset of REAL
for D1, D2 being Division of A
for f being Function of A,REAL st D1 <= D2 & f | A is bounded_above holds
for i being non zero Element of NAT st i in dom D1 holds
Sum ((upper_volume (f,D1)) | i) >= Sum ((upper_volume (f,D2)) | (indx (D2,D1,i)))
proof end;

theorem Th37: :: INTEGRA1:39
for A being non empty closed_interval Subset of REAL
for D1, D2 being Division of A
for f being Function of A,REAL st D1 <= D2 & f | A is bounded_below holds
for i being non zero Element of NAT st i in dom D1 holds
Sum ((lower_volume (f,D1)) | i) <= Sum ((lower_volume (f,D2)) | (indx (D2,D1,i)))
proof end;

theorem Th38: :: INTEGRA1:40
for i being Element of NAT
for A being non empty closed_interval Subset of REAL
for D1, D2 being Division of A
for f being Function of A,REAL st D1 <= D2 & i in dom D1 & f | A is bounded_above holds
(PartSums (upper_volume (f,D1))) . i >= (PartSums (upper_volume (f,D2))) . (indx (D2,D1,i))
proof end;

theorem Th39: :: INTEGRA1:41
for i being Element of NAT
for A being non empty closed_interval Subset of REAL
for D1, D2 being Division of A
for f being Function of A,REAL st D1 <= D2 & i in dom D1 & f | A is bounded_below holds
(PartSums (lower_volume (f,D1))) . i <= (PartSums (lower_volume (f,D2))) . (indx (D2,D1,i))
proof end;

theorem Th40: :: INTEGRA1:42
for A being non empty closed_interval Subset of REAL
for D being Division of A
for f being Function of A,REAL holds (PartSums (upper_volume (f,D))) . (len D) = upper_sum (f,D)
proof end;

theorem Th41: :: INTEGRA1:43
for A being non empty closed_interval Subset of REAL
for D being Division of A
for f being Function of A,REAL holds (PartSums (lower_volume (f,D))) . (len D) = lower_sum (f,D)
proof end;

theorem Th42: :: INTEGRA1:44
for A being non empty closed_interval Subset of REAL
for D1, D2 being Division of A st D1 <= D2 holds
indx (D2,D1,(len D1)) = len D2
proof end;

theorem Th43: :: INTEGRA1:45
for A being non empty closed_interval Subset of REAL
for D1, D2 being Division of A
for f being Function of A,REAL st D1 <= D2 & f | A is bounded_above holds
upper_sum (f,D2) <= upper_sum (f,D1)
proof end;

theorem Th44: :: INTEGRA1:46
for A being non empty closed_interval Subset of REAL
for D1, D2 being Division of A
for f being Function of A,REAL st D1 <= D2 & f | A is bounded_below holds
lower_sum (f,D2) >= lower_sum (f,D1)
proof end;

theorem Th45: :: INTEGRA1:47
for A being non empty closed_interval Subset of REAL
for D1, D2 being Division of A ex D being Division of A st
( D1 <= D & D2 <= D )
proof end;

theorem Th46: :: INTEGRA1:48
for A being non empty closed_interval Subset of REAL
for D1, D2 being Division of A
for f being Function of A,REAL st f | A is bounded holds
lower_sum (f,D1) <= upper_sum (f,D2)
proof end;

theorem Th47: :: INTEGRA1:49
for A being non empty closed_interval Subset of REAL
for f being Function of A,REAL st f | A is bounded holds
upper_integral f >= lower_integral f
proof end;

theorem Th48: :: INTEGRA1:50
for X, Y being Subset of REAL holds (-- X) ++ (-- Y) = -- (X ++ Y)
proof end;

theorem Th49: :: INTEGRA1:51
for X, Y being Subset of REAL st X is bounded_above & Y is bounded_above holds
X ++ Y is bounded_above
proof end;

theorem Th50: :: INTEGRA1:52
for X, Y being non empty Subset of REAL st X is bounded_above & Y is bounded_above holds
upper_bound (X ++ Y) = (upper_bound X) + (upper_bound Y)
proof end;

theorem Th51: :: INTEGRA1:53
for i being Element of NAT
for A being non empty closed_interval Subset of REAL
for D being Division of A
for f, g being Function of A,REAL st i in dom D & f | A is bounded_above & g | A is bounded_above holds
(upper_volume ((f + g),D)) . i <= ((upper_volume (f,D)) . i) + ((upper_volume (g,D)) . i)
proof end;

theorem Th52: :: INTEGRA1:54
for i being Element of NAT
for A being non empty closed_interval Subset of REAL
for D being Division of A
for f, g being Function of A,REAL st i in dom D & f | A is bounded_below & g | A is bounded_below holds
((lower_volume (f,D)) . i) + ((lower_volume (g,D)) . i) <= (lower_volume ((f + g),D)) . i
proof end;

theorem Th53: :: INTEGRA1:55
for A being non empty closed_interval Subset of REAL
for D being Division of A
for f, g being Function of A,REAL st f | A is bounded_above & g | A is bounded_above holds
upper_sum ((f + g),D) <= (upper_sum (f,D)) + (upper_sum (g,D))
proof end;

theorem Th54: :: INTEGRA1:56
for A being non empty closed_interval Subset of REAL
for D being Division of A
for f, g being Function of A,REAL st f | A is bounded_below & g | A is bounded_below holds
(lower_sum (f,D)) + (lower_sum (g,D)) <= lower_sum ((f + g),D)
proof end;

theorem :: INTEGRA1:57
for A being non empty closed_interval Subset of REAL
for f, g being Function of A,REAL st f | A is bounded & g | A is bounded & f is integrable & g is integrable holds
( f + g is integrable & integral (f + g) = (integral f) + (integral g) )
proof end;

theorem :: INTEGRA1:58
for i, j being Element of NAT
for f being FinSequence st i in dom f & j in dom f & i <= j holds
len (mid (f,i,j)) = (j - i) + 1 by Lm1;