diff -r e5366291d6aa -r 8d5e5ec1cac3 src/HOL/Probability/Caratheodory.thy --- a/src/HOL/Probability/Caratheodory.thy Mon May 19 12:04:45 2014 +0200 +++ b/src/HOL/Probability/Caratheodory.thy Mon May 19 13:44:13 2014 +0200 @@ -45,7 +45,7 @@ SUP_ereal_setsum[symmetric] incseq_setsumI setsum_nonneg) qed -subsection {* Measure Spaces *} +subsection {* Characterizations of Measures *} definition subadditive where "subadditive M f \ (\x\M. \y\M. x \ y = {} \ f (x \ y) \ f x + f y)" @@ -54,9 +54,6 @@ (\A. range A \ M \ disjoint_family A \ (\i. A i) \ M \ (f (\i. A i) \ (\i. f (A i))))" -definition lambda_system where "lambda_system \ M f = {l \ M. - \x \ M. f (l \ x) + f ((\ - l) \ x) = f x}" - definition outer_measure_space where "outer_measure_space M f \ positive M f \ increasing M f \ countably_subadditive M f" @@ -67,7 +64,10 @@ "subadditive M f \ x \ y = {} \ x \ M \ y \ M \ f (x \ y) \ f x + f y" by (auto simp add: subadditive_def) -subsection {* Lambda Systems *} +subsubsection {* Lambda Systems *} + +definition lambda_system where "lambda_system \ M f = {l \ M. + \x \ M. f (l \ x) + f ((\ - l) \ x) = f x}" lemma (in algebra) lambda_system_eq: shows "lambda_system \ M f = {l \ M. \x \ M. f (x \ l) + f (x - l) = f x}" @@ -643,6 +643,8 @@ by (simp add: measure_space_def positive_def countably_additive_def) blast +subsection {* Caratheodory's theorem *} + theorem (in ring_of_sets) caratheodory': assumes posf: "positive M f" and ca: "countably_additive M f" shows "\\ :: 'a set \ ereal. (\s \ M. \ s = f s) \ measure_space \ (sigma_sets \ M) \" @@ -670,8 +672,6 @@ by (intro exI[of _ ?infm]) auto qed -subsubsection {*Alternative instances of caratheodory*} - lemma (in ring_of_sets) caratheodory_empty_continuous: assumes f: "positive M f" "additive M f" and fin: "\A. A \ M \ f A \ \" assumes cont: "\A. range A \ M \ decseq A \ (\i. A i) = {} \ (\i. f (A i)) ----> 0" @@ -680,7 +680,7 @@ show "\A\M. f A \ \" using fin by auto qed (rule cont) -section {* Volumes *} +subsection {* Volumes *} definition volume :: "'a set set \ ('a set \ ereal) \ bool" where "volume M f \ @@ -818,7 +818,7 @@ qed qed -section {* Caratheodory on semirings *} +subsubsection {* Caratheodory on semirings *} theorem (in semiring_of_sets) caratheodory: assumes pos: "positive M \" and ca: "countably_additive M \"