:: Integral of Measurable Function :: by Noboru Endou and Yasunari Shidama :: :: Received May 24, 2006 :: Copyright (c) 2006-2018 Association of Mizar Users :: (Stowarzyszenie Uzytkownikow Mizara, Bialystok, Poland). :: This code can be distributed under the GNU General Public Licence :: version 3.0 or later, or the Creative Commons Attribution-ShareAlike :: License version 3.0 or later, subject to the binding interpretation :: detailed in file COPYING.interpretation. :: See COPYING.GPL and COPYING.CC-BY-SA for the full text of these :: licenses, or see http://www.gnu.org/licenses/gpl.html and :: http://creativecommons.org/licenses/by-sa/3.0/. environ vocabularies NUMBERS, SUPINF_1, COMPLEX1, ARYTM_1, XXREAL_0, ARYTM_3, MEASURE6, NAT_1, CARD_1, NEWTON, RELAT_1, REAL_1, SUBSET_1, INT_1, PARTFUN1, FUNCT_1, SUPINF_2, XBOOLE_0, PROB_1, MESFUNC2, VALUED_0, RAT_1, MESFUNC1, TARSKI, MEASURE1, VALUED_1, RFUNCT_3, FINSEQ_1, MESFUNC3, CARD_3, ZFMISC_1, SEQ_2, ORDINAL2, XCMPLX_0, XXREAL_2, SEQFUNC, PBOOLE, INTEGRA1, FINSET_1, MEMBERED, SETLIM_1, INTEGRA5, FUNCT_3, MESFUNC5, FUNCT_7; notations TARSKI, XBOOLE_0, ZFMISC_1, SUBSET_1, ORDINAL1, NUMBERS, XCMPLX_0, XXREAL_0, XREAL_0, COMPLEX1, XXREAL_2, XXREAL_3, REAL_1, RELAT_1, FUNCT_1, FINSEQ_1, RELSET_1, FINSET_1, PARTFUN1, FUNCT_2, INT_1, NAT_D, RAT_1, MEMBERED, SUPINF_1, SUPINF_2, MEASURE1, EXTREAL1, NAT_1, MESFUNC1, MESFUNC2, MESFUNC3, FINSEQOP, RFUNCT_3, PROB_1, MEASURE6, NEWTON, SEQFUNC, SETLIM_1, VALUED_0; constructors WELLORD2, REAL_1, SQUARE_1, NAT_D, FINSEQOP, LIMFUNC1, SEQFUNC, NEWTON, RFUNCT_3, MEASURE6, EXTREAL1, MESFUNC1, BINARITH, MESFUNC2, KURATO_0, MESFUNC3, SETLIM_1, SUPINF_1, RELSET_1, XREAL_0; registrations SUBSET_1, ORDINAL1, RELSET_1, PARTFUN1, FINSET_1, NUMBERS, XXREAL_0, XREAL_0, NAT_1, INT_1, RAT_1, MEMBERED, COMPLEX1, FINSEQ_1, MEASURE1, VALUED_0, XXREAL_2, XXREAL_3, XCMPLX_0, NEWTON, EXTREAL1, SUPINF_2; requirements NUMERALS, REAL, BOOLE, SUBSET, ARITHM; definitions TARSKI, XBOOLE_0, FINSEQ_1, SUPINF_2; equalities XBOOLE_0, FINSEQ_1, CARD_3, XXREAL_3, SUPINF_2; expansions TARSKI, XBOOLE_0, SUPINF_2; theorems MEASURE1, MEASURE2, TARSKI, PARTFUN1, FUNCT_1, FUNCT_2, MEASURE6, NAT_1, SUPINF_2, RELSET_1, EXTREAL1, INT_1, MESFUNC1, FINSEQ_1, XREAL_0, PROB_1, PROB_2, XBOOLE_0, NAT_2, XBOOLE_1, XCMPLX_1, MESFUNC2, MESFUNC3, NEWTON, XREAL_1, PREPOWER, XXREAL_0, MESFUNC4, ZFMISC_1, SETLIM_1, FINSEQ_3, RAT_1, FUNCT_3, RELAT_1, FINSUB_1, GRFUNC_1, FINSEQ_2, ORDINAL1, NAT_D, NUMBERS, CARD_1, XXREAL_2, VALUED_0, XXREAL_3, KURATO_0; schemes CLASSES1, FUNCT_2, FINSEQ_1, FINSEQ_2, NAT_1, SEQFUNC, PARTFUN1; begin :: Lemmas for extended real numbers reconsider jj=1 as Real; theorem Th1: for x,y be R_eal holds |.x-y.| = |.y-x.| proof let x,y be R_eal; |.y-x.| = |.-(x-y).| by XXREAL_3:26; hence thesis by EXTREAL1:29; end; theorem Th2: for x,y be R_eal holds y-x <= |.x-y.| proof let x,y be R_eal; -|.x-y.| <= (x-y) by EXTREAL1:20; then -(x-y) <= |.x-y.| by XXREAL_3:60; hence thesis by XXREAL_3:26; end; theorem Th3: for x,y be R_eal, e be Real st |.x-y.| < e & not ( x = +infty & y = +infty or x = -infty & y = -infty ) holds x <> +infty & x <> -infty & y <> +infty & y <>-infty proof let x,y be R_eal, e be Real; assume A1: |.x-y.| < e; y-x <= |.x-y.| by Th2; then A2: y-x < e by A1,XXREAL_0:2; x-y <= |.x-y.| by EXTREAL1:20; then x-y < e by A1,XXREAL_0:2; hence thesis by A2,XXREAL_3:54; end; theorem Th4: for n be Nat, p be ExtReal st 0 <= p & p < n ex k be Nat st 1 <= k & k <= 2|^n*n & (k-1)/(2|^n) <= p & p < k/(2|^n) proof let n be Nat; let p be ExtReal; assume that A1: 0 <= p and A2: p < n; 0 in REAL by XREAL_0:def 1; then reconsider p1 = p as Element of REAL by A1,A2,XXREAL_0:46; reconsider 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 < k by A1,INT_1:def 6; then reconsider k 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 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; 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; end; take k; 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 thesis by A6,A7,A5,A11,XREAL_1:79,81; end; theorem Th5: for n,k be Nat, p be ExtReal st k <= 2|^n*n & n <= p holds k/(2|^n) <= p proof let n,k be Nat; let p be ExtReal; assume that A1: k <= 2|^n*n and A2: n <= p; assume p < k/(2|^n); then n < k/(2|^n) by A2,XXREAL_0:2; hence contradiction by A1,PREPOWER:6,XREAL_1:79; end; theorem Th6: for x,y,k being ExtReal 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 ExtReal; assume A1: 0 <= k; now per cases by XXREAL_0:16; suppose A2: max(x,y) = x; 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; end; suppose A3: max(x,y) = 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; end; end; hence k*max(x,y) = max(k*x,k*y); per cases by XXREAL_0:15; suppose A4: min(x,y) = x; then x <= y by XXREAL_0:def 9; then k*x <= k*y by A1,XXREAL_3:71; hence thesis by A4,XXREAL_0:def 9; end; suppose A5: min(x,y) = y; then y <= x by XXREAL_0:def 9; then k*y <= k*x by A1,XXREAL_3:71; hence thesis by A5,XXREAL_0:def 9; end; end; theorem 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; assume A1: k <= 0; hereby per cases by XXREAL_0:16; suppose max(x,y) = x; 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; end; suppose max(x,y) = 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; end; end; per cases by XXREAL_0:15; suppose min(x,y) = x; 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 thesis by A4,XXREAL_0:def 10; end; suppose min(x,y) = 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 thesis by A5,XXREAL_0:def 10; end; end; begin :: Lemmas for partial function of non empty set,extended real numbers definition let IT be set; attr IT is nonpositive means for x being R_eal holds x in IT implies x <= 0; end; definition let R be Relation; attr R is nonpositive means rng R is nonpositive; end; theorem Th8: for X being set, F being PartFunc of X,ExtREAL holds F is nonpositive iff for n being object holds F.n <= 0. proof let X be set, F be PartFunc of X,ExtREAL; hereby assume F is nonpositive; then A1: rng F is nonpositive; let n be object; per cases; suppose n in dom F; then F.n in rng F by FUNCT_1:def 3; hence F.n <= 0. by A1; end; suppose not n in dom F; hence F.n <= 0. by FUNCT_1:def 2; end; end; assume A2: for n being object holds F.n <= 0.; let y be R_eal; assume y in rng F; then ex x being object st x in dom F & y = F.x by FUNCT_1:def 3; hence thesis by A2; end; theorem Th9: for X being set, 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, F be PartFunc of X,ExtREAL such that A1: for n being set st n in dom F holds F.n <= 0.; let y be R_eal; assume y in rng F; then ex x being object st x in dom F & y = F.x by FUNCT_1:def 3; hence thesis by A1; end; definition let R be Relation; attr R is without-infty means :Def3: not -infty in rng R; attr R is without+infty means not +infty in rng R; end; definition let X be non empty set, f be PartFunc of X,ExtREAL; redefine attr f is without-infty means :Def5: for x being object holds -infty < f.x; compatibility proof hereby assume f is without-infty; then A1: not -infty in rng f; hereby let x be object; per cases; suppose x in dom f; then f.x <> -infty by A1,FUNCT_1:def 3; hence -infty < f.x by XXREAL_0:6; end; suppose not x in dom f; hence -infty < f.x by FUNCT_1:def 2; end; end; end; assume A2: for x being object holds -infty < f.x; now assume -infty in rng f; then ex x be object st x in dom f & f.x = -infty by FUNCT_1:def 3; hence contradiction by A2; end; hence thesis; end; redefine attr f is without+infty means for x being object holds f.x < +infty; compatibility proof hereby assume f is without+infty; then A3: not +infty in rng f; hereby let x be object; per cases; suppose x in dom f; then f.x <> +infty by A3,FUNCT_1:def 3; hence f.x < +infty by XXREAL_0:4; end; suppose not x in dom f; hence f.x < +infty by FUNCT_1:def 2; end; end; end; assume A4: for x being object holds f.x < +infty; now assume +infty in rng f; then ex x be object st x in dom f & f.x = +infty by FUNCT_1:def 3; hence contradiction by A4; end; hence thesis; end; end; theorem Th10: for X be non empty set, f be PartFunc of X,ExtREAL holds (for x be set st x in dom f holds -infty < f.x) iff f is without-infty proof let X be non empty set, f be PartFunc of X,ExtREAL; hereby assume A1: for x be set st x in dom f holds -infty < f.x; now let x be object; per cases; suppose x in dom f; hence -infty < f.x by A1; end; suppose not x in dom f; hence -infty < f.x by FUNCT_1:def 2; end; end; hence f is without-infty; end; assume f is without-infty; hence thesis; end; theorem Th11: for X be non empty set, f be PartFunc of X,ExtREAL holds (for x be set st x in dom f holds f.x < +infty) iff f is without+infty proof let X be non empty set, f be PartFunc of X,ExtREAL; hereby assume A1: for x be set st x in dom f holds f.x < +infty; now let x be object; per cases; suppose x in dom f; hence f.x < +infty by A1; end; suppose not x in dom f; hence f.x < +infty by FUNCT_1:def 2; end; end; hence f is without+infty; end; assume f is without+infty; hence thesis; end; theorem Th12: for X be non empty set, f be PartFunc of X,ExtREAL st f is nonnegative holds f is without-infty by SUPINF_2:51; theorem Th13: for X be non empty set, f be PartFunc of X,ExtREAL st f is nonpositive holds f is without+infty by Th8; registration let X be non empty set; cluster nonnegative -> without-infty for PartFunc of X,ExtREAL; coherence by Th12; cluster nonpositive -> without+infty for PartFunc of X,ExtREAL; coherence by Th13; end; theorem Th14: for X be non empty set, S be SigmaField of X, f be PartFunc of X ,ExtREAL st f is_simple_func_in S holds f is without+infty & f is without-infty proof let X be non empty set, S be SigmaField of X, f be PartFunc of X,ExtREAL; assume A1: f is_simple_func_in S; hereby assume not f is without+infty; then +infty in rng f; then f"{+infty} <> {} by FUNCT_1:72; then consider x be object such that A2: x in f"{+infty} by XBOOLE_0:def 1; A3: f is real-valued by A1,MESFUNC2:def 4; f.x in {+infty} by A2,FUNCT_1:def 7; hence contradiction by A3,TARSKI:def 1; end; hereby assume not f is without-infty; then -infty in rng f; then f"{-infty} <> {} by FUNCT_1:72; then consider x be object such that A4: x in f"{-infty} by XBOOLE_0:def 1; A5: f is real-valued by A1,MESFUNC2:def 4; f.x in {-infty} by A4,FUNCT_1:def 7; hence contradiction by A5,TARSKI:def 1; end; end; theorem Th15: for X be non empty set, Y be set, f be PartFunc of X,ExtREAL st f is nonnegative holds f|Y is nonnegative proof let X be non empty set, Y be set, f be PartFunc of X,ExtREAL; assume A1: f is nonnegative; now let x be object; assume A2: x in dom(f|Y); then (f|Y).x = f.x by FUNCT_1:47; hence 0 <= (f|Y).x by A1,A2,SUPINF_2:39; end; hence thesis by SUPINF_2:52; end; theorem Th16: for X be non empty set, f,g be PartFunc of X,ExtREAL st f is without-infty & g is without-infty holds dom(f+g)=dom f /\ dom g proof let X be non empty set; let f,g be PartFunc of X,ExtREAL; assume that A1: f is without-infty and A2: g is without-infty; not -infty in rng g by A2; then A3: g"{-infty} = {} by FUNCT_1:72; not -infty in rng f by A1; 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 thesis; end; theorem for X be non empty set, f,g be PartFunc of X,ExtREAL st f is without-infty & g is without+infty holds dom(f-g)=dom f /\ dom g proof let X be non empty set; let f,g be PartFunc of X,ExtREAL; assume that A1: f is without-infty and A2: g is without+infty; not +infty in rng g by A2; then A3: g"{+infty} = {} by FUNCT_1:72; not -infty in rng f by A1; 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 thesis; end; theorem Th18: for X be non empty set, S be SigmaField of X, f,g be PartFunc of X,ExtREAL, F be Function of RAT,S, r be Real, A be Element of S st f is without-infty & g is without-infty & (for p be Rational holds F.p = A /\ less_dom(f,p) /\ (A /\ less_dom(g, (r-(p qua Complex))))) holds A /\ less_dom(f+g,r) = union rng F proof let X be non empty set; let S be SigmaField of X; let f,g be PartFunc of X,ExtREAL; let F be Function of RAT,S; let r be Real; let A be Element of S; assume that A1: f is without-infty and A2: g is without-infty and A3: for p be Rational holds F.p = A /\ less_dom(f,p) /\ (A /\ less_dom(g,(r-(p qua Complex)))); A4: dom(f+g) = dom f /\ dom g by A1,A2,Th16; A5: union rng F c= A /\ less_dom(f+g,r) proof let x be object; assume x in union rng F; then consider Y being set such that A6: x in Y and A7: Y in rng F by TARSKI:def 4; consider p being object such that A8: p in dom F and A9: Y = F.p by A7,FUNCT_1:def 3; reconsider p as Rational by A8; A10: x in A /\ less_dom(f,p)/\(A /\ less_dom(g,r-p)) by A3,A6,A9; then A11: x in A /\ less_dom(f,p) by XBOOLE_0:def 4; then A12: x in A by XBOOLE_0:def 4; A13: x in less_dom(f,p) by A11,XBOOLE_0:def 4; x in A /\ less_dom(g,r-p) by A10,XBOOLE_0:def 4; then A14: x in less_dom(g,r-p) by XBOOLE_0:def 4; reconsider x as Element of X by A10; f.x < p by A13,MESFUNC1:def 11; then A15: f.x <> +infty by XXREAL_0:4; A16: -infty < f.x by A1; A17: -infty < g.x by A2; A18: g.x < r-p by A14,MESFUNC1:def 11; then g.x <> +infty by XXREAL_0:4; then reconsider f1 = f.x,g1 = g.x as Element of 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) by A20,A22,MESFUNC1:def 11; hence thesis by A12,XBOOLE_0:def 4; end; A /\ less_dom(f+g,r) c= union rng F proof let x be object; assume A23: x in A /\ less_dom(f+g,r); then A24: x in A by XBOOLE_0:def 4; A25: x in less_dom(f+g,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 by A25,MESFUNC1:def 11; A29: x in dom g by A4,A26,XBOOLE_0:def 4; reconsider x as Element of X by A23; A30: -infty < f.x by A1; A31: f.x + g.x < r by A26,A28,MESFUNC1:def 3; then A32: g.x <> +infty by A30,XXREAL_3:52; A33: -infty < g.x by A2; then f.x <> +infty by A31,XXREAL_3:52; then reconsider f1 = f.x, g1 = g.x as Element of REAL by A30,A33,A32,XXREAL_0:14; f.x < r - g.x by A31,A30,A33,XXREAL_3:52; then consider p being Rational such that A34: f1 < p and A35: p < r - g1 by RAT_1:7; not r - p <= g1 by A35,XREAL_1:12; then x in less_dom(g,r-p) by A29,MESFUNC1:def 11; then A36: x in A /\ less_dom(g,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 A37: F.p in rng F by FUNCT_1:def 3; x in less_dom(f,p) by A27,A34,MESFUNC1:def 11; then x in A /\ less_dom(f,p) by A24,XBOOLE_0:def 4; then x in A /\ less_dom(f,p)/\(A /\ less_dom(g,r-p)) by A36, XBOOLE_0:def 4; then x in F.p by A3; hence thesis by A37,TARSKI:def 4; end; hence thesis by A5; end; definition let X be non empty set; let f be PartFunc of X,REAL; func R_EAL f -> PartFunc of X,ExtREAL equals f; coherence by NUMBERS:31,RELSET_1:7; end; theorem Th19: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL st f is nonnegative & g is nonnegative holds f+g is nonnegative proof let X be non empty set; let S be SigmaField of X; let M be sigma_Measure of S; let f,g be PartFunc of X,ExtREAL; assume that A1: f is nonnegative and A2: g is nonnegative; for x be object st x in dom(f+g) holds 0 <= (f+g).x proof let x be object; assume A3: x in dom(f+g); 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 thesis by A3,A4,MESFUNC1:def 3; end; hence thesis by SUPINF_2:52; end; theorem Th20: for X be non empty set, f be PartFunc of X,ExtREAL,c be 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; let f be PartFunc of X,ExtREAL; let c be Real; set g = c(#)f; assume A1: f is nonnegative; hereby set g = c(#)f; assume A2: 0 <= c; for x be object st x in dom g holds 0 <= g.x proof let x be object; 0 <= f.x by A1,SUPINF_2:51; then A3: 0 <= c*f.x by A2; assume x in dom g; hence thesis by A3,MESFUNC1:def 6; end; hence c(#)f is nonnegative by SUPINF_2:52; end; assume A4: c <= 0; now let x be set; 0 <= f.x by A1,SUPINF_2:51; then A5: c*f.x <= 0 by A4; assume x in dom g; hence g.x <= 0 by A5,MESFUNC1:def 6; end; hence thesis by Th9; end; theorem Th21: for X be non empty set, f,g be PartFunc of X,ExtREAL st (for x be 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, f,g be PartFunc of X,ExtREAL; assume A1: for x be set st x in dom f /\ dom g holds g.x <= f.x & -infty < g.x & f.x < +infty; now let x be object; assume A2: x in dom(f-g); 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; end; hence thesis by SUPINF_2:52; end; Lm1: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL holds max+f is nonnegative & max-f is nonnegative & |. f.| is nonnegative proof 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; A1: for x be object st x in dom max- f holds 0<= (max-f).x by MESFUNC2:13; for x be object 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; now let x be object; assume x in dom |.f.|; then (|.f.|).x =|. f.x .| by MESFUNC1:def 10; hence 0 <= (|.f.|).x by EXTREAL1:14; end; hence thesis by SUPINF_2:52; end; theorem Th22: for X be non empty set, f,g be 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, f,g be PartFunc of X,ExtREAL; assume that A1: f is nonnegative and A2: g is nonnegative; thus A3: dom(f+g)=dom f /\ dom g by A1,A2,Th16; now let x be object; assume A4: x in dom f /\ dom g; 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; end; hence thesis by A3,SUPINF_2:52; end; theorem Th23: for X be non empty set, f,g,h be 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 be 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; let f,g,h be PartFunc of X,ExtREAL; assume that A1: f is nonnegative and A2: g is nonnegative and A3: h is nonnegative; 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; thus f+g+h is nonnegative by A3,A4,Th22; hereby let x be set; assume x in dom f /\ dom g /\ dom h; 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; end; end; theorem Th24: for X be non empty set, f,g being PartFunc of X,ExtREAL st f is without-infty & g is without-infty 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; let f,g be PartFunc of X,ExtREAL; assume that A1: f is without-infty and A2: g is without-infty; 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 be 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 ,13; then A6: max+f is without-infty by Th10; A7: max-f is without-infty by A5,Th10; A8: for x be 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,13; then max+(f+g) is without-infty by Th10; then A9: dom(max+(f+g) + max- f) = dom(max+(f+g)) /\ dom(max- f) by A7,Th16; max-(f+g) is without-infty 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 without-infty by A5,Th10; A12: dom(max- f) = dom f by MESFUNC2:def 3; A13: max+g is without-infty 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; A18: dom(max+ g) = dom g by MESFUNC2:def 2; A19: for x be object 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 object; hereby assume A20: x in dom(max+(f+g) + max- f); 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; end; assume A22: x in dom(max-(f+g) + max+ f); 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 thesis by A22,MESFUNC1:def 3; end; then A24: for x be 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 without-infty 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; max-(f+g) + max+f is without-infty 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; thus thesis by A19,SUPINF_2:52; end; theorem Th25: for X being non empty set, f,g being PartFunc of X,ExtREAL st f is without-infty & f is without+infty & g is without-infty & g is without+infty holds max+(f+g) + max- f + max- g = max-(f+g) + max+ f + max+ g proof let X be non empty set, f,g be PartFunc of X,ExtREAL; assume that A1: f is without-infty and A2: f is without+infty and A3: g is without-infty and A4: g is without+infty; A5: dom(max-(f+g))= dom (f+g) by MESFUNC2:def 3; for x be object 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 be object 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 be object 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 be object 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 be object 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 be object 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 be object 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 object; assume A20: x in dom(max+(f+g) + max-f + max-g); 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; 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; then A30: -f.x <> +infty by XXREAL_3:23; A31: f.x <> +infty by A2; 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; then A35: -g.x <> +infty by XXREAL_3:23; A36: now per cases; suppose A37: 0 <= f.x; then A38: max-f.x =0 by A32,XXREAL_0:def 10; per cases; suppose A39: 0 <= 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 + 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 = 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; end; suppose A42: g.x < 0; 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; per cases; suppose A45: 0 <= f.x + 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 = 0 + f.x + 0 by A27,A37,A43,XXREAL_0:def 10; (max+(f+g)).x + (max-f).x + (max-g).x =f.x+g.x + 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 + 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; end; suppose A47: f.x + g.x < 0; then (max+(f+g)).x = 0 by A25,XXREAL_0:def 10; then (max+(f+g)).x + (max-f).x + (max-g).x = 0 + 0 + -(g.x) by A38,A44; 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 + 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; end; end; end; suppose A49: f.x < 0; then A50: (max-f).x = -f.x by A32,XXREAL_0:def 10; per cases; suppose A51: 0 <= 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; per cases; suppose A54: 0 <= f.x + g.x; then A55: (max-(f+g)).x = 0 by A28,XXREAL_0:def 10; (max+f).x = 0 by A27,A49,XXREAL_0:def 10; then A56: (max-(f+g)).x + (max+f).x + (max+g).x = 0 + 0 + g.x by A52,A55; (max+(f+g)).x + (max-f).x + (max-g).x =f.x + g.x+ -f.x + 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 A56; end; suppose A57: f.x + g.x < 0; then (max-(f+g)).x = -(f.x + g.x) by A28,XXREAL_0:def 10; then A58: (max-(f+g)).x + (max+f).x + (max+g).x = -(f.x + g.x) + 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 = 0 + -f.x + 0 by A25,A50,A53,A57,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 A58,XXREAL_3:7; end; end; suppose A59: g.x < 0; then (max-g).x=-g.x by A33,XXREAL_0:def 10; then A60: (max+(f+g)).x+ (max-f).x + (max-g).x = 0 + -f.x + -g.x by A25 ,A49,A50,A59,XXREAL_0:def 10 .=-f.x -g.x by XXREAL_3:4; A61: (max+g).x = 0 by A22,A59,XXREAL_0:def 10; (max-(f+g)).x=-(f.x + g.x) by A28,A49,A59,XXREAL_0:def 10; then (max-(f+g)).x + (max+f).x + (max+g).x = -(f.x + g.x) + 0 + 0 by A27,A49,A61,XXREAL_0:def 10 .= -(f.x + g.x) + 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 A60,XXREAL_3:25; end; end; end; A62: 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 thesis by A9,A14,A6,A20,A62,A36,Th23; 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 thesis by A19,FUNCT_1:2; end; theorem Th26: for C being non empty set, f being PartFunc of C,ExtREAL, c be Real st 0 <= c holds max+(c(#)f) = c(#)max+f & max-(c(#)f) = c(#)max-f proof let C be non empty set; let f be PartFunc of C,ExtREAL; let c be Real; assume A1: 0 <= c; 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 be 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; assume A3: x in dom max+(c(#)f); 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( c * f.x,0) by A4,MESFUNC1:def 6; (c(#)max+f).x = 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 qua ExtReal)) by A1,Th6; hence thesis by A6; end; hence max+(c(#)f) = c(#)max+f by A2,PARTFUN1:5; 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 be 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; assume A8: x in dom max-(c(#)f); 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(-( c)*f.x,0) by A9,MESFUNC1:def 6; (c(#)max-f).x = 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 qua ExtReal)) by A1,Th6 .= max(-(c)*f.x, c * (0 qua ExtReal)) by XXREAL_3:92; hence thesis by A11; end; hence thesis by A7,PARTFUN1:5; end; theorem Th27: for C being non empty set, f being PartFunc of C,ExtREAL, c be Real st 0 <= c holds max+((-c)(#)f) = c(#)max-f & max-((-c)(#)f) = c(#)max+f proof let C be non empty set; let f be PartFunc of C,ExtREAL; let c be Real; assume A1: 0 <= c; A2: dom max+((-c)(#)f) = dom((-c)(#)f) by MESFUNC2:def 2; then dom max+((-c)(#)f) = dom f by MESFUNC1:def 6; then A3: dom max+((-c)(#)f) = dom max-f by MESFUNC2:def 3; then A4: dom max+((-c)(#)f) = dom(c(#)max-f) by MESFUNC1:def 6; for x be 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; assume A5: x in dom max+((-c)(#)f); then A6: (max+((-c)(#)f)).x = max(((-c)(#)f).x,0) by MESFUNC2:def 2 .= max((-c)*f.x,0) by A2,A5,MESFUNC1:def 6 .= max(-( c * f.x),0) by XXREAL_3:92; (c(#)max-f).x = c * max-f.x by A4,A5,MESFUNC1:def 6 .= c * max(-f.x, 0) by A3,A5,MESFUNC2:def 3 .= max( c * (-f.x), c * 0) by A1,Th6 .= max(-( c * f.x), c * (0 qua ExtReal)) by XXREAL_3:92; hence thesis by A6; end; hence max+((-c)(#)f) = c(#)max-f by A4,PARTFUN1:5; A7: dom max-((-c)(#)f) = dom((-c)(#)f) by MESFUNC2:def 3; then dom max-((-c)(#)f) = dom f by MESFUNC1:def 6; then A8: dom max-((-c)(#)f) = dom max+f by MESFUNC2:def 2; then A9: dom max-((-c)(#)f) = dom(c(#) max+ f) by MESFUNC1:def 6; for x be 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; assume A10: x in dom max-((-c)(#)f); then A11: max-((-c)(#)f).x = max(-((-c)(#)f).x,0) by MESFUNC2:def 3 .= max(-((-c))*f.x,0) by A7,A10,MESFUNC1:def 6 .= max((-(-( c)))*f.x,0) by XXREAL_3:92; (c(#)max+f).x = c * max+f.x by A9,A10,MESFUNC1:def 6 .= c * max(f.x, 0) by A8,A10,MESFUNC2:def 2 .= max( c * f.x, c * (0 qua ExtReal)) by A1,Th6; hence thesis by A11; end; hence thesis by A9,PARTFUN1:5; end; theorem Th28: for X be non empty set, f be PartFunc of X,ExtREAL, A be set holds max+(f|A)=max+f|A & max-(f|A)=max-f|A proof let X be non empty set; let f be PartFunc of X,ExtREAL; let A be set; 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; assume A2: x in dom max+(f|A); 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 thesis by A5,A3,MESFUNC2:def 2; end; hence max+(f|A) = max+f|A by A1,PARTFUN1:5; 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; assume A8: x in dom max-(f|A); 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 thesis by A11,A9,MESFUNC2:def 3; end; hence thesis by A7,PARTFUN1:5; end; theorem Th29: for X be non empty set, f,g be PartFunc of X,ExtREAL, B be 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, f,g be PartFunc of X,ExtREAL, B be set such that A1: B c= dom(f+g); for x be object st x in dom g holds g.x in ExtREAL by XXREAL_0:def 1; then reconsider gg = g as Function of dom g,ExtREAL by FUNCT_2:3; for x be object st x in dom(g|B) holds (g|B).x in ExtREAL by XXREAL_0:def 1; then reconsider gb = g|B as Function of dom (g|B), ExtREAL by FUNCT_2:3; now let x be object; assume A2: x in g"{+infty} /\ B; 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; end; then A6: g"{+infty} /\ B c= (g|B)"{+infty}; now let x be object; assume A7: x in (g|B)"{+infty}; 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; end; then (g|B)"{+infty} c= g"{+infty} /\ B; then A12: (g|B)"{+infty}=g"{+infty} /\ B by A6; now let x be object; assume A13: x in g"{-infty} /\ B; 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; end; then A17: g"{-infty} /\ B c= (g|B)"{-infty}; now let x be object; assume A18: x in (g|B)"{-infty}; 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; end; then (g|B)"{-infty} c= g"{-infty} /\ B; then A23: (g|B)"{-infty}=g"{-infty} /\B by A17; for x be object st x in dom f holds f.x in ExtREAL by XXREAL_0:def 1; then reconsider ff = f as Function of dom f,ExtREAL by FUNCT_2:3; for x be object st x in dom(f|B) holds (f|B).x in ExtREAL by XXREAL_0:def 1; then reconsider fb = f|B as Function of dom(f|B),ExtREAL by FUNCT_2:3; now let x be object; assume A24: x in f"{+infty} /\ B; 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; end; then A28: f"{+infty} /\ B c= (f|B)"{+infty}; now let x be object; assume A29: x in f"{-infty} /\ B; 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; end; then A33: f"{-infty} /\ B c= (f|B)"{-infty}; now let x be object; assume A34: x in (f|B)"{-infty}; 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; end; then (f|B)"{-infty} c= f"{-infty} /\ B; then (f|B)"{-infty}=f"{-infty} /\B by A33; 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 let x be object; assume A40: x in (f|B)"{+infty}; 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; end; then (f|B)"{+infty} c= f"{+infty} /\ B; then (f|B)"{+infty}=f"{+infty} /\ B by A28; 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; dom(g|B) = dom g /\ B by RELAT_1:61; then A48: dom(g|B) = B by A47,XBOOLE_1:18,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,28; now let x be object; assume A52: x in dom ((f+g)|B); 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; end; hence thesis by A1,A49,A46,FUNCT_1:2,XBOOLE_1:28; end; theorem Th30: for X be non empty set, f be PartFunc of X,ExtREAL, a be R_eal holds eq_dom(f,a) = f"{a} proof let X be non empty set; let f be PartFunc of X,ExtREAL; let a be R_eal; now let x be object; assume A1: x in 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; end; then A3: f"{a} c= eq_dom(f,a); now let x be object; assume A4: x in eq_dom(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; end; then eq_dom(f,a) c= f"{a}; hence thesis by A3; end; begin :: Lemmas for measurable function and simple valued function theorem Th31: for X be non empty set, S be SigmaField of X, f,g be PartFunc of X,ExtREAL, A be Element of S st f is without-infty & g is without-infty & f is_measurable_on A & g is_measurable_on A holds f+g is_measurable_on A proof let X be non empty set, S be SigmaField of X, f,g be PartFunc of X,ExtREAL, A be Element of S; assume that A1: f is without-infty and A2: g is without-infty and A3: f is_measurable_on A and A4: g is_measurable_on A; for r be Real holds A /\ less_dom(f+g, r) in S proof let r be Real; consider F being Function of RAT,S such that A5: for p being Rational holds F.p = A /\ less_dom(f, p) /\ (A /\ less_dom(g,(r-(p qua Complex)))) by A3,A4,MESFUNC2:6; ex G be sequence of 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) = union rng F by A1,A2,A5,Th18; hence thesis by A6,MEASURE1:def 5; end; hence thesis by MESFUNC1:def 16; end; theorem for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL st f is_simple_func_in S & dom f = {} holds ex F be Finite_Sep_Sequence of S, a,x be FinSequence of ExtREAL st F,a are_Re-presentation_of f & a.1 = 0 & (for n be Nat st 2 <= n & n in dom a holds 0 < a.n & a.n < +infty) & dom x = dom F & (for n be Nat st n in dom x holds x.n = a.n*(M*F).n) & Sum x = 0 proof 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: f is_simple_func_in S and A2: dom f = {}; for x be object st x in dom f holds 0 <= f.x by A2; then f is nonnegative by SUPINF_2:52; then consider F be Finite_Sep_Sequence of S, a be FinSequence of ExtREAL such that A3: F,a are_Re-presentation_of f and A4: a.1 = 0 and A5: for n be Nat st 2 <= n & n in dom a holds 0 < a.n & a.n < +infty by A1, MESFUNC3:14; deffunc F(Nat) = a.$1*(M*F).$1; consider x be FinSequence of ExtREAL such that A6: len x = len F and A7: for n be Nat st n in dom x holds x.n = F(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,a,x; consider sumx be sequence of ExtREAL such that A10: Sum x = sumx.(len x) and A11: sumx.0 = 0 and A12: for i be Nat st i < len x holds sumx.(i+1)=sumx.i+x.(i+1 ) by EXTREAL1:def 2; defpred P[Nat] means $1 <= len x implies sumx.$1 = 0; A13: union rng F = {} by A2,A3,MESFUNC3:def 1; A14: for n be Nat st n in dom F holds F.n = {} proof let n be Nat; assume n in dom F; then A15: F.n in rng F by FUNCT_1:3; assume F.n <> {}; then ex v be object st v in F.n by XBOOLE_0:def 1; hence contradiction by A13,A15,TARSKI:def 4; end; A16: for i be Nat st P[i] holds P[i+1] proof let i be Nat; assume A17: P[i]; assume A18: i+1 <= len x; reconsider 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 thesis by A17,A18,A19,NAT_1:13; end; A22: P[ 0 ] by A11; for i be Nat holds P[i] from NAT_1:sch 2(A22,A16); hence thesis by A3,A4,A5,A7,A8,A10,FINSEQ_1:def 3; end; theorem Th33: for X be non empty set, S be SigmaField of X, f be PartFunc of X ,ExtREAL, A be Element of S, r,s be Real st f is_measurable_on A & A c= dom f holds A /\ great_eq_dom(f, r) /\ less_dom(f, s) in S proof let X be non empty set; let S be SigmaField of X; let f be PartFunc of X,ExtREAL; let A be Element of S; let r,s be Real; assume that A1: f is_measurable_on A and A2: A c= dom f; A3: A /\ less_dom(f, s) in S by A1,MESFUNC1:def 16; A4: A /\ great_eq_dom(f, r) /\ (A /\ less_dom(f, s)) = A /\ great_eq_dom(f, r) /\ A /\ less_dom(f, s) by XBOOLE_1:16 .= great_eq_dom(f, r) /\ (A/\A) /\ less_dom(f, s) by XBOOLE_1:16; A /\ great_eq_dom(f, r) in S by A1,A2,MESFUNC1:27; hence thesis by A3,A4,FINSUB_1:def 2; end; theorem Th34: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A be 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; let S be SigmaField of X; let M be sigma_Measure of S; let f be PartFunc of X,ExtREAL; let A be Element of S; assume A1: f is_simple_func_in S; then consider F be Finite_Sep_Sequence of S such that A2: dom f = union rng F and A3: for n be Nat, 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 FA(Nat) = F.$1 /\ A; consider G be FinSequence such that A4: len G = len F & for n be Nat st n in dom G holds G.n = FA(n) from FINSEQ_1:sch 2; A5: rng G c= S proof let P be object; assume P in rng G; then consider k be 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 thesis by A7,A8,FINSUB_1:def 2; end; A9: dom G = Seg len F by A4,FINSEQ_1:def 3; reconsider G as FinSequence of S by A5,FINSEQ_1:def 4; for i,j be Nat st i in dom G & j in dom G & i <> j holds G.i misses G.j proof let i,j be Nat; assume that A10: i in dom G and A11: j in dom G and A12: i <> 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; hence thesis; end; then reconsider G as Finite_Sep_Sequence of S by MESFUNC3:4; for v be object st v in union rng G holds v in dom(f|A) proof let v be object; assume v in union rng G; then consider W be set such that A16: v in W and A17: W in rng G by TARSKI:def 4; consider k be 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 thesis by RELAT_1:61; end; then A23: union rng G c= dom(f|A); for v be object st v in dom(f|A) holds v in union rng G proof let v be object; assume v in dom(f|A); 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 be set such that A26: v in W and A27: W in rng F by A2,TARSKI:def 4; consider k be 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 thesis by A31,TARSKI:def 4; end; then dom(f|A) c= union rng G; then A32: dom(f|A) = union rng G by A23; A33: for n be Nat, x,y be 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; let x,y be Element of X; assume that A34: n in dom G and A35: x in G.n and A36: y in G.n; 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 thesis by A42,FUNCT_1:47; end; f is real-valued by A1,MESFUNC2:def 4; hence thesis by A32,A33,MESFUNC2:def 4; end; theorem Th35: for X be non empty set, S be SigmaField of X, A be Element of S, F be Finite_Sep_Sequence of S, G be FinSequence st dom F = dom G & (for n be 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; let S be SigmaField of X; let A be Element of S; let F be Finite_Sep_Sequence of S; let G be FinSequence; assume that A1: dom F = dom G and A2: for n be Nat st n in dom F holds G.n = F.n /\ A; rng G c= S proof let v be object; assume v in rng G; then consider k be object 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 thesis by A4,A5,FINSUB_1:def 2; end; then reconsider G as FinSequence of S by FINSEQ_1:def 4; now let i,j be Nat; assume that A6: i in dom G and A7: j in dom G and A8: i <> 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; end; hence thesis by MESFUNC3:4; end; theorem Th36: for X be non empty set, S be SigmaField of X, f be PartFunc of X ,ExtREAL, A be Element of S, F,G be Finite_Sep_Sequence of S, a be FinSequence of ExtREAL st dom F = dom G & (for n be 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; let S be SigmaField of X; let f be PartFunc of X,ExtREAL; let A be Element of S; let F,G be Finite_Sep_Sequence of S; let a be FinSequence of ExtREAL; assume that A1: dom F = dom G and A2: for n be Nat st n in dom F holds G.n = F.n /\ A and A3: F,a are_Re-presentation_of f; A4: dom G = dom a by A1,A3,MESFUNC3:def 1; now let v be object; assume v in union rng G; then consider C be set such that A5: v in C and A6: C in rng G by TARSKI:def 4; consider k be object 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; end; then A12: union rng G c= dom(f|A); A13: for k be Nat st k in dom G for x be object st x in G.k holds (f|A).x = a.k proof A14: for k be Nat st k in dom G for x be set st x in G.k holds f.x = a.k proof let k be Nat; assume A15: k in dom G; let x be set; assume x in G.k; then x in F.k /\ A by A1,A2,A15; then x in F.k by XBOOLE_0:def 4; hence thesis by A1,A3,A15,MESFUNC3:def 1; end; let k be Nat; assume A16: k in dom G; let x be object; assume A17: x in G.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 thesis by A16,A17,A14; end; now let v be object; assume v in dom(f|A); 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 be set such that A19: v in C and A20: C in rng F by TARSKI:def 4; consider k be 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; end; then dom(f|A) c= union rng G; then dom(f|A) = union rng G by A12; hence thesis by A4,A13,MESFUNC3:def 1; end; theorem Th37: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be 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; let S be SigmaField of X; let M be sigma_Measure of S; let f be PartFunc of X,ExtREAL; assume f is_simple_func_in S; then ex F be Finite_Sep_Sequence of S st dom f = union rng F & for n being Nat,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 thesis by MESFUNC2:31; end; Lm2: for Y being set, p be FinSequence holds (for i be Nat st i in dom p holds p.i in Y) implies p is FinSequence of Y proof let Y be set; let p be FinSequence; assume A1: for i be Nat st i in dom p holds p.i in Y; let b be object; assume b in rng p; then ex i be Nat st i in dom p & p.i = b by FINSEQ_2:10; hence thesis by A1; end; Lm3: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f, g be 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, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL such 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; consider F be Finite_Sep_Sequence of S, a be 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 be Finite_Sep_Sequence of S, b be FinSequence of ExtREAL such that A7: G,b are_Re-presentation_of g by A3,MESFUNC3:12; set lb = len G; deffunc FG1(Nat) = F.(($1-'1) div lb + 1) /\ G.(($1-'1) mod lb + 1); consider FG be FinSequence such that A8: len FG = la*lb and A9: for k be Nat st k in dom FG holds FG.k=FG1(k) from FINSEQ_1:sch 2; A10: dom FG = Seg(la*lb) by A8,FINSEQ_1:def 3; now reconsider lb9=lb as Nat; let k be Nat; set i=(k-'1) div lb + 1; set j=(k-'1) mod lb + 1; A11: lb9 divides (la*lb) by NAT_D:def 3; assume A12: k in dom FG; then A13: k in Seg(la*lb) by A8,FINSEQ_1:def 3; then A14: k <= la*lb by FINSEQ_1:1; then (k -' 1) <= (la*lb -' 1) by NAT_D:42; then A15: (k -' 1) div lb <= (la*lb -' 1) div lb by NAT_2:24; 1 <= k by A13,FINSEQ_1:1; then A16: 1 <= la*lb by A14,XXREAL_0:2; A17: lb <> 0 by A13; then (k -' 1) mod lb < lb by NAT_D:1; then A18: j <= lb by NAT_1:13; lb >= 0+1 by A17,NAT_1:13; then ((la*lb) -' 1) div lb = ((la*lb) div lb) - 1 by A11,A16,NAT_2:15; then (k -' 1) div lb + 1 <= la*lb div lb by A15,XREAL_1:19; then A19: i <= la by A17,NAT_D:18; i >= 0+1 by XREAL_1:6; then i in Seg la by A19; then i in dom F by FINSEQ_1:def 3; then A20: F.i in rng F by FUNCT_1:3; j >= 0+1 by XREAL_1:6; then j in dom G by A18,FINSEQ_3:25; then A21: G.j in rng G by FUNCT_1:3; FG.k = F.((k-'1) div lb + 1) /\ G.((k-'1) mod lb + 1) by A9,A12; hence FG.k in S by A20,A21,MEASURE1:34; end; then reconsider FG as FinSequence of S by Lm2; A22: for k,l be Nat st k in dom FG & l in dom FG & k <> l holds FG.k misses FG. l proof A23: lb divides (la *lb) by NAT_D:def 3; let k,l be Nat; assume that A24: k in dom FG and A25: l in dom FG and A26: k <> l; A27: k in Seg(la*lb) by A8,A24,FINSEQ_1:def 3; then A28: 1 <= k by FINSEQ_1:1; set m=(l-'1) mod lb + 1; set n=(l-'1) div lb + 1; set j=(k-'1) mod lb + 1; set i=(k-'1) div lb + 1; A29: FG.k = F.i /\ G.j by A9,A24; A30: k <= la*lb by A27,FINSEQ_1:1; then A31: 1 <= la*lb by A28,XXREAL_0:2; A32: lb <> 0 by A27; then lb >= 0+1 by NAT_1:13; then A33: ((la*lb) -' 1) div lb = ((la*lb) div lb) - 1 by A23,A31,NAT_2:15; A34: l in Seg (la*lb) by A8,A25,FINSEQ_1:def 3; then A35: 1 <= l by FINSEQ_1:1; A36: now (l-'1)+1=(n-1)*lb+(m-1)+1 by A32,NAT_D:2; then A37: l - 1 + 1 = (n-1)*lb+m by A35,XREAL_1:233; assume that A38: i=n and A39: j=m; (k-'1)+1=(i-1)*lb+(j-1)+1 by A32,NAT_D:2; then k - 1 + 1 = (i-1)*lb + j by A28,XREAL_1:233; hence contradiction by A26,A38,A39,A37; end; (k -' 1) <= (la*lb -' 1) by A30,NAT_D:42; then (k -' 1) div lb <= (la*lb div lb) - 1 by A33,NAT_2:24; then (k -' 1) div lb + 1 <= la*lb div lb by XREAL_1:19; then A40: i <= la by A32,NAT_D:18; i >= 0+1 by XREAL_1:6; then i in Seg la by A40; then A41: i in dom F by FINSEQ_1:def 3; A42: j >= 0+1 by XREAL_1:6; (k -' 1) mod lb < lb by A32,NAT_D:1; then j <= lb by NAT_1:13; then A43: j in dom G by A42,FINSEQ_3:25; A44: m >= 0+1 by XREAL_1:6; (l -' 1) mod lb < lb by A32,NAT_D:1; then m <= lb by NAT_1:13; then A45: m in dom G by A44,FINSEQ_3:25; A46: n >= 0+1 by XREAL_1:6; l <= la*lb by A34,FINSEQ_1:1; then (l -' 1) <= (la*lb -' 1) by NAT_D:42; then (l -' 1) div lb <= (la*lb div lb) - 1 by A33,NAT_2:24; then (l -' 1) div lb + 1 <= la*lb div lb by XREAL_1:19; then n <= la by A32,NAT_D:18; then n in Seg la by A46; then A47: n in dom F by FINSEQ_1:def 3; per cases by A36; suppose A48: i <> n; FG.k /\ FG.l= F.i /\ G.j /\ (F.n /\ G.m) by A9,A25,A29; then FG.k /\ FG.l= F.i /\ (G.j /\ (F.n /\ G.m)) by XBOOLE_1:16; then FG.k /\ FG.l= F.i /\ (F.n /\ (G.j /\ G.m)) by XBOOLE_1:16; then A49: FG.k /\ FG.l= F.i /\ F.n /\ (G.j /\ G.m) by XBOOLE_1:16; F.i misses F.n by A41,A47,A48,MESFUNC3:4; then FG.k /\ FG.l= {} /\ (G.j /\ G.m) by A49; hence thesis; end; suppose A50: j <> m; FG.k /\ FG.l= F.i /\ G.j /\ (F.n /\ G.m) by A9,A25,A29; then FG.k /\ FG.l= F.i /\ (G.j /\ (F.n /\ G.m)) by XBOOLE_1:16; then FG.k /\ FG.l= F.i /\ (F.n /\ (G.j /\ G.m)) by XBOOLE_1:16; then A51: FG.k /\ FG.l= F.i /\ F.n /\ (G.j /\ G.m) by XBOOLE_1:16; G.j misses G.m by A43,A45,A50,MESFUNC3:4; then FG.k /\ FG.l= F.i /\ F.n /\ {} by A51; hence thesis; end; end; A52: g is real-valued by A3,MESFUNC2:def 4; then A53: dom(f+g) = dom f /\ dom g by MESFUNC2:2; reconsider 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 let z be object; assume A56: z in dom f; then consider Y be set such that A57: z in Y and A58: Y in rng F by A6,TARSKI:def 4; consider i be 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 qua Complex) + i9 by NAT_1:10; consider Z be set such that A63: z in Z and A64: Z in rng G by A4,A54,A56,TARSKI:def 4; consider j be 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 qua Complex) + j9 by NAT_1:10; i9*lb + j in NAT by ORDINAL1:def 12; then reconsider k=(i-1)*lb+j as Element of NAT by A62; i <= la by A61,FINSEQ_1:1; then i-1 <= la-1 by XREAL_1:9; then (i-1)*lb <= (la - 1)*lb by XREAL_1:64; then A70: k <= (la - 1) * lb + j by XREAL_1:6; A71: j <= lb by A67,FINSEQ_1:1; then A72: j9 < lb 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*lb + j9 by A62,A69; then A75: i = (k-'1) div lb +1 by A62,A72,NAT_D:def 1; (la - 1) * lb + j <= (la - 1) * lb + lb by A71,XREAL_1:6; then A76: k <= la*lb by A70,XXREAL_0:2; k >= 1 by A68,A73,XXREAL_0:2; then A77: k in Seg (la*lb) 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 lb +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; end; hence dom f c= union rng FG; reconsider lb9=lb as Nat; let z be object; A80: lb9 divides (la*lb) by NAT_D:def 3; assume z in union rng FG; then consider Y be set such that A81: z in Y and A82: Y in rng FG by TARSKI:def 4; consider k be 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 <= la*lb by A8,FINSEQ_1:1; then A87: (k -' 1) <= (la*lb -' 1) by NAT_D:42; set j=(k-'1) mod lb +1; set i=(k-'1) div lb +1; A88: i >= 0+1 by NAT_1:13; 1 <= k by A85,FINSEQ_1:1; then A89: 1 <= la*lb by A86,XXREAL_0:2; A90: lb <> 0 by A8,A85; then lb >= 0+1 by NAT_1:13; then ((la*lb) -' 1) div lb9 = ((la*lb) div lb) - 1 by A80,A89,NAT_2:15; then (k -' 1) div lb <= (la*lb div lb) - 1 by A87,NAT_2:24; then A91: i <= la*lb div lb by XREAL_1:19; la*lb div lb = la by A90,NAT_D:18; then i in Seg la by A91,A88; then i in dom F by FINSEQ_1:def 3; then A92: F.i in rng F by FUNCT_1:def 3; FG.k=F.i /\ G.j by A9,A83; then z in F.i by A81,A84,XBOOLE_0:def 4; hence thesis by A6,A92,TARSKI:def 4; end; A93: for k being Nat,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: lb divides (la*lb) by NAT_D:def 3; let k be Nat; let x,y be Element of X; assume that A95: k in dom FG and A96: x in FG.k and A97: y in FG.k; set j=(k-'1) mod lb + 1; A98: FG.k = F.( (k-'1) div lb + 1 ) /\ G.( (k-'1) mod lb + 1 ) by A9,A95; then A99: y in G.j by A97,XBOOLE_0:def 4; set i=(k-'1) div lb + 1; A100: i >= 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: lb > 0 by A8,A101; then A104: lb >= 0+1 by NAT_1:13; A105: k <= la*lb by A8,A101,FINSEQ_1:1; then A106: (k -' 1) <= (la*lb -' 1) by NAT_D:42; 1 <= la*lb by A102,A105,XXREAL_0:2; then ((la*lb) -' 1) div lb = ((la*lb) div lb) - 1 by A104,A94,NAT_2:15; then (k -' 1) div lb <= (la*lb div lb) - 1 by A106,NAT_2:24; then A107: (k -' 1) div lb + 1 <= la*lb div lb by XREAL_1:19; lb <> 0 by A8,A101; then i <= la by A107,NAT_D:18; then i in Seg la by A100; then A108: i in dom F by FINSEQ_1:def 3; x in F.i by A96,A98,XBOOLE_0:def 4; then A109: f.x=a.i by A5,A108,MESFUNC3:def 1; A110: j >= 0+1 by XREAL_1:6; (k -' 1) mod lb < lb by A103,NAT_D:1; then j <= lb by NAT_1:13; then j in Seg lb by A110; then A111: j in dom G by FINSEQ_1:def 3; y in F.i by A97,A98,XBOOLE_0:def 4; then A112: f.y=a.i 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.j by A96,A98,XBOOLE_0:def 4; then (f+g).x= a.i+b.j 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 thesis by A115,MESFUNC1:def 3; end; now let x be Element of X; assume A116: x in dom(f+g); 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 EXTREAL1:24; f is real-valued 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,4; end; then f+g is real-valued by MESFUNC2:def 1; hence f+g is_simple_func_in S by A4,A55,A53,A93,MESFUNC2:def 4; thus thesis by A2,A4,A53; end; theorem Th38: for X be non empty set, S be SigmaField of X, f,g be 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; let S be SigmaField of X; let f,g be PartFunc of X,ExtREAL; assume that A1: f is_simple_func_in S and A2: g is_simple_func_in S; per cases; suppose A3: dom(f+g) = {}; reconsider EMPTY = {} as Element of S by PROB_1:4; set F = <*EMPTY*>; A4: dom F = Seg 1 by FINSEQ_1:38; A5: now let i,j be Nat; assume that A6: i in dom F and A7: j in dom F and A8: i <> j; i = 1 by A4,A6,FINSEQ_1:2,TARSKI:def 1; hence F.i misses F.j by A4,A7,A8,FINSEQ_1:2,TARSKI:def 1; end; A9: for n be Nat st n in dom F holds F.n = EMPTY proof let n be Nat; assume n in dom F; then n = 1 by A4,FINSEQ_1:2,TARSKI:def 1; hence thesis by FINSEQ_1:40; end; reconsider F 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 be Element of X st x in dom(f+g) holds |. (f+g).x .| < +infty by A3; then A11: f+g is real-valued by MESFUNC2:def 1; for n being Nat for x,y be 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 thesis by A11,A10,MESFUNC2:def 4; end; suppose A12: dom(f+g) <> {}; A13: (f|dom(f+g))"{+infty} = dom(f+g) /\ f"{+infty} by FUNCT_1:70; g is without+infty by A2,Th14; then not +infty in rng g; 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 without+infty by A1,Th14; then not +infty in rng f; 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 be 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; assume A28: x in dom(f|dom(f+g) + g|dom(f+g)); 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 thesis by A26,A28,MESFUNC1:def 3; 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 thesis by A26,A27,PARTFUN1:5; end; end; theorem Th39: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL,c be Real st f is_simple_func_in S holds c(#)f is_simple_func_in S proof 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 c be Real; set g = c(#)f; assume A1: f is_simple_func_in S; then consider G be Finite_Sep_Sequence of S such that A2: dom f = union rng G and A3: for n be Nat, x,y be 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 real-valued by A1,MESFUNC2:def 4; now let x be Element of X; assume A5: x in dom g; c * f.x <> -infty by A4; then g.x <> -infty by A5,MESFUNC1:def 6; then -infty < g.x by XXREAL_0:6; then A6: -(+infty) < g.x by XXREAL_3:def 3; c * f.x <> +infty by A4; then g.x <> +infty by A5,MESFUNC1:def 6; then g.x < +infty by XXREAL_0:4; hence |. g.x .| < +infty by A6,EXTREAL1:22; end; then A7: g is real-valued by MESFUNC2:def 1; A8: dom g = dom f by MESFUNC1:def 6; now let n be Nat; let x,y be Element of X; assume that A9: n in dom G and A10: x in G.n and A11: y in G.n; A12: G.n in rng G by A9,FUNCT_1:3; then y in dom g by A8,A2,A11,TARSKI:def 4; then A13: g.y = ( c)*f.y by MESFUNC1:def 6; x in dom g by A8,A2,A10,A12,TARSKI:def 4; then g.x = ( c)*f.x by MESFUNC1:def 6; hence g.x = g.y by A3,A9,A10,A11,A13; end; hence thesis by A8,A7,A2,MESFUNC2:def 4; end; theorem Th40: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL st f is_simple_func_in S & g is_simple_func_in S & (for x be object st x in dom(f-g) holds g.x <= f.x) holds f-g is nonnegative proof let X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL such that A1: f is_simple_func_in S and A2: g is_simple_func_in S and A3: for x be object st x in dom(f-g) holds g.x <= f.x; g is without-infty by A2,Th14; then not -infty in rng g; then A4: g"{-infty} = {} by FUNCT_1:72; f is without+infty by A1,Th14; then not +infty in rng f; 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 be set st x in dom f /\ dom g holds g.x <= f.x & -infty < g.x & f .x < +infty proof let x be set; assume A7: x in dom f /\ dom g; hence g.x <= f.x by A3,A6; 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; 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 thesis by XXREAL_0:4; end; hence thesis by Th21; end; theorem Th41: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, A be Element of S, c be R_eal st c <> +infty & c <> -infty holds ex f be PartFunc of X,ExtREAL st f is_simple_func_in S & dom f = A & for x be object st x in A holds f.x=c proof let X be non empty set; let S be SigmaField of X; let M be sigma_Measure of S; let A be Element of S; let c be R_eal; assume that A1: c <> +infty and A2: c <> -infty; -infty < c by A2,XXREAL_0:6; then A3: -(+infty) < c by XXREAL_3:def 3; deffunc F(object) = c; defpred P[object] means $1 in A; A4: for x be object st P[x] holds F(x) in ExtREAL; consider f be PartFunc of X,ExtREAL such that A5: (for x be object holds x in dom f iff x in X & P[x]) & for x be object st x in dom f holds f.x = F(x) from PARTFUN1:sch 3(A4); c < +infty by A1,XXREAL_0:4; then |.c.| < +infty by A3,EXTREAL1:22; then for x be Element of X st x in dom f holds |. f.x .| < +infty by A5; then A6: f is real-valued by MESFUNC2:def 1; take f; A7: A c= dom f by A5; set F = <* dom f *>; A8: dom f c= A by A5; A9: rng F = {dom f} by FINSEQ_1:38; then A10: rng F = {A} by A8,A7,XBOOLE_0:def 10; rng F c= S proof let z be object; assume z in rng F; then z = A by A10,TARSKI:def 1; hence thesis; end; then reconsider F as FinSequence of S by FINSEQ_1:def 4; now let i,j be Nat; assume that A11: i in dom F and A12: j in dom F and A13: i <> 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; end; then reconsider F as Finite_Sep_Sequence of S by MESFUNC3:4; A15: now let n be Nat; let x,y be Element of X; assume that A16: n in dom F and A17: x in F.n and A18: y in F.n; 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; end; dom f = union rng F by A9,ZFMISC_1:25; hence f is_simple_func_in S by A6,A15,MESFUNC2:def 4; thus dom f = A by A8,A7; thus thesis by A5; end; theorem Th42: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, B,BF be 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, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, B,BF be Element of S such that A1: f is_measurable_on B and A2: BF = dom f /\ B; now let r be Real; A3: now let x be object; reconsider xx=x as set by TARSKI:1; (x in dom(f|B) & ex y being R_eal st y=f|B.x & y < r) iff x in dom f /\ B & ex y being R_eal st y=f|B.x & y < r by RELAT_1:61; then A4: x in BF & x in less_dom(f|B, r) iff x in B & x in dom f & (f|B) .xx < r by A2,MESFUNC1:def 11,XBOOLE_0:def 4; x in B & x in dom f implies (f.x < r iff (f|B).x < r) by FUNCT_1:49; then x in BF /\ less_dom(f|B, r) iff x in B & x in less_dom(f, r) by A4,MESFUNC1:def 11,XBOOLE_0:def 4; hence x in BF /\ less_dom(f|B, r) iff x in B /\ less_dom(f, r) by XBOOLE_0:def 4; end; then A5: B /\ less_dom(f, r) c= BF /\ less_dom(f|B, r); BF /\ less_dom(f|B, r) c= B /\ less_dom(f, r) by A3; then BF /\ less_dom(f|B, r) = B /\ less_dom(f, r) by A5; hence BF /\ less_dom(f|B, r) in S by A1,MESFUNC1:def 16; end; hence thesis by MESFUNC1:def 16; end; theorem Th43: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, A be Element of S, f,g being PartFunc of X,ExtREAL st A c= dom f & f is_measurable_on A & g is_measurable_on A & f is without-infty & g is without-infty holds max+(f+g) + max-f is_measurable_on A proof let X be non empty set, S be SigmaField of X, M be sigma_Measure of S, A be Element of S, f,g be PartFunc of X,ExtREAL; 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 without-infty and A5: g is without-infty; 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 thesis by A6,A8,A7,Th31; end; theorem Th44: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, A be Element of S, 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 without-infty & g is without-infty holds max-(f+g) + max+f is_measurable_on A proof let X be non empty set, S be SigmaField of X, M be sigma_Measure of S, A be Element of S, f,g be PartFunc of X,ExtREAL; 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 without-infty and A5: g is without-infty; 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 thesis by A7,A8,A9,Th31; end; theorem Th45: for X be non empty set, S being SigmaField of X, M being sigma_Measure of S, A being set st A in S holds 0 <= M.A proof let X be non empty set; let S be SigmaField of X; let M be sigma_Measure of S; let A be set; reconsider E = {} as Element of S by PROB_1:4; assume A in S; then reconsider A as Element of S; M.E <= M.A by MEASURE1:31,XBOOLE_1:2; hence thesis by VALUED_0:def 19; end; Lm4: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, r be Real st dom f in S & (for x be object st x in dom f holds f.x = r) holds f is_simple_func_in S proof 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 r be Real; assume that A1: dom f in S and A2: for x be object st x in dom f holds f.x = r; reconsider Df = dom f as Element of S by A1; set F = <*Df*>; A3: dom F = Seg 1 by FINSEQ_1:38; A4: now let i,j be Nat; assume that A5: i in dom F and A6: j in dom F and A7: i <> j; i = 1 by A3,A5,FINSEQ_1:2,TARSKI:def 1; hence F.i misses F.j by A3,A6,A7,FINSEQ_1:2,TARSKI:def 1; end; A8: for n be Nat st n in dom F holds F.n = Df proof let n be Nat; assume n in dom F; then n = 1 by A3,FINSEQ_1:2,TARSKI:def 1; hence thesis by FINSEQ_1:40; end; reconsider F as Finite_Sep_Sequence of S by A4,MESFUNC3:4; A9: now let n be Nat; let x,y be Element of X; assume that A10: n in dom F and A11: x in F.n and A12: y in F.n; A13: F.n = Df by A8,A10; then f.x = r by A2,A11; hence f.x = f.y by A2,A12,A13; end; F = <* F.1 *> by FINSEQ_1:40; then A14: rng F = {F.1} by FINSEQ_1:38; A15: r in REAL by XREAL_0:def 1; now let x be Element of X; assume x in dom f; then A16: f.x = r by A2; then -infty < f.x by XXREAL_0:12,A15; then A17: -(+infty) < f.x by XXREAL_3:def 3; f.x < +infty by A16,XXREAL_0:9,A15; hence |. f.x .| < +infty by A17,EXTREAL1:22; end; then A18: f is real-valued 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 thesis by A18,A9,MESFUNC2:def 4; end; theorem Th46: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL st (ex E1 be Element of S st E1=dom f & f is_measurable_on E1) & (ex E2 be 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, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL; assume that A1: ex E1 be Element of S st E1=dom f & f is_measurable_on E1 and A2: ex E2 be 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; 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 be Element of S such that A9: E2=dom g and g is_measurable_on E2 by A2; consider E1 be 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 thesis by A8,A11,MEASURE1:34; end; Lm5: for X be non empty set, S be SigmaField of X, A be Element of S, f be PartFunc of X,ExtREAL, r be Real holds A /\ less_dom(f, r) = less_dom(f|A, r) proof let X be non empty set; let S be SigmaField of X; let A be Element of S; let f be PartFunc of X,ExtREAL; let r be Real; now let v be object; assume A1: v in A /\ less_dom(f, r); then A2: v in less_dom(f, 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 by A2,MESFUNC1:def 11; hence v in less_dom(f|A, r) by A5,A4,MESFUNC1:def 11; end; hence A /\ less_dom(f, r) c= less_dom(f|A, r); let v be object; reconsider vv=v as set by TARSKI:1; assume A6: v in less_dom(f|A, 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).vv < r by A6,MESFUNC1:def 11; then ex w be R_eal st w=f.vv & w< r by A7,FUNCT_1:47; then A10: v in less_dom(f, r) by A9,MESFUNC1:def 11; v in A by A8,XBOOLE_0:def 4; hence thesis by A10,XBOOLE_0:def 4; end; Lm6: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, A be Element of S, f be PartFunc of X,ExtREAL holds f|A is_measurable_on A iff f is_measurable_on A proof let X be non empty set; let S be SigmaField of X; let M be sigma_Measure of S; let A be Element of S; let f be PartFunc of X,ExtREAL; now assume A1: f|A is_measurable_on A; now let r be Real; A /\ less_dom(f|A, r) in S by A1,MESFUNC1:def 16; then A /\ (A /\ less_dom(f, r)) in S by Lm5; then A /\ A /\ less_dom(f, r) in S by XBOOLE_1:16; hence A /\ less_dom(f, r) in S; end; hence f is_measurable_on A by MESFUNC1:def 16; end; hence f|A is_measurable_on A implies f is_measurable_on A; assume A2: f is_measurable_on A; now let r be Real; A/\A /\ less_dom(f, r) in S by A2,MESFUNC1:def 16; then A /\ (A /\ less_dom(f, r)) in S by XBOOLE_1:16; hence A /\ less_dom(f|A, r) in S by Lm5; end; hence thesis by MESFUNC1:def 16; end; Lm7: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f, g be PartFunc of X,ExtREAL st (ex E1 be Element of S st E1=dom f & f is_measurable_on E1) & (ex E2 be Element of S st E2=dom g & g is_measurable_on E2) & dom f = dom g holds ex DFPG be Element of S st DFPG=dom(f+g) & f+g is_measurable_on DFPG proof let X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL such that A1: ex E1 be Element of S st E1=dom f & f is_measurable_on E1 and A2: ex E2 be Element of S st E2=dom g & g is_measurable_on E2 and A3: dom f = dom g; consider Gf be Element of S such that A4: Gf=dom g and A5: g is_measurable_on Gf by A2; now let v be object; assume A6: v in 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; end; then A8: g"{-infty} c= Gf /\ eq_dom(g,-infty); now let v be object; assume v in Gf /\ eq_dom(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; end; then A11: Gf /\ eq_dom(g,-infty) c= g"{-infty}; 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 let v be object; assume A13: v in 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; end; then A15: g"{+infty} c= Gf /\ eq_dom(g,+infty); now let v be object; assume v in Gf /\ eq_dom(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; end; then A18: Gf /\ eq_dom(g,+infty) c= g"{+infty}; 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 be 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 let v be object; assume v in E0 /\ eq_dom(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; end; then A28: E0 /\ eq_dom(f,+infty) c= f"{+infty}; now let v be object; assume A29: v in 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; end; then f"{+infty} c= E0 /\ eq_dom(f,+infty); then A31: f"{+infty} = E0 /\ eq_dom(f,+infty) by A28; now let v be object; assume v in E0 /\ eq_dom(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; end; then A34: E0 /\ eq_dom(f,-infty) c= f"{-infty}; now let v be object; assume A35: v in 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; end; then A37: f"{-infty} c= E0 /\ eq_dom(f,-infty); then A38: f"{-infty} = E0 /\ eq_dom(f,-infty) by A34; 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,36; g"{-infty} c= NG by XBOOLE_1:7; then A42: f"{+infty}/\g"{-infty} c= NG by A19; f"{-infty} c= NF by XBOOLE_1:7; then f"{-infty}/\g"{+infty} c= NF by A20; 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 = f1+g1 by Th29; A45: dom(f1+g1)=E by A43,Th29; A46: E =dom(f1+g1) by A43,Th29; A47: for r being Real holds DFPG /\ less_dom(f+g, r) = E/\ less_dom(f1+g1, r) \/ f"{-infty}/\(DFPG \ g"{+infty}) \/ g"{-infty}/\(DFPG \ f"{+infty}) proof let r be Real; set SL= DFPG/\less_dom(f+g, r); set SR= E/\less_dom(f1+g1, r) \/ f"{-infty}/\(DFPG \ g"{+infty}) \/ g "{-infty}/\(DFPG \ f"{+infty}); A48: now let x be object; reconsider xx=x as set by TARSKI:1; assume x in SR; then A49: x in E /\ less_dom(f1+g1, r) \/ f"{-infty} /\ (DFPG\ g"{ +infty} ) or x in g"{-infty} /\ (DFPG \ f"{+infty}) by XBOOLE_0:def 3; per cases by A49,XBOOLE_0:def 3; suppose A50: x in E /\ less_dom(f1+g1, r); then A51: x in E by XBOOLE_0:def 4; x in less_dom(f1+g1, r) by A50,XBOOLE_0:def 4; then (f1+g1).xx < r by MESFUNC1:def 11; then (f+g).xx < r by A44,A45,A51,FUNCT_1:47; then x in less_dom(f+g, r) by A43,A51,MESFUNC1:def 11; hence x in SL by A43,A51,XBOOLE_0:def 4; end; suppose A52: x in f"{-infty} /\ (DFPG \ g"{+infty}) or x in g"{-infty} /\ (DFPG \ f"{+infty}); per cases by A52; suppose A53: x in f"{-infty} /\ (DFPG \ g"{+infty}); r in REAL by XREAL_0:def 1; then A54: -infty < 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.xx + g.xx = -infty by A53,A55,A58,XBOOLE_0:def 4,def 5 ,XXREAL_3:def 2; then (f+g).xx < r by A56,A54,MESFUNC1:def 3; then x in less_dom(f+g, r) by A56,MESFUNC1:def 11; hence x in DFPG /\ less_dom(f+g, r) by A56,XBOOLE_0:def 4; end; suppose A59: x in g"{-infty}/\(DFPG \ f"{+infty}); r in REAL by XREAL_0:def 1; then A60: -infty < 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.xx + g.xx = -infty by A59,A61,A65,XBOOLE_0:def 4,def 5 ,XXREAL_3:def 2; then (f+g).xx < r by A62,A60,MESFUNC1:def 3; then x in less_dom(f+g, r) by A62,MESFUNC1:def 11; hence x in DFPG /\ less_dom(f+g, r) by A63,XBOOLE_0:def 4; end; end; end; now let x be object; reconsider xx=x as set by TARSKI:1; assume A66: x in SL; 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) by A66,XBOOLE_0:def 4; then A71: (f+g).xx < r by MESFUNC1:def 11; then A72: f.xx + g.xx < 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; per cases; suppose A83: f.x =-infty; then x in DFPG \ g"{+infty} by A67,A76,A81,A73,TARSKI:def 1 ,XBOOLE_0:def 4,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(f1+g1, r) \/ f"{-infty} /\ (DFPG \ g"{ +infty}) by XBOOLE_0:def 3; hence x in SR by XBOOLE_0:def 3; end; suppose A84: f.x <>-infty; per cases; suppose A85: g.x = -infty; then x in DFPG \ f"{+infty} by A67,A77,A80,A70,TARSKI:def 1 ,XBOOLE_0:def 4,def 5; then x in g"{-infty}/\(DFPG \ f"{+infty}) by A80,A85,TARSKI:def 1 ,XBOOLE_0:def 4; hence x in SR by XBOOLE_0:def 3; end; suppose A86: g.x <> -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 (f1+g1).x = (f+g).x by A44,A45,FUNCT_1:47; then x in less_dom(f1+g1, r) by A46,A71,A87,MESFUNC1:def 11; then x in E/\less_dom(f1+g1, r) by A87,XBOOLE_0:def 4; then x in E/\less_dom(f1+g1, r) \/ f"{-infty}/\(DFPG \ g"{ +infty}) by XBOOLE_0:def 3; hence x in SR by XBOOLE_0:def 3; end; end; end; hence thesis by A48,TARSKI:2; end; A88: now let x be set; for x be object st x in dom f holds f.x in ExtREAL by XXREAL_0:def 1; then reconsider ff=f as Function of dom f,ExtREAL by FUNCT_2:3; assume A89: x in dom f1; 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 assume f1.x = -infty; 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; end; now assume f1.x = +infty; 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; end; hence -infty < f1.x & f1.x < +infty by A93,XXREAL_0:4,6; end; now let x be Element of X; A96: -(+infty) = -infty by XXREAL_3:def 3; assume A97: x in dom f1; then A98: f1.x < +infty by A88; -infty < f1.x by A88,A97; hence |.f1.x.| < +infty by A98,A96,EXTREAL1:22; end; then A99: f1 is real-valued by MESFUNC2:def 1; A100: now let x be set; for x be object st x in dom g holds g.x in ExtREAL by XXREAL_0:def 1; then reconsider gg=g as Function of dom g, ExtREAL by FUNCT_2:3; assume A101: x in dom g1; 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 assume g1.x = -infty; 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; end; now assume g1.x = +infty; 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; end; hence -infty < g1.x & g1.x < +infty by A105,XXREAL_0:4,6; end; now let x be Element of X; A108: -(+infty) = -infty by XXREAL_3:def 3; assume A109: x in dom g1; then A110: g1.x < +infty by A100; -infty < g1.x by A100,A109; hence |.g1.x.| < +infty by A110,A108,EXTREAL1:22; end; then A111: g1 is real-valued by MESFUNC2:def 1; f is_measurable_on E by A1,A23,MESFUNC1:30,XBOOLE_1:36; then A112: f1 is_measurable_on E by A41,Th42; A113: dom g /\ E = E by A3,A23,XBOOLE_1:28,36; g is_measurable_on E by A2,A3,A23,MESFUNC1:30,XBOOLE_1:36; then g1 is_measurable_on E by A113,Th42; then A114: f1+g1 is_measurable_on E by A112,A99,A111,MESFUNC2:7; now let r be Real; A115: E/\less_dom(f1+g1, 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(f1+g1, r) \/ (f"{-infty}/\(DFPG \ g"{+infty}) \/ g" {-infty}/\(DFPG \ f"{+infty})) in S by A115,PROB_1:3; DFPG /\ less_dom(f+g, r) = E /\ less_dom(f1+g1, r) \/ f"{ -infty}/\(DFPG \ g"{+infty}) \/ g"{-infty}/\(DFPG \ f"{+infty}) by A47; hence DFPG /\ less_dom(f+g, r) in S by A117,XBOOLE_1:4; end; then f+g is_measurable_on DFPG by MESFUNC1:def 16; hence thesis; end; theorem Th47: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL st (ex E1 be Element of S st E1=dom f & f is_measurable_on E1) & (ex E2 be Element of S st E2=dom g & g is_measurable_on E2) holds ex E be Element of S st E=dom(f+g) & (f+g) is_measurable_on E proof let X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL; assume that A1: ex E1 be Element of S st E1=dom f & f is_measurable_on E1 and A2: ex E2 be Element of S st E2=dom g & g is_measurable_on E2; 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|E3; A7: g1"{-infty} = E3 /\ g"{-infty} by FUNCT_1:70; set f1 = f|E3; dom f1 = dom f /\ E3 by RELAT_1:61; then A8: dom f1 = E3 by A3,XBOOLE_1:17,28; g is_measurable_on E3 by A6,MESFUNC1:30,XBOOLE_1:17; then A9: g1 is_measurable_on E3 by Lm6; A10: g1"{+infty} = E3 /\ g"{+infty} by FUNCT_1:70; dom g1 = dom g /\ E3 by RELAT_1:61; then A11: dom g1 = E3 by A5,XBOOLE_1:17,28; f1"{+infty} = E3 /\ f"{+infty} by FUNCT_1:70; then A12: f1"{+infty} /\ g1"{-infty} = f"{+infty} /\ (E3 /\ (E3 /\ g"{-infty})) by A7,XBOOLE_1:16 .= f"{+infty} /\ (E3 /\ E3 /\ g"{-infty}) by XBOOLE_1:16 .= f"{+infty} /\ g"{-infty} /\ E3 by XBOOLE_1:16; A13: dom(f1+g1) = (dom f1 /\ dom g1) \ (f1"{-infty}/\g1"{+infty} \/ f1"{ +infty}/\g1"{-infty}) by MESFUNC1:def 3; f is_measurable_on E3 by A4,MESFUNC1:30,XBOOLE_1:17; then f1 is_measurable_on E3 by Lm6; then consider E be Element of S such that A14: E = dom(f1+g1) and A15: (f1+g1) is_measurable_on E by A9,A8,A11,Lm7; take E; A16: dom((f+g)|E) = dom(f+g) /\ E by RELAT_1:61; f1"{-infty} = E3 /\ f"{-infty} by FUNCT_1:70; then f1"{-infty} /\ g1"{+infty} = f"{-infty} /\ (E3 /\ (E3 /\ g"{+infty})) by A10,XBOOLE_1:16 .= f"{-infty} /\ (E3 /\ E3 /\ g"{+infty}) by XBOOLE_1:16 .= f"{-infty} /\ g"{+infty} /\ E3 by XBOOLE_1:16; then A17: f1"{-infty}/\g1"{+infty} \/ f1"{+infty}/\g1"{-infty} = E3 /\ (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 let v be Element of X; assume A20: v in dom((f+g)|E); 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= E3 by A8,A11,A14,A13,XBOOLE_1:36; (f1+g1).v = f1.v + g1.v by A14,A19,A16,A20,MESFUNC1:def 3 .= f.v + g1.v by A8,A24,A25,FUNCT_1:47; hence ((f+g)|E).v = (f1+g1).v by A11,A24,A25,A23,FUNCT_1:47; end; then (f+g)|E = f1+g1 by A14,A19,A16,PARTFUN1:5; hence thesis by A3,A5,A8,A11,A14,A15,A13,A17,A18,Lm6,XBOOLE_1:47; end; theorem Th48: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A,B be 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, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A,B be Element of S such that A1: dom f = A; A2: now let r be Real; A3: now let x be object; x in A /\ less_dom(f, r) iff x in A & x in less_dom(f, r) by XBOOLE_0:def 4; hence x in A /\ less_dom(f, r) iff x in less_dom(f, r) by A1, MESFUNC1:def 11; end; then A4: less_dom(f, r) c= A /\ less_dom(f, r); A /\ less_dom(f, r) c= less_dom(f, r) by A3; hence A /\ less_dom(f, r) = less_dom(f, r) by A4; end; hereby assume A5: f is_measurable_on B; now let r be Real; A /\ B /\ less_dom(f, r) = B /\ (A /\ less_dom(f, r)) by XBOOLE_1:16 .= B /\ less_dom(f, r) by A2; hence A/\B/\less_dom(f, r) in S by A5,MESFUNC1:def 16; end; hence f is_measurable_on A/\B by MESFUNC1:def 16; end; assume A6: f is_measurable_on (A/\B); now let r be Real; A /\ B /\ less_dom(f, r) = B /\ (A /\ less_dom(f, r)) by XBOOLE_1:16 .=B /\ less_dom(f, r) by A2; hence B /\ less_dom(f, r) in S by A6,MESFUNC1:def 16; end; hence thesis by MESFUNC1:def 16; end; theorem Th49: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL st ( ex A be Element of S st dom f = A ) for c be Real, B be Element of S st f is_measurable_on B holds c(#)f is_measurable_on B proof let X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL; assume ex A be Element of S st A = dom f; then consider A be Element of S such that A1: A = dom f; let c be Real, B be Element of S; assume 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 thesis by A2,Th48; end; begin :: Sequence of extended real numbers definition mode ExtREAL_sequence is sequence of ExtREAL; end; definition let seq be ExtREAL_sequence; attr seq is convergent_to_finite_number means ex g be Real st for p be Real st 0

0; then consider n1 be Nat such that A6: for m be Nat st n1 <= m holds |.seq.m- g.| < g by A4; A7: now let m be Nat; assume n1 <= m; then |.seq.m- g qua ExtReal.| < g by A6; then seq.m - g1 < g by EXTREAL1:21; then seq.m < (g+g) by XXREAL_3:54; hence seq.m < (2*g); end; consider n2 be Nat such that A8: for m be Nat st n2 <= m holds (2*g) <= seq.m by A1,A5; reconsider n1,n2 as Element of NAT by ORDINAL1:def 12; set m = max(n1,n2); seq.m < (2*g) by A7,XXREAL_0:25; hence contradiction by A8,XXREAL_0:25; end; suppose A9: g = 0; consider n1 be Nat such that A10: for m be Nat st n1 <= m holds |. seq.m- g .| < 1 by A4; consider n2 be Nat such that A11: for m be Nat st n2 <= m holds 1 <= seq.m by A1; reconsider n1,n2 as Element of NAT by ORDINAL1:def 12; reconsider jj =1 as R_eal by XXREAL_0:def 1; set m = max(n1,n2); |. seq.m- g1 .| < jj by A10,XXREAL_0:25; then seq.m - g1 < jj by EXTREAL1:21; then seq.m < 1 + g by XXREAL_3:54; then seq.m < 1 by A9; hence contradiction by A11,XXREAL_0:25; end; suppose A12: g < 0; consider n1 be Nat such that A13: for m be Nat st n1 <= m holds |.seq.m- g.| < -g1 by A4,A12; A14: now let m be Element of NAT; assume n1 <= m; then |.seq.m- g1.| < -g1 by A13; then seq.m - g1 < -g1 by EXTREAL1:21; then seq.m < g-g1 by XXREAL_3:54; hence seq.m < 0 by XXREAL_3:7; end; consider n2 be Nat such that A15: for m be Nat st n2 <= m holds 1 <= seq.m by A1; reconsider n1,n2 as Element of NAT by ORDINAL1:def 12; set m = max(n1,n2); seq.m < 0 by A14,XXREAL_0:25; hence contradiction by A15,XXREAL_0:25; end; end; theorem Th51: for seq be 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; assume A1: seq is convergent_to_-infty; hereby assume seq is convergent_to_+infty; then consider n1 being Nat such that A2: for m be Nat st n1 <= m holds 1 <= seq.m; consider n2 being Nat such that A3: for m be Nat st n2 <= m holds seq.m <= -1 by A1; reconsider n1,n2 as Element of NAT by ORDINAL1:def 12; set m = max(n1,n2); seq.m <= -1 by A3,XXREAL_0:25; hence contradiction by A2,XXREAL_0:25; end; assume seq is convergent_to_finite_number; then consider g being Real such that A4: for p be Real st 0 < p ex n be Nat st for m be Nat st n<=m holds |.seq.m- g.| < p; reconsider g1 = g as R_eal by XXREAL_0:def 1; per cases; suppose A5: g > 0; then consider n1 be Nat such that A6: for m be Nat st n1 <= m holds |.seq.m- g.| < (g) by A4; A7: now let m be Element of NAT; assume n1 <= m; then |.seq.m- g1.| < g by A6; then -g1 < seq.m - g by EXTREAL1:21; then -g + g < seq.m by XXREAL_3:53; hence 0 < seq.m; end; consider n2 be Nat such that A8: for m be Nat st n2 <= m holds seq.m <= -g1 by A1,A5; reconsider n1,n2 as Element of NAT by ORDINAL1:def 12; set m = max(n1,n2); 0 < seq.m by A7,XXREAL_0:25; hence contradiction by A5,A8,XXREAL_0:25; end; suppose A9: g = 0; consider n1 be Nat such that A10: for m be Nat st n1 <= m holds |. seq.m- g .| < 1 by A4; consider n2 be Nat such that A11: for m be Nat st n2 <= m holds seq.m <= -1 by A1; reconsider n1,n2 as Element of NAT by ORDINAL1:def 12; reconsider jj=1 as R_eal by XXREAL_0:def 1; set m = max(n1,n2); |. seq.m- g1 .| < 1 by A10,XXREAL_0:25; then - jj < seq.m - g1 by EXTREAL1:21; then - 1 + g < seq.m by XXREAL_3:53; then - 1 < seq.m by A9; then (-1) < seq.m; hence contradiction by A11,XXREAL_0:25; end; suppose A12: g < 0; then consider n1 be Nat such that A13: for m be Nat st n1 <= m holds |.seq.m- g.| < -g1 by A4; A14: now let m be Element of NAT; assume n1 <= m; then |.seq.m- g1.| < -g1 by A13; then --g1 < seq.m - g by EXTREAL1:21; then g1 + g < seq.m by XXREAL_3:53; then (g+g) < seq.m; hence (2*g) < seq.m; end; consider n2 be Nat such that A15: for m be Nat st n2 <= m holds seq.m <= 2*g by A1,A12; reconsider n1,n2 as Element of NAT by ORDINAL1:def 12; set m = max(n1,n2); seq.m <= (2*g) by A15,XXREAL_0:25; hence contradiction by A14,XXREAL_0:25; end; end; definition let seq be ExtREAL_sequence; attr seq is convergent means seq is convergent_to_finite_number or seq is convergent_to_+infty or seq is convergent_to_-infty; end; definition let seq be ExtREAL_sequence; assume A1: seq is convergent; func lim seq -> R_eal means :Def12: ( ex g be Real st it = g & (for p be Real st 0

g2; per cases by A1; suppose A7: seq is convergent_to_finite_number; then consider g being Real such that A8: g1 = g and A9: for p be Real st 0 < p ex n be Nat st for m be Nat st n <= m holds |.seq.m-g1.| < p and seq is convergent_to_finite_number by A4,Th50,Th51; consider h being Real such that A10: g2 = h and A11: for p be Real st 0 < p ex n be Nat st for m be Nat st n <= m holds |.seq.m-g2.| < p and seq is convergent_to_finite_number by A5,A7,Th50,Th51; reconsider g,h as Complex; g - h <> 0 by A6,A8,A10; then A12: |.g-h.| > 0; then consider n1 being Nat such that A13: for m be Nat st n1 <= m holds |.seq.m-g1.| < |.g-h.|/2 by A9; consider n2 being Nat such that A14: for m be Nat st n2 <= m holds |.seq.m-g2.| < |.g-h.|/2 by A11,A12; reconsider n1,n2 as Element of NAT by ORDINAL1:def 12; set m = max(n1,n2); A15: |.seq.m-g1.| < |.g-h.|/2 by A13,XXREAL_0:25; A16: |.seq.m-g2.| < |.g-h.|/2 by A14,XXREAL_0:25; reconsider g,h as Complex; A17: seq.m-g2 < |.g-h.|/2 by A16,EXTREAL1:21; A18: -(|.g-h.|/2 qua ExtReal) < seq.m-g2 by A16,EXTREAL1:21; then reconsider w = seq.m - g2 as Element of REAL by A17,XXREAL_0:48; A19: seq.m-g2 in REAL by A18,A17,XXREAL_0:48; then A20: seq.m <> +infty by A10; A21: -seq.m + g1 = -(seq.m - g1) by XXREAL_3:26; then A22: |.-seq.m+g1.| < |.g-h.|/2 by A15,EXTREAL1:29; then A23: -seq.m+g1 < |.g-h.|/2 by EXTREAL1:21; -(|.g-h.|/2 qua ExtReal) < -seq.m+g1 by A22,EXTREAL1:21; then A24: -seq.m + g1 in REAL by A23,XXREAL_0:48; A25: seq.m <> -infty by A10,A19; |.g1-g2.| = |.g1 + 0. - g2.| by XXREAL_3:4 .= |.g1 + (seq.m + -seq.m) - g2.| by XXREAL_3:7 .= |.-seq.m + g1 + seq.m - g2.| by A8,A20,A25,XXREAL_3:29 .= |.-seq.m + g1 +(seq.m - g2).| by A10,A24,XXREAL_3:30; then |.g1-g2.| <= |.-seq.m + g1.| + |.seq.m - g2.| by EXTREAL1:24; then A26: |.g1-g2.| <= |.seq.m-g1.| + |.seq.m-g2.| by A21,EXTREAL1:29; |.w.| in REAL by XREAL_0:def 1; then |.seq.m-g2.| in REAL; then A27: |.seq.m-g1.| + |.seq.m-g2.| < (|.g-h.|/2 qua ExtReal) + |.seq.m-g2.| by A15,XXREAL_3:43; |.g-h.|/2 in REAL by XREAL_0:def 1; then |.g-h.|/2 in REAL; then (|.g-h.|/2 qua ExtReal) + |.seq.m-g2.| < (|.g-h.|/2 qua ExtReal) + |.g-h.|/2 by A16,XXREAL_3:43; then A28: |.seq.m-g1.| + |.seq.m-g2.| < (|.g-h.|/2 qua ExtReal) + |.g-h.| /2 by A27,XXREAL_0:2; g-h = g1-g2 by A8,A10,SUPINF_2:3; then |.g-h.| = |.g1-g2.| by EXTREAL1:12; then |.g-h.| < (|.g-h.|/2) + |.g-h.| /2 by A28,A26; hence contradiction; end; suppose seq is convergent_to_+infty or seq is convergent_to_-infty; hence contradiction by A4,A5,A6,Th50,Th51; end; end; end; theorem Th52: for seq be ExtREAL_sequence, r be Real st (for n be Nat holds seq.n = r) holds seq is convergent_to_finite_number & lim seq = r proof let seq be ExtREAL_sequence; let r be Real; assume A1: for n be Nat holds seq.n = r; A2: now reconsider n=1 as Nat; let p be Real; assume A3: 0 < p; take n; let m be Nat such that n <= m; seq.m = r by A1; then seq.m - r = 0 by XXREAL_3:7; then |. seq.m - r.| = 0 by EXTREAL1:16; hence |. seq.m - r.| < p by A3; end; hence A4: seq is convergent_to_finite_number; then A5: seq is convergent; reconsider r as R_eal by XXREAL_0:def 1; ( ex g be Real st r = g & (for p be Real st 0

+infty by A10,XXREAL_0:9; A12: sup rng L <> -infty by A2,A8; then reconsider h=sup rng L as Element of REAL by A11,XXREAL_0:14; A13: for p be Real st 0

+infty; L.k0 <= sup(rng L) by A2; then reconsider h=sup rng L as Element of REAL by A8,A30,XXREAL_0:14; set K=max(0,h); 0 <=K by XXREAL_0:25; then consider N0 be Nat such that A31: K+1 <= L.N0 by A26; h+0 < K+1 by XREAL_1:8,XXREAL_0:25; then sup rng L < L.N0 by A31,XXREAL_0:2; hence contradiction by A2; end; hence thesis by A28,A29,Def12; end; end; end; theorem Th55: for L,G be ExtREAL_sequence st (for n be Nat holds L.n <= G.n) holds sup rng L <= sup rng G proof let L,G be ExtREAL_sequence; assume A1: for n be Nat holds L.n <= G.n; A2: now let n be Element of NAT; 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; end; now let x be ExtReal; assume x in rng L; then ex z be object st z in dom L & x=L.z by FUNCT_1:def 3; hence x <= sup rng G by A2; end; then sup rng G is UpperBound of rng L by XXREAL_2:def 1; hence thesis by XXREAL_2:def 3; end; theorem Th56: for L be ExtREAL_sequence holds for n be Nat holds L.n <= sup rng L proof let L be ExtREAL_sequence; let n be Nat; reconsider 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 thesis by A1,XXREAL_2:def 1; end; theorem Th57: for L be ExtREAL_sequence, K be R_eal st (for n be Nat holds L.n <= K) holds sup rng L <= K proof let L be ExtREAL_sequence, K be R_eal; assume A1: for n be Nat holds L.n <= K; now let x be ExtReal; assume x in rng L; then ex z be object st z in dom L & x=L.z by FUNCT_1:def 3; hence x <= K by A1; end; then K is UpperBound of rng L by XXREAL_2:def 1; hence thesis by XXREAL_2:def 3; end; theorem for L be ExtREAL_sequence, K be R_eal st K <> +infty & (for n be Nat holds L.n <= K) holds sup rng L < +infty proof let L be ExtREAL_sequence, K be R_eal; assume that A1: K <> +infty and A2: for n be Nat holds L.n <= K; now let x be ExtReal; assume x in rng L; then ex z be object st z in dom L & x=L.z by FUNCT_1:def 3; hence x <= K by A2; end; then K is UpperBound of rng L by XXREAL_2:def 1; then sup rng L <= K by XXREAL_2:def 3; hence thesis by A1,XXREAL_0:2,4; end; theorem Th59: for L be ExtREAL_sequence st L is without-infty holds sup rng L <> +infty iff ex K be Real st 0 +infty; 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; take K = max(S,1); thus 0 < K by XXREAL_0:25; let n be Nat; 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 <= K by A5,XXREAL_0:2; end; now given K being Real such that 0 < K and A6: for n be Nat holds L.n <= K; now let w be ExtReal; assume w in rng L; then ex v be object st v in dom L & w = L.v by FUNCT_1:def 3; hence w <= K by A6; end; then K is UpperBound of rng L by XXREAL_2:def 1; then A7: sup rng L <= K by XXREAL_2:def 3; K in REAL by XREAL_0:def 1; hence sup rng L <> +infty by A7,XXREAL_0:9; end; hence thesis by A4; end; theorem Th60: for L be ExtREAL_sequence,c be ExtReal st (for n be Nat holds L.n=c) holds L is convergent & lim L = c & lim L = sup rng L proof let L be ExtREAL_sequence; let c be ExtReal; reconsider cc = c as R_eal by XXREAL_0:def 1; 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 be Nat holds L.n = c; then A4: L.1 = c; now let v be ExtReal; assume v in rng L; then ex n be object st n in dom L & v = L.n by FUNCT_1:def 3; hence v <= c by A3; end; then A5: c is UpperBound of rng L by XXREAL_2:def 1; per cases by A2,TARSKI:def 2; suppose c in REAL; then reconsider rc = c as Real; A6: now reconsider n=0 as Nat; let p be Real; assume A7: 0 < p; take n; let m be Nat such that n <= m; L.m - rc = L.m - L.m by A3; then L.m - rc = 0 by XXREAL_3:7; hence |. L.m - rc .| < p by A7,EXTREAL1:16; end; then A8: L is convergent_to_finite_number; hence L is convergent; then lim L = cc by A6,A8,Def12; hence lim L = c; hence thesis by A5,A1,A4,FUNCT_1:3,XXREAL_2:55; end; suppose A9: c = -infty; for p be Real st p < 0 ex n be Nat st for m be Nat st n<=m holds L.m <= p proof let p be Real such that p < 0; take 0; A10: p in REAL by XREAL_0:def 1; now let m be Nat such that 0 <= m; L.m = -infty by A3,A9; hence L.m < p by A10,XXREAL_0:12; end; hence thesis; end; then A11: L is convergent_to_-infty; hence L is convergent; hence lim(L) = c by A9,A11,Def12; hence thesis by A5,A1,A4,FUNCT_1:3,XXREAL_2:55; end; suppose A12: c = +infty; for p be Real st 0 < p ex n be Nat st for m be Nat st n <= m holds p <= L.m proof let p be Real such that 0 < p; take 0; A13: p in REAL by XREAL_0:def 1; now let m be Nat such that 0 <= m; L.m = +infty by A3,A12; hence p < L.m by A13,XXREAL_0:9; end; hence thesis; end; then A14: L is convergent_to_+infty; hence L is convergent; hence lim L = c by A12,A14,Def12; hence thesis by A5,A1,A4,FUNCT_1:3,XXREAL_2:55; end; end; Lm8: for J be ExtREAL_sequence st J is without-infty holds sup rng J in REAL or sup rng J = +infty proof let J be ExtREAL_sequence; assume J is without-infty; then A1: -infty < J.1; dom J = NAT by FUNCT_2:def 1; then not -infty is UpperBound of rng J by A1,FUNCT_1:3,XXREAL_2:def 1; then sup rng J <> -infty by XXREAL_2:def 3; hence thesis by XXREAL_0:14; end; theorem Th61: for J,K,L be ExtREAL_sequence st (for n,m be Nat st n <=m holds J.n <= J.m) & (for n,m be Nat st n <=m holds K.n <= K.m) & J is without-infty & K is without-infty & (for n be 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; assume that A1: for n,m be Nat st n <=m holds J.n <= J.m and A2: for n,m be Nat st n <=m holds K.n <= K.m and A3: J is without-infty and A4: K is without-infty and A5: for n be Nat holds J.n+K.n =L.n; A6: dom K = NAT by FUNCT_2:def 1; A7: dom J = NAT by FUNCT_2:def 1; A8: now per cases by A3,Lm8; suppose A9: sup rng J in REAL; then reconsider SJ = sup rng J as Real; per cases by A4,Lm8; suppose A10: sup rng K in REAL; then reconsider SK = sup rng K as Real; A11: now let p be Real; assume A12: 0 < p; then consider SJ9 be ExtReal such that A13: SJ9 in rng J and A14: sup rng J - p/2 < SJ9 by A9,MEASURE6:6; consider nj be object such that A15: nj in dom J and A16: SJ9 = J.nj by A13,FUNCT_1:def 3; reconsider nj as Element of NAT by A15; consider SK9 be ExtReal such that A17: SK9 in rng K and A18: sup rng K - p/2 < SK9 by A10,A12,MEASURE6:6; consider nk be object such that A19: nk in dom K and A20: SK9 = K.nk by A17,FUNCT_1:def 3; reconsider nk as Element of NAT by A19; reconsider n = max(nj,nk) as Nat; take n; hereby reconsider SJ9, SK9 as R_eal by XXREAL_0:def 1; let m be Nat; assume A21: n <= m; nk <=n by XXREAL_0:25; then nk <= m by A21,XXREAL_0:2; then SK9 <= K.m by A2,A20; then A22: SK - K.m <= SK - SK9 by XXREAL_3:37; nj <= n by XXREAL_0:25; then nj <= m by A21,XXREAL_0:2; then SJ9 <= J.m by A1,A16; then SJ - J.m <= SJ - SJ9 by XXREAL_3:37; then A23: (SJ - J.m) + (SK - K.m) <= (SJ - SJ9) + ( SK - SK9) by A22,XXREAL_3:36; SJ in REAL by XREAL_0:def 1; then A24: SJ < +infty by XXREAL_0:9; reconsider s1 = SK as Element of REAL by XREAL_0:def 1; reconsider m1=m as Element of NAT by ORDINAL1:def 12; A25: -(L.m - (SJ+SK)) = (SJ+SK) - L.m by XXREAL_3:26; A26: (p/2) in REAL by XREAL_0:def 1; SK < (p/2) + SK9 by A18,XXREAL_3:54; then SK - SK9 < (p/2) by XXREAL_3:55; then A27: (p/2) + (SK - SK9) < (p/2) + (p/2) by XXREAL_3:43,A26; SJ < (p/2) + SJ9 by A14,XXREAL_3:54; then A28: SJ - SJ9 < (p/2) by XXREAL_3:55; nk <= n by XXREAL_0:25; then nk <= m by A21,XXREAL_0:2; then A29: K.nk <= K.m by A2; A30: SK in REAL by XREAL_0:def 1; then A31: SK < +infty by XXREAL_0:9; K.m1 in rng K by A6,FUNCT_1:3; then A32: K.m <= SK by XXREAL_2:4; then A33: K.m < +infty by A30,XXREAL_0:2,9; -infty < SK9 by A4,A20; then reconsider s0 = SK9 as Element of 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 <= SJ by XXREAL_2:4; then J.m + K.m <= SJ + SK by A32,XXREAL_3:36; then L.m - (SJ+SK) <= 0 by A34,A25,XXREAL_3:40; then A36: |. L.m - (SJ+SK) .| = (SJ+SK) - L.m by A25,EXTREAL1:18; SK - SK9 = s1 - s0 by SUPINF_2:3; then (SJ - SJ9) + (SK - SK9) < (p/2) + (SK - SK9) by A28,XXREAL_3:43; then A37: (SJ - SJ9) + (SK - SK9) < (p/2) + (p/2) by A27,XXREAL_0:2; -infty < K.m by A4; then SJ - J.m + (SK - K.m) = SJ - J.m + SK - K.m by A33,XXREAL_3:30 .= SK + SJ - J.m - K.m by XXREAL_3:30 .= (SK+SJ) - (J.m + K.m) by A24,A31,A35,A32,XXREAL_3:31 .= (SK+SJ) - L.m by A5; hence |. L.m - (SJ+SK) .| < p by A36,A37,A23,XXREAL_0:2; end; end; then A38: L is convergent_to_finite_number; hence L is convergent; hence lim L = sup rng J + sup rng K by A11,A38,Def12; end; suppose A39: sup rng K = +infty; for p be Real st 0 < p ex n be Nat st for m be Nat st n<=m holds p <= L.m proof reconsider supj = sup rng J as Element of REAL by A9; let p be Real; reconsider p92 = p/2, p9 = p as Element of REAL by XREAL_0:def 1; assume 0 < p; then consider j be ExtReal such that A40: j in rng J and A41: sup rng J - (p/2) < j by A9,MEASURE6:6; consider n1 being object such that A42: n1 in dom J and A43: j = J.n1 by A40,FUNCT_1:def 3; A44: supj - p92 = sup rng J - (p/2) by SUPINF_2:3; then A45: p9 - (supj - p92) = p - (sup rng J - (p/2)) by SUPINF_2:3; then p - (sup rng J -(p/2)) < sup rng K by A39,XXREAL_0:9; then consider k being Element of ExtREAL such that A46: k in rng K and A47: p - (sup rng J - (p/2)) < k by XXREAL_2:94; p9 = (p9 - (supj - p92)) + (supj - p92); then A48: p - (sup rng J - (p/2)) + (sup rng J - (p/2)) = p9 by A44,A45,SUPINF_2:1; reconsider n1 as Element of NAT by A42; consider n2 being object such that A49: n2 in dom K and A50: k = K.n2 by A46,FUNCT_1:def 3; reconsider n2 as Element of NAT by A49; set n = max(n1,n2); J.n1 <= J.n by A1,XXREAL_0:25; then A51: sup rng J - (p/2) < J.n by A41,A43,XXREAL_0:2; K.n2 <= K.n by A2,XXREAL_0:25; then A52: p - (sup rng J - (p/2)) < K.n by A47,A50,XXREAL_0:2; A53: p < J.n + K.n by A51,A52,A48,XXREAL_3:64; now let m be Nat; assume A54: n <= m; then A55: K.n <= K.m by A2; J.n <= J.m by A1,A54; then J.n + K.n <= J.m + K.m by A55,XXREAL_3:36; then J.n + K.n <= L.m by A5; hence p <= L.m by A53,XXREAL_0:2; end; hence thesis; end; then A56: L is convergent_to_+infty; hence L is convergent; then lim L = +infty by A56,Def12; hence lim L = sup rng J + sup rng K by A9,A39,XXREAL_3:def 2; end; end; suppose A57: sup rng J = +infty; now let p be Real; assume A58: 0 < p; per cases by A4,Lm8; suppose A59: sup rng K in REAL; then reconsider supk = sup rng K as Element of REAL; reconsider p92 = p/2, p9 = p as Element of REAL by XREAL_0:def 1; A60: supk - p92 = sup rng K - (p/2) by SUPINF_2:3; then A61: p9 - (supk - p92) = p - (sup rng K - (p/2)) by SUPINF_2:3; then p - (sup rng K - (p/2)) < sup rng J by A57,XXREAL_0:9; then consider j being Element of ExtREAL such that A62: j in rng J and A63: p - (sup rng K - (p/2)) < j by XXREAL_2:94; p9 = p9 - (supk - p92) + (supk - p92); then A64: p - (sup rng K - (p/2)) + (sup rng K - (p/2)) = p9 by A60,A61,SUPINF_2:1; consider k be ExtReal such that A65: k in rng K and A66: sup rng K - (p/2) < k by A58,A59,MEASURE6:6; consider n1 being object such that A67: n1 in dom K and A68: k = K.n1 by A65,FUNCT_1:def 3; consider n2 being object such that A69: n2 in dom J and A70: j = J.n2 by A62,FUNCT_1:def 3; reconsider n1 as Element of NAT by A67; reconsider n2 as Element of NAT by A69; set n = max(n1,n2); J.n2 <= J.n by A1,XXREAL_0:25; then A71: p - (sup rng K - (p/2)) < J.n by A63,A70,XXREAL_0:2; K.n1 <= K.n by A2,XXREAL_0:25; then A72: sup rng K - (p/2) < K.n by A66,A68,XXREAL_0:2; A73: p < J.n + K.n by A72,A71,A64,XXREAL_3:64; now let m be Nat; assume A74: n <= m; then A75: K.n <= K.m by A2; J.n <= J.m by A1,A74; then J.n + K.n <= J.m + K.m by A75,XXREAL_3:36; then J.n + K.n <= L.m by A5; hence p <= L.m by A73,XXREAL_0:2; end; hence ex n be Nat st for m be Nat st n <= m holds p <= L.m; end; suppose sup rng K = +infty; then consider n1 be Nat such that A76: (p/2) < K.n1 by A4,A58,Th59; consider n2 be Nat such that A77: (p/2) < J.n2 by A3,A57,A58,Th59; reconsider n1,n2 as Element of NAT by ORDINAL1:def 12; set n = max(n1,n2); K.n1 <= K.n by A2,XXREAL_0:25; then A78: (p/2) < K.n by A76,XXREAL_0:2; J.n2 <= J.n by A1,XXREAL_0:25; then A79: (p/2) < J.n by A77,XXREAL_0:2; (p/2) + (p/2) < J.n + K.n by A79,A78,XXREAL_3:64; then p < J.n + K.n; then A80: p < L.n by A5; now let m be Nat; assume A81: n <= m; then A82: K.n <= K.m by A2; J.n <= J.m by A1,A81; then J.n + K.n <= J.m + K.m by A82,XXREAL_3:36; then J.n + K.n <= L.m by A5; then L.n <= L.m by A5; hence p <= L.m by A80,XXREAL_0:2; end; hence ex n be Nat st for m be Nat st n <= m holds p <= L.m; end; end; then A83: L is convergent_to_+infty; hence L is convergent; then A84: lim L = +infty by A83,Def12; A85: K.0 <= sup rng K by A6,FUNCT_1:3,XXREAL_2:4; -infty < K.0 by A4; hence lim L = sup rng J + sup rng K by A57,A84,A85,XXREAL_3:def 2; end; end; hence L is convergent; A86: now let n,m be Nat; assume A87: n <= m; then A88: K.n <= K.m by A2; J.n <= J.m by A1,A87; then J.n + K.n <= J.m + K.m by A88,XXREAL_3:36; then L.n <= J.m + K.m by A5; hence L.n <= L.m by A5; end; hence lim L= sup rng L by Th54; lim J = sup rng J by A1,Th54; hence thesis by A2,A8,A86,Th54; end; theorem Th62: for L,K be ExtREAL_sequence,c be Real st 0 <= c & L is without-infty & (for n be Nat holds K.n = c * L.n ) holds sup rng K = c * sup rng L & K is without-infty proof let L,K be ExtREAL_sequence; let c be Real; assume that A1: 0 <= c and A2: L is without-infty and A3: for n be Nat holds K.n = c * L.n; now per cases by A2,Lm8; suppose A4: sup rng L in REAL; A5: for y being UpperBound of rng K holds c * sup rng L <= y proof let y be UpperBound of rng K; reconsider 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; per cases; suppose A8: c = 0; A9: K.1 <= y by A7,FUNCT_1:3,XXREAL_2:def 1; K.1 = c * L.1 by A3; hence thesis by A8,A9; end; suppose A10: c <> 0; now let x be ExtReal; assume x in rng L; then consider n be object such that A11: n in dom L and A12: x = L.n by FUNCT_1:def 3; reconsider n as Element of NAT by A11; A13: K.n in rng K by A7,FUNCT_1:3; K.n = c * L.n by A3; then c * L.n / c <= y/ c by A1,A10,A13, XXREAL_2:def 1,XXREAL_3:79; hence x <= y / c by A10,A12,XXREAL_3:88; end; then y/c is UpperBound of rng L by XXREAL_2:def 1; then A14: sup rng L <= y/c by XXREAL_2:def 3; A15: now assume A16: y = -infty; 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: 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,9; -infty < L.1 by A2; hence contradiction by A17,A18,XXREAL_3:70; end; per cases by A15,XXREAL_0:14; suppose y = +infty; hence thesis by XXREAL_0:4; end; suppose y in REAL; then reconsider ry = y as Real; reconsider sl = sup rng L as Real by A4; y/c = ry/c; then sl*c <= ry by A1,A10,A14,XREAL_1:83; hence thesis; end; end; end; now let x be ExtReal; A19: sup rng L is UpperBound of rng L by XXREAL_2:def 3; assume x in rng K; then consider m be object such that A20: m in dom K and A21: x = K.m by FUNCT_1:def 3; reconsider 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 = c * L.m by A3,A21; hence x <= c*sup rng L by A1,A22,XXREAL_3:71; end; then 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; end; suppose A23: sup rng L = +infty; per cases; suppose A24: c = 0; A25: now let n be Nat; K.n = c * L.n by A3; hence K.n = 0 by A24; end; then lim K = sup rng K by Th60; hence sup rng K = c * sup rng L by A24,A25,Th60; end; suppose A26: c <> 0; now let n be object; -infty < L.n by A2; then A27: -infty* c < c * L.n by A1,A26,XXREAL_3:72; per cases; suppose n in dom K; then reconsider n1=n as Element of NAT; -infty* c = -infty by A1,A26,XXREAL_3:def 5; then -infty < K.n1 by A3,A27; hence -infty < K.n; end; suppose not n in dom K; hence -infty < K.n by FUNCT_1:def 2; end; end; then A28: K is without-infty; A29: now let k be Real; reconsider k1 = k as Real; A30: (k/c)* c = k1/c * c .= k1/(c/c) by XCMPLX_1:82 .= k by A26,XCMPLX_1:51; assume 0 < k; then consider n be Nat such that A31: (k/c) < L.n by A1,A2,A23,A26,Th59; (k/c) * c < c * L.n by A1,A26,A31,XXREAL_3:72; then k < K.n by A3,A30; hence ex n be Nat st not K.n <= k; end; 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; end; end; end; hence sup rng K = c * sup rng L; now let n be object; A32: L.n = +infty implies c * L.n <> -infty by A1; -infty < L.n by A2; then A33: -infty <> c * L.n by A32,XXREAL_3:70; per cases; suppose n in dom K; then reconsider n1=n as Element of NAT; K.n1 <> -infty by A3,A33; hence -infty < K.n by XXREAL_0:6; end; suppose not n in dom K; hence -infty < K.n by FUNCT_1:def 2; end; end; hence thesis; end; theorem Th63: for L,K be ExtREAL_sequence, c be Real st 0 <=c & (for n,m be Nat st n <=m holds L.n <= L.m) & (for n be Nat holds K.n = c * L.n) & L is without-infty holds (for n,m be Nat st n <=m holds K.n <= K.m) & K is without-infty & K is convergent & lim K = sup rng K & lim K = c * lim L proof let L,K be ExtREAL_sequence, c be Real; assume that A1: 0 <= c and A2: for n,m be Nat st n <=m holds L.n <= L.m and A3: for n be Nat holds K.n = c * L.n and A4: L is without-infty; A5: sup rng L = lim L by A2,Th54; thus A6: now let n,m be Nat; assume n <= m; then c * L.n <= c * L.m by A1,A2,XXREAL_3:71; then K.n <= c * L.m by A3; hence K.n <= K.m by A3; end; thus K is without-infty by A1,A3,A4,Th62; thus K is convergent & lim K= sup rng K by A6,Th54; sup rng K = lim K by A6,Th54; hence thesis by A1,A3,A4,A5,Th62; end; begin :: Sequence of extended real valued functions definition let X be non empty set, H be Functional_Sequence of X,ExtREAL, x be Element of X; func H#x -> ExtREAL_sequence means :Def13: for n be Nat holds it.n = (H.n).x; existence proof deffunc U(Nat) = (H.$1).x; consider f be sequence of ExtREAL such that A1: for n be Element of NAT holds f.n = U(n) from FUNCT_2:sch 4; take f; let n be Nat; n in NAT by ORDINAL1:def 12; hence thesis by A1; end; uniqueness proof let S1,S2 be ExtREAL_sequence such that A2: for n be Nat holds S1.n = (H.n).x and A3: for n be Nat holds S2.n = (H.n).x; now let n be Element of NAT; S1.n = (H.n).x by A2; hence S1.n = S2.n by A3; end; hence thesis by FUNCT_2:63; end; end; definition let D1,D2 be set, F be sequence of PFuncs(D1,D2), n be Nat; redefine func F.n -> PartFunc of D1,D2; coherence 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 thesis; end; end; theorem Th64: for X be non empty set, S be SigmaField of X, f be PartFunc of X,ExtREAL st (ex A be Element of S st A = dom f & f is_measurable_on A) & f qua ext-real-valued Function is nonnegative ex F be Functional_Sequence of X,ExtREAL st (for n be Nat holds F.n is_simple_func_in S & dom(F.n) = dom f) & (for n be Nat holds F.n is nonnegative) & (for n,m be Nat st n <=m holds for x be Element of X st x in dom f holds (F.n).x <= (F.m).x ) & for x be 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, S be SigmaField of X, f be PartFunc of X,ExtREAL such that A1: ex A be Element of S st A = dom f & f is_measurable_on A and A2: f is nonnegative; defpred PF[Element of NAT,PartFunc of X,ExtREAL] means dom $2 = dom f & (for x be Element of X st x in dom f holds ( for k be 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 be Element of NAT ex y being Element of PFuncs(X,ExtREAL) st PF[n, y] proof let n be Element of NAT; reconsider nn=n as Nat; defpred PP[object,object] means ( for k be 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 be object st x in dom f ex y be object st PP[x,y] proof let x be object; assume x in dom f; per cases; suppose A5: f.x < n; 0 <= f.x by A2,SUPINF_2:51; then consider k be Nat such that 1 <= k and k <= 2|^nn*nn and A6: (k-1)/(2|^nn) <= (f.x qua ExtReal) and A7: f.x < k/(2|^nn) by A5,Th4; take y = (k-1)/(2|^n); now let k1 be Nat; assume that 1 <= k1 and k1 <= 2|^n*n and A8: (k1-1)/(2|^n) <= f.x and A9: f.x < k1/(2|^n); A10: now assume k1 < k; 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; end; now assume k < k1; 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; end; hence y=(k1-1)/(2|^n) by A10,XXREAL_0:1; end; hence thesis by A5; end; suppose A11: n <= f.x; reconsider y = nn as Real; take y; thus for k be Nat st 1 <= k & k <= 2|^n*n & (k-1)/( 2|^n) <= f.x & f.x < k/(2|^n) holds y=(k-1)/(2|^n) proof let k be Nat such that 1 <= k and A12: k <= 2|^n*n and (k-1)/(2|^n) <= f.x and A13: f.x < k/(2|^n); reconsider p = f.x as ExtReal; k <= 2|^nn*nn by A12; then k/(2|^nn) <= p by A11,Th5; then k/(2|^n) <= p; hence y=(k-1)/(2|^n) by A13; end; thus thesis; end; end; consider fn be Function such that A14: dom fn = dom f & for x be object st x in dom f holds PP[x,fn.x] from CLASSES1:sch 1(A4); now let w be object; assume w in rng fn; then consider v be object such that A15: v in dom fn and A16: w = fn.v by FUNCT_1:def 3; per cases; suppose n <= f.v; then fn.v = n by A14,A15; hence w in ExtREAL by A16,XXREAL_0:def 1; end; suppose A17: f.v < n; 0 <= f.v by A2,SUPINF_2:51; then consider k be Nat such that A18: 1 <= k and A19: k <= 2|^nn*nn and A20: (k-1)/(2|^nn) <= f.v and A21: f.v < k/(2|^nn) by A17,Th4; fn.v = (k-1)/(2|^n) by A14,A15,A18,A19,A20,A21; hence w in ExtREAL by A16; end; end; then rng fn c= ExtREAL; then reconsider fn as PartFunc of dom f,ExtREAL by A14,RELSET_1:4; reconsider fn as PartFunc of X,ExtREAL by A14,RELSET_1:5; reconsider y = fn as Element of PFuncs(X,ExtREAL) by PARTFUN1:45; take y; thus thesis by A14; end; consider F be sequence of PFuncs(X,ExtREAL) such that A22: for n be Element of NAT holds PF[n,F.n] from FUNCT_2:sch 3(A3); A23: for n be Element of NAT holds dom (F.n) = dom f by A22; A24: for n be Element of NAT, x be Element of X st x in dom f holds (for k be 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) by A22; A25: now let n be Nat; n in NAT by ORDINAL1:def 12; hence dom (F.n) = dom f by A23; end; reconsider F as Functional_Sequence of X,ExtREAL; consider A be Element of S such that A26: A = dom f and A27: f is_measurable_on A by A1; A28: for n,m be Nat st n <=m holds for x be Element of X st x in dom f holds (F.n).x <= (F.m).x proof let n,m be Nat such that A29: n <=m; reconsider nn=n, mm=m as Element of NAT by ORDINAL1:def 12; let x be Element of X such that A30: x in dom f; per cases; suppose A31: m <= f.x; then A32: nn <= f.x by A29,XXREAL_0:2; (F.mm).x = m by A24,A30,A31; hence thesis by A24,A29,A30,A32; end; suppose A33: f.x < m; A34: 0 <= f.x by A2,SUPINF_2:51; then consider M be Nat such that A35: 1 <= M and A36: M <= 2|^m*m and A37: (M-1)/(2|^m) <= f.x and A38: f.x < M/(2|^m) by A33,Th4; reconsider M as Element of NAT by ORDINAL1:def 12; A39: (F.mm).x = (M-1)/(2|^m) by A24,A30,A35,A36,A37,A38; per cases; suppose A40: n <= f.x; reconsider M1 = 2|^mm as Element of NAT; n < M/(2|^m) by A38,A40,XXREAL_0:2; then 2|^m*n < M by PREPOWER:6,XREAL_1:79; then M1*n+1 <= M by NAT_1:13; then A41: M1*n <= M-1 by XREAL_1:19; A42: 0 < 2|^m by PREPOWER:6; (F.n).x = nn by A24,A30,A40; hence thesis by A39,A42,A41,XREAL_1:77; end; suppose A43: f.x < n; consider k being Nat such that A44: m = (nn qua Complex) + k by A29,NAT_1:10; reconsider k as Element of NAT by ORDINAL1:def 12; reconsider K = 2|^k as Element of NAT; consider N1 be Nat such that A45: 1 <= N1 and A46: N1 <= 2|^n*n and A47: (N1-1)/(2|^n) <= f.x and A48: f.x < N1/(2|^n) by A34,A43,Th4; reconsider N1 as Element of NAT by ORDINAL1:def 12; A49: (F.nn).x = (N1-1)/(2|^nn) by A24,A30,A45,A46,A47,A48; ((N1-1)/(2|^n)) < (M/(2|^(n+k))) by A38,A47,A44,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 A50: K*(N1-1)/(2|^n*2|^k) <= (M-1)/(2|^(n+k)) by NEWTON:8; 2|^k > 0 by PREPOWER:6; hence thesis by A39,A49,A44,A50,XCMPLX_1:91; end; end; end; A51: for n be Nat holds F.n is_simple_func_in S proof let n be Nat; reconsider nn=n as Element of NAT by ORDINAL1:def 12; reconsider N = 2|^nn as Element of NAT; defpred PG[Nat,set] means ( $1 <= N*n implies $2 = A /\ great_eq_dom(f, (($1-1)/(2|^n))) /\ less_dom(f,($1/(2|^n))) ) & ( $1 = N*n + 1 implies $2 = A /\ great_eq_dom(f,(n)) ); now let x be Element of X; assume x in dom(F.n); then A52: x in dom f by A25; per cases; suppose A53: n <= f.x; then (F.nn).x = n by A24,A52; then (F.n).x in REAL by XREAL_0:def 1; then A54: (F.n).x < +infty by XXREAL_0:9; - +infty < (F.nn).x by A24,A52,A53; hence |. (F.n).x .| < +infty by A54,EXTREAL1:22; end; suppose A55: f.x < n; A56: 0 <= f.x by A2,SUPINF_2:51; nn in REAL by XREAL_0:def 1; then nn < +infty by XXREAL_0:9; then reconsider y = f.x as Element of REAL by A55,A56,XXREAL_0:14; set k = [\ 2|^n*y /] + 1; A57: [\ 2|^n*y /] <= 2|^n*y by INT_1:def 6; 2|^n*y -1 < [\ 2|^n*y /] by INT_1:def 6; then A58: 2|^n*y < k by XREAL_1:19; A59: 0 < 2|^n by PREPOWER:6; then 2|^n*y < 2|^n*n by A55,XREAL_1:68; then [\ 2|^n*y /] < 2|^n*n by A57,XXREAL_0:2; then A60: k <= 2|^n*n by INT_1:7; A61: 0 <= 2|^n*y by A56; then A62: 0+1 <=k by A58,INT_1:7; reconsider k as Element of NAT by A61,A58,INT_1:3; reconsider k as Nat; k - 1 <= 2|^n*y by INT_1:def 6; then A63: (k-1)/(2|^nn) <= y by PREPOWER:6,XREAL_1:79; A64: (k-1)/(2|^nn) in REAL by XREAL_0:def 1; y < k/(2|^nn) by A59,INT_1:29,XREAL_1:81; then A65: (F.nn).x = (k-1)/(2|^nn) by A24,A52,A62,A60,A63; then -infty < (F.n).x by XXREAL_0:12,A64; then A66: - +infty < (F.n).x by XXREAL_3:def 3; (F.n).x < +infty by A65,XXREAL_0:9,A64; hence |. (F.n).x .| < +infty by A66,EXTREAL1:22; end; end; then A67: F.n is real-valued by MESFUNC2:def 1; A68: now let k be Nat; assume k in Seg(N*n+1); reconsider k1=k as Element of NAT by ORDINAL1:def 12; per cases; suppose A69: k <> N*n+1; set B = A /\ great_eq_dom(f,((k1-1)/(2|^n))) /\ less_dom(f, (k1/(2|^n))); reconsider B as Element of S by A26,A27,Th33; take B; thus PG[k,B] by A69; end; suppose A70: k = N*n+1; set B = A /\ great_eq_dom(f,n); reconsider B as Element of S by A26,A27,MESFUNC1:27; take B; thus PG[k,B] by A70,NAT_1:13; end; end; consider G be FinSequence of S such that A71: dom G = Seg(N*n +1) & for k be Nat st k in Seg(N*n +1) holds PG[k ,G.k] from FINSEQ_1:sch 5(A68); A72: now let k be Nat; assume that A73: 1 <= k and A74: k <= 2|^n*n; k <= N*n + 1 by A74,NAT_1:12; then k in Seg(N*n+1) by A73; hence G.k= A /\ great_eq_dom(f,((k-1)/(2|^n))) /\ less_dom(f,(k /(2|^n))) by A71,A74; end; A75: len G = 2|^n*n +1 by A71,FINSEQ_1:def 3; now let x,y be object; assume A76: x <> y; per cases; suppose not x in dom G or not y in dom G; then G.x = {} or G.y = {} by FUNCT_1:def 2; hence G.x misses G.y; end; suppose A77: x in dom G & y in dom G; then reconsider x1=x,y1=y as Nat;::Element of NAT; A78: x1 in Seg len G by A77,FINSEQ_1:def 3; then A79: 1 <= x1 by FINSEQ_1:1; A80: y1 in Seg len G by A77,FINSEQ_1:def 3; then A81: 1 <= y1 by FINSEQ_1:1; A82: y1 <= 2|^n*n + 1 by A75,A80,FINSEQ_1:1; A83: x1 <= 2|^n*n + 1 by A75,A78,FINSEQ_1:1; now per cases by A76,XXREAL_0:1; case A84: x1 < y1; hereby assume A85: y1 = 2|^n*n+1; then A86: G.y = A /\ great_eq_dom(f,n) by A71,A77; A87: x1 <= N*n by A84,A85,NAT_1:13; then A88: G.x = A /\ great_eq_dom(f,((x1-1)/(2 |^n))) /\ less_dom(f,(x1/(2|^n))) by A72,A79; now given v be object such that A89: v in G.x /\ G.y; v in G.y by A89,XBOOLE_0:def 4; then v in great_eq_dom(f,n) by A86,XBOOLE_0:def 4; then A90: n <= f.v by MESFUNC1:def 14; v in G.x by A89,XBOOLE_0:def 4; then v in less_dom(f,(x1/(2|^n))) by A88,XBOOLE_0:def 4; then f.v < (x1/(2|^n)) by MESFUNC1:def 11; then n < (x1/(2|^n)) by A90,XXREAL_0:2; hence contradiction by A87,PREPOWER:6,XREAL_1:79; end; then G.x /\ G.y = {} by XBOOLE_0:def 1; hence G.x misses G.y; end; assume y1 <> 2|^n*n+1; then y1 < N*n+1 by A82,XXREAL_0:1; then A91: y1 <= N*n by NAT_1:13; then x1 <= 2|^n*n by A84,XXREAL_0:2; then A92: G.x = A /\ great_eq_dom(f,((x1-1)/(2|^n)) ) /\ less_dom (f,( x1/(2|^n))) by A72,A79; A93: G.y = A /\ great_eq_dom(f,((y1-1)/(2|^n))) /\ less_dom( f,(y1 /(2|^n))) by A72,A81,A91; now given v be object such that A94: v in G.x /\ G.y; v in G.y by A94,XBOOLE_0:def 4; then v in A /\ great_eq_dom(f,((y1-1)/(2|^n))) by A93, XBOOLE_0:def 4; then v in great_eq_dom(f,((y1-1)/(2|^n))) by XBOOLE_0:def 4; then A95: ((y1-1)/(2|^n)) <= f.v by MESFUNC1:def 14; v in G.x by A94,XBOOLE_0:def 4; then v in less_dom(f,(x1/(2|^n))) by A92,XBOOLE_0:def 4; then f.v < (x1/(2|^n)) by MESFUNC1:def 11; then ((y1-1)/(2|^n)) < (x1/(2|^n)) by A95,XXREAL_0:2; then y1-1 < x1 by XREAL_1:72; then y1 < x1 + 1 by XREAL_1:19; hence contradiction by A84,NAT_1:13; end; then G.x /\ G.y = {} by XBOOLE_0:def 1; hence G.x misses G.y; end; case A96: y1 < x1; hereby assume x1 <> 2|^n*n+1; then x1 < N*n+1 by A83,XXREAL_0:1; then A97: x1 <= N*n by NAT_1:13; then y1 <= 2|^n*n by A96,XXREAL_0:2; then A98: G.y = A /\ great_eq_dom(f,((y1-1)/(2 |^n))) /\ less_dom(f,( y1/(2|^n))) by A72,A81; A99: G.x = A /\ great_eq_dom(f,((x1-1)/(2|^n))) /\ less_dom(f,(x1 /(2|^n))) by A72,A79,A97; now given v be object such that A100: v in G.x /\ G.y; v in G.x by A100,XBOOLE_0:def 4; then v in A /\ great_eq_dom(f,((x1-1)/(2|^n))) by A99, XBOOLE_0:def 4; then v in great_eq_dom(f,((x1-1)/(2|^n))) by XBOOLE_0:def 4; then A101: ((x1-1)/(2|^n)) <= f.v by MESFUNC1:def 14; v in G.y by A100,XBOOLE_0:def 4; then v in less_dom(f,(y1/(2|^n))) by A98,XBOOLE_0:def 4; then f.v < (y1/(2|^n)) by MESFUNC1:def 11; then ((x1-1)/(2|^n)) < (y1/(2|^n)) by A101,XXREAL_0:2; then x1-1 < y1 by XREAL_1:72; then x1 < y1 + 1 by XREAL_1:19; hence contradiction by A96,NAT_1:13; end; then G.x /\ G.y = {} by XBOOLE_0:def 1; hence G.x misses G.y; end; assume A102: x1 = 2|^n*n+1; then A103: G.x = A /\ great_eq_dom(f,n) by A71,A77; A104: y1 <= N*n by A96,A102,NAT_1:13; then A105: G.y = A /\ great_eq_dom(f,((y1-1)/(2|^n)) ) /\ less_dom (f,(y1/(2|^n))) by A72,A81; now given v be object such that A106: v in G.x /\ G.y; v in G.y by A106,XBOOLE_0:def 4; then v in less_dom(f,(y1/(2|^n))) by A105,XBOOLE_0:def 4; then A107: f.v < (y1/(2|^n)) by MESFUNC1:def 11; v in G.x by A106,XBOOLE_0:def 4; then v in great_eq_dom(f,n) by A103,XBOOLE_0:def 4; then n <= f.v by MESFUNC1:def 14; then n < (y1/(2|^n)) by A107,XXREAL_0:2; hence contradiction by A104,PREPOWER:6,XREAL_1:79; end; then G.x /\ G.y = {} by XBOOLE_0:def 1; hence G.x misses G.y; end; end; hence G.x misses G.y; end; end; then reconsider G as Finite_Sep_Sequence of S by PROB_2:def 2; A108: for k be Nat, x,y be 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, x,y be Element of X; assume that A109: k in dom G and A110: x in G.k and A111: y in G.k; A112: 1 <= k by A71,A109,FINSEQ_1:1; A113: k <= N*n+1 by A71,A109,FINSEQ_1:1; now per cases; suppose k = N*n + 1; then A114: G.k = A /\ great_eq_dom(f,n) by A71,A109; then x in great_eq_dom(f,n) by A110,XBOOLE_0:def 4; then A115: n <= f.x by MESFUNC1:def 14; y in great_eq_dom(f,n) by A111,A114,XBOOLE_0:def 4; then A116: n <= f.y by MESFUNC1:def 14; x in A by A110,A114,XBOOLE_0:def 4; then A117: (F.nn).x = nn by A26,A24,A115; y in A by A111,A114,XBOOLE_0:def 4; hence thesis by A26,A24,A117,A116; end; suppose k <> N*n + 1; then k < N*n+1 by A113,XXREAL_0:1; then A118: k <= N*n by NAT_1:13; then A119: G.k = A /\ great_eq_dom(f, ((k-1)/(2|^n))) /\ less_dom(f, (k/(2|^n))) by A71,A109; then x in less_dom(f,(k/(2|^n))) by A110,XBOOLE_0:def 4; then A120: f.x < (k/(2|^n)) by MESFUNC1:def 11; A121: x in A /\ great_eq_dom(f,((k-1)/(2|^n))) by A110,A119, XBOOLE_0:def 4; then x in great_eq_dom(f,((k-1)/(2|^n))) by XBOOLE_0:def 4; then A122: ((k-1)/(2|^n)) <= f.x by MESFUNC1:def 14; x in A by A121,XBOOLE_0:def 4; then A123: (F.n).x = ((k-1)/(2|^n)) by A26,A24,A112,A118,A120,A122; y in less_dom(f,(k/(2|^n))) by A111,A119,XBOOLE_0:def 4; then A124: f.y < (k/(2|^n)) by MESFUNC1:def 11; A125: y in A /\ great_eq_dom(f,((k-1)/(2|^n))) by A111,A119, XBOOLE_0:def 4; then y in great_eq_dom(f,((k-1)/(2|^n))) by XBOOLE_0:def 4; then A126: ((k-1)/(2|^n)) <= f.y by MESFUNC1:def 14; y in A by A125,XBOOLE_0:def 4; hence thesis by A26,A24,A112,A118,A124,A126,A123; end; end; hence thesis; end; for v be object st v in dom f holds v in union rng G proof let v be object; reconsider vv=v as set by TARSKI:1; assume A127: v in dom f; ex B being set st v in B & B in rng G proof per cases; suppose A128: f.v < n; 0 <= f.v by A2,SUPINF_2:51; then consider k be Nat such that A129: 1 <= k and A130: k <= 2|^n*n and A131: (k-1)/(2|^n) <= f.vv and A132: f.vv < k/(2|^n) by A128,Th4; v in great_eq_dom(f,((k-1)/(2|^n))) by A127,A131,MESFUNC1:def 14; then A133: v in A /\ great_eq_dom(f,((k-1)/(2|^n))) by A26,A127, XBOOLE_0:def 4; v in less_dom(f,(k/(2|^n))) by A127,A132,MESFUNC1:def 11; then A134: v in A /\ great_eq_dom(f,((k-1)/(2|^n))) /\ less_dom(f, (k/(2|^n))) by A133,XBOOLE_0:def 4; take G.k; N*n <= N*n + 1 by NAT_1:11; then k <= N*n + 1 by A130,XXREAL_0:2; then k in Seg(N*n+1) by A129; hence thesis by A71,A130,A134,FUNCT_1:3; end; suppose A135: n <= f.v; set k = N*n+1; take G.k; 1 <= k by NAT_1:11; then A136: k in Seg(N*n+1); v in great_eq_dom(f,n) by A127,A135,MESFUNC1:def 14; then v in A /\ great_eq_dom(f,n) by A26,A127,XBOOLE_0:def 4; hence thesis by A71,A136,FUNCT_1:3; end; end; hence thesis by TARSKI:def 4; end; then A137: dom f c= union rng G; for v be object st v in union rng G holds v in dom f proof let v be object; assume v in union rng G; then consider B be set such that A138: v in B and A139: B in rng G by TARSKI:def 4; consider m be object such that A140: m in dom G and A141: B = G.m by A139,FUNCT_1:def 3; reconsider m as Element of NAT by A140; reconsider mm=m as Nat; A142: m <= N*n+1 by A71,A140,FINSEQ_1:1; now per cases; suppose m = N*n+1; then B = A /\ great_eq_dom(f,n) by A71,A140,A141; hence v in A by A138,XBOOLE_0:def 4; end; suppose m <> N*n+1; then m < N*n+1 by A142,XXREAL_0:1; then m <= N*n by NAT_1:13; then B = A /\ great_eq_dom(f,((mm-1)/(2|^n))) /\ less_dom(f, (mm/(2|^n))) by A71,A140,A141; then v in A /\ great_eq_dom(f,((m-1)/(2|^n))) by A138, XBOOLE_0:def 4; hence v in A by XBOOLE_0:def 4; end; end; hence thesis by A26; end; then union rng G c= dom f; then union rng G = dom f by A137; then dom (F.n) = union rng G by A25; hence thesis by A67,A108,MESFUNC2:def 4; end; A143: now let x be Element of X such that A144: x in dom f; per cases; suppose A145: |.f.x.| =+infty; now assume -f.x = +infty; then f.x < 0; hence contradiction by A2,SUPINF_2:51; end; then A146: f.x = +infty by A145,EXTREAL1:13; for g be Real st 0 < g ex n be Nat st for m be Nat st n <= m holds g <= (F#x).m proof let g be Real; assume 0 < g; then reconsider n = [/ g \] as Nat by INT_1:53; A147: g <= n by INT_1:def 7; for m be Nat st n <= m holds g <= (F#x).m proof let m be Nat; assume n <= m; then A148: g <= m by A147,XXREAL_0:2; reconsider m as Element of NAT by ORDINAL1:def 12; m <= f.x by A146,XXREAL_0:4; then (F.m).x = m by A24,A144; hence thesis by A148,Def13; end; hence thesis; end; then A149: F#x is convergent_to_+infty; then F#x is convergent; hence F#x is convergent & lim(F#x) = f.x by A146,A149,Def12; end; suppose |.f.x.| <> +infty; then reconsider g=f.x as Element of REAL by EXTREAL1:30,XXREAL_0:14; A150: for p be Real st 0

=k holds |.(F#x).j - f.x.| < p proof set N2 = [/g\] + 1; let p be Real; A151: g <= [/g\] by INT_1:def 7; [/g\] < [/g\] + 1 by XREAL_1:29; then A152: g < N2 by A151,XXREAL_0:2; 0 <= g by A2,SUPINF_2:51; then reconsider N2 as Element of NAT by A151,INT_1:3; A153: for N be Nat st N >= N2 holds |.(F#x).N - f.x.| < 1/(2|^N) proof let N be Nat; assume A154: N >= N2; reconsider NN=N as Element of NAT by ORDINAL1:def 12; A155: 0 <= f.x by A2,SUPINF_2:51; f.x < N by A152,A154,XXREAL_0:2; then consider m be Nat such that A156: 1 <= m and A157: m <= 2|^N*N and A158: (m-1)/(2|^N) <= f.x and A159: f.x < m/(2|^N) by A155,Th4; reconsider m as Element of NAT by ORDINAL1:def 12; A160: (F#x).N = (F.NN).x by Def13 .= (m-1)/(2|^NN) by A24,A144,A156,A157,A158,A159; then A161: (F#x).N in REAL by XREAL_0:def 1; (m/(2|^N)) - ((m-1)/(2|^N)) = m/(2|^N) - (m-1)/(2|^N ) .= m/(2|^N) + -((m-1)/(2|^N)) .= m/(2|^N) + (-(m-1))/(2|^N) .= (m + -(m-1))/(2|^N); then A162: f.x - (F#x).N < 1/(2|^N) by A159,A160,XXREAL_3:43,A161; -((F#x).N - f.x) = f.x - (F#x).N by XXREAL_3:26; then A163: |.(F#x).N - f.x.| = |.f.x - (F#x).N .| by EXTREAL1:29; 2|^N > 0 by PREPOWER:6; then A164: -(1/(2|^N)) < 0; 0 <= f.x - (F#x).N by A158,A160,XXREAL_3:40; hence thesis by A163,A162,A164,EXTREAL1:22; end; assume 0 < p; then consider N1 be Nat such that A165: (1 qua Complex)/(2|^N1) <= p by PREPOWER:92; reconsider k=max(N2,N1) as Element of NAT by ORDINAL1:def 12; A166: for k be Nat st k >= N1 holds 1/(2|^k) <= p proof let k be Nat; assume k >= N1; then consider i being Nat such that A167: k = (N1 qua Complex) + i by NAT_1:10; (2|^N1)*(2|^i) >= 2|^N1 by PREPOWER:11,XREAL_1:151; then A168: 2|^k >= 2|^N1 by A167,NEWTON:8; 2|^N1 > 0 by PREPOWER:11; then (2|^k)" <= (2|^N1)" by A168,XREAL_1:85; then 1/(2|^k) <= (2|^N1)"; then 1/(2|^k) <= 1/(2|^N1); hence thesis by A165,XXREAL_0:2; end; now let j be Nat; assume A169: j >= k; k >= N2 by XXREAL_0:25; then j >= N2 by A169,XXREAL_0:2; then A170: |.(F#x).j - f.x.| < (1/(2|^j)) by A153; k >= N1 by XXREAL_0:25; then j >= N1 by A169,XXREAL_0:2; then 1/(2|^j) <= p by A166; hence |.(F#x).j - f.x.| < p by A170,XXREAL_0:2; end; hence thesis; end; A171: f.x=g; then A172: F#x is convergent_to_finite_number by A150; then F#x is convergent; hence F#x is convergent & lim(F#x) = f.x by A171,A150,A172,Def12; end; end; for n be Nat holds F.n is nonnegative proof let n be Nat; reconsider nn=n as Element of NAT by ORDINAL1:def 12; now let x be object; assume x in dom (F.n); then A173: x in dom f by A25; per cases; suppose n <= f.x; hence 0 <= (F.nn).x by A24,A173; end; suppose A174: f.x < n; 0 <= f.x by A2,SUPINF_2:51; then consider k be Nat such that A175: 1 <= k and A176: k <= 2|^n*n and A177: (k-1)/(2|^n) <= f.x and A178: f.x < k/(2|^n) by A174,Th4; thus 0 <= (F.nn).x by A24,A173,A175,A176,A177,A178; end; end; hence thesis by SUPINF_2:52; end; hence thesis by A25,A51,A28,A143; end; begin :: Integral of non negative simple valued function 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: integral(M,f) if dom f <> {} otherwise 0.; correctness; end; theorem Th65: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f,g be 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; let S be SigmaField of X; let M be sigma_Measure of S; let f,g be PartFunc of X,ExtREAL; 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; A5: g|dom(f+g) is nonnegative by A4,Th15; A: f|dom(f+g) is nonnegative by A3,Th15; 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; 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; per cases; suppose A16: 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 thesis by A18,A17; end; suppose A19: 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 be 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; assume A23: x in dom(f|dom(f+g) + g|dom(f+g)); 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 thesis by A21,A23,MESFUNC1:def 3; end; integral(M,(f|dom(f+g) + g|dom(f+g))) = integral(M,f|dom(f+g) )+integral(M,g|dom(f+g)) by A10,A12,A13,A14,A15,A19,MESFUNC4:5,A,A5; then A24: integral(M,f+g) = integral(M,f|dom(f+g)) + integral(M,g| dom(f+g)) by A21,A22,PARTFUN1:5; A25: integral(M,g|dom(f+g)) = integral'(M,g|dom(f+g)) by A10,A15,A19,Def14; integral(M,f|dom(f+g)) = integral'(M,f|dom(f+g)) by A10,A14,A19,Def14; hence thesis by A19,A24,A25,Def14; end; end; theorem Th66: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL,c be Real st f is_simple_func_in S & f is nonnegative & 0 <= c holds integral'(M,c(#)f) = (c)*integral'(M,f) proof 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 c be Real; assume that A1: f is_simple_func_in S and A2: f is nonnegative and A3: 0 <= c; set g = c(#)f; A5: dom g = dom f by MESFUNC1:def 6; A6: for x be set st x in dom g holds g.x = ( c)*f.x by MESFUNC1:def 6; per cases; suppose A7: dom g = {}; then integral'(M,f) = 0 by A5,Def14; then c*integral'(M,f) = 0; hence thesis by A7,Def14; end; suppose A8: dom g <> {}; then A9: integral'(M,f) = integral(M,f) by A5,Def14; reconsider cc = c as R_eal by XXREAL_0:def 1; c in REAL by XREAL_0:def 1; then c < +infty by XXREAL_0:9; then integral(M,g) = cc*integral'(M,f) by A1,A3,A5,A2,A6,A8,MESFUNC4:6,A9; hence thesis by A8,Def14; end; end; theorem Th67: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A,B be 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; let S be SigmaField of X; let M be sigma_Measure of S; let f be PartFunc of X,ExtREAL; let A,B be Element of S; assume that A1: f is_simple_func_in S and A2: f is nonnegative and A3: A misses B; set g2 = f|B; set g1 = f|A; set g = f|(A\/B); a4: g is nonnegative by A2,Th15; consider G be Finite_Sep_Sequence of S, b be FinSequence of ExtREAL such that A5: G,b are_Re-presentation_of g and A6: b.1 = 0 and A7: for n be Nat st 2 <= n & n in dom b holds 0 < b.n & b.n < +infty by A1,Th34 ,MESFUNC3:14,a4; deffunc G1(Nat) = G.$1 /\ A; consider G1 be FinSequence such that A8: len G1 = len G & for k be Nat st k in dom G1 holds G1.k = G1(k) from FINSEQ_1:sch 2; A9: dom G1 = Seg len G by A8,FINSEQ_1:def 3; A10: for k be Nat st k in dom G holds G1.k = G.k /\ A proof let k be Nat; assume k in dom G; then k in Seg len G by FINSEQ_1:def 3; hence thesis by A8,A9; end; deffunc G2(Nat) = G.$1 /\ B; consider G2 be FinSequence such that A11: len G2 = len G & for k be Nat st k in dom G2 holds G2.k = G2(k) from FINSEQ_1:sch 2; A12: dom G2 = Seg len G by A11,FINSEQ_1:def 3; A13: for k be Nat st k in dom G holds G2.k = G.k /\ B proof let k be Nat; assume k in dom G; then k in Seg len G by FINSEQ_1:def 3; hence thesis by A11,A12; end; A14: dom G = Seg len G2 by A11,FINSEQ_1:def 3 .= dom G2 by FINSEQ_1:def 3; then reconsider G2 as Finite_Sep_Sequence of S by A13,Th35; A15: dom(g|B) = dom g /\ 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 g2 by RELAT_1:61; for x be object st x in dom(g|B) holds (g|B).x = g2.x proof let x be object; assume A16: x in dom(g|B); then x in dom g /\ B by RELAT_1:61; then A17: x in dom g by XBOOLE_0:def 4; (g|B).x = g.x by A16,FUNCT_1:47 .= f.x by A17,FUNCT_1:47; hence thesis by A15,A16,FUNCT_1:47; end; then g|B = g2 by A15,FUNCT_1:2; then A18: G2,b are_Re-presentation_of g2 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 as Finite_Sep_Sequence of S by A10,Th35; A20: dom(g|A) = dom g /\ 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 g1 by RELAT_1:61; for x be object st x in dom(g|A) holds (g|A).x = g1.x proof let x be object; assume A21: x in dom(g|A); then x in dom g /\ A by RELAT_1:61; then A22: x in dom g by XBOOLE_0:def 4; (g|A).x = g.x by A21,FUNCT_1:47 .= f.x by A22,FUNCT_1:47; hence thesis by A20,A21,FUNCT_1:47; end; then g|A = g1 by A20,FUNCT_1:2; then A23: G1,b are_Re-presentation_of g1 by A5,A19,A10,Th36; deffunc Fy(Nat) = b.$1*(M*G).$1; consider y be FinSequence of ExtREAL such that A24: len y = len G & for j be Nat st j in dom y holds y.j = Fy(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; per cases; suppose A26: dom(f|(A\/B)) = {}; dom f /\ B c= dom f /\ (A\/B) by XBOOLE_1:7,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,g2) = 0 by Def14; dom f /\ A c= dom f /\ (A\/B) by XBOOLE_1:7,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,g1) = 0 by Def14; integral'(M,g) = 0 by A26,Def14; hence thesis by A28,A27; end; suppose A29: dom(f|(A\/B)) <> {}; then integral(M,g) = Sum y by A1,a4,A5,A24,A25,Th34,MESFUNC4:3; then A30: integral'(M,g) = Sum y by A29,Def14; per cases; suppose A31: dom(f|A) = {}; 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 let x be object; assume A33: x in dom g; then g.x = f.x by FUNCT_1:47; hence g.x = g2.x by A32,A33,FUNCT_1:47; end; then A34: g = g2 by A32,FUNCT_1:2; integral'(M,g1) = 0 by A31,Def14; hence thesis by A34,XXREAL_3:4; end; suppose A35: dom(f|A) <> {}; per cases; suppose A36: dom(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 let x be object; assume A38: x in dom g; then g.x = f.x by FUNCT_1:47; hence g.x = g1.x by A37,A38,FUNCT_1:47; end; then A39: g = g1 by A37,FUNCT_1:2; integral'(M,g2) = 0 by A36,Def14; hence thesis by A39,XXREAL_3:4; end; suppose A40: dom(f|B) <> {}; aa: g2 is nonnegative by A2,Th15; deffunc Fy2(Nat) = b.$1*(M*G2).$1; consider y2 be FinSequence of ExtREAL such that A42: len y2 = len G2 & for j be Nat st j in dom y2 holds y2.j = Fy2(j) from FINSEQ_2:sch 1; A43: for k be Nat st k in dom y2 holds 0 <= y2.k proof let k be Nat; assume A44: k in dom y2; 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 per cases; suppose k = 1; hence 0 <= b.k by A6; end; suppose k <> 1; then 1 < k by A45,XXREAL_0:1; then 1+1 <= k by NAT_1:13; hence 0 <= b.k by A7,A44,A46; 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 thesis by A47,A49,A50; end; then for k be object st k in dom y2 holds 0 <= y2.k; then cc: y2 is nonnegative by SUPINF_2:52; A51: for x be set st x in dom y2 holds not y2.x in {-infty} proof let x be set; assume A52: x in dom y2; then reconsider x as Element of NAT; 0 <= y2.x by A43,A52; hence thesis by TARSKI:def 1; end; for x be object holds not x in y2"{-infty} proof let x be object; not (x in dom y2 & y2.x in {-infty}) by A51; hence thesis by FUNCT_1:def 7; 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(M,g2) = Sum y2 by A1,A18,A40,A42,Th34,MESFUNC4:3,aa; then A54: integral'(M,g2) = Sum y2 by A40,Def14; ac: g1 is nonnegative by A2,Th15; deffunc Fy1(Nat) = b.$1*(M*G1).$1; consider y1 be FinSequence of ExtREAL such that A56: len y1 = len G1 & for j be Nat st j in dom y1 holds y1.j = Fy1(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 be Element of NAT st n in dom y holds y.n = y1.n + y2.n proof let n be Element of NAT; assume A59: n in dom y; 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 let v be object; assume A62: v in G.n; 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 g 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; end; then G.n c= A \/ B; 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 assume G1.n meets G2.n; then consider v be object 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; 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 thesis by A42,A64; end; A74: for k be Nat st k in dom y1 holds 0 <= y1.k proof let k be Nat; assume A75: k in dom y1; 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 per cases; suppose k = 1; hence 0 <= b.k by A6; end; suppose k <> 1; then 1 < k by A76,XXREAL_0:1; then 1+1 <= k by NAT_1:13; hence 0 <= b.k by A7,A75,A77; 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 thesis by A78,A80,A81; end; then for x being object st x in dom y1 holds 0. <= y1.x; then ab: y1 is nonnegative by SUPINF_2:52; A82: for x be set st x in dom y1 holds not y1.x in {-infty} proof let x be set; assume A83: x in dom y1; then reconsider x as Element of NAT; 0 <= y1.x by A74,A83; hence thesis by TARSKI:def 1; end; for x be object holds not x in y1"{-infty} proof let x be object; not (x in dom y1 & y1.x in {-infty}) by A82; hence thesis by FUNCT_1:def 7; 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(M,g1) = Sum y1 by A1,A23,A35,A56,Th34,MESFUNC4:3, ac; then A85: integral'(M,g1) = 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 thesis by A30,A85,A54,A84,MESFUNC4:1,ab,cc; end; end; end; end; theorem Th68: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be 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; let S be SigmaField of X; let M be sigma_Measure of S; let f be PartFunc of X,ExtREAL; assume that A1: f is_simple_func_in S and A2: f is nonnegative; per cases; suppose dom f = {}; hence thesis by Def14; end; suppose A4: dom f <> {}; then consider F be Finite_Sep_Sequence of S, a,x be FinSequence of ExtREAL such that A5: F,a are_Re-presentation_of f and A6: dom x = dom F and A7: for n be Nat st n in dom x holds x.n=a.n*(M*F).n and A8: integral(M,f)=Sum x by A1,A2,MESFUNC4:4; A9: for n be Nat st n in dom x holds 0 <= x.n proof let n be Nat; assume A10: n in dom x; per cases; suppose F.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 thesis by A7,A10; end; suppose F.n <> {}; then consider v be object 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 thesis by A7,A10; end; end; integral'(M,f) = integral(M,f) by A4,Def14; hence thesis by A8,A9,Th53; end; end; Lm9: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f, g be 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 be 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(M,f)=integral(M,f-g)+integral(M,g) proof let X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL such 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 be set st x in dom f holds g.x <= f.x; consider G be Finite_Sep_Sequence of S, b,y be FinSequence of ExtREAL such that A9: G,b are_Re-presentation_of g and dom y = dom G and for n be Nat st n in dom y holds y.n=b.n*(M*G).n and integral(M,g)=Sum(y) by A2,A4,A5,A6,MESFUNC4:4; g is real-valued by A4,MESFUNC2:def 4; then A10: dom(f-g) = dom f /\ dom g by MESFUNC2:2; consider F be Finite_Sep_Sequence of S, a,x be FinSequence of ExtREAL such that A12: F,a are_Re-presentation_of f and dom x = dom F and for n be Nat st n in dom x holds x.n=a.n*(M*F).n and integral(M,f)=Sum(x) by A1,A2,MESFUNC4:4,A3; set la = len a; A13: dom F = dom a by A12,MESFUNC3:def 1; set lb = len b; deffunc FG1(Nat) = F.(($1-'1) div lb + 1) /\ G.(($1-'1) mod lb + 1); consider FG be FinSequence such that A14: len FG = la*lb and A15: for k be Nat st k in dom FG holds FG.k=FG1(k) from FINSEQ_1:sch 2; A16: dom FG = Seg(la*lb) 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 object; A18: now let k be Element of NAT; set i=(k -' 1) div lb + 1; set j=(k -' 1) mod lb + 1; A19: lb divides la*lb by NAT_D:def 3; assume A20: k in dom FG; then A21: k in Seg (la*lb) by A14,FINSEQ_1:def 3; then A22: k <= la*lb by FINSEQ_1:1; then (k -' 1) <= (la*lb -' 1) by NAT_D:42; then A23: (k -' 1) div lb <= (la*lb -' 1) div lb by NAT_2:24; 1 <= k by A21,FINSEQ_1:1; then A24: 1 <= la*lb by A22,XXREAL_0:2; A25: lb <> 0 by A21; then (k -' 1) mod lb < lb by NAT_D:1; then A26: j <= lb by NAT_1:13; lb >= 0+1 by A25,NAT_1:13; then ((la*lb) -' 1) div lb = ((la*lb) div lb) - 1 by A19,A24,NAT_2:15; then (k -' 1) div lb + 1 <= la*lb div lb by A23,XREAL_1:19; then A27: i <= la by A25,NAT_D:18; i >= 0+1 by XREAL_1:6; then i in Seg la by A27; then i in dom F by A13,FINSEQ_1:def 3; then A28: F.i in rng F by FUNCT_1:3; j >= 0+1 by XREAL_1:6; then j in Seg lb by A26; then j in dom G by A17,FINSEQ_1:def 3; then A29: G.j in rng G by FUNCT_1:3; FG.k = F.((k -' 1) div lb + 1) /\ G.((k-'1) mod lb + 1) by A15,A20; hence FG.k in S by A28,A29,MEASURE1:34; end; assume b in rng FG; then ex a be object st a in dom FG & b = FG.a by FUNCT_1:def 3; hence thesis by A18; end; then reconsider FG as FinSequence of S; for k,l be Nat st k in dom FG & l in dom FG & k <> l holds FG.k misses FG.l proof let k,l be Nat; assume that A30: k in dom FG and A31: l in dom FG and A32: k <> l; A33: k in Seg (la*lb) by A14,A30,FINSEQ_1:def 3; then A34: 1 <= k by FINSEQ_1:1; set m=(l-'1) mod lb + 1; set n=(l-'1) div lb + 1; set j=(k-'1) mod lb + 1; set i=(k-'1) div lb + 1; A35: lb divides la*lb by NAT_D:def 3; FG.k = F.i /\ G.j by A15,A30; then A36: FG.k /\ FG.l= F.i /\ G.j /\ (F.n /\ G.m) by A15,A31 .= F.i /\ (G.j /\ (F.n /\ G.m)) by XBOOLE_1:16 .= F.i /\ (F.n /\ (G.j /\ G.m)) by XBOOLE_1:16 .= F.i /\ F.n /\ (G.j /\ G.m) by XBOOLE_1:16; A37: k <= la*lb by A33,FINSEQ_1:1; then A38: 1 <= la*lb by A34,XXREAL_0:2; A39: lb <> 0 by A33; then lb >= 0+1 by NAT_1:13; then A40: ((la*lb) -' 1) div lb = ((la*lb) div lb) - 1 by A35,A38,NAT_2:15; k -' 1 <= la*lb -' 1 by A37,NAT_D:42; then (k -' 1) div lb <= (la*lb div lb) - 1 by A40,NAT_2:24; then (k -' 1) div lb + 1 <= la*lb div lb by XREAL_1:19; then A41: i <= la by A39,NAT_D:18; i >= 0+1 by XREAL_1:6; then i in Seg la by A41; then A42: i in dom F by A13,FINSEQ_1:def 3; (l -' 1) mod lb < lb by A39,NAT_D:1; then A43: m <= lb by NAT_1:13; m >= 0+1 by XREAL_1:6; then m in Seg lb by A43; then A44: m in dom G by A17,FINSEQ_1:def 3; (k -' 1) mod lb < lb by A39,NAT_D:1; then A45: j <= lb by NAT_1:13; j >= 0+1 by XREAL_1:6; then j in Seg lb by A45; then A46: j in dom G by A17,FINSEQ_1:def 3; A47: l in Seg (la*lb) by A14,A31,FINSEQ_1:def 3; then A48: 1 <= l by FINSEQ_1:1; A49: now (l-'1)+1=(n-1)*lb+(m-1)+1 by A39,NAT_D:2; then A50: l - 1 + 1 = (n-1)*lb+m by A48,XREAL_1:233; assume that A51: i=n and A52: j=m; (k-'1)+1=(i-1)*lb+(j-1)+1 by A39,NAT_D:2; then k - 1 + 1 = (i-1)*lb + j by A34,XREAL_1:233; hence contradiction by A32,A51,A52,A50; end; l <= la*lb by A47,FINSEQ_1:1; then l -' 1 <= la*lb -' 1 by NAT_D:42; then (l -' 1) div lb <= (la*lb div lb) - 1 by A40,NAT_2:24; then (l -' 1) div lb + 1 <= la*lb div lb by XREAL_1:19; then A53: n <= la by A39,NAT_D:18; n >= 0+1 by XREAL_1:6; then n in Seg la by A53; then A54: n in dom F by A13,FINSEQ_1:def 3; per cases by A49; suppose i <> n; then F.i misses F.n by A42,A54,MESFUNC3:4; then FG.k /\ FG.l= {} /\ (G.j /\ G.m) by A36; hence thesis; end; suppose j <> m; then G.j misses G.m by A46,A44,MESFUNC3:4; then FG.k /\ FG.l= (F.i /\ F.n) /\ {} by A36; hence thesis; end; end; then reconsider FG as Finite_Sep_Sequence of S by MESFUNC3:4; A55: dom f = union rng F by A12,MESFUNC3:def 1; defpred PB[Nat,set] means (G.(($1 -' 1) mod lb + 1) = {} implies $2 = 0) & ( G.(($1 -' 1) mod lb + 1) <> {} implies $2 = b.(($1 -' 1) mod lb + 1)); defpred PA[Nat,set] means (F.(($1 -' 1) div lb + 1) = {} implies $2 = 0) & ( F.(($1 -' 1) div lb + 1) <> {} implies $2 = a.(($1 -' 1) div lb + 1)); A56: for k be Nat st k in Seg len FG ex v be Element of ExtREAL st PA[k,v] proof let k be Nat; assume k in Seg len FG; per cases; suppose A57: F.((k-'1) div lb + 1)={}; take 0.; thus thesis by A57; end; suppose A58: F.((k-'1) div lb + 1)<>{}; take a.((k-'1) div lb + 1); thus thesis by A58; end; end; consider a1 be FinSequence of ExtREAL such that A59: dom a1 = Seg len FG & for k be Nat st k in Seg len FG holds PA[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 proof let z be object; assume A62: z in dom f; then consider Y be set such that A63: z in Y and A64: Y in rng F by A55,TARSKI:def 4; consider i be 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 qua Complex) + i9 by NAT_1:10; consider Z be set such that A69: z in Z and A70: Z in rng G by A5,A60,A62,TARSKI:def 4; consider j be 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 qua Complex) + j9 by NAT_1:10; i9*lb + j in NAT by ORDINAL1:def 12; then reconsider k=(i-1)*lb+j as Element of NAT by A68; i <= la by A67,FINSEQ_1:1; then i-1 <= la-1 by XREAL_1:9; then (i-1)*lb <= (la - 1)*lb by XREAL_1:64; then A76: k <= (la - 1) * lb + 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*lb + j9 by A68,A75; A79: j <= lb by A73,FINSEQ_1:1; then (la - 1) * lb + j <= (la - 1) * lb + lb by XREAL_1:6; then A80: k <= la*lb by A76,XXREAL_0:2; k >= 1 by A74,A77,XXREAL_0:2; then A81: k in Seg (la*lb) 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 < lb by A79,A75,NAT_1:13; then A84: j = (k-'1) mod lb +1 by A75,A78,NAT_D:def 2; A85: i = (k-'1) div lb +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 thesis by A82,TARSKI:def 4; end; let z be object; A86: lb divides la*lb by NAT_D:def 3; assume z in union rng FG; then consider Y be set such that A87: z in Y and A88: Y in rng FG by TARSKI:def 4; consider k be Nat such that A89: k in dom FG and A90: Y = FG.k by A88,FINSEQ_2:10; set i=(k-'1) div lb +1; A91: k in Seg len FG by A89,FINSEQ_1:def 3; then A92: k <= la*lb by A14,FINSEQ_1:1; then A93: (k -' 1) <= (la*lb -' 1) by NAT_D:42; 1 <= k by A91,FINSEQ_1:1; then A94: 1 <= la*lb by A92,XXREAL_0:2; A95: lb <> 0 by A14,A91; then lb >= 0+1 by NAT_1:13; then ((la*lb) -' 1) div lb = ((la*lb) div lb) - 1 by A86,A94,NAT_2:15; then (k -' 1) div lb <= (la*lb div lb) - 1 by A93,NAT_2:24; then A96: i <= la*lb div lb by XREAL_1:19; set j=(k-'1) mod lb +1; A97: i >= 0+1 by XREAL_1:6; la*lb div lb = la by A95,NAT_D:18; then i in Seg la by A97,A96; then i in dom F by A13,FINSEQ_1:def 3; then A98: F.i in rng F by FUNCT_1:def 3; FG.k=F.i /\ G.j by A15,A89; then z in F.i by A87,A90,XBOOLE_0:def 4; hence thesis by A55,A98,TARSKI:def 4; end; A99: for k being Nat,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; let x,y be Element of X; assume that A100: k in dom FG and A101: x in FG.k and A102: y in FG.k; set j=(k-'1) mod lb + 1; A103: FG.k = F.( (k-'1) div lb + 1 ) /\ G.( (k-'1) mod lb + 1) by A15,A100; then A104: x in G.j by A101,XBOOLE_0:def 4; set i=(k-'1) div lb + 1; A105: i >= 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: lb > 0 by A14,A106; then A109: lb >= 0+1 by NAT_1:13; A110: y in G.j by A102,A103,XBOOLE_0:def 4; A111: lb divides la*lb by NAT_D:def 3; A112: k <= la*lb by A14,A106,FINSEQ_1:1; then A113: (k -' 1) <= (la*lb -' 1) by NAT_D:42; 1 <= la*lb by A107,A112,XXREAL_0:2; then ((la*lb) -' 1) div lb = ((la*lb) div lb) - 1 by A109,A111,NAT_2:15; then (k -' 1) div lb <= (la*lb div lb) - 1 by A113,NAT_2:24; then A114: (k -' 1) div lb + 1 <= la*lb div lb by XREAL_1:19; lb <> 0 by A14,A106; then i <= la by A114,NAT_D:18; then i in Seg la by A105; then A115: i in dom F by A13,FINSEQ_1:def 3; (k -' 1) mod lb < lb by A108,NAT_D:1; then A116: j <= lb by NAT_1:13; j >= 0+1 by XREAL_1:6; then j in Seg lb by A116; then A117: j in dom G by A17,FINSEQ_1:def 3; y in F.i by A102,A103,XBOOLE_0:def 4; then A118: f.y=a.i by A12,A115,MESFUNC3:def 1; x in F.i by A101,A103,XBOOLE_0:def 4; then A119: f.x=a.i 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.i-b.j by A9,A117,A104,A119,MESFUNC3:def 1 .= f.y-g.y by A9,A117,A110,A118,MESFUNC3:def 1; hence thesis by A121,MESFUNC1:def 4; end; deffunc X1(Nat) = a1.$1*(M*FG).$1; consider x1 be FinSequence of ExtREAL such that A122: len x1 = len FG & for k be Nat st k in dom x1 holds x1.k=X1(k) from FINSEQ_2:sch 1; A123: for k be Nat st k in dom FG for x be object st x in FG.k holds f.x=a1.k proof let k be Nat; set i=(k-'1) div lb + 1; A124: i >= 0+1 by XREAL_1:6; assume A125: k in dom FG; then A126: k in Seg len FG by FINSEQ_1:def 3; let x be object; assume A127: x in FG.k; FG.k = F.((k-'1) div lb + 1) /\ G.((k-'1) mod lb + 1) by A15,A125; then A128: x in F.i by A127,XBOOLE_0:def 4; A129: k in Seg len FG by A125,FINSEQ_1:def 3; then A130: k <= la*lb by A14,FINSEQ_1:1; then (k -' 1) <= (la*lb -' 1) by NAT_D:42; then A131: (k -' 1) div lb <= (la*lb -' 1) div lb by NAT_2:24; A132: lb divides la*lb by NAT_D:def 3; 1 <= k by A129,FINSEQ_1:1; then A133: 1 <= la*lb by A130,XXREAL_0:2; A134: lb <> 0 by A14,A129; then lb >= 0+1 by NAT_1:13; then ((la*lb) -' 1) div lb = ((la*lb) div lb) - 1 by A132,A133,NAT_2:15; then A135: i <= la*lb div lb by A131,XREAL_1:19; la*lb div lb = la by A134,NAT_D:18; then i in Seg la by A124,A135; then i in dom F by A13,FINSEQ_1:def 3; then f.x = a.i by A12,A128,MESFUNC3:def 1; hence thesis by A59,A126,A128; end; A136: for k be Nat st k in Seg len FG ex v be Element of ExtREAL st PB[k,v] proof let k be Nat; assume k in Seg len FG; per cases; suppose A137: G.((k-'1) mod lb + 1)={}; reconsider z = 0 as R_eal by XXREAL_0:def 1; take z; thus thesis by A137; end; suppose A138: G.((k-'1) mod lb + 1)<>{}; take b.((k-'1) mod lb + 1); thus thesis by A138; end; end; consider b1 be FinSequence of ExtREAL such that A139: dom b1 = Seg len FG & for k be Nat st k in Seg len FG holds PB[k, b1.k] from FINSEQ_1:sch 5(A136); deffunc C1(Nat) = a1.$1-b1.$1; consider c1 be FinSequence of ExtREAL such that A140: len c1 = len FG and A141: for k be Nat st k in dom c1 holds c1.k=C1(k) from FINSEQ_2:sch 1; A142: dom c1 = Seg len FG by A140,FINSEQ_1:def 3; A143: for k be Nat st k in dom FG for x be object st x in FG.k holds g.x=b1.k proof let k be Nat; set j=(k-'1) mod lb + 1; assume A144: k in dom FG; then A145: k in Seg len FG by FINSEQ_1:def 3; k in Seg len FG by A144,FINSEQ_1:def 3; then lb <> 0 by A14; then (k -' 1) mod lb < lb by NAT_D:1; then A146: j <= lb by NAT_1:13; let x be object; assume A147: x in FG.k; FG.k = F.( (k-'1) div lb + 1 ) /\ G.((k-'1) mod lb + 1) by A15,A144; then A148: x in G.j by A147,XBOOLE_0:def 4; j >= 0+1 by XREAL_1:6; then j in Seg lb by A146; then j in dom G by A17,FINSEQ_1:def 3; hence g.x=b.j by A9,A148,MESFUNC3:def 1 .=b1.k by A139,A148,A145; end; A149: for k be Nat st k in dom FG for x be object st x in FG.k holds (f-g).x=c1 .k proof let k be Nat; assume A150: k in dom FG; let x be object; assume A151: x in FG.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 thesis by A123,A150,A151,A152; end; deffunc Z1(Nat) = c1.$1*(M*FG).$1; consider z1 be FinSequence of ExtREAL such that A153: len z1 = len FG & for k be Nat st k in dom z1 holds z1.k=Z1(k) from FINSEQ_2:sch 1; deffunc Y1(Nat) = b1.$1*(M*FG).$1; consider y1 be FinSequence of ExtREAL such that A154: len y1 = len FG & for k be Nat st k in dom y1 holds y1.k=Y1(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 be 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; assume A158: i in dom x1; 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; per cases; suppose FG.i <> {}; then consider x be object 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 thesis by A155,A153,A156,A158; end; suppose FG.i = {}; then (M*FG).i = 0 by A159,VALUED_0:def 19; then c1.i * (M*FG).i = 0; hence thesis by A155,A153,A156,A158; end; end; then for i be object st i in dom z1 holds 0 <= z1.i by A156,A155; then cd: z1 is nonnegative by SUPINF_2:52; not -infty in rng z1 proof assume -infty in rng z1; then ex i be object st i in dom z1 & z1.i = -infty by FUNCT_1:def 3; hence contradiction by A155,A156,A157; 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 be Nat st i in dom y1 holds 0 <= y1.i proof let i be Nat; set i9 = (i -' 1) mod lb + 1; A166: i9 >= 0+1 by XREAL_1:6; assume A167: i in dom y1; 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 lb <> 0 by A14; then (i -' 1) mod lb < lb by NAT_D:1; then i9 <= lb by NAT_1:13; then i9 in Seg lb by A166; then A170: i9 in dom G by A17,FINSEQ_1:def 3; per cases; suppose A171: G.i9 <> {}; 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; then A172: M.({}) <= M.FGi by MEASURE1:31; consider x9 be object such that A173: x9 in G.i9 by A171,XBOOLE_0:def 1; g.x9 = b.i9 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 thesis by A168,A174; end; suppose A175: G.i9 = {}; FG.i = F.((i-'1) div lb + 1) /\ G.i9 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 thesis by A168; end; end; then for i be object st i in dom y1 holds 0 <= y1.i; then ag: y1 is nonnegative by SUPINF_2:52; not -infty in rng y1 proof assume -infty in rng y1; then ex i be object st i in dom y1 & y1.i = -infty by FUNCT_1:def 3; hence contradiction by A165; 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 be Nat st k in dom x1 holds x1.k = (z1+y1).k proof A178: lb divides la*lb by NAT_D:def 3; let k be Nat; set p=(k-'1) div lb + 1; set q=(k-'1) mod lb + 1; A179: p >= 0+1 by XREAL_1:6; assume A180: k in dom x1; then A181: k in Seg len FG by A122,FINSEQ_1:def 3; then A182: 1 <= k by FINSEQ_1:1; A183: lb > 0 by A14,A181; then A184: lb >= 0+1 by NAT_1:13; A185: k <= la*lb by A14,A181,FINSEQ_1:1; then A186: (k -' 1) <= (la*lb -' 1) by NAT_D:42; 1 <= la*lb by A182,A185,XXREAL_0:2; then ((la*lb) -' 1) div lb = ((la*lb) div lb) - 1 by A184,A178,NAT_2:15; then (k -' 1) div lb <= (la*lb div lb) - 1 by A186,NAT_2:24; then A187: p <= la*lb div lb by XREAL_1:19; lb <> 0 by A14,A181; then p <= la by A187,NAT_D:18; then p in Seg la by A179; then A188: p in dom F by A13,FINSEQ_1:def 3; A189: q >= 0+1 by XREAL_1:6; (k -' 1) mod lb < lb by A183,NAT_D:1; then q <= lb by NAT_1:13; then q in Seg lb by A189; then A190: q 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 per cases; suppose FG.k <> {}; then F.p /\ G.q <> {} by A14,A15,A16,A181; then consider v be object such that A192: v in F.p /\ G.q by XBOOLE_0:def 1; A193: G.q <> {} by A192; A194: v in F.p by A192,XBOOLE_0:def 4; v in G.q by A192,XBOOLE_0:def 4; then A195: b.q = g.v by A9,A190,MESFUNC3:def 1; F.p in rng F by A188,FUNCT_1:3; then A196: v in dom f by A55,A194,TARSKI:def 4; a.p = f.v by A12,A188,A194,MESFUNC3:def 1; then b.q <= a.p by A7,A195,A196; then A197: b1.k <= a.p by A139,A181,A193; F.p <> {} 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.q by A6,A195,SUPINF_2:51; then 0 = b1.k or 0 < b1.k by A139,A181,A192; hence thesis by A198,XXREAL_3:96; end; suppose 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; end; end; A200: a1.k <> +infty & a1.k <> -infty & b1.k <> +infty & b1.k <> -infty proof now per cases; suppose A201: F.p <> {}; then consider v be object such that A202: v in F.p by XBOOLE_0:def 1; A203: f is real-valued by A1,MESFUNC2:def 4; a1.k = a.p 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; end; suppose F.p = {}; hence a1.k <> +infty & -infty <> a1.k by A59,A181; end; end; hence +infty <> a1.k & a1.k <> -infty; now per cases; suppose A204: G.q <> {}; then consider v be object such that A205: v in G.q by XBOOLE_0:def 1; A206: g is real-valued by A4,MESFUNC2:def 4; b1.k = b.q by A139,A181,A204; then b1.k = g.v by A9,A190,A205,MESFUNC3:def 1; hence thesis by A206; end; suppose G.q = {}; hence thesis by A139,A181; end; end; hence thesis; 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; end; now let x be Element of X; assume A208: x in dom (f-g); g is real-valued by A4,MESFUNC2:def 4; then A209: |. g.x .| < +infty by A5,A10,A208,MESFUNC2:def 1; f is real-valued 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 EXTREAL1:32; hence |. (f-g).x .| < +infty by A210,XXREAL_0:2,4; end; then f-g is real-valued by MESFUNC2:def 1; hence A211: f-g is_simple_func_in S by A5,A61,A10,A99,MESFUNC2:def 4; 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(M,f)=Sum x1 by A1,A2,A3,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(M,g)=Sum y1 by A2,A4,A5,A6,A154,A164,MESFUNC4:3; thus dom (f-g) <> {} by A2,A5,A10; for x be object st x in dom (f-g) holds 0 <= (f-g).x proof let x be object; assume A216: x in dom (f-g); then 0 <= f.x - g.x by A5,A7,A10,XXREAL_3:40; hence thesis by A216,MESFUNC1:def 4; end; hence aa:f-g is nonnegative by SUPINF_2:52; 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(M,f-g)=Sum z1 by aa,A2,A5,A153,A156,A10,A211,MESFUNC4:3; hence thesis by A164,A156,A212,A214,A213,MESFUNC4:1,ag,cd; end; theorem Th69: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f,g be 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 be object 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, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL such 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 be object st x in dom(f-g) holds g.x <= f.x; A6: f|dom(f-g) is nonnegative by A2,Th15; (-jj)(#)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 without-infty by A3,Th14; then not -infty in rng g; then A10: g"{-infty} = {} by FUNCT_1:72; f is without+infty by A1,Th14; then not +infty in rng f; 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; 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 be 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; assume A16: x in dom(f|dom(f-g)); 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 thesis by A13,A14,A16,FUNCT_1:49; 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)) proof per cases; suppose A19: 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 thesis by A21,A20; end; suppose A22: 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 be 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; assume A26: x in dom(f|dom(f-g) - g|dom(f-g)); 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 thesis by A24,A26,MESFUNC1:def 4; end; integral(M,f|dom(f-g)) = integral(M,(f|dom(f-g) - g| dom(f- g))) +integral(M,g|dom(f-g)) by A13,A18,A8,A6,A9,A14,A17,A15,A22,Lm9; then A27: integral(M,f|dom(f-g)) = integral(M,f-g) + integral(M,g |dom(f-g)) by A24,A25,PARTFUN1:5; A28: integral(M,g|dom(f-g)) = integral'(M,g|dom(f-g)) by A13,A17,A22,Def14 ; integral(M,f|dom(f-g)) = integral'(M,f|dom(f-g)) by A13,A14,A22,Def14 ; hence thesis by A22,A27,A28,Def14; end; end; end; theorem Th70: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f,g be 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 be object 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; let S be SigmaField of X; let M be sigma_Measure of S; let f,g be PartFunc of X,ExtREAL; 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 be object st x in dom(f-g) holds g.x <= f.x; (-jj)(#)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 assume integral'(M,f|dom(f-g)) <> +infty; 0 <= integral'(M,f-g) by A1,A2,A5,A6,Th40,Th68; hence thesis by A7,XXREAL_3:39; end; hence thesis by XXREAL_0:4; end; theorem Th71: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, c be R_eal st 0 <= c & f is_simple_func_in S & (for x be object st x in dom f holds f.x=c) holds integral'(M,f) = c*(M.(dom f)) proof 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 c be R_eal; assume that A1: 0 <= c and A2: f is_simple_func_in S and A3: for x be object st x in dom f holds f.x = c; for x be object st x in dom f holds 0 <= f.x by A1,A3; then a4: f is nonnegative by SUPINF_2:52; reconsider A = dom f as Element of S by A2,Th37; per cases; suppose A5: dom f = {}; then A6: M.A = 0 by VALUED_0:def 19; integral'(M,f) = 0 by A5,Def14; hence thesis by A6; end; suppose A7: dom f <> {}; set x = <* c * M.A *>; reconsider a = <* c *> as FinSequence of ExtREAL; set F = <* dom f *>; reconsider x as FinSequence of ExtREAL; A8: rng F = {A} by FINSEQ_1:38; rng F c= S proof let z be object; assume z in rng F; then z = A by A8,TARSKI:def 1; hence thesis; end; then reconsider F as FinSequence of S by FINSEQ_1:def 4; for i,j be Nat st i in dom F & j in dom F & i <> j holds F.i misses F .j proof let i,j be Nat; assume that A9: i in dom F and A10: j in dom F and A11: i <> j; A12: dom F = {1} by FINSEQ_1:2,38; then i = 1 by A9,TARSKI:def 1; hence thesis by A10,A11,A12,TARSKI:def 1; end; then reconsider 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 be Nat st n in dom F for x be object st x in F.n holds f.x = a.n proof let n be Nat; assume n in dom F; then n in {1} by FINSEQ_1:2,38; then A15: n = 1 by TARSKI:def 1; let x be object; assume x in F.n; then x in dom f by A15,FINSEQ_1:40; then f.x = c by A3; hence thesis by A15,FINSEQ_1:40; end; A16: for n be Nat st n in dom x holds x.n = c*M.A proof let n be Nat; assume n in dom x; then n in {1} by FINSEQ_1:2,38; then n = 1 by TARSKI:def 1; hence thesis by FINSEQ_1:40; end; A17: dom x = Seg 1 by FINSEQ_1:38 .= dom F by FINSEQ_1:38; A18: for n be Nat st n in dom x holds x.n = a.n*(M*F).n proof let n be Nat; assume A19: n in dom x; then n in {1} by FINSEQ_1:2,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 thesis by A20,A21,FINSEQ_1:40; 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(M,f) = Sum x by A2,a4,A7,A17,A18,MESFUNC4:3; then A22: integral'(M,f) = Sum x by A7,Def14; reconsider j = 1 as R_eal by XXREAL_0:def 1; 1 = len x by FINSEQ_1:40; then Sum x = j *(c*M.A) by A16,MESFUNC3:18; hence thesis by A22,XXREAL_3:81; end; end; theorem Th72: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL st f is_simple_func_in S & f is nonnegative holds integral'(M,f|eq_dom(f,0)) = 0 proof 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: f is_simple_func_in S and A2: f is nonnegative; set A = dom f; set g = f|(A /\ eq_dom(f,0)); for x be object st x in eq_dom(f,0) holds x in A by MESFUNC1:def 15; then eq_dom(f,0) c= A; then A3: f|(A/\eq_dom(f,0)) = f|eq_dom(f,0) by XBOOLE_1:28; A4: ex G be Finite_Sep_Sequence of S st (dom g = union rng G & for n be Nat, x,y be Element of X st n in dom G & x in G.n & y in G.n holds g.x = g.y) proof consider F be Finite_Sep_Sequence of S such that A5: dom f = union rng F and A6: for n be Nat, x,y be 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 G(Nat) = F.$1 /\ (A/\eq_dom(f,0)); reconsider A as Element of S by A5,MESFUNC2:31; consider G be FinSequence such that A7: len G = len F & for n be Nat st n in dom G holds G.n = G(n) from FINSEQ_1:sch 2; f is_measurable_on A by A1,MESFUNC2:34; then A /\ less_dom(f,0) in S by MESFUNC1:def 16; then A\(A /\ less_dom(f,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,0))/\less_eq_dom(f,0) in S by MESFUNC1:28; then reconsider A2 = A /\ eq_dom(f,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 be 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 as Finite_Sep_Sequence of S by A9,Th35; take G; for i be 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 g by RELAT_1:61; for i be Nat, x,y be Element of X st i in dom G & x in G.i & y in G.i holds g.x = g.y proof let i be Nat; let x,y be Element of X; assume that A12: i in dom G and A13: x in G.i and A14: y in G.i; 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 g by A11,A13,TARSKI:def 4; then A18: g.x = f.x by FUNCT_1:47; y in dom g by A11,A14,A17,TARSKI:def 4; then A19: g.y = f.y by FUNCT_1:47; x in F.i by A13,A15,XBOOLE_0:def 4; hence thesis by A6,A10,A12,A16,A18,A19; end; hence thesis by A11; end; for x be object st x in dom g holds 0 <= g.x proof let x be object; assume A21: x in dom g; 0 <= f.x by A2,SUPINF_2:51; hence thesis by A21,FUNCT_1:47; end; then a2: g is nonnegative by SUPINF_2:52; f is real-valued by A1,MESFUNC2:def 4; then A22: g is_simple_func_in S by A4,MESFUNC2:def 4; now consider F be Finite_Sep_Sequence of S, a,x be FinSequence of ExtREAL such that A23: F,a are_Re-presentation_of g and a.1 =0 and for n be 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 be Nat st n in dom x holds x.n=a.n*(M*F).n and A26: integral(M,g)=Sum(x) by a2,A22,MESFUNC3:def 2; A27: for x be set st x in dom g holds g.x = 0 proof let x be set; assume A28: x in dom g; then x in dom f /\ (A /\ eq_dom(f,0)) by RELAT_1:61; then x in A /\ eq_dom(f,0) by XBOOLE_0:def 4; then x in eq_dom(f,0) by XBOOLE_0:def 4; then 0 = f.x by MESFUNC1:def 15; hence thesis by A28,FUNCT_1:47; end; A29: for n be Nat st n in dom F holds a.n = 0 or F.n = {} proof let n be Nat; assume A30: n in dom F; now assume F.n <> {}; then consider x be object 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 g by A23,MESFUNC3:def 1; then g.x = 0 by A27; hence thesis by A23,A30,A31,MESFUNC3:def 1; end; hence thesis; end; A32: for n be Nat st n in dom x holds x.n = 0 proof let n be Nat; assume A33: n in dom x; per cases by A24,A29,A33; suppose a.n = 0; then a.n*(M*F).n = 0; hence thesis by A25,A33; end; suppose F.n = {}; 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 thesis by A25,A33; end; end; A34: Sum x = 0 proof consider sumx be sequence of ExtREAL such that A35: Sum x = sumx.(len x) and A36: sumx.0 = 0 and A37: for i be Nat st i < len x holds sumx.(i+1)=sumx.i + x.(i +1) by EXTREAL1:def 2; now defpred P[Nat] means $1 <= len x implies sumx.$1 = 0; assume x <> {}; A38: for k be Nat st P[k] holds P[k+1] proof let k be Nat; assume A39: P[k]; assume A40: k+1 <= len x; reconsider 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 thesis by A39,A40,A41,NAT_1:13; end; A42: P[ 0 ] by A36; for i be Nat holds P[i] from NAT_1:sch 2(A42,A38); hence thesis by A35; end; hence thesis by A35,A36,CARD_1:27; end; assume dom g <> {}; hence thesis by A3,A26,A34,Def14; end; hence thesis by A3,Def14; end; theorem Th73: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, B be Element of S, f be 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; let S be SigmaField of X; let M be sigma_Measure of S; let B be Element of S; let f be PartFunc of X,ExtREAL; assume that A1: f is_simple_func_in S and A2: M.B=0 and A3: f is nonnegative; set A = dom f; set g = f|(A/\B); for x be object st x in dom g holds 0 <= g.x proof let x be object; assume A5: x in dom g; 0 <= f.x by A3,SUPINF_2:51; hence thesis by A5,FUNCT_1:47; end; then a4: g is nonnegative by SUPINF_2:52; A6: ex G be Finite_Sep_Sequence of S st (dom g = union rng G & for n be Nat, x,y be Element of X st n in dom G & x in G.n & y in G.n holds g.x = g.y) proof consider F be Finite_Sep_Sequence of S such that A7: dom f = union rng F and A8: for n be Nat, x,y be 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 G(Nat) = F.$1 /\ (A/\ B); reconsider A as Element of S by A7,MESFUNC2:31; reconsider A2 = A/\B as Element of S; consider G be FinSequence such that A9: len G = len F & for n be Nat st n in dom G holds G.n = G(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 be 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 as Finite_Sep_Sequence of S by A11,Th35; take G; for i be 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 g by RELAT_1:61; for i be Nat, x,y be Element of X st i in dom G & x in G.i & y in G.i holds g.x = g.y proof let i be Nat; let x,y be Element of X; assume that A14: i in dom G and A15: x in G.i and A16: y in G.i; 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 g by A13,A15,TARSKI:def 4; then A20: g.x = f.x by FUNCT_1:47; y in dom g by A13,A16,A19,TARSKI:def 4; then A21: g.y = f.y by FUNCT_1:47; x in F.i by A15,A17,XBOOLE_0:def 4; hence thesis by A8,A12,A14,A18,A20,A21; end; hence thesis by A13; end; dom(f|(A/\B)) = A/\(A/\B) by RELAT_1:61; then A22: dom(f|(A/\B)) = (A/\A)/\B by XBOOLE_1:16; then A23: dom(f|(A/\B)) = dom(f|B) by RELAT_1:61; for x be object st x in dom(f|(A/\B)) holds (f|(A/\B)).x = (f|B).x proof let x be object; assume A24: x in dom(f|(A/\B)); then (f|(A/\B)).x = f.x by FUNCT_1:47; hence thesis by A23,A24,FUNCT_1:47; end; then A25: f|(A/\B) = f|B by A23,FUNCT_1:2; f is real-valued by A1,MESFUNC2:def 4; then A26: g is_simple_func_in S by A6,MESFUNC2:def 4; now per cases; suppose dom g = {}; hence thesis by A23,Def14; end; suppose A27: dom g <> {}; consider F be Finite_Sep_Sequence of S, a,x be FinSequence of ExtREAL such that A28: F,a are_Re-presentation_of g and a.1 =0 and for n be 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 be Nat st n in dom x holds x.n=a.n*(M*F).n and A31: integral(M,g)=Sum(x) by A26,MESFUNC3:def 2,a4; A32: for n be 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; A33: dom g c= B by A22,XBOOLE_1:17; assume A34: n in dom F; then F.n in rng F by FUNCT_1:3; then reconsider FF=F.n as Element of S; for v be object st v in F.n holds v in union rng F proof let v be object; assume A35: v in F.n; F.n in rng F by A34,FUNCT_1:3; hence thesis by A35,TARSKI:def 4; end; then A36: F.n c= union rng F; union rng F = dom g by A28,MESFUNC3:def 1; then FF c= BB by A36,A33; then F.n is measure_zero of M by MEASURE1:36; hence thesis by MEASURE1:def 7; end; A37: for n be Nat st n in dom x holds x.n = 0 proof let n be Nat; assume A38: n in dom x; 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 thesis by A30,A38; end; Sum(x) = 0 proof consider sumx be sequence of ExtREAL such that A39: Sum(x) = sumx.(len x) and A40: sumx.0 = 0 and A41: for i be Nat st i < len x holds sumx.(i+1)=sumx.i + x.(i +1) by EXTREAL1:def 2; now defpred P[Nat] means $1 <= len x implies sumx.$1 = 0; assume x <> {}; A42: for k be Nat st P[k] holds P[k+1] proof let k be Nat; assume A43: P[k]; assume A44: k+1 <= len x; reconsider 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 thesis by A43,A44,A45,NAT_1:13; end; A46: P[ 0 ] by A40; for i be Nat holds P[i] from NAT_1:sch 2(A46,A42); hence thesis by A39; end; hence thesis by A39,A40,CARD_1:27; end; hence thesis by A25,A27,A31,Def14; end; end; hence thesis; end; theorem Th74: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, g be PartFunc of X,ExtREAL, F be Functional_Sequence of X,ExtREAL, L be ExtREAL_sequence st g is_simple_func_in S & (for x be object st x in dom g holds 0 < g.x) & (for n be Nat holds F.n is_simple_func_in S) & (for n be Nat holds dom (F.n) = dom g) & (for n be Nat holds F.n is nonnegative) & (for n,m be Nat st n <=m holds for x be Element of X st x in dom g holds (F.n).x <= (F.m).x ) & (for x be Element of X st x in dom g holds (F#x) is convergent & g.x <= lim(F#x) ) & (for n be Nat holds L.n = integral'(M,F.n)) holds L is convergent & integral'(M ,g) <= lim(L) proof let X be non empty set; let S be SigmaField of X; let M be sigma_Measure of S; let g be PartFunc of X,ExtREAL; let F be Functional_Sequence of X,ExtREAL; let L be ExtREAL_sequence; assume that A1: g is_simple_func_in S and A2: for x be object st x in dom g holds 0 < g.x and A3: for n be Nat holds F.n is_simple_func_in S and A4: for n be Nat holds dom(F.n) = dom g and A5: for n be Nat holds F.n is nonnegative and A6: for n,m be Nat st n <= m holds for x be Element of X st x in dom g holds (F.n).x <= (F.m).x and A7: for x be Element of X st x in dom g holds (F#x) is convergent & g.x <= lim(F#x) and A8: for n be Nat holds L.n = integral'(M,F.n); per cases; suppose A9: dom g = {}; A10: now let n be Nat; dom(F.n) = {} by A4,A9; then integral'(M,F.n) = 0 by Def14; hence L.n = 0 by A8; end; then L is convergent_to_finite_number by Th52; hence L is convergent; lim L = 0 by A10,Th52; hence thesis by A9,Def14; end; suppose A11: dom g <> {}; for v be object st v in dom g holds 0 <= g.v by A2; then g is nonnegative by SUPINF_2:52; then consider G be Finite_Sep_Sequence of S, a be FinSequence of ExtREAL such that A12: G,a are_Re-presentation_of g and A13: a.1 = 0 and A14: for n be Nat st 2 <= n & n in dom a holds 0 < a.n & a.n < +infty by A1,MESFUNC3:14; defpred PP1[Nat,set] means $2 = a.$1; A15: for k be Nat st k in Seg len a ex x be Element of REAL st PP1[k,x] proof let k be Nat; assume A16: k in Seg len a; then A17: 1 <= k by FINSEQ_1:1; A18: k in dom a by A16,FINSEQ_1:def 3; per cases; suppose A19: k = 1; take In(0,REAL); thus thesis by A13,A19; end; suppose k <> 1; 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; thus thesis; end; end; consider a1 be FinSequence of REAL such that A22: dom a1 = Seg len a & for k be Nat st k in Seg len a holds PP1[k, a1.k] from FINSEQ_1:sch 5(A15); A23: len a <> 0 proof assume len a = 0; then A24: dom a = Seg 0 by FINSEQ_1:def 3; A25: rng G = {} proof assume rng G <> {}; then consider y be object such that A26: y in rng G by XBOOLE_0:def 1; ex x be object st x in dom G & y=G.x by A26,FUNCT_1:def 3; hence contradiction by A12,A24,MESFUNC3:def 1; end; union rng G <> {} by A11,A12,MESFUNC3:def 1; then consider x be object such that A27: x in union rng G by XBOOLE_0:def 1; ex Y be set st x in Y & Y in rng G by A27,TARSKI:def 4; hence contradiction by A25; end; A28: 2 <= len a proof assume not 2<=len a; then len a = 1 by A23,NAT_1:23; then dom a = {1} by FINSEQ_1:2,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 be object 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; 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 finite non empty real-membered set by A35,XBOOLE_0:def 5; reconsider alpha = min RINGA as R_eal by XXREAL_0:def 1; reconsider beta1=max RINGA as Element of 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 be object such that A37: i in dom a1 and A38: min RINGA = a1.i by FUNCT_1:def 3; reconsider i as Element of NAT by A37; A39: a.i = 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 < alpha by A14,A38,A41,A39; reconsider beta = max RINGA as R_eal by XXREAL_0:def 1; A44: for x be set st x in dom g holds alpha <= g.x & g.x <= beta proof let x be set; assume A45: x in dom g; then x in union rng G by A12,MESFUNC3:def 1; then consider Y be set such that A46: x in Y and A47: Y in rng G by TARSKI:def 4; consider k be object such that A48: k in dom G and A49: Y = G.k by A47,FUNCT_1:def 3; reconsider 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 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; a.k=a1.k by A22,A50; then a.k=a.1 by A22,A53,A51; hence contradiction by A2,A13,A45,A52; 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 thesis by A55,XXREAL_2:def 7,def 8; end; A56: for n be Nat holds dom(g - F.n) = dom g proof g is without-infty by A1,Th14; then not -infty in rng g; then A57: g"{-infty} = {} by FUNCT_1:72; g is without+infty by A1,Th14; then not +infty in rng g; then A58: g"{+infty} = {} by FUNCT_1:72; let n be Nat; 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 thesis by A58,A57,A59; end; A60: g is real-valued by A1,MESFUNC2:def 4; A61: for e be R_eal st 0 < e & e < alpha holds ex H be SetSequence of X, MF be ExtREAL_sequence st (for n be Nat holds H.n = less_dom(g-(F.n),e)) & (for n,m be Nat st n <= m holds H.n c= H.m) & (for n be Nat holds H.n c= dom g) & (for n be Nat holds MF.n = M.(H.n)) & M.(dom g) = sup rng MF & for n be Nat holds H.n in S proof let e be R_eal; assume that A62: 0 < e and A63: e < alpha; deffunc FFH(Nat) = less_dom(g-(F.$1),e); consider H be SetSequence of X such that A64: for n be Element of NAT holds H.n = FFH(n) from FUNCT_2:sch 4; A65: now let n be Nat; n in NAT by ORDINAL1:def 12; hence H.n = FFH(n) by A64; end; A66: for n be Nat holds H.n c= dom g proof let n be Nat; now let x be object; assume x in H.n; 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; end; hence thesis; end; A67: Union H c= dom g proof let x be object; assume x in Union H; then consider n be Nat such that A68: x in H.n by PROB_1:12; H.n c= dom g by A66; hence thesis by A68; end; now let x be object; assume A69: x in dom g; then reconsider x1=x as Element of X; A70: F#x1 is convergent by A7,A69; A71: now reconsider E = e as Element of REAL by A62,A63,XXREAL_0:48; assume F#x1 is convergent_to_-infty; then consider N be Nat such that A72: for m be Nat st N <= m holds (F#x1).m <= -E by A62; 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; end; now per cases by A70,A71; suppose A74: F#x1 is convergent_to_finite_number; reconsider E = e as Element of REAL by A62,A63,XXREAL_0:48; A75: ( ex limFx be Real st lim(F#x1) = limFx & (for p be Real st 0

-infty by A74,A75,Th3,Th51; then A85: (F.m).x <> -infty by Def13; (F#x1).m <> +infty by A74,A75,A84,Th3,Th50; then (F.m).x <> +infty by Def13; then lim(F#x1) < (F.m).x + (E/2) by A85,A83,XXREAL_3:54; then A86: lim(F#x1) + E/2 < (F.m).x + (E/2)+ E/2 by XXREAL_3:62; g.x <= lim(F#x1)+(E/2) by A77,XXREAL_3:41; then g.x < (F.m).x1 +(E/2) +(E/2) by A86,XXREAL_0:2; then g.x < (F.m).x1 +((E/2) +(E/2) ) by XXREAL_3:29; then g.x < (F.m).x1 +(E/2+E/2); then g.x - (F.m).x1 < e by XXREAL_3:55; then (g-F.m).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 qua Complex)) by A65; 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 be Nat st N in dom inferior_setsequence H & x in ( inferior_setsequence H).N by A87; end; suppose A88: F#x1 is convergent_to_+infty; ex N be Nat st for m be Nat st N <=m holds g.x1 - (F.m).x1 < e proof A89: e in REAL by A62,A63,XXREAL_0:48; per cases; suppose A90: g.x1-e <= 0; consider N be Nat such that A91: for m be Nat st N <= m holds 1 <= (F#x1).m by A88; now let m be Nat; assume N <=m; 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; end; hence thesis; end; suppose A92: 0 < g.x1-e; reconsider e1=e as Element of REAL by A62,A63,XXREAL_0:48; reconsider gx1=g.x as Real by A60; g.x-e = gx1-e1; then reconsider ee=g.x1-e as Real; consider N be Nat such that A93: for m be Nat st N <= m holds (ee+1) <= (F#x1).m by A88,A92; A94: ee < (ee+1) by XREAL_1:29; now let m be Nat; assume N <=m; then (ee+1) <= (F#x1).m by A93; then 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; end; hence thesis; end; end; then consider N be Nat such that A95: for m be Nat st N <=m holds g.x1 - (F.m).x1 < e; reconsider N as Element of NAT by ORDINAL1:def 12; A96: now let m be Nat; A97: x1 in dom(g - F.m) by A56,A69; assume N <= m; 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; end; now let k be Nat; x in less_dom((g-F.(N+k)),e) by A96,NAT_1:11; hence x in H.(N+(k qua Complex)) by A65; 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 be Nat st N in dom inferior_setsequence H & x in ( inferior_setsequence H).N by A98; end; end; then consider N be 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; end; then A101: dom g c= lim_inf H; deffunc U(Nat) = M.(H.$1); A102: lim_inf H c= lim_sup H by KURATO_0:6; consider MF be ExtREAL_sequence such that A103: for n be Element of NAT holds MF.n = U(n) from FUNCT_2:sch 4; A104: for n,m be Nat st n <= m holds H.n c= H.m proof let n,m be Nat; assume A105: n <= m; now let x be object; assume x in H.n; 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; end; hence thesis; end; then for n,m be Nat st n <= m holds H.n c= H.m; then A113: H is non-descending by PROB_1:def 5; A114: now let n be Nat; n in NAT by ORDINAL1:def 12; hence MF.n = U(n) by A103; end; now let x be object; assume x in lim_inf H; then x in Union inferior_setsequence H by SETLIM_1:def 4; then consider V be set such that A115: x in V and A116: V in rng inferior_setsequence H by TARSKI:def 4; consider n be object such that A117: n in dom inferior_setsequence H and A118: V = (inferior_setsequence H).n by A116,FUNCT_1:def 3; reconsider 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; end; then lim_inf H c= dom g; then A119: lim_inf H = dom g by A101; A120: M.(dom g) = sup rng MF & for n be Element of NAT holds H.n in S proof A121: now reconsider E = e as Element of REAL by A62,A63,XXREAL_0:48; let x be object; assume x in NAT; then reconsider n=x as Element of NAT; A122: less_dom(g-F.n,E) c= dom(g-F.n) by MESFUNC1:def 11; 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,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 real-valued 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,E) = dom(g-F.n) /\ less_dom(g-F.n, E) by A56,A124; then A126: DGH /\ less_dom(g-F.n,E) = less_dom(g-F.n,E) by A122, XBOOLE_1:28; A127: F.n is_measurable_on DGH by A3,MESFUNC2:34; A128: g is real-valued 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,E) in S by MESFUNC1:def 16; hence H.x in S by A65,A126; end; dom H = NAT by FUNCT_2:def 1; then reconsider HH= H as sequence of S by A121,FUNCT_2:3; A129: for n being 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 be object holds x in dom MF iff x in dom H & H.x in dom M proof let x be object; now assume A135: x in dom MF; 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; end; hence thesis by A133; end; for x be object st x in dom MF holds MF.x = M.(H.x) by A103; then M*H =MF by A134,FUNCT_1:10; hence thesis by A121,A129,A131,MEASURE2:23; end; now let n be Nat; n in NAT by ORDINAL1:def 12; hence H.n in S by A120; end; hence thesis by A65,A104,A66,A114,A120; end; per cases; suppose A136: M.(dom g) <> +infty; A137: 0 < beta proof consider x be object such that A138: x in dom g by A11,XBOOLE_0:def 1; A139: g.x <= beta by A44,A138; alpha <= g.x by A44,A138; hence thesis by A14,A38,A41,A42,A39,A139; 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 Element of REAL by A136,XXREAL_0:14; reconsider DG=dom g as Element of S by A1,Th37; A143: for x be object 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) <= beta*(M.DG) proof consider GP be PartFunc of X,ExtREAL such that A147: GP is_simple_func_in S and A148: dom GP = DG and A149: for x be object st x in DG holds GP.x = beta by Th41; A150: for x be object st x in dom(GP-g) holds g.x <= GP.x proof let x be object; assume x in dom(GP-g); 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 = beta by A148,A149; hence thesis by A44,A148,A151; end; for x be object 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 thesis by A137,A147,A148,A149,A154,A155,Th71; end; beta*(M.DG)=beta1*MG by EXTREAL1:1; then A156: integral'(M,g) <> +infty by A146,XXREAL_0:9; A157: for e be R_eal st 0 < e & e < alpha holds ex N0 be Nat st for n be Nat st N0<= n holds integral'(M,g) - e*(beta + M.(dom g)) < integral'(M,F.n) proof let e be R_eal; assume that A158: 0 < e and A159: e < alpha; A160: e <> +infty by A159,XXREAL_0:4; consider H be SetSequence of X, MF be ExtREAL_sequence such that A161: for n be Nat holds H.n = less_dom(g-F.n,e) and A162: for n,m be Nat st n <= m holds H.n c= H.m and A163: for n be Nat holds H.n c= dom g and A164: for n be Nat holds MF.n = M.(H.n) and A165: M.(dom g) = sup rng MF and A166: for n be Nat holds H.n in S by A61,A158,A159; sup rng MF in REAL by A136,A142,A165,XXREAL_0:14; then consider y being ExtReal such that A167: y in rng MF and A168: sup rng MF - e < y by A158,MEASURE6:6; consider N0 be object such that A169: N0 in dom MF and A170: y=MF.N0 by A167,FUNCT_1:def 3; reconsider 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,4; then A171: M.(DG \ B0) = M.DG - M.B0 by A163,MEASURE1:32; take N0; 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 let n be Nat; reconsider BN=H.n as Element of S by A166; assume N0 <= n; 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) +infty by A159,XXREAL_0:4; then reconsider ee=e as Element of REAL by A158,XXREAL_0:14; A176: A misses B by XBOOLE_1:79; beta*e = beta1*ee by EXTREAL1:1; then reconsider betae=beta*e as Real; A177: for x be object 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) in 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,4; then beta*(M.A) < beta*(+infty) by A137,XXREAL_3:72; then A183: beta*(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 be PartFunc of X,ExtREAL such that A187: GP is_simple_func_in S and A188: dom GP = A and A189: for x be object st x in A holds GP.x = beta by Th41; A190: integral'(M,GP) = beta*(M.A) by A137,A187,A188,A189,Th71; A191: dom(g|A) = A by A181,RELAT_1:62,XBOOLE_1:7; A192: for x be object st x in dom(GP-(g|A)) holds (g|A).x <= GP.x proof let x be object; assume x in dom(GP-(g|A)); 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 <= beta by A44; (g|A).x=g.x by A191,A188,A193,FUNCT_1:47; hence thesis by A188,A189,A194,A195; end; for x be object 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 Element of 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; then M.A < e by A173; then A205: beta*(M.A) < beta*e by A137,XXREAL_3:72; then A206: integral'(M,(g|A)) <> +infty by A203,A190,XXREAL_0:2,4; then reconsider XSMgA =integral'(M,g|A) as Element of 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 be PartFunc of X,ExtREAL such that A211: EP is_simple_func_in S and A212: dom EP= B and A213: for x be object 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 be object 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,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 Element of 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 in REAL by A221,A178,XXREAL_0:14; then consider MB,MG be 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; 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 be object 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 object; assume x in dom( ((F.n)|B+EP) - g|B ); 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: f.x < e by MESFUNC1:def 11; x in dom f 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 thesis by A158,A160,A239,A242,XXREAL_3:41; 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 = beta + M.(dom g) by A223; then ee*(beta1+MG) =e*(beta + M.(dom g)); then A248: XSMg-ee*(beta1+MG) = integral'(M,g)-e*(beta + M.(dom g )); integral'(M,g) - e*(beta+M.(dom g)) < integral'(M,(F.n)|B) by A247,A228,A248,XXREAL_0:2; hence integral'(M,g) - e*(beta + M.(dom g)) < integral'(M,F.n) by A220,XXREAL_0:2; end; hence thesis; end; A249: for e be R_eal st 0 < e & e < alpha holds ex N0 be Nat st for n be Nat st N0<= n holds integral'(M,g) - e*(beta + M.(dom g)) < L.n proof let e be R_eal; assume that A250: 0 < e and A251: e < alpha; consider N0 be Nat such that A252: for n be Nat st N0<= n holds integral'(M,g) - e*(beta + M.( dom g)) < integral'(M,F.n) by A157,A250,A251; now let n be Nat; assume N0 <=n; then integral'(M,g)-e*(beta + M.(dom g)) < integral'(M,F.n) by A252; hence integral'(M,g) - e*(beta + M.(dom g)) < L.n by A8; end; hence thesis; end; A253: for e1 be R_eal st 0 < e1 ex e be R_eal st 0 < e & e < alpha & e*( beta + M.(dom g)) <= e1 proof reconsider ralpha=alpha as Real; reconsider rdomg=M.(dom g) as Element of REAL by A136,A142,XXREAL_0:14; let e1 be R_eal; assume A254: 0 < e1; {} c= DG; then A255: 0 <= rdomg by A141,A140,MEASURE1:31; per cases; suppose A256: e1 = +infty; reconsider z=ralpha/2 as R_eal; z*(beta+M.(dom g)) <= +infty by XXREAL_0:4; hence thesis by A43,A256,XREAL_1:216; end; suppose e1 <> +infty; then reconsider re1=e1 as Element of REAL by A254,XXREAL_0:14; set x=re1/(beta1 + rdomg); set y=ralpha/2; set z=min(x,y); A257: z <= y by XXREAL_0:17; y < ralpha by A43,XREAL_1:216; then A258: z < ralpha by A257,XXREAL_0:2; beta1 + rdomg = beta1 + (rdomg qua ExtReal); then A259: z*(beta1 + rdomg) = (z qua ExtReal)*(beta+(M.(dom g))); A260: z*(beta1 + rdomg) <= re1/(beta1 + rdomg)*(beta1 + rdomg) by A137,A255,XREAL_1:64,XXREAL_0:17; reconsider z as R_eal; take z; A261: now per cases by XXREAL_0:15; suppose min(x,y)=x; hence 0 < z by A137,A254,A255; end; suppose min(x,y)=y; hence 0 < z by A43; end; end; z*(beta1 + rdomg) <= re1 by A137,A255,XCMPLX_1:87,A260; hence thesis by A261,A258,A259; end; end; A262: for e1 be R_eal st 0 < e1 holds ex N0 be Nat st for n be Nat st N0 <= n holds integral'(M,g) - e1 < L.n proof let e1 be R_eal; assume 0 < e1; then consider e be R_eal such that A263: 0 < e and A264: e < alpha and A265: e*(beta + M.(dom g)) <= e1 by A253; consider N0 be Nat such that A266: for n be Nat st N0<= n holds integral'(M,g) - e*(beta + M.( dom g)) < L.n by A249,A263,A264; take N0; now let n be Nat; assume N0 <= n; then A267: integral'(M,g) - e*(beta + M.(dom g)) < L.n by A266; integral'(M,g) - e1 <= integral'(M,g) - e*(beta + M.(dom g)) by A265,XXREAL_3:37; hence integral'(M,g) - e1 < L.n by A267,XXREAL_0:2; end; hence thesis; end; A268: for n be Nat holds 0 <= L.n proof let n be Nat; F.n is nonnegative by A5; then 0 <= integral'(M,F.n) by A3,Th68; hence thesis by A8; end; A269: for n,m be Nat st n<=m holds L.n <= L.m proof let n,m be Nat; A270: dom (F.n) = dom g by A4; A271: F.m is_simple_func_in S by A3; A272: dom (F.m) = dom g by A4; assume A273: n <=m; A274: for x be object st x in dom(F.m - F.n) holds (F.n).x <= (F.m).x proof let x be object; assume x in dom(F.m - F.n); 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 thesis by A6,A273,A270,A272; end; A275: F.n is_simple_func_in S by A3; A276: F.m is nonnegative by A5; A277: F.n is nonnegative by A5; then A278: dom(F.m - F.n) = dom(F.m) /\ dom(F.n) by A276,A275,A271,A274,Th69; then A279: (F.m)|dom(F.m - F.n) = F.m by A270,A272,GRFUNC_1:23; A280: (F.n)|dom(F.m - F.n) = F.n by A270,A272,A278,GRFUNC_1:23; integral'(M,(F.n)|dom(F.m - F.n)) <= integral'(M,(F.m)|dom (F.m - F .n)) by A277,A276,A275,A271,A274,Th70; then L.n <= integral'(M,(F.m) ) by A8,A279,A280; hence thesis by A8; end; per cases; suppose ex K be Real st 0 < K & for n be Nat holds L.n < K; then consider K be Real such that 0 < K and A281: for n be Nat holds L.n < K; now let x be ExtReal; assume x in rng L; then ex z be object st z in dom L & x=L.z by FUNCT_1:def 3; hence x <= K by A281; end; then K is UpperBound of rng L by XXREAL_2:def 1; then A282: sup rng L <= K by XXREAL_2:def 3; K in REAL by XREAL_0:def 1; then A283: sup rng L <>+infty by A282,XXREAL_0:9; A284: for n be Nat holds L.n <= sup rng L proof let n be Nat; reconsider n as Element of NAT by ORDINAL1:def 12; dom L = NAT by FUNCT_2:def 1; then A285: L.n in rng L by FUNCT_1:def 3; sup rng L is UpperBound of rng L by XXREAL_2:def 3; hence thesis by A285,XXREAL_2:def 1; end; then L.1 <= sup rng L; then A286: sup rng L <> -infty by A268; then reconsider h=sup rng L as Element of REAL by A283,XXREAL_0:14; A287: for p be Real st 0

sup rng L + p proof assume A290: sup rng L = sup rng L + (p qua ExtReal); p +sup rng L + -sup rng L = p +(sup rng L + -sup rng L) by A286,A283,XXREAL_3:29 .= p + 0 by XXREAL_3:7 .= p; hence contradiction by A288,A290,XXREAL_3:7; end; sup rng L in REAL by A286,A283,XXREAL_0:14; then consider y being ExtReal such that A291: y in rng L and A292: sup rng L - p < y by A288,MEASURE6:6; consider x be object such that A293: x in dom L and A294: y=L.x by A291,FUNCT_1:def 3; reconsider N0=x as Element of NAT by A293; take N0; let n be Nat; assume N0 <= n; then L.N0 <= L.n by A269; then sup rng L - p < L.n by A292,A294,XXREAL_0:2; then sup rng L < L.n + p by XXREAL_3:54; then sup rng L - L.n < p by XXREAL_3:55; then - p < - (sup rng L - L.n) by XXREAL_3:38; then A295: - p < L.n -sup rng L by XXREAL_3:26; A296: L.n <= sup rng L by A284; sup rng L + 0. <= sup rng L + p by A288,XXREAL_3:36; then sup rng L <= sup rng L + p by XXREAL_3:4; then sup rng L < sup rng L + p by A289,XXREAL_0:1; then L.n < sup rng L + p by A296,XXREAL_0:2; then L.n -sup rng L < p by XXREAL_3:55; hence thesis by A295,EXTREAL1:22; end; A297: h=sup rng L; then A298: L is convergent_to_finite_number by A287; hence L is convergent; then A299: lim L = sup rng L by A287,A297,A298,Def12; now let e be Real; assume A300: 0 < e; reconsider ee =e as R_eal by XXREAL_0:def 1; consider N0 be Nat such that A301: for n be Nat st N0<= n holds integral'(M,g) - ee < L. n by A262,A300; A302: L.N0 <= sup rng L by A284; integral'(M,g) - ee < L.N0 by A301; then integral'(M,g) - ee < sup rng L by A302,XXREAL_0:2; hence integral'(M,g) < lim L+ e by A299,XXREAL_3:54; end; hence thesis by XXREAL_3:61; end; suppose A303: not (ex K be Real st 0 < K & for n be Nat holds L.n < K); now let K be Real; assume 0 < K; then consider N0 be Nat such that A304: K <= L.N0 by A303; now let n be Nat; assume N0 <=n; then L.N0 <= L.n by A269; hence K <= L.n by A304,XXREAL_0:2; end; hence ex N0 be Nat st for n be Nat st N0<=n holds K <= L.n; end; then A305: L is convergent_to_+infty; hence L is convergent; then lim L = +infty by A305,Def12; hence thesis by XXREAL_0:4; end; end; suppose A306: M.(dom g) = +infty; reconsider DG=dom g as Element of S by A1,Th37; A307: for e be R_eal st 0 < e & e < alpha holds for n be Nat holds ( alpha - e)*M.less_dom(g-F.n,e) <= integral'(M,F.n) proof let e be R_eal; assume that A308: 0 < e and A309: e < alpha; A310: 0<= alpha-e by A309,XXREAL_3:40; consider H be SetSequence of X, MF be ExtREAL_sequence such that A311: for n be Nat holds H.n = less_dom(g-(F.n),e) and for n,m be Nat st n <= m holds H.n c= H.m and A312: for n be Nat holds H.n c= dom g and for n be Nat holds MF.n = M.(H.n) and M.(dom g) = sup(rng MF) and A313: for n be Nat holds H.n in S by A61,A308,A309; A314: e <> +infty by A309,XXREAL_0:4; now let n be Nat; reconsider B=H.n as Element of S by A313; A315: for x be object st x in dom(F.n) holds (F.n).x=(F.n).x; H.n in S by A313; then A316: 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 A316,MEASURE1:34; A317: dom(F.n) = dom g by A4; A318: DG =DG \/ H.n by A312,XBOOLE_1:12 .=(DG \ H.n) \/ H.n by XBOOLE_1:39; then dom (F.n) = (A \/ B) /\ dom(F.n) by A317; then A319: F.n = (F.n)|(A \/ B) by A315,FUNCT_1:46; consider EP be PartFunc of X,ExtREAL such that A320: EP is_simple_func_in S and A321: dom EP= B and A322: for x be object st x in B holds EP.x = alpha- e by A308,A310,Th41, XXREAL_3:18; for x be object st x in dom EP holds 0 <= EP.x by A310,A321,A322; then A323: EP is nonnegative by SUPINF_2:52; A324: dom((F.n)|B) =dom(F.n) /\ B by RELAT_1:61 .= B by A318,A317,XBOOLE_1:7,28; A325: for x be object st x in dom((F.n)|B - EP) holds EP.x <= ((F.n)|B) .x proof set f=g-F.n; let x be object; assume x in dom((F.n)|B - EP); 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 A326: x in dom((F.n)|B) /\ dom EP by XBOOLE_0:def 5; then A327: x in dom((F.n)|B) by XBOOLE_0:def 4; then A328: ((F.n)|B).x =(F.n).x by FUNCT_1:47; A329: x in less_dom(g-F.n,e) by A311,A324,A327; then A330: x in dom f by MESFUNC1:def 11; f.x < e by A329,MESFUNC1:def 11; then g.x - (F.n).x <= e by A330,MESFUNC1:def 4; then g.x <= (F.n).x + e by A308,A314,XXREAL_3:41; then A331: g.x-e <= (F.n).x by A308,A314,XXREAL_3:42; dom f = dom g by A56; then alpha <= g.x by A44,A330; then alpha-e <= g.x-e by XXREAL_3:37; then alpha -e <= (F.n).x by A331,XXREAL_0:2; hence thesis by A324,A321,A322,A326,A328; end; A332: F.n is_simple_func_in S by A3; (F.n)|A is nonnegative by A5,Th15; then A333: 0 <= integral'(M,(F.n)|A) by A332,Th34,Th68; A334: 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,A319,A334,Th67; then A335: integral'(M,(F.n)|B) <= integral'(M,F.n) by A333,XXREAL_3:39; A336: (F.n)|B is_simple_func_in S by A3,Th34; A337: (F.n)|B is nonnegative by A5,Th15; then A338: dom((F.n)|B - EP) = dom((F.n)|B) /\ dom EP by A336,A320,A323,A325 ,Th69; then A339: EP|dom((F.n)|B - EP) = EP by A324,A321,GRFUNC_1:23; A340: ((F.n)|B)|dom((F.n)|B - EP) = (F.n)|B by A324,A321,A338,GRFUNC_1:23; integral'(M,EP|dom((F.n)|B - EP)) <= integral'(M,((F.n)|B)|dom ((F.n)|B - EP) ) by A337,A336,A320,A323,A325,Th70; then A341: integral'(M,EP) <= integral'(M,F.n) by A335,A339,A340,XXREAL_0:2; integral'(M,EP) = (alpha-e)* (M.B) by A309,A320,A321,A322,Th71, XXREAL_3:40; hence (alpha-e)* M.less_dom(g-F.n,e) <= integral'(M,F.n) by A311,A341 ; end; hence thesis; end; for y be Real st 0 < y ex n be Nat st for m be Nat st n<=m holds y <= L.m proof reconsider ralpha=alpha as Real; reconsider e=alpha/2 as R_eal; let y be Real; assume 0 < y; set a2=ralpha/2; reconsider y1=y as Real; y =(ralpha - a2) * (y1/(ralpha - a2)) by A43,XCMPLX_1:87; then A342: y =(ralpha - a2) *(y1/(ralpha - a2)); A343: e =a2; then consider H be SetSequence of X, MF be ExtREAL_sequence such that A344: for n be Nat holds H.n = less_dom(g-(F.n),e) and A345: for n,m be Nat st n <= m holds H.n c= H.m and for n be Nat holds H.n c= dom g and A346: for n be Nat holds MF.n = M.(H.n) and A347: M.(dom g) = sup rng MF and A348: for n be Nat holds H.n in S by A61,A43,XREAL_1:216; A349: y/(ralpha - a2) in REAL by XREAL_0:def 1; A350: y / (alpha - e) < +infty by XXREAL_0:9,A349; ex z be ExtReal st z in rng MF & (y/ (alpha - e)) <= z proof assume not (ex z be ExtReal st z in rng MF & (y / ( alpha - e)) <= z); then for z be ExtReal st z in rng MF holds z <= (y / (alpha - e)); then y / (alpha - e) is UpperBound of rng MF by XXREAL_2:def 1; hence contradiction by A306,A350,A347,XXREAL_2:def 3; end; then consider z be R_eal such that A351: z in rng MF and A352: y / (alpha - e) <= z; a2-a2 < ralpha - a2 by A43; then A353: 0 < alpha - e; consider x be object such that A354: x in dom MF and A355: z=MF.x by A351,FUNCT_1:def 3; reconsider N0=x as Element of NAT by A354; take N0; A356: (alpha - e)*(y / (alpha - e)) = y by A342; thus for m be Nat st N0 <= m holds y <= L.m proof y / (alpha - e) <= M.(H.N0) by A346,A352,A355; then A357: y <= (alpha - e)*M.(H.N0) by A353,A356,XXREAL_3:71; let m be Nat; A358: H.m in S by A348; assume N0 <= m; then A359: H.N0 c= H.m by A345; H.N0 in S by A348; then (alpha - e)*M.(H.N0) <= (alpha - e)*M.(H.m) by A353,A359,A358, MEASURE1:31,XXREAL_3:71; then y <= (alpha - e)*M.(H.m) by A357,XXREAL_0:2; then A360: y <= (alpha - e)*M.less_dom(g-(F.m),e) by A344; (alpha - e)*M.less_dom(g-(F.m),e) <= integral'(M,F.m) by A43,A307 ,A343,XREAL_1:216; then y <= integral'(M,F.m) by A360,XXREAL_0:2; hence thesis by A8; end; end; then A361: L is convergent_to_+infty; hence L is convergent; then ( ex g be Real st lim L = g & (for p be Real st 0

0 by A16,MESFUNC1:def 15; then 0 < g.x by A2,SUPINF_2:51; hence thesis by A15,FUNCT_1:47; end; deffunc V(Nat) = integral'(M,(F.$1)|E9); deffunc U(Nat) = integral'(M,F.$1); deffunc W(Nat) = (F.$1)|E9; consider F9 be Functional_Sequence of X,ExtREAL such that A17: for n be Nat holds F9.n=W(n) from SEQFUNC:sch 1; consider L be ExtREAL_sequence such that A18: for n be Element of NAT holds L.n = V(n) from FUNCT_2:sch 4; A19: now let n be Nat; n in NAT by ORDINAL1:def 12; hence L.n=V(n) by A18; end; A20: for n be Nat holds L.n = integral'(M,F9.n) proof let n be Nat; thus L.n = integral'(M,F.n|E9) by A19 .=integral'(M,F9.n) by A17; end; consider G be ExtREAL_sequence such that A21: for n be Element of NAT holds G.n = U(n) from FUNCT_2:sch 4; take G; A22: for x be object 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 be Nat holds (F.n)|E9 is_simple_func_in S & F9.n is_simple_func_in S proof let n be Nat; thus F.n|E9 is_simple_func_in S by A3,Th34; hence thesis by A17; end; A27: for n be Nat holds dom(F.n|E9) = dom(g|E9) & dom(F9.n) = dom(g|E9) proof let n be Nat; 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,28; hence thesis by A17; end; A29: for x be 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; assume A30: x in dom(g|E9); now let n be Element of NAT; 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; 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 thesis by A7,A30,A33,A32,FUNCT_1:47; end; A34: for n be Nat holds F9.n is nonnegative proof let n be Nat; F.n|E9 is nonnegative by A5,Th15; hence thesis by A17; end; A35: E9 c= dom g by A9,XBOOLE_1:7; A36: for n,m be Nat st n <= m holds for x be 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; assume A37: n<=m; thus for x be 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 x be Element of X; assume A38: x in dom(g|E9); 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; then (F9.n).x <= (F.m|E9).x by A17; hence thesis by A17; end; end; then for n,m be Nat st n <= m holds for x be 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 be Nat st n <= m holds L.n <= L.m proof let n,m be Nat; 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; A48: for x be object st x in dom(F9.m - F9.n) holds (F9.n).x <= (F9.m).x proof let x be object; assume x in dom(F9.m - F9.n); 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 thesis by A36,A47,A46,A43; 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 thesis by A51,A42,A50,A49,A48,A53,A45,A44,Th70; end; then A54: lim L = sup rng L by Th54; A55: now let n be Nat; n in NAT by ORDINAL1:def 12; hence G.n = U(n) by A21; end; for n be Nat holds L.n <= G.n proof let n be Nat; 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 be object 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 thesis by A19,A60; end; then A61: sup rng L <= sup rng G by Th55; A62: for n,m be Nat st n <=m holds G.n <= G.m proof let n,m be Nat; 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; A69: for x be object st x in dom (F.m - F.n) holds (F.n).x <= (F.m).x proof let x be object; assume x in dom(F.m - F.n); 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 thesis by A6,A68,A67,A64; 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 thesis by A72,A63,A71,A70,A69,A74,A66,A65,Th70; end; then lim G = sup rng G by Th54; hence thesis by A24,A55,A62,A41,A54,A61,Th54,XXREAL_0:2; end; end; theorem Th76: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, A be Element of S, F,G be Functional_Sequence of X,ExtREAL, K,L be ExtREAL_sequence st (for n be Nat holds F.n is_simple_func_in S & dom(F.n)=A ) & (for n be Nat holds F.n is nonnegative) & (for n,m be Nat st n <=m holds for x be Element of X st x in A holds (F.n).x <= (F.m).x ) & (for n be Nat holds G. n is_simple_func_in S & dom(G.n)=A) & (for n be Nat holds G.n is nonnegative) & (for n,m be Nat st n <=m holds for x be Element of X st x in A holds (G.n).x <= (G.m).x ) & (for x be Element of X st x in A holds F#x is convergent & G#x is convergent & lim(F#x) = lim(G#x)) & (for n be 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, S be SigmaField of X, M be sigma_Measure of S, A be Element of S, F,G be Functional_Sequence of X,ExtREAL, K,L be ExtREAL_sequence such that A1: for n be Nat holds F.n is_simple_func_in S & dom(F.n)=A and A2: for n be Nat holds F.n is nonnegative and A3: for n,m be Nat st n <=m holds for x be Element of X st x in A holds (F.n).x <= (F.m).x and A4: for n be Nat holds G.n is_simple_func_in S & dom(G.n)=A and A5: for n be Nat holds G.n is nonnegative and A6: for n,m be Nat st n <=m holds for x be Element of X st x in A holds (G.n).x <= (G.m).x and A7: for x be 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 be Nat holds K.n=integral'(M,F.n) & L.n=integral'(M,G.n); A9: for n0 be Nat holds L is convergent & sup rng L=lim L & K.n0 <= lim L proof let n0 be Nat; 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 be 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; A13: (F#x).n0 <= sup rng (F#x) by Th56; assume x in dom f; then A14: x in A by A1; now let n,m be Nat; assume A15: n<=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; end; then A17: lim(F#x)=sup rng(F#x) by Th54; f.x=(F#x).n0 by Def13; hence thesis by A7,A14,A17,A13; end; dom f = A by A1; then consider FF be ExtREAL_sequence such that A18: for n be 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 let n be Element of NAT; FF.n = integral'(M,G.n) by A18; hence FF.n = L.n by A8; end; then FF=L by FUNCT_2:63; hence thesis by A8,A19,A20,A21; end; A22: for n0 be Nat holds K is convergent & sup rng K = lim K & L.n0 <= lim K proof let n0 be Nat; 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 be 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; A26: (G#x).n0 <= sup rng(G#x) by Th56; assume x in dom g; then A27: x in A by A4; now let n,m be Nat; assume A28: n<=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; end; then A30: lim(G#x)=sup rng(G#x) by Th54; g.x=(G#x).n0 by Def13; hence thesis by A7,A27,A30,A26; end; dom g = A by A4; then consider GG be ExtREAL_sequence such that A31: for n be 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 let n be Element of NAT; GG.n = integral'(M,F.n) by A31; hence GG.n = K.n by A8; end; then GG=K by FUNCT_2:63; hence thesis by A8,A32,A33,A34; end; hence K is convergent & L is convergent by A9; A35: lim K <= lim L by A22,A9,Th57; lim L <= lim K by A22,A9,Th57; hence thesis by A35,XXREAL_0:1; 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 be 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: ex F be Functional_Sequence of X,ExtREAL, K be ExtREAL_sequence st (for n be Nat holds F.n is_simple_func_in S & dom(F.n) = dom f) & (for n be Nat holds F.n is nonnegative) & (for n,m be Nat st n <=m holds for x be Element of X st x in dom f holds (F.n).x <= (F.m).x ) & (for x be Element of X st x in dom f holds F#x is convergent & lim(F#x) = f.x) & (for n be Nat holds K.n=integral'(M,F.n)) & K is convergent & it=lim K; existence proof consider F be Functional_Sequence of X,ExtREAL such that A3: for n be Nat holds F.n is_simple_func_in S & dom(F.n) = dom f and A4: for n be Nat holds F.n is nonnegative and A5: for n,m be Nat st n <=m holds for x be Element of X st x in dom f holds (F.n).x <= (F.m).x and A6: for x be 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 be 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 such that A9: x in dom f; A10: now let n,m be Nat; assume A11: n<=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; end; A13: g.x=(F#x).0 by Def13; lim(F#x)=sup rng(F#x) by A10,Th54; hence thesis by A10,A13,Th54,Th56; end; dom g = dom f by A3; then ex G be ExtREAL_sequence st (for n be 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 be ExtREAL_sequence such that A14: for n be Nat holds G.n = integral'(M,F.n) and A15: G is convergent and integral'(M,g) <= lim G; take lim G; thus thesis by A3,A4,A5,A6,A14,A15; end; uniqueness proof let s1,s2 be Element of ExtREAL such that A16: ex F1 be Functional_Sequence of X,ExtREAL, K1 be ExtREAL_sequence st (for n be Nat holds F1.n is_simple_func_in S & dom(F1.n) = dom f) & (for n be Nat holds F1.n is nonnegative) & (for n,m be Nat st n <=m holds for x be Element of X st x in dom f holds (F1.n).x <= (F1.m).x ) & (for x be Element of X st x in dom f holds F1#x is convergent & lim(F1#x) = f.x) & (for n be Nat holds K1.n=integral'(M,F1.n)) & K1 is convergent & s1=lim K1 and A17: ex F2 be Functional_Sequence of X,ExtREAL, K2 be ExtREAL_sequence st (for n be Nat holds F2.n is_simple_func_in S & dom(F2.n) = dom f) & (for n be Nat holds F2.n is nonnegative) & (for n,m be Nat st n <=m holds for x be Element of X st x in dom f holds (F2.n).x <= (F2.m).x ) & (for x be Element of X st x in dom f holds F2#x is convergent & lim(F2#x) = f.x) & (for n be Nat holds K2.n=integral'(M,F2.n)) & K2 is convergent & s2=lim K2; consider F1 be Functional_Sequence of X,ExtREAL, K1 be ExtREAL_sequence such that A18: for n be Nat holds F1.n is_simple_func_in S & dom(F1.n) = dom f and A19: for n be Nat holds F1.n is nonnegative and A20: for n,m be Nat st n <=m holds for x be Element of X st x in dom f holds (F1.n).x <= (F1.m).x and A21: for x be Element of X st x in dom f holds F1#x is convergent & lim (F1#x) = f.x and A22: for n be Nat holds K1.n=integral'(M,F1.n) and K1 is convergent and A23: s1=lim(K1) by A16; consider F2 be Functional_Sequence of X,ExtREAL, K2 be ExtREAL_sequence such that A24: for n be Nat holds F2.n is_simple_func_in S & dom(F2.n) = dom f and A25: for n be Nat holds F2.n is nonnegative and A26: for n,m be Nat st n <=m holds for x be Element of X st x in dom f holds (F2.n).x <= (F2.m).x and A27: for x be Element of X st x in dom f holds (F2#x) is convergent & lim(F2#x) = f.x and A28: for n be Nat holds K2.n=integral'(M,F2.n) and K2 is convergent and A29: s2=lim K2 by A17; for x be 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; assume A30: x in dom f; then lim(F1#x) = f.x by A21 .= lim(F2#x) by A27,A30; hence thesis by A21,A27,A30; end; hence thesis by A1,A18,A19,A20,A22,A23,A24,A25,A26,A28,A29,Th76; end; end; theorem Th77: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be 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, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL such that A1: f is_simple_func_in S and A2: f is nonnegative; deffunc PF(Nat) = f; consider F be Functional_Sequence of X,ExtREAL such that A3: for n be Nat holds F.n=PF(n) from SEQFUNC:sch 1; A4: for n,m be Nat st n <=m holds for x be Element of X st x in dom f holds (F.n).x <= (F.m).x proof let n,m be Nat; assume n<=m; let x be Element of X; assume x in dom f; (F.n).x=f.x by A3; hence thesis by A3; end; deffunc PK(Nat) = integral'(M,(F.$1)); consider K be sequence of ExtREAL such that A5: for n be Element of NAT holds K.n = PK(n) from FUNCT_2:sch 4; A6: now let n be Nat; n in NAT by ORDINAL1:def 12; hence K.n=PK(n) by A5; end; A7: for n be Nat holds K.n=integral'(M,f) proof let n be Nat; thus K.n=integral'(M,F.n) by A6 .=integral'(M,f) by A3; end; then A8: lim K=integral'(M,f) by Th60; ex GF be Finite_Sep_Sequence of S st dom f = union rng GF & for n being Nat,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 be 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; assume x in dom f; now let n be Nat; thus (F#x).n = (F.n).x by Def13 .=f.x by A3; end; hence thesis by Th60; end; A11: for n be Nat holds F.n is nonnegative by A2,A3; A12: for n be Nat holds F.n is_simple_func_in S & dom(F.n) = dom f by A1,A3; K is convergent by A7,Th60; hence thesis by A2,A9,A6,A12,A11,A4,A10,A8,Def15; end; Lm10: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f ,g be PartFunc of X,ExtREAL st ( ex A be 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, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL such that A1: ex A be 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; consider F1 be Functional_Sequence of X,ExtREAL, K1 be ExtREAL_sequence such that A4: for n be Nat holds F1.n is_simple_func_in S & dom(F1.n) = dom f and A5: for n be Nat holds F1.n is nonnegative and A6: for n,m be Nat st n <=m holds for x be Element of X st x in dom f holds (F1.n).x <= (F1.m).x and A7: for x be Element of X st x in dom f holds F1#x is convergent & lim( F1#x) = f.x and A8: for n be 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 be 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 be Nat st n<=m holds K1.n <= K1.m proof let n,m be Nat such that A17: n<=m; A18: dom(F1.m) = dom f by A4; A19: dom(F1.n) = dom f by A4; A20: now let x be object; assume x in dom(F1.m - F1.n); 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; 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 thesis by A24,A23,A26,A25,A22,A21,A20,A28,Th70; end; consider F2 be Functional_Sequence of X,ExtREAL, K2 be ExtREAL_sequence such that A29: for n be Nat holds F2.n is_simple_func_in S & dom(F2.n) = dom g and A30: for n be Nat holds F2.n is nonnegative and A31: for n,m be Nat st n <=m holds for x be Element of X st x in dom g holds (F2.n).x <= (F2.m).x and A32: for x be Element of X st x in dom g holds F2#x is convergent & lim( F2#x) = g.x and A33: for n be Nat holds K2.n=integral'(M,F2.n) and K2 is convergent and A34: integral+(M,g)=lim K2 by A1,A3,Def15; deffunc PF(Nat) = F1.$1+F2.$1; consider F be Functional_Sequence of X,ExtREAL such that A35: for n be Nat holds F.n=PF(n) from SEQFUNC:sch 1; A36: for n be Nat holds F.n is_simple_func_in S & dom(F.n) = dom(f+g) & F.n is nonnegative proof let n be Nat; 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; 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; A43: F2.n is nonnegative by A30; A44: F.n=F1.n+F2.n by A35; F1.n is nonnegative by A5; hence thesis by A43,A44,Th19; end; A45: for n,m be Nat st n <=m holds for x be Element of X st x in dom (f+g) holds (F.n).x <= (F.m).x proof let n,m be Nat; assume A46: n<=m; 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; assume A49: x in dom (f+g); 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 thesis by A52,A51,A50,XXREAL_3:36; end; now let n be set; assume n in dom K2; 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; end; then A54: K2 is without-infty by Th10; deffunc PK(Nat) = integral'(M,F.$1); consider K be ExtREAL_sequence such that A55: for n be Element of NAT holds K.n = PK(n) from FUNCT_2:sch 4; A56: now let n be Nat; n in NAT by ORDINAL1:def 12; hence K.n = PK(n) by A55; end; A57: for n be Nat holds K.n=K1.n+K2.n proof let n be Nat; 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 thesis by A8,A65; end; A66: for n,m be Nat st n<=m holds K2.n <= K2.m proof let n,m be Nat such that A67: n<=m; A68: dom(F2.m) = dom g by A29; A69: dom(F2.n) = dom g by A29; A70: now let x be object; assume x in dom(F2.m - F2.n); 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; 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 thesis by A74,A73,A76,A75,A72,A71,A70,A78,Th70; end; now let n be set; assume n in dom K1; 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; end; then A80: K1 is without-infty by Th10; then A81: lim K=lim K1+lim K2 by A16,A54,A66,A57,Th61; A82: for x be 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; A83: now let n be set; hereby assume n in dom(F1#x); 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; end; assume n in dom(F2#x); 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; end; then A86: F2#x is without-infty by Th10; assume A87: x in dom(f+g); 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 let n,m be Nat; assume A90: n <=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; end; A92: now let n,m be Nat; assume A93: n <=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; end; A95: now let n be Nat; (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; end; F1#x is without-infty by A83,Th10; hence thesis by A95,A86,A92,A89,A88,Th61; 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 thesis by A9,A34,A97,A15,A10,A56,A36,A45,A82,A81,Def15; end; theorem Th78: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL st ( ex A be Element of S st A = dom f & f is_measurable_on A ) & ( ex B be Element of S st B = dom g & g is_measurable_on B ) & f is nonnegative & g is nonnegative holds ex C be 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, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL; assume that A1: ex A be Element of S st A = dom f & f is_measurable_on A and A2: ex B be Element of S st B = dom g & g is_measurable_on B and A3: f is nonnegative and A4: g is nonnegative; set g1 = g|(dom f /\ dom g); A5: g1 is without-infty by A4,Th12,Th15; A6: g1 is nonnegative by A4,Th15; dom g1 = dom g /\ (dom f /\ dom g) by RELAT_1:61; then A7: dom g1 = dom g /\ dom g /\ dom f by XBOOLE_1:16; consider B be Element of S such that A8: B = dom g and A9: g is_measurable_on B by A2; consider A be Element of S such that A10: A = dom f and A11: f is_measurable_on A by A1; take C = A /\ B; A12: C = dom(f+g) by A3,A4,A10,A8,Th16; A13: C = dom g /\ C by A8,XBOOLE_1:17,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,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 f1 = dom f /\ (dom f /\ dom g) by RELAT_1:61; then A17: dom f1 = dom f /\ dom f /\ dom g by XBOOLE_1:16; A18: f1 is without-infty by A3,Th12,Th15; then A19: dom(f1+g1) = C /\ C by A10,A8,A17,A7,A5,Th16; A20: dom(f1+g1) = dom f1 /\ dom g1 by A18,A5,Th16; A21: for x be object st x in dom(f1+g1) holds (f1+g1).x = (f+g).x proof let x be object; assume A22: x in dom(f1+g1); then A23: x in dom f1 by A20,XBOOLE_0:def 4; A24: x in dom g1 by A20,A22,XBOOLE_0:def 4; (f1+g1).x = f1.x + g1.x by A22,MESFUNC1:def 3 .= f.x + g1.x by A23,FUNCT_1:47 .= f.x + g.x by A24,FUNCT_1:47; hence thesis by A12,A19,A22,MESFUNC1:def 3; end; f1 is nonnegative by A3,Th15; then integral+(M,f1+g1) = integral+(M,f1) + integral+(M,g1) by A10,A8,A17,A7,A16 ,A14,A6,Lm10; hence thesis by A10,A8,A12,A19,A21,FUNCT_1:2; end; theorem Th79: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL st (ex A be 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, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL; assume that A1: ex A be Element of S st A = dom f & f is_measurable_on A and A2: f is nonnegative; consider F1 be Functional_Sequence of X,ExtREAL, K1 be ExtREAL_sequence such that A3: for n be Nat holds F1.n is_simple_func_in S & dom(F1.n) = dom f and A4: for n be Nat holds F1.n is nonnegative and A5: for n,m be Nat st n <=m holds for x be Element of X st x in dom f holds (F1.n).x <= (F1.m).x and for x be Element of X st x in dom f holds F1#x is convergent & lim( F1#x ) = f.x and A6: for n be 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 be Nat st n<=m holds K1.n <= K1.m proof let n,m be Nat; 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; A13: for x be object st x in dom(F1.m - F1.n) holds (F1.n).x <= (F1.m).x proof let x be object; assume x in dom(F1.m - F1.n); 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 thesis by A5,A12,A11,A9; 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 thesis by A6,A10,A18,A19; end; then lim K1 =sup rng K1 by Th54; then A20: K1.0 <= lim K1 by Th56; for n be Nat holds 0 <= K1.n proof let n be Nat; A21: F1.n is_simple_func_in S by A3; K1.n = integral'(M,F1.n) by A6; hence thesis by A4,A21,Th68; end; hence thesis by A7,A20; end; theorem Th80: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A be Element of S st (ex E be 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, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A be Element of S; assume that A1: ex E be Element of S st E = dom f & f is_measurable_on E and A2: f is nonnegative; consider E be Element of S such that A3: E = dom f and A4: f is_measurable_on E by A1; set C = E/\A; A5: C = dom(f|A) by A3,RELAT_1:61; A6: dom(f|A) = C by A3,RELAT_1:61 .= dom f /\ C by A3,XBOOLE_1:17,28 .= dom(f|C) by RELAT_1:61; A7: for x be object st x in dom(f|A) holds (f|A).x = (f|C).x proof let x be object; assume A8: x in dom(f|A); then (f|A).x = f.x by FUNCT_1:47; hence thesis by A6,A8,FUNCT_1:47; end; A9: dom f /\ C = C by A3,XBOOLE_1:17,28; f is_measurable_on C by A4,MESFUNC1:30,XBOOLE_1:17; then f|C is_measurable_on C by A9,Th42; then f|A is_measurable_on C by A6,A7,FUNCT_1:2; hence thesis by A2,A5,Th15,Th79; end; theorem Th81: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A,B be Element of S st (ex E be 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, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A,B be Element of S; assume that A1: ex E be Element of S st E = dom f & f is_measurable_on E and A2: f is nonnegative and A3: A misses B; consider F0 be Functional_Sequence of X,ExtREAL, K0 be ExtREAL_sequence such that A4: for n be Nat holds F0.n is_simple_func_in S & dom(F0.n) = dom f and A5: for n be Nat holds F0.n is nonnegative and A6: for n,m be Nat st n <=m holds for x be Element of X st x in dom f holds (F0.n).x <= (F0.m).x and A7: for x be Element of X st x in dom f holds F0#x is convergent & lim (F0#x) = f.x and for n be Nat holds K0.n=integral'(M,F0.n) and K0 is convergent and integral+(M,f)=lim K0 by A1,A2,Def15; deffunc PFB(Nat) = F0.$1|B; deffunc PFA(Nat) = F0.$1|A; consider FA be Functional_Sequence of X,ExtREAL such that A8: for n be Nat holds FA.n=PFA(n) from SEQFUNC:sch 1; consider E be Element of S such that A9: E = dom f and A10: f is_measurable_on E by A1; consider FB be Functional_Sequence of X,ExtREAL such that A11: for n be Nat holds FB.n=PFB(n) from SEQFUNC:sch 1; set DB = E /\ B; A12: DB = dom(f|B) by A9,RELAT_1:61; then A13: dom f /\ DB = DB by RELAT_1:60,XBOOLE_1:28; then A14: dom(f|DB) = dom(f|B) by A12,RELAT_1:61; for x be object st x in dom(f|DB) holds (f|DB).x = (f|B).x proof let x be object; assume A15: x in dom(f|DB); then f|B.x = f.x by A14,FUNCT_1:47; hence thesis by A15,FUNCT_1:47; end; then A16: f|DB = f|B by A12,A13,FUNCT_1:2,RELAT_1:61; set DA = E /\ A; A17: DA = dom(f|A) by A9,RELAT_1:61; then A18: dom f /\ DA = DA by RELAT_1:60,XBOOLE_1:28; then A19: dom(f|DA) = dom(f|A) by A17,RELAT_1:61; for x be object st x in dom(f|DA) holds (f|DA).x = (f|A).x proof let x be object; assume A20: x in dom(f|DA); then f|A.x = f.x by A19,FUNCT_1:47; hence thesis by A20,FUNCT_1:47; end; then A21: f|DA = f|A by A17,A18,FUNCT_1:2,RELAT_1:61; A22: for n be 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; 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; dom(FA.n)=dom(F0.n) /\ A by A25,RELAT_1:61; hence thesis by A9,A4,A17,A12,A24; end; A26: for x be 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; assume A27: x in dom(f|A); now let n be Element of NAT; (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; 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 thesis by A7,A27,A30,A29,FUNCT_1:47; end; A31: for x be 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; assume A32: x in dom(f|B); now let n be Element of NAT; 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; 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 thesis by A7,A32,A35,A34,FUNCT_1:47; end; set C = E/\(A\/B); A36: C = dom f /\ C by A9,XBOOLE_1:17,28; A37: dom(f|(A\/B)) = C by A9,RELAT_1:61; then A38: dom(f|(A\/B)) = dom(f|C) by A36,RELAT_1:61; for x be object st x in dom(f|(A\/B)) holds f|(A\/B).x = f|C.x proof let x be object; assume A39: x in dom(f|(A\/B)); then f|(A\/B).x = f.x by FUNCT_1:47; hence thesis by A38,A39,FUNCT_1:47; end; then A40: f|(A\/B) = f|C by A38,FUNCT_1:2; f is_measurable_on C by A10,MESFUNC1:30,XBOOLE_1:17; then A41: f|(A\/B) is_measurable_on C by A36,A40,Th42; f is_measurable_on DB by A10,MESFUNC1:30,XBOOLE_1:17; then A42: f|B is_measurable_on DB by A13,A16,Th42; A43: f|B is nonnegative by A2,Th15; f is_measurable_on DA by A10,MESFUNC1:30,XBOOLE_1:17; then A44: f|A is_measurable_on DA by A18,A21,Th42; A45: f|A is nonnegative by A2,Th15; deffunc PFAB(Nat) = (F0.$1)|(A\/B); consider FAB be Functional_Sequence of X,ExtREAL such that A46: for n be Nat holds FAB.n=PFAB(n) from SEQFUNC:sch 1; A47: for n be Nat holds FA.n is nonnegative & FB.n is nonnegative proof let n be Nat; reconsider 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 thesis by A8,A11,A48; end; A49: for n,m be Nat st n <=m holds for x be Element of X st x in dom(f|B) holds (FB.n).x <= (FB.m).x proof let n,m be Nat; assume A50: n<=m; reconsider n,m as Element of NAT by ORDINAL1:def 12; let x be Element of X; assume A51: x in dom(f|B); 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 thesis by A6,A50,A52,A54; end; deffunc PKA(Nat) = integral'(M,FA.$1); consider KA be ExtREAL_sequence such that A56: for n be Element of NAT holds KA.n = PKA(n) from FUNCT_2:sch 4; deffunc PKB(Nat) = integral'(M,FB.$1); consider KB be ExtREAL_sequence such that A57: for n be Element of NAT holds KB.n = PKB(n) from FUNCT_2:sch 4; A58: now let n be Nat; n in NAT by ORDINAL1:def 12; hence KB.n = PKB(n) by A57; end; A59: now let n be Nat; n in NAT by ORDINAL1:def 12; hence KA.n = PKA(n) by A56; end; A60: for n be set holds (n in dom KA implies -infty < KA.n) & (n in dom KB implies -infty < KB.n) proof let n be set; hereby assume n in dom KA; 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; end; assume n in dom KB; 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 thesis by A47,A62,Th68; end; then A63: KB is without-infty by Th10; deffunc PK(Nat) = integral'(M,FAB.$1); consider KAB be ExtREAL_sequence such that A64: for n be Element of NAT holds KAB.n = PK(n) from FUNCT_2:sch 4; A65: now let n be Nat; n in NAT by ORDINAL1:def 12; hence KAB.n = PK(n) by A64; end; A66: for n be Nat holds KAB.n=KA.n + KB.n proof let n be Nat; reconsider 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 thesis by A58,A69,A70; end; A71: for n,m be Nat st n <=m holds for x be Element of X st x in dom(f|A) holds (FA.n).x <= (FA.m).x proof let n,m be Nat; assume A72: n<=m; reconsider n,m as Element of NAT by ORDINAL1:def 12; let x be Element of X; assume A73: x in dom(f|A); 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 thesis by A6,A72,A74,A76; end; A78: for n,m be Nat st n<=m holds KA.n <= KA.m & KB.n <= KB.m proof let n,m be Nat; 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; A84: for x be object st x in dom(FA.m - FA.n) holds (FA.n).x <= (FA.m).x proof let x be object; assume x in dom(FA.m - FA.n); 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 thesis by A71,A83,A82,A80; 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; 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 be object st x in dom(FB.m - FB.n) holds (FB.n).x <= (FB.m).x proof let x be object; assume x in dom(FB.m - FB.n); 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 thesis by A49,A83,A96,A95; 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 thesis by A58,A94,A100,A101; end; then A102: for n,m be 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 be Nat holds FAB.n is_simple_func_in S & dom(FAB.n) = dom(f|(A\/ B))) & (for n be Nat holds for x be Element of X st x in dom(f|(A\/B)) holds ( FAB.n).x = (F0.n).x) & (for n be Nat holds FAB.n is nonnegative) & (for n,m be Nat st n <=m holds for x be Element of X st x in dom(f|(A\/B)) holds (FAB.n).x <= (FAB.m).x ) & for x be Element of X st x in dom(f|(A\/B)) holds FAB#x is convergent & lim(FAB#x) = f|(A\/B).x proof thus A105: now let n be Nat; FAB.n=(F0.n)|(A \/ B) by A46; hence FAB.n is_simple_func_in S by A4,Th34; 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; end; thus A106: now let n be Nat, x be Element of X; assume x in dom(f|(A\/B)); 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; end; hereby let n be Nat; 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; end; hereby let n,m be Nat such that A108: n <= m; now let x be Element of X; assume A109: x in dom(f|(A\/B)); 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; end; hence for x be Element of X st x in dom(f|(A\/B)) holds (FAB.n).x <= (FAB . m).x; end; hereby let x be Element of X; assume A112: x in dom(f|(A\/B)); then x in dom f /\ (A\/B) by RELAT_1:61; then A113: x in dom f by XBOOLE_0:def 4; A114: now let n be Element of NAT; thus (FAB#x).n = (FAB.n).x by Def13 .=(F0.n).x by A106,A112 .=(F0#x).n by Def13; end; then FAB#x=F0#x by FUNCT_2:63; hence FAB#x is convergent by A7,A113; 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; end; end; for n,m be Nat st n<=m holds KAB.n <= KAB.m proof let n,m be Nat; assume A115: n<=m; reconsider n,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 be object st x in dom(FAB.m - FAB.n) holds (FAB.n).x <= (FAB.m).x proof let x be object; assume x in dom(FAB.m - FAB.n); 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 thesis by A104,A115,A117,A116; 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 thesis by A65,A119,A125,A126; end; then A127: KAB is convergent by Th54; A128: for n,m be 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 without-infty by A60,Th10; hence thesis by A130,A102,A128,A103,A129,A66,A63,Th61; end; theorem Th82: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A be Element of S st (ex E be 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, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A be Element of S; assume that A1: ex E be Element of S st E = dom f & f is_measurable_on E and A2: f is nonnegative and A3: M.A = 0; consider F0 be Functional_Sequence of X,ExtREAL, K0 be ExtREAL_sequence such that A4: for n be Nat holds F0.n is_simple_func_in S & dom(F0.n) = dom f and A5: for n be Nat holds F0.n is nonnegative and A6: for n,m be Nat st n <=m holds for x be Element of X st x in dom f holds (F0.n).x <= (F0.m).x and A7: for x be Element of X st x in dom f holds F0#x is convergent & lim( F0#x) = f.x and for n be Nat holds K0.n=integral'(M,F0.n) and K0 is convergent and integral+(M,f)=lim K0 by A1,A2,Def15; deffunc PFA(Nat) = (F0.$1)|A; consider FA be Functional_Sequence of X,ExtREAL such that A8: for n be Nat holds FA.n=PFA(n) from SEQFUNC:sch 1; consider E be 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 /\ C = C by A9,XBOOLE_1:17,28; then A13: dom(f|C) = C by RELAT_1:61; then A14: dom(f|C) = dom(f|A) by A9,RELAT_1:61; for x be object st x in dom(f|A) holds f|A.x = f|C.x proof let x be object; assume A15: x in dom(f|A); then (f|A).x = f.x by FUNCT_1:47; hence thesis by A14,A15,FUNCT_1:47; end; then A16: f|A = f|C by A9,A13,FUNCT_1:2,RELAT_1:61; f is_measurable_on C by A10,MESFUNC1:30,XBOOLE_1:17; then A17: f|A is_measurable_on C by A12,A16,Th42; A18: for n be Nat holds FA.n is nonnegative proof let n be Nat; reconsider n as Element of NAT by ORDINAL1:def 12; F0.n|A is nonnegative by A5,Th15; hence thesis by A8; end; deffunc PK(Nat) = integral'(M,FA.$1); consider KA be ExtREAL_sequence such that A19: for n be Element of NAT holds KA.n = PK(n) from FUNCT_2:sch 4; A20: now let n be Nat; n in NAT by ORDINAL1:def 12; hence KA.n = PK(n) by A19; end; A21: for n be Nat holds KA.n =0 proof let n be Nat; reconsider n as Element of NAT by ORDINAL1:def 12; F0.n is_simple_func_in S by A4; then integral'(M,F0.n|A) = 0 by A3,A5,Th73; then integral'(M,FA.n) = 0 by A8; hence thesis by A20; end; then A22: lim KA = 0 by Th60; A23: C = dom(f|A) by A9,RELAT_1:61; A24: for n be Nat holds FA.n is_simple_func_in S & dom(FA.n) = dom(f|A) proof let n be Nat; 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; dom(FA.n1)=dom(F0.n1|A) by A8; then dom(FA.n)=dom(F0.n) /\ A by RELAT_1:61; hence thesis by A9,A4,A23; end; A25: for x be 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; assume A26: x in dom(f|A); now let n be Element of NAT; 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; 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 thesis by A7,A26,A29,A28,FUNCT_1:47; end; A30: for n,m be Nat st n <=m holds for x be Element of X st x in dom(f|A) holds (FA.n).x <= (FA.m).x proof let n,m be Nat; assume A31: n<=m; let x be Element of X; reconsider n,m as Element of NAT by ORDINAL1:def 12; assume A32: x in dom(f|A); 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 thesis by A6,A31,A33,A35; end; KA is convergent by A21,Th60; hence thesis by A17,A20,A23,A11,A24,A18,A30,A25,A22,Def15; end; theorem Th83: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A,B be Element of S st (ex E be 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, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A,B be Element of S; assume that A1: ex E be Element of S st E = dom f & f is_measurable_on E and A2: f is nonnegative and A3: A c= B; set A9 = A /\ B; A4: A9 = A by A3,XBOOLE_1:28; set B9 = B \ A; A5: A9\/B9 = B by XBOOLE_1:51; integral+(M,f|(A9\/B9)) =integral+(M,f|A9)+integral+(M,f|B9) by A1,A2,Th81, XBOOLE_1:89; hence thesis by A1,A2,A4,A5,Th80,XXREAL_3:39; end; theorem Th84: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, E,A be 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, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, E,A be Element of S such that A1: f is nonnegative and A2: E = dom f and A3: f is_measurable_on E and A4: M.A =0; set B = E\A; A \/ B = A \/ E by XBOOLE_1:39; then A5: dom f = dom f /\ (A\/B) by A2,XBOOLE_1:7,28 .= dom(f|(A\/B)) by RELAT_1:61; for x be object st x in dom(f|(A\/B)) holds (f|(A\/B)).x = f.x by FUNCT_1:47; then A6: f|(A\/B) =f by A5,FUNCT_1:2; integral+(M,f|(A\/B)) =integral+(M,f|A)+integral+(M,f|B) by A1,A2,A3,Th81, XBOOLE_1:79; then integral+(M,f) = 0.+ integral+(M,f|B) by A1,A2,A3,A4,A6,Th82; hence thesis by XXREAL_3:4; end; theorem Th85: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL st (ex E be 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 be 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, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL such that A1: ex A be 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 be Element of X st x in dom g holds g.x <= f.x; consider G be Functional_Sequence of X,ExtREAL, L be ExtREAL_sequence such that A5: for n be Nat holds G.n is_simple_func_in S & dom(G.n) = dom g and A6: for n be Nat holds G.n is nonnegative and A7: for n,m be Nat st n <=m holds for x be Element of X st x in dom g holds (G.n).x <= (G.m).x and A8: for x be Element of X st x in dom g holds G#x is convergent & lim(G #x) = g.x and A9: for n be 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 be Functional_Sequence of X,ExtREAL, K be ExtREAL_sequence such that A11: for n be Nat holds F.n is_simple_func_in S & dom(F.n) = dom f and A12: for n be Nat holds F.n is nonnegative and A13: for n,m be Nat st n <=m holds for x be Element of X st x in dom f holds (F.n).x <= (F.m).x and A14: for x be Element of X st x in dom f holds (F#x) is convergent & lim (F#x) = f.x and A15: for n be 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 be 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 be Element of X st x in A holds lim(G#x)=sup rng(G#x) proof let x be Element of X; assume A20: x in A; now let n,m be Nat; assume A21: n<=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; end; hence thesis by Th54; end; A23: for n0 be Nat holds L is convergent & sup rng L=lim L proof let n0 be Nat; set ff = G.n0; A24: dom ff = A by A18,A5; A25: for x be Element of X st x in dom ff holds G#x is convergent & ff.x <= lim(G#x) proof let x be Element of X such that A26: x in dom ff; A27: (G#x).n0 <= sup rng (G#x) by Th56; ff.x =(G#x).n0 by Def13; hence thesis by A18,A8,A19,A24,A26,A27; end; ff is_simple_func_in S by A5; then consider FF be ExtREAL_sequence such that A28: for n be Nat holds FF.n = integral'(M,G.n) and A29: FF is convergent and A30: sup rng FF = lim FF and integral'(M,ff) <= lim FF by A18,A5,A6,A7,A24,A25,Th75; now let n be Element of NAT; thus FF.n = integral'(M,G.n) by A28 .=L.n by A9; end; then FF=L by FUNCT_2:63; hence thesis by A29,A30; end; for n0 be Nat holds K is convergent & sup rng K = lim K & L.n0 <= lim K proof let n0 be Nat; set gg = G.n0; A31: gg is nonnegative by A6; A32: dom gg = A by A18,A5; A33: for x be Element of X st x in dom gg holds F#x is convergent & gg.x <= lim(F#x) proof let x be Element of X such that A34: x in dom gg; A35: (G#x).n0 <= sup rng (G#x) by Th56; gg.x =(G#x).n0 by Def13; then gg.x <= lim(G#x) by A19,A32,A34,A35; then A36: gg.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 thesis by A17,A14,A32,A34,A36,XXREAL_0:2; end; gg is_simple_func_in S by A5; then consider GG be ExtREAL_sequence such that A37: for n be Nat holds GG.n = integral'(M,F.n) and A38: GG is convergent and A39: sup rng GG =lim GG and A40: integral'(M,gg) <= lim GG by A17,A11,A12,A13,A32,A31,A33,Th75; now let n be Element of NAT; GG.n = integral'(M,F.n) by A37; hence GG.n = K.n by A15; end; then GG=K by FUNCT_2:63; hence thesis by A9,A38,A39,A40; end; hence thesis by A16,A10,A23,Th57; end; theorem Th86: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, c be Real st 0 <= c & (ex A be Element of S st A = dom f & f is_measurable_on A) & f is nonnegative holds integral+(M,c(#)f ) = c * integral+(M,f) proof let X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, c be Real such that A1: 0 <= c and A2: ex A be Element of S st A = dom f & f is_measurable_on A and A3: f is nonnegative; consider F1 be Functional_Sequence of X,ExtREAL, K1 be ExtREAL_sequence such that A4: for n be Nat holds F1.n is_simple_func_in S & dom(F1.n) = dom f and A5: for n be Nat holds F1.n is nonnegative and A6: for n,m be Nat st n <=m holds for x be Element of X st x in dom f holds (F1.n).x <= (F1.m).x and A7: for x be Element of X st x in dom f holds F1#x is convergent & lim( F1#x) = f.x and A8: for n be Nat holds K1.n=integral'(M,F1.n) and K1 is convergent and A9: integral+(M,f)=lim K1 by A2,A3,Def15; deffunc PF(Nat) = c(#)(F1.$1); consider F be Functional_Sequence of X,ExtREAL such that A10: for n be Nat holds F.n=PF(n) from SEQFUNC:sch 1; A11: c(#)f is nonnegative by A1,A3,Th20; A12: for n be Nat holds F.n is nonnegative proof let n be Nat; reconsider 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 thesis by A10; end; consider A be 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 be Nat holds F.n is_simple_func_in S & dom(F.n) = dom(c(#)f) proof let n be Nat; 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; thus dom(F.n) = dom(F1.n) by A17,MESFUNC1:def 6 .=A by A4,A13 .=dom(c(#)f) by A13,MESFUNC1:def 6; end; A18: for n,m be Nat st n<=m holds K1.n <= K1.m proof let n,m be Nat; 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; A27: for x be object st x in dom(F1.m - F1.n) holds (F1.n).x <= (F1.m).x proof let x be object; assume x in dom(F1.m - F1.n); 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 thesis by A6,A26,A23,A25; 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 thesis by A19,A20,A28,A21,A22,A24,A27,A30,Th70; end; deffunc PK(Nat) = integral'(M,F.$1); consider K be ExtREAL_sequence such that A31: for n be Element of NAT holds K.n = PK(n) from FUNCT_2:sch 4; A32: now let n be Nat; n in NAT by ORDINAL1:def 12; hence K.n = PK(n) by A31; end; A33: for n be Nat holds K.n=(c)*(K1.n) proof let n be Nat; 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 .= c * integral'(M,F1.n) by A1,A5,A34,A35,Th66 .= c * K1.n by A8; end; A36: A = dom(c(#)f) by A13,MESFUNC1:def 6; A37: for x be 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; now let n1 be set; assume n1 in dom(F1#x); 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; end; then A39: F1#x is without-infty by Th10; assume A40: x in dom(c(#)f); A41: now let n be Nat; 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 .= c * (F1.n).x by A40,A42,MESFUNC1:def 6 .= c * (F1#x).n by Def13; end; A43: now let n,m be Nat; assume A44: n <=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; end; c *lim(F1#x) = c * f.x by A7,A13,A36,A40 .=(c(#)f).x by A40,MESFUNC1:def 6; hence thesis by A1,A41,A39,A43,Th63; end; now let n1 be set; assume n1 in dom K1; 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; end; then A47: K1 is without-infty by Th10; then A48: lim K = c * lim K1 by A1,A18,A33,Th63; A49: for n,m be Nat st n <=m holds for x be Element of X st x in dom (c(#)f) holds (F.n).x <= (F.m).x proof let n,m be Nat; assume A50: n<=m; reconsider n,m as Element of NAT by ORDINAL1:def 12; let x be Element of X; assume A51: x in dom(c(#)f); 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 =(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 =(c)*(F1.n).x by A51,A54,MESFUNC1:def 6; hence thesis by A1,A6,A13,A36,A50,A51,A53,XXREAL_3:71; end; K is convergent by A1,A47,A18,A33,Th63; hence thesis by A9,A36,A15,A11,A32,A16,A12,A49,A37,A48,Def15; end; theorem Th87: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL st (ex A be Element of S st A = dom f & f is_measurable_on A) & (for x be Element of X st x in dom f holds 0= f.x) holds integral+(M,f) = 0 proof let X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL such that A1: ex A be Element of S st A = dom f & f is_measurable_on A and A2: for x be Element of X st x in dom f holds 0 = f.x; A3: for x be object st x in dom f holds 0 <= f.x by A2; A4: dom(0(#)f) =dom f by MESFUNC1:def 6; now let x be object; assume A5: x in dom(0(#)f); hence (0(#)f).x = 0 * f.x by MESFUNC1:def 6 .= 0 .= f.x by A2,A4,A5; end; hence integral+(M,f) = integral+(M,0(#)f) by A4,FUNCT_1:2 .= 0 * integral+(M,f) by A1,A3,Th86,SUPINF_2:52 .= 0; end; begin :: Integral of measurable function 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; ::$N Lebesgue Measure and Integration ::$N Integral of Measurable Function func Integral(M,f) -> Element of ExtREAL equals integral+(M,max+f)-integral+ (M,max-f); coherence; end; theorem Th88: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL st (ex A be 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, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL; assume that A1: ex A be Element of S st A = dom f & f is_measurable_on A and A2: f is nonnegative; A3: dom f = dom max+ f by MESFUNC2:def 2; A4: now let x be object; A5: 0 <= f.x by A2,SUPINF_2:51; assume x in dom f; hence max+f.x = max(f.x,0) by A3,MESFUNC2:def 2 .=f.x by A5,XXREAL_0:def 10; end; A6: dom f = dom max-f by MESFUNC2:def 3; A7: now let x be Element of X; assume x in dom max- f; then max+f.x=f.x by A4,A6; hence max-f.x=0 by MESFUNC2:19; 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) - 0 by A1,A7,A8,Th87,MESFUNC2:26 .=integral+(M,f) by XXREAL_3:15; end; theorem for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be 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, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL; assume that A1: f is_simple_func_in S and A2: f is nonnegative; 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; hence thesis by A1,A2,Th77; end; theorem for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL st (ex A be 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, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL; assume that A1: ex A be Element of S st A = dom f & f is_measurable_on A and A2: f is nonnegative; 0 <= integral+(M,f) by A1,A2,Th79; hence thesis by A1,A2,Th88; end; theorem for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A,B be Element of S st (ex E be 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, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A,B be Element of S; assume that A1: ex E be Element of S st E = dom f & f is_measurable_on E and A2: f is nonnegative and A3: A misses B; consider E be Element of S such that A4: E = dom f and A5: f is_measurable_on E by A1; ex C be Element of S st C = dom(f|A) & f|A is_measurable_on C proof take C=E/\A; thus dom(f|A) = C by A4,RELAT_1:61; A6: C = dom f /\ C by A4,XBOOLE_1:17,28; A7: dom(f|A) = C by A4,RELAT_1:61 .= dom(f|C) by A6,RELAT_1:61; for x be object st x in dom(f|A) holds (f|A).x = (f|C).x proof let x be object; assume A8: x in dom(f|A); then (f|A).x = f.x by FUNCT_1:47; hence thesis by A7,A8,FUNCT_1:47; 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 thesis by A6,A9,Th42; end; then A10: Integral(M,f|A)=integral+(M,f|A) by A2,Th15,Th88; ex C be 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; thus dom(f|(A\/B)) = C by A4,RELAT_1:61; A11: C = dom f /\ C by A4,XBOOLE_1:17,28; A12: dom(f|(A\/B)) = C by A4,RELAT_1:61 .= dom(f|C) by A11,RELAT_1:61; A13: for x be object st x in dom(f|(A\/B)) holds (f|(A\/B)).x = (f|C).x proof let x be object; assume A14: x in dom(f|(A\/B)); then (f|(A\/B)).x = f.x by FUNCT_1:47; hence thesis by A12,A14,FUNCT_1:47; end; f is_measurable_on C by A5,MESFUNC1:30,XBOOLE_1:17; then f|C is_measurable_on C by A11,Th42; hence thesis by A12,A13,FUNCT_1:2; end; then A15: Integral(M,f|(A\/B))=integral+(M,f|(A\/B)) by A2,Th15,Th88; A16: ex C be Element of S st C = dom(f|B) & f|B is_measurable_on C proof take C=E/\B; thus dom(f|B) = C by A4,RELAT_1:61; A17: C = dom f /\ C by A4,XBOOLE_1:17,28; A18: dom(f|B) = C by A4,RELAT_1:61 .= dom(f|C) by A17,RELAT_1:61; for x be object st x in dom(f|B) holds (f|B).x = (f|C).x proof let x be object; assume A19: x in dom(f|B); then (f|B).x = f.x by FUNCT_1:47; hence thesis by A18,A19,FUNCT_1:47; 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 thesis by A17,A20,Th42; end; integral+(M,f|(A\/B)) = integral+(M,f|A)+integral+(M,f|B) by A1,A2,A3,Th81; hence thesis by A2,A15,A10,A16,Th15,Th88; end; theorem for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A be Element of S st (ex E be 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, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A be Element of S; assume that A1: ex E be Element of S st E = dom f & f is_measurable_on E and A2: f is nonnegative; consider E be Element of S such that A3: E = dom f and A4: f is_measurable_on E by A1; A5: ex C be Element of S st C = dom(f|A) & f|A is_measurable_on C proof take C = E /\ A; thus dom(f|A) = C by A3,RELAT_1:61; A6: C = dom f /\ C by A3,XBOOLE_1:17,28; A7: dom(f|A) = C by A3,RELAT_1:61 .= dom(f|C) by A6,RELAT_1:61; A8: for x be object st x in dom(f|A) holds (f|A).x = (f|C).x proof let x be object; assume A9: x in dom(f|A); then (f|A).x = f.x by FUNCT_1:47; hence thesis by A7,A9,FUNCT_1:47; end; f is_measurable_on C by A4,MESFUNC1:30,XBOOLE_1:17; then f|C is_measurable_on C by A6,Th42; hence thesis by A7,A8,FUNCT_1:2; end; then 0<= integral+(M,f|A) by A2,Th15,Th79; hence thesis by A2,A5,Th15,Th88; end; theorem for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A,B be Element of S st (ex E be 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, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A,B be Element of S; assume that A1: ex E be Element of S st E = dom f & f is_measurable_on E and A2: f is nonnegative and A3: A c= B; consider E be Element of S such that A4: E = dom f and A5: f is_measurable_on E by A1; A6: ex C be Element of S st C = dom(f|A) & f|A is_measurable_on C proof take C = E /\ A; thus dom(f|A) = C by A4,RELAT_1:61; A7: C = dom f /\ C by A4,XBOOLE_1:17,28; A8: dom(f|A) = C by A4,RELAT_1:61 .= dom(f|C) by A7,RELAT_1:61; A9: for x be object st x in dom(f|A) holds (f|A).x = (f|C).x proof let x be object; assume A10: x in dom(f|A); then (f|A).x = f.x by FUNCT_1:47; hence thesis by A8,A10,FUNCT_1:47; end; f is_measurable_on C by A5,MESFUNC1:30,XBOOLE_1:17; then f|C is_measurable_on C by A7,Th42; hence thesis by A8,A9,FUNCT_1:2; end; A11: ex C be Element of S st C = dom(f|B) & f|B is_measurable_on C proof take C = E /\ B; thus dom(f|B) = C by A4,RELAT_1:61; A12: C = dom f /\ C by A4,XBOOLE_1:17,28; A13: dom(f|B) = C by A4,RELAT_1:61 .= dom(f|C) by A12,RELAT_1:61; A14: for x be object st x in dom(f|B) holds (f|B).x = (f|C).x proof let x be object; assume A15: x in dom(f|B); then (f|B).x = f.x by FUNCT_1:47; hence thesis by A13,A15,FUNCT_1:47; end; f is_measurable_on C by A5,MESFUNC1:30,XBOOLE_1:17; then f|C is_measurable_on C by A12,Th42; hence thesis by A13,A14,FUNCT_1:2; 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 thesis by A2,A11,Th15,Th88; end; theorem for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A be Element of S st (ex E be 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, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A be Element of S such that A1: ex E be Element of S st E = dom f & f is_measurable_on E and A2: M.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 thesis; end; theorem Th95: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, E,A be 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, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, E,A be Element of S such that A1: E = dom f and A2: f is_measurable_on E and A3: M.A =0; 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|B) = integral+(M,(max+f)|B) -integral+(M,max-(f|B)) by Th28 .= integral+(M,(max+f)|B) -integral+(M,(max-f)|B) by Th28 .=integral+(M,max+f) -integral+(M,(max-f)|B) by A1,A2,A3,A4,A6,Th84, MESFUNC2:25; hence thesis by A1,A2,A3,A7,A5,Th84,MESFUNC2:26; 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; pred f is_integrable_on M means (ex A be Element of S st A = dom f & f is_measurable_on A ) & integral+(M,max+ f) < +infty & integral+(M,max- f) < +infty; end; theorem Th96: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be 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, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL such that A1: f is_integrable_on M; consider A be Element of S such that A2: A = dom f and A3: f is_measurable_on A by A1; A4: integral+(M,max+f) <> +infty by A1; 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 Element of 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; 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 Element of REAL by A8,XXREAL_0:14; integral+(M,max+f)-integral+(M,max-f) = maxf1-maxf2 by SUPINF_2:3; hence thesis by A2,A3,A5,A9,A6,A10,A7,Th79,MESFUNC2:26,XXREAL_0:9,12; end; theorem Th97: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A be 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, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A be Element of S; A1: max+f is nonnegative by Lm1; assume A2: f is_integrable_on M; then consider E be Element of S such that A3: E = dom f and A4: f is_measurable_on E; 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,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; A9: max-f is nonnegative by Lm1; A10: integral+(M,max+f) < +infty by A2; 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 thesis by A13,A18,A7,A14,A19,Th28; end; theorem Th98: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A,B be 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, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A,B be Element of S; assume that A1: f is_integrable_on M and A2: A misses B; consider E be Element of S such that A3: E = dom f and A4: f is_measurable_on E by A1; 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|AB 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; then reconsider g2 = integral+(M,max-(f|B)) as Element of REAL by A13,XXREAL_0:14; integral+(M,max+(f|B)) < +infty by A11; then reconsider g1 = integral+(M,max+(f|B)) as Element of 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|AB 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; then reconsider f1 = integral+(M,max+(f|A)) as Element of 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; then reconsider f2 = integral+(M,max-(f|A)) as Element of 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|AB) by RELAT_1:71 .= integral+(M,max+f|AB) - integral+(M,max-(f|AB)) by Th28 .= integral+(M,max+f|AB) - integral+(M,max-f|AB) 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 thesis by A24,A14,SUPINF_2:1; end; theorem for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A,B be 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, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, A,B be Element of S such that A1: f is_integrable_on M and A2: B = dom f \ A; A \/ B = A\/dom f by A2,XBOOLE_1:39; then A3: dom f /\ (A \/ B) = dom f by XBOOLE_1:7,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 thesis by A1,A4,Th97,Th98; end; theorem Th100: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL st (ex A be 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, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL; A1: dom |.f.| = dom max-|.f.| by MESFUNC2:def 3; A2: dom f = dom max-f by MESFUNC2:def 3; A3: now let x be object; assume x in dom |.f.|; then (|.f.|).x =|. f.x .| by MESFUNC1:def 10; hence 0 <= (|.f.|).x by EXTREAL1:14; 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 be Element of S st A = dom f & f is_measurable_on A; then consider A be 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 A15: now let x be object; assume A16: x in dom |.f.|; then (|.f.|).x =|. f.x .| by MESFUNC1:def 10; then A17: 0 <= (|.f.|).x by EXTREAL1:14; (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; end; now let x be Element of X; assume x in dom max-|.f.|; then (max+|.f.|).x=(|.f.|).x by A1,A15; hence (max-|.f.|).x=0 by MESFUNC2:19; 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; then A21: integral+(M,max+f) < +infty; A22: integral+(M,max-f) < +infty by A20; |.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; end; assume |.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 thesis by A7,A25; end; theorem for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL st f is_integrable_on M holds |. Integral(M,f) .| <= Integral(M,|.f.|) proof let X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL such that A1: f is_integrable_on M; A2: |.integral+(M,max+f)-integral+(M,max-f).| <= |.integral+(M,max+f).| + |.integral+(M,max-f).| by EXTREAL1:32; A3: dom f=dom max+f by MESFUNC2:def 2; A4: now let x be object; assume x in dom (|.f.|); then (|.f.|).x =|. f.x .| by MESFUNC1:def 10; hence 0 <= (|.f.|).x by EXTREAL1:14; end; A5: dom f = dom max-f by MESFUNC2:def 3; A6: |.f.| = max+f + max-f by MESFUNC2:24; consider A be Element of S such that A7: A = dom f and A8: f is_measurable_on A by A1; 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 thesis by A15,A11,EXTREAL1:def 1; end; theorem Th102: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL st ( ex A be Element of S st A = dom f & f is_measurable_on A ) & dom f = dom g & g is_integrable_on M & ( for x be 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, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL; assume that A1: ex A be 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 be Element of X st x in dom f holds |. f.x .| <= g.x; A5: ex AA be Element of S st AA = dom g & g is_measurable_on AA by A3; A6: now let x be object; assume x in dom g; then |. f.x .| <= g.x by A2,A4; hence 0 <= g.x by EXTREAL1:14; end; then A7: g is nonnegative by SUPINF_2:52; A8: dom g = dom max+ g by MESFUNC2:def 2; now let x be object; A9: 0 <= g.x by A7,SUPINF_2:51; assume x in dom g; hence max+g.x = max(g.x,0) by A8,MESFUNC2:def 2 .=g.x by A9,XXREAL_0:def 10; end; then A10: g = max+g by A8,FUNCT_1:2; A11: dom |.f.| = dom max+|.f.| by MESFUNC2:def 2; A12: now let x be object; assume A13: x in dom |.f.|; then (|.f.|).x =|. f.x .| by MESFUNC1:def 10; then A14: 0 <= (|.f.|).x by EXTREAL1:14; thus (max+|.f.|).x = max((|.f.|).x,0) by A11,A13,MESFUNC2:def 2 .=(|.f.|).x by A14,XXREAL_0:def 10; end; then A15: |.f.| = max+|.f.| by A11,FUNCT_1:2; consider A be 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 be Element of X st x in dom |.f.| holds (|.f.|).x <= g.x proof let x be Element of X; assume A21: x in dom |.f.|; then (|.f.|).x =|. f.x .| by MESFUNC1:def 10; hence thesis by A4,A16,A19,A21; end; A22: now let x be object; assume x in dom |.f.|; then (|.f.|).x =|. f.x .| by MESFUNC1:def 10; hence 0 <= (|.f.|).x by EXTREAL1:14; 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 let x be Element of X; assume x in dom max-(|.f.|); then max+(|.f.|).x=(|.f.|).x by A24,A12; hence max-(|.f.|).x=0 by MESFUNC2:19; end; then A25: integral+(M,max-|.f.|) = 0 by A19,A18,A24,Th87,MESFUNC2:26; integral+(M,max+g) < +infty by A3; then integral+(M,max+(|.f.|)) < +infty by A15,A10,A23,XXREAL_0:2; then |.f.| is_integrable_on M by A19,A18,A25; hence f is_integrable_on M by A1,Th100; Integral(M,g) =integral+(M,g) by A5,A6,Th88,SUPINF_2:52; hence thesis by A19,A18,A22,A23,Th88,SUPINF_2:52; end; theorem Th103: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, r be Real st dom f in S & 0 <= r & dom f <> {} & (for x be object st x in dom f holds f.x = r) holds integral(M,f) = r * M.(dom f) proof 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 r be Real; assume that A1: dom f in S and A2: 0 <= r and A3: dom f <> {} and A4: for x be object st x in dom f holds f.x = r; for x be object st x in dom f holds 0 <= f.x by A2,A4; then a5: f is nonnegative by SUPINF_2:52; f is_simple_func_in S by A1,A4,Lm4; then consider F be Finite_Sep_Sequence of S, a,v be FinSequence of ExtREAL such that A6: F,a are_Re-presentation_of f and A7: dom v = dom F and A8: for n be Nat st n in dom v holds v.n = a.n*(M*F).n and A9: integral(M,f) = Sum v by A3,a5,MESFUNC4:4; A10: dom f = union rng F by A6,MESFUNC3:def 1; A11: for n be Nat st n in dom v holds v.n = r * (M*F).n proof let n be Nat; assume A12: n in dom v; 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; per cases; suppose 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 thesis by A15; end; suppose F.n <> {}; then consider x be object 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 thesis by A4,A10,A13,A14,A16; end; end; reconsider rr=r as R_eal by XXREAL_0:def 1; dom v = dom(M*F) by A7,MESFUNC3:8; then integral(M,f) = rr * Sum(M*F) by A9,A11,MESFUNC3:10 .= rr * M.(union rng F) by MESFUNC3:9; hence thesis by A6,MESFUNC3:def 1; end; theorem Th104: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, r be Real st dom f in S & 0 <= r & (for x be object st x in dom f holds f.x = r) holds integral'(M,f) = r * M.(dom f) proof 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 r be Real; assume that A1: dom f in S and A2: 0 <= r and A3: for x be object st x in dom f holds f.x = r; per cases; suppose A4: dom f = {}; then A5: M.(dom f) = 0 by VALUED_0:def 19; integral'(M,f) = 0 by A4,Def14; hence thesis by A5; end; suppose A6: dom f <> {}; then integral'(M,f) = integral(M,f) by Def14; hence thesis by A1,A2,A3,A6,Th103; end; end; theorem Th105: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be 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; let S be SigmaField of X; let M be sigma_Measure of S; let f be PartFunc of X,ExtREAL; A1: max+f is nonnegative by Lm1; assume A2: f is_integrable_on M; then A3: integral+(M,max+f) < +infty; consider A be Element of S such that A4: A = dom f and A5: f is_measurable_on A by A2; A6: for x be object 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; then A8: A /\ eq_dom(f,+infty) = eq_dom(f,+infty) by XBOOLE_1:28; A9: eq_dom(f,-infty) c= A by A6; 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; set C2 = A \ B2; A15: integral+(M,max-f) < +infty by A2; 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 be Element of X holds ( x in dom(max+f|(B1\/C1)) implies max+f|( B1\/C1).x = max+f.x ) & ( x in dom(max-f|(B2\/C2)) implies max-f|(B2\/C2).x = max-f.x ) by FUNCT_1:47; B1\/C1 = A by A16,A17,XBOOLE_1:45; then dom(max+f|(B1\/C1)) = dom max+f /\ dom max+f by A16,RELAT_1:61; then max+f|(B1\/C1) = max+f by A21,PARTFUN1:5; then integral+(M,max+f) = integral+(M,max+f|B1) + integral+(M,max+f|C1) 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 ; thus now A23: for r be Real st 0 < r holds r * M.B1 <= integral+(M,max+f) proof defpred P[object] means $1 in dom(max+f|B1); let r be Real; deffunc F(object) = In(r,ExtREAL); A24: for x be object st P[x] holds F(x) in ExtREAL; consider g be PartFunc of X,ExtREAL such that A25: (for x be object holds x in dom g iff x in X & P[x]) & for x be object st x in dom g holds g.x = F(x) from PARTFUN1:sch 3(A24); assume A26: 0 < r; then for x be object 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; for x be object holds x in dom g iff x in X & x in dom(max+f|B1) by A25; then dom g = X /\ dom(max+f|B1) by XBOOLE_0:def 4; then A29: dom g = dom(max+f|B1) by XBOOLE_1:28; then A30: integral'(M,g) = r * M.(dom g) by A26,A25,A28,Th104; A31: for x be Element of X st x in dom g holds g.x <= max+f|B1.x proof let x be Element of X; assume A32: x in dom g; 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 thesis by XXREAL_0:4; 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 be Element of X st x in dom g holds g.x = (r(#)(chi(B1,X)|B1) ).x proof let x be Element of X; assume A39: x in dom g; 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 * (chi(B1,X)|B1).x by MESFUNC1:def 6 .= 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 by A40,XXREAL_3:81; hence thesis by A25,A39; 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 thesis by A25,A29,A28,A27,A30,Lm4,Th77; end; assume A42: M.(f"{+infty}) <> 0; then A43: 0 < M.(f"{+infty}) by A12,Th45; per cases; suppose A44: M.B1 = +infty; jj * M.B1 <= integral+(M,max+f) by A23; hence contradiction by A3,A44,XXREAL_3:81; end; suppose M.B1 <> +infty; then reconsider MB = M.B1 as Element of REAL by A43,XXREAL_0:14; jj * 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 Element of REAL by A3,XXREAL_0:14; A46: (2*I/MB) * M.B1 = 2*I/MB * MB; (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; end; end; then reconsider 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\/C2 = A by A48,A49,XBOOLE_1:45; then dom(max-f|(B2\/C2)) = dom max-f /\ dom max-f by A48,RELAT_1:61; then max-f|(B2\/C2) = max-f by A21,PARTFUN1:5; then integral+(M,max-f) = integral+(M,max-f|B2) + integral+(M,max-f|C2) 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; thus A54: now A55: for r be Real st 0 < r holds r * M.B2 <= integral+(M,max-f) proof defpred P[object] means $1 in dom(max-f|B2); let r be Real; deffunc F(object) = In(r,ExtREAL); A56: for x be object st P[x] holds F(x) in ExtREAL; consider g be PartFunc of X,ExtREAL such that A57: (for x be object holds x in dom g iff x in X & P[x]) & for x be object st x in dom g holds g.x = F(x) from PARTFUN1:sch 3(A56); assume A58: 0 < r; then for x be object 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; for x be object holds x in dom g iff x in X & x in dom(max-f|B2) by A57; then dom g = X /\ dom(max-f|B2) by XBOOLE_0:def 4; then A61: dom g = dom(max- f|B2) by XBOOLE_1:28; then A62: integral'(M,g) = 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 be Element of X st x in dom g holds g.x = (r(#)(chi(B2,X)|B2 )).x proof let x be Element of X; assume A66: x in dom g; 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 * (chi(B2,X)|B2).x by MESFUNC1:def 6 .= 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 by A67,XXREAL_3:81; hence thesis by A57,A66; end; A68: for x be Element of X st x in dom g holds g.x <= (max-f|B2).x proof let x be Element of X; assume A69: x in dom g; 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 thesis by XXREAL_0:4; 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 thesis by A57,A61,A60,A59,A62,Lm4,Th77; end; assume A74: M.(f"{-infty}) <> 0; A75: 0 <= M.(f"{-infty}) by A14,Th45; per cases; suppose A76: M.B2 = +infty; jj * M.B2 <= integral+(M,max-f) by A55; hence contradiction by A15,A76,XXREAL_3:81; end; suppose M.B2 <> +infty; then reconsider MB = M.B2 as Element of REAL by A75,XXREAL_0:14; jj * 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 Element of REAL by A15,XXREAL_0:14; A78: (2*I/MB) * M.B2 = 2*I/MB * MB; (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; end; end; thus f"{+infty} \/ f"{-infty} in S by A12,A14,PROB_1:3; thus M.(f"{+infty} \/ f"{-infty}) = M.(B1 \/ B2) .= 0 by A54,MEASURE1:38; end; theorem Th106: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f,g be 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, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL; 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; A5: integral+(M,max+g) < +infty by A2; A6: dom g = dom max+g by MESFUNC2:def 2; now let x be object; assume x in dom g; 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; end; then A8: g = max+g by A6,FUNCT_1:2; consider B be Element of S such that A9: B = dom g and A10: g is_measurable_on B by A2; consider A be Element of S such that A11: A = dom f and A12: f is_measurable_on A by A1; 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 be 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; 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 let x be object; assume x in dom f; 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; 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 let x be object; assume A25: x in dom(f+g); 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; end; then A29: f+g = max+(f+g) by A23,FUNCT_1:2; A30: now let x be Element of X; assume x in dom max-(f+g); 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; 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 thesis by A17,A14,A31; end; theorem Th107: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f,g be 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, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL; assume that A1: f is_integrable_on M and A2: g is_integrable_on M; A3: f"{-infty} in S by A1,Th105; A4: ex E2 be Element of S st E2=dom g & g is_measurable_on E2 by A2; A5: ex E1 be Element of S st E1=dom f & f is_measurable_on E1 by A1; A6: g"{-infty} in S by A2,Th105; A7: g"{+infty} in S by A2,Th105; f"{+infty} in S by A1,Th105; hence thesis by A3,A7,A6,A5,A4,Th46; end; Lm11: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f ,g be PartFunc of X,ExtREAL st (ex E0 be 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, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL; assume that A1: ex E0 be 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; consider E be 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; |.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,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,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 let x be Element of X; A18: |.f.x+g.x.| <= |.f.x.| +|.g.x.| by EXTREAL1:24; assume A19: x in dom(f+g); 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; end; |.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 thesis by A4,A5,A17,A15,Th102; end; theorem Th108: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f,g be 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, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL such that A1: f is_integrable_on M and A2: g is_integrable_on M; A3: ex E2 be Element of S st E2=dom g & g is_measurable_on E2 by A2; ex E1 be Element of S st E1=dom f & f is_measurable_on E1 by A1; then ex K0 be Element of S st K0 = dom(f+g) & f+g is_measurable_on K0 by A3 ,Th47; hence thesis by A1,A2,Lm11; end; Lm12: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f ,g be PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M & dom f = dom g holds ex E,NFG,NFPG be Element of S st E c= dom f & NFG c= dom f & E = dom f \ NFG & f|E is real-valued & 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 real-valued & 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, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL; assume that A1: f is_integrable_on M and A2: g is_integrable_on M and A3: dom f = dom g; 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 be Element of S such that A5: E0=dom f and A6: f is_measurable_on E0 by A1; set E = E0 \ NFG; set f1=f|E; set g1=g|E; A7: E 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 be object 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; set NFPG=DFPG \ E; A11: dom(f+g) = (dom f /\ dom g)\( f"{-infty}/\g"{+infty} \/ f"{+infty}/\g"{ -infty} ) by MESFUNC1:def 3; then DFPG \ (E0 \ NFG) c=E0 \ (E0 \ NFG) by A3,A5,XBOOLE_1:33,36; then A12: NFPG c= E0 /\ NFG by XBOOLE_1:48; g"{-infty} c= NG by XBOOLE_1:7; then A13: f"{+infty}/\g"{-infty} c= NG by A4; f"{-infty} c= NF by XBOOLE_1:7; then f"{-infty}/\g"{+infty} c= NF by A8; then A14: f"{-infty}/\g"{+infty} \/ f"{+infty}/\g"{-infty} c= NF \/ NG by A13, XBOOLE_1:13; then A15: E c= dom(f+g) by A3,A5,A11,XBOOLE_1:34; then A16: (f+g)|E = f1+g1 by Th29; DFPG \ NFPG = DFPG /\ E by XBOOLE_1:48; then A17: E= DFPG \ NFPG by A3,A5,A11,A14,XBOOLE_1:28,34; A18: dom(f1+g1)=E by A15,Th29; A19: for x be set st x in dom g1 holds -infty < g1.x & g1.x < +infty proof let x be set; for x be object st x in dom g holds g.x in ExtREAL by XXREAL_0:def 1; then reconsider gg=g as Function of dom g, ExtREAL by FUNCT_2:3; assume A20: x in dom g1; then A21: x in dom g /\ E by RELAT_1:61; then A22: x in dom g by XBOOLE_0:def 4; x in E by A21,XBOOLE_0:def 4; then A23: not x in NFG by XBOOLE_0:def 5; A24: now assume g1.x = -infty; 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; end; now assume g1.x = +infty; 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; end; hence thesis by A24,XXREAL_0:4,6; end; then for x be set st x in dom g1 holds -infty < g1.x; then A27: g1 is without-infty by Th10; now let x be Element of X; A28: -(+infty) = -infty by XXREAL_3:def 3; assume A29: x in dom g1; then A30: g1.x < +infty by A19; -infty < g1.x by A19,A29; hence |.g1.x.| < +infty by A30,A28,EXTREAL1:22; end; then A31: g1 is real-valued by MESFUNC2:def 1; A32: for x be set st x in dom f1 holds f1.x < +infty & -infty < f1.x proof let x be set; for x be object st x in dom f holds f.x in ExtREAL by XXREAL_0:def 1; then reconsider ff=f as Function of dom f, ExtREAL by FUNCT_2:3; assume A33: x in dom f1; then A34: x in dom f /\ E by RELAT_1:61; then A35: x in dom f by XBOOLE_0:def 4; x in E by A34,XBOOLE_0:def 4; then A36: not x in NFG by XBOOLE_0:def 5; A37: now assume f1.x = -infty; 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; end; now assume f1.x = +infty; 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; end; hence thesis by A37,XXREAL_0:4,6; end; then for x be set st x in dom f1 holds -infty < f1.x; then A40: f1 is without-infty by Th10; then A41: dom(max-(f1+g1) + max+f1) = dom f1 /\ dom g1 by A27,Th24; A42: max+(f1+g1) + max-f1 is nonnegative by A40,A27,Th24; A43: dom(max+(f1+g1) + max-f1) = dom f1 /\ dom g1 by A40,A27,Th24; A44: max-(f1+g1) + max+f1 is nonnegative by A40,A27,Th24; A45: max+f1 is nonnegative by Lm1; A46: dom(max+(f1+g1))= dom(f1+g1) by MESFUNC2:def 2; A47: dom g1 = dom g /\ E by RELAT_1:61; then A48: E = dom g1 by A3,A5,XBOOLE_1:28,36; then A49: dom(max-g1) = E by MESFUNC2:def 3; A50: ex Gf be Element of S st Gf=dom g & g is_measurable_on Gf by A2; then g is_measurable_on E by A3,A5,MESFUNC1:30,XBOOLE_1:36; then A51: g1 is_measurable_on E by A47,A48,Th42; then A52: max-g1 is_measurable_on E by A48,MESFUNC2:26; A53: dom(max+g1) = E by A48,MESFUNC2:def 2; A54: max+g1 is nonnegative by Lm1; A55: max-g1 is nonnegative by Lm1; A56: dom f1 = dom f /\ E by RELAT_1:61; then A57: E = dom f1 by A5,XBOOLE_1:28,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: NFG is measure_zero of M by A58,MEASURE1:37; then A60: M.NFG=0 by MEASURE1:def 7; then A61: Integral(M,f) =Integral(M,f1) by A5,A6,Th95; E0 /\ NFG c= NFG by XBOOLE_1:17; then NFPG is measure_zero of M by A59,A12,MEASURE1:36,XBOOLE_1:1; then A62: M.NFPG=0 by MEASURE1:def 7; A63: max-(f1+g1) is nonnegative by Lm1; A64: max+(f1+g1) is nonnegative by Lm1; for x be set st x in dom g1 holds g1.x < +infty by A19; then A65: g1 is without+infty by Th11; A66: dom(max+f1)= dom f1 by MESFUNC2:def 2; for x be set st x in dom g1 holds -infty < g1.x by A19; then A67: g1 is without-infty by Th10; A68: dom(max-f1)= dom f1 by MESFUNC2:def 3; A69: f"{-infty} c= dom f by A9; g"{+infty} c= dom g by A9; then A70: NG c= dom g by A10,XBOOLE_1:8; f"{+infty} c= dom f by A9; then NF c= dom g by A3,A69,XBOOLE_1:8; then A71: NF \/ NG c= dom g by A70,XBOOLE_1:8; A72: NFPG c= dom(f+g) by XBOOLE_1:36; A73: g1 is_integrable_on M by A2,Th97; then A74: 0 <= integral+(M,max+g1) by Th96; for x be set st x in dom f1 holds f1.x < +infty by A32; then A75: f1 is without+infty by Th11; for x be set st x in dom f1 holds -infty < f1.x by A32; then f1 is without-infty by Th10; then A76: max+(f1+g1) + max-f1 + max-g1 = max-(f1+g1) + max+f1 + max+g1 by A75,A67 ,A65,Th25; A77: max-f1 is nonnegative by Lm1; A78: dom(max-(f1+g1))= dom(f1+g1) by MESFUNC2:def 3; A79: integral+(M,max+g1) <> +infty by A73; A80: 0 <= integral+(M,max-g1) by A73,Th96; f is_measurable_on E by A6,MESFUNC1:30,XBOOLE_1:36; then A81: f1 is_measurable_on E by A56,A57,Th42; then A82: max-f1 is_measurable_on E by A57,MESFUNC2:26; now let x be Element of X; A83: -(+infty) = -infty by XXREAL_3:def 3; assume A84: x in dom f1; then A85: f1.x < +infty by A32; -infty < f1.x by A32,A84; hence |.f1.x.| < +infty by A85,A83,EXTREAL1:22; end; then A86: f1 is real-valued by MESFUNC2:def 1; then A87: f1+g1 is_measurable_on E by A81,A51,A31,MESFUNC2:7; then A88: max+(f1+g1) is_measurable_on E by MESFUNC2:25; dom f1 /\ dom g1 = E by A3,A5,A56,A47,XBOOLE_1:28,36; then A89: max-(f1+g1) + max+f1 is_measurable_on E by A81,A51,A40,A27,Th44; E =dom(f1+g1) by A15,Th29; then A90: max-(f1+g1) is_measurable_on E by A87,MESFUNC2:26; A91: max+f1 is_measurable_on E by A81,MESFUNC2:25; A92: integral+(M,max-g1) <> +infty by A73; max+(f1+g1) + max-f1 is_measurable_on E by A57,A81,A51,A40,A27,Th43; then A93: integral+(M,max+(f1+g1)+max-f1+max-g1) =integral+(M,max+(f1+g1)+max-f1 ) + integral+(M,max-g1) by A57,A48,A43,A49,A42,A55,A52,Lm10 .=integral+(M,max+(f1+g1)) + integral+(M,max-f1) + integral+(M,max-g1) by A18,A57,A68,A46,A77,A64,A88,A82,Lm10; max+g1 is_measurable_on E by A51,MESFUNC2:25; then integral+(M,max-(f1+g1)+max+f1+max+g1) =integral+(M,max-(f1+g1) + max+ f1) + integral+(M,max+g1) by A57,A48,A41,A53,A44,A54,A89,Lm10 .=integral+(M,max-(f1+g1)) + integral+(M,max+f1) + integral+(M,max+g1) by A18,A57,A66,A78,A45,A63,A90,A91,Lm10; then integral+(M,max+(f1+g1))+integral+(M,max-f1)+ (integral+(M,max-g1) - integral+(M,max-g1) ) = integral+(M,max-(f1+g1))+integral+(M,max+f1) + integral+(M,max+g1) - integral+(M,max-g1) by A76,A80,A92,A93,XXREAL_3:30; then integral+(M,max+(f1+g1))+integral+(M,max-f1)+ (integral+(M,max-g1) - integral+(M,max-g1) ) = integral+(M,max-(f1+g1))+integral+(M,max+f1) + ( integral+(M,max+g1) -integral+(M,max-g1)) by A74,A79,A80,A92,XXREAL_3:30; then integral+(M,max+(f1+g1))+integral+(M,max-f1)+ 0. = integral+(M,max-(f1 +g1))+integral+(M,max+f1) + (integral+(M,max+g1) -integral+(M,max-g1)) by XXREAL_3:7; then A94: integral+(M,max+(f1+g1))+integral+(M,max-f1) = integral+(M,max-(f1+g1) )+integral+(M,max+f1)+(integral+(M,max+g1) - integral+(M,max-g1)) by XXREAL_3:4; A95: f1 is_integrable_on M by A1,Th97; then A96: 0 <= integral+(M,max+f1) by Th96; A97: f1+g1 is_integrable_on M by A95,A73,Th108; then A98: integral+(M,max+(f1+g1)) <> +infty; A99: integral+(M,max-(f1+g1)) <> +infty by A97; then A100: -integral+(M,max-(f1+g1)) <> -infty by XXREAL_3:23; A101: 0 <= integral+(M,max-(f1+g1)) by A97,Th96; A102: integral+(M,max-f1) <> +infty by A95; then A103: -integral+(M,max-f1) <> -infty by XXREAL_3:23; A104: integral+(M,max+f1) <> +infty by A95; A105: 0 <= integral+(M,max-f1) by A95,Th96; 0 <= integral+(M,max+(f1+g1)) by A97,Th96; then -integral+(M,max-(f1+g1))+ integral+(M,max+(f1+g1))+integral+(M,max-f1 ) = -integral+(M,max-(f1+g1))+ ( integral+(M,max-(f1+g1)) +integral+(M,max+f1)+ (integral+(M,max+g1) -integral+(M,max-g1)) ) by A105,A102,A98,A94,XXREAL_3:29 ; then -integral+(M,max-(f1+g1))+ integral+(M,max+(f1+g1))+ integral+(M,max- f1) = -integral+(M,max-(f1+g1))+(integral+(M,max-(f1+g1)) +(integral+(M,max+f1) +(integral+(M,max+g1)-integral+(M,max-g1)))) by A96,A104,A101,A99,XXREAL_3:29 ; then -integral+(M,max-(f1+g1))+ integral+(M,max+(f1+g1))+ integral+(M,max- f1) = -integral+(M,max-(f1+g1))+integral+(M,max-(f1+g1)) +(integral+(M,max+f1)+ (integral+(M,max+g1)-integral+(M,max-g1))) by A101,A99,A100,XXREAL_3:29; then -integral+(M,max-(f1+g1))+ integral+(M,max+(f1+g1))+ integral+(M,max- f1) = 0 + (integral+(M,max+f1) + (integral+(M,max+g1) -integral+(M,max-g1 ))) by XXREAL_3:7; then -integral+(M,max-(f1+g1)) + integral+(M,max+(f1+g1))+ integral+(M,max- f1) = integral+(M,max+f1) + (integral+(M,max+g1) -integral+(M,max-g1)) by XXREAL_3:4; then -integral+(M,max-f1)+ integral+(M,max-f1) +(-integral+(M,max-(f1+g1)) + integral+(M,max+(f1+g1))) = -integral+(M,max-f1)+(integral+(M,max+f1) +( integral+(M,max+g1)-integral+(M,max-g1))) by A105,A102,A103,XXREAL_3:29; then -integral+(M,max-f1) + integral+(M,max-f1) + (-integral+(M,max-(f1+g1) ) + integral+(M,max+(f1+g1))) = -integral+(M,max-f1)+ integral+(M,max+f1)+( integral+(M,max+g1) -integral+(M,max-g1)) by A96,A104,A105,A103,XXREAL_3:29; then 0 + (-integral+(M,max-(f1+g1)) + integral+(M,max+(f1+g1))) = (- integral+(M,max-f1)+integral+(M,max+f1))+(integral+(M,max+g1) -integral+(M,max- g1)) by XXREAL_3:7; then A106: Integral(M,(f1+g1))=Integral(M,f1)+Integral(M,g1) by XXREAL_3:4; Integral(M,g) =Integral(M,g1) by A3,A5,A50,A60,Th95; hence thesis by A3,A5,A60,A71,A15,A16,A18,A62,A17,A72,A7,A57,A48,A86 ,A31,A87,A95,A73,A106,A61,Th108; end; Lm13: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f ,g be 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, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL; assume that A1: f is_integrable_on M and A2: g is_integrable_on M and A3: dom f = dom g; thus f+g is_integrable_on M by A1,A2,Th108; then A4: ex K0 be Element of S st K0 = dom(f+g) & f+g is_measurable_on K0; ex E,NFG,NFPG be Element of S st E c= dom f & NFG c= dom f & E = (dom f) \ NFG & f|E is real-valued & 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 real-valued & 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 thesis by A4,Th95; end; theorem for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL st f is_integrable_on M & g is_integrable_on M holds ex E be 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, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL; assume that A1: f is_integrable_on M and A2: g is_integrable_on M; consider B be Element of S such that A3: B = dom g and g is_measurable_on B by A2; consider A be Element of S such that A4: A = dom f and f is_measurable_on A by A1; set E = A /\ B; set g1 = g|E; set f1 = f|E; take E = A /\ B; A5: dom f1 = dom f /\ (A/\B) by RELAT_1:61 .= A /\ A /\ B by A4,XBOOLE_1:16; A6: f1"{+infty} = E /\ (f"{+infty}) by FUNCT_1:70; g1"{-infty} = E /\ (g"{-infty}) by FUNCT_1:70; then A7: f1"{+infty} /\ g1"{-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: g1"{+infty} = E /\ (g"{+infty}) by FUNCT_1:70; f1"{-infty} = E /\ (f"{-infty}) by FUNCT_1:70; then f1"{-infty} /\ g1"{+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: f1"{-infty}/\g1"{+infty} \/ f1"{+infty}/\g1"{-infty} = E /\ (f"{-infty} /\g"{+infty} \/ f"{+infty}/\g"{-infty}) by A7,XBOOLE_1:23; A10: dom g1 = dom g /\ (A/\B) by RELAT_1:61 .= B /\ B /\ A by A3,XBOOLE_1:16; A11: dom(f1+g1) = (dom f1 /\ dom g1) \ (f1"{-infty}/\g1"{+infty} \/ f1"{ +infty}/\g1"{-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 be object st x in dom(f1+g1) holds (f1+g1).x = (f+g).x proof let x be object; assume A13: x in dom(f1+g1); then x in (dom f1 /\ dom g1) \ (f1"{-infty} /\ g1"{+infty} \/ f1"{+infty} /\g1"{-infty}) by MESFUNC1:def 3; then A14: x in dom f1 /\ dom g1 by XBOOLE_0:def 5; then A15: x in dom f1 by XBOOLE_0:def 4; A16: x in dom g1 by A14,XBOOLE_0:def 4; (f1+g1).x = f1.x + g1.x by A13,MESFUNC1:def 3 .= f.x + g1.x by A15,FUNCT_1:47 .= f.x + g.x by A16,FUNCT_1:47; hence thesis by A11,A13,MESFUNC1:def 3; end; thus E = dom f /\ dom g by A4,A3; A17: g1 is_integrable_on M by A2,Th97; f1 is_integrable_on M by A1,Th97; then Integral(M,f1+g1) = Integral(M,f1) + Integral(M,g1) by A17,A5,A10,Lm13; hence thesis by A11,A12,FUNCT_1:2; end; theorem Th110: for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, c be Real st f is_integrable_on M holds c(#)f is_integrable_on M & Integral(M,c(#)f) = c * Integral(M,f) proof let X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, c be Real such that A1: f is_integrable_on M; A2: integral+(M,max+f) <>+infty by A1; consider A be Element of S such that A3: A = dom f and A4: f is_measurable_on A by A1; 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; 0 <= integral+(M,max-f) by A1,Th96; then reconsider I2 = integral+(M,max-f) as Element of 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 Element of 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; per cases; suppose A12: 0 <= c; c*I1 in REAL by XREAL_0:def 1; then A13: c * integral+(M,max+f) in REAL; A14: max+(c(#)f)=c(#)max+f by A12,Th26; integral+(M,c(#)max+f) = 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 in REAL by XREAL_0:def 1; then c * integral+(M,max-f) is Element of REAL; then A16: 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) = 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; 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 .= c * integral+(M,max+f) - integral+(M,c(#)max-f) by A4,A9,A11,A12 ,Th86,MESFUNC2:25 .= c * integral+(M,max+f) - c *integral+(M,max-f) by A3,A4,A8 ,A6,A12,Th86,MESFUNC2:26 .= c * Integral(M,f) by XXREAL_3:100; end; suppose A18: c < 0; -(-c)=c; then consider a be 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 in REAL by XREAL_0:def 1; then A23: a *integral+(M,max-f) in REAL; integral+(M,a(#)max-f) = 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 in REAL by XREAL_0:def 1; then (a)*integral+(M,max+f) is Element of REAL; then A25: (a)*integral+(M,max+f) < +infty by XXREAL_0:9; integral+(M,a(#)max+f) = 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; thus Integral(M,c(#)f) = a * integral+(M,max-f) -integral+(M,a(#)max+ f) by A3,A4,A8,A6,A20,A21,A22,Th86,MESFUNC2:26 .= a * integral+(M,max-f)- a * integral+(M,max+f) by A4,A9,A11 ,A20,Th86,MESFUNC2:25 .= a * (integral+(M,max-f)-integral+(M,max+f)) by XXREAL_3:100 .= a * (-(integral+(M,max+f)-integral+(M,max-f))) by XXREAL_3:26 .=-( a * (integral+(M,max+f)-integral+(M,max-f))) by XXREAL_3:92 .= c * Integral(M,f) by A19,XXREAL_3:92; 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 Integral(M,f|B); coherence; end; theorem for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL, B be 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, S be SigmaField of X, M be sigma_Measure of S, f,g be PartFunc of X,ExtREAL, B be Element of S such that A1: f is_integrable_on M and A2: g is_integrable_on M and A3: B c= dom(f+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; then A6: dom(f|B) = B by A3,A4,XBOOLE_1: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; then A10: dom(g|B) = B by A3,A9,XBOOLE_1:1,28; f|B is_integrable_on M by A1,Th97; hence thesis by A1,A2,A6,A8,A10,A7,Lm13,Th108; end; theorem for X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, c be Real, B be Element of S st f is_integrable_on M holds f|B is_integrable_on M & Integral_on(M,B,c(#)f) = c * Integral_on (M,B,f) proof let X be non empty set, S be SigmaField of X, M be sigma_Measure of S, f be PartFunc of X,ExtREAL, c be Real, B be Element of S; assume f is_integrable_on M; then A1: f|B is_integrable_on M by Th97; A2: for x be object st x in dom((c(#)f)|B) holds (c(#)f)|B.x = (c(#)(f|B)).x proof let x be object; assume A3: x in dom ((c(#)f)|B); 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= c * f.x by A4,MESFUNC1:def 6; then A7: (c(#)f)|B.x= c * f|B.x by A6,FUNCT_1:47; x in dom (c(#)(f|B)) by A6,MESFUNC1:def 6; hence thesis by A7,MESFUNC1:def 6; 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 thesis by A1,Th110; end;