:: Series of Positive Real Numbers. Measure Theory
:: by J\'ozef Bia{\l}as
::
:: Received September 27, 1990
:: Copyright (c) 1990 Association of Mizar Users


notation
synonym 0. for 0 ;
end;

definition
canceled;
:: original: 0.
redefine func 0. -> R_eal;
coherence
0. is R_eal
by XXREAL_0:def 1;
end;

:: deftheorem SUPINF_2:def 1 :
canceled;

definition
let x, y be R_eal;
:: original: +
redefine func x + y -> R_eal means :Def2: :: SUPINF_2:def 2
ex a, b being Real st
( x = a & y = b & it = a + b ) if ( x in REAL & y in REAL )
it = +infty if ( ( x = +infty & y <> -infty ) or ( y = +infty & x <> -infty ) )
it = -infty if ( ( x = -infty & y <> +infty ) or ( y = -infty & x <> +infty ) )
otherwise it = 0. ;
coherence
x + y is R_eal
by XXREAL_0:def 1;
compatibility
for b1 being R_eal holds
( ( x in REAL & y in REAL implies ( b1 = x + y iff ex a, b being Real st
( x = a & y = b & b1 = a + b ) ) ) & ( ( ( x = +infty & y <> -infty ) or ( y = +infty & x <> -infty ) ) implies ( b1 = x + y iff b1 = +infty ) ) & ( ( ( x = -infty & y <> +infty ) or ( y = -infty & x <> +infty ) ) implies ( b1 = x + y iff b1 = -infty ) ) & ( ( x in REAL & y in REAL ) or ( x = +infty & y <> -infty ) or ( y = +infty & x <> -infty ) or ( x = -infty & y <> +infty ) or ( y = -infty & x <> +infty ) or ( b1 = x + y iff b1 = 0. ) ) )
proof end;
correctness
consistency
for b1 being R_eal holds
( ( x in REAL & y in REAL & ( ( x = +infty & y <> -infty ) or ( y = +infty & x <> -infty ) ) implies ( ex a, b being Real st
( x = a & y = b & b1 = a + b ) iff b1 = +infty ) ) & ( x in REAL & y in REAL & ( ( x = -infty & y <> +infty ) or ( y = -infty & x <> +infty ) ) implies ( ex a, b being Real st
( x = a & y = b & b1 = a + b ) iff b1 = -infty ) ) & ( ( ( x = +infty & y <> -infty ) or ( y = +infty & x <> -infty ) ) & ( ( x = -infty & y <> +infty ) or ( y = -infty & x <> +infty ) ) implies ( b1 = +infty iff b1 = -infty ) ) )
;
;
end;

:: deftheorem Def2 defines + SUPINF_2:def 2 :
for x, y, b3 being R_eal holds
( ( x in REAL & y in REAL implies ( b3 = x + y iff ex a, b being Real st
( x = a & y = b & b3 = a + b ) ) ) & ( ( ( x = +infty & y <> -infty ) or ( y = +infty & x <> -infty ) ) implies ( b3 = x + y iff b3 = +infty ) ) & ( ( ( x = -infty & y <> +infty ) or ( y = -infty & x <> +infty ) ) implies ( b3 = x + y iff b3 = -infty ) ) & ( ( x in REAL & y in REAL ) or ( x = +infty & y <> -infty ) or ( y = +infty & x <> -infty ) or ( x = -infty & y <> +infty ) or ( y = -infty & x <> +infty ) or ( b3 = x + y iff b3 = 0. ) ) );

theorem Th1: :: SUPINF_2:1
for x, y being R_eal
for a, b being Real st x = a & y = b holds
x + y = a + b
proof end;

definition
let x be R_eal;
:: original: -
redefine func - x -> R_eal means :Def3: :: SUPINF_2:def 3
ex a being Real st
( x = a & it = - a ) if x in REAL
it = -infty if x = +infty
otherwise it = +infty ;
coherence
- x is R_eal
by XXREAL_0:def 1;
compatibility
for b1 being R_eal holds
( ( x in REAL implies ( b1 = - x iff ex a being Real st
( x = a & b1 = - a ) ) ) & ( x = +infty implies ( b1 = - x iff b1 = -infty ) ) & ( not x in REAL & not x = +infty implies ( b1 = - x iff b1 = +infty ) ) )
proof end;
correctness
consistency
for b1 being R_eal st x in REAL & x = +infty holds
( ex a being Real st
( x = a & b1 = - a ) iff b1 = -infty )
;
;
end;

:: deftheorem Def3 defines - SUPINF_2:def 3 :
for x, b2 being R_eal holds
( ( x in REAL implies ( b2 = - x iff ex a being Real st
( x = a & b2 = - a ) ) ) & ( x = +infty implies ( b2 = - x iff b2 = -infty ) ) & ( not x in REAL & not x = +infty implies ( b2 = - x iff b2 = +infty ) ) );

definition
let x, y be R_eal;
:: original: -
redefine func x - y -> R_eal equals :: SUPINF_2:def 4
x + (- y);
coherence
x - y is R_eal
by XXREAL_0:def 1;
correctness
compatibility
for b1 being R_eal holds
( b1 = x - y iff b1 = x + (- y) )
;
;
end;

:: deftheorem defines - SUPINF_2:def 4 :
for x, y being R_eal holds x - y = x + (- y);

theorem :: SUPINF_2:2
canceled;

theorem Th3: :: SUPINF_2:3
for x being R_eal
for a being Real st x = a holds
- x = - a
proof end;

Lm1: ( not +infty in REAL & not -infty in REAL )
;

theorem Th4: :: SUPINF_2:4
- -infty = +infty by Def3, Lm1;

theorem Th5: :: SUPINF_2:5
for x, y being R_eal
for a, b being Real st x = a & y = b holds
x - y = a - b
proof end;

theorem Th6: :: SUPINF_2:6
for x being R_eal st x <> +infty holds
( +infty - x = +infty & x - +infty = -infty )
proof end;

theorem Th7: :: SUPINF_2:7
for x being R_eal st x <> -infty holds
( -infty - x = -infty & x - -infty = +infty )
proof end;

theorem Th8: :: SUPINF_2:8
for x, s being R_eal holds
( not x + s = +infty or x = +infty or s = +infty )
proof end;

theorem Th9: :: SUPINF_2:9
for x, s being R_eal holds
( not x + s = -infty or x = -infty or s = -infty )
proof end;

theorem :: SUPINF_2:10
for x, s being R_eal holds
( not x - s = +infty or x = +infty or s = -infty )
proof end;

theorem Th11: :: SUPINF_2:11
for x, s being R_eal holds
( not x - s = -infty or x = -infty or s = +infty )
proof end;

theorem Th12: :: SUPINF_2:12
for x, s being R_eal holds
( ( x = +infty & s = -infty ) or ( x = -infty & s = +infty ) or not x + s in REAL or ( x in REAL & s in REAL ) )
proof end;

theorem Th13: :: SUPINF_2:13
for x, s being R_eal holds
( ( x = +infty & s = +infty ) or ( x = -infty & s = -infty ) or not x - s in REAL or ( x in REAL & s in REAL ) )
proof end;

theorem Th14: :: SUPINF_2:14
for x, y, s, t being R_eal st x <= y & s <= t holds
x + s <= y + t
proof end;

Lm2: for x being R_eal holds
( - x in REAL iff x in REAL )
proof end;

Lm3: for x, y being R_eal st x <= y holds
- y <= - x
proof end;

theorem :: SUPINF_2:15
for x, y, s, t being R_eal st x <= y & s <= t holds
x - t <= y - s
proof end;

theorem Th16: :: SUPINF_2:16
for x, y being R_eal holds
( x <= y iff - y <= - x )
proof end;

theorem :: SUPINF_2:17
for x, y being R_eal holds
( x < y iff - y < - x ) by Th16;

theorem Th18: :: SUPINF_2:18
for x being R_eal holds
( x + 0. = x & 0. + x = x )
proof end;

theorem :: SUPINF_2:19
( -infty < 0. & 0. < +infty )
proof end;

theorem Th20: :: SUPINF_2:20
for x, y, z being R_eal st 0. <= z & 0. <= x & y = x + z holds
x <= y
proof end;

definition
let X, Y be non empty Subset of ExtREAL ;
func X + Y -> Subset of ExtREAL means :Def5: :: SUPINF_2:def 5
for z being R_eal holds
( z in it iff ex x, y being R_eal st
( x in X & y in Y & z = x + y ) );
existence
ex b1 being Subset of ExtREAL st
for z being R_eal holds
( z in b1 iff ex x, y being R_eal st
( x in X & y in Y & z = x + y ) )
proof end;
uniqueness
for b1, b2 being Subset of ExtREAL st ( for z being R_eal holds
( z in b1 iff ex x, y being R_eal st
( x in X & y in Y & z = x + y ) ) ) & ( for z being R_eal holds
( z in b2 iff ex x, y being R_eal st
( x in X & y in Y & z = x + y ) ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def5 defines + SUPINF_2:def 5 :
for X, Y being non empty Subset of ExtREAL
for b3 being Subset of ExtREAL holds
( b3 = X + Y iff for z being R_eal holds
( z in b3 iff ex x, y being R_eal st
( x in X & y in Y & z = x + y ) ) );

registration
let X, Y be non empty Subset of ExtREAL ;
cluster X + Y -> non empty ;
coherence
not X + Y is empty
proof end;
end;

definition
let X be non empty Subset of ExtREAL ;
func - X -> Subset of ExtREAL means :Def6: :: SUPINF_2:def 6
for z being R_eal holds
( z in it iff ex x being R_eal st
( x in X & z = - x ) );
existence
ex b1 being Subset of ExtREAL st
for z being R_eal holds
( z in b1 iff ex x being R_eal st
( x in X & z = - x ) )
proof end;
uniqueness
for b1, b2 being Subset of ExtREAL st ( for z being R_eal holds
( z in b1 iff ex x being R_eal st
( x in X & z = - x ) ) ) & ( for z being R_eal holds
( z in b2 iff ex x being R_eal st
( x in X & z = - x ) ) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def6 defines - SUPINF_2:def 6 :
for X being non empty Subset of ExtREAL
for b2 being Subset of ExtREAL holds
( b2 = - X iff for z being R_eal holds
( z in b2 iff ex x being R_eal st
( x in X & z = - x ) ) );

registration
let X be non empty Subset of ExtREAL ;
cluster - X -> non empty ;
coherence
not - X is empty
proof end;
end;

theorem :: SUPINF_2:21
canceled;

theorem Th22: :: SUPINF_2:22
for X being non empty Subset of ExtREAL holds - (- X) = X
proof end;

theorem Th23: :: SUPINF_2:23
for X being non empty Subset of ExtREAL
for z being R_eal holds
( z in X iff - z in - X )
proof end;

theorem :: SUPINF_2:24
for X, Y being non empty Subset of ExtREAL holds
( X c= Y iff - X c= - Y )
proof end;

theorem :: SUPINF_2:25
for z being R_eal holds
( z in REAL iff - z in REAL )
proof end;

definition
let X be ext-real-membered set ;
:: original: inf
redefine func inf X -> R_eal;
coherence
inf X is R_eal
by XXREAL_0:def 1;
:: original: sup
redefine func sup X -> R_eal;
coherence
sup X is R_eal
by XXREAL_0:def 1;
end;

theorem Th26: :: SUPINF_2:26
for X, Y being non empty Subset of ExtREAL holds sup (X + Y) <= (sup X) + (sup Y)
proof end;

theorem Th27: :: SUPINF_2:27
for X, Y being non empty Subset of ExtREAL holds (inf X) + (inf Y) <= inf (X + Y)
proof end;

theorem :: SUPINF_2:28
for X, Y being non empty Subset of ExtREAL st X is bounded_above & Y is bounded_above holds
sup (X + Y) <= (sup X) + (sup Y) by Th26;

theorem :: SUPINF_2:29
for X, Y being non empty Subset of ExtREAL st X is bounded_below & Y is bounded_below holds
(inf X) + (inf Y) <= inf (X + Y) by Th27;

theorem Th30: :: SUPINF_2:30
for X being non empty Subset of ExtREAL
for a being R_eal holds
( a is UpperBound of X iff - a is LowerBound of - X )
proof end;

theorem Th31: :: SUPINF_2:31
for X being non empty Subset of ExtREAL
for a being R_eal holds
( a is LowerBound of X iff - a is UpperBound of - X )
proof end;

theorem Th32: :: SUPINF_2:32
for X being non empty Subset of ExtREAL holds inf (- X) = - (sup X)
proof end;

theorem Th33: :: SUPINF_2:33
for X being non empty Subset of ExtREAL holds sup (- X) = - (inf X)
proof end;

definition
let X be non empty set ;
let Y be non empty Subset of ExtREAL ;
let F be Function of X,Y;
:: original: rng
redefine func rng F -> non empty Subset of ExtREAL ;
coherence
rng F is non empty Subset of ExtREAL
proof end;
end;

definition
let X be non empty set ;
let Y be non empty Subset of ExtREAL ;
let F be Function of X,Y;
func sup F -> R_eal equals :: SUPINF_2:def 7
sup (rng F);
coherence
sup (rng F) is R_eal
;
func inf F -> R_eal equals :: SUPINF_2:def 8
inf (rng F);
coherence
inf (rng F) is R_eal
;
end;

:: deftheorem defines sup SUPINF_2:def 7 :
for X being non empty set
for Y being non empty Subset of ExtREAL
for F being Function of X,Y holds sup F = sup (rng F);

:: deftheorem defines inf SUPINF_2:def 8 :
for X being non empty set
for Y being non empty Subset of ExtREAL
for F being Function of X,Y holds inf F = inf (rng F);

definition
let F be ext-real-valued Function;
let x be set ;
:: original: .
redefine func F . x -> R_eal;
coherence
F . x is R_eal
by XXREAL_0:def 1;
end;

definition
let X be non empty set ;
let Y, Z be non empty Subset of ExtREAL ;
let F be Function of X,Y;
let G be Function of X,Z;
func F + G -> Function of X,Y + Z means :Def9: :: SUPINF_2:def 9
for x being Element of X holds it . x = (F . x) + (G . x);
existence
ex b1 being Function of X,Y + Z st
for x being Element of X holds b1 . x = (F . x) + (G . x)
proof end;
uniqueness
for b1, b2 being Function of X,Y + Z st ( for x being Element of X holds b1 . x = (F . x) + (G . x) ) & ( for x being Element of X holds b2 . x = (F . x) + (G . x) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def9 defines + SUPINF_2:def 9 :
for X being non empty set
for Y, Z being non empty Subset of ExtREAL
for F being Function of X,Y
for G being Function of X,Z
for b6 being Function of X,Y + Z holds
( b6 = F + G iff for x being Element of X holds b6 . x = (F . x) + (G . x) );

theorem Th34: :: SUPINF_2:34
for X being non empty set
for Y, Z being non empty Subset of ExtREAL
for F being Function of X,Y
for G being Function of X,Z holds rng (F + G) c= (rng F) + (rng G)
proof end;

theorem Th35: :: SUPINF_2:35
for X being non empty set
for Y, Z being non empty Subset of ExtREAL
for F being Function of X,Y
for G being Function of X,Z holds sup (F + G) <= (sup F) + (sup G)
proof end;

theorem Th36: :: SUPINF_2:36
for X being non empty set
for Y, Z being non empty Subset of ExtREAL
for F being Function of X,Y
for G being Function of X,Z holds (inf F) + (inf G) <= inf (F + G)
proof end;

definition
let X be non empty set ;
let Y be non empty Subset of ExtREAL ;
let F be Function of X,Y;
func - F -> Function of X, - Y means :Def10: :: SUPINF_2:def 10
for x being Element of X holds it . x = - (F . x);
existence
ex b1 being Function of X, - Y st
for x being Element of X holds b1 . x = - (F . x)
proof end;
uniqueness
for b1, b2 being Function of X, - Y st ( for x being Element of X holds b1 . x = - (F . x) ) & ( for x being Element of X holds b2 . x = - (F . x) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def10 defines - SUPINF_2:def 10 :
for X being non empty set
for Y being non empty Subset of ExtREAL
for F being Function of X,Y
for b4 being Function of X, - Y holds
( b4 = - F iff for x being Element of X holds b4 . x = - (F . x) );

theorem Th37: :: SUPINF_2:37
for X being non empty set
for Y being non empty Subset of ExtREAL
for F being Function of X,Y holds rng (- F) = - (rng F)
proof end;

theorem Th38: :: SUPINF_2:38
for X being non empty set
for Y being non empty Subset of ExtREAL
for F being Function of X,Y holds
( inf (- F) = - (sup F) & sup (- F) = - (inf F) )
proof end;

definition
let X be non empty set ;
let Y be non empty Subset of ExtREAL ;
let F be Function of X,Y;
attr F is bounded_above means :Def11: :: SUPINF_2:def 11
sup F < +infty ;
attr F is bounded_below means :Def12: :: SUPINF_2:def 12
-infty < inf F;
end;

:: deftheorem Def11 defines bounded_above SUPINF_2:def 11 :
for X being non empty set
for Y being non empty Subset of ExtREAL
for F being Function of X,Y holds
( F is bounded_above iff sup F < +infty );

:: deftheorem Def12 defines bounded_below SUPINF_2:def 12 :
for X being non empty set
for Y being non empty Subset of ExtREAL
for F being Function of X,Y holds
( F is bounded_below iff -infty < inf F );

definition
let X be non empty set ;
let Y be non empty Subset of ExtREAL ;
let F be Function of X,Y;
attr F is bounded means :Def13: :: SUPINF_2:def 13
( F is bounded_above & F is bounded_below );
end;

:: deftheorem Def13 defines bounded SUPINF_2:def 13 :
for X being non empty set
for Y being non empty Subset of ExtREAL
for F being Function of X,Y holds
( F is bounded iff ( F is bounded_above & F is bounded_below ) );

registration
let X be non empty set ;
let Y be non empty Subset of ExtREAL ;
cluster bounded -> bounded_above bounded_below Relation of X,Y;
coherence
for b1 being Function of X,Y st b1 is bounded holds
( b1 is bounded_above & b1 is bounded_below )
by Def13;
cluster bounded_above bounded_below -> bounded Relation of X,Y;
coherence
for b1 being Function of X,Y st b1 is bounded_above & b1 is bounded_below holds
b1 is bounded
by Def13;
end;

theorem :: SUPINF_2:39
for X being non empty set
for Y being non empty Subset of ExtREAL
for F being Function of X,Y holds
( F is bounded iff ( sup F < +infty & -infty < inf F ) )
proof end;

theorem Th40: :: SUPINF_2:40
for X being non empty set
for Y being non empty Subset of ExtREAL
for F being Function of X,Y holds
( F is bounded_above iff - F is bounded_below )
proof end;

theorem Th41: :: SUPINF_2:41
for X being non empty set
for Y being non empty Subset of ExtREAL
for F being Function of X,Y holds
( F is bounded_below iff - F is bounded_above )
proof end;

theorem :: SUPINF_2:42
for X being non empty set
for Y being non empty Subset of ExtREAL
for F being Function of X,Y holds
( F is bounded iff - F is bounded )
proof end;

theorem :: SUPINF_2:43
for X being non empty set
for Y being non empty Subset of ExtREAL
for F being Function of X,Y
for x being Element of X holds
( -infty <= F . x & F . x <= +infty ) by XXREAL_0:4, XXREAL_0:5;

theorem :: SUPINF_2:44
for X being non empty set
for Y being non empty Subset of ExtREAL
for F being Function of X,Y
for x being Element of X st Y c= REAL holds
( -infty < F . x & F . x < +infty )
proof end;

theorem Th45: :: SUPINF_2:45
for X being non empty set
for Y being non empty Subset of ExtREAL
for F being Function of X,Y
for x being Element of X holds
( inf F <= F . x & F . x <= sup F )
proof end;

theorem Th46: :: SUPINF_2:46
for X being non empty set
for Y being non empty Subset of ExtREAL
for F being Function of X,Y st Y c= REAL holds
( F is bounded_above iff sup F in REAL )
proof end;

theorem Th47: :: SUPINF_2:47
for X being non empty set
for Y being non empty Subset of ExtREAL
for F being Function of X,Y st Y c= REAL holds
( F is bounded_below iff inf F in REAL )
proof end;

theorem :: SUPINF_2:48
for X being non empty set
for Y being non empty Subset of ExtREAL
for F being Function of X,Y st Y c= REAL holds
( F is bounded iff ( inf F in REAL & sup F in REAL ) )
proof end;

theorem Th49: :: SUPINF_2:49
for X being non empty set
for Y, Z being non empty Subset of ExtREAL st Y c= REAL & Z c= REAL holds
for F1 being Function of X,Y
for F2 being Function of X,Z st F1 is bounded_above & F2 is bounded_above holds
F1 + F2 is bounded_above
proof end;

theorem Th50: :: SUPINF_2:50
for X being non empty set
for Y, Z being non empty Subset of ExtREAL st Y c= REAL & Z c= REAL holds
for F1 being Function of X,Y
for F2 being Function of X,Z st F1 is bounded_below & F2 is bounded_below holds
F1 + F2 is bounded_below
proof end;

theorem :: SUPINF_2:51
for X being non empty set
for Y, Z being non empty Subset of ExtREAL st Y c= REAL & Z c= REAL holds
for F1 being Function of X,Y
for F2 being Function of X,Z st F1 is bounded & F2 is bounded holds
F1 + F2 is bounded
proof end;

theorem Th52: :: SUPINF_2:52
ex F being Function of NAT , ExtREAL st
( F is one-to-one & NAT = rng F & rng F is non empty Subset of ExtREAL )
proof end;

definition
let D be non empty set ;
let IT be Subset of D;
:: original: countable
redefine attr IT is countable means :Def14: :: SUPINF_2:def 14
( IT is empty or ex F being Function of NAT ,D st IT = rng F );
compatibility
( IT is countable iff ( IT is empty or ex F being Function of NAT ,D st IT = rng F ) )
proof end;
end;

:: deftheorem Def14 defines countable SUPINF_2:def 14 :
for D being non empty set
for IT being Subset of D holds
( IT is countable iff ( IT is empty or ex F being Function of NAT ,D st IT = rng F ) );

registration
cluster non empty V70 Element of K24(ExtREAL );
existence
ex b1 being non empty Subset of ExtREAL st b1 is countable
proof end;
end;

definition
mode Denum_Set_of_R_EAL is non empty V70 Subset of ExtREAL ;
end;

definition
let IT be set ;
attr IT is nonnegative means :Def15: :: SUPINF_2:def 15
for x being R_eal st x in IT holds
0. <= x;
end;

:: deftheorem Def15 defines nonnegative SUPINF_2:def 15 :
for IT being set holds
( IT is nonnegative iff for x being R_eal st x in IT holds
0. <= x );

registration
cluster nonnegative Element of K24(ExtREAL );
existence
ex b1 being Denum_Set_of_R_EAL st b1 is nonnegative
proof end;
end;

definition
mode Pos_Denum_Set_of_R_EAL is nonnegative Denum_Set_of_R_EAL;
end;

definition
let D be Denum_Set_of_R_EAL;
mode Num of D -> Function of NAT , ExtREAL means :Def16: :: SUPINF_2:def 16
D = rng it;
existence
ex b1 being Function of NAT , ExtREAL st D = rng b1
by Def14;
end;

:: deftheorem Def16 defines Num SUPINF_2:def 16 :
for D being Denum_Set_of_R_EAL
for b2 being Function of NAT , ExtREAL holds
( b2 is Num of D iff D = rng b2 );

theorem Th53: :: SUPINF_2:53
for D being Denum_Set_of_R_EAL
for N being Num of D ex F being Function of NAT , ExtREAL st
( F . 0 = N . 0 & ( for n being Element of NAT
for y being R_eal st y = F . n holds
F . (n + 1) = y + (N . (n + 1)) ) )
proof end;

definition
let D be Denum_Set_of_R_EAL;
let N be Num of D;
func Ser D,N -> Function of NAT , ExtREAL means :Def17: :: SUPINF_2:def 17
( it . 0 = N . 0 & ( for n being Element of NAT
for y being R_eal st y = it . n holds
it . (n + 1) = y + (N . (n + 1)) ) );
existence
ex b1 being Function of NAT , ExtREAL st
( b1 . 0 = N . 0 & ( for n being Element of NAT
for y being R_eal st y = b1 . n holds
b1 . (n + 1) = y + (N . (n + 1)) ) )
by Th53;
uniqueness
for b1, b2 being Function of NAT , ExtREAL st b1 . 0 = N . 0 & ( for n being Element of NAT
for y being R_eal st y = b1 . n holds
b1 . (n + 1) = y + (N . (n + 1)) ) & b2 . 0 = N . 0 & ( for n being Element of NAT
for y being R_eal st y = b2 . n holds
b2 . (n + 1) = y + (N . (n + 1)) ) holds
b1 = b2
proof end;
end;

:: deftheorem Def17 defines Ser SUPINF_2:def 17 :
for D being Denum_Set_of_R_EAL
for N being Num of D
for b3 being Function of NAT , ExtREAL holds
( b3 = Ser D,N iff ( b3 . 0 = N . 0 & ( for n being Element of NAT
for y being R_eal st y = b3 . n holds
b3 . (n + 1) = y + (N . (n + 1)) ) ) );

theorem Th54: :: SUPINF_2:54
for D being Pos_Denum_Set_of_R_EAL
for N being Num of D
for n being Element of NAT holds 0. <= N . n
proof end;

theorem Th55: :: SUPINF_2:55
for D being Pos_Denum_Set_of_R_EAL
for N being Num of D
for n being Element of NAT holds
( (Ser D,N) . n <= (Ser D,N) . (n + 1) & 0. <= (Ser D,N) . n )
proof end;

theorem Th56: :: SUPINF_2:56
for D being Pos_Denum_Set_of_R_EAL
for N being Num of D
for n, m being Element of NAT holds (Ser D,N) . n <= (Ser D,N) . (n + m)
proof end;

definition
let D be Denum_Set_of_R_EAL;
mode Set_of_Series of D -> non empty Subset of ExtREAL means :: SUPINF_2:def 18
ex N being Num of D st it = rng (Ser D,N);
existence
ex b1 being non empty Subset of ExtREAL ex N being Num of D st b1 = rng (Ser D,N)
proof end;
end;

:: deftheorem defines Set_of_Series SUPINF_2:def 18 :
for D being Denum_Set_of_R_EAL
for b2 being non empty Subset of ExtREAL holds
( b2 is Set_of_Series of D iff ex N being Num of D st b2 = rng (Ser D,N) );

Lm5: for F being Function of NAT , ExtREAL holds rng F is non empty Subset of ExtREAL
;

definition
let F be Function of NAT , ExtREAL ;
:: original: rng
redefine func rng F -> non empty Subset of ExtREAL ;
coherence
rng F is non empty Subset of ExtREAL
by Lm5;
end;

definition
let D be Pos_Denum_Set_of_R_EAL;
let N be Num of D;
func SUM D,N -> R_eal equals :: SUPINF_2:def 19
sup (rng (Ser D,N));
coherence
sup (rng (Ser D,N)) is R_eal
;
end;

:: deftheorem defines SUM SUPINF_2:def 19 :
for D being Pos_Denum_Set_of_R_EAL
for N being Num of D holds SUM D,N = sup (rng (Ser D,N));

definition
let D be Pos_Denum_Set_of_R_EAL;
let N be Num of D;
pred D is_sumable N means :: SUPINF_2:def 20
SUM D,N in REAL ;
end;

:: deftheorem defines is_sumable SUPINF_2:def 20 :
for D being Pos_Denum_Set_of_R_EAL
for N being Num of D holds
( D is_sumable N iff SUM D,N in REAL );

theorem Th57: :: SUPINF_2:57
for F being Function of NAT , ExtREAL holds rng F is Denum_Set_of_R_EAL by Def14;

definition
let F be Function of NAT , ExtREAL ;
:: original: rng
redefine func rng F -> Denum_Set_of_R_EAL;
coherence
rng F is Denum_Set_of_R_EAL
by Th57;
end;

definition
let F be Function of NAT , ExtREAL ;
func Ser F -> Function of NAT , ExtREAL means :Def21: :: SUPINF_2:def 21
for N being Num of rng F st N = F holds
it = Ser (rng F),N;
existence
ex b1 being Function of NAT , ExtREAL st
for N being Num of rng F st N = F holds
b1 = Ser (rng F),N
proof end;
uniqueness
for b1, b2 being Function of NAT , ExtREAL st ( for N being Num of rng F st N = F holds
b1 = Ser (rng F),N ) & ( for N being Num of rng F st N = F holds
b2 = Ser (rng F),N ) holds
b1 = b2
proof end;
end;

:: deftheorem Def21 defines Ser SUPINF_2:def 21 :
for F, b2 being Function of NAT , ExtREAL holds
( b2 = Ser F iff for N being Num of rng F st N = F holds
b2 = Ser (rng F),N );

definition
let R be Relation;
attr R is nonnegative means :Def22: :: SUPINF_2:def 22
rng R is nonnegative;
end;

:: deftheorem Def22 defines nonnegative SUPINF_2:def 22 :
for R being Relation holds
( R is nonnegative iff rng R is nonnegative );

definition
let F be Function of NAT , ExtREAL ;
func SUM F -> R_eal equals :: SUPINF_2:def 23
sup (rng (Ser F));
coherence
sup (rng (Ser F)) is R_eal
;
end;

:: deftheorem defines SUM SUPINF_2:def 23 :
for F being Function of NAT , ExtREAL holds SUM F = sup (rng (Ser F));

theorem Th58: :: SUPINF_2:58
for X being set
for F being PartFunc of X, ExtREAL holds
( F is nonnegative iff for n being Element of X holds 0. <= F . n )
proof end;

theorem Th59: :: SUPINF_2:59
for F being Function of NAT , ExtREAL
for n being Element of NAT st F is nonnegative holds
( (Ser F) . n <= (Ser F) . (n + 1) & 0. <= (Ser F) . n )
proof end;

theorem Th60: :: SUPINF_2:60
for F being Function of NAT , ExtREAL st F is nonnegative holds
for n, m being Element of NAT holds (Ser F) . n <= (Ser F) . (n + m)
proof end;

theorem Th61: :: SUPINF_2:61
for F1, F2 being Function of NAT , ExtREAL st ( for n being Element of NAT holds F1 . n <= F2 . n ) holds
for n being Element of NAT holds (Ser F1) . n <= (Ser F2) . n
proof end;

theorem Th62: :: SUPINF_2:62
for F1, F2 being Function of NAT , ExtREAL st ( for n being Element of NAT holds F1 . n <= F2 . n ) holds
SUM F1 <= SUM F2
proof end;

theorem Th63: :: SUPINF_2:63
for F being Function of NAT , ExtREAL holds
( (Ser F) . 0 = F . 0 & ( for n being Element of NAT
for y being R_eal st y = (Ser F) . n holds
(Ser F) . (n + 1) = y + (F . (n + 1)) ) )
proof end;

theorem Th64: :: SUPINF_2:64
for F being Function of NAT , ExtREAL st F is nonnegative & ex n being Element of NAT st F . n = +infty holds
SUM F = +infty
proof end;

definition
let F be Function of NAT , ExtREAL ;
attr F is summable means :Def24: :: SUPINF_2:def 24
SUM F in REAL ;
end;

:: deftheorem Def24 defines summable SUPINF_2:def 24 :
for F being Function of NAT , ExtREAL holds
( F is summable iff SUM F in REAL );

theorem :: SUPINF_2:65
for F being Function of NAT , ExtREAL st F is nonnegative & ex n being Element of NAT st F . n = +infty holds
not F is summable
proof end;

theorem :: SUPINF_2:66
for F1, F2 being Function of NAT , ExtREAL st F1 is nonnegative & ( for n being Element of NAT holds F1 . n <= F2 . n ) & F2 is summable holds
F1 is summable
proof end;

theorem Th67: :: SUPINF_2:67
for F being Function of NAT , ExtREAL st F is nonnegative holds
for n being Nat st ( for r being Element of NAT st n <= r holds
F . r = 0. ) holds
SUM F = (Ser F) . n
proof end;

theorem Th68: :: SUPINF_2:68
for F being Function of NAT , ExtREAL st ( for n being Element of NAT holds F . n in REAL ) holds
for n being Element of NAT holds (Ser F) . n in REAL
proof end;

theorem :: SUPINF_2:69
for F being Function of NAT , ExtREAL st F is nonnegative & ex n being Element of NAT st
( ( for k being Element of NAT st n <= k holds
F . k = 0. ) & ( for k being Element of NAT st k <= n holds
F . k <> +infty ) ) holds
F is summable
proof end;

theorem :: SUPINF_2:70
for X being set
for F being PartFunc of X, ExtREAL holds
( F is nonnegative iff for n being set holds 0. <= F . n )
proof end;

theorem :: SUPINF_2:71
for X being set
for F being PartFunc of X, ExtREAL st ( for n being set st n in dom F holds
0. <= F . n ) holds
F is nonnegative
proof end;

theorem :: SUPINF_2:72
for x, y being R_eal st x <= y holds
y - x >= 0
proof end;

theorem Th73: :: SUPINF_2:73
for x being R_eal holds x - x = 0
proof end;

theorem :: SUPINF_2:74
for z, y, x being R_eal st ( z = -infty & y = +infty implies x <= 0 ) & ( z = +infty & y = -infty implies x <= 0 ) & x - y <= z holds
x <= z + y
proof end;

theorem :: SUPINF_2:75
for x, y, z being R_eal st ( x = +infty & y = +infty implies 0 <= z ) & ( x = -infty & y = -infty implies 0 <= z ) & x <= z + y holds
x - y <= z
proof end;