:: MEASURE4 semantic presentation

theorem Th1: :: MEASURE4:1
for b1, b2, b3 being R_eal holds
( 0. <= b1 & 0. <= b2 & 0. <= b3 implies (b1 + b2) + b3 = b1 + (b2 + b3) )
proof end;

theorem Th2: :: MEASURE4:2
for b1, b2, b3 being R_eal holds
( not b1 = -infty & not b1 = +infty implies ( b2 + b1 <= b3 iff b2 <= b3 - b1 ) )
proof end;

theorem Th3: :: MEASURE4:3
canceled;

theorem Th4: :: MEASURE4:4
for b1 being set
for b2 being non empty Subset-Family of b1
for b3, b4 being Function of NAT ,b2
for b5 being Element of b2 holds
( ( for b6 being Element of NAT holds b4 . b6 = b5 /\ (b3 . b6) ) implies union (rng b4) = b5 /\ (union (rng b3)) )
proof end;

theorem Th5: :: MEASURE4:5
for b1 being set
for b2 being non empty Subset-Family of b1
for b3, b4 being Function of NAT ,b2 holds
( b4 . 0 = b3 . 0 & ( for b5 being Element of NAT holds b4 . (b5 + 1) = (b3 . (b5 + 1)) \/ (b4 . b5) ) implies for b5 being Function of NAT ,b2 holds
( b5 . 0 = b3 . 0 & ( for b6 being Element of NAT holds b5 . (b6 + 1) = (b3 . (b6 + 1)) \ (b4 . b6) ) implies union (rng b3) = union (rng b5) ) )
proof end;

theorem Th6: :: MEASURE4:6
for b1 being set holds
bool b1 is SigmaField of b1
proof end;

theorem Th7: :: MEASURE4:7
for b1 being set
for b2, b3 being R_eal holds
ex b4 being Function of bool b1, ExtREAL st
for b5 being Subset of b1 holds
( ( b5 = {} implies b4 . b5 = b2 ) & ( b5 <> {} implies b4 . b5 = b3 ) )
proof end;

theorem Th8: :: MEASURE4:8
for b1 being set holds
ex b2 being Function of bool b1, ExtREAL st
for b3 being Subset of b1 holds b2 . b3 = 0.
proof end;

theorem Th9: :: MEASURE4:9
canceled;

theorem Th10: :: MEASURE4:10
canceled;

theorem Th11: :: MEASURE4:11
for b1 being set holds
ex b2 being Function of bool b1, ExtREAL st
( b2 is nonnegative & b2 . {} = 0. & ( for b3, b4 being Subset of b1 holds
( b3 c= b4 implies ( b2 . b3 <= b2 . b4 & ( for b5 being Function of NAT , bool b1 holds b2 . (union (rng b5)) <= SUM (b2 * b5) ) ) ) ) )
proof end;

definition
let c1 be set ;
canceled;
mode C_Measure of c1 -> Function of bool a1, ExtREAL means :Def2: :: MEASURE4:def 2
( a2 is nonnegative & a2 . {} = 0. & ( for b1, b2 being Subset of a1 holds
( b1 c= b2 implies ( a2 . b1 <= a2 . b2 & ( for b3 being Function of NAT , bool a1 holds a2 . (union (rng b3)) <= SUM (a2 * b3) ) ) ) ) );
existence
ex b1 being Function of bool c1, ExtREAL st
( b1 is nonnegative & b1 . {} = 0. & ( for b2, b3 being Subset of c1 holds
( b2 c= b3 implies ( b1 . b2 <= b1 . b3 & ( for b4 being Function of NAT , bool c1 holds b1 . (union (rng b4)) <= SUM (b1 * b4) ) ) ) ) )
by Th11;
end;

:: deftheorem Def1 MEASURE4:def 1 :
canceled;

:: deftheorem Def2 defines C_Measure MEASURE4:def 2 :
for b1 being set
for b2 being Function of bool b1, ExtREAL holds
( b2 is C_Measure of b1 iff ( b2 is nonnegative & b2 . {} = 0. & ( for b3, b4 being Subset of b1 holds
( b3 c= b4 implies ( b2 . b3 <= b2 . b4 & ( for b5 being Function of NAT , bool b1 holds b2 . (union (rng b5)) <= SUM (b2 * b5) ) ) ) ) ) );

definition
let c1 be set ;
let c2 be C_Measure of c1;
func sigma_Field c2 -> non empty Subset-Family of a1 means :Def3: :: MEASURE4:def 3
for b1 being Subset of a1 holds
( b1 in a3 iff for b2, b3 being Subset of a1 holds
( b2 c= b1 & b3 c= a1 \ b1 implies (a2 . b2) + (a2 . b3) <= a2 . (b2 \/ b3) ) );
existence
ex b1 being non empty Subset-Family of c1 st
for b2 being Subset of c1 holds
( b2 in b1 iff for b3, b4 being Subset of c1 holds
( b3 c= b2 & b4 c= c1 \ b2 implies (c2 . b3) + (c2 . b4) <= c2 . (b3 \/ b4) ) )
proof end;
uniqueness
for b1, b2 being non empty Subset-Family of c1 holds
( ( for b3 being Subset of c1 holds
( b3 in b1 iff for b4, b5 being Subset of c1 holds
( b4 c= b3 & b5 c= c1 \ b3 implies (c2 . b4) + (c2 . b5) <= c2 . (b4 \/ b5) ) ) ) & ( for b3 being Subset of c1 holds
( b3 in b2 iff for b4, b5 being Subset of c1 holds
( b4 c= b3 & b5 c= c1 \ b3 implies (c2 . b4) + (c2 . b5) <= c2 . (b4 \/ b5) ) ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def3 defines sigma_Field MEASURE4:def 3 :
for b1 being set
for b2 being C_Measure of b1
for b3 being non empty Subset-Family of b1 holds
( b3 = sigma_Field b2 iff for b4 being Subset of b1 holds
( b4 in b3 iff for b5, b6 being Subset of b1 holds
( b5 c= b4 & b6 c= b1 \ b4 implies (b2 . b5) + (b2 . b6) <= b2 . (b5 \/ b6) ) ) );

theorem Th12: :: MEASURE4:12
for b1 being set
for b2 being C_Measure of b1
for b3, b4 being Subset of b1 holds b2 . (b3 \/ b4) <= (b2 . b3) + (b2 . b4)
proof end;

theorem Th13: :: MEASURE4:13
canceled;

theorem Th14: :: MEASURE4:14
for b1 being set
for b2 being C_Measure of b1
for b3 being Subset of b1 holds
( b3 in sigma_Field b2 iff for b4, b5 being Subset of b1 holds
( b4 c= b3 & b5 c= b1 \ b3 implies (b2 . b4) + (b2 . b5) = b2 . (b4 \/ b5) ) )
proof end;

theorem Th15: :: MEASURE4:15
for b1 being set
for b2 being C_Measure of b1
for b3, b4 being Subset of b1 holds
( b3 in sigma_Field b2 & b4 in sigma_Field b2 & b4 misses b3 implies b2 . (b3 \/ b4) = (b2 . b3) + (b2 . b4) )
proof end;

theorem Th16: :: MEASURE4:16
for b1, b2 being set
for b3 being C_Measure of b2 holds
( b1 in sigma_Field b3 implies b2 \ b1 in sigma_Field b3 )
proof end;

theorem Th17: :: MEASURE4:17
for b1, b2, b3 being set
for b4 being C_Measure of b1 holds
( b2 in sigma_Field b4 & b3 in sigma_Field b4 implies b2 \/ b3 in sigma_Field b4 )
proof end;

theorem Th18: :: MEASURE4:18
for b1, b2, b3 being set
for b4 being C_Measure of b1 holds
( b2 in sigma_Field b4 & b3 in sigma_Field b4 implies b2 /\ b3 in sigma_Field b4 )
proof end;

theorem Th19: :: MEASURE4:19
for b1, b2, b3 being set
for b4 being C_Measure of b1 holds
( b2 in sigma_Field b4 & b3 in sigma_Field b4 implies b2 \ b3 in sigma_Field b4 )
proof end;

theorem Th20: :: MEASURE4:20
for b1 being set
for b2 being SigmaField of b1
for b3 being Function of NAT ,b2
for b4 being Element of b2 holds
ex b5 being Function of NAT ,b2 st
for b6 being Element of NAT holds b5 . b6 = b4 /\ (b3 . b6)
proof end;

theorem Th21: :: MEASURE4:21
for b1 being set
for b2 being C_Measure of b1 holds
sigma_Field b2 is SigmaField of b1
proof end;

registration
let c1 be set ;
let c2 be C_Measure of c1;
cluster sigma_Field a2 -> non empty compl-closed sigma-additive ;
coherence
( sigma_Field c2 is sigma-additive & sigma_Field c2 is compl-closed & not sigma_Field c2 is empty )
by Th21;
end;

definition
let c1 be set ;
let c2 be SigmaField of c1;
let c3 be N_Sub_fam of c2;
redefine func union as union c3 -> Element of a2;
coherence
union c3 is Element of c2
proof end;
end;

definition
let c1 be set ;
let c2 be C_Measure of c1;
func sigma_Meas c2 -> Function of sigma_Field a2, ExtREAL means :Def4: :: MEASURE4:def 4
for b1 being Subset of a1 holds
( b1 in sigma_Field a2 implies a3 . b1 = a2 . b1 );
existence
ex b1 being Function of sigma_Field c2, ExtREAL st
for b2 being Subset of c1 holds
( b2 in sigma_Field c2 implies b1 . b2 = c2 . b2 )
proof end;
uniqueness
for b1, b2 being Function of sigma_Field c2, ExtREAL holds
( ( for b3 being Subset of c1 holds
( b3 in sigma_Field c2 implies b1 . b3 = c2 . b3 ) ) & ( for b3 being Subset of c1 holds
( b3 in sigma_Field c2 implies b2 . b3 = c2 . b3 ) ) implies b1 = b2 )
proof end;
end;

:: deftheorem Def4 defines sigma_Meas MEASURE4:def 4 :
for b1 being set
for b2 being C_Measure of b1
for b3 being Function of sigma_Field b2, ExtREAL holds
( b3 = sigma_Meas b2 iff for b4 being Subset of b1 holds
( b4 in sigma_Field b2 implies b3 . b4 = b2 . b4 ) );

theorem Th22: :: MEASURE4:22
for b1 being set
for b2 being C_Measure of b1 holds
sigma_Meas b2 is Measure of sigma_Field b2
proof end;

definition
let c1 be set ;
let c2 be C_Measure of c1;
let c3 be Element of sigma_Field c2;
redefine func . as c2 . c3 -> R_eal;
coherence
c2 . c3 is R_eal
by FUNCT_2:7;
end;

theorem Th23: :: MEASURE4:23
for b1 being set
for b2 being C_Measure of b1 holds
sigma_Meas b2 is sigma_Measure of sigma_Field b2
proof end;

definition
let c1 be set ;
let c2 be C_Measure of c1;
redefine func sigma_Meas as sigma_Meas c2 -> sigma_Measure of sigma_Field a2;
coherence
sigma_Meas c2 is sigma_Measure of sigma_Field c2
by Th23;
end;

theorem Th24: :: MEASURE4:24
for b1 being set
for b2 being C_Measure of b1
for b3 being Subset of b1 holds
( b2 . b3 = 0. implies b3 in sigma_Field b2 )
proof end;

theorem Th25: :: MEASURE4:25
for b1 being set
for b2 being C_Measure of b1 holds sigma_Meas b2 is_complete sigma_Field b2
proof end;