:: MESFUNC5 semantic presentation begin theorem Th1: :: MESFUNC5:1 for x, y being R_eal holds |.(x - y).| = |.(y - x).| proof let x, y be R_eal; ::_thesis: |.(x - y).| = |.(y - x).| |.(y - x).| = |.(- (x - y)).| by XXREAL_3:26; hence |.(x - y).| = |.(y - x).| by EXTREAL2:18; ::_thesis: verum end; theorem Th2: :: MESFUNC5:2 for x, y being R_eal holds y - x <= |.(x - y).| proof let x, y be R_eal; ::_thesis: y - x <= |.(x - y).| - |.(x - y).| <= x - y by EXTREAL2:9; then - (x - y) <= |.(x - y).| by XXREAL_3:60; hence y - x <= |.(x - y).| by XXREAL_3:26; ::_thesis: verum end; theorem Th3: :: MESFUNC5:3 for x, y being R_eal for e being real number holds ( not |.(x - y).| < e or ( x = +infty & y = +infty ) or ( x = -infty & y = -infty ) or ( x <> +infty & x <> -infty & y <> +infty & y <> -infty ) ) proof let x, y be R_eal; ::_thesis: for e being real number holds ( not |.(x - y).| < e or ( x = +infty & y = +infty ) or ( x = -infty & y = -infty ) or ( x <> +infty & x <> -infty & y <> +infty & y <> -infty ) ) let e be real number ; ::_thesis: ( not |.(x - y).| < e or ( x = +infty & y = +infty ) or ( x = -infty & y = -infty ) or ( x <> +infty & x <> -infty & y <> +infty & y <> -infty ) ) assume A1: |.(x - y).| < e ; ::_thesis: ( ( x = +infty & y = +infty ) or ( x = -infty & y = -infty ) or ( x <> +infty & x <> -infty & y <> +infty & y <> -infty ) ) y - x <= |.(x - y).| by Th2; then A2: y - x < e by A1, XXREAL_0:2; x - y <= |.(x - y).| by EXTREAL2:9; then x - y < R_EAL e by A1, XXREAL_0:2; hence ( ( x = +infty & y = +infty ) or ( x = -infty & y = -infty ) or ( x <> +infty & x <> -infty & y <> +infty & y <> -infty ) ) by A2, XXREAL_3:54; ::_thesis: verum end; theorem Th4: :: MESFUNC5:4 for n being Nat for p being R_eal st 0 <= p & p < n holds ex k being Nat st ( 1 <= k & k <= (2 |^ n) * n & (k - 1) / (2 |^ n) <= p & p < k / (2 |^ n) ) proof let n be Nat; ::_thesis: for p being R_eal st 0 <= p & p < n holds ex k being Nat st ( 1 <= k & k <= (2 |^ n) * n & (k - 1) / (2 |^ n) <= p & p < k / (2 |^ n) ) let p be R_eal; ::_thesis: ( 0 <= p & p < n implies ex k being Nat st ( 1 <= k & k <= (2 |^ n) * n & (k - 1) / (2 |^ n) <= p & p < k / (2 |^ n) ) ) assume that A1: 0 <= p and A2: p < n ; ::_thesis: ex k being Nat st ( 1 <= k & k <= (2 |^ n) * n & (k - 1) / (2 |^ n) <= p & p < k / (2 |^ n) ) reconsider p1 = p as Real by A1, A2, XXREAL_0:46; reconsider n = n as Element of NAT by ORDINAL1:def_12; set k = [\((p1 * (2 |^ n)) + 1)/]; A3: ((p1 * (2 |^ n)) + 1) - 1 = p1 * (2 |^ n) ; then A4: 0 < [\((p1 * (2 |^ n)) + 1)/] by A1, INT_1:def_6; then reconsider k = [\((p1 * (2 |^ n)) + 1)/] as Element of NAT by INT_1:3; A5: p1 * (2 |^ n) < k by A3, INT_1:def_6; A6: 0 < 2 |^ n by PREPOWER:6; A7: now__::_thesis:_not_k_>_(2_|^_n)_*_n p1 * (2 |^ n) < (2 |^ n) * n by A2, A6, XREAL_1:68; then A8: (p1 * (2 |^ n)) + 1 < ((2 |^ n) * n) + 1 by XREAL_1:6; reconsider N = (2 |^ n) * n as Integer ; assume A9: k > (2 |^ n) * n ; ::_thesis: contradiction A10: [\N/] = N by INT_1:25; k <= (p1 * (2 |^ n)) + 1 by INT_1:def_6; then (2 |^ n) * n < (p1 * (2 |^ n)) + 1 by A9, XXREAL_0:2; hence contradiction by A9, A8, A10, INT_1:67; ::_thesis: verum end; take k ; ::_thesis: ( 1 <= k & k <= (2 |^ n) * n & (k - 1) / (2 |^ n) <= p & p < k / (2 |^ n) ) k <= (p1 * (2 |^ n)) + 1 by INT_1:def_6; then A11: k - 1 <= p1 * (2 |^ n) by XREAL_1:20; 0 + 1 <= k by A4, NAT_1:13; hence ( 1 <= k & k <= (2 |^ n) * n & (k - 1) / (2 |^ n) <= p & p < k / (2 |^ n) ) by A6, A7, A5, A11, XREAL_1:79, XREAL_1:81; ::_thesis: verum end; theorem Th5: :: MESFUNC5:5 for n, k being Nat for p being R_eal st k <= (2 |^ n) * n & n <= p holds k / (2 |^ n) <= p proof let n, k be Nat; ::_thesis: for p being R_eal st k <= (2 |^ n) * n & n <= p holds k / (2 |^ n) <= p let p be R_eal; ::_thesis: ( k <= (2 |^ n) * n & n <= p implies k / (2 |^ n) <= p ) assume that A1: k <= (2 |^ n) * n and A2: n <= p ; ::_thesis: k / (2 |^ n) <= p assume p < k / (2 |^ n) ; ::_thesis: contradiction then n < k / (2 |^ n) by A2, XXREAL_0:2; hence contradiction by A1, PREPOWER:6, XREAL_1:79; ::_thesis: verum end; theorem Th6: :: MESFUNC5:6 for x, y, k being ext-real number st 0 <= k holds ( k * (max (x,y)) = max ((k * x),(k * y)) & k * (min (x,y)) = min ((k * x),(k * y)) ) proof let x, y, k be ext-real number ; ::_thesis: ( 0 <= k implies ( k * (max (x,y)) = max ((k * x),(k * y)) & k * (min (x,y)) = min ((k * x),(k * y)) ) ) assume A1: 0 <= k ; ::_thesis: ( k * (max (x,y)) = max ((k * x),(k * y)) & k * (min (x,y)) = min ((k * x),(k * y)) ) now__::_thesis:_k_*_(max_(x,y))_=_max_((k_*_x),(k_*_y)) percases ( max (x,y) = x or max (x,y) = y ) by XXREAL_0:16; supposeA2: max (x,y) = x ; ::_thesis: k * (max (x,y)) = max ((k * x),(k * y)) then y <= x by XXREAL_0:def_10; then k * y <= k * x by A1, XXREAL_3:71; hence k * (max (x,y)) = max ((k * x),(k * y)) by A2, XXREAL_0:def_10; ::_thesis: verum end; supposeA3: max (x,y) = y ; ::_thesis: k * (max (x,y)) = max ((k * x),(k * y)) then x <= y by XXREAL_0:def_10; then k * x <= k * y by A1, XXREAL_3:71; hence k * (max (x,y)) = max ((k * x),(k * y)) by A3, XXREAL_0:def_10; ::_thesis: verum end; end; end; hence k * (max (x,y)) = max ((k * x),(k * y)) ; ::_thesis: k * (min (x,y)) = min ((k * x),(k * y)) percases ( min (x,y) = x or min (x,y) = y ) by XXREAL_0:15; supposeA4: min (x,y) = x ; ::_thesis: k * (min (x,y)) = min ((k * x),(k * y)) then x <= y by XXREAL_0:def_9; then k * x <= k * y by A1, XXREAL_3:71; hence k * (min (x,y)) = min ((k * x),(k * y)) by A4, XXREAL_0:def_9; ::_thesis: verum end; supposeA5: min (x,y) = y ; ::_thesis: k * (min (x,y)) = min ((k * x),(k * y)) then y <= x by XXREAL_0:def_9; then k * y <= k * x by A1, XXREAL_3:71; hence k * (min (x,y)) = min ((k * x),(k * y)) by A5, XXREAL_0:def_9; ::_thesis: verum end; end; end; theorem :: MESFUNC5:7 for x, y, k being R_eal st k <= 0 holds ( k * (min (x,y)) = max ((k * x),(k * y)) & k * (max (x,y)) = min ((k * x),(k * y)) ) proof let x, y, k be R_eal; ::_thesis: ( k <= 0 implies ( k * (min (x,y)) = max ((k * x),(k * y)) & k * (max (x,y)) = min ((k * x),(k * y)) ) ) assume A1: k <= 0 ; ::_thesis: ( k * (min (x,y)) = max ((k * x),(k * y)) & k * (max (x,y)) = min ((k * x),(k * y)) ) hereby ::_thesis: k * (max (x,y)) = min ((k * x),(k * y)) percases ( max (x,y) = x or max (x,y) = y ) by XXREAL_0:16; suppose max (x,y) = x ; ::_thesis: k * (min (x,y)) = max ((k * x),(k * y)) then A2: y <= x by XXREAL_0:def_10; then k * x <= k * y by A1, XXREAL_3:101; then max ((k * x),(k * y)) = k * y by XXREAL_0:def_10; hence k * (min (x,y)) = max ((k * x),(k * y)) by A2, XXREAL_0:def_9; ::_thesis: verum end; suppose max (x,y) = y ; ::_thesis: k * (min (x,y)) = max ((k * x),(k * y)) then A3: x <= y by XXREAL_0:def_10; then k * y <= k * x by A1, XXREAL_3:101; then max ((k * x),(k * y)) = k * x by XXREAL_0:def_10; hence k * (min (x,y)) = max ((k * x),(k * y)) by A3, XXREAL_0:def_9; ::_thesis: verum end; end; end; percases ( min (x,y) = x or min (x,y) = y ) by XXREAL_0:15; suppose min (x,y) = x ; ::_thesis: k * (max (x,y)) = min ((k * x),(k * y)) then A4: x <= y by XXREAL_0:def_9; then k * y <= k * x by A1, XXREAL_3:101; then min ((k * x),(k * y)) = k * y by XXREAL_0:def_9; hence k * (max (x,y)) = min ((k * x),(k * y)) by A4, XXREAL_0:def_10; ::_thesis: verum end; suppose min (x,y) = y ; ::_thesis: k * (max (x,y)) = min ((k * x),(k * y)) then A5: y <= x by XXREAL_0:def_9; then k * x <= k * y by A1, XXREAL_3:101; then min ((k * y),(k * x)) = k * x by XXREAL_0:def_9; hence k * (max (x,y)) = min ((k * x),(k * y)) by A5, XXREAL_0:def_10; ::_thesis: verum end; end; end; begin definition let IT be set ; attrIT is nonpositive means :Def1: :: MESFUNC5:def 1 for x being R_eal st x in IT holds x <= 0 ; end; :: deftheorem Def1 defines nonpositive MESFUNC5:def_1_:_ for IT being set holds ( IT is nonpositive iff for x being R_eal st x in IT holds x <= 0 ); definition let R be Relation; attrR is nonpositive means :Def2: :: MESFUNC5:def 2 rng R is nonpositive ; end; :: deftheorem Def2 defines nonpositive MESFUNC5:def_2_:_ for R being Relation holds ( R is nonpositive iff rng R is nonpositive ); theorem Th8: :: MESFUNC5:8 for X being set for F being PartFunc of X,ExtREAL holds ( F is nonpositive iff for n being set holds F . n <= 0. ) proof let X be set ; ::_thesis: for F being PartFunc of X,ExtREAL holds ( F is nonpositive iff for n being set holds F . n <= 0. ) let F be PartFunc of X,ExtREAL; ::_thesis: ( F is nonpositive iff for n being set holds F . n <= 0. ) hereby ::_thesis: ( ( for n being set holds F . n <= 0. ) implies F is nonpositive ) assume F is nonpositive ; ::_thesis: for n being set holds F . b2 <= 0. then A1: rng F is nonpositive by Def2; let n be set ; ::_thesis: F . b1 <= 0. percases ( n in dom F or not n in dom F ) ; suppose n in dom F ; ::_thesis: F . b1 <= 0. then F . n in rng F by FUNCT_1:def_3; hence F . n <= 0. by A1, Def1; ::_thesis: verum end; suppose not n in dom F ; ::_thesis: F . b1 <= 0. hence F . n <= 0. by FUNCT_1:def_2; ::_thesis: verum end; end; end; assume A2: for n being set holds F . n <= 0. ; ::_thesis: F is nonpositive let y be R_eal; :: according to MESFUNC5:def_1,MESFUNC5:def_2 ::_thesis: ( y in rng F implies y <= 0 ) assume y in rng F ; ::_thesis: y <= 0 then ex x being set st ( x in dom F & y = F . x ) by FUNCT_1:def_3; hence y <= 0 by A2; ::_thesis: verum end; theorem Th9: :: MESFUNC5:9 for X being set for F being PartFunc of X,ExtREAL st ( for n being set st n in dom F holds F . n <= 0. ) holds F is nonpositive proof let X be set ; ::_thesis: for F being PartFunc of X,ExtREAL st ( for n being set st n in dom F holds F . n <= 0. ) holds F is nonpositive let F be PartFunc of X,ExtREAL; ::_thesis: ( ( for n being set st n in dom F holds F . n <= 0. ) implies F is nonpositive ) assume A1: for n being set st n in dom F holds F . n <= 0. ; ::_thesis: F is nonpositive let y be R_eal; :: according to MESFUNC5:def_1,MESFUNC5:def_2 ::_thesis: ( y in rng F implies y <= 0 ) assume y in rng F ; ::_thesis: y <= 0 then ex x being set st ( x in dom F & y = F . x ) by FUNCT_1:def_3; hence y <= 0 by A1; ::_thesis: verum end; definition let R be Relation; attrR is without-infty means :Def3: :: MESFUNC5:def 3 not -infty in rng R; attrR is without+infty means :Def4: :: MESFUNC5:def 4 not +infty in rng R; end; :: deftheorem Def3 defines without-infty MESFUNC5:def_3_:_ for R being Relation holds ( R is without-infty iff not -infty in rng R ); :: deftheorem Def4 defines without+infty MESFUNC5:def_4_:_ for R being Relation holds ( R is without+infty iff not +infty in rng R ); definition let X be non empty set ; let f be PartFunc of X,ExtREAL; :: original: without-infty redefine attrf is without-infty means :Def5: :: MESFUNC5:def 5 for x being set holds -infty < f . x; compatibility ( f is without-infty iff for x being set holds -infty < f . x ) proof hereby ::_thesis: ( ( for x being set holds -infty < f . x ) implies f is without-infty ) assume f is without-infty ; ::_thesis: for x being set holds -infty < f . x then A1: not -infty in rng f by Def3; hereby ::_thesis: verum let x be set ; ::_thesis: -infty < f . b1 percases ( x in dom f or not x in dom f ) ; suppose x in dom f ; ::_thesis: -infty < f . b1 then f . x <> -infty by A1, FUNCT_1:def_3; hence -infty < f . x by XXREAL_0:6; ::_thesis: verum end; suppose not x in dom f ; ::_thesis: -infty < f . b1 hence -infty < f . x by FUNCT_1:def_2; ::_thesis: verum end; end; end; end; assume A2: for x being set holds -infty < f . x ; ::_thesis: f is without-infty now__::_thesis:_not_-infty_in_rng_f assume -infty in rng f ; ::_thesis: contradiction then ex x being set st ( x in dom f & f . x = -infty ) by FUNCT_1:def_3; hence contradiction by A2; ::_thesis: verum end; hence f is without-infty by Def3; ::_thesis: verum end; :: original: without+infty redefine attrf is without+infty means :Def6: :: MESFUNC5:def 6 for x being set holds f . x < +infty ; compatibility ( f is without+infty iff for x being set holds f . x < +infty ) proof hereby ::_thesis: ( ( for x being set holds f . x < +infty ) implies f is without+infty ) assume f is without+infty ; ::_thesis: for x being set holds f . x < +infty then A3: not +infty in rng f by Def4; hereby ::_thesis: verum let x be set ; ::_thesis: f . b1 < +infty percases ( x in dom f or not x in dom f ) ; suppose x in dom f ; ::_thesis: f . b1 < +infty then f . x <> +infty by A3, FUNCT_1:def_3; hence f . x < +infty by XXREAL_0:4; ::_thesis: verum end; suppose not x in dom f ; ::_thesis: f . b1 < +infty hence f . x < +infty by FUNCT_1:def_2; ::_thesis: verum end; end; end; end; assume A4: for x being set holds f . x < +infty ; ::_thesis: f is without+infty now__::_thesis:_not_+infty_in_rng_f assume +infty in rng f ; ::_thesis: contradiction then ex x being set st ( x in dom f & f . x = +infty ) by FUNCT_1:def_3; hence contradiction by A4; ::_thesis: verum end; hence f is without+infty by Def4; ::_thesis: verum end; end; :: deftheorem Def5 defines without-infty MESFUNC5:def_5_:_ for X being non empty set for f being PartFunc of X,ExtREAL holds ( f is without-infty iff for x being set holds -infty < f . x ); :: deftheorem Def6 defines without+infty MESFUNC5:def_6_:_ for X being non empty set for f being PartFunc of X,ExtREAL holds ( f is without+infty iff for x being set holds f . x < +infty ); theorem Th10: :: MESFUNC5:10 for X being non empty set for f being PartFunc of X,ExtREAL holds ( ( for x being set st x in dom f holds -infty < f . x ) iff f is () ) proof let X be non empty set ; ::_thesis: for f being PartFunc of X,ExtREAL holds ( ( for x being set st x in dom f holds -infty < f . x ) iff f is () ) let f be PartFunc of X,ExtREAL; ::_thesis: ( ( for x being set st x in dom f holds -infty < f . x ) iff f is () ) hereby ::_thesis: ( f is () implies for x being set st x in dom f holds -infty < f . x ) assume A1: for x being set st x in dom f holds -infty < f . x ; ::_thesis: f is () now__::_thesis:_for_x_being_set_holds_-infty_<_f_._x let x be set ; ::_thesis: -infty < f . b1 percases ( x in dom f or not x in dom f ) ; suppose x in dom f ; ::_thesis: -infty < f . b1 hence -infty < f . x by A1; ::_thesis: verum end; suppose not x in dom f ; ::_thesis: -infty < f . b1 hence -infty < f . x by FUNCT_1:def_2; ::_thesis: verum end; end; end; hence f is () by Def5; ::_thesis: verum end; assume f is () ; ::_thesis: for x being set st x in dom f holds -infty < f . x hence for x being set st x in dom f holds -infty < f . x by Def5; ::_thesis: verum end; theorem Th11: :: MESFUNC5:11 for X being non empty set for f being PartFunc of X,ExtREAL holds ( ( for x being set st x in dom f holds f . x < +infty ) iff f is () ) proof let X be non empty set ; ::_thesis: for f being PartFunc of X,ExtREAL holds ( ( for x being set st x in dom f holds f . x < +infty ) iff f is () ) let f be PartFunc of X,ExtREAL; ::_thesis: ( ( for x being set st x in dom f holds f . x < +infty ) iff f is () ) hereby ::_thesis: ( f is () implies for x being set st x in dom f holds f . x < +infty ) assume A1: for x being set st x in dom f holds f . x < +infty ; ::_thesis: f is () now__::_thesis:_for_x_being_set_holds_f_._x_<_+infty let x be set ; ::_thesis: f . b1 < +infty percases ( x in dom f or not x in dom f ) ; suppose x in dom f ; ::_thesis: f . b1 < +infty hence f . x < +infty by A1; ::_thesis: verum end; suppose not x in dom f ; ::_thesis: f . b1 < +infty hence f . x < +infty by FUNCT_1:def_2; ::_thesis: verum end; end; end; hence f is () by Def6; ::_thesis: verum end; assume f is () ; ::_thesis: for x being set st x in dom f holds f . x < +infty hence for x being set st x in dom f holds f . x < +infty by Def6; ::_thesis: verum end; theorem Th12: :: MESFUNC5:12 for X being non empty set for f being PartFunc of X,ExtREAL st f is nonnegative holds f is () proof let X be non empty set ; ::_thesis: for f being PartFunc of X,ExtREAL st f is nonnegative holds f is () let f be PartFunc of X,ExtREAL; ::_thesis: ( f is nonnegative implies f is () ) assume f is nonnegative ; ::_thesis: f is () then for x being set holds -infty < f . x by SUPINF_2:51; hence f is () by Def5; ::_thesis: verum end; theorem Th13: :: MESFUNC5:13 for X being non empty set for f being PartFunc of X,ExtREAL st f is nonpositive holds f is () proof let X be non empty set ; ::_thesis: for f being PartFunc of X,ExtREAL st f is nonpositive holds f is () let f be PartFunc of X,ExtREAL; ::_thesis: ( f is nonpositive implies f is () ) assume f is nonpositive ; ::_thesis: f is () then for x being set holds f . x < +infty by Th8; hence f is () by Def6; ::_thesis: verum end; registration let X be non empty set ; cluster Function-like nonnegative -> () for Element of bool [:X,ExtREAL:]; coherence for b1 being PartFunc of X,ExtREAL st b1 is nonnegative holds b1 is () by Th12; cluster Function-like nonpositive -> () for Element of bool [:X,ExtREAL:]; coherence for b1 being PartFunc of X,ExtREAL st b1 is nonpositive holds b1 is () by Th13; end; theorem Th14: :: MESFUNC5:14 for X being non empty set for S being SigmaField of X for f being PartFunc of X,ExtREAL st f is_simple_func_in S holds ( f is () & f is () ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for f being PartFunc of X,ExtREAL st f is_simple_func_in S holds ( f is () & f is () ) let S be SigmaField of X; ::_thesis: for f being PartFunc of X,ExtREAL st f is_simple_func_in S holds ( f is () & f is () ) let f be PartFunc of X,ExtREAL; ::_thesis: ( f is_simple_func_in S implies ( f is () & f is () ) ) assume A1: f is_simple_func_in S ; ::_thesis: ( f is () & f is () ) hereby ::_thesis: f is () assume f is () ; ::_thesis: contradiction then +infty in rng f by Def4; then f " {+infty} <> {} by FUNCT_1:72; then consider x being set such that A2: x in f " {+infty} by XBOOLE_0:def_1; A3: f is V60() by A1, MESFUNC2:def_4; f . x in {+infty} by A2, FUNCT_1:def_7; hence contradiction by A3, TARSKI:def_1; ::_thesis: verum end; hereby ::_thesis: verum assume f is () ; ::_thesis: contradiction then -infty in rng f by Def3; then f " {-infty} <> {} by FUNCT_1:72; then consider x being set such that A4: x in f " {-infty} by XBOOLE_0:def_1; A5: f is V60() by A1, MESFUNC2:def_4; f . x in {-infty} by A4, FUNCT_1:def_7; hence contradiction by A5, TARSKI:def_1; ::_thesis: verum end; end; theorem Th15: :: MESFUNC5:15 for X being non empty set for Y being set for f being PartFunc of X,ExtREAL st f is nonnegative holds f | Y is nonnegative proof let X be non empty set ; ::_thesis: for Y being set for f being PartFunc of X,ExtREAL st f is nonnegative holds f | Y is nonnegative let Y be set ; ::_thesis: for f being PartFunc of X,ExtREAL st f is nonnegative holds f | Y is nonnegative let f be PartFunc of X,ExtREAL; ::_thesis: ( f is nonnegative implies f | Y is nonnegative ) assume A1: f is nonnegative ; ::_thesis: f | Y is nonnegative now__::_thesis:_for_x_being_set_st_x_in_dom_(f_|_Y)_holds_ 0_<=_(f_|_Y)_._x let x be set ; ::_thesis: ( x in dom (f | Y) implies 0 <= (f | Y) . x ) assume A2: x in dom (f | Y) ; ::_thesis: 0 <= (f | Y) . x then (f | Y) . x = f . x by FUNCT_1:47; hence 0 <= (f | Y) . x by A1, A2, SUPINF_2:39; ::_thesis: verum end; hence f | Y is nonnegative by SUPINF_2:52; ::_thesis: verum end; theorem Th16: :: MESFUNC5:16 for X being non empty set for f, g being PartFunc of X,ExtREAL st f is () & g is () holds dom (f + g) = (dom f) /\ (dom g) proof let X be non empty set ; ::_thesis: for f, g being PartFunc of X,ExtREAL st f is () & g is () holds dom (f + g) = (dom f) /\ (dom g) let f, g be PartFunc of X,ExtREAL; ::_thesis: ( f is () & g is () implies dom (f + g) = (dom f) /\ (dom g) ) assume that A1: f is () and A2: g is () ; ::_thesis: dom (f + g) = (dom f) /\ (dom g) not -infty in rng g by A2, Def3; then A3: g " {-infty} = {} by FUNCT_1:72; not -infty in rng f by A1, Def3; then f " {-infty} = {} by FUNCT_1:72; then ((f " {+infty}) /\ (g " {-infty})) \/ ((f " {-infty}) /\ (g " {+infty})) = {} by A3; then dom (f + g) = ((dom f) /\ (dom g)) \ {} by MESFUNC1:def_3; hence dom (f + g) = (dom f) /\ (dom g) ; ::_thesis: verum end; theorem :: MESFUNC5:17 for X being non empty set for f, g being PartFunc of X,ExtREAL st f is () & g is () holds dom (f - g) = (dom f) /\ (dom g) proof let X be non empty set ; ::_thesis: for f, g being PartFunc of X,ExtREAL st f is () & g is () holds dom (f - g) = (dom f) /\ (dom g) let f, g be PartFunc of X,ExtREAL; ::_thesis: ( f is () & g is () implies dom (f - g) = (dom f) /\ (dom g) ) assume that A1: f is () and A2: g is () ; ::_thesis: dom (f - g) = (dom f) /\ (dom g) not +infty in rng g by A2, Def4; then A3: g " {+infty} = {} by FUNCT_1:72; not -infty in rng f by A1, Def3; then f " {-infty} = {} by FUNCT_1:72; then ((f " {+infty}) /\ (g " {+infty})) \/ ((f " {-infty}) /\ (g " {-infty})) = {} by A3; then dom (f - g) = ((dom f) /\ (dom g)) \ {} by MESFUNC1:def_4; hence dom (f - g) = (dom f) /\ (dom g) ; ::_thesis: verum end; theorem Th18: :: MESFUNC5:18 for X being non empty set for S being SigmaField of X for f, g being PartFunc of X,ExtREAL for F being Function of RAT,S for r being Real for A being Element of S st f is () & g is () & ( for p being Rational holds F . p = (A /\ (less_dom (f,(R_EAL p)))) /\ (A /\ (less_dom (g,(R_EAL (r - p))))) ) holds A /\ (less_dom ((f + g),(R_EAL r))) = union (rng F) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for f, g being PartFunc of X,ExtREAL for F being Function of RAT,S for r being Real for A being Element of S st f is () & g is () & ( for p being Rational holds F . p = (A /\ (less_dom (f,(R_EAL p)))) /\ (A /\ (less_dom (g,(R_EAL (r - p))))) ) holds A /\ (less_dom ((f + g),(R_EAL r))) = union (rng F) let S be SigmaField of X; ::_thesis: for f, g being PartFunc of X,ExtREAL for F being Function of RAT,S for r being Real for A being Element of S st f is () & g is () & ( for p being Rational holds F . p = (A /\ (less_dom (f,(R_EAL p)))) /\ (A /\ (less_dom (g,(R_EAL (r - p))))) ) holds A /\ (less_dom ((f + g),(R_EAL r))) = union (rng F) let f, g be PartFunc of X,ExtREAL; ::_thesis: for F being Function of RAT,S for r being Real for A being Element of S st f is () & g is () & ( for p being Rational holds F . p = (A /\ (less_dom (f,(R_EAL p)))) /\ (A /\ (less_dom (g,(R_EAL (r - p))))) ) holds A /\ (less_dom ((f + g),(R_EAL r))) = union (rng F) let F be Function of RAT,S; ::_thesis: for r being Real for A being Element of S st f is () & g is () & ( for p being Rational holds F . p = (A /\ (less_dom (f,(R_EAL p)))) /\ (A /\ (less_dom (g,(R_EAL (r - p))))) ) holds A /\ (less_dom ((f + g),(R_EAL r))) = union (rng F) let r be Real; ::_thesis: for A being Element of S st f is () & g is () & ( for p being Rational holds F . p = (A /\ (less_dom (f,(R_EAL p)))) /\ (A /\ (less_dom (g,(R_EAL (r - p))))) ) holds A /\ (less_dom ((f + g),(R_EAL r))) = union (rng F) let A be Element of S; ::_thesis: ( f is () & g is () & ( for p being Rational holds F . p = (A /\ (less_dom (f,(R_EAL p)))) /\ (A /\ (less_dom (g,(R_EAL (r - p))))) ) implies A /\ (less_dom ((f + g),(R_EAL r))) = union (rng F) ) assume that A1: f is () and A2: g is () and A3: for p being Rational holds F . p = (A /\ (less_dom (f,(R_EAL p)))) /\ (A /\ (less_dom (g,(R_EAL (r - p))))) ; ::_thesis: A /\ (less_dom ((f + g),(R_EAL r))) = union (rng F) A4: dom (f + g) = (dom f) /\ (dom g) by A1, A2, Th16; A5: union (rng F) c= A /\ (less_dom ((f + g),(R_EAL r))) proof let x be set ; :: according to TARSKI:def_3 ::_thesis: ( not x in union (rng F) or x in A /\ (less_dom ((f + g),(R_EAL r))) ) assume x in union (rng F) ; ::_thesis: x in A /\ (less_dom ((f + g),(R_EAL r))) then consider Y being set such that A6: x in Y and A7: Y in rng F by TARSKI:def_4; consider p being set such that A8: p in dom F and A9: Y = F . p by A7, FUNCT_1:def_3; reconsider p = p as Rational by A8; A10: x in (A /\ (less_dom (f,(R_EAL p)))) /\ (A /\ (less_dom (g,(R_EAL (r - p))))) by A3, A6, A9; then A11: x in A /\ (less_dom (f,(R_EAL p))) by XBOOLE_0:def_4; then A12: x in A by XBOOLE_0:def_4; A13: x in less_dom (f,(R_EAL p)) by A11, XBOOLE_0:def_4; x in A /\ (less_dom (g,(R_EAL (r - p)))) by A10, XBOOLE_0:def_4; then A14: x in less_dom (g,(R_EAL (r - p))) by XBOOLE_0:def_4; reconsider x = x as Element of X by A10; f . x < R_EAL p by A13, MESFUNC1:def_11; then A15: f . x <> +infty by XXREAL_0:4; A16: -infty < f . x by A1, Def5; A17: -infty < g . x by A2, Def5; A18: g . x < R_EAL (r - p) by A14, MESFUNC1:def_11; then g . x <> +infty by XXREAL_0:4; then reconsider f1 = f . x, g1 = g . x as Real by A16, A17, A15, XXREAL_0:14; A19: p < r - g1 by A18, XREAL_1:12; f1 < p by A13, MESFUNC1:def_11; then f1 < r - g1 by A19, XXREAL_0:2; then A20: f1 + g1 < r by XREAL_1:20; A21: x in dom g by A14, MESFUNC1:def_11; x in dom f by A13, MESFUNC1:def_11; then A22: x in dom (f + g) by A4, A21, XBOOLE_0:def_4; then (f + g) . x = (f . x) + (g . x) by MESFUNC1:def_3 .= f1 + g1 by SUPINF_2:1 ; then x in less_dom ((f + g),(R_EAL r)) by A20, A22, MESFUNC1:def_11; hence x in A /\ (less_dom ((f + g),(R_EAL r))) by A12, XBOOLE_0:def_4; ::_thesis: verum end; A /\ (less_dom ((f + g),(R_EAL r))) c= union (rng F) proof let x be set ; :: according to TARSKI:def_3 ::_thesis: ( not x in A /\ (less_dom ((f + g),(R_EAL r))) or x in union (rng F) ) assume A23: x in A /\ (less_dom ((f + g),(R_EAL r))) ; ::_thesis: x in union (rng F) then A24: x in A by XBOOLE_0:def_4; A25: x in less_dom ((f + g),(R_EAL r)) by A23, XBOOLE_0:def_4; then A26: x in dom (f + g) by MESFUNC1:def_11; then A27: x in dom f by A4, XBOOLE_0:def_4; A28: (f + g) . x < R_EAL r by A25, MESFUNC1:def_11; A29: x in dom g by A4, A26, XBOOLE_0:def_4; reconsider x = x as Element of X by A23; A30: -infty < f . x by A1, Def5; A31: (f . x) + (g . x) < R_EAL r by A26, A28, MESFUNC1:def_3; then A32: g . x <> +infty by A30, XXREAL_3:52; A33: -infty < g . x by A2, Def5; then f . x <> +infty by A31, XXREAL_3:52; then reconsider f1 = f . x, g1 = g . x as Real by A30, A33, A32, XXREAL_0:14; A34: (R_EAL r) - (g . x) = r - g1 by SUPINF_2:3; f . x < (R_EAL r) - (g . x) by A31, A30, A33, XXREAL_3:52; then consider p being Rational such that A35: f1 < p and A36: p < r - g1 by A34, RAT_1:7; not r - p <= g1 by A36, XREAL_1:12; then x in less_dom (g,(R_EAL (r - p))) by A29, MESFUNC1:def_11; then A37: x in A /\ (less_dom (g,(R_EAL (r - p)))) by A24, XBOOLE_0:def_4; p in RAT by RAT_1:def_2; then p in dom F by FUNCT_2:def_1; then A38: F . p in rng F by FUNCT_1:def_3; x in less_dom (f,(R_EAL p)) by A27, A35, MESFUNC1:def_11; then x in A /\ (less_dom (f,(R_EAL p))) by A24, XBOOLE_0:def_4; then x in (A /\ (less_dom (f,(R_EAL p)))) /\ (A /\ (less_dom (g,(R_EAL (r - p))))) by A37, XBOOLE_0:def_4; then x in F . p by A3; hence x in union (rng F) by A38, TARSKI:def_4; ::_thesis: verum end; hence A /\ (less_dom ((f + g),(R_EAL r))) = union (rng F) by A5, XBOOLE_0:def_10; ::_thesis: verum end; definition let X be non empty set ; let f be PartFunc of X,REAL; func R_EAL f -> PartFunc of X,ExtREAL equals :: MESFUNC5:def 7 f; coherence f is PartFunc of X,ExtREAL by NUMBERS:31, RELSET_1:7; end; :: deftheorem defines R_EAL MESFUNC5:def_7_:_ for X being non empty set for f being PartFunc of X,REAL holds R_EAL f = f; theorem Th19: :: MESFUNC5:19 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is nonnegative & g is nonnegative holds f + g is nonnegative proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is nonnegative & g is nonnegative holds f + g is nonnegative let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is nonnegative & g is nonnegative holds f + g is nonnegative let M be sigma_Measure of S; ::_thesis: for f, g being PartFunc of X,ExtREAL st f is nonnegative & g is nonnegative holds f + g is nonnegative let f, g be PartFunc of X,ExtREAL; ::_thesis: ( f is nonnegative & g is nonnegative implies f + g is nonnegative ) assume that A1: f is nonnegative and A2: g is nonnegative ; ::_thesis: f + g is nonnegative for x being set st x in dom (f + g) holds 0 <= (f + g) . x proof let x be set ; ::_thesis: ( x in dom (f + g) implies 0 <= (f + g) . x ) assume A3: x in dom (f + g) ; ::_thesis: 0 <= (f + g) . x 0 <= f . x by A1, SUPINF_2:51; then A4: g . x <= (f . x) + (g . x) by XXREAL_3:39; 0 <= g . x by A2, SUPINF_2:51; hence 0 <= (f + g) . x by A3, A4, MESFUNC1:def_3; ::_thesis: verum end; hence f + g is nonnegative by SUPINF_2:52; ::_thesis: verum end; theorem Th20: :: MESFUNC5:20 for X being non empty set for f being PartFunc of X,ExtREAL for c being Real st f is nonnegative holds ( ( 0 <= c implies c (#) f is nonnegative ) & ( c <= 0 implies c (#) f is nonpositive ) ) proof let X be non empty set ; ::_thesis: for f being PartFunc of X,ExtREAL for c being Real st f is nonnegative holds ( ( 0 <= c implies c (#) f is nonnegative ) & ( c <= 0 implies c (#) f is nonpositive ) ) let f be PartFunc of X,ExtREAL; ::_thesis: for c being Real st f is nonnegative holds ( ( 0 <= c implies c (#) f is nonnegative ) & ( c <= 0 implies c (#) f is nonpositive ) ) let c be Real; ::_thesis: ( f is nonnegative implies ( ( 0 <= c implies c (#) f is nonnegative ) & ( c <= 0 implies c (#) f is nonpositive ) ) ) set g = c (#) f; assume A1: f is nonnegative ; ::_thesis: ( ( 0 <= c implies c (#) f is nonnegative ) & ( c <= 0 implies c (#) f is nonpositive ) ) hereby ::_thesis: ( c <= 0 implies c (#) f is nonpositive ) set g = c (#) f; assume A2: 0 <= c ; ::_thesis: c (#) f is nonnegative for x being set st x in dom (c (#) f) holds 0 <= (c (#) f) . x proof let x be set ; ::_thesis: ( x in dom (c (#) f) implies 0 <= (c (#) f) . x ) 0 <= f . x by A1, SUPINF_2:51; then A3: 0 <= (R_EAL c) * (f . x) by A2; assume x in dom (c (#) f) ; ::_thesis: 0 <= (c (#) f) . x hence 0 <= (c (#) f) . x by A3, MESFUNC1:def_6; ::_thesis: verum end; hence c (#) f is nonnegative by SUPINF_2:52; ::_thesis: verum end; assume A4: c <= 0 ; ::_thesis: c (#) f is nonpositive now__::_thesis:_for_x_being_set_st_x_in_dom_(c_(#)_f)_holds_ (c_(#)_f)_._x_<=_0 let x be set ; ::_thesis: ( x in dom (c (#) f) implies (c (#) f) . x <= 0 ) 0 <= f . x by A1, SUPINF_2:51; then A5: (R_EAL c) * (f . x) <= 0 by A4; assume x in dom (c (#) f) ; ::_thesis: (c (#) f) . x <= 0 hence (c (#) f) . x <= 0 by A5, MESFUNC1:def_6; ::_thesis: verum end; hence c (#) f is nonpositive by Th9; ::_thesis: verum end; theorem Th21: :: MESFUNC5:21 for X being non empty set for f, g being PartFunc of X,ExtREAL st ( for x being set st x in (dom f) /\ (dom g) holds ( g . x <= f . x & -infty < g . x & f . x < +infty ) ) holds f - g is nonnegative proof let X be non empty set ; ::_thesis: for f, g being PartFunc of X,ExtREAL st ( for x being set st x in (dom f) /\ (dom g) holds ( g . x <= f . x & -infty < g . x & f . x < +infty ) ) holds f - g is nonnegative let f, g be PartFunc of X,ExtREAL; ::_thesis: ( ( for x being set st x in (dom f) /\ (dom g) holds ( g . x <= f . x & -infty < g . x & f . x < +infty ) ) implies f - g is nonnegative ) assume A1: for x being set st x in (dom f) /\ (dom g) holds ( g . x <= f . x & -infty < g . x & f . x < +infty ) ; ::_thesis: f - g is nonnegative now__::_thesis:_for_x_being_set_st_x_in_dom_(f_-_g)_holds_ 0_<=_(f_-_g)_._x let x be set ; ::_thesis: ( x in dom (f - g) implies 0 <= (f - g) . x ) assume A2: x in dom (f - g) ; ::_thesis: 0 <= (f - g) . x dom (f - g) = ((dom f) /\ (dom g)) \ (((f " {+infty}) /\ (g " {+infty})) \/ ((f " {-infty}) /\ (g " {-infty}))) by MESFUNC1:def_4; then dom (f - g) c= (dom f) /\ (dom g) by XBOOLE_1:36; then 0 <= (f . x) - (g . x) by A1, A2, XXREAL_3:40; hence 0 <= (f - g) . x by A2, MESFUNC1:def_4; ::_thesis: verum end; hence f - g is nonnegative by SUPINF_2:52; ::_thesis: verum end; Lm1: for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL holds ( max+ f is nonnegative & max- f is nonnegative & |.f.| is nonnegative ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL holds ( max+ f is nonnegative & max- f is nonnegative & |.f.| is nonnegative ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL holds ( max+ f is nonnegative & max- f is nonnegative & |.f.| is nonnegative ) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL holds ( max+ f is nonnegative & max- f is nonnegative & |.f.| is nonnegative ) let f be PartFunc of X,ExtREAL; ::_thesis: ( max+ f is nonnegative & max- f is nonnegative & |.f.| is nonnegative ) A1: for x being set st x in dom (max- f) holds 0 <= (max- f) . x by MESFUNC2:13; for x being set st x in dom (max+ f) holds 0 <= (max+ f) . x by MESFUNC2:12; hence ( max+ f is nonnegative & max- f is nonnegative ) by A1, SUPINF_2:52; ::_thesis: |.f.| is nonnegative now__::_thesis:_for_x_being_set_st_x_in_dom_|.f.|_holds_ 0_<=_|.f.|_._x let x be set ; ::_thesis: ( x in dom |.f.| implies 0 <= |.f.| . x ) assume x in dom |.f.| ; ::_thesis: 0 <= |.f.| . x then |.f.| . x = |.(f . x).| by MESFUNC1:def_10; hence 0 <= |.f.| . x by EXTREAL2:3; ::_thesis: verum end; hence |.f.| is nonnegative by SUPINF_2:52; ::_thesis: verum end; theorem Th22: :: MESFUNC5:22 for X being non empty set for f, g being PartFunc of X,ExtREAL st f is nonnegative & g is nonnegative holds ( dom (f + g) = (dom f) /\ (dom g) & f + g is nonnegative ) proof let X be non empty set ; ::_thesis: for f, g being PartFunc of X,ExtREAL st f is nonnegative & g is nonnegative holds ( dom (f + g) = (dom f) /\ (dom g) & f + g is nonnegative ) let f, g be PartFunc of X,ExtREAL; ::_thesis: ( f is nonnegative & g is nonnegative implies ( dom (f + g) = (dom f) /\ (dom g) & f + g is nonnegative ) ) assume that A1: f is nonnegative and A2: g is nonnegative ; ::_thesis: ( dom (f + g) = (dom f) /\ (dom g) & f + g is nonnegative ) thus A3: dom (f + g) = (dom f) /\ (dom g) by A1, A2, Th16; ::_thesis: f + g is nonnegative now__::_thesis:_for_x_being_set_st_x_in_(dom_f)_/\_(dom_g)_holds_ 0_<=_(f_+_g)_._x let x be set ; ::_thesis: ( x in (dom f) /\ (dom g) implies 0 <= (f + g) . x ) assume A4: x in (dom f) /\ (dom g) ; ::_thesis: 0 <= (f + g) . x A5: 0 <= g . x by A2, SUPINF_2:51; 0 <= f . x by A1, SUPINF_2:51; then 0 <= (f . x) + (g . x) by A5; hence 0 <= (f + g) . x by A3, A4, MESFUNC1:def_3; ::_thesis: verum end; hence f + g is nonnegative by A3, SUPINF_2:52; ::_thesis: verum end; theorem Th23: :: MESFUNC5:23 for X being non empty set for f, g, h being PartFunc of X,ExtREAL st f is nonnegative & g is nonnegative & h is nonnegative holds ( dom ((f + g) + h) = ((dom f) /\ (dom g)) /\ (dom h) & (f + g) + h is nonnegative & ( for x being set st x in ((dom f) /\ (dom g)) /\ (dom h) holds ((f + g) + h) . x = ((f . x) + (g . x)) + (h . x) ) ) proof let X be non empty set ; ::_thesis: for f, g, h being PartFunc of X,ExtREAL st f is nonnegative & g is nonnegative & h is nonnegative holds ( dom ((f + g) + h) = ((dom f) /\ (dom g)) /\ (dom h) & (f + g) + h is nonnegative & ( for x being set st x in ((dom f) /\ (dom g)) /\ (dom h) holds ((f + g) + h) . x = ((f . x) + (g . x)) + (h . x) ) ) let f, g, h be PartFunc of X,ExtREAL; ::_thesis: ( f is nonnegative & g is nonnegative & h is nonnegative implies ( dom ((f + g) + h) = ((dom f) /\ (dom g)) /\ (dom h) & (f + g) + h is nonnegative & ( for x being set st x in ((dom f) /\ (dom g)) /\ (dom h) holds ((f + g) + h) . x = ((f . x) + (g . x)) + (h . x) ) ) ) assume that A1: f is nonnegative and A2: g is nonnegative and A3: h is nonnegative ; ::_thesis: ( dom ((f + g) + h) = ((dom f) /\ (dom g)) /\ (dom h) & (f + g) + h is nonnegative & ( for x being set st x in ((dom f) /\ (dom g)) /\ (dom h) holds ((f + g) + h) . x = ((f . x) + (g . x)) + (h . x) ) ) A4: f + g is nonnegative by A1, A2, Th22; then A5: dom ((f + g) + h) = (dom (f + g)) /\ (dom h) by A3, Th22; hence dom ((f + g) + h) = ((dom f) /\ (dom g)) /\ (dom h) by A1, A2, Th22; ::_thesis: ( (f + g) + h is nonnegative & ( for x being set st x in ((dom f) /\ (dom g)) /\ (dom h) holds ((f + g) + h) . x = ((f . x) + (g . x)) + (h . x) ) ) thus (f + g) + h is nonnegative by A3, A4, Th22; ::_thesis: for x being set st x in ((dom f) /\ (dom g)) /\ (dom h) holds ((f + g) + h) . x = ((f . x) + (g . x)) + (h . x) hereby ::_thesis: verum let x be set ; ::_thesis: ( x in ((dom f) /\ (dom g)) /\ (dom h) implies ((f + g) + h) . x = ((f . x) + (g . x)) + (h . x) ) assume x in ((dom f) /\ (dom g)) /\ (dom h) ; ::_thesis: ((f + g) + h) . x = ((f . x) + (g . x)) + (h . x) then A6: x in (dom (f + g)) /\ (dom h) by A1, A2, Th22; then A7: x in dom (f + g) by XBOOLE_0:def_4; thus ((f + g) + h) . x = ((f + g) . x) + (h . x) by A5, A6, MESFUNC1:def_3 .= ((f . x) + (g . x)) + (h . x) by A7, MESFUNC1:def_3 ; ::_thesis: verum end; end; theorem Th24: :: MESFUNC5:24 for X being non empty set for f, g being PartFunc of X,ExtREAL st f is () & g is () holds ( dom ((max+ (f + g)) + (max- f)) = (dom f) /\ (dom g) & dom ((max- (f + g)) + (max+ f)) = (dom f) /\ (dom g) & dom (((max+ (f + g)) + (max- f)) + (max- g)) = (dom f) /\ (dom g) & dom (((max- (f + g)) + (max+ f)) + (max+ g)) = (dom f) /\ (dom g) & (max+ (f + g)) + (max- f) is nonnegative & (max- (f + g)) + (max+ f) is nonnegative ) proof let X be non empty set ; ::_thesis: for f, g being PartFunc of X,ExtREAL st f is () & g is () holds ( dom ((max+ (f + g)) + (max- f)) = (dom f) /\ (dom g) & dom ((max- (f + g)) + (max+ f)) = (dom f) /\ (dom g) & dom (((max+ (f + g)) + (max- f)) + (max- g)) = (dom f) /\ (dom g) & dom (((max- (f + g)) + (max+ f)) + (max+ g)) = (dom f) /\ (dom g) & (max+ (f + g)) + (max- f) is nonnegative & (max- (f + g)) + (max+ f) is nonnegative ) let f, g be PartFunc of X,ExtREAL; ::_thesis: ( f is () & g is () implies ( dom ((max+ (f + g)) + (max- f)) = (dom f) /\ (dom g) & dom ((max- (f + g)) + (max+ f)) = (dom f) /\ (dom g) & dom (((max+ (f + g)) + (max- f)) + (max- g)) = (dom f) /\ (dom g) & dom (((max- (f + g)) + (max+ f)) + (max+ g)) = (dom f) /\ (dom g) & (max+ (f + g)) + (max- f) is nonnegative & (max- (f + g)) + (max+ f) is nonnegative ) ) assume that A1: f is () and A2: g is () ; ::_thesis: ( dom ((max+ (f + g)) + (max- f)) = (dom f) /\ (dom g) & dom ((max- (f + g)) + (max+ f)) = (dom f) /\ (dom g) & dom (((max+ (f + g)) + (max- f)) + (max- g)) = (dom f) /\ (dom g) & dom (((max- (f + g)) + (max+ f)) + (max+ g)) = (dom f) /\ (dom g) & (max+ (f + g)) + (max- f) is nonnegative & (max- (f + g)) + (max+ f) is nonnegative ) A3: dom (f + g) = (dom f) /\ (dom g) by A1, A2, Th16; then A4: dom (max- (f + g)) = (dom f) /\ (dom g) by MESFUNC2:def_3; A5: for x being set holds ( ( x in dom (max- f) implies -infty < (max- f) . x ) & ( x in dom (max+ f) implies -infty < (max+ f) . x ) & ( x in dom (max+ g) implies -infty < (max+ g) . x ) & ( x in dom (max- g) implies -infty < (max- g) . x ) ) by MESFUNC2:12, MESFUNC2:13; then A6: max+ f is () by Th10; A7: max- f is () by A5, Th10; A8: for x being set holds ( ( x in dom (max+ (f + g)) implies -infty < (max+ (f + g)) . x ) & ( x in dom (max- (f + g)) implies -infty < (max- (f + g)) . x ) ) by MESFUNC2:12, MESFUNC2:13; then max+ (f + g) is () by Th10; then A9: dom ((max+ (f + g)) + (max- f)) = (dom (max+ (f + g))) /\ (dom (max- f)) by A7, Th16; max- (f + g) is () by A8, Th10; then A10: dom ((max- (f + g)) + (max+ f)) = (dom (max- (f + g))) /\ (dom (max+ f)) by A6, Th16; A11: max- g is () by A5, Th10; A12: dom (max- f) = dom f by MESFUNC2:def_3; A13: max+ g is () by A5, Th10; A14: dom (max- g) = dom g by MESFUNC2:def_3; A15: dom (max+ f) = dom f by MESFUNC2:def_2; then A16: dom ((max- (f + g)) + (max+ f)) = (dom g) /\ ((dom f) /\ (dom f)) by A4, A10, XBOOLE_1:16; dom (max+ (f + g)) = (dom f) /\ (dom g) by A3, MESFUNC2:def_2; then A17: dom ((max+ (f + g)) + (max- f)) = (dom g) /\ ((dom f) /\ (dom f)) by A12, A9, XBOOLE_1:16; hence ( dom ((max+ (f + g)) + (max- f)) = (dom f) /\ (dom g) & dom ((max- (f + g)) + (max+ f)) = (dom f) /\ (dom g) ) by A4, A15, A10, XBOOLE_1:16; ::_thesis: ( dom (((max+ (f + g)) + (max- f)) + (max- g)) = (dom f) /\ (dom g) & dom (((max- (f + g)) + (max+ f)) + (max+ g)) = (dom f) /\ (dom g) & (max+ (f + g)) + (max- f) is nonnegative & (max- (f + g)) + (max+ f) is nonnegative ) A18: dom (max+ g) = dom g by MESFUNC2:def_2; A19: for x being set holds ( ( x in dom ((max+ (f + g)) + (max- f)) implies 0 <= ((max+ (f + g)) + (max- f)) . x ) & ( x in dom ((max- (f + g)) + (max+ f)) implies 0 <= ((max- (f + g)) + (max+ f)) . x ) ) proof let x be set ; ::_thesis: ( ( x in dom ((max+ (f + g)) + (max- f)) implies 0 <= ((max+ (f + g)) + (max- f)) . x ) & ( x in dom ((max- (f + g)) + (max+ f)) implies 0 <= ((max- (f + g)) + (max+ f)) . x ) ) hereby ::_thesis: ( x in dom ((max- (f + g)) + (max+ f)) implies 0 <= ((max- (f + g)) + (max+ f)) . x ) assume A20: x in dom ((max+ (f + g)) + (max- f)) ; ::_thesis: 0 <= ((max+ (f + g)) + (max- f)) . x then A21: 0 <= (max- f) . x by MESFUNC2:13; 0 <= (max+ (f + g)) . x by A20, MESFUNC2:12; then 0 <= ((max+ (f + g)) . x) + ((max- f) . x) by A21; hence 0 <= ((max+ (f + g)) + (max- f)) . x by A20, MESFUNC1:def_3; ::_thesis: verum end; assume A22: x in dom ((max- (f + g)) + (max+ f)) ; ::_thesis: 0 <= ((max- (f + g)) + (max+ f)) . x then A23: 0 <= (max+ f) . x by MESFUNC2:12; 0 <= (max- (f + g)) . x by A22, MESFUNC2:13; then 0 <= ((max- (f + g)) . x) + ((max+ f) . x) by A23; hence 0 <= ((max- (f + g)) + (max+ f)) . x by A22, MESFUNC1:def_3; ::_thesis: verum end; then A24: for x being set holds ( ( x in dom ((max+ (f + g)) + (max- f)) implies -infty < ((max+ (f + g)) + (max- f)) . x ) & ( x in dom ((max- (f + g)) + (max+ f)) implies -infty < ((max- (f + g)) + (max+ f)) . x ) ) ; then (max+ (f + g)) + (max- f) is () by Th10; then dom (((max+ (f + g)) + (max- f)) + (max- g)) = ((dom f) /\ (dom g)) /\ (dom g) by A14, A11, A17, Th16 .= (dom f) /\ ((dom g) /\ (dom g)) by XBOOLE_1:16 ; hence dom (((max+ (f + g)) + (max- f)) + (max- g)) = (dom f) /\ (dom g) ; ::_thesis: ( dom (((max- (f + g)) + (max+ f)) + (max+ g)) = (dom f) /\ (dom g) & (max+ (f + g)) + (max- f) is nonnegative & (max- (f + g)) + (max+ f) is nonnegative ) (max- (f + g)) + (max+ f) is () by A24, Th10; then dom (((max- (f + g)) + (max+ f)) + (max+ g)) = ((dom f) /\ (dom g)) /\ (dom g) by A18, A13, A16, Th16; then dom (((max- (f + g)) + (max+ f)) + (max+ g)) = (dom f) /\ ((dom g) /\ (dom g)) by XBOOLE_1:16; hence dom (((max- (f + g)) + (max+ f)) + (max+ g)) = (dom f) /\ (dom g) ; ::_thesis: ( (max+ (f + g)) + (max- f) is nonnegative & (max- (f + g)) + (max+ f) is nonnegative ) thus ( (max+ (f + g)) + (max- f) is nonnegative & (max- (f + g)) + (max+ f) is nonnegative ) by A19, SUPINF_2:52; ::_thesis: verum end; theorem Th25: :: MESFUNC5:25 for X being non empty set for f, g being PartFunc of X,ExtREAL st f is () & f is () & g is () & g is () holds ((max+ (f + g)) + (max- f)) + (max- g) = ((max- (f + g)) + (max+ f)) + (max+ g) proof let X be non empty set ; ::_thesis: for f, g being PartFunc of X,ExtREAL st f is () & f is () & g is () & g is () holds ((max+ (f + g)) + (max- f)) + (max- g) = ((max- (f + g)) + (max+ f)) + (max+ g) let f, g be PartFunc of X,ExtREAL; ::_thesis: ( f is () & f is () & g is () & g is () implies ((max+ (f + g)) + (max- f)) + (max- g) = ((max- (f + g)) + (max+ f)) + (max+ g) ) assume that A1: f is () and A2: f is () and A3: g is () and A4: g is () ; ::_thesis: ((max+ (f + g)) + (max- f)) + (max- g) = ((max- (f + g)) + (max+ f)) + (max+ g) A5: dom (max- (f + g)) = dom (f + g) by MESFUNC2:def_3; for x being set st x in dom (max- g) holds 0 <= (max- g) . x by MESFUNC2:13; then A6: max- g is nonnegative by SUPINF_2:52; for x being set st x in dom (max+ g) holds 0 <= (max+ g) . x by MESFUNC2:12; then A7: max+ g is nonnegative by SUPINF_2:52; A8: dom (max- f) = dom f by MESFUNC2:def_3; for x being set st x in dom (max+ (f + g)) holds 0 <= (max+ (f + g)) . x by MESFUNC2:12; then A9: max+ (f + g) is nonnegative by SUPINF_2:52; for x being set st x in dom (max+ f) holds 0 <= (max+ f) . x by MESFUNC2:12; then A10: max+ f is nonnegative by SUPINF_2:52; A11: dom (max+ f) = dom f by MESFUNC2:def_2; A12: dom (max+ g) = dom g by MESFUNC2:def_2; A13: dom (max- g) = dom g by MESFUNC2:def_3; for x being set st x in dom (max- f) holds 0 <= (max- f) . x by MESFUNC2:13; then A14: max- f is nonnegative by SUPINF_2:52; A15: dom (max+ (f + g)) = dom (f + g) by MESFUNC2:def_2; then A16: dom (((max+ (f + g)) + (max- f)) + (max- g)) = ((dom (f + g)) /\ (dom f)) /\ (dom g) by A8, A13, A9, A14, A6, Th23; then A17: dom (((max+ (f + g)) + (max- f)) + (max- g)) = (dom (f + g)) /\ ((dom f) /\ (dom g)) by XBOOLE_1:16; for x being set st x in dom (max- (f + g)) holds 0 <= (max- (f + g)) . x by MESFUNC2:13; then A18: max- (f + g) is nonnegative by SUPINF_2:52; A19: for x being set st x in dom (((max+ (f + g)) + (max- f)) + (max- g)) holds (((max+ (f + g)) + (max- f)) + (max- g)) . x = (((max- (f + g)) + (max+ f)) + (max+ g)) . x proof let x be set ; ::_thesis: ( x in dom (((max+ (f + g)) + (max- f)) + (max- g)) implies (((max+ (f + g)) + (max- f)) + (max- g)) . x = (((max- (f + g)) + (max+ f)) + (max+ g)) . x ) assume A20: x in dom (((max+ (f + g)) + (max- f)) + (max- g)) ; ::_thesis: (((max+ (f + g)) + (max- f)) + (max- g)) . x = (((max- (f + g)) + (max+ f)) + (max+ g)) . x then A21: x in dom g by A16, XBOOLE_0:def_4; then A22: (max+ g) . x = max ((g . x),0) by A12, MESFUNC2:def_2; A23: g . x <> +infty by A4, A21, Th11; A24: dom (f + g) = (dom f) /\ (dom g) by A1, A3, Th16; then A25: (max+ (f + g)) . x = max (((f + g) . x),0) by A15, A17, A20, MESFUNC2:def_2 .= max (((f . x) + (g . x)),0) by A17, A20, A24, MESFUNC1:def_3 ; A26: x in dom f by A17, A20, A24, XBOOLE_0:def_4; then A27: (max+ f) . x = max ((f . x),0) by A11, MESFUNC2:def_2; A28: (max- (f + g)) . x = max ((- ((f + g) . x)),0) by A5, A17, A20, A24, MESFUNC2:def_3 .= max ((- ((f . x) + (g . x))),0) by A17, A20, A24, MESFUNC1:def_3 ; A29: f . x <> -infty by A1, A26, Th10; then A30: - (f . x) <> +infty by XXREAL_3:23; A31: f . x <> +infty by A2, A26, Th11; A32: (max- f) . x = max ((- (f . x)),0) by A8, A26, MESFUNC2:def_3; A33: (max- g) . x = max ((- (g . x)),0) by A13, A21, MESFUNC2:def_3; A34: g . x <> -infty by A3, A21, Th10; then A35: - (g . x) <> +infty by XXREAL_3:23; A36: now__::_thesis:_(((max+_(f_+_g))_._x)_+_((max-_f)_._x))_+_((max-_g)_._x)_=_(((max-_(f_+_g))_._x)_+_((max+_f)_._x))_+_((max+_g)_._x) percases ( 0 <= f . x or f . x < 0 ) ; supposeA37: 0 <= f . x ; ::_thesis: (((max+ (f + g)) . x) + ((max- f) . x)) + ((max- g) . x) = (((max- (f + g)) . x) + ((max+ f) . x)) + ((max+ g) . x) then A38: (max- f) . x = 0 by A32, XXREAL_0:def_10; percases ( 0 <= g . x or g . x < 0 ) ; supposeA39: 0 <= g . x ; ::_thesis: (((max+ (f + g)) . x) + ((max- f) . x)) + ((max- g) . x) = (((max- (f + g)) . x) + ((max+ f) . x)) + ((max+ g) . x) then (max- g) . x = 0 by A33, XXREAL_0:def_10; then A40: (((max+ (f + g)) . x) + ((max- f) . x)) + ((max- g) . x) = (((f . x) + (g . x)) + 0) + 0 by A25, A37, A38, A39, XXREAL_0:def_10 .= ((f . x) + (g . x)) + (R_EAL 0) by XXREAL_3:4 .= (f . x) + (g . x) by XXREAL_3:4 ; A41: (max+ g) . x = g . x by A22, A39, XXREAL_0:def_10; (max- (f + g)) . x = 0 by A28, A37, A39, XXREAL_0:def_10; then (((max- (f + g)) . x) + ((max+ f) . x)) + ((max+ g) . x) = ((R_EAL 0) + (f . x)) + (g . x) by A27, A37, A41, XXREAL_0:def_10; hence (((max+ (f + g)) . x) + ((max- f) . x)) + ((max- g) . x) = (((max- (f + g)) . x) + ((max+ f) . x)) + ((max+ g) . x) by A40, XXREAL_3:4; ::_thesis: verum end; supposeA42: g . x < 0 ; ::_thesis: (((max+ (f + g)) . x) + ((max- f) . x)) + ((max- g) . x) = (((max- (f + g)) . x) + ((max+ f) . x)) + ((max+ g) . x) then A43: (max+ g) . x = 0 by A22, XXREAL_0:def_10; A44: (max- g) . x = - (g . x) by A33, A42, XXREAL_0:def_10; percases ( 0 <= (f . x) + (g . x) or (f . x) + (g . x) < 0 ) ; supposeA45: 0 <= (f . x) + (g . x) ; ::_thesis: (((max+ (f + g)) . x) + ((max- f) . x)) + ((max- g) . x) = (((max- (f + g)) . x) + ((max+ f) . x)) + ((max+ g) . x) then (max- (f + g)) . x = 0 by A28, XXREAL_0:def_10; then A46: (((max- (f + g)) . x) + ((max+ f) . x)) + ((max+ g) . x) = ((R_EAL 0) + (f . x)) + (R_EAL 0) by A27, A37, A43, XXREAL_0:def_10; (((max+ (f + g)) . x) + ((max- f) . x)) + ((max- g) . x) = (((f . x) + (g . x)) + (R_EAL 0)) + (- (g . x)) by A25, A38, A44, A45, XXREAL_0:def_10 .= ((f . x) + (g . x)) - (g . x) by XXREAL_3:4 .= (f . x) + ((g . x) - (g . x)) by A23, A34, XXREAL_3:30 .= (f . x) + (R_EAL 0) by XXREAL_3:7 ; hence (((max+ (f + g)) . x) + ((max- f) . x)) + ((max- g) . x) = (((max- (f + g)) . x) + ((max+ f) . x)) + ((max+ g) . x) by A46, XXREAL_3:4; ::_thesis: verum end; supposeA47: (f . x) + (g . x) < 0 ; ::_thesis: (((max+ (f + g)) . x) + ((max- f) . x)) + ((max- g) . x) = (((max- (f + g)) . x) + ((max+ f) . x)) + ((max+ g) . x) then (((max+ (f + g)) . x) + ((max- f) . x)) + ((max- g) . x) = ((R_EAL 0) + (R_EAL 0)) + (- (g . x)) by A25, A38, A44, XXREAL_0:def_10; then A48: (((max+ (f + g)) . x) + ((max- f) . x)) + ((max- g) . x) = 0 + (- (g . x)) ; (max- (f + g)) . x = - ((f . x) + (g . x)) by A28, A47, XXREAL_0:def_10; then (((max- (f + g)) . x) + ((max+ f) . x)) + ((max+ g) . x) = ((- ((f . x) + (g . x))) + (f . x)) + (R_EAL 0) by A27, A37, A43, XXREAL_0:def_10 .= (- ((f . x) + (g . x))) + (f . x) by XXREAL_3:4 .= ((- (g . x)) - (f . x)) + (f . x) by XXREAL_3:25 .= (- (g . x)) + ((- (f . x)) + (f . x)) by A31, A30, A35, XXREAL_3:29 ; hence (((max+ (f + g)) . x) + ((max- f) . x)) + ((max- g) . x) = (((max- (f + g)) . x) + ((max+ f) . x)) + ((max+ g) . x) by A48, XXREAL_3:7; ::_thesis: verum end; end; end; end; end; supposeA49: f . x < 0 ; ::_thesis: (((max+ (f + g)) . x) + ((max- f) . x)) + ((max- g) . x) = (((max- (f + g)) . x) + ((max+ f) . x)) + ((max+ g) . x) then A50: (max- f) . x = - (f . x) by A32, XXREAL_0:def_10; percases ( 0 <= g . x or g . x < 0 ) ; supposeA51: 0 <= g . x ; ::_thesis: (((max+ (f + g)) . x) + ((max- f) . x)) + ((max- g) . x) = (((max- (f + g)) . x) + ((max+ f) . x)) + ((max+ g) . x) then A52: (max+ g) . x = g . x by A22, XXREAL_0:def_10; A53: (max- g) . x = 0 by A33, A51, XXREAL_0:def_10; percases ( 0 <= (f . x) + (g . x) or (f . x) + (g . x) < 0 ) ; supposeA54: 0 <= (f . x) + (g . x) ; ::_thesis: (((max+ (f + g)) . x) + ((max- f) . x)) + ((max- g) . x) = (((max- (f + g)) . x) + ((max+ f) . x)) + ((max+ g) . x) then (max- (f + g)) . x = 0 by A28, XXREAL_0:def_10; then A55: (((max- (f + g)) . x) + ((max+ f) . x)) + ((max+ g) . x) = ((R_EAL 0) + (R_EAL 0)) + (g . x) by A27, A49, A52, XXREAL_0:def_10; (((max+ (f + g)) . x) + ((max- f) . x)) + ((max- g) . x) = (((f . x) + (g . x)) + (- (f . x))) + (R_EAL 0) by A25, A50, A53, A54, XXREAL_0:def_10 .= ((f . x) + (g . x)) + (- (f . x)) by XXREAL_3:4 .= (g . x) + ((f . x) - (f . x)) by A31, A29, A23, A34, XXREAL_3:29 .= (g . x) + 0 by XXREAL_3:7 ; hence (((max+ (f + g)) . x) + ((max- f) . x)) + ((max- g) . x) = (((max- (f + g)) . x) + ((max+ f) . x)) + ((max+ g) . x) by A55; ::_thesis: verum end; supposeA56: (f . x) + (g . x) < 0 ; ::_thesis: (((max+ (f + g)) . x) + ((max- f) . x)) + ((max- g) . x) = (((max- (f + g)) . x) + ((max+ f) . x)) + ((max+ g) . x) then (max- (f + g)) . x = - ((f . x) + (g . x)) by A28, XXREAL_0:def_10; then A57: (((max- (f + g)) . x) + ((max+ f) . x)) + ((max+ g) . x) = ((- ((f . x) + (g . x))) + (R_EAL 0)) + (g . x) by A27, A49, A52, XXREAL_0:def_10 .= (- ((f . x) + (g . x))) + (g . x) by XXREAL_3:4 .= ((- (f . x)) - (g . x)) + (g . x) by XXREAL_3:25 .= (- (f . x)) + ((- (g . x)) + (g . x)) by A23, A30, A35, XXREAL_3:29 ; (((max+ (f + g)) . x) + ((max- f) . x)) + ((max- g) . x) = ((R_EAL 0) + (- (f . x))) + (R_EAL 0) by A25, A50, A53, A56, XXREAL_0:def_10 .= 0 + (- (f . x)) by XXREAL_3:4 ; hence (((max+ (f + g)) . x) + ((max- f) . x)) + ((max- g) . x) = (((max- (f + g)) . x) + ((max+ f) . x)) + ((max+ g) . x) by A57, XXREAL_3:7; ::_thesis: verum end; end; end; supposeA58: g . x < 0 ; ::_thesis: (((max+ (f + g)) . x) + ((max- f) . x)) + ((max- g) . x) = (((max- (f + g)) . x) + ((max+ f) . x)) + ((max+ g) . x) then (max- g) . x = - (g . x) by A33, XXREAL_0:def_10; then A59: (((max+ (f + g)) . x) + ((max- f) . x)) + ((max- g) . x) = ((R_EAL 0) + (- (f . x))) + (- (g . x)) by A25, A49, A50, A58, XXREAL_0:def_10 .= (- (f . x)) - (g . x) by XXREAL_3:4 ; A60: (max+ g) . x = 0 by A22, A58, XXREAL_0:def_10; (max- (f + g)) . x = - ((f . x) + (g . x)) by A28, A49, A58, XXREAL_0:def_10; then (((max- (f + g)) . x) + ((max+ f) . x)) + ((max+ g) . x) = ((- ((f . x) + (g . x))) + (R_EAL 0)) + (R_EAL 0) by A27, A49, A60, XXREAL_0:def_10 .= (- ((f . x) + (g . x))) + (R_EAL 0) by XXREAL_3:4 .= - ((f . x) + (g . x)) by XXREAL_3:4 ; hence (((max+ (f + g)) . x) + ((max- f) . x)) + ((max- g) . x) = (((max- (f + g)) . x) + ((max+ f) . x)) + ((max+ g) . x) by A59, XXREAL_3:25; ::_thesis: verum end; end; end; end; end; A61: dom (((max+ (f + g)) + (max- f)) + (max- g)) = ((dom (max+ (f + g))) /\ (dom (max- f))) /\ (dom (max- g)) by A9, A14, A6, Th23; (((max- (f + g)) + (max+ f)) + (max+ g)) . x = (((max- (f + g)) . x) + ((max+ f) . x)) + ((max+ g) . x) by A5, A11, A12, A18, A10, A7, A16, A20, Th23; hence (((max+ (f + g)) + (max- f)) + (max- g)) . x = (((max- (f + g)) + (max+ f)) + (max+ g)) . x by A9, A14, A6, A20, A61, A36, Th23; ::_thesis: verum end; dom (((max+ (f + g)) + (max- f)) + (max- g)) = (dom f) /\ (dom g) by A1, A3, Th24; then dom (((max+ (f + g)) + (max- f)) + (max- g)) = dom (((max- (f + g)) + (max+ f)) + (max+ g)) by A1, A3, Th24; hence ((max+ (f + g)) + (max- f)) + (max- g) = ((max- (f + g)) + (max+ f)) + (max+ g) by A19, FUNCT_1:2; ::_thesis: verum end; theorem Th26: :: MESFUNC5:26 for C being non empty set for f being PartFunc of C,ExtREAL for c being Real st 0 <= c holds ( max+ (c (#) f) = c (#) (max+ f) & max- (c (#) f) = c (#) (max- f) ) proof let C be non empty set ; ::_thesis: for f being PartFunc of C,ExtREAL for c being Real st 0 <= c holds ( max+ (c (#) f) = c (#) (max+ f) & max- (c (#) f) = c (#) (max- f) ) let f be PartFunc of C,ExtREAL; ::_thesis: for c being Real st 0 <= c holds ( max+ (c (#) f) = c (#) (max+ f) & max- (c (#) f) = c (#) (max- f) ) let c be Real; ::_thesis: ( 0 <= c implies ( max+ (c (#) f) = c (#) (max+ f) & max- (c (#) f) = c (#) (max- f) ) ) assume A1: 0 <= c ; ::_thesis: ( max+ (c (#) f) = c (#) (max+ f) & max- (c (#) f) = c (#) (max- f) ) A2: dom (max+ (c (#) f)) = dom (c (#) f) by MESFUNC2:def_2 .= dom f by MESFUNC1:def_6 .= dom (max+ f) by MESFUNC2:def_2 .= dom (c (#) (max+ f)) by MESFUNC1:def_6 ; for x being Element of C st x in dom (max+ (c (#) f)) holds (max+ (c (#) f)) . x = (c (#) (max+ f)) . x proof let x be Element of C; ::_thesis: ( x in dom (max+ (c (#) f)) implies (max+ (c (#) f)) . x = (c (#) (max+ f)) . x ) assume A3: x in dom (max+ (c (#) f)) ; ::_thesis: (max+ (c (#) f)) . x = (c (#) (max+ f)) . x then A4: x in dom (c (#) f) by MESFUNC2:def_2; then x in dom f by MESFUNC1:def_6; then A5: x in dom (max+ f) by MESFUNC2:def_2; A6: (max+ (c (#) f)) . x = max (((c (#) f) . x),0) by A3, MESFUNC2:def_2 .= max (((R_EAL c) * (f . x)),0) by A4, MESFUNC1:def_6 ; (c (#) (max+ f)) . x = (R_EAL c) * ((max+ f) . x) by A2, A3, MESFUNC1:def_6 .= c * (max ((f . x),0)) by A5, MESFUNC2:def_2 .= max ((c * (f . x)),(c * 0)) by A1, Th6 ; hence (max+ (c (#) f)) . x = (c (#) (max+ f)) . x by A6; ::_thesis: verum end; hence max+ (c (#) f) = c (#) (max+ f) by A2, PARTFUN1:5; ::_thesis: max- (c (#) f) = c (#) (max- f) A7: dom (max- (c (#) f)) = dom (c (#) f) by MESFUNC2:def_3 .= dom f by MESFUNC1:def_6 .= dom (max- f) by MESFUNC2:def_3 .= dom (c (#) (max- f)) by MESFUNC1:def_6 ; for x being Element of C st x in dom (max- (c (#) f)) holds (max- (c (#) f)) . x = (c (#) (max- f)) . x proof let x be Element of C; ::_thesis: ( x in dom (max- (c (#) f)) implies (max- (c (#) f)) . x = (c (#) (max- f)) . x ) assume A8: x in dom (max- (c (#) f)) ; ::_thesis: (max- (c (#) f)) . x = (c (#) (max- f)) . x then A9: x in dom (c (#) f) by MESFUNC2:def_3; then x in dom f by MESFUNC1:def_6; then A10: x in dom (max- f) by MESFUNC2:def_3; A11: (max- (c (#) f)) . x = max ((- ((c (#) f) . x)),0) by A8, MESFUNC2:def_3 .= max ((- ((R_EAL c) * (f . x))),0) by A9, MESFUNC1:def_6 ; (c (#) (max- f)) . x = (R_EAL c) * ((max- f) . x) by A7, A8, MESFUNC1:def_6 .= c * (max ((- (f . x)),0)) by A10, MESFUNC2:def_3 .= max ((c * (- (f . x))),(c * 0)) by A1, Th6 .= max ((- (c * (f . x))),(c * 0)) by XXREAL_3:92 ; hence (max- (c (#) f)) . x = (c (#) (max- f)) . x by A11; ::_thesis: verum end; hence max- (c (#) f) = c (#) (max- f) by A7, PARTFUN1:5; ::_thesis: verum end; theorem Th27: :: MESFUNC5:27 for C being non empty set for f being PartFunc of C,ExtREAL for c being Real st 0 <= c holds ( max+ ((- c) (#) f) = c (#) (max- f) & max- ((- c) (#) f) = c (#) (max+ f) ) proof let C be non empty set ; ::_thesis: for f being PartFunc of C,ExtREAL for c being Real st 0 <= c holds ( max+ ((- c) (#) f) = c (#) (max- f) & max- ((- c) (#) f) = c (#) (max+ f) ) let f be PartFunc of C,ExtREAL; ::_thesis: for c being Real st 0 <= c holds ( max+ ((- c) (#) f) = c (#) (max- f) & max- ((- c) (#) f) = c (#) (max+ f) ) let c be Real; ::_thesis: ( 0 <= c implies ( max+ ((- c) (#) f) = c (#) (max- f) & max- ((- c) (#) f) = c (#) (max+ f) ) ) assume A1: 0 <= c ; ::_thesis: ( max+ ((- c) (#) f) = c (#) (max- f) & max- ((- c) (#) f) = c (#) (max+ f) ) A2: - (R_EAL c) = - c by SUPINF_2:2; A3: dom (max+ ((- c) (#) f)) = dom ((- c) (#) f) by MESFUNC2:def_2; then dom (max+ ((- c) (#) f)) = dom f by MESFUNC1:def_6; then A4: dom (max+ ((- c) (#) f)) = dom (max- f) by MESFUNC2:def_3; then A5: dom (max+ ((- c) (#) f)) = dom (c (#) (max- f)) by MESFUNC1:def_6; for x being Element of C st x in dom (max+ ((- c) (#) f)) holds (max+ ((- c) (#) f)) . x = (c (#) (max- f)) . x proof let x be Element of C; ::_thesis: ( x in dom (max+ ((- c) (#) f)) implies (max+ ((- c) (#) f)) . x = (c (#) (max- f)) . x ) assume A6: x in dom (max+ ((- c) (#) f)) ; ::_thesis: (max+ ((- c) (#) f)) . x = (c (#) (max- f)) . x then A7: (max+ ((- c) (#) f)) . x = max ((((- c) (#) f) . x),0) by MESFUNC2:def_2 .= max (((R_EAL (- c)) * (f . x)),0) by A3, A6, MESFUNC1:def_6 .= max ((- ((R_EAL c) * (f . x))),0) by A2, XXREAL_3:92 ; (c (#) (max- f)) . x = (R_EAL c) * ((max- f) . x) by A5, A6, MESFUNC1:def_6 .= (R_EAL c) * (max ((- (f . x)),(R_EAL 0))) by A4, A6, MESFUNC2:def_3 .= max (((R_EAL c) * (- (f . x))),((R_EAL c) * (R_EAL 0))) by A1, Th6 .= max ((- ((R_EAL c) * (f . x))),(c * 0)) by XXREAL_3:92 ; hence (max+ ((- c) (#) f)) . x = (c (#) (max- f)) . x by A7; ::_thesis: verum end; hence max+ ((- c) (#) f) = c (#) (max- f) by A5, PARTFUN1:5; ::_thesis: max- ((- c) (#) f) = c (#) (max+ f) A8: dom (max- ((- c) (#) f)) = dom ((- c) (#) f) by MESFUNC2:def_3; then dom (max- ((- c) (#) f)) = dom f by MESFUNC1:def_6; then A9: dom (max- ((- c) (#) f)) = dom (max+ f) by MESFUNC2:def_2; then A10: dom (max- ((- c) (#) f)) = dom (c (#) (max+ f)) by MESFUNC1:def_6; for x being Element of C st x in dom (max- ((- c) (#) f)) holds (max- ((- c) (#) f)) . x = (c (#) (max+ f)) . x proof let x be Element of C; ::_thesis: ( x in dom (max- ((- c) (#) f)) implies (max- ((- c) (#) f)) . x = (c (#) (max+ f)) . x ) assume A11: x in dom (max- ((- c) (#) f)) ; ::_thesis: (max- ((- c) (#) f)) . x = (c (#) (max+ f)) . x then A12: (max- ((- c) (#) f)) . x = max ((- (((- c) (#) f) . x)),0) by MESFUNC2:def_3 .= max ((- ((R_EAL (- c)) * (f . x))),0) by A8, A11, MESFUNC1:def_6 .= max (((- (- (R_EAL c))) * (f . x)),0) by A2, XXREAL_3:92 ; (c (#) (max+ f)) . x = (R_EAL c) * ((max+ f) . x) by A10, A11, MESFUNC1:def_6 .= (R_EAL c) * (max ((f . x),(R_EAL 0))) by A9, A11, MESFUNC2:def_2 .= max (((R_EAL c) * (f . x)),(c * 0)) by A1, Th6 ; hence (max- ((- c) (#) f)) . x = (c (#) (max+ f)) . x by A12; ::_thesis: verum end; hence max- ((- c) (#) f) = c (#) (max+ f) by A10, PARTFUN1:5; ::_thesis: verum end; theorem Th28: :: MESFUNC5:28 for X being non empty set for f being PartFunc of X,ExtREAL for A being set holds ( max+ (f | A) = (max+ f) | A & max- (f | A) = (max- f) | A ) proof let X be non empty set ; ::_thesis: for f being PartFunc of X,ExtREAL for A being set holds ( max+ (f | A) = (max+ f) | A & max- (f | A) = (max- f) | A ) let f be PartFunc of X,ExtREAL; ::_thesis: for A being set holds ( max+ (f | A) = (max+ f) | A & max- (f | A) = (max- f) | A ) let A be set ; ::_thesis: ( max+ (f | A) = (max+ f) | A & max- (f | A) = (max- f) | A ) A1: dom (max+ (f | A)) = dom (f | A) by MESFUNC2:def_2 .= (dom f) /\ A by RELAT_1:61 .= (dom (max+ f)) /\ A by MESFUNC2:def_2 .= dom ((max+ f) | A) by RELAT_1:61 ; for x being Element of X st x in dom (max+ (f | A)) holds (max+ (f | A)) . x = ((max+ f) | A) . x proof let x be Element of X; ::_thesis: ( x in dom (max+ (f | A)) implies (max+ (f | A)) . x = ((max+ f) | A) . x ) assume A2: x in dom (max+ (f | A)) ; ::_thesis: (max+ (f | A)) . x = ((max+ f) | A) . x then A3: ((max+ f) | A) . x = (max+ f) . x by A1, FUNCT_1:47; A4: x in (dom (max+ f)) /\ A by A1, A2, RELAT_1:61; then A5: x in dom (max+ f) by XBOOLE_0:def_4; A6: x in A by A4, XBOOLE_0:def_4; (max+ (f | A)) . x = max (((f | A) . x),0) by A2, MESFUNC2:def_2 .= max ((f . x),0) by A6, FUNCT_1:49 ; hence (max+ (f | A)) . x = ((max+ f) | A) . x by A5, A3, MESFUNC2:def_2; ::_thesis: verum end; hence max+ (f | A) = (max+ f) | A by A1, PARTFUN1:5; ::_thesis: max- (f | A) = (max- f) | A A7: dom (max- (f | A)) = dom (f | A) by MESFUNC2:def_3 .= (dom f) /\ A by RELAT_1:61 .= (dom (max- f)) /\ A by MESFUNC2:def_3 .= dom ((max- f) | A) by RELAT_1:61 ; for x being Element of X st x in dom (max- (f | A)) holds (max- (f | A)) . x = ((max- f) | A) . x proof let x be Element of X; ::_thesis: ( x in dom (max- (f | A)) implies (max- (f | A)) . x = ((max- f) | A) . x ) assume A8: x in dom (max- (f | A)) ; ::_thesis: (max- (f | A)) . x = ((max- f) | A) . x then A9: ((max- f) | A) . x = (max- f) . x by A7, FUNCT_1:47; A10: x in (dom (max- f)) /\ A by A7, A8, RELAT_1:61; then A11: x in dom (max- f) by XBOOLE_0:def_4; A12: x in A by A10, XBOOLE_0:def_4; (max- (f | A)) . x = max ((- ((f | A) . x)),0) by A8, MESFUNC2:def_3 .= max ((- (f . x)),0) by A12, FUNCT_1:49 ; hence (max- (f | A)) . x = ((max- f) | A) . x by A11, A9, MESFUNC2:def_3; ::_thesis: verum end; hence max- (f | A) = (max- f) | A by A7, PARTFUN1:5; ::_thesis: verum end; theorem Th29: :: MESFUNC5:29 for X being non empty set for f, g being PartFunc of X,ExtREAL for B being set st B c= dom (f + g) holds ( dom ((f + g) | B) = B & dom ((f | B) + (g | B)) = B & (f + g) | B = (f | B) + (g | B) ) proof let X be non empty set ; ::_thesis: for f, g being PartFunc of X,ExtREAL for B being set st B c= dom (f + g) holds ( dom ((f + g) | B) = B & dom ((f | B) + (g | B)) = B & (f + g) | B = (f | B) + (g | B) ) let f, g be PartFunc of X,ExtREAL; ::_thesis: for B being set st B c= dom (f + g) holds ( dom ((f + g) | B) = B & dom ((f | B) + (g | B)) = B & (f + g) | B = (f | B) + (g | B) ) let B be set ; ::_thesis: ( B c= dom (f + g) implies ( dom ((f + g) | B) = B & dom ((f | B) + (g | B)) = B & (f + g) | B = (f | B) + (g | B) ) ) assume A1: B c= dom (f + g) ; ::_thesis: ( dom ((f + g) | B) = B & dom ((f | B) + (g | B)) = B & (f + g) | B = (f | B) + (g | B) ) for x being set st x in dom g holds g . x in ExtREAL ; then reconsider gg = g as Function of (dom g),ExtREAL by FUNCT_2:3; for x being set st x in dom (g | B) holds (g | B) . x in ExtREAL ; then reconsider gb = g | B as Function of (dom (g | B)),ExtREAL by FUNCT_2:3; now__::_thesis:_for_x_being_set_st_x_in_(g_"_{+infty})_/\_B_holds_ x_in_(g_|_B)_"_{+infty} let x be set ; ::_thesis: ( x in (g " {+infty}) /\ B implies x in (g | B) " {+infty} ) assume A2: x in (g " {+infty}) /\ B ; ::_thesis: x in (g | B) " {+infty} then A3: x in B by XBOOLE_0:def_4; A4: x in g " {+infty} by A2, XBOOLE_0:def_4; then x in dom gg by FUNCT_2:38; then x in (dom gg) /\ B by A3, XBOOLE_0:def_4; then A5: x in dom (gg | B) by RELAT_1:61; gg . x in {+infty} by A4, FUNCT_2:38; then gb . x in {+infty} by A5, FUNCT_1:47; hence x in (g | B) " {+infty} by A5, FUNCT_2:38; ::_thesis: verum end; then A6: (g " {+infty}) /\ B c= (g | B) " {+infty} by TARSKI:def_3; now__::_thesis:_for_x_being_set_st_x_in_(g_|_B)_"_{+infty}_holds_ x_in_(g_"_{+infty})_/\_B let x be set ; ::_thesis: ( x in (g | B) " {+infty} implies x in (g " {+infty}) /\ B ) assume A7: x in (g | B) " {+infty} ; ::_thesis: x in (g " {+infty}) /\ B then A8: x in dom gb by FUNCT_2:38; then A9: x in (dom g) /\ B by RELAT_1:61; then A10: x in dom g by XBOOLE_0:def_4; gb . x in {+infty} by A7, FUNCT_2:38; then g . x in {+infty} by A8, FUNCT_1:47; then A11: x in gg " {+infty} by A10, FUNCT_2:38; x in B by A9, XBOOLE_0:def_4; hence x in (g " {+infty}) /\ B by A11, XBOOLE_0:def_4; ::_thesis: verum end; then (g | B) " {+infty} c= (g " {+infty}) /\ B by TARSKI:def_3; then A12: (g | B) " {+infty} = (g " {+infty}) /\ B by A6, XBOOLE_0:def_10; now__::_thesis:_for_x_being_set_st_x_in_(g_"_{-infty})_/\_B_holds_ x_in_(g_|_B)_"_{-infty} let x be set ; ::_thesis: ( x in (g " {-infty}) /\ B implies x in (g | B) " {-infty} ) assume A13: x in (g " {-infty}) /\ B ; ::_thesis: x in (g | B) " {-infty} then A14: x in B by XBOOLE_0:def_4; A15: x in g " {-infty} by A13, XBOOLE_0:def_4; then x in dom gg by FUNCT_2:38; then x in (dom gg) /\ B by A14, XBOOLE_0:def_4; then A16: x in dom (gg | B) by RELAT_1:61; gg . x in {-infty} by A15, FUNCT_2:38; then gb . x in {-infty} by A16, FUNCT_1:47; hence x in (g | B) " {-infty} by A16, FUNCT_2:38; ::_thesis: verum end; then A17: (g " {-infty}) /\ B c= (g | B) " {-infty} by TARSKI:def_3; now__::_thesis:_for_x_being_set_st_x_in_(g_|_B)_"_{-infty}_holds_ x_in_(g_"_{-infty})_/\_B let x be set ; ::_thesis: ( x in (g | B) " {-infty} implies x in (g " {-infty}) /\ B ) assume A18: x in (g | B) " {-infty} ; ::_thesis: x in (g " {-infty}) /\ B then A19: x in dom gb by FUNCT_2:38; then A20: x in (dom g) /\ B by RELAT_1:61; then A21: x in dom g by XBOOLE_0:def_4; gb . x in {-infty} by A18, FUNCT_2:38; then g . x in {-infty} by A19, FUNCT_1:47; then A22: x in gg " {-infty} by A21, FUNCT_2:38; x in B by A20, XBOOLE_0:def_4; hence x in (g " {-infty}) /\ B by A22, XBOOLE_0:def_4; ::_thesis: verum end; then (g | B) " {-infty} c= (g " {-infty}) /\ B by TARSKI:def_3; then A23: (g | B) " {-infty} = (g " {-infty}) /\ B by A17, XBOOLE_0:def_10; for x being set st x in dom f holds f . x in ExtREAL ; then reconsider ff = f as Function of (dom f),ExtREAL by FUNCT_2:3; for x being set st x in dom (f | B) holds (f | B) . x in ExtREAL ; then reconsider fb = f | B as Function of (dom (f | B)),ExtREAL by FUNCT_2:3; now__::_thesis:_for_x_being_set_st_x_in_(f_"_{+infty})_/\_B_holds_ x_in_(f_|_B)_"_{+infty} let x be set ; ::_thesis: ( x in (f " {+infty}) /\ B implies x in (f | B) " {+infty} ) assume A24: x in (f " {+infty}) /\ B ; ::_thesis: x in (f | B) " {+infty} then A25: x in B by XBOOLE_0:def_4; A26: x in f " {+infty} by A24, XBOOLE_0:def_4; then x in dom ff by FUNCT_2:38; then x in (dom ff) /\ B by A25, XBOOLE_0:def_4; then A27: x in dom (ff | B) by RELAT_1:61; ff . x in {+infty} by A26, FUNCT_2:38; then fb . x in {+infty} by A27, FUNCT_1:47; hence x in (f | B) " {+infty} by A27, FUNCT_2:38; ::_thesis: verum end; then A28: (f " {+infty}) /\ B c= (f | B) " {+infty} by TARSKI:def_3; now__::_thesis:_for_x_being_set_st_x_in_(f_"_{-infty})_/\_B_holds_ x_in_(f_|_B)_"_{-infty} let x be set ; ::_thesis: ( x in (f " {-infty}) /\ B implies x in (f | B) " {-infty} ) assume A29: x in (f " {-infty}) /\ B ; ::_thesis: x in (f | B) " {-infty} then A30: x in B by XBOOLE_0:def_4; A31: x in f " {-infty} by A29, XBOOLE_0:def_4; then x in dom ff by FUNCT_2:38; then x in (dom ff) /\ B by A30, XBOOLE_0:def_4; then A32: x in dom (ff | B) by RELAT_1:61; ff . x in {-infty} by A31, FUNCT_2:38; then fb . x in {-infty} by A32, FUNCT_1:47; hence x in (f | B) " {-infty} by A32, FUNCT_2:38; ::_thesis: verum end; then A33: (f " {-infty}) /\ B c= (f | B) " {-infty} by TARSKI:def_3; now__::_thesis:_for_x_being_set_st_x_in_(f_|_B)_"_{-infty}_holds_ x_in_(f_"_{-infty})_/\_B let x be set ; ::_thesis: ( x in (f | B) " {-infty} implies x in (f " {-infty}) /\ B ) assume A34: x in (f | B) " {-infty} ; ::_thesis: x in (f " {-infty}) /\ B then A35: x in dom fb by FUNCT_2:38; then A36: x in (dom f) /\ B by RELAT_1:61; then A37: x in dom f by XBOOLE_0:def_4; fb . x in {-infty} by A34, FUNCT_2:38; then f . x in {-infty} by A35, FUNCT_1:47; then A38: x in ff " {-infty} by A37, FUNCT_2:38; x in B by A36, XBOOLE_0:def_4; hence x in (f " {-infty}) /\ B by A38, XBOOLE_0:def_4; ::_thesis: verum end; then (f | B) " {-infty} c= (f " {-infty}) /\ B by TARSKI:def_3; then (f | B) " {-infty} = (f " {-infty}) /\ B by A33, XBOOLE_0:def_10; then A39: ((f | B) " {-infty}) /\ ((g | B) " {+infty}) = (((f " {-infty}) /\ B) /\ (g " {+infty})) /\ B by A12, XBOOLE_1:16 .= (((f " {-infty}) /\ (g " {+infty})) /\ B) /\ B by XBOOLE_1:16 .= ((f " {-infty}) /\ (g " {+infty})) /\ (B /\ B) by XBOOLE_1:16 ; now__::_thesis:_for_x_being_set_st_x_in_(f_|_B)_"_{+infty}_holds_ x_in_(f_"_{+infty})_/\_B let x be set ; ::_thesis: ( x in (f | B) " {+infty} implies x in (f " {+infty}) /\ B ) assume A40: x in (f | B) " {+infty} ; ::_thesis: x in (f " {+infty}) /\ B then A41: x in dom fb by FUNCT_2:38; then A42: x in (dom f) /\ B by RELAT_1:61; then A43: x in dom f by XBOOLE_0:def_4; fb . x in {+infty} by A40, FUNCT_2:38; then f . x in {+infty} by A41, FUNCT_1:47; then A44: x in ff " {+infty} by A43, FUNCT_2:38; x in B by A42, XBOOLE_0:def_4; hence x in (f " {+infty}) /\ B by A44, XBOOLE_0:def_4; ::_thesis: verum end; then (f | B) " {+infty} c= (f " {+infty}) /\ B by TARSKI:def_3; then (f | B) " {+infty} = (f " {+infty}) /\ B by A28, XBOOLE_0:def_10; then ((f | B) " {+infty}) /\ ((g | B) " {-infty}) = (((f " {+infty}) /\ B) /\ (g " {-infty})) /\ B by A23, XBOOLE_1:16 .= (((f " {+infty}) /\ (g " {-infty})) /\ B) /\ B by XBOOLE_1:16 .= ((f " {+infty}) /\ (g " {-infty})) /\ (B /\ B) by XBOOLE_1:16 ; then A45: (((f | B) " {-infty}) /\ ((g | B) " {+infty})) \/ (((f | B) " {+infty}) /\ ((g | B) " {-infty})) = (((f " {-infty}) /\ (g " {+infty})) \/ ((f " {+infty}) /\ (g " {-infty}))) /\ B by A39, XBOOLE_1:23; (dom (f | B)) /\ (dom (g | B)) = ((dom f) /\ B) /\ (dom (g | B)) by RELAT_1:61 .= ((dom f) /\ B) /\ ((dom g) /\ B) by RELAT_1:61 .= (((dom f) /\ B) /\ (dom g)) /\ B by XBOOLE_1:16 .= (((dom f) /\ (dom g)) /\ B) /\ B by XBOOLE_1:16 .= ((dom f) /\ (dom g)) /\ (B /\ B) by XBOOLE_1:16 ; then A46: dom ((f | B) + (g | B)) = (((dom f) /\ (dom g)) /\ B) \ ((((f | B) " {-infty}) /\ ((g | B) " {+infty})) \/ (((f | B) " {+infty}) /\ ((g | B) " {-infty}))) by MESFUNC1:def_3 .= (((dom f) /\ (dom g)) \ (((f " {-infty}) /\ (g " {+infty})) \/ ((f " {+infty}) /\ (g " {-infty})))) /\ B by A45, XBOOLE_1:50 .= (dom (f + g)) /\ B by MESFUNC1:def_3 .= B by A1, XBOOLE_1:28 ; dom (f + g) = ((dom f) /\ (dom g)) \ (((f " {-infty}) /\ (g " {+infty})) \/ ((f " {+infty}) /\ (g " {-infty}))) by MESFUNC1:def_3; then dom (f + g) c= (dom f) /\ (dom g) by XBOOLE_1:36; then A47: B c= (dom f) /\ (dom g) by A1, XBOOLE_1:1; dom (g | B) = (dom g) /\ B by RELAT_1:61; then A48: dom (g | B) = B by A47, XBOOLE_1:18, XBOOLE_1:28; A49: dom ((f + g) | B) = (dom (f + g)) /\ B by RELAT_1:61; then A50: dom ((f + g) | B) = B by A1, XBOOLE_1:28; dom (f | B) = (dom f) /\ B by RELAT_1:61; then A51: dom (f | B) = B by A47, XBOOLE_1:18, XBOOLE_1:28; now__::_thesis:_for_x_being_set_st_x_in_dom_((f_+_g)_|_B)_holds_ ((f_+_g)_|_B)_._x_=_((f_|_B)_+_(g_|_B))_._x let x be set ; ::_thesis: ( x in dom ((f + g) | B) implies ((f + g) | B) . x = ((f | B) + (g | B)) . x ) assume A52: x in dom ((f + g) | B) ; ::_thesis: ((f + g) | B) . x = ((f | B) + (g | B)) . x hence ((f + g) | B) . x = (f + g) . x by FUNCT_1:47 .= (f . x) + (g . x) by A1, A50, A52, MESFUNC1:def_3 .= ((f | B) . x) + (g . x) by A50, A51, A52, FUNCT_1:47 .= ((f | B) . x) + ((g | B) . x) by A50, A48, A52, FUNCT_1:47 .= ((f | B) + (g | B)) . x by A50, A46, A52, MESFUNC1:def_3 ; ::_thesis: verum end; hence ( dom ((f + g) | B) = B & dom ((f | B) + (g | B)) = B & (f + g) | B = (f | B) + (g | B) ) by A1, A49, A46, FUNCT_1:2, XBOOLE_1:28; ::_thesis: verum end; theorem Th30: :: MESFUNC5:30 for X being non empty set for f being PartFunc of X,ExtREAL for a being R_eal holds eq_dom (f,a) = f " {a} proof let X be non empty set ; ::_thesis: for f being PartFunc of X,ExtREAL for a being R_eal holds eq_dom (f,a) = f " {a} let f be PartFunc of X,ExtREAL; ::_thesis: for a being R_eal holds eq_dom (f,a) = f " {a} let a be R_eal; ::_thesis: eq_dom (f,a) = f " {a} now__::_thesis:_for_x_being_set_st_x_in_f_"_{a}_holds_ x_in_eq_dom_(f,a) let x be set ; ::_thesis: ( x in f " {a} implies x in eq_dom (f,a) ) assume A1: x in f " {a} ; ::_thesis: x in eq_dom (f,a) then f . x in {a} by FUNCT_1:def_7; then A2: f . x = a by TARSKI:def_1; x in dom f by A1, FUNCT_1:def_7; hence x in eq_dom (f,a) by A2, MESFUNC1:def_15; ::_thesis: verum end; then A3: f " {a} c= eq_dom (f,a) by TARSKI:def_3; now__::_thesis:_for_x_being_set_st_x_in_eq_dom_(f,a)_holds_ x_in_f_"_{a} let x be set ; ::_thesis: ( x in eq_dom (f,a) implies x in f " {a} ) assume A4: x in eq_dom (f,a) ; ::_thesis: x in f " {a} then f . x = a by MESFUNC1:def_15; then A5: f . x in {a} by TARSKI:def_1; x in dom f by A4, MESFUNC1:def_15; hence x in f " {a} by A5, FUNCT_1:def_7; ::_thesis: verum end; then eq_dom (f,a) c= f " {a} by TARSKI:def_3; hence eq_dom (f,a) = f " {a} by A3, XBOOLE_0:def_10; ::_thesis: verum end; begin theorem Th31: :: MESFUNC5:31 for X being non empty set for S being SigmaField of X for f, g being PartFunc of X,ExtREAL for A being Element of S st f is () & g is () & f is_measurable_on A & g is_measurable_on A holds f + g is_measurable_on A proof let X be non empty set ; ::_thesis: for S being SigmaField of X for f, g being PartFunc of X,ExtREAL for A being Element of S st f is () & g is () & f is_measurable_on A & g is_measurable_on A holds f + g is_measurable_on A let S be SigmaField of X; ::_thesis: for f, g being PartFunc of X,ExtREAL for A being Element of S st f is () & g is () & f is_measurable_on A & g is_measurable_on A holds f + g is_measurable_on A let f, g be PartFunc of X,ExtREAL; ::_thesis: for A being Element of S st f is () & g is () & f is_measurable_on A & g is_measurable_on A holds f + g is_measurable_on A let A be Element of S; ::_thesis: ( f is () & g is () & f is_measurable_on A & g is_measurable_on A implies f + g is_measurable_on A ) assume that A1: f is () and A2: g is () and A3: f is_measurable_on A and A4: g is_measurable_on A ; ::_thesis: f + g is_measurable_on A for r being real number holds A /\ (less_dom ((f + g),(R_EAL r))) in S proof let r be real number ; ::_thesis: A /\ (less_dom ((f + g),(R_EAL r))) in S reconsider r = r as Real by XREAL_0:def_1; consider F being Function of RAT,S such that A5: for p being Rational holds F . p = (A /\ (less_dom (f,(R_EAL p)))) /\ (A /\ (less_dom (g,(R_EAL (r - p))))) by A3, A4, MESFUNC2:6; ex G being Function of NAT,S st rng F = rng G by MESFUNC1:5, MESFUNC2:5; then A6: rng F is N_Sub_set_fam of X by MEASURE1:23; A /\ (less_dom ((f + g),(R_EAL r))) = union (rng F) by A1, A2, A5, Th18; hence A /\ (less_dom ((f + g),(R_EAL r))) in S by A6, MEASURE1:def_5; ::_thesis: verum end; hence f + g is_measurable_on A by MESFUNC1:def_16; ::_thesis: verum end; theorem :: MESFUNC5:32 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_simple_func_in S & dom f = {} holds ex F being Finite_Sep_Sequence of S ex a, x being FinSequence of ExtREAL st ( F,a are_Re-presentation_of f & a . 1 = 0 & ( for n being Nat st 2 <= n & n in dom a holds ( 0 < a . n & a . n < +infty ) ) & dom x = dom F & ( for n being Nat st n in dom x holds x . n = (a . n) * ((M * F) . n) ) & Sum x = 0 ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_simple_func_in S & dom f = {} holds ex F being Finite_Sep_Sequence of S ex a, x being FinSequence of ExtREAL st ( F,a are_Re-presentation_of f & a . 1 = 0 & ( for n being Nat st 2 <= n & n in dom a holds ( 0 < a . n & a . n < +infty ) ) & dom x = dom F & ( for n being Nat st n in dom x holds x . n = (a . n) * ((M * F) . n) ) & Sum x = 0 ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_simple_func_in S & dom f = {} holds ex F being Finite_Sep_Sequence of S ex a, x being FinSequence of ExtREAL st ( F,a are_Re-presentation_of f & a . 1 = 0 & ( for n being Nat st 2 <= n & n in dom a holds ( 0 < a . n & a . n < +infty ) ) & dom x = dom F & ( for n being Nat st n in dom x holds x . n = (a . n) * ((M * F) . n) ) & Sum x = 0 ) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL st f is_simple_func_in S & dom f = {} holds ex F being Finite_Sep_Sequence of S ex a, x being FinSequence of ExtREAL st ( F,a are_Re-presentation_of f & a . 1 = 0 & ( for n being Nat st 2 <= n & n in dom a holds ( 0 < a . n & a . n < +infty ) ) & dom x = dom F & ( for n being Nat st n in dom x holds x . n = (a . n) * ((M * F) . n) ) & Sum x = 0 ) let f be PartFunc of X,ExtREAL; ::_thesis: ( f is_simple_func_in S & dom f = {} implies ex F being Finite_Sep_Sequence of S ex a, x being FinSequence of ExtREAL st ( F,a are_Re-presentation_of f & a . 1 = 0 & ( for n being Nat st 2 <= n & n in dom a holds ( 0 < a . n & a . n < +infty ) ) & dom x = dom F & ( for n being Nat st n in dom x holds x . n = (a . n) * ((M * F) . n) ) & Sum x = 0 ) ) assume that A1: f is_simple_func_in S and A2: dom f = {} ; ::_thesis: ex F being Finite_Sep_Sequence of S ex a, x being FinSequence of ExtREAL st ( F,a are_Re-presentation_of f & a . 1 = 0 & ( for n being Nat st 2 <= n & n in dom a holds ( 0 < a . n & a . n < +infty ) ) & dom x = dom F & ( for n being Nat st n in dom x holds x . n = (a . n) * ((M * F) . n) ) & Sum x = 0 ) for x being set st x in dom f holds 0 <= f . x by A2; then consider F being Finite_Sep_Sequence of S, a being FinSequence of ExtREAL such that A3: F,a are_Re-presentation_of f and A4: a . 1 = 0 and A5: for n being Nat st 2 <= n & n in dom a holds ( 0 < a . n & a . n < +infty ) by A1, MESFUNC3:14; deffunc H1( Nat) -> Element of ExtREAL = (a . $1) * ((M * F) . $1); consider x being FinSequence of ExtREAL such that A6: len x = len F and A7: for n being Nat st n in dom x holds x . n = H1(n) from FINSEQ_2:sch_1(); A8: dom x = Seg (len F) by A6, FINSEQ_1:def_3; then A9: dom x = dom F by FINSEQ_1:def_3; take F ; ::_thesis: ex a, x being FinSequence of ExtREAL st ( F,a are_Re-presentation_of f & a . 1 = 0 & ( for n being Nat st 2 <= n & n in dom a holds ( 0 < a . n & a . n < +infty ) ) & dom x = dom F & ( for n being Nat st n in dom x holds x . n = (a . n) * ((M * F) . n) ) & Sum x = 0 ) take a ; ::_thesis: ex x being FinSequence of ExtREAL st ( F,a are_Re-presentation_of f & a . 1 = 0 & ( for n being Nat st 2 <= n & n in dom a holds ( 0 < a . n & a . n < +infty ) ) & dom x = dom F & ( for n being Nat st n in dom x holds x . n = (a . n) * ((M * F) . n) ) & Sum x = 0 ) take x ; ::_thesis: ( F,a are_Re-presentation_of f & a . 1 = 0 & ( for n being Nat st 2 <= n & n in dom a holds ( 0 < a . n & a . n < +infty ) ) & dom x = dom F & ( for n being Nat st n in dom x holds x . n = (a . n) * ((M * F) . n) ) & Sum x = 0 ) consider sumx being Function of NAT,ExtREAL such that A10: Sum x = sumx . (len x) and A11: sumx . 0 = 0 and A12: for i being Element of NAT st i < len x holds sumx . (i + 1) = (sumx . i) + (x . (i + 1)) by EXTREAL1:def_2; defpred S1[ Nat] means ( $1 <= len x implies sumx . $1 = 0 ); A13: union (rng F) = {} by A2, A3, MESFUNC3:def_1; A14: for n being Nat st n in dom F holds F . n = {} proof let n be Nat; ::_thesis: ( n in dom F implies F . n = {} ) assume n in dom F ; ::_thesis: F . n = {} then A15: F . n in rng F by FUNCT_1:3; assume F . n <> {} ; ::_thesis: contradiction then ex v being set st v in F . n by XBOOLE_0:def_1; hence contradiction by A13, A15, TARSKI:def_4; ::_thesis: verum end; A16: for i being Nat st S1[i] holds S1[i + 1] proof let i be Nat; ::_thesis: ( S1[i] implies S1[i + 1] ) assume A17: S1[i] ; ::_thesis: S1[i + 1] assume A18: i + 1 <= len x ; ::_thesis: sumx . (i + 1) = 0 reconsider i = i as Element of NAT by ORDINAL1:def_12; i < len x by A18, NAT_1:13; then A19: sumx . (i + 1) = (sumx . i) + (x . (i + 1)) by A12; 1 <= i + 1 by NAT_1:11; then A20: i + 1 in dom x by A18, FINSEQ_3:25; then F . (i + 1) = {} by A9, A14; then M . (F . (i + 1)) = 0 by VALUED_0:def_19; then A21: (M * F) . (i + 1) = 0 by A9, A20, FUNCT_1:13; x . (i + 1) = (a . (i + 1)) * ((M * F) . (i + 1)) by A7, A20 .= 0 by A21 ; hence sumx . (i + 1) = 0 by A17, A18, A19, NAT_1:13; ::_thesis: verum end; A22: S1[ 0 ] by A11; for i being Nat holds S1[i] from NAT_1:sch_2(A22, A16); hence ( F,a are_Re-presentation_of f & a . 1 = 0 & ( for n being Nat st 2 <= n & n in dom a holds ( 0 < a . n & a . n < +infty ) ) & dom x = dom F & ( for n being Nat st n in dom x holds x . n = (a . n) * ((M * F) . n) ) & Sum x = 0 ) by A3, A4, A5, A7, A8, A10, FINSEQ_1:def_3; ::_thesis: verum end; theorem Th33: :: MESFUNC5:33 for X being non empty set for S being SigmaField of X for f being PartFunc of X,ExtREAL for A being Element of S for r, s being Real st f is_measurable_on A & A c= dom f holds (A /\ (great_eq_dom (f,(R_EAL r)))) /\ (less_dom (f,(R_EAL s))) in S proof let X be non empty set ; ::_thesis: for S being SigmaField of X for f being PartFunc of X,ExtREAL for A being Element of S for r, s being Real st f is_measurable_on A & A c= dom f holds (A /\ (great_eq_dom (f,(R_EAL r)))) /\ (less_dom (f,(R_EAL s))) in S let S be SigmaField of X; ::_thesis: for f being PartFunc of X,ExtREAL for A being Element of S for r, s being Real st f is_measurable_on A & A c= dom f holds (A /\ (great_eq_dom (f,(R_EAL r)))) /\ (less_dom (f,(R_EAL s))) in S let f be PartFunc of X,ExtREAL; ::_thesis: for A being Element of S for r, s being Real st f is_measurable_on A & A c= dom f holds (A /\ (great_eq_dom (f,(R_EAL r)))) /\ (less_dom (f,(R_EAL s))) in S let A be Element of S; ::_thesis: for r, s being Real st f is_measurable_on A & A c= dom f holds (A /\ (great_eq_dom (f,(R_EAL r)))) /\ (less_dom (f,(R_EAL s))) in S let r, s be Real; ::_thesis: ( f is_measurable_on A & A c= dom f implies (A /\ (great_eq_dom (f,(R_EAL r)))) /\ (less_dom (f,(R_EAL s))) in S ) assume that A1: f is_measurable_on A and A2: A c= dom f ; ::_thesis: (A /\ (great_eq_dom (f,(R_EAL r)))) /\ (less_dom (f,(R_EAL s))) in S A3: A /\ (less_dom (f,(R_EAL s))) in S by A1, MESFUNC1:def_16; A4: (A /\ (great_eq_dom (f,(R_EAL r)))) /\ (A /\ (less_dom (f,(R_EAL s)))) = ((A /\ (great_eq_dom (f,(R_EAL r)))) /\ A) /\ (less_dom (f,(R_EAL s))) by XBOOLE_1:16 .= ((great_eq_dom (f,(R_EAL r))) /\ (A /\ A)) /\ (less_dom (f,(R_EAL s))) by XBOOLE_1:16 ; A /\ (great_eq_dom (f,(R_EAL r))) in S by A1, A2, MESFUNC1:27; hence (A /\ (great_eq_dom (f,(R_EAL r)))) /\ (less_dom (f,(R_EAL s))) in S by A3, A4, FINSUB_1:def_2; ::_thesis: verum end; theorem Th34: :: MESFUNC5:34 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A being Element of S st f is_simple_func_in S holds f | A is_simple_func_in S proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A being Element of S st f is_simple_func_in S holds f | A is_simple_func_in S let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A being Element of S st f is_simple_func_in S holds f | A is_simple_func_in S let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for A being Element of S st f is_simple_func_in S holds f | A is_simple_func_in S let f be PartFunc of X,ExtREAL; ::_thesis: for A being Element of S st f is_simple_func_in S holds f | A is_simple_func_in S let A be Element of S; ::_thesis: ( f is_simple_func_in S implies f | A is_simple_func_in S ) assume A1: f is_simple_func_in S ; ::_thesis: f | A is_simple_func_in S then consider F being Finite_Sep_Sequence of S such that A2: dom f = union (rng F) and A3: for n being Nat for x, y being Element of X st n in dom F & x in F . n & y in F . n holds f . x = f . y by MESFUNC2:def_4; deffunc H1( Nat) -> Element of bool X = (F . $1) /\ A; consider G being FinSequence such that A4: ( len G = len F & ( for n being Nat st n in dom G holds G . n = H1(n) ) ) from FINSEQ_1:sch_2(); A5: rng G c= S proof let P be set ; :: according to TARSKI:def_3 ::_thesis: ( not P in rng G or P in S ) assume P in rng G ; ::_thesis: P in S then consider k being Nat such that A6: k in dom G and A7: P = G . k by FINSEQ_2:10; k in Seg (len F) by A4, A6, FINSEQ_1:def_3; then k in dom F by FINSEQ_1:def_3; then A8: F . k in rng F by FUNCT_1:3; G . k = (F . k) /\ A by A4, A6; hence P in S by A7, A8, FINSUB_1:def_2; ::_thesis: verum end; A9: dom G = Seg (len F) by A4, FINSEQ_1:def_3; reconsider G = G as FinSequence of S by A5, FINSEQ_1:def_4; for i, j being Nat st i in dom G & j in dom G & i <> j holds G . i misses G . j proof let i, j be Nat; ::_thesis: ( i in dom G & j in dom G & i <> j implies G . i misses G . j ) assume that A10: i in dom G and A11: j in dom G and A12: i <> j ; ::_thesis: G . i misses G . j j in Seg (len F) by A4, A11, FINSEQ_1:def_3; then A13: j in dom F by FINSEQ_1:def_3; i in Seg (len F) by A4, A10, FINSEQ_1:def_3; then i in dom F by FINSEQ_1:def_3; then A14: F . i misses F . j by A12, A13, MESFUNC3:4; A15: G . j = (F . j) /\ A by A4, A11; G . i = (F . i) /\ A by A4, A10; then (G . i) /\ (G . j) = (((F . i) /\ A) /\ (F . j)) /\ A by A15, XBOOLE_1:16 .= (((F . i) /\ (F . j)) /\ A) /\ A by XBOOLE_1:16 .= ({} /\ A) /\ A by A14, XBOOLE_0:def_7 ; hence G . i misses G . j by XBOOLE_0:def_7; ::_thesis: verum end; then reconsider G = G as Finite_Sep_Sequence of S by MESFUNC3:4; for v being set st v in union (rng G) holds v in dom (f | A) proof let v be set ; ::_thesis: ( v in union (rng G) implies v in dom (f | A) ) assume v in union (rng G) ; ::_thesis: v in dom (f | A) then consider W being set such that A16: v in W and A17: W in rng G by TARSKI:def_4; consider k being Nat such that A18: k in dom G and A19: W = G . k by A17, FINSEQ_2:10; k in Seg (len F) by A4, A18, FINSEQ_1:def_3; then k in dom F by FINSEQ_1:def_3; then A20: F . k in rng F by FUNCT_1:3; A21: G . k = (F . k) /\ A by A4, A18; then v in F . k by A16, A19, XBOOLE_0:def_4; then A22: v in union (rng F) by A20, TARSKI:def_4; v in A by A16, A19, A21, XBOOLE_0:def_4; then v in (dom f) /\ A by A2, A22, XBOOLE_0:def_4; hence v in dom (f | A) by RELAT_1:61; ::_thesis: verum end; then A23: union (rng G) c= dom (f | A) by TARSKI:def_3; for v being set st v in dom (f | A) holds v in union (rng G) proof let v be set ; ::_thesis: ( v in dom (f | A) implies v in union (rng G) ) assume v in dom (f | A) ; ::_thesis: v in union (rng G) then A24: v in (dom f) /\ A by RELAT_1:61; then A25: v in A by XBOOLE_0:def_4; v in dom f by A24, XBOOLE_0:def_4; then consider W being set such that A26: v in W and A27: W in rng F by A2, TARSKI:def_4; consider k being Nat such that A28: k in dom F and A29: W = F . k by A27, FINSEQ_2:10; A30: k in Seg (len F) by A28, FINSEQ_1:def_3; then k in dom G by A4, FINSEQ_1:def_3; then A31: G . k in rng G by FUNCT_1:3; G . k = (F . k) /\ A by A4, A9, A30; then v in G . k by A25, A26, A29, XBOOLE_0:def_4; hence v in union (rng G) by A31, TARSKI:def_4; ::_thesis: verum end; then dom (f | A) c= union (rng G) by TARSKI:def_3; then A32: dom (f | A) = union (rng G) by A23, XBOOLE_0:def_10; A33: for n being Nat for x, y being Element of X st n in dom G & x in G . n & y in G . n holds (f | A) . x = (f | A) . y proof let n be Nat; ::_thesis: for x, y being Element of X st n in dom G & x in G . n & y in G . n holds (f | A) . x = (f | A) . y let x, y be Element of X; ::_thesis: ( n in dom G & x in G . n & y in G . n implies (f | A) . x = (f | A) . y ) assume that A34: n in dom G and A35: x in G . n and A36: y in G . n ; ::_thesis: (f | A) . x = (f | A) . y A37: G . n in rng G by A34, FUNCT_1:3; then A38: x in dom (f | A) by A32, A35, TARSKI:def_4; A39: G . n = (F . n) /\ A by A4, A34; then A40: y in F . n by A36, XBOOLE_0:def_4; n in Seg (len F) by A4, A34, FINSEQ_1:def_3; then A41: n in dom F by FINSEQ_1:def_3; x in F . n by A35, A39, XBOOLE_0:def_4; then f . x = f . y by A3, A40, A41; then A42: (f | A) . x = f . y by A38, FUNCT_1:47; y in dom (f | A) by A32, A36, A37, TARSKI:def_4; hence (f | A) . x = (f | A) . y by A42, FUNCT_1:47; ::_thesis: verum end; f is V60() by A1, MESFUNC2:def_4; hence f | A is_simple_func_in S by A32, A33, MESFUNC2:def_4; ::_thesis: verum end; theorem Th35: :: MESFUNC5:35 for X being non empty set for S being SigmaField of X for A being Element of S for F being Finite_Sep_Sequence of S for G being FinSequence st dom F = dom G & ( for n being Nat st n in dom F holds G . n = (F . n) /\ A ) holds G is Finite_Sep_Sequence of S proof let X be non empty set ; ::_thesis: for S being SigmaField of X for A being Element of S for F being Finite_Sep_Sequence of S for G being FinSequence st dom F = dom G & ( for n being Nat st n in dom F holds G . n = (F . n) /\ A ) holds G is Finite_Sep_Sequence of S let S be SigmaField of X; ::_thesis: for A being Element of S for F being Finite_Sep_Sequence of S for G being FinSequence st dom F = dom G & ( for n being Nat st n in dom F holds G . n = (F . n) /\ A ) holds G is Finite_Sep_Sequence of S let A be Element of S; ::_thesis: for F being Finite_Sep_Sequence of S for G being FinSequence st dom F = dom G & ( for n being Nat st n in dom F holds G . n = (F . n) /\ A ) holds G is Finite_Sep_Sequence of S let F be Finite_Sep_Sequence of S; ::_thesis: for G being FinSequence st dom F = dom G & ( for n being Nat st n in dom F holds G . n = (F . n) /\ A ) holds G is Finite_Sep_Sequence of S let G be FinSequence; ::_thesis: ( dom F = dom G & ( for n being Nat st n in dom F holds G . n = (F . n) /\ A ) implies G is Finite_Sep_Sequence of S ) assume that A1: dom F = dom G and A2: for n being Nat st n in dom F holds G . n = (F . n) /\ A ; ::_thesis: G is Finite_Sep_Sequence of S rng G c= S proof let v be set ; :: according to TARSKI:def_3 ::_thesis: ( not v in rng G or v in S ) assume v in rng G ; ::_thesis: v in S then consider k being set such that A3: k in dom G and A4: v = G . k by FUNCT_1:def_3; A5: F . k in rng F by A1, A3, FUNCT_1:3; G . k = (F . k) /\ A by A1, A2, A3; hence v in S by A4, A5, FINSUB_1:def_2; ::_thesis: verum end; then reconsider G = G as FinSequence of S by FINSEQ_1:def_4; now__::_thesis:_for_i,_j_being_Nat_st_i_in_dom_G_&_j_in_dom_G_&_i_<>_j_holds_ G_._i_misses_G_._j let i, j be Nat; ::_thesis: ( i in dom G & j in dom G & i <> j implies G . i misses G . j ) assume that A6: i in dom G and A7: j in dom G and A8: i <> j ; ::_thesis: G . i misses G . j A9: F . i misses F . j by A8, PROB_2:def_2; A10: G . j = (F . j) /\ A by A1, A2, A7; G . i = (F . i) /\ A by A1, A2, A6; hence G . i misses G . j by A10, A9, XBOOLE_1:76; ::_thesis: verum end; hence G is Finite_Sep_Sequence of S by MESFUNC3:4; ::_thesis: verum end; theorem Th36: :: MESFUNC5:36 for X being non empty set for S being SigmaField of X for f being PartFunc of X,ExtREAL for A being Element of S for F, G being Finite_Sep_Sequence of S for a being FinSequence of ExtREAL st dom F = dom G & ( for n being Nat st n in dom F holds G . n = (F . n) /\ A ) & F,a are_Re-presentation_of f holds G,a are_Re-presentation_of f | A proof let X be non empty set ; ::_thesis: for S being SigmaField of X for f being PartFunc of X,ExtREAL for A being Element of S for F, G being Finite_Sep_Sequence of S for a being FinSequence of ExtREAL st dom F = dom G & ( for n being Nat st n in dom F holds G . n = (F . n) /\ A ) & F,a are_Re-presentation_of f holds G,a are_Re-presentation_of f | A let S be SigmaField of X; ::_thesis: for f being PartFunc of X,ExtREAL for A being Element of S for F, G being Finite_Sep_Sequence of S for a being FinSequence of ExtREAL st dom F = dom G & ( for n being Nat st n in dom F holds G . n = (F . n) /\ A ) & F,a are_Re-presentation_of f holds G,a are_Re-presentation_of f | A let f be PartFunc of X,ExtREAL; ::_thesis: for A being Element of S for F, G being Finite_Sep_Sequence of S for a being FinSequence of ExtREAL st dom F = dom G & ( for n being Nat st n in dom F holds G . n = (F . n) /\ A ) & F,a are_Re-presentation_of f holds G,a are_Re-presentation_of f | A let A be Element of S; ::_thesis: for F, G being Finite_Sep_Sequence of S for a being FinSequence of ExtREAL st dom F = dom G & ( for n being Nat st n in dom F holds G . n = (F . n) /\ A ) & F,a are_Re-presentation_of f holds G,a are_Re-presentation_of f | A let F, G be Finite_Sep_Sequence of S; ::_thesis: for a being FinSequence of ExtREAL st dom F = dom G & ( for n being Nat st n in dom F holds G . n = (F . n) /\ A ) & F,a are_Re-presentation_of f holds G,a are_Re-presentation_of f | A let a be FinSequence of ExtREAL ; ::_thesis: ( dom F = dom G & ( for n being Nat st n in dom F holds G . n = (F . n) /\ A ) & F,a are_Re-presentation_of f implies G,a are_Re-presentation_of f | A ) assume that A1: dom F = dom G and A2: for n being Nat st n in dom F holds G . n = (F . n) /\ A and A3: F,a are_Re-presentation_of f ; ::_thesis: G,a are_Re-presentation_of f | A A4: dom G = dom a by A1, A3, MESFUNC3:def_1; now__::_thesis:_for_v_being_set_st_v_in_union_(rng_G)_holds_ v_in_dom_(f_|_A) let v be set ; ::_thesis: ( v in union (rng G) implies v in dom (f | A) ) assume v in union (rng G) ; ::_thesis: v in dom (f | A) then consider C being set such that A5: v in C and A6: C in rng G by TARSKI:def_4; consider k being set such that A7: k in dom G and A8: C = G . k by A6, FUNCT_1:def_3; A9: F . k in rng F by A1, A7, FUNCT_1:3; A10: G . k = (F . k) /\ A by A1, A2, A7; then v in F . k by A5, A8, XBOOLE_0:def_4; then v in union (rng F) by A9, TARSKI:def_4; then A11: v in dom f by A3, MESFUNC3:def_1; v in A by A5, A8, A10, XBOOLE_0:def_4; then v in (dom f) /\ A by A11, XBOOLE_0:def_4; hence v in dom (f | A) by RELAT_1:61; ::_thesis: verum end; then A12: union (rng G) c= dom (f | A) by TARSKI:def_3; A13: for k being Nat st k in dom G holds for x being set st x in G . k holds (f | A) . x = a . k proof A14: for k being Nat st k in dom G holds for x being set st x in G . k holds f . x = a . k proof let k be Nat; ::_thesis: ( k in dom G implies for x being set st x in G . k holds f . x = a . k ) assume A15: k in dom G ; ::_thesis: for x being set st x in G . k holds f . x = a . k let x be set ; ::_thesis: ( x in G . k implies f . x = a . k ) assume x in G . k ; ::_thesis: f . x = a . k then x in (F . k) /\ A by A1, A2, A15; then x in F . k by XBOOLE_0:def_4; hence f . x = a . k by A1, A3, A15, MESFUNC3:def_1; ::_thesis: verum end; let k be Nat; ::_thesis: ( k in dom G implies for x being set st x in G . k holds (f | A) . x = a . k ) assume A16: k in dom G ; ::_thesis: for x being set st x in G . k holds (f | A) . x = a . k let x be set ; ::_thesis: ( x in G . k implies (f | A) . x = a . k ) assume A17: x in G . k ; ::_thesis: (f | A) . x = a . k G . k in rng G by A16, FUNCT_1:3; then x in union (rng G) by A17, TARSKI:def_4; then (f | A) . x = f . x by A12, FUNCT_1:47; hence (f | A) . x = a . k by A16, A17, A14; ::_thesis: verum end; now__::_thesis:_for_v_being_set_st_v_in_dom_(f_|_A)_holds_ v_in_union_(rng_G) let v be set ; ::_thesis: ( v in dom (f | A) implies v in union (rng G) ) assume v in dom (f | A) ; ::_thesis: v in union (rng G) then A18: v in (dom f) /\ A by RELAT_1:61; then v in dom f by XBOOLE_0:def_4; then v in union (rng F) by A3, MESFUNC3:def_1; then consider C being set such that A19: v in C and A20: C in rng F by TARSKI:def_4; consider k being Nat such that A21: k in dom F and A22: C = F . k by A20, FINSEQ_2:10; A23: G . k = (F . k) /\ A by A2, A21; A24: G . k in rng G by A1, A21, FUNCT_1:3; v in A by A18, XBOOLE_0:def_4; then v in (F . k) /\ A by A19, A22, XBOOLE_0:def_4; hence v in union (rng G) by A23, A24, TARSKI:def_4; ::_thesis: verum end; then dom (f | A) c= union (rng G) by TARSKI:def_3; then dom (f | A) = union (rng G) by A12, XBOOLE_0:def_10; hence G,a are_Re-presentation_of f | A by A4, A13, MESFUNC3:def_1; ::_thesis: verum end; theorem Th37: :: MESFUNC5:37 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_simple_func_in S holds dom f is Element of S proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_simple_func_in S holds dom f is Element of S let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_simple_func_in S holds dom f is Element of S let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL st f is_simple_func_in S holds dom f is Element of S let f be PartFunc of X,ExtREAL; ::_thesis: ( f is_simple_func_in S implies dom f is Element of S ) assume f is_simple_func_in S ; ::_thesis: dom f is Element of S then ex F being Finite_Sep_Sequence of S st ( dom f = union (rng F) & ( for n being Nat for x, y being Element of X st n in dom F & x in F . n & y in F . n holds f . x = f . y ) ) by MESFUNC2:def_4; hence dom f is Element of S by MESFUNC2:31; ::_thesis: verum end; Lm2: for Y being set for p being FinSequence st ( for i being Nat st i in dom p holds p . i in Y ) holds p is FinSequence of Y proof let Y be set ; ::_thesis: for p being FinSequence st ( for i being Nat st i in dom p holds p . i in Y ) holds p is FinSequence of Y let p be FinSequence; ::_thesis: ( ( for i being Nat st i in dom p holds p . i in Y ) implies p is FinSequence of Y ) assume A1: for i being Nat st i in dom p holds p . i in Y ; ::_thesis: p is FinSequence of Y let b be set ; :: according to TARSKI:def_3,FINSEQ_1:def_4 ::_thesis: ( not b in rng p or b in Y ) assume b in rng p ; ::_thesis: b in Y then ex i being Nat st ( i in dom p & p . i = b ) by FINSEQ_2:10; hence b in Y by A1; ::_thesis: verum end; Lm3: for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & dom f <> {} & g is_simple_func_in S & dom g = dom f holds ( f + g is_simple_func_in S & dom (f + g) <> {} ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & dom f <> {} & g is_simple_func_in S & dom g = dom f holds ( f + g is_simple_func_in S & dom (f + g) <> {} ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & dom f <> {} & g is_simple_func_in S & dom g = dom f holds ( f + g is_simple_func_in S & dom (f + g) <> {} ) let M be sigma_Measure of S; ::_thesis: for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & dom f <> {} & g is_simple_func_in S & dom g = dom f holds ( f + g is_simple_func_in S & dom (f + g) <> {} ) let f, g be PartFunc of X,ExtREAL; ::_thesis: ( f is_simple_func_in S & dom f <> {} & g is_simple_func_in S & dom g = dom f implies ( f + g is_simple_func_in S & dom (f + g) <> {} ) ) assume that A1: f is_simple_func_in S and A2: dom f <> {} and A3: g is_simple_func_in S and A4: dom g = dom f ; ::_thesis: ( f + g is_simple_func_in S & dom (f + g) <> {} ) consider F being Finite_Sep_Sequence of S, a being FinSequence of ExtREAL such that A5: F,a are_Re-presentation_of f by A1, MESFUNC3:12; set la = len F; A6: dom f = union (rng F) by A5, MESFUNC3:def_1; consider G being Finite_Sep_Sequence of S, b being FinSequence of ExtREAL such that A7: G,b are_Re-presentation_of g by A3, MESFUNC3:12; set lb = len G; deffunc H1( Nat) -> set = (F . ((($1 -' 1) div (len G)) + 1)) /\ (G . ((($1 -' 1) mod (len G)) + 1)); consider FG being FinSequence such that A8: len FG = (len F) * (len G) and A9: for k being Nat st k in dom FG holds FG . k = H1(k) from FINSEQ_1:sch_2(); A10: dom FG = Seg ((len F) * (len G)) by A8, FINSEQ_1:def_3; now__::_thesis:_for_k_being_Nat_st_k_in_dom_FG_holds_ FG_._k_in_S reconsider lb9 = len G as Nat ; let k be Nat; ::_thesis: ( k in dom FG implies FG . k in S ) set i = ((k -' 1) div (len G)) + 1; set j = ((k -' 1) mod (len G)) + 1; A11: lb9 divides (len F) * (len G) by NAT_D:def_3; assume A12: k in dom FG ; ::_thesis: FG . k in S then A13: k in Seg ((len F) * (len G)) by A8, FINSEQ_1:def_3; then A14: k <= (len F) * (len G) by FINSEQ_1:1; then k -' 1 <= ((len F) * (len G)) -' 1 by NAT_D:42; then A15: (k -' 1) div (len G) <= (((len F) * (len G)) -' 1) div (len G) by NAT_2:24; 1 <= k by A13, FINSEQ_1:1; then A16: 1 <= (len F) * (len G) by A14, XXREAL_0:2; A17: len G <> 0 by A13; then (k -' 1) mod (len G) < len G by NAT_D:1; then A18: ((k -' 1) mod (len G)) + 1 <= len G by NAT_1:13; len G >= 0 + 1 by A17, NAT_1:13; then (((len F) * (len G)) -' 1) div (len G) = (((len F) * (len G)) div (len G)) - 1 by A11, A16, NAT_2:15; then ((k -' 1) div (len G)) + 1 <= ((len F) * (len G)) div (len G) by A15, XREAL_1:19; then A19: ((k -' 1) div (len G)) + 1 <= len F by A17, NAT_D:18; ((k -' 1) div (len G)) + 1 >= 0 + 1 by XREAL_1:6; then ((k -' 1) div (len G)) + 1 in Seg (len F) by A19; then ((k -' 1) div (len G)) + 1 in dom F by FINSEQ_1:def_3; then A20: F . (((k -' 1) div (len G)) + 1) in rng F by FUNCT_1:3; ((k -' 1) mod (len G)) + 1 >= 0 + 1 by XREAL_1:6; then ((k -' 1) mod (len G)) + 1 in dom G by A18, FINSEQ_3:25; then A21: G . (((k -' 1) mod (len G)) + 1) in rng G by FUNCT_1:3; FG . k = (F . (((k -' 1) div (len G)) + 1)) /\ (G . (((k -' 1) mod (len G)) + 1)) by A9, A12; hence FG . k in S by A20, A21, MEASURE1:34; ::_thesis: verum end; then reconsider FG = FG as FinSequence of S by Lm2; A22: for k, l being Nat st k in dom FG & l in dom FG & k <> l holds FG . k misses FG . l proof A23: len G divides (len F) * (len G) by NAT_D:def_3; let k, l be Nat; ::_thesis: ( k in dom FG & l in dom FG & k <> l implies FG . k misses FG . l ) assume that A24: k in dom FG and A25: l in dom FG and A26: k <> l ; ::_thesis: FG . k misses FG . l A27: k in Seg ((len F) * (len G)) by A8, A24, FINSEQ_1:def_3; then A28: 1 <= k by FINSEQ_1:1; set m = ((l -' 1) mod (len G)) + 1; set n = ((l -' 1) div (len G)) + 1; set j = ((k -' 1) mod (len G)) + 1; set i = ((k -' 1) div (len G)) + 1; A29: FG . k = (F . (((k -' 1) div (len G)) + 1)) /\ (G . (((k -' 1) mod (len G)) + 1)) by A9, A24; A30: k <= (len F) * (len G) by A27, FINSEQ_1:1; then A31: 1 <= (len F) * (len G) by A28, XXREAL_0:2; A32: len G <> 0 by A27; then len G >= 0 + 1 by NAT_1:13; then A33: (((len F) * (len G)) -' 1) div (len G) = (((len F) * (len G)) div (len G)) - 1 by A23, A31, NAT_2:15; A34: l in Seg ((len F) * (len G)) by A8, A25, FINSEQ_1:def_3; then A35: 1 <= l by FINSEQ_1:1; A36: now__::_thesis:_(_((k_-'_1)_div_(len_G))_+_1_=_((l_-'_1)_div_(len_G))_+_1_implies_not_((k_-'_1)_mod_(len_G))_+_1_=_((l_-'_1)_mod_(len_G))_+_1_) (l -' 1) + 1 = ((((((l -' 1) div (len G)) + 1) - 1) * (len G)) + ((((l -' 1) mod (len G)) + 1) - 1)) + 1 by A32, NAT_D:2; then A37: (l - 1) + 1 = (((((l -' 1) div (len G)) + 1) - 1) * (len G)) + (((l -' 1) mod (len G)) + 1) by A35, XREAL_1:233; assume that A38: ((k -' 1) div (len G)) + 1 = ((l -' 1) div (len G)) + 1 and A39: ((k -' 1) mod (len G)) + 1 = ((l -' 1) mod (len G)) + 1 ; ::_thesis: contradiction (k -' 1) + 1 = ((((((k -' 1) div (len G)) + 1) - 1) * (len G)) + ((((k -' 1) mod (len G)) + 1) - 1)) + 1 by A32, NAT_D:2; then (k - 1) + 1 = (((((k -' 1) div (len G)) + 1) - 1) * (len G)) + (((k -' 1) mod (len G)) + 1) by A28, XREAL_1:233; hence contradiction by A26, A38, A39, A37; ::_thesis: verum end; k -' 1 <= ((len F) * (len G)) -' 1 by A30, NAT_D:42; then (k -' 1) div (len G) <= (((len F) * (len G)) div (len G)) - 1 by A33, NAT_2:24; then ((k -' 1) div (len G)) + 1 <= ((len F) * (len G)) div (len G) by XREAL_1:19; then A40: ((k -' 1) div (len G)) + 1 <= len F by A32, NAT_D:18; ((k -' 1) div (len G)) + 1 >= 0 + 1 by XREAL_1:6; then ((k -' 1) div (len G)) + 1 in Seg (len F) by A40; then A41: ((k -' 1) div (len G)) + 1 in dom F by FINSEQ_1:def_3; A42: ((k -' 1) mod (len G)) + 1 >= 0 + 1 by XREAL_1:6; (k -' 1) mod (len G) < len G by A32, NAT_D:1; then ((k -' 1) mod (len G)) + 1 <= len G by NAT_1:13; then A43: ((k -' 1) mod (len G)) + 1 in dom G by A42, FINSEQ_3:25; A44: ((l -' 1) mod (len G)) + 1 >= 0 + 1 by XREAL_1:6; (l -' 1) mod (len G) < len G by A32, NAT_D:1; then ((l -' 1) mod (len G)) + 1 <= len G by NAT_1:13; then A45: ((l -' 1) mod (len G)) + 1 in dom G by A44, FINSEQ_3:25; A46: ((l -' 1) div (len G)) + 1 >= 0 + 1 by XREAL_1:6; l <= (len F) * (len G) by A34, FINSEQ_1:1; then l -' 1 <= ((len F) * (len G)) -' 1 by NAT_D:42; then (l -' 1) div (len G) <= (((len F) * (len G)) div (len G)) - 1 by A33, NAT_2:24; then ((l -' 1) div (len G)) + 1 <= ((len F) * (len G)) div (len G) by XREAL_1:19; then ((l -' 1) div (len G)) + 1 <= len F by A32, NAT_D:18; then ((l -' 1) div (len G)) + 1 in Seg (len F) by A46; then A47: ((l -' 1) div (len G)) + 1 in dom F by FINSEQ_1:def_3; percases ( ((k -' 1) div (len G)) + 1 <> ((l -' 1) div (len G)) + 1 or ((k -' 1) mod (len G)) + 1 <> ((l -' 1) mod (len G)) + 1 ) by A36; supposeA48: ((k -' 1) div (len G)) + 1 <> ((l -' 1) div (len G)) + 1 ; ::_thesis: FG . k misses FG . l (FG . k) /\ (FG . l) = ((F . (((k -' 1) div (len G)) + 1)) /\ (G . (((k -' 1) mod (len G)) + 1))) /\ ((F . (((l -' 1) div (len G)) + 1)) /\ (G . (((l -' 1) mod (len G)) + 1))) by A9, A25, A29; then (FG . k) /\ (FG . l) = (F . (((k -' 1) div (len G)) + 1)) /\ ((G . (((k -' 1) mod (len G)) + 1)) /\ ((F . (((l -' 1) div (len G)) + 1)) /\ (G . (((l -' 1) mod (len G)) + 1)))) by XBOOLE_1:16; then (FG . k) /\ (FG . l) = (F . (((k -' 1) div (len G)) + 1)) /\ ((F . (((l -' 1) div (len G)) + 1)) /\ ((G . (((k -' 1) mod (len G)) + 1)) /\ (G . (((l -' 1) mod (len G)) + 1)))) by XBOOLE_1:16; then A49: (FG . k) /\ (FG . l) = ((F . (((k -' 1) div (len G)) + 1)) /\ (F . (((l -' 1) div (len G)) + 1))) /\ ((G . (((k -' 1) mod (len G)) + 1)) /\ (G . (((l -' 1) mod (len G)) + 1))) by XBOOLE_1:16; F . (((k -' 1) div (len G)) + 1) misses F . (((l -' 1) div (len G)) + 1) by A41, A47, A48, MESFUNC3:4; then (FG . k) /\ (FG . l) = {} /\ ((G . (((k -' 1) mod (len G)) + 1)) /\ (G . (((l -' 1) mod (len G)) + 1))) by A49, XBOOLE_0:def_7; hence FG . k misses FG . l by XBOOLE_0:def_7; ::_thesis: verum end; supposeA50: ((k -' 1) mod (len G)) + 1 <> ((l -' 1) mod (len G)) + 1 ; ::_thesis: FG . k misses FG . l (FG . k) /\ (FG . l) = ((F . (((k -' 1) div (len G)) + 1)) /\ (G . (((k -' 1) mod (len G)) + 1))) /\ ((F . (((l -' 1) div (len G)) + 1)) /\ (G . (((l -' 1) mod (len G)) + 1))) by A9, A25, A29; then (FG . k) /\ (FG . l) = (F . (((k -' 1) div (len G)) + 1)) /\ ((G . (((k -' 1) mod (len G)) + 1)) /\ ((F . (((l -' 1) div (len G)) + 1)) /\ (G . (((l -' 1) mod (len G)) + 1)))) by XBOOLE_1:16; then (FG . k) /\ (FG . l) = (F . (((k -' 1) div (len G)) + 1)) /\ ((F . (((l -' 1) div (len G)) + 1)) /\ ((G . (((k -' 1) mod (len G)) + 1)) /\ (G . (((l -' 1) mod (len G)) + 1)))) by XBOOLE_1:16; then A51: (FG . k) /\ (FG . l) = ((F . (((k -' 1) div (len G)) + 1)) /\ (F . (((l -' 1) div (len G)) + 1))) /\ ((G . (((k -' 1) mod (len G)) + 1)) /\ (G . (((l -' 1) mod (len G)) + 1))) by XBOOLE_1:16; G . (((k -' 1) mod (len G)) + 1) misses G . (((l -' 1) mod (len G)) + 1) by A43, A45, A50, MESFUNC3:4; then (FG . k) /\ (FG . l) = ((F . (((k -' 1) div (len G)) + 1)) /\ (F . (((l -' 1) div (len G)) + 1))) /\ {} by A51, XBOOLE_0:def_7; hence FG . k misses FG . l by XBOOLE_0:def_7; ::_thesis: verum end; end; end; A52: g is V60() by A3, MESFUNC2:def_4; then A53: dom (f + g) = (dom f) /\ (dom g) by MESFUNC2:2; reconsider FG = FG as Finite_Sep_Sequence of S by A22, MESFUNC3:4; A54: dom g = union (rng G) by A7, MESFUNC3:def_1; A55: dom f = union (rng FG) proof now__::_thesis:_for_z_being_set_st_z_in_dom_f_holds_ z_in_union_(rng_FG) let z be set ; ::_thesis: ( z in dom f implies z in union (rng FG) ) assume A56: z in dom f ; ::_thesis: z in union (rng FG) then consider Y being set such that A57: z in Y and A58: Y in rng F by A6, TARSKI:def_4; consider i being Nat such that A59: i in dom F and A60: Y = F . i by A58, FINSEQ_2:10; A61: i in Seg (len F) by A59, FINSEQ_1:def_3; then 1 <= i by FINSEQ_1:1; then consider i9 being Nat such that A62: i = 1 + i9 by NAT_1:10; consider Z being set such that A63: z in Z and A64: Z in rng G by A4, A54, A56, TARSKI:def_4; consider j being Nat such that A65: j in dom G and A66: Z = G . j by A64, FINSEQ_2:10; A67: j in Seg (len G) by A65, FINSEQ_1:def_3; then A68: 1 <= j by FINSEQ_1:1; then consider j9 being Nat such that A69: j = 1 + j9 by NAT_1:10; (i9 * (len G)) + j is Nat ; then reconsider k = ((i - 1) * (len G)) + j as Element of NAT by A62; i <= len F by A61, FINSEQ_1:1; then i - 1 <= (len F) - 1 by XREAL_1:9; then (i - 1) * (len G) <= ((len F) - 1) * (len G) by XREAL_1:64; then A70: k <= (((len F) - 1) * (len G)) + j by XREAL_1:6; A71: j <= len G by A67, FINSEQ_1:1; then A72: j9 < len G by A69, NAT_1:13; A73: k >= 0 + j by A62, XREAL_1:6; then A74: k -' 1 = k - 1 by A68, XREAL_1:233, XXREAL_0:2 .= (i9 * (len G)) + j9 by A62, A69 ; then A75: i = ((k -' 1) div (len G)) + 1 by A62, A72, NAT_D:def_1; (((len F) - 1) * (len G)) + j <= (((len F) - 1) * (len G)) + (len G) by A71, XREAL_1:6; then A76: k <= (len F) * (len G) by A70, XXREAL_0:2; k >= 1 by A68, A73, XXREAL_0:2; then A77: k in Seg ((len F) * (len G)) by A76; then k in dom FG by A8, FINSEQ_1:def_3; then A78: FG . k in rng FG by FUNCT_1:def_3; A79: j = ((k -' 1) mod (len G)) + 1 by A69, A74, A72, NAT_D:def_2; z in (F . i) /\ (G . j) by A57, A60, A63, A66, XBOOLE_0:def_4; then z in FG . k by A9, A10, A75, A79, A77; hence z in union (rng FG) by A78, TARSKI:def_4; ::_thesis: verum end; hence dom f c= union (rng FG) by TARSKI:def_3; :: according to XBOOLE_0:def_10 ::_thesis: union (rng FG) c= dom f reconsider lb9 = len G as Nat ; let z be set ; :: according to TARSKI:def_3 ::_thesis: ( not z in union (rng FG) or z in dom f ) A80: lb9 divides (len F) * (len G) by NAT_D:def_3; assume z in union (rng FG) ; ::_thesis: z in dom f then consider Y being set such that A81: z in Y and A82: Y in rng FG by TARSKI:def_4; consider k being Nat such that A83: k in dom FG and A84: Y = FG . k by A82, FINSEQ_2:10; A85: k in Seg (len FG) by A83, FINSEQ_1:def_3; then A86: k <= (len F) * (len G) by A8, FINSEQ_1:1; then A87: k -' 1 <= ((len F) * (len G)) -' 1 by NAT_D:42; set j = ((k -' 1) mod (len G)) + 1; set i = ((k -' 1) div (len G)) + 1; A88: ((k -' 1) div (len G)) + 1 >= 0 + 1 by NAT_1:13; 1 <= k by A85, FINSEQ_1:1; then A89: 1 <= (len F) * (len G) by A86, XXREAL_0:2; A90: len G <> 0 by A8, A85; then len G >= 0 + 1 by NAT_1:13; then (((len F) * (len G)) -' 1) div lb9 = (((len F) * (len G)) div (len G)) - 1 by A80, A89, NAT_2:15; then (k -' 1) div (len G) <= (((len F) * (len G)) div (len G)) - 1 by A87, NAT_2:24; then A91: ((k -' 1) div (len G)) + 1 <= ((len F) * (len G)) div (len G) by XREAL_1:19; ((len F) * (len G)) div (len G) = len F by A90, NAT_D:18; then ((k -' 1) div (len G)) + 1 in Seg (len F) by A91, A88; then ((k -' 1) div (len G)) + 1 in dom F by FINSEQ_1:def_3; then A92: F . (((k -' 1) div (len G)) + 1) in rng F by FUNCT_1:def_3; FG . k = (F . (((k -' 1) div (len G)) + 1)) /\ (G . (((k -' 1) mod (len G)) + 1)) by A9, A83; then z in F . (((k -' 1) div (len G)) + 1) by A81, A84, XBOOLE_0:def_4; hence z in dom f by A6, A92, TARSKI:def_4; ::_thesis: verum end; A93: for k being Nat for x, y being Element of X st k in dom FG & x in FG . k & y in FG . k holds (f + g) . x = (f + g) . y proof A94: len G divides (len F) * (len G) by NAT_D:def_3; let k be Nat; ::_thesis: for x, y being Element of X st k in dom FG & x in FG . k & y in FG . k holds (f + g) . x = (f + g) . y let x, y be Element of X; ::_thesis: ( k in dom FG & x in FG . k & y in FG . k implies (f + g) . x = (f + g) . y ) assume that A95: k in dom FG and A96: x in FG . k and A97: y in FG . k ; ::_thesis: (f + g) . x = (f + g) . y set j = ((k -' 1) mod (len G)) + 1; A98: FG . k = (F . (((k -' 1) div (len G)) + 1)) /\ (G . (((k -' 1) mod (len G)) + 1)) by A9, A95; then A99: y in G . (((k -' 1) mod (len G)) + 1) by A97, XBOOLE_0:def_4; set i = ((k -' 1) div (len G)) + 1; A100: ((k -' 1) div (len G)) + 1 >= 0 + 1 by XREAL_1:6; A101: k in Seg (len FG) by A95, FINSEQ_1:def_3; then A102: 1 <= k by FINSEQ_1:1; A103: len G > 0 by A8, A101; then A104: len G >= 0 + 1 by NAT_1:13; A105: k <= (len F) * (len G) by A8, A101, FINSEQ_1:1; then A106: k -' 1 <= ((len F) * (len G)) -' 1 by NAT_D:42; 1 <= (len F) * (len G) by A102, A105, XXREAL_0:2; then (((len F) * (len G)) -' 1) div (len G) = (((len F) * (len G)) div (len G)) - 1 by A104, A94, NAT_2:15; then (k -' 1) div (len G) <= (((len F) * (len G)) div (len G)) - 1 by A106, NAT_2:24; then A107: ((k -' 1) div (len G)) + 1 <= ((len F) * (len G)) div (len G) by XREAL_1:19; len G <> 0 by A8, A101; then ((k -' 1) div (len G)) + 1 <= len F by A107, NAT_D:18; then ((k -' 1) div (len G)) + 1 in Seg (len F) by A100; then A108: ((k -' 1) div (len G)) + 1 in dom F by FINSEQ_1:def_3; x in F . (((k -' 1) div (len G)) + 1) by A96, A98, XBOOLE_0:def_4; then A109: f . x = a . (((k -' 1) div (len G)) + 1) by A5, A108, MESFUNC3:def_1; A110: ((k -' 1) mod (len G)) + 1 >= 0 + 1 by XREAL_1:6; (k -' 1) mod (len G) < len G by A103, NAT_D:1; then ((k -' 1) mod (len G)) + 1 <= len G by NAT_1:13; then ((k -' 1) mod (len G)) + 1 in Seg (len G) by A110; then A111: ((k -' 1) mod (len G)) + 1 in dom G by FINSEQ_1:def_3; y in F . (((k -' 1) div (len G)) + 1) by A97, A98, XBOOLE_0:def_4; then A112: f . y = a . (((k -' 1) div (len G)) + 1) by A5, A108, MESFUNC3:def_1; A113: FG . k in rng FG by A95, FUNCT_1:def_3; then x in dom (f + g) by A4, A55, A53, A96, TARSKI:def_4; then A114: (f + g) . x = (f . x) + (g . x) by MESFUNC1:def_3; x in G . (((k -' 1) mod (len G)) + 1) by A96, A98, XBOOLE_0:def_4; then (f + g) . x = (a . (((k -' 1) div (len G)) + 1)) + (b . (((k -' 1) mod (len G)) + 1)) by A7, A109, A111, A114, MESFUNC3:def_1; then A115: (f + g) . x = (f . y) + (g . y) by A7, A99, A111, A112, MESFUNC3:def_1; y in dom (f + g) by A4, A55, A53, A97, A113, TARSKI:def_4; hence (f + g) . x = (f + g) . y by A115, MESFUNC1:def_3; ::_thesis: verum end; now__::_thesis:_for_x_being_Element_of_X_st_x_in_dom_(f_+_g)_holds_ |.((f_+_g)_._x).|_<_+infty let x be Element of X; ::_thesis: ( x in dom (f + g) implies |.((f + g) . x).| < +infty ) assume A116: x in dom (f + g) ; ::_thesis: |.((f + g) . x).| < +infty then A117: |.(g . x).| < +infty by A4, A52, A53, MESFUNC2:def_1; |.((f + g) . x).| = |.((f . x) + (g . x)).| by A116, MESFUNC1:def_3; then A118: |.((f + g) . x).| <= |.(f . x).| + |.(g . x).| by EXTREAL2:13; f is V60() by A1, MESFUNC2:def_4; then |.(f . x).| < +infty by A4, A53, A116, MESFUNC2:def_1; then |.(f . x).| + |.(g . x).| <> +infty by A117, XXREAL_3:16; hence |.((f + g) . x).| < +infty by A118, XXREAL_0:2, XXREAL_0:4; ::_thesis: verum end; then f + g is V60() by MESFUNC2:def_1; hence f + g is_simple_func_in S by A4, A55, A53, A93, MESFUNC2:def_4; ::_thesis: dom (f + g) <> {} thus dom (f + g) <> {} by A2, A4, A53; ::_thesis: verum end; theorem Th38: :: MESFUNC5:38 for X being non empty set for S being SigmaField of X for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & g is_simple_func_in S holds f + g is_simple_func_in S proof let X be non empty set ; ::_thesis: for S being SigmaField of X for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & g is_simple_func_in S holds f + g is_simple_func_in S let S be SigmaField of X; ::_thesis: for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & g is_simple_func_in S holds f + g is_simple_func_in S let f, g be PartFunc of X,ExtREAL; ::_thesis: ( f is_simple_func_in S & g is_simple_func_in S implies f + g is_simple_func_in S ) assume that A1: f is_simple_func_in S and A2: g is_simple_func_in S ; ::_thesis: f + g is_simple_func_in S percases ( dom (f + g) = {} or dom (f + g) <> {} ) ; supposeA3: dom (f + g) = {} ; ::_thesis: f + g is_simple_func_in S reconsider EMPTY = {} as Element of S by PROB_1:4; set F = <*EMPTY*>; A4: dom <*EMPTY*> = Seg 1 by FINSEQ_1:38; A5: now__::_thesis:_for_i,_j_being_Nat_st_i_in_dom_<*EMPTY*>_&_j_in_dom_<*EMPTY*>_&_i_<>_j_holds_ <*EMPTY*>_._i_misses_<*EMPTY*>_._j let i, j be Nat; ::_thesis: ( i in dom <*EMPTY*> & j in dom <*EMPTY*> & i <> j implies <*EMPTY*> . i misses <*EMPTY*> . j ) assume that A6: i in dom <*EMPTY*> and A7: j in dom <*EMPTY*> and A8: i <> j ; ::_thesis: <*EMPTY*> . i misses <*EMPTY*> . j i = 1 by A4, A6, FINSEQ_1:2, TARSKI:def_1; hence <*EMPTY*> . i misses <*EMPTY*> . j by A4, A7, A8, FINSEQ_1:2, TARSKI:def_1; ::_thesis: verum end; A9: for n being Nat st n in dom <*EMPTY*> holds <*EMPTY*> . n = EMPTY proof let n be Nat; ::_thesis: ( n in dom <*EMPTY*> implies <*EMPTY*> . n = EMPTY ) assume n in dom <*EMPTY*> ; ::_thesis: <*EMPTY*> . n = EMPTY then n = 1 by A4, FINSEQ_1:2, TARSKI:def_1; hence <*EMPTY*> . n = EMPTY by FINSEQ_1:40; ::_thesis: verum end; reconsider F = <*EMPTY*> as Finite_Sep_Sequence of S by A5, MESFUNC3:4; union (rng F) = union (bool {}) by FINSEQ_1:39, ZFMISC_1:1; then A10: dom (f + g) = union (rng F) by A3, ZFMISC_1:81; for x being Element of X st x in dom (f + g) holds |.((f + g) . x).| < +infty by A3; then A11: f + g is V60() by MESFUNC2:def_1; for n being Nat for x, y being Element of X st n in dom F & x in F . n & y in F . n holds (f + g) . x = (f + g) . y by A9; hence f + g is_simple_func_in S by A11, A10, MESFUNC2:def_4; ::_thesis: verum end; supposeA12: dom (f + g) <> {} ; ::_thesis: f + g is_simple_func_in S A13: (f | (dom (f + g))) " {+infty} = (dom (f + g)) /\ (f " {+infty}) by FUNCT_1:70; g is () by A2, Th14; then not +infty in rng g by Def4; then A14: g " {+infty} = {} by FUNCT_1:72; A15: (g | (dom (f + g))) " {+infty} = (dom (f + g)) /\ (g " {+infty}) by FUNCT_1:70; f is () by A1, Th14; then not +infty in rng f by Def4; then A16: f " {+infty} = {} by FUNCT_1:72; then A17: ((dom f) /\ (dom g)) \ (((f " {+infty}) /\ (g " {-infty})) \/ ((f " {-infty}) /\ (g " {+infty}))) = (dom f) /\ (dom g) by A14; then A18: dom (f + g) = (dom f) /\ (dom g) by MESFUNC1:def_3; dom (f | (dom (f + g))) = (dom f) /\ (dom (f + g)) by RELAT_1:61; then A19: dom (f | (dom (f + g))) = ((dom f) /\ (dom f)) /\ (dom g) by A18, XBOOLE_1:16; then A20: dom (f | (dom (f + g))) = dom (f + g) by A17, MESFUNC1:def_3; A21: dom g is Element of S by A2, Th37; dom f is Element of S by A1, Th37; then A22: dom (f + g) in S by A18, A21, FINSUB_1:def_2; then A23: g | (dom (f + g)) is_simple_func_in S by A2, Th34; dom (g | (dom (f + g))) = (dom g) /\ (dom (f + g)) by RELAT_1:61; then A24: dom (g | (dom (f + g))) = ((dom g) /\ (dom g)) /\ (dom f) by A18, XBOOLE_1:16; then A25: dom (g | (dom (f + g))) = dom (f + g) by A17, MESFUNC1:def_3; A26: dom ((f | (dom (f + g))) + (g | (dom (f + g)))) = ((dom (f | (dom (f + g)))) /\ (dom (g | (dom (f + g))))) \ ((((f | (dom (f + g))) " {+infty}) /\ ((g | (dom (f + g))) " {-infty})) \/ (((f | (dom (f + g))) " {-infty}) /\ ((g | (dom (f + g))) " {+infty}))) by MESFUNC1:def_3 .= dom (f + g) by A16, A14, A17, A19, A24, A13, A15, MESFUNC1:def_3 ; A27: for x being Element of X st x in dom ((f | (dom (f + g))) + (g | (dom (f + g)))) holds ((f | (dom (f + g))) + (g | (dom (f + g)))) . x = (f + g) . x proof let x be Element of X; ::_thesis: ( x in dom ((f | (dom (f + g))) + (g | (dom (f + g)))) implies ((f | (dom (f + g))) + (g | (dom (f + g)))) . x = (f + g) . x ) assume A28: x in dom ((f | (dom (f + g))) + (g | (dom (f + g)))) ; ::_thesis: ((f | (dom (f + g))) + (g | (dom (f + g)))) . x = (f + g) . x then ((f | (dom (f + g))) + (g | (dom (f + g)))) . x = ((f | (dom (f + g))) . x) + ((g | (dom (f + g))) . x) by MESFUNC1:def_3 .= (f . x) + ((g | (dom (f + g))) . x) by A26, A28, FUNCT_1:49 .= (f . x) + (g . x) by A26, A28, FUNCT_1:49 ; hence ((f | (dom (f + g))) + (g | (dom (f + g)))) . x = (f + g) . x by A26, A28, MESFUNC1:def_3; ::_thesis: verum end; f | (dom (f + g)) is_simple_func_in S by A1, A22, Th34; then (f | (dom (f + g))) + (g | (dom (f + g))) is_simple_func_in S by A12, A23, A20, A25, Lm3; hence f + g is_simple_func_in S by A26, A27, PARTFUN1:5; ::_thesis: verum end; end; end; theorem Th39: :: MESFUNC5:39 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for c being Real st f is_simple_func_in S holds c (#) f is_simple_func_in S proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for c being Real st f is_simple_func_in S holds c (#) f is_simple_func_in S let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for c being Real st f is_simple_func_in S holds c (#) f is_simple_func_in S let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for c being Real st f is_simple_func_in S holds c (#) f is_simple_func_in S let f be PartFunc of X,ExtREAL; ::_thesis: for c being Real st f is_simple_func_in S holds c (#) f is_simple_func_in S let c be Real; ::_thesis: ( f is_simple_func_in S implies c (#) f is_simple_func_in S ) set g = c (#) f; assume A1: f is_simple_func_in S ; ::_thesis: c (#) f is_simple_func_in S then consider G being Finite_Sep_Sequence of S such that A2: dom f = union (rng G) and A3: for n being Nat for x, y being Element of X st n in dom G & x in G . n & y in G . n holds f . x = f . y by MESFUNC2:def_4; A4: f is V60() by A1, MESFUNC2:def_4; now__::_thesis:_for_x_being_Element_of_X_st_x_in_dom_(c_(#)_f)_holds_ |.((c_(#)_f)_._x).|_<_+infty let x be Element of X; ::_thesis: ( x in dom (c (#) f) implies |.((c (#) f) . x).| < +infty ) assume A5: x in dom (c (#) f) ; ::_thesis: |.((c (#) f) . x).| < +infty (R_EAL c) * (f . x) <> -infty by A4; then (c (#) f) . x <> -infty by A5, MESFUNC1:def_6; then -infty < (c (#) f) . x by XXREAL_0:6; then A6: - +infty < (c (#) f) . x by XXREAL_3:def_3; (R_EAL c) * (f . x) <> +infty by A4; then (c (#) f) . x <> +infty by A5, MESFUNC1:def_6; then (c (#) f) . x < +infty by XXREAL_0:4; hence |.((c (#) f) . x).| < +infty by A6, EXTREAL2:11; ::_thesis: verum end; then A7: c (#) f is V60() by MESFUNC2:def_1; A8: dom (c (#) f) = dom f by MESFUNC1:def_6; now__::_thesis:_for_n_being_Nat for_x,_y_being_Element_of_X_st_n_in_dom_G_&_x_in_G_._n_&_y_in_G_._n_holds_ (c_(#)_f)_._x_=_(c_(#)_f)_._y let n be Nat; ::_thesis: for x, y being Element of X st n in dom G & x in G . n & y in G . n holds (c (#) f) . x = (c (#) f) . y let x, y be Element of X; ::_thesis: ( n in dom G & x in G . n & y in G . n implies (c (#) f) . x = (c (#) f) . y ) assume that A9: n in dom G and A10: x in G . n and A11: y in G . n ; ::_thesis: (c (#) f) . x = (c (#) f) . y A12: G . n in rng G by A9, FUNCT_1:3; then y in dom (c (#) f) by A8, A2, A11, TARSKI:def_4; then A13: (c (#) f) . y = (R_EAL c) * (f . y) by MESFUNC1:def_6; x in dom (c (#) f) by A8, A2, A10, A12, TARSKI:def_4; then (c (#) f) . x = (R_EAL c) * (f . x) by MESFUNC1:def_6; hence (c (#) f) . x = (c (#) f) . y by A3, A9, A10, A11, A13; ::_thesis: verum end; hence c (#) f is_simple_func_in S by A8, A7, A2, MESFUNC2:def_4; ::_thesis: verum end; theorem Th40: :: MESFUNC5:40 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & g is_simple_func_in S & ( for x being set st x in dom (f - g) holds g . x <= f . x ) holds f - g is nonnegative proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & g is_simple_func_in S & ( for x being set st x in dom (f - g) holds g . x <= f . x ) holds f - g is nonnegative let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & g is_simple_func_in S & ( for x being set st x in dom (f - g) holds g . x <= f . x ) holds f - g is nonnegative let M be sigma_Measure of S; ::_thesis: for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & g is_simple_func_in S & ( for x being set st x in dom (f - g) holds g . x <= f . x ) holds f - g is nonnegative let f, g be PartFunc of X,ExtREAL; ::_thesis: ( f is_simple_func_in S & g is_simple_func_in S & ( for x being set st x in dom (f - g) holds g . x <= f . x ) implies f - g is nonnegative ) assume that A1: f is_simple_func_in S and A2: g is_simple_func_in S and A3: for x being set st x in dom (f - g) holds g . x <= f . x ; ::_thesis: f - g is nonnegative g is () by A2, Th14; then not -infty in rng g by Def3; then A4: g " {-infty} = {} by FUNCT_1:72; f is () by A1, Th14; then not +infty in rng f by Def4; then A5: f " {+infty} = {} by FUNCT_1:72; then ((dom f) /\ (dom g)) \ (((f " {+infty}) /\ (g " {+infty})) \/ ((f " {-infty}) /\ (g " {-infty}))) = (dom f) /\ (dom g) by A4; then A6: dom (f - g) = (dom f) /\ (dom g) by MESFUNC1:def_4; for x being set st x in (dom f) /\ (dom g) holds ( g . x <= f . x & -infty < g . x & f . x < +infty ) proof let x be set ; ::_thesis: ( x in (dom f) /\ (dom g) implies ( g . x <= f . x & -infty < g . x & f . x < +infty ) ) assume A7: x in (dom f) /\ (dom g) ; ::_thesis: ( g . x <= f . x & -infty < g . x & f . x < +infty ) hence g . x <= f . x by A3, A6; ::_thesis: ( -infty < g . x & f . x < +infty ) x in dom g by A7, XBOOLE_0:def_4; then not g . x in {-infty} by A4, FUNCT_1:def_7; then not g . x = -infty by TARSKI:def_1; hence -infty < g . x by XXREAL_0:6; ::_thesis: f . x < +infty x in dom f by A7, XBOOLE_0:def_4; then not f . x in {+infty} by A5, FUNCT_1:def_7; then not f . x = +infty by TARSKI:def_1; hence f . x < +infty by XXREAL_0:4; ::_thesis: verum end; hence f - g is nonnegative by Th21; ::_thesis: verum end; theorem Th41: :: MESFUNC5:41 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for A being Element of S for c being R_eal st c <> +infty & c <> -infty holds ex f being PartFunc of X,ExtREAL st ( f is_simple_func_in S & dom f = A & ( for x being set st x in A holds f . x = c ) ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for A being Element of S for c being R_eal st c <> +infty & c <> -infty holds ex f being PartFunc of X,ExtREAL st ( f is_simple_func_in S & dom f = A & ( for x being set st x in A holds f . x = c ) ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for A being Element of S for c being R_eal st c <> +infty & c <> -infty holds ex f being PartFunc of X,ExtREAL st ( f is_simple_func_in S & dom f = A & ( for x being set st x in A holds f . x = c ) ) let M be sigma_Measure of S; ::_thesis: for A being Element of S for c being R_eal st c <> +infty & c <> -infty holds ex f being PartFunc of X,ExtREAL st ( f is_simple_func_in S & dom f = A & ( for x being set st x in A holds f . x = c ) ) let A be Element of S; ::_thesis: for c being R_eal st c <> +infty & c <> -infty holds ex f being PartFunc of X,ExtREAL st ( f is_simple_func_in S & dom f = A & ( for x being set st x in A holds f . x = c ) ) let c be R_eal; ::_thesis: ( c <> +infty & c <> -infty implies ex f being PartFunc of X,ExtREAL st ( f is_simple_func_in S & dom f = A & ( for x being set st x in A holds f . x = c ) ) ) assume that A1: c <> +infty and A2: c <> -infty ; ::_thesis: ex f being PartFunc of X,ExtREAL st ( f is_simple_func_in S & dom f = A & ( for x being set st x in A holds f . x = c ) ) -infty < c by A2, XXREAL_0:6; then A3: - +infty < c by XXREAL_3:def_3; deffunc H1( set ) -> R_eal = c; defpred S1[ set ] means $1 in A; A4: for x being set st S1[x] holds H1(x) in ExtREAL ; consider f being PartFunc of X,ExtREAL such that A5: ( ( for x being set holds ( x in dom f iff ( x in X & S1[x] ) ) ) & ( for x being set st x in dom f holds f . x = H1(x) ) ) from PARTFUN1:sch_3(A4); c < +infty by A1, XXREAL_0:4; then |.c.| < +infty by A3, EXTREAL2:11; then for x being Element of X st x in dom f holds |.(f . x).| < +infty by A5; then A6: f is V60() by MESFUNC2:def_1; take f ; ::_thesis: ( f is_simple_func_in S & dom f = A & ( for x being set st x in A holds f . x = c ) ) A7: A c= dom f proof let x be set ; :: according to TARSKI:def_3 ::_thesis: ( not x in A or x in dom f ) assume x in A ; ::_thesis: x in dom f hence x in dom f by A5; ::_thesis: verum end; set F = <*(dom f)*>; A8: dom f c= A proof let x be set ; :: according to TARSKI:def_3 ::_thesis: ( not x in dom f or x in A ) assume x in dom f ; ::_thesis: x in A hence x in A by A5; ::_thesis: verum end; A9: rng <*(dom f)*> = {(dom f)} by FINSEQ_1:38; then A10: rng <*(dom f)*> = {A} by A8, A7, XBOOLE_0:def_10; rng <*(dom f)*> c= S proof let z be set ; :: according to TARSKI:def_3 ::_thesis: ( not z in rng <*(dom f)*> or z in S ) assume z in rng <*(dom f)*> ; ::_thesis: z in S then z = A by A10, TARSKI:def_1; hence z in S ; ::_thesis: verum end; then reconsider F = <*(dom f)*> as FinSequence of S by FINSEQ_1:def_4; now__::_thesis:_for_i,_j_being_Nat_st_i_in_dom_F_&_j_in_dom_F_&_i_<>_j_holds_ F_._i_misses_F_._j let i, j be Nat; ::_thesis: ( i in dom F & j in dom F & i <> j implies F . i misses F . j ) assume that A11: i in dom F and A12: j in dom F and A13: i <> j ; ::_thesis: F . i misses F . j A14: dom F = Seg 1 by FINSEQ_1:38; then i = 1 by A11, FINSEQ_1:2, TARSKI:def_1; hence F . i misses F . j by A12, A13, A14, FINSEQ_1:2, TARSKI:def_1; ::_thesis: verum end; then reconsider F = F as Finite_Sep_Sequence of S by MESFUNC3:4; A15: now__::_thesis:_for_n_being_Nat for_x,_y_being_Element_of_X_st_n_in_dom_F_&_x_in_F_._n_&_y_in_F_._n_holds_ f_._x_=_f_._y let n be Nat; ::_thesis: for x, y being Element of X st n in dom F & x in F . n & y in F . n holds f . x = f . y let x, y be Element of X; ::_thesis: ( n in dom F & x in F . n & y in F . n implies f . x = f . y ) assume that A16: n in dom F and A17: x in F . n and A18: y in F . n ; ::_thesis: f . x = f . y dom F = Seg 1 by FINSEQ_1:38; then A19: n = 1 by A16, FINSEQ_1:2, TARSKI:def_1; then x in dom f by A17, FINSEQ_1:40; then A20: f . x = c by A5; y in dom f by A18, A19, FINSEQ_1:40; hence f . x = f . y by A5, A20; ::_thesis: verum end; dom f = union (rng F) by A9, ZFMISC_1:25; hence f is_simple_func_in S by A6, A15, MESFUNC2:def_4; ::_thesis: ( dom f = A & ( for x being set st x in A holds f . x = c ) ) thus dom f = A by A8, A7, XBOOLE_0:def_10; ::_thesis: for x being set st x in A holds f . x = c thus for x being set st x in A holds f . x = c by A5; ::_thesis: verum end; theorem Th42: :: MESFUNC5:42 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for B, BF being Element of S st f is_measurable_on B & BF = (dom f) /\ B holds f | B is_measurable_on BF proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for B, BF being Element of S st f is_measurable_on B & BF = (dom f) /\ B holds f | B is_measurable_on BF let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for B, BF being Element of S st f is_measurable_on B & BF = (dom f) /\ B holds f | B is_measurable_on BF let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for B, BF being Element of S st f is_measurable_on B & BF = (dom f) /\ B holds f | B is_measurable_on BF let f be PartFunc of X,ExtREAL; ::_thesis: for B, BF being Element of S st f is_measurable_on B & BF = (dom f) /\ B holds f | B is_measurable_on BF let B, BF be Element of S; ::_thesis: ( f is_measurable_on B & BF = (dom f) /\ B implies f | B is_measurable_on BF ) assume that A1: f is_measurable_on B and A2: BF = (dom f) /\ B ; ::_thesis: f | B is_measurable_on BF now__::_thesis:_for_r_being_real_number_holds_BF_/\_(less_dom_((f_|_B),(R_EAL_r)))_in_S let r be real number ; ::_thesis: BF /\ (less_dom ((f | B),(R_EAL r))) in S A3: now__::_thesis:_for_x_being_set_holds_ (_x_in_BF_/\_(less_dom_((f_|_B),(R_EAL_r)))_iff_x_in_B_/\_(less_dom_(f,(R_EAL_r)))_) let x be set ; ::_thesis: ( x in BF /\ (less_dom ((f | B),(R_EAL r))) iff x in B /\ (less_dom (f,(R_EAL r))) ) ( x in dom (f | B) & ex y being R_eal st ( y = (f | B) . x & y < R_EAL r ) iff ( x in (dom f) /\ B & ex y being R_eal st ( y = (f | B) . x & y < R_EAL r ) ) ) by RELAT_1:61; then A4: ( x in BF & x in less_dom ((f | B),(R_EAL r)) iff ( x in B & x in dom f & (f | B) . x < R_EAL r ) ) by A2, MESFUNC1:def_11, XBOOLE_0:def_4; ( x in B & x in dom f implies ( f . x < R_EAL r iff (f | B) . x < R_EAL r ) ) by FUNCT_1:49; then ( x in BF /\ (less_dom ((f | B),(R_EAL r))) iff ( x in B & x in less_dom (f,(R_EAL r)) ) ) by A4, MESFUNC1:def_11, XBOOLE_0:def_4; hence ( x in BF /\ (less_dom ((f | B),(R_EAL r))) iff x in B /\ (less_dom (f,(R_EAL r))) ) by XBOOLE_0:def_4; ::_thesis: verum end; then A5: B /\ (less_dom (f,(R_EAL r))) c= BF /\ (less_dom ((f | B),(R_EAL r))) by TARSKI:def_3; BF /\ (less_dom ((f | B),(R_EAL r))) c= B /\ (less_dom (f,(R_EAL r))) by A3, TARSKI:def_3; then BF /\ (less_dom ((f | B),(R_EAL r))) = B /\ (less_dom (f,(R_EAL r))) by A5, XBOOLE_0:def_10; hence BF /\ (less_dom ((f | B),(R_EAL r))) in S by A1, MESFUNC1:def_16; ::_thesis: verum end; hence f | B is_measurable_on BF by MESFUNC1:def_16; ::_thesis: verum end; theorem Th43: :: MESFUNC5:43 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for A being Element of S for f, g being PartFunc of X,ExtREAL st A c= dom f & f is_measurable_on A & g is_measurable_on A & f is () & g is () holds (max+ (f + g)) + (max- f) is_measurable_on A proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for A being Element of S for f, g being PartFunc of X,ExtREAL st A c= dom f & f is_measurable_on A & g is_measurable_on A & f is () & g is () holds (max+ (f + g)) + (max- f) is_measurable_on A let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for A being Element of S for f, g being PartFunc of X,ExtREAL st A c= dom f & f is_measurable_on A & g is_measurable_on A & f is () & g is () holds (max+ (f + g)) + (max- f) is_measurable_on A let M be sigma_Measure of S; ::_thesis: for A being Element of S for f, g being PartFunc of X,ExtREAL st A c= dom f & f is_measurable_on A & g is_measurable_on A & f is () & g is () holds (max+ (f + g)) + (max- f) is_measurable_on A let A be Element of S; ::_thesis: for f, g being PartFunc of X,ExtREAL st A c= dom f & f is_measurable_on A & g is_measurable_on A & f is () & g is () holds (max+ (f + g)) + (max- f) is_measurable_on A let f, g be PartFunc of X,ExtREAL; ::_thesis: ( A c= dom f & f is_measurable_on A & g is_measurable_on A & f is () & g is () implies (max+ (f + g)) + (max- f) is_measurable_on A ) assume that A1: A c= dom f and A2: f is_measurable_on A and A3: g is_measurable_on A and A4: f is () and A5: g is () ; ::_thesis: (max+ (f + g)) + (max- f) is_measurable_on A f + g is_measurable_on A by A2, A3, A4, A5, Th31; then A6: max+ (f + g) is_measurable_on A by MESFUNC2:25; A7: max- f is nonnegative by Lm1; A8: max+ (f + g) is nonnegative by Lm1; max- f is_measurable_on A by A1, A2, MESFUNC2:26; hence (max+ (f + g)) + (max- f) is_measurable_on A by A6, A8, A7, Th31; ::_thesis: verum end; theorem Th44: :: MESFUNC5:44 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for A being Element of S for f, g being PartFunc of X,ExtREAL st A c= (dom f) /\ (dom g) & f is_measurable_on A & g is_measurable_on A & f is () & g is () holds (max- (f + g)) + (max+ f) is_measurable_on A proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for A being Element of S for f, g being PartFunc of X,ExtREAL st A c= (dom f) /\ (dom g) & f is_measurable_on A & g is_measurable_on A & f is () & g is () holds (max- (f + g)) + (max+ f) is_measurable_on A let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for A being Element of S for f, g being PartFunc of X,ExtREAL st A c= (dom f) /\ (dom g) & f is_measurable_on A & g is_measurable_on A & f is () & g is () holds (max- (f + g)) + (max+ f) is_measurable_on A let M be sigma_Measure of S; ::_thesis: for A being Element of S for f, g being PartFunc of X,ExtREAL st A c= (dom f) /\ (dom g) & f is_measurable_on A & g is_measurable_on A & f is () & g is () holds (max- (f + g)) + (max+ f) is_measurable_on A let A be Element of S; ::_thesis: for f, g being PartFunc of X,ExtREAL st A c= (dom f) /\ (dom g) & f is_measurable_on A & g is_measurable_on A & f is () & g is () holds (max- (f + g)) + (max+ f) is_measurable_on A let f, g be PartFunc of X,ExtREAL; ::_thesis: ( A c= (dom f) /\ (dom g) & f is_measurable_on A & g is_measurable_on A & f is () & g is () implies (max- (f + g)) + (max+ f) is_measurable_on A ) assume that A1: A c= (dom f) /\ (dom g) and A2: f is_measurable_on A and A3: g is_measurable_on A and A4: f is () and A5: g is () ; ::_thesis: (max- (f + g)) + (max+ f) is_measurable_on A A6: dom (f + g) = (dom f) /\ (dom g) by A4, A5, Th16; f + g is_measurable_on A by A2, A3, A4, A5, Th31; then A7: max- (f + g) is_measurable_on A by A1, A6, MESFUNC2:26; A8: max- (f + g) is nonnegative by Lm1; A9: max+ f is nonnegative by Lm1; max+ f is_measurable_on A by A2, MESFUNC2:25; hence (max- (f + g)) + (max+ f) is_measurable_on A by A7, A8, A9, Th31; ::_thesis: verum end; theorem Th45: :: MESFUNC5:45 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for A being set st A in S holds 0 <= M . A proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for A being set st A in S holds 0 <= M . A let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for A being set st A in S holds 0 <= M . A let M be sigma_Measure of S; ::_thesis: for A being set st A in S holds 0 <= M . A let A be set ; ::_thesis: ( A in S implies 0 <= M . A ) reconsider E = {} as Element of S by PROB_1:4; assume A in S ; ::_thesis: 0 <= M . A then reconsider A = A as Element of S ; M . E <= M . A by MEASURE1:31, XBOOLE_1:2; hence 0 <= M . A by VALUED_0:def_19; ::_thesis: verum end; Lm4: for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for r being Real st dom f in S & ( for x being set st x in dom f holds f . x = r ) holds f is_simple_func_in S proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for r being Real st dom f in S & ( for x being set st x in dom f holds f . x = r ) holds f is_simple_func_in S let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for r being Real st dom f in S & ( for x being set st x in dom f holds f . x = r ) holds f is_simple_func_in S let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for r being Real st dom f in S & ( for x being set st x in dom f holds f . x = r ) holds f is_simple_func_in S let f be PartFunc of X,ExtREAL; ::_thesis: for r being Real st dom f in S & ( for x being set st x in dom f holds f . x = r ) holds f is_simple_func_in S let r be Real; ::_thesis: ( dom f in S & ( for x being set st x in dom f holds f . x = r ) implies f is_simple_func_in S ) assume that A1: dom f in S and A2: for x being set st x in dom f holds f . x = r ; ::_thesis: f is_simple_func_in S reconsider Df = dom f as Element of S by A1; set F = <*Df*>; A3: dom <*Df*> = Seg 1 by FINSEQ_1:38; A4: now__::_thesis:_for_i,_j_being_Nat_st_i_in_dom_<*Df*>_&_j_in_dom_<*Df*>_&_i_<>_j_holds_ <*Df*>_._i_misses_<*Df*>_._j let i, j be Nat; ::_thesis: ( i in dom <*Df*> & j in dom <*Df*> & i <> j implies <*Df*> . i misses <*Df*> . j ) assume that A5: i in dom <*Df*> and A6: j in dom <*Df*> and A7: i <> j ; ::_thesis: <*Df*> . i misses <*Df*> . j i = 1 by A3, A5, FINSEQ_1:2, TARSKI:def_1; hence <*Df*> . i misses <*Df*> . j by A3, A6, A7, FINSEQ_1:2, TARSKI:def_1; ::_thesis: verum end; A8: for n being Nat st n in dom <*Df*> holds <*Df*> . n = Df proof let n be Nat; ::_thesis: ( n in dom <*Df*> implies <*Df*> . n = Df ) assume n in dom <*Df*> ; ::_thesis: <*Df*> . n = Df then n = 1 by A3, FINSEQ_1:2, TARSKI:def_1; hence <*Df*> . n = Df by FINSEQ_1:40; ::_thesis: verum end; reconsider F = <*Df*> as Finite_Sep_Sequence of S by A4, MESFUNC3:4; A9: now__::_thesis:_for_n_being_Nat for_x,_y_being_Element_of_X_st_n_in_dom_F_&_x_in_F_._n_&_y_in_F_._n_holds_ f_._x_=_f_._y let n be Nat; ::_thesis: for x, y being Element of X st n in dom F & x in F . n & y in F . n holds f . x = f . y let x, y be Element of X; ::_thesis: ( n in dom F & x in F . n & y in F . n implies f . x = f . y ) assume that A10: n in dom F and A11: x in F . n and A12: y in F . n ; ::_thesis: f . x = f . y A13: F . n = Df by A8, A10; then f . x = R_EAL r by A2, A11; hence f . x = f . y by A2, A12, A13; ::_thesis: verum end; F = <*(F . 1)*> by FINSEQ_1:40; then A14: rng F = {(F . 1)} by FINSEQ_1:38; now__::_thesis:_for_x_being_Element_of_X_st_x_in_dom_f_holds_ |.(f_._x).|_<_+infty let x be Element of X; ::_thesis: ( x in dom f implies |.(f . x).| < +infty ) assume x in dom f ; ::_thesis: |.(f . x).| < +infty then A15: f . x = R_EAL r by A2; then -infty < f . x by XXREAL_0:12; then A16: - +infty < f . x by XXREAL_3:def_3; f . x < +infty by A15, XXREAL_0:9; hence |.(f . x).| < +infty by A16, EXTREAL2:11; ::_thesis: verum end; then A17: f is V60() by MESFUNC2:def_1; 1 in Seg 1 ; then F . 1 = Df by A3, A8; then dom f = union (rng F) by A14, ZFMISC_1:25; hence f is_simple_func_in S by A17, A9, MESFUNC2:def_4; ::_thesis: verum end; theorem Th46: :: MESFUNC5:46 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st ex E1 being Element of S st ( E1 = dom f & f is_measurable_on E1 ) & ex E2 being Element of S st ( E2 = dom g & g is_measurable_on E2 ) & f " {+infty} in S & f " {-infty} in S & g " {+infty} in S & g " {-infty} in S holds dom (f + g) in S proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st ex E1 being Element of S st ( E1 = dom f & f is_measurable_on E1 ) & ex E2 being Element of S st ( E2 = dom g & g is_measurable_on E2 ) & f " {+infty} in S & f " {-infty} in S & g " {+infty} in S & g " {-infty} in S holds dom (f + g) in S let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st ex E1 being Element of S st ( E1 = dom f & f is_measurable_on E1 ) & ex E2 being Element of S st ( E2 = dom g & g is_measurable_on E2 ) & f " {+infty} in S & f " {-infty} in S & g " {+infty} in S & g " {-infty} in S holds dom (f + g) in S let M be sigma_Measure of S; ::_thesis: for f, g being PartFunc of X,ExtREAL st ex E1 being Element of S st ( E1 = dom f & f is_measurable_on E1 ) & ex E2 being Element of S st ( E2 = dom g & g is_measurable_on E2 ) & f " {+infty} in S & f " {-infty} in S & g " {+infty} in S & g " {-infty} in S holds dom (f + g) in S let f, g be PartFunc of X,ExtREAL; ::_thesis: ( ex E1 being Element of S st ( E1 = dom f & f is_measurable_on E1 ) & ex E2 being Element of S st ( E2 = dom g & g is_measurable_on E2 ) & f " {+infty} in S & f " {-infty} in S & g " {+infty} in S & g " {-infty} in S implies dom (f + g) in S ) assume that A1: ex E1 being Element of S st ( E1 = dom f & f is_measurable_on E1 ) and A2: ex E2 being Element of S st ( E2 = dom g & g is_measurable_on E2 ) and A3: f " {+infty} in S and A4: f " {-infty} in S and A5: g " {+infty} in S and A6: g " {-infty} in S ; ::_thesis: dom (f + g) in S A7: (f " {+infty}) /\ (g " {-infty}) in S by A3, A6, MEASURE1:34; (f " {-infty}) /\ (g " {+infty}) in S by A4, A5, MEASURE1:34; then ((f " {-infty}) /\ (g " {+infty})) \/ ((f " {+infty}) /\ (g " {-infty})) in S by A7, MEASURE1:34; then A8: X \ (((f " {-infty}) /\ (g " {+infty})) \/ ((f " {+infty}) /\ (g " {-infty}))) in S by MEASURE1:34; consider E2 being Element of S such that A9: E2 = dom g and g is_measurable_on E2 by A2; consider E1 being Element of S such that A10: E1 = dom f and f is_measurable_on E1 by A1; A11: (E1 /\ E2) /\ (X \ (((f " {-infty}) /\ (g " {+infty})) \/ ((f " {+infty}) /\ (g " {-infty})))) = ((E1 /\ E2) /\ X) \ (((f " {-infty}) /\ (g " {+infty})) \/ ((f " {+infty}) /\ (g " {-infty}))) by XBOOLE_1:49 .= (E1 /\ E2) \ (((f " {-infty}) /\ (g " {+infty})) \/ ((f " {+infty}) /\ (g " {-infty}))) by XBOOLE_1:28 ; dom (f + g) = (E1 /\ E2) \ (((f " {-infty}) /\ (g " {+infty})) \/ ((f " {+infty}) /\ (g " {-infty}))) by A10, A9, MESFUNC1:def_3; hence dom (f + g) in S by A8, A11, MEASURE1:34; ::_thesis: verum end; Lm5: for X being non empty set for S being SigmaField of X for A being Element of S for f being PartFunc of X,ExtREAL for r being real number holds A /\ (less_dom (f,(R_EAL r))) = less_dom ((f | A),(R_EAL r)) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for A being Element of S for f being PartFunc of X,ExtREAL for r being real number holds A /\ (less_dom (f,(R_EAL r))) = less_dom ((f | A),(R_EAL r)) let S be SigmaField of X; ::_thesis: for A being Element of S for f being PartFunc of X,ExtREAL for r being real number holds A /\ (less_dom (f,(R_EAL r))) = less_dom ((f | A),(R_EAL r)) let A be Element of S; ::_thesis: for f being PartFunc of X,ExtREAL for r being real number holds A /\ (less_dom (f,(R_EAL r))) = less_dom ((f | A),(R_EAL r)) let f be PartFunc of X,ExtREAL; ::_thesis: for r being real number holds A /\ (less_dom (f,(R_EAL r))) = less_dom ((f | A),(R_EAL r)) let r be real number ; ::_thesis: A /\ (less_dom (f,(R_EAL r))) = less_dom ((f | A),(R_EAL r)) now__::_thesis:_for_v_being_set_st_v_in_A_/\_(less_dom_(f,(R_EAL_r)))_holds_ v_in_less_dom_((f_|_A),(R_EAL_r)) let v be set ; ::_thesis: ( v in A /\ (less_dom (f,(R_EAL r))) implies v in less_dom ((f | A),(R_EAL r)) ) assume A1: v in A /\ (less_dom (f,(R_EAL r))) ; ::_thesis: v in less_dom ((f | A),(R_EAL r)) then A2: v in less_dom (f,(R_EAL r)) by XBOOLE_0:def_4; A3: v in A by A1, XBOOLE_0:def_4; then A4: f . v = (f | A) . v by FUNCT_1:49; v in dom f by A2, MESFUNC1:def_11; then v in A /\ (dom f) by A3, XBOOLE_0:def_4; then A5: v in dom (f | A) by RELAT_1:61; f . v < R_EAL r by A2, MESFUNC1:def_11; hence v in less_dom ((f | A),(R_EAL r)) by A5, A4, MESFUNC1:def_11; ::_thesis: verum end; hence A /\ (less_dom (f,(R_EAL r))) c= less_dom ((f | A),(R_EAL r)) by TARSKI:def_3; :: according to XBOOLE_0:def_10 ::_thesis: less_dom ((f | A),(R_EAL r)) c= A /\ (less_dom (f,(R_EAL r))) let v be set ; :: according to TARSKI:def_3 ::_thesis: ( not v in less_dom ((f | A),(R_EAL r)) or v in A /\ (less_dom (f,(R_EAL r))) ) assume A6: v in less_dom ((f | A),(R_EAL r)) ; ::_thesis: v in A /\ (less_dom (f,(R_EAL r))) then A7: v in dom (f | A) by MESFUNC1:def_11; then A8: v in (dom f) /\ A by RELAT_1:61; then A9: v in dom f by XBOOLE_0:def_4; (f | A) . v < R_EAL r by A6, MESFUNC1:def_11; then ex w being R_eal st ( w = f . v & w < R_EAL r ) by A7, FUNCT_1:47; then A10: v in less_dom (f,(R_EAL r)) by A9, MESFUNC1:def_11; v in A by A8, XBOOLE_0:def_4; hence v in A /\ (less_dom (f,(R_EAL r))) by A10, XBOOLE_0:def_4; ::_thesis: verum end; Lm6: for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for A being Element of S for f being PartFunc of X,ExtREAL holds ( f | A is_measurable_on A iff f is_measurable_on A ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for A being Element of S for f being PartFunc of X,ExtREAL holds ( f | A is_measurable_on A iff f is_measurable_on A ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for A being Element of S for f being PartFunc of X,ExtREAL holds ( f | A is_measurable_on A iff f is_measurable_on A ) let M be sigma_Measure of S; ::_thesis: for A being Element of S for f being PartFunc of X,ExtREAL holds ( f | A is_measurable_on A iff f is_measurable_on A ) let A be Element of S; ::_thesis: for f being PartFunc of X,ExtREAL holds ( f | A is_measurable_on A iff f is_measurable_on A ) let f be PartFunc of X,ExtREAL; ::_thesis: ( f | A is_measurable_on A iff f is_measurable_on A ) now__::_thesis:_(_f_|_A_is_measurable_on_A_implies_f_is_measurable_on_A_) assume A1: f | A is_measurable_on A ; ::_thesis: f is_measurable_on A now__::_thesis:_for_r_being_real_number_holds_A_/\_(less_dom_(f,(R_EAL_r)))_in_S let r be real number ; ::_thesis: A /\ (less_dom (f,(R_EAL r))) in S A /\ (less_dom ((f | A),(R_EAL r))) in S by A1, MESFUNC1:def_16; then A /\ (A /\ (less_dom (f,(R_EAL r)))) in S by Lm5; then (A /\ A) /\ (less_dom (f,(R_EAL r))) in S by XBOOLE_1:16; hence A /\ (less_dom (f,(R_EAL r))) in S ; ::_thesis: verum end; hence f is_measurable_on A by MESFUNC1:def_16; ::_thesis: verum end; hence ( f | A is_measurable_on A implies f is_measurable_on A ) ; ::_thesis: ( f is_measurable_on A implies f | A is_measurable_on A ) assume A2: f is_measurable_on A ; ::_thesis: f | A is_measurable_on A now__::_thesis:_for_r_being_real_number_holds_A_/\_(less_dom_((f_|_A),(R_EAL_r)))_in_S let r be real number ; ::_thesis: A /\ (less_dom ((f | A),(R_EAL r))) in S (A /\ A) /\ (less_dom (f,(R_EAL r))) in S by A2, MESFUNC1:def_16; then A /\ (A /\ (less_dom (f,(R_EAL r)))) in S by XBOOLE_1:16; hence A /\ (less_dom ((f | A),(R_EAL r))) in S by Lm5; ::_thesis: verum end; hence f | A is_measurable_on A by MESFUNC1:def_16; ::_thesis: verum end; Lm7: for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st ex E1 being Element of S st ( E1 = dom f & f is_measurable_on E1 ) & ex E2 being Element of S st ( E2 = dom g & g is_measurable_on E2 ) & dom f = dom g holds ex DFPG being Element of S st ( DFPG = dom (f + g) & f + g is_measurable_on DFPG ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st ex E1 being Element of S st ( E1 = dom f & f is_measurable_on E1 ) & ex E2 being Element of S st ( E2 = dom g & g is_measurable_on E2 ) & dom f = dom g holds ex DFPG being Element of S st ( DFPG = dom (f + g) & f + g is_measurable_on DFPG ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st ex E1 being Element of S st ( E1 = dom f & f is_measurable_on E1 ) & ex E2 being Element of S st ( E2 = dom g & g is_measurable_on E2 ) & dom f = dom g holds ex DFPG being Element of S st ( DFPG = dom (f + g) & f + g is_measurable_on DFPG ) let M be sigma_Measure of S; ::_thesis: for f, g being PartFunc of X,ExtREAL st ex E1 being Element of S st ( E1 = dom f & f is_measurable_on E1 ) & ex E2 being Element of S st ( E2 = dom g & g is_measurable_on E2 ) & dom f = dom g holds ex DFPG being Element of S st ( DFPG = dom (f + g) & f + g is_measurable_on DFPG ) let f, g be PartFunc of X,ExtREAL; ::_thesis: ( ex E1 being Element of S st ( E1 = dom f & f is_measurable_on E1 ) & ex E2 being Element of S st ( E2 = dom g & g is_measurable_on E2 ) & dom f = dom g implies ex DFPG being Element of S st ( DFPG = dom (f + g) & f + g is_measurable_on DFPG ) ) assume that A1: ex E1 being Element of S st ( E1 = dom f & f is_measurable_on E1 ) and A2: ex E2 being Element of S st ( E2 = dom g & g is_measurable_on E2 ) and A3: dom f = dom g ; ::_thesis: ex DFPG being Element of S st ( DFPG = dom (f + g) & f + g is_measurable_on DFPG ) consider Gf being Element of S such that A4: Gf = dom g and A5: g is_measurable_on Gf by A2; now__::_thesis:_for_v_being_set_st_v_in_g_"_{-infty}_holds_ v_in_Gf_/\_(eq_dom_(g,-infty)) let v be set ; ::_thesis: ( v in g " {-infty} implies v in Gf /\ (eq_dom (g,-infty)) ) assume A6: v in g " {-infty} ; ::_thesis: v in Gf /\ (eq_dom (g,-infty)) then A7: v in dom g by FUNCT_1:def_7; g . v in {-infty} by A6, FUNCT_1:def_7; then g . v = -infty by TARSKI:def_1; then v in eq_dom (g,-infty) by A7, MESFUNC1:def_15; hence v in Gf /\ (eq_dom (g,-infty)) by A4, A7, XBOOLE_0:def_4; ::_thesis: verum end; then A8: g " {-infty} c= Gf /\ (eq_dom (g,-infty)) by TARSKI:def_3; now__::_thesis:_for_v_being_set_st_v_in_Gf_/\_(eq_dom_(g,-infty))_holds_ v_in_g_"_{-infty} let v be set ; ::_thesis: ( v in Gf /\ (eq_dom (g,-infty)) implies v in g " {-infty} ) assume v in Gf /\ (eq_dom (g,-infty)) ; ::_thesis: v in g " {-infty} then A9: v in eq_dom (g,-infty) by XBOOLE_0:def_4; then g . v = -infty by MESFUNC1:def_15; then A10: g . v in {-infty} by TARSKI:def_1; v in dom g by A9, MESFUNC1:def_15; hence v in g " {-infty} by A10, FUNCT_1:def_7; ::_thesis: verum end; then A11: Gf /\ (eq_dom (g,-infty)) c= g " {-infty} by TARSKI:def_3; Gf /\ (eq_dom (g,-infty)) in S by A5, MESFUNC1:34; then A12: g " {-infty} in S by A8, A11, XBOOLE_0:def_10; now__::_thesis:_for_v_being_set_st_v_in_g_"_{+infty}_holds_ v_in_Gf_/\_(eq_dom_(g,+infty)) let v be set ; ::_thesis: ( v in g " {+infty} implies v in Gf /\ (eq_dom (g,+infty)) ) assume A13: v in g " {+infty} ; ::_thesis: v in Gf /\ (eq_dom (g,+infty)) then A14: v in dom g by FUNCT_1:def_7; g . v in {+infty} by A13, FUNCT_1:def_7; then g . v = +infty by TARSKI:def_1; then v in eq_dom (g,+infty) by A14, MESFUNC1:def_15; hence v in Gf /\ (eq_dom (g,+infty)) by A4, A14, XBOOLE_0:def_4; ::_thesis: verum end; then A15: g " {+infty} c= Gf /\ (eq_dom (g,+infty)) by TARSKI:def_3; now__::_thesis:_for_v_being_set_st_v_in_Gf_/\_(eq_dom_(g,+infty))_holds_ v_in_g_"_{+infty} let v be set ; ::_thesis: ( v in Gf /\ (eq_dom (g,+infty)) implies v in g " {+infty} ) assume v in Gf /\ (eq_dom (g,+infty)) ; ::_thesis: v in g " {+infty} then A16: v in eq_dom (g,+infty) by XBOOLE_0:def_4; then g . v = +infty by MESFUNC1:def_15; then A17: g . v in {+infty} by TARSKI:def_1; v in dom g by A16, MESFUNC1:def_15; hence v in g " {+infty} by A17, FUNCT_1:def_7; ::_thesis: verum end; then A18: Gf /\ (eq_dom (g,+infty)) c= g " {+infty} by TARSKI:def_3; A19: (f " {+infty}) /\ (g " {-infty}) c= g " {-infty} by XBOOLE_1:17; A20: (f " {-infty}) /\ (g " {+infty}) c= f " {-infty} by XBOOLE_1:17; A21: dom (f + g) = ((dom f) /\ (dom g)) \ (((f " {-infty}) /\ (g " {+infty})) \/ ((f " {+infty}) /\ (g " {-infty}))) by MESFUNC1:def_3; Gf /\ (eq_dom (g,+infty)) in S by A4, A5, MESFUNC1:33; then A22: g " {+infty} in S by A15, A18, XBOOLE_0:def_10; then reconsider NG = (g " {+infty}) \/ (g " {-infty}) as Element of S by A12, PROB_1:3; consider E0 being Element of S such that A23: E0 = dom f and A24: f is_measurable_on E0 by A1; A25: E0 /\ (eq_dom (f,+infty)) in S by A23, A24, MESFUNC1:33; now__::_thesis:_for_v_being_set_st_v_in_E0_/\_(eq_dom_(f,+infty))_holds_ v_in_f_"_{+infty} let v be set ; ::_thesis: ( v in E0 /\ (eq_dom (f,+infty)) implies v in f " {+infty} ) assume v in E0 /\ (eq_dom (f,+infty)) ; ::_thesis: v in f " {+infty} then A26: v in eq_dom (f,+infty) by XBOOLE_0:def_4; then f . v = +infty by MESFUNC1:def_15; then A27: f . v in {+infty} by TARSKI:def_1; v in dom f by A26, MESFUNC1:def_15; hence v in f " {+infty} by A27, FUNCT_1:def_7; ::_thesis: verum end; then A28: E0 /\ (eq_dom (f,+infty)) c= f " {+infty} by TARSKI:def_3; now__::_thesis:_for_v_being_set_st_v_in_f_"_{+infty}_holds_ v_in_E0_/\_(eq_dom_(f,+infty)) let v be set ; ::_thesis: ( v in f " {+infty} implies v in E0 /\ (eq_dom (f,+infty)) ) assume A29: v in f " {+infty} ; ::_thesis: v in E0 /\ (eq_dom (f,+infty)) then A30: v in dom f by FUNCT_1:def_7; f . v in {+infty} by A29, FUNCT_1:def_7; then f . v = +infty by TARSKI:def_1; then v in eq_dom (f,+infty) by A30, MESFUNC1:def_15; hence v in E0 /\ (eq_dom (f,+infty)) by A23, A30, XBOOLE_0:def_4; ::_thesis: verum end; then f " {+infty} c= E0 /\ (eq_dom (f,+infty)) by TARSKI:def_3; then A31: f " {+infty} = E0 /\ (eq_dom (f,+infty)) by A28, XBOOLE_0:def_10; now__::_thesis:_for_v_being_set_st_v_in_E0_/\_(eq_dom_(f,-infty))_holds_ v_in_f_"_{-infty} let v be set ; ::_thesis: ( v in E0 /\ (eq_dom (f,-infty)) implies v in f " {-infty} ) assume v in E0 /\ (eq_dom (f,-infty)) ; ::_thesis: v in f " {-infty} then A32: v in eq_dom (f,-infty) by XBOOLE_0:def_4; then f . v = -infty by MESFUNC1:def_15; then A33: f . v in {-infty} by TARSKI:def_1; v in dom f by A32, MESFUNC1:def_15; hence v in f " {-infty} by A33, FUNCT_1:def_7; ::_thesis: verum end; then A34: E0 /\ (eq_dom (f,-infty)) c= f " {-infty} by TARSKI:def_3; now__::_thesis:_for_v_being_set_st_v_in_f_"_{-infty}_holds_ v_in_E0_/\_(eq_dom_(f,-infty)) let v be set ; ::_thesis: ( v in f " {-infty} implies v in E0 /\ (eq_dom (f,-infty)) ) assume A35: v in f " {-infty} ; ::_thesis: v in E0 /\ (eq_dom (f,-infty)) then A36: v in dom f by FUNCT_1:def_7; f . v in {-infty} by A35, FUNCT_1:def_7; then f . v = -infty by TARSKI:def_1; then v in eq_dom (f,-infty) by A36, MESFUNC1:def_15; hence v in E0 /\ (eq_dom (f,-infty)) by A23, A36, XBOOLE_0:def_4; ::_thesis: verum end; then A37: f " {-infty} c= E0 /\ (eq_dom (f,-infty)) by TARSKI:def_3; then A38: f " {-infty} = E0 /\ (eq_dom (f,-infty)) by A34, XBOOLE_0:def_10; A39: E0 /\ (eq_dom (f,-infty)) in S by A24, MESFUNC1:34; then A40: f " {-infty} in S by A37, A34, XBOOLE_0:def_10; then reconsider NF = (f " {+infty}) \/ (f " {-infty}) as Element of S by A25, A31, PROB_1:3; reconsider NFG = NF \/ NG as Element of S ; reconsider E = E0 \ NFG as Element of S ; reconsider DFPG = dom (f + g) as Element of S by A1, A2, A25, A31, A40, A22, A12, Th46; set g1 = g | E; set f1 = f | E; A41: (dom f) /\ E = E by A23, XBOOLE_1:28, XBOOLE_1:36; g " {-infty} c= NG by XBOOLE_1:7; then A42: (f " {+infty}) /\ (g " {-infty}) c= NG by A19, XBOOLE_1:1; f " {-infty} c= NF by XBOOLE_1:7; then (f " {-infty}) /\ (g " {+infty}) c= NF by A20, XBOOLE_1:1; then ((f " {-infty}) /\ (g " {+infty})) \/ ((f " {+infty}) /\ (g " {-infty})) c= NF \/ NG by A42, XBOOLE_1:13; then A43: E c= dom (f + g) by A3, A23, A21, XBOOLE_1:34; then A44: (f + g) | E = (f | E) + (g | E) by Th29; A45: dom ((f | E) + (g | E)) = E by A43, Th29; A46: E = dom ((f | E) + (g | E)) by A43, Th29; A47: for r being real number holds DFPG /\ (less_dom ((f + g),(R_EAL r))) = ((E /\ (less_dom (((f | E) + (g | E)),(R_EAL r)))) \/ ((f " {-infty}) /\ (DFPG \ (g " {+infty})))) \/ ((g " {-infty}) /\ (DFPG \ (f " {+infty}))) proof let r be real number ; ::_thesis: DFPG /\ (less_dom ((f + g),(R_EAL r))) = ((E /\ (less_dom (((f | E) + (g | E)),(R_EAL r)))) \/ ((f " {-infty}) /\ (DFPG \ (g " {+infty})))) \/ ((g " {-infty}) /\ (DFPG \ (f " {+infty}))) set SL = DFPG /\ (less_dom ((f + g),(R_EAL r))); set SR = ((E /\ (less_dom (((f | E) + (g | E)),(R_EAL r)))) \/ ((f " {-infty}) /\ (DFPG \ (g " {+infty})))) \/ ((g " {-infty}) /\ (DFPG \ (f " {+infty}))); A48: now__::_thesis:_for_x_being_set_st_x_in_((E_/\_(less_dom_(((f_|_E)_+_(g_|_E)),(R_EAL_r))))_\/_((f_"_{-infty})_/\_(DFPG_\_(g_"_{+infty}))))_\/_((g_"_{-infty})_/\_(DFPG_\_(f_"_{+infty})))_holds_ x_in_DFPG_/\_(less_dom_((f_+_g),(R_EAL_r))) let x be set ; ::_thesis: ( x in ((E /\ (less_dom (((f | E) + (g | E)),(R_EAL r)))) \/ ((f " {-infty}) /\ (DFPG \ (g " {+infty})))) \/ ((g " {-infty}) /\ (DFPG \ (f " {+infty}))) implies b1 in DFPG /\ (less_dom ((f + g),(R_EAL r))) ) assume x in ((E /\ (less_dom (((f | E) + (g | E)),(R_EAL r)))) \/ ((f " {-infty}) /\ (DFPG \ (g " {+infty})))) \/ ((g " {-infty}) /\ (DFPG \ (f " {+infty}))) ; ::_thesis: b1 in DFPG /\ (less_dom ((f + g),(R_EAL r))) then A49: ( x in (E /\ (less_dom (((f | E) + (g | E)),(R_EAL r)))) \/ ((f " {-infty}) /\ (DFPG \ (g " {+infty}))) or x in (g " {-infty}) /\ (DFPG \ (f " {+infty})) ) by XBOOLE_0:def_3; percases ( x in E /\ (less_dom (((f | E) + (g | E)),(R_EAL r))) or x in (f " {-infty}) /\ (DFPG \ (g " {+infty})) or x in (g " {-infty}) /\ (DFPG \ (f " {+infty})) ) by A49, XBOOLE_0:def_3; supposeA50: x in E /\ (less_dom (((f | E) + (g | E)),(R_EAL r))) ; ::_thesis: b1 in DFPG /\ (less_dom ((f + g),(R_EAL r))) then A51: x in E by XBOOLE_0:def_4; x in less_dom (((f | E) + (g | E)),(R_EAL r)) by A50, XBOOLE_0:def_4; then ((f | E) + (g | E)) . x < R_EAL r by MESFUNC1:def_11; then (f + g) . x < R_EAL r by A44, A45, A51, FUNCT_1:47; then x in less_dom ((f + g),(R_EAL r)) by A43, A51, MESFUNC1:def_11; hence x in DFPG /\ (less_dom ((f + g),(R_EAL r))) by A43, A51, XBOOLE_0:def_4; ::_thesis: verum end; supposeA52: ( x in (f " {-infty}) /\ (DFPG \ (g " {+infty})) or x in (g " {-infty}) /\ (DFPG \ (f " {+infty})) ) ; ::_thesis: b1 in DFPG /\ (less_dom ((f + g),(R_EAL r))) percases ( x in (f " {-infty}) /\ (DFPG \ (g " {+infty})) or x in (g " {-infty}) /\ (DFPG \ (f " {+infty})) ) by A52; supposeA53: x in (f " {-infty}) /\ (DFPG \ (g " {+infty})) ; ::_thesis: b1 in DFPG /\ (less_dom ((f + g),(R_EAL r))) R_EAL r in REAL by XREAL_0:def_1; then A54: -infty < R_EAL r by XXREAL_0:12; A55: x in DFPG \ (g " {+infty}) by A53, XBOOLE_0:def_4; then A56: x in DFPG by XBOOLE_0:def_5; then x in ((dom f) /\ (dom g)) \ (((f " {-infty}) /\ (g " {+infty})) \/ ((f " {+infty}) /\ (g " {-infty}))) by MESFUNC1:def_3; then A57: x in (dom f) /\ (dom g) by XBOOLE_0:def_5; then x in dom g by XBOOLE_0:def_4; then ( x in g " {+infty} iff g . x in {+infty} ) by FUNCT_1:def_7; then A58: ( x in g " {+infty} iff g . x = +infty ) by TARSKI:def_1; x in dom f by A57, XBOOLE_0:def_4; then ( x in f " {-infty} iff f . x in {-infty} ) by FUNCT_1:def_7; then ( x in f " {-infty} iff f . x = -infty ) by TARSKI:def_1; then (f . x) + (g . x) = -infty by A53, A55, A58, XBOOLE_0:def_4, XBOOLE_0:def_5, XXREAL_3:def_2; then (f + g) . x < R_EAL r by A56, A54, MESFUNC1:def_3; then x in less_dom ((f + g),(R_EAL r)) by A56, MESFUNC1:def_11; hence x in DFPG /\ (less_dom ((f + g),(R_EAL r))) by A56, XBOOLE_0:def_4; ::_thesis: verum end; supposeA59: x in (g " {-infty}) /\ (DFPG \ (f " {+infty})) ; ::_thesis: b1 in DFPG /\ (less_dom ((f + g),(R_EAL r))) R_EAL r in REAL by XREAL_0:def_1; then A60: -infty < R_EAL r by XXREAL_0:12; A61: x in DFPG \ (f " {+infty}) by A59, XBOOLE_0:def_4; then A62: x in DFPG by XBOOLE_0:def_5; A63: x in DFPG by A61, XBOOLE_0:def_5; then x in ((dom f) /\ (dom g)) \ (((f " {-infty}) /\ (g " {+infty})) \/ ((f " {+infty}) /\ (g " {-infty}))) by MESFUNC1:def_3; then A64: x in (dom f) /\ (dom g) by XBOOLE_0:def_5; then x in dom g by XBOOLE_0:def_4; then ( x in g " {-infty} iff g . x in {-infty} ) by FUNCT_1:def_7; then A65: ( x in g " {-infty} iff g . x = -infty ) by TARSKI:def_1; x in dom f by A64, XBOOLE_0:def_4; then ( x in f " {+infty} iff f . x in {+infty} ) by FUNCT_1:def_7; then ( x in f " {+infty} iff f . x = +infty ) by TARSKI:def_1; then (f . x) + (g . x) = -infty by A59, A61, A65, XBOOLE_0:def_4, XBOOLE_0:def_5, XXREAL_3:def_2; then (f + g) . x < R_EAL r by A62, A60, MESFUNC1:def_3; then x in less_dom ((f + g),(R_EAL r)) by A62, MESFUNC1:def_11; hence x in DFPG /\ (less_dom ((f + g),(R_EAL r))) by A63, XBOOLE_0:def_4; ::_thesis: verum end; end; end; end; end; now__::_thesis:_for_x_being_set_st_x_in_DFPG_/\_(less_dom_((f_+_g),(R_EAL_r)))_holds_ x_in_((E_/\_(less_dom_(((f_|_E)_+_(g_|_E)),(R_EAL_r))))_\/_((f_"_{-infty})_/\_(DFPG_\_(g_"_{+infty}))))_\/_((g_"_{-infty})_/\_(DFPG_\_(f_"_{+infty}))) let x be set ; ::_thesis: ( x in DFPG /\ (less_dom ((f + g),(R_EAL r))) implies b1 in ((E /\ (less_dom (((f | E) + (g | E)),(R_EAL r)))) \/ ((f " {-infty}) /\ (DFPG \ (g " {+infty})))) \/ ((g " {-infty}) /\ (DFPG \ (f " {+infty}))) ) assume A66: x in DFPG /\ (less_dom ((f + g),(R_EAL r))) ; ::_thesis: b1 in ((E /\ (less_dom (((f | E) + (g | E)),(R_EAL r)))) \/ ((f " {-infty}) /\ (DFPG \ (g " {+infty})))) \/ ((g " {-infty}) /\ (DFPG \ (f " {+infty}))) then A67: x in DFPG by XBOOLE_0:def_4; then A68: x in ((dom f) /\ (dom g)) \ (((f " {-infty}) /\ (g " {+infty})) \/ ((f " {+infty}) /\ (g " {-infty}))) by MESFUNC1:def_3; then A69: not x in ((f " {+infty}) /\ (g " {-infty})) \/ ((f " {-infty}) /\ (g " {+infty})) by XBOOLE_0:def_5; then A70: not x in (f " {+infty}) /\ (g " {-infty}) by XBOOLE_0:def_3; x in less_dom ((f + g),(R_EAL r)) by A66, XBOOLE_0:def_4; then A71: (f + g) . x < R_EAL r by MESFUNC1:def_11; then A72: (f . x) + (g . x) < R_EAL r by A67, MESFUNC1:def_3; A73: not x in (f " {-infty}) /\ (g " {+infty}) by A69, XBOOLE_0:def_3; A74: x in (dom f) /\ (dom g) by A68, XBOOLE_0:def_5; then A75: x in dom f by XBOOLE_0:def_4; then A76: ( x in f " {-infty} iff f . x in {-infty} ) by FUNCT_1:def_7; A77: ( x in f " {+infty} iff f . x in {+infty} ) by A75, FUNCT_1:def_7; then A78: ( x in f " {+infty} iff f . x = +infty ) by TARSKI:def_1; A79: x in dom g by A74, XBOOLE_0:def_4; then A80: ( x in g " {-infty} iff g . x in {-infty} ) by FUNCT_1:def_7; A81: ( x in g " {+infty} iff g . x in {+infty} ) by A79, FUNCT_1:def_7; then A82: ( x in g " {+infty} iff g . x = +infty ) by TARSKI:def_1; percases ( f . x = -infty or f . x <> -infty ) ; supposeA83: f . x = -infty ; ::_thesis: b1 in ((E /\ (less_dom (((f | E) + (g | E)),(R_EAL r)))) \/ ((f " {-infty}) /\ (DFPG \ (g " {+infty})))) \/ ((g " {-infty}) /\ (DFPG \ (f " {+infty}))) then x in DFPG \ (g " {+infty}) by A67, A76, A81, A73, TARSKI:def_1, XBOOLE_0:def_4, XBOOLE_0:def_5; then x in (f " {-infty}) /\ (DFPG \ (g " {+infty})) by A76, A83, TARSKI:def_1, XBOOLE_0:def_4; then x in (E /\ (less_dom (((f | E) + (g | E)),(R_EAL r)))) \/ ((f " {-infty}) /\ (DFPG \ (g " {+infty}))) by XBOOLE_0:def_3; hence x in ((E /\ (less_dom (((f | E) + (g | E)),(R_EAL r)))) \/ ((f " {-infty}) /\ (DFPG \ (g " {+infty})))) \/ ((g " {-infty}) /\ (DFPG \ (f " {+infty}))) by XBOOLE_0:def_3; ::_thesis: verum end; supposeA84: f . x <> -infty ; ::_thesis: b1 in ((E /\ (less_dom (((f | E) + (g | E)),(R_EAL r)))) \/ ((f " {-infty}) /\ (DFPG \ (g " {+infty})))) \/ ((g " {-infty}) /\ (DFPG \ (f " {+infty}))) percases ( g . x = -infty or g . x <> -infty ) ; supposeA85: g . x = -infty ; ::_thesis: b1 in ((E /\ (less_dom (((f | E) + (g | E)),(R_EAL r)))) \/ ((f " {-infty}) /\ (DFPG \ (g " {+infty})))) \/ ((g " {-infty}) /\ (DFPG \ (f " {+infty}))) then x in DFPG \ (f " {+infty}) by A67, A77, A80, A70, TARSKI:def_1, XBOOLE_0:def_4, XBOOLE_0:def_5; then x in (g " {-infty}) /\ (DFPG \ (f " {+infty})) by A80, A85, TARSKI:def_1, XBOOLE_0:def_4; hence x in ((E /\ (less_dom (((f | E) + (g | E)),(R_EAL r)))) \/ ((f " {-infty}) /\ (DFPG \ (g " {+infty})))) \/ ((g " {-infty}) /\ (DFPG \ (f " {+infty}))) by XBOOLE_0:def_3; ::_thesis: verum end; supposeA86: g . x <> -infty ; ::_thesis: b1 in ((E /\ (less_dom (((f | E) + (g | E)),(R_EAL r)))) \/ ((f " {-infty}) /\ (DFPG \ (g " {+infty})))) \/ ((g " {-infty}) /\ (DFPG \ (f " {+infty}))) then not x in (f " {-infty}) \/ (f " {+infty}) by A76, A78, A72, A84, TARSKI:def_1, XBOOLE_0:def_3, XXREAL_3:52; then not x in ((f " {-infty}) \/ (f " {+infty})) \/ (g " {-infty}) by A80, A86, TARSKI:def_1, XBOOLE_0:def_3; then not x in (((f " {-infty}) \/ (f " {+infty})) \/ (g " {-infty})) \/ (g " {+infty}) by A82, A72, A84, XBOOLE_0:def_3, XXREAL_3:52; then not x in NFG by XBOOLE_1:4; then A87: x in E by A23, A75, XBOOLE_0:def_5; then ((f | E) + (g | E)) . x = (f + g) . x by A44, A45, FUNCT_1:47; then x in less_dom (((f | E) + (g | E)),(R_EAL r)) by A46, A71, A87, MESFUNC1:def_11; then x in E /\ (less_dom (((f | E) + (g | E)),(R_EAL r))) by A87, XBOOLE_0:def_4; then x in (E /\ (less_dom (((f | E) + (g | E)),(R_EAL r)))) \/ ((f " {-infty}) /\ (DFPG \ (g " {+infty}))) by XBOOLE_0:def_3; hence x in ((E /\ (less_dom (((f | E) + (g | E)),(R_EAL r)))) \/ ((f " {-infty}) /\ (DFPG \ (g " {+infty})))) \/ ((g " {-infty}) /\ (DFPG \ (f " {+infty}))) by XBOOLE_0:def_3; ::_thesis: verum end; end; end; end; end; hence DFPG /\ (less_dom ((f + g),(R_EAL r))) = ((E /\ (less_dom (((f | E) + (g | E)),(R_EAL r)))) \/ ((f " {-infty}) /\ (DFPG \ (g " {+infty})))) \/ ((g " {-infty}) /\ (DFPG \ (f " {+infty}))) by A48, TARSKI:1; ::_thesis: verum end; A88: now__::_thesis:_for_x_being_set_st_x_in_dom_(f_|_E)_holds_ (_-infty_<_(f_|_E)_._x_&_(f_|_E)_._x_<_+infty_) let x be set ; ::_thesis: ( x in dom (f | E) implies ( -infty < (f | E) . x & (f | E) . x < +infty ) ) for x being set st x in dom f holds f . x in ExtREAL ; then reconsider ff = f as Function of (dom f),ExtREAL by FUNCT_2:3; assume A89: x in dom (f | E) ; ::_thesis: ( -infty < (f | E) . x & (f | E) . x < +infty ) then A90: x in (dom f) /\ E by RELAT_1:61; then A91: x in dom f by XBOOLE_0:def_4; x in E by A90, XBOOLE_0:def_4; then A92: not x in NFG by XBOOLE_0:def_5; A93: now__::_thesis:_not_(f_|_E)_._x_=_-infty assume (f | E) . x = -infty ; ::_thesis: contradiction then f . x = -infty by A89, FUNCT_1:47; then ff . x in {-infty} by TARSKI:def_1; then A94: x in ff " {-infty} by A91, FUNCT_2:38; f " {-infty} c= NF by XBOOLE_1:7; hence contradiction by A92, A94, XBOOLE_0:def_3; ::_thesis: verum end; now__::_thesis:_not_(f_|_E)_._x_=_+infty assume (f | E) . x = +infty ; ::_thesis: contradiction then f . x = +infty by A89, FUNCT_1:47; then f . x in {+infty} by TARSKI:def_1; then A95: x in ff " {+infty} by A91, FUNCT_2:38; f " {+infty} c= NF by XBOOLE_1:7; hence contradiction by A92, A95, XBOOLE_0:def_3; ::_thesis: verum end; hence ( -infty < (f | E) . x & (f | E) . x < +infty ) by A93, XXREAL_0:4, XXREAL_0:6; ::_thesis: verum end; now__::_thesis:_for_x_being_Element_of_X_st_x_in_dom_(f_|_E)_holds_ |.((f_|_E)_._x).|_<_+infty let x be Element of X; ::_thesis: ( x in dom (f | E) implies |.((f | E) . x).| < +infty ) A96: - +infty = -infty by XXREAL_3:def_3; assume A97: x in dom (f | E) ; ::_thesis: |.((f | E) . x).| < +infty then A98: (f | E) . x < +infty by A88; -infty < (f | E) . x by A88, A97; hence |.((f | E) . x).| < +infty by A98, A96, EXTREAL2:11; ::_thesis: verum end; then A99: f | E is V60() by MESFUNC2:def_1; A100: now__::_thesis:_for_x_being_set_st_x_in_dom_(g_|_E)_holds_ (_-infty_<_(g_|_E)_._x_&_(g_|_E)_._x_<_+infty_) let x be set ; ::_thesis: ( x in dom (g | E) implies ( -infty < (g | E) . x & (g | E) . x < +infty ) ) for x being set st x in dom g holds g . x in ExtREAL ; then reconsider gg = g as Function of (dom g),ExtREAL by FUNCT_2:3; assume A101: x in dom (g | E) ; ::_thesis: ( -infty < (g | E) . x & (g | E) . x < +infty ) then A102: x in (dom g) /\ E by RELAT_1:61; then A103: x in dom g by XBOOLE_0:def_4; x in E by A102, XBOOLE_0:def_4; then A104: not x in NFG by XBOOLE_0:def_5; A105: now__::_thesis:_not_(g_|_E)_._x_=_-infty assume (g | E) . x = -infty ; ::_thesis: contradiction then g . x = -infty by A101, FUNCT_1:47; then gg . x in {-infty} by TARSKI:def_1; then A106: x in gg " {-infty} by A103, FUNCT_2:38; g " {-infty} c= NG by XBOOLE_1:7; hence contradiction by A104, A106, XBOOLE_0:def_3; ::_thesis: verum end; now__::_thesis:_not_(g_|_E)_._x_=_+infty assume (g | E) . x = +infty ; ::_thesis: contradiction then g . x = +infty by A101, FUNCT_1:47; then gg . x in {+infty} by TARSKI:def_1; then A107: x in gg " {+infty} by A103, FUNCT_2:38; g " {+infty} c= NG by XBOOLE_1:7; hence contradiction by A104, A107, XBOOLE_0:def_3; ::_thesis: verum end; hence ( -infty < (g | E) . x & (g | E) . x < +infty ) by A105, XXREAL_0:4, XXREAL_0:6; ::_thesis: verum end; now__::_thesis:_for_x_being_Element_of_X_st_x_in_dom_(g_|_E)_holds_ |.((g_|_E)_._x).|_<_+infty let x be Element of X; ::_thesis: ( x in dom (g | E) implies |.((g | E) . x).| < +infty ) A108: - +infty = -infty by XXREAL_3:def_3; assume A109: x in dom (g | E) ; ::_thesis: |.((g | E) . x).| < +infty then A110: (g | E) . x < +infty by A100; -infty < (g | E) . x by A100, A109; hence |.((g | E) . x).| < +infty by A110, A108, EXTREAL2:11; ::_thesis: verum end; then A111: g | E is V60() by MESFUNC2:def_1; f is_measurable_on E by A1, A23, MESFUNC1:30, XBOOLE_1:36; then A112: f | E is_measurable_on E by A41, Th42; A113: (dom g) /\ E = E by A3, A23, XBOOLE_1:28, XBOOLE_1:36; g is_measurable_on E by A2, A3, A23, MESFUNC1:30, XBOOLE_1:36; then g | E is_measurable_on E by A113, Th42; then A114: (f | E) + (g | E) is_measurable_on E by A112, A99, A111, MESFUNC2:7; now__::_thesis:_for_r_being_real_number_holds_DFPG_/\_(less_dom_((f_+_g),(R_EAL_r)))_in_S let r be real number ; ::_thesis: DFPG /\ (less_dom ((f + g),(R_EAL r))) in S A115: E /\ (less_dom (((f | E) + (g | E)),(R_EAL r))) in S by A114, MESFUNC1:def_16; DFPG \ (f " {+infty}) in S by A25, A31, PROB_1:6; then A116: (g " {-infty}) /\ (DFPG \ (f " {+infty})) in S by A12, FINSUB_1:def_2; DFPG \ (g " {+infty}) in S by A22, PROB_1:6; then (f " {-infty}) /\ (DFPG \ (g " {+infty})) in S by A39, A38, FINSUB_1:def_2; then ((f " {-infty}) /\ (DFPG \ (g " {+infty}))) \/ ((g " {-infty}) /\ (DFPG \ (f " {+infty}))) in S by A116, PROB_1:3; then A117: (E /\ (less_dom (((f | E) + (g | E)),(R_EAL r)))) \/ (((f " {-infty}) /\ (DFPG \ (g " {+infty}))) \/ ((g " {-infty}) /\ (DFPG \ (f " {+infty})))) in S by A115, PROB_1:3; DFPG /\ (less_dom ((f + g),(R_EAL r))) = ((E /\ (less_dom (((f | E) + (g | E)),(R_EAL r)))) \/ ((f " {-infty}) /\ (DFPG \ (g " {+infty})))) \/ ((g " {-infty}) /\ (DFPG \ (f " {+infty}))) by A47; hence DFPG /\ (less_dom ((f + g),(R_EAL r))) in S by A117, XBOOLE_1:4; ::_thesis: verum end; then f + g is_measurable_on DFPG by MESFUNC1:def_16; hence ex DFPG being Element of S st ( DFPG = dom (f + g) & f + g is_measurable_on DFPG ) ; ::_thesis: verum end; theorem Th47: :: MESFUNC5:47 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st ex E1 being Element of S st ( E1 = dom f & f is_measurable_on E1 ) & ex E2 being Element of S st ( E2 = dom g & g is_measurable_on E2 ) holds ex E being Element of S st ( E = dom (f + g) & f + g is_measurable_on E ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st ex E1 being Element of S st ( E1 = dom f & f is_measurable_on E1 ) & ex E2 being Element of S st ( E2 = dom g & g is_measurable_on E2 ) holds ex E being Element of S st ( E = dom (f + g) & f + g is_measurable_on E ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st ex E1 being Element of S st ( E1 = dom f & f is_measurable_on E1 ) & ex E2 being Element of S st ( E2 = dom g & g is_measurable_on E2 ) holds ex E being Element of S st ( E = dom (f + g) & f + g is_measurable_on E ) let M be sigma_Measure of S; ::_thesis: for f, g being PartFunc of X,ExtREAL st ex E1 being Element of S st ( E1 = dom f & f is_measurable_on E1 ) & ex E2 being Element of S st ( E2 = dom g & g is_measurable_on E2 ) holds ex E being Element of S st ( E = dom (f + g) & f + g is_measurable_on E ) let f, g be PartFunc of X,ExtREAL; ::_thesis: ( ex E1 being Element of S st ( E1 = dom f & f is_measurable_on E1 ) & ex E2 being Element of S st ( E2 = dom g & g is_measurable_on E2 ) implies ex E being Element of S st ( E = dom (f + g) & f + g is_measurable_on E ) ) assume that A1: ex E1 being Element of S st ( E1 = dom f & f is_measurable_on E1 ) and A2: ex E2 being Element of S st ( E2 = dom g & g is_measurable_on E2 ) ; ::_thesis: ex E being Element of S st ( E = dom (f + g) & f + g is_measurable_on E ) consider E1 being Element of S such that A3: E1 = dom f and A4: f is_measurable_on E1 by A1; consider E2 being Element of S such that A5: E2 = dom g and A6: g is_measurable_on E2 by A2; set E3 = E1 /\ E2; set g1 = g | (E1 /\ E2); A7: (g | (E1 /\ E2)) " {-infty} = (E1 /\ E2) /\ (g " {-infty}) by FUNCT_1:70; set f1 = f | (E1 /\ E2); dom (f | (E1 /\ E2)) = (dom f) /\ (E1 /\ E2) by RELAT_1:61; then A8: dom (f | (E1 /\ E2)) = E1 /\ E2 by A3, XBOOLE_1:17, XBOOLE_1:28; g is_measurable_on E1 /\ E2 by A6, MESFUNC1:30, XBOOLE_1:17; then A9: g | (E1 /\ E2) is_measurable_on E1 /\ E2 by Lm6; A10: (g | (E1 /\ E2)) " {+infty} = (E1 /\ E2) /\ (g " {+infty}) by FUNCT_1:70; dom (g | (E1 /\ E2)) = (dom g) /\ (E1 /\ E2) by RELAT_1:61; then A11: dom (g | (E1 /\ E2)) = E1 /\ E2 by A5, XBOOLE_1:17, XBOOLE_1:28; (f | (E1 /\ E2)) " {+infty} = (E1 /\ E2) /\ (f " {+infty}) by FUNCT_1:70; then A12: ((f | (E1 /\ E2)) " {+infty}) /\ ((g | (E1 /\ E2)) " {-infty}) = (f " {+infty}) /\ ((E1 /\ E2) /\ ((E1 /\ E2) /\ (g " {-infty}))) by A7, XBOOLE_1:16 .= (f " {+infty}) /\ (((E1 /\ E2) /\ (E1 /\ E2)) /\ (g " {-infty})) by XBOOLE_1:16 .= ((f " {+infty}) /\ (g " {-infty})) /\ (E1 /\ E2) by XBOOLE_1:16 ; A13: dom ((f | (E1 /\ E2)) + (g | (E1 /\ E2))) = ((dom (f | (E1 /\ E2))) /\ (dom (g | (E1 /\ E2)))) \ ((((f | (E1 /\ E2)) " {-infty}) /\ ((g | (E1 /\ E2)) " {+infty})) \/ (((f | (E1 /\ E2)) " {+infty}) /\ ((g | (E1 /\ E2)) " {-infty}))) by MESFUNC1:def_3; f is_measurable_on E1 /\ E2 by A4, MESFUNC1:30, XBOOLE_1:17; then f | (E1 /\ E2) is_measurable_on E1 /\ E2 by Lm6; then consider E being Element of S such that A14: E = dom ((f | (E1 /\ E2)) + (g | (E1 /\ E2))) and A15: (f | (E1 /\ E2)) + (g | (E1 /\ E2)) is_measurable_on E by A9, A8, A11, Lm7; take E ; ::_thesis: ( E = dom (f + g) & f + g is_measurable_on E ) A16: dom ((f + g) | E) = (dom (f + g)) /\ E by RELAT_1:61; (f | (E1 /\ E2)) " {-infty} = (E1 /\ E2) /\ (f " {-infty}) by FUNCT_1:70; then ((f | (E1 /\ E2)) " {-infty}) /\ ((g | (E1 /\ E2)) " {+infty}) = (f " {-infty}) /\ ((E1 /\ E2) /\ ((E1 /\ E2) /\ (g " {+infty}))) by A10, XBOOLE_1:16 .= (f " {-infty}) /\ (((E1 /\ E2) /\ (E1 /\ E2)) /\ (g " {+infty})) by XBOOLE_1:16 .= ((f " {-infty}) /\ (g " {+infty})) /\ (E1 /\ E2) by XBOOLE_1:16 ; then A17: (((f | (E1 /\ E2)) " {-infty}) /\ ((g | (E1 /\ E2)) " {+infty})) \/ (((f | (E1 /\ E2)) " {+infty}) /\ ((g | (E1 /\ E2)) " {-infty})) = (E1 /\ E2) /\ (((f " {-infty}) /\ (g " {+infty})) \/ ((f " {+infty}) /\ (g " {-infty}))) by A12, XBOOLE_1:23; A18: dom (f + g) = ((dom f) /\ (dom g)) \ (((f " {-infty}) /\ (g " {+infty})) \/ ((f " {+infty}) /\ (g " {-infty}))) by MESFUNC1:def_3; then A19: dom (f + g) = E by A3, A5, A8, A11, A14, A13, A17, XBOOLE_1:47; now__::_thesis:_for_v_being_Element_of_X_st_v_in_dom_((f_+_g)_|_E)_holds_ ((f_+_g)_|_E)_._v_=_((f_|_(E1_/\_E2))_+_(g_|_(E1_/\_E2)))_._v let v be Element of X; ::_thesis: ( v in dom ((f + g) | E) implies ((f + g) | E) . v = ((f | (E1 /\ E2)) + (g | (E1 /\ E2))) . v ) assume A20: v in dom ((f + g) | E) ; ::_thesis: ((f + g) | E) . v = ((f | (E1 /\ E2)) + (g | (E1 /\ E2))) . v then A21: v in (dom (f + g)) /\ E by RELAT_1:61; then A22: v in dom (f + g) by XBOOLE_0:def_4; A23: ((f + g) | E) . v = (f + g) . v by A20, FUNCT_1:47 .= (f . v) + (g . v) by A22, MESFUNC1:def_3 ; A24: v in E by A21, XBOOLE_0:def_4; A25: E c= E1 /\ E2 by A8, A11, A14, A13, XBOOLE_1:36; ((f | (E1 /\ E2)) + (g | (E1 /\ E2))) . v = ((f | (E1 /\ E2)) . v) + ((g | (E1 /\ E2)) . v) by A14, A19, A16, A20, MESFUNC1:def_3 .= (f . v) + ((g | (E1 /\ E2)) . v) by A8, A24, A25, FUNCT_1:47 ; hence ((f + g) | E) . v = ((f | (E1 /\ E2)) + (g | (E1 /\ E2))) . v by A11, A24, A25, A23, FUNCT_1:47; ::_thesis: verum end; then (f + g) | E = (f | (E1 /\ E2)) + (g | (E1 /\ E2)) by A14, A19, A16, PARTFUN1:5; hence ( E = dom (f + g) & f + g is_measurable_on E ) by A3, A5, A8, A11, A14, A15, A13, A17, A18, Lm6, XBOOLE_1:47; ::_thesis: verum end; theorem Th48: :: MESFUNC5:48 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A, B being Element of S st dom f = A holds ( f is_measurable_on B iff f is_measurable_on A /\ B ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A, B being Element of S st dom f = A holds ( f is_measurable_on B iff f is_measurable_on A /\ B ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A, B being Element of S st dom f = A holds ( f is_measurable_on B iff f is_measurable_on A /\ B ) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for A, B being Element of S st dom f = A holds ( f is_measurable_on B iff f is_measurable_on A /\ B ) let f be PartFunc of X,ExtREAL; ::_thesis: for A, B being Element of S st dom f = A holds ( f is_measurable_on B iff f is_measurable_on A /\ B ) let A, B be Element of S; ::_thesis: ( dom f = A implies ( f is_measurable_on B iff f is_measurable_on A /\ B ) ) assume A1: dom f = A ; ::_thesis: ( f is_measurable_on B iff f is_measurable_on A /\ B ) A2: now__::_thesis:_for_r_being_real_number_holds_A_/\_(less_dom_(f,(R_EAL_r)))_=_less_dom_(f,(R_EAL_r)) let r be real number ; ::_thesis: A /\ (less_dom (f,(R_EAL r))) = less_dom (f,(R_EAL r)) A3: now__::_thesis:_for_x_being_set_holds_ (_x_in_A_/\_(less_dom_(f,(R_EAL_r)))_iff_x_in_less_dom_(f,(R_EAL_r))_) let x be set ; ::_thesis: ( x in A /\ (less_dom (f,(R_EAL r))) iff x in less_dom (f,(R_EAL r)) ) ( x in A /\ (less_dom (f,(R_EAL r))) iff ( x in A & x in less_dom (f,(R_EAL r)) ) ) by XBOOLE_0:def_4; hence ( x in A /\ (less_dom (f,(R_EAL r))) iff x in less_dom (f,(R_EAL r)) ) by A1, MESFUNC1:def_11; ::_thesis: verum end; then A4: less_dom (f,(R_EAL r)) c= A /\ (less_dom (f,(R_EAL r))) by TARSKI:def_3; A /\ (less_dom (f,(R_EAL r))) c= less_dom (f,(R_EAL r)) by A3, TARSKI:def_3; hence A /\ (less_dom (f,(R_EAL r))) = less_dom (f,(R_EAL r)) by A4, XBOOLE_0:def_10; ::_thesis: verum end; hereby ::_thesis: ( f is_measurable_on A /\ B implies f is_measurable_on B ) assume A5: f is_measurable_on B ; ::_thesis: f is_measurable_on A /\ B now__::_thesis:_for_r_being_real_number_holds_(A_/\_B)_/\_(less_dom_(f,(R_EAL_r)))_in_S let r be real number ; ::_thesis: (A /\ B) /\ (less_dom (f,(R_EAL r))) in S (A /\ B) /\ (less_dom (f,(R_EAL r))) = B /\ (A /\ (less_dom (f,(R_EAL r)))) by XBOOLE_1:16 .= B /\ (less_dom (f,(R_EAL r))) by A2 ; hence (A /\ B) /\ (less_dom (f,(R_EAL r))) in S by A5, MESFUNC1:def_16; ::_thesis: verum end; hence f is_measurable_on A /\ B by MESFUNC1:def_16; ::_thesis: verum end; assume A6: f is_measurable_on A /\ B ; ::_thesis: f is_measurable_on B now__::_thesis:_for_r_being_real_number_holds_B_/\_(less_dom_(f,(R_EAL_r)))_in_S let r be real number ; ::_thesis: B /\ (less_dom (f,(R_EAL r))) in S (A /\ B) /\ (less_dom (f,(R_EAL r))) = B /\ (A /\ (less_dom (f,(R_EAL r)))) by XBOOLE_1:16 .= B /\ (less_dom (f,(R_EAL r))) by A2 ; hence B /\ (less_dom (f,(R_EAL r))) in S by A6, MESFUNC1:def_16; ::_thesis: verum end; hence f is_measurable_on B by MESFUNC1:def_16; ::_thesis: verum end; theorem Th49: :: MESFUNC5:49 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st ex A being Element of S st dom f = A holds for c being Real for B being Element of S st f is_measurable_on B holds c (#) f is_measurable_on B proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st ex A being Element of S st dom f = A holds for c being Real for B being Element of S st f is_measurable_on B holds c (#) f is_measurable_on B let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st ex A being Element of S st dom f = A holds for c being Real for B being Element of S st f is_measurable_on B holds c (#) f is_measurable_on B let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL st ex A being Element of S st dom f = A holds for c being Real for B being Element of S st f is_measurable_on B holds c (#) f is_measurable_on B let f be PartFunc of X,ExtREAL; ::_thesis: ( ex A being Element of S st dom f = A implies for c being Real for B being Element of S st f is_measurable_on B holds c (#) f is_measurable_on B ) assume ex A being Element of S st A = dom f ; ::_thesis: for c being Real for B being Element of S st f is_measurable_on B holds c (#) f is_measurable_on B then consider A being Element of S such that A1: A = dom f ; let c be Real; ::_thesis: for B being Element of S st f is_measurable_on B holds c (#) f is_measurable_on B let B be Element of S; ::_thesis: ( f is_measurable_on B implies c (#) f is_measurable_on B ) assume f is_measurable_on B ; ::_thesis: c (#) f is_measurable_on B then f is_measurable_on A /\ B by A1, Th48; then A2: c (#) f is_measurable_on A /\ B by A1, MESFUNC1:37, XBOOLE_1:17; dom (c (#) f) = A by A1, MESFUNC1:def_6; hence c (#) f is_measurable_on B by A2, Th48; ::_thesis: verum end; begin definition mode ExtREAL_sequence is Function of NAT,ExtREAL; end; definition let seq be ExtREAL_sequence; attrseq is convergent_to_finite_number means :Def8: :: MESFUNC5:def 8 ex g being real number st for p being real number st 0 < p holds ex n being Nat st for m being Nat st n <= m holds |.((seq . m) - (R_EAL g)).| < p; end; :: deftheorem Def8 defines convergent_to_finite_number MESFUNC5:def_8_:_ for seq being ExtREAL_sequence holds ( seq is convergent_to_finite_number iff ex g being real number st for p being real number st 0 < p holds ex n being Nat st for m being Nat st n <= m holds |.((seq . m) - (R_EAL g)).| < p ); definition let seq be ExtREAL_sequence; attrseq is convergent_to_+infty means :Def9: :: MESFUNC5:def 9 for g being real number st 0 < g holds ex n being Nat st for m being Nat st n <= m holds g <= seq . m; end; :: deftheorem Def9 defines convergent_to_+infty MESFUNC5:def_9_:_ for seq being ExtREAL_sequence holds ( seq is convergent_to_+infty iff for g being real number st 0 < g holds ex n being Nat st for m being Nat st n <= m holds g <= seq . m ); definition let seq be ExtREAL_sequence; attrseq is convergent_to_-infty means :Def10: :: MESFUNC5:def 10 for g being real number st g < 0 holds ex n being Nat st for m being Nat st n <= m holds seq . m <= g; end; :: deftheorem Def10 defines convergent_to_-infty MESFUNC5:def_10_:_ for seq being ExtREAL_sequence holds ( seq is convergent_to_-infty iff for g being real number st g < 0 holds ex n being Nat st for m being Nat st n <= m holds seq . m <= g ); theorem Th50: :: MESFUNC5:50 for seq being ExtREAL_sequence st seq is convergent_to_+infty holds ( not seq is convergent_to_-infty & not seq is convergent_to_finite_number ) proof let seq be ExtREAL_sequence; ::_thesis: ( seq is convergent_to_+infty implies ( not seq is convergent_to_-infty & not seq is convergent_to_finite_number ) ) assume A1: seq is convergent_to_+infty ; ::_thesis: ( not seq is convergent_to_-infty & not seq is convergent_to_finite_number ) hereby ::_thesis: not seq is convergent_to_finite_number assume seq is convergent_to_-infty ; ::_thesis: contradiction then consider n1 being Nat such that A2: for m being Nat st n1 <= m holds seq . m <= R_EAL (- 1) by Def10; consider n2 being Nat such that A3: for m being Nat st n2 <= m holds R_EAL 1 <= seq . m by A1, Def9; reconsider n1 = n1, n2 = n2 as Element of NAT by ORDINAL1:def_12; set m = max (n1,n2); seq . (max (n1,n2)) <= R_EAL (- 1) by A2, XXREAL_0:25; hence contradiction by A3, XXREAL_0:25; ::_thesis: verum end; assume seq is convergent_to_finite_number ; ::_thesis: contradiction then consider g being real number such that A4: for p being real number st 0 < p holds ex n being Nat st for m being Nat st n <= m holds |.((seq . m) - (R_EAL g)).| < p by Def8; percases ( g > 0 or g = 0 or g < 0 ) ; supposeA5: g > 0 ; ::_thesis: contradiction then consider n1 being Nat such that A6: for m being Nat st n1 <= m holds |.((seq . m) - (R_EAL g)).| < R_EAL g by A4; A7: now__::_thesis:_for_m_being_Nat_st_n1_<=_m_holds_ seq_._m_<_R_EAL_(2_*_g) let m be Nat; ::_thesis: ( n1 <= m implies seq . m < R_EAL (2 * g) ) assume n1 <= m ; ::_thesis: seq . m < R_EAL (2 * g) then |.((seq . m) - (R_EAL g)).| < R_EAL g by A6; then (seq . m) - (R_EAL g) < R_EAL g by EXTREAL2:10; then seq . m < R_EAL (g + g) by XXREAL_3:54; hence seq . m < R_EAL (2 * g) ; ::_thesis: verum end; consider n2 being Nat such that A8: for m being Nat st n2 <= m holds R_EAL (2 * g) <= seq . m by A1, A5, Def9; reconsider n1 = n1, n2 = n2 as Element of NAT by ORDINAL1:def_12; set m = max (n1,n2); seq . (max (n1,n2)) < R_EAL (2 * g) by A7, XXREAL_0:25; hence contradiction by A8, XXREAL_0:25; ::_thesis: verum end; supposeA9: g = 0 ; ::_thesis: contradiction consider n1 being Nat such that A10: for m being Nat st n1 <= m holds |.((seq . m) - (R_EAL g)).| < R_EAL 1 by A4; consider n2 being Nat such that A11: for m being Nat st n2 <= m holds R_EAL 1 <= seq . m by A1, Def9; reconsider n1 = n1, n2 = n2 as Element of NAT by ORDINAL1:def_12; set m = max (n1,n2); |.((seq . (max (n1,n2))) - (R_EAL g)).| < R_EAL 1 by A10, XXREAL_0:25; then (seq . (max (n1,n2))) - (R_EAL g) < R_EAL 1 by EXTREAL2:10; then seq . (max (n1,n2)) < (R_EAL 1) + (R_EAL g) by XXREAL_3:54; then seq . (max (n1,n2)) < R_EAL 1 by A9, XXREAL_3:4; hence contradiction by A11, XXREAL_0:25; ::_thesis: verum end; supposeA12: g < 0 ; ::_thesis: contradiction set g1 = - g; consider n1 being Nat such that A13: for m being Nat st n1 <= m holds |.((seq . m) - (R_EAL g)).| < R_EAL (- g) by A4, A12; A14: now__::_thesis:_for_m_being_Element_of_NAT_st_n1_<=_m_holds_ seq_._m_<_0 let m be Element of NAT ; ::_thesis: ( n1 <= m implies seq . m < 0 ) assume n1 <= m ; ::_thesis: seq . m < 0 then |.((seq . m) - (R_EAL g)).| < R_EAL (- g) by A13; then (seq . m) - (R_EAL g) < R_EAL (- g) by EXTREAL2:10; then seq . m < (- g) + g by XXREAL_3:54; hence seq . m < 0 ; ::_thesis: verum end; consider n2 being Nat such that A15: for m being Nat st n2 <= m holds 1 <= seq . m by A1, Def9; reconsider n1 = n1, n2 = n2 as Element of NAT by ORDINAL1:def_12; set m = max (n1,n2); seq . (max (n1,n2)) < 0 by A14, XXREAL_0:25; hence contradiction by A15, XXREAL_0:25; ::_thesis: verum end; end; end; theorem Th51: :: MESFUNC5:51 for seq being ExtREAL_sequence st seq is convergent_to_-infty holds ( not seq is convergent_to_+infty & not seq is convergent_to_finite_number ) proof let seq be ExtREAL_sequence; ::_thesis: ( seq is convergent_to_-infty implies ( not seq is convergent_to_+infty & not seq is convergent_to_finite_number ) ) assume A1: seq is convergent_to_-infty ; ::_thesis: ( not seq is convergent_to_+infty & not seq is convergent_to_finite_number ) hereby ::_thesis: not seq is convergent_to_finite_number assume seq is convergent_to_+infty ; ::_thesis: contradiction then consider n1 being Nat such that A2: for m being Nat st n1 <= m holds R_EAL 1 <= seq . m by Def9; consider n2 being Nat such that A3: for m being Nat st n2 <= m holds seq . m <= R_EAL (- 1) by A1, Def10; reconsider n1 = n1, n2 = n2 as Element of NAT by ORDINAL1:def_12; set m = max (n1,n2); seq . (max (n1,n2)) <= R_EAL (- 1) by A3, XXREAL_0:25; hence contradiction by A2, XXREAL_0:25; ::_thesis: verum end; assume seq is convergent_to_finite_number ; ::_thesis: contradiction then consider g being real number such that A4: for p being real number st 0 < p holds ex n being Nat st for m being Nat st n <= m holds |.((seq . m) - (R_EAL g)).| < p by Def8; reconsider g1 = - g as Real by XREAL_0:def_1; percases ( g > 0 or g = 0 or g < 0 ) ; supposeA5: g > 0 ; ::_thesis: contradiction then consider n1 being Nat such that A6: for m being Nat st n1 <= m holds |.((seq . m) - (R_EAL g)).| < R_EAL g by A4; A7: now__::_thesis:_for_m_being_Element_of_NAT_st_n1_<=_m_holds_ 0_<_seq_._m let m be Element of NAT ; ::_thesis: ( n1 <= m implies 0 < seq . m ) assume n1 <= m ; ::_thesis: 0 < seq . m then |.((seq . m) - (R_EAL g)).| < R_EAL g by A6; then - (R_EAL g) < (seq . m) - (R_EAL g) by EXTREAL2:10; then (- g) + g < seq . m by XXREAL_3:53; hence 0 < seq . m ; ::_thesis: verum end; consider n2 being Nat such that A8: for m being Nat st n2 <= m holds seq . m <= R_EAL g1 by A1, A5, Def10; reconsider n1 = n1, n2 = n2 as Element of NAT by ORDINAL1:def_12; set m = max (n1,n2); 0 < seq . (max (n1,n2)) by A7, XXREAL_0:25; hence contradiction by A5, A8, XXREAL_0:25; ::_thesis: verum end; supposeA9: g = 0 ; ::_thesis: contradiction consider n1 being Nat such that A10: for m being Nat st n1 <= m holds |.((seq . m) - (R_EAL g)).| < R_EAL 1 by A4; consider n2 being Nat such that A11: for m being Nat st n2 <= m holds seq . m <= R_EAL (- 1) by A1, Def10; reconsider n1 = n1, n2 = n2 as Element of NAT by ORDINAL1:def_12; set m = max (n1,n2); |.((seq . (max (n1,n2))) - (R_EAL g)).| < R_EAL 1 by A10, XXREAL_0:25; then - (R_EAL 1) < (seq . (max (n1,n2))) - (R_EAL g) by EXTREAL2:10; then (- (R_EAL 1)) + (R_EAL g) < seq . (max (n1,n2)) by XXREAL_3:53; then - (R_EAL 1) < seq . (max (n1,n2)) by A9, XXREAL_3:4; then R_EAL (- 1) < seq . (max (n1,n2)) by SUPINF_2:2; hence contradiction by A11, XXREAL_0:25; ::_thesis: verum end; supposeA12: g < 0 ; ::_thesis: contradiction then consider n1 being Nat such that A13: for m being Nat st n1 <= m holds |.((seq . m) - (R_EAL g)).| < R_EAL g1 by A4; A14: now__::_thesis:_for_m_being_Element_of_NAT_st_n1_<=_m_holds_ R_EAL_(2_*_g)_<_seq_._m let m be Element of NAT ; ::_thesis: ( n1 <= m implies R_EAL (2 * g) < seq . m ) assume n1 <= m ; ::_thesis: R_EAL (2 * g) < seq . m then |.((seq . m) - (R_EAL g)).| < R_EAL g1 by A13; then - (R_EAL g1) < (seq . m) - (R_EAL g) by EXTREAL2:10; then (- (R_EAL g1)) + (R_EAL g) < seq . m by XXREAL_3:53; then R_EAL ((- g1) + g) < seq . m by SUPINF_2:1; hence R_EAL (2 * g) < seq . m ; ::_thesis: verum end; consider n2 being Nat such that A15: for m being Nat st n2 <= m holds seq . m <= R_EAL (2 * g) by A1, A12, Def10; reconsider n1 = n1, n2 = n2 as Element of NAT by ORDINAL1:def_12; set m = max (n1,n2); seq . (max (n1,n2)) <= R_EAL (2 * g) by A15, XXREAL_0:25; hence contradiction by A14, XXREAL_0:25; ::_thesis: verum end; end; end; definition let seq be ExtREAL_sequence; attrseq is convergent means :Def11: :: MESFUNC5:def 11 ( seq is convergent_to_finite_number or seq is convergent_to_+infty or seq is convergent_to_-infty ); end; :: deftheorem Def11 defines convergent MESFUNC5:def_11_:_ for seq being ExtREAL_sequence holds ( seq is convergent iff ( seq is convergent_to_finite_number or seq is convergent_to_+infty or seq is convergent_to_-infty ) ); definition let seq be ExtREAL_sequence; assume A1: seq is convergent ; func lim seq -> R_eal means :Def12: :: MESFUNC5:def 12 ( ex g being real number st ( it = g & ( for p being real number st 0 < p holds ex n being Nat st for m being Nat st n <= m holds |.((seq . m) - it).| < p ) & seq is convergent_to_finite_number ) or ( it = +infty & seq is convergent_to_+infty ) or ( it = -infty & seq is convergent_to_-infty ) ); existence ex b1 being R_eal st ( ex g being real number st ( b1 = g & ( for p being real number st 0 < p holds ex n being Nat st for m being Nat st n <= m holds |.((seq . m) - b1).| < p ) & seq is convergent_to_finite_number ) or ( b1 = +infty & seq is convergent_to_+infty ) or ( b1 = -infty & seq is convergent_to_-infty ) ) proof percases ( seq is convergent_to_finite_number or seq is convergent_to_+infty or seq is convergent_to_-infty ) by A1, Def11; supposeA2: seq is convergent_to_finite_number ; ::_thesis: ex b1 being R_eal st ( ex g being real number st ( b1 = g & ( for p being real number st 0 < p holds ex n being Nat st for m being Nat st n <= m holds |.((seq . m) - b1).| < p ) & seq is convergent_to_finite_number ) or ( b1 = +infty & seq is convergent_to_+infty ) or ( b1 = -infty & seq is convergent_to_-infty ) ) then ex g being real number st for p being real number st 0 < p holds ex n being Nat st for m being Nat st n <= m holds |.((seq . m) - (R_EAL g)).| < p by Def8; hence ex b1 being R_eal st ( ex g being real number st ( b1 = g & ( for p being real number st 0 < p holds ex n being Nat st for m being Nat st n <= m holds |.((seq . m) - b1).| < p ) & seq is convergent_to_finite_number ) or ( b1 = +infty & seq is convergent_to_+infty ) or ( b1 = -infty & seq is convergent_to_-infty ) ) by A2; ::_thesis: verum end; suppose ( seq is convergent_to_+infty or seq is convergent_to_-infty ) ; ::_thesis: ex b1 being R_eal st ( ex g being real number st ( b1 = g & ( for p being real number st 0 < p holds ex n being Nat st for m being Nat st n <= m holds |.((seq . m) - b1).| < p ) & seq is convergent_to_finite_number ) or ( b1 = +infty & seq is convergent_to_+infty ) or ( b1 = -infty & seq is convergent_to_-infty ) ) hence ex b1 being R_eal st ( ex g being real number st ( b1 = g & ( for p being real number st 0 < p holds ex n being Nat st for m being Nat st n <= m holds |.((seq . m) - b1).| < p ) & seq is convergent_to_finite_number ) or ( b1 = +infty & seq is convergent_to_+infty ) or ( b1 = -infty & seq is convergent_to_-infty ) ) ; ::_thesis: verum end; end; end; uniqueness for b1, b2 being R_eal st ( ex g being real number st ( b1 = g & ( for p being real number st 0 < p holds ex n being Nat st for m being Nat st n <= m holds |.((seq . m) - b1).| < p ) & seq is convergent_to_finite_number ) or ( b1 = +infty & seq is convergent_to_+infty ) or ( b1 = -infty & seq is convergent_to_-infty ) ) & ( ex g being real number st ( b2 = g & ( for p being real number st 0 < p holds ex n being Nat st for m being Nat st n <= m holds |.((seq . m) - b2).| < p ) & seq is convergent_to_finite_number ) or ( b2 = +infty & seq is convergent_to_+infty ) or ( b2 = -infty & seq is convergent_to_-infty ) ) holds b1 = b2 proof defpred S1[ R_eal] means ( ex g being real number st ( $1 = g & ( for p being real number st 0 < p holds ex n being Nat st for m being Nat st n <= m holds |.((seq . m) - $1).| < p ) & seq is convergent_to_finite_number ) or ( $1 = +infty & seq is convergent_to_+infty ) or ( $1 = -infty & seq is convergent_to_-infty ) ); given g1, g2 being R_eal such that A3: S1[g1] and A4: S1[g2] and A5: g1 <> g2 ; ::_thesis: contradiction percases ( seq is convergent_to_finite_number or seq is convergent_to_+infty or seq is convergent_to_-infty ) by A1, Def11; supposeA6: seq is convergent_to_finite_number ; ::_thesis: contradiction then consider g being real number such that A7: g1 = g and A8: for p being real number st 0 < p holds ex n being Nat st for m being Nat st n <= m holds |.((seq . m) - g1).| < p and seq is convergent_to_finite_number by A3, Th50, Th51; consider h being real number such that A9: g2 = h and A10: for p being real number st 0 < p holds ex n being Nat st for m being Nat st n <= m holds |.((seq . m) - g2).| < p and seq is convergent_to_finite_number by A4, A6, Th50, Th51; reconsider g = g, h = h as complex number ; g - h <> 0 by A5, A7, A9; then A11: |.(g - h).| > 0 ; then consider n1 being Nat such that A12: for m being Nat st n1 <= m holds |.((seq . m) - g1).| < R_EAL (|.(g - h).| / 2) by A8; consider n2 being Nat such that A13: for m being Nat st n2 <= m holds |.((seq . m) - g2).| < R_EAL (|.(g - h).| / 2) by A10, A11; reconsider n1 = n1, n2 = n2 as Element of NAT by ORDINAL1:def_12; set m = max (n1,n2); A14: |.((seq . (max (n1,n2))) - g1).| < R_EAL (|.(g - h).| / 2) by A12, XXREAL_0:25; A15: |.((seq . (max (n1,n2))) - g2).| < R_EAL (|.(g - h).| / 2) by A13, XXREAL_0:25; reconsider g = g, h = h as Real by XREAL_0:def_1; A16: (seq . (max (n1,n2))) - g2 < R_EAL (|.(g - h).| / 2) by A15, EXTREAL2:10; A17: - (R_EAL (|.(g - h).| / 2)) < (seq . (max (n1,n2))) - g2 by A15, EXTREAL2:10; then reconsider w = (seq . (max (n1,n2))) - g2 as Real by A16, XXREAL_0:48; A18: (seq . (max (n1,n2))) - g2 in REAL by A17, A16, XXREAL_0:48; then A19: seq . (max (n1,n2)) <> +infty by A9; A20: (- (seq . (max (n1,n2)))) + g1 = - ((seq . (max (n1,n2))) - g1) by XXREAL_3:26; then A21: |.((- (seq . (max (n1,n2)))) + g1).| < R_EAL (|.(g - h).| / 2) by A14, EXTREAL2:18; then A22: (- (seq . (max (n1,n2)))) + g1 < R_EAL (|.(g - h).| / 2) by EXTREAL2:10; - (R_EAL (|.(g - h).| / 2)) < (- (seq . (max (n1,n2)))) + g1 by A21, EXTREAL2:10; then A23: (- (seq . (max (n1,n2)))) + g1 in REAL by A22, XXREAL_0:48; A24: seq . (max (n1,n2)) <> -infty by A9, A18; |.(g1 - g2).| = |.((g1 + 0.) - g2).| by XXREAL_3:4 .= |.((g1 + ((seq . (max (n1,n2))) + (- (seq . (max (n1,n2)))))) - g2).| by XXREAL_3:7 .= |.((((- (seq . (max (n1,n2)))) + g1) + (seq . (max (n1,n2)))) - g2).| by A7, A19, A24, XXREAL_3:29 .= |.(((- (seq . (max (n1,n2)))) + g1) + ((seq . (max (n1,n2))) - g2)).| by A9, A23, XXREAL_3:30 ; then |.(g1 - g2).| <= |.((- (seq . (max (n1,n2)))) + g1).| + |.((seq . (max (n1,n2))) - g2).| by EXTREAL2:13; then A25: |.(g1 - g2).| <= |.((seq . (max (n1,n2))) - g1).| + |.((seq . (max (n1,n2))) - g2).| by A20, EXTREAL2:18; A26: (R_EAL (|.(g - h).| / 2)) + (R_EAL (|.(g - h).| / 2)) = (|.(g - h).| / 2) + (|.(g - h).| / 2) by SUPINF_2:1; |.w.| in REAL ; then |.((seq . (max (n1,n2))) - g2).| in REAL by EXTREAL2:1; then A27: |.((seq . (max (n1,n2))) - g1).| + |.((seq . (max (n1,n2))) - g2).| < (R_EAL (|.(g - h).| / 2)) + |.((seq . (max (n1,n2))) - g2).| by A14, XXREAL_3:43; (R_EAL (|.(g - h).| / 2)) + |.((seq . (max (n1,n2))) - g2).| < (R_EAL (|.(g - h).| / 2)) + (R_EAL (|.(g - h).| / 2)) by A15, XXREAL_3:43; then A28: |.((seq . (max (n1,n2))) - g1).| + |.((seq . (max (n1,n2))) - g2).| < (R_EAL (|.(g - h).| / 2)) + (R_EAL (|.(g - h).| / 2)) by A27, XXREAL_0:2; g - h = g1 - g2 by A7, A9, SUPINF_2:3; hence contradiction by A28, A25, A26, EXTREAL2:1; ::_thesis: verum end; suppose ( seq is convergent_to_+infty or seq is convergent_to_-infty ) ; ::_thesis: contradiction hence contradiction by A3, A4, A5, Th50, Th51; ::_thesis: verum end; end; end; end; :: deftheorem Def12 defines lim MESFUNC5:def_12_:_ for seq being ExtREAL_sequence st seq is convergent holds for b2 being R_eal holds ( b2 = lim seq iff ( ex g being real number st ( b2 = g & ( for p being real number st 0 < p holds ex n being Nat st for m being Nat st n <= m holds |.((seq . m) - b2).| < p ) & seq is convergent_to_finite_number ) or ( b2 = +infty & seq is convergent_to_+infty ) or ( b2 = -infty & seq is convergent_to_-infty ) ) ); theorem Th52: :: MESFUNC5:52 for seq being ExtREAL_sequence for r being real number st ( for n being Nat holds seq . n = r ) holds ( seq is convergent_to_finite_number & lim seq = r ) proof let seq be ExtREAL_sequence; ::_thesis: for r being real number st ( for n being Nat holds seq . n = r ) holds ( seq is convergent_to_finite_number & lim seq = r ) let r be real number ; ::_thesis: ( ( for n being Nat holds seq . n = r ) implies ( seq is convergent_to_finite_number & lim seq = r ) ) assume A1: for n being Nat holds seq . n = r ; ::_thesis: ( seq is convergent_to_finite_number & lim seq = r ) A2: now__::_thesis:_for_p_being_real_number_st_0_<_p_holds_ ex_n_being_Nat_st_ for_m_being_Nat_st_n_<=_m_holds_ |.((seq_._m)_-_(R_EAL_r)).|_<_p reconsider n = 1 as Nat ; let p be real number ; ::_thesis: ( 0 < p implies ex n being Nat st for m being Nat st n <= m holds |.((seq . m) - (R_EAL r)).| < p ) assume A3: 0 < p ; ::_thesis: ex n being Nat st for m being Nat st n <= m holds |.((seq . m) - (R_EAL r)).| < p take n = n; ::_thesis: for m being Nat st n <= m holds |.((seq . m) - (R_EAL r)).| < p let m be Nat; ::_thesis: ( n <= m implies |.((seq . m) - (R_EAL r)).| < p ) assume n <= m ; ::_thesis: |.((seq . m) - (R_EAL r)).| < p seq . m = R_EAL r by A1; then (seq . m) - (R_EAL r) = 0 by XXREAL_3:7; hence |.((seq . m) - (R_EAL r)).| < p by A3, EXTREAL2:5; ::_thesis: verum end; hence A4: seq is convergent_to_finite_number by Def8; ::_thesis: lim seq = r then seq is convergent by Def11; hence lim seq = r by A2, A4, Def12; ::_thesis: verum end; theorem Th53: :: MESFUNC5:53 for F being FinSequence of ExtREAL st ( for n being Nat st n in dom F holds 0 <= F . n ) holds 0 <= Sum F proof let F be FinSequence of ExtREAL ; ::_thesis: ( ( for n being Nat st n in dom F holds 0 <= F . n ) implies 0 <= Sum F ) consider sumf being Function of NAT,ExtREAL such that A1: Sum F = sumf . (len F) and A2: sumf . 0 = 0 and A3: for n being Element of NAT st n < len F holds sumf . (n + 1) = (sumf . n) + (F . (n + 1)) by EXTREAL1:def_2; defpred S1[ Nat] means ( $1 <= len F implies 0 <= sumf . $1 ); assume A4: for n being Nat st n in dom F holds 0 <= F . n ; ::_thesis: 0 <= Sum F A5: for n being Nat st S1[n] holds S1[n + 1] proof let n be Nat; ::_thesis: ( S1[n] implies S1[n + 1] ) assume A6: S1[n] ; ::_thesis: S1[n + 1] assume A7: n + 1 <= len F ; ::_thesis: 0 <= sumf . (n + 1) reconsider n = n as Element of NAT by ORDINAL1:def_12; 1 <= n + 1 by NAT_1:11; then n + 1 in Seg (len F) by A7; then n + 1 in dom F by FINSEQ_1:def_3; then A8: 0 <= F . (n + 1) by A4; n < len F by A7, NAT_1:13; then sumf . (n + 1) = (sumf . n) + (F . (n + 1)) by A3; hence 0 <= sumf . (n + 1) by A6, A7, A8, NAT_1:13; ::_thesis: verum end; A9: S1[ 0 ] by A2; for n being Nat holds S1[n] from NAT_1:sch_2(A9, A5); hence 0 <= Sum F by A1; ::_thesis: verum end; theorem Th54: :: MESFUNC5:54 for L being ExtREAL_sequence st ( for n, m being Nat st n <= m holds L . n <= L . m ) holds ( L is convergent & lim L = sup (rng L) ) proof let L be ExtREAL_sequence; ::_thesis: ( ( for n, m being Nat st n <= m holds L . n <= L . m ) implies ( L is convergent & lim L = sup (rng L) ) ) assume A1: for n, m being Nat st n <= m holds L . n <= L . m ; ::_thesis: ( L is convergent & lim L = sup (rng L) ) A2: now__::_thesis:_for_n_being_Nat_holds_L_._n_<=_sup_(rng_L) let n be Nat; ::_thesis: L . n <= sup (rng L) reconsider n1 = n as Element of NAT by ORDINAL1:def_12; dom L = NAT by FUNCT_2:def_1; then A3: L . n1 in rng L by FUNCT_1:def_3; sup (rng L) is UpperBound of rng L by XXREAL_2:def_3; hence L . n <= sup (rng L) by A3, XXREAL_2:def_1; ::_thesis: verum end; percases ( for k0 being Nat holds not -infty < L . k0 or ex k0 being Nat st -infty < L . k0 ) ; supposeA4: for k0 being Nat holds not -infty < L . k0 ; ::_thesis: ( L is convergent & lim L = sup (rng L) ) now__::_thesis:_for_x_being_ext-real_number_st_x_in_rng_L_holds_ x_<=_-infty let x be ext-real number ; ::_thesis: ( x in rng L implies x <= -infty ) assume x in rng L ; ::_thesis: x <= -infty then ex n being set st ( n in dom L & x = L . n ) by FUNCT_1:def_3; hence x <= -infty by A4; ::_thesis: verum end; then A5: -infty is UpperBound of rng L by XXREAL_2:def_1; for y being UpperBound of rng L holds -infty <= y by XXREAL_0:5; then A6: -infty = sup (rng L) by A5, XXREAL_2:def_3; now__::_thesis:_for_K_being_real_number_st_K_<_0_holds_ ex_N0_being_Nat_st_ for_n_being_Nat_st_N0_<=_n_holds_ L_._n_<=_K reconsider N0 = 0 as Nat ; let K be real number ; ::_thesis: ( K < 0 implies ex N0 being Nat st for n being Nat st N0 <= n holds L . n <= K ) assume K < 0 ; ::_thesis: ex N0 being Nat st for n being Nat st N0 <= n holds L . n <= K take N0 = N0; ::_thesis: for n being Nat st N0 <= n holds L . n <= K let n be Nat; ::_thesis: ( N0 <= n implies L . n <= K ) assume N0 <= n ; ::_thesis: L . n <= K L . n = -infty by A4, XXREAL_0:6; hence L . n <= K by XXREAL_0:5; ::_thesis: verum end; then A7: L is convergent_to_-infty by Def10; then L is convergent by Def11; hence ( L is convergent & lim L = sup (rng L) ) by A7, A6, Def12; ::_thesis: verum end; suppose ex k0 being Nat st -infty < L . k0 ; ::_thesis: ( L is convergent & lim L = sup (rng L) ) then consider k0 being Nat such that A8: -infty < L . k0 ; reconsider k0 = k0 as Element of NAT by ORDINAL1:def_12; percases ( ex K being real number st for n being Nat holds L . n < K or for K being real number holds ( not 0 < K or ex n being Nat st not L . n < K ) ) ; suppose ex K being real number st for n being Nat holds L . n < K ; ::_thesis: ( L is convergent & lim L = sup (rng L) ) then consider K being real number such that A9: for n being Nat holds L . n < K ; now__::_thesis:_for_x_being_ext-real_number_st_x_in_rng_L_holds_ x_<=_R_EAL_K let x be ext-real number ; ::_thesis: ( x in rng L implies x <= R_EAL K ) assume x in rng L ; ::_thesis: x <= R_EAL K then ex z being set st ( z in dom L & x = L . z ) by FUNCT_1:def_3; hence x <= R_EAL K by A9; ::_thesis: verum end; then R_EAL K is UpperBound of rng L by XXREAL_2:def_1; then A10: sup (rng L) <= R_EAL K by XXREAL_2:def_3; R_EAL K is Real by XREAL_0:def_1; then A11: sup (rng L) <> +infty by A10, XXREAL_0:9; A12: sup (rng L) <> -infty by A2, A8; then reconsider h = sup (rng L) as Real by A11, XXREAL_0:14; A13: for p being real number st 0 < p holds ex N0 being Nat st for m being Nat st N0 <= m holds |.((L . m) - (sup (rng L))).| < p proof let p be real number ; ::_thesis: ( 0 < p implies ex N0 being Nat st for m being Nat st N0 <= m holds |.((L . m) - (sup (rng L))).| < p ) assume A14: 0 < p ; ::_thesis: ex N0 being Nat st for m being Nat st N0 <= m holds |.((L . m) - (sup (rng L))).| < p reconsider e = p as Real by XREAL_0:def_1; sup (rng L) is Real by A12, A11, XXREAL_0:14; then consider y being ext-real number such that A15: y in rng L and A16: (sup (rng L)) - (R_EAL e) < y by A14, MEASURE6:6; consider x being set such that A17: x in dom L and A18: y = L . x by A15, FUNCT_1:def_3; reconsider N1 = x as Element of NAT by A17; set N0 = max (N1,k0); take max (N1,k0) ; ::_thesis: for m being Nat st max (N1,k0) <= m holds |.((L . m) - (sup (rng L))).| < p hereby ::_thesis: verum let n be Nat; ::_thesis: ( max (N1,k0) <= n implies |.((L . n) - (sup (rng L))).| < p ) A19: N1 <= max (N1,k0) by XXREAL_0:25; assume max (N1,k0) <= n ; ::_thesis: |.((L . n) - (sup (rng L))).| < p then N1 <= n by A19, XXREAL_0:2; then L . N1 <= L . n by A1; then (sup (rng L)) - (R_EAL e) < L . n by A16, A18, XXREAL_0:2; then sup (rng L) < (L . n) + (R_EAL e) by XXREAL_3:54; then (sup (rng L)) - (L . n) < R_EAL e by XXREAL_3:55; then - (R_EAL e) < - ((sup (rng L)) - (L . n)) by XXREAL_3:38; then A20: - (R_EAL e) < (L . n) - (sup (rng L)) by XXREAL_3:26; A21: L . n <= sup (rng L) by A2; A22: now__::_thesis:_not_sup_(rng_L)_=_(sup_(rng_L))_+_e assume A23: sup (rng L) = (sup (rng L)) + e ; ::_thesis: contradiction ((R_EAL e) + (sup (rng L))) + (- (sup (rng L))) = (R_EAL e) + ((sup (rng L)) + (- (sup (rng L)))) by A12, A11, XXREAL_3:29 .= (R_EAL e) + 0 by XXREAL_3:7 .= e by XXREAL_3:4 ; hence contradiction by A14, A23, XXREAL_3:7; ::_thesis: verum end; (sup (rng L)) + (R_EAL 0) <= (sup (rng L)) + (R_EAL e) by A14, XXREAL_3:36; then sup (rng L) <= (sup (rng L)) + (R_EAL e) by XXREAL_3:4; then sup (rng L) < (sup (rng L)) + e by A22, XXREAL_0:1; then L . n < (sup (rng L)) + (R_EAL e) by A21, XXREAL_0:2; then (L . n) - (sup (rng L)) < R_EAL e by XXREAL_3:55; hence |.((L . n) - (sup (rng L))).| < p by A20, EXTREAL2:11; ::_thesis: verum end; end; A24: R_EAL h = sup (rng L) ; then A25: L is convergent_to_finite_number by A13, Def8; hence L is convergent by Def11; ::_thesis: lim L = sup (rng L) hence lim L = sup (rng L) by A13, A24, A25, Def12; ::_thesis: verum end; supposeA26: for K being real number holds ( not 0 < K or ex n being Nat st not L . n < K ) ; ::_thesis: ( L is convergent & lim L = sup (rng L) ) now__::_thesis:_for_K_being_real_number_st_0_<_K_holds_ ex_N0_being_Nat_st_ for_n_being_Nat_st_N0_<=_n_holds_ K_<=_L_._n let K be real number ; ::_thesis: ( 0 < K implies ex N0 being Nat st for n being Nat st N0 <= n holds K <= L . n ) assume 0 < K ; ::_thesis: ex N0 being Nat st for n being Nat st N0 <= n holds K <= L . n then consider N0 being Nat such that A27: R_EAL K <= L . N0 by A26; now__::_thesis:_for_n_being_Nat_st_N0_<=_n_holds_ R_EAL_K_<=_L_._n let n be Nat; ::_thesis: ( N0 <= n implies R_EAL K <= L . n ) assume N0 <= n ; ::_thesis: R_EAL K <= L . n then L . N0 <= L . n by A1; hence R_EAL K <= L . n by A27, XXREAL_0:2; ::_thesis: verum end; hence ex N0 being Nat st for n being Nat st N0 <= n holds K <= L . n ; ::_thesis: verum end; then A28: L is convergent_to_+infty by Def9; hence A29: L is convergent by Def11; ::_thesis: lim L = sup (rng L) now__::_thesis:_not_sup_(rng_L)_<>_+infty assume A30: sup (rng L) <> +infty ; ::_thesis: contradiction L . k0 <= sup (rng L) by A2; then reconsider h = sup (rng L) as Real by A8, A30, XXREAL_0:14; set K = max (0,h); 0 <= max (0,h) by XXREAL_0:25; then consider N0 being Nat such that A31: R_EAL ((max (0,h)) + 1) <= L . N0 by A26; h + 0 < (max (0,h)) + 1 by XREAL_1:8, XXREAL_0:25; then sup (rng L) < L . N0 by A31, XXREAL_0:2; hence contradiction by A2; ::_thesis: verum end; hence lim L = sup (rng L) by A28, A29, Def12; ::_thesis: verum end; end; end; end; end; theorem Th55: :: MESFUNC5:55 for L, G being ExtREAL_sequence st ( for n being Nat holds L . n <= G . n ) holds sup (rng L) <= sup (rng G) proof let L, G be ExtREAL_sequence; ::_thesis: ( ( for n being Nat holds L . n <= G . n ) implies sup (rng L) <= sup (rng G) ) assume A1: for n being Nat holds L . n <= G . n ; ::_thesis: sup (rng L) <= sup (rng G) A2: now__::_thesis:_for_n_being_Element_of_NAT_holds_L_._n_<=_sup_(rng_G) let n be Element of NAT ; ::_thesis: L . n <= sup (rng G) dom G = NAT by FUNCT_2:def_1; then A3: G . n in rng G by FUNCT_1:def_3; A4: L . n <= G . n by A1; sup (rng G) is UpperBound of rng G by XXREAL_2:def_3; then G . n <= sup (rng G) by A3, XXREAL_2:def_1; hence L . n <= sup (rng G) by A4, XXREAL_0:2; ::_thesis: verum end; now__::_thesis:_for_x_being_ext-real_number_st_x_in_rng_L_holds_ x_<=_sup_(rng_G) let x be ext-real number ; ::_thesis: ( x in rng L implies x <= sup (rng G) ) assume x in rng L ; ::_thesis: x <= sup (rng G) then ex z being set st ( z in dom L & x = L . z ) by FUNCT_1:def_3; hence x <= sup (rng G) by A2; ::_thesis: verum end; then sup (rng G) is UpperBound of rng L by XXREAL_2:def_1; hence sup (rng L) <= sup (rng G) by XXREAL_2:def_3; ::_thesis: verum end; theorem Th56: :: MESFUNC5:56 for L being ExtREAL_sequence for n being Nat holds L . n <= sup (rng L) proof let L be ExtREAL_sequence; ::_thesis: for n being Nat holds L . n <= sup (rng L) let n be Nat; ::_thesis: L . n <= sup (rng L) reconsider n = n as Element of NAT by ORDINAL1:def_12; dom L = NAT by FUNCT_2:def_1; then A1: L . n in rng L by FUNCT_1:def_3; sup (rng L) is UpperBound of rng L by XXREAL_2:def_3; hence L . n <= sup (rng L) by A1, XXREAL_2:def_1; ::_thesis: verum end; theorem Th57: :: MESFUNC5:57 for L being ExtREAL_sequence for K being R_eal st ( for n being Nat holds L . n <= K ) holds sup (rng L) <= K proof let L be ExtREAL_sequence; ::_thesis: for K being R_eal st ( for n being Nat holds L . n <= K ) holds sup (rng L) <= K let K be R_eal; ::_thesis: ( ( for n being Nat holds L . n <= K ) implies sup (rng L) <= K ) assume A1: for n being Nat holds L . n <= K ; ::_thesis: sup (rng L) <= K now__::_thesis:_for_x_being_ext-real_number_st_x_in_rng_L_holds_ x_<=_K let x be ext-real number ; ::_thesis: ( x in rng L implies x <= K ) assume x in rng L ; ::_thesis: x <= K then ex z being set st ( z in dom L & x = L . z ) by FUNCT_1:def_3; hence x <= K by A1; ::_thesis: verum end; then K is UpperBound of rng L by XXREAL_2:def_1; hence sup (rng L) <= K by XXREAL_2:def_3; ::_thesis: verum end; theorem :: MESFUNC5:58 for L being ExtREAL_sequence for K being R_eal st K <> +infty & ( for n being Nat holds L . n <= K ) holds sup (rng L) < +infty proof let L be ExtREAL_sequence; ::_thesis: for K being R_eal st K <> +infty & ( for n being Nat holds L . n <= K ) holds sup (rng L) < +infty let K be R_eal; ::_thesis: ( K <> +infty & ( for n being Nat holds L . n <= K ) implies sup (rng L) < +infty ) assume that A1: K <> +infty and A2: for n being Nat holds L . n <= K ; ::_thesis: sup (rng L) < +infty now__::_thesis:_for_x_being_ext-real_number_st_x_in_rng_L_holds_ x_<=_K let x be ext-real number ; ::_thesis: ( x in rng L implies x <= K ) assume x in rng L ; ::_thesis: x <= K then ex z being set st ( z in dom L & x = L . z ) by FUNCT_1:def_3; hence x <= K by A2; ::_thesis: verum end; then K is UpperBound of rng L by XXREAL_2:def_1; then sup (rng L) <= K by XXREAL_2:def_3; hence sup (rng L) < +infty by A1, XXREAL_0:2, XXREAL_0:4; ::_thesis: verum end; theorem Th59: :: MESFUNC5:59 for L being ExtREAL_sequence st L is () holds ( sup (rng L) <> +infty iff ex K being real number st ( 0 < K & ( for n being Nat holds L . n <= K ) ) ) proof let L be ExtREAL_sequence; ::_thesis: ( L is () implies ( sup (rng L) <> +infty iff ex K being real number st ( 0 < K & ( for n being Nat holds L . n <= K ) ) ) ) assume L is () ; ::_thesis: ( sup (rng L) <> +infty iff ex K being real number st ( 0 < K & ( for n being Nat holds L . n <= K ) ) ) then A1: -infty < L . 1 by Def5; A2: dom L = NAT by FUNCT_2:def_1; then A3: L . 1 <= sup (rng L) by FUNCT_1:3, XXREAL_2:4; A4: now__::_thesis:_(_sup_(rng_L)_<>_+infty_implies_ex_K_being_set_st_ (_0_<_K_&_(_for_n_being_Nat_holds_L_._n_<=_R_EAL_K_)_)_) assume sup (rng L) <> +infty ; ::_thesis: ex K being set st ( 0 < K & ( for n being Nat holds L . n <= R_EAL K ) ) then not sup (rng L) in {-infty,+infty} by A1, A3, TARSKI:def_2; then sup (rng L) in REAL by XBOOLE_0:def_3, XXREAL_0:def_4; then reconsider S = sup (rng L) as real number ; take K = max (S,1); ::_thesis: ( 0 < K & ( for n being Nat holds L . n <= R_EAL K ) ) thus 0 < K by XXREAL_0:25; ::_thesis: for n being Nat holds L . n <= R_EAL K let n be Nat; ::_thesis: L . n <= R_EAL K n in NAT by ORDINAL1:def_12; then A5: L . n <= sup (rng L) by A2, FUNCT_1:3, XXREAL_2:4; S <= K by XXREAL_0:25; hence L . n <= R_EAL K by A5, XXREAL_0:2; ::_thesis: verum end; now__::_thesis:_(_ex_K_being_real_number_st_ (_0_<_K_&_(_for_n_being_Nat_holds_L_._n_<=_K_)_)_implies_sup_(rng_L)_<>_+infty_) given K being real number such that 0 < K and A6: for n being Nat holds L . n <= K ; ::_thesis: sup (rng L) <> +infty now__::_thesis:_for_w_being_ext-real_number_st_w_in_rng_L_holds_ w_<=_R_EAL_K let w be ext-real number ; ::_thesis: ( w in rng L implies w <= R_EAL K ) assume w in rng L ; ::_thesis: w <= R_EAL K then ex v being set st ( v in dom L & w = L . v ) by FUNCT_1:def_3; hence w <= R_EAL K by A6; ::_thesis: verum end; then R_EAL K is UpperBound of rng L by XXREAL_2:def_1; then A7: sup (rng L) <= R_EAL K by XXREAL_2:def_3; K in REAL by XREAL_0:def_1; hence sup (rng L) <> +infty by A7, XXREAL_0:9; ::_thesis: verum end; hence ( sup (rng L) <> +infty iff ex K being real number st ( 0 < K & ( for n being Nat holds L . n <= K ) ) ) by A4; ::_thesis: verum end; theorem Th60: :: MESFUNC5:60 for L being ExtREAL_sequence for c being ext-real number st ( for n being Nat holds L . n = c ) holds ( L is convergent & lim L = c & lim L = sup (rng L) ) proof let L be ExtREAL_sequence; ::_thesis: for c being ext-real number st ( for n being Nat holds L . n = c ) holds ( L is convergent & lim L = c & lim L = sup (rng L) ) let c be ext-real number ; ::_thesis: ( ( for n being Nat holds L . n = c ) implies ( L is convergent & lim L = c & lim L = sup (rng L) ) ) A1: dom L = NAT by FUNCT_2:def_1; c in ExtREAL by XXREAL_0:def_1; then A2: ( c in REAL or c in {-infty,+infty} ) by XBOOLE_0:def_3, XXREAL_0:def_4; assume A3: for n being Nat holds L . n = c ; ::_thesis: ( L is convergent & lim L = c & lim L = sup (rng L) ) then A4: L . 1 = c ; now__::_thesis:_for_v_being_ext-real_number_st_v_in_rng_L_holds_ v_<=_c let v be ext-real number ; ::_thesis: ( v in rng L implies v <= c ) assume v in rng L ; ::_thesis: v <= c then ex n being set st ( n in dom L & v = L . n ) by FUNCT_1:def_3; hence v <= c by A3; ::_thesis: verum end; then A5: c is UpperBound of rng L by XXREAL_2:def_1; percases ( c in REAL or c = -infty or c = +infty ) by A2, TARSKI:def_2; suppose c in REAL ; ::_thesis: ( L is convergent & lim L = c & lim L = sup (rng L) ) then reconsider rc = c as real number ; A6: now__::_thesis:_for_p_being_real_number_st_0_<_p_holds_ ex_n_being_Nat_st_ for_m_being_Nat_st_n_<=_m_holds_ |.((L_._m)_-_(R_EAL_rc)).|_<_p reconsider n = 0 as Nat ; let p be real number ; ::_thesis: ( 0 < p implies ex n being Nat st for m being Nat st n <= m holds |.((L . m) - (R_EAL rc)).| < p ) assume A7: 0 < p ; ::_thesis: ex n being Nat st for m being Nat st n <= m holds |.((L . m) - (R_EAL rc)).| < p take n = n; ::_thesis: for m being Nat st n <= m holds |.((L . m) - (R_EAL rc)).| < p let m be Nat; ::_thesis: ( n <= m implies |.((L . m) - (R_EAL rc)).| < p ) assume n <= m ; ::_thesis: |.((L . m) - (R_EAL rc)).| < p (L . m) - (R_EAL rc) = (L . m) - (L . m) by A3; then (L . m) - rc = 0 by XXREAL_3:7; hence |.((L . m) - (R_EAL rc)).| < p by A7, EXTREAL2:5; ::_thesis: verum end; then A8: L is convergent_to_finite_number by Def8; hence L is convergent by Def11; ::_thesis: ( lim L = c & lim L = sup (rng L) ) hence lim L = c by A6, A8, Def12; ::_thesis: lim L = sup (rng L) hence lim L = sup (rng L) by A5, A1, A4, FUNCT_1:3, XXREAL_2:55; ::_thesis: verum end; supposeA9: c = -infty ; ::_thesis: ( L is convergent & lim L = c & lim L = sup (rng L) ) for p being real number st p < 0 holds ex n being Nat st for m being Nat st n <= m holds L . m <= p proof let p be real number ; ::_thesis: ( p < 0 implies ex n being Nat st for m being Nat st n <= m holds L . m <= p ) assume p < 0 ; ::_thesis: ex n being Nat st for m being Nat st n <= m holds L . m <= p take 0 ; ::_thesis: for m being Nat st 0 <= m holds L . m <= p A10: p in REAL by XREAL_0:def_1; now__::_thesis:_for_m_being_Nat_st_0_<=_m_holds_ L_._m_<_p let m be Nat; ::_thesis: ( 0 <= m implies L . m < p ) assume 0 <= m ; ::_thesis: L . m < p L . m = -infty by A3, A9; hence L . m < p by A10, XXREAL_0:12; ::_thesis: verum end; hence for m being Nat st 0 <= m holds L . m <= p ; ::_thesis: verum end; then A11: L is convergent_to_-infty by Def10; hence L is convergent by Def11; ::_thesis: ( lim L = c & lim L = sup (rng L) ) hence lim L = c by A9, A11, Def12; ::_thesis: lim L = sup (rng L) hence lim L = sup (rng L) by A5, A1, A4, FUNCT_1:3, XXREAL_2:55; ::_thesis: verum end; supposeA12: c = +infty ; ::_thesis: ( L is convergent & lim L = c & lim L = sup (rng L) ) for p being real number st 0 < p holds ex n being Nat st for m being Nat st n <= m holds p <= L . m proof let p be real number ; ::_thesis: ( 0 < p implies ex n being Nat st for m being Nat st n <= m holds p <= L . m ) assume 0 < p ; ::_thesis: ex n being Nat st for m being Nat st n <= m holds p <= L . m take 0 ; ::_thesis: for m being Nat st 0 <= m holds p <= L . m A13: p in REAL by XREAL_0:def_1; now__::_thesis:_for_m_being_Nat_st_0_<=_m_holds_ p_<_L_._m let m be Nat; ::_thesis: ( 0 <= m implies p < L . m ) assume 0 <= m ; ::_thesis: p < L . m L . m = +infty by A3, A12; hence p < L . m by A13, XXREAL_0:9; ::_thesis: verum end; hence for m being Nat st 0 <= m holds p <= L . m ; ::_thesis: verum end; then A14: L is convergent_to_+infty by Def9; hence L is convergent by Def11; ::_thesis: ( lim L = c & lim L = sup (rng L) ) hence lim L = c by A12, A14, Def12; ::_thesis: lim L = sup (rng L) hence lim L = sup (rng L) by A5, A1, A4, FUNCT_1:3, XXREAL_2:55; ::_thesis: verum end; end; end; Lm8: for J being ExtREAL_sequence holds ( not J is () or sup (rng J) in REAL or sup (rng J) = +infty ) proof let J be ExtREAL_sequence; ::_thesis: ( not J is () or sup (rng J) in REAL or sup (rng J) = +infty ) assume J is () ; ::_thesis: ( sup (rng J) in REAL or sup (rng J) = +infty ) then A1: -infty < J . 1 by Def5; dom J = NAT by FUNCT_2:def_1; then -infty is not UpperBound of rng J by A1, FUNCT_1:3, XXREAL_2:def_1; then sup (rng J) <> -infty by XXREAL_2:def_3; hence ( sup (rng J) in REAL or sup (rng J) = +infty ) by XXREAL_0:14; ::_thesis: verum end; theorem Th61: :: MESFUNC5:61 for J, K, L being ExtREAL_sequence st ( for n, m being Nat st n <= m holds J . n <= J . m ) & ( for n, m being Nat st n <= m holds K . n <= K . m ) & J is () & K is () & ( for n being Nat holds (J . n) + (K . n) = L . n ) holds ( L is convergent & lim L = sup (rng L) & lim L = (lim J) + (lim K) & sup (rng L) = (sup (rng K)) + (sup (rng J)) ) proof let J, K, L be ExtREAL_sequence; ::_thesis: ( ( for n, m being Nat st n <= m holds J . n <= J . m ) & ( for n, m being Nat st n <= m holds K . n <= K . m ) & J is () & K is () & ( for n being Nat holds (J . n) + (K . n) = L . n ) implies ( L is convergent & lim L = sup (rng L) & lim L = (lim J) + (lim K) & sup (rng L) = (sup (rng K)) + (sup (rng J)) ) ) assume that A1: for n, m being Nat st n <= m holds J . n <= J . m and A2: for n, m being Nat st n <= m holds K . n <= K . m and A3: J is () and A4: K is () and A5: for n being Nat holds (J . n) + (K . n) = L . n ; ::_thesis: ( L is convergent & lim L = sup (rng L) & lim L = (lim J) + (lim K) & sup (rng L) = (sup (rng K)) + (sup (rng J)) ) A6: dom K = NAT by FUNCT_2:def_1; A7: dom J = NAT by FUNCT_2:def_1; A8: now__::_thesis:_(_L_is_convergent_&_lim_L_=_(sup_(rng_J))_+_(sup_(rng_K))_) percases ( sup (rng J) in REAL or sup (rng J) = +infty ) by A3, Lm8; supposeA9: sup (rng J) in REAL ; ::_thesis: ( L is convergent & lim L = (sup (rng J)) + (sup (rng K)) ) then reconsider SJ = sup (rng J) as real number ; percases ( sup (rng K) in REAL or sup (rng K) = +infty ) by A4, Lm8; supposeA10: sup (rng K) in REAL ; ::_thesis: ( L is convergent & lim L = (sup (rng J)) + (sup (rng K)) ) then reconsider SK = sup (rng K) as real number ; A11: now__::_thesis:_for_p_being_real_number_st_0_<_p_holds_ ex_n_being_Nat_st_ for_m_being_Nat_st_n_<=_m_holds_ |.((L_._m)_-_(R_EAL_(SJ_+_SK))).|_<_p let p be real number ; ::_thesis: ( 0 < p implies ex n being Nat st for m being Nat st n <= m holds |.((L . m) - (R_EAL (SJ + SK))).| < p ) assume A12: 0 < p ; ::_thesis: ex n being Nat st for m being Nat st n <= m holds |.((L . m) - (R_EAL (SJ + SK))).| < p then consider SJ9 being ext-real number such that A13: SJ9 in rng J and A14: (sup (rng J)) - (R_EAL (p / 2)) < SJ9 by A9, MEASURE6:6; consider nj being set such that A15: nj in dom J and A16: SJ9 = J . nj by A13, FUNCT_1:def_3; reconsider nj = nj as Element of NAT by A15; consider SK9 being ext-real number such that A17: SK9 in rng K and A18: (sup (rng K)) - (R_EAL (p / 2)) < SK9 by A10, A12, MEASURE6:6; consider nk being set such that A19: nk in dom K and A20: SK9 = K . nk by A17, FUNCT_1:def_3; reconsider nk = nk as Element of NAT by A19; reconsider n = max (nj,nk) as Nat ; take n = n; ::_thesis: for m being Nat st n <= m holds |.((L . m) - (R_EAL (SJ + SK))).| < p (p / 2) + (p / 2) = p ; then A21: R_EAL p = (R_EAL (p / 2)) + (R_EAL (p / 2)) by SUPINF_2:1; hereby ::_thesis: verum reconsider SJ9 = SJ9, SK9 = SK9 as R_eal by XXREAL_0:def_1; let m be Nat; ::_thesis: ( n <= m implies |.((L . m) - (R_EAL (SJ + SK))).| < p ) assume A22: n <= m ; ::_thesis: |.((L . m) - (R_EAL (SJ + SK))).| < p nk <= n by XXREAL_0:25; then nk <= m by A22, XXREAL_0:2; then SK9 <= K . m by A2, A20; then A23: (R_EAL SK) - (K . m) <= (R_EAL SK) - SK9 by XXREAL_3:37; nj <= n by XXREAL_0:25; then nj <= m by A22, XXREAL_0:2; then SJ9 <= J . m by A1, A16; then (R_EAL SJ) - (J . m) <= (R_EAL SJ) - SJ9 by XXREAL_3:37; then A24: ((R_EAL SJ) - (J . m)) + ((R_EAL SK) - (K . m)) <= ((R_EAL SJ) - SJ9) + ((R_EAL SK) - SK9) by A23, XXREAL_3:36; R_EAL SJ in REAL by XREAL_0:def_1; then A25: R_EAL SJ < +infty by XXREAL_0:9; reconsider s1 = R_EAL SK as Real by XREAL_0:def_1; reconsider m1 = m as Element of NAT by ORDINAL1:def_12; A26: - ((L . m) - (R_EAL (SJ + SK))) = (R_EAL (SJ + SK)) - (L . m) by XXREAL_3:26; R_EAL SK < (R_EAL (p / 2)) + SK9 by A18, XXREAL_3:54; then (R_EAL SK) - SK9 < R_EAL (p / 2) by XXREAL_3:55; then A27: (R_EAL (p / 2)) + ((R_EAL SK) - SK9) < (R_EAL (p / 2)) + (R_EAL (p / 2)) by XXREAL_3:43; R_EAL SJ < (R_EAL (p / 2)) + SJ9 by A14, XXREAL_3:54; then A28: (R_EAL SJ) - SJ9 < R_EAL (p / 2) by XXREAL_3:55; nk <= n by XXREAL_0:25; then nk <= m by A22, XXREAL_0:2; then A29: K . nk <= K . m by A2; A30: R_EAL SK in REAL by XREAL_0:def_1; then A31: R_EAL SK < +infty by XXREAL_0:9; K . m1 in rng K by A6, FUNCT_1:3; then A32: K . m <= R_EAL SK by XXREAL_2:4; then A33: K . m < +infty by A30, XXREAL_0:2, XXREAL_0:9; -infty < SK9 by A4, A20, Def5; then reconsider s0 = SK9 as Real by A20, A33, A29, XXREAL_0:14; A34: L . m = (J . m) + (K . m) by A5; J . m1 in rng J by A7, FUNCT_1:3; then A35: J . m <= R_EAL SJ by XXREAL_2:4; then (J . m) + (K . m) <= (R_EAL SJ) + (R_EAL SK) by A32, XXREAL_3:36; then (L . m) - (R_EAL (SJ + SK)) <= 0 by A34, A26, XXREAL_3:40; then A36: |.((L . m) - (R_EAL (SJ + SK))).| = (R_EAL (SJ + SK)) - (L . m) by A26, EXTREAL2:7; (R_EAL SK) - SK9 = s1 - s0 by SUPINF_2:3; then ((R_EAL SJ) - SJ9) + ((R_EAL SK) - SK9) < (R_EAL (p / 2)) + ((R_EAL SK) - SK9) by A28, XXREAL_3:43; then A37: ((R_EAL SJ) - SJ9) + ((R_EAL SK) - SK9) < (R_EAL (p / 2)) + (R_EAL (p / 2)) by A27, XXREAL_0:2; -infty < K . m by A4, Def5; then ((R_EAL SJ) - (J . m)) + ((R_EAL SK) - (K . m)) = (((R_EAL SJ) - (J . m)) + (R_EAL SK)) - (K . m) by A33, XXREAL_3:30 .= (((R_EAL SK) + (R_EAL SJ)) - (J . m)) - (K . m) by XXREAL_3:30 .= (R_EAL (SK + SJ)) - ((J . m) + (K . m)) by A25, A31, A35, A32, XXREAL_3:31 .= (R_EAL (SK + SJ)) - (L . m) by A5 ; hence |.((L . m) - (R_EAL (SJ + SK))).| < p by A21, A36, A37, A24, XXREAL_0:2; ::_thesis: verum end; end; then A38: L is convergent_to_finite_number by Def8; hence L is convergent by Def11; ::_thesis: lim L = (sup (rng J)) + (sup (rng K)) hence lim L = (sup (rng J)) + (sup (rng K)) by A11, A38, Def12; ::_thesis: verum end; supposeA39: sup (rng K) = +infty ; ::_thesis: ( L is convergent & lim L = (sup (rng J)) + (sup (rng K)) ) for p being real number st 0 < p holds ex n being Nat st for m being Nat st n <= m holds p <= L . m proof reconsider supj = sup (rng J) as Real by A9; let p be real number ; ::_thesis: ( 0 < p implies ex n being Nat st for m being Nat st n <= m holds p <= L . m ) reconsider p92 = p / 2, p9 = p as Real by XREAL_0:def_1; assume 0 < p ; ::_thesis: ex n being Nat st for m being Nat st n <= m holds p <= L . m then consider j being ext-real number such that A40: j in rng J and A41: (sup (rng J)) - (R_EAL (p / 2)) < j by A9, MEASURE6:6; consider n1 being set such that A42: n1 in dom J and A43: j = J . n1 by A40, FUNCT_1:def_3; A44: supj - p92 = (sup (rng J)) - (R_EAL (p / 2)) by SUPINF_2:3; then A45: p9 - (supj - p92) = (R_EAL p) - ((sup (rng J)) - (R_EAL (p / 2))) by SUPINF_2:3; then (R_EAL p) - ((sup (rng J)) - (R_EAL (p / 2))) < sup (rng K) by A39, XXREAL_0:9; then consider k being R_eal such that A46: k in rng K and A47: (R_EAL p) - ((sup (rng J)) - (R_EAL (p / 2))) < k by XXREAL_2:94; p9 = (p9 - (supj - p92)) + (supj - p92) ; then A48: ((R_EAL p) - ((sup (rng J)) - (R_EAL (p / 2)))) + ((sup (rng J)) - (R_EAL (p / 2))) = p9 by A44, A45, SUPINF_2:1; reconsider n1 = n1 as Element of NAT by A42; consider n2 being set such that A49: n2 in dom K and A50: k = K . n2 by A46, FUNCT_1:def_3; reconsider n2 = n2 as Element of NAT by A49; set n = max (n1,n2); J . n1 <= J . (max (n1,n2)) by A1, XXREAL_0:25; then A51: (sup (rng J)) - (R_EAL (p / 2)) < J . (max (n1,n2)) by A41, A43, XXREAL_0:2; K . n2 <= K . (max (n1,n2)) by A2, XXREAL_0:25; then A52: (R_EAL p) - ((sup (rng J)) - (R_EAL (p / 2))) < K . (max (n1,n2)) by A47, A50, XXREAL_0:2; A53: R_EAL p < (J . (max (n1,n2))) + (K . (max (n1,n2))) by A51, A52, A48, XXREAL_3:64; now__::_thesis:_for_m_being_Nat_st_max_(n1,n2)_<=_m_holds_ p_<=_L_._m let m be Nat; ::_thesis: ( max (n1,n2) <= m implies p <= L . m ) assume A54: max (n1,n2) <= m ; ::_thesis: p <= L . m then A55: K . (max (n1,n2)) <= K . m by A2; J . (max (n1,n2)) <= J . m by A1, A54; then (J . (max (n1,n2))) + (K . (max (n1,n2))) <= (J . m) + (K . m) by A55, XXREAL_3:36; then (J . (max (n1,n2))) + (K . (max (n1,n2))) <= L . m by A5; hence p <= L . m by A53, XXREAL_0:2; ::_thesis: verum end; hence ex n being Nat st for m being Nat st n <= m holds p <= L . m ; ::_thesis: verum end; then A56: L is convergent_to_+infty by Def9; hence L is convergent by Def11; ::_thesis: lim L = (sup (rng J)) + (sup (rng K)) then lim L = +infty by A56, Def12; hence lim L = (sup (rng J)) + (sup (rng K)) by A9, A39, XXREAL_3:def_2; ::_thesis: verum end; end; end; supposeA57: sup (rng J) = +infty ; ::_thesis: ( L is convergent & lim L = (sup (rng J)) + (sup (rng K)) ) now__::_thesis:_for_p_being_real_number_st_0_<_p_holds_ ex_n_being_Nat_st_ for_m_being_Nat_st_n_<=_m_holds_ p_<=_L_._m let p be real number ; ::_thesis: ( 0 < p implies ex n being Nat st for m being Nat st m <= b3 holds n <= L . b3 ) assume A58: 0 < p ; ::_thesis: ex n being Nat st for m being Nat st m <= b3 holds n <= L . b3 percases ( sup (rng K) in REAL or sup (rng K) = +infty ) by A4, Lm8; supposeA59: sup (rng K) in REAL ; ::_thesis: ex n being Nat st for m being Nat st m <= b3 holds n <= L . b3 then reconsider supk = sup (rng K) as Real ; reconsider p92 = p / 2, p9 = p as Real by XREAL_0:def_1; A60: supk - p92 = (sup (rng K)) - (R_EAL (p / 2)) by SUPINF_2:3; then A61: p9 - (supk - p92) = (R_EAL p) - ((sup (rng K)) - (R_EAL (p / 2))) by SUPINF_2:3; then (R_EAL p) - ((sup (rng K)) - (R_EAL (p / 2))) < sup (rng J) by A57, XXREAL_0:9; then consider j being R_eal such that A62: j in rng J and A63: (R_EAL p) - ((sup (rng K)) - (R_EAL (p / 2))) < j by XXREAL_2:94; p9 = (p9 - (supk - p92)) + (supk - p92) ; then A64: ((R_EAL p) - ((sup (rng K)) - (R_EAL (p / 2)))) + ((sup (rng K)) - (R_EAL (p / 2))) = p9 by A60, A61, SUPINF_2:1; consider k being ext-real number such that A65: k in rng K and A66: (sup (rng K)) - (R_EAL (p / 2)) < k by A58, A59, MEASURE6:6; consider n1 being set such that A67: n1 in dom K and A68: k = K . n1 by A65, FUNCT_1:def_3; consider n2 being set such that A69: n2 in dom J and A70: j = J . n2 by A62, FUNCT_1:def_3; reconsider n1 = n1 as Element of NAT by A67; reconsider n2 = n2 as Element of NAT by A69; set n = max (n1,n2); J . n2 <= J . (max (n1,n2)) by A1, XXREAL_0:25; then A71: (R_EAL p) - ((sup (rng K)) - (R_EAL (p / 2))) < J . (max (n1,n2)) by A63, A70, XXREAL_0:2; K . n1 <= K . (max (n1,n2)) by A2, XXREAL_0:25; then A72: (sup (rng K)) - (R_EAL (p / 2)) < K . (max (n1,n2)) by A66, A68, XXREAL_0:2; A73: R_EAL p < (J . (max (n1,n2))) + (K . (max (n1,n2))) by A72, A71, A64, XXREAL_3:64; now__::_thesis:_for_m_being_Nat_st_max_(n1,n2)_<=_m_holds_ p_<=_L_._m let m be Nat; ::_thesis: ( max (n1,n2) <= m implies p <= L . m ) assume A74: max (n1,n2) <= m ; ::_thesis: p <= L . m then A75: K . (max (n1,n2)) <= K . m by A2; J . (max (n1,n2)) <= J . m by A1, A74; then (J . (max (n1,n2))) + (K . (max (n1,n2))) <= (J . m) + (K . m) by A75, XXREAL_3:36; then (J . (max (n1,n2))) + (K . (max (n1,n2))) <= L . m by A5; hence p <= L . m by A73, XXREAL_0:2; ::_thesis: verum end; hence ex n being Nat st for m being Nat st n <= m holds p <= L . m ; ::_thesis: verum end; suppose sup (rng K) = +infty ; ::_thesis: ex n being Nat st for m being Nat st m <= b3 holds n <= L . b3 then consider n1 being Nat such that A76: R_EAL (p / 2) < K . n1 by A4, A58, Th59; consider n2 being Nat such that A77: R_EAL (p / 2) < J . n2 by A3, A57, A58, Th59; reconsider n1 = n1, n2 = n2 as Element of NAT by ORDINAL1:def_12; set n = max (n1,n2); A78: (p / 2) + (p / 2) = p ; K . n1 <= K . (max (n1,n2)) by A2, XXREAL_0:25; then A79: R_EAL (p / 2) < K . (max (n1,n2)) by A76, XXREAL_0:2; J . n2 <= J . (max (n1,n2)) by A1, XXREAL_0:25; then A80: R_EAL (p / 2) < J . (max (n1,n2)) by A77, XXREAL_0:2; (R_EAL (p / 2)) + (R_EAL (p / 2)) < (J . (max (n1,n2))) + (K . (max (n1,n2))) by A80, A79, XXREAL_3:64; then p < (J . (max (n1,n2))) + (K . (max (n1,n2))) by A78, SUPINF_2:1; then A81: R_EAL p < L . (max (n1,n2)) by A5; now__::_thesis:_for_m_being_Nat_st_max_(n1,n2)_<=_m_holds_ p_<=_L_._m let m be Nat; ::_thesis: ( max (n1,n2) <= m implies p <= L . m ) assume A82: max (n1,n2) <= m ; ::_thesis: p <= L . m then A83: K . (max (n1,n2)) <= K . m by A2; J . (max (n1,n2)) <= J . m by A1, A82; then (J . (max (n1,n2))) + (K . (max (n1,n2))) <= (J . m) + (K . m) by A83, XXREAL_3:36; then (J . (max (n1,n2))) + (K . (max (n1,n2))) <= L . m by A5; then L . (max (n1,n2)) <= L . m by A5; hence p <= L . m by A81, XXREAL_0:2; ::_thesis: verum end; hence ex n being Nat st for m being Nat st n <= m holds p <= L . m ; ::_thesis: verum end; end; end; then A84: L is convergent_to_+infty by Def9; hence L is convergent by Def11; ::_thesis: lim L = (sup (rng J)) + (sup (rng K)) then A85: lim L = +infty by A84, Def12; A86: K . 0 <= sup (rng K) by A6, FUNCT_1:3, XXREAL_2:4; -infty < K . 0 by A4, Def5; hence lim L = (sup (rng J)) + (sup (rng K)) by A57, A85, A86, XXREAL_3:def_2; ::_thesis: verum end; end; end; hence L is convergent ; ::_thesis: ( lim L = sup (rng L) & lim L = (lim J) + (lim K) & sup (rng L) = (sup (rng K)) + (sup (rng J)) ) A87: now__::_thesis:_for_n,_m_being_Nat_st_n_<=_m_holds_ L_._n_<=_L_._m let n, m be Nat; ::_thesis: ( n <= m implies L . n <= L . m ) assume A88: n <= m ; ::_thesis: L . n <= L . m then A89: K . n <= K . m by A2; J . n <= J . m by A1, A88; then (J . n) + (K . n) <= (J . m) + (K . m) by A89, XXREAL_3:36; then L . n <= (J . m) + (K . m) by A5; hence L . n <= L . m by A5; ::_thesis: verum end; hence lim L = sup (rng L) by Th54; ::_thesis: ( lim L = (lim J) + (lim K) & sup (rng L) = (sup (rng K)) + (sup (rng J)) ) lim J = sup (rng J) by A1, Th54; hence ( lim L = (lim J) + (lim K) & sup (rng L) = (sup (rng K)) + (sup (rng J)) ) by A2, A8, A87, Th54; ::_thesis: verum end; theorem Th62: :: MESFUNC5:62 for L, K being ExtREAL_sequence for c being Real st 0 <= c & L is () & ( for n being Nat holds K . n = (R_EAL c) * (L . n) ) holds ( sup (rng K) = (R_EAL c) * (sup (rng L)) & K is () ) proof let L, K be ExtREAL_sequence; ::_thesis: for c being Real st 0 <= c & L is () & ( for n being Nat holds K . n = (R_EAL c) * (L . n) ) holds ( sup (rng K) = (R_EAL c) * (sup (rng L)) & K is () ) let c be Real; ::_thesis: ( 0 <= c & L is () & ( for n being Nat holds K . n = (R_EAL c) * (L . n) ) implies ( sup (rng K) = (R_EAL c) * (sup (rng L)) & K is () ) ) assume that A1: 0 <= c and A2: L is () and A3: for n being Nat holds K . n = (R_EAL c) * (L . n) ; ::_thesis: ( sup (rng K) = (R_EAL c) * (sup (rng L)) & K is () ) now__::_thesis:_sup_(rng_K)_=_c_*_(sup_(rng_L)) percases ( sup (rng L) in REAL or sup (rng L) = +infty ) by A2, Lm8; supposeA4: sup (rng L) in REAL ; ::_thesis: sup (rng K) = c * (sup (rng L)) A5: for y being UpperBound of rng K holds (R_EAL c) * (sup (rng L)) <= y proof let y be UpperBound of rng K; ::_thesis: (R_EAL c) * (sup (rng L)) <= y reconsider y = y as R_eal by XXREAL_0:def_1; A6: dom L = NAT by FUNCT_2:def_1; A7: dom K = NAT by FUNCT_2:def_1; percases ( c = 0 or c <> 0 ) ; supposeA8: c = 0 ; ::_thesis: (R_EAL c) * (sup (rng L)) <= y A9: K . 1 <= y by A7, FUNCT_1:3, XXREAL_2:def_1; K . 1 = c * (L . 1) by A3; hence (R_EAL c) * (sup (rng L)) <= y by A8, A9; ::_thesis: verum end; supposeA10: c <> 0 ; ::_thesis: (R_EAL c) * (sup (rng L)) <= y now__::_thesis:_for_x_being_ext-real_number_st_x_in_rng_L_holds_ x_<=_y_/_(R_EAL_c) let x be ext-real number ; ::_thesis: ( x in rng L implies x <= y / (R_EAL c) ) assume x in rng L ; ::_thesis: x <= y / (R_EAL c) then consider n being set such that A11: n in dom L and A12: x = L . n by FUNCT_1:def_3; reconsider n = n as Element of NAT by A11; A13: K . n in rng K by A7, FUNCT_1:3; K . n = (R_EAL c) * (L . n) by A3; then ((R_EAL c) * (L . n)) / (R_EAL c) <= y / (R_EAL c) by A1, A10, A13, XXREAL_2:def_1, XXREAL_3:79; hence x <= y / (R_EAL c) by A10, A12, XXREAL_3:88; ::_thesis: verum end; then y / (R_EAL c) is UpperBound of rng L by XXREAL_2:def_1; then A14: sup (rng L) <= y / (R_EAL c) by XXREAL_2:def_3; A15: now__::_thesis:_not_y_=_-infty assume A16: y = -infty ; ::_thesis: contradiction K . 1 in rng K by A7, FUNCT_1:3; then K . 1 = -infty by A16, XXREAL_0:6, XXREAL_2:def_1; then A17: (R_EAL c) * (L . 1) = -infty by A3; L . 1 <= sup (rng L) by A6, FUNCT_1:3, XXREAL_2:4; then A18: L . 1 < +infty by A4, XXREAL_0:2, XXREAL_0:9; -infty < L . 1 by A2, Def5; hence contradiction by A17, A18, XXREAL_3:70; ::_thesis: verum end; percases ( y = +infty or y in REAL ) by A15, XXREAL_0:14; suppose y = +infty ; ::_thesis: (R_EAL c) * (sup (rng L)) <= y hence (R_EAL c) * (sup (rng L)) <= y by XXREAL_0:4; ::_thesis: verum end; suppose y in REAL ; ::_thesis: (R_EAL c) * (sup (rng L)) <= y then reconsider ry = y as Real ; reconsider sl = sup (rng L) as Real by A4; y / (R_EAL c) = ry / c by EXTREAL1:2; then sl * c <= ry by A1, A10, A14, XREAL_1:83; hence (R_EAL c) * (sup (rng L)) <= y by EXTREAL1:1; ::_thesis: verum end; end; end; end; end; now__::_thesis:_for_x_being_ext-real_number_st_x_in_rng_K_holds_ x_<=_(R_EAL_c)_*_(sup_(rng_L)) let x be ext-real number ; ::_thesis: ( x in rng K implies x <= (R_EAL c) * (sup (rng L)) ) A19: sup (rng L) is UpperBound of rng L by XXREAL_2:def_3; assume x in rng K ; ::_thesis: x <= (R_EAL c) * (sup (rng L)) then consider m being set such that A20: m in dom K and A21: x = K . m by FUNCT_1:def_3; reconsider m = m as Element of NAT by A20; dom L = NAT by FUNCT_2:def_1; then A22: L . m <= sup (rng L) by A19, FUNCT_1:3, XXREAL_2:def_1; x = (R_EAL c) * (L . m) by A3, A21; hence x <= (R_EAL c) * (sup (rng L)) by A1, A22, XXREAL_3:71; ::_thesis: verum end; then (R_EAL c) * (sup (rng L)) is UpperBound of rng K by XXREAL_2:def_1; hence sup (rng K) = c * (sup (rng L)) by A5, XXREAL_2:def_3; ::_thesis: verum end; supposeA23: sup (rng L) = +infty ; ::_thesis: sup (rng K) = c * (sup (rng L)) percases ( c = 0 or c <> 0 ) ; supposeA24: c = 0 ; ::_thesis: sup (rng K) = c * (sup (rng L)) A25: now__::_thesis:_for_n_being_Nat_holds_K_._n_=_0 let n be Nat; ::_thesis: K . n = 0 K . n = c * (L . n) by A3; hence K . n = 0 by A24; ::_thesis: verum end; then lim K = sup (rng K) by Th60; hence sup (rng K) = c * (sup (rng L)) by A24, A25, Th60; ::_thesis: verum end; supposeA26: c <> 0 ; ::_thesis: sup (rng K) = c * (sup (rng L)) now__::_thesis:_for_n_being_set_holds_-infty_<_K_._n let n be set ; ::_thesis: -infty < K . b1 -infty < L . n by A2, Def5; then A27: -infty * (R_EAL c) < (R_EAL c) * (L . n) by A1, A26, XXREAL_3:72; percases ( n in dom K or not n in dom K ) ; suppose n in dom K ; ::_thesis: -infty < K . b1 then reconsider n1 = n as Element of NAT ; -infty * (R_EAL c) = -infty by A1, A26, XXREAL_3:def_5; then -infty < K . n1 by A3, A27; hence -infty < K . n ; ::_thesis: verum end; suppose not n in dom K ; ::_thesis: -infty < K . b1 hence -infty < K . n by FUNCT_1:def_2; ::_thesis: verum end; end; end; then A28: K is () by Def5; A29: now__::_thesis:_for_k_being_real_number_st_0_<_k_holds_ ex_n_being_Nat_st_not_K_._n_<=_k let k be real number ; ::_thesis: ( 0 < k implies ex n being Nat st not K . n <= k ) reconsider k1 = k as Real by XREAL_0:def_1; A30: (R_EAL (k / c)) * (R_EAL c) = (k1 / c) * c by EXTREAL1:1 .= k1 / (c / c) by XCMPLX_1:82 .= k by A26, XCMPLX_1:51 ; assume 0 < k ; ::_thesis: not for n being Nat holds K . n <= k then consider n being Nat such that A31: R_EAL (k / c) < L . n by A1, A2, A23, A26, Th59; (R_EAL (k / c)) * (R_EAL c) < (R_EAL c) * (L . n) by A1, A26, A31, XXREAL_3:72; then R_EAL k < K . n by A3, A30; hence not for n being Nat holds K . n <= k ; ::_thesis: verum end; (R_EAL c) * (sup (rng L)) = +infty by A1, A23, A26, XXREAL_3:def_5; hence sup (rng K) = c * (sup (rng L)) by A28, A29, Th59; ::_thesis: verum end; end; end; end; end; hence sup (rng K) = (R_EAL c) * (sup (rng L)) ; ::_thesis: K is () now__::_thesis:_for_n_being_set_holds_-infty_<_K_._n let n be set ; ::_thesis: -infty < K . b1 A32: ( L . n = +infty implies (R_EAL c) * (L . n) <> -infty ) by A1; -infty < L . n by A2, Def5; then A33: -infty <> (R_EAL c) * (L . n) by A32, XXREAL_3:70; percases ( n in dom K or not n in dom K ) ; suppose n in dom K ; ::_thesis: -infty < K . b1 then reconsider n1 = n as Element of NAT ; K . n1 <> -infty by A3, A33; hence -infty < K . n by XXREAL_0:6; ::_thesis: verum end; suppose not n in dom K ; ::_thesis: -infty < K . b1 hence -infty < K . n by FUNCT_1:def_2; ::_thesis: verum end; end; end; hence K is () by Def5; ::_thesis: verum end; theorem Th63: :: MESFUNC5:63 for L, K being ExtREAL_sequence for c being Real st 0 <= c & ( for n, m being Nat st n <= m holds L . n <= L . m ) & ( for n being Nat holds K . n = (R_EAL c) * (L . n) ) & L is () holds ( ( for n, m being Nat st n <= m holds K . n <= K . m ) & K is () & K is convergent & lim K = sup (rng K) & lim K = (R_EAL c) * (lim L) ) proof let L, K be ExtREAL_sequence; ::_thesis: for c being Real st 0 <= c & ( for n, m being Nat st n <= m holds L . n <= L . m ) & ( for n being Nat holds K . n = (R_EAL c) * (L . n) ) & L is () holds ( ( for n, m being Nat st n <= m holds K . n <= K . m ) & K is () & K is convergent & lim K = sup (rng K) & lim K = (R_EAL c) * (lim L) ) let c be Real; ::_thesis: ( 0 <= c & ( for n, m being Nat st n <= m holds L . n <= L . m ) & ( for n being Nat holds K . n = (R_EAL c) * (L . n) ) & L is () implies ( ( for n, m being Nat st n <= m holds K . n <= K . m ) & K is () & K is convergent & lim K = sup (rng K) & lim K = (R_EAL c) * (lim L) ) ) assume that A1: 0 <= c and A2: for n, m being Nat st n <= m holds L . n <= L . m and A3: for n being Nat holds K . n = (R_EAL c) * (L . n) and A4: L is () ; ::_thesis: ( ( for n, m being Nat st n <= m holds K . n <= K . m ) & K is () & K is convergent & lim K = sup (rng K) & lim K = (R_EAL c) * (lim L) ) A5: sup (rng L) = lim L by A2, Th54; hereby ::_thesis: ( K is () & K is convergent & lim K = sup (rng K) & lim K = (R_EAL c) * (lim L) ) let n, m be Nat; ::_thesis: ( n <= m implies K . n <= K . m ) assume n <= m ; ::_thesis: K . n <= K . m then (R_EAL c) * (L . n) <= (R_EAL c) * (L . m) by A1, A2, XXREAL_3:71; then K . n <= (R_EAL c) * (L . m) by A3; hence K . n <= K . m by A3; ::_thesis: verum end; thus K is () by A1, A3, A4, Th62; ::_thesis: ( K is convergent & lim K = sup (rng K) & lim K = (R_EAL c) * (lim L) ) thus ( K is convergent & lim K = sup (rng K) ) by A6, Th54; ::_thesis: lim K = (R_EAL c) * (lim L) sup (rng K) = lim K by A6, Th54; hence lim K = (R_EAL c) * (lim L) by A1, A3, A4, A5, Th62; ::_thesis: verum end; begin definition let X be non empty set ; let H be Functional_Sequence of X,ExtREAL; let x be Element of X; funcH # x -> ExtREAL_sequence means :Def13: :: MESFUNC5:def 13 for n being Nat holds it . n = (H . n) . x; existence ex b1 being ExtREAL_sequence st for n being Nat holds b1 . n = (H . n) . x proof deffunc H1( Nat) -> Element of ExtREAL = (H . $1) . x; consider f being Function of NAT,ExtREAL such that A1: for n being Element of NAT holds f . n = H1(n) from FUNCT_2:sch_4(); take f ; ::_thesis: for n being Nat holds f . n = (H . n) . x let n be Nat; ::_thesis: f . n = (H . n) . x n in NAT by ORDINAL1:def_12; hence f . n = (H . n) . x by A1; ::_thesis: verum end; uniqueness for b1, b2 being ExtREAL_sequence st ( for n being Nat holds b1 . n = (H . n) . x ) & ( for n being Nat holds b2 . n = (H . n) . x ) holds b1 = b2 proof let S1, S2 be ExtREAL_sequence; ::_thesis: ( ( for n being Nat holds S1 . n = (H . n) . x ) & ( for n being Nat holds S2 . n = (H . n) . x ) implies S1 = S2 ) assume that A2: for n being Nat holds S1 . n = (H . n) . x and A3: for n being Nat holds S2 . n = (H . n) . x ; ::_thesis: S1 = S2 now__::_thesis:_for_n_being_Element_of_NAT_holds_S1_._n_=_S2_._n let n be Element of NAT ; ::_thesis: S1 . n = S2 . n S1 . n = (H . n) . x by A2; hence S1 . n = S2 . n by A3; ::_thesis: verum end; hence S1 = S2 by FUNCT_2:63; ::_thesis: verum end; end; :: deftheorem Def13 defines # MESFUNC5:def_13_:_ for X being non empty set for H being Functional_Sequence of X,ExtREAL for x being Element of X for b4 being ExtREAL_sequence holds ( b4 = H # x iff for n being Nat holds b4 . n = (H . n) . x ); definition let D1, D2 be set ; let F be Function of NAT,(PFuncs (D1,D2)); let n be Nat; :: original: . redefine funcF . n -> PartFunc of D1,D2; coherence F . n is PartFunc of D1,D2 proof n in NAT by ORDINAL1:def_12; then n in dom F by FUNCT_2:def_1; then F . n in rng F by FUNCT_1:def_3; hence F . n is PartFunc of D1,D2 ; ::_thesis: verum end; end; theorem Th64: :: MESFUNC5:64 for X being non empty set for S being SigmaField of X for f being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative holds ex F being Functional_Sequence of X,ExtREAL st ( ( for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = dom f ) ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom f holds ( F # x is convergent & lim (F # x) = f . x ) ) ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for f being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative holds ex F being Functional_Sequence of X,ExtREAL st ( ( for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = dom f ) ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom f holds ( F # x is convergent & lim (F # x) = f . x ) ) ) let S be SigmaField of X; ::_thesis: for f being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative holds ex F being Functional_Sequence of X,ExtREAL st ( ( for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = dom f ) ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom f holds ( F # x is convergent & lim (F # x) = f . x ) ) ) let f be PartFunc of X,ExtREAL; ::_thesis: ( ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative implies ex F being Functional_Sequence of X,ExtREAL st ( ( for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = dom f ) ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom f holds ( F # x is convergent & lim (F # x) = f . x ) ) ) ) assume that A1: ex A being Element of S st ( A = dom f & f is_measurable_on A ) and A2: f is nonnegative ; ::_thesis: ex F being Functional_Sequence of X,ExtREAL st ( ( for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = dom f ) ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom f holds ( F # x is convergent & lim (F # x) = f . x ) ) ) defpred S1[ Nat, PartFunc of X,ExtREAL] means ( dom $2 = dom f & ( for x being Element of X st x in dom f holds ( ( for k being Nat st 1 <= k & k <= (2 |^ $1) * $1 & (k - 1) / (2 |^ $1) <= f . x & f . x < k / (2 |^ $1) holds $2 . x = (k - 1) / (2 |^ $1) ) & ( $1 <= f . x implies $2 . x = $1 ) ) ) ); A3: for n being Element of NAT ex y being Element of PFuncs (X,ExtREAL) st S1[n,y] proof let n be Element of NAT ; ::_thesis: ex y being Element of PFuncs (X,ExtREAL) st S1[n,y] defpred S2[ set , set ] means ( ( for k being Nat st 1 <= k & k <= (2 |^ n) * n & (k - 1) / (2 |^ n) <= f . $1 & f . $1 < k / (2 |^ n) holds $2 = (k - 1) / (2 |^ n) ) & ( n <= f . $1 implies $2 = n ) ); A4: for x being set st x in dom f holds ex y being set st S2[x,y] proof let x be set ; ::_thesis: ( x in dom f implies ex y being set st S2[x,y] ) assume x in dom f ; ::_thesis: ex y being set st S2[x,y] percases ( f . x < n or n <= f . x ) ; supposeA5: f . x < n ; ::_thesis: ex y being set st S2[x,y] 0 <= f . x by A2, SUPINF_2:51; then consider k being Nat such that 1 <= k and k <= (2 |^ n) * n and A6: (k - 1) / (2 |^ n) <= f . x and A7: f . x < k / (2 |^ n) by A5, Th4; take y = (k - 1) / (2 |^ n); ::_thesis: S2[x,y] now__::_thesis:_for_k1_being_Nat_st_1_<=_k1_&_k1_<=_(2_|^_n)_*_n_&_(k1_-_1)_/_(2_|^_n)_<=_f_._x_&_f_._x_<_k1_/_(2_|^_n)_holds_ y_=_(k1_-_1)_/_(2_|^_n) let k1 be Nat; ::_thesis: ( 1 <= k1 & k1 <= (2 |^ n) * n & (k1 - 1) / (2 |^ n) <= f . x & f . x < k1 / (2 |^ n) implies y = (k1 - 1) / (2 |^ n) ) assume that 1 <= k1 and k1 <= (2 |^ n) * n and A8: (k1 - 1) / (2 |^ n) <= f . x and A9: f . x < k1 / (2 |^ n) ; ::_thesis: y = (k1 - 1) / (2 |^ n) A10: now__::_thesis:_not_k1_<_k assume k1 < k ; ::_thesis: contradiction then k1 + 1 <= k by NAT_1:13; then k1 <= k - 1 by XREAL_1:19; then k1 / (2 |^ n) <= (k - 1) / (2 |^ n) by XREAL_1:72; hence contradiction by A6, A9, XXREAL_0:2; ::_thesis: verum end; now__::_thesis:_not_k_<_k1 assume k < k1 ; ::_thesis: contradiction then k + 1 <= k1 by NAT_1:13; then k <= k1 - 1 by XREAL_1:19; then k / (2 |^ n) <= (k1 - 1) / (2 |^ n) by XREAL_1:72; hence contradiction by A7, A8, XXREAL_0:2; ::_thesis: verum end; hence y = (k1 - 1) / (2 |^ n) by A10, XXREAL_0:1; ::_thesis: verum end; hence S2[x,y] by A5; ::_thesis: verum end; supposeA11: n <= f . x ; ::_thesis: ex y being set st S2[x,y] reconsider y = n as Real ; take y ; ::_thesis: S2[x,y] thus S2[x,y] by A11, Th5; ::_thesis: verum end; end; end; consider fn being Function such that A12: ( dom fn = dom f & ( for x being set st x in dom f holds S2[x,fn . x] ) ) from CLASSES1:sch_1(A4); now__::_thesis:_for_w_being_set_st_w_in_rng_fn_holds_ w_in_ExtREAL let w be set ; ::_thesis: ( w in rng fn implies b1 in ExtREAL ) assume w in rng fn ; ::_thesis: b1 in ExtREAL then consider v being set such that A13: v in dom fn and A14: w = fn . v by FUNCT_1:def_3; percases ( R_EAL n <= f . v or f . v < R_EAL n ) ; suppose R_EAL n <= f . v ; ::_thesis: b1 in ExtREAL then fn . v = R_EAL n by A12, A13; hence w in ExtREAL by A14; ::_thesis: verum end; supposeA15: f . v < R_EAL n ; ::_thesis: b1 in ExtREAL 0 <= f . v by A2, SUPINF_2:51; then consider k being Nat such that A16: 1 <= k and A17: k <= (2 |^ n) * n and A18: (k - 1) / (2 |^ n) <= f . v and A19: f . v < k / (2 |^ n) by A15, Th4; fn . v = (k - 1) / (2 |^ n) by A12, A13, A16, A17, A18, A19; hence w in ExtREAL by A14, XXREAL_0:def_1; ::_thesis: verum end; end; end; then rng fn c= ExtREAL by TARSKI:def_3; then reconsider fn = fn as PartFunc of (dom f),ExtREAL by A12, RELSET_1:4; reconsider fn = fn as PartFunc of X,ExtREAL by A12, RELSET_1:5; reconsider y = fn as Element of PFuncs (X,ExtREAL) by PARTFUN1:45; take y ; ::_thesis: S1[n,y] thus S1[n,y] by A12; ::_thesis: verum end; ex F being Function of NAT,(PFuncs (X,ExtREAL)) st for n being Element of NAT holds S1[n,F . n] from FUNCT_2:sch_3(A3); then consider F being Function of NAT,(PFuncs (X,ExtREAL)) such that A20: for n being Element of NAT holds dom (F . n) = dom f and A21: for n being Element of NAT for x being Element of X st x in dom f holds ( ( for k being Nat st 1 <= k & k <= (2 |^ n) * n & (k - 1) / (2 |^ n) <= f . x & f . x < k / (2 |^ n) holds (F . n) . x = (k - 1) / (2 |^ n) ) & ( n <= f . x implies (F . n) . x = n ) ) ; A22: now__::_thesis:_for_n_being_Nat_holds_dom_(F_._n)_=_dom_f let n be Nat; ::_thesis: dom (F . n) = dom f n in NAT by ORDINAL1:def_12; hence dom (F . n) = dom f by A20; ::_thesis: verum end; reconsider F = F as Functional_Sequence of X,ExtREAL ; consider A being Element of S such that A23: A = dom f and A24: f is_measurable_on A by A1; A25: for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F . n) . x <= (F . m) . x proof let n, m be Nat; ::_thesis: ( n <= m implies for x being Element of X st x in dom f holds (F . n) . x <= (F . m) . x ) assume A26: n <= m ; ::_thesis: for x being Element of X st x in dom f holds (F . n) . x <= (F . m) . x reconsider n = n, m = m as Element of NAT by ORDINAL1:def_12; let x be Element of X; ::_thesis: ( x in dom f implies (F . n) . x <= (F . m) . x ) assume A27: x in dom f ; ::_thesis: (F . n) . x <= (F . m) . x percases ( R_EAL m <= f . x or f . x < R_EAL m ) ; supposeA28: R_EAL m <= f . x ; ::_thesis: (F . n) . x <= (F . m) . x then A29: R_EAL n <= f . x by A26, XXREAL_0:2; (F . m) . x = R_EAL m by A21, A27, A28; hence (F . n) . x <= (F . m) . x by A21, A26, A27, A29; ::_thesis: verum end; supposeA30: f . x < R_EAL m ; ::_thesis: (F . n) . x <= (F . m) . x A31: 0 <= f . x by A2, SUPINF_2:51; then consider M being Nat such that A32: 1 <= M and A33: M <= (2 |^ m) * m and A34: (M - 1) / (2 |^ m) <= f . x and A35: f . x < M / (2 |^ m) by A30, Th4; reconsider M = M as Element of NAT by ORDINAL1:def_12; A36: (F . m) . x = (M - 1) / (2 |^ m) by A21, A27, A32, A33, A34, A35; percases ( R_EAL n <= f . x or f . x < R_EAL n ) ; supposeA37: R_EAL n <= f . x ; ::_thesis: (F . n) . x <= (F . m) . x reconsider M1 = 2 |^ m as Element of NAT ; n < M / (2 |^ m) by A35, A37, XXREAL_0:2; then (2 |^ m) * n < M by PREPOWER:6, XREAL_1:79; then (M1 * n) + 1 <= M by NAT_1:13; then A38: M1 * n <= M - 1 by XREAL_1:19; A39: 0 < 2 |^ m by PREPOWER:6; (F . n) . x = R_EAL n by A21, A27, A37; hence (F . n) . x <= (F . m) . x by A36, A39, A38, XREAL_1:77; ::_thesis: verum end; supposeA40: f . x < R_EAL n ; ::_thesis: (F . n) . x <= (F . m) . x consider k being Nat such that A41: m = n + k by A26, NAT_1:10; reconsider k = k as Element of NAT by ORDINAL1:def_12; reconsider K = 2 |^ k as Element of NAT ; consider N1 being Nat such that A42: 1 <= N1 and A43: N1 <= (2 |^ n) * n and A44: (N1 - 1) / (2 |^ n) <= f . x and A45: f . x < N1 / (2 |^ n) by A31, A40, Th4; reconsider N1 = N1 as Element of NAT by ORDINAL1:def_12; A46: (F . n) . x = (N1 - 1) / (2 |^ n) by A21, A27, A42, A43, A44, A45; R_EAL ((N1 - 1) / (2 |^ n)) < R_EAL (M / (2 |^ (n + k))) by A35, A44, A41, XXREAL_0:2; then (N1 - 1) / (2 |^ n) < M / ((2 |^ n) * (2 |^ k)) by NEWTON:8; then (N1 - 1) / (2 |^ n) < (M / (2 |^ k)) / (2 |^ n) by XCMPLX_1:78; then N1 - 1 < M / (2 |^ k) by XREAL_1:72; then K * (N1 - 1) < M by PREPOWER:6, XREAL_1:79; then (K * (N1 - 1)) + 1 <= M by INT_1:7; then K * (N1 - 1) <= M - 1 by XREAL_1:19; then (K * (N1 - 1)) / (2 |^ (n + k)) <= (M - 1) / (2 |^ (n + k)) by XREAL_1:72; then A47: (K * (N1 - 1)) / ((2 |^ n) * (2 |^ k)) <= (M - 1) / (2 |^ (n + k)) by NEWTON:8; 2 |^ k > 0 by PREPOWER:6; hence (F . n) . x <= (F . m) . x by A36, A46, A41, A47, XCMPLX_1:91; ::_thesis: verum end; end; end; end; end; A48: for n being Nat holds F . n is_simple_func_in S proof let n be Nat; ::_thesis: F . n is_simple_func_in S reconsider n = n as Element of NAT by ORDINAL1:def_12; reconsider N = 2 |^ n as Element of NAT ; defpred S2[ Nat, set ] means ( ( $1 <= N * n implies $2 = (A /\ (great_eq_dom (f,(R_EAL (($1 - 1) / (2 |^ n)))))) /\ (less_dom (f,(R_EAL ($1 / (2 |^ n))))) ) & ( $1 = (N * n) + 1 implies $2 = A /\ (great_eq_dom (f,(R_EAL n))) ) ); now__::_thesis:_for_x_being_Element_of_X_st_x_in_dom_(F_._n)_holds_ |.((F_._n)_._x).|_<_+infty let x be Element of X; ::_thesis: ( x in dom (F . n) implies |.((F . n) . b1).| < +infty ) assume x in dom (F . n) ; ::_thesis: |.((F . n) . b1).| < +infty then A49: x in dom f by A22; percases ( R_EAL n <= f . x or f . x < R_EAL n ) ; supposeA50: R_EAL n <= f . x ; ::_thesis: |.((F . n) . b1).| < +infty then (F . n) . x = R_EAL n by A21, A49; then A51: (F . n) . x < +infty by XXREAL_0:9; - +infty < (F . n) . x by A21, A49, A50; hence |.((F . n) . x).| < +infty by A51, EXTREAL2:11; ::_thesis: verum end; supposeA52: f . x < R_EAL n ; ::_thesis: |.((F . n) . b1).| < +infty A53: 0 <= f . x by A2, SUPINF_2:51; R_EAL n < +infty by XXREAL_0:9; then reconsider y = f . x as Real by A52, A53, XXREAL_0:14; set k = [\((2 |^ n) * y)/] + 1; A54: [\((2 |^ n) * y)/] <= (2 |^ n) * y by INT_1:def_6; ((2 |^ n) * y) - 1 < [\((2 |^ n) * y)/] by INT_1:def_6; then A55: (2 |^ n) * y < [\((2 |^ n) * y)/] + 1 by XREAL_1:19; A56: 0 < 2 |^ n by PREPOWER:6; then (2 |^ n) * y < (2 |^ n) * n by A52, XREAL_1:68; then [\((2 |^ n) * y)/] < (2 |^ n) * n by A54, XXREAL_0:2; then A57: [\((2 |^ n) * y)/] + 1 <= (2 |^ n) * n by INT_1:7; A58: 0 <= (2 |^ n) * y by A53; then A59: 0 + 1 <= [\((2 |^ n) * y)/] + 1 by A55, INT_1:7; reconsider k = [\((2 |^ n) * y)/] + 1 as Element of NAT by A58, A55, INT_1:3; k - 1 <= (2 |^ n) * y by INT_1:def_6; then A60: (k - 1) / (2 |^ n) <= y by PREPOWER:6, XREAL_1:79; y < k / (2 |^ n) by A56, INT_1:29, XREAL_1:81; then A61: (F . n) . x = (k - 1) / (2 |^ n) by A21, A49, A59, A57, A60; then -infty < (F . n) . x by XXREAL_0:12; then A62: - +infty < (F . n) . x by XXREAL_3:def_3; (F . n) . x < +infty by A61, XXREAL_0:9; hence |.((F . n) . x).| < +infty by A62, EXTREAL2:11; ::_thesis: verum end; end; end; then A63: F . n is V60() by MESFUNC2:def_1; A64: now__::_thesis:_for_k_being_Nat_st_k_in_Seg_((N_*_n)_+_1)_holds_ ex_B_being_Element_of_S_st_S2[k,B] let k be Nat; ::_thesis: ( k in Seg ((N * n) + 1) implies ex B being Element of S st S2[B,b2] ) assume k in Seg ((N * n) + 1) ; ::_thesis: ex B being Element of S st S2[B,b2] reconsider k1 = k as Element of NAT by ORDINAL1:def_12; percases ( k <> (N * n) + 1 or k = (N * n) + 1 ) ; supposeA65: k <> (N * n) + 1 ; ::_thesis: ex B being Element of S st S2[B,b2] set B = (A /\ (great_eq_dom (f,(R_EAL ((k1 - 1) / (2 |^ n)))))) /\ (less_dom (f,(R_EAL (k1 / (2 |^ n))))); reconsider B = (A /\ (great_eq_dom (f,(R_EAL ((k1 - 1) / (2 |^ n)))))) /\ (less_dom (f,(R_EAL (k1 / (2 |^ n))))) as Element of S by A23, A24, Th33; take B = B; ::_thesis: S2[k,B] thus S2[k,B] by A65; ::_thesis: verum end; supposeA66: k = (N * n) + 1 ; ::_thesis: ex B being Element of S st S2[B,b2] set B = A /\ (great_eq_dom (f,(R_EAL n))); reconsider B = A /\ (great_eq_dom (f,(R_EAL n))) as Element of S by A23, A24, MESFUNC1:27; take B = B; ::_thesis: S2[k,B] thus S2[k,B] by A66, NAT_1:13; ::_thesis: verum end; end; end; consider G being FinSequence of S such that A67: ( dom G = Seg ((N * n) + 1) & ( for k being Nat st k in Seg ((N * n) + 1) holds S2[k,G . k] ) ) from FINSEQ_1:sch_5(A64); A68: now__::_thesis:_for_k_being_Nat_st_1_<=_k_&_k_<=_(2_|^_n)_*_n_holds_ G_._k_=_(A_/\_(great_eq_dom_(f,(R_EAL_((k_-_1)_/_(2_|^_n))))))_/\_(less_dom_(f,(R_EAL_(k_/_(2_|^_n))))) let k be Nat; ::_thesis: ( 1 <= k & k <= (2 |^ n) * n implies G . k = (A /\ (great_eq_dom (f,(R_EAL ((k - 1) / (2 |^ n)))))) /\ (less_dom (f,(R_EAL (k / (2 |^ n))))) ) assume that A69: 1 <= k and A70: k <= (2 |^ n) * n ; ::_thesis: G . k = (A /\ (great_eq_dom (f,(R_EAL ((k - 1) / (2 |^ n)))))) /\ (less_dom (f,(R_EAL (k / (2 |^ n))))) A71: k in NAT by ORDINAL1:def_12; k <= (N * n) + 1 by A70, NAT_1:12; then k in Seg ((N * n) + 1) by A69, A71; hence G . k = (A /\ (great_eq_dom (f,(R_EAL ((k - 1) / (2 |^ n)))))) /\ (less_dom (f,(R_EAL (k / (2 |^ n))))) by A67, A70; ::_thesis: verum end; A72: len G = ((2 |^ n) * n) + 1 by A67, FINSEQ_1:def_3; now__::_thesis:_for_x,_y_being_set_st_x_<>_y_holds_ G_._x_misses_G_._y let x, y be set ; ::_thesis: ( x <> y implies G . b1 misses G . b2 ) assume A73: x <> y ; ::_thesis: G . b1 misses G . b2 percases ( not x in dom G or not y in dom G or ( x in dom G & y in dom G ) ) ; suppose ( not x in dom G or not y in dom G ) ; ::_thesis: G . b1 misses G . b2 then ( G . x = {} or G . y = {} ) by FUNCT_1:def_2; hence G . x misses G . y by XBOOLE_1:65; ::_thesis: verum end; supposeA74: ( x in dom G & y in dom G ) ; ::_thesis: G . b1 misses G . b2 then reconsider x1 = x, y1 = y as Element of NAT ; A75: x1 in Seg (len G) by A74, FINSEQ_1:def_3; then A76: 1 <= x1 by FINSEQ_1:1; A77: y1 in Seg (len G) by A74, FINSEQ_1:def_3; then A78: 1 <= y1 by FINSEQ_1:1; A79: y1 <= ((2 |^ n) * n) + 1 by A72, A77, FINSEQ_1:1; A80: x1 <= ((2 |^ n) * n) + 1 by A72, A75, FINSEQ_1:1; now__::_thesis:_(_(_x1_<_y1_&_(_y1_=_((2_|^_n)_*_n)_+_1_implies_G_._x_misses_G_._y_)_&_(_y1_<>_((2_|^_n)_*_n)_+_1_implies_G_._x_misses_G_._y_)_)_or_(_y1_<_x1_&_(_x1_<>_((2_|^_n)_*_n)_+_1_implies_G_._x_misses_G_._y_)_&_(_x1_=_((2_|^_n)_*_n)_+_1_implies_G_._x_misses_G_._y_)_)_) percases ( x1 < y1 or y1 < x1 ) by A73, XXREAL_0:1; caseA81: x1 < y1 ; ::_thesis: ( ( y1 = ((2 |^ n) * n) + 1 implies G . x misses G . y ) & ( y1 <> ((2 |^ n) * n) + 1 implies G . x misses G . y ) ) hereby ::_thesis: ( y1 <> ((2 |^ n) * n) + 1 implies G . x misses G . y ) assume A82: y1 = ((2 |^ n) * n) + 1 ; ::_thesis: G . x misses G . y then A83: G . y = A /\ (great_eq_dom (f,(R_EAL n))) by A67, A74; A84: x1 <= N * n by A81, A82, NAT_1:13; then A85: G . x = (A /\ (great_eq_dom (f,(R_EAL ((x1 - 1) / (2 |^ n)))))) /\ (less_dom (f,(R_EAL (x1 / (2 |^ n))))) by A68, A76; now__::_thesis:_for_v_being_set_holds_not_v_in_(G_._x)_/\_(G_._y) assume ex v being set st v in (G . x) /\ (G . y) ; ::_thesis: contradiction then consider v being set such that A86: v in (G . x) /\ (G . y) ; v in G . y by A86, XBOOLE_0:def_4; then v in great_eq_dom (f,(R_EAL n)) by A83, XBOOLE_0:def_4; then A87: R_EAL n <= f . v by MESFUNC1:def_14; v in G . x by A86, XBOOLE_0:def_4; then v in less_dom (f,(R_EAL (x1 / (2 |^ n)))) by A85, XBOOLE_0:def_4; then f . v < R_EAL (x1 / (2 |^ n)) by MESFUNC1:def_11; then R_EAL n < R_EAL (x1 / (2 |^ n)) by A87, XXREAL_0:2; hence contradiction by A84, PREPOWER:6, XREAL_1:79; ::_thesis: verum end; then (G . x) /\ (G . y) = {} by XBOOLE_0:def_1; hence G . x misses G . y by XBOOLE_0:def_7; ::_thesis: verum end; assume y1 <> ((2 |^ n) * n) + 1 ; ::_thesis: G . x misses G . y then y1 < (N * n) + 1 by A79, XXREAL_0:1; then A88: y1 <= N * n by NAT_1:13; then x1 <= (2 |^ n) * n by A81, XXREAL_0:2; then A89: G . x = (A /\ (great_eq_dom (f,(R_EAL ((x1 - 1) / (2 |^ n)))))) /\ (less_dom (f,(R_EAL (x1 / (2 |^ n))))) by A68, A76; A90: G . y = (A /\ (great_eq_dom (f,(R_EAL ((y1 - 1) / (2 |^ n)))))) /\ (less_dom (f,(R_EAL (y1 / (2 |^ n))))) by A68, A78, A88; now__::_thesis:_for_v_being_set_holds_not_v_in_(G_._x)_/\_(G_._y) assume ex v being set st v in (G . x) /\ (G . y) ; ::_thesis: contradiction then consider v being set such that A91: v in (G . x) /\ (G . y) ; v in G . y by A91, XBOOLE_0:def_4; then v in A /\ (great_eq_dom (f,(R_EAL ((y1 - 1) / (2 |^ n))))) by A90, XBOOLE_0:def_4; then v in great_eq_dom (f,(R_EAL ((y1 - 1) / (2 |^ n)))) by XBOOLE_0:def_4; then A92: R_EAL ((y1 - 1) / (2 |^ n)) <= f . v by MESFUNC1:def_14; v in G . x by A91, XBOOLE_0:def_4; then v in less_dom (f,(R_EAL (x1 / (2 |^ n)))) by A89, XBOOLE_0:def_4; then f . v < R_EAL (x1 / (2 |^ n)) by MESFUNC1:def_11; then R_EAL ((y1 - 1) / (2 |^ n)) < R_EAL (x1 / (2 |^ n)) by A92, XXREAL_0:2; then y1 - 1 < x1 by XREAL_1:72; then y1 < x1 + 1 by XREAL_1:19; hence contradiction by A81, NAT_1:13; ::_thesis: verum end; then (G . x) /\ (G . y) = {} by XBOOLE_0:def_1; hence G . x misses G . y by XBOOLE_0:def_7; ::_thesis: verum end; caseA93: y1 < x1 ; ::_thesis: ( ( x1 <> ((2 |^ n) * n) + 1 implies G . x misses G . y ) & ( x1 = ((2 |^ n) * n) + 1 implies G . x misses G . y ) ) hereby ::_thesis: ( x1 = ((2 |^ n) * n) + 1 implies G . x misses G . y ) assume x1 <> ((2 |^ n) * n) + 1 ; ::_thesis: G . x misses G . y then x1 < (N * n) + 1 by A80, XXREAL_0:1; then A94: x1 <= N * n by NAT_1:13; then y1 <= (2 |^ n) * n by A93, XXREAL_0:2; then A95: G . y = (A /\ (great_eq_dom (f,(R_EAL ((y1 - 1) / (2 |^ n)))))) /\ (less_dom (f,(R_EAL (y1 / (2 |^ n))))) by A68, A78; A96: G . x = (A /\ (great_eq_dom (f,(R_EAL ((x1 - 1) / (2 |^ n)))))) /\ (less_dom (f,(R_EAL (x1 / (2 |^ n))))) by A68, A76, A94; now__::_thesis:_for_v_being_set_holds_not_v_in_(G_._x)_/\_(G_._y) assume ex v being set st v in (G . x) /\ (G . y) ; ::_thesis: contradiction then consider v being set such that A97: v in (G . x) /\ (G . y) ; v in G . x by A97, XBOOLE_0:def_4; then v in A /\ (great_eq_dom (f,(R_EAL ((x1 - 1) / (2 |^ n))))) by A96, XBOOLE_0:def_4; then v in great_eq_dom (f,(R_EAL ((x1 - 1) / (2 |^ n)))) by XBOOLE_0:def_4; then A98: R_EAL ((x1 - 1) / (2 |^ n)) <= f . v by MESFUNC1:def_14; v in G . y by A97, XBOOLE_0:def_4; then v in less_dom (f,(R_EAL (y1 / (2 |^ n)))) by A95, XBOOLE_0:def_4; then f . v < R_EAL (y1 / (2 |^ n)) by MESFUNC1:def_11; then R_EAL ((x1 - 1) / (2 |^ n)) < R_EAL (y1 / (2 |^ n)) by A98, XXREAL_0:2; then x1 - 1 < y1 by XREAL_1:72; then x1 < y1 + 1 by XREAL_1:19; hence contradiction by A93, NAT_1:13; ::_thesis: verum end; then (G . x) /\ (G . y) = {} by XBOOLE_0:def_1; hence G . x misses G . y by XBOOLE_0:def_7; ::_thesis: verum end; assume A99: x1 = ((2 |^ n) * n) + 1 ; ::_thesis: G . x misses G . y then A100: G . x = A /\ (great_eq_dom (f,(R_EAL n))) by A67, A74; A101: y1 <= N * n by A93, A99, NAT_1:13; then A102: G . y = (A /\ (great_eq_dom (f,(R_EAL ((y1 - 1) / (2 |^ n)))))) /\ (less_dom (f,(R_EAL (y1 / (2 |^ n))))) by A68, A78; now__::_thesis:_for_v_being_set_holds_not_v_in_(G_._x)_/\_(G_._y) assume ex v being set st v in (G . x) /\ (G . y) ; ::_thesis: contradiction then consider v being set such that A103: v in (G . x) /\ (G . y) ; v in G . y by A103, XBOOLE_0:def_4; then v in less_dom (f,(R_EAL (y1 / (2 |^ n)))) by A102, XBOOLE_0:def_4; then A104: f . v < R_EAL (y1 / (2 |^ n)) by MESFUNC1:def_11; v in G . x by A103, XBOOLE_0:def_4; then v in great_eq_dom (f,(R_EAL n)) by A100, XBOOLE_0:def_4; then R_EAL n <= f . v by MESFUNC1:def_14; then R_EAL n < R_EAL (y1 / (2 |^ n)) by A104, XXREAL_0:2; hence contradiction by A101, PREPOWER:6, XREAL_1:79; ::_thesis: verum end; then (G . x) /\ (G . y) = {} by XBOOLE_0:def_1; hence G . x misses G . y by XBOOLE_0:def_7; ::_thesis: verum end; end; end; hence G . x misses G . y ; ::_thesis: verum end; end; end; then reconsider G = G as Finite_Sep_Sequence of S by PROB_2:def_2; A105: for k being Nat for x, y being Element of X st k in dom G & x in G . k & y in G . k holds (F . n) . x = (F . n) . y proof let k be Nat; ::_thesis: for x, y being Element of X st k in dom G & x in G . k & y in G . k holds (F . n) . x = (F . n) . y let x, y be Element of X; ::_thesis: ( k in dom G & x in G . k & y in G . k implies (F . n) . x = (F . n) . y ) assume that A106: k in dom G and A107: x in G . k and A108: y in G . k ; ::_thesis: (F . n) . x = (F . n) . y A109: 1 <= k by A67, A106, FINSEQ_1:1; A110: k <= (N * n) + 1 by A67, A106, FINSEQ_1:1; now__::_thesis:_(F_._n)_._x_=_(F_._n)_._y percases ( k = (N * n) + 1 or k <> (N * n) + 1 ) ; suppose k = (N * n) + 1 ; ::_thesis: (F . n) . x = (F . n) . y then A111: G . k = A /\ (great_eq_dom (f,(R_EAL n))) by A67, A106; then x in great_eq_dom (f,(R_EAL n)) by A107, XBOOLE_0:def_4; then A112: R_EAL n <= f . x by MESFUNC1:def_14; y in great_eq_dom (f,(R_EAL n)) by A108, A111, XBOOLE_0:def_4; then A113: R_EAL n <= f . y by MESFUNC1:def_14; x in A by A107, A111, XBOOLE_0:def_4; then A114: (F . n) . x = R_EAL n by A23, A21, A112; y in A by A108, A111, XBOOLE_0:def_4; hence (F . n) . x = (F . n) . y by A23, A21, A114, A113; ::_thesis: verum end; suppose k <> (N * n) + 1 ; ::_thesis: (F . n) . x = (F . n) . y then k < (N * n) + 1 by A110, XXREAL_0:1; then A115: k <= N * n by NAT_1:13; then A116: G . k = (A /\ (great_eq_dom (f,(R_EAL ((k - 1) / (2 |^ n)))))) /\ (less_dom (f,(R_EAL (k / (2 |^ n))))) by A67, A106; then x in less_dom (f,(R_EAL (k / (2 |^ n)))) by A107, XBOOLE_0:def_4; then A117: f . x < R_EAL (k / (2 |^ n)) by MESFUNC1:def_11; A118: x in A /\ (great_eq_dom (f,(R_EAL ((k - 1) / (2 |^ n))))) by A107, A116, XBOOLE_0:def_4; then x in great_eq_dom (f,(R_EAL ((k - 1) / (2 |^ n)))) by XBOOLE_0:def_4; then A119: R_EAL ((k - 1) / (2 |^ n)) <= f . x by MESFUNC1:def_14; x in A by A118, XBOOLE_0:def_4; then A120: (F . n) . x = R_EAL ((k - 1) / (2 |^ n)) by A23, A21, A109, A115, A117, A119; y in less_dom (f,(R_EAL (k / (2 |^ n)))) by A108, A116, XBOOLE_0:def_4; then A121: f . y < R_EAL (k / (2 |^ n)) by MESFUNC1:def_11; A122: y in A /\ (great_eq_dom (f,(R_EAL ((k - 1) / (2 |^ n))))) by A108, A116, XBOOLE_0:def_4; then y in great_eq_dom (f,(R_EAL ((k - 1) / (2 |^ n)))) by XBOOLE_0:def_4; then A123: R_EAL ((k - 1) / (2 |^ n)) <= f . y by MESFUNC1:def_14; y in A by A122, XBOOLE_0:def_4; hence (F . n) . x = (F . n) . y by A23, A21, A109, A115, A121, A123, A120; ::_thesis: verum end; end; end; hence (F . n) . x = (F . n) . y ; ::_thesis: verum end; for v being set st v in dom f holds v in union (rng G) proof let v be set ; ::_thesis: ( v in dom f implies v in union (rng G) ) assume A124: v in dom f ; ::_thesis: v in union (rng G) ex B being set st ( v in B & B in rng G ) proof percases ( f . v < R_EAL n or R_EAL n <= f . v ) ; supposeA125: f . v < R_EAL n ; ::_thesis: ex B being set st ( v in B & B in rng G ) 0 <= f . v by A2, SUPINF_2:51; then consider k being Nat such that A126: 1 <= k and A127: k <= (2 |^ n) * n and A128: (k - 1) / (2 |^ n) <= f . v and A129: f . v < k / (2 |^ n) by A125, Th4; v in great_eq_dom (f,(R_EAL ((k - 1) / (2 |^ n)))) by A124, A128, MESFUNC1:def_14; then A130: v in A /\ (great_eq_dom (f,(R_EAL ((k - 1) / (2 |^ n))))) by A23, A124, XBOOLE_0:def_4; v in less_dom (f,(R_EAL (k / (2 |^ n)))) by A124, A129, MESFUNC1:def_11; then A131: v in (A /\ (great_eq_dom (f,(R_EAL ((k - 1) / (2 |^ n)))))) /\ (less_dom (f,(R_EAL (k / (2 |^ n))))) by A130, XBOOLE_0:def_4; take G . k ; ::_thesis: ( v in G . k & G . k in rng G ) A132: k in NAT by ORDINAL1:def_12; N * n <= (N * n) + 1 by NAT_1:11; then k <= (N * n) + 1 by A127, XXREAL_0:2; then k in Seg ((N * n) + 1) by A126, A132; hence ( v in G . k & G . k in rng G ) by A67, A127, A131, FUNCT_1:3; ::_thesis: verum end; supposeA133: R_EAL n <= f . v ; ::_thesis: ex B being set st ( v in B & B in rng G ) set k = (N * n) + 1; take G . ((N * n) + 1) ; ::_thesis: ( v in G . ((N * n) + 1) & G . ((N * n) + 1) in rng G ) 1 <= (N * n) + 1 by NAT_1:11; then A134: (N * n) + 1 in Seg ((N * n) + 1) ; v in great_eq_dom (f,(R_EAL n)) by A124, A133, MESFUNC1:def_14; then v in A /\ (great_eq_dom (f,(R_EAL n))) by A23, A124, XBOOLE_0:def_4; hence ( v in G . ((N * n) + 1) & G . ((N * n) + 1) in rng G ) by A67, A134, FUNCT_1:3; ::_thesis: verum end; end; end; hence v in union (rng G) by TARSKI:def_4; ::_thesis: verum end; then A135: dom f c= union (rng G) by TARSKI:def_3; for v being set st v in union (rng G) holds v in dom f proof let v be set ; ::_thesis: ( v in union (rng G) implies v in dom f ) assume v in union (rng G) ; ::_thesis: v in dom f then consider B being set such that A136: v in B and A137: B in rng G by TARSKI:def_4; consider m being set such that A138: m in dom G and A139: B = G . m by A137, FUNCT_1:def_3; reconsider m = m as Element of NAT by A138; A140: m <= (N * n) + 1 by A67, A138, FINSEQ_1:1; now__::_thesis:_v_in_A percases ( m = (N * n) + 1 or m <> (N * n) + 1 ) ; suppose m = (N * n) + 1 ; ::_thesis: v in A then B = A /\ (great_eq_dom (f,(R_EAL n))) by A67, A138, A139; hence v in A by A136, XBOOLE_0:def_4; ::_thesis: verum end; suppose m <> (N * n) + 1 ; ::_thesis: v in A then m < (N * n) + 1 by A140, XXREAL_0:1; then m <= N * n by NAT_1:13; then B = (A /\ (great_eq_dom (f,(R_EAL ((m - 1) / (2 |^ n)))))) /\ (less_dom (f,(R_EAL (m / (2 |^ n))))) by A67, A138, A139; then v in A /\ (great_eq_dom (f,(R_EAL ((m - 1) / (2 |^ n))))) by A136, XBOOLE_0:def_4; hence v in A by XBOOLE_0:def_4; ::_thesis: verum end; end; end; hence v in dom f by A23; ::_thesis: verum end; then union (rng G) c= dom f by TARSKI:def_3; then union (rng G) = dom f by A135, XBOOLE_0:def_10; then dom (F . n) = union (rng G) by A22; hence F . n is_simple_func_in S by A63, A105, MESFUNC2:def_4; ::_thesis: verum end; A141: now__::_thesis:_for_x_being_Element_of_X_st_x_in_dom_f_holds_ (_F_#_x_is_convergent_&_lim_(F_#_x)_=_f_._x_) let x be Element of X; ::_thesis: ( x in dom f implies ( F # b1 is convergent & lim (F # b1) = f . b1 ) ) assume A142: x in dom f ; ::_thesis: ( F # b1 is convergent & lim (F # b1) = f . b1 ) percases ( |.(f . x).| = +infty or |.(f . x).| <> +infty ) ; supposeA143: |.(f . x).| = +infty ; ::_thesis: ( F # b1 is convergent & lim (F # b1) = f . b1 ) now__::_thesis:_not_-_(f_._x)_=_+infty assume - (f . x) = +infty ; ::_thesis: contradiction then f . x < 0 ; hence contradiction by A2, SUPINF_2:51; ::_thesis: verum end; then A144: f . x = +infty by A143, EXTREAL2:2; for g being real number st 0 < g holds ex n being Nat st for m being Nat st n <= m holds g <= (F # x) . m proof let g be real number ; ::_thesis: ( 0 < g implies ex n being Nat st for m being Nat st n <= m holds g <= (F # x) . m ) assume 0 < g ; ::_thesis: ex n being Nat st for m being Nat st n <= m holds g <= (F # x) . m then reconsider n = [/g\] as Nat by INT_1:53; A145: g <= n by INT_1:def_7; for m being Nat st n <= m holds g <= (F # x) . m proof let m be Nat; ::_thesis: ( n <= m implies g <= (F # x) . m ) assume n <= m ; ::_thesis: g <= (F # x) . m then A146: g <= m by A145, XXREAL_0:2; reconsider m = m as Element of NAT by ORDINAL1:def_12; R_EAL m <= f . x by A144, XXREAL_0:4; then (F . m) . x = R_EAL m by A21, A142; hence g <= (F # x) . m by A146, Def13; ::_thesis: verum end; hence ex n being Nat st for m being Nat st n <= m holds g <= (F # x) . m ; ::_thesis: verum end; then A147: F # x is convergent_to_+infty by Def9; then F # x is convergent by Def11; hence ( F # x is convergent & lim (F # x) = f . x ) by A144, A147, Def12; ::_thesis: verum end; suppose |.(f . x).| <> +infty ; ::_thesis: ( F # b1 is convergent & lim (F # b1) = f . b1 ) then reconsider g = f . x as Real by EXTREAL2:19, XXREAL_0:14; A148: for p being real number st 0 < p holds ex k being Nat st for j being Nat st j >= k holds |.(((F # x) . j) - (f . x)).| < p proof set N2 = [/g\] + 1; let p be real number ; ::_thesis: ( 0 < p implies ex k being Nat st for j being Nat st j >= k holds |.(((F # x) . j) - (f . x)).| < p ) A149: g <= [/g\] by INT_1:def_7; [/g\] < [/g\] + 1 by XREAL_1:29; then A150: g < [/g\] + 1 by A149, XXREAL_0:2; 0 <= g by A2, SUPINF_2:51; then reconsider N2 = [/g\] + 1 as Element of NAT by A149, INT_1:3; A151: for N being Nat st N >= N2 holds |.(((F # x) . N) - (f . x)).| < 1 / (2 |^ N) proof let N be Nat; ::_thesis: ( N >= N2 implies |.(((F # x) . N) - (f . x)).| < 1 / (2 |^ N) ) assume A152: N >= N2 ; ::_thesis: |.(((F # x) . N) - (f . x)).| < 1 / (2 |^ N) reconsider N = N as Element of NAT by ORDINAL1:def_12; A153: 0 <= f . x by A2, SUPINF_2:51; f . x < N by A150, A152, XXREAL_0:2; then consider m being Nat such that A154: 1 <= m and A155: m <= (2 |^ N) * N and A156: (m - 1) / (2 |^ N) <= f . x and A157: f . x < m / (2 |^ N) by A153, Th4; reconsider m = m as Element of NAT by ORDINAL1:def_12; A158: (F # x) . N = (F . N) . x by Def13 .= (m - 1) / (2 |^ N) by A21, A142, A154, A155, A156, A157 ; (R_EAL (m / (2 |^ N))) - (R_EAL ((m - 1) / (2 |^ N))) = (m / (2 |^ N)) - ((m - 1) / (2 |^ N)) by SUPINF_2:3 .= (m / (2 |^ N)) + (- ((m - 1) / (2 |^ N))) .= (m / (2 |^ N)) + ((- (m - 1)) / (2 |^ N)) by XCMPLX_1:187 .= (m + (- (m - 1))) / (2 |^ N) by XCMPLX_1:62 ; then A159: (f . x) - ((F # x) . N) < 1 / (2 |^ N) by A157, A158, XXREAL_3:43; - (((F # x) . N) - (f . x)) = (f . x) - ((F # x) . N) by XXREAL_3:26; then A160: |.(((F # x) . N) - (f . x)).| = |.((f . x) - ((F # x) . N)).| by EXTREAL2:18; 2 |^ N > 0 by PREPOWER:6; then A161: - (R_EAL (1 / (2 |^ N))) < 0 ; 0 <= (f . x) - ((F # x) . N) by A156, A158, XXREAL_3:40; hence |.(((F # x) . N) - (f . x)).| < 1 / (2 |^ N) by A160, A159, A161, EXTREAL2:11; ::_thesis: verum end; assume 0 < p ; ::_thesis: ex k being Nat st for j being Nat st j >= k holds |.(((F # x) . j) - (f . x)).| < p then consider N1 being Element of NAT such that A162: 1 / (2 |^ N1) <= p by PREPOWER:92; reconsider k = max (N2,N1) as Element of NAT ; A163: for k being Nat st k >= N1 holds 1 / (2 |^ k) <= p proof let k be Nat; ::_thesis: ( k >= N1 implies 1 / (2 |^ k) <= p ) assume k >= N1 ; ::_thesis: 1 / (2 |^ k) <= p then consider i being Nat such that A164: k = N1 + i by NAT_1:10; (2 |^ N1) * (2 |^ i) >= 2 |^ N1 by PREPOWER:11, XREAL_1:151; then A165: 2 |^ k >= 2 |^ N1 by A164, NEWTON:8; 2 |^ N1 > 0 by PREPOWER:11; then (2 |^ k) " <= (2 |^ N1) " by A165, XREAL_1:85; then 1 / (2 |^ k) <= (2 |^ N1) " by XCMPLX_1:215; then 1 / (2 |^ k) <= 1 / (2 |^ N1) by XCMPLX_1:215; hence 1 / (2 |^ k) <= p by A162, XXREAL_0:2; ::_thesis: verum end; now__::_thesis:_for_j_being_Nat_st_j_>=_k_holds_ |.(((F_#_x)_._j)_-_(f_._x)).|_<_p let j be Nat; ::_thesis: ( j >= k implies |.(((F # x) . j) - (f . x)).| < p ) assume A166: j >= k ; ::_thesis: |.(((F # x) . j) - (f . x)).| < p k >= N2 by XXREAL_0:25; then j >= N2 by A166, XXREAL_0:2; then A167: |.(((F # x) . j) - (f . x)).| < R_EAL (1 / (2 |^ j)) by A151; k >= N1 by XXREAL_0:25; then j >= N1 by A166, XXREAL_0:2; then 1 / (2 |^ j) <= p by A163; hence |.(((F # x) . j) - (f . x)).| < p by A167, XXREAL_0:2; ::_thesis: verum end; hence ex k being Nat st for j being Nat st j >= k holds |.(((F # x) . j) - (f . x)).| < p ; ::_thesis: verum end; A168: f . x = R_EAL g ; then A169: F # x is convergent_to_finite_number by A148, Def8; then F # x is convergent by Def11; hence ( F # x is convergent & lim (F # x) = f . x ) by A168, A148, A169, Def12; ::_thesis: verum end; end; end; for n being Nat holds F . n is nonnegative proof let n be Nat; ::_thesis: F . n is nonnegative reconsider n = n as Element of NAT by ORDINAL1:def_12; now__::_thesis:_for_x_being_set_st_x_in_dom_(F_._n)_holds_ 0_<=_(F_._n)_._x let x be set ; ::_thesis: ( x in dom (F . n) implies 0 <= (F . n) . b1 ) assume x in dom (F . n) ; ::_thesis: 0 <= (F . n) . b1 then A170: x in dom f by A22; percases ( n <= f . x or f . x < n ) ; suppose n <= f . x ; ::_thesis: 0 <= (F . n) . b1 hence 0 <= (F . n) . x by A21, A170; ::_thesis: verum end; supposeA171: f . x < n ; ::_thesis: 0 <= (F . n) . b1 0 <= f . x by A2, SUPINF_2:51; then consider k being Nat such that A172: 1 <= k and A173: k <= (2 |^ n) * n and A174: (k - 1) / (2 |^ n) <= f . x and A175: f . x < k / (2 |^ n) by A171, Th4; 0 <= k - 1 by A172, XREAL_1:48; hence 0 <= (F . n) . x by A21, A170, A172, A173, A174, A175; ::_thesis: verum end; end; end; hence F . n is nonnegative by SUPINF_2:52; ::_thesis: verum end; hence ex F being Functional_Sequence of X,ExtREAL st ( ( for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = dom f ) ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom f holds ( F # x is convergent & lim (F # x) = f . x ) ) ) by A22, A48, A25, A141; ::_thesis: verum end; begin definition let X be non empty set ; let S be SigmaField of X; let M be sigma_Measure of S; let f be PartFunc of X,ExtREAL; func integral' (M,f) -> Element of ExtREAL equals :Def14: :: MESFUNC5:def 14 integral (X,S,M,f) if dom f <> {} otherwise 0. ; correctness coherence ( ( dom f <> {} implies integral (X,S,M,f) is Element of ExtREAL ) & ( not dom f <> {} implies 0. is Element of ExtREAL ) ); consistency for b1 being Element of ExtREAL holds verum; ; end; :: deftheorem Def14 defines integral' MESFUNC5:def_14_:_ for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL holds ( ( dom f <> {} implies integral' (M,f) = integral (X,S,M,f) ) & ( not dom f <> {} implies integral' (M,f) = 0. ) ); theorem Th65: :: MESFUNC5:65 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & g is_simple_func_in S & f is nonnegative & g is nonnegative holds ( dom (f + g) = (dom f) /\ (dom g) & integral' (M,(f + g)) = (integral' (M,(f | (dom (f + g))))) + (integral' (M,(g | (dom (f + g))))) ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & g is_simple_func_in S & f is nonnegative & g is nonnegative holds ( dom (f + g) = (dom f) /\ (dom g) & integral' (M,(f + g)) = (integral' (M,(f | (dom (f + g))))) + (integral' (M,(g | (dom (f + g))))) ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & g is_simple_func_in S & f is nonnegative & g is nonnegative holds ( dom (f + g) = (dom f) /\ (dom g) & integral' (M,(f + g)) = (integral' (M,(f | (dom (f + g))))) + (integral' (M,(g | (dom (f + g))))) ) let M be sigma_Measure of S; ::_thesis: for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & g is_simple_func_in S & f is nonnegative & g is nonnegative holds ( dom (f + g) = (dom f) /\ (dom g) & integral' (M,(f + g)) = (integral' (M,(f | (dom (f + g))))) + (integral' (M,(g | (dom (f + g))))) ) let f, g be PartFunc of X,ExtREAL; ::_thesis: ( f is_simple_func_in S & g is_simple_func_in S & f is nonnegative & g is nonnegative implies ( dom (f + g) = (dom f) /\ (dom g) & integral' (M,(f + g)) = (integral' (M,(f | (dom (f + g))))) + (integral' (M,(g | (dom (f + g))))) ) ) assume that A1: f is_simple_func_in S and A2: g is_simple_func_in S and A3: f is nonnegative and A4: g is nonnegative ; ::_thesis: ( dom (f + g) = (dom f) /\ (dom g) & integral' (M,(f + g)) = (integral' (M,(f | (dom (f + g))))) + (integral' (M,(g | (dom (f + g))))) ) A5: g | (dom (f + g)) is nonnegative by A4, Th15; f | (dom (f + g)) is nonnegative by A3, Th15; then A6: for x being set holds ( ( x in dom (f | (dom (f + g))) implies 0 <= (f | (dom (f + g))) . x ) & ( x in dom (g | (dom (f + g))) implies 0 <= (g | (dom (f + g))) . x ) ) by A5, SUPINF_2:51; not -infty in rng g by A4, Def3; then A7: g " {-infty} = {} by FUNCT_1:72; not -infty in rng f by A3, Def3; then A8: f " {-infty} = {} by FUNCT_1:72; then A9: ((dom f) /\ (dom g)) \ (((f " {-infty}) /\ (g " {+infty})) \/ ((f " {+infty}) /\ (g " {-infty}))) = (dom f) /\ (dom g) by A7; hence A10: dom (f + g) = (dom f) /\ (dom g) by MESFUNC1:def_3; ::_thesis: integral' (M,(f + g)) = (integral' (M,(f | (dom (f + g))))) + (integral' (M,(g | (dom (f + g))))) A11: dom (f + g) is Element of S by A1, A2, Th37, Th38; then A12: f | (dom (f + g)) is_simple_func_in S by A1, Th34; A13: g | (dom (f + g)) is_simple_func_in S by A2, A11, Th34; dom (f | (dom (f + g))) = (dom f) /\ (dom (f + g)) by RELAT_1:61; then A14: dom (f | (dom (f + g))) = ((dom f) /\ (dom f)) /\ (dom g) by A10, XBOOLE_1:16; dom (g | (dom (f + g))) = (dom g) /\ (dom (f + g)) by RELAT_1:61; then A15: dom (g | (dom (f + g))) = ((dom g) /\ (dom g)) /\ (dom f) by A10, XBOOLE_1:16; percases ( dom (f + g) = {} or dom (f + g) <> {} ) ; supposeA16: dom (f + g) = {} ; ::_thesis: integral' (M,(f + g)) = (integral' (M,(f | (dom (f + g))))) + (integral' (M,(g | (dom (f + g))))) dom (g | (dom (f + g))) = (dom g) /\ (dom (f + g)) by RELAT_1:61; then A17: integral' (M,(g | (dom (f + g)))) = 0 by A16, Def14; dom (f | (dom (f + g))) = (dom f) /\ (dom (f + g)) by RELAT_1:61; then A18: integral' (M,(f | (dom (f + g)))) = 0 by A16, Def14; integral' (M,(f + g)) = 0 by A16, Def14; hence integral' (M,(f + g)) = (integral' (M,(f | (dom (f + g))))) + (integral' (M,(g | (dom (f + g))))) by A18, A17; ::_thesis: verum end; supposeA19: dom (f + g) <> {} ; ::_thesis: integral' (M,(f + g)) = (integral' (M,(f | (dom (f + g))))) + (integral' (M,(g | (dom (f + g))))) A20: (g | (dom (f + g))) " {-infty} = (dom (f + g)) /\ (g " {-infty}) by FUNCT_1:70 .= {} by A7 ; (f | (dom (f + g))) " {-infty} = (dom (f + g)) /\ (f " {-infty}) by FUNCT_1:70 .= {} by A8 ; then ((dom (f | (dom (f + g)))) /\ (dom (g | (dom (f + g))))) \ ((((f | (dom (f + g))) " {-infty}) /\ ((g | (dom (f + g))) " {+infty})) \/ (((f | (dom (f + g))) " {+infty}) /\ ((g | (dom (f + g))) " {-infty}))) = dom (f + g) by A9, A14, A15, A20, MESFUNC1:def_3; then A21: dom ((f | (dom (f + g))) + (g | (dom (f + g)))) = dom (f + g) by MESFUNC1:def_3; A22: for x being Element of X st x in dom ((f | (dom (f + g))) + (g | (dom (f + g)))) holds ((f | (dom (f + g))) + (g | (dom (f + g)))) . x = (f + g) . x proof let x be Element of X; ::_thesis: ( x in dom ((f | (dom (f + g))) + (g | (dom (f + g)))) implies ((f | (dom (f + g))) + (g | (dom (f + g)))) . x = (f + g) . x ) assume A23: x in dom ((f | (dom (f + g))) + (g | (dom (f + g)))) ; ::_thesis: ((f | (dom (f + g))) + (g | (dom (f + g)))) . x = (f + g) . x then ((f | (dom (f + g))) + (g | (dom (f + g)))) . x = ((f | (dom (f + g))) . x) + ((g | (dom (f + g))) . x) by MESFUNC1:def_3 .= (f . x) + ((g | (dom (f + g))) . x) by A21, A23, FUNCT_1:49 .= (f . x) + (g . x) by A21, A23, FUNCT_1:49 ; hence ((f | (dom (f + g))) + (g | (dom (f + g)))) . x = (f + g) . x by A21, A23, MESFUNC1:def_3; ::_thesis: verum end; integral (X,S,M,((f | (dom (f + g))) + (g | (dom (f + g))))) = (integral (X,S,M,(f | (dom (f + g))))) + (integral (X,S,M,(g | (dom (f + g))))) by A10, A12, A13, A6, A14, A15, A19, MESFUNC4:5; then A24: integral (X,S,M,(f + g)) = (integral (X,S,M,(f | (dom (f + g))))) + (integral (X,S,M,(g | (dom (f + g))))) by A21, A22, PARTFUN1:5; A25: integral (X,S,M,(g | (dom (f + g)))) = integral' (M,(g | (dom (f + g)))) by A10, A15, A19, Def14; integral (X,S,M,(f | (dom (f + g)))) = integral' (M,(f | (dom (f + g)))) by A10, A14, A19, Def14; hence integral' (M,(f + g)) = (integral' (M,(f | (dom (f + g))))) + (integral' (M,(g | (dom (f + g))))) by A19, A24, A25, Def14; ::_thesis: verum end; end; end; theorem Th66: :: MESFUNC5:66 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for c being Real st f is_simple_func_in S & f is nonnegative & 0 <= c holds integral' (M,(c (#) f)) = (R_EAL c) * (integral' (M,f)) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for c being Real st f is_simple_func_in S & f is nonnegative & 0 <= c holds integral' (M,(c (#) f)) = (R_EAL c) * (integral' (M,f)) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for c being Real st f is_simple_func_in S & f is nonnegative & 0 <= c holds integral' (M,(c (#) f)) = (R_EAL c) * (integral' (M,f)) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for c being Real st f is_simple_func_in S & f is nonnegative & 0 <= c holds integral' (M,(c (#) f)) = (R_EAL c) * (integral' (M,f)) let f be PartFunc of X,ExtREAL; ::_thesis: for c being Real st f is_simple_func_in S & f is nonnegative & 0 <= c holds integral' (M,(c (#) f)) = (R_EAL c) * (integral' (M,f)) let c be Real; ::_thesis: ( f is_simple_func_in S & f is nonnegative & 0 <= c implies integral' (M,(c (#) f)) = (R_EAL c) * (integral' (M,f)) ) assume that A1: f is_simple_func_in S and A2: f is nonnegative and A3: 0 <= c ; ::_thesis: integral' (M,(c (#) f)) = (R_EAL c) * (integral' (M,f)) A4: for x being set st x in dom f holds 0 <= f . x by A2, SUPINF_2:51; set g = c (#) f; A5: dom (c (#) f) = dom f by MESFUNC1:def_6; A6: for x being set st x in dom (c (#) f) holds (c (#) f) . x = (R_EAL c) * (f . x) by MESFUNC1:def_6; percases ( dom (c (#) f) = {} or dom (c (#) f) <> {} ) ; supposeA7: dom (c (#) f) = {} ; ::_thesis: integral' (M,(c (#) f)) = (R_EAL c) * (integral' (M,f)) then integral' (M,f) = 0 by A5, Def14; then c * (integral' (M,f)) = 0 ; hence integral' (M,(c (#) f)) = (R_EAL c) * (integral' (M,f)) by A7, Def14; ::_thesis: verum end; supposeA8: dom (c (#) f) <> {} ; ::_thesis: integral' (M,(c (#) f)) = (R_EAL c) * (integral' (M,f)) then integral' (M,f) = integral (X,S,M,f) by A5, Def14; then integral (X,S,M,(c (#) f)) = (R_EAL c) * (integral' (M,f)) by A1, A3, A5, A4, A6, A8, MESFUNC4:6, XXREAL_0:9; hence integral' (M,(c (#) f)) = (R_EAL c) * (integral' (M,f)) by A8, Def14; ::_thesis: verum end; end; end; theorem Th67: :: MESFUNC5:67 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A, B being Element of S st f is_simple_func_in S & f is nonnegative & A misses B holds integral' (M,(f | (A \/ B))) = (integral' (M,(f | A))) + (integral' (M,(f | B))) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A, B being Element of S st f is_simple_func_in S & f is nonnegative & A misses B holds integral' (M,(f | (A \/ B))) = (integral' (M,(f | A))) + (integral' (M,(f | B))) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A, B being Element of S st f is_simple_func_in S & f is nonnegative & A misses B holds integral' (M,(f | (A \/ B))) = (integral' (M,(f | A))) + (integral' (M,(f | B))) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for A, B being Element of S st f is_simple_func_in S & f is nonnegative & A misses B holds integral' (M,(f | (A \/ B))) = (integral' (M,(f | A))) + (integral' (M,(f | B))) let f be PartFunc of X,ExtREAL; ::_thesis: for A, B being Element of S st f is_simple_func_in S & f is nonnegative & A misses B holds integral' (M,(f | (A \/ B))) = (integral' (M,(f | A))) + (integral' (M,(f | B))) let A, B be Element of S; ::_thesis: ( f is_simple_func_in S & f is nonnegative & A misses B implies integral' (M,(f | (A \/ B))) = (integral' (M,(f | A))) + (integral' (M,(f | B))) ) assume that A1: f is_simple_func_in S and A2: f is nonnegative and A3: A misses B ; ::_thesis: integral' (M,(f | (A \/ B))) = (integral' (M,(f | A))) + (integral' (M,(f | B))) set g2 = f | B; set g1 = f | A; set g = f | (A \/ B); f | (A \/ B) is nonnegative by A2, Th15; then A4: for x being set st x in dom (f | (A \/ B)) holds 0 <= (f | (A \/ B)) . x by SUPINF_2:51; then consider G being Finite_Sep_Sequence of S, b being FinSequence of ExtREAL such that A5: G,b are_Re-presentation_of f | (A \/ B) and A6: b . 1 = 0 and A7: for n being Nat st 2 <= n & n in dom b holds ( 0 < b . n & b . n < +infty ) by A1, Th34, MESFUNC3:14; deffunc H1( Nat) -> Element of bool X = (G . $1) /\ A; consider G1 being FinSequence such that A8: ( len G1 = len G & ( for k being Nat st k in dom G1 holds G1 . k = H1(k) ) ) from FINSEQ_1:sch_2(); A9: dom G1 = Seg (len G) by A8, FINSEQ_1:def_3; A10: for k being Nat st k in dom G holds G1 . k = (G . k) /\ A proof let k be Nat; ::_thesis: ( k in dom G implies G1 . k = (G . k) /\ A ) assume k in dom G ; ::_thesis: G1 . k = (G . k) /\ A then k in Seg (len G) by FINSEQ_1:def_3; hence G1 . k = (G . k) /\ A by A8, A9; ::_thesis: verum end; deffunc H2( Nat) -> Element of bool X = (G . $1) /\ B; consider G2 being FinSequence such that A11: ( len G2 = len G & ( for k being Nat st k in dom G2 holds G2 . k = H2(k) ) ) from FINSEQ_1:sch_2(); A12: dom G2 = Seg (len G) by A11, FINSEQ_1:def_3; A13: for k being Nat st k in dom G holds G2 . k = (G . k) /\ B proof let k be Nat; ::_thesis: ( k in dom G implies G2 . k = (G . k) /\ B ) assume k in dom G ; ::_thesis: G2 . k = (G . k) /\ B then k in Seg (len G) by FINSEQ_1:def_3; hence G2 . k = (G . k) /\ B by A11, A12; ::_thesis: verum end; A14: dom G = Seg (len G2) by A11, FINSEQ_1:def_3 .= dom G2 by FINSEQ_1:def_3 ; then reconsider G2 = G2 as Finite_Sep_Sequence of S by A13, Th35; A15: dom ((f | (A \/ B)) | B) = (dom (f | (A \/ B))) /\ B by RELAT_1:61 .= ((dom f) /\ (A \/ B)) /\ B by RELAT_1:61 .= (dom f) /\ ((A \/ B) /\ B) by XBOOLE_1:16 .= (dom f) /\ B by XBOOLE_1:21 .= dom (f | B) by RELAT_1:61 ; for x being set st x in dom ((f | (A \/ B)) | B) holds ((f | (A \/ B)) | B) . x = (f | B) . x proof let x be set ; ::_thesis: ( x in dom ((f | (A \/ B)) | B) implies ((f | (A \/ B)) | B) . x = (f | B) . x ) assume A16: x in dom ((f | (A \/ B)) | B) ; ::_thesis: ((f | (A \/ B)) | B) . x = (f | B) . x then x in (dom (f | (A \/ B))) /\ B by RELAT_1:61; then A17: x in dom (f | (A \/ B)) by XBOOLE_0:def_4; ((f | (A \/ B)) | B) . x = (f | (A \/ B)) . x by A16, FUNCT_1:47 .= f . x by A17, FUNCT_1:47 ; hence ((f | (A \/ B)) | B) . x = (f | B) . x by A15, A16, FUNCT_1:47; ::_thesis: verum end; then (f | (A \/ B)) | B = f | B by A15, FUNCT_1:2; then A18: G2,b are_Re-presentation_of f | B by A5, A14, A13, Th36; A19: dom G = Seg (len G1) by A8, FINSEQ_1:def_3 .= dom G1 by FINSEQ_1:def_3 ; then reconsider G1 = G1 as Finite_Sep_Sequence of S by A10, Th35; A20: dom ((f | (A \/ B)) | A) = (dom (f | (A \/ B))) /\ A by RELAT_1:61 .= ((dom f) /\ (A \/ B)) /\ A by RELAT_1:61 .= (dom f) /\ ((A \/ B) /\ A) by XBOOLE_1:16 .= (dom f) /\ A by XBOOLE_1:21 .= dom (f | A) by RELAT_1:61 ; for x being set st x in dom ((f | (A \/ B)) | A) holds ((f | (A \/ B)) | A) . x = (f | A) . x proof let x be set ; ::_thesis: ( x in dom ((f | (A \/ B)) | A) implies ((f | (A \/ B)) | A) . x = (f | A) . x ) assume A21: x in dom ((f | (A \/ B)) | A) ; ::_thesis: ((f | (A \/ B)) | A) . x = (f | A) . x then x in (dom (f | (A \/ B))) /\ A by RELAT_1:61; then A22: x in dom (f | (A \/ B)) by XBOOLE_0:def_4; ((f | (A \/ B)) | A) . x = (f | (A \/ B)) . x by A21, FUNCT_1:47 .= f . x by A22, FUNCT_1:47 ; hence ((f | (A \/ B)) | A) . x = (f | A) . x by A20, A21, FUNCT_1:47; ::_thesis: verum end; then (f | (A \/ B)) | A = f | A by A20, FUNCT_1:2; then A23: G1,b are_Re-presentation_of f | A by A5, A19, A10, Th36; deffunc H3( Nat) -> Element of ExtREAL = (b . $1) * ((M * G) . $1); consider y being FinSequence of ExtREAL such that A24: ( len y = len G & ( for j being Nat st j in dom y holds y . j = H3(j) ) ) from FINSEQ_2:sch_1(); A25: dom y = Seg (len y) by FINSEQ_1:def_3 .= dom G by A24, FINSEQ_1:def_3 ; percases ( dom (f | (A \/ B)) = {} or dom (f | (A \/ B)) <> {} ) ; supposeA26: dom (f | (A \/ B)) = {} ; ::_thesis: integral' (M,(f | (A \/ B))) = (integral' (M,(f | A))) + (integral' (M,(f | B))) (dom f) /\ B c= (dom f) /\ (A \/ B) by XBOOLE_1:7, XBOOLE_1:26; then dom (f | B) c= (dom f) /\ (A \/ B) by RELAT_1:61; then dom (f | B) c= dom (f | (A \/ B)) by RELAT_1:61; then dom (f | B) = {} by A26; then A27: integral' (M,(f | B)) = 0 by Def14; (dom f) /\ A c= (dom f) /\ (A \/ B) by XBOOLE_1:7, XBOOLE_1:26; then dom (f | A) c= (dom f) /\ (A \/ B) by RELAT_1:61; then dom (f | A) c= dom (f | (A \/ B)) by RELAT_1:61; then dom (f | A) = {} by A26; then A28: integral' (M,(f | A)) = 0 by Def14; integral' (M,(f | (A \/ B))) = 0 by A26, Def14; hence integral' (M,(f | (A \/ B))) = (integral' (M,(f | A))) + (integral' (M,(f | B))) by A28, A27; ::_thesis: verum end; supposeA29: dom (f | (A \/ B)) <> {} ; ::_thesis: integral' (M,(f | (A \/ B))) = (integral' (M,(f | A))) + (integral' (M,(f | B))) then integral (X,S,M,(f | (A \/ B))) = Sum y by A1, A4, A5, A24, A25, Th34, MESFUNC4:3; then A30: integral' (M,(f | (A \/ B))) = Sum y by A29, Def14; percases ( dom (f | A) = {} or dom (f | A) <> {} ) ; supposeA31: dom (f | A) = {} ; ::_thesis: integral' (M,(f | (A \/ B))) = (integral' (M,(f | A))) + (integral' (M,(f | B))) A32: dom (f | (A \/ B)) = (dom f) /\ (A \/ B) by RELAT_1:61 .= ((dom f) /\ A) \/ ((dom f) /\ B) by XBOOLE_1:23 .= (dom (f | A)) \/ ((dom f) /\ B) by RELAT_1:61 .= dom (f | B) by A31, RELAT_1:61 ; now__::_thesis:_for_x_being_set_st_x_in_dom_(f_|_(A_\/_B))_holds_ (f_|_(A_\/_B))_._x_=_(f_|_B)_._x let x be set ; ::_thesis: ( x in dom (f | (A \/ B)) implies (f | (A \/ B)) . x = (f | B) . x ) assume A33: x in dom (f | (A \/ B)) ; ::_thesis: (f | (A \/ B)) . x = (f | B) . x then (f | (A \/ B)) . x = f . x by FUNCT_1:47; hence (f | (A \/ B)) . x = (f | B) . x by A32, A33, FUNCT_1:47; ::_thesis: verum end; then A34: f | (A \/ B) = f | B by A32, FUNCT_1:2; integral' (M,(f | A)) = 0 by A31, Def14; hence integral' (M,(f | (A \/ B))) = (integral' (M,(f | A))) + (integral' (M,(f | B))) by A34, XXREAL_3:4; ::_thesis: verum end; supposeA35: dom (f | A) <> {} ; ::_thesis: integral' (M,(f | (A \/ B))) = (integral' (M,(f | A))) + (integral' (M,(f | B))) percases ( dom (f | B) = {} or dom (f | B) <> {} ) ; supposeA36: dom (f | B) = {} ; ::_thesis: integral' (M,(f | (A \/ B))) = (integral' (M,(f | A))) + (integral' (M,(f | B))) A37: dom (f | (A \/ B)) = (dom f) /\ (A \/ B) by RELAT_1:61 .= ((dom f) /\ B) \/ ((dom f) /\ A) by XBOOLE_1:23 .= (dom (f | B)) \/ ((dom f) /\ A) by RELAT_1:61 .= dom (f | A) by A36, RELAT_1:61 ; now__::_thesis:_for_x_being_set_st_x_in_dom_(f_|_(A_\/_B))_holds_ (f_|_(A_\/_B))_._x_=_(f_|_A)_._x let x be set ; ::_thesis: ( x in dom (f | (A \/ B)) implies (f | (A \/ B)) . x = (f | A) . x ) assume A38: x in dom (f | (A \/ B)) ; ::_thesis: (f | (A \/ B)) . x = (f | A) . x then (f | (A \/ B)) . x = f . x by FUNCT_1:47; hence (f | (A \/ B)) . x = (f | A) . x by A37, A38, FUNCT_1:47; ::_thesis: verum end; then A39: f | (A \/ B) = f | A by A37, FUNCT_1:2; integral' (M,(f | B)) = 0 by A36, Def14; hence integral' (M,(f | (A \/ B))) = (integral' (M,(f | A))) + (integral' (M,(f | B))) by A39, XXREAL_3:4; ::_thesis: verum end; supposeA40: dom (f | B) <> {} ; ::_thesis: integral' (M,(f | (A \/ B))) = (integral' (M,(f | A))) + (integral' (M,(f | B))) f | B is nonnegative by A2, Th15; then A41: for x being set st x in dom (f | B) holds 0 <= (f | B) . x by SUPINF_2:51; deffunc H4( Nat) -> Element of ExtREAL = (b . $1) * ((M * G2) . $1); consider y2 being FinSequence of ExtREAL such that A42: ( len y2 = len G2 & ( for j being Nat st j in dom y2 holds y2 . j = H4(j) ) ) from FINSEQ_2:sch_1(); A43: for k being Nat st k in dom y2 holds 0 <= y2 . k proof let k be Nat; ::_thesis: ( k in dom y2 implies 0 <= y2 . k ) assume A44: k in dom y2 ; ::_thesis: 0 <= y2 . k then k in Seg (len y2) by FINSEQ_1:def_3; then A45: 1 <= k by FINSEQ_1:1; A46: dom b = dom G by A5, MESFUNC3:def_1 .= Seg (len y2) by A11, A42, FINSEQ_1:def_3 .= dom y2 by FINSEQ_1:def_3 ; A47: now__::_thesis:_0_<=_b_._k percases ( k = 1 or k <> 1 ) ; suppose k = 1 ; ::_thesis: 0 <= b . k hence 0 <= b . k by A6; ::_thesis: verum end; suppose k <> 1 ; ::_thesis: 0 <= b . k then 1 < k by A45, XXREAL_0:1; then 1 + 1 <= k by NAT_1:13; hence 0 <= b . k by A7, A44, A46; ::_thesis: verum end; end; end; k in Seg (len G2) by A42, A44, FINSEQ_1:def_3; then A48: k in dom G2 by FINSEQ_1:def_3; then A49: (M * G2) . k = M . (G2 . k) by FUNCT_1:13; G2 . k in rng G2 by A48, FUNCT_1:3; then reconsider G2k = G2 . k as Element of S ; A50: 0 <= M . G2k by SUPINF_2:39; y2 . k = (b . k) * ((M * G2) . k) by A42, A44; hence 0 <= y2 . k by A47, A49, A50; ::_thesis: verum end; A51: for x being set st x in dom y2 holds not y2 . x in {-infty} proof let x be set ; ::_thesis: ( x in dom y2 implies not y2 . x in {-infty} ) assume A52: x in dom y2 ; ::_thesis: not y2 . x in {-infty} then reconsider x = x as Element of NAT ; 0 <= y2 . x by A43, A52; hence not y2 . x in {-infty} by TARSKI:def_1; ::_thesis: verum end; for x being set holds not x in y2 " {-infty} proof let x be set ; ::_thesis: not x in y2 " {-infty} ( not x in dom y2 or not y2 . x in {-infty} ) by A51; hence not x in y2 " {-infty} by FUNCT_1:def_7; ::_thesis: verum end; then A53: y2 " {-infty} = {} by XBOOLE_0:def_1; dom y2 = Seg (len G2) by A42, FINSEQ_1:def_3 .= dom G2 by FINSEQ_1:def_3 ; then integral (X,S,M,(f | B)) = Sum y2 by A1, A18, A40, A41, A42, Th34, MESFUNC4:3; then A54: integral' (M,(f | B)) = Sum y2 by A40, Def14; f | A is nonnegative by A2, Th15; then A55: for x being set st x in dom (f | A) holds 0 <= (f | A) . x by SUPINF_2:51; deffunc H5( Nat) -> Element of ExtREAL = (b . $1) * ((M * G1) . $1); consider y1 being FinSequence of ExtREAL such that A56: ( len y1 = len G1 & ( for j being Nat st j in dom y1 holds y1 . j = H5(j) ) ) from FINSEQ_2:sch_1(); A57: dom y = (Seg (len G)) /\ (Seg (len G)) by A25, FINSEQ_1:def_3 .= (dom y1) /\ (Seg (len G2)) by A8, A11, A56, FINSEQ_1:def_3 .= (dom y1) /\ (dom y2) by A42, FINSEQ_1:def_3 ; A58: for n being Element of NAT st n in dom y holds y . n = (y1 . n) + (y2 . n) proof let n be Element of NAT ; ::_thesis: ( n in dom y implies y . n = (y1 . n) + (y2 . n) ) assume A59: n in dom y ; ::_thesis: y . n = (y1 . n) + (y2 . n) then n in Seg (len G) by A24, FINSEQ_1:def_3; then A60: n in dom G by FINSEQ_1:def_3; then A61: G2 . n = (G . n) /\ B by A13; now__::_thesis:_for_v_being_set_st_v_in_G_._n_holds_ v_in_A_\/_B let v be set ; ::_thesis: ( v in G . n implies v in A \/ B ) assume A62: v in G . n ; ::_thesis: v in A \/ B G . n in rng G by A60, FUNCT_1:3; then v in union (rng G) by A62, TARSKI:def_4; then v in dom (f | (A \/ B)) by A5, MESFUNC3:def_1; then v in (dom f) /\ (A \/ B) by RELAT_1:61; hence v in A \/ B by XBOOLE_0:def_4; ::_thesis: verum end; then G . n c= A \/ B by TARSKI:def_3; then A63: G . n = (G . n) /\ (A \/ B) by XBOOLE_1:28 .= ((G . n) /\ A) \/ ((G . n) /\ B) by XBOOLE_1:23 .= (G1 . n) \/ (G2 . n) by A10, A60, A61 ; A64: n in dom y2 by A57, A59, XBOOLE_0:def_4; then n in Seg (len G2) by A42, FINSEQ_1:def_3; then A65: n in dom G2 by FINSEQ_1:def_3; then G2 . n in rng G2 by FUNCT_1:3; then reconsider G2n = G2 . n as Element of S ; 0 <= M . G2n by MEASURE1:def_2; then A66: ( 0 = (M * G2) . n or 0 < (M * G2) . n ) by A65, FUNCT_1:13; A67: now__::_thesis:_not_G1_._n_meets_G2_._n assume G1 . n meets G2 . n ; ::_thesis: contradiction then consider v being set such that A68: v in G1 . n and A69: v in G2 . n by XBOOLE_0:3; v in (G . n) /\ B by A13, A60, A69; then A70: v in B by XBOOLE_0:def_4; v in (G . n) /\ A by A10, A60, A68; then v in A by XBOOLE_0:def_4; hence contradiction by A3, A70, XBOOLE_0:3; ::_thesis: verum end; A71: n in dom y1 by A57, A59, XBOOLE_0:def_4; then n in Seg (len G1) by A56, FINSEQ_1:def_3; then A72: n in dom G1 by FINSEQ_1:def_3; then G1 . n in rng G1 by FUNCT_1:3; then reconsider G1n = G1 . n as Element of S ; 0 <= M . G1n by MEASURE1:def_2; then A73: ( 0 = (M * G1) . n or 0 < (M * G1) . n ) by A72, FUNCT_1:13; (M * G) . n = M . (G . n) by A60, FUNCT_1:13 .= (M . G1n) + (M . G2n) by A63, A67, MEASURE1:30 .= ((M * G1) . n) + (M . (G2 . n)) by A72, FUNCT_1:13 .= ((M * G1) . n) + ((M * G2) . n) by A65, FUNCT_1:13 ; then (b . n) * ((M * G) . n) = ((b . n) * ((M * G1) . n)) + ((b . n) * ((M * G2) . n)) by A73, A66, XXREAL_3:96; then y . n = ((b . n) * ((M * G1) . n)) + ((b . n) * ((M * G2) . n)) by A24, A59; then y . n = (y1 . n) + ((b . n) * ((M * G2) . n)) by A56, A71; hence y . n = (y1 . n) + (y2 . n) by A42, A64; ::_thesis: verum end; A74: for k being Nat st k in dom y1 holds 0 <= y1 . k proof let k be Nat; ::_thesis: ( k in dom y1 implies 0 <= y1 . k ) assume A75: k in dom y1 ; ::_thesis: 0 <= y1 . k then k in Seg (len y1) by FINSEQ_1:def_3; then A76: 1 <= k by FINSEQ_1:1; A77: dom b = dom G by A5, MESFUNC3:def_1 .= Seg (len y1) by A8, A56, FINSEQ_1:def_3 .= dom y1 by FINSEQ_1:def_3 ; A78: now__::_thesis:_0_<=_b_._k percases ( k = 1 or k <> 1 ) ; suppose k = 1 ; ::_thesis: 0 <= b . k hence 0 <= b . k by A6; ::_thesis: verum end; suppose k <> 1 ; ::_thesis: 0 <= b . k then 1 < k by A76, XXREAL_0:1; then 1 + 1 <= k by NAT_1:13; hence 0 <= b . k by A7, A75, A77; ::_thesis: verum end; end; end; k in Seg (len G1) by A56, A75, FINSEQ_1:def_3; then A79: k in dom G1 by FINSEQ_1:def_3; then A80: (M * G1) . k = M . (G1 . k) by FUNCT_1:13; G1 . k in rng G1 by A79, FUNCT_1:3; then reconsider G1k = G1 . k as Element of S ; A81: 0 <= M . G1k by SUPINF_2:39; y1 . k = (b . k) * ((M * G1) . k) by A56, A75; hence 0 <= y1 . k by A78, A80, A81; ::_thesis: verum end; A82: for x being set st x in dom y1 holds not y1 . x in {-infty} proof let x be set ; ::_thesis: ( x in dom y1 implies not y1 . x in {-infty} ) assume A83: x in dom y1 ; ::_thesis: not y1 . x in {-infty} then reconsider x = x as Element of NAT ; 0 <= y1 . x by A74, A83; hence not y1 . x in {-infty} by TARSKI:def_1; ::_thesis: verum end; for x being set holds not x in y1 " {-infty} proof let x be set ; ::_thesis: not x in y1 " {-infty} ( not x in dom y1 or not y1 . x in {-infty} ) by A82; hence not x in y1 " {-infty} by FUNCT_1:def_7; ::_thesis: verum end; then y1 " {-infty} = {} by XBOOLE_0:def_1; then dom y = ((dom y1) /\ (dom y2)) \ (((y1 " {-infty}) /\ (y2 " {+infty})) \/ ((y1 " {+infty}) /\ (y2 " {-infty}))) by A53, A57; then A84: y = y1 + y2 by A58, MESFUNC1:def_3; dom y1 = Seg (len G1) by A56, FINSEQ_1:def_3 .= dom G1 by FINSEQ_1:def_3 ; then integral (X,S,M,(f | A)) = Sum y1 by A1, A23, A35, A55, A56, Th34, MESFUNC4:3; then A85: integral' (M,(f | A)) = Sum y1 by A35, Def14; dom y1 = Seg (len y2) by A8, A11, A56, A42, FINSEQ_1:def_3 .= dom y2 by FINSEQ_1:def_3 ; hence integral' (M,(f | (A \/ B))) = (integral' (M,(f | A))) + (integral' (M,(f | B))) by A30, A85, A54, A74, A43, A84, MESFUNC4:1; ::_thesis: verum end; end; end; end; end; end; end; theorem Th68: :: MESFUNC5:68 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_simple_func_in S & f is nonnegative holds 0 <= integral' (M,f) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_simple_func_in S & f is nonnegative holds 0 <= integral' (M,f) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_simple_func_in S & f is nonnegative holds 0 <= integral' (M,f) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL st f is_simple_func_in S & f is nonnegative holds 0 <= integral' (M,f) let f be PartFunc of X,ExtREAL; ::_thesis: ( f is_simple_func_in S & f is nonnegative implies 0 <= integral' (M,f) ) assume that A1: f is_simple_func_in S and A2: f is nonnegative ; ::_thesis: 0 <= integral' (M,f) A3: for x being set st x in dom f holds 0 <= f . x by A2, SUPINF_2:51; percases ( dom f = {} or dom f <> {} ) ; suppose dom f = {} ; ::_thesis: 0 <= integral' (M,f) hence 0 <= integral' (M,f) by Def14; ::_thesis: verum end; supposeA4: dom f <> {} ; ::_thesis: 0 <= integral' (M,f) then consider F being Finite_Sep_Sequence of S, a, x being FinSequence of ExtREAL such that A5: F,a are_Re-presentation_of f and A6: dom x = dom F and A7: for n being Nat st n in dom x holds x . n = (a . n) * ((M * F) . n) and A8: integral (X,S,M,f) = Sum x by A1, A3, MESFUNC4:4; A9: for n being Nat st n in dom x holds 0 <= x . n proof let n be Nat; ::_thesis: ( n in dom x implies 0 <= x . n ) assume A10: n in dom x ; ::_thesis: 0 <= x . n percases ( F . n = {} or F . n <> {} ) ; suppose F . n = {} ; ::_thesis: 0 <= x . n then M . (F . n) = 0 by VALUED_0:def_19; then (M * F) . n = 0 by A6, A10, FUNCT_1:13; then (a . n) * ((M * F) . n) = 0 ; hence 0 <= x . n by A7, A10; ::_thesis: verum end; suppose F . n <> {} ; ::_thesis: 0 <= x . n then consider v being set such that A11: v in F . n by XBOOLE_0:def_1; F . n in rng F by A6, A10, FUNCT_1:3; then reconsider Fn = F . n as Element of S ; 0 <= M . Fn by MEASURE1:def_2; then A12: 0 <= (M * F) . n by A6, A10, FUNCT_1:13; f . v = a . n by A5, A6, A10, A11, MESFUNC3:def_1; then 0 <= a . n by A2, SUPINF_2:51; then 0 <= (a . n) * ((M * F) . n) by A12; hence 0 <= x . n by A7, A10; ::_thesis: verum end; end; end; integral' (M,f) = integral (X,S,M,f) by A4, Def14; hence 0 <= integral' (M,f) by A8, A9, Th53; ::_thesis: verum end; end; end; Lm9: for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & dom f <> {} & f is nonnegative & g is_simple_func_in S & dom g = dom f & g is nonnegative & ( for x being set st x in dom f holds g . x <= f . x ) holds ( f - g is_simple_func_in S & dom (f - g) <> {} & f - g is nonnegative & integral (X,S,M,f) = (integral (X,S,M,(f - g))) + (integral (X,S,M,g)) ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & dom f <> {} & f is nonnegative & g is_simple_func_in S & dom g = dom f & g is nonnegative & ( for x being set st x in dom f holds g . x <= f . x ) holds ( f - g is_simple_func_in S & dom (f - g) <> {} & f - g is nonnegative & integral (X,S,M,f) = (integral (X,S,M,(f - g))) + (integral (X,S,M,g)) ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & dom f <> {} & f is nonnegative & g is_simple_func_in S & dom g = dom f & g is nonnegative & ( for x being set st x in dom f holds g . x <= f . x ) holds ( f - g is_simple_func_in S & dom (f - g) <> {} & f - g is nonnegative & integral (X,S,M,f) = (integral (X,S,M,(f - g))) + (integral (X,S,M,g)) ) let M be sigma_Measure of S; ::_thesis: for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & dom f <> {} & f is nonnegative & g is_simple_func_in S & dom g = dom f & g is nonnegative & ( for x being set st x in dom f holds g . x <= f . x ) holds ( f - g is_simple_func_in S & dom (f - g) <> {} & f - g is nonnegative & integral (X,S,M,f) = (integral (X,S,M,(f - g))) + (integral (X,S,M,g)) ) let f, g be PartFunc of X,ExtREAL; ::_thesis: ( f is_simple_func_in S & dom f <> {} & f is nonnegative & g is_simple_func_in S & dom g = dom f & g is nonnegative & ( for x being set st x in dom f holds g . x <= f . x ) implies ( f - g is_simple_func_in S & dom (f - g) <> {} & f - g is nonnegative & integral (X,S,M,f) = (integral (X,S,M,(f - g))) + (integral (X,S,M,g)) ) ) assume that A1: f is_simple_func_in S and A2: dom f <> {} and A3: f is nonnegative and A4: g is_simple_func_in S and A5: dom g = dom f and A6: g is nonnegative and A7: for x being set st x in dom f holds g . x <= f . x ; ::_thesis: ( f - g is_simple_func_in S & dom (f - g) <> {} & f - g is nonnegative & integral (X,S,M,f) = (integral (X,S,M,(f - g))) + (integral (X,S,M,g)) ) A8: for x being set st x in dom g holds 0 <= g . x by A6, SUPINF_2:51; then consider G being Finite_Sep_Sequence of S, b, y being FinSequence of ExtREAL such that A9: G,b are_Re-presentation_of g and dom y = dom G and for n being Nat st n in dom y holds y . n = (b . n) * ((M * G) . n) and integral (X,S,M,g) = Sum y by A2, A4, A5, MESFUNC4:4; g is V60() by A4, MESFUNC2:def_4; then A10: dom (f - g) = (dom f) /\ (dom g) by MESFUNC2:2; A11: for x being set st x in dom f holds 0 <= f . x by A3, SUPINF_2:51; then consider F being Finite_Sep_Sequence of S, a, x being FinSequence of ExtREAL such that A12: F,a are_Re-presentation_of f and dom x = dom F and for n being Nat st n in dom x holds x . n = (a . n) * ((M * F) . n) and integral (X,S,M,f) = Sum x by A1, A2, MESFUNC4:4; set la = len a; A13: dom F = dom a by A12, MESFUNC3:def_1; set lb = len b; deffunc H1( Nat) -> set = (F . ((($1 -' 1) div (len b)) + 1)) /\ (G . ((($1 -' 1) mod (len b)) + 1)); consider FG being FinSequence such that A14: len FG = (len a) * (len b) and A15: for k being Nat st k in dom FG holds FG . k = H1(k) from FINSEQ_1:sch_2(); A16: dom FG = Seg ((len a) * (len b)) by A14, FINSEQ_1:def_3; A17: dom G = dom b by A9, MESFUNC3:def_1; FG is FinSequence of S proof let b be set ; :: according to TARSKI:def_3,FINSEQ_1:def_4 ::_thesis: ( not b in rng FG or b in S ) A18: now__::_thesis:_for_k_being_Element_of_NAT_st_k_in_dom_FG_holds_ FG_._k_in_S let k be Element of NAT ; ::_thesis: ( k in dom FG implies FG . k in S ) set i = ((k -' 1) div (len b)) + 1; set j = ((k -' 1) mod (len b)) + 1; A19: len b divides (len a) * (len b) by NAT_D:def_3; assume A20: k in dom FG ; ::_thesis: FG . k in S then A21: k in Seg ((len a) * (len b)) by A14, FINSEQ_1:def_3; then A22: k <= (len a) * (len b) by FINSEQ_1:1; then k -' 1 <= ((len a) * (len b)) -' 1 by NAT_D:42; then A23: (k -' 1) div (len b) <= (((len a) * (len b)) -' 1) div (len b) by NAT_2:24; 1 <= k by A21, FINSEQ_1:1; then A24: 1 <= (len a) * (len b) by A22, XXREAL_0:2; A25: len b <> 0 by A21; then (k -' 1) mod (len b) < len b by NAT_D:1; then A26: ((k -' 1) mod (len b)) + 1 <= len b by NAT_1:13; len b >= 0 + 1 by A25, NAT_1:13; then (((len a) * (len b)) -' 1) div (len b) = (((len a) * (len b)) div (len b)) - 1 by A19, A24, NAT_2:15; then ((k -' 1) div (len b)) + 1 <= ((len a) * (len b)) div (len b) by A23, XREAL_1:19; then A27: ((k -' 1) div (len b)) + 1 <= len a by A25, NAT_D:18; ((k -' 1) div (len b)) + 1 >= 0 + 1 by XREAL_1:6; then ((k -' 1) div (len b)) + 1 in Seg (len a) by A27; then ((k -' 1) div (len b)) + 1 in dom F by A13, FINSEQ_1:def_3; then A28: F . (((k -' 1) div (len b)) + 1) in rng F by FUNCT_1:3; ((k -' 1) mod (len b)) + 1 >= 0 + 1 by XREAL_1:6; then ((k -' 1) mod (len b)) + 1 in Seg (len b) by A26; then ((k -' 1) mod (len b)) + 1 in dom G by A17, FINSEQ_1:def_3; then A29: G . (((k -' 1) mod (len b)) + 1) in rng G by FUNCT_1:3; FG . k = (F . (((k -' 1) div (len b)) + 1)) /\ (G . (((k -' 1) mod (len b)) + 1)) by A15, A20; hence FG . k in S by A28, A29, MEASURE1:34; ::_thesis: verum end; assume b in rng FG ; ::_thesis: b in S then ex a being set st ( a in dom FG & b = FG . a ) by FUNCT_1:def_3; hence b in S by A18; ::_thesis: verum end; then reconsider FG = FG as FinSequence of S ; for k, l being Nat st k in dom FG & l in dom FG & k <> l holds FG . k misses FG . l proof let k, l be Nat; ::_thesis: ( k in dom FG & l in dom FG & k <> l implies FG . k misses FG . l ) assume that A30: k in dom FG and A31: l in dom FG and A32: k <> l ; ::_thesis: FG . k misses FG . l A33: k in Seg ((len a) * (len b)) by A14, A30, FINSEQ_1:def_3; then A34: 1 <= k by FINSEQ_1:1; set m = ((l -' 1) mod (len b)) + 1; set n = ((l -' 1) div (len b)) + 1; set j = ((k -' 1) mod (len b)) + 1; set i = ((k -' 1) div (len b)) + 1; A35: len b divides (len a) * (len b) by NAT_D:def_3; FG . k = (F . (((k -' 1) div (len b)) + 1)) /\ (G . (((k -' 1) mod (len b)) + 1)) by A15, A30; then A36: (FG . k) /\ (FG . l) = ((F . (((k -' 1) div (len b)) + 1)) /\ (G . (((k -' 1) mod (len b)) + 1))) /\ ((F . (((l -' 1) div (len b)) + 1)) /\ (G . (((l -' 1) mod (len b)) + 1))) by A15, A31 .= (F . (((k -' 1) div (len b)) + 1)) /\ ((G . (((k -' 1) mod (len b)) + 1)) /\ ((F . (((l -' 1) div (len b)) + 1)) /\ (G . (((l -' 1) mod (len b)) + 1)))) by XBOOLE_1:16 .= (F . (((k -' 1) div (len b)) + 1)) /\ ((F . (((l -' 1) div (len b)) + 1)) /\ ((G . (((k -' 1) mod (len b)) + 1)) /\ (G . (((l -' 1) mod (len b)) + 1)))) by XBOOLE_1:16 .= ((F . (((k -' 1) div (len b)) + 1)) /\ (F . (((l -' 1) div (len b)) + 1))) /\ ((G . (((k -' 1) mod (len b)) + 1)) /\ (G . (((l -' 1) mod (len b)) + 1))) by XBOOLE_1:16 ; A37: k <= (len a) * (len b) by A33, FINSEQ_1:1; then A38: 1 <= (len a) * (len b) by A34, XXREAL_0:2; A39: len b <> 0 by A33; then len b >= 0 + 1 by NAT_1:13; then A40: (((len a) * (len b)) -' 1) div (len b) = (((len a) * (len b)) div (len b)) - 1 by A35, A38, NAT_2:15; k -' 1 <= ((len a) * (len b)) -' 1 by A37, NAT_D:42; then (k -' 1) div (len b) <= (((len a) * (len b)) div (len b)) - 1 by A40, NAT_2:24; then ((k -' 1) div (len b)) + 1 <= ((len a) * (len b)) div (len b) by XREAL_1:19; then A41: ((k -' 1) div (len b)) + 1 <= len a by A39, NAT_D:18; ((k -' 1) div (len b)) + 1 >= 0 + 1 by XREAL_1:6; then ((k -' 1) div (len b)) + 1 in Seg (len a) by A41; then A42: ((k -' 1) div (len b)) + 1 in dom F by A13, FINSEQ_1:def_3; (l -' 1) mod (len b) < len b by A39, NAT_D:1; then A43: ((l -' 1) mod (len b)) + 1 <= len b by NAT_1:13; ((l -' 1) mod (len b)) + 1 >= 0 + 1 by XREAL_1:6; then ((l -' 1) mod (len b)) + 1 in Seg (len b) by A43; then A44: ((l -' 1) mod (len b)) + 1 in dom G by A17, FINSEQ_1:def_3; (k -' 1) mod (len b) < len b by A39, NAT_D:1; then A45: ((k -' 1) mod (len b)) + 1 <= len b by NAT_1:13; ((k -' 1) mod (len b)) + 1 >= 0 + 1 by XREAL_1:6; then ((k -' 1) mod (len b)) + 1 in Seg (len b) by A45; then A46: ((k -' 1) mod (len b)) + 1 in dom G by A17, FINSEQ_1:def_3; A47: l in Seg ((len a) * (len b)) by A14, A31, FINSEQ_1:def_3; then A48: 1 <= l by FINSEQ_1:1; A49: now__::_thesis:_(_((k_-'_1)_div_(len_b))_+_1_=_((l_-'_1)_div_(len_b))_+_1_implies_not_((k_-'_1)_mod_(len_b))_+_1_=_((l_-'_1)_mod_(len_b))_+_1_) (l -' 1) + 1 = ((((((l -' 1) div (len b)) + 1) - 1) * (len b)) + ((((l -' 1) mod (len b)) + 1) - 1)) + 1 by A39, NAT_D:2; then A50: (l - 1) + 1 = (((((l -' 1) div (len b)) + 1) - 1) * (len b)) + (((l -' 1) mod (len b)) + 1) by A48, XREAL_1:233; assume that A51: ((k -' 1) div (len b)) + 1 = ((l -' 1) div (len b)) + 1 and A52: ((k -' 1) mod (len b)) + 1 = ((l -' 1) mod (len b)) + 1 ; ::_thesis: contradiction (k -' 1) + 1 = ((((((k -' 1) div (len b)) + 1) - 1) * (len b)) + ((((k -' 1) mod (len b)) + 1) - 1)) + 1 by A39, NAT_D:2; then (k - 1) + 1 = (((((k -' 1) div (len b)) + 1) - 1) * (len b)) + (((k -' 1) mod (len b)) + 1) by A34, XREAL_1:233; hence contradiction by A32, A51, A52, A50; ::_thesis: verum end; l <= (len a) * (len b) by A47, FINSEQ_1:1; then l -' 1 <= ((len a) * (len b)) -' 1 by NAT_D:42; then (l -' 1) div (len b) <= (((len a) * (len b)) div (len b)) - 1 by A40, NAT_2:24; then ((l -' 1) div (len b)) + 1 <= ((len a) * (len b)) div (len b) by XREAL_1:19; then A53: ((l -' 1) div (len b)) + 1 <= len a by A39, NAT_D:18; ((l -' 1) div (len b)) + 1 >= 0 + 1 by XREAL_1:6; then ((l -' 1) div (len b)) + 1 in Seg (len a) by A53; then A54: ((l -' 1) div (len b)) + 1 in dom F by A13, FINSEQ_1:def_3; percases ( ((k -' 1) div (len b)) + 1 <> ((l -' 1) div (len b)) + 1 or ((k -' 1) mod (len b)) + 1 <> ((l -' 1) mod (len b)) + 1 ) by A49; suppose ((k -' 1) div (len b)) + 1 <> ((l -' 1) div (len b)) + 1 ; ::_thesis: FG . k misses FG . l then F . (((k -' 1) div (len b)) + 1) misses F . (((l -' 1) div (len b)) + 1) by A42, A54, MESFUNC3:4; then (FG . k) /\ (FG . l) = {} /\ ((G . (((k -' 1) mod (len b)) + 1)) /\ (G . (((l -' 1) mod (len b)) + 1))) by A36, XBOOLE_0:def_7; hence FG . k misses FG . l by XBOOLE_0:def_7; ::_thesis: verum end; suppose ((k -' 1) mod (len b)) + 1 <> ((l -' 1) mod (len b)) + 1 ; ::_thesis: FG . k misses FG . l then G . (((k -' 1) mod (len b)) + 1) misses G . (((l -' 1) mod (len b)) + 1) by A46, A44, MESFUNC3:4; then (FG . k) /\ (FG . l) = ((F . (((k -' 1) div (len b)) + 1)) /\ (F . (((l -' 1) div (len b)) + 1))) /\ {} by A36, XBOOLE_0:def_7; hence FG . k misses FG . l by XBOOLE_0:def_7; ::_thesis: verum end; end; end; then reconsider FG = FG as Finite_Sep_Sequence of S by MESFUNC3:4; A55: dom f = union (rng F) by A12, MESFUNC3:def_1; defpred S1[ Nat, set ] means ( ( G . ((($1 -' 1) mod (len b)) + 1) = {} implies $2 = 0 ) & ( G . ((($1 -' 1) mod (len b)) + 1) <> {} implies $2 = b . ((($1 -' 1) mod (len b)) + 1) ) ); defpred S2[ Nat, set ] means ( ( F . ((($1 -' 1) div (len b)) + 1) = {} implies $2 = 0 ) & ( F . ((($1 -' 1) div (len b)) + 1) <> {} implies $2 = a . ((($1 -' 1) div (len b)) + 1) ) ); A56: for k being Nat st k in Seg (len FG) holds ex v being Element of ExtREAL st S2[k,v] proof let k be Nat; ::_thesis: ( k in Seg (len FG) implies ex v being Element of ExtREAL st S2[k,v] ) assume k in Seg (len FG) ; ::_thesis: ex v being Element of ExtREAL st S2[k,v] percases ( F . (((k -' 1) div (len b)) + 1) = {} or F . (((k -' 1) div (len b)) + 1) <> {} ) ; supposeA57: F . (((k -' 1) div (len b)) + 1) = {} ; ::_thesis: ex v being Element of ExtREAL st S2[k,v] take 0. ; ::_thesis: S2[k, 0. ] thus S2[k, 0. ] by A57; ::_thesis: verum end; supposeA58: F . (((k -' 1) div (len b)) + 1) <> {} ; ::_thesis: ex v being Element of ExtREAL st S2[k,v] take a . (((k -' 1) div (len b)) + 1) ; ::_thesis: S2[k,a . (((k -' 1) div (len b)) + 1)] thus S2[k,a . (((k -' 1) div (len b)) + 1)] by A58; ::_thesis: verum end; end; end; consider a1 being FinSequence of ExtREAL such that A59: ( dom a1 = Seg (len FG) & ( for k being Nat st k in Seg (len FG) holds S2[k,a1 . k] ) ) from FINSEQ_1:sch_5(A56); A60: dom g = union (rng G) by A9, MESFUNC3:def_1; A61: dom f = union (rng FG) proof thus dom f c= union (rng FG) :: according to XBOOLE_0:def_10 ::_thesis: union (rng FG) c= dom f proof let z be set ; :: according to TARSKI:def_3 ::_thesis: ( not z in dom f or z in union (rng FG) ) assume A62: z in dom f ; ::_thesis: z in union (rng FG) then consider Y being set such that A63: z in Y and A64: Y in rng F by A55, TARSKI:def_4; consider i being Nat such that A65: i in dom F and A66: Y = F . i by A64, FINSEQ_2:10; A67: i in Seg (len a) by A13, A65, FINSEQ_1:def_3; then 1 <= i by FINSEQ_1:1; then consider i9 being Nat such that A68: i = 1 + i9 by NAT_1:10; consider Z being set such that A69: z in Z and A70: Z in rng G by A5, A60, A62, TARSKI:def_4; consider j being Nat such that A71: j in dom G and A72: Z = G . j by A70, FINSEQ_2:10; A73: j in Seg (len b) by A17, A71, FINSEQ_1:def_3; then A74: 1 <= j by FINSEQ_1:1; then consider j9 being Nat such that A75: j = 1 + j9 by NAT_1:10; (i9 * (len b)) + j is Nat ; then reconsider k = ((i - 1) * (len b)) + j as Element of NAT by A68; i <= len a by A67, FINSEQ_1:1; then i - 1 <= (len a) - 1 by XREAL_1:9; then (i - 1) * (len b) <= ((len a) - 1) * (len b) by XREAL_1:64; then A76: k <= (((len a) - 1) * (len b)) + j by XREAL_1:6; A77: k >= 0 + j by A68, XREAL_1:6; then k -' 1 = k - 1 by A74, XREAL_1:233, XXREAL_0:2; then A78: k -' 1 = (i9 * (len b)) + j9 by A68, A75; A79: j <= len b by A73, FINSEQ_1:1; then (((len a) - 1) * (len b)) + j <= (((len a) - 1) * (len b)) + (len b) by XREAL_1:6; then A80: k <= (len a) * (len b) by A76, XXREAL_0:2; k >= 1 by A74, A77, XXREAL_0:2; then A81: k in Seg ((len a) * (len b)) by A80; then k in dom FG by A14, FINSEQ_1:def_3; then A82: FG . k in rng FG by FUNCT_1:def_3; A83: j9 < len b by A79, A75, NAT_1:13; then A84: j = ((k -' 1) mod (len b)) + 1 by A75, A78, NAT_D:def_2; A85: i = ((k -' 1) div (len b)) + 1 by A68, A78, A83, NAT_D:def_1; z in (F . i) /\ (G . j) by A63, A66, A69, A72, XBOOLE_0:def_4; then z in FG . k by A15, A16, A85, A84, A81; hence z in union (rng FG) by A82, TARSKI:def_4; ::_thesis: verum end; let z be set ; :: according to TARSKI:def_3 ::_thesis: ( not z in union (rng FG) or z in dom f ) A86: len b divides (len a) * (len b) by NAT_D:def_3; assume z in union (rng FG) ; ::_thesis: z in dom f then consider Y being set such that A87: z in Y and A88: Y in rng FG by TARSKI:def_4; consider k being Nat such that A89: k in dom FG and A90: Y = FG . k by A88, FINSEQ_2:10; set i = ((k -' 1) div (len b)) + 1; A91: k in Seg (len FG) by A89, FINSEQ_1:def_3; then A92: k <= (len a) * (len b) by A14, FINSEQ_1:1; then A93: k -' 1 <= ((len a) * (len b)) -' 1 by NAT_D:42; 1 <= k by A91, FINSEQ_1:1; then A94: 1 <= (len a) * (len b) by A92, XXREAL_0:2; A95: len b <> 0 by A14, A91; then len b >= 0 + 1 by NAT_1:13; then (((len a) * (len b)) -' 1) div (len b) = (((len a) * (len b)) div (len b)) - 1 by A86, A94, NAT_2:15; then (k -' 1) div (len b) <= (((len a) * (len b)) div (len b)) - 1 by A93, NAT_2:24; then A96: ((k -' 1) div (len b)) + 1 <= ((len a) * (len b)) div (len b) by XREAL_1:19; set j = ((k -' 1) mod (len b)) + 1; A97: ((k -' 1) div (len b)) + 1 >= 0 + 1 by XREAL_1:6; ((len a) * (len b)) div (len b) = len a by A95, NAT_D:18; then ((k -' 1) div (len b)) + 1 in Seg (len a) by A97, A96; then ((k -' 1) div (len b)) + 1 in dom F by A13, FINSEQ_1:def_3; then A98: F . (((k -' 1) div (len b)) + 1) in rng F by FUNCT_1:def_3; FG . k = (F . (((k -' 1) div (len b)) + 1)) /\ (G . (((k -' 1) mod (len b)) + 1)) by A15, A89; then z in F . (((k -' 1) div (len b)) + 1) by A87, A90, XBOOLE_0:def_4; hence z in dom f by A55, A98, TARSKI:def_4; ::_thesis: verum end; A99: for k being Nat for x, y being Element of X st k in dom FG & x in FG . k & y in FG . k holds (f - g) . x = (f - g) . y proof let k be Nat; ::_thesis: for x, y being Element of X st k in dom FG & x in FG . k & y in FG . k holds (f - g) . x = (f - g) . y let x, y be Element of X; ::_thesis: ( k in dom FG & x in FG . k & y in FG . k implies (f - g) . x = (f - g) . y ) assume that A100: k in dom FG and A101: x in FG . k and A102: y in FG . k ; ::_thesis: (f - g) . x = (f - g) . y set j = ((k -' 1) mod (len b)) + 1; A103: FG . k = (F . (((k -' 1) div (len b)) + 1)) /\ (G . (((k -' 1) mod (len b)) + 1)) by A15, A100; then A104: x in G . (((k -' 1) mod (len b)) + 1) by A101, XBOOLE_0:def_4; set i = ((k -' 1) div (len b)) + 1; A105: ((k -' 1) div (len b)) + 1 >= 0 + 1 by XREAL_1:6; A106: k in Seg (len FG) by A100, FINSEQ_1:def_3; then A107: 1 <= k by FINSEQ_1:1; A108: len b > 0 by A14, A106; then A109: len b >= 0 + 1 by NAT_1:13; A110: y in G . (((k -' 1) mod (len b)) + 1) by A102, A103, XBOOLE_0:def_4; A111: len b divides (len a) * (len b) by NAT_D:def_3; A112: k <= (len a) * (len b) by A14, A106, FINSEQ_1:1; then A113: k -' 1 <= ((len a) * (len b)) -' 1 by NAT_D:42; 1 <= (len a) * (len b) by A107, A112, XXREAL_0:2; then (((len a) * (len b)) -' 1) div (len b) = (((len a) * (len b)) div (len b)) - 1 by A109, A111, NAT_2:15; then (k -' 1) div (len b) <= (((len a) * (len b)) div (len b)) - 1 by A113, NAT_2:24; then A114: ((k -' 1) div (len b)) + 1 <= ((len a) * (len b)) div (len b) by XREAL_1:19; len b <> 0 by A14, A106; then ((k -' 1) div (len b)) + 1 <= len a by A114, NAT_D:18; then ((k -' 1) div (len b)) + 1 in Seg (len a) by A105; then A115: ((k -' 1) div (len b)) + 1 in dom F by A13, FINSEQ_1:def_3; (k -' 1) mod (len b) < len b by A108, NAT_D:1; then A116: ((k -' 1) mod (len b)) + 1 <= len b by NAT_1:13; ((k -' 1) mod (len b)) + 1 >= 0 + 1 by XREAL_1:6; then ((k -' 1) mod (len b)) + 1 in Seg (len b) by A116; then A117: ((k -' 1) mod (len b)) + 1 in dom G by A17, FINSEQ_1:def_3; y in F . (((k -' 1) div (len b)) + 1) by A102, A103, XBOOLE_0:def_4; then A118: f . y = a . (((k -' 1) div (len b)) + 1) by A12, A115, MESFUNC3:def_1; x in F . (((k -' 1) div (len b)) + 1) by A101, A103, XBOOLE_0:def_4; then A119: f . x = a . (((k -' 1) div (len b)) + 1) by A12, A115, MESFUNC3:def_1; A120: FG . k in rng FG by A100, FUNCT_1:def_3; then A121: y in dom (f - g) by A5, A61, A10, A102, TARSKI:def_4; x in dom (f - g) by A5, A61, A10, A101, A120, TARSKI:def_4; then (f - g) . x = (f . x) - (g . x) by MESFUNC1:def_4 .= (a . (((k -' 1) div (len b)) + 1)) - (b . (((k -' 1) mod (len b)) + 1)) by A9, A117, A104, A119, MESFUNC3:def_1 .= (f . y) - (g . y) by A9, A117, A110, A118, MESFUNC3:def_1 ; hence (f - g) . x = (f - g) . y by A121, MESFUNC1:def_4; ::_thesis: verum end; deffunc H2( Nat) -> Element of ExtREAL = (a1 . $1) * ((M * FG) . $1); consider x1 being FinSequence of ExtREAL such that A122: ( len x1 = len FG & ( for k being Nat st k in dom x1 holds x1 . k = H2(k) ) ) from FINSEQ_2:sch_1(); A123: for k being Nat st k in dom FG holds for x being set st x in FG . k holds f . x = a1 . k proof let k be Nat; ::_thesis: ( k in dom FG implies for x being set st x in FG . k holds f . x = a1 . k ) set i = ((k -' 1) div (len b)) + 1; A124: ((k -' 1) div (len b)) + 1 >= 0 + 1 by XREAL_1:6; assume A125: k in dom FG ; ::_thesis: for x being set st x in FG . k holds f . x = a1 . k then A126: k in Seg (len FG) by FINSEQ_1:def_3; let x be set ; ::_thesis: ( x in FG . k implies f . x = a1 . k ) assume A127: x in FG . k ; ::_thesis: f . x = a1 . k FG . k = (F . (((k -' 1) div (len b)) + 1)) /\ (G . (((k -' 1) mod (len b)) + 1)) by A15, A125; then A128: x in F . (((k -' 1) div (len b)) + 1) by A127, XBOOLE_0:def_4; A129: k in Seg (len FG) by A125, FINSEQ_1:def_3; then A130: k <= (len a) * (len b) by A14, FINSEQ_1:1; then k -' 1 <= ((len a) * (len b)) -' 1 by NAT_D:42; then A131: (k -' 1) div (len b) <= (((len a) * (len b)) -' 1) div (len b) by NAT_2:24; A132: len b divides (len a) * (len b) by NAT_D:def_3; 1 <= k by A129, FINSEQ_1:1; then A133: 1 <= (len a) * (len b) by A130, XXREAL_0:2; A134: len b <> 0 by A14, A129; then len b >= 0 + 1 by NAT_1:13; then (((len a) * (len b)) -' 1) div (len b) = (((len a) * (len b)) div (len b)) - 1 by A132, A133, NAT_2:15; then A135: ((k -' 1) div (len b)) + 1 <= ((len a) * (len b)) div (len b) by A131, XREAL_1:19; ((len a) * (len b)) div (len b) = len a by A134, NAT_D:18; then ((k -' 1) div (len b)) + 1 in Seg (len a) by A124, A135; then ((k -' 1) div (len b)) + 1 in dom F by A13, FINSEQ_1:def_3; then f . x = a . (((k -' 1) div (len b)) + 1) by A12, A128, MESFUNC3:def_1; hence f . x = a1 . k by A59, A126, A128; ::_thesis: verum end; A136: for k being Nat st k in Seg (len FG) holds ex v being Element of ExtREAL st S1[k,v] proof let k be Nat; ::_thesis: ( k in Seg (len FG) implies ex v being Element of ExtREAL st S1[k,v] ) assume k in Seg (len FG) ; ::_thesis: ex v being Element of ExtREAL st S1[k,v] percases ( G . (((k -' 1) mod (len b)) + 1) = {} or G . (((k -' 1) mod (len b)) + 1) <> {} ) ; supposeA137: G . (((k -' 1) mod (len b)) + 1) = {} ; ::_thesis: ex v being Element of ExtREAL st S1[k,v] take R_EAL 0 ; ::_thesis: S1[k, R_EAL 0] thus S1[k, R_EAL 0] by A137; ::_thesis: verum end; supposeA138: G . (((k -' 1) mod (len b)) + 1) <> {} ; ::_thesis: ex v being Element of ExtREAL st S1[k,v] take b . (((k -' 1) mod (len b)) + 1) ; ::_thesis: S1[k,b . (((k -' 1) mod (len b)) + 1)] thus S1[k,b . (((k -' 1) mod (len b)) + 1)] by A138; ::_thesis: verum end; end; end; consider b1 being FinSequence of ExtREAL such that A139: ( dom b1 = Seg (len FG) & ( for k being Nat st k in Seg (len FG) holds S1[k,b1 . k] ) ) from FINSEQ_1:sch_5(A136); deffunc H3( Nat) -> Element of ExtREAL = (a1 . $1) - (b1 . $1); consider c1 being FinSequence of ExtREAL such that A140: len c1 = len FG and A141: for k being Nat st k in dom c1 holds c1 . k = H3(k) from FINSEQ_2:sch_1(); A142: dom c1 = Seg (len FG) by A140, FINSEQ_1:def_3; A143: for k being Nat st k in dom FG holds for x being set st x in FG . k holds g . x = b1 . k proof let k be Nat; ::_thesis: ( k in dom FG implies for x being set st x in FG . k holds g . x = b1 . k ) set j = ((k -' 1) mod (len b)) + 1; assume A144: k in dom FG ; ::_thesis: for x being set st x in FG . k holds g . x = b1 . k then A145: k in Seg (len FG) by FINSEQ_1:def_3; k in Seg (len FG) by A144, FINSEQ_1:def_3; then len b <> 0 by A14; then (k -' 1) mod (len b) < len b by NAT_D:1; then A146: ((k -' 1) mod (len b)) + 1 <= len b by NAT_1:13; let x be set ; ::_thesis: ( x in FG . k implies g . x = b1 . k ) assume A147: x in FG . k ; ::_thesis: g . x = b1 . k FG . k = (F . (((k -' 1) div (len b)) + 1)) /\ (G . (((k -' 1) mod (len b)) + 1)) by A15, A144; then A148: x in G . (((k -' 1) mod (len b)) + 1) by A147, XBOOLE_0:def_4; ((k -' 1) mod (len b)) + 1 >= 0 + 1 by XREAL_1:6; then ((k -' 1) mod (len b)) + 1 in Seg (len b) by A146; then ((k -' 1) mod (len b)) + 1 in dom G by A17, FINSEQ_1:def_3; hence g . x = b . (((k -' 1) mod (len b)) + 1) by A9, A148, MESFUNC3:def_1 .= b1 . k by A139, A148, A145 ; ::_thesis: verum end; A149: for k being Nat st k in dom FG holds for x being set st x in FG . k holds (f - g) . x = c1 . k proof let k be Nat; ::_thesis: ( k in dom FG implies for x being set st x in FG . k holds (f - g) . x = c1 . k ) assume A150: k in dom FG ; ::_thesis: for x being set st x in FG . k holds (f - g) . x = c1 . k let x be set ; ::_thesis: ( x in FG . k implies (f - g) . x = c1 . k ) assume A151: x in FG . k ; ::_thesis: (f - g) . x = c1 . k FG . k in rng FG by A150, FUNCT_1:def_3; then x in dom (f - g) by A5, A61, A10, A151, TARSKI:def_4; then A152: (f - g) . x = (f . x) - (g . x) by MESFUNC1:def_4; k in Seg (len FG) by A150, FINSEQ_1:def_3; then (a1 . k) - (b1 . k) = c1 . k by A141, A142; then (a1 . k) - (g . x) = c1 . k by A143, A150, A151; hence (f - g) . x = c1 . k by A123, A150, A151, A152; ::_thesis: verum end; deffunc H4( Nat) -> Element of ExtREAL = (c1 . $1) * ((M * FG) . $1); consider z1 being FinSequence of ExtREAL such that A153: ( len z1 = len FG & ( for k being Nat st k in dom z1 holds z1 . k = H4(k) ) ) from FINSEQ_2:sch_1(); deffunc H5( Nat) -> Element of ExtREAL = (b1 . $1) * ((M * FG) . $1); consider y1 being FinSequence of ExtREAL such that A154: ( len y1 = len FG & ( for k being Nat st k in dom y1 holds y1 . k = H5(k) ) ) from FINSEQ_2:sch_1(); A155: dom x1 = dom FG by A122, FINSEQ_3:29; A156: dom z1 = dom FG by A153, FINSEQ_3:29; A157: for i being Nat st i in dom x1 holds 0 <= z1 . i proof reconsider EMPTY = {} as Element of S by PROB_1:4; let i be Nat; ::_thesis: ( i in dom x1 implies 0 <= z1 . i ) assume A158: i in dom x1 ; ::_thesis: 0 <= z1 . i then A159: (M * FG) . i = M . (FG . i) by A155, FUNCT_1:13; FG . i in rng FG by A155, A158, FUNCT_1:3; then reconsider V = FG . i as Element of S ; M . EMPTY <= M . V by MEASURE1:31, XBOOLE_1:2; then A160: 0 <= (M * FG) . i by A159, VALUED_0:def_19; A161: i in Seg (len FG) by A122, A158, FINSEQ_1:def_3; percases ( FG . i <> {} or FG . i = {} ) ; suppose FG . i <> {} ; ::_thesis: 0 <= z1 . i then consider x being set such that A162: x in FG . i by XBOOLE_0:def_1; FG . i in rng FG by A155, A158, FUNCT_1:3; then x in union (rng FG) by A162, TARSKI:def_4; then g . x <= f . x by A7, A61; then g . x <= a1 . i by A155, A123, A158, A162; then b1 . i <= a1 . i by A155, A143, A158, A162; then 0 <= (a1 . i) - (b1 . i) by XXREAL_3:40; then 0 <= c1 . i by A141, A142, A161; then 0 <= (c1 . i) * ((M * FG) . i) by A160; hence 0 <= z1 . i by A155, A153, A156, A158; ::_thesis: verum end; suppose FG . i = {} ; ::_thesis: 0 <= z1 . i then (M * FG) . i = 0 by A159, VALUED_0:def_19; then (c1 . i) * ((M * FG) . i) = 0 ; hence 0 <= z1 . i by A155, A153, A156, A158; ::_thesis: verum end; end; end; not -infty in rng z1 proof assume -infty in rng z1 ; ::_thesis: contradiction then ex i being set st ( i in dom z1 & z1 . i = -infty ) by FUNCT_1:def_3; hence contradiction by A155, A156, A157; ::_thesis: verum end; then A163: (z1 " {-infty}) /\ (y1 " {+infty}) = {} /\ (y1 " {+infty}) by FUNCT_1:72 .= {} ; A164: dom y1 = dom FG by A154, FINSEQ_3:29; A165: for i being Nat st i in dom y1 holds 0 <= y1 . i proof let i be Nat; ::_thesis: ( i in dom y1 implies 0 <= y1 . i ) set i9 = ((i -' 1) mod (len b)) + 1; A166: ((i -' 1) mod (len b)) + 1 >= 0 + 1 by XREAL_1:6; assume A167: i in dom y1 ; ::_thesis: 0 <= y1 . i then A168: y1 . i = (b1 . i) * ((M * FG) . i) by A154; A169: i in Seg (len FG) by A154, A167, FINSEQ_1:def_3; then len b <> 0 by A14; then (i -' 1) mod (len b) < len b by NAT_D:1; then ((i -' 1) mod (len b)) + 1 <= len b by NAT_1:13; then ((i -' 1) mod (len b)) + 1 in Seg (len b) by A166; then A170: ((i -' 1) mod (len b)) + 1 in dom G by A17, FINSEQ_1:def_3; percases ( G . (((i -' 1) mod (len b)) + 1) <> {} or G . (((i -' 1) mod (len b)) + 1) = {} ) ; supposeA171: G . (((i -' 1) mod (len b)) + 1) <> {} ; ::_thesis: 0 <= y1 . i FG . i in rng FG by A164, A167, FUNCT_1:3; then reconsider FGi = FG . i as Element of S ; reconsider EMPTY = {} as Element of S by MEASURE1:7; EMPTY c= FGi by XBOOLE_1:2; then A172: M . {} <= M . FGi by MEASURE1:31; consider x9 being set such that A173: x9 in G . (((i -' 1) mod (len b)) + 1) by A171, XBOOLE_0:def_1; g . x9 = b . (((i -' 1) mod (len b)) + 1) by A9, A170, A173, MESFUNC3:def_1 .= b1 . i by A139, A169, A171 ; then A174: 0 <= b1 . i by A6, SUPINF_2:51; M . {} = 0 by VALUED_0:def_19; then 0 <= (M * FG) . i by A164, A167, A172, FUNCT_1:13; hence 0 <= y1 . i by A168, A174; ::_thesis: verum end; supposeA175: G . (((i -' 1) mod (len b)) + 1) = {} ; ::_thesis: 0 <= y1 . i FG . i = (F . (((i -' 1) div (len b)) + 1)) /\ (G . (((i -' 1) mod (len b)) + 1)) by A14, A15, A16, A169; then M . (FG . i) = 0 by A175, VALUED_0:def_19; then (M * FG) . i = 0 by A164, A167, FUNCT_1:13; hence 0 <= y1 . i by A168; ::_thesis: verum end; end; end; not -infty in rng y1 proof assume -infty in rng y1 ; ::_thesis: contradiction then ex i being set st ( i in dom y1 & y1 . i = -infty ) by FUNCT_1:def_3; hence contradiction by A165; ::_thesis: verum end; then (z1 " {+infty}) /\ (y1 " {-infty}) = (z1 " {+infty}) /\ {} by FUNCT_1:72 .= {} ; then A176: dom (z1 + y1) = ((dom z1) /\ (dom y1)) \ ({} \/ {}) by A163, MESFUNC1:def_3 .= dom x1 by A122, A164, A156, FINSEQ_3:29 ; A177: for k being Nat st k in dom x1 holds x1 . k = (z1 + y1) . k proof A178: len b divides (len a) * (len b) by NAT_D:def_3; let k be Nat; ::_thesis: ( k in dom x1 implies x1 . k = (z1 + y1) . k ) set p = ((k -' 1) div (len b)) + 1; set q = ((k -' 1) mod (len b)) + 1; A179: ((k -' 1) div (len b)) + 1 >= 0 + 1 by XREAL_1:6; assume A180: k in dom x1 ; ::_thesis: x1 . k = (z1 + y1) . k then A181: k in Seg (len FG) by A122, FINSEQ_1:def_3; then A182: 1 <= k by FINSEQ_1:1; A183: len b > 0 by A14, A181; then A184: len b >= 0 + 1 by NAT_1:13; A185: k <= (len a) * (len b) by A14, A181, FINSEQ_1:1; then A186: k -' 1 <= ((len a) * (len b)) -' 1 by NAT_D:42; 1 <= (len a) * (len b) by A182, A185, XXREAL_0:2; then (((len a) * (len b)) -' 1) div (len b) = (((len a) * (len b)) div (len b)) - 1 by A184, A178, NAT_2:15; then (k -' 1) div (len b) <= (((len a) * (len b)) div (len b)) - 1 by A186, NAT_2:24; then A187: ((k -' 1) div (len b)) + 1 <= ((len a) * (len b)) div (len b) by XREAL_1:19; len b <> 0 by A14, A181; then ((k -' 1) div (len b)) + 1 <= len a by A187, NAT_D:18; then ((k -' 1) div (len b)) + 1 in Seg (len a) by A179; then A188: ((k -' 1) div (len b)) + 1 in dom F by A13, FINSEQ_1:def_3; A189: ((k -' 1) mod (len b)) + 1 >= 0 + 1 by XREAL_1:6; (k -' 1) mod (len b) < len b by A183, NAT_D:1; then ((k -' 1) mod (len b)) + 1 <= len b by NAT_1:13; then ((k -' 1) mod (len b)) + 1 in Seg (len b) by A189; then A190: ((k -' 1) mod (len b)) + 1 in dom G by A17, FINSEQ_1:def_3; A191: ((c1 . k) + (b1 . k)) * ((M * FG) . k) = ((c1 . k) * ((M * FG) . k)) + ((b1 . k) * ((M * FG) . k)) proof percases ( FG . k <> {} or FG . k = {} ) ; suppose FG . k <> {} ; ::_thesis: ((c1 . k) + (b1 . k)) * ((M * FG) . k) = ((c1 . k) * ((M * FG) . k)) + ((b1 . k) * ((M * FG) . k)) then (F . (((k -' 1) div (len b)) + 1)) /\ (G . (((k -' 1) mod (len b)) + 1)) <> {} by A14, A15, A16, A181; then consider v being set such that A192: v in (F . (((k -' 1) div (len b)) + 1)) /\ (G . (((k -' 1) mod (len b)) + 1)) by XBOOLE_0:def_1; A193: G . (((k -' 1) mod (len b)) + 1) <> {} by A192; A194: v in F . (((k -' 1) div (len b)) + 1) by A192, XBOOLE_0:def_4; v in G . (((k -' 1) mod (len b)) + 1) by A192, XBOOLE_0:def_4; then A195: b . (((k -' 1) mod (len b)) + 1) = g . v by A9, A190, MESFUNC3:def_1; F . (((k -' 1) div (len b)) + 1) in rng F by A188, FUNCT_1:3; then A196: v in dom f by A55, A194, TARSKI:def_4; a . (((k -' 1) div (len b)) + 1) = f . v by A12, A188, A194, MESFUNC3:def_1; then b . (((k -' 1) mod (len b)) + 1) <= a . (((k -' 1) div (len b)) + 1) by A7, A195, A196; then A197: b1 . k <= a . (((k -' 1) div (len b)) + 1) by A139, A181, A193; F . (((k -' 1) div (len b)) + 1) <> {} by A192; then b1 . k <= a1 . k by A59, A181, A197; then 0 <= (a1 . k) - (b1 . k) by XXREAL_3:40; then A198: ( 0 = c1 . k or 0 < c1 . k ) by A141, A142, A181; 0 <= b . (((k -' 1) mod (len b)) + 1) by A6, A195, SUPINF_2:51; then ( 0 = b1 . k or 0 < b1 . k ) by A139, A181, A192; hence ((c1 . k) + (b1 . k)) * ((M * FG) . k) = ((c1 . k) * ((M * FG) . k)) + ((b1 . k) * ((M * FG) . k)) by A198, XXREAL_3:96; ::_thesis: verum end; suppose FG . k = {} ; ::_thesis: ((c1 . k) + (b1 . k)) * ((M * FG) . k) = ((c1 . k) * ((M * FG) . k)) + ((b1 . k) * ((M * FG) . k)) then M . (FG . k) = 0 by VALUED_0:def_19; then A199: (M * FG) . k = 0 by A155, A180, FUNCT_1:13; hence ((c1 . k) + (b1 . k)) * ((M * FG) . k) = 0 .= ((c1 . k) * ((M * FG) . k)) + ((b1 . k) * ((M * FG) . k)) by A199 ; ::_thesis: verum end; end; end; A200: ( a1 . k <> +infty & a1 . k <> -infty & b1 . k <> +infty & b1 . k <> -infty ) proof now__::_thesis:_(_a1_._k_<>_+infty_&_-infty_<>_a1_._k_) percases ( F . (((k -' 1) div (len b)) + 1) <> {} or F . (((k -' 1) div (len b)) + 1) = {} ) ; supposeA201: F . (((k -' 1) div (len b)) + 1) <> {} ; ::_thesis: ( a1 . k <> +infty & -infty <> a1 . k ) then consider v being set such that A202: v in F . (((k -' 1) div (len b)) + 1) by XBOOLE_0:def_1; A203: f is V60() by A1, MESFUNC2:def_4; a1 . k = a . (((k -' 1) div (len b)) + 1) by A59, A181, A201; then a1 . k = f . v by A12, A188, A202, MESFUNC3:def_1; hence ( a1 . k <> +infty & -infty <> a1 . k ) by A203; ::_thesis: verum end; suppose F . (((k -' 1) div (len b)) + 1) = {} ; ::_thesis: ( a1 . k <> +infty & -infty <> a1 . k ) hence ( a1 . k <> +infty & -infty <> a1 . k ) by A59, A181; ::_thesis: verum end; end; end; hence ( +infty <> a1 . k & a1 . k <> -infty ) ; ::_thesis: ( b1 . k <> +infty & b1 . k <> -infty ) now__::_thesis:_(_b1_._k_<>_+infty_&_b1_._k_<>_-infty_) percases ( G . (((k -' 1) mod (len b)) + 1) <> {} or G . (((k -' 1) mod (len b)) + 1) = {} ) ; supposeA204: G . (((k -' 1) mod (len b)) + 1) <> {} ; ::_thesis: ( b1 . k <> +infty & b1 . k <> -infty ) then consider v being set such that A205: v in G . (((k -' 1) mod (len b)) + 1) by XBOOLE_0:def_1; A206: g is V60() by A4, MESFUNC2:def_4; b1 . k = b . (((k -' 1) mod (len b)) + 1) by A139, A181, A204; then b1 . k = g . v by A9, A190, A205, MESFUNC3:def_1; hence ( b1 . k <> +infty & b1 . k <> -infty ) by A206; ::_thesis: verum end; suppose G . (((k -' 1) mod (len b)) + 1) = {} ; ::_thesis: ( b1 . k <> +infty & b1 . k <> -infty ) hence ( b1 . k <> +infty & b1 . k <> -infty ) by A139, A181; ::_thesis: verum end; end; end; hence ( b1 . k <> +infty & b1 . k <> -infty ) ; ::_thesis: verum end; A207: (b1 . k) - (b1 . k) = - 0 by XXREAL_3:7; c1 . k = (a1 . k) - (b1 . k) by A141, A142, A181; then (c1 . k) + (b1 . k) = (a1 . k) - ((b1 . k) - (b1 . k)) by A200, XXREAL_3:32 .= (a1 . k) + (- 0) by A207 .= a1 . k by XXREAL_3:4 ; hence x1 . k = ((c1 . k) + (b1 . k)) * ((M * FG) . k) by A122, A180 .= (z1 . k) + ((b1 . k) * ((M * FG) . k)) by A155, A153, A156, A180, A191 .= (z1 . k) + (y1 . k) by A155, A154, A164, A180 .= (z1 + y1) . k by A176, A180, MESFUNC1:def_3 ; ::_thesis: verum end; now__::_thesis:_for_x_being_Element_of_X_st_x_in_dom_(f_-_g)_holds_ |.((f_-_g)_._x).|_<_+infty let x be Element of X; ::_thesis: ( x in dom (f - g) implies |.((f - g) . x).| < +infty ) assume A208: x in dom (f - g) ; ::_thesis: |.((f - g) . x).| < +infty g is V60() by A4, MESFUNC2:def_4; then A209: |.(g . x).| < +infty by A5, A10, A208, MESFUNC2:def_1; f is V60() by A1, MESFUNC2:def_4; then |.(f . x).| < +infty by A5, A10, A208, MESFUNC2:def_1; then A210: |.(f . x).| + |.(g . x).| <> +infty by A209, XXREAL_3:16; |.((f - g) . x).| = |.((f . x) - (g . x)).| by A208, MESFUNC1:def_4; then |.((f - g) . x).| <= |.(f . x).| + |.(g . x).| by EXTREAL2:21; hence |.((f - g) . x).| < +infty by A210, XXREAL_0:2, XXREAL_0:4; ::_thesis: verum end; then f - g is V60() by MESFUNC2:def_1; hence A211: f - g is_simple_func_in S by A5, A61, A10, A99, MESFUNC2:def_4; ::_thesis: ( dom (f - g) <> {} & f - g is nonnegative & integral (X,S,M,f) = (integral (X,S,M,(f - g))) + (integral (X,S,M,g)) ) dom FG = dom a1 by A59, FINSEQ_1:def_3; then FG,a1 are_Re-presentation_of f by A61, A123, MESFUNC3:def_1; then A212: integral (X,S,M,f) = Sum x1 by A1, A2, A11, A122, A155, MESFUNC4:3; dom (z1 + y1) = Seg (len x1) by A176, FINSEQ_1:def_3; then z1 + y1 is FinSequence by FINSEQ_1:def_2; then A213: x1 = z1 + y1 by A176, A177, FINSEQ_1:13; dom FG = dom b1 by A139, FINSEQ_1:def_3; then FG,b1 are_Re-presentation_of g by A5, A61, A143, MESFUNC3:def_1; then A214: integral (X,S,M,g) = Sum y1 by A2, A4, A5, A8, A154, A164, MESFUNC4:3; thus dom (f - g) <> {} by A2, A5, A10; ::_thesis: ( f - g is nonnegative & integral (X,S,M,f) = (integral (X,S,M,(f - g))) + (integral (X,S,M,g)) ) A215: for x being set st x in dom (f - g) holds 0 <= (f - g) . x proof let x be set ; ::_thesis: ( x in dom (f - g) implies 0 <= (f - g) . x ) assume A216: x in dom (f - g) ; ::_thesis: 0 <= (f - g) . x then 0 <= (f . x) - (g . x) by A5, A7, A10, XXREAL_3:40; hence 0 <= (f - g) . x by A216, MESFUNC1:def_4; ::_thesis: verum end; hence f - g is nonnegative by SUPINF_2:52; ::_thesis: integral (X,S,M,f) = (integral (X,S,M,(f - g))) + (integral (X,S,M,g)) dom FG = dom c1 by A140, FINSEQ_3:29; then FG,c1 are_Re-presentation_of f - g by A5, A61, A10, A149, MESFUNC3:def_1; then integral (X,S,M,(f - g)) = Sum z1 by A2, A5, A153, A156, A10, A211, A215, MESFUNC4:3; hence integral (X,S,M,f) = (integral (X,S,M,(f - g))) + (integral (X,S,M,g)) by A155, A164, A156, A212, A214, A165, A157, A213, MESFUNC4:1; ::_thesis: verum end; theorem Th69: :: MESFUNC5:69 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & f is nonnegative & g is_simple_func_in S & g is nonnegative & ( for x being set st x in dom (f - g) holds g . x <= f . x ) holds ( dom (f - g) = (dom f) /\ (dom g) & integral' (M,(f | (dom (f - g)))) = (integral' (M,(f - g))) + (integral' (M,(g | (dom (f - g))))) ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & f is nonnegative & g is_simple_func_in S & g is nonnegative & ( for x being set st x in dom (f - g) holds g . x <= f . x ) holds ( dom (f - g) = (dom f) /\ (dom g) & integral' (M,(f | (dom (f - g)))) = (integral' (M,(f - g))) + (integral' (M,(g | (dom (f - g))))) ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & f is nonnegative & g is_simple_func_in S & g is nonnegative & ( for x being set st x in dom (f - g) holds g . x <= f . x ) holds ( dom (f - g) = (dom f) /\ (dom g) & integral' (M,(f | (dom (f - g)))) = (integral' (M,(f - g))) + (integral' (M,(g | (dom (f - g))))) ) let M be sigma_Measure of S; ::_thesis: for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & f is nonnegative & g is_simple_func_in S & g is nonnegative & ( for x being set st x in dom (f - g) holds g . x <= f . x ) holds ( dom (f - g) = (dom f) /\ (dom g) & integral' (M,(f | (dom (f - g)))) = (integral' (M,(f - g))) + (integral' (M,(g | (dom (f - g))))) ) let f, g be PartFunc of X,ExtREAL; ::_thesis: ( f is_simple_func_in S & f is nonnegative & g is_simple_func_in S & g is nonnegative & ( for x being set st x in dom (f - g) holds g . x <= f . x ) implies ( dom (f - g) = (dom f) /\ (dom g) & integral' (M,(f | (dom (f - g)))) = (integral' (M,(f - g))) + (integral' (M,(g | (dom (f - g))))) ) ) assume that A1: f is_simple_func_in S and A2: f is nonnegative and A3: g is_simple_func_in S and A4: g is nonnegative and A5: for x being set st x in dom (f - g) holds g . x <= f . x ; ::_thesis: ( dom (f - g) = (dom f) /\ (dom g) & integral' (M,(f | (dom (f - g)))) = (integral' (M,(f - g))) + (integral' (M,(g | (dom (f - g))))) ) A6: f | (dom (f - g)) is nonnegative by A2, Th15; (- 1) (#) g is_simple_func_in S by A3, Th39; then - g is_simple_func_in S by MESFUNC2:9; then f + (- g) is_simple_func_in S by A1, Th38; then f - g is_simple_func_in S by MESFUNC2:8; then A7: dom (f - g) is Element of S by Th37; then A8: g | (dom (f - g)) is_simple_func_in S by A3, Th34; A9: g | (dom (f - g)) is nonnegative by A4, Th15; g is () by A3, Th14; then not -infty in rng g by Def3; then A10: g " {-infty} = {} by FUNCT_1:72; f is () by A1, Th14; then not +infty in rng f by Def4; then A11: f " {+infty} = {} by FUNCT_1:72; then A12: ((dom f) /\ (dom g)) \ (((f " {+infty}) /\ (g " {+infty})) \/ ((f " {-infty}) /\ (g " {-infty}))) = (dom f) /\ (dom g) by A10; hence A13: dom (f - g) = (dom f) /\ (dom g) by MESFUNC1:def_4; ::_thesis: integral' (M,(f | (dom (f - g)))) = (integral' (M,(f - g))) + (integral' (M,(g | (dom (f - g))))) dom (f | (dom (f - g))) = (dom f) /\ (dom (f - g)) by RELAT_1:61; then A14: dom (f | (dom (f - g))) = ((dom f) /\ (dom f)) /\ (dom g) by A13, XBOOLE_1:16; A15: for x being set st x in dom (f | (dom (f - g))) holds (g | (dom (f - g))) . x <= (f | (dom (f - g))) . x proof let x be set ; ::_thesis: ( x in dom (f | (dom (f - g))) implies (g | (dom (f - g))) . x <= (f | (dom (f - g))) . x ) assume A16: x in dom (f | (dom (f - g))) ; ::_thesis: (g | (dom (f - g))) . x <= (f | (dom (f - g))) . x then g . x <= f . x by A5, A13, A14; then (g | (dom (f - g))) . x <= f . x by A13, A14, A16, FUNCT_1:49; hence (g | (dom (f - g))) . x <= (f | (dom (f - g))) . x by A13, A14, A16, FUNCT_1:49; ::_thesis: verum end; dom (g | (dom (f - g))) = (dom g) /\ (dom (f - g)) by RELAT_1:61; then A17: dom (g | (dom (f - g))) = ((dom g) /\ (dom g)) /\ (dom f) by A13, XBOOLE_1:16; A18: f | (dom (f - g)) is_simple_func_in S by A1, A7, Th34; thus integral' (M,(f | (dom (f - g)))) = (integral' (M,(f - g))) + (integral' (M,(g | (dom (f - g))))) ::_thesis: verum proof percases ( dom (f - g) = {} or dom (f - g) <> {} ) ; supposeA19: dom (f - g) = {} ; ::_thesis: integral' (M,(f | (dom (f - g)))) = (integral' (M,(f - g))) + (integral' (M,(g | (dom (f - g))))) dom (g | (dom (f - g))) = (dom g) /\ (dom (f - g)) by RELAT_1:61; then A20: integral' (M,(g | (dom (f - g)))) = 0 by A19, Def14; dom (f | (dom (f - g))) = (dom f) /\ (dom (f - g)) by RELAT_1:61; then A21: integral' (M,(f | (dom (f - g)))) = 0 by A19, Def14; integral' (M,(f - g)) = 0 by A19, Def14; hence integral' (M,(f | (dom (f - g)))) = (integral' (M,(f - g))) + (integral' (M,(g | (dom (f - g))))) by A21, A20; ::_thesis: verum end; supposeA22: dom (f - g) <> {} ; ::_thesis: integral' (M,(f | (dom (f - g)))) = (integral' (M,(f - g))) + (integral' (M,(g | (dom (f - g))))) A23: (g | (dom (f - g))) " {-infty} = (dom (f - g)) /\ (g " {-infty}) by FUNCT_1:70; (f | (dom (f - g))) " {+infty} = (dom (f - g)) /\ (f " {+infty}) by FUNCT_1:70; then ((dom (f | (dom (f - g)))) /\ (dom (g | (dom (f - g))))) \ ((((f | (dom (f - g))) " {+infty}) /\ ((g | (dom (f - g))) " {+infty})) \/ (((f | (dom (f - g))) " {-infty}) /\ ((g | (dom (f - g))) " {-infty}))) = dom (f - g) by A11, A10, A12, A14, A17, A23, MESFUNC1:def_4; then A24: dom ((f | (dom (f - g))) - (g | (dom (f - g)))) = dom (f - g) by MESFUNC1:def_4; A25: for x being Element of X st x in dom ((f | (dom (f - g))) - (g | (dom (f - g)))) holds ((f | (dom (f - g))) - (g | (dom (f - g)))) . x = (f - g) . x proof let x be Element of X; ::_thesis: ( x in dom ((f | (dom (f - g))) - (g | (dom (f - g)))) implies ((f | (dom (f - g))) - (g | (dom (f - g)))) . x = (f - g) . x ) assume A26: x in dom ((f | (dom (f - g))) - (g | (dom (f - g)))) ; ::_thesis: ((f | (dom (f - g))) - (g | (dom (f - g)))) . x = (f - g) . x then ((f | (dom (f - g))) - (g | (dom (f - g)))) . x = ((f | (dom (f - g))) . x) - ((g | (dom (f - g))) . x) by MESFUNC1:def_4 .= (f . x) - ((g | (dom (f - g))) . x) by A24, A26, FUNCT_1:49 .= (f . x) - (g . x) by A24, A26, FUNCT_1:49 ; hence ((f | (dom (f - g))) - (g | (dom (f - g)))) . x = (f - g) . x by A24, A26, MESFUNC1:def_4; ::_thesis: verum end; integral (X,S,M,(f | (dom (f - g)))) = (integral (X,S,M,((f | (dom (f - g))) - (g | (dom (f - g)))))) + (integral (X,S,M,(g | (dom (f - g))))) by A13, A18, A8, A6, A9, A14, A17, A15, A22, Lm9; then A27: integral (X,S,M,(f | (dom (f - g)))) = (integral (X,S,M,(f - g))) + (integral (X,S,M,(g | (dom (f - g))))) by A24, A25, PARTFUN1:5; A28: integral (X,S,M,(g | (dom (f - g)))) = integral' (M,(g | (dom (f - g)))) by A13, A17, A22, Def14; integral (X,S,M,(f | (dom (f - g)))) = integral' (M,(f | (dom (f - g)))) by A13, A14, A22, Def14; hence integral' (M,(f | (dom (f - g)))) = (integral' (M,(f - g))) + (integral' (M,(g | (dom (f - g))))) by A22, A27, A28, Def14; ::_thesis: verum end; end; end; end; theorem Th70: :: MESFUNC5:70 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & g is_simple_func_in S & f is nonnegative & g is nonnegative & ( for x being set st x in dom (f - g) holds g . x <= f . x ) holds integral' (M,(g | (dom (f - g)))) <= integral' (M,(f | (dom (f - g)))) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & g is_simple_func_in S & f is nonnegative & g is nonnegative & ( for x being set st x in dom (f - g) holds g . x <= f . x ) holds integral' (M,(g | (dom (f - g)))) <= integral' (M,(f | (dom (f - g)))) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & g is_simple_func_in S & f is nonnegative & g is nonnegative & ( for x being set st x in dom (f - g) holds g . x <= f . x ) holds integral' (M,(g | (dom (f - g)))) <= integral' (M,(f | (dom (f - g)))) let M be sigma_Measure of S; ::_thesis: for f, g being PartFunc of X,ExtREAL st f is_simple_func_in S & g is_simple_func_in S & f is nonnegative & g is nonnegative & ( for x being set st x in dom (f - g) holds g . x <= f . x ) holds integral' (M,(g | (dom (f - g)))) <= integral' (M,(f | (dom (f - g)))) let f, g be PartFunc of X,ExtREAL; ::_thesis: ( f is_simple_func_in S & g is_simple_func_in S & f is nonnegative & g is nonnegative & ( for x being set st x in dom (f - g) holds g . x <= f . x ) implies integral' (M,(g | (dom (f - g)))) <= integral' (M,(f | (dom (f - g)))) ) assume that A1: f is_simple_func_in S and A2: g is_simple_func_in S and A3: f is nonnegative and A4: g is nonnegative and A5: for x being set st x in dom (f - g) holds g . x <= f . x ; ::_thesis: integral' (M,(g | (dom (f - g)))) <= integral' (M,(f | (dom (f - g)))) (- 1) (#) g is_simple_func_in S by A2, Th39; then - g is_simple_func_in S by MESFUNC2:9; then f + (- g) is_simple_func_in S by A1, Th38; then A6: f - g is_simple_func_in S by MESFUNC2:8; A7: integral' (M,(f | (dom (f - g)))) = (integral' (M,(f - g))) + (integral' (M,(g | (dom (f - g))))) by A1, A2, A3, A4, A5, Th69; now__::_thesis:_(_integral'_(M,(f_|_(dom_(f_-_g))))_<>_+infty_implies_integral'_(M,(g_|_(dom_(f_-_g))))_<=_integral'_(M,(f_|_(dom_(f_-_g))))_) assume integral' (M,(f | (dom (f - g)))) <> +infty ; ::_thesis: integral' (M,(g | (dom (f - g)))) <= integral' (M,(f | (dom (f - g)))) 0 <= integral' (M,(f - g)) by A1, A2, A5, A6, Th40, Th68; hence integral' (M,(g | (dom (f - g)))) <= integral' (M,(f | (dom (f - g)))) by A7, XXREAL_3:39; ::_thesis: verum end; hence integral' (M,(g | (dom (f - g)))) <= integral' (M,(f | (dom (f - g)))) by XXREAL_0:4; ::_thesis: verum end; theorem Th71: :: MESFUNC5:71 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for c being R_eal st 0 <= c & f is_simple_func_in S & ( for x being set st x in dom f holds f . x = c ) holds integral' (M,f) = c * (M . (dom f)) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for c being R_eal st 0 <= c & f is_simple_func_in S & ( for x being set st x in dom f holds f . x = c ) holds integral' (M,f) = c * (M . (dom f)) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for c being R_eal st 0 <= c & f is_simple_func_in S & ( for x being set st x in dom f holds f . x = c ) holds integral' (M,f) = c * (M . (dom f)) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for c being R_eal st 0 <= c & f is_simple_func_in S & ( for x being set st x in dom f holds f . x = c ) holds integral' (M,f) = c * (M . (dom f)) let f be PartFunc of X,ExtREAL; ::_thesis: for c being R_eal st 0 <= c & f is_simple_func_in S & ( for x being set st x in dom f holds f . x = c ) holds integral' (M,f) = c * (M . (dom f)) let c be R_eal; ::_thesis: ( 0 <= c & f is_simple_func_in S & ( for x being set st x in dom f holds f . x = c ) implies integral' (M,f) = c * (M . (dom f)) ) assume that A1: 0 <= c and A2: f is_simple_func_in S and A3: for x being set st x in dom f holds f . x = c ; ::_thesis: integral' (M,f) = c * (M . (dom f)) A4: for x being set st x in dom f holds 0 <= f . x by A1, A3; reconsider A = dom f as Element of S by A2, Th37; percases ( dom f = {} or dom f <> {} ) ; supposeA5: dom f = {} ; ::_thesis: integral' (M,f) = c * (M . (dom f)) then A6: M . A = 0 by VALUED_0:def_19; integral' (M,f) = 0 by A5, Def14; hence integral' (M,f) = c * (M . (dom f)) by A6; ::_thesis: verum end; supposeA7: dom f <> {} ; ::_thesis: integral' (M,f) = c * (M . (dom f)) set x = <*(c * (M . A))*>; reconsider a = <*c*> as FinSequence of ExtREAL ; set F = <*(dom f)*>; reconsider x = <*(c * (M . A))*> as FinSequence of ExtREAL ; A8: rng <*(dom f)*> = {A} by FINSEQ_1:38; rng <*(dom f)*> c= S proof let z be set ; :: according to TARSKI:def_3 ::_thesis: ( not z in rng <*(dom f)*> or z in S ) assume z in rng <*(dom f)*> ; ::_thesis: z in S then z = A by A8, TARSKI:def_1; hence z in S ; ::_thesis: verum end; then reconsider F = <*(dom f)*> as FinSequence of S by FINSEQ_1:def_4; for i, j being Nat st i in dom F & j in dom F & i <> j holds F . i misses F . j proof let i, j be Nat; ::_thesis: ( i in dom F & j in dom F & i <> j implies F . i misses F . j ) assume that A9: i in dom F and A10: j in dom F and A11: i <> j ; ::_thesis: F . i misses F . j A12: dom F = {1} by FINSEQ_1:2, FINSEQ_1:38; then i = 1 by A9, TARSKI:def_1; hence F . i misses F . j by A10, A11, A12, TARSKI:def_1; ::_thesis: verum end; then reconsider F = F as Finite_Sep_Sequence of S by MESFUNC3:4; A13: dom F = Seg 1 by FINSEQ_1:38 .= dom a by FINSEQ_1:38 ; A14: for n being Nat st n in dom F holds for x being set st x in F . n holds f . x = a . n proof let n be Nat; ::_thesis: ( n in dom F implies for x being set st x in F . n holds f . x = a . n ) assume n in dom F ; ::_thesis: for x being set st x in F . n holds f . x = a . n then n in {1} by FINSEQ_1:2, FINSEQ_1:38; then A15: n = 1 by TARSKI:def_1; let x be set ; ::_thesis: ( x in F . n implies f . x = a . n ) assume x in F . n ; ::_thesis: f . x = a . n then x in dom f by A15, FINSEQ_1:40; then f . x = c by A3; hence f . x = a . n by A15, FINSEQ_1:40; ::_thesis: verum end; A16: for n being Nat st n in dom x holds x . n = c * (M . A) proof let n be Nat; ::_thesis: ( n in dom x implies x . n = c * (M . A) ) assume n in dom x ; ::_thesis: x . n = c * (M . A) then n in {1} by FINSEQ_1:2, FINSEQ_1:38; then n = 1 by TARSKI:def_1; hence x . n = c * (M . A) by FINSEQ_1:40; ::_thesis: verum end; A17: dom x = Seg 1 by FINSEQ_1:38 .= dom F by FINSEQ_1:38 ; A18: for n being Nat st n in dom x holds x . n = (a . n) * ((M * F) . n) proof let n be Nat; ::_thesis: ( n in dom x implies x . n = (a . n) * ((M * F) . n) ) assume A19: n in dom x ; ::_thesis: x . n = (a . n) * ((M * F) . n) then n in {1} by FINSEQ_1:2, FINSEQ_1:38; then A20: n = 1 by TARSKI:def_1; then A21: x . n = c * (M . A) by FINSEQ_1:40; (M * F) . n = M . (F . n) by A17, A19, FUNCT_1:13 .= M . A by A20, FINSEQ_1:40 ; hence x . n = (a . n) * ((M * F) . n) by A20, A21, FINSEQ_1:40; ::_thesis: verum end; dom f = union (rng F) by A8, ZFMISC_1:25; then F,a are_Re-presentation_of f by A13, A14, MESFUNC3:def_1; then integral (X,S,M,f) = Sum x by A2, A4, A7, A17, A18, MESFUNC4:3; then A22: integral' (M,f) = Sum x by A7, Def14; 1 = len x by FINSEQ_1:40; then Sum x = (R_EAL 1) * (c * (M . A)) by A16, MESFUNC3:18; hence integral' (M,f) = c * (M . (dom f)) by A22, XXREAL_3:81; ::_thesis: verum end; end; end; theorem Th72: :: MESFUNC5:72 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_simple_func_in S & f is nonnegative holds integral' (M,(f | (eq_dom (f,(R_EAL 0))))) = 0 proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_simple_func_in S & f is nonnegative holds integral' (M,(f | (eq_dom (f,(R_EAL 0))))) = 0 let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_simple_func_in S & f is nonnegative holds integral' (M,(f | (eq_dom (f,(R_EAL 0))))) = 0 let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL st f is_simple_func_in S & f is nonnegative holds integral' (M,(f | (eq_dom (f,(R_EAL 0))))) = 0 let f be PartFunc of X,ExtREAL; ::_thesis: ( f is_simple_func_in S & f is nonnegative implies integral' (M,(f | (eq_dom (f,(R_EAL 0))))) = 0 ) assume that A1: f is_simple_func_in S and A2: f is nonnegative ; ::_thesis: integral' (M,(f | (eq_dom (f,(R_EAL 0))))) = 0 set A = dom f; set g = f | ((dom f) /\ (eq_dom (f,(R_EAL 0)))); for x being set st x in eq_dom (f,(R_EAL 0)) holds x in dom f by MESFUNC1:def_15; then eq_dom (f,(R_EAL 0)) c= dom f by TARSKI:def_3; then A3: f | ((dom f) /\ (eq_dom (f,(R_EAL 0)))) = f | (eq_dom (f,(R_EAL 0))) by XBOOLE_1:28; A4: ex G being Finite_Sep_Sequence of S st ( dom (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) = union (rng G) & ( for n being Nat for x, y being Element of X st n in dom G & x in G . n & y in G . n holds (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . x = (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . y ) ) proof consider F being Finite_Sep_Sequence of S such that A5: dom f = union (rng F) and A6: for n being Nat for x, y being Element of X st n in dom F & x in F . n & y in F . n holds f . x = f . y by A1, MESFUNC2:def_4; deffunc H1( Nat) -> Element of bool X = (F . $1) /\ ((dom f) /\ (eq_dom (f,(R_EAL 0)))); reconsider A = dom f as Element of S by A5, MESFUNC2:31; consider G being FinSequence such that A7: ( len G = len F & ( for n being Nat st n in dom G holds G . n = H1(n) ) ) from FINSEQ_1:sch_2(); f is_measurable_on A by A1, MESFUNC2:34; then A /\ (less_dom (f,(R_EAL 0))) in S by MESFUNC1:def_16; then A \ (A /\ (less_dom (f,(R_EAL 0)))) in S by PROB_1:6; then reconsider A1 = A /\ (great_eq_dom (f,0.)) as Element of S by MESFUNC1:14; f is_measurable_on A1 by A1, MESFUNC2:34; then (A /\ (great_eq_dom (f,(R_EAL 0)))) /\ (less_eq_dom (f,(R_EAL 0))) in S by MESFUNC1:28; then reconsider A2 = A /\ (eq_dom (f,(R_EAL 0))) as Element of S by MESFUNC1:18; A8: dom F = Seg (len F) by FINSEQ_1:def_3; dom G = Seg (len F) by A7, FINSEQ_1:def_3; then A9: for i being Nat st i in dom F holds G . i = (F . i) /\ A2 by A7, A8; dom G = Seg (len F) by A7, FINSEQ_1:def_3; then A10: dom G = dom F by FINSEQ_1:def_3; then reconsider G = G as Finite_Sep_Sequence of S by A9, Th35; take G ; ::_thesis: ( dom (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) = union (rng G) & ( for n being Nat for x, y being Element of X st n in dom G & x in G . n & y in G . n holds (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . x = (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . y ) ) for i being Nat st i in dom G holds G . i = A2 /\ (F . i) by A7; then A11: union (rng G) = A2 /\ (dom f) by A5, A10, MESFUNC3:6 .= dom (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) by RELAT_1:61 ; for i being Nat for x, y being Element of X st i in dom G & x in G . i & y in G . i holds (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . x = (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . y proof let i be Nat; ::_thesis: for x, y being Element of X st i in dom G & x in G . i & y in G . i holds (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . x = (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . y let x, y be Element of X; ::_thesis: ( i in dom G & x in G . i & y in G . i implies (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . x = (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . y ) assume that A12: i in dom G and A13: x in G . i and A14: y in G . i ; ::_thesis: (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . x = (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . y A15: G . i = (F . i) /\ A2 by A7, A12; then A16: y in F . i by A14, XBOOLE_0:def_4; A17: G . i in rng G by A12, FUNCT_1:3; then x in dom (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) by A11, A13, TARSKI:def_4; then A18: (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . x = f . x by FUNCT_1:47; y in dom (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) by A11, A14, A17, TARSKI:def_4; then A19: (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . y = f . y by FUNCT_1:47; x in F . i by A13, A15, XBOOLE_0:def_4; hence (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . x = (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . y by A6, A10, A12, A16, A18, A19; ::_thesis: verum end; hence ( dom (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) = union (rng G) & ( for n being Nat for x, y being Element of X st n in dom G & x in G . n & y in G . n holds (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . x = (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . y ) ) by A11; ::_thesis: verum end; A20: for x being set st x in dom (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) holds 0 <= (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . x proof let x be set ; ::_thesis: ( x in dom (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) implies 0 <= (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . x ) assume A21: x in dom (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) ; ::_thesis: 0 <= (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . x 0 <= f . x by A2, SUPINF_2:51; hence 0 <= (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . x by A21, FUNCT_1:47; ::_thesis: verum end; f is V60() by A1, MESFUNC2:def_4; then A22: f | ((dom f) /\ (eq_dom (f,(R_EAL 0)))) is_simple_func_in S by A4, MESFUNC2:def_4; now__::_thesis:_(_dom_(f_|_((dom_f)_/\_(eq_dom_(f,(R_EAL_0)))))_<>_{}_implies_integral'_(M,(f_|_(eq_dom_(f,(R_EAL_0)))))_=_0_) consider F being Finite_Sep_Sequence of S, a, x being FinSequence of ExtREAL such that A23: F,a are_Re-presentation_of f | ((dom f) /\ (eq_dom (f,(R_EAL 0)))) and a . 1 = 0 and for n being Nat st 2 <= n & n in dom a holds ( 0 < a . n & a . n < +infty ) and A24: dom x = dom F and A25: for n being Nat st n in dom x holds x . n = (a . n) * ((M * F) . n) and A26: integral (X,S,M,(f | ((dom f) /\ (eq_dom (f,(R_EAL 0)))))) = Sum x by A22, A20, MESFUNC3:def_2; A27: for x being set st x in dom (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) holds (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . x = 0 proof let x be set ; ::_thesis: ( x in dom (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) implies (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . x = 0 ) assume A28: x in dom (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) ; ::_thesis: (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . x = 0 then x in (dom f) /\ ((dom f) /\ (eq_dom (f,(R_EAL 0)))) by RELAT_1:61; then x in (dom f) /\ (eq_dom (f,(R_EAL 0))) by XBOOLE_0:def_4; then x in eq_dom (f,(R_EAL 0)) by XBOOLE_0:def_4; then 0 = f . x by MESFUNC1:def_15; hence (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . x = 0 by A28, FUNCT_1:47; ::_thesis: verum end; A29: for n being Nat holds ( not n in dom F or a . n = 0 or F . n = {} ) proof let n be Nat; ::_thesis: ( not n in dom F or a . n = 0 or F . n = {} ) assume A30: n in dom F ; ::_thesis: ( a . n = 0 or F . n = {} ) now__::_thesis:_(_not_F_._n_<>_{}_or_a_._n_=_0_or_F_._n_=_{}_) assume F . n <> {} ; ::_thesis: ( a . n = 0 or F . n = {} ) then consider x being set such that A31: x in F . n by XBOOLE_0:def_1; F . n in rng F by A30, FUNCT_1:3; then x in union (rng F) by A31, TARSKI:def_4; then x in dom (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) by A23, MESFUNC3:def_1; then (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) . x = 0 by A27; hence ( a . n = 0 or F . n = {} ) by A23, A30, A31, MESFUNC3:def_1; ::_thesis: verum end; hence ( a . n = 0 or F . n = {} ) ; ::_thesis: verum end; A32: for n being Nat st n in dom x holds x . n = 0 proof let n be Nat; ::_thesis: ( n in dom x implies x . n = 0 ) assume A33: n in dom x ; ::_thesis: x . n = 0 percases ( a . n = 0 or F . n = {} ) by A24, A29, A33; suppose a . n = 0 ; ::_thesis: x . n = 0 then (a . n) * ((M * F) . n) = 0 ; hence x . n = 0 by A25, A33; ::_thesis: verum end; suppose F . n = {} ; ::_thesis: x . n = 0 then M . (F . n) = 0 by VALUED_0:def_19; then (M * F) . n = 0 by A24, A33, FUNCT_1:13; then (a . n) * ((M * F) . n) = 0 ; hence x . n = 0 by A25, A33; ::_thesis: verum end; end; end; A34: Sum x = 0 proof consider sumx being Function of NAT,ExtREAL such that A35: Sum x = sumx . (len x) and A36: sumx . 0 = 0 and A37: for i being Element of NAT st i < len x holds sumx . (i + 1) = (sumx . i) + (x . (i + 1)) by EXTREAL1:def_2; now__::_thesis:_(_x_<>_{}_implies_Sum_x_=_0_) defpred S1[ Nat] means ( $1 <= len x implies sumx . $1 = 0 ); assume x <> {} ; ::_thesis: Sum x = 0 A38: for k being Nat st S1[k] holds S1[k + 1] proof let k be Nat; ::_thesis: ( S1[k] implies S1[k + 1] ) assume A39: S1[k] ; ::_thesis: S1[k + 1] assume A40: k + 1 <= len x ; ::_thesis: sumx . (k + 1) = 0 reconsider k = k as Element of NAT by ORDINAL1:def_12; 1 <= k + 1 by NAT_1:11; then k + 1 in Seg (len x) by A40; then k + 1 in dom x by FINSEQ_1:def_3; then A41: x . (k + 1) = 0 by A32; k < len x by A40, NAT_1:13; then sumx . (k + 1) = (sumx . k) + (x . (k + 1)) by A37; hence sumx . (k + 1) = 0 by A39, A40, A41, NAT_1:13; ::_thesis: verum end; A42: S1[ 0 ] by A36; for i being Nat holds S1[i] from NAT_1:sch_2(A42, A38); hence Sum x = 0 by A35; ::_thesis: verum end; hence Sum x = 0 by A35, A36, CARD_1:27; ::_thesis: verum end; assume dom (f | ((dom f) /\ (eq_dom (f,(R_EAL 0))))) <> {} ; ::_thesis: integral' (M,(f | (eq_dom (f,(R_EAL 0))))) = 0 hence integral' (M,(f | (eq_dom (f,(R_EAL 0))))) = 0 by A3, A26, A34, Def14; ::_thesis: verum end; hence integral' (M,(f | (eq_dom (f,(R_EAL 0))))) = 0 by A3, Def14; ::_thesis: verum end; theorem Th73: :: MESFUNC5:73 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for B being Element of S for f being PartFunc of X,ExtREAL st f is_simple_func_in S & M . B = 0 & f is nonnegative holds integral' (M,(f | B)) = 0 proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for B being Element of S for f being PartFunc of X,ExtREAL st f is_simple_func_in S & M . B = 0 & f is nonnegative holds integral' (M,(f | B)) = 0 let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for B being Element of S for f being PartFunc of X,ExtREAL st f is_simple_func_in S & M . B = 0 & f is nonnegative holds integral' (M,(f | B)) = 0 let M be sigma_Measure of S; ::_thesis: for B being Element of S for f being PartFunc of X,ExtREAL st f is_simple_func_in S & M . B = 0 & f is nonnegative holds integral' (M,(f | B)) = 0 let B be Element of S; ::_thesis: for f being PartFunc of X,ExtREAL st f is_simple_func_in S & M . B = 0 & f is nonnegative holds integral' (M,(f | B)) = 0 let f be PartFunc of X,ExtREAL; ::_thesis: ( f is_simple_func_in S & M . B = 0 & f is nonnegative implies integral' (M,(f | B)) = 0 ) assume that A1: f is_simple_func_in S and A2: M . B = 0 and A3: f is nonnegative ; ::_thesis: integral' (M,(f | B)) = 0 set A = dom f; set g = f | ((dom f) /\ B); A4: for x being set st x in dom (f | ((dom f) /\ B)) holds 0 <= (f | ((dom f) /\ B)) . x proof let x be set ; ::_thesis: ( x in dom (f | ((dom f) /\ B)) implies 0 <= (f | ((dom f) /\ B)) . x ) assume A5: x in dom (f | ((dom f) /\ B)) ; ::_thesis: 0 <= (f | ((dom f) /\ B)) . x 0 <= f . x by A3, SUPINF_2:51; hence 0 <= (f | ((dom f) /\ B)) . x by A5, FUNCT_1:47; ::_thesis: verum end; A6: ex G being Finite_Sep_Sequence of S st ( dom (f | ((dom f) /\ B)) = union (rng G) & ( for n being Nat for x, y being Element of X st n in dom G & x in G . n & y in G . n holds (f | ((dom f) /\ B)) . x = (f | ((dom f) /\ B)) . y ) ) proof consider F being Finite_Sep_Sequence of S such that A7: dom f = union (rng F) and A8: for n being Nat for x, y being Element of X st n in dom F & x in F . n & y in F . n holds f . x = f . y by A1, MESFUNC2:def_4; deffunc H1( Nat) -> Element of bool X = (F . $1) /\ ((dom f) /\ B); reconsider A = dom f as Element of S by A7, MESFUNC2:31; reconsider A2 = A /\ B as Element of S ; consider G being FinSequence such that A9: ( len G = len F & ( for n being Nat st n in dom G holds G . n = H1(n) ) ) from FINSEQ_1:sch_2(); A10: dom F = Seg (len F) by FINSEQ_1:def_3; dom G = Seg (len F) by A9, FINSEQ_1:def_3; then A11: for i being Nat st i in dom F holds G . i = (F . i) /\ A2 by A9, A10; dom G = Seg (len F) by A9, FINSEQ_1:def_3; then A12: dom G = dom F by FINSEQ_1:def_3; then reconsider G = G as Finite_Sep_Sequence of S by A11, Th35; take G ; ::_thesis: ( dom (f | ((dom f) /\ B)) = union (rng G) & ( for n being Nat for x, y being Element of X st n in dom G & x in G . n & y in G . n holds (f | ((dom f) /\ B)) . x = (f | ((dom f) /\ B)) . y ) ) for i being Nat st i in dom G holds G . i = A2 /\ (F . i) by A9; then A13: union (rng G) = A2 /\ (dom f) by A7, A12, MESFUNC3:6 .= dom (f | ((dom f) /\ B)) by RELAT_1:61 ; for i being Nat for x, y being Element of X st i in dom G & x in G . i & y in G . i holds (f | ((dom f) /\ B)) . x = (f | ((dom f) /\ B)) . y proof let i be Nat; ::_thesis: for x, y being Element of X st i in dom G & x in G . i & y in G . i holds (f | ((dom f) /\ B)) . x = (f | ((dom f) /\ B)) . y let x, y be Element of X; ::_thesis: ( i in dom G & x in G . i & y in G . i implies (f | ((dom f) /\ B)) . x = (f | ((dom f) /\ B)) . y ) assume that A14: i in dom G and A15: x in G . i and A16: y in G . i ; ::_thesis: (f | ((dom f) /\ B)) . x = (f | ((dom f) /\ B)) . y A17: G . i = (F . i) /\ A2 by A9, A14; then A18: y in F . i by A16, XBOOLE_0:def_4; A19: G . i in rng G by A14, FUNCT_1:3; then x in dom (f | ((dom f) /\ B)) by A13, A15, TARSKI:def_4; then A20: (f | ((dom f) /\ B)) . x = f . x by FUNCT_1:47; y in dom (f | ((dom f) /\ B)) by A13, A16, A19, TARSKI:def_4; then A21: (f | ((dom f) /\ B)) . y = f . y by FUNCT_1:47; x in F . i by A15, A17, XBOOLE_0:def_4; hence (f | ((dom f) /\ B)) . x = (f | ((dom f) /\ B)) . y by A8, A12, A14, A18, A20, A21; ::_thesis: verum end; hence ( dom (f | ((dom f) /\ B)) = union (rng G) & ( for n being Nat for x, y being Element of X st n in dom G & x in G . n & y in G . n holds (f | ((dom f) /\ B)) . x = (f | ((dom f) /\ B)) . y ) ) by A13; ::_thesis: verum end; dom (f | ((dom f) /\ B)) = (dom f) /\ ((dom f) /\ B) by RELAT_1:61; then A22: dom (f | ((dom f) /\ B)) = ((dom f) /\ (dom f)) /\ B by XBOOLE_1:16; then A23: dom (f | ((dom f) /\ B)) = dom (f | B) by RELAT_1:61; for x being set st x in dom (f | ((dom f) /\ B)) holds (f | ((dom f) /\ B)) . x = (f | B) . x proof let x be set ; ::_thesis: ( x in dom (f | ((dom f) /\ B)) implies (f | ((dom f) /\ B)) . x = (f | B) . x ) assume A24: x in dom (f | ((dom f) /\ B)) ; ::_thesis: (f | ((dom f) /\ B)) . x = (f | B) . x then (f | ((dom f) /\ B)) . x = f . x by FUNCT_1:47; hence (f | ((dom f) /\ B)) . x = (f | B) . x by A23, A24, FUNCT_1:47; ::_thesis: verum end; then A25: f | ((dom f) /\ B) = f | B by A23, FUNCT_1:2; f is V60() by A1, MESFUNC2:def_4; then A26: f | ((dom f) /\ B) is_simple_func_in S by A6, MESFUNC2:def_4; now__::_thesis:_integral'_(M,(f_|_B))_=_0 percases ( dom (f | ((dom f) /\ B)) = {} or dom (f | ((dom f) /\ B)) <> {} ) ; suppose dom (f | ((dom f) /\ B)) = {} ; ::_thesis: integral' (M,(f | B)) = 0 hence integral' (M,(f | B)) = 0 by A23, Def14; ::_thesis: verum end; supposeA27: dom (f | ((dom f) /\ B)) <> {} ; ::_thesis: integral' (M,(f | B)) = 0 consider F being Finite_Sep_Sequence of S, a, x being FinSequence of ExtREAL such that A28: F,a are_Re-presentation_of f | ((dom f) /\ B) and a . 1 = 0 and for n being Nat st 2 <= n & n in dom a holds ( 0 < a . n & a . n < +infty ) and A29: dom x = dom F and A30: for n being Nat st n in dom x holds x . n = (a . n) * ((M * F) . n) and A31: integral (X,S,M,(f | ((dom f) /\ B))) = Sum x by A26, A4, MESFUNC3:def_2; A32: for n being Nat st n in dom F holds M . (F . n) = 0 proof reconsider BB = B as measure_zero of M by A2, MEASURE1:def_7; let n be Nat; ::_thesis: ( n in dom F implies M . (F . n) = 0 ) A33: dom (f | ((dom f) /\ B)) c= B by A22, XBOOLE_1:17; assume A34: n in dom F ; ::_thesis: M . (F . n) = 0 then F . n in rng F by FUNCT_1:3; then reconsider FF = F . n as Element of S ; for v being set st v in F . n holds v in union (rng F) proof let v be set ; ::_thesis: ( v in F . n implies v in union (rng F) ) assume A35: v in F . n ; ::_thesis: v in union (rng F) F . n in rng F by A34, FUNCT_1:3; hence v in union (rng F) by A35, TARSKI:def_4; ::_thesis: verum end; then A36: F . n c= union (rng F) by TARSKI:def_3; union (rng F) = dom (f | ((dom f) /\ B)) by A28, MESFUNC3:def_1; then FF c= BB by A36, A33, XBOOLE_1:1; then F . n is measure_zero of M by MEASURE1:36; hence M . (F . n) = 0 by MEASURE1:def_7; ::_thesis: verum end; A37: for n being Nat st n in dom x holds x . n = 0 proof let n be Nat; ::_thesis: ( n in dom x implies x . n = 0 ) assume A38: n in dom x ; ::_thesis: x . n = 0 then M . (F . n) = 0 by A29, A32; then (M * F) . n = 0 by A29, A38, FUNCT_1:13; then (a . n) * ((M * F) . n) = 0 ; hence x . n = 0 by A30, A38; ::_thesis: verum end; Sum x = 0 proof consider sumx being Function of NAT,ExtREAL such that A39: Sum x = sumx . (len x) and A40: sumx . 0 = 0 and A41: for i being Element of NAT st i < len x holds sumx . (i + 1) = (sumx . i) + (x . (i + 1)) by EXTREAL1:def_2; now__::_thesis:_(_x_<>_{}_implies_Sum_x_=_0_) defpred S1[ Nat] means ( $1 <= len x implies sumx . $1 = 0 ); assume x <> {} ; ::_thesis: Sum x = 0 A42: for k being Nat st S1[k] holds S1[k + 1] proof let k be Nat; ::_thesis: ( S1[k] implies S1[k + 1] ) assume A43: S1[k] ; ::_thesis: S1[k + 1] assume A44: k + 1 <= len x ; ::_thesis: sumx . (k + 1) = 0 reconsider k = k as Element of NAT by ORDINAL1:def_12; 1 <= k + 1 by NAT_1:11; then k + 1 in Seg (len x) by A44; then k + 1 in dom x by FINSEQ_1:def_3; then A45: x . (k + 1) = 0 by A37; k < len x by A44, NAT_1:13; then sumx . (k + 1) = (sumx . k) + (x . (k + 1)) by A41; hence sumx . (k + 1) = 0 by A43, A44, A45, NAT_1:13; ::_thesis: verum end; A46: S1[ 0 ] by A40; for i being Nat holds S1[i] from NAT_1:sch_2(A46, A42); hence Sum x = 0 by A39; ::_thesis: verum end; hence Sum x = 0 by A39, A40, CARD_1:27; ::_thesis: verum end; hence integral' (M,(f | B)) = 0 by A25, A27, A31, Def14; ::_thesis: verum end; end; end; hence integral' (M,(f | B)) = 0 ; ::_thesis: verum end; theorem Th74: :: MESFUNC5:74 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for g being PartFunc of X,ExtREAL for F being Functional_Sequence of X,ExtREAL for L being ExtREAL_sequence st g is_simple_func_in S & ( for x being set st x in dom g holds 0 < g . x ) & ( for n being Nat holds F . n is_simple_func_in S ) & ( for n being Nat holds dom (F . n) = dom g ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom g holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom g holds ( F # x is convergent & g . x <= lim (F # x) ) ) & ( for n being Nat holds L . n = integral' (M,(F . n)) ) holds ( L is convergent & integral' (M,g) <= lim L ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for g being PartFunc of X,ExtREAL for F being Functional_Sequence of X,ExtREAL for L being ExtREAL_sequence st g is_simple_func_in S & ( for x being set st x in dom g holds 0 < g . x ) & ( for n being Nat holds F . n is_simple_func_in S ) & ( for n being Nat holds dom (F . n) = dom g ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom g holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom g holds ( F # x is convergent & g . x <= lim (F # x) ) ) & ( for n being Nat holds L . n = integral' (M,(F . n)) ) holds ( L is convergent & integral' (M,g) <= lim L ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for g being PartFunc of X,ExtREAL for F being Functional_Sequence of X,ExtREAL for L being ExtREAL_sequence st g is_simple_func_in S & ( for x being set st x in dom g holds 0 < g . x ) & ( for n being Nat holds F . n is_simple_func_in S ) & ( for n being Nat holds dom (F . n) = dom g ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom g holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom g holds ( F # x is convergent & g . x <= lim (F # x) ) ) & ( for n being Nat holds L . n = integral' (M,(F . n)) ) holds ( L is convergent & integral' (M,g) <= lim L ) let M be sigma_Measure of S; ::_thesis: for g being PartFunc of X,ExtREAL for F being Functional_Sequence of X,ExtREAL for L being ExtREAL_sequence st g is_simple_func_in S & ( for x being set st x in dom g holds 0 < g . x ) & ( for n being Nat holds F . n is_simple_func_in S ) & ( for n being Nat holds dom (F . n) = dom g ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom g holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom g holds ( F # x is convergent & g . x <= lim (F # x) ) ) & ( for n being Nat holds L . n = integral' (M,(F . n)) ) holds ( L is convergent & integral' (M,g) <= lim L ) let g be PartFunc of X,ExtREAL; ::_thesis: for F being Functional_Sequence of X,ExtREAL for L being ExtREAL_sequence st g is_simple_func_in S & ( for x being set st x in dom g holds 0 < g . x ) & ( for n being Nat holds F . n is_simple_func_in S ) & ( for n being Nat holds dom (F . n) = dom g ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom g holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom g holds ( F # x is convergent & g . x <= lim (F # x) ) ) & ( for n being Nat holds L . n = integral' (M,(F . n)) ) holds ( L is convergent & integral' (M,g) <= lim L ) let F be Functional_Sequence of X,ExtREAL; ::_thesis: for L being ExtREAL_sequence st g is_simple_func_in S & ( for x being set st x in dom g holds 0 < g . x ) & ( for n being Nat holds F . n is_simple_func_in S ) & ( for n being Nat holds dom (F . n) = dom g ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom g holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom g holds ( F # x is convergent & g . x <= lim (F # x) ) ) & ( for n being Nat holds L . n = integral' (M,(F . n)) ) holds ( L is convergent & integral' (M,g) <= lim L ) let L be ExtREAL_sequence; ::_thesis: ( g is_simple_func_in S & ( for x being set st x in dom g holds 0 < g . x ) & ( for n being Nat holds F . n is_simple_func_in S ) & ( for n being Nat holds dom (F . n) = dom g ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom g holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom g holds ( F # x is convergent & g . x <= lim (F # x) ) ) & ( for n being Nat holds L . n = integral' (M,(F . n)) ) implies ( L is convergent & integral' (M,g) <= lim L ) ) assume that A1: g is_simple_func_in S and A2: for x being set st x in dom g holds 0 < g . x and A3: for n being Nat holds F . n is_simple_func_in S and A4: for n being Nat holds dom (F . n) = dom g and A5: for n being Nat holds F . n is nonnegative and A6: for n, m being Nat st n <= m holds for x being Element of X st x in dom g holds (F . n) . x <= (F . m) . x and A7: for x being Element of X st x in dom g holds ( F # x is convergent & g . x <= lim (F # x) ) and A8: for n being Nat holds L . n = integral' (M,(F . n)) ; ::_thesis: ( L is convergent & integral' (M,g) <= lim L ) percases ( dom g = {} or dom g <> {} ) ; supposeA9: dom g = {} ; ::_thesis: ( L is convergent & integral' (M,g) <= lim L ) A10: now__::_thesis:_for_n_being_Nat_holds_L_._n_=_R_EAL_0 let n be Nat; ::_thesis: L . n = R_EAL 0 dom (F . n) = {} by A4, A9; then integral' (M,(F . n)) = 0 by Def14; hence L . n = R_EAL 0 by A8; ::_thesis: verum end; then L is convergent_to_finite_number by Th52; hence L is convergent by Def11; ::_thesis: integral' (M,g) <= lim L lim L = R_EAL 0 by A10, Th52; hence integral' (M,g) <= lim L by A9, Def14; ::_thesis: verum end; supposeA11: dom g <> {} ; ::_thesis: ( L is convergent & integral' (M,g) <= lim L ) for v being set st v in dom g holds 0 <= g . v by A2; then consider G being Finite_Sep_Sequence of S, a being FinSequence of ExtREAL such that A12: G,a are_Re-presentation_of g and A13: a . 1 = 0 and A14: for n being Nat st 2 <= n & n in dom a holds ( 0 < a . n & a . n < +infty ) by A1, MESFUNC3:14; defpred S1[ Nat, set ] means $2 = a . $1; A15: for k being Nat st k in Seg (len a) holds ex x being Element of REAL st S1[k,x] proof let k be Nat; ::_thesis: ( k in Seg (len a) implies ex x being Element of REAL st S1[k,x] ) assume A16: k in Seg (len a) ; ::_thesis: ex x being Element of REAL st S1[k,x] then A17: 1 <= k by FINSEQ_1:1; A18: k in dom a by A16, FINSEQ_1:def_3; percases ( k = 1 or k <> 1 ) ; supposeA19: k = 1 ; ::_thesis: ex x being Element of REAL st S1[k,x] take 0 ; ::_thesis: S1[k, 0 ] thus S1[k, 0 ] by A13, A19; ::_thesis: verum end; suppose k <> 1 ; ::_thesis: ex x being Element of REAL st S1[k,x] then k > 1 by A17, XXREAL_0:1; then A20: k >= 1 + 1 by NAT_1:13; then A21: a . k < +infty by A14, A18; 0 < a . k by A14, A18, A20; then reconsider x = a . k as Element of REAL by A21, XXREAL_0:14; take x ; ::_thesis: S1[k,x] thus S1[k,x] ; ::_thesis: verum end; end; end; consider a1 being FinSequence of REAL such that A22: ( dom a1 = Seg (len a) & ( for k being Nat st k in Seg (len a) holds S1[k,a1 . k] ) ) from FINSEQ_1:sch_5(A15); A23: len a <> 0 proof assume len a = 0 ; ::_thesis: contradiction then A24: dom a = Seg 0 by FINSEQ_1:def_3; A25: rng G = {} proof assume rng G <> {} ; ::_thesis: contradiction then consider y being set such that A26: y in rng G by XBOOLE_0:def_1; ex x being set st ( x in dom G & y = G . x ) by A26, FUNCT_1:def_3; hence contradiction by A12, A24, MESFUNC3:def_1; ::_thesis: verum end; union (rng G) <> {} by A11, A12, MESFUNC3:def_1; then consider x being set such that A27: x in union (rng G) by XBOOLE_0:def_1; ex Y being set st ( x in Y & Y in rng G ) by A27, TARSKI:def_4; hence contradiction by A25; ::_thesis: verum end; A28: 2 <= len a proof assume not 2 <= len a ; ::_thesis: contradiction then len a = 1 by A23, NAT_1:23; then dom a = {1} by FINSEQ_1:2, FINSEQ_1:def_3; then A29: dom G = {1} by A12, MESFUNC3:def_1; A30: dom g = union (rng G) by A12, MESFUNC3:def_1 .= union {(G . 1)} by A29, FUNCT_1:4 .= G . 1 by ZFMISC_1:25 ; then consider x being set such that A31: x in G . 1 by A11, XBOOLE_0:def_1; 1 in dom G by A29, TARSKI:def_1; then g . x = 0 by A12, A13, A31, MESFUNC3:def_1; hence contradiction by A2, A30, A31; ::_thesis: verum end; then 1 <= len a by XXREAL_0:2; then 1 in Seg (len a) ; then A32: a . 1 = a1 . 1 by A22; A33: 2 in dom a1 by A22, A28; then A34: 2 in dom a by A22, FINSEQ_1:def_3; a1 . 2 = a . 2 by A22, A33; then a1 . 2 <> a . 1 by A13, A14, A34; then A35: not a1 . 2 in {(a1 . 1)} by A32, TARSKI:def_1; a1 . 2 in rng a1 by A33, FUNCT_1:3; then reconsider RINGA = (rng a1) \ {(a1 . 1)} as non empty finite real-membered set by A35, XBOOLE_0:def_5; set alpha = R_EAL (min RINGA); reconsider beta1 = max RINGA as Real by XREAL_0:def_1; A36: min RINGA in RINGA by XXREAL_2:def_7; then min RINGA in rng a1 by XBOOLE_0:def_5; then consider i being set such that A37: i in dom a1 and A38: min RINGA = a1 . i by FUNCT_1:def_3; reconsider i = i as Element of NAT by A37; A39: a . i = R_EAL (a1 . i) by A22, A37; i in Seg (len a1) by A37, FINSEQ_1:def_3; then A40: 1 <= i by FINSEQ_1:1; not min RINGA in {(a1 . 1)} by A36, XBOOLE_0:def_5; then i <> 1 by A38, TARSKI:def_1; then 1 < i by A40, XXREAL_0:1; then A41: 1 + 1 <= i by NAT_1:13; A42: i in dom a by A22, A37, FINSEQ_1:def_3; then A43: 0 < R_EAL (min RINGA) by A14, A38, A41, A39; set beta = R_EAL (max RINGA); A44: for x being set st x in dom g holds ( R_EAL (min RINGA) <= g . x & g . x <= R_EAL (max RINGA) ) proof let x be set ; ::_thesis: ( x in dom g implies ( R_EAL (min RINGA) <= g . x & g . x <= R_EAL (max RINGA) ) ) assume A45: x in dom g ; ::_thesis: ( R_EAL (min RINGA) <= g . x & g . x <= R_EAL (max RINGA) ) then x in union (rng G) by A12, MESFUNC3:def_1; then consider Y being set such that A46: x in Y and A47: Y in rng G by TARSKI:def_4; consider k being set such that A48: k in dom G and A49: Y = G . k by A47, FUNCT_1:def_3; reconsider k = k as Element of NAT by A48; k in dom a by A12, A48, MESFUNC3:def_1; then A50: k in Seg (len a) by FINSEQ_1:def_3; now__::_thesis:_not_a1_._k_=_a1_._1 1 <= len a by A28, XXREAL_0:2; then A51: 1 in dom a1 by A22; A52: g . x = a . k by A12, A46, A48, A49, MESFUNC3:def_1; assume A53: a1 . k = a1 . 1 ; ::_thesis: contradiction a . k = a1 . k by A22, A50; then a . k = a . 1 by A22, A53, A51; hence contradiction by A2, A13, A45, A52; ::_thesis: verum end; then A54: not a1 . k in {(a1 . 1)} by TARSKI:def_1; a1 . k in rng a1 by A22, A50, FUNCT_1:3; then A55: a1 . k in RINGA by A54, XBOOLE_0:def_5; g . x = a . k by A12, A46, A48, A49, MESFUNC3:def_1 .= a1 . k by A22, A50 ; hence ( R_EAL (min RINGA) <= g . x & g . x <= R_EAL (max RINGA) ) by A55, XXREAL_2:def_7, XXREAL_2:def_8; ::_thesis: verum end; A56: for n being Nat holds dom (g - (F . n)) = dom g proof g is () by A1, Th14; then not -infty in rng g by Def3; then A57: g " {-infty} = {} by FUNCT_1:72; g is () by A1, Th14; then not +infty in rng g by Def4; then A58: g " {+infty} = {} by FUNCT_1:72; let n be Nat; ::_thesis: dom (g - (F . n)) = dom g A59: dom (g - (F . n)) = ((dom (F . n)) /\ (dom g)) \ ((((F . n) " {+infty}) /\ (g " {+infty})) \/ (((F . n) " {-infty}) /\ (g " {-infty}))) by MESFUNC1:def_4; dom (F . n) = dom g by A4; hence dom (g - (F . n)) = dom g by A58, A57, A59; ::_thesis: verum end; A60: g is V60() by A1, MESFUNC2:def_4; A61: for e being R_eal st R_EAL 0 < e & e < R_EAL (min RINGA) holds ex H being SetSequence of X ex MF being ExtREAL_sequence st ( ( for n being Nat holds H . n = less_dom ((g - (F . n)),e) ) & ( for n, m being Nat st n <= m holds H . n c= H . m ) & ( for n being Nat holds H . n c= dom g ) & ( for n being Nat holds MF . n = M . (H . n) ) & M . (dom g) = sup (rng MF) & ( for n being Nat holds H . n in S ) ) proof let e be R_eal; ::_thesis: ( R_EAL 0 < e & e < R_EAL (min RINGA) implies ex H being SetSequence of X ex MF being ExtREAL_sequence st ( ( for n being Nat holds H . n = less_dom ((g - (F . n)),e) ) & ( for n, m being Nat st n <= m holds H . n c= H . m ) & ( for n being Nat holds H . n c= dom g ) & ( for n being Nat holds MF . n = M . (H . n) ) & M . (dom g) = sup (rng MF) & ( for n being Nat holds H . n in S ) ) ) assume that A62: R_EAL 0 < e and A63: e < R_EAL (min RINGA) ; ::_thesis: ex H being SetSequence of X ex MF being ExtREAL_sequence st ( ( for n being Nat holds H . n = less_dom ((g - (F . n)),e) ) & ( for n, m being Nat st n <= m holds H . n c= H . m ) & ( for n being Nat holds H . n c= dom g ) & ( for n being Nat holds MF . n = M . (H . n) ) & M . (dom g) = sup (rng MF) & ( for n being Nat holds H . n in S ) ) deffunc H1( Nat) -> Element of bool X = less_dom ((g - (F . $1)),e); consider H being SetSequence of X such that A64: for n being Element of NAT holds H . n = H1(n) from FUNCT_2:sch_4(); A65: now__::_thesis:_for_n_being_Nat_holds_H_._n_=_H1(n) let n be Nat; ::_thesis: H . n = H1(n) n in NAT by ORDINAL1:def_12; hence H . n = H1(n) by A64; ::_thesis: verum end; A66: for n being Nat holds H . n c= dom g proof let n be Nat; ::_thesis: H . n c= dom g now__::_thesis:_for_x_being_set_st_x_in_H_._n_holds_ x_in_dom_g let x be set ; ::_thesis: ( x in H . n implies x in dom g ) assume x in H . n ; ::_thesis: x in dom g then x in less_dom ((g - (F . n)),e) by A65; then x in dom (g - (F . n)) by MESFUNC1:def_11; hence x in dom g by A56; ::_thesis: verum end; hence H . n c= dom g by TARSKI:def_3; ::_thesis: verum end; A67: Union H c= dom g proof let x be set ; :: according to TARSKI:def_3 ::_thesis: ( not x in Union H or x in dom g ) assume x in Union H ; ::_thesis: x in dom g then consider n being Element of NAT such that A68: x in H . n by PROB_1:12; H . n c= dom g by A66; hence x in dom g by A68; ::_thesis: verum end; now__::_thesis:_for_x_being_set_st_x_in_dom_g_holds_ x_in_lim_inf_H let x be set ; ::_thesis: ( x in dom g implies x in lim_inf H ) assume A69: x in dom g ; ::_thesis: x in lim_inf H then reconsider x1 = x as Element of X ; A70: F # x1 is convergent by A7, A69; A71: now__::_thesis:_not_F_#_x1_is_convergent_to_-infty reconsider E = e as Real by A62, A63, XXREAL_0:48; assume F # x1 is convergent_to_-infty ; ::_thesis: contradiction then consider N being Nat such that A72: for m being Nat st N <= m holds (F # x1) . m <= R_EAL (- E) by A62, Def10; F . N is nonnegative by A5; then A73: 0 <= (F . N) . x by SUPINF_2:51; (F # x1) . N < 0 by A62, A72; hence contradiction by A73, Def13; ::_thesis: verum end; now__::_thesis:_ex_N_being_Nat_st_ (_N_in_dom_(inferior_setsequence_H)_&_x_in_(inferior_setsequence_H)_._N_) percases ( F # x1 is convergent_to_finite_number or F # x1 is convergent_to_+infty ) by A70, A71, Def11; supposeA74: F # x1 is convergent_to_finite_number ; ::_thesis: ex N being Nat st ( N in dom (inferior_setsequence H) & x in (inferior_setsequence H) . N ) reconsider E = e as Real by A62, A63, XXREAL_0:48; A75: ( ex limFx being real number st ( lim (F # x1) = limFx & ( for p being real number st 0 < p holds ex n being Nat st for m being Nat st n <= m holds |.(((F # x1) . m) - (lim (F # x1))).| < p ) & F # x1 is convergent_to_finite_number ) or ( lim (F # x1) = +infty & F # x1 is convergent_to_+infty ) or ( lim (F # x1) = -infty & F # x1 is convergent_to_-infty ) ) by A70, Def12; then consider N being Nat such that A76: for m being Nat st N <= m holds |.(((F # x1) . m) - (lim (F # x1))).| < R_EAL (E / 2) by A62, A74, Th50, Th51; reconsider N = N as Element of NAT by ORDINAL1:def_12; g . x <= lim (F # x1) by A7, A69; then (g . x) - (R_EAL (E / 2)) <= (lim (F # x1)) - 0. by A62, XXREAL_3:37; then A77: (g . x) - (R_EAL (E / 2)) <= lim (F # x1) by XXREAL_3:4; now__::_thesis:_for_k_being_Element_of_NAT_holds_x_in_H_._(N_+_k) let k be Element of NAT ; ::_thesis: x in H . (N + k) set m = N + k; A78: x1 in dom (g - (F . (N + k))) by A56, A69; now__::_thesis:_for_e_being_real_number_st_0_<_e_holds_ (F_#_x1)_._(N_+_k)_<_(lim_(F_#_x1))_+_e let e be real number ; ::_thesis: ( 0 < e implies (F # x1) . (N + k) < (lim (F # x1)) + e ) assume 0 < e ; ::_thesis: (F # x1) . (N + k) < (lim (F # x1)) + e then consider N0 being Nat such that A79: for M being Nat st N0 <= M holds |.(((F # x1) . M) - (lim (F # x1))).| < R_EAL e by A74, A75, Th50, Th51; reconsider N0 = N0, n1 = N + k as Element of NAT by ORDINAL1:def_12; set M = max (N0,n1); A80: ((F # x1) . (max (N0,n1))) - (lim (F # x1)) <= |.(((F # x1) . (max (N0,n1))) - (lim (F # x1))).| by EXTREAL2:9; (F . (N + k)) . x1 <= (F . (max (N0,n1))) . x1 by A6, A69, XXREAL_0:25; then (F . (N + k)) . x1 <= (F # x1) . (max (N0,n1)) by Def13; then A81: (F # x1) . (N + k) <= (F # x1) . (max (N0,n1)) by Def13; |.(((F # x1) . (max (N0,n1))) - (lim (F # x1))).| < R_EAL e by A79, XXREAL_0:25; then ((F # x1) . (max (N0,n1))) - (lim (F # x1)) < R_EAL e by A80, XXREAL_0:2; then (F # x1) . (max (N0,n1)) < (R_EAL e) + (lim (F # x1)) by A74, A75, Th50, Th51, XXREAL_3:54; hence (F # x1) . (N + k) < (lim (F # x1)) + e by A81, XXREAL_0:2; ::_thesis: verum end; then (F # x1) . (N + k) <= lim (F # x1) by XXREAL_3:61; then A82: 0 <= (lim (F # x1)) - ((F # x1) . (N + k)) by XXREAL_3:40; |.(((F # x1) . (N + k)) - (lim (F # x1))).| = |.((lim (F # x1)) - ((F # x1) . (N + k))).| by Th1 .= (lim (F # x1)) - ((F # x1) . (N + k)) by A82, EXTREAL1:def_1 ; then (lim (F # x1)) - ((F # x1) . (N + k)) < R_EAL (E / 2) by A76, NAT_1:11; then A83: (lim (F # x1)) - ((F . (N + k)) . x1) < R_EAL (E / 2) by Def13; A84: |.(((F # x1) . (N + k)) - (lim (F # x1))).| < R_EAL (E / 2) by A76, NAT_1:11; then (F # x1) . (N + k) <> -infty by A74, A75, Th3, Th51; then A85: (F . (N + k)) . x <> -infty by Def13; (F # x1) . (N + k) <> +infty by A74, A75, A84, Th3, Th50; then (F . (N + k)) . x <> +infty by Def13; then lim (F # x1) < ((F . (N + k)) . x) + (R_EAL (E / 2)) by A85, A83, XXREAL_3:54; then A86: (lim (F # x1)) + (E / 2) < (((F . (N + k)) . x) + (R_EAL (E / 2))) + (E / 2) by XXREAL_3:62; g . x <= (lim (F # x1)) + (R_EAL (E / 2)) by A77, XXREAL_3:41; then g . x < (((F . (N + k)) . x1) + (R_EAL (E / 2))) + (R_EAL (E / 2)) by A86, XXREAL_0:2; then g . x < ((F . (N + k)) . x1) + ((R_EAL (E / 2)) + (R_EAL (E / 2))) by XXREAL_3:29; then g . x < ((F . (N + k)) . x1) + (R_EAL ((E / 2) + (E / 2))) by SUPINF_2:1; then (g . x) - ((F . (N + k)) . x1) < e by XXREAL_3:55; then (g - (F . (N + k))) . x1 < e by A78, MESFUNC1:def_4; then x in less_dom ((g - (F . (N + k))),e) by A78, MESFUNC1:def_11; hence x in H . (N + k) by A65; ::_thesis: verum end; then A87: x in (inferior_setsequence H) . N by SETLIM_1:19; dom (inferior_setsequence H) = NAT by FUNCT_2:def_1; hence ex N being Nat st ( N in dom (inferior_setsequence H) & x in (inferior_setsequence H) . N ) by A87; ::_thesis: verum end; supposeA88: F # x1 is convergent_to_+infty ; ::_thesis: ex N being Nat st ( N in dom (inferior_setsequence H) & x in (inferior_setsequence H) . N ) ex N being Nat st for m being Nat st N <= m holds (g . x1) - ((F . m) . x1) < e proof A89: e in REAL by A62, A63, XXREAL_0:48; percases ( (g . x1) - e <= 0 or 0 < (g . x1) - e ) ; supposeA90: (g . x1) - e <= 0 ; ::_thesis: ex N being Nat st for m being Nat st N <= m holds (g . x1) - ((F . m) . x1) < e consider N being Nat such that A91: for m being Nat st N <= m holds R_EAL 1 <= (F # x1) . m by A88, Def9; now__::_thesis:_for_m_being_Nat_st_N_<=_m_holds_ (g_._x1)_-_((F_._m)_._x1)_<_e let m be Nat; ::_thesis: ( N <= m implies (g . x1) - ((F . m) . x1) < e ) assume N <= m ; ::_thesis: (g . x1) - ((F . m) . x1) < e then (g . x1) - e < (F # x1) . m by A90, A91; then g . x1 < ((F # x1) . m) + e by A89, XXREAL_3:54; then (g . x1) - ((F # x1) . m) < e by A89, XXREAL_3:55; hence (g . x1) - ((F . m) . x1) < e by Def13; ::_thesis: verum end; hence ex N being Nat st for m being Nat st N <= m holds (g . x1) - ((F . m) . x1) < e ; ::_thesis: verum end; supposeA92: 0 < (g . x1) - e ; ::_thesis: ex N being Nat st for m being Nat st N <= m holds (g . x1) - ((F . m) . x1) < e reconsider e1 = e as Real by A62, A63, XXREAL_0:48; reconsider gx1 = g . x as Real by A60, XXREAL_0:14; (g . x) - e = gx1 - e1 by SUPINF_2:3; then reconsider ee = (g . x1) - e as Real ; consider N being Nat such that A93: for m being Nat st N <= m holds R_EAL (ee + 1) <= (F # x1) . m by A88, A92, Def9; A94: R_EAL ee < R_EAL (ee + 1) by XREAL_1:29; now__::_thesis:_for_m_being_Nat_st_N_<=_m_holds_ (g_._x1)_-_((F_._m)_._x1)_<_e let m be Nat; ::_thesis: ( N <= m implies (g . x1) - ((F . m) . x1) < e ) assume N <= m ; ::_thesis: (g . x1) - ((F . m) . x1) < e then R_EAL (ee + 1) <= (F # x1) . m by A93; then R_EAL ee < (F # x1) . m by A94, XXREAL_0:2; then g . x1 < ((F # x1) . m) + e by A89, XXREAL_3:54; then (g . x1) - ((F # x1) . m) < e by A89, XXREAL_3:55; hence (g . x1) - ((F . m) . x1) < e by Def13; ::_thesis: verum end; hence ex N being Nat st for m being Nat st N <= m holds (g . x1) - ((F . m) . x1) < e ; ::_thesis: verum end; end; end; then consider N being Nat such that A95: for m being Nat st N <= m holds (g . x1) - ((F . m) . x1) < e ; reconsider N = N as Element of NAT by ORDINAL1:def_12; A96: now__::_thesis:_for_m_being_Nat_st_N_<=_m_holds_ x1_in_less_dom_((g_-_(F_._m)),e) let m be Nat; ::_thesis: ( N <= m implies x1 in less_dom ((g - (F . m)),e) ) A97: x1 in dom (g - (F . m)) by A56, A69; assume N <= m ; ::_thesis: x1 in less_dom ((g - (F . m)),e) then (g . x1) - ((F . m) . x1) < e by A95; then (g - (F . m)) . x1 < e by A97, MESFUNC1:def_4; hence x1 in less_dom ((g - (F . m)),e) by A97, MESFUNC1:def_11; ::_thesis: verum end; now__::_thesis:_for_k_being_Element_of_NAT_holds_x_in_H_._(N_+_k) let k be Element of NAT ; ::_thesis: x in H . (N + k) x in less_dom ((g - (F . (N + k))),e) by A96, NAT_1:11; hence x in H . (N + k) by A65; ::_thesis: verum end; then A98: x in (inferior_setsequence H) . N by SETLIM_1:19; dom (inferior_setsequence H) = NAT by FUNCT_2:def_1; hence ex N being Nat st ( N in dom (inferior_setsequence H) & x in (inferior_setsequence H) . N ) by A98; ::_thesis: verum end; end; end; then consider N being Nat such that A99: N in dom (inferior_setsequence H) and A100: x in (inferior_setsequence H) . N ; (inferior_setsequence H) . N in rng (inferior_setsequence H) by A99, FUNCT_1:3; then x in Union (inferior_setsequence H) by A100, TARSKI:def_4; hence x in lim_inf H by SETLIM_1:def_4; ::_thesis: verum end; then A101: dom g c= lim_inf H by TARSKI:def_3; deffunc H2( Nat) -> Element of ExtREAL = M . (H . $1); A102: lim_inf H c= lim_sup H by KURATO_0:6; consider MF being ExtREAL_sequence such that A103: for n being Element of NAT holds MF . n = H2(n) from FUNCT_2:sch_4(); A104: for n, m being Nat st n <= m holds H . n c= H . m proof let n, m be Nat; ::_thesis: ( n <= m implies H . n c= H . m ) assume A105: n <= m ; ::_thesis: H . n c= H . m now__::_thesis:_for_x_being_set_st_x_in_H_._n_holds_ x_in_H_._m let x be set ; ::_thesis: ( x in H . n implies x in H . m ) assume x in H . n ; ::_thesis: x in H . m then A106: x in less_dom ((g - (F . n)),e) by A65; then A107: x in dom (g - (F . n)) by MESFUNC1:def_11; then A108: (g - (F . n)) . x = (g . x) - ((F . n) . x) by MESFUNC1:def_4; A109: (g - (F . n)) . x < e by A106, MESFUNC1:def_11; A110: dom (g - (F . n)) = dom g by A56; then A111: (F . n) . x <= (F . m) . x by A6, A105, A107; A112: dom (g - (F . m)) = dom g by A56; then (g - (F . m)) . x = (g . x) - ((F . m) . x) by A107, A110, MESFUNC1:def_4; then (g - (F . m)) . x <= (g - (F . n)) . x by A108, A111, XXREAL_3:37; then (g - (F . m)) . x < e by A109, XXREAL_0:2; then x in less_dom ((g - (F . m)),e) by A107, A110, A112, MESFUNC1:def_11; hence x in H . m by A65; ::_thesis: verum end; hence H . n c= H . m by TARSKI:def_3; ::_thesis: verum end; then for n, m being Element of NAT st n <= m holds H . n c= H . m ; then A113: H is non-descending by PROB_1:def_5; A114: now__::_thesis:_for_n_being_Nat_holds_MF_._n_=_H2(n) let n be Nat; ::_thesis: MF . n = H2(n) n in NAT by ORDINAL1:def_12; hence MF . n = H2(n) by A103; ::_thesis: verum end; now__::_thesis:_for_x_being_set_st_x_in_lim_inf_H_holds_ x_in_dom_g let x be set ; ::_thesis: ( x in lim_inf H implies x in dom g ) assume x in lim_inf H ; ::_thesis: x in dom g then x in Union (inferior_setsequence H) by SETLIM_1:def_4; then consider V being set such that A115: x in V and A116: V in rng (inferior_setsequence H) by TARSKI:def_4; consider n being set such that A117: n in dom (inferior_setsequence H) and A118: V = (inferior_setsequence H) . n by A116, FUNCT_1:def_3; reconsider n = n as Element of NAT by A117; x in H . (n + 0) by A115, A118, SETLIM_1:19; then x in less_dom ((g - (F . n)),e) by A65; then x in dom (g - (F . n)) by MESFUNC1:def_11; hence x in dom g by A56; ::_thesis: verum end; then lim_inf H c= dom g by TARSKI:def_3; then A119: lim_inf H = dom g by A101, XBOOLE_0:def_10; A120: ( M . (dom g) = sup (rng MF) & ( for n being Element of NAT holds H . n in S ) ) proof A121: now__::_thesis:_for_x_being_set_st_x_in_NAT_holds_ H_._x_in_S reconsider E = e as Real by A62, A63, XXREAL_0:48; let x be set ; ::_thesis: ( x in NAT implies H . x in S ) assume x in NAT ; ::_thesis: H . x in S then reconsider n = x as Element of NAT ; A122: less_dom ((g - (F . n)),(R_EAL E)) c= dom (g - (F . n)) proof let x be set ; :: according to TARSKI:def_3 ::_thesis: ( not x in less_dom ((g - (F . n)),(R_EAL E)) or x in dom (g - (F . n)) ) assume x in less_dom ((g - (F . n)),(R_EAL E)) ; ::_thesis: x in dom (g - (F . n)) hence x in dom (g - (F . n)) by MESFUNC1:def_11; ::_thesis: verum end; A123: F . n is_simple_func_in S by A3; then consider GF being Finite_Sep_Sequence of S such that A124: dom (F . n) = union (rng GF) and for m being Nat for x, y being Element of X st m in dom GF & x in GF . m & y in GF . m holds (F . n) . x = (F . n) . y by MESFUNC2:def_4; A125: F . n is V60() by A123, MESFUNC2:def_4; reconsider DGH = union (rng GF) as Element of S by MESFUNC2:31; dom (F . n) = dom g by A4; then DGH /\ (less_dom ((g - (F . n)),(R_EAL E))) = (dom (g - (F . n))) /\ (less_dom ((g - (F . n)),(R_EAL E))) by A56, A124; then A126: DGH /\ (less_dom ((g - (F . n)),(R_EAL E))) = less_dom ((g - (F . n)),(R_EAL E)) by A122, XBOOLE_1:28; A127: F . n is_measurable_on DGH by A3, MESFUNC2:34; A128: g is V60() by A1, MESFUNC2:def_4; g is_measurable_on DGH by A1, MESFUNC2:34; then g - (F . n) is_measurable_on DGH by A124, A128, A125, A127, MESFUNC2:11; then DGH /\ (less_dom ((g - (F . n)),(R_EAL E))) in S by MESFUNC1:def_16; hence H . x in S by A65, A126; ::_thesis: verum end; dom H = NAT by FUNCT_2:def_1; then reconsider HH = H as Function of NAT,S by A121, FUNCT_2:3; A129: for n being Element of NAT holds HH . n c= HH . (n + 1) by A104, NAT_1:11; rng HH c= S by RELAT_1:def_19; then A130: rng H c= dom M by FUNCT_2:def_1; lim_sup H = Union H by A113, SETLIM_1:59; then A131: M . (union (rng H)) = M . (dom g) by A119, A67, A102, XBOOLE_0:def_10; A132: dom H = NAT by FUNCT_2:def_1; A133: dom MF = NAT by FUNCT_2:def_1; A134: for x being set holds ( x in dom MF iff ( x in dom H & H . x in dom M ) ) proof let x be set ; ::_thesis: ( x in dom MF iff ( x in dom H & H . x in dom M ) ) now__::_thesis:_(_x_in_dom_MF_implies_(_x_in_dom_H_&_H_._x_in_dom_M_)_) assume A135: x in dom MF ; ::_thesis: ( x in dom H & H . x in dom M ) then H . x in rng H by A132, FUNCT_1:3; hence ( x in dom H & H . x in dom M ) by A132, A130, A135; ::_thesis: verum end; hence ( x in dom MF iff ( x in dom H & H . x in dom M ) ) by A133; ::_thesis: verum end; for x being set st x in dom MF holds MF . x = M . (H . x) by A103; then M * H = MF by A134, FUNCT_1:10; hence ( M . (dom g) = sup (rng MF) & ( for n being Element of NAT holds H . n in S ) ) by A121, A129, A131, MEASURE2:23; ::_thesis: verum end; now__::_thesis:_for_n_being_Nat_holds_H_._n_in_S let n be Nat; ::_thesis: H . n in S n in NAT by ORDINAL1:def_12; hence H . n in S by A120; ::_thesis: verum end; hence ex H being SetSequence of X ex MF being ExtREAL_sequence st ( ( for n being Nat holds H . n = less_dom ((g - (F . n)),e) ) & ( for n, m being Nat st n <= m holds H . n c= H . m ) & ( for n being Nat holds H . n c= dom g ) & ( for n being Nat holds MF . n = M . (H . n) ) & M . (dom g) = sup (rng MF) & ( for n being Nat holds H . n in S ) ) by A65, A104, A66, A114, A120; ::_thesis: verum end; percases ( M . (dom g) <> +infty or M . (dom g) = +infty ) ; supposeA136: M . (dom g) <> +infty ; ::_thesis: ( L is convergent & integral' (M,g) <= lim L ) A137: 0 < R_EAL (max RINGA) proof consider x being set such that A138: x in dom g by A11, XBOOLE_0:def_1; A139: g . x <= R_EAL (max RINGA) by A44, A138; R_EAL (min RINGA) <= g . x by A44, A138; hence 0 < R_EAL (max RINGA) by A14, A38, A41, A42, A39, A139; ::_thesis: verum end; A140: {} in S by MEASURE1:34; A141: M . {} = 0 by VALUED_0:def_19; dom g is Element of S by A1, Th37; then A142: M . (dom g) <> -infty by A141, A140, MEASURE1:31, XBOOLE_1:2; then reconsider MG = M . (dom g) as Real by A136, XXREAL_0:14; reconsider DG = dom g as Element of S by A1, Th37; A143: for x being set st x in dom g holds 0 <= g . x by A2; then A144: integral' (M,g) <> -infty by A1, Th68, SUPINF_2:52; A145: g is nonnegative by A143, SUPINF_2:52; A146: integral' (M,g) <= (R_EAL (max RINGA)) * (M . DG) proof consider GP being PartFunc of X,ExtREAL such that A147: GP is_simple_func_in S and A148: dom GP = DG and A149: for x being set st x in DG holds GP . x = R_EAL (max RINGA) by Th41; A150: for x being set st x in dom (GP - g) holds g . x <= GP . x proof let x be set ; ::_thesis: ( x in dom (GP - g) implies g . x <= GP . x ) assume x in dom (GP - g) ; ::_thesis: g . x <= GP . x then x in ((dom GP) /\ (dom g)) \ (((GP " {+infty}) /\ (g " {+infty})) \/ ((GP " {-infty}) /\ (g " {-infty}))) by MESFUNC1:def_4; then A151: x in (dom GP) /\ (dom g) by XBOOLE_0:def_5; then GP . x = R_EAL (max RINGA) by A148, A149; hence g . x <= GP . x by A44, A148, A151; ::_thesis: verum end; for x being set st x in dom GP holds 0 <= GP . x by A137, A148, A149; then A152: GP is nonnegative by SUPINF_2:52; then A153: dom (GP - g) = (dom GP) /\ (dom g) by A1, A145, A147, A150, Th69; then A154: g | (dom (GP - g)) = g by A148, GRFUNC_1:23; A155: GP | (dom (GP - g)) = GP by A148, A153, GRFUNC_1:23; integral' (M,(g | (dom (GP - g)))) <= integral' (M,(GP | (dom (GP - g)))) by A1, A145, A147, A152, A150, Th70; hence integral' (M,g) <= (R_EAL (max RINGA)) * (M . DG) by A137, A147, A148, A149, A154, A155, Th71; ::_thesis: verum end; (R_EAL (max RINGA)) * (M . DG) = beta1 * MG by EXTREAL1:1; then A156: integral' (M,g) <> +infty by A146, XXREAL_0:9; A157: for e being R_eal st 0 < e & e < R_EAL (min RINGA) holds ex N0 being Nat st for n being Nat st N0 <= n holds (integral' (M,g)) - (e * ((R_EAL (max RINGA)) + (M . (dom g)))) < integral' (M,(F . n)) proof let e be R_eal; ::_thesis: ( 0 < e & e < R_EAL (min RINGA) implies ex N0 being Nat st for n being Nat st N0 <= n holds (integral' (M,g)) - (e * ((R_EAL (max RINGA)) + (M . (dom g)))) < integral' (M,(F . n)) ) assume that A158: 0 < e and A159: e < R_EAL (min RINGA) ; ::_thesis: ex N0 being Nat st for n being Nat st N0 <= n holds (integral' (M,g)) - (e * ((R_EAL (max RINGA)) + (M . (dom g)))) < integral' (M,(F . n)) A160: e <> +infty by A159, XXREAL_0:4; consider H being SetSequence of X, MF being ExtREAL_sequence such that A161: for n being Nat holds H . n = less_dom ((g - (F . n)),e) and A162: for n, m being Nat st n <= m holds H . n c= H . m and A163: for n being Nat holds H . n c= dom g and A164: for n being Nat holds MF . n = M . (H . n) and A165: M . (dom g) = sup (rng MF) and A166: for n being Nat holds H . n in S by A61, A158, A159; sup (rng MF) is Real by A136, A142, A165, XXREAL_0:14; then consider y being ext-real number such that A167: y in rng MF and A168: (sup (rng MF)) - e < y by A158, MEASURE6:6; consider N0 being set such that A169: N0 in dom MF and A170: y = MF . N0 by A167, FUNCT_1:def_3; reconsider N0 = N0 as Element of NAT by A169; reconsider B0 = H . N0 as Element of S by A166; M . B0 <= M . DG by A163, MEASURE1:31; then M . B0 < +infty by A136, XXREAL_0:2, XXREAL_0:4; then A171: M . (DG \ B0) = (M . DG) - (M . B0) by A163, MEASURE1:32; take N0 ; ::_thesis: for n being Nat st N0 <= n holds (integral' (M,g)) - (e * ((R_EAL (max RINGA)) + (M . (dom g)))) < integral' (M,(F . n)) (M . (dom g)) - e < M . (H . N0) by A164, A165, A168, A170; then M . (dom g) < (M . (H . N0)) + e by A158, A160, XXREAL_3:54; then A172: (M . (dom g)) - (M . (H . N0)) < e by A158, A160, XXREAL_3:55; A173: now__::_thesis:_for_n_being_Nat_st_N0_<=_n_holds_ M_._((dom_g)_\_(H_._n))_<_e let n be Nat; ::_thesis: ( N0 <= n implies M . ((dom g) \ (H . n)) < e ) reconsider BN = H . n as Element of S by A166; assume N0 <= n ; ::_thesis: M . ((dom g) \ (H . n)) < e then H . N0 c= H . n by A162; then M . (DG \ BN) <= M . (DG \ B0) by MEASURE1:31, XBOOLE_1:34; hence M . ((dom g) \ (H . n)) < e by A172, A171, XXREAL_0:2; ::_thesis: verum end; now__::_thesis:_for_n_being_Nat_st_N0_<=_n_holds_ (integral'_(M,g))_-_(e_*_((R_EAL_(max_RINGA))_+_(M_._(dom_g))))_<_integral'_(M,(F_._n)) reconsider XSMg = integral' (M,g) as Real by A156, A144, XXREAL_0:14; let n be Nat; ::_thesis: ( N0 <= n implies (integral' (M,g)) - (e * ((R_EAL (max RINGA)) + (M . (dom g)))) < integral' (M,(F . n)) ) A174: for x being set st x in dom (F . n) holds (F . n) . x = (F . n) . x ; reconsider B = H . n as Element of S by A166; H . n in S by A166; then X \ (H . n) in S by MEASURE1:34; then A175: DG /\ (X \ (H . n)) in S by MEASURE1:34; DG /\ (X \ (H . n)) = (DG /\ X) \ (H . n) by XBOOLE_1:49; then reconsider A = DG \ (H . n) as Element of S by A175, XBOOLE_1:28; e <> +infty by A159, XXREAL_0:4; then reconsider ee = e as Real by A158, XXREAL_0:14; A176: A misses B by XBOOLE_1:79; (R_EAL (max RINGA)) * e = beta1 * ee by EXTREAL1:1; then reconsider betae = (R_EAL (max RINGA)) * e as Real ; A177: for x being set st x in dom g holds g . x = g . x ; A178: M . B <= M . DG by A163, MEASURE1:31; then M . (dom g) <> -infty by A141, A140, MEASURE1:31, XBOOLE_1:2; then A179: M . (dom g) is Real by A136, XXREAL_0:14; A180: DG = DG \/ (H . n) by A163, XBOOLE_1:12; then A181: DG = (DG \ (H . n)) \/ (H . n) by XBOOLE_1:39; then dom g = (A \/ B) /\ (dom g) ; then g = g | (A \/ B) by A177, FUNCT_1:46; then A182: integral' (M,g) = (integral' (M,(g | A))) + (integral' (M,(g | B))) by A1, A145, Th67, XBOOLE_1:79; M . A <= M . DG by A181, MEASURE1:31, XBOOLE_1:7; then M . A < +infty by A136, XXREAL_0:2, XXREAL_0:4; then (R_EAL (max RINGA)) * (M . A) < (R_EAL (max RINGA)) * +infty by A137, XXREAL_3:72; then A183: (R_EAL (max RINGA)) * (M . A) <> +infty by A137, XXREAL_3:def_5; A184: g | B is nonnegative by A143, Th15, SUPINF_2:52; A185: dom (F . n) = dom g by A4; then dom (F . n) = (A \/ B) /\ (dom (F . n)) by A181; then A186: F . n = (F . n) | (A \/ B) by A174, FUNCT_1:46; consider GP being PartFunc of X,ExtREAL such that A187: GP is_simple_func_in S and A188: dom GP = A and A189: for x being set st x in A holds GP . x = R_EAL (max RINGA) by Th41; A190: integral' (M,GP) = (R_EAL (max RINGA)) * (M . A) by A137, A187, A188, A189, Th71; A191: dom (g | A) = A by A181, RELAT_1:62, XBOOLE_1:7; A192: for x being set st x in dom (GP - (g | A)) holds (g | A) . x <= GP . x proof let x be set ; ::_thesis: ( x in dom (GP - (g | A)) implies (g | A) . x <= GP . x ) assume x in dom (GP - (g | A)) ; ::_thesis: (g | A) . x <= GP . x then x in ((dom GP) /\ (dom (g | A))) \ (((GP " {+infty}) /\ ((g | A) " {+infty})) \/ ((GP " {-infty}) /\ ((g | A) " {-infty}))) by MESFUNC1:def_4; then A193: x in (dom GP) /\ (dom (g | A)) by XBOOLE_0:def_5; then A194: x in dom GP by XBOOLE_0:def_4; x in (dom g) /\ A by A191, A188, A193, RELAT_1:61; then x in dom g by XBOOLE_0:def_4; then A195: g . x <= R_EAL (max RINGA) by A44; (g | A) . x = g . x by A191, A188, A193, FUNCT_1:47; hence (g | A) . x <= GP . x by A188, A189, A194, A195; ::_thesis: verum end; for x being set st x in dom GP holds 0 <= GP . x by A137, A188, A189; then A196: GP is nonnegative by SUPINF_2:52; 0 <= M . A by A141, A140, MEASURE1:31, XBOOLE_1:2; then reconsider XSMGP = integral' (M,GP) as Real by A137, A190, A183, XXREAL_0:14; A197: (integral' (M,g)) - (integral' (M,GP)) = XSMg - XSMGP by SUPINF_2:3; A198: g | A is_simple_func_in S by A1, Th34; then A199: integral' (M,(g | A)) <> -infty by A145, Th15, Th68; A200: g | A is nonnegative by A143, Th15, SUPINF_2:52; then A201: dom (GP - (g | A)) = (dom GP) /\ (dom (g | A)) by A198, A187, A196, A192, Th69; then A202: GP | (dom (GP - (g | A))) = GP by A191, A188, GRFUNC_1:23; (g | A) | (dom (GP - (g | A))) = g | A by A191, A188, A201, GRFUNC_1:23; then A203: integral' (M,(g | A)) <= integral' (M,GP) by A198, A200, A187, A196, A192, A202, Th70; then A204: (integral' (M,g)) - (integral' (M,GP)) <= (integral' (M,g)) - (integral' (M,(g | A))) by XXREAL_3:37; assume N0 <= n ; ::_thesis: (integral' (M,g)) - (e * ((R_EAL (max RINGA)) + (M . (dom g)))) < integral' (M,(F . n)) then M . A < e by A173; then A205: (R_EAL (max RINGA)) * (M . A) < (R_EAL (max RINGA)) * e by A137, XXREAL_3:72; then A206: integral' (M,(g | A)) <> +infty by A203, A190, XXREAL_0:2, XXREAL_0:4; then reconsider XSMgA = integral' (M,(g | A)) as Real by A199, XXREAL_0:14; A207: integral' (M,(g | A)) is Element of REAL by A199, A206, XXREAL_0:14; XSMg - XSMgA = (integral' (M,g)) - (integral' (M,(g | A))) by SUPINF_2:3 .= integral' (M,(g | B)) by A182, A207, XXREAL_3:24 ; then reconsider XSMgB = integral' (M,(g | B)) as Real ; A208: H . n c= DG by A163; integral' (M,(g | A)) is Element of REAL by A199, A206, XXREAL_0:14; then A209: (integral' (M,g)) - (integral' (M,(g | A))) = integral' (M,(g | B)) by A182, XXREAL_3:24; XSMg - betae < XSMg - XSMGP by A190, A205, XREAL_1:15; then A210: XSMg - betae < XSMgB by A209, A204, A197, XXREAL_0:2; consider EP being PartFunc of X,ExtREAL such that A211: EP is_simple_func_in S and A212: dom EP = B and A213: for x being set st x in B holds EP . x = e by A158, A160, Th41; A214: integral' (M,EP) = e * (M . B) by A158, A211, A212, A213, Th71; for x being set st x in dom EP holds 0 <= EP . x by A158, A212, A213; then A215: EP is nonnegative by SUPINF_2:52; M . B < +infty by A136, A178, XXREAL_0:2, XXREAL_0:4; then e * (M . B) < e * +infty by A158, A160, XXREAL_3:72; then A216: integral' (M,EP) <> +infty by A214, XXREAL_0:4; A217: 0 <= M . B by A141, A140, MEASURE1:31, XBOOLE_1:2; then reconsider XSMEP = integral' (M,EP) as Real by A158, A214, A216, XXREAL_0:14; A218: F . n is_simple_func_in S by A3; (F . n) | A is nonnegative by A5, Th15; then A219: 0 <= integral' (M,((F . n) | A)) by A218, Th34, Th68; F . n is nonnegative by A5; then integral' (M,(F . n)) = (integral' (M,((F . n) | A))) + (integral' (M,((F . n) | B))) by A3, A186, A176, Th67; then A220: integral' (M,((F . n) | B)) <= integral' (M,(F . n)) by A219, XXREAL_3:39; A221: M . (dom g) < +infty by A136, XXREAL_0:4; M . B <> -infty by A141, A140, MEASURE1:31, XBOOLE_1:2; then M . B is Real by A221, A178, XXREAL_0:14; then consider MB, MG being Real such that A222: MB = M . B and A223: MG = M . (dom g) and A224: MB <= MG by A208, A179, MEASURE1:31; A225: g | B is_simple_func_in S by A1, Th34; ee * MB <= ee * MG by A158, A224, XREAL_1:64; then A226: (XSMg - betae) - (ee * MG) <= (XSMg - betae) - (ee * MB) by XREAL_1:13; XSMEP = e * (M . B) by A158, A211, A212, A213, Th71 .= ee * MB by A222, EXTREAL1:1 ; then A227: (XSMg - betae) - (ee * MB) < XSMgB - XSMEP by A210, XREAL_1:14; betae = ee * beta1 by EXTREAL1:1; then A228: XSMg - (ee * (beta1 + MG)) < XSMgB - XSMEP by A227, A226, XXREAL_0:2; dom ((F . n) | B) = (dom (F . n)) /\ B by RELAT_1:61; then A229: dom ((F . n) | B) = B by A163, A185, XBOOLE_1:28; A230: (F . n) | B is_simple_func_in S by A3, Th34; then A231: ((F . n) | B) + EP is_simple_func_in S by A211, Th38; A232: (F . n) | B is nonnegative by A5, Th15; then A233: dom (((F . n) | B) + EP) = (dom ((F . n) | B)) /\ (dom EP) by A230, A211, A215, Th65; A234: dom (((F . n) | B) + EP) = (dom ((F . n) | B)) /\ (dom EP) by A232, A230, A211, A215, Th65 .= B by A229, A212 ; A235: dom (g | B) = B by A180, RELAT_1:62, XBOOLE_1:7; A236: for x being set st x in dom ((((F . n) | B) + EP) - (g | B)) holds (g | B) . x <= (((F . n) | B) + EP) . x proof set f = g - (F . n); let x be set ; ::_thesis: ( x in dom ((((F . n) | B) + EP) - (g | B)) implies (g | B) . x <= (((F . n) | B) + EP) . x ) assume x in dom ((((F . n) | B) + EP) - (g | B)) ; ::_thesis: (g | B) . x <= (((F . n) | B) + EP) . x then x in ((dom (((F . n) | B) + EP)) /\ (dom (g | B))) \ ((((((F . n) | B) + EP) " {+infty}) /\ ((g | B) " {+infty})) \/ (((((F . n) | B) + EP) " {-infty}) /\ ((g | B) " {-infty}))) by MESFUNC1:def_4; then A237: x in (dom (((F . n) | B) + EP)) /\ (dom (g | B)) by XBOOLE_0:def_5; then A238: x in dom (((F . n) | B) + EP) by XBOOLE_0:def_4; then (((F . n) | B) + EP) . x = (((F . n) | B) . x) + (EP . x) by MESFUNC1:def_3; then (((F . n) | B) + EP) . x = ((F . n) . x) + (EP . x) by A229, A234, A238, FUNCT_1:47; then A239: (((F . n) | B) + EP) . x = ((F . n) . x) + e by A213, A234, A238; A240: x in less_dom ((g - (F . n)),e) by A161, A234, A238; then A241: (g - (F . n)) . x < e by MESFUNC1:def_11; x in dom (g - (F . n)) by A240, MESFUNC1:def_11; then A242: (g . x) - ((F . n) . x) <= e by A241, MESFUNC1:def_4; (g | B) . x = g . x by A235, A234, A237, FUNCT_1:47; hence (g | B) . x <= (((F . n) | B) + EP) . x by A158, A160, A239, A242, XXREAL_3:41; ::_thesis: verum end; A243: ((F . n) | B) + EP is nonnegative by A232, A215, Th19; then A244: dom ((((F . n) | B) + EP) - (g | B)) = (dom (((F . n) | B) + EP)) /\ (dom (g | B)) by A225, A184, A231, A236, Th69; then A245: g | B = (g | B) | (dom ((((F . n) | B) + EP) - (g | B))) by A229, A212, A235, A233, GRFUNC_1:23; ((F . n) | B) + EP = (((F . n) | B) + EP) | (dom ((((F . n) | B) + EP) - (g | B))) by A229, A212, A235, A244, A233, GRFUNC_1:23; then A246: integral' (M,(g | B)) <= integral' (M,(((F . n) | B) + EP)) by A225, A184, A231, A243, A236, A245, Th70; integral' (M,(((F . n) | B) + EP)) = (integral' (M,(((F . n) | B) | B))) + (integral' (M,(EP | B))) by A232, A230, A211, A215, A234, Th65 .= (integral' (M,(((F . n) | B) | B))) + (integral' (M,EP)) by A212, GRFUNC_1:23 .= (integral' (M,((F . n) | B))) + (integral' (M,EP)) by FUNCT_1:51 ; then A247: (integral' (M,(g | B))) - (integral' (M,EP)) <= integral' (M,((F . n) | B)) by A158, A214, A217, A216, A246, XXREAL_3:42; beta1 + MG = (R_EAL (max RINGA)) + (M . (dom g)) by A223, SUPINF_2:1; then ee * (beta1 + MG) = e * ((R_EAL (max RINGA)) + (M . (dom g))) by EXTREAL1:1; then A248: XSMg - (ee * (beta1 + MG)) = (integral' (M,g)) - (e * ((R_EAL (max RINGA)) + (M . (dom g)))) by SUPINF_2:3; XSMgB - XSMEP = (integral' (M,(g | B))) - (integral' (M,EP)) by SUPINF_2:3; then (integral' (M,g)) - (e * ((R_EAL (max RINGA)) + (M . (dom g)))) < integral' (M,((F . n) | B)) by A247, A228, A248, XXREAL_0:2; hence (integral' (M,g)) - (e * ((R_EAL (max RINGA)) + (M . (dom g)))) < integral' (M,(F . n)) by A220, XXREAL_0:2; ::_thesis: verum end; hence for n being Nat st N0 <= n holds (integral' (M,g)) - (e * ((R_EAL (max RINGA)) + (M . (dom g)))) < integral' (M,(F . n)) ; ::_thesis: verum end; A249: for e being R_eal st 0 < e & e < R_EAL (min RINGA) holds ex N0 being Nat st for n being Nat st N0 <= n holds (integral' (M,g)) - (e * ((R_EAL (max RINGA)) + (M . (dom g)))) < L . n proof let e be R_eal; ::_thesis: ( 0 < e & e < R_EAL (min RINGA) implies ex N0 being Nat st for n being Nat st N0 <= n holds (integral' (M,g)) - (e * ((R_EAL (max RINGA)) + (M . (dom g)))) < L . n ) assume that A250: 0 < e and A251: e < R_EAL (min RINGA) ; ::_thesis: ex N0 being Nat st for n being Nat st N0 <= n holds (integral' (M,g)) - (e * ((R_EAL (max RINGA)) + (M . (dom g)))) < L . n consider N0 being Nat such that A252: for n being Nat st N0 <= n holds (integral' (M,g)) - (e * ((R_EAL (max RINGA)) + (M . (dom g)))) < integral' (M,(F . n)) by A157, A250, A251; now__::_thesis:_for_n_being_Nat_st_N0_<=_n_holds_ (integral'_(M,g))_-_(e_*_((R_EAL_(max_RINGA))_+_(M_._(dom_g))))_<_L_._n let n be Nat; ::_thesis: ( N0 <= n implies (integral' (M,g)) - (e * ((R_EAL (max RINGA)) + (M . (dom g)))) < L . n ) assume N0 <= n ; ::_thesis: (integral' (M,g)) - (e * ((R_EAL (max RINGA)) + (M . (dom g)))) < L . n then (integral' (M,g)) - (e * ((R_EAL (max RINGA)) + (M . (dom g)))) < integral' (M,(F . n)) by A252; hence (integral' (M,g)) - (e * ((R_EAL (max RINGA)) + (M . (dom g)))) < L . n by A8; ::_thesis: verum end; hence ex N0 being Nat st for n being Nat st N0 <= n holds (integral' (M,g)) - (e * ((R_EAL (max RINGA)) + (M . (dom g)))) < L . n ; ::_thesis: verum end; A253: for e1 being R_eal st 0 < e1 holds ex e being R_eal st ( 0 < e & e < R_EAL (min RINGA) & e * ((R_EAL (max RINGA)) + (M . (dom g))) <= e1 ) proof reconsider ralpha = R_EAL (min RINGA) as Real by XXREAL_0:14; reconsider rdomg = M . (dom g) as Real by A136, A142, XXREAL_0:14; let e1 be R_eal; ::_thesis: ( 0 < e1 implies ex e being R_eal st ( 0 < e & e < R_EAL (min RINGA) & e * ((R_EAL (max RINGA)) + (M . (dom g))) <= e1 ) ) assume A254: 0 < e1 ; ::_thesis: ex e being R_eal st ( 0 < e & e < R_EAL (min RINGA) & e * ((R_EAL (max RINGA)) + (M . (dom g))) <= e1 ) {} c= DG by XBOOLE_1:2; then A255: 0 <= rdomg by A141, A140, MEASURE1:31; percases ( e1 = +infty or e1 <> +infty ) ; supposeA256: e1 = +infty ; ::_thesis: ex e being R_eal st ( 0 < e & e < R_EAL (min RINGA) & e * ((R_EAL (max RINGA)) + (M . (dom g))) <= e1 ) set z = ralpha / 2; (R_EAL (ralpha / 2)) * ((R_EAL (max RINGA)) + (M . (dom g))) <= +infty by XXREAL_0:4; hence ex e being R_eal st ( 0 < e & e < R_EAL (min RINGA) & e * ((R_EAL (max RINGA)) + (M . (dom g))) <= e1 ) by A43, A256, XREAL_1:216; ::_thesis: verum end; suppose e1 <> +infty ; ::_thesis: ex e being R_eal st ( 0 < e & e < R_EAL (min RINGA) & e * ((R_EAL (max RINGA)) + (M . (dom g))) <= e1 ) then reconsider re1 = e1 as Real by A254, XXREAL_0:14; set x = re1 / (beta1 + rdomg); set y = ralpha / 2; set z = min ((re1 / (beta1 + rdomg)),(ralpha / 2)); A257: min ((re1 / (beta1 + rdomg)),(ralpha / 2)) <= ralpha / 2 by XXREAL_0:17; ralpha / 2 < ralpha by A43, XREAL_1:216; then A258: min ((re1 / (beta1 + rdomg)),(ralpha / 2)) < ralpha by A257, XXREAL_0:2; R_EAL (beta1 + rdomg) = (R_EAL beta1) + (R_EAL rdomg) by SUPINF_2:1; then A259: (min ((re1 / (beta1 + rdomg)),(ralpha / 2))) * (beta1 + rdomg) = (R_EAL (min ((re1 / (beta1 + rdomg)),(ralpha / 2)))) * ((R_EAL (max RINGA)) + (M . (dom g))) by EXTREAL1:5; A260: now__::_thesis:_0_<_min_((re1_/_(beta1_+_rdomg)),(ralpha_/_2)) percases ( min ((re1 / (beta1 + rdomg)),(ralpha / 2)) = re1 / (beta1 + rdomg) or min ((re1 / (beta1 + rdomg)),(ralpha / 2)) = ralpha / 2 ) by XXREAL_0:15; suppose min ((re1 / (beta1 + rdomg)),(ralpha / 2)) = re1 / (beta1 + rdomg) ; ::_thesis: 0 < min ((re1 / (beta1 + rdomg)),(ralpha / 2)) hence 0 < min ((re1 / (beta1 + rdomg)),(ralpha / 2)) by A137, A254, A255; ::_thesis: verum end; suppose min ((re1 / (beta1 + rdomg)),(ralpha / 2)) = ralpha / 2 ; ::_thesis: 0 < min ((re1 / (beta1 + rdomg)),(ralpha / 2)) hence 0 < min ((re1 / (beta1 + rdomg)),(ralpha / 2)) by A43; ::_thesis: verum end; end; end; (min ((re1 / (beta1 + rdomg)),(ralpha / 2))) * (beta1 + rdomg) <= (re1 / (beta1 + rdomg)) * (beta1 + rdomg) by A137, A255, XREAL_1:64, XXREAL_0:17; then (min ((re1 / (beta1 + rdomg)),(ralpha / 2))) * (beta1 + rdomg) <= re1 by A137, A255, XCMPLX_1:87; hence ex e being R_eal st ( 0 < e & e < R_EAL (min RINGA) & e * ((R_EAL (max RINGA)) + (M . (dom g))) <= e1 ) by A260, A258, A259; ::_thesis: verum end; end; end; A261: for e1 being R_eal st 0 < e1 holds ex N0 being Nat st for n being Nat st N0 <= n holds (integral' (M,g)) - e1 < L . n proof let e1 be R_eal; ::_thesis: ( 0 < e1 implies ex N0 being Nat st for n being Nat st N0 <= n holds (integral' (M,g)) - e1 < L . n ) assume 0 < e1 ; ::_thesis: ex N0 being Nat st for n being Nat st N0 <= n holds (integral' (M,g)) - e1 < L . n then consider e being R_eal such that A262: 0 < e and A263: e < R_EAL (min RINGA) and A264: e * ((R_EAL (max RINGA)) + (M . (dom g))) <= e1 by A253; consider N0 being Nat such that A265: for n being Nat st N0 <= n holds (integral' (M,g)) - (e * ((R_EAL (max RINGA)) + (M . (dom g)))) < L . n by A249, A262, A263; take N0 ; ::_thesis: for n being Nat st N0 <= n holds (integral' (M,g)) - e1 < L . n now__::_thesis:_for_n_being_Nat_st_N0_<=_n_holds_ (integral'_(M,g))_-_e1_<_L_._n let n be Nat; ::_thesis: ( N0 <= n implies (integral' (M,g)) - e1 < L . n ) assume N0 <= n ; ::_thesis: (integral' (M,g)) - e1 < L . n then A266: (integral' (M,g)) - (e * ((R_EAL (max RINGA)) + (M . (dom g)))) < L . n by A265; (integral' (M,g)) - e1 <= (integral' (M,g)) - (e * ((R_EAL (max RINGA)) + (M . (dom g)))) by A264, XXREAL_3:37; hence (integral' (M,g)) - e1 < L . n by A266, XXREAL_0:2; ::_thesis: verum end; hence for n being Nat st N0 <= n holds (integral' (M,g)) - e1 < L . n ; ::_thesis: verum end; A267: for n being Nat holds 0 <= L . n proof let n be Nat; ::_thesis: 0 <= L . n F . n is nonnegative by A5; then 0 <= integral' (M,(F . n)) by A3, Th68; hence 0 <= L . n by A8; ::_thesis: verum end; A268: for n, m being Nat st n <= m holds L . n <= L . m proof let n, m be Nat; ::_thesis: ( n <= m implies L . n <= L . m ) A269: dom (F . n) = dom g by A4; A270: F . m is_simple_func_in S by A3; A271: dom (F . m) = dom g by A4; assume A272: n <= m ; ::_thesis: L . n <= L . m A273: for x being set st x in dom ((F . m) - (F . n)) holds (F . n) . x <= (F . m) . x proof let x be set ; ::_thesis: ( x in dom ((F . m) - (F . n)) implies (F . n) . x <= (F . m) . x ) assume x in dom ((F . m) - (F . n)) ; ::_thesis: (F . n) . x <= (F . m) . x then x in ((dom (F . m)) /\ (dom (F . n))) \ ((((F . m) " {+infty}) /\ ((F . n) " {+infty})) \/ (((F . m) " {-infty}) /\ ((F . n) " {-infty}))) by MESFUNC1:def_4; then x in (dom (F . m)) /\ (dom (F . n)) by XBOOLE_0:def_5; hence (F . n) . x <= (F . m) . x by A6, A272, A269, A271; ::_thesis: verum end; A274: F . n is_simple_func_in S by A3; A275: F . m is nonnegative by A5; A276: F . n is nonnegative by A5; then A277: dom ((F . m) - (F . n)) = (dom (F . m)) /\ (dom (F . n)) by A275, A274, A270, A273, Th69; then A278: (F . m) | (dom ((F . m) - (F . n))) = F . m by A269, A271, GRFUNC_1:23; A279: (F . n) | (dom ((F . m) - (F . n))) = F . n by A269, A271, A277, GRFUNC_1:23; integral' (M,((F . n) | (dom ((F . m) - (F . n))))) <= integral' (M,((F . m) | (dom ((F . m) - (F . n))))) by A276, A275, A274, A270, A273, Th70; then L . n <= integral' (M,(F . m)) by A8, A278, A279; hence L . n <= L . m by A8; ::_thesis: verum end; percases ( ex K being real number st ( 0 < K & ( for n being Nat holds L . n < K ) ) or for K being real number holds ( not 0 < K or ex n being Nat st not L . n < K ) ) ; suppose ex K being real number st ( 0 < K & ( for n being Nat holds L . n < K ) ) ; ::_thesis: ( L is convergent & integral' (M,g) <= lim L ) then consider K being real number such that 0 < K and A280: for n being Nat holds L . n < K ; now__::_thesis:_for_x_being_ext-real_number_st_x_in_rng_L_holds_ x_<=_R_EAL_K let x be ext-real number ; ::_thesis: ( x in rng L implies x <= R_EAL K ) assume x in rng L ; ::_thesis: x <= R_EAL K then ex z being set st ( z in dom L & x = L . z ) by FUNCT_1:def_3; hence x <= R_EAL K by A280; ::_thesis: verum end; then R_EAL K is UpperBound of rng L by XXREAL_2:def_1; then A281: sup (rng L) <= R_EAL K by XXREAL_2:def_3; R_EAL K is Real by XREAL_0:def_1; then A282: sup (rng L) <> +infty by A281, XXREAL_0:9; A283: for n being Nat holds L . n <= sup (rng L) proof let n be Nat; ::_thesis: L . n <= sup (rng L) reconsider n = n as Element of NAT by ORDINAL1:def_12; dom L = NAT by FUNCT_2:def_1; then A284: L . n in rng L by FUNCT_1:def_3; sup (rng L) is UpperBound of rng L by XXREAL_2:def_3; hence L . n <= sup (rng L) by A284, XXREAL_2:def_1; ::_thesis: verum end; then L . 1 <= sup (rng L) ; then A285: sup (rng L) <> -infty by A267; then reconsider h = sup (rng L) as Real by A282, XXREAL_0:14; A286: for p being real number st 0 < p holds ex N0 being Nat st for m being Nat st N0 <= m holds |.((L . m) - (sup (rng L))).| < p proof let p be real number ; ::_thesis: ( 0 < p implies ex N0 being Nat st for m being Nat st N0 <= m holds |.((L . m) - (sup (rng L))).| < p ) assume A287: 0 < p ; ::_thesis: ex N0 being Nat st for m being Nat st N0 <= m holds |.((L . m) - (sup (rng L))).| < p A288: sup (rng L) <> (sup (rng L)) + (R_EAL p) proof assume A289: sup (rng L) = (sup (rng L)) + (R_EAL p) ; ::_thesis: contradiction ((R_EAL p) + (sup (rng L))) + (- (sup (rng L))) = (R_EAL p) + ((sup (rng L)) + (- (sup (rng L)))) by A285, A282, XXREAL_3:29 .= (R_EAL p) + 0 by XXREAL_3:7 .= R_EAL p by XXREAL_3:4 ; hence contradiction by A287, A289, XXREAL_3:7; ::_thesis: verum end; sup (rng L) is Real by A285, A282, XXREAL_0:14; then consider y being ext-real number such that A290: y in rng L and A291: (sup (rng L)) - (R_EAL p) < y by A287, MEASURE6:6; consider x being set such that A292: x in dom L and A293: y = L . x by A290, FUNCT_1:def_3; reconsider N0 = x as Element of NAT by A292; take N0 ; ::_thesis: for m being Nat st N0 <= m holds |.((L . m) - (sup (rng L))).| < p let n be Nat; ::_thesis: ( N0 <= n implies |.((L . n) - (sup (rng L))).| < p ) assume N0 <= n ; ::_thesis: |.((L . n) - (sup (rng L))).| < p then L . N0 <= L . n by A268; then (sup (rng L)) - (R_EAL p) < L . n by A291, A293, XXREAL_0:2; then sup (rng L) < (L . n) + (R_EAL p) by XXREAL_3:54; then (sup (rng L)) - (L . n) < R_EAL p by XXREAL_3:55; then - (R_EAL p) < - ((sup (rng L)) - (L . n)) by XXREAL_3:38; then A294: - (R_EAL p) < (L . n) - (sup (rng L)) by XXREAL_3:26; A295: L . n <= sup (rng L) by A283; (sup (rng L)) + 0. <= (sup (rng L)) + (R_EAL p) by A287, XXREAL_3:36; then sup (rng L) <= (sup (rng L)) + (R_EAL p) by XXREAL_3:4; then sup (rng L) < (sup (rng L)) + (R_EAL p) by A288, XXREAL_0:1; then L . n < (sup (rng L)) + (R_EAL p) by A295, XXREAL_0:2; then (L . n) - (sup (rng L)) < R_EAL p by XXREAL_3:55; hence |.((L . n) - (sup (rng L))).| < p by A294, EXTREAL2:11; ::_thesis: verum end; A296: R_EAL h = sup (rng L) ; then A297: L is convergent_to_finite_number by A286, Def8; hence L is convergent by Def11; ::_thesis: integral' (M,g) <= lim L then A298: lim L = sup (rng L) by A286, A296, A297, Def12; now__::_thesis:_for_e_being_real_number_st_0_<_e_holds_ integral'_(M,g)_<_(lim_L)_+_e let e be real number ; ::_thesis: ( 0 < e implies integral' (M,g) < (lim L) + e ) assume 0 < e ; ::_thesis: integral' (M,g) < (lim L) + e then consider N0 being Nat such that A299: for n being Nat st N0 <= n holds (integral' (M,g)) - (R_EAL e) < L . n by A261; A300: L . N0 <= sup (rng L) by A283; (integral' (M,g)) - (R_EAL e) < L . N0 by A299; then (integral' (M,g)) - (R_EAL e) < sup (rng L) by A300, XXREAL_0:2; hence integral' (M,g) < (lim L) + e by A298, XXREAL_3:54; ::_thesis: verum end; hence integral' (M,g) <= lim L by XXREAL_3:61; ::_thesis: verum end; supposeA301: for K being real number holds ( not 0 < K or ex n being Nat st not L . n < K ) ; ::_thesis: ( L is convergent & integral' (M,g) <= lim L ) now__::_thesis:_for_K_being_real_number_st_0_<_K_holds_ ex_N0_being_Nat_st_ for_n_being_Nat_st_N0_<=_n_holds_ K_<=_L_._n let K be real number ; ::_thesis: ( 0 < K implies ex N0 being Nat st for n being Nat st N0 <= n holds K <= L . n ) assume 0 < K ; ::_thesis: ex N0 being Nat st for n being Nat st N0 <= n holds K <= L . n then consider N0 being Nat such that A302: R_EAL K <= L . N0 by A301; now__::_thesis:_for_n_being_Nat_st_N0_<=_n_holds_ R_EAL_K_<=_L_._n let n be Nat; ::_thesis: ( N0 <= n implies R_EAL K <= L . n ) assume N0 <= n ; ::_thesis: R_EAL K <= L . n then L . N0 <= L . n by A268; hence R_EAL K <= L . n by A302, XXREAL_0:2; ::_thesis: verum end; hence ex N0 being Nat st for n being Nat st N0 <= n holds K <= L . n ; ::_thesis: verum end; then A303: L is convergent_to_+infty by Def9; hence L is convergent by Def11; ::_thesis: integral' (M,g) <= lim L then lim L = +infty by A303, Def12; hence integral' (M,g) <= lim L by XXREAL_0:4; ::_thesis: verum end; end; end; supposeA304: M . (dom g) = +infty ; ::_thesis: ( L is convergent & integral' (M,g) <= lim L ) reconsider DG = dom g as Element of S by A1, Th37; A305: for e being R_eal st 0 < e & e < R_EAL (min RINGA) holds for n being Nat holds ((R_EAL (min RINGA)) - e) * (M . (less_dom ((g - (F . n)),e))) <= integral' (M,(F . n)) proof let e be R_eal; ::_thesis: ( 0 < e & e < R_EAL (min RINGA) implies for n being Nat holds ((R_EAL (min RINGA)) - e) * (M . (less_dom ((g - (F . n)),e))) <= integral' (M,(F . n)) ) assume that A306: 0 < e and A307: e < R_EAL (min RINGA) ; ::_thesis: for n being Nat holds ((R_EAL (min RINGA)) - e) * (M . (less_dom ((g - (F . n)),e))) <= integral' (M,(F . n)) A308: 0 <= (R_EAL (min RINGA)) - e by A307, XXREAL_3:40; consider H being SetSequence of X, MF being ExtREAL_sequence such that A309: for n being Nat holds H . n = less_dom ((g - (F . n)),e) and for n, m being Nat st n <= m holds H . n c= H . m and A310: for n being Nat holds H . n c= dom g and for n being Nat holds MF . n = M . (H . n) and M . (dom g) = sup (rng MF) and A311: for n being Nat holds H . n in S by A61, A306, A307; A312: e <> +infty by A307, XXREAL_0:4; now__::_thesis:_for_n_being_Nat_holds_((R_EAL_(min_RINGA))_-_e)_*_(M_._(less_dom_((g_-_(F_._n)),e)))_<=_integral'_(M,(F_._n)) let n be Nat; ::_thesis: ((R_EAL (min RINGA)) - e) * (M . (less_dom ((g - (F . n)),e))) <= integral' (M,(F . n)) reconsider B = H . n as Element of S by A311; A313: for x being set st x in dom (F . n) holds (F . n) . x = (F . n) . x ; H . n in S by A311; then A314: X \ (H . n) in S by MEASURE1:34; DG /\ (X \ (H . n)) = (DG /\ X) \ (H . n) by XBOOLE_1:49 .= DG \ (H . n) by XBOOLE_1:28 ; then reconsider A = DG \ (H . n) as Element of S by A314, MEASURE1:34; A315: dom (F . n) = dom g by A4; A316: DG = DG \/ (H . n) by A310, XBOOLE_1:12 .= (DG \ (H . n)) \/ (H . n) by XBOOLE_1:39 ; then dom (F . n) = (A \/ B) /\ (dom (F . n)) by A315; then A317: F . n = (F . n) | (A \/ B) by A313, FUNCT_1:46; consider EP being PartFunc of X,ExtREAL such that A318: EP is_simple_func_in S and A319: dom EP = B and A320: for x being set st x in B holds EP . x = (R_EAL (min RINGA)) - e by A306, A308, Th41, XXREAL_3:18; for x being set st x in dom EP holds 0 <= EP . x by A308, A319, A320; then A321: EP is nonnegative by SUPINF_2:52; A322: dom ((F . n) | B) = (dom (F . n)) /\ B by RELAT_1:61 .= B by A316, A315, XBOOLE_1:7, XBOOLE_1:28 ; A323: for x being set st x in dom (((F . n) | B) - EP) holds EP . x <= ((F . n) | B) . x proof set f = g - (F . n); let x be set ; ::_thesis: ( x in dom (((F . n) | B) - EP) implies EP . x <= ((F . n) | B) . x ) assume x in dom (((F . n) | B) - EP) ; ::_thesis: EP . x <= ((F . n) | B) . x then x in ((dom ((F . n) | B)) /\ (dom EP)) \ (((((F . n) | B) " {+infty}) /\ (EP " {+infty})) \/ ((((F . n) | B) " {-infty}) /\ (EP " {-infty}))) by MESFUNC1:def_4; then A324: x in (dom ((F . n) | B)) /\ (dom EP) by XBOOLE_0:def_5; then A325: x in dom ((F . n) | B) by XBOOLE_0:def_4; then A326: ((F . n) | B) . x = (F . n) . x by FUNCT_1:47; A327: x in less_dom ((g - (F . n)),e) by A309, A322, A325; then A328: x in dom (g - (F . n)) by MESFUNC1:def_11; (g - (F . n)) . x < e by A327, MESFUNC1:def_11; then (g . x) - ((F . n) . x) <= e by A328, MESFUNC1:def_4; then g . x <= ((F . n) . x) + e by A306, A312, XXREAL_3:41; then A329: (g . x) - e <= (F . n) . x by A306, A312, XXREAL_3:42; dom (g - (F . n)) = dom g by A56; then R_EAL (min RINGA) <= g . x by A44, A328; then (R_EAL (min RINGA)) - e <= (g . x) - e by XXREAL_3:37; then (R_EAL (min RINGA)) - e <= (F . n) . x by A329, XXREAL_0:2; hence EP . x <= ((F . n) | B) . x by A322, A319, A320, A324, A326; ::_thesis: verum end; A330: F . n is_simple_func_in S by A3; (F . n) | A is nonnegative by A5, Th15; then A331: 0 <= integral' (M,((F . n) | A)) by A330, Th34, Th68; A332: A misses B by XBOOLE_1:79; F . n is nonnegative by A5; then integral' (M,(F . n)) = (integral' (M,((F . n) | A))) + (integral' (M,((F . n) | B))) by A3, A317, A332, Th67; then A333: integral' (M,((F . n) | B)) <= integral' (M,(F . n)) by A331, XXREAL_3:39; A334: (F . n) | B is_simple_func_in S by A3, Th34; A335: (F . n) | B is nonnegative by A5, Th15; then A336: dom (((F . n) | B) - EP) = (dom ((F . n) | B)) /\ (dom EP) by A334, A318, A321, A323, Th69; then A337: EP | (dom (((F . n) | B) - EP)) = EP by A322, A319, GRFUNC_1:23; A338: ((F . n) | B) | (dom (((F . n) | B) - EP)) = (F . n) | B by A322, A319, A336, GRFUNC_1:23; integral' (M,(EP | (dom (((F . n) | B) - EP)))) <= integral' (M,(((F . n) | B) | (dom (((F . n) | B) - EP)))) by A335, A334, A318, A321, A323, Th70; then A339: integral' (M,EP) <= integral' (M,(F . n)) by A333, A337, A338, XXREAL_0:2; integral' (M,EP) = ((R_EAL (min RINGA)) - e) * (M . B) by A307, A318, A319, A320, Th71, XXREAL_3:40; hence ((R_EAL (min RINGA)) - e) * (M . (less_dom ((g - (F . n)),e))) <= integral' (M,(F . n)) by A309, A339; ::_thesis: verum end; hence for n being Nat holds ((R_EAL (min RINGA)) - e) * (M . (less_dom ((g - (F . n)),e))) <= integral' (M,(F . n)) ; ::_thesis: verum end; for y being real number st 0 < y holds ex n being Nat st for m being Nat st n <= m holds y <= L . m proof reconsider ralpha = R_EAL (min RINGA) as Real by XXREAL_0:14; set e = (R_EAL (min RINGA)) / (R_EAL 2); let y be real number ; ::_thesis: ( 0 < y implies ex n being Nat st for m being Nat st n <= m holds y <= L . m ) assume 0 < y ; ::_thesis: ex n being Nat st for m being Nat st n <= m holds y <= L . m set a2 = ralpha / 2; reconsider y1 = y as Real by XREAL_0:def_1; y = (ralpha - (ralpha / 2)) * (y1 / (ralpha - (ralpha / 2))) by A43, XCMPLX_1:87; then A340: y = (R_EAL (ralpha - (ralpha / 2))) * (R_EAL (y1 / (ralpha - (ralpha / 2)))) by EXTREAL1:5; A341: (R_EAL (min RINGA)) / (R_EAL 2) = ralpha / 2 by EXTREAL1:6; then consider H being SetSequence of X, MF being ExtREAL_sequence such that A342: for n being Nat holds H . n = less_dom ((g - (F . n)),((R_EAL (min RINGA)) / (R_EAL 2))) and A343: for n, m being Nat st n <= m holds H . n c= H . m and for n being Nat holds H . n c= dom g and A344: for n being Nat holds MF . n = M . (H . n) and A345: M . (dom g) = sup (rng MF) and A346: for n being Nat holds H . n in S by A61, A43, XREAL_1:216; y / (ralpha - (ralpha / 2)) = (R_EAL y1) / (R_EAL (ralpha - (ralpha / 2))) by EXTREAL1:6 .= (R_EAL y) / ((R_EAL (min RINGA)) - ((R_EAL (min RINGA)) / (R_EAL 2))) by A341, SUPINF_2:3 ; then A347: (R_EAL y) / ((R_EAL (min RINGA)) - ((R_EAL (min RINGA)) / (R_EAL 2))) < +infty by XXREAL_0:9; ex z being ext-real number st ( z in rng MF & (R_EAL y) / ((R_EAL (min RINGA)) - ((R_EAL (min RINGA)) / (R_EAL 2))) <= z ) proof assume for z being ext-real number holds ( not z in rng MF or not (R_EAL y) / ((R_EAL (min RINGA)) - ((R_EAL (min RINGA)) / (R_EAL 2))) <= z ) ; ::_thesis: contradiction then for z being ext-real number st z in rng MF holds z <= (R_EAL y) / ((R_EAL (min RINGA)) - ((R_EAL (min RINGA)) / (R_EAL 2))) ; then (R_EAL y) / ((R_EAL (min RINGA)) - ((R_EAL (min RINGA)) / (R_EAL 2))) is UpperBound of rng MF by XXREAL_2:def_1; hence contradiction by A304, A347, A345, XXREAL_2:def_3; ::_thesis: verum end; then consider z being R_eal such that A348: z in rng MF and A349: (R_EAL y) / ((R_EAL (min RINGA)) - ((R_EAL (min RINGA)) / (R_EAL 2))) <= z ; (ralpha / 2) - (ralpha / 2) < ralpha - (ralpha / 2) by A43; then A350: 0 < (R_EAL (min RINGA)) - ((R_EAL (min RINGA)) / (R_EAL 2)) by A341, SUPINF_2:3; consider x being set such that A351: x in dom MF and A352: z = MF . x by A348, FUNCT_1:def_3; reconsider N0 = x as Element of NAT by A351; take N0 ; ::_thesis: for m being Nat st N0 <= m holds y <= L . m (R_EAL (min RINGA)) - ((R_EAL (min RINGA)) / (R_EAL 2)) = ralpha - (ralpha / 2) by A341, SUPINF_2:3; then A353: ((R_EAL (min RINGA)) - ((R_EAL (min RINGA)) / (R_EAL 2))) * ((R_EAL y) / ((R_EAL (min RINGA)) - ((R_EAL (min RINGA)) / (R_EAL 2)))) = y by A340, EXTREAL1:6; thus for m being Nat st N0 <= m holds y <= L . m ::_thesis: verum proof (R_EAL y) / ((R_EAL (min RINGA)) - ((R_EAL (min RINGA)) / (R_EAL 2))) <= M . (H . N0) by A344, A349, A352; then A354: R_EAL y <= ((R_EAL (min RINGA)) - ((R_EAL (min RINGA)) / (R_EAL 2))) * (M . (H . N0)) by A350, A353, XXREAL_3:71; let m be Nat; ::_thesis: ( N0 <= m implies y <= L . m ) A355: H . m in S by A346; assume N0 <= m ; ::_thesis: y <= L . m then A356: H . N0 c= H . m by A343; H . N0 in S by A346; then ((R_EAL (min RINGA)) - ((R_EAL (min RINGA)) / (R_EAL 2))) * (M . (H . N0)) <= ((R_EAL (min RINGA)) - ((R_EAL (min RINGA)) / (R_EAL 2))) * (M . (H . m)) by A350, A356, A355, MEASURE1:31, XXREAL_3:71; then R_EAL y <= ((R_EAL (min RINGA)) - ((R_EAL (min RINGA)) / (R_EAL 2))) * (M . (H . m)) by A354, XXREAL_0:2; then A357: R_EAL y <= ((R_EAL (min RINGA)) - ((R_EAL (min RINGA)) / (R_EAL 2))) * (M . (less_dom ((g - (F . m)),((R_EAL (min RINGA)) / (R_EAL 2))))) by A342; ((R_EAL (min RINGA)) - ((R_EAL (min RINGA)) / (R_EAL 2))) * (M . (less_dom ((g - (F . m)),((R_EAL (min RINGA)) / (R_EAL 2))))) <= integral' (M,(F . m)) by A43, A305, A341, XREAL_1:216; then R_EAL y <= integral' (M,(F . m)) by A357, XXREAL_0:2; hence y <= L . m by A8; ::_thesis: verum end; end; then A358: L is convergent_to_+infty by Def9; hence L is convergent by Def11; ::_thesis: integral' (M,g) <= lim L then ( ex g being real number st ( lim L = g & ( for p being real number st 0 < p holds ex n being Nat st for m being Nat st n <= m holds |.((L . m) - (lim L)).| < p ) & L is convergent_to_finite_number ) or ( lim L = +infty & L is convergent_to_+infty ) or ( lim L = -infty & L is convergent_to_-infty ) ) by Def12; hence integral' (M,g) <= lim L by A358, Th50, XXREAL_0:4; ::_thesis: verum end; end; end; end; end; theorem Th75: :: MESFUNC5:75 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for g being PartFunc of X,ExtREAL for F being Functional_Sequence of X,ExtREAL st g is_simple_func_in S & g is nonnegative & ( for n being Nat holds F . n is_simple_func_in S ) & ( for n being Nat holds dom (F . n) = dom g ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom g holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom g holds ( F # x is convergent & g . x <= lim (F # x) ) ) holds ex G being ExtREAL_sequence st ( ( for n being Nat holds G . n = integral' (M,(F . n)) ) & G is convergent & sup (rng G) = lim G & integral' (M,g) <= lim G ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for g being PartFunc of X,ExtREAL for F being Functional_Sequence of X,ExtREAL st g is_simple_func_in S & g is nonnegative & ( for n being Nat holds F . n is_simple_func_in S ) & ( for n being Nat holds dom (F . n) = dom g ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom g holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom g holds ( F # x is convergent & g . x <= lim (F # x) ) ) holds ex G being ExtREAL_sequence st ( ( for n being Nat holds G . n = integral' (M,(F . n)) ) & G is convergent & sup (rng G) = lim G & integral' (M,g) <= lim G ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for g being PartFunc of X,ExtREAL for F being Functional_Sequence of X,ExtREAL st g is_simple_func_in S & g is nonnegative & ( for n being Nat holds F . n is_simple_func_in S ) & ( for n being Nat holds dom (F . n) = dom g ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom g holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom g holds ( F # x is convergent & g . x <= lim (F # x) ) ) holds ex G being ExtREAL_sequence st ( ( for n being Nat holds G . n = integral' (M,(F . n)) ) & G is convergent & sup (rng G) = lim G & integral' (M,g) <= lim G ) let M be sigma_Measure of S; ::_thesis: for g being PartFunc of X,ExtREAL for F being Functional_Sequence of X,ExtREAL st g is_simple_func_in S & g is nonnegative & ( for n being Nat holds F . n is_simple_func_in S ) & ( for n being Nat holds dom (F . n) = dom g ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom g holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom g holds ( F # x is convergent & g . x <= lim (F # x) ) ) holds ex G being ExtREAL_sequence st ( ( for n being Nat holds G . n = integral' (M,(F . n)) ) & G is convergent & sup (rng G) = lim G & integral' (M,g) <= lim G ) let g be PartFunc of X,ExtREAL; ::_thesis: for F being Functional_Sequence of X,ExtREAL st g is_simple_func_in S & g is nonnegative & ( for n being Nat holds F . n is_simple_func_in S ) & ( for n being Nat holds dom (F . n) = dom g ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom g holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom g holds ( F # x is convergent & g . x <= lim (F # x) ) ) holds ex G being ExtREAL_sequence st ( ( for n being Nat holds G . n = integral' (M,(F . n)) ) & G is convergent & sup (rng G) = lim G & integral' (M,g) <= lim G ) let F be Functional_Sequence of X,ExtREAL; ::_thesis: ( g is_simple_func_in S & g is nonnegative & ( for n being Nat holds F . n is_simple_func_in S ) & ( for n being Nat holds dom (F . n) = dom g ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom g holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom g holds ( F # x is convergent & g . x <= lim (F # x) ) ) implies ex G being ExtREAL_sequence st ( ( for n being Nat holds G . n = integral' (M,(F . n)) ) & G is convergent & sup (rng G) = lim G & integral' (M,g) <= lim G ) ) assume that A1: g is_simple_func_in S and A2: g is nonnegative and A3: for n being Nat holds F . n is_simple_func_in S and A4: for n being Nat holds dom (F . n) = dom g and A5: for n being Nat holds F . n is nonnegative and A6: for n, m being Nat st n <= m holds for x being Element of X st x in dom g holds (F . n) . x <= (F . m) . x and A7: for x being Element of X st x in dom g holds ( F # x is convergent & g . x <= lim (F # x) ) ; ::_thesis: ex G being ExtREAL_sequence st ( ( for n being Nat holds G . n = integral' (M,(F . n)) ) & G is convergent & sup (rng G) = lim G & integral' (M,g) <= lim G ) set E0 = eq_dom (g,(R_EAL 0)); reconsider DG = dom g as Element of S by A1, Th37; g is_measurable_on DG by A1, MESFUNC2:34; then reconsider GG = DG /\ (great_eq_dom (g,(R_EAL 0))) as Element of S by MESFUNC1:27; for x being set st x in eq_dom (g,(R_EAL 0)) holds x in dom g by MESFUNC1:def_15; then A8: eq_dom (g,(R_EAL 0)) c= DG by TARSKI:def_3; then A9: DG = DG \/ (eq_dom (g,(R_EAL 0))) by XBOOLE_1:12 .= (DG \ (eq_dom (g,(R_EAL 0)))) \/ (eq_dom (g,(R_EAL 0))) by XBOOLE_1:39 ; set E9 = (dom g) \ (eq_dom (g,(R_EAL 0))); g is_measurable_on GG by A1, MESFUNC2:34; then GG /\ (less_eq_dom (g,(R_EAL 0))) in S by MESFUNC1:28; then A10: DG /\ (eq_dom (g,(R_EAL 0))) in S by MESFUNC1:18; then eq_dom (g,(R_EAL 0)) in S by A8, XBOOLE_1:28; then A11: X \ (eq_dom (g,(R_EAL 0))) in S by MEASURE1:34; DG /\ (X \ (eq_dom (g,(R_EAL 0)))) = (DG /\ X) \ (eq_dom (g,(R_EAL 0))) by XBOOLE_1:49 .= DG \ (eq_dom (g,(R_EAL 0))) by XBOOLE_1:28 ; then reconsider E9 = (dom g) \ (eq_dom (g,(R_EAL 0))) as Element of S by A11, MEASURE1:34; reconsider E0 = eq_dom (g,(R_EAL 0)) as Element of S by A8, A10, XBOOLE_1:28; A12: E0 misses E9 by XBOOLE_1:79; thus ex G being ExtREAL_sequence st ( ( for n being Nat holds G . n = integral' (M,(F . n)) ) & G is convergent & sup (rng G) = lim G & integral' (M,g) <= lim G ) ::_thesis: verum proof A13: dom (g | E9) = (dom g) /\ E9 by RELAT_1:61 .= E9 by A9, XBOOLE_1:7, XBOOLE_1:28 ; A14: for x being set st x in dom (g | E9) holds 0 < (g | E9) . x proof let x be set ; ::_thesis: ( x in dom (g | E9) implies 0 < (g | E9) . x ) assume A15: x in dom (g | E9) ; ::_thesis: 0 < (g | E9) . x then A16: not x in E0 by A13, XBOOLE_0:def_5; x in DG by A13, A15, XBOOLE_0:def_5; then g . x <> 0 by A16, MESFUNC1:def_15; then 0 < g . x by A2, SUPINF_2:51; hence 0 < (g | E9) . x by A15, FUNCT_1:47; ::_thesis: verum end; deffunc H1( Nat) -> Element of ExtREAL = integral' (M,((F . $1) | E9)); deffunc H2( Nat) -> Element of ExtREAL = integral' (M,(F . $1)); deffunc H3( Nat) -> Element of bool [:X,ExtREAL:] = (F . $1) | E9; consider F9 being Functional_Sequence of X,ExtREAL such that A17: for n being Nat holds F9 . n = H3(n) from SEQFUNC:sch_1(); consider L being ExtREAL_sequence such that A18: for n being Element of NAT holds L . n = H1(n) from FUNCT_2:sch_4(); A19: now__::_thesis:_for_n_being_Nat_holds_L_._n_=_H1(n) let n be Nat; ::_thesis: L . n = H1(n) n in NAT by ORDINAL1:def_12; hence L . n = H1(n) by A18; ::_thesis: verum end; A20: for n being Nat holds L . n = integral' (M,(F9 . n)) proof let n be Nat; ::_thesis: L . n = integral' (M,(F9 . n)) thus L . n = integral' (M,((F . n) | E9)) by A19 .= integral' (M,(F9 . n)) by A17 ; ::_thesis: verum end; consider G being ExtREAL_sequence such that A21: for n being Element of NAT holds G . n = H2(n) from FUNCT_2:sch_4(); take G ; ::_thesis: ( ( for n being Nat holds G . n = integral' (M,(F . n)) ) & G is convergent & sup (rng G) = lim G & integral' (M,g) <= lim G ) A22: for x being set st x in dom g holds g . x = g . x ; dom g = (E0 \/ E9) /\ (dom g) by A9; then g | (E0 \/ E9) = g by A22, FUNCT_1:46; then A23: integral' (M,g) = (integral' (M,(g | E0))) + (integral' (M,(g | E9))) by A1, A2, Th67, XBOOLE_1:79; integral' (M,(g | E0)) = 0 by A1, A2, Th72; then A24: integral' (M,g) = integral' (M,(g | E9)) by A23, XXREAL_3:4; A25: g | E9 is_simple_func_in S by A1, Th34; A26: for n being Nat holds ( (F . n) | E9 is_simple_func_in S & F9 . n is_simple_func_in S ) proof let n be Nat; ::_thesis: ( (F . n) | E9 is_simple_func_in S & F9 . n is_simple_func_in S ) thus (F . n) | E9 is_simple_func_in S by A3, Th34; ::_thesis: F9 . n is_simple_func_in S hence F9 . n is_simple_func_in S by A17; ::_thesis: verum end; A27: for n being Nat holds ( dom ((F . n) | E9) = dom (g | E9) & dom (F9 . n) = dom (g | E9) ) proof let n be Nat; ::_thesis: ( dom ((F . n) | E9) = dom (g | E9) & dom (F9 . n) = dom (g | E9) ) A28: dom (F . n) = E9 \/ E0 by A4, A9; thus dom ((F . n) | E9) = (dom (F . n)) /\ E9 by RELAT_1:61 .= dom (g | E9) by A13, A28, XBOOLE_1:7, XBOOLE_1:28 ; ::_thesis: dom (F9 . n) = dom (g | E9) hence dom (F9 . n) = dom (g | E9) by A17; ::_thesis: verum end; A29: for x being Element of X st x in dom (g | E9) holds ( F9 # x is convergent & (g | E9) . x <= lim (F9 # x) ) proof let x be Element of X; ::_thesis: ( x in dom (g | E9) implies ( F9 # x is convergent & (g | E9) . x <= lim (F9 # x) ) ) assume A30: x in dom (g | E9) ; ::_thesis: ( F9 # x is convergent & (g | E9) . x <= lim (F9 # x) ) now__::_thesis:_for_n_being_Element_of_NAT_holds_(F9_#_x)_._n_=_(F_#_x)_._n let n be Element of NAT ; ::_thesis: (F9 # x) . n = (F # x) . n A31: x in dom ((F . n) | E9) by A27, A30; thus (F9 # x) . n = (F9 . n) . x by Def13 .= ((F . n) | E9) . x by A17 .= (F . n) . x by A31, FUNCT_1:47 .= (F # x) . n by Def13 ; ::_thesis: verum end; then A32: F9 # x = F # x by FUNCT_2:63; x in (dom g) /\ E9 by A30, RELAT_1:61; then A33: x in dom g by XBOOLE_0:def_4; then g . x <= lim (F # x) by A7; hence ( F9 # x is convergent & (g | E9) . x <= lim (F9 # x) ) by A7, A30, A33, A32, FUNCT_1:47; ::_thesis: verum end; A34: for n being Nat holds F9 . n is nonnegative proof let n be Nat; ::_thesis: F9 . n is nonnegative (F . n) | E9 is nonnegative by A5, Th15; hence F9 . n is nonnegative by A17; ::_thesis: verum end; A35: E9 c= dom g by A9, XBOOLE_1:7; A36: for n, m being Nat st n <= m holds for x being Element of X st x in dom (g | E9) holds ( ((F . n) | E9) . x <= ((F . m) | E9) . x & (F9 . n) . x <= (F9 . m) . x ) proof let n, m be Nat; ::_thesis: ( n <= m implies for x being Element of X st x in dom (g | E9) holds ( ((F . n) | E9) . x <= ((F . m) | E9) . x & (F9 . n) . x <= (F9 . m) . x ) ) assume A37: n <= m ; ::_thesis: for x being Element of X st x in dom (g | E9) holds ( ((F . n) | E9) . x <= ((F . m) | E9) . x & (F9 . n) . x <= (F9 . m) . x ) thus for x being Element of X st x in dom (g | E9) holds ( ((F . n) | E9) . x <= ((F . m) | E9) . x & (F9 . n) . x <= (F9 . m) . x ) ::_thesis: verum proof let x be Element of X; ::_thesis: ( x in dom (g | E9) implies ( ((F . n) | E9) . x <= ((F . m) | E9) . x & (F9 . n) . x <= (F9 . m) . x ) ) assume A38: x in dom (g | E9) ; ::_thesis: ( ((F . n) | E9) . x <= ((F . m) | E9) . x & (F9 . n) . x <= (F9 . m) . x ) then A39: x in dom ((F . n) | E9) by A27; (F . n) . x <= (F . m) . x by A6, A35, A13, A37, A38; then A40: ((F . n) | E9) . x <= (F . m) . x by A39, FUNCT_1:47; x in dom ((F . m) | E9) by A27, A38; hence ((F . n) | E9) . x <= ((F . m) | E9) . x by A40, FUNCT_1:47; ::_thesis: (F9 . n) . x <= (F9 . m) . x then (F9 . n) . x <= ((F . m) | E9) . x by A17; hence (F9 . n) . x <= (F9 . m) . x by A17; ::_thesis: verum end; end; then for n, m being Nat st n <= m holds for x being Element of X st x in dom (g | E9) holds (F9 . n) . x <= (F9 . m) . x ; then A41: integral' (M,(g | E9)) <= lim L by A25, A14, A27, A26, A29, A34, A20, Th74; for n, m being Nat st n <= m holds L . n <= L . m proof let n, m be Nat; ::_thesis: ( n <= m implies L . n <= L . m ) A42: F9 . m is_simple_func_in S by A26; A43: dom (F9 . m) = dom (g | E9) by A27; A44: L . m = integral' (M,(F9 . m)) by A20; A45: L . n = integral' (M,(F9 . n)) by A20; A46: dom (F9 . n) = dom (g | E9) by A27; assume A47: n <= m ; ::_thesis: L . n <= L . m A48: for x being set st x in dom ((F9 . m) - (F9 . n)) holds (F9 . n) . x <= (F9 . m) . x proof let x be set ; ::_thesis: ( x in dom ((F9 . m) - (F9 . n)) implies (F9 . n) . x <= (F9 . m) . x ) assume x in dom ((F9 . m) - (F9 . n)) ; ::_thesis: (F9 . n) . x <= (F9 . m) . x then x in ((dom (F9 . m)) /\ (dom (F9 . n))) \ ((((F9 . m) " {+infty}) /\ ((F9 . n) " {+infty})) \/ (((F9 . m) " {-infty}) /\ ((F9 . n) " {-infty}))) by MESFUNC1:def_4; then x in (dom (F9 . m)) /\ (dom (F9 . n)) by XBOOLE_0:def_5; hence (F9 . n) . x <= (F9 . m) . x by A36, A47, A46, A43; ::_thesis: verum end; A49: F9 . m is nonnegative by A34; A50: F9 . n is nonnegative by A34; A51: F9 . n is_simple_func_in S by A26; then A52: dom ((F9 . m) - (F9 . n)) = (dom (F9 . m)) /\ (dom (F9 . n)) by A42, A50, A49, A48, Th69; then A53: (F9 . m) | (dom ((F9 . m) - (F9 . n))) = F9 . m by A46, A43, GRFUNC_1:23; (F9 . n) | (dom ((F9 . m) - (F9 . n))) = F9 . n by A46, A43, A52, GRFUNC_1:23; hence L . n <= L . m by A51, A42, A50, A49, A48, A53, A45, A44, Th70; ::_thesis: verum end; then A54: lim L = sup (rng L) by Th54; A55: now__::_thesis:_for_n_being_Nat_holds_G_._n_=_H2(n) let n be Nat; ::_thesis: G . n = H2(n) n in NAT by ORDINAL1:def_12; hence G . n = H2(n) by A21; ::_thesis: verum end; for n being Nat holds L . n <= G . n proof let n be Nat; ::_thesis: L . n <= G . n A56: F . n is_simple_func_in S by A3; dom (F . n) = E9 \/ E0 by A4, A9; then A57: dom (F . n) = (E0 \/ E9) /\ (dom (F . n)) ; for x being set st x in dom (F . n) holds (F . n) . x = (F . n) . x ; then A58: F . n = (F . n) | (E0 \/ E9) by A57, FUNCT_1:46; then (F . n) | (E0 \/ E9) is nonnegative by A5; then A59: integral' (M,(F . n)) = (integral' (M,((F . n) | E0))) + (integral' (M,((F . n) | E9))) by A3, A12, A58, Th67; (F . n) | E0 is nonnegative by A5, Th15; then 0 <= integral' (M,((F . n) | E0)) by A56, Th34, Th68; then A60: integral' (M,((F . n) | E9)) <= integral' (M,(F . n)) by A59, XXREAL_3:39; G . n = integral' (M,(F . n)) by A55; hence L . n <= G . n by A19, A60; ::_thesis: verum end; then A61: sup (rng L) <= sup (rng G) by Th55; A62: for n, m being Nat st n <= m holds G . n <= G . m proof let n, m be Nat; ::_thesis: ( n <= m implies G . n <= G . m ) A63: F . m is_simple_func_in S by A3; A64: dom (F . m) = dom g by A4; A65: G . m = integral' (M,(F . m)) by A55; A66: G . n = integral' (M,(F . n)) by A55; A67: dom (F . n) = dom g by A4; assume A68: n <= m ; ::_thesis: G . n <= G . m A69: for x being set st x in dom ((F . m) - (F . n)) holds (F . n) . x <= (F . m) . x proof let x be set ; ::_thesis: ( x in dom ((F . m) - (F . n)) implies (F . n) . x <= (F . m) . x ) assume x in dom ((F . m) - (F . n)) ; ::_thesis: (F . n) . x <= (F . m) . x then x in ((dom (F . m)) /\ (dom (F . n))) \ ((((F . m) " {+infty}) /\ ((F . n) " {+infty})) \/ (((F . m) " {-infty}) /\ ((F . n) " {-infty}))) by MESFUNC1:def_4; then x in (dom (F . m)) /\ (dom (F . n)) by XBOOLE_0:def_5; hence (F . n) . x <= (F . m) . x by A6, A68, A67, A64; ::_thesis: verum end; A70: F . m is nonnegative by A5; A71: F . n is nonnegative by A5; A72: F . n is_simple_func_in S by A3; then A73: dom ((F . m) - (F . n)) = (dom (F . m)) /\ (dom (F . n)) by A63, A71, A70, A69, Th69; then A74: (F . m) | (dom ((F . m) - (F . n))) = F . m by A67, A64, GRFUNC_1:23; (F . n) | (dom ((F . m) - (F . n))) = F . n by A67, A64, A73, GRFUNC_1:23; hence G . n <= G . m by A72, A63, A71, A70, A69, A74, A66, A65, Th70; ::_thesis: verum end; then lim G = sup (rng G) by Th54; hence ( ( for n being Nat holds G . n = integral' (M,(F . n)) ) & G is convergent & sup (rng G) = lim G & integral' (M,g) <= lim G ) by A24, A55, A62, A41, A54, A61, Th54, XXREAL_0:2; ::_thesis: verum end; end; theorem Th76: :: MESFUNC5:76 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for A being Element of S for F, G being Functional_Sequence of X,ExtREAL for K, L being ExtREAL_sequence st ( for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = A ) ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in A holds (F . n) . x <= (F . m) . x ) & ( for n being Nat holds ( G . n is_simple_func_in S & dom (G . n) = A ) ) & ( for n being Nat holds G . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in A holds (G . n) . x <= (G . m) . x ) & ( for x being Element of X st x in A holds ( F # x is convergent & G # x is convergent & lim (F # x) = lim (G # x) ) ) & ( for n being Nat holds ( K . n = integral' (M,(F . n)) & L . n = integral' (M,(G . n)) ) ) holds ( K is convergent & L is convergent & lim K = lim L ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for A being Element of S for F, G being Functional_Sequence of X,ExtREAL for K, L being ExtREAL_sequence st ( for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = A ) ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in A holds (F . n) . x <= (F . m) . x ) & ( for n being Nat holds ( G . n is_simple_func_in S & dom (G . n) = A ) ) & ( for n being Nat holds G . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in A holds (G . n) . x <= (G . m) . x ) & ( for x being Element of X st x in A holds ( F # x is convergent & G # x is convergent & lim (F # x) = lim (G # x) ) ) & ( for n being Nat holds ( K . n = integral' (M,(F . n)) & L . n = integral' (M,(G . n)) ) ) holds ( K is convergent & L is convergent & lim K = lim L ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for A being Element of S for F, G being Functional_Sequence of X,ExtREAL for K, L being ExtREAL_sequence st ( for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = A ) ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in A holds (F . n) . x <= (F . m) . x ) & ( for n being Nat holds ( G . n is_simple_func_in S & dom (G . n) = A ) ) & ( for n being Nat holds G . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in A holds (G . n) . x <= (G . m) . x ) & ( for x being Element of X st x in A holds ( F # x is convergent & G # x is convergent & lim (F # x) = lim (G # x) ) ) & ( for n being Nat holds ( K . n = integral' (M,(F . n)) & L . n = integral' (M,(G . n)) ) ) holds ( K is convergent & L is convergent & lim K = lim L ) let M be sigma_Measure of S; ::_thesis: for A being Element of S for F, G being Functional_Sequence of X,ExtREAL for K, L being ExtREAL_sequence st ( for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = A ) ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in A holds (F . n) . x <= (F . m) . x ) & ( for n being Nat holds ( G . n is_simple_func_in S & dom (G . n) = A ) ) & ( for n being Nat holds G . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in A holds (G . n) . x <= (G . m) . x ) & ( for x being Element of X st x in A holds ( F # x is convergent & G # x is convergent & lim (F # x) = lim (G # x) ) ) & ( for n being Nat holds ( K . n = integral' (M,(F . n)) & L . n = integral' (M,(G . n)) ) ) holds ( K is convergent & L is convergent & lim K = lim L ) let A be Element of S; ::_thesis: for F, G being Functional_Sequence of X,ExtREAL for K, L being ExtREAL_sequence st ( for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = A ) ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in A holds (F . n) . x <= (F . m) . x ) & ( for n being Nat holds ( G . n is_simple_func_in S & dom (G . n) = A ) ) & ( for n being Nat holds G . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in A holds (G . n) . x <= (G . m) . x ) & ( for x being Element of X st x in A holds ( F # x is convergent & G # x is convergent & lim (F # x) = lim (G # x) ) ) & ( for n being Nat holds ( K . n = integral' (M,(F . n)) & L . n = integral' (M,(G . n)) ) ) holds ( K is convergent & L is convergent & lim K = lim L ) let F, G be Functional_Sequence of X,ExtREAL; ::_thesis: for K, L being ExtREAL_sequence st ( for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = A ) ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in A holds (F . n) . x <= (F . m) . x ) & ( for n being Nat holds ( G . n is_simple_func_in S & dom (G . n) = A ) ) & ( for n being Nat holds G . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in A holds (G . n) . x <= (G . m) . x ) & ( for x being Element of X st x in A holds ( F # x is convergent & G # x is convergent & lim (F # x) = lim (G # x) ) ) & ( for n being Nat holds ( K . n = integral' (M,(F . n)) & L . n = integral' (M,(G . n)) ) ) holds ( K is convergent & L is convergent & lim K = lim L ) let K, L be ExtREAL_sequence; ::_thesis: ( ( for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = A ) ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in A holds (F . n) . x <= (F . m) . x ) & ( for n being Nat holds ( G . n is_simple_func_in S & dom (G . n) = A ) ) & ( for n being Nat holds G . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in A holds (G . n) . x <= (G . m) . x ) & ( for x being Element of X st x in A holds ( F # x is convergent & G # x is convergent & lim (F # x) = lim (G # x) ) ) & ( for n being Nat holds ( K . n = integral' (M,(F . n)) & L . n = integral' (M,(G . n)) ) ) implies ( K is convergent & L is convergent & lim K = lim L ) ) assume that A1: for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = A ) and A2: for n being Nat holds F . n is nonnegative and A3: for n, m being Nat st n <= m holds for x being Element of X st x in A holds (F . n) . x <= (F . m) . x and A4: for n being Nat holds ( G . n is_simple_func_in S & dom (G . n) = A ) and A5: for n being Nat holds G . n is nonnegative and A6: for n, m being Nat st n <= m holds for x being Element of X st x in A holds (G . n) . x <= (G . m) . x and A7: for x being Element of X st x in A holds ( F # x is convergent & G # x is convergent & lim (F # x) = lim (G # x) ) and A8: for n being Nat holds ( K . n = integral' (M,(F . n)) & L . n = integral' (M,(G . n)) ) ; ::_thesis: ( K is convergent & L is convergent & lim K = lim L ) A9: for n0 being Nat holds ( L is convergent & sup (rng L) = lim L & K . n0 <= lim L ) proof let n0 be Nat; ::_thesis: ( L is convergent & sup (rng L) = lim L & K . n0 <= lim L ) reconsider f = F . n0 as PartFunc of X,ExtREAL ; A10: f is_simple_func_in S by A1; A11: f is nonnegative by A2; A12: for x being Element of X st x in dom f holds ( G # x is convergent & f . x <= lim (G # x) ) proof let x be Element of X; ::_thesis: ( x in dom f implies ( G # x is convergent & f . x <= lim (G # x) ) ) A13: (F # x) . n0 <= sup (rng (F # x)) by Th56; assume x in dom f ; ::_thesis: ( G # x is convergent & f . x <= lim (G # x) ) then A14: x in A by A1; now__::_thesis:_for_n,_m_being_Nat_st_n_<=_m_holds_ (F_#_x)_._n_<=_(F_#_x)_._m let n, m be Nat; ::_thesis: ( n <= m implies (F # x) . n <= (F # x) . m ) assume A15: n <= m ; ::_thesis: (F # x) . n <= (F # x) . m A16: (F # x) . m = (F . m) . x by Def13; (F # x) . n = (F . n) . x by Def13; hence (F # x) . n <= (F # x) . m by A3, A14, A15, A16; ::_thesis: verum end; then A17: lim (F # x) = sup (rng (F # x)) by Th54; f . x = (F # x) . n0 by Def13; hence ( G # x is convergent & f . x <= lim (G # x) ) by A7, A14, A17, A13; ::_thesis: verum end; dom f = A by A1; then consider FF being ExtREAL_sequence such that A18: for n being Nat holds FF . n = integral' (M,(G . n)) and A19: FF is convergent and A20: sup (rng FF) = lim FF and A21: integral' (M,f) <= lim FF by A4, A5, A6, A12, A10, A11, Th75; now__::_thesis:_for_n_being_Element_of_NAT_holds_FF_._n_=_L_._n let n be Element of NAT ; ::_thesis: FF . n = L . n FF . n = integral' (M,(G . n)) by A18; hence FF . n = L . n by A8; ::_thesis: verum end; then FF = L by FUNCT_2:63; hence ( L is convergent & sup (rng L) = lim L & K . n0 <= lim L ) by A8, A19, A20, A21; ::_thesis: verum end; A22: for n0 being Nat holds ( K is convergent & sup (rng K) = lim K & L . n0 <= lim K ) proof let n0 be Nat; ::_thesis: ( K is convergent & sup (rng K) = lim K & L . n0 <= lim K ) reconsider g = G . n0 as PartFunc of X,ExtREAL ; A23: g is_simple_func_in S by A4; A24: g is nonnegative by A5; A25: for x being Element of X st x in dom g holds ( F # x is convergent & g . x <= lim (F # x) ) proof let x be Element of X; ::_thesis: ( x in dom g implies ( F # x is convergent & g . x <= lim (F # x) ) ) A26: (G # x) . n0 <= sup (rng (G # x)) by Th56; assume x in dom g ; ::_thesis: ( F # x is convergent & g . x <= lim (F # x) ) then A27: x in A by A4; now__::_thesis:_for_n,_m_being_Nat_st_n_<=_m_holds_ (G_#_x)_._n_<=_(G_#_x)_._m let n, m be Nat; ::_thesis: ( n <= m implies (G # x) . n <= (G # x) . m ) assume A28: n <= m ; ::_thesis: (G # x) . n <= (G # x) . m A29: (G # x) . m = (G . m) . x by Def13; (G # x) . n = (G . n) . x by Def13; hence (G # x) . n <= (G # x) . m by A6, A27, A28, A29; ::_thesis: verum end; then A30: lim (G # x) = sup (rng (G # x)) by Th54; g . x = (G # x) . n0 by Def13; hence ( F # x is convergent & g . x <= lim (F # x) ) by A7, A27, A30, A26; ::_thesis: verum end; dom g = A by A4; then consider GG being ExtREAL_sequence such that A31: for n being Nat holds GG . n = integral' (M,(F . n)) and A32: GG is convergent and A33: sup (rng GG) = lim GG and A34: integral' (M,g) <= lim GG by A1, A2, A3, A25, A23, A24, Th75; now__::_thesis:_for_n_being_Element_of_NAT_holds_GG_._n_=_K_._n let n be Element of NAT ; ::_thesis: GG . n = K . n GG . n = integral' (M,(F . n)) by A31; hence GG . n = K . n by A8; ::_thesis: verum end; then GG = K by FUNCT_2:63; hence ( K is convergent & sup (rng K) = lim K & L . n0 <= lim K ) by A8, A32, A33, A34; ::_thesis: verum end; hence ( K is convergent & L is convergent ) by A9; ::_thesis: lim K = lim L A35: lim K <= lim L by A22, A9, Th57; lim L <= lim K by A22, A9, Th57; hence lim K = lim L by A35, XXREAL_0:1; ::_thesis: verum end; definition let X be non empty set ; let S be SigmaField of X; let M be sigma_Measure of S; let f be PartFunc of X,ExtREAL; assume that A1: ex A being Element of S st ( A = dom f & f is_measurable_on A ) and A2: f is nonnegative ; func integral+ (M,f) -> Element of ExtREAL means :Def15: :: MESFUNC5:def 15 ex F being Functional_Sequence of X,ExtREAL ex K being ExtREAL_sequence st ( ( for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = dom f ) ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom f holds ( F # x is convergent & lim (F # x) = f . x ) ) & ( for n being Nat holds K . n = integral' (M,(F . n)) ) & K is convergent & it = lim K ); existence ex b1 being Element of ExtREAL ex F being Functional_Sequence of X,ExtREAL ex K being ExtREAL_sequence st ( ( for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = dom f ) ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom f holds ( F # x is convergent & lim (F # x) = f . x ) ) & ( for n being Nat holds K . n = integral' (M,(F . n)) ) & K is convergent & b1 = lim K ) proof consider F being Functional_Sequence of X,ExtREAL such that A3: for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = dom f ) and A4: for n being Nat holds F . n is nonnegative and A5: for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F . n) . x <= (F . m) . x and A6: for x being Element of X st x in dom f holds ( F # x is convergent & lim (F # x) = f . x ) by A1, A2, Th64; reconsider g = F . 0 as PartFunc of X,ExtREAL ; A7: g is_simple_func_in S by A3; A8: for x being Element of X st x in dom f holds ( F # x is convergent & g . x <= lim (F # x) ) proof let x be Element of X; ::_thesis: ( x in dom f implies ( F # x is convergent & g . x <= lim (F # x) ) ) assume A9: x in dom f ; ::_thesis: ( F # x is convergent & g . x <= lim (F # x) ) A10: now__::_thesis:_for_n,_m_being_Nat_st_n_<=_m_holds_ (F_#_x)_._n_<=_(F_#_x)_._m let n, m be Nat; ::_thesis: ( n <= m implies (F # x) . n <= (F # x) . m ) assume A11: n <= m ; ::_thesis: (F # x) . n <= (F # x) . m A12: (F # x) . m = (F . m) . x by Def13; (F # x) . n = (F . n) . x by Def13; hence (F # x) . n <= (F # x) . m by A5, A9, A11, A12; ::_thesis: verum end; A13: g . x = (F # x) . 0 by Def13; lim (F # x) = sup (rng (F # x)) by A10, Th54; hence ( F # x is convergent & g . x <= lim (F # x) ) by A10, A13, Th54, Th56; ::_thesis: verum end; dom g = dom f by A3; then ex G being ExtREAL_sequence st ( ( for n being Nat holds G . n = integral' (M,(F . n)) ) & G is convergent & sup (rng G) = lim G & integral' (M,g) <= lim G ) by A3, A4, A5, A8, A7, Th75; then consider G being ExtREAL_sequence such that A14: for n being Nat holds G . n = integral' (M,(F . n)) and A15: G is convergent and integral' (M,g) <= lim G ; take lim G ; ::_thesis: ex F being Functional_Sequence of X,ExtREAL ex K being ExtREAL_sequence st ( ( for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = dom f ) ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom f holds ( F # x is convergent & lim (F # x) = f . x ) ) & ( for n being Nat holds K . n = integral' (M,(F . n)) ) & K is convergent & lim G = lim K ) thus ex F being Functional_Sequence of X,ExtREAL ex K being ExtREAL_sequence st ( ( for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = dom f ) ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom f holds ( F # x is convergent & lim (F # x) = f . x ) ) & ( for n being Nat holds K . n = integral' (M,(F . n)) ) & K is convergent & lim G = lim K ) by A3, A4, A5, A6, A14, A15; ::_thesis: verum end; uniqueness for b1, b2 being Element of ExtREAL st ex F being Functional_Sequence of X,ExtREAL ex K being ExtREAL_sequence st ( ( for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = dom f ) ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom f holds ( F # x is convergent & lim (F # x) = f . x ) ) & ( for n being Nat holds K . n = integral' (M,(F . n)) ) & K is convergent & b1 = lim K ) & ex F being Functional_Sequence of X,ExtREAL ex K being ExtREAL_sequence st ( ( for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = dom f ) ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom f holds ( F # x is convergent & lim (F # x) = f . x ) ) & ( for n being Nat holds K . n = integral' (M,(F . n)) ) & K is convergent & b2 = lim K ) holds b1 = b2 proof let s1, s2 be Element of ExtREAL ; ::_thesis: ( ex F being Functional_Sequence of X,ExtREAL ex K being ExtREAL_sequence st ( ( for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = dom f ) ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom f holds ( F # x is convergent & lim (F # x) = f . x ) ) & ( for n being Nat holds K . n = integral' (M,(F . n)) ) & K is convergent & s1 = lim K ) & ex F being Functional_Sequence of X,ExtREAL ex K being ExtREAL_sequence st ( ( for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = dom f ) ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom f holds ( F # x is convergent & lim (F # x) = f . x ) ) & ( for n being Nat holds K . n = integral' (M,(F . n)) ) & K is convergent & s2 = lim K ) implies s1 = s2 ) assume that A16: ex F1 being Functional_Sequence of X,ExtREAL ex K1 being ExtREAL_sequence st ( ( for n being Nat holds ( F1 . n is_simple_func_in S & dom (F1 . n) = dom f ) ) & ( for n being Nat holds F1 . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F1 . n) . x <= (F1 . m) . x ) & ( for x being Element of X st x in dom f holds ( F1 # x is convergent & lim (F1 # x) = f . x ) ) & ( for n being Nat holds K1 . n = integral' (M,(F1 . n)) ) & K1 is convergent & s1 = lim K1 ) and A17: ex F2 being Functional_Sequence of X,ExtREAL ex K2 being ExtREAL_sequence st ( ( for n being Nat holds ( F2 . n is_simple_func_in S & dom (F2 . n) = dom f ) ) & ( for n being Nat holds F2 . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F2 . n) . x <= (F2 . m) . x ) & ( for x being Element of X st x in dom f holds ( F2 # x is convergent & lim (F2 # x) = f . x ) ) & ( for n being Nat holds K2 . n = integral' (M,(F2 . n)) ) & K2 is convergent & s2 = lim K2 ) ; ::_thesis: s1 = s2 consider F1 being Functional_Sequence of X,ExtREAL, K1 being ExtREAL_sequence such that A18: for n being Nat holds ( F1 . n is_simple_func_in S & dom (F1 . n) = dom f ) and A19: for n being Nat holds F1 . n is nonnegative and A20: for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F1 . n) . x <= (F1 . m) . x and A21: for x being Element of X st x in dom f holds ( F1 # x is convergent & lim (F1 # x) = f . x ) and A22: for n being Nat holds K1 . n = integral' (M,(F1 . n)) and K1 is convergent and A23: s1 = lim K1 by A16; consider F2 being Functional_Sequence of X,ExtREAL, K2 being ExtREAL_sequence such that A24: for n being Nat holds ( F2 . n is_simple_func_in S & dom (F2 . n) = dom f ) and A25: for n being Nat holds F2 . n is nonnegative and A26: for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F2 . n) . x <= (F2 . m) . x and A27: for x being Element of X st x in dom f holds ( F2 # x is convergent & lim (F2 # x) = f . x ) and A28: for n being Nat holds K2 . n = integral' (M,(F2 . n)) and K2 is convergent and A29: s2 = lim K2 by A17; for x being Element of X st x in dom f holds ( F1 # x is convergent & F2 # x is convergent & lim (F1 # x) = lim (F2 # x) ) proof let x be Element of X; ::_thesis: ( x in dom f implies ( F1 # x is convergent & F2 # x is convergent & lim (F1 # x) = lim (F2 # x) ) ) assume A30: x in dom f ; ::_thesis: ( F1 # x is convergent & F2 # x is convergent & lim (F1 # x) = lim (F2 # x) ) then lim (F1 # x) = f . x by A21 .= lim (F2 # x) by A27, A30 ; hence ( F1 # x is convergent & F2 # x is convergent & lim (F1 # x) = lim (F2 # x) ) by A21, A27, A30; ::_thesis: verum end; hence s1 = s2 by A1, A18, A19, A20, A22, A23, A24, A25, A26, A28, A29, Th76; ::_thesis: verum end; end; :: deftheorem Def15 defines integral+ MESFUNC5:def_15_:_ for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative holds for b5 being Element of ExtREAL holds ( b5 = integral+ (M,f) iff ex F being Functional_Sequence of X,ExtREAL ex K being ExtREAL_sequence st ( ( for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = dom f ) ) & ( for n being Nat holds F . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F . n) . x <= (F . m) . x ) & ( for x being Element of X st x in dom f holds ( F # x is convergent & lim (F # x) = f . x ) ) & ( for n being Nat holds K . n = integral' (M,(F . n)) ) & K is convergent & b5 = lim K ) ); theorem Th77: :: MESFUNC5:77 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_simple_func_in S & f is nonnegative holds integral+ (M,f) = integral' (M,f) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_simple_func_in S & f is nonnegative holds integral+ (M,f) = integral' (M,f) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_simple_func_in S & f is nonnegative holds integral+ (M,f) = integral' (M,f) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL st f is_simple_func_in S & f is nonnegative holds integral+ (M,f) = integral' (M,f) let f be PartFunc of X,ExtREAL; ::_thesis: ( f is_simple_func_in S & f is nonnegative implies integral+ (M,f) = integral' (M,f) ) assume that A1: f is_simple_func_in S and A2: f is nonnegative ; ::_thesis: integral+ (M,f) = integral' (M,f) deffunc H1( Nat) -> PartFunc of X,ExtREAL = f; consider F being Functional_Sequence of X,ExtREAL such that A3: for n being Nat holds F . n = H1(n) from SEQFUNC:sch_1(); A4: for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F . n) . x <= (F . m) . x proof let n, m be Nat; ::_thesis: ( n <= m implies for x being Element of X st x in dom f holds (F . n) . x <= (F . m) . x ) assume n <= m ; ::_thesis: for x being Element of X st x in dom f holds (F . n) . x <= (F . m) . x let x be Element of X; ::_thesis: ( x in dom f implies (F . n) . x <= (F . m) . x ) assume x in dom f ; ::_thesis: (F . n) . x <= (F . m) . x (F . n) . x = f . x by A3; hence (F . n) . x <= (F . m) . x by A3; ::_thesis: verum end; deffunc H2( Nat) -> Element of ExtREAL = integral' (M,(F . $1)); consider K being Function of NAT,ExtREAL such that A5: for n being Element of NAT holds K . n = H2(n) from FUNCT_2:sch_4(); A6: now__::_thesis:_for_n_being_Nat_holds_K_._n_=_H2(n) let n be Nat; ::_thesis: K . n = H2(n) n in NAT by ORDINAL1:def_12; hence K . n = H2(n) by A5; ::_thesis: verum end; A7: for n being Nat holds K . n = integral' (M,f) proof let n be Nat; ::_thesis: K . n = integral' (M,f) thus K . n = integral' (M,(F . n)) by A6 .= integral' (M,f) by A3 ; ::_thesis: verum end; then A8: lim K = integral' (M,f) by Th60; ex GF being Finite_Sep_Sequence of S st ( dom f = union (rng GF) & ( for n being Nat for x, y being Element of X st n in dom GF & x in GF . n & y in GF . n holds f . x = f . y ) ) by A1, MESFUNC2:def_4; then reconsider A = dom f as Element of S by MESFUNC2:31; A9: f is_measurable_on A by A1, MESFUNC2:34; A10: for x being Element of X st x in dom f holds ( F # x is convergent & lim (F # x) = f . x ) proof let x be Element of X; ::_thesis: ( x in dom f implies ( F # x is convergent & lim (F # x) = f . x ) ) assume x in dom f ; ::_thesis: ( F # x is convergent & lim (F # x) = f . x ) now__::_thesis:_for_n_being_Nat_holds_(F_#_x)_._n_=_f_._x let n be Nat; ::_thesis: (F # x) . n = f . x thus (F # x) . n = (F . n) . x by Def13 .= f . x by A3 ; ::_thesis: verum end; hence ( F # x is convergent & lim (F # x) = f . x ) by Th60; ::_thesis: verum end; A11: for n being Nat holds F . n is nonnegative by A2, A3; A12: for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = dom f ) by A1, A3; K is convergent by A7, Th60; hence integral+ (M,f) = integral' (M,f) by A2, A9, A6, A12, A11, A4, A10, A8, Def15; ::_thesis: verum end; Lm10: for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & A = dom g & f is_measurable_on A & g is_measurable_on A ) & f is nonnegative & g is nonnegative holds integral+ (M,(f + g)) = (integral+ (M,f)) + (integral+ (M,g)) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & A = dom g & f is_measurable_on A & g is_measurable_on A ) & f is nonnegative & g is nonnegative holds integral+ (M,(f + g)) = (integral+ (M,f)) + (integral+ (M,g)) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & A = dom g & f is_measurable_on A & g is_measurable_on A ) & f is nonnegative & g is nonnegative holds integral+ (M,(f + g)) = (integral+ (M,f)) + (integral+ (M,g)) let M be sigma_Measure of S; ::_thesis: for f, g being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & A = dom g & f is_measurable_on A & g is_measurable_on A ) & f is nonnegative & g is nonnegative holds integral+ (M,(f + g)) = (integral+ (M,f)) + (integral+ (M,g)) let f, g be PartFunc of X,ExtREAL; ::_thesis: ( ex A being Element of S st ( A = dom f & A = dom g & f is_measurable_on A & g is_measurable_on A ) & f is nonnegative & g is nonnegative implies integral+ (M,(f + g)) = (integral+ (M,f)) + (integral+ (M,g)) ) assume that A1: ex A being Element of S st ( A = dom f & A = dom g & f is_measurable_on A & g is_measurable_on A ) and A2: f is nonnegative and A3: g is nonnegative ; ::_thesis: integral+ (M,(f + g)) = (integral+ (M,f)) + (integral+ (M,g)) consider F1 being Functional_Sequence of X,ExtREAL, K1 being ExtREAL_sequence such that A4: for n being Nat holds ( F1 . n is_simple_func_in S & dom (F1 . n) = dom f ) and A5: for n being Nat holds F1 . n is nonnegative and A6: for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F1 . n) . x <= (F1 . m) . x and A7: for x being Element of X st x in dom f holds ( F1 # x is convergent & lim (F1 # x) = f . x ) and A8: for n being Nat holds K1 . n = integral' (M,(F1 . n)) and K1 is convergent and A9: integral+ (M,f) = lim K1 by A1, A2, Def15; A10: f + g is nonnegative by A2, A3, Th19; consider A being Element of S such that A11: A = dom f and A12: A = dom g and A13: f is_measurable_on A and A14: g is_measurable_on A by A1; A = (dom f) /\ (dom g) by A11, A12; then A15: A = dom (f + g) by A2, A3, Th16; A16: for n, m being Nat st n <= m holds K1 . n <= K1 . m proof let n, m be Nat; ::_thesis: ( n <= m implies K1 . n <= K1 . m ) assume A17: n <= m ; ::_thesis: K1 . n <= K1 . m A18: dom (F1 . m) = dom f by A4; A19: dom (F1 . n) = dom f by A4; A20: now__::_thesis:_for_x_being_set_st_x_in_dom_((F1_._m)_-_(F1_._n))_holds_ (F1_._n)_._x_<=_(F1_._m)_._x let x be set ; ::_thesis: ( x in dom ((F1 . m) - (F1 . n)) implies (F1 . n) . x <= (F1 . m) . x ) assume x in dom ((F1 . m) - (F1 . n)) ; ::_thesis: (F1 . n) . x <= (F1 . m) . x then x in ((dom (F1 . m)) /\ (dom (F1 . n))) \ ((((F1 . m) " {+infty}) /\ ((F1 . n) " {+infty})) \/ (((F1 . m) " {-infty}) /\ ((F1 . n) " {-infty}))) by MESFUNC1:def_4; then x in (dom (F1 . m)) /\ (dom (F1 . n)) by XBOOLE_0:def_5; hence (F1 . n) . x <= (F1 . m) . x by A6, A17, A19, A18; ::_thesis: verum end; A21: F1 . m is nonnegative by A5; A22: F1 . n is nonnegative by A5; A23: K1 . m = integral' (M,(F1 . m)) by A8; A24: K1 . n = integral' (M,(F1 . n)) by A8; A25: F1 . m is_simple_func_in S by A4; A26: F1 . n is_simple_func_in S by A4; then A27: dom ((F1 . m) - (F1 . n)) = (dom (F1 . m)) /\ (dom (F1 . n)) by A25, A22, A21, A20, Th69; then A28: (F1 . m) | (dom ((F1 . m) - (F1 . n))) = F1 . m by A19, A18, GRFUNC_1:23; (F1 . n) | (dom ((F1 . m) - (F1 . n))) = F1 . n by A19, A18, A27, GRFUNC_1:23; hence K1 . n <= K1 . m by A24, A23, A26, A25, A22, A21, A20, A28, Th70; ::_thesis: verum end; consider F2 being Functional_Sequence of X,ExtREAL, K2 being ExtREAL_sequence such that A29: for n being Nat holds ( F2 . n is_simple_func_in S & dom (F2 . n) = dom g ) and A30: for n being Nat holds F2 . n is nonnegative and A31: for n, m being Nat st n <= m holds for x being Element of X st x in dom g holds (F2 . n) . x <= (F2 . m) . x and A32: for x being Element of X st x in dom g holds ( F2 # x is convergent & lim (F2 # x) = g . x ) and A33: for n being Nat holds K2 . n = integral' (M,(F2 . n)) and K2 is convergent and A34: integral+ (M,g) = lim K2 by A1, A3, Def15; deffunc H1( Nat) -> Element of bool [:X,ExtREAL:] = (F1 . $1) + (F2 . $1); consider F being Functional_Sequence of X,ExtREAL such that A35: for n being Nat holds F . n = H1(n) from SEQFUNC:sch_1(); A36: for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = dom (f + g) & F . n is nonnegative ) proof let n be Nat; ::_thesis: ( F . n is_simple_func_in S & dom (F . n) = dom (f + g) & F . n is nonnegative ) A37: dom (F1 . n) = dom f by A4; A38: F2 . n is_simple_func_in S by A29; A39: F2 . n is nonnegative by A30; A40: F . n = (F1 . n) + (F2 . n) by A35; A41: F1 . n is_simple_func_in S by A4; hence F . n is_simple_func_in S by A38, A40, Th38; ::_thesis: ( dom (F . n) = dom (f + g) & F . n is nonnegative ) A42: dom (F2 . n) = dom g by A29; F1 . n is nonnegative by A5; then dom (F . n) = (dom (F1 . n)) /\ (dom (F2 . n)) by A41, A38, A39, A40, Th65; hence dom (F . n) = dom (f + g) by A2, A3, A37, A42, Th16; ::_thesis: F . n is nonnegative A43: F2 . n is nonnegative by A30; A44: F . n = (F1 . n) + (F2 . n) by A35; F1 . n is nonnegative by A5; hence F . n is nonnegative by A43, A44, Th19; ::_thesis: verum end; A45: for n, m being Nat st n <= m holds for x being Element of X st x in dom (f + g) holds (F . n) . x <= (F . m) . x proof let n, m be Nat; ::_thesis: ( n <= m implies for x being Element of X st x in dom (f + g) holds (F . n) . x <= (F . m) . x ) assume A46: n <= m ; ::_thesis: for x being Element of X st x in dom (f + g) holds (F . n) . x <= (F . m) . x dom ((F1 . m) + (F2 . m)) = dom (F . m) by A35; then A47: dom ((F1 . m) + (F2 . m)) = dom (f + g) by A36; dom ((F1 . n) + (F2 . n)) = dom (F . n) by A35; then A48: dom ((F1 . n) + (F2 . n)) = dom (f + g) by A36; let x be Element of X; ::_thesis: ( x in dom (f + g) implies (F . n) . x <= (F . m) . x ) assume A49: x in dom (f + g) ; ::_thesis: (F . n) . x <= (F . m) . x then A50: (F2 . n) . x <= (F2 . m) . x by A31, A12, A15, A46; (F . m) . x = ((F1 . m) + (F2 . m)) . x by A35; then A51: (F . m) . x = ((F1 . m) . x) + ((F2 . m) . x) by A49, A47, MESFUNC1:def_3; (F . n) . x = ((F1 . n) + (F2 . n)) . x by A35; then A52: (F . n) . x = ((F1 . n) . x) + ((F2 . n) . x) by A49, A48, MESFUNC1:def_3; (F1 . n) . x <= (F1 . m) . x by A6, A11, A15, A46, A49; hence (F . n) . x <= (F . m) . x by A52, A51, A50, XXREAL_3:36; ::_thesis: verum end; now__::_thesis:_for_n_being_set_st_n_in_dom_K2_holds_ -infty_<_K2_._n let n be set ; ::_thesis: ( n in dom K2 implies -infty < K2 . n ) assume n in dom K2 ; ::_thesis: -infty < K2 . n then reconsider n1 = n as Element of NAT ; A53: F2 . n1 is_simple_func_in S by A29; K2 . n1 = integral' (M,(F2 . n1)) by A33; hence -infty < K2 . n by A30, A53, Th68; ::_thesis: verum end; then A54: K2 is () by Th10; deffunc H2( Nat) -> Element of ExtREAL = integral' (M,(F . $1)); consider K being ExtREAL_sequence such that A55: for n being Element of NAT holds K . n = H2(n) from FUNCT_2:sch_4(); A56: now__::_thesis:_for_n_being_Nat_holds_K_._n_=_H2(n) let n be Nat; ::_thesis: K . n = H2(n) n in NAT by ORDINAL1:def_12; hence K . n = H2(n) by A55; ::_thesis: verum end; A57: for n being Nat holds K . n = (K1 . n) + (K2 . n) proof let n be Nat; ::_thesis: K . n = (K1 . n) + (K2 . n) A58: F1 . n is nonnegative by A5; A59: F . n = (F1 . n) + (F2 . n) by A35; A60: dom (F1 . n) = dom f by A4 .= dom (F2 . n) by A29, A11, A12 ; A61: F2 . n is_simple_func_in S by A29; A62: K . n = integral' (M,(F . n)) by A56; A63: F2 . n is nonnegative by A30; A64: F1 . n is_simple_func_in S by A4; then dom (F . n) = (dom (F1 . n)) /\ (dom (F2 . n)) by A58, A61, A63, A59, Th65; then K . n = (integral' (M,((F1 . n) | (dom (F1 . n))))) + (integral' (M,((F2 . n) | (dom (F2 . n))))) by A64, A58, A61, A63, A59, A60, A62, Th65; then K . n = (integral' (M,(F1 . n))) + (integral' (M,((F2 . n) | (dom (F2 . n))))) by GRFUNC_1:23; then A65: K . n = (integral' (M,(F1 . n))) + (integral' (M,(F2 . n))) by GRFUNC_1:23; K2 . n = integral' (M,(F2 . n)) by A33; hence K . n = (K1 . n) + (K2 . n) by A8, A65; ::_thesis: verum end; A66: for n, m being Nat st n <= m holds K2 . n <= K2 . m proof let n, m be Nat; ::_thesis: ( n <= m implies K2 . n <= K2 . m ) assume A67: n <= m ; ::_thesis: K2 . n <= K2 . m A68: dom (F2 . m) = dom g by A29; A69: dom (F2 . n) = dom g by A29; A70: now__::_thesis:_for_x_being_set_st_x_in_dom_((F2_._m)_-_(F2_._n))_holds_ (F2_._n)_._x_<=_(F2_._m)_._x let x be set ; ::_thesis: ( x in dom ((F2 . m) - (F2 . n)) implies (F2 . n) . x <= (F2 . m) . x ) assume x in dom ((F2 . m) - (F2 . n)) ; ::_thesis: (F2 . n) . x <= (F2 . m) . x then x in ((dom (F2 . m)) /\ (dom (F2 . n))) \ ((((F2 . m) " {+infty}) /\ ((F2 . n) " {+infty})) \/ (((F2 . m) " {-infty}) /\ ((F2 . n) " {-infty}))) by MESFUNC1:def_4; then x in (dom (F2 . m)) /\ (dom (F2 . n)) by XBOOLE_0:def_5; hence (F2 . n) . x <= (F2 . m) . x by A31, A67, A69, A68; ::_thesis: verum end; A71: F2 . m is nonnegative by A30; A72: F2 . n is nonnegative by A30; A73: K2 . m = integral' (M,(F2 . m)) by A33; A74: K2 . n = integral' (M,(F2 . n)) by A33; A75: F2 . m is_simple_func_in S by A29; A76: F2 . n is_simple_func_in S by A29; then A77: dom ((F2 . m) - (F2 . n)) = (dom (F2 . m)) /\ (dom (F2 . n)) by A75, A72, A71, A70, Th69; then A78: (F2 . m) | (dom ((F2 . m) - (F2 . n))) = F2 . m by A69, A68, GRFUNC_1:23; (F2 . n) | (dom ((F2 . m) - (F2 . n))) = F2 . n by A69, A68, A77, GRFUNC_1:23; hence K2 . n <= K2 . m by A74, A73, A76, A75, A72, A71, A70, A78, Th70; ::_thesis: verum end; now__::_thesis:_for_n_being_set_st_n_in_dom_K1_holds_ -infty_<_K1_._n let n be set ; ::_thesis: ( n in dom K1 implies -infty < K1 . n ) assume n in dom K1 ; ::_thesis: -infty < K1 . n then reconsider n1 = n as Element of NAT ; A79: F1 . n1 is_simple_func_in S by A4; K1 . n1 = integral' (M,(F1 . n1)) by A8; hence -infty < K1 . n by A5, A79, Th68; ::_thesis: verum end; then A80: K1 is () by Th10; then A81: lim K = (lim K1) + (lim K2) by A16, A54, A66, A57, Th61; A82: for x being Element of X st x in dom (f + g) holds ( F # x is convergent & lim (F # x) = (f + g) . x ) proof let x be Element of X; ::_thesis: ( x in dom (f + g) implies ( F # x is convergent & lim (F # x) = (f + g) . x ) ) A83: now__::_thesis:_for_n_being_set_holds_ (_(_n_in_dom_(F1_#_x)_implies_-infty_<_(F1_#_x)_._n_)_&_(_n_in_dom_(F2_#_x)_implies_-infty_<_(F2_#_x)_._n_)_) let n be set ; ::_thesis: ( ( n in dom (F1 # x) implies -infty < (F1 # x) . n ) & ( n in dom (F2 # x) implies -infty < (F2 # x) . n ) ) hereby ::_thesis: ( n in dom (F2 # x) implies -infty < (F2 # x) . n ) assume n in dom (F1 # x) ; ::_thesis: -infty < (F1 # x) . n then reconsider n1 = n as Element of NAT ; A84: (F1 # x) . n1 = (F1 . n1) . x by Def13; F1 . n1 is nonnegative by A5; hence -infty < (F1 # x) . n by A84, Def5; ::_thesis: verum end; assume n in dom (F2 # x) ; ::_thesis: -infty < (F2 # x) . n then reconsider n1 = n as Element of NAT ; A85: (F2 # x) . n1 = (F2 . n1) . x by Def13; F2 . n1 is nonnegative by A30; hence -infty < (F2 # x) . n by A85, Def5; ::_thesis: verum end; then A86: F2 # x is () by Th10; assume A87: x in dom (f + g) ; ::_thesis: ( F # x is convergent & lim (F # x) = (f + g) . x ) then (lim (F1 # x)) + (lim (F2 # x)) = (f . x) + (lim (F2 # x)) by A7, A11, A15; then (lim (F1 # x)) + (lim (F2 # x)) = (f . x) + (g . x) by A32, A12, A15, A87; then A88: (lim (F1 # x)) + (lim (F2 # x)) = (f + g) . x by A87, MESFUNC1:def_3; A89: now__::_thesis:_for_n,_m_being_Nat_st_n_<=_m_holds_ (F2_#_x)_._n_<=_(F2_#_x)_._m let n, m be Nat; ::_thesis: ( n <= m implies (F2 # x) . n <= (F2 # x) . m ) assume A90: n <= m ; ::_thesis: (F2 # x) . n <= (F2 # x) . m A91: (F2 # x) . m = (F2 . m) . x by Def13; (F2 # x) . n = (F2 . n) . x by Def13; hence (F2 # x) . n <= (F2 # x) . m by A31, A12, A15, A87, A90, A91; ::_thesis: verum end; A92: now__::_thesis:_for_n,_m_being_Nat_st_n_<=_m_holds_ (F1_#_x)_._n_<=_(F1_#_x)_._m let n, m be Nat; ::_thesis: ( n <= m implies (F1 # x) . n <= (F1 # x) . m ) assume A93: n <= m ; ::_thesis: (F1 # x) . n <= (F1 # x) . m A94: (F1 # x) . m = (F1 . m) . x by Def13; (F1 # x) . n = (F1 . n) . x by Def13; hence (F1 # x) . n <= (F1 # x) . m by A6, A11, A15, A87, A93, A94; ::_thesis: verum end; A95: now__::_thesis:_for_n_being_Nat_holds_(F_#_x)_._n_=_((F1_#_x)_._n)_+_((F2_#_x)_._n) let n be Nat; ::_thesis: (F # x) . n = ((F1 # x) . n) + ((F2 # x) . n) (F # x) . n = (F . n) . x by Def13; then A96: (F # x) . n = ((F1 . n) + (F2 . n)) . x by A35; dom ((F1 . n) + (F2 . n)) = dom (F . n) by A35 .= dom (f + g) by A36 ; then (F # x) . n = ((F1 . n) . x) + ((F2 . n) . x) by A87, A96, MESFUNC1:def_3; then (F # x) . n = ((F1 # x) . n) + ((F2 . n) . x) by Def13; hence (F # x) . n = ((F1 # x) . n) + ((F2 # x) . n) by Def13; ::_thesis: verum end; F1 # x is () by A83, Th10; hence ( F # x is convergent & lim (F # x) = (f + g) . x ) by A95, A86, A92, A89, A88, Th61; ::_thesis: verum end; A97: f + g is_measurable_on A by A2, A3, A13, A14, Th31; K is convergent by A80, A16, A54, A66, A57, Th61; hence integral+ (M,(f + g)) = (integral+ (M,f)) + (integral+ (M,g)) by A9, A34, A97, A15, A10, A56, A36, A45, A82, A81, Def15; ::_thesis: verum end; theorem Th78: :: MESFUNC5:78 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & ex B being Element of S st ( B = dom g & g is_measurable_on B ) & f is nonnegative & g is nonnegative holds ex C being Element of S st ( C = dom (f + g) & integral+ (M,(f + g)) = (integral+ (M,(f | C))) + (integral+ (M,(g | C))) ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & ex B being Element of S st ( B = dom g & g is_measurable_on B ) & f is nonnegative & g is nonnegative holds ex C being Element of S st ( C = dom (f + g) & integral+ (M,(f + g)) = (integral+ (M,(f | C))) + (integral+ (M,(g | C))) ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & ex B being Element of S st ( B = dom g & g is_measurable_on B ) & f is nonnegative & g is nonnegative holds ex C being Element of S st ( C = dom (f + g) & integral+ (M,(f + g)) = (integral+ (M,(f | C))) + (integral+ (M,(g | C))) ) let M be sigma_Measure of S; ::_thesis: for f, g being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & ex B being Element of S st ( B = dom g & g is_measurable_on B ) & f is nonnegative & g is nonnegative holds ex C being Element of S st ( C = dom (f + g) & integral+ (M,(f + g)) = (integral+ (M,(f | C))) + (integral+ (M,(g | C))) ) let f, g be PartFunc of X,ExtREAL; ::_thesis: ( ex A being Element of S st ( A = dom f & f is_measurable_on A ) & ex B being Element of S st ( B = dom g & g is_measurable_on B ) & f is nonnegative & g is nonnegative implies ex C being Element of S st ( C = dom (f + g) & integral+ (M,(f + g)) = (integral+ (M,(f | C))) + (integral+ (M,(g | C))) ) ) assume that A1: ex A being Element of S st ( A = dom f & f is_measurable_on A ) and A2: ex B being Element of S st ( B = dom g & g is_measurable_on B ) and A3: f is nonnegative and A4: g is nonnegative ; ::_thesis: ex C being Element of S st ( C = dom (f + g) & integral+ (M,(f + g)) = (integral+ (M,(f | C))) + (integral+ (M,(g | C))) ) set g1 = g | ((dom f) /\ (dom g)); A5: g | ((dom f) /\ (dom g)) is () by A4, Th12, Th15; A6: g | ((dom f) /\ (dom g)) is nonnegative by A4, Th15; dom (g | ((dom f) /\ (dom g))) = (dom g) /\ ((dom f) /\ (dom g)) by RELAT_1:61; then A7: dom (g | ((dom f) /\ (dom g))) = ((dom g) /\ (dom g)) /\ (dom f) by XBOOLE_1:16; consider B being Element of S such that A8: B = dom g and A9: g is_measurable_on B by A2; consider A being Element of S such that A10: A = dom f and A11: f is_measurable_on A by A1; take C = A /\ B; ::_thesis: ( C = dom (f + g) & integral+ (M,(f + g)) = (integral+ (M,(f | C))) + (integral+ (M,(g | C))) ) A12: C = dom (f + g) by A3, A4, A10, A8, Th16; A13: C = (dom g) /\ C by A8, XBOOLE_1:17, XBOOLE_1:28; g is_measurable_on C by A9, MESFUNC1:30, XBOOLE_1:17; then A14: g | C is_measurable_on C by A13, Th42; A15: C = (dom f) /\ C by A10, XBOOLE_1:17, XBOOLE_1:28; f is_measurable_on C by A11, MESFUNC1:30, XBOOLE_1:17; then A16: f | C is_measurable_on C by A15, Th42; set f1 = f | ((dom f) /\ (dom g)); dom (f | ((dom f) /\ (dom g))) = (dom f) /\ ((dom f) /\ (dom g)) by RELAT_1:61; then A17: dom (f | ((dom f) /\ (dom g))) = ((dom f) /\ (dom f)) /\ (dom g) by XBOOLE_1:16; A18: f | ((dom f) /\ (dom g)) is () by A3, Th12, Th15; then A19: dom ((f | ((dom f) /\ (dom g))) + (g | ((dom f) /\ (dom g)))) = C /\ C by A10, A8, A17, A7, A5, Th16; A20: dom ((f | ((dom f) /\ (dom g))) + (g | ((dom f) /\ (dom g)))) = (dom (f | ((dom f) /\ (dom g)))) /\ (dom (g | ((dom f) /\ (dom g)))) by A18, A5, Th16; A21: for x being set st x in dom ((f | ((dom f) /\ (dom g))) + (g | ((dom f) /\ (dom g)))) holds ((f | ((dom f) /\ (dom g))) + (g | ((dom f) /\ (dom g)))) . x = (f + g) . x proof let x be set ; ::_thesis: ( x in dom ((f | ((dom f) /\ (dom g))) + (g | ((dom f) /\ (dom g)))) implies ((f | ((dom f) /\ (dom g))) + (g | ((dom f) /\ (dom g)))) . x = (f + g) . x ) assume A22: x in dom ((f | ((dom f) /\ (dom g))) + (g | ((dom f) /\ (dom g)))) ; ::_thesis: ((f | ((dom f) /\ (dom g))) + (g | ((dom f) /\ (dom g)))) . x = (f + g) . x then A23: x in dom (f | ((dom f) /\ (dom g))) by A20, XBOOLE_0:def_4; A24: x in dom (g | ((dom f) /\ (dom g))) by A20, A22, XBOOLE_0:def_4; ((f | ((dom f) /\ (dom g))) + (g | ((dom f) /\ (dom g)))) . x = ((f | ((dom f) /\ (dom g))) . x) + ((g | ((dom f) /\ (dom g))) . x) by A22, MESFUNC1:def_3 .= (f . x) + ((g | ((dom f) /\ (dom g))) . x) by A23, FUNCT_1:47 .= (f . x) + (g . x) by A24, FUNCT_1:47 ; hence ((f | ((dom f) /\ (dom g))) + (g | ((dom f) /\ (dom g)))) . x = (f + g) . x by A12, A19, A22, MESFUNC1:def_3; ::_thesis: verum end; f | ((dom f) /\ (dom g)) is nonnegative by A3, Th15; then integral+ (M,((f | ((dom f) /\ (dom g))) + (g | ((dom f) /\ (dom g))))) = (integral+ (M,(f | ((dom f) /\ (dom g))))) + (integral+ (M,(g | ((dom f) /\ (dom g))))) by A10, A8, A17, A7, A16, A14, A6, Lm10; hence ( C = dom (f + g) & integral+ (M,(f + g)) = (integral+ (M,(f | C))) + (integral+ (M,(g | C))) ) by A10, A8, A12, A19, A21, FUNCT_1:2; ::_thesis: verum end; theorem Th79: :: MESFUNC5:79 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative holds 0 <= integral+ (M,f) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative holds 0 <= integral+ (M,f) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative holds 0 <= integral+ (M,f) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative holds 0 <= integral+ (M,f) let f be PartFunc of X,ExtREAL; ::_thesis: ( ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative implies 0 <= integral+ (M,f) ) assume that A1: ex A being Element of S st ( A = dom f & f is_measurable_on A ) and A2: f is nonnegative ; ::_thesis: 0 <= integral+ (M,f) consider F1 being Functional_Sequence of X,ExtREAL, K1 being ExtREAL_sequence such that A3: for n being Nat holds ( F1 . n is_simple_func_in S & dom (F1 . n) = dom f ) and A4: for n being Nat holds F1 . n is nonnegative and A5: for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F1 . n) . x <= (F1 . m) . x and for x being Element of X st x in dom f holds ( F1 # x is convergent & lim (F1 # x) = f . x ) and A6: for n being Nat holds K1 . n = integral' (M,(F1 . n)) and K1 is convergent and A7: integral+ (M,f) = lim K1 by A1, A2, Def15; for n, m being Nat st n <= m holds K1 . n <= K1 . m proof let n, m be Nat; ::_thesis: ( n <= m implies K1 . n <= K1 . m ) A8: F1 . m is_simple_func_in S by A3; A9: dom (F1 . m) = dom f by A3; A10: K1 . m = integral' (M,(F1 . m)) by A6; A11: dom (F1 . n) = dom f by A3; assume A12: n <= m ; ::_thesis: K1 . n <= K1 . m A13: for x being set st x in dom ((F1 . m) - (F1 . n)) holds (F1 . n) . x <= (F1 . m) . x proof let x be set ; ::_thesis: ( x in dom ((F1 . m) - (F1 . n)) implies (F1 . n) . x <= (F1 . m) . x ) assume x in dom ((F1 . m) - (F1 . n)) ; ::_thesis: (F1 . n) . x <= (F1 . m) . x then x in ((dom (F1 . m)) /\ (dom (F1 . n))) \ ((((F1 . m) " {+infty}) /\ ((F1 . n) " {+infty})) \/ (((F1 . m) " {-infty}) /\ ((F1 . n) " {-infty}))) by MESFUNC1:def_4; then x in (dom (F1 . m)) /\ (dom (F1 . n)) by XBOOLE_0:def_5; hence (F1 . n) . x <= (F1 . m) . x by A5, A12, A11, A9; ::_thesis: verum end; A14: F1 . m is nonnegative by A4; A15: F1 . n is nonnegative by A4; A16: F1 . n is_simple_func_in S by A3; then A17: dom ((F1 . m) - (F1 . n)) = (dom (F1 . m)) /\ (dom (F1 . n)) by A8, A15, A14, A13, Th69; then A18: (F1 . n) | (dom ((F1 . m) - (F1 . n))) = F1 . n by A11, A9, GRFUNC_1:23; A19: (F1 . m) | (dom ((F1 . m) - (F1 . n))) = F1 . m by A11, A9, A17, GRFUNC_1:23; integral' (M,((F1 . n) | (dom ((F1 . m) - (F1 . n))))) <= integral' (M,((F1 . m) | (dom ((F1 . m) - (F1 . n))))) by A16, A8, A15, A14, A13, Th70; hence K1 . n <= K1 . m by A6, A10, A18, A19; ::_thesis: verum end; then lim K1 = sup (rng K1) by Th54; then A20: K1 . 0 <= lim K1 by Th56; for n being Nat holds 0 <= K1 . n proof let n be Nat; ::_thesis: 0 <= K1 . n A21: F1 . n is_simple_func_in S by A3; K1 . n = integral' (M,(F1 . n)) by A6; hence 0 <= K1 . n by A4, A21, Th68; ::_thesis: verum end; hence 0 <= integral+ (M,f) by A7, A20; ::_thesis: verum end; theorem Th80: :: MESFUNC5:80 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative holds 0 <= integral+ (M,(f | A)) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative holds 0 <= integral+ (M,(f | A)) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative holds 0 <= integral+ (M,(f | A)) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for A being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative holds 0 <= integral+ (M,(f | A)) let f be PartFunc of X,ExtREAL; ::_thesis: for A being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative holds 0 <= integral+ (M,(f | A)) let A be Element of S; ::_thesis: ( ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative implies 0 <= integral+ (M,(f | A)) ) assume that A1: ex E being Element of S st ( E = dom f & f is_measurable_on E ) and A2: f is nonnegative ; ::_thesis: 0 <= integral+ (M,(f | A)) consider E being Element of S such that A3: E = dom f and A4: f is_measurable_on E by A1; set C = E /\ A; A5: E /\ A = dom (f | A) by A3, RELAT_1:61; A6: dom (f | A) = E /\ A by A3, RELAT_1:61 .= (dom f) /\ (E /\ A) by A3, XBOOLE_1:17, XBOOLE_1:28 .= dom (f | (E /\ A)) by RELAT_1:61 ; A7: for x being set st x in dom (f | A) holds (f | A) . x = (f | (E /\ A)) . x proof let x be set ; ::_thesis: ( x in dom (f | A) implies (f | A) . x = (f | (E /\ A)) . x ) assume A8: x in dom (f | A) ; ::_thesis: (f | A) . x = (f | (E /\ A)) . x then (f | A) . x = f . x by FUNCT_1:47; hence (f | A) . x = (f | (E /\ A)) . x by A6, A8, FUNCT_1:47; ::_thesis: verum end; A9: (dom f) /\ (E /\ A) = E /\ A by A3, XBOOLE_1:17, XBOOLE_1:28; f is_measurable_on E /\ A by A4, MESFUNC1:30, XBOOLE_1:17; then f | (E /\ A) is_measurable_on E /\ A by A9, Th42; then f | A is_measurable_on E /\ A by A6, A7, FUNCT_1:2; hence 0 <= integral+ (M,(f | A)) by A2, A5, Th15, Th79; ::_thesis: verum end; theorem Th81: :: MESFUNC5:81 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A, B being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & A misses B holds integral+ (M,(f | (A \/ B))) = (integral+ (M,(f | A))) + (integral+ (M,(f | B))) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A, B being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & A misses B holds integral+ (M,(f | (A \/ B))) = (integral+ (M,(f | A))) + (integral+ (M,(f | B))) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A, B being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & A misses B holds integral+ (M,(f | (A \/ B))) = (integral+ (M,(f | A))) + (integral+ (M,(f | B))) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for A, B being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & A misses B holds integral+ (M,(f | (A \/ B))) = (integral+ (M,(f | A))) + (integral+ (M,(f | B))) let f be PartFunc of X,ExtREAL; ::_thesis: for A, B being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & A misses B holds integral+ (M,(f | (A \/ B))) = (integral+ (M,(f | A))) + (integral+ (M,(f | B))) let A, B be Element of S; ::_thesis: ( ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & A misses B implies integral+ (M,(f | (A \/ B))) = (integral+ (M,(f | A))) + (integral+ (M,(f | B))) ) assume that A1: ex E being Element of S st ( E = dom f & f is_measurable_on E ) and A2: f is nonnegative and A3: A misses B ; ::_thesis: integral+ (M,(f | (A \/ B))) = (integral+ (M,(f | A))) + (integral+ (M,(f | B))) consider F0 being Functional_Sequence of X,ExtREAL, K0 being ExtREAL_sequence such that A4: for n being Nat holds ( F0 . n is_simple_func_in S & dom (F0 . n) = dom f ) and A5: for n being Nat holds F0 . n is nonnegative and A6: for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F0 . n) . x <= (F0 . m) . x and A7: for x being Element of X st x in dom f holds ( F0 # x is convergent & lim (F0 # x) = f . x ) and for n being Nat holds K0 . n = integral' (M,(F0 . n)) and K0 is convergent and integral+ (M,f) = lim K0 by A1, A2, Def15; deffunc H1( Nat) -> Element of bool [:X,ExtREAL:] = (F0 . $1) | B; deffunc H2( Nat) -> Element of bool [:X,ExtREAL:] = (F0 . $1) | A; consider FA being Functional_Sequence of X,ExtREAL such that A8: for n being Nat holds FA . n = H2(n) from SEQFUNC:sch_1(); consider E being Element of S such that A9: E = dom f and A10: f is_measurable_on E by A1; consider FB being Functional_Sequence of X,ExtREAL such that A11: for n being Nat holds FB . n = H1(n) from SEQFUNC:sch_1(); set DB = E /\ B; A12: E /\ B = dom (f | B) by A9, RELAT_1:61; then A13: (dom f) /\ (E /\ B) = E /\ B by RELAT_1:60, XBOOLE_1:28; then A14: dom (f | (E /\ B)) = dom (f | B) by A12, RELAT_1:61; for x being set st x in dom (f | (E /\ B)) holds (f | (E /\ B)) . x = (f | B) . x proof let x be set ; ::_thesis: ( x in dom (f | (E /\ B)) implies (f | (E /\ B)) . x = (f | B) . x ) assume A15: x in dom (f | (E /\ B)) ; ::_thesis: (f | (E /\ B)) . x = (f | B) . x then (f | B) . x = f . x by A14, FUNCT_1:47; hence (f | (E /\ B)) . x = (f | B) . x by A15, FUNCT_1:47; ::_thesis: verum end; then A16: f | (E /\ B) = f | B by A12, A13, FUNCT_1:2, RELAT_1:61; set DA = E /\ A; A17: E /\ A = dom (f | A) by A9, RELAT_1:61; then A18: (dom f) /\ (E /\ A) = E /\ A by RELAT_1:60, XBOOLE_1:28; then A19: dom (f | (E /\ A)) = dom (f | A) by A17, RELAT_1:61; for x being set st x in dom (f | (E /\ A)) holds (f | (E /\ A)) . x = (f | A) . x proof let x be set ; ::_thesis: ( x in dom (f | (E /\ A)) implies (f | (E /\ A)) . x = (f | A) . x ) assume A20: x in dom (f | (E /\ A)) ; ::_thesis: (f | (E /\ A)) . x = (f | A) . x then (f | A) . x = f . x by A19, FUNCT_1:47; hence (f | (E /\ A)) . x = (f | A) . x by A20, FUNCT_1:47; ::_thesis: verum end; then A21: f | (E /\ A) = f | A by A17, A18, FUNCT_1:2, RELAT_1:61; A22: for n being Nat holds ( FA . n is_simple_func_in S & FB . n is_simple_func_in S & dom (FA . n) = dom (f | A) & dom (FB . n) = dom (f | B) ) proof let n be Nat; ::_thesis: ( FA . n is_simple_func_in S & FB . n is_simple_func_in S & dom (FA . n) = dom (f | A) & dom (FB . n) = dom (f | B) ) reconsider n1 = n as Element of NAT by ORDINAL1:def_12; A23: FB . n1 = (F0 . n1) | B by A11; then A24: dom (FB . n) = (dom (F0 . n)) /\ B by RELAT_1:61; A25: FA . n1 = (F0 . n1) | A by A8; hence ( FA . n is_simple_func_in S & FB . n is_simple_func_in S ) by A4, A23, Th34; ::_thesis: ( dom (FA . n) = dom (f | A) & dom (FB . n) = dom (f | B) ) dom (FA . n) = (dom (F0 . n)) /\ A by A25, RELAT_1:61; hence ( dom (FA . n) = dom (f | A) & dom (FB . n) = dom (f | B) ) by A9, A4, A17, A12, A24; ::_thesis: verum end; A26: for x being Element of X st x in dom (f | A) holds ( FA # x is convergent & lim (FA # x) = (f | A) . x ) proof let x be Element of X; ::_thesis: ( x in dom (f | A) implies ( FA # x is convergent & lim (FA # x) = (f | A) . x ) ) assume A27: x in dom (f | A) ; ::_thesis: ( FA # x is convergent & lim (FA # x) = (f | A) . x ) now__::_thesis:_for_n_being_Element_of_NAT_holds_(FA_#_x)_._n_=_(F0_#_x)_._n let n be Element of NAT ; ::_thesis: (FA # x) . n = (F0 # x) . n (FA # x) . n = (FA . n) . x by Def13; then A28: (FA # x) . n = ((F0 . n) | A) . x by A8; dom ((F0 . n) | A) = dom (FA . n) by A8 .= dom (f | A) by A22 ; then (FA # x) . n = (F0 . n) . x by A27, A28, FUNCT_1:47; hence (FA # x) . n = (F0 # x) . n by Def13; ::_thesis: verum end; then A29: FA # x = F0 # x by FUNCT_2:63; x in (dom f) /\ A by A27, RELAT_1:61; then A30: x in dom f by XBOOLE_0:def_4; then lim (F0 # x) = f . x by A7; hence ( FA # x is convergent & lim (FA # x) = (f | A) . x ) by A7, A27, A30, A29, FUNCT_1:47; ::_thesis: verum end; A31: for x being Element of X st x in dom (f | B) holds ( FB # x is convergent & lim (FB # x) = (f | B) . x ) proof let x be Element of X; ::_thesis: ( x in dom (f | B) implies ( FB # x is convergent & lim (FB # x) = (f | B) . x ) ) assume A32: x in dom (f | B) ; ::_thesis: ( FB # x is convergent & lim (FB # x) = (f | B) . x ) now__::_thesis:_for_n_being_Element_of_NAT_holds_(FB_#_x)_._n_=_(F0_#_x)_._n let n be Element of NAT ; ::_thesis: (FB # x) . n = (F0 # x) . n A33: dom ((F0 . n) | B) = dom (FB . n) by A11 .= dom (f | B) by A22 ; thus (FB # x) . n = (FB . n) . x by Def13 .= ((F0 . n) | B) . x by A11 .= (F0 . n) . x by A32, A33, FUNCT_1:47 .= (F0 # x) . n by Def13 ; ::_thesis: verum end; then A34: FB # x = F0 # x by FUNCT_2:63; x in (dom f) /\ B by A32, RELAT_1:61; then A35: x in dom f by XBOOLE_0:def_4; then lim (F0 # x) = f . x by A7; hence ( FB # x is convergent & lim (FB # x) = (f | B) . x ) by A7, A32, A35, A34, FUNCT_1:47; ::_thesis: verum end; set C = E /\ (A \/ B); A36: E /\ (A \/ B) = (dom f) /\ (E /\ (A \/ B)) by A9, XBOOLE_1:17, XBOOLE_1:28; A37: dom (f | (A \/ B)) = E /\ (A \/ B) by A9, RELAT_1:61; then A38: dom (f | (A \/ B)) = dom (f | (E /\ (A \/ B))) by A36, RELAT_1:61; for x being set st x in dom (f | (A \/ B)) holds (f | (A \/ B)) . x = (f | (E /\ (A \/ B))) . x proof let x be set ; ::_thesis: ( x in dom (f | (A \/ B)) implies (f | (A \/ B)) . x = (f | (E /\ (A \/ B))) . x ) assume A39: x in dom (f | (A \/ B)) ; ::_thesis: (f | (A \/ B)) . x = (f | (E /\ (A \/ B))) . x then (f | (A \/ B)) . x = f . x by FUNCT_1:47; hence (f | (A \/ B)) . x = (f | (E /\ (A \/ B))) . x by A38, A39, FUNCT_1:47; ::_thesis: verum end; then A40: f | (A \/ B) = f | (E /\ (A \/ B)) by A38, FUNCT_1:2; f is_measurable_on E /\ (A \/ B) by A10, MESFUNC1:30, XBOOLE_1:17; then A41: f | (A \/ B) is_measurable_on E /\ (A \/ B) by A36, A40, Th42; f is_measurable_on E /\ B by A10, MESFUNC1:30, XBOOLE_1:17; then A42: f | B is_measurable_on E /\ B by A13, A16, Th42; A43: f | B is nonnegative by A2, Th15; f is_measurable_on E /\ A by A10, MESFUNC1:30, XBOOLE_1:17; then A44: f | A is_measurable_on E /\ A by A18, A21, Th42; A45: f | A is nonnegative by A2, Th15; deffunc H3( Nat) -> Element of bool [:X,ExtREAL:] = (F0 . $1) | (A \/ B); consider FAB being Functional_Sequence of X,ExtREAL such that A46: for n being Nat holds FAB . n = H3(n) from SEQFUNC:sch_1(); A47: for n being Nat holds ( FA . n is nonnegative & FB . n is nonnegative ) proof let n be Nat; ::_thesis: ( FA . n is nonnegative & FB . n is nonnegative ) reconsider n = n as Element of NAT by ORDINAL1:def_12; A48: (F0 . n) | B is nonnegative by A5, Th15; (F0 . n) | A is nonnegative by A5, Th15; hence ( FA . n is nonnegative & FB . n is nonnegative ) by A8, A11, A48; ::_thesis: verum end; A49: for n, m being Nat st n <= m holds for x being Element of X st x in dom (f | B) holds (FB . n) . x <= (FB . m) . x proof let n, m be Nat; ::_thesis: ( n <= m implies for x being Element of X st x in dom (f | B) holds (FB . n) . x <= (FB . m) . x ) assume A50: n <= m ; ::_thesis: for x being Element of X st x in dom (f | B) holds (FB . n) . x <= (FB . m) . x reconsider n = n, m = m as Element of NAT by ORDINAL1:def_12; let x be Element of X; ::_thesis: ( x in dom (f | B) implies (FB . n) . x <= (FB . m) . x ) assume A51: x in dom (f | B) ; ::_thesis: (FB . n) . x <= (FB . m) . x then x in (dom f) /\ B by RELAT_1:61; then A52: x in dom f by XBOOLE_0:def_4; dom ((F0 . m) | B) = dom (FB . m) by A11; then A53: dom ((F0 . m) | B) = dom (f | B) by A22; (FB . m) . x = ((F0 . m) | B) . x by A11; then A54: (FB . m) . x = (F0 . m) . x by A51, A53, FUNCT_1:47; dom ((F0 . n) | B) = dom (FB . n) by A11; then A55: dom ((F0 . n) | B) = dom (f | B) by A22; (FB . n) . x = ((F0 . n) | B) . x by A11; then (FB . n) . x = (F0 . n) . x by A51, A55, FUNCT_1:47; hence (FB . n) . x <= (FB . m) . x by A6, A50, A52, A54; ::_thesis: verum end; deffunc H4( Nat) -> Element of ExtREAL = integral' (M,(FA . $1)); consider KA being ExtREAL_sequence such that A56: for n being Element of NAT holds KA . n = H4(n) from FUNCT_2:sch_4(); deffunc H5( Nat) -> Element of ExtREAL = integral' (M,(FB . $1)); consider KB being ExtREAL_sequence such that A57: for n being Element of NAT holds KB . n = H5(n) from FUNCT_2:sch_4(); A58: now__::_thesis:_for_n_being_Nat_holds_KB_._n_=_H5(n) let n be Nat; ::_thesis: KB . n = H5(n) n in NAT by ORDINAL1:def_12; hence KB . n = H5(n) by A57; ::_thesis: verum end; A59: now__::_thesis:_for_n_being_Nat_holds_KA_._n_=_H4(n) let n be Nat; ::_thesis: KA . n = H4(n) n in NAT by ORDINAL1:def_12; hence KA . n = H4(n) by A56; ::_thesis: verum end; A60: for n being set holds ( ( n in dom KA implies -infty < KA . n ) & ( n in dom KB implies -infty < KB . n ) ) proof let n be set ; ::_thesis: ( ( n in dom KA implies -infty < KA . n ) & ( n in dom KB implies -infty < KB . n ) ) hereby ::_thesis: ( n in dom KB implies -infty < KB . n ) assume n in dom KA ; ::_thesis: -infty < KA . n then reconsider n1 = n as Element of NAT ; A61: FA . n1 is_simple_func_in S by A22; KA . n1 = integral' (M,(FA . n1)) by A59; hence -infty < KA . n by A47, A61, Th68; ::_thesis: verum end; assume n in dom KB ; ::_thesis: -infty < KB . n then reconsider n1 = n as Element of NAT ; A62: FB . n1 is_simple_func_in S by A22; KB . n1 = integral' (M,(FB . n1)) by A58; hence -infty < KB . n by A47, A62, Th68; ::_thesis: verum end; then A63: KB is () by Th10; deffunc H6( Nat) -> Element of ExtREAL = integral' (M,(FAB . $1)); consider KAB being ExtREAL_sequence such that A64: for n being Element of NAT holds KAB . n = H6(n) from FUNCT_2:sch_4(); A65: now__::_thesis:_for_n_being_Nat_holds_KAB_._n_=_H6(n) let n be Nat; ::_thesis: KAB . n = H6(n) n in NAT by ORDINAL1:def_12; hence KAB . n = H6(n) by A64; ::_thesis: verum end; A66: for n being Nat holds KAB . n = (KA . n) + (KB . n) proof let n be Nat; ::_thesis: KAB . n = (KA . n) + (KB . n) reconsider n = n as Element of NAT by ORDINAL1:def_12; A67: FA . n = (F0 . n) | A by A8; A68: FB . n = (F0 . n) | B by A11; A69: KAB . n = integral' (M,(FAB . n)) by A65 .= integral' (M,((F0 . n) | (A \/ B))) by A46 ; A70: KA . n = integral' (M,(FA . n)) by A59; F0 . n is_simple_func_in S by A4; then integral' (M,((F0 . n) | (A \/ B))) = (integral' (M,(FA . n))) + (integral' (M,(FB . n))) by A3, A5, A67, A68, Th67; hence KAB . n = (KA . n) + (KB . n) by A58, A69, A70; ::_thesis: verum end; A71: for n, m being Nat st n <= m holds for x being Element of X st x in dom (f | A) holds (FA . n) . x <= (FA . m) . x proof let n, m be Nat; ::_thesis: ( n <= m implies for x being Element of X st x in dom (f | A) holds (FA . n) . x <= (FA . m) . x ) assume A72: n <= m ; ::_thesis: for x being Element of X st x in dom (f | A) holds (FA . n) . x <= (FA . m) . x reconsider n = n, m = m as Element of NAT by ORDINAL1:def_12; let x be Element of X; ::_thesis: ( x in dom (f | A) implies (FA . n) . x <= (FA . m) . x ) assume A73: x in dom (f | A) ; ::_thesis: (FA . n) . x <= (FA . m) . x then x in (dom f) /\ A by RELAT_1:61; then A74: x in dom f by XBOOLE_0:def_4; dom ((F0 . m) | A) = dom (FA . m) by A8; then A75: dom ((F0 . m) | A) = dom (f | A) by A22; (FA . m) . x = ((F0 . m) | A) . x by A8; then A76: (FA . m) . x = (F0 . m) . x by A73, A75, FUNCT_1:47; dom ((F0 . n) | A) = dom (FA . n) by A8; then A77: dom ((F0 . n) | A) = dom (f | A) by A22; (FA . n) . x = ((F0 . n) | A) . x by A8; then (FA . n) . x = (F0 . n) . x by A73, A77, FUNCT_1:47; hence (FA . n) . x <= (FA . m) . x by A6, A72, A74, A76; ::_thesis: verum end; A78: for n, m being Nat st n <= m holds ( KA . n <= KA . m & KB . n <= KB . m ) proof let n, m be Nat; ::_thesis: ( n <= m implies ( KA . n <= KA . m & KB . n <= KB . m ) ) A79: FA . m is_simple_func_in S by A22; A80: dom (FA . m) = dom (f | A) by A22; A81: KA . m = integral' (M,(FA . m)) by A59; A82: dom (FA . n) = dom (f | A) by A22; assume A83: n <= m ; ::_thesis: ( KA . n <= KA . m & KB . n <= KB . m ) A84: for x being set st x in dom ((FA . m) - (FA . n)) holds (FA . n) . x <= (FA . m) . x proof let x be set ; ::_thesis: ( x in dom ((FA . m) - (FA . n)) implies (FA . n) . x <= (FA . m) . x ) assume x in dom ((FA . m) - (FA . n)) ; ::_thesis: (FA . n) . x <= (FA . m) . x then x in ((dom (FA . m)) /\ (dom (FA . n))) \ ((((FA . m) " {+infty}) /\ ((FA . n) " {+infty})) \/ (((FA . m) " {-infty}) /\ ((FA . n) " {-infty}))) by MESFUNC1:def_4; then x in (dom (FA . m)) /\ (dom (FA . n)) by XBOOLE_0:def_5; hence (FA . n) . x <= (FA . m) . x by A71, A83, A82, A80; ::_thesis: verum end; A85: FA . m is nonnegative by A47; A86: FA . n is nonnegative by A47; A87: FA . n is_simple_func_in S by A22; then A88: dom ((FA . m) - (FA . n)) = (dom (FA . m)) /\ (dom (FA . n)) by A79, A86, A85, A84, Th69; then A89: (FA . m) | (dom ((FA . m) - (FA . n))) = FA . m by A82, A80, GRFUNC_1:23; A90: (FA . n) | (dom ((FA . m) - (FA . n))) = FA . n by A82, A80, A88, GRFUNC_1:23; integral' (M,((FA . n) | (dom ((FA . m) - (FA . n))))) <= integral' (M,((FA . m) | (dom ((FA . m) - (FA . n))))) by A87, A79, A86, A85, A84, Th70; hence KA . n <= KA . m by A59, A81, A89, A90; ::_thesis: KB . n <= KB . m A91: FB . m is_simple_func_in S by A22; A92: FB . n is nonnegative by A47; A93: FB . m is nonnegative by A47; A94: KB . m = integral' (M,(FB . m)) by A58; A95: dom (FB . m) = dom (f | B) by A22; A96: dom (FB . n) = dom (f | B) by A22; A97: for x being set st x in dom ((FB . m) - (FB . n)) holds (FB . n) . x <= (FB . m) . x proof let x be set ; ::_thesis: ( x in dom ((FB . m) - (FB . n)) implies (FB . n) . x <= (FB . m) . x ) assume x in dom ((FB . m) - (FB . n)) ; ::_thesis: (FB . n) . x <= (FB . m) . x then x in ((dom (FB . m)) /\ (dom (FB . n))) \ ((((FB . m) " {+infty}) /\ ((FB . n) " {+infty})) \/ (((FB . m) " {-infty}) /\ ((FB . n) " {-infty}))) by MESFUNC1:def_4; then x in (dom (FB . m)) /\ (dom (FB . n)) by XBOOLE_0:def_5; hence (FB . n) . x <= (FB . m) . x by A49, A83, A96, A95; ::_thesis: verum end; A98: FB . n is_simple_func_in S by A22; then A99: dom ((FB . m) - (FB . n)) = (dom (FB . m)) /\ (dom (FB . n)) by A91, A92, A93, A97, Th69; then A100: (FB . m) | (dom ((FB . m) - (FB . n))) = FB . m by A96, A95, GRFUNC_1:23; A101: (FB . n) | (dom ((FB . m) - (FB . n))) = FB . n by A96, A95, A99, GRFUNC_1:23; integral' (M,((FB . n) | (dom ((FB . m) - (FB . n))))) <= integral' (M,((FB . m) | (dom ((FB . m) - (FB . n))))) by A98, A91, A92, A93, A97, Th70; hence KB . n <= KB . m by A58, A94, A100, A101; ::_thesis: verum end; then A102: for n, m being Nat st n <= m holds KA . n <= KA . m ; then KA is convergent by Th54; then A103: integral+ (M,(f | A)) = lim KA by A17, A44, A45, A22, A47, A71, A26, A59, Def15; A104: ( ( for n being Nat holds ( FAB . n is_simple_func_in S & dom (FAB . n) = dom (f | (A \/ B)) ) ) & ( for n being Nat for x being Element of X st x in dom (f | (A \/ B)) holds (FAB . n) . x = (F0 . n) . x ) & ( for n being Nat holds FAB . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom (f | (A \/ B)) holds (FAB . n) . x <= (FAB . m) . x ) & ( for x being Element of X st x in dom (f | (A \/ B)) holds ( FAB # x is convergent & lim (FAB # x) = (f | (A \/ B)) . x ) ) ) proof hereby ::_thesis: ( ( for n being Nat for x being Element of X st x in dom (f | (A \/ B)) holds (FAB . n) . x = (F0 . n) . x ) & ( for n being Nat holds FAB . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom (f | (A \/ B)) holds (FAB . n) . x <= (FAB . m) . x ) & ( for x being Element of X st x in dom (f | (A \/ B)) holds ( FAB # x is convergent & lim (FAB # x) = (f | (A \/ B)) . x ) ) ) let n be Nat; ::_thesis: ( FAB . n is_simple_func_in S & dom (FAB . n) = dom (f | (A \/ B)) ) FAB . n = (F0 . n) | (A \/ B) by A46; hence FAB . n is_simple_func_in S by A4, Th34; ::_thesis: dom (FAB . n) = dom (f | (A \/ B)) thus dom (FAB . n) = dom ((F0 . n) | (A \/ B)) by A46 .= (dom (F0 . n)) /\ (A \/ B) by RELAT_1:61 .= (dom f) /\ (A \/ B) by A4 .= dom (f | (A \/ B)) by RELAT_1:61 ; ::_thesis: verum end; hereby ::_thesis: ( ( for n being Nat holds FAB . n is nonnegative ) & ( for n, m being Nat st n <= m holds for x being Element of X st x in dom (f | (A \/ B)) holds (FAB . n) . x <= (FAB . m) . x ) & ( for x being Element of X st x in dom (f | (A \/ B)) holds ( FAB # x is convergent & lim (FAB # x) = (f | (A \/ B)) . x ) ) ) let n be Nat; ::_thesis: for x being Element of X st x in dom (f | (A \/ B)) holds (FAB . n) . x = (F0 . n) . x let x be Element of X; ::_thesis: ( x in dom (f | (A \/ B)) implies (FAB . n) . x = (F0 . n) . x ) assume x in dom (f | (A \/ B)) ; ::_thesis: (FAB . n) . x = (F0 . n) . x then A107: x in dom (FAB . n) by A105; FAB . n = (F0 . n) | (A \/ B) by A46; hence (FAB . n) . x = (F0 . n) . x by A107, FUNCT_1:47; ::_thesis: verum end; hereby ::_thesis: ( ( for n, m being Nat st n <= m holds for x being Element of X st x in dom (f | (A \/ B)) holds (FAB . n) . x <= (FAB . m) . x ) & ( for x being Element of X st x in dom (f | (A \/ B)) holds ( FAB # x is convergent & lim (FAB # x) = (f | (A \/ B)) . x ) ) ) let n be Nat; ::_thesis: FAB . n is nonnegative reconsider n1 = n as Element of NAT by ORDINAL1:def_12; (F0 . n1) | (A \/ B) is nonnegative by A5, Th15; hence FAB . n is nonnegative by A46; ::_thesis: verum end; hereby ::_thesis: for x being Element of X st x in dom (f | (A \/ B)) holds ( FAB # x is convergent & lim (FAB # x) = (f | (A \/ B)) . x ) let n, m be Nat; ::_thesis: ( n <= m implies for x being Element of X st x in dom (f | (A \/ B)) holds (FAB . n) . x <= (FAB . m) . x ) assume A108: n <= m ; ::_thesis: for x being Element of X st x in dom (f | (A \/ B)) holds (FAB . n) . x <= (FAB . m) . x now__::_thesis:_for_x_being_Element_of_X_st_x_in_dom_(f_|_(A_\/_B))_holds_ (FAB_._n)_._x_<=_(FAB_._m)_._x let x be Element of X; ::_thesis: ( x in dom (f | (A \/ B)) implies (FAB . n) . x <= (FAB . m) . x ) assume A109: x in dom (f | (A \/ B)) ; ::_thesis: (FAB . n) . x <= (FAB . m) . x then A110: (FAB . m) . x = (F0 . m) . x by A106; x in (dom f) /\ (A \/ B) by A109, RELAT_1:61; then A111: x in dom f by XBOOLE_0:def_4; (FAB . n) . x = (F0 . n) . x by A106, A109; hence (FAB . n) . x <= (FAB . m) . x by A6, A108, A111, A110; ::_thesis: verum end; hence for x being Element of X st x in dom (f | (A \/ B)) holds (FAB . n) . x <= (FAB . m) . x ; ::_thesis: verum end; hereby ::_thesis: verum let x be Element of X; ::_thesis: ( x in dom (f | (A \/ B)) implies ( FAB # x is convergent & lim (FAB # x) = (f | (A \/ B)) . x ) ) assume A112: x in dom (f | (A \/ B)) ; ::_thesis: ( FAB # x is convergent & lim (FAB # x) = (f | (A \/ B)) . x ) then x in (dom f) /\ (A \/ B) by RELAT_1:61; then A113: x in dom f by XBOOLE_0:def_4; A114: now__::_thesis:_for_n_being_Element_of_NAT_holds_(FAB_#_x)_._n_=_(F0_#_x)_._n let n be Element of NAT ; ::_thesis: (FAB # x) . n = (F0 # x) . n thus (FAB # x) . n = (FAB . n) . x by Def13 .= (F0 . n) . x by A106, A112 .= (F0 # x) . n by Def13 ; ::_thesis: verum end; then FAB # x = F0 # x by FUNCT_2:63; hence FAB # x is convergent by A7, A113; ::_thesis: lim (FAB # x) = (f | (A \/ B)) . x thus lim (FAB # x) = lim (F0 # x) by A114, FUNCT_2:63 .= f . x by A7, A113 .= (f | (A \/ B)) . x by A112, FUNCT_1:47 ; ::_thesis: verum end; end; for n, m being Nat st n <= m holds KAB . n <= KAB . m proof let n, m be Nat; ::_thesis: ( n <= m implies KAB . n <= KAB . m ) assume A115: n <= m ; ::_thesis: KAB . n <= KAB . m reconsider n = n, m = m as Element of NAT by ORDINAL1:def_12; A116: dom (FAB . m) = dom (f | (A \/ B)) by A104; A117: dom (FAB . n) = dom (f | (A \/ B)) by A104; A118: for x being set st x in dom ((FAB . m) - (FAB . n)) holds (FAB . n) . x <= (FAB . m) . x proof let x be set ; ::_thesis: ( x in dom ((FAB . m) - (FAB . n)) implies (FAB . n) . x <= (FAB . m) . x ) assume x in dom ((FAB . m) - (FAB . n)) ; ::_thesis: (FAB . n) . x <= (FAB . m) . x then x in ((dom (FAB . m)) /\ (dom (FAB . n))) \ ((((FAB . m) " {+infty}) /\ ((FAB . n) " {+infty})) \/ (((FAB . m) " {-infty}) /\ ((FAB . n) " {-infty}))) by MESFUNC1:def_4; then x in (dom (FAB . m)) /\ (dom (FAB . n)) by XBOOLE_0:def_5; hence (FAB . n) . x <= (FAB . m) . x by A104, A115, A117, A116; ::_thesis: verum end; A119: KAB . m = integral' (M,(FAB . m)) by A65; A120: FAB . m is_simple_func_in S by A104; A121: FAB . m is nonnegative by A104; A122: FAB . n is nonnegative by A104; A123: FAB . n is_simple_func_in S by A104; then A124: dom ((FAB . m) - (FAB . n)) = (dom (FAB . m)) /\ (dom (FAB . n)) by A120, A122, A121, A118, Th69; then A125: (FAB . m) | (dom ((FAB . m) - (FAB . n))) = FAB . m by A117, A116, GRFUNC_1:23; A126: (FAB . n) | (dom ((FAB . m) - (FAB . n))) = FAB . n by A117, A116, A124, GRFUNC_1:23; integral' (M,((FAB . n) | (dom ((FAB . m) - (FAB . n))))) <= integral' (M,((FAB . m) | (dom ((FAB . m) - (FAB . n))))) by A123, A120, A122, A121, A118, Th70; hence KAB . n <= KAB . m by A65, A119, A125, A126; ::_thesis: verum end; then A127: KAB is convergent by Th54; A128: for n, m being Nat st n <= m holds KB . n <= KB . m by A78; then KB is convergent by Th54; then A129: integral+ (M,(f | B)) = lim KB by A12, A42, A43, A22, A47, A49, A31, A58, Def15; f | (A \/ B) is nonnegative by A2, Th15; then A130: integral+ (M,(f | (A \/ B))) = lim KAB by A37, A41, A65, A104, A127, Def15; KA is () by A60, Th10; hence integral+ (M,(f | (A \/ B))) = (integral+ (M,(f | A))) + (integral+ (M,(f | B))) by A130, A102, A128, A103, A129, A66, A63, Th61; ::_thesis: verum end; theorem Th82: :: MESFUNC5:82 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & M . A = 0 holds integral+ (M,(f | A)) = 0 proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & M . A = 0 holds integral+ (M,(f | A)) = 0 let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & M . A = 0 holds integral+ (M,(f | A)) = 0 let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for A being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & M . A = 0 holds integral+ (M,(f | A)) = 0 let f be PartFunc of X,ExtREAL; ::_thesis: for A being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & M . A = 0 holds integral+ (M,(f | A)) = 0 let A be Element of S; ::_thesis: ( ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & M . A = 0 implies integral+ (M,(f | A)) = 0 ) assume that A1: ex E being Element of S st ( E = dom f & f is_measurable_on E ) and A2: f is nonnegative and A3: M . A = 0 ; ::_thesis: integral+ (M,(f | A)) = 0 consider F0 being Functional_Sequence of X,ExtREAL, K0 being ExtREAL_sequence such that A4: for n being Nat holds ( F0 . n is_simple_func_in S & dom (F0 . n) = dom f ) and A5: for n being Nat holds F0 . n is nonnegative and A6: for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F0 . n) . x <= (F0 . m) . x and A7: for x being Element of X st x in dom f holds ( F0 # x is convergent & lim (F0 # x) = f . x ) and for n being Nat holds K0 . n = integral' (M,(F0 . n)) and K0 is convergent and integral+ (M,f) = lim K0 by A1, A2, Def15; deffunc H1( Nat) -> Element of bool [:X,ExtREAL:] = (F0 . $1) | A; consider FA being Functional_Sequence of X,ExtREAL such that A8: for n being Nat holds FA . n = H1(n) from SEQFUNC:sch_1(); consider E being Element of S such that A9: E = dom f and A10: f is_measurable_on E by A1; set C = E /\ A; A11: f | A is nonnegative by A2, Th15; A12: (dom f) /\ (E /\ A) = E /\ A by A9, XBOOLE_1:17, XBOOLE_1:28; then A13: dom (f | (E /\ A)) = E /\ A by RELAT_1:61; then A14: dom (f | (E /\ A)) = dom (f | A) by A9, RELAT_1:61; for x being set st x in dom (f | A) holds (f | A) . x = (f | (E /\ A)) . x proof let x be set ; ::_thesis: ( x in dom (f | A) implies (f | A) . x = (f | (E /\ A)) . x ) assume A15: x in dom (f | A) ; ::_thesis: (f | A) . x = (f | (E /\ A)) . x then (f | A) . x = f . x by FUNCT_1:47; hence (f | A) . x = (f | (E /\ A)) . x by A14, A15, FUNCT_1:47; ::_thesis: verum end; then A16: f | A = f | (E /\ A) by A9, A13, FUNCT_1:2, RELAT_1:61; f is_measurable_on E /\ A by A10, MESFUNC1:30, XBOOLE_1:17; then A17: f | A is_measurable_on E /\ A by A12, A16, Th42; A18: for n being Nat holds FA . n is nonnegative proof let n be Nat; ::_thesis: FA . n is nonnegative reconsider n = n as Element of NAT by ORDINAL1:def_12; (F0 . n) | A is nonnegative by A5, Th15; hence FA . n is nonnegative by A8; ::_thesis: verum end; deffunc H2( Nat) -> Element of ExtREAL = integral' (M,(FA . $1)); consider KA being ExtREAL_sequence such that A19: for n being Element of NAT holds KA . n = H2(n) from FUNCT_2:sch_4(); A20: now__::_thesis:_for_n_being_Nat_holds_KA_._n_=_H2(n) let n be Nat; ::_thesis: KA . n = H2(n) n in NAT by ORDINAL1:def_12; hence KA . n = H2(n) by A19; ::_thesis: verum end; A21: for n being Nat holds KA . n = R_EAL 0 proof let n be Nat; ::_thesis: KA . n = R_EAL 0 reconsider n = n as Element of NAT by ORDINAL1:def_12; F0 . n is_simple_func_in S by A4; then integral' (M,((F0 . n) | A)) = R_EAL 0 by A3, A5, Th73; then integral' (M,(FA . n)) = R_EAL 0 by A8; hence KA . n = R_EAL 0 by A20; ::_thesis: verum end; then A22: lim KA = R_EAL 0 by Th60; A23: E /\ A = dom (f | A) by A9, RELAT_1:61; A24: for n being Nat holds ( FA . n is_simple_func_in S & dom (FA . n) = dom (f | A) ) proof let n be Nat; ::_thesis: ( FA . n is_simple_func_in S & dom (FA . n) = dom (f | A) ) reconsider n1 = n as Element of NAT by ORDINAL1:def_12; FA . n1 = (F0 . n1) | A by A8; hence FA . n is_simple_func_in S by A4, Th34; ::_thesis: dom (FA . n) = dom (f | A) dom (FA . n1) = dom ((F0 . n1) | A) by A8; then dom (FA . n) = (dom (F0 . n)) /\ A by RELAT_1:61; hence dom (FA . n) = dom (f | A) by A9, A4, A23; ::_thesis: verum end; A25: for x being Element of X st x in dom (f | A) holds ( FA # x is convergent & lim (FA # x) = (f | A) . x ) proof let x be Element of X; ::_thesis: ( x in dom (f | A) implies ( FA # x is convergent & lim (FA # x) = (f | A) . x ) ) assume A26: x in dom (f | A) ; ::_thesis: ( FA # x is convergent & lim (FA # x) = (f | A) . x ) now__::_thesis:_for_n_being_Element_of_NAT_holds_(FA_#_x)_._n_=_(F0_#_x)_._n let n be Element of NAT ; ::_thesis: (FA # x) . n = (F0 # x) . n A27: dom ((F0 . n) | A) = dom (FA . n) by A8 .= dom (f | A) by A24 ; thus (FA # x) . n = (FA . n) . x by Def13 .= ((F0 . n) | A) . x by A8 .= (F0 . n) . x by A26, A27, FUNCT_1:47 .= (F0 # x) . n by Def13 ; ::_thesis: verum end; then A28: FA # x = F0 # x by FUNCT_2:63; x in (dom f) /\ A by A26, RELAT_1:61; then A29: x in dom f by XBOOLE_0:def_4; then lim (F0 # x) = f . x by A7; hence ( FA # x is convergent & lim (FA # x) = (f | A) . x ) by A7, A26, A29, A28, FUNCT_1:47; ::_thesis: verum end; A30: for n, m being Nat st n <= m holds for x being Element of X st x in dom (f | A) holds (FA . n) . x <= (FA . m) . x proof let n, m be Nat; ::_thesis: ( n <= m implies for x being Element of X st x in dom (f | A) holds (FA . n) . x <= (FA . m) . x ) assume A31: n <= m ; ::_thesis: for x being Element of X st x in dom (f | A) holds (FA . n) . x <= (FA . m) . x let x be Element of X; ::_thesis: ( x in dom (f | A) implies (FA . n) . x <= (FA . m) . x ) reconsider n = n, m = m as Element of NAT by ORDINAL1:def_12; assume A32: x in dom (f | A) ; ::_thesis: (FA . n) . x <= (FA . m) . x then x in (dom f) /\ A by RELAT_1:61; then A33: x in dom f by XBOOLE_0:def_4; dom ((F0 . m) | A) = dom (FA . m) by A8; then A34: dom ((F0 . m) | A) = dom (f | A) by A24; (FA . m) . x = ((F0 . m) | A) . x by A8; then A35: (FA . m) . x = (F0 . m) . x by A32, A34, FUNCT_1:47; dom ((F0 . n) | A) = dom (FA . n) by A8; then A36: dom ((F0 . n) | A) = dom (f | A) by A24; (FA . n) . x = ((F0 . n) | A) . x by A8; then (FA . n) . x = (F0 . n) . x by A32, A36, FUNCT_1:47; hence (FA . n) . x <= (FA . m) . x by A6, A31, A33, A35; ::_thesis: verum end; KA is convergent by A21, Th60; hence integral+ (M,(f | A)) = 0 by A17, A20, A23, A11, A24, A18, A30, A25, A22, Def15; ::_thesis: verum end; theorem Th83: :: MESFUNC5:83 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A, B being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & A c= B holds integral+ (M,(f | A)) <= integral+ (M,(f | B)) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A, B being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & A c= B holds integral+ (M,(f | A)) <= integral+ (M,(f | B)) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A, B being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & A c= B holds integral+ (M,(f | A)) <= integral+ (M,(f | B)) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for A, B being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & A c= B holds integral+ (M,(f | A)) <= integral+ (M,(f | B)) let f be PartFunc of X,ExtREAL; ::_thesis: for A, B being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & A c= B holds integral+ (M,(f | A)) <= integral+ (M,(f | B)) let A, B be Element of S; ::_thesis: ( ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & A c= B implies integral+ (M,(f | A)) <= integral+ (M,(f | B)) ) assume that A1: ex E being Element of S st ( E = dom f & f is_measurable_on E ) and A2: f is nonnegative and A3: A c= B ; ::_thesis: integral+ (M,(f | A)) <= integral+ (M,(f | B)) set A9 = A /\ B; A4: A /\ B = A by A3, XBOOLE_1:28; set B9 = B \ A; A5: (A /\ B) \/ (B \ A) = B by XBOOLE_1:51; integral+ (M,(f | ((A /\ B) \/ (B \ A)))) = (integral+ (M,(f | (A /\ B)))) + (integral+ (M,(f | (B \ A)))) by A1, A2, Th81, XBOOLE_1:89; hence integral+ (M,(f | A)) <= integral+ (M,(f | B)) by A1, A2, A4, A5, Th80, XXREAL_3:39; ::_thesis: verum end; theorem Th84: :: MESFUNC5:84 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for E, A being Element of S st f is nonnegative & E = dom f & f is_measurable_on E & M . A = 0 holds integral+ (M,(f | (E \ A))) = integral+ (M,f) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for E, A being Element of S st f is nonnegative & E = dom f & f is_measurable_on E & M . A = 0 holds integral+ (M,(f | (E \ A))) = integral+ (M,f) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for E, A being Element of S st f is nonnegative & E = dom f & f is_measurable_on E & M . A = 0 holds integral+ (M,(f | (E \ A))) = integral+ (M,f) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for E, A being Element of S st f is nonnegative & E = dom f & f is_measurable_on E & M . A = 0 holds integral+ (M,(f | (E \ A))) = integral+ (M,f) let f be PartFunc of X,ExtREAL; ::_thesis: for E, A being Element of S st f is nonnegative & E = dom f & f is_measurable_on E & M . A = 0 holds integral+ (M,(f | (E \ A))) = integral+ (M,f) let E, A be Element of S; ::_thesis: ( f is nonnegative & E = dom f & f is_measurable_on E & M . A = 0 implies integral+ (M,(f | (E \ A))) = integral+ (M,f) ) assume that A1: f is nonnegative and A2: E = dom f and A3: f is_measurable_on E and A4: M . A = 0 ; ::_thesis: integral+ (M,(f | (E \ A))) = integral+ (M,f) set B = E \ A; A \/ (E \ A) = A \/ E by XBOOLE_1:39; then A5: dom f = (dom f) /\ (A \/ (E \ A)) by A2, XBOOLE_1:7, XBOOLE_1:28 .= dom (f | (A \/ (E \ A))) by RELAT_1:61 ; for x being set st x in dom (f | (A \/ (E \ A))) holds (f | (A \/ (E \ A))) . x = f . x by FUNCT_1:47; then A6: f | (A \/ (E \ A)) = f by A5, FUNCT_1:2; integral+ (M,(f | (A \/ (E \ A)))) = (integral+ (M,(f | A))) + (integral+ (M,(f | (E \ A)))) by A1, A2, A3, Th81, XBOOLE_1:79; then integral+ (M,f) = 0. + (integral+ (M,(f | (E \ A)))) by A1, A2, A3, A4, A6, Th82; hence integral+ (M,(f | (E \ A))) = integral+ (M,f) by XXREAL_3:4; ::_thesis: verum end; theorem Th85: :: MESFUNC5:85 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st ex E being Element of S st ( E = dom f & E = dom g & f is_measurable_on E & g is_measurable_on E ) & f is nonnegative & g is nonnegative & ( for x being Element of X st x in dom g holds g . x <= f . x ) holds integral+ (M,g) <= integral+ (M,f) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st ex E being Element of S st ( E = dom f & E = dom g & f is_measurable_on E & g is_measurable_on E ) & f is nonnegative & g is nonnegative & ( for x being Element of X st x in dom g holds g . x <= f . x ) holds integral+ (M,g) <= integral+ (M,f) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st ex E being Element of S st ( E = dom f & E = dom g & f is_measurable_on E & g is_measurable_on E ) & f is nonnegative & g is nonnegative & ( for x being Element of X st x in dom g holds g . x <= f . x ) holds integral+ (M,g) <= integral+ (M,f) let M be sigma_Measure of S; ::_thesis: for f, g being PartFunc of X,ExtREAL st ex E being Element of S st ( E = dom f & E = dom g & f is_measurable_on E & g is_measurable_on E ) & f is nonnegative & g is nonnegative & ( for x being Element of X st x in dom g holds g . x <= f . x ) holds integral+ (M,g) <= integral+ (M,f) let f, g be PartFunc of X,ExtREAL; ::_thesis: ( ex E being Element of S st ( E = dom f & E = dom g & f is_measurable_on E & g is_measurable_on E ) & f is nonnegative & g is nonnegative & ( for x being Element of X st x in dom g holds g . x <= f . x ) implies integral+ (M,g) <= integral+ (M,f) ) assume that A1: ex A being Element of S st ( A = dom f & A = dom g & f is_measurable_on A & g is_measurable_on A ) and A2: f is nonnegative and A3: g is nonnegative and A4: for x being Element of X st x in dom g holds g . x <= f . x ; ::_thesis: integral+ (M,g) <= integral+ (M,f) consider G being Functional_Sequence of X,ExtREAL, L being ExtREAL_sequence such that A5: for n being Nat holds ( G . n is_simple_func_in S & dom (G . n) = dom g ) and A6: for n being Nat holds G . n is nonnegative and A7: for n, m being Nat st n <= m holds for x being Element of X st x in dom g holds (G . n) . x <= (G . m) . x and A8: for x being Element of X st x in dom g holds ( G # x is convergent & lim (G # x) = g . x ) and A9: for n being Nat holds L . n = integral' (M,(G . n)) and L is convergent and A10: integral+ (M,g) = lim L by A1, A3, Def15; consider F being Functional_Sequence of X,ExtREAL, K being ExtREAL_sequence such that A11: for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = dom f ) and A12: for n being Nat holds F . n is nonnegative and A13: for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F . n) . x <= (F . m) . x and A14: for x being Element of X st x in dom f holds ( F # x is convergent & lim (F # x) = f . x ) and A15: for n being Nat holds K . n = integral' (M,(F . n)) and K is convergent and A16: integral+ (M,f) = lim K by A1, A2, Def15; consider A being Element of S such that A17: A = dom f and A18: A = dom g and f is_measurable_on A and g is_measurable_on A by A1; A19: for x being Element of X st x in A holds lim (G # x) = sup (rng (G # x)) proof let x be Element of X; ::_thesis: ( x in A implies lim (G # x) = sup (rng (G # x)) ) assume A20: x in A ; ::_thesis: lim (G # x) = sup (rng (G # x)) now__::_thesis:_for_n,_m_being_Nat_st_n_<=_m_holds_ (G_#_x)_._n_<=_(G_#_x)_._m let n, m be Nat; ::_thesis: ( n <= m implies (G # x) . n <= (G # x) . m ) assume A21: n <= m ; ::_thesis: (G # x) . n <= (G # x) . m A22: (G # x) . m = (G . m) . x by Def13; (G # x) . n = (G . n) . x by Def13; hence (G # x) . n <= (G # x) . m by A18, A7, A20, A21, A22; ::_thesis: verum end; hence lim (G # x) = sup (rng (G # x)) by Th54; ::_thesis: verum end; A23: for n0 being Nat holds ( L is convergent & sup (rng L) = lim L ) proof let n0 be Nat; ::_thesis: ( L is convergent & sup (rng L) = lim L ) set ff = G . n0; A24: dom (G . n0) = A by A18, A5; A25: for x being Element of X st x in dom (G . n0) holds ( G # x is convergent & (G . n0) . x <= lim (G # x) ) proof let x be Element of X; ::_thesis: ( x in dom (G . n0) implies ( G # x is convergent & (G . n0) . x <= lim (G # x) ) ) assume A26: x in dom (G . n0) ; ::_thesis: ( G # x is convergent & (G . n0) . x <= lim (G # x) ) A27: (G # x) . n0 <= sup (rng (G # x)) by Th56; (G . n0) . x = (G # x) . n0 by Def13; hence ( G # x is convergent & (G . n0) . x <= lim (G # x) ) by A18, A8, A19, A24, A26, A27; ::_thesis: verum end; G . n0 is_simple_func_in S by A5; then consider FF being ExtREAL_sequence such that A28: for n being Nat holds FF . n = integral' (M,(G . n)) and A29: FF is convergent and A30: sup (rng FF) = lim FF and integral' (M,(G . n0)) <= lim FF by A18, A5, A6, A7, A24, A25, Th75; now__::_thesis:_for_n_being_Element_of_NAT_holds_FF_._n_=_L_._n let n be Element of NAT ; ::_thesis: FF . n = L . n thus FF . n = integral' (M,(G . n)) by A28 .= L . n by A9 ; ::_thesis: verum end; then FF = L by FUNCT_2:63; hence ( L is convergent & sup (rng L) = lim L ) by A29, A30; ::_thesis: verum end; for n0 being Nat holds ( K is convergent & sup (rng K) = lim K & L . n0 <= lim K ) proof let n0 be Nat; ::_thesis: ( K is convergent & sup (rng K) = lim K & L . n0 <= lim K ) set gg = G . n0; A31: G . n0 is nonnegative by A6; A32: dom (G . n0) = A by A18, A5; A33: for x being Element of X st x in dom (G . n0) holds ( F # x is convergent & (G . n0) . x <= lim (F # x) ) proof let x be Element of X; ::_thesis: ( x in dom (G . n0) implies ( F # x is convergent & (G . n0) . x <= lim (F # x) ) ) assume A34: x in dom (G . n0) ; ::_thesis: ( F # x is convergent & (G . n0) . x <= lim (F # x) ) A35: (G # x) . n0 <= sup (rng (G # x)) by Th56; (G . n0) . x = (G # x) . n0 by Def13; then (G . n0) . x <= lim (G # x) by A19, A32, A34, A35; then A36: (G . n0) . x <= g . x by A18, A8, A32, A34; g . x <= f . x by A1, A4, A17, A32, A34; then g . x <= lim (F # x) by A17, A14, A32, A34; hence ( F # x is convergent & (G . n0) . x <= lim (F # x) ) by A17, A14, A32, A34, A36, XXREAL_0:2; ::_thesis: verum end; G . n0 is_simple_func_in S by A5; then consider GG being ExtREAL_sequence such that A37: for n being Nat holds GG . n = integral' (M,(F . n)) and A38: GG is convergent and A39: sup (rng GG) = lim GG and A40: integral' (M,(G . n0)) <= lim GG by A17, A11, A12, A13, A32, A31, A33, Th75; now__::_thesis:_for_n_being_Element_of_NAT_holds_GG_._n_=_K_._n let n be Element of NAT ; ::_thesis: GG . n = K . n GG . n = integral' (M,(F . n)) by A37; hence GG . n = K . n by A15; ::_thesis: verum end; then GG = K by FUNCT_2:63; hence ( K is convergent & sup (rng K) = lim K & L . n0 <= lim K ) by A9, A38, A39, A40; ::_thesis: verum end; hence integral+ (M,g) <= integral+ (M,f) by A16, A10, A23, Th57; ::_thesis: verum end; theorem Th86: :: MESFUNC5:86 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for c being Real st 0 <= c & ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative holds integral+ (M,(c (#) f)) = (R_EAL c) * (integral+ (M,f)) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for c being Real st 0 <= c & ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative holds integral+ (M,(c (#) f)) = (R_EAL c) * (integral+ (M,f)) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for c being Real st 0 <= c & ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative holds integral+ (M,(c (#) f)) = (R_EAL c) * (integral+ (M,f)) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for c being Real st 0 <= c & ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative holds integral+ (M,(c (#) f)) = (R_EAL c) * (integral+ (M,f)) let f be PartFunc of X,ExtREAL; ::_thesis: for c being Real st 0 <= c & ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative holds integral+ (M,(c (#) f)) = (R_EAL c) * (integral+ (M,f)) let c be Real; ::_thesis: ( 0 <= c & ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative implies integral+ (M,(c (#) f)) = (R_EAL c) * (integral+ (M,f)) ) assume that A1: 0 <= c and A2: ex A being Element of S st ( A = dom f & f is_measurable_on A ) and A3: f is nonnegative ; ::_thesis: integral+ (M,(c (#) f)) = (R_EAL c) * (integral+ (M,f)) consider F1 being Functional_Sequence of X,ExtREAL, K1 being ExtREAL_sequence such that A4: for n being Nat holds ( F1 . n is_simple_func_in S & dom (F1 . n) = dom f ) and A5: for n being Nat holds F1 . n is nonnegative and A6: for n, m being Nat st n <= m holds for x being Element of X st x in dom f holds (F1 . n) . x <= (F1 . m) . x and A7: for x being Element of X st x in dom f holds ( F1 # x is convergent & lim (F1 # x) = f . x ) and A8: for n being Nat holds K1 . n = integral' (M,(F1 . n)) and K1 is convergent and A9: integral+ (M,f) = lim K1 by A2, A3, Def15; deffunc H1( Nat) -> Element of bool [:X,ExtREAL:] = c (#) (F1 . $1); consider F being Functional_Sequence of X,ExtREAL such that A10: for n being Nat holds F . n = H1(n) from SEQFUNC:sch_1(); A11: c (#) f is nonnegative by A1, A3, Th20; A12: for n being Nat holds F . n is nonnegative proof let n be Nat; ::_thesis: F . n is nonnegative reconsider n = n as Element of NAT by ORDINAL1:def_12; F1 . n is nonnegative by A5; then c (#) (F1 . n) is nonnegative by A1, Th20; hence F . n is nonnegative by A10; ::_thesis: verum end; consider A being Element of S such that A13: A = dom f and A14: f is_measurable_on A by A2; A15: c (#) f is_measurable_on A by A13, A14, MESFUNC1:37; A16: for n being Nat holds ( F . n is_simple_func_in S & dom (F . n) = dom (c (#) f) ) proof let n be Nat; ::_thesis: ( F . n is_simple_func_in S & dom (F . n) = dom (c (#) f) ) reconsider n1 = n as Element of NAT by ORDINAL1:def_12; A17: F . n1 = c (#) (F1 . n1) by A10; hence F . n is_simple_func_in S by A4, Th39; ::_thesis: dom (F . n) = dom (c (#) f) thus dom (F . n) = dom (F1 . n) by A17, MESFUNC1:def_6 .= A by A4, A13 .= dom (c (#) f) by A13, MESFUNC1:def_6 ; ::_thesis: verum end; A18: for n, m being Nat st n <= m holds K1 . n <= K1 . m proof let n, m be Nat; ::_thesis: ( n <= m implies K1 . n <= K1 . m ) A19: K1 . n = integral' (M,(F1 . n)) by A8; A20: K1 . m = integral' (M,(F1 . m)) by A8; A21: F1 . m is_simple_func_in S by A4; A22: F1 . n is nonnegative by A5; A23: dom (F1 . n) = dom f by A4; A24: F1 . m is nonnegative by A5; A25: dom (F1 . m) = dom f by A4; assume A26: n <= m ; ::_thesis: K1 . n <= K1 . m A27: for x being set st x in dom ((F1 . m) - (F1 . n)) holds (F1 . n) . x <= (F1 . m) . x proof let x be set ; ::_thesis: ( x in dom ((F1 . m) - (F1 . n)) implies (F1 . n) . x <= (F1 . m) . x ) assume x in dom ((F1 . m) - (F1 . n)) ; ::_thesis: (F1 . n) . x <= (F1 . m) . x then x in ((dom (F1 . m)) /\ (dom (F1 . n))) \ ((((F1 . m) " {+infty}) /\ ((F1 . n) " {+infty})) \/ (((F1 . m) " {-infty}) /\ ((F1 . n) " {-infty}))) by MESFUNC1:def_4; then x in (dom (F1 . m)) /\ (dom (F1 . n)) by XBOOLE_0:def_5; hence (F1 . n) . x <= (F1 . m) . x by A6, A26, A23, A25; ::_thesis: verum end; A28: F1 . n is_simple_func_in S by A4; then A29: dom ((F1 . m) - (F1 . n)) = (dom (F1 . m)) /\ (dom (F1 . n)) by A21, A22, A24, A27, Th69; then A30: (F1 . m) | (dom ((F1 . m) - (F1 . n))) = F1 . m by A23, A25, GRFUNC_1:23; (F1 . n) | (dom ((F1 . m) - (F1 . n))) = F1 . n by A23, A25, A29, GRFUNC_1:23; hence K1 . n <= K1 . m by A19, A20, A28, A21, A22, A24, A27, A30, Th70; ::_thesis: verum end; deffunc H2( Nat) -> Element of ExtREAL = integral' (M,(F . $1)); consider K being ExtREAL_sequence such that A31: for n being Element of NAT holds K . n = H2(n) from FUNCT_2:sch_4(); A32: now__::_thesis:_for_n_being_Nat_holds_K_._n_=_H2(n) let n be Nat; ::_thesis: K . n = H2(n) n in NAT by ORDINAL1:def_12; hence K . n = H2(n) by A31; ::_thesis: verum end; A33: for n being Nat holds K . n = (R_EAL c) * (K1 . n) proof let n be Nat; ::_thesis: K . n = (R_EAL c) * (K1 . n) reconsider n1 = n as Element of NAT by ORDINAL1:def_12; A34: F1 . n is_simple_func_in S by A4; A35: F . n1 = c (#) (F1 . n1) by A10; thus K . n = integral' (M,(F . n1)) by A32 .= (R_EAL c) * (integral' (M,(F1 . n))) by A1, A5, A34, A35, Th66 .= (R_EAL c) * (K1 . n) by A8 ; ::_thesis: verum end; A36: A = dom (c (#) f) by A13, MESFUNC1:def_6; A37: for x being Element of X st x in dom (c (#) f) holds ( F # x is convergent & lim (F # x) = (c (#) f) . x ) proof let x be Element of X; ::_thesis: ( x in dom (c (#) f) implies ( F # x is convergent & lim (F # x) = (c (#) f) . x ) ) now__::_thesis:_for_n1_being_set_st_n1_in_dom_(F1_#_x)_holds_ -infty_<_(F1_#_x)_._n1 let n1 be set ; ::_thesis: ( n1 in dom (F1 # x) implies -infty < (F1 # x) . n1 ) assume n1 in dom (F1 # x) ; ::_thesis: -infty < (F1 # x) . n1 then reconsider n = n1 as Element of NAT ; A38: (F1 # x) . n = (F1 . n) . x by Def13; F1 . n is nonnegative by A5; hence -infty < (F1 # x) . n1 by A38, Def5; ::_thesis: verum end; then A39: F1 # x is () by Th10; assume A40: x in dom (c (#) f) ; ::_thesis: ( F # x is convergent & lim (F # x) = (c (#) f) . x ) A41: now__::_thesis:_for_n_being_Nat_holds_(F_#_x)_._n_=_(R_EAL_c)_*_((F1_#_x)_._n) let n be Nat; ::_thesis: (F # x) . n = (R_EAL c) * ((F1 # x) . n) reconsider n1 = n as Element of NAT by ORDINAL1:def_12; A42: dom (c (#) (F1 . n1)) = dom (F . n1) by A10 .= dom (c (#) f) by A16 ; thus (F # x) . n = (F . n) . x by Def13 .= (c (#) (F1 . n1)) . x by A10 .= (R_EAL c) * ((F1 . n) . x) by A40, A42, MESFUNC1:def_6 .= (R_EAL c) * ((F1 # x) . n) by Def13 ; ::_thesis: verum end; A43: now__::_thesis:_for_n,_m_being_Nat_st_n_<=_m_holds_ (F1_#_x)_._n_<=_(F1_#_x)_._m let n, m be Nat; ::_thesis: ( n <= m implies (F1 # x) . n <= (F1 # x) . m ) assume A44: n <= m ; ::_thesis: (F1 # x) . n <= (F1 # x) . m A45: (F1 # x) . m = (F1 . m) . x by Def13; (F1 # x) . n = (F1 . n) . x by Def13; hence (F1 # x) . n <= (F1 # x) . m by A6, A13, A36, A40, A44, A45; ::_thesis: verum end; (R_EAL c) * (lim (F1 # x)) = (R_EAL c) * (f . x) by A7, A13, A36, A40 .= (c (#) f) . x by A40, MESFUNC1:def_6 ; hence ( F # x is convergent & lim (F # x) = (c (#) f) . x ) by A1, A41, A39, A43, Th63; ::_thesis: verum end; now__::_thesis:_for_n1_being_set_st_n1_in_dom_K1_holds_ -infty_<_K1_._n1 let n1 be set ; ::_thesis: ( n1 in dom K1 implies -infty < K1 . n1 ) assume n1 in dom K1 ; ::_thesis: -infty < K1 . n1 then reconsider n = n1 as Element of NAT ; A46: F1 . n is_simple_func_in S by A4; K1 . n = integral' (M,(F1 . n)) by A8; hence -infty < K1 . n1 by A5, A46, Th68; ::_thesis: verum end; then A47: K1 is () by Th10; then A48: lim K = (R_EAL c) * (lim K1) by A1, A18, A33, Th63; A49: for n, m being Nat st n <= m holds for x being Element of X st x in dom (c (#) f) holds (F . n) . x <= (F . m) . x proof let n, m be Nat; ::_thesis: ( n <= m implies for x being Element of X st x in dom (c (#) f) holds (F . n) . x <= (F . m) . x ) assume A50: n <= m ; ::_thesis: for x being Element of X st x in dom (c (#) f) holds (F . n) . x <= (F . m) . x reconsider n = n, m = m as Element of NAT by ORDINAL1:def_12; let x be Element of X; ::_thesis: ( x in dom (c (#) f) implies (F . n) . x <= (F . m) . x ) assume A51: x in dom (c (#) f) ; ::_thesis: (F . n) . x <= (F . m) . x dom (c (#) (F1 . m)) = dom (F . m) by A10; then A52: dom (c (#) (F1 . m)) = dom (c (#) f) by A16; (F . m) . x = (c (#) (F1 . m)) . x by A10; then A53: (F . m) . x = (R_EAL c) * ((F1 . m) . x) by A51, A52, MESFUNC1:def_6; dom (c (#) (F1 . n)) = dom (F . n) by A10; then A54: dom (c (#) (F1 . n)) = dom (c (#) f) by A16; (F . n) . x = (c (#) (F1 . n)) . x by A10; then (F . n) . x = (R_EAL c) * ((F1 . n) . x) by A51, A54, MESFUNC1:def_6; hence (F . n) . x <= (F . m) . x by A1, A6, A13, A36, A50, A51, A53, XXREAL_3:71; ::_thesis: verum end; K is convergent by A1, A47, A18, A33, Th63; hence integral+ (M,(c (#) f)) = (R_EAL c) * (integral+ (M,f)) by A9, A36, A15, A11, A32, A16, A12, A49, A37, A48, Def15; ::_thesis: verum end; theorem Th87: :: MESFUNC5:87 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & ( for x being Element of X st x in dom f holds 0 = f . x ) holds integral+ (M,f) = 0 proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & ( for x being Element of X st x in dom f holds 0 = f . x ) holds integral+ (M,f) = 0 let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & ( for x being Element of X st x in dom f holds 0 = f . x ) holds integral+ (M,f) = 0 let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & ( for x being Element of X st x in dom f holds 0 = f . x ) holds integral+ (M,f) = 0 let f be PartFunc of X,ExtREAL; ::_thesis: ( ex A being Element of S st ( A = dom f & f is_measurable_on A ) & ( for x being Element of X st x in dom f holds 0 = f . x ) implies integral+ (M,f) = 0 ) assume that A1: ex A being Element of S st ( A = dom f & f is_measurable_on A ) and A2: for x being Element of X st x in dom f holds 0 = f . x ; ::_thesis: integral+ (M,f) = 0 A3: for x being set st x in dom f holds 0 <= f . x by A2; A4: dom (0 (#) f) = dom f by MESFUNC1:def_6; now__::_thesis:_for_x_being_set_st_x_in_dom_(0_(#)_f)_holds_ (0_(#)_f)_._x_=_f_._x let x be set ; ::_thesis: ( x in dom (0 (#) f) implies (0 (#) f) . x = f . x ) assume A5: x in dom (0 (#) f) ; ::_thesis: (0 (#) f) . x = f . x hence (0 (#) f) . x = (R_EAL 0) * (f . x) by MESFUNC1:def_6 .= 0 .= f . x by A2, A4, A5 ; ::_thesis: verum end; hence integral+ (M,f) = integral+ (M,(0 (#) f)) by A4, FUNCT_1:2 .= (R_EAL 0) * (integral+ (M,f)) by A1, A3, Th86, SUPINF_2:52 .= 0 ; ::_thesis: verum end; begin definition let X be non empty set ; let S be SigmaField of X; let M be sigma_Measure of S; let f be PartFunc of X,ExtREAL; func Integral (M,f) -> Element of ExtREAL equals :: MESFUNC5:def 16 (integral+ (M,(max+ f))) - (integral+ (M,(max- f))); coherence (integral+ (M,(max+ f))) - (integral+ (M,(max- f))) is Element of ExtREAL ; end; :: deftheorem defines Integral MESFUNC5:def_16_:_ for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL holds Integral (M,f) = (integral+ (M,(max+ f))) - (integral+ (M,(max- f))); theorem Th88: :: MESFUNC5:88 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative holds Integral (M,f) = integral+ (M,f) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative holds Integral (M,f) = integral+ (M,f) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative holds Integral (M,f) = integral+ (M,f) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative holds Integral (M,f) = integral+ (M,f) let f be PartFunc of X,ExtREAL; ::_thesis: ( ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative implies Integral (M,f) = integral+ (M,f) ) assume that A1: ex A being Element of S st ( A = dom f & f is_measurable_on A ) and A2: f is nonnegative ; ::_thesis: Integral (M,f) = integral+ (M,f) A3: dom f = dom (max+ f) by MESFUNC2:def_2; A4: now__::_thesis:_for_x_being_set_st_x_in_dom_f_holds_ (max+_f)_._x_=_f_._x let x be set ; ::_thesis: ( x in dom f implies (max+ f) . x = f . x ) A5: 0 <= f . x by A2, SUPINF_2:51; assume x in dom f ; ::_thesis: (max+ f) . x = f . x hence (max+ f) . x = max ((f . x),0) by A3, MESFUNC2:def_2 .= f . x by A5, XXREAL_0:def_10 ; ::_thesis: verum end; A6: dom f = dom (max- f) by MESFUNC2:def_3; A7: now__::_thesis:_for_x_being_Element_of_X_st_x_in_dom_(max-_f)_holds_ (max-_f)_._x_=_0 let x be Element of X; ::_thesis: ( x in dom (max- f) implies (max- f) . x = 0 ) assume x in dom (max- f) ; ::_thesis: (max- f) . x = 0 then (max+ f) . x = f . x by A4, A6; hence (max- f) . x = 0 by MESFUNC2:19; ::_thesis: verum end; A8: dom f = dom (max- f) by MESFUNC2:def_3; f = max+ f by A3, A4, FUNCT_1:2; hence Integral (M,f) = (integral+ (M,f)) - (R_EAL 0) by A1, A7, A8, Th87, MESFUNC2:26 .= integral+ (M,f) by XXREAL_3:15 ; ::_thesis: verum end; theorem :: MESFUNC5:89 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_simple_func_in S & f is nonnegative holds ( Integral (M,f) = integral+ (M,f) & Integral (M,f) = integral' (M,f) ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_simple_func_in S & f is nonnegative holds ( Integral (M,f) = integral+ (M,f) & Integral (M,f) = integral' (M,f) ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_simple_func_in S & f is nonnegative holds ( Integral (M,f) = integral+ (M,f) & Integral (M,f) = integral' (M,f) ) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL st f is_simple_func_in S & f is nonnegative holds ( Integral (M,f) = integral+ (M,f) & Integral (M,f) = integral' (M,f) ) let f be PartFunc of X,ExtREAL; ::_thesis: ( f is_simple_func_in S & f is nonnegative implies ( Integral (M,f) = integral+ (M,f) & Integral (M,f) = integral' (M,f) ) ) assume that A1: f is_simple_func_in S and A2: f is nonnegative ; ::_thesis: ( Integral (M,f) = integral+ (M,f) & Integral (M,f) = integral' (M,f) ) reconsider A = dom f as Element of S by A1, Th37; f is_measurable_on A by A1, MESFUNC2:34; hence Integral (M,f) = integral+ (M,f) by A2, Th88; ::_thesis: Integral (M,f) = integral' (M,f) hence Integral (M,f) = integral' (M,f) by A1, A2, Th77; ::_thesis: verum end; theorem :: MESFUNC5:90 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative holds 0 <= Integral (M,f) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative holds 0 <= Integral (M,f) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative holds 0 <= Integral (M,f) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative holds 0 <= Integral (M,f) let f be PartFunc of X,ExtREAL; ::_thesis: ( ex A being Element of S st ( A = dom f & f is_measurable_on A ) & f is nonnegative implies 0 <= Integral (M,f) ) assume that A1: ex A being Element of S st ( A = dom f & f is_measurable_on A ) and A2: f is nonnegative ; ::_thesis: 0 <= Integral (M,f) 0 <= integral+ (M,f) by A1, A2, Th79; hence 0 <= Integral (M,f) by A1, A2, Th88; ::_thesis: verum end; theorem :: MESFUNC5:91 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A, B being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & A misses B holds Integral (M,(f | (A \/ B))) = (Integral (M,(f | A))) + (Integral (M,(f | B))) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A, B being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & A misses B holds Integral (M,(f | (A \/ B))) = (Integral (M,(f | A))) + (Integral (M,(f | B))) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A, B being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & A misses B holds Integral (M,(f | (A \/ B))) = (Integral (M,(f | A))) + (Integral (M,(f | B))) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for A, B being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & A misses B holds Integral (M,(f | (A \/ B))) = (Integral (M,(f | A))) + (Integral (M,(f | B))) let f be PartFunc of X,ExtREAL; ::_thesis: for A, B being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & A misses B holds Integral (M,(f | (A \/ B))) = (Integral (M,(f | A))) + (Integral (M,(f | B))) let A, B be Element of S; ::_thesis: ( ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & A misses B implies Integral (M,(f | (A \/ B))) = (Integral (M,(f | A))) + (Integral (M,(f | B))) ) assume that A1: ex E being Element of S st ( E = dom f & f is_measurable_on E ) and A2: f is nonnegative and A3: A misses B ; ::_thesis: Integral (M,(f | (A \/ B))) = (Integral (M,(f | A))) + (Integral (M,(f | B))) consider E being Element of S such that A4: E = dom f and A5: f is_measurable_on E by A1; ex C being Element of S st ( C = dom (f | A) & f | A is_measurable_on C ) proof take C = E /\ A; ::_thesis: ( C = dom (f | A) & f | A is_measurable_on C ) thus dom (f | A) = C by A4, RELAT_1:61; ::_thesis: f | A is_measurable_on C A6: C = (dom f) /\ C by A4, XBOOLE_1:17, XBOOLE_1:28; A7: dom (f | A) = C by A4, RELAT_1:61 .= dom (f | C) by A6, RELAT_1:61 ; for x being set st x in dom (f | A) holds (f | A) . x = (f | C) . x proof let x be set ; ::_thesis: ( x in dom (f | A) implies (f | A) . x = (f | C) . x ) assume A8: x in dom (f | A) ; ::_thesis: (f | A) . x = (f | C) . x then (f | A) . x = f . x by FUNCT_1:47; hence (f | A) . x = (f | C) . x by A7, A8, FUNCT_1:47; ::_thesis: verum end; then A9: f | C = f | A by A7, FUNCT_1:2; f is_measurable_on C by A5, MESFUNC1:30, XBOOLE_1:17; hence f | A is_measurable_on C by A6, A9, Th42; ::_thesis: verum end; then A10: Integral (M,(f | A)) = integral+ (M,(f | A)) by A2, Th15, Th88; ex C being Element of S st ( C = dom (f | (A \/ B)) & f | (A \/ B) is_measurable_on C ) proof reconsider C = E /\ (A \/ B) as Element of S ; take C ; ::_thesis: ( C = dom (f | (A \/ B)) & f | (A \/ B) is_measurable_on C ) thus dom (f | (A \/ B)) = C by A4, RELAT_1:61; ::_thesis: f | (A \/ B) is_measurable_on C A11: C = (dom f) /\ C by A4, XBOOLE_1:17, XBOOLE_1:28; A12: dom (f | (A \/ B)) = C by A4, RELAT_1:61 .= dom (f | C) by A11, RELAT_1:61 ; A13: for x being set st x in dom (f | (A \/ B)) holds (f | (A \/ B)) . x = (f | C) . x proof let x be set ; ::_thesis: ( x in dom (f | (A \/ B)) implies (f | (A \/ B)) . x = (f | C) . x ) assume A14: x in dom (f | (A \/ B)) ; ::_thesis: (f | (A \/ B)) . x = (f | C) . x then (f | (A \/ B)) . x = f . x by FUNCT_1:47; hence (f | (A \/ B)) . x = (f | C) . x by A12, A14, FUNCT_1:47; ::_thesis: verum end; f is_measurable_on C by A5, MESFUNC1:30, XBOOLE_1:17; then f | C is_measurable_on C by A11, Th42; hence f | (A \/ B) is_measurable_on C by A12, A13, FUNCT_1:2; ::_thesis: verum end; then A15: Integral (M,(f | (A \/ B))) = integral+ (M,(f | (A \/ B))) by A2, Th15, Th88; A16: ex C being Element of S st ( C = dom (f | B) & f | B is_measurable_on C ) proof take C = E /\ B; ::_thesis: ( C = dom (f | B) & f | B is_measurable_on C ) thus dom (f | B) = C by A4, RELAT_1:61; ::_thesis: f | B is_measurable_on C A17: C = (dom f) /\ C by A4, XBOOLE_1:17, XBOOLE_1:28; A18: dom (f | B) = C by A4, RELAT_1:61 .= dom (f | C) by A17, RELAT_1:61 ; for x being set st x in dom (f | B) holds (f | B) . x = (f | C) . x proof let x be set ; ::_thesis: ( x in dom (f | B) implies (f | B) . x = (f | C) . x ) assume A19: x in dom (f | B) ; ::_thesis: (f | B) . x = (f | C) . x then (f | B) . x = f . x by FUNCT_1:47; hence (f | B) . x = (f | C) . x by A18, A19, FUNCT_1:47; ::_thesis: verum end; then A20: f | C = f | B by A18, FUNCT_1:2; f is_measurable_on C by A5, MESFUNC1:30, XBOOLE_1:17; hence f | B is_measurable_on C by A17, A20, Th42; ::_thesis: verum end; integral+ (M,(f | (A \/ B))) = (integral+ (M,(f | A))) + (integral+ (M,(f | B))) by A1, A2, A3, Th81; hence Integral (M,(f | (A \/ B))) = (Integral (M,(f | A))) + (Integral (M,(f | B))) by A2, A15, A10, A16, Th15, Th88; ::_thesis: verum end; theorem :: MESFUNC5:92 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative holds 0 <= Integral (M,(f | A)) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative holds 0 <= Integral (M,(f | A)) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative holds 0 <= Integral (M,(f | A)) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for A being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative holds 0 <= Integral (M,(f | A)) let f be PartFunc of X,ExtREAL; ::_thesis: for A being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative holds 0 <= Integral (M,(f | A)) let A be Element of S; ::_thesis: ( ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative implies 0 <= Integral (M,(f | A)) ) assume that A1: ex E being Element of S st ( E = dom f & f is_measurable_on E ) and A2: f is nonnegative ; ::_thesis: 0 <= Integral (M,(f | A)) consider E being Element of S such that A3: E = dom f and A4: f is_measurable_on E by A1; A5: ex C being Element of S st ( C = dom (f | A) & f | A is_measurable_on C ) proof take C = E /\ A; ::_thesis: ( C = dom (f | A) & f | A is_measurable_on C ) thus dom (f | A) = C by A3, RELAT_1:61; ::_thesis: f | A is_measurable_on C A6: C = (dom f) /\ C by A3, XBOOLE_1:17, XBOOLE_1:28; A7: dom (f | A) = C by A3, RELAT_1:61 .= dom (f | C) by A6, RELAT_1:61 ; A8: for x being set st x in dom (f | A) holds (f | A) . x = (f | C) . x proof let x be set ; ::_thesis: ( x in dom (f | A) implies (f | A) . x = (f | C) . x ) assume A9: x in dom (f | A) ; ::_thesis: (f | A) . x = (f | C) . x then (f | A) . x = f . x by FUNCT_1:47; hence (f | A) . x = (f | C) . x by A7, A9, FUNCT_1:47; ::_thesis: verum end; f is_measurable_on C by A4, MESFUNC1:30, XBOOLE_1:17; then f | C is_measurable_on C by A6, Th42; hence f | A is_measurable_on C by A7, A8, FUNCT_1:2; ::_thesis: verum end; then 0 <= integral+ (M,(f | A)) by A2, Th15, Th79; hence 0 <= Integral (M,(f | A)) by A2, A5, Th15, Th88; ::_thesis: verum end; theorem :: MESFUNC5:93 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A, B being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & A c= B holds Integral (M,(f | A)) <= Integral (M,(f | B)) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A, B being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & A c= B holds Integral (M,(f | A)) <= Integral (M,(f | B)) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A, B being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & A c= B holds Integral (M,(f | A)) <= Integral (M,(f | B)) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for A, B being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & A c= B holds Integral (M,(f | A)) <= Integral (M,(f | B)) let f be PartFunc of X,ExtREAL; ::_thesis: for A, B being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & A c= B holds Integral (M,(f | A)) <= Integral (M,(f | B)) let A, B be Element of S; ::_thesis: ( ex E being Element of S st ( E = dom f & f is_measurable_on E ) & f is nonnegative & A c= B implies Integral (M,(f | A)) <= Integral (M,(f | B)) ) assume that A1: ex E being Element of S st ( E = dom f & f is_measurable_on E ) and A2: f is nonnegative and A3: A c= B ; ::_thesis: Integral (M,(f | A)) <= Integral (M,(f | B)) consider E being Element of S such that A4: E = dom f and A5: f is_measurable_on E by A1; A6: ex C being Element of S st ( C = dom (f | A) & f | A is_measurable_on C ) proof take C = E /\ A; ::_thesis: ( C = dom (f | A) & f | A is_measurable_on C ) thus dom (f | A) = C by A4, RELAT_1:61; ::_thesis: f | A is_measurable_on C A7: C = (dom f) /\ C by A4, XBOOLE_1:17, XBOOLE_1:28; A8: dom (f | A) = C by A4, RELAT_1:61 .= dom (f | C) by A7, RELAT_1:61 ; A9: for x being set st x in dom (f | A) holds (f | A) . x = (f | C) . x proof let x be set ; ::_thesis: ( x in dom (f | A) implies (f | A) . x = (f | C) . x ) assume A10: x in dom (f | A) ; ::_thesis: (f | A) . x = (f | C) . x then (f | A) . x = f . x by FUNCT_1:47; hence (f | A) . x = (f | C) . x by A8, A10, FUNCT_1:47; ::_thesis: verum end; f is_measurable_on C by A5, MESFUNC1:30, XBOOLE_1:17; then f | C is_measurable_on C by A7, Th42; hence f | A is_measurable_on C by A8, A9, FUNCT_1:2; ::_thesis: verum end; A11: ex C being Element of S st ( C = dom (f | B) & f | B is_measurable_on C ) proof take C = E /\ B; ::_thesis: ( C = dom (f | B) & f | B is_measurable_on C ) thus dom (f | B) = C by A4, RELAT_1:61; ::_thesis: f | B is_measurable_on C A12: C = (dom f) /\ C by A4, XBOOLE_1:17, XBOOLE_1:28; A13: dom (f | B) = C by A4, RELAT_1:61 .= dom (f | C) by A12, RELAT_1:61 ; A14: for x being set st x in dom (f | B) holds (f | B) . x = (f | C) . x proof let x be set ; ::_thesis: ( x in dom (f | B) implies (f | B) . x = (f | C) . x ) assume A15: x in dom (f | B) ; ::_thesis: (f | B) . x = (f | C) . x then (f | B) . x = f . x by FUNCT_1:47; hence (f | B) . x = (f | C) . x by A13, A15, FUNCT_1:47; ::_thesis: verum end; f is_measurable_on C by A5, MESFUNC1:30, XBOOLE_1:17; then f | C is_measurable_on C by A12, Th42; hence f | B is_measurable_on C by A13, A14, FUNCT_1:2; ::_thesis: verum end; integral+ (M,(f | A)) <= integral+ (M,(f | B)) by A1, A2, A3, Th83; then Integral (M,(f | A)) <= integral+ (M,(f | B)) by A2, A6, Th15, Th88; hence Integral (M,(f | A)) <= Integral (M,(f | B)) by A2, A11, Th15, Th88; ::_thesis: verum end; theorem :: MESFUNC5:94 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & M . A = 0 holds Integral (M,(f | A)) = 0 proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & M . A = 0 holds Integral (M,(f | A)) = 0 let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & M . A = 0 holds Integral (M,(f | A)) = 0 let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for A being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & M . A = 0 holds Integral (M,(f | A)) = 0 let f be PartFunc of X,ExtREAL; ::_thesis: for A being Element of S st ex E being Element of S st ( E = dom f & f is_measurable_on E ) & M . A = 0 holds Integral (M,(f | A)) = 0 let A be Element of S; ::_thesis: ( ex E being Element of S st ( E = dom f & f is_measurable_on E ) & M . A = 0 implies Integral (M,(f | A)) = 0 ) assume that A1: ex E being Element of S st ( E = dom f & f is_measurable_on E ) and A2: M . A = 0 ; ::_thesis: Integral (M,(f | A)) = 0 A3: dom f = dom (max+ f) by MESFUNC2:def_2; max+ f is nonnegative by Lm1; then A4: integral+ (M,((max+ f) | A)) = 0 by A1, A2, A3, Th82, MESFUNC2:25; A5: dom f = dom (max- f) by MESFUNC2:def_3; A6: max- f is nonnegative by Lm1; Integral (M,(f | A)) = (integral+ (M,((max+ f) | A))) - (integral+ (M,(max- (f | A)))) by Th28 .= (integral+ (M,((max+ f) | A))) - (integral+ (M,((max- f) | A))) by Th28 .= 0. - 0. by A1, A2, A5, A6, A4, Th82, MESFUNC2:26 ; hence Integral (M,(f | A)) = 0 ; ::_thesis: verum end; theorem Th95: :: MESFUNC5:95 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for E, A being Element of S st E = dom f & f is_measurable_on E & M . A = 0 holds Integral (M,(f | (E \ A))) = Integral (M,f) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for E, A being Element of S st E = dom f & f is_measurable_on E & M . A = 0 holds Integral (M,(f | (E \ A))) = Integral (M,f) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for E, A being Element of S st E = dom f & f is_measurable_on E & M . A = 0 holds Integral (M,(f | (E \ A))) = Integral (M,f) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for E, A being Element of S st E = dom f & f is_measurable_on E & M . A = 0 holds Integral (M,(f | (E \ A))) = Integral (M,f) let f be PartFunc of X,ExtREAL; ::_thesis: for E, A being Element of S st E = dom f & f is_measurable_on E & M . A = 0 holds Integral (M,(f | (E \ A))) = Integral (M,f) let E, A be Element of S; ::_thesis: ( E = dom f & f is_measurable_on E & M . A = 0 implies Integral (M,(f | (E \ A))) = Integral (M,f) ) assume that A1: E = dom f and A2: f is_measurable_on E and A3: M . A = 0 ; ::_thesis: Integral (M,(f | (E \ A))) = Integral (M,f) set B = E \ A; A4: dom f = dom (max+ f) by MESFUNC2:def_2; A5: max- f is nonnegative by Lm1; A6: max+ f is nonnegative by Lm1; A7: dom f = dom (max- f) by MESFUNC2:def_3; Integral (M,(f | (E \ A))) = (integral+ (M,((max+ f) | (E \ A)))) - (integral+ (M,(max- (f | (E \ A))))) by Th28 .= (integral+ (M,((max+ f) | (E \ A)))) - (integral+ (M,((max- f) | (E \ A)))) by Th28 .= (integral+ (M,(max+ f))) - (integral+ (M,((max- f) | (E \ A)))) by A1, A2, A3, A4, A6, Th84, MESFUNC2:25 ; hence Integral (M,(f | (E \ A))) = Integral (M,f) by A1, A2, A3, A7, A5, Th84, MESFUNC2:26; ::_thesis: verum end; definition let X be non empty set ; let S be SigmaField of X; let M be sigma_Measure of S; let f be PartFunc of X,ExtREAL; predf is_integrable_on M means :Def17: :: MESFUNC5:def 17 ( ex A being Element of S st ( A = dom f & f is_measurable_on A ) & integral+ (M,(max+ f)) < +infty & integral+ (M,(max- f)) < +infty ); end; :: deftheorem Def17 defines is_integrable_on MESFUNC5:def_17_:_ for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL holds ( f is_integrable_on M iff ( ex A being Element of S st ( A = dom f & f is_measurable_on A ) & integral+ (M,(max+ f)) < +infty & integral+ (M,(max- f)) < +infty ) ); theorem Th96: :: MESFUNC5:96 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_integrable_on M holds ( 0 <= integral+ (M,(max+ f)) & 0 <= integral+ (M,(max- f)) & -infty < Integral (M,f) & Integral (M,f) < +infty ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_integrable_on M holds ( 0 <= integral+ (M,(max+ f)) & 0 <= integral+ (M,(max- f)) & -infty < Integral (M,f) & Integral (M,f) < +infty ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_integrable_on M holds ( 0 <= integral+ (M,(max+ f)) & 0 <= integral+ (M,(max- f)) & -infty < Integral (M,f) & Integral (M,f) < +infty ) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL st f is_integrable_on M holds ( 0 <= integral+ (M,(max+ f)) & 0 <= integral+ (M,(max- f)) & -infty < Integral (M,f) & Integral (M,f) < +infty ) let f be PartFunc of X,ExtREAL; ::_thesis: ( f is_integrable_on M implies ( 0 <= integral+ (M,(max+ f)) & 0 <= integral+ (M,(max- f)) & -infty < Integral (M,f) & Integral (M,f) < +infty ) ) assume A1: f is_integrable_on M ; ::_thesis: ( 0 <= integral+ (M,(max+ f)) & 0 <= integral+ (M,(max- f)) & -infty < Integral (M,f) & Integral (M,f) < +infty ) consider A being Element of S such that A2: A = dom f and A3: f is_measurable_on A by A1, Def17; A4: integral+ (M,(max+ f)) <> +infty by A1, Def17; A5: dom f = dom (max+ f) by MESFUNC2:def_2; A6: max+ f is nonnegative by Lm1; then -infty <> integral+ (M,(max+ f)) by A2, A3, A5, Th79, MESFUNC2:25; then reconsider maxf1 = integral+ (M,(max+ f)) as Real by A4, XXREAL_0:14; A7: max+ f is_measurable_on A by A3, MESFUNC2:25; A8: integral+ (M,(max- f)) <> +infty by A1, Def17; A9: dom f = dom (max- f) by MESFUNC2:def_3; A10: max- f is nonnegative by Lm1; then -infty <> integral+ (M,(max- f)) by A2, A3, A9, Th79, MESFUNC2:26; then reconsider maxf2 = integral+ (M,(max- f)) as Real by A8, XXREAL_0:14; (integral+ (M,(max+ f))) - (integral+ (M,(max- f))) = maxf1 - maxf2 by SUPINF_2:3; hence ( 0 <= integral+ (M,(max+ f)) & 0 <= integral+ (M,(max- f)) & -infty < Integral (M,f) & Integral (M,f) < +infty ) by A2, A3, A5, A9, A6, A10, A7, Th79, MESFUNC2:26, XXREAL_0:9, XXREAL_0:12; ::_thesis: verum end; theorem Th97: :: MESFUNC5:97 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A being Element of S st f is_integrable_on M holds ( integral+ (M,(max+ (f | A))) <= integral+ (M,(max+ f)) & integral+ (M,(max- (f | A))) <= integral+ (M,(max- f)) & f | A is_integrable_on M ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A being Element of S st f is_integrable_on M holds ( integral+ (M,(max+ (f | A))) <= integral+ (M,(max+ f)) & integral+ (M,(max- (f | A))) <= integral+ (M,(max- f)) & f | A is_integrable_on M ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A being Element of S st f is_integrable_on M holds ( integral+ (M,(max+ (f | A))) <= integral+ (M,(max+ f)) & integral+ (M,(max- (f | A))) <= integral+ (M,(max- f)) & f | A is_integrable_on M ) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for A being Element of S st f is_integrable_on M holds ( integral+ (M,(max+ (f | A))) <= integral+ (M,(max+ f)) & integral+ (M,(max- (f | A))) <= integral+ (M,(max- f)) & f | A is_integrable_on M ) let f be PartFunc of X,ExtREAL; ::_thesis: for A being Element of S st f is_integrable_on M holds ( integral+ (M,(max+ (f | A))) <= integral+ (M,(max+ f)) & integral+ (M,(max- (f | A))) <= integral+ (M,(max- f)) & f | A is_integrable_on M ) let A be Element of S; ::_thesis: ( f is_integrable_on M implies ( integral+ (M,(max+ (f | A))) <= integral+ (M,(max+ f)) & integral+ (M,(max- (f | A))) <= integral+ (M,(max- f)) & f | A is_integrable_on M ) ) A1: max+ f is nonnegative by Lm1; assume A2: f is_integrable_on M ; ::_thesis: ( integral+ (M,(max+ (f | A))) <= integral+ (M,(max+ f)) & integral+ (M,(max- (f | A))) <= integral+ (M,(max- f)) & f | A is_integrable_on M ) then consider E being Element of S such that A3: E = dom f and A4: f is_measurable_on E by Def17; A5: max+ f is_measurable_on E by A4, MESFUNC2:25; A6: f is_measurable_on E /\ A by A4, MESFUNC1:30, XBOOLE_1:17; (dom f) /\ (E /\ A) = E /\ A by A3, XBOOLE_1:17, XBOOLE_1:28; then f | (E /\ A) is_measurable_on E /\ A by A6, Th42; then (f | E) | A is_measurable_on E /\ A by RELAT_1:71; then A7: f | A is_measurable_on E /\ A by A3, GRFUNC_1:23; A8: integral+ (M,(max- f)) < +infty by A2, Def17; A9: max- f is nonnegative by Lm1; A10: integral+ (M,(max+ f)) < +infty by A2, Def17; A11: (max+ f) | (E /\ A) = ((max+ f) | E) | A by RELAT_1:71; A12: dom f = dom (max+ f) by MESFUNC2:def_2; then (max+ f) | E = max+ f by A3, GRFUNC_1:23; then A13: integral+ (M,((max+ f) | A)) <= integral+ (M,(max+ f)) by A3, A5, A12, A1, A11, Th83, XBOOLE_1:17; then integral+ (M,(max+ (f | A))) <= integral+ (M,(max+ f)) by Th28; then A14: integral+ (M,(max+ (f | A))) < +infty by A10, XXREAL_0:2; A15: max- f is_measurable_on E by A3, A4, MESFUNC2:26; A16: (max- f) | (E /\ A) = ((max- f) | E) | A by RELAT_1:71; A17: dom f = dom (max- f) by MESFUNC2:def_3; then (max- f) | E = max- f by A3, GRFUNC_1:23; then A18: integral+ (M,((max- f) | A)) <= integral+ (M,(max- f)) by A3, A15, A17, A9, A16, Th83, XBOOLE_1:17; then integral+ (M,(max- (f | A))) <= integral+ (M,(max- f)) by Th28; then A19: integral+ (M,(max- (f | A))) < +infty by A8, XXREAL_0:2; E /\ A = dom (f | A) by A3, RELAT_1:61; hence ( integral+ (M,(max+ (f | A))) <= integral+ (M,(max+ f)) & integral+ (M,(max- (f | A))) <= integral+ (M,(max- f)) & f | A is_integrable_on M ) by A13, A18, A7, A14, A19, Def17, Th28; ::_thesis: verum end; theorem Th98: :: MESFUNC5:98 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A, B being Element of S st f is_integrable_on M & A misses B holds Integral (M,(f | (A \/ B))) = (Integral (M,(f | A))) + (Integral (M,(f | B))) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A, B being Element of S st f is_integrable_on M & A misses B holds Integral (M,(f | (A \/ B))) = (Integral (M,(f | A))) + (Integral (M,(f | B))) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A, B being Element of S st f is_integrable_on M & A misses B holds Integral (M,(f | (A \/ B))) = (Integral (M,(f | A))) + (Integral (M,(f | B))) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for A, B being Element of S st f is_integrable_on M & A misses B holds Integral (M,(f | (A \/ B))) = (Integral (M,(f | A))) + (Integral (M,(f | B))) let f be PartFunc of X,ExtREAL; ::_thesis: for A, B being Element of S st f is_integrable_on M & A misses B holds Integral (M,(f | (A \/ B))) = (Integral (M,(f | A))) + (Integral (M,(f | B))) let A, B be Element of S; ::_thesis: ( f is_integrable_on M & A misses B implies Integral (M,(f | (A \/ B))) = (Integral (M,(f | A))) + (Integral (M,(f | B))) ) assume that A1: f is_integrable_on M and A2: A misses B ; ::_thesis: Integral (M,(f | (A \/ B))) = (Integral (M,(f | A))) + (Integral (M,(f | B))) consider E being Element of S such that A3: E = dom f and A4: f is_measurable_on E by A1, Def17; set AB = E /\ (A \/ B); A5: max+ (f | A) = (max+ f) | A by Th28; A6: dom f = dom (max- f) by MESFUNC2:def_3; then (max- f) | (A \/ B) = ((max- f) | E) | (A \/ B) by A3, GRFUNC_1:23; then A7: (max- f) | (A \/ B) = (max- f) | (E /\ (A \/ B)) by RELAT_1:71; max- f is nonnegative by Lm1; then A8: integral+ (M,((max- f) | (A \/ B))) = (integral+ (M,((max- f) | A))) + (integral+ (M,((max- f) | B))) by A2, A3, A4, A6, Th81, MESFUNC2:26; A9: f | A is_integrable_on M by A1, Th97; then A10: 0 <= integral+ (M,(max+ (f | A))) by Th96; A11: f | B is_integrable_on M by A1, Th97; then A12: 0 <= integral+ (M,(max+ (f | B))) by Th96; A13: 0 <= integral+ (M,(max- (f | B))) by A11, Th96; integral+ (M,(max- (f | B))) < +infty by A11, Def17; then reconsider g2 = integral+ (M,(max- (f | B))) as Real by A13, XXREAL_0:14; integral+ (M,(max+ (f | B))) < +infty by A11, Def17; then reconsider g1 = integral+ (M,(max+ (f | B))) as Real by A12, XXREAL_0:14; A14: (integral+ (M,(max+ (f | B)))) - (integral+ (M,(max- (f | B)))) = g1 - g2 by SUPINF_2:3; A15: max- (f | A) = (max- f) | A by Th28; A16: dom f = dom (max+ f) by MESFUNC2:def_2; then (max+ f) | (A \/ B) = ((max+ f) | E) | (A \/ B) by A3, GRFUNC_1:23; then A17: (max+ f) | (A \/ B) = (max+ f) | (E /\ (A \/ B)) by RELAT_1:71; max+ f is nonnegative by Lm1; then A18: integral+ (M,((max+ f) | (A \/ B))) = (integral+ (M,((max+ f) | A))) + (integral+ (M,((max+ f) | B))) by A2, A3, A4, A16, Th81, MESFUNC2:25; A19: max- (f | B) = (max- f) | B by Th28; A20: max+ (f | B) = (max+ f) | B by Th28; integral+ (M,(max+ (f | A))) < +infty by A9, Def17; then reconsider f1 = integral+ (M,(max+ (f | A))) as Real by A10, XXREAL_0:14; A21: (integral+ (M,(max+ (f | A)))) + (integral+ (M,(max+ (f | B)))) = f1 + g1 by SUPINF_2:1; A22: 0 <= integral+ (M,(max- (f | A))) by A9, Th96; integral+ (M,(max- (f | A))) < +infty by A9, Def17; then reconsider f2 = integral+ (M,(max- (f | A))) as Real by A22, XXREAL_0:14; A23: (integral+ (M,(max- (f | A)))) + (integral+ (M,(max- (f | B)))) = f2 + g2 by SUPINF_2:1; Integral (M,(f | (A \/ B))) = Integral (M,((f | E) | (A \/ B))) by A3, GRFUNC_1:23 .= Integral (M,(f | (E /\ (A \/ B)))) by RELAT_1:71 .= (integral+ (M,((max+ f) | (E /\ (A \/ B))))) - (integral+ (M,(max- (f | (E /\ (A \/ B)))))) by Th28 .= (integral+ (M,((max+ f) | (E /\ (A \/ B))))) - (integral+ (M,((max- f) | (E /\ (A \/ B))))) by Th28 ; then Integral (M,(f | (A \/ B))) = (f1 + g1) - (f2 + g2) by A18, A8, A17, A7, A5, A15, A20, A19, A21, A23, SUPINF_2:3; then A24: Integral (M,(f | (A \/ B))) = (f1 - f2) + (g1 - g2) ; (integral+ (M,(max+ (f | A)))) - (integral+ (M,(max- (f | A)))) = f1 - f2 by SUPINF_2:3; hence Integral (M,(f | (A \/ B))) = (Integral (M,(f | A))) + (Integral (M,(f | B))) by A24, A14, SUPINF_2:1; ::_thesis: verum end; theorem :: MESFUNC5:99 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A, B being Element of S st f is_integrable_on M & B = (dom f) \ A holds ( f | A is_integrable_on M & Integral (M,f) = (Integral (M,(f | A))) + (Integral (M,(f | B))) ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A, B being Element of S st f is_integrable_on M & B = (dom f) \ A holds ( f | A is_integrable_on M & Integral (M,f) = (Integral (M,(f | A))) + (Integral (M,(f | B))) ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for A, B being Element of S st f is_integrable_on M & B = (dom f) \ A holds ( f | A is_integrable_on M & Integral (M,f) = (Integral (M,(f | A))) + (Integral (M,(f | B))) ) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for A, B being Element of S st f is_integrable_on M & B = (dom f) \ A holds ( f | A is_integrable_on M & Integral (M,f) = (Integral (M,(f | A))) + (Integral (M,(f | B))) ) let f be PartFunc of X,ExtREAL; ::_thesis: for A, B being Element of S st f is_integrable_on M & B = (dom f) \ A holds ( f | A is_integrable_on M & Integral (M,f) = (Integral (M,(f | A))) + (Integral (M,(f | B))) ) let A, B be Element of S; ::_thesis: ( f is_integrable_on M & B = (dom f) \ A implies ( f | A is_integrable_on M & Integral (M,f) = (Integral (M,(f | A))) + (Integral (M,(f | B))) ) ) assume that A1: f is_integrable_on M and A2: B = (dom f) \ A ; ::_thesis: ( f | A is_integrable_on M & Integral (M,f) = (Integral (M,(f | A))) + (Integral (M,(f | B))) ) A \/ B = A \/ (dom f) by A2, XBOOLE_1:39; then A3: (dom f) /\ (A \/ B) = dom f by XBOOLE_1:7, XBOOLE_1:28; A4: f | (A \/ B) = (f | (dom f)) | (A \/ B) by GRFUNC_1:23 .= f | ((dom f) /\ (A \/ B)) by RELAT_1:71 .= f by A3, GRFUNC_1:23 ; A misses B by A2, XBOOLE_1:79; hence ( f | A is_integrable_on M & Integral (M,f) = (Integral (M,(f | A))) + (Integral (M,(f | B))) ) by A1, A4, Th97, Th98; ::_thesis: verum end; theorem Th100: :: MESFUNC5:100 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) holds ( f is_integrable_on M iff |.f.| is_integrable_on M ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) holds ( f is_integrable_on M iff |.f.| is_integrable_on M ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) holds ( f is_integrable_on M iff |.f.| is_integrable_on M ) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) holds ( f is_integrable_on M iff |.f.| is_integrable_on M ) let f be PartFunc of X,ExtREAL; ::_thesis: ( ex A being Element of S st ( A = dom f & f is_measurable_on A ) implies ( f is_integrable_on M iff |.f.| is_integrable_on M ) ) A1: dom |.f.| = dom (max- |.f.|) by MESFUNC2:def_3; A2: dom f = dom (max- f) by MESFUNC2:def_3; A3: now__::_thesis:_for_x_being_set_st_x_in_dom_|.f.|_holds_ 0_<=_|.f.|_._x let x be set ; ::_thesis: ( x in dom |.f.| implies 0 <= |.f.| . x ) assume x in dom |.f.| ; ::_thesis: 0 <= |.f.| . x then |.f.| . x = |.(f . x).| by MESFUNC1:def_10; hence 0 <= |.f.| . x by EXTREAL2:3; ::_thesis: verum end; A4: dom f = dom (max+ f) by MESFUNC2:def_2; A5: |.f.| = (max+ f) + (max- f) by MESFUNC2:24; A6: max+ f is nonnegative by Lm1; assume A7: ex A being Element of S st ( A = dom f & f is_measurable_on A ) ; ::_thesis: ( f is_integrable_on M iff |.f.| is_integrable_on M ) then consider A being Element of S such that A8: A = dom f and A9: f is_measurable_on A ; A10: max- f is_measurable_on A by A8, A9, MESFUNC2:26; A11: |.f.| is_measurable_on A by A8, A9, MESFUNC2:27; A12: A = dom |.f.| by A8, MESFUNC1:def_10; A13: max+ f is_measurable_on A by A9, MESFUNC2:25; A14: dom |.f.| = dom (max+ |.f.|) by MESFUNC2:def_2; hereby ::_thesis: ( |.f.| is_integrable_on M implies f is_integrable_on M ) A15: now__::_thesis:_for_x_being_set_st_x_in_dom_|.f.|_holds_ (max+_|.f.|)_._x_=_|.f.|_._x let x be set ; ::_thesis: ( x in dom |.f.| implies (max+ |.f.|) . x = |.f.| . x ) assume A16: x in dom |.f.| ; ::_thesis: (max+ |.f.|) . x = |.f.| . x then |.f.| . x = |.(f . x).| by MESFUNC1:def_10; then A17: 0 <= |.f.| . x by EXTREAL2:3; (max+ |.f.|) . x = max ((|.f.| . x),0) by A14, A16, MESFUNC2:def_2; hence (max+ |.f.|) . x = |.f.| . x by A17, XXREAL_0:def_10; ::_thesis: verum end; now__::_thesis:_for_x_being_Element_of_X_st_x_in_dom_(max-_|.f.|)_holds_ (max-_|.f.|)_._x_=_0 let x be Element of X; ::_thesis: ( x in dom (max- |.f.|) implies (max- |.f.|) . x = 0 ) assume x in dom (max- |.f.|) ; ::_thesis: (max- |.f.|) . x = 0 then (max+ |.f.|) . x = |.f.| . x by A1, A15; hence (max- |.f.|) . x = 0 by MESFUNC2:19; ::_thesis: verum end; then A18: integral+ (M,(max- |.f.|)) = 0 by A1, A12, A11, Th87, MESFUNC2:26; max- f is nonnegative by Lm1; then A19: integral+ (M,((max+ f) + (max- f))) = (integral+ (M,(max+ f))) + (integral+ (M,(max- f))) by A8, A4, A2, A13, A10, A6, Lm10; assume A20: f is_integrable_on M ; ::_thesis: |.f.| is_integrable_on M then A21: integral+ (M,(max+ f)) < +infty by Def17; A22: integral+ (M,(max- f)) < +infty by A20, Def17; |.f.| = max+ |.f.| by A14, A15, FUNCT_1:2; then integral+ (M,(max+ |.f.|)) < +infty by A5, A21, A22, A19, XXREAL_0:4, XXREAL_3:16; hence |.f.| is_integrable_on M by A12, A11, A18, Def17; ::_thesis: verum end; assume |.f.| is_integrable_on M ; ::_thesis: f is_integrable_on M then Integral (M,|.f.|) < +infty by Th96; then A23: integral+ (M,((max+ f) + (max- f))) < +infty by A12, A11, A5, A3, Th88, SUPINF_2:52; max- f is nonnegative by Lm1; then A24: integral+ (M,((max+ f) + (max- f))) = (integral+ (M,(max+ f))) + (integral+ (M,(max- f))) by A8, A4, A2, A13, A10, A6, Lm10; -infty <> integral+ (M,(max- f)) by A8, A2, A10, Lm1, Th79; then integral+ (M,(max+ f)) <> +infty by A24, A23, XXREAL_3:def_2; then A25: integral+ (M,(max+ f)) < +infty by XXREAL_0:4; -infty <> integral+ (M,(max+ f)) by A8, A4, A13, Lm1, Th79; then integral+ (M,(max- f)) <> +infty by A24, A23, XXREAL_3:def_2; then integral+ (M,(max- f)) < +infty by XXREAL_0:4; hence f is_integrable_on M by A7, A25, Def17; ::_thesis: verum end; theorem :: MESFUNC5:101 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_integrable_on M holds |.(Integral (M,f)).| <= Integral (M,|.f.|) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_integrable_on M holds |.(Integral (M,f)).| <= Integral (M,|.f.|) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_integrable_on M holds |.(Integral (M,f)).| <= Integral (M,|.f.|) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL st f is_integrable_on M holds |.(Integral (M,f)).| <= Integral (M,|.f.|) let f be PartFunc of X,ExtREAL; ::_thesis: ( f is_integrable_on M implies |.(Integral (M,f)).| <= Integral (M,|.f.|) ) assume A1: f is_integrable_on M ; ::_thesis: |.(Integral (M,f)).| <= Integral (M,|.f.|) A2: |.((integral+ (M,(max+ f))) - (integral+ (M,(max- f)))).| <= |.(integral+ (M,(max+ f))).| + |.(integral+ (M,(max- f))).| by EXTREAL2:21; A3: dom f = dom (max+ f) by MESFUNC2:def_2; A4: now__::_thesis:_for_x_being_set_st_x_in_dom_|.f.|_holds_ 0_<=_|.f.|_._x let x be set ; ::_thesis: ( x in dom |.f.| implies 0 <= |.f.| . x ) assume x in dom |.f.| ; ::_thesis: 0 <= |.f.| . x then |.f.| . x = |.(f . x).| by MESFUNC1:def_10; hence 0 <= |.f.| . x by EXTREAL2:3; ::_thesis: verum end; A5: dom f = dom (max- f) by MESFUNC2:def_3; A6: |.f.| = (max+ f) + (max- f) by MESFUNC2:24; consider A being Element of S such that A7: A = dom f and A8: f is_measurable_on A by A1, Def17; A9: max- f is_measurable_on A by A7, A8, MESFUNC2:26; A10: max+ f is nonnegative by Lm1; then 0 <= integral+ (M,(max+ f)) by A7, A8, A3, Th79, MESFUNC2:25; then A11: |.(Integral (M,f)).| <= (integral+ (M,(max+ f))) + |.(integral+ (M,(max- f))).| by A2, EXTREAL1:def_1; A12: max+ f is_measurable_on A by A8, MESFUNC2:25; A13: A = dom |.f.| by A7, MESFUNC1:def_10; A14: max- f is nonnegative by Lm1; then A15: 0 <= integral+ (M,(max- f)) by A7, A8, A5, Th79, MESFUNC2:26; |.f.| is_measurable_on A by A7, A8, MESFUNC2:27; then Integral (M,|.f.|) = integral+ (M,((max+ f) + (max- f))) by A13, A4, A6, Th88, SUPINF_2:52 .= (integral+ (M,(max+ f))) + (integral+ (M,(max- f))) by A7, A3, A5, A10, A14, A12, A9, Lm10 ; hence |.(Integral (M,f)).| <= Integral (M,|.f.|) by A15, A11, EXTREAL1:def_1; ::_thesis: verum end; theorem Th102: :: MESFUNC5:102 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & dom f = dom g & g is_integrable_on M & ( for x being Element of X st x in dom f holds |.(f . x).| <= g . x ) holds ( f is_integrable_on M & Integral (M,|.f.|) <= Integral (M,g) ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & dom f = dom g & g is_integrable_on M & ( for x being Element of X st x in dom f holds |.(f . x).| <= g . x ) holds ( f is_integrable_on M & Integral (M,|.f.|) <= Integral (M,g) ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & dom f = dom g & g is_integrable_on M & ( for x being Element of X st x in dom f holds |.(f . x).| <= g . x ) holds ( f is_integrable_on M & Integral (M,|.f.|) <= Integral (M,g) ) let M be sigma_Measure of S; ::_thesis: for f, g being PartFunc of X,ExtREAL st ex A being Element of S st ( A = dom f & f is_measurable_on A ) & dom f = dom g & g is_integrable_on M & ( for x being Element of X st x in dom f holds |.(f . x).| <= g . x ) holds ( f is_integrable_on M & Integral (M,|.f.|) <= Integral (M,g) ) let f, g be PartFunc of X,ExtREAL; ::_thesis: ( ex A being Element of S st ( A = dom f & f is_measurable_on A ) & dom f = dom g & g is_integrable_on M & ( for x being Element of X st x in dom f holds |.(f . x).| <= g . x ) implies ( f is_integrable_on M & Integral (M,|.f.|) <= Integral (M,g) ) ) assume that A1: ex A being Element of S st ( A = dom f & f is_measurable_on A ) and A2: dom f = dom g and A3: g is_integrable_on M and A4: for x being Element of X st x in dom f holds |.(f . x).| <= g . x ; ::_thesis: ( f is_integrable_on M & Integral (M,|.f.|) <= Integral (M,g) ) A5: ex AA being Element of S st ( AA = dom g & g is_measurable_on AA ) by A3, Def17; A6: now__::_thesis:_for_x_being_set_st_x_in_dom_g_holds_ 0_<=_g_._x let x be set ; ::_thesis: ( x in dom g implies 0 <= g . x ) assume x in dom g ; ::_thesis: 0 <= g . x then |.(f . x).| <= g . x by A2, A4; hence 0 <= g . x by EXTREAL2:3; ::_thesis: verum end; then A7: g is nonnegative by SUPINF_2:52; A8: dom g = dom (max+ g) by MESFUNC2:def_2; now__::_thesis:_for_x_being_set_st_x_in_dom_g_holds_ (max+_g)_._x_=_g_._x let x be set ; ::_thesis: ( x in dom g implies (max+ g) . x = g . x ) A9: 0 <= g . x by A7, SUPINF_2:51; assume x in dom g ; ::_thesis: (max+ g) . x = g . x hence (max+ g) . x = max ((g . x),0) by A8, MESFUNC2:def_2 .= g . x by A9, XXREAL_0:def_10 ; ::_thesis: verum end; then A10: g = max+ g by A8, FUNCT_1:2; A11: dom |.f.| = dom (max+ |.f.|) by MESFUNC2:def_2; A12: now__::_thesis:_for_x_being_set_st_x_in_dom_|.f.|_holds_ (max+_|.f.|)_._x_=_|.f.|_._x let x be set ; ::_thesis: ( x in dom |.f.| implies (max+ |.f.|) . x = |.f.| . x ) assume A13: x in dom |.f.| ; ::_thesis: (max+ |.f.|) . x = |.f.| . x then |.f.| . x = |.(f . x).| by MESFUNC1:def_10; then A14: 0 <= |.f.| . x by EXTREAL2:3; thus (max+ |.f.|) . x = max ((|.f.| . x),0) by A11, A13, MESFUNC2:def_2 .= |.f.| . x by A14, XXREAL_0:def_10 ; ::_thesis: verum end; then A15: |.f.| = max+ |.f.| by A11, FUNCT_1:2; consider A being Element of S such that A16: A = dom f and A17: f is_measurable_on A by A1; A18: |.f.| is_measurable_on A by A16, A17, MESFUNC2:27; A19: A = dom |.f.| by A16, MESFUNC1:def_10; A20: for x being Element of X st x in dom |.f.| holds |.f.| . x <= g . x proof let x be Element of X; ::_thesis: ( x in dom |.f.| implies |.f.| . x <= g . x ) assume A21: x in dom |.f.| ; ::_thesis: |.f.| . x <= g . x then |.f.| . x = |.(f . x).| by MESFUNC1:def_10; hence |.f.| . x <= g . x by A4, A16, A19, A21; ::_thesis: verum end; A22: now__::_thesis:_for_x_being_set_st_x_in_dom_|.f.|_holds_ 0_<=_|.f.|_._x let x be set ; ::_thesis: ( x in dom |.f.| implies 0 <= |.f.| . x ) assume x in dom |.f.| ; ::_thesis: 0 <= |.f.| . x then |.f.| . x = |.(f . x).| by MESFUNC1:def_10; hence 0 <= |.f.| . x by EXTREAL2:3; ::_thesis: verum end; then |.f.| is nonnegative by SUPINF_2:52; then A23: integral+ (M,|.f.|) <= integral+ (M,g) by A2, A16, A5, A19, A18, A7, A20, Th85; A24: dom |.f.| = dom (max- |.f.|) by MESFUNC2:def_3; now__::_thesis:_for_x_being_Element_of_X_st_x_in_dom_(max-_|.f.|)_holds_ (max-_|.f.|)_._x_=_0 let x be Element of X; ::_thesis: ( x in dom (max- |.f.|) implies (max- |.f.|) . x = 0 ) assume x in dom (max- |.f.|) ; ::_thesis: (max- |.f.|) . x = 0 then (max+ |.f.|) . x = |.f.| . x by A24, A12; hence (max- |.f.|) . x = 0 by MESFUNC2:19; ::_thesis: verum end; then A25: integral+ (M,(max- |.f.|)) = 0 by A19, A18, A24, Th87, MESFUNC2:26; integral+ (M,(max+ g)) < +infty by A3, Def17; then integral+ (M,(max+ |.f.|)) < +infty by A15, A10, A23, XXREAL_0:2; then |.f.| is_integrable_on M by A19, A18, A25, Def17; hence f is_integrable_on M by A1, Th100; ::_thesis: Integral (M,|.f.|) <= Integral (M,g) Integral (M,g) = integral+ (M,g) by A5, A6, Th88, SUPINF_2:52; hence Integral (M,|.f.|) <= Integral (M,g) by A19, A18, A22, A23, Th88, SUPINF_2:52; ::_thesis: verum end; theorem Th103: :: MESFUNC5:103 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for r being Real st dom f in S & 0 <= r & dom f <> {} & ( for x being set st x in dom f holds f . x = r ) holds integral (X,S,M,f) = r * (M . (dom f)) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for r being Real st dom f in S & 0 <= r & dom f <> {} & ( for x being set st x in dom f holds f . x = r ) holds integral (X,S,M,f) = r * (M . (dom f)) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for r being Real st dom f in S & 0 <= r & dom f <> {} & ( for x being set st x in dom f holds f . x = r ) holds integral (X,S,M,f) = r * (M . (dom f)) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for r being Real st dom f in S & 0 <= r & dom f <> {} & ( for x being set st x in dom f holds f . x = r ) holds integral (X,S,M,f) = r * (M . (dom f)) let f be PartFunc of X,ExtREAL; ::_thesis: for r being Real st dom f in S & 0 <= r & dom f <> {} & ( for x being set st x in dom f holds f . x = r ) holds integral (X,S,M,f) = r * (M . (dom f)) let r be Real; ::_thesis: ( dom f in S & 0 <= r & dom f <> {} & ( for x being set st x in dom f holds f . x = r ) implies integral (X,S,M,f) = r * (M . (dom f)) ) assume that A1: dom f in S and A2: 0 <= r and A3: dom f <> {} and A4: for x being set st x in dom f holds f . x = r ; ::_thesis: integral (X,S,M,f) = r * (M . (dom f)) A5: for x being set st x in dom f holds 0 <= f . x by A2, A4; f is_simple_func_in S by A1, A4, Lm4; then consider F being Finite_Sep_Sequence of S, a, v being FinSequence of ExtREAL such that A6: F,a are_Re-presentation_of f and A7: dom v = dom F and A8: for n being Nat st n in dom v holds v . n = (a . n) * ((M * F) . n) and A9: integral (X,S,M,f) = Sum v by A3, A5, MESFUNC4:4; A10: dom f = union (rng F) by A6, MESFUNC3:def_1; A11: for n being Nat st n in dom v holds v . n = r * ((M * F) . n) proof let n be Nat; ::_thesis: ( n in dom v implies v . n = r * ((M * F) . n) ) assume A12: n in dom v ; ::_thesis: v . n = r * ((M * F) . n) then A13: F . n c= union (rng F) by A7, FUNCT_1:3, ZFMISC_1:74; A14: v . n = (a . n) * ((M * F) . n) by A8, A12; percases ( F . n = {} or F . n <> {} ) ; suppose F . n = {} ; ::_thesis: v . n = r * ((M * F) . n) then M . (F . n) = 0 by VALUED_0:def_19; then A15: (M * F) . n = 0 by A7, A12, FUNCT_1:13; then v . n = 0 by A14; hence v . n = r * ((M * F) . n) by A15; ::_thesis: verum end; suppose F . n <> {} ; ::_thesis: v . n = r * ((M * F) . n) then consider x being set such that A16: x in F . n by XBOOLE_0:def_1; a . n = f . x by A6, A7, A12, A16, MESFUNC3:def_1; hence v . n = r * ((M * F) . n) by A4, A10, A13, A14, A16; ::_thesis: verum end; end; end; dom v = dom (M * F) by A7, MESFUNC3:8; then integral (X,S,M,f) = (R_EAL r) * (Sum (M * F)) by A9, A11, MESFUNC3:10 .= (R_EAL r) * (M . (union (rng F))) by MESFUNC3:9 ; hence integral (X,S,M,f) = r * (M . (dom f)) by A6, MESFUNC3:def_1; ::_thesis: verum end; theorem Th104: :: MESFUNC5:104 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for r being Real st dom f in S & 0 <= r & ( for x being set st x in dom f holds f . x = r ) holds integral' (M,f) = (R_EAL r) * (M . (dom f)) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for r being Real st dom f in S & 0 <= r & ( for x being set st x in dom f holds f . x = r ) holds integral' (M,f) = (R_EAL r) * (M . (dom f)) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for r being Real st dom f in S & 0 <= r & ( for x being set st x in dom f holds f . x = r ) holds integral' (M,f) = (R_EAL r) * (M . (dom f)) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for r being Real st dom f in S & 0 <= r & ( for x being set st x in dom f holds f . x = r ) holds integral' (M,f) = (R_EAL r) * (M . (dom f)) let f be PartFunc of X,ExtREAL; ::_thesis: for r being Real st dom f in S & 0 <= r & ( for x being set st x in dom f holds f . x = r ) holds integral' (M,f) = (R_EAL r) * (M . (dom f)) let r be Real; ::_thesis: ( dom f in S & 0 <= r & ( for x being set st x in dom f holds f . x = r ) implies integral' (M,f) = (R_EAL r) * (M . (dom f)) ) assume that A1: dom f in S and A2: 0 <= r and A3: for x being set st x in dom f holds f . x = r ; ::_thesis: integral' (M,f) = (R_EAL r) * (M . (dom f)) percases ( dom f = {} or dom f <> {} ) ; supposeA4: dom f = {} ; ::_thesis: integral' (M,f) = (R_EAL r) * (M . (dom f)) then A5: M . (dom f) = 0 by VALUED_0:def_19; integral' (M,f) = 0 by A4, Def14; hence integral' (M,f) = (R_EAL r) * (M . (dom f)) by A5; ::_thesis: verum end; supposeA6: dom f <> {} ; ::_thesis: integral' (M,f) = (R_EAL r) * (M . (dom f)) then integral' (M,f) = integral (X,S,M,f) by Def14; hence integral' (M,f) = (R_EAL r) * (M . (dom f)) by A1, A2, A3, A6, Th103; ::_thesis: verum end; end; end; theorem Th105: :: MESFUNC5:105 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_integrable_on M holds ( f " {+infty} in S & f " {-infty} in S & M . (f " {+infty}) = 0 & M . (f " {-infty}) = 0 & (f " {+infty}) \/ (f " {-infty}) in S & M . ((f " {+infty}) \/ (f " {-infty})) = 0 ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_integrable_on M holds ( f " {+infty} in S & f " {-infty} in S & M . (f " {+infty}) = 0 & M . (f " {-infty}) = 0 & (f " {+infty}) \/ (f " {-infty}) in S & M . ((f " {+infty}) \/ (f " {-infty})) = 0 ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL st f is_integrable_on M holds ( f " {+infty} in S & f " {-infty} in S & M . (f " {+infty}) = 0 & M . (f " {-infty}) = 0 & (f " {+infty}) \/ (f " {-infty}) in S & M . ((f " {+infty}) \/ (f " {-infty})) = 0 ) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL st f is_integrable_on M holds ( f " {+infty} in S & f " {-infty} in S & M . (f " {+infty}) = 0 & M . (f " {-infty}) = 0 & (f " {+infty}) \/ (f " {-infty}) in S & M . ((f " {+infty}) \/ (f " {-infty})) = 0 ) let f be PartFunc of X,ExtREAL; ::_thesis: ( f is_integrable_on M implies ( f " {+infty} in S & f " {-infty} in S & M . (f " {+infty}) = 0 & M . (f " {-infty}) = 0 & (f " {+infty}) \/ (f " {-infty}) in S & M . ((f " {+infty}) \/ (f " {-infty})) = 0 ) ) A1: max+ f is nonnegative by Lm1; assume A2: f is_integrable_on M ; ::_thesis: ( f " {+infty} in S & f " {-infty} in S & M . (f " {+infty}) = 0 & M . (f " {-infty}) = 0 & (f " {+infty}) \/ (f " {-infty}) in S & M . ((f " {+infty}) \/ (f " {-infty})) = 0 ) then A3: integral+ (M,(max+ f)) < +infty by Def17; consider A being Element of S such that A4: A = dom f and A5: f is_measurable_on A by A2, Def17; A6: for x being set holds ( ( x in eq_dom (f,+infty) implies x in A ) & ( x in eq_dom (f,-infty) implies x in A ) ) by A4, MESFUNC1:def_15; then A7: eq_dom (f,+infty) c= A by TARSKI:def_3; then A8: A /\ (eq_dom (f,+infty)) = eq_dom (f,+infty) by XBOOLE_1:28; A9: eq_dom (f,-infty) c= A by A6, TARSKI:def_3; then A10: A /\ (eq_dom (f,-infty)) = eq_dom (f,-infty) by XBOOLE_1:28; A11: A /\ (eq_dom (f,+infty)) in S by A4, A5, MESFUNC1:33; then A12: f " {+infty} in S by A8, Th30; A13: A /\ (eq_dom (f,-infty)) in S by A5, MESFUNC1:34; then reconsider B2 = f " {-infty} as Element of S by A10, Th30; A14: f " {-infty} in S by A13, A10, Th30; thus ( f " {+infty} in S & f " {-infty} in S ) by A11, A13, A8, A10, Th30; ::_thesis: ( M . (f " {+infty}) = 0 & M . (f " {-infty}) = 0 & (f " {+infty}) \/ (f " {-infty}) in S & M . ((f " {+infty}) \/ (f " {-infty})) = 0 ) set C2 = A \ B2; A15: integral+ (M,(max- f)) < +infty by A2, Def17; reconsider B1 = f " {+infty} as Element of S by A11, A8, Th30; A16: A = dom (max+ f) by A4, MESFUNC2:def_2; then A17: B1 c= dom (max+ f) by A7, Th30; then A18: B1 = (dom (max+ f)) /\ B1 by XBOOLE_1:28; A19: max+ f is_measurable_on A by A5, MESFUNC2:25; then max+ f is_measurable_on B1 by A16, A17, MESFUNC1:30; then A20: (max+ f) | B1 is_measurable_on B1 by A18, Th42; set C1 = A \ B1; A21: for x being Element of X holds ( ( x in dom ((max+ f) | (B1 \/ (A \ B1))) implies ((max+ f) | (B1 \/ (A \ B1))) . x = (max+ f) . x ) & ( x in dom ((max- f) | (B2 \/ (A \ B2))) implies ((max- f) | (B2 \/ (A \ B2))) . x = (max- f) . x ) ) by FUNCT_1:47; B1 \/ (A \ B1) = A by A16, A17, XBOOLE_1:45; then dom ((max+ f) | (B1 \/ (A \ B1))) = (dom (max+ f)) /\ (dom (max+ f)) by A16, RELAT_1:61; then (max+ f) | (B1 \/ (A \ B1)) = max+ f by A21, PARTFUN1:5; then integral+ (M,(max+ f)) = (integral+ (M,((max+ f) | B1))) + (integral+ (M,((max+ f) | (A \ B1)))) by A1, A16, A19, Th81, XBOOLE_1:106; then A22: integral+ (M,((max+ f) | B1)) <= integral+ (M,(max+ f)) by A1, A16, A19, Th80, XXREAL_3:65; hereby ::_thesis: ( M . (f " {-infty}) = 0 & (f " {+infty}) \/ (f " {-infty}) in S & M . ((f " {+infty}) \/ (f " {-infty})) = 0 ) A23: for r being Real st 0 < r holds (R_EAL r) * (M . B1) <= integral+ (M,(max+ f)) proof defpred S1[ set ] means $1 in dom ((max+ f) | B1); let r be Real; ::_thesis: ( 0 < r implies (R_EAL r) * (M . B1) <= integral+ (M,(max+ f)) ) deffunc H1( set ) -> Element of ExtREAL = R_EAL r; A24: for x being set st S1[x] holds H1(x) in ExtREAL ; consider g being PartFunc of X,ExtREAL such that A25: ( ( for x being set holds ( x in dom g iff ( x in X & S1[x] ) ) ) & ( for x being set st x in dom g holds g . x = H1(x) ) ) from PARTFUN1:sch_3(A24); assume A26: 0 < r ; ::_thesis: (R_EAL r) * (M . B1) <= integral+ (M,(max+ f)) then for x being set st x in dom g holds 0 <= g . x by A25; then A27: g is nonnegative by SUPINF_2:52; dom ((max+ f) | B1) = (dom (max+ f)) /\ B1 by RELAT_1:61; then A28: dom ((max+ f) | B1) = B1 by A17, XBOOLE_1:28; dom g = X /\ (dom ((max+ f) | B1)) by A25, XBOOLE_0:def_4; then A29: dom g = dom ((max+ f) | B1) by XBOOLE_1:28; then A30: integral' (M,g) = (R_EAL r) * (M . (dom g)) by A26, A25, A28, Th104; A31: for x being Element of X st x in dom g holds g . x <= ((max+ f) | B1) . x proof let x be Element of X; ::_thesis: ( x in dom g implies g . x <= ((max+ f) | B1) . x ) assume A32: x in dom g ; ::_thesis: g . x <= ((max+ f) | B1) . x then x in dom f by A29, A28, FUNCT_1:def_7; then A33: x in dom (max+ f) by MESFUNC2:def_2; f . x in {+infty} by A29, A28, A32, FUNCT_1:def_7; then A34: f . x = +infty by TARSKI:def_1; then max ((f . x),0) = f . x by XXREAL_0:def_10; then (max+ f) . x = +infty by A34, A33, MESFUNC2:def_2; then ((max+ f) | B1) . x = +infty by A29, A28, A32, FUNCT_1:49; hence g . x <= ((max+ f) | B1) . x by XXREAL_0:4; ::_thesis: verum end; dom (chi (B1,X)) = X by FUNCT_3:def_3; then A35: B1 = (dom (chi (B1,X))) /\ B1 by XBOOLE_1:28; then A36: (chi (B1,X)) | B1 is_measurable_on B1 by Th42, MESFUNC2:29; A37: B1 = dom ((chi (B1,X)) | B1) by A35, RELAT_1:61; A38: for x being Element of X st x in dom g holds g . x = (r (#) ((chi (B1,X)) | B1)) . x proof let x be Element of X; ::_thesis: ( x in dom g implies g . x = (r (#) ((chi (B1,X)) | B1)) . x ) assume A39: x in dom g ; ::_thesis: g . x = (r (#) ((chi (B1,X)) | B1)) . x then x in dom ((chi (B1,X)) | B1) by A29, A28, A35, RELAT_1:61; then x in dom (r (#) ((chi (B1,X)) | B1)) by MESFUNC1:def_6; then A40: (r (#) ((chi (B1,X)) | B1)) . x = (R_EAL r) * (((chi (B1,X)) | B1) . x) by MESFUNC1:def_6 .= (R_EAL r) * ((chi (B1,X)) . x) by A29, A28, A37, A39, FUNCT_1:47 ; (chi (B1,X)) . x = 1 by A29, A28, A39, FUNCT_3:def_3; then (r (#) ((chi (B1,X)) | B1)) . x = R_EAL r by A40, XXREAL_3:81; hence g . x = (r (#) ((chi (B1,X)) | B1)) . x by A25, A39; ::_thesis: verum end; dom g = dom (r (#) ((chi (B1,X)) | B1)) by A29, A28, A37, MESFUNC1:def_6; then g = r (#) ((chi (B1,X)) | B1) by A38, PARTFUN1:5; then A41: g is_measurable_on B1 by A37, A36, MESFUNC1:37; (max+ f) | B1 is nonnegative by Lm1, Th15; then integral+ (M,g) <= integral+ (M,((max+ f) | B1)) by A20, A29, A28, A41, A27, A31, Th85; then integral+ (M,g) <= integral+ (M,(max+ f)) by A22, XXREAL_0:2; hence (R_EAL r) * (M . B1) <= integral+ (M,(max+ f)) by A25, A29, A28, A27, A30, Lm4, Th77; ::_thesis: verum end; assume A42: M . (f " {+infty}) <> 0 ; ::_thesis: contradiction then A43: 0 < M . (f " {+infty}) by A12, Th45; percases ( M . B1 = +infty or M . B1 <> +infty ) ; supposeA44: M . B1 = +infty ; ::_thesis: contradiction (R_EAL 1) * (M . B1) <= integral+ (M,(max+ f)) by A23; hence contradiction by A3, A44, XXREAL_3:81; ::_thesis: verum end; suppose M . B1 <> +infty ; ::_thesis: contradiction then reconsider MB = M . B1 as Real by A43, XXREAL_0:14; (R_EAL 1) * (M . B1) <= integral+ (M,(max+ f)) by A23; then A45: 0 < integral+ (M,(max+ f)) by A43; then reconsider I = integral+ (M,(max+ f)) as Real by A3, XXREAL_0:14; A46: (R_EAL ((2 * I) / MB)) * (M . B1) = ((2 * I) / MB) * MB by EXTREAL1:1; (R_EAL ((2 * I) / MB)) * (M . B1) <= integral+ (M,(max+ f)) by A43, A23, A45; then 2 * I <= I by A42, A46, XCMPLX_1:87; hence contradiction by A45, XREAL_1:155; ::_thesis: verum end; end; end; then reconsider B1 = B1 as measure_zero of M by MEASURE1:def_7; A47: max- f is nonnegative by Lm1; A48: A = dom (max- f) by A4, MESFUNC2:def_3; then A49: B2 c= dom (max- f) by A9, Th30; then A50: B2 = (dom (max- f)) /\ B2 by XBOOLE_1:28; A51: max- f is_measurable_on A by A4, A5, MESFUNC2:26; then max- f is_measurable_on B2 by A48, A49, MESFUNC1:30; then A52: (max- f) | B2 is_measurable_on B2 by A50, Th42; B2 \/ (A \ B2) = A by A48, A49, XBOOLE_1:45; then dom ((max- f) | (B2 \/ (A \ B2))) = (dom (max- f)) /\ (dom (max- f)) by A48, RELAT_1:61; then (max- f) | (B2 \/ (A \ B2)) = max- f by A21, PARTFUN1:5; then integral+ (M,(max- f)) = (integral+ (M,((max- f) | B2))) + (integral+ (M,((max- f) | (A \ B2)))) by A47, A48, A51, Th81, XBOOLE_1:106; then A53: integral+ (M,((max- f) | B2)) <= integral+ (M,(max- f)) by A47, A48, A51, Th80, XXREAL_3:65; hereby ::_thesis: ( (f " {+infty}) \/ (f " {-infty}) in S & M . ((f " {+infty}) \/ (f " {-infty})) = 0 ) A55: for r being Real st 0 < r holds (R_EAL r) * (M . B2) <= integral+ (M,(max- f)) proof defpred S1[ set ] means $1 in dom ((max- f) | B2); let r be Real; ::_thesis: ( 0 < r implies (R_EAL r) * (M . B2) <= integral+ (M,(max- f)) ) deffunc H1( set ) -> Element of ExtREAL = R_EAL r; A56: for x being set st S1[x] holds H1(x) in ExtREAL ; consider g being PartFunc of X,ExtREAL such that A57: ( ( for x being set holds ( x in dom g iff ( x in X & S1[x] ) ) ) & ( for x being set st x in dom g holds g . x = H1(x) ) ) from PARTFUN1:sch_3(A56); assume A58: 0 < r ; ::_thesis: (R_EAL r) * (M . B2) <= integral+ (M,(max- f)) then for x being set st x in dom g holds 0 <= g . x by A57; then A59: g is nonnegative by SUPINF_2:52; dom ((max- f) | B2) = (dom (max- f)) /\ B2 by RELAT_1:61; then A60: dom ((max- f) | B2) = B2 by A49, XBOOLE_1:28; dom g = X /\ (dom ((max- f) | B2)) by A57, XBOOLE_0:def_4; then A61: dom g = dom ((max- f) | B2) by XBOOLE_1:28; then A62: integral' (M,g) = (R_EAL r) * (M . (dom g)) by A58, A57, A60, Th104; dom (chi (B2,X)) = X by FUNCT_3:def_3; then A63: B2 = (dom (chi (B2,X))) /\ B2 by XBOOLE_1:28; then A64: B2 = dom ((chi (B2,X)) | B2) by RELAT_1:61; A65: for x being Element of X st x in dom g holds g . x = (r (#) ((chi (B2,X)) | B2)) . x proof let x be Element of X; ::_thesis: ( x in dom g implies g . x = (r (#) ((chi (B2,X)) | B2)) . x ) assume A66: x in dom g ; ::_thesis: g . x = (r (#) ((chi (B2,X)) | B2)) . x then x in dom (r (#) ((chi (B2,X)) | B2)) by A61, A60, A64, MESFUNC1:def_6; then A67: (r (#) ((chi (B2,X)) | B2)) . x = (R_EAL r) * (((chi (B2,X)) | B2) . x) by MESFUNC1:def_6 .= (R_EAL r) * ((chi (B2,X)) . x) by A61, A60, A64, A66, FUNCT_1:47 ; (chi (B2,X)) . x = 1 by A61, A60, A66, FUNCT_3:def_3; then (r (#) ((chi (B2,X)) | B2)) . x = R_EAL r by A67, XXREAL_3:81; hence g . x = (r (#) ((chi (B2,X)) | B2)) . x by A57, A66; ::_thesis: verum end; A68: for x being Element of X st x in dom g holds g . x <= ((max- f) | B2) . x proof let x be Element of X; ::_thesis: ( x in dom g implies g . x <= ((max- f) | B2) . x ) assume A69: x in dom g ; ::_thesis: g . x <= ((max- f) | B2) . x then x in dom f by A61, A60, FUNCT_1:def_7; then A70: x in dom (max- f) by MESFUNC2:def_3; f . x in {-infty} by A61, A60, A69, FUNCT_1:def_7; then A71: - (f . x) = +infty by TARSKI:def_1, XXREAL_3:5; then max ((- (f . x)),0) = - (f . x) by XXREAL_0:def_10; then (max- f) . x = +infty by A71, A70, MESFUNC2:def_3; then ((max- f) | B2) . x = +infty by A61, A60, A69, FUNCT_1:49; hence g . x <= ((max- f) | B2) . x by XXREAL_0:4; ::_thesis: verum end; A72: (chi (B2,X)) | B2 is_measurable_on B2 by A63, Th42, MESFUNC2:29; dom g = dom (r (#) ((chi (B2,X)) | B2)) by A61, A60, A64, MESFUNC1:def_6; then g = r (#) ((chi (B2,X)) | B2) by A65, PARTFUN1:5; then A73: g is_measurable_on B2 by A64, A72, MESFUNC1:37; (max- f) | B2 is nonnegative by Lm1, Th15; then integral+ (M,g) <= integral+ (M,((max- f) | B2)) by A52, A61, A60, A73, A59, A68, Th85; then integral+ (M,g) <= integral+ (M,(max- f)) by A53, XXREAL_0:2; hence (R_EAL r) * (M . B2) <= integral+ (M,(max- f)) by A57, A61, A60, A59, A62, Lm4, Th77; ::_thesis: verum end; assume A74: M . (f " {-infty}) <> 0 ; ::_thesis: contradiction A75: 0 <= M . (f " {-infty}) by A14, Th45; percases ( M . B2 = +infty or M . B2 <> +infty ) ; supposeA76: M . B2 = +infty ; ::_thesis: contradiction (R_EAL 1) * (M . B2) <= integral+ (M,(max- f)) by A55; hence contradiction by A15, A76, XXREAL_3:81; ::_thesis: verum end; suppose M . B2 <> +infty ; ::_thesis: contradiction then reconsider MB = M . B2 as Real by A75, XXREAL_0:14; (R_EAL 1) * (M . B2) <= integral+ (M,(max- f)) by A55; then A77: 0 < integral+ (M,(max- f)) by A74, A75; then reconsider I = integral+ (M,(max- f)) as Real by A15, XXREAL_0:14; A78: (R_EAL ((2 * I) / MB)) * (M . B2) = ((2 * I) / MB) * MB by EXTREAL1:1; (R_EAL ((2 * I) / MB)) * (M . B2) <= integral+ (M,(max- f)) by A74, A75, A55, A77; then 2 * I <= I by A74, A78, XCMPLX_1:87; hence contradiction by A77, XREAL_1:155; ::_thesis: verum end; end; end; thus (f " {+infty}) \/ (f " {-infty}) in S by A12, A14, PROB_1:3; ::_thesis: M . ((f " {+infty}) \/ (f " {-infty})) = 0 thus M . ((f " {+infty}) \/ (f " {-infty})) = M . (B1 \/ B2) .= 0 by A54, MEASURE1:38 ; ::_thesis: verum end; theorem Th106: :: MESFUNC5:106 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M & f is nonnegative & g is nonnegative holds f + g is_integrable_on M proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M & f is nonnegative & g is nonnegative holds f + g is_integrable_on M let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M & f is nonnegative & g is nonnegative holds f + g is_integrable_on M let M be sigma_Measure of S; ::_thesis: for f, g being PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M & f is nonnegative & g is nonnegative holds f + g is_integrable_on M let f, g be PartFunc of X,ExtREAL; ::_thesis: ( f is_integrable_on M & g is_integrable_on M & f is nonnegative & g is nonnegative implies f + g is_integrable_on M ) assume that A1: f is_integrable_on M and A2: g is_integrable_on M and A3: f is nonnegative and A4: g is nonnegative ; ::_thesis: f + g is_integrable_on M A5: integral+ (M,(max+ g)) < +infty by A2, Def17; A6: dom g = dom (max+ g) by MESFUNC2:def_2; now__::_thesis:_for_x_being_set_st_x_in_dom_g_holds_ (max+_g)_._x_=_g_._x let x be set ; ::_thesis: ( x in dom g implies (max+ g) . x = g . x ) assume x in dom g ; ::_thesis: (max+ g) . x = g . x then A7: (max+ g) . x = max ((g . x),0) by A6, MESFUNC2:def_2; 0 <= g . x by A4, SUPINF_2:51; hence (max+ g) . x = g . x by A7, XXREAL_0:def_10; ::_thesis: verum end; then A8: g = max+ g by A6, FUNCT_1:2; consider B being Element of S such that A9: B = dom g and A10: g is_measurable_on B by A2, Def17; consider A being Element of S such that A11: A = dom f and A12: f is_measurable_on A by A1, Def17; A13: g is_measurable_on A /\ B by A10, MESFUNC1:30, XBOOLE_1:17; f is_measurable_on A /\ B by A12, MESFUNC1:30, XBOOLE_1:17; then A14: f + g is_measurable_on A /\ B by A3, A4, A13, Th31; consider C being Element of S such that A15: C = dom (f + g) and A16: integral+ (M,(f + g)) = (integral+ (M,(f | C))) + (integral+ (M,(g | C))) by A3, A4, A11, A12, A9, A10, Th78; A17: A /\ B = dom (f + g) by A3, A4, A11, A9, Th16; then integral+ (M,(g | C)) <= integral+ (M,(g | B)) by A4, A9, A10, A15, Th83, XBOOLE_1:17; then A18: integral+ (M,(g | C)) <= integral+ (M,(max+ g)) by A9, A8, GRFUNC_1:23; integral+ (M,(max+ f)) < +infty by A1, Def17; then A19: (integral+ (M,(max+ f))) + (integral+ (M,(max+ g))) < +infty by A5, XXREAL_0:4, XXREAL_3:16; A20: dom f = dom (max+ f) by MESFUNC2:def_2; now__::_thesis:_for_x_being_set_st_x_in_dom_f_holds_ (max+_f)_._x_=_f_._x let x be set ; ::_thesis: ( x in dom f implies (max+ f) . x = f . x ) assume x in dom f ; ::_thesis: (max+ f) . x = f . x then A21: (max+ f) . x = max ((f . x),0) by A20, MESFUNC2:def_2; 0 <= f . x by A3, SUPINF_2:51; hence (max+ f) . x = f . x by A21, XXREAL_0:def_10; ::_thesis: verum end; then A22: f = max+ f by A20, FUNCT_1:2; A23: dom (f + g) = dom (max+ (f + g)) by MESFUNC2:def_2; A24: now__::_thesis:_for_x_being_set_st_x_in_dom_(f_+_g)_holds_ (max+_(f_+_g))_._x_=_(f_+_g)_._x let x be set ; ::_thesis: ( x in dom (f + g) implies (max+ (f + g)) . x = (f + g) . x ) assume A25: x in dom (f + g) ; ::_thesis: (max+ (f + g)) . x = (f + g) . x then A26: (f + g) . x = (f . x) + (g . x) by MESFUNC1:def_3; A27: 0 <= g . x by A4, SUPINF_2:51; A28: 0 <= f . x by A3, SUPINF_2:51; (max+ (f + g)) . x = max (((f + g) . x),0) by A23, A25, MESFUNC2:def_2; hence (max+ (f + g)) . x = (f + g) . x by A26, A28, A27, XXREAL_0:def_10; ::_thesis: verum end; then A29: f + g = max+ (f + g) by A23, FUNCT_1:2; A30: now__::_thesis:_for_x_being_Element_of_X_st_x_in_dom_(max-_(f_+_g))_holds_ (max-_(f_+_g))_._x_=_0 let x be Element of X; ::_thesis: ( x in dom (max- (f + g)) implies (max- (f + g)) . x = 0 ) assume x in dom (max- (f + g)) ; ::_thesis: (max- (f + g)) . x = 0 then x in dom (f + g) by MESFUNC2:def_3; then (max+ (f + g)) . x = (f + g) . x by A24; hence (max- (f + g)) . x = 0 by MESFUNC2:19; ::_thesis: verum end; integral+ (M,(f | C)) <= integral+ (M,(f | A)) by A3, A11, A12, A17, A15, Th83, XBOOLE_1:17; then integral+ (M,(f | C)) <= integral+ (M,(max+ f)) by A11, A22, GRFUNC_1:23; then integral+ (M,(max+ (f + g))) <= (integral+ (M,(max+ f))) + (integral+ (M,(max+ g))) by A29, A16, A18, XXREAL_3:36; then A31: integral+ (M,(max+ (f + g))) < +infty by A19, XXREAL_0:4; dom (f + g) = dom (max- (f + g)) by MESFUNC2:def_3; then integral+ (M,(max- (f + g))) = 0 by A17, A14, A30, Th87, MESFUNC2:26; hence f + g is_integrable_on M by A17, A14, A31, Def17; ::_thesis: verum end; theorem Th107: :: MESFUNC5:107 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M holds dom (f + g) in S proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M holds dom (f + g) in S let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M holds dom (f + g) in S let M be sigma_Measure of S; ::_thesis: for f, g being PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M holds dom (f + g) in S let f, g be PartFunc of X,ExtREAL; ::_thesis: ( f is_integrable_on M & g is_integrable_on M implies dom (f + g) in S ) assume that A1: f is_integrable_on M and A2: g is_integrable_on M ; ::_thesis: dom (f + g) in S A3: f " {-infty} in S by A1, Th105; A4: ex E2 being Element of S st ( E2 = dom g & g is_measurable_on E2 ) by A2, Def17; A5: ex E1 being Element of S st ( E1 = dom f & f is_measurable_on E1 ) by A1, Def17; A6: g " {-infty} in S by A2, Th105; A7: g " {+infty} in S by A2, Th105; f " {+infty} in S by A1, Th105; hence dom (f + g) in S by A3, A7, A6, A5, A4, Th46; ::_thesis: verum end; Lm11: for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st ex E0 being Element of S st ( dom (f + g) = E0 & f + g is_measurable_on E0 ) & f is_integrable_on M & g is_integrable_on M holds f + g is_integrable_on M proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st ex E0 being Element of S st ( dom (f + g) = E0 & f + g is_measurable_on E0 ) & f is_integrable_on M & g is_integrable_on M holds f + g is_integrable_on M let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st ex E0 being Element of S st ( dom (f + g) = E0 & f + g is_measurable_on E0 ) & f is_integrable_on M & g is_integrable_on M holds f + g is_integrable_on M let M be sigma_Measure of S; ::_thesis: for f, g being PartFunc of X,ExtREAL st ex E0 being Element of S st ( dom (f + g) = E0 & f + g is_measurable_on E0 ) & f is_integrable_on M & g is_integrable_on M holds f + g is_integrable_on M let f, g be PartFunc of X,ExtREAL; ::_thesis: ( ex E0 being Element of S st ( dom (f + g) = E0 & f + g is_measurable_on E0 ) & f is_integrable_on M & g is_integrable_on M implies f + g is_integrable_on M ) assume that A1: ex E0 being Element of S st ( dom (f + g) = E0 & f + g is_measurable_on E0 ) and A2: f is_integrable_on M and A3: g is_integrable_on M ; ::_thesis: f + g is_integrable_on M consider E being Element of S such that A4: dom (f + g) = E and A5: f + g is_measurable_on E by A1; A6: |.f.| | E is nonnegative by Lm1, Th15; ex E2 being Element of S st ( E2 = dom g & g is_measurable_on E2 ) by A3, Def17; then |.g.| is_integrable_on M by A3, Th100; then A7: |.g.| | E is_integrable_on M by Th97; A8: |.g.| | E is nonnegative by Lm1, Th15; A9: dom (f + g) = ((dom f) /\ (dom g)) \ (((f " {-infty}) /\ (g " {+infty})) \/ ((f " {+infty}) /\ (g " {-infty}))) by MESFUNC1:def_3; then dom (f + g) c= dom g by XBOOLE_1:18, XBOOLE_1:36; then A10: E c= dom |.g.| by A4, MESFUNC1:def_10; then A11: (dom |.g.|) /\ E = E by XBOOLE_1:28; dom (f + g) c= dom f by A9, XBOOLE_1:18, XBOOLE_1:36; then A12: E c= dom |.f.| by A4, MESFUNC1:def_10; then (dom |.f.|) /\ E = E by XBOOLE_1:28; then A13: E = dom (|.f.| | E) by RELAT_1:61; then A14: (dom (|.f.| | E)) /\ (dom (|.g.| | E)) = E /\ E by A11, RELAT_1:61; then A15: dom ((|.f.| | E) + (|.g.| | E)) = E by A6, A8, Th22; A16: E = dom (|.g.| | E) by A11, RELAT_1:61; A17: now__::_thesis:_for_x_being_Element_of_X_st_x_in_dom_(f_+_g)_holds_ |.((f_+_g)_._x).|_<=_((|.f.|_|_E)_+_(|.g.|_|_E))_._x let x be Element of X; ::_thesis: ( x in dom (f + g) implies |.((f + g) . x).| <= ((|.f.| | E) + (|.g.| | E)) . x ) A18: |.((f . x) + (g . x)).| <= |.(f . x).| + |.(g . x).| by EXTREAL2:13; assume A19: x in dom (f + g) ; ::_thesis: |.((f + g) . x).| <= ((|.f.| | E) + (|.g.| | E)) . x then A20: x in dom ((|.f.| | E) + (|.g.| | E)) by A4, A6, A8, A14, Th22; |.(f . x).| + |.(g . x).| = (|.f.| . x) + |.(g . x).| by A4, A12, A19, MESFUNC1:def_10 .= (|.f.| . x) + (|.g.| . x) by A4, A10, A19, MESFUNC1:def_10 .= ((|.f.| | E) . x) + (|.g.| . x) by A4, A13, A19, FUNCT_1:47 .= ((|.f.| | E) . x) + ((|.g.| | E) . x) by A4, A16, A19, FUNCT_1:47 .= ((|.f.| | E) + (|.g.| | E)) . x by A20, MESFUNC1:def_3 ; hence |.((f + g) . x).| <= ((|.f.| | E) + (|.g.| | E)) . x by A19, A18, MESFUNC1:def_3; ::_thesis: verum end; ex E1 being Element of S st ( E1 = dom f & f is_measurable_on E1 ) by A2, Def17; then |.f.| is_integrable_on M by A2, Th100; then |.f.| | E is_integrable_on M by Th97; then (|.f.| | E) + (|.g.| | E) is_integrable_on M by A7, A6, A8, Th106; hence f + g is_integrable_on M by A4, A5, A17, A15, Th102; ::_thesis: verum end; theorem Th108: :: MESFUNC5:108 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M holds f + g is_integrable_on M proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M holds f + g is_integrable_on M let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M holds f + g is_integrable_on M let M be sigma_Measure of S; ::_thesis: for f, g being PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M holds f + g is_integrable_on M let f, g be PartFunc of X,ExtREAL; ::_thesis: ( f is_integrable_on M & g is_integrable_on M implies f + g is_integrable_on M ) assume that A1: f is_integrable_on M and A2: g is_integrable_on M ; ::_thesis: f + g is_integrable_on M A3: ex E2 being Element of S st ( E2 = dom g & g is_measurable_on E2 ) by A2, Def17; ex E1 being Element of S st ( E1 = dom f & f is_measurable_on E1 ) by A1, Def17; then ex K0 being Element of S st ( K0 = dom (f + g) & f + g is_measurable_on K0 ) by A3, Th47; hence f + g is_integrable_on M by A1, A2, Lm11; ::_thesis: verum end; Lm12: for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M & dom f = dom g holds ex E, NFG, NFPG being Element of S st ( E c= dom f & NFG c= dom f & E = (dom f) \ NFG & f | E is V60() & E = dom (f | E) & f | E is_measurable_on E & f | E is_integrable_on M & Integral (M,f) = Integral (M,(f | E)) & E c= dom g & NFG c= dom g & E = (dom g) \ NFG & g | E is V60() & E = dom (g | E) & g | E is_measurable_on E & g | E is_integrable_on M & Integral (M,g) = Integral (M,(g | E)) & E c= dom (f + g) & NFPG c= dom (f + g) & E = (dom (f + g)) \ NFPG & M . NFG = 0 & M . NFPG = 0 & E = dom ((f + g) | E) & (f + g) | E is_measurable_on E & (f + g) | E is_integrable_on M & (f + g) | E = (f | E) + (g | E) & Integral (M,((f + g) | E)) = (Integral (M,(f | E))) + (Integral (M,(g | E))) ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M & dom f = dom g holds ex E, NFG, NFPG being Element of S st ( E c= dom f & NFG c= dom f & E = (dom f) \ NFG & f | E is V60() & E = dom (f | E) & f | E is_measurable_on E & f | E is_integrable_on M & Integral (M,f) = Integral (M,(f | E)) & E c= dom g & NFG c= dom g & E = (dom g) \ NFG & g | E is V60() & E = dom (g | E) & g | E is_measurable_on E & g | E is_integrable_on M & Integral (M,g) = Integral (M,(g | E)) & E c= dom (f + g) & NFPG c= dom (f + g) & E = (dom (f + g)) \ NFPG & M . NFG = 0 & M . NFPG = 0 & E = dom ((f + g) | E) & (f + g) | E is_measurable_on E & (f + g) | E is_integrable_on M & (f + g) | E = (f | E) + (g | E) & Integral (M,((f + g) | E)) = (Integral (M,(f | E))) + (Integral (M,(g | E))) ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M & dom f = dom g holds ex E, NFG, NFPG being Element of S st ( E c= dom f & NFG c= dom f & E = (dom f) \ NFG & f | E is V60() & E = dom (f | E) & f | E is_measurable_on E & f | E is_integrable_on M & Integral (M,f) = Integral (M,(f | E)) & E c= dom g & NFG c= dom g & E = (dom g) \ NFG & g | E is V60() & E = dom (g | E) & g | E is_measurable_on E & g | E is_integrable_on M & Integral (M,g) = Integral (M,(g | E)) & E c= dom (f + g) & NFPG c= dom (f + g) & E = (dom (f + g)) \ NFPG & M . NFG = 0 & M . NFPG = 0 & E = dom ((f + g) | E) & (f + g) | E is_measurable_on E & (f + g) | E is_integrable_on M & (f + g) | E = (f | E) + (g | E) & Integral (M,((f + g) | E)) = (Integral (M,(f | E))) + (Integral (M,(g | E))) ) let M be sigma_Measure of S; ::_thesis: for f, g being PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M & dom f = dom g holds ex E, NFG, NFPG being Element of S st ( E c= dom f & NFG c= dom f & E = (dom f) \ NFG & f | E is V60() & E = dom (f | E) & f | E is_measurable_on E & f | E is_integrable_on M & Integral (M,f) = Integral (M,(f | E)) & E c= dom g & NFG c= dom g & E = (dom g) \ NFG & g | E is V60() & E = dom (g | E) & g | E is_measurable_on E & g | E is_integrable_on M & Integral (M,g) = Integral (M,(g | E)) & E c= dom (f + g) & NFPG c= dom (f + g) & E = (dom (f + g)) \ NFPG & M . NFG = 0 & M . NFPG = 0 & E = dom ((f + g) | E) & (f + g) | E is_measurable_on E & (f + g) | E is_integrable_on M & (f + g) | E = (f | E) + (g | E) & Integral (M,((f + g) | E)) = (Integral (M,(f | E))) + (Integral (M,(g | E))) ) let f, g be PartFunc of X,ExtREAL; ::_thesis: ( f is_integrable_on M & g is_integrable_on M & dom f = dom g implies ex E, NFG, NFPG being Element of S st ( E c= dom f & NFG c= dom f & E = (dom f) \ NFG & f | E is V60() & E = dom (f | E) & f | E is_measurable_on E & f | E is_integrable_on M & Integral (M,f) = Integral (M,(f | E)) & E c= dom g & NFG c= dom g & E = (dom g) \ NFG & g | E is V60() & E = dom (g | E) & g | E is_measurable_on E & g | E is_integrable_on M & Integral (M,g) = Integral (M,(g | E)) & E c= dom (f + g) & NFPG c= dom (f + g) & E = (dom (f + g)) \ NFPG & M . NFG = 0 & M . NFPG = 0 & E = dom ((f + g) | E) & (f + g) | E is_measurable_on E & (f + g) | E is_integrable_on M & (f + g) | E = (f | E) + (g | E) & Integral (M,((f + g) | E)) = (Integral (M,(f | E))) + (Integral (M,(g | E))) ) ) assume that A1: f is_integrable_on M and A2: g is_integrable_on M and A3: dom f = dom g ; ::_thesis: ex E, NFG, NFPG being Element of S st ( E c= dom f & NFG c= dom f & E = (dom f) \ NFG & f | E is V60() & E = dom (f | E) & f | E is_measurable_on E & f | E is_integrable_on M & Integral (M,f) = Integral (M,(f | E)) & E c= dom g & NFG c= dom g & E = (dom g) \ NFG & g | E is V60() & E = dom (g | E) & g | E is_measurable_on E & g | E is_integrable_on M & Integral (M,g) = Integral (M,(g | E)) & E c= dom (f + g) & NFPG c= dom (f + g) & E = (dom (f + g)) \ NFPG & M . NFG = 0 & M . NFPG = 0 & E = dom ((f + g) | E) & (f + g) | E is_measurable_on E & (f + g) | E is_integrable_on M & (f + g) | E = (f | E) + (g | E) & Integral (M,((f + g) | E)) = (Integral (M,(f | E))) + (Integral (M,(g | E))) ) A4: (f " {+infty}) /\ (g " {-infty}) c= g " {-infty} by XBOOLE_1:17; reconsider NG = (g " {+infty}) \/ (g " {-infty}) as Element of S by A2, Th105; reconsider NF = (f " {+infty}) \/ (f " {-infty}) as Element of S by A1, Th105; set NFG = NF \/ NG; consider E0 being Element of S such that A5: E0 = dom f and A6: f is_measurable_on E0 by A1, Def17; set E = E0 \ (NF \/ NG); set f1 = f | (E0 \ (NF \/ NG)); set g1 = g | (E0 \ (NF \/ NG)); A7: E0 \ (NF \/ NG) c= dom f by A5, XBOOLE_1:36; reconsider DFPG = dom (f + g) as Element of S by A1, A2, Th107; A8: (f " {-infty}) /\ (g " {+infty}) c= f " {-infty} by XBOOLE_1:17; A9: for x being set holds ( ( x in f " {+infty} implies x in dom f ) & ( x in f " {-infty} implies x in dom f ) & ( x in g " {+infty} implies x in dom g ) & ( x in g " {-infty} implies x in dom g ) ) by FUNCT_1:def_7; then A10: g " {-infty} c= dom g by TARSKI:def_3; set NFPG = DFPG \ (E0 \ (NF \/ NG)); A11: dom (f + g) = ((dom f) /\ (dom g)) \ (((f " {-infty}) /\ (g " {+infty})) \/ ((f " {+infty}) /\ (g " {-infty}))) by MESFUNC1:def_3; then DFPG \ (E0 \ (NF \/ NG)) c= E0 \ (E0 \ (NF \/ NG)) by A3, A5, XBOOLE_1:33, XBOOLE_1:36; then A12: DFPG \ (E0 \ (NF \/ NG)) c= E0 /\ (NF \/ NG) by XBOOLE_1:48; g " {-infty} c= NG by XBOOLE_1:7; then A13: (f " {+infty}) /\ (g " {-infty}) c= NG by A4, XBOOLE_1:1; f " {-infty} c= NF by XBOOLE_1:7; then (f " {-infty}) /\ (g " {+infty}) c= NF by A8, XBOOLE_1:1; then A14: ((f " {-infty}) /\ (g " {+infty})) \/ ((f " {+infty}) /\ (g " {-infty})) c= NF \/ NG by A13, XBOOLE_1:13; then A15: E0 \ (NF \/ NG) c= dom (f + g) by A3, A5, A11, XBOOLE_1:34; then A16: (f + g) | (E0 \ (NF \/ NG)) = (f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))) by Th29; DFPG \ (DFPG \ (E0 \ (NF \/ NG))) = DFPG /\ (E0 \ (NF \/ NG)) by XBOOLE_1:48; then A17: E0 \ (NF \/ NG) = DFPG \ (DFPG \ (E0 \ (NF \/ NG))) by A3, A5, A11, A14, XBOOLE_1:28, XBOOLE_1:34; A18: dom ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))) = E0 \ (NF \/ NG) by A15, Th29; A19: for x being set st x in dom (g | (E0 \ (NF \/ NG))) holds ( -infty < (g | (E0 \ (NF \/ NG))) . x & (g | (E0 \ (NF \/ NG))) . x < +infty ) proof let x be set ; ::_thesis: ( x in dom (g | (E0 \ (NF \/ NG))) implies ( -infty < (g | (E0 \ (NF \/ NG))) . x & (g | (E0 \ (NF \/ NG))) . x < +infty ) ) for x being set st x in dom g holds g . x in ExtREAL ; then reconsider gg = g as Function of (dom g),ExtREAL by FUNCT_2:3; assume A20: x in dom (g | (E0 \ (NF \/ NG))) ; ::_thesis: ( -infty < (g | (E0 \ (NF \/ NG))) . x & (g | (E0 \ (NF \/ NG))) . x < +infty ) then A21: x in (dom g) /\ (E0 \ (NF \/ NG)) by RELAT_1:61; then A22: x in dom g by XBOOLE_0:def_4; x in E0 \ (NF \/ NG) by A21, XBOOLE_0:def_4; then A23: not x in NF \/ NG by XBOOLE_0:def_5; A24: now__::_thesis:_not_(g_|_(E0_\_(NF_\/_NG)))_._x_=_-infty assume (g | (E0 \ (NF \/ NG))) . x = -infty ; ::_thesis: contradiction then g . x = -infty by A20, FUNCT_1:47; then gg . x in {-infty} by TARSKI:def_1; then A25: x in gg " {-infty} by A22, FUNCT_2:38; g " {-infty} c= NG by XBOOLE_1:7; hence contradiction by A23, A25, XBOOLE_0:def_3; ::_thesis: verum end; now__::_thesis:_not_(g_|_(E0_\_(NF_\/_NG)))_._x_=_+infty assume (g | (E0 \ (NF \/ NG))) . x = +infty ; ::_thesis: contradiction then g . x = +infty by A20, FUNCT_1:47; then gg . x in {+infty} by TARSKI:def_1; then A26: x in gg " {+infty} by A22, FUNCT_2:38; g " {+infty} c= NG by XBOOLE_1:7; hence contradiction by A23, A26, XBOOLE_0:def_3; ::_thesis: verum end; hence ( -infty < (g | (E0 \ (NF \/ NG))) . x & (g | (E0 \ (NF \/ NG))) . x < +infty ) by A24, XXREAL_0:4, XXREAL_0:6; ::_thesis: verum end; then for x being set st x in dom (g | (E0 \ (NF \/ NG))) holds -infty < (g | (E0 \ (NF \/ NG))) . x ; then A27: g | (E0 \ (NF \/ NG)) is () by Th10; now__::_thesis:_for_x_being_Element_of_X_st_x_in_dom_(g_|_(E0_\_(NF_\/_NG)))_holds_ |.((g_|_(E0_\_(NF_\/_NG)))_._x).|_<_+infty let x be Element of X; ::_thesis: ( x in dom (g | (E0 \ (NF \/ NG))) implies |.((g | (E0 \ (NF \/ NG))) . x).| < +infty ) A28: - +infty = -infty by XXREAL_3:def_3; assume A29: x in dom (g | (E0 \ (NF \/ NG))) ; ::_thesis: |.((g | (E0 \ (NF \/ NG))) . x).| < +infty then A30: (g | (E0 \ (NF \/ NG))) . x < +infty by A19; -infty < (g | (E0 \ (NF \/ NG))) . x by A19, A29; hence |.((g | (E0 \ (NF \/ NG))) . x).| < +infty by A30, A28, EXTREAL2:11; ::_thesis: verum end; then A31: g | (E0 \ (NF \/ NG)) is V60() by MESFUNC2:def_1; A32: for x being set st x in dom (f | (E0 \ (NF \/ NG))) holds ( (f | (E0 \ (NF \/ NG))) . x < +infty & -infty < (f | (E0 \ (NF \/ NG))) . x ) proof let x be set ; ::_thesis: ( x in dom (f | (E0 \ (NF \/ NG))) implies ( (f | (E0 \ (NF \/ NG))) . x < +infty & -infty < (f | (E0 \ (NF \/ NG))) . x ) ) for x being set st x in dom f holds f . x in ExtREAL ; then reconsider ff = f as Function of (dom f),ExtREAL by FUNCT_2:3; assume A33: x in dom (f | (E0 \ (NF \/ NG))) ; ::_thesis: ( (f | (E0 \ (NF \/ NG))) . x < +infty & -infty < (f | (E0 \ (NF \/ NG))) . x ) then A34: x in (dom f) /\ (E0 \ (NF \/ NG)) by RELAT_1:61; then A35: x in dom f by XBOOLE_0:def_4; x in E0 \ (NF \/ NG) by A34, XBOOLE_0:def_4; then A36: not x in NF \/ NG by XBOOLE_0:def_5; A37: now__::_thesis:_not_(f_|_(E0_\_(NF_\/_NG)))_._x_=_-infty assume (f | (E0 \ (NF \/ NG))) . x = -infty ; ::_thesis: contradiction then f . x = -infty by A33, FUNCT_1:47; then ff . x in {-infty} by TARSKI:def_1; then A38: x in ff " {-infty} by A35, FUNCT_2:38; f " {-infty} c= NF by XBOOLE_1:7; hence contradiction by A36, A38, XBOOLE_0:def_3; ::_thesis: verum end; now__::_thesis:_not_(f_|_(E0_\_(NF_\/_NG)))_._x_=_+infty assume (f | (E0 \ (NF \/ NG))) . x = +infty ; ::_thesis: contradiction then f . x = +infty by A33, FUNCT_1:47; then ff . x in {+infty} by TARSKI:def_1; then A39: x in ff " {+infty} by A35, FUNCT_2:38; f " {+infty} c= NF by XBOOLE_1:7; hence contradiction by A36, A39, XBOOLE_0:def_3; ::_thesis: verum end; hence ( (f | (E0 \ (NF \/ NG))) . x < +infty & -infty < (f | (E0 \ (NF \/ NG))) . x ) by A37, XXREAL_0:4, XXREAL_0:6; ::_thesis: verum end; then for x being set st x in dom (f | (E0 \ (NF \/ NG))) holds -infty < (f | (E0 \ (NF \/ NG))) . x ; then A40: f | (E0 \ (NF \/ NG)) is () by Th10; then A41: dom ((max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))) + (max+ (f | (E0 \ (NF \/ NG))))) = (dom (f | (E0 \ (NF \/ NG)))) /\ (dom (g | (E0 \ (NF \/ NG)))) by A27, Th24; A42: (max+ ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))) + (max- (f | (E0 \ (NF \/ NG)))) is nonnegative by A40, A27, Th24; A43: dom ((max+ ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))) + (max- (f | (E0 \ (NF \/ NG))))) = (dom (f | (E0 \ (NF \/ NG)))) /\ (dom (g | (E0 \ (NF \/ NG)))) by A40, A27, Th24; A44: (max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))) + (max+ (f | (E0 \ (NF \/ NG)))) is nonnegative by A40, A27, Th24; A45: max+ (f | (E0 \ (NF \/ NG))) is nonnegative by Lm1; A46: dom (max+ ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))) = dom ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))) by MESFUNC2:def_2; A47: dom (g | (E0 \ (NF \/ NG))) = (dom g) /\ (E0 \ (NF \/ NG)) by RELAT_1:61; then A48: E0 \ (NF \/ NG) = dom (g | (E0 \ (NF \/ NG))) by A3, A5, XBOOLE_1:28, XBOOLE_1:36; then A49: dom (max- (g | (E0 \ (NF \/ NG)))) = E0 \ (NF \/ NG) by MESFUNC2:def_3; A50: ex Gf being Element of S st ( Gf = dom g & g is_measurable_on Gf ) by A2, Def17; then g is_measurable_on E0 \ (NF \/ NG) by A3, A5, MESFUNC1:30, XBOOLE_1:36; then A51: g | (E0 \ (NF \/ NG)) is_measurable_on E0 \ (NF \/ NG) by A47, A48, Th42; then A52: max- (g | (E0 \ (NF \/ NG))) is_measurable_on E0 \ (NF \/ NG) by A48, MESFUNC2:26; A53: dom (max+ (g | (E0 \ (NF \/ NG)))) = E0 \ (NF \/ NG) by A48, MESFUNC2:def_2; A54: max+ (g | (E0 \ (NF \/ NG))) is nonnegative by Lm1; A55: max- (g | (E0 \ (NF \/ NG))) is nonnegative by Lm1; A56: dom (f | (E0 \ (NF \/ NG))) = (dom f) /\ (E0 \ (NF \/ NG)) by RELAT_1:61; then A57: E0 \ (NF \/ NG) = dom (f | (E0 \ (NF \/ NG))) by A5, XBOOLE_1:28, XBOOLE_1:36; M . NG = 0 by A2, Th105; then A58: NG is measure_zero of M by MEASURE1:def_7; M . NF = 0 by A1, Th105; then NF is measure_zero of M by MEASURE1:def_7; then A59: NF \/ NG is measure_zero of M by A58, MEASURE1:37; then A60: M . (NF \/ NG) = 0 by MEASURE1:def_7; then A61: Integral (M,f) = Integral (M,(f | (E0 \ (NF \/ NG)))) by A5, A6, Th95; E0 /\ (NF \/ NG) c= NF \/ NG by XBOOLE_1:17; then DFPG \ (E0 \ (NF \/ NG)) is measure_zero of M by A59, A12, MEASURE1:36, XBOOLE_1:1; then A62: M . (DFPG \ (E0 \ (NF \/ NG))) = 0 by MEASURE1:def_7; A63: max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))) is nonnegative by Lm1; A64: max+ ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))) is nonnegative by Lm1; for x being set st x in dom (g | (E0 \ (NF \/ NG))) holds (g | (E0 \ (NF \/ NG))) . x < +infty by A19; then A65: g | (E0 \ (NF \/ NG)) is () by Th11; A66: dom (max+ (f | (E0 \ (NF \/ NG)))) = dom (f | (E0 \ (NF \/ NG))) by MESFUNC2:def_2; for x being set st x in dom (g | (E0 \ (NF \/ NG))) holds -infty < (g | (E0 \ (NF \/ NG))) . x by A19; then A67: g | (E0 \ (NF \/ NG)) is () by Th10; A68: dom (max- (f | (E0 \ (NF \/ NG)))) = dom (f | (E0 \ (NF \/ NG))) by MESFUNC2:def_3; A69: f " {-infty} c= dom f by A9, TARSKI:def_3; g " {+infty} c= dom g by A9, TARSKI:def_3; then A70: NG c= dom g by A10, XBOOLE_1:8; f " {+infty} c= dom f by A9, TARSKI:def_3; then NF c= dom g by A3, A69, XBOOLE_1:8; then A71: NF \/ NG c= dom g by A70, XBOOLE_1:8; A72: DFPG \ (E0 \ (NF \/ NG)) c= dom (f + g) by XBOOLE_1:36; A73: g | (E0 \ (NF \/ NG)) is_integrable_on M by A2, Th97; then A74: 0 <= integral+ (M,(max+ (g | (E0 \ (NF \/ NG))))) by Th96; for x being set st x in dom (f | (E0 \ (NF \/ NG))) holds (f | (E0 \ (NF \/ NG))) . x < +infty by A32; then A75: f | (E0 \ (NF \/ NG)) is () by Th11; for x being set st x in dom (f | (E0 \ (NF \/ NG))) holds -infty < (f | (E0 \ (NF \/ NG))) . x by A32; then f | (E0 \ (NF \/ NG)) is () by Th10; then A76: ((max+ ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))) + (max- (f | (E0 \ (NF \/ NG))))) + (max- (g | (E0 \ (NF \/ NG)))) = ((max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))) + (max+ (f | (E0 \ (NF \/ NG))))) + (max+ (g | (E0 \ (NF \/ NG)))) by A75, A67, A65, Th25; A77: max- (f | (E0 \ (NF \/ NG))) is nonnegative by Lm1; A78: dom (max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))) = dom ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))) by MESFUNC2:def_3; A79: integral+ (M,(max+ (g | (E0 \ (NF \/ NG))))) <> +infty by A73, Def17; A80: 0 <= integral+ (M,(max- (g | (E0 \ (NF \/ NG))))) by A73, Th96; f is_measurable_on E0 \ (NF \/ NG) by A6, MESFUNC1:30, XBOOLE_1:36; then A81: f | (E0 \ (NF \/ NG)) is_measurable_on E0 \ (NF \/ NG) by A56, A57, Th42; then A82: max- (f | (E0 \ (NF \/ NG))) is_measurable_on E0 \ (NF \/ NG) by A57, MESFUNC2:26; now__::_thesis:_for_x_being_Element_of_X_st_x_in_dom_(f_|_(E0_\_(NF_\/_NG)))_holds_ |.((f_|_(E0_\_(NF_\/_NG)))_._x).|_<_+infty let x be Element of X; ::_thesis: ( x in dom (f | (E0 \ (NF \/ NG))) implies |.((f | (E0 \ (NF \/ NG))) . x).| < +infty ) A83: - +infty = -infty by XXREAL_3:def_3; assume A84: x in dom (f | (E0 \ (NF \/ NG))) ; ::_thesis: |.((f | (E0 \ (NF \/ NG))) . x).| < +infty then A85: (f | (E0 \ (NF \/ NG))) . x < +infty by A32; -infty < (f | (E0 \ (NF \/ NG))) . x by A32, A84; hence |.((f | (E0 \ (NF \/ NG))) . x).| < +infty by A85, A83, EXTREAL2:11; ::_thesis: verum end; then A86: f | (E0 \ (NF \/ NG)) is V60() by MESFUNC2:def_1; then A87: (f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))) is_measurable_on E0 \ (NF \/ NG) by A81, A51, A31, MESFUNC2:7; then A88: max+ ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))) is_measurable_on E0 \ (NF \/ NG) by MESFUNC2:25; (dom (f | (E0 \ (NF \/ NG)))) /\ (dom (g | (E0 \ (NF \/ NG)))) = E0 \ (NF \/ NG) by A3, A5, A56, A47, XBOOLE_1:28, XBOOLE_1:36; then A89: (max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))) + (max+ (f | (E0 \ (NF \/ NG)))) is_measurable_on E0 \ (NF \/ NG) by A81, A51, A40, A27, Th44; E0 \ (NF \/ NG) = dom ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))) by A15, Th29; then A90: max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))) is_measurable_on E0 \ (NF \/ NG) by A87, MESFUNC2:26; A91: max+ (f | (E0 \ (NF \/ NG))) is_measurable_on E0 \ (NF \/ NG) by A81, MESFUNC2:25; A92: integral+ (M,(max- (g | (E0 \ (NF \/ NG))))) <> +infty by A73, Def17; (max+ ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))) + (max- (f | (E0 \ (NF \/ NG)))) is_measurable_on E0 \ (NF \/ NG) by A57, A81, A51, A40, A27, Th43; then A93: integral+ (M,(((max+ ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))) + (max- (f | (E0 \ (NF \/ NG))))) + (max- (g | (E0 \ (NF \/ NG)))))) = (integral+ (M,((max+ ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))) + (max- (f | (E0 \ (NF \/ NG))))))) + (integral+ (M,(max- (g | (E0 \ (NF \/ NG)))))) by A57, A48, A43, A49, A42, A55, A52, Lm10 .= ((integral+ (M,(max+ ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))))) + (integral+ (M,(max- (f | (E0 \ (NF \/ NG))))))) + (integral+ (M,(max- (g | (E0 \ (NF \/ NG)))))) by A18, A57, A68, A46, A77, A64, A88, A82, Lm10 ; max+ (g | (E0 \ (NF \/ NG))) is_measurable_on E0 \ (NF \/ NG) by A51, MESFUNC2:25; then integral+ (M,(((max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))) + (max+ (f | (E0 \ (NF \/ NG))))) + (max+ (g | (E0 \ (NF \/ NG)))))) = (integral+ (M,((max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))) + (max+ (f | (E0 \ (NF \/ NG))))))) + (integral+ (M,(max+ (g | (E0 \ (NF \/ NG)))))) by A57, A48, A41, A53, A44, A54, A89, Lm10 .= ((integral+ (M,(max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))))) + (integral+ (M,(max+ (f | (E0 \ (NF \/ NG))))))) + (integral+ (M,(max+ (g | (E0 \ (NF \/ NG)))))) by A18, A57, A66, A78, A45, A63, A90, A91, Lm10 ; then ((integral+ (M,(max+ ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))))) + (integral+ (M,(max- (f | (E0 \ (NF \/ NG))))))) + ((integral+ (M,(max- (g | (E0 \ (NF \/ NG)))))) - (integral+ (M,(max- (g | (E0 \ (NF \/ NG))))))) = (((integral+ (M,(max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))))) + (integral+ (M,(max+ (f | (E0 \ (NF \/ NG))))))) + (integral+ (M,(max+ (g | (E0 \ (NF \/ NG))))))) - (integral+ (M,(max- (g | (E0 \ (NF \/ NG)))))) by A76, A80, A92, A93, XXREAL_3:30; then ((integral+ (M,(max+ ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))))) + (integral+ (M,(max- (f | (E0 \ (NF \/ NG))))))) + ((integral+ (M,(max- (g | (E0 \ (NF \/ NG)))))) - (integral+ (M,(max- (g | (E0 \ (NF \/ NG))))))) = ((integral+ (M,(max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))))) + (integral+ (M,(max+ (f | (E0 \ (NF \/ NG))))))) + ((integral+ (M,(max+ (g | (E0 \ (NF \/ NG)))))) - (integral+ (M,(max- (g | (E0 \ (NF \/ NG))))))) by A74, A79, A80, A92, XXREAL_3:30; then ((integral+ (M,(max+ ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))))) + (integral+ (M,(max- (f | (E0 \ (NF \/ NG))))))) + 0. = ((integral+ (M,(max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))))) + (integral+ (M,(max+ (f | (E0 \ (NF \/ NG))))))) + ((integral+ (M,(max+ (g | (E0 \ (NF \/ NG)))))) - (integral+ (M,(max- (g | (E0 \ (NF \/ NG))))))) by XXREAL_3:7; then A94: (integral+ (M,(max+ ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))))) + (integral+ (M,(max- (f | (E0 \ (NF \/ NG)))))) = ((integral+ (M,(max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))))) + (integral+ (M,(max+ (f | (E0 \ (NF \/ NG))))))) + ((integral+ (M,(max+ (g | (E0 \ (NF \/ NG)))))) - (integral+ (M,(max- (g | (E0 \ (NF \/ NG))))))) by XXREAL_3:4; A95: f | (E0 \ (NF \/ NG)) is_integrable_on M by A1, Th97; then A96: 0 <= integral+ (M,(max+ (f | (E0 \ (NF \/ NG))))) by Th96; A97: (f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))) is_integrable_on M by A95, A73, Th108; then A98: integral+ (M,(max+ ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))))) <> +infty by Def17; A99: integral+ (M,(max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))))) <> +infty by A97, Def17; then A100: - (integral+ (M,(max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))))) <> -infty by XXREAL_3:23; A101: 0 <= integral+ (M,(max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))))) by A97, Th96; A102: integral+ (M,(max- (f | (E0 \ (NF \/ NG))))) <> +infty by A95, Def17; then A103: - (integral+ (M,(max- (f | (E0 \ (NF \/ NG)))))) <> -infty by XXREAL_3:23; A104: integral+ (M,(max+ (f | (E0 \ (NF \/ NG))))) <> +infty by A95, Def17; A105: 0 <= integral+ (M,(max- (f | (E0 \ (NF \/ NG))))) by A95, Th96; 0 <= integral+ (M,(max+ ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))))) by A97, Th96; then ((- (integral+ (M,(max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))))))) + (integral+ (M,(max+ ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))))))) + (integral+ (M,(max- (f | (E0 \ (NF \/ NG)))))) = (- (integral+ (M,(max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))))))) + (((integral+ (M,(max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))))) + (integral+ (M,(max+ (f | (E0 \ (NF \/ NG))))))) + ((integral+ (M,(max+ (g | (E0 \ (NF \/ NG)))))) - (integral+ (M,(max- (g | (E0 \ (NF \/ NG)))))))) by A105, A102, A98, A94, XXREAL_3:29; then ((- (integral+ (M,(max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))))))) + (integral+ (M,(max+ ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))))))) + (integral+ (M,(max- (f | (E0 \ (NF \/ NG)))))) = (- (integral+ (M,(max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))))))) + ((integral+ (M,(max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))))) + ((integral+ (M,(max+ (f | (E0 \ (NF \/ NG)))))) + ((integral+ (M,(max+ (g | (E0 \ (NF \/ NG)))))) - (integral+ (M,(max- (g | (E0 \ (NF \/ NG))))))))) by A96, A104, A101, A99, XXREAL_3:29; then ((- (integral+ (M,(max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))))))) + (integral+ (M,(max+ ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))))))) + (integral+ (M,(max- (f | (E0 \ (NF \/ NG)))))) = ((- (integral+ (M,(max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))))))) + (integral+ (M,(max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))))))) + ((integral+ (M,(max+ (f | (E0 \ (NF \/ NG)))))) + ((integral+ (M,(max+ (g | (E0 \ (NF \/ NG)))))) - (integral+ (M,(max- (g | (E0 \ (NF \/ NG)))))))) by A101, A99, A100, XXREAL_3:29; then ((- (integral+ (M,(max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))))))) + (integral+ (M,(max+ ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))))))) + (integral+ (M,(max- (f | (E0 \ (NF \/ NG)))))) = (R_EAL 0) + ((integral+ (M,(max+ (f | (E0 \ (NF \/ NG)))))) + ((integral+ (M,(max+ (g | (E0 \ (NF \/ NG)))))) - (integral+ (M,(max- (g | (E0 \ (NF \/ NG)))))))) by XXREAL_3:7; then ((- (integral+ (M,(max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))))))) + (integral+ (M,(max+ ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))))))) + (integral+ (M,(max- (f | (E0 \ (NF \/ NG)))))) = (integral+ (M,(max+ (f | (E0 \ (NF \/ NG)))))) + ((integral+ (M,(max+ (g | (E0 \ (NF \/ NG)))))) - (integral+ (M,(max- (g | (E0 \ (NF \/ NG))))))) by XXREAL_3:4; then ((- (integral+ (M,(max- (f | (E0 \ (NF \/ NG))))))) + (integral+ (M,(max- (f | (E0 \ (NF \/ NG))))))) + ((- (integral+ (M,(max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))))))) + (integral+ (M,(max+ ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))))))) = (- (integral+ (M,(max- (f | (E0 \ (NF \/ NG))))))) + ((integral+ (M,(max+ (f | (E0 \ (NF \/ NG)))))) + ((integral+ (M,(max+ (g | (E0 \ (NF \/ NG)))))) - (integral+ (M,(max- (g | (E0 \ (NF \/ NG)))))))) by A105, A102, A103, XXREAL_3:29; then ((- (integral+ (M,(max- (f | (E0 \ (NF \/ NG))))))) + (integral+ (M,(max- (f | (E0 \ (NF \/ NG))))))) + ((- (integral+ (M,(max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))))))) + (integral+ (M,(max+ ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))))))) = ((- (integral+ (M,(max- (f | (E0 \ (NF \/ NG))))))) + (integral+ (M,(max+ (f | (E0 \ (NF \/ NG))))))) + ((integral+ (M,(max+ (g | (E0 \ (NF \/ NG)))))) - (integral+ (M,(max- (g | (E0 \ (NF \/ NG))))))) by A96, A104, A105, A103, XXREAL_3:29; then (R_EAL 0) + ((- (integral+ (M,(max- ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))))))) + (integral+ (M,(max+ ((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG)))))))) = ((- (integral+ (M,(max- (f | (E0 \ (NF \/ NG))))))) + (integral+ (M,(max+ (f | (E0 \ (NF \/ NG))))))) + ((integral+ (M,(max+ (g | (E0 \ (NF \/ NG)))))) - (integral+ (M,(max- (g | (E0 \ (NF \/ NG))))))) by XXREAL_3:7; then A106: Integral (M,((f | (E0 \ (NF \/ NG))) + (g | (E0 \ (NF \/ NG))))) = (Integral (M,(f | (E0 \ (NF \/ NG))))) + (Integral (M,(g | (E0 \ (NF \/ NG))))) by XXREAL_3:4; Integral (M,g) = Integral (M,(g | (E0 \ (NF \/ NG)))) by A3, A5, A50, A60, Th95; hence ex E, NFG, NFPG being Element of S st ( E c= dom f & NFG c= dom f & E = (dom f) \ NFG & f | E is V60() & E = dom (f | E) & f | E is_measurable_on E & f | E is_integrable_on M & Integral (M,f) = Integral (M,(f | E)) & E c= dom g & NFG c= dom g & E = (dom g) \ NFG & g | E is V60() & E = dom (g | E) & g | E is_measurable_on E & g | E is_integrable_on M & Integral (M,g) = Integral (M,(g | E)) & E c= dom (f + g) & NFPG c= dom (f + g) & E = (dom (f + g)) \ NFPG & M . NFG = 0 & M . NFPG = 0 & E = dom ((f + g) | E) & (f + g) | E is_measurable_on E & (f + g) | E is_integrable_on M & (f + g) | E = (f | E) + (g | E) & Integral (M,((f + g) | E)) = (Integral (M,(f | E))) + (Integral (M,(g | E))) ) by A3, A5, A60, A71, A15, A16, A18, A62, A17, A72, A7, A57, A48, A81, A51, A86, A31, A87, A95, A73, A106, A61, Th108; ::_thesis: verum end; Lm13: for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M & dom f = dom g holds ( f + g is_integrable_on M & Integral (M,(f + g)) = (Integral (M,f)) + (Integral (M,g)) ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M & dom f = dom g holds ( f + g is_integrable_on M & Integral (M,(f + g)) = (Integral (M,f)) + (Integral (M,g)) ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M & dom f = dom g holds ( f + g is_integrable_on M & Integral (M,(f + g)) = (Integral (M,f)) + (Integral (M,g)) ) let M be sigma_Measure of S; ::_thesis: for f, g being PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M & dom f = dom g holds ( f + g is_integrable_on M & Integral (M,(f + g)) = (Integral (M,f)) + (Integral (M,g)) ) let f, g be PartFunc of X,ExtREAL; ::_thesis: ( f is_integrable_on M & g is_integrable_on M & dom f = dom g implies ( f + g is_integrable_on M & Integral (M,(f + g)) = (Integral (M,f)) + (Integral (M,g)) ) ) assume that A1: f is_integrable_on M and A2: g is_integrable_on M and A3: dom f = dom g ; ::_thesis: ( f + g is_integrable_on M & Integral (M,(f + g)) = (Integral (M,f)) + (Integral (M,g)) ) thus f + g is_integrable_on M by A1, A2, Th108; ::_thesis: Integral (M,(f + g)) = (Integral (M,f)) + (Integral (M,g)) then A4: ex K0 being Element of S st ( K0 = dom (f + g) & f + g is_measurable_on K0 ) by Def17; ex E, NFG, NFPG being Element of S st ( E c= dom f & NFG c= dom f & E = (dom f) \ NFG & f | E is V60() & E = dom (f | E) & f | E is_measurable_on E & f | E is_integrable_on M & Integral (M,f) = Integral (M,(f | E)) & E c= dom g & NFG c= dom g & E = (dom g) \ NFG & g | E is V60() & E = dom (g | E) & g | E is_measurable_on E & g | E is_integrable_on M & Integral (M,g) = Integral (M,(g | E)) & E c= dom (f + g) & NFPG c= dom (f + g) & E = (dom (f + g)) \ NFPG & M . NFG = 0 & M . NFPG = 0 & E = dom ((f + g) | E) & (f + g) | E is_measurable_on E & (f + g) | E is_integrable_on M & (f + g) | E = (f | E) + (g | E) & Integral (M,((f + g) | E)) = (Integral (M,(f | E))) + (Integral (M,(g | E))) ) by A1, A2, A3, Lm12; hence Integral (M,(f + g)) = (Integral (M,f)) + (Integral (M,g)) by A4, Th95; ::_thesis: verum end; theorem :: MESFUNC5:109 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M holds ex E being Element of S st ( E = (dom f) /\ (dom g) & Integral (M,(f + g)) = (Integral (M,(f | E))) + (Integral (M,(g | E))) ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M holds ex E being Element of S st ( E = (dom f) /\ (dom g) & Integral (M,(f + g)) = (Integral (M,(f | E))) + (Integral (M,(g | E))) ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M holds ex E being Element of S st ( E = (dom f) /\ (dom g) & Integral (M,(f + g)) = (Integral (M,(f | E))) + (Integral (M,(g | E))) ) let M be sigma_Measure of S; ::_thesis: for f, g being PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M holds ex E being Element of S st ( E = (dom f) /\ (dom g) & Integral (M,(f + g)) = (Integral (M,(f | E))) + (Integral (M,(g | E))) ) let f, g be PartFunc of X,ExtREAL; ::_thesis: ( f is_integrable_on M & g is_integrable_on M implies ex E being Element of S st ( E = (dom f) /\ (dom g) & Integral (M,(f + g)) = (Integral (M,(f | E))) + (Integral (M,(g | E))) ) ) assume that A1: f is_integrable_on M and A2: g is_integrable_on M ; ::_thesis: ex E being Element of S st ( E = (dom f) /\ (dom g) & Integral (M,(f + g)) = (Integral (M,(f | E))) + (Integral (M,(g | E))) ) consider B being Element of S such that A3: B = dom g and g is_measurable_on B by A2, Def17; consider A being Element of S such that A4: A = dom f and f is_measurable_on A by A1, Def17; set E = A /\ B; set g1 = g | (A /\ B); set f1 = f | (A /\ B); take E = A /\ B; ::_thesis: ( E = (dom f) /\ (dom g) & Integral (M,(f + g)) = (Integral (M,(f | E))) + (Integral (M,(g | E))) ) A5: dom (f | (A /\ B)) = (dom f) /\ (A /\ B) by RELAT_1:61 .= (A /\ A) /\ B by A4, XBOOLE_1:16 ; A6: (f | (A /\ B)) " {+infty} = E /\ (f " {+infty}) by FUNCT_1:70; (g | (A /\ B)) " {-infty} = E /\ (g " {-infty}) by FUNCT_1:70; then A7: ((f | (A /\ B)) " {+infty}) /\ ((g | (A /\ B)) " {-infty}) = (((f " {+infty}) /\ E) /\ E) /\ (g " {-infty}) by A6, XBOOLE_1:16 .= ((f " {+infty}) /\ (E /\ E)) /\ (g " {-infty}) by XBOOLE_1:16 .= E /\ ((f " {+infty}) /\ (g " {-infty})) by XBOOLE_1:16 ; A8: (g | (A /\ B)) " {+infty} = E /\ (g " {+infty}) by FUNCT_1:70; (f | (A /\ B)) " {-infty} = E /\ (f " {-infty}) by FUNCT_1:70; then ((f | (A /\ B)) " {-infty}) /\ ((g | (A /\ B)) " {+infty}) = (((f " {-infty}) /\ E) /\ E) /\ (g " {+infty}) by A8, XBOOLE_1:16 .= ((f " {-infty}) /\ (E /\ E)) /\ (g " {+infty}) by XBOOLE_1:16 .= E /\ ((f " {-infty}) /\ (g " {+infty})) by XBOOLE_1:16 ; then A9: (((f | (A /\ B)) " {-infty}) /\ ((g | (A /\ B)) " {+infty})) \/ (((f | (A /\ B)) " {+infty}) /\ ((g | (A /\ B)) " {-infty})) = E /\ (((f " {-infty}) /\ (g " {+infty})) \/ ((f " {+infty}) /\ (g " {-infty}))) by A7, XBOOLE_1:23; A10: dom (g | (A /\ B)) = (dom g) /\ (A /\ B) by RELAT_1:61 .= (B /\ B) /\ A by A3, XBOOLE_1:16 ; A11: dom ((f | (A /\ B)) + (g | (A /\ B))) = ((dom (f | (A /\ B))) /\ (dom (g | (A /\ B)))) \ ((((f | (A /\ B)) " {-infty}) /\ ((g | (A /\ B)) " {+infty})) \/ (((f | (A /\ B)) " {+infty}) /\ ((g | (A /\ B)) " {-infty}))) by MESFUNC1:def_3 .= E \ (((f " {-infty}) /\ (g " {+infty})) \/ ((f " {+infty}) /\ (g " {-infty}))) by A5, A10, A9, XBOOLE_1:47 .= dom (f + g) by A4, A3, MESFUNC1:def_3 ; A12: for x being set st x in dom ((f | (A /\ B)) + (g | (A /\ B))) holds ((f | (A /\ B)) + (g | (A /\ B))) . x = (f + g) . x proof let x be set ; ::_thesis: ( x in dom ((f | (A /\ B)) + (g | (A /\ B))) implies ((f | (A /\ B)) + (g | (A /\ B))) . x = (f + g) . x ) assume A13: x in dom ((f | (A /\ B)) + (g | (A /\ B))) ; ::_thesis: ((f | (A /\ B)) + (g | (A /\ B))) . x = (f + g) . x then x in ((dom (f | (A /\ B))) /\ (dom (g | (A /\ B)))) \ ((((f | (A /\ B)) " {-infty}) /\ ((g | (A /\ B)) " {+infty})) \/ (((f | (A /\ B)) " {+infty}) /\ ((g | (A /\ B)) " {-infty}))) by MESFUNC1:def_3; then A14: x in (dom (f | (A /\ B))) /\ (dom (g | (A /\ B))) by XBOOLE_0:def_5; then A15: x in dom (f | (A /\ B)) by XBOOLE_0:def_4; A16: x in dom (g | (A /\ B)) by A14, XBOOLE_0:def_4; ((f | (A /\ B)) + (g | (A /\ B))) . x = ((f | (A /\ B)) . x) + ((g | (A /\ B)) . x) by A13, MESFUNC1:def_3 .= (f . x) + ((g | (A /\ B)) . x) by A15, FUNCT_1:47 .= (f . x) + (g . x) by A16, FUNCT_1:47 ; hence ((f | (A /\ B)) + (g | (A /\ B))) . x = (f + g) . x by A11, A13, MESFUNC1:def_3; ::_thesis: verum end; thus E = (dom f) /\ (dom g) by A4, A3; ::_thesis: Integral (M,(f + g)) = (Integral (M,(f | E))) + (Integral (M,(g | E))) A17: g | (A /\ B) is_integrable_on M by A2, Th97; f | (A /\ B) is_integrable_on M by A1, Th97; then Integral (M,((f | (A /\ B)) + (g | (A /\ B)))) = (Integral (M,(f | (A /\ B)))) + (Integral (M,(g | (A /\ B)))) by A17, A5, A10, Lm13; hence Integral (M,(f + g)) = (Integral (M,(f | E))) + (Integral (M,(g | E))) by A11, A12, FUNCT_1:2; ::_thesis: verum end; theorem Th110: :: MESFUNC5:110 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for c being Real st f is_integrable_on M holds ( c (#) f is_integrable_on M & Integral (M,(c (#) f)) = (R_EAL c) * (Integral (M,f)) ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for c being Real st f is_integrable_on M holds ( c (#) f is_integrable_on M & Integral (M,(c (#) f)) = (R_EAL c) * (Integral (M,f)) ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for c being Real st f is_integrable_on M holds ( c (#) f is_integrable_on M & Integral (M,(c (#) f)) = (R_EAL c) * (Integral (M,f)) ) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for c being Real st f is_integrable_on M holds ( c (#) f is_integrable_on M & Integral (M,(c (#) f)) = (R_EAL c) * (Integral (M,f)) ) let f be PartFunc of X,ExtREAL; ::_thesis: for c being Real st f is_integrable_on M holds ( c (#) f is_integrable_on M & Integral (M,(c (#) f)) = (R_EAL c) * (Integral (M,f)) ) let c be Real; ::_thesis: ( f is_integrable_on M implies ( c (#) f is_integrable_on M & Integral (M,(c (#) f)) = (R_EAL c) * (Integral (M,f)) ) ) assume A1: f is_integrable_on M ; ::_thesis: ( c (#) f is_integrable_on M & Integral (M,(c (#) f)) = (R_EAL c) * (Integral (M,f)) ) A2: integral+ (M,(max+ f)) <> +infty by A1, Def17; consider A being Element of S such that A3: A = dom f and A4: f is_measurable_on A by A1, Def17; A5: c (#) f is_measurable_on A by A3, A4, Th49; A6: dom (max- f) = A by A3, MESFUNC2:def_3; A7: integral+ (M,(max- f)) <> +infty by A1, Def17; 0 <= integral+ (M,(max- f)) by A1, Th96; then reconsider I2 = integral+ (M,(max- f)) as Real by A7, XXREAL_0:14; A8: max- f is nonnegative by Lm1; 0 <= integral+ (M,(max+ f)) by A1, Th96; then reconsider I1 = integral+ (M,(max+ f)) as Real by A2, XXREAL_0:14; A9: max+ f is nonnegative by Lm1; A10: dom (c (#) f) = A by A3, MESFUNC1:def_6; A11: dom (max+ f) = A by A3, MESFUNC2:def_2; percases ( 0 <= c or c < 0 ) ; supposeA12: 0 <= c ; ::_thesis: ( c (#) f is_integrable_on M & Integral (M,(c (#) f)) = (R_EAL c) * (Integral (M,f)) ) c * I1 is Real ; then A13: (R_EAL c) * (integral+ (M,(max+ f))) is Real by EXTREAL1:1; A14: max+ (c (#) f) = c (#) (max+ f) by A12, Th26; integral+ (M,(c (#) (max+ f))) = (R_EAL c) * (integral+ (M,(max+ f))) by A4, A9, A11, A12, Th86, MESFUNC2:25; then A15: integral+ (M,(max+ (c (#) f))) < +infty by A14, A13, XXREAL_0:9; c * I2 is Real ; then (R_EAL c) * (integral+ (M,(max- f))) is Real by EXTREAL1:1; then A16: (R_EAL c) * (integral+ (M,(max- f))) < +infty by XXREAL_0:9; A17: max- (c (#) f) = c (#) (max- f) by A12, Th26; integral+ (M,(c (#) (max- f))) = (R_EAL c) * (integral+ (M,(max- f))) by A3, A4, A8, A6, A12, Th86, MESFUNC2:26; hence c (#) f is_integrable_on M by A5, A10, A17, A15, A16, Def17; ::_thesis: Integral (M,(c (#) f)) = (R_EAL c) * (Integral (M,f)) thus Integral (M,(c (#) f)) = (integral+ (M,(c (#) (max+ f)))) - (integral+ (M,(max- (c (#) f)))) by A12, Th26 .= (integral+ (M,(c (#) (max+ f)))) - (integral+ (M,(c (#) (max- f)))) by A12, Th26 .= ((R_EAL c) * (integral+ (M,(max+ f)))) - (integral+ (M,(c (#) (max- f)))) by A4, A9, A11, A12, Th86, MESFUNC2:25 .= ((R_EAL c) * (integral+ (M,(max+ f)))) - ((R_EAL c) * (integral+ (M,(max- f)))) by A3, A4, A8, A6, A12, Th86, MESFUNC2:26 .= (R_EAL c) * (Integral (M,f)) by XXREAL_3:100 ; ::_thesis: verum end; supposeA18: c < 0 ; ::_thesis: ( c (#) f is_integrable_on M & Integral (M,(c (#) f)) = (R_EAL c) * (Integral (M,f)) ) - (- c) = c ; then consider a being Real such that A19: c = - a and A20: a > 0 by A18; A21: max+ (c (#) f) = a (#) (max- f) by A19, A20, Th27; A22: max- (c (#) f) = a (#) (max+ f) by A19, A20, Th27; a * I2 is Real ; then A23: (R_EAL a) * (integral+ (M,(max- f))) is Real by EXTREAL1:1; integral+ (M,(a (#) (max- f))) = (R_EAL a) * (integral+ (M,(max- f))) by A3, A4, A8, A6, A20, Th86, MESFUNC2:26; then A24: integral+ (M,(max+ (c (#) f))) < +infty by A21, A23, XXREAL_0:9; a * I1 is Real ; then (R_EAL a) * (integral+ (M,(max+ f))) is Real by EXTREAL1:1; then A25: (R_EAL a) * (integral+ (M,(max+ f))) < +infty by XXREAL_0:9; integral+ (M,(a (#) (max+ f))) = (R_EAL a) * (integral+ (M,(max+ f))) by A4, A9, A11, A20, Th86, MESFUNC2:25; hence c (#) f is_integrable_on M by A5, A10, A22, A24, A25, Def17; ::_thesis: Integral (M,(c (#) f)) = (R_EAL c) * (Integral (M,f)) A26: - (R_EAL a) = - a by SUPINF_2:2; thus Integral (M,(c (#) f)) = ((R_EAL a) * (integral+ (M,(max- f)))) - (integral+ (M,(a (#) (max+ f)))) by A3, A4, A8, A6, A20, A21, A22, Th86, MESFUNC2:26 .= ((R_EAL a) * (integral+ (M,(max- f)))) - ((R_EAL a) * (integral+ (M,(max+ f)))) by A4, A9, A11, A20, Th86, MESFUNC2:25 .= (R_EAL a) * ((integral+ (M,(max- f))) - (integral+ (M,(max+ f)))) by XXREAL_3:100 .= (R_EAL a) * (- ((integral+ (M,(max+ f))) - (integral+ (M,(max- f))))) by XXREAL_3:26 .= - ((R_EAL a) * ((integral+ (M,(max+ f))) - (integral+ (M,(max- f))))) by XXREAL_3:92 .= (R_EAL c) * (Integral (M,f)) by A19, A26, XXREAL_3:92 ; ::_thesis: verum end; end; end; definition let X be non empty set ; let S be SigmaField of X; let M be sigma_Measure of S; let f be PartFunc of X,ExtREAL; let B be Element of S; func Integral_on (M,B,f) -> Element of ExtREAL equals :: MESFUNC5:def 18 Integral (M,(f | B)); coherence Integral (M,(f | B)) is Element of ExtREAL ; end; :: deftheorem defines Integral_on MESFUNC5:def_18_:_ for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for B being Element of S holds Integral_on (M,B,f) = Integral (M,(f | B)); theorem :: MESFUNC5:111 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL for B being Element of S st f is_integrable_on M & g is_integrable_on M & B c= dom (f + g) holds ( f + g is_integrable_on M & Integral_on (M,B,(f + g)) = (Integral_on (M,B,f)) + (Integral_on (M,B,g)) ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL for B being Element of S st f is_integrable_on M & g is_integrable_on M & B c= dom (f + g) holds ( f + g is_integrable_on M & Integral_on (M,B,(f + g)) = (Integral_on (M,B,f)) + (Integral_on (M,B,g)) ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f, g being PartFunc of X,ExtREAL for B being Element of S st f is_integrable_on M & g is_integrable_on M & B c= dom (f + g) holds ( f + g is_integrable_on M & Integral_on (M,B,(f + g)) = (Integral_on (M,B,f)) + (Integral_on (M,B,g)) ) let M be sigma_Measure of S; ::_thesis: for f, g being PartFunc of X,ExtREAL for B being Element of S st f is_integrable_on M & g is_integrable_on M & B c= dom (f + g) holds ( f + g is_integrable_on M & Integral_on (M,B,(f + g)) = (Integral_on (M,B,f)) + (Integral_on (M,B,g)) ) let f, g be PartFunc of X,ExtREAL; ::_thesis: for B being Element of S st f is_integrable_on M & g is_integrable_on M & B c= dom (f + g) holds ( f + g is_integrable_on M & Integral_on (M,B,(f + g)) = (Integral_on (M,B,f)) + (Integral_on (M,B,g)) ) let B be Element of S; ::_thesis: ( f is_integrable_on M & g is_integrable_on M & B c= dom (f + g) implies ( f + g is_integrable_on M & Integral_on (M,B,(f + g)) = (Integral_on (M,B,f)) + (Integral_on (M,B,g)) ) ) assume that A1: f is_integrable_on M and A2: g is_integrable_on M and A3: B c= dom (f + g) ; ::_thesis: ( f + g is_integrable_on M & Integral_on (M,B,(f + g)) = (Integral_on (M,B,f)) + (Integral_on (M,B,g)) ) A4: dom (f | B) = (dom f) /\ B by RELAT_1:61; dom (f + g) = ((dom f) /\ (dom g)) \ (((f " {-infty}) /\ (g " {+infty})) \/ ((f " {+infty}) /\ (g " {-infty}))) by MESFUNC1:def_3; then A5: dom (f + g) c= (dom f) /\ (dom g) by XBOOLE_1:36; (dom f) /\ (dom g) c= dom f by XBOOLE_1:17; then dom (f + g) c= dom f by A5, XBOOLE_1:1; then A6: dom (f | B) = B by A3, A4, XBOOLE_1:1, XBOOLE_1:28; A7: Integral_on (M,B,(f + g)) = Integral (M,((f | B) + (g | B))) by A3, Th29; A8: g | B is_integrable_on M by A2, Th97; A9: dom (g | B) = (dom g) /\ B by RELAT_1:61; (dom f) /\ (dom g) c= dom g by XBOOLE_1:17; then dom (f + g) c= dom g by A5, XBOOLE_1:1; then A10: dom (g | B) = B by A3, A9, XBOOLE_1:1, XBOOLE_1:28; f | B is_integrable_on M by A1, Th97; hence ( f + g is_integrable_on M & Integral_on (M,B,(f + g)) = (Integral_on (M,B,f)) + (Integral_on (M,B,g)) ) by A1, A2, A6, A8, A10, A7, Lm13, Th108; ::_thesis: verum end; theorem :: MESFUNC5:112 for X being non empty set for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for c being Real for B being Element of S st f is_integrable_on M holds ( f | B is_integrable_on M & Integral_on (M,B,(c (#) f)) = (R_EAL c) * (Integral_on (M,B,f)) ) proof let X be non empty set ; ::_thesis: for S being SigmaField of X for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for c being Real for B being Element of S st f is_integrable_on M holds ( f | B is_integrable_on M & Integral_on (M,B,(c (#) f)) = (R_EAL c) * (Integral_on (M,B,f)) ) let S be SigmaField of X; ::_thesis: for M being sigma_Measure of S for f being PartFunc of X,ExtREAL for c being Real for B being Element of S st f is_integrable_on M holds ( f | B is_integrable_on M & Integral_on (M,B,(c (#) f)) = (R_EAL c) * (Integral_on (M,B,f)) ) let M be sigma_Measure of S; ::_thesis: for f being PartFunc of X,ExtREAL for c being Real for B being Element of S st f is_integrable_on M holds ( f | B is_integrable_on M & Integral_on (M,B,(c (#) f)) = (R_EAL c) * (Integral_on (M,B,f)) ) let f be PartFunc of X,ExtREAL; ::_thesis: for c being Real for B being Element of S st f is_integrable_on M holds ( f | B is_integrable_on M & Integral_on (M,B,(c (#) f)) = (R_EAL c) * (Integral_on (M,B,f)) ) let c be Real; ::_thesis: for B being Element of S st f is_integrable_on M holds ( f | B is_integrable_on M & Integral_on (M,B,(c (#) f)) = (R_EAL c) * (Integral_on (M,B,f)) ) let B be Element of S; ::_thesis: ( f is_integrable_on M implies ( f | B is_integrable_on M & Integral_on (M,B,(c (#) f)) = (R_EAL c) * (Integral_on (M,B,f)) ) ) assume f is_integrable_on M ; ::_thesis: ( f | B is_integrable_on M & Integral_on (M,B,(c (#) f)) = (R_EAL c) * (Integral_on (M,B,f)) ) then A1: f | B is_integrable_on M by Th97; A2: for x being set st x in dom ((c (#) f) | B) holds ((c (#) f) | B) . x = (c (#) (f | B)) . x proof let x be set ; ::_thesis: ( x in dom ((c (#) f) | B) implies ((c (#) f) | B) . x = (c (#) (f | B)) . x ) assume A3: x in dom ((c (#) f) | B) ; ::_thesis: ((c (#) f) | B) . x = (c (#) (f | B)) . x then A4: ((c (#) f) | B) . x = (c (#) f) . x by FUNCT_1:47; A5: x in (dom (c (#) f)) /\ B by A3, RELAT_1:61; then x in (dom f) /\ B by MESFUNC1:def_6; then A6: x in dom (f | B) by RELAT_1:61; x in dom (c (#) f) by A5, XBOOLE_0:def_4; then ((c (#) f) | B) . x = (R_EAL c) * (f . x) by A4, MESFUNC1:def_6; then A7: ((c (#) f) | B) . x = (R_EAL c) * ((f | B) . x) by A6, FUNCT_1:47; x in dom (c (#) (f | B)) by A6, MESFUNC1:def_6; hence ((c (#) f) | B) . x = (c (#) (f | B)) . x by A7, MESFUNC1:def_6; ::_thesis: verum end; dom ((c (#) f) | B) = (dom (c (#) f)) /\ B by RELAT_1:61; then dom ((c (#) f) | B) = (dom f) /\ B by MESFUNC1:def_6; then dom ((c (#) f) | B) = dom (f | B) by RELAT_1:61; then dom ((c (#) f) | B) = dom (c (#) (f | B)) by MESFUNC1:def_6; then (c (#) f) | B = c (#) (f | B) by A2, FUNCT_1:2; hence ( f | B is_integrable_on M & Integral_on (M,B,(c (#) f)) = (R_EAL c) * (Integral_on (M,B,f)) ) by A1, Th110; ::_thesis: verum end;