# HG changeset patch # User hoelzl # Date 1400499853 -7200 # Node ID 8d5e5ec1cac3ab4f697e632f9df2fbacbb0e2027 # Parent e5366291d6aaa0cf08f1e2ec23f0b9e9ae32889f fixed document generation for HOL-Probability diff -r e5366291d6aa -r 8d5e5ec1cac3 src/HOL/Probability/Binary_Product_Measure.thy --- a/src/HOL/Probability/Binary_Product_Measure.thy Mon May 19 12:04:45 2014 +0200 +++ b/src/HOL/Probability/Binary_Product_Measure.thy Mon May 19 13:44:13 2014 +0200 @@ -14,7 +14,7 @@ lemma rev_Pair_vimage_times[simp]: "(\x. (x, y)) -` (A \ B) = (if y \ B then A else {})" by auto -section "Binary products" +subsection "Binary products" definition pair_measure (infixr "\\<^sub>M" 80) where "A \\<^sub>M B = measure_of (space A \ space B) @@ -480,7 +480,7 @@ done qed -section "Fubinis theorem" +subsection "Fubinis theorem" lemma measurable_compose_Pair1: "x \ space M1 \ g \ measurable (M1 \\<^sub>M M2) L \ (\y. g (x, y)) \ measurable M2 L" @@ -557,7 +557,7 @@ shows "(\\<^sup>+ y. (\\<^sup>+ x. f (x, y) \M1) \M2) = (\\<^sup>+ x. (\\<^sup>+ y. f (x, y) \M2) \M1)" unfolding positive_integral_snd[OF assms] M2.positive_integral_fst[OF assms] .. -section {* Products on counting spaces, densities and distributions *} +subsection {* Products on counting spaces, densities and distributions *} lemma sigma_sets_pair_measure_generator_finite: assumes "finite A" and "finite B" diff -r e5366291d6aa -r 8d5e5ec1cac3 src/HOL/Probability/Bochner_Integration.thy --- a/src/HOL/Probability/Bochner_Integration.thy Mon May 19 12:04:45 2014 +0200 +++ b/src/HOL/Probability/Bochner_Integration.thy Mon May 19 13:44:13 2014 +0200 @@ -1519,7 +1519,7 @@ integral\<^sup>L M f \ integral\<^sup>L M g" by (intro integral_mono_AE) auto -section {* Measure spaces with an associated density *} +subsection {* Measure spaces with an associated density *} lemma integrable_density: fixes f :: "'a \ 'b::{banach, second_countable_topology}" and g :: "'a \ real" @@ -1673,7 +1673,7 @@ has_bochner_integral M (\x. f (g x)) x \ has_bochner_integral (distr M N g) f x" by (simp add: has_bochner_integral_iff integrable_distr_eq integral_distr) -section {* Lebesgue integration on @{const count_space} *} +subsection {* Lebesgue integration on @{const count_space} *} lemma integrable_count_space: fixes f :: "'a \ 'b::{banach,second_countable_topology}" @@ -1703,7 +1703,7 @@ by (subst lebesgue_integral_count_space_finite_support) (auto intro!: setsum_mono_zero_cong_left) -section {* Point measure *} +subsection {* Point measure *} lemma lebesgue_integral_point_measure_finite: fixes g :: "'a \ 'b::{banach, second_countable_topology}" @@ -1720,7 +1720,7 @@ apply (auto simp: AE_count_space integrable_count_space) done -subsection {* Legacy lemmas for the real-valued Lebesgue integral\<^sup>L *} +subsection {* Legacy lemmas for the real-valued Lebesgue integral *} lemma real_lebesgue_integral_def: assumes f: "integrable M f" @@ -2080,7 +2080,7 @@ qed -section "Lebesgue integration on countable spaces" +subsection {* Lebesgue integration on countable spaces *} lemma integral_on_countable: fixes f :: "real \ real" diff -r e5366291d6aa -r 8d5e5ec1cac3 src/HOL/Probability/Borel_Space.thy --- a/src/HOL/Probability/Borel_Space.thy Mon May 19 12:04:45 2014 +0200 +++ b/src/HOL/Probability/Borel_Space.thy Mon May 19 13:44:13 2014 +0200 @@ -11,7 +11,7 @@ "~~/src/HOL/Multivariate_Analysis/Multivariate_Analysis" begin -section "Generic Borel spaces" +subsection {* Generic Borel spaces *} definition borel :: "'a::topological_space measure" where "borel = sigma UNIV {S. open S}" @@ -213,7 +213,7 @@ unfolding eq by (rule borel_measurable_continuous_on[OF H]) auto qed -section "Borel spaces on euclidean spaces" +subsection {* Borel spaces on euclidean spaces *} lemma borel_measurable_inner[measurable (raw)]: fixes f g :: "'a \ 'b::{second_countable_topology, real_inner}" @@ -1140,7 +1140,7 @@ shows "(\x. (\i. f i x)) \ borel_measurable M" unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp -section "LIMSEQ is borel measurable" +subsection {* LIMSEQ is borel measurable *} lemma borel_measurable_LIMSEQ: fixes u :: "nat \ 'a \ real" 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 \" diff -r e5366291d6aa -r 8d5e5ec1cac3 src/HOL/Probability/Finite_Product_Measure.thy --- a/src/HOL/Probability/Finite_Product_Measure.thy Mon May 19 12:04:45 2014 +0200 +++ b/src/HOL/Probability/Finite_Product_Measure.thy Mon May 19 13:44:13 2014 +0200 @@ -11,7 +11,7 @@ lemma split_const: "(\(i, j). c) = (\_. c)" by auto -subsubsection {* Merge two extensional functions *} +subsubsection {* More about Function restricted by @{const extensional} *} definition "merge I J = (\(x, y) i. if i \ I then x i else if i \ J then y i else undefined)" @@ -105,9 +105,9 @@ "I \ J = {} \ merge I J -` Pi\<^sub>E (I \ J) E = Pi I E \ Pi J E" by (auto simp: restrict_Pi_cancel PiE_def) -section "Finite product spaces" +subsection {* Finite product spaces *} -section "Products" +subsubsection {* Products *} definition prod_emb where "prod_emb I M K X = (\x. restrict x K) -` X \ (PIE i:I. space (M i))" diff -r e5366291d6aa -r 8d5e5ec1cac3 src/HOL/Probability/Information.thy --- a/src/HOL/Probability/Information.thy Mon May 19 12:04:45 2014 +0200 +++ b/src/HOL/Probability/Information.thy Mon May 19 13:44:13 2014 +0200 @@ -26,7 +26,7 @@ "((A, B) = X) \ (fst X = A \ snd X = B)" and "(X = (A, B)) \ (fst X = A \ snd X = B)" by auto -section "Information theory" +subsection "Information theory" locale information_space = prob_space + fixes b :: real assumes b_gt_1: "1 < b" diff -r e5366291d6aa -r 8d5e5ec1cac3 src/HOL/Probability/Measure_Space.thy --- a/src/HOL/Probability/Measure_Space.thy Mon May 19 12:04:45 2014 +0200 +++ b/src/HOL/Probability/Measure_Space.thy Mon May 19 13:44:13 2014 +0200 @@ -73,7 +73,7 @@ represent sigma algebras (with an arbitrary emeasure). *} -section "Extend binary sets" +subsection "Extend binary sets" lemma LIMSEQ_binaryset: assumes f: "f {} = 0" @@ -105,7 +105,7 @@ shows "f {} = 0 \ (\n. f (binaryset A B n)) = f A + f B" by (metis binaryset_sums sums_unique) -section {* Properties of a premeasure @{term \} *} +subsection {* Properties of a premeasure @{term \} *} text {* The definitions for @{const positive} and @{const countably_additive} should be here, by they are @@ -429,7 +429,7 @@ using empty_continuous_imp_continuous_from_below[OF f fin] cont by blast -section {* Properties of @{const emeasure} *} +subsection {* Properties of @{const emeasure} *} lemma emeasure_positive: "positive (sets M) (emeasure M)" by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def) @@ -761,7 +761,7 @@ by (simp add: emeasure_countably_additive) qed simp_all -section "@{text \}-null sets" +subsection {* @{text \}-null sets *} definition null_sets :: "'a measure \ 'a set set" where "null_sets M = {N\sets M. emeasure M N = 0}" @@ -853,7 +853,7 @@ by (subst plus_emeasure[symmetric]) auto qed -section "Formalize almost everywhere" +subsection {* The almost everywhere filter (i.e.\ quantifier) *} definition ae_filter :: "'a measure \ 'a filter" where "ae_filter M = Abs_filter (\P. \N\null_sets M. {x \ space M. \ P x} \ N)" @@ -1053,7 +1053,7 @@ shows "emeasure M A = emeasure M B" using assms by (safe intro!: antisym emeasure_mono_AE) auto -section {* @{text \}-finite Measures *} +subsection {* @{text \}-finite Measures *} locale sigma_finite_measure = fixes M :: "'a measure" @@ -1103,7 +1103,7 @@ qed (force simp: incseq_def)+ qed -section {* Measure space induced by distribution of @{const measurable}-functions *} +subsection {* Measure space induced by distribution of @{const measurable}-functions *} definition distr :: "'a measure \ 'b measure \ ('a \ 'b) \ 'b measure" where "distr M N f = measure_of (space N) (sets N) (\A. emeasure M (f -` A \ space M))" @@ -1187,7 +1187,7 @@ by (auto simp add: emeasure_distr measurable_space intro!: arg_cong[where f="emeasure M"] measure_eqI) -section {* Real measure values *} +subsection {* Real measure values *} lemma measure_nonneg: "0 \ measure M A" using emeasure_nonneg[of M A] unfolding measure_def by (auto intro: real_of_ereal_pos) @@ -1324,7 +1324,7 @@ by (intro lim_real_of_ereal) simp qed -section {* Measure spaces with @{term "emeasure M (space M) < \"} *} +subsection {* Measure spaces with @{term "emeasure M (space M) < \"} *} locale finite_measure = sigma_finite_measure M for M + assumes finite_emeasure_space: "emeasure M (space M) \ \" @@ -1527,7 +1527,7 @@ shows "measure M B = 0" using measure_space_inter[of B A] assms by (auto simp: ac_simps) -section {* Counting space *} +subsection {* Counting space *} lemma strict_monoI_Suc: assumes ord [simp]: "(\n. f n < f (Suc n))" shows "strict_mono f" @@ -1653,7 +1653,7 @@ show "sigma_finite_measure (count_space A)" .. qed -section {* Measure restricted to space *} +subsection {* Measure restricted to space *} lemma emeasure_restrict_space: assumes "\ \ sets M" "A \ \" diff -r e5366291d6aa -r 8d5e5ec1cac3 src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy --- a/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy Mon May 19 12:04:45 2014 +0200 +++ b/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy Mon May 19 13:44:13 2014 +0200 @@ -13,7 +13,7 @@ "indicator A x \ (indicator B x::ereal) \ (x \ A \ x \ B)" by (simp add: indicator_def not_le) -section "Simple function" +subsection "Simple function" text {* @@ -501,7 +501,7 @@ finally show "(\x. f (T x)) -` {i} \ space M \ sets M" . qed -section "Simple integral" +subsection "Simple integral" definition simple_integral :: "'a measure \ ('a \ ereal) \ ereal" ("integral\<^sup>S") where "integral\<^sup>S M f = (\x \ f ` space M. x * emeasure M (f -` {x} \ space M))" @@ -734,7 +734,7 @@ then show ?thesis by simp qed -section "Continuous positive integration" +subsection {* Integral on nonnegative functions *} definition positive_integral :: "'a measure \ ('a \ ereal) \ ereal" ("integral\<^sup>P") where "integral\<^sup>P M f = (SUP g : {g. simple_function M g \ g \ max 0 \ f}. integral\<^sup>S M g)" @@ -1556,7 +1556,9 @@ by (simp add: F_def) qed -section {* Distributions *} +subsection {* Integral under concrete measures *} + +subsubsection {* Distributions *} lemma positive_integral_distr': assumes T: "T \ measurable M M'" @@ -1587,7 +1589,7 @@ by (subst (1 2) positive_integral_max_0[symmetric]) (simp add: positive_integral_distr') -section {* Lebesgue integration on @{const count_space} *} +subsubsection {* Counting space *} lemma simple_function_count_space[simp]: "simple_function (count_space A) f \ finite (f ` A)" @@ -1666,7 +1668,7 @@ finally show ?thesis . qed -section {* Measures with Restricted Space *} +subsubsection {* Measures with Restricted Space *} lemma positive_integral_restrict_space: assumes \: "\ \ sets M" and f: "f \ borel_measurable M" "\x. 0 \ f x" "\x. x \ space M - \ \ f x = 0" @@ -1696,7 +1698,7 @@ by (auto simp add: SUP_eq_iff measurable_restrict_space1 \ positive_integral_monotone_convergence_SUP) qed -section {* Measure spaces with an associated density *} +subsubsection {* Measure spaces with an associated density *} definition density :: "'a measure \ ('a \ ereal) \ 'a measure" where "density M f = measure_of (space M) (sets M) (\A. \\<^sup>+ x. f x * indicator A x \M)" @@ -1918,7 +1920,7 @@ using f g by (subst density_density_eq) auto qed -section {* Point measure *} +subsubsection {* Point measure *} definition point_measure :: "'a set \ ('a \ ereal) \ 'a measure" where "point_measure A f = density (count_space A) f" @@ -1991,7 +1993,7 @@ integral\<^sup>P (point_measure A f) g = (\a\A. f a * g a)" by (subst positive_integral_point_measure) (auto intro!: setsum_mono_zero_left simp: less_le) -section {* Uniform measure *} +subsubsection {* Uniform measure *} definition "uniform_measure M A = density M (\x. indicator A x / emeasure M A)" @@ -2019,7 +2021,7 @@ using emeasure_uniform_measure[OF emeasure_neq_0_sets[OF A(1)] B] A by (cases "emeasure M A" "emeasure M (A \ B)" rule: ereal2_cases) (simp_all add: measure_def) -section {* Uniform count measure *} +subsubsection {* Uniform count measure *} definition "uniform_count_measure A = point_measure A (\x. 1 / card A)" diff -r e5366291d6aa -r 8d5e5ec1cac3 src/HOL/Probability/Probability_Measure.thy --- a/src/HOL/Probability/Probability_Measure.thy Mon May 19 12:04:45 2014 +0200 +++ b/src/HOL/Probability/Probability_Measure.thy Mon May 19 13:44:13 2014 +0200 @@ -349,7 +349,7 @@ finally show ?thesis by simp qed -section {* Distributions *} +subsection {* Distributions *} definition "distributed M N X f \ distr M N X = density N f \ f \ borel_measurable N \ (AE x in N. 0 \ f x) \ X \ measurable M N" diff -r e5366291d6aa -r 8d5e5ec1cac3 src/HOL/Probability/Radon_Nikodym.thy --- a/src/HOL/Probability/Radon_Nikodym.thy Mon May 19 12:04:45 2014 +0200 +++ b/src/HOL/Probability/Radon_Nikodym.thy Mon May 19 13:44:13 2014 +0200 @@ -769,7 +769,7 @@ by (auto intro!: bexI[of _ "\x. h x * f x"] simp: density_density_eq) qed -section "Uniqueness of densities" +subsection {* Uniqueness of densities *} lemma finite_density_unique: assumes borel: "f \ borel_measurable M" "g \ borel_measurable M" @@ -1060,7 +1060,7 @@ apply (auto simp: max_def intro!: measurable_If) done -section "Radon-Nikodym derivative" +subsection {* Radon-Nikodym derivative *} definition RN_deriv :: "'a measure \ 'a measure \ 'a \ ereal" where "RN_deriv M N = diff -r e5366291d6aa -r 8d5e5ec1cac3 src/HOL/Probability/Sigma_Algebra.thy --- a/src/HOL/Probability/Sigma_Algebra.thy Mon May 19 12:04:45 2014 +0200 +++ b/src/HOL/Probability/Sigma_Algebra.thy Mon May 19 13:44:13 2014 +0200 @@ -5,7 +5,7 @@ translated by Lawrence Paulson. *) -header {* Sigma Algebras *} +header {* Describing measurable sets *} theory Sigma_Algebra imports @@ -33,9 +33,7 @@ lemma (in subset_class) sets_into_space: "x \ M \ x \ \" by (metis PowD contra_subsetD space_closed) -subsection {* Semiring of sets *} - -subsubsection {* Disjoint sets *} +subsubsection {* Semiring of sets *} definition "disjoint A \ (\a\A. \b\A. a \ b \ a \ b = {})" @@ -255,7 +253,7 @@ "X \ S \ algebra S { {}, X, S - X, S }" by (auto simp: algebra_iff_Int) -subsection {* Restricted algebras *} +subsubsection {* Restricted algebras *} abbreviation (in algebra) "restricted_space A \ (op \ A) ` M" @@ -264,7 +262,7 @@ assumes "A \ M" shows "algebra A (restricted_space A)" using assms by (auto simp: algebra_iff_Int) -subsection {* Sigma Algebras *} +subsubsection {* Sigma Algebras *} locale sigma_algebra = algebra + assumes countable_nat_UN [intro]: "\A. range A \ M \ (\i::nat. A i) \ M" @@ -446,7 +444,7 @@ shows "sigma_algebra S { {}, X, S - X, S }" using algebra.is_sigma_algebra[OF algebra_single_set[OF `X \ S`]] by simp -subsection {* Binary Unions *} +subsubsection {* Binary Unions *} definition binary :: "'a \ 'a \ nat \ 'a" where "binary a b = (\x. b)(0 := a)" @@ -468,7 +466,7 @@ by (auto simp add: range_binary_eq sigma_algebra_def sigma_algebra_axioms_def algebra_iff_Un Un_range_binary) -subsection {* Initial Sigma Algebra *} +subsubsection {* Initial Sigma Algebra *} text {*Sigma algebras can naturally be created as the closure of any set of M with regard to the properties just postulated. *} @@ -775,7 +773,7 @@ qed qed -subsection "Disjoint families" +subsubsection "Disjoint families" definition disjoint_family_on where @@ -934,7 +932,7 @@ by (intro disjointD[OF d]) auto qed -subsection {* Ring generated by a semiring *} +subsubsection {* Ring generated by a semiring *} definition (in semiring_of_sets) "generated_ring = { \C | C. C \ M \ finite C \ disjoint C }" @@ -1064,759 +1062,7 @@ by (blast intro!: sigma_sets_mono elim: generated_ringE) qed (auto intro!: generated_ringI_Basic sigma_sets_mono) -subsection {* Measure type *} - -definition positive :: "'a set set \ ('a set \ ereal) \ bool" where - "positive M \ \ \ {} = 0 \ (\A\M. 0 \ \ A)" - -definition countably_additive :: "'a set set \ ('a set \ ereal) \ bool" where - "countably_additive M f \ (\A. range A \ M \ disjoint_family A \ (\i. A i) \ M \ - (\i. f (A i)) = f (\i. A i))" - -definition measure_space :: "'a set \ 'a set set \ ('a set \ ereal) \ bool" where - "measure_space \ A \ \ sigma_algebra \ A \ positive A \ \ countably_additive A \" - -typedef 'a measure = "{(\::'a set, A, \). (\a\-A. \ a = 0) \ measure_space \ A \ }" -proof - have "sigma_algebra UNIV {{}, UNIV}" - by (auto simp: sigma_algebra_iff2) - then show "(UNIV, {{}, UNIV}, \A. 0) \ {(\, A, \). (\a\-A. \ a = 0) \ measure_space \ A \} " - by (auto simp: measure_space_def positive_def countably_additive_def) -qed - -definition space :: "'a measure \ 'a set" where - "space M = fst (Rep_measure M)" - -definition sets :: "'a measure \ 'a set set" where - "sets M = fst (snd (Rep_measure M))" - -definition emeasure :: "'a measure \ 'a set \ ereal" where - "emeasure M = snd (snd (Rep_measure M))" - -definition measure :: "'a measure \ 'a set \ real" where - "measure M A = real (emeasure M A)" - -declare [[coercion sets]] - -declare [[coercion measure]] - -declare [[coercion emeasure]] - -lemma measure_space: "measure_space (space M) (sets M) (emeasure M)" - by (cases M) (auto simp: space_def sets_def emeasure_def Abs_measure_inverse) - -interpretation sets!: sigma_algebra "space M" "sets M" for M :: "'a measure" - using measure_space[of M] by (auto simp: measure_space_def) - -definition measure_of :: "'a set \ 'a set set \ ('a set \ ereal) \ 'a measure" where - "measure_of \ A \ = Abs_measure (\, if A \ Pow \ then sigma_sets \ A else {{}, \}, - \a. if a \ sigma_sets \ A \ measure_space \ (sigma_sets \ A) \ then \ a else 0)" - -abbreviation "sigma \ A \ measure_of \ A (\x. 0)" - -lemma measure_space_0: "A \ Pow \ \ measure_space \ (sigma_sets \ A) (\x. 0)" - unfolding measure_space_def - by (auto intro!: sigma_algebra_sigma_sets simp: positive_def countably_additive_def) - -lemma sigma_algebra_trivial: "sigma_algebra \ {{}, \}" -by unfold_locales(fastforce intro: exI[where x="{{}}"] exI[where x="{\}"])+ - -lemma measure_space_0': "measure_space \ {{}, \} (\x. 0)" -by(simp add: measure_space_def positive_def countably_additive_def sigma_algebra_trivial) - -lemma measure_space_closed: - assumes "measure_space \ M \" - shows "M \ Pow \" -proof - - interpret sigma_algebra \ M using assms by(simp add: measure_space_def) - show ?thesis by(rule space_closed) -qed - -lemma (in ring_of_sets) positive_cong_eq: - "(\a. a \ M \ \' a = \ a) \ positive M \' = positive M \" - by (auto simp add: positive_def) - -lemma (in sigma_algebra) countably_additive_eq: - "(\a. a \ M \ \' a = \ a) \ countably_additive M \' = countably_additive M \" - unfolding countably_additive_def - by (intro arg_cong[where f=All] ext) (auto simp add: countably_additive_def subset_eq) - -lemma measure_space_eq: - assumes closed: "A \ Pow \" and eq: "\a. a \ sigma_sets \ A \ \ a = \' a" - shows "measure_space \ (sigma_sets \ A) \ = measure_space \ (sigma_sets \ A) \'" -proof - - interpret sigma_algebra \ "sigma_sets \ A" using closed by (rule sigma_algebra_sigma_sets) - from positive_cong_eq[OF eq, of "\i. i"] countably_additive_eq[OF eq, of "\i. i"] show ?thesis - by (auto simp: measure_space_def) -qed - -lemma measure_of_eq: - assumes closed: "A \ Pow \" and eq: "(\a. a \ sigma_sets \ A \ \ a = \' a)" - shows "measure_of \ A \ = measure_of \ A \'" -proof - - have "measure_space \ (sigma_sets \ A) \ = measure_space \ (sigma_sets \ A) \'" - using assms by (rule measure_space_eq) - with eq show ?thesis - by (auto simp add: measure_of_def intro!: arg_cong[where f=Abs_measure]) -qed - -lemma - shows space_measure_of_conv: "space (measure_of \ A \) = \" (is ?space) - and sets_measure_of_conv: - "sets (measure_of \ A \) = (if A \ Pow \ then sigma_sets \ A else {{}, \})" (is ?sets) - and emeasure_measure_of_conv: - "emeasure (measure_of \ A \) = - (\B. if B \ sigma_sets \ A \ measure_space \ (sigma_sets \ A) \ then \ B else 0)" (is ?emeasure) -proof - - have "?space \ ?sets \ ?emeasure" - proof(cases "measure_space \ (sigma_sets \ A) \") - case True - from measure_space_closed[OF this] sigma_sets_superset_generator[of A \] - have "A \ Pow \" by simp - hence "measure_space \ (sigma_sets \ A) \ = measure_space \ (sigma_sets \ A) - (\a. if a \ sigma_sets \ A then \ a else 0)" - by(rule measure_space_eq) auto - with True `A \ Pow \` show ?thesis - by(simp add: measure_of_def space_def sets_def emeasure_def Abs_measure_inverse) - next - case False thus ?thesis - by(cases "A \ Pow \")(simp_all add: Abs_measure_inverse measure_of_def sets_def space_def emeasure_def measure_space_0 measure_space_0') - qed - thus ?space ?sets ?emeasure by simp_all -qed - -lemma [simp]: - assumes A: "A \ Pow \" - shows sets_measure_of: "sets (measure_of \ A \) = sigma_sets \ A" - and space_measure_of: "space (measure_of \ A \) = \" -using assms -by(simp_all add: sets_measure_of_conv space_measure_of_conv) - -lemma (in sigma_algebra) sets_measure_of_eq[simp]: "sets (measure_of \ M \) = M" - using space_closed by (auto intro!: sigma_sets_eq) - -lemma (in sigma_algebra) space_measure_of_eq[simp]: "space (measure_of \ M \) = \" - by (rule space_measure_of_conv) - -lemma measure_of_subset: "M \ Pow \ \ M' \ M \ sets (measure_of \ M' \) \ sets (measure_of \ M \')" - by (auto intro!: sigma_sets_subseteq) - -lemma sigma_sets_mono'': - assumes "A \ sigma_sets C D" - assumes "B \ D" - assumes "D \ Pow C" - shows "sigma_sets A B \ sigma_sets C D" -proof - fix x assume "x \ sigma_sets A B" - thus "x \ sigma_sets C D" - proof induct - case (Basic a) with assms have "a \ D" by auto - thus ?case .. - next - case Empty show ?case by (rule sigma_sets.Empty) - next - from assms have "A \ sets (sigma C D)" by (subst sets_measure_of[OF `D \ Pow C`]) - moreover case (Compl a) hence "a \ sets (sigma C D)" by (subst sets_measure_of[OF `D \ Pow C`]) - ultimately have "A - a \ sets (sigma C D)" .. - thus ?case by (subst (asm) sets_measure_of[OF `D \ Pow C`]) - next - case (Union a) - thus ?case by (intro sigma_sets.Union) - qed -qed - -lemma in_measure_of[intro, simp]: "M \ Pow \ \ A \ M \ A \ sets (measure_of \ M \)" - by auto - -subsection {* Constructing simple @{typ "'a measure"} *} - -lemma emeasure_measure_of: - assumes M: "M = measure_of \ A \" - assumes ms: "A \ Pow \" "positive (sets M) \" "countably_additive (sets M) \" - assumes X: "X \ sets M" - shows "emeasure M X = \ X" -proof - - interpret sigma_algebra \ "sigma_sets \ A" by (rule sigma_algebra_sigma_sets) fact - have "measure_space \ (sigma_sets \ A) \" - using ms M by (simp add: measure_space_def sigma_algebra_sigma_sets) - thus ?thesis using X ms - by(simp add: M emeasure_measure_of_conv sets_measure_of_conv) -qed - -lemma emeasure_measure_of_sigma: - assumes ms: "sigma_algebra \ M" "positive M \" "countably_additive M \" - assumes A: "A \ M" - shows "emeasure (measure_of \ M \) A = \ A" -proof - - interpret sigma_algebra \ M by fact - have "measure_space \ (sigma_sets \ M) \" - using ms sigma_sets_eq by (simp add: measure_space_def) - thus ?thesis by(simp add: emeasure_measure_of_conv A) -qed - -lemma measure_cases[cases type: measure]: - obtains (measure) \ A \ where "x = Abs_measure (\, A, \)" "\a\-A. \ a = 0" "measure_space \ A \" - by atomize_elim (cases x, auto) - -lemma sets_eq_imp_space_eq: - "sets M = sets M' \ space M = space M'" - using sets.top[of M] sets.top[of M'] sets.space_closed[of M] sets.space_closed[of M'] - by blast - -lemma emeasure_notin_sets: "A \ sets M \ emeasure M A = 0" - by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def) - -lemma emeasure_neq_0_sets: "emeasure M A \ 0 \ A \ sets M" - using emeasure_notin_sets[of A M] by blast - -lemma measure_notin_sets: "A \ sets M \ measure M A = 0" - by (simp add: measure_def emeasure_notin_sets) - -lemma measure_eqI: - fixes M N :: "'a measure" - assumes "sets M = sets N" and eq: "\A. A \ sets M \ emeasure M A = emeasure N A" - shows "M = N" -proof (cases M N rule: measure_cases[case_product measure_cases]) - case (measure_measure \ A \ \' A' \') - interpret M: sigma_algebra \ A using measure_measure by (auto simp: measure_space_def) - interpret N: sigma_algebra \' A' using measure_measure by (auto simp: measure_space_def) - have "A = sets M" "A' = sets N" - using measure_measure by (simp_all add: sets_def Abs_measure_inverse) - with `sets M = sets N` have AA': "A = A'" by simp - moreover from M.top N.top M.space_closed N.space_closed AA' have "\ = \'" by auto - moreover { fix B have "\ B = \' B" - proof cases - assume "B \ A" - with eq `A = sets M` have "emeasure M B = emeasure N B" by simp - with measure_measure show "\ B = \' B" - by (simp add: emeasure_def Abs_measure_inverse) - next - assume "B \ A" - with `A = sets M` `A' = sets N` `A = A'` have "B \ sets M" "B \ sets N" - by auto - then have "emeasure M B = 0" "emeasure N B = 0" - by (simp_all add: emeasure_notin_sets) - with measure_measure show "\ B = \' B" - by (simp add: emeasure_def Abs_measure_inverse) - qed } - then have "\ = \'" by auto - ultimately show "M = N" - by (simp add: measure_measure) -qed - -lemma emeasure_sigma: "A \ Pow \ \ emeasure (sigma \ A) = (\_. 0)" - using measure_space_0[of A \] - by (simp add: measure_of_def emeasure_def Abs_measure_inverse) - -lemma sigma_eqI: - assumes [simp]: "M \ Pow \" "N \ Pow \" "sigma_sets \ M = sigma_sets \ N" - shows "sigma \ M = sigma \ N" - by (rule measure_eqI) (simp_all add: emeasure_sigma) - -subsection {* Measurable functions *} - -definition measurable :: "'a measure \ 'b measure \ ('a \ 'b) set" where - "measurable A B = {f \ space A -> space B. \y \ sets B. f -` y \ space A \ sets A}" - -lemma measurable_space: - "f \ measurable M A \ x \ space M \ f x \ space A" - unfolding measurable_def by auto - -lemma measurable_sets: - "f \ measurable M A \ S \ sets A \ f -` S \ space M \ sets M" - unfolding measurable_def by auto - -lemma measurable_sets_Collect: - assumes f: "f \ measurable M N" and P: "{x\space N. P x} \ sets N" shows "{x\space M. P (f x)} \ sets M" -proof - - have "f -` {x \ space N. P x} \ space M = {x\space M. P (f x)}" - using measurable_space[OF f] by auto - with measurable_sets[OF f P] show ?thesis - by simp -qed - -lemma measurable_sigma_sets: - assumes B: "sets N = sigma_sets \ A" "A \ Pow \" - and f: "f \ space M \ \" - and ba: "\y. y \ A \ (f -` y) \ space M \ sets M" - shows "f \ measurable M N" -proof - - interpret A: sigma_algebra \ "sigma_sets \ A" using B(2) by (rule sigma_algebra_sigma_sets) - from B sets.top[of N] A.top sets.space_closed[of N] A.space_closed have \: "\ = space N" by force - - { fix X assume "X \ sigma_sets \ A" - then have "f -` X \ space M \ sets M \ X \ \" - proof induct - case (Basic a) then show ?case - by (auto simp add: ba) (metis B(2) subsetD PowD) - next - case (Compl a) - have [simp]: "f -` \ \ space M = space M" - by (auto simp add: funcset_mem [OF f]) - then show ?case - by (auto simp add: vimage_Diff Diff_Int_distrib2 sets.compl_sets Compl) - next - case (Union a) - then show ?case - by (simp add: vimage_UN, simp only: UN_extend_simps(4)) blast - qed auto } - with f show ?thesis - by (auto simp add: measurable_def B \) -qed - -lemma measurable_measure_of: - assumes B: "N \ Pow \" - and f: "f \ space M \ \" - and ba: "\y. y \ N \ (f -` y) \ space M \ sets M" - shows "f \ measurable M (measure_of \ N \)" -proof - - have "sets (measure_of \ N \) = sigma_sets \ N" - using B by (rule sets_measure_of) - from this assms show ?thesis by (rule measurable_sigma_sets) -qed - -lemma measurable_iff_measure_of: - assumes "N \ Pow \" "f \ space M \ \" - shows "f \ measurable M (measure_of \ N \) \ (\A\N. f -` A \ space M \ sets M)" - by (metis assms in_measure_of measurable_measure_of assms measurable_sets) - -lemma measurable_cong_sets: - assumes sets: "sets M = sets M'" "sets N = sets N'" - shows "measurable M N = measurable M' N'" - using sets[THEN sets_eq_imp_space_eq] sets by (simp add: measurable_def) - -lemma measurable_cong: - assumes "\ w. w \ space M \ f w = g w" - shows "f \ measurable M M' \ g \ measurable M M'" - unfolding measurable_def using assms - by (simp cong: vimage_inter_cong Pi_cong) - -lemma measurable_cong_strong: - "M = N \ M' = N' \ (\w. w \ space M \ f w = g w) \ - f \ measurable M M' \ g \ measurable N N'" - by (metis measurable_cong) - -lemma measurable_eqI: - "\ space m1 = space m1' ; space m2 = space m2' ; - sets m1 = sets m1' ; sets m2 = sets m2' \ - \ measurable m1 m2 = measurable m1' m2'" - by (simp add: measurable_def sigma_algebra_iff2) - -lemma measurable_compose: - assumes f: "f \ measurable M N" and g: "g \ measurable N L" - shows "(\x. g (f x)) \ measurable M L" -proof - - have "\A. (\x. g (f x)) -` A \ space M = f -` (g -` A \ space N) \ space M" - using measurable_space[OF f] by auto - with measurable_space[OF f] measurable_space[OF g] show ?thesis - by (auto intro: measurable_sets[OF f] measurable_sets[OF g] - simp del: vimage_Int simp add: measurable_def) -qed - -lemma measurable_comp: - "f \ measurable M N \ g \ measurable N L \ g \ f \ measurable M L" - using measurable_compose[of f M N g L] by (simp add: comp_def) - -lemma measurable_const: - "c \ space M' \ (\x. c) \ measurable M M'" - by (auto simp add: measurable_def) - -lemma measurable_If: - assumes measure: "f \ measurable M M'" "g \ measurable M M'" - assumes P: "{x\space M. P x} \ sets M" - shows "(\x. if P x then f x else g x) \ measurable M M'" - unfolding measurable_def -proof safe - fix x assume "x \ space M" - thus "(if P x then f x else g x) \ space M'" - using measure unfolding measurable_def by auto -next - fix A assume "A \ sets M'" - hence *: "(\x. if P x then f x else g x) -` A \ space M = - ((f -` A \ space M) \ {x\space M. P x}) \ - ((g -` A \ space M) \ (space M - {x\space M. P x}))" - using measure unfolding measurable_def by (auto split: split_if_asm) - show "(\x. if P x then f x else g x) -` A \ space M \ sets M" - using `A \ sets M'` measure P unfolding * measurable_def - by (auto intro!: sets.Un) -qed - -lemma measurable_If_set: - assumes measure: "f \ measurable M M'" "g \ measurable M M'" - assumes P: "A \ space M \ sets M" - shows "(\x. if x \ A then f x else g x) \ measurable M M'" -proof (rule measurable_If[OF measure]) - have "{x \ space M. x \ A} = A \ space M" by auto - thus "{x \ space M. x \ A} \ sets M" using `A \ space M \ sets M` by auto -qed - -lemma measurable_ident: "id \ measurable M M" - by (auto simp add: measurable_def) - -lemma measurable_ident_sets: - assumes eq: "sets M = sets M'" shows "(\x. x) \ measurable M M'" - using measurable_ident[of M] - unfolding id_def measurable_def eq sets_eq_imp_space_eq[OF eq] . - -lemma sets_Least: - assumes meas: "\i::nat. {x\space M. P i x} \ M" - shows "(\x. LEAST j. P j x) -` A \ space M \ sets M" -proof - - { fix i have "(\x. LEAST j. P j x) -` {i} \ space M \ sets M" - proof cases - assume i: "(LEAST j. False) = i" - have "(\x. LEAST j. P j x) -` {i} \ space M = - {x\space M. P i x} \ (space M - (\jspace M. P j x})) \ (space M - (\i. {x\space M. P i x}))" - by (simp add: set_eq_iff, safe) - (insert i, auto dest: Least_le intro: LeastI intro!: Least_equality) - with meas show ?thesis - by (auto intro!: sets.Int) - next - assume i: "(LEAST j. False) \ i" - then have "(\x. LEAST j. P j x) -` {i} \ space M = - {x\space M. P i x} \ (space M - (\jspace M. P j x}))" - proof (simp add: set_eq_iff, safe) - fix x assume neq: "(LEAST j. False) \ (LEAST j. P j x)" - have "\j. P j x" - by (rule ccontr) (insert neq, auto) - then show "P (LEAST j. P j x) x" by (rule LeastI_ex) - qed (auto dest: Least_le intro!: Least_equality) - with meas show ?thesis - by auto - qed } - then have "(\i\A. (\x. LEAST j. P j x) -` {i} \ space M) \ sets M" - by (intro sets.countable_UN) auto - moreover have "(\i\A. (\x. LEAST j. P j x) -` {i} \ space M) = - (\x. LEAST j. P j x) -` A \ space M" by auto - ultimately show ?thesis by auto -qed - -lemma measurable_strong: - fixes f :: "'a \ 'b" and g :: "'b \ 'c" - assumes f: "f \ measurable a b" and g: "g \ space b \ space c" - and t: "f ` (space a) \ t" - and cb: "\s. s \ sets c \ (g -` s) \ t \ sets b" - shows "(g o f) \ measurable a c" -proof - - have fab: "f \ (space a -> space b)" - and ba: "\y. y \ sets b \ (f -` y) \ (space a) \ sets a" using f - by (auto simp add: measurable_def) - have eq: "\y. (g \ f) -` y \ space a = f -` (g -` y \ t) \ space a" using t - by force - show ?thesis - apply (auto simp add: measurable_def vimage_comp) - apply (metis funcset_mem fab g) - apply (subst eq) - apply (metis ba cb) - done -qed - -lemma measurable_mono1: - "M' \ Pow \ \ M \ M' \ - measurable (measure_of \ M \) N \ measurable (measure_of \ M' \') N" - using measure_of_subset[of M' \ M] by (auto simp add: measurable_def) - -subsection {* Counting space *} - -definition count_space :: "'a set \ 'a measure" where - "count_space \ = measure_of \ (Pow \) (\A. if finite A then ereal (card A) else \)" - -lemma - shows space_count_space[simp]: "space (count_space \) = \" - and sets_count_space[simp]: "sets (count_space \) = Pow \" - using sigma_sets_into_sp[of "Pow \" \] - by (auto simp: count_space_def) - -lemma measurable_count_space_eq1[simp]: - "f \ measurable (count_space A) M \ f \ A \ space M" - unfolding measurable_def by simp - -lemma measurable_count_space_eq2: - assumes "finite A" - shows "f \ measurable M (count_space A) \ (f \ space M \ A \ (\a\A. f -` {a} \ space M \ sets M))" -proof - - { fix X assume "X \ A" "f \ space M \ A" - with `finite A` have "f -` X \ space M = (\a\X. f -` {a} \ space M)" "finite X" - by (auto dest: finite_subset) - moreover assume "\a\A. f -` {a} \ space M \ sets M" - ultimately have "f -` X \ space M \ sets M" - using `X \ A` by (auto intro!: sets.finite_UN simp del: UN_simps) } - then show ?thesis - unfolding measurable_def by auto -qed - -lemma measurable_count_space_eq2_countable: - fixes f :: "'a => 'c::countable" - shows "f \ measurable M (count_space A) \ (f \ space M \ A \ (\a\A. f -` {a} \ space M \ sets M))" -proof - - { fix X assume "X \ A" "f \ space M \ A" - assume *: "\a. a\A \ f -` {a} \ space M \ sets M" - have "f -` X \ space M = (\a\X. f -` {a} \ space M)" - by auto - also have "\ \ sets M" - using * `X \ A` by (intro sets.countable_UN) auto - finally have "f -` X \ space M \ sets M" . } - then show ?thesis - unfolding measurable_def by auto -qed - -lemma measurable_compose_countable: - assumes f: "\i::'i::countable. (\x. f i x) \ measurable M N" and g: "g \ measurable M (count_space UNIV)" - shows "(\x. f (g x) x) \ measurable M N" - unfolding measurable_def -proof safe - fix x assume "x \ space M" then show "f (g x) x \ space N" - using f[THEN measurable_space] g[THEN measurable_space] by auto -next - fix A assume A: "A \ sets N" - have "(\x. f (g x) x) -` A \ space M = (\i. (g -` {i} \ space M) \ (f i -` A \ space M))" - by auto - also have "\ \ sets M" using f[THEN measurable_sets, OF A] g[THEN measurable_sets] - by (auto intro!: sets.countable_UN measurable_sets) - finally show "(\x. f (g x) x) -` A \ space M \ sets M" . -qed - -lemma measurable_count_space_const: - "(\x. c) \ measurable M (count_space UNIV)" - by (simp add: measurable_const) - -lemma measurable_count_space: - "f \ measurable (count_space A) (count_space UNIV)" - by simp - -lemma measurable_compose_rev: - assumes f: "f \ measurable L N" and g: "g \ measurable M L" - shows "(\x. f (g x)) \ measurable M N" - using measurable_compose[OF g f] . - -lemma measurable_count_space_eq_countable: - assumes "countable A" - shows "f \ measurable M (count_space A) \ (f \ space M \ A \ (\a\A. f -` {a} \ space M \ sets M))" -proof - - { fix X assume "X \ A" "f \ space M \ A" - with `countable A` have "f -` X \ space M = (\a\X. f -` {a} \ space M)" "countable X" - by (auto dest: countable_subset) - moreover assume "\a\A. f -` {a} \ space M \ sets M" - ultimately have "f -` X \ space M \ sets M" - using `X \ A` by (auto intro!: sets.countable_UN' simp del: UN_simps) } - then show ?thesis - unfolding measurable_def by auto -qed - -subsection {* Extend measure *} - -definition "extend_measure \ I G \ = - (if (\\'. (\i\I. \' (G i) = \ i) \ measure_space \ (sigma_sets \ (G`I)) \') \ \ (\i\I. \ i = 0) - then measure_of \ (G`I) (SOME \'. (\i\I. \' (G i) = \ i) \ measure_space \ (sigma_sets \ (G`I)) \') - else measure_of \ (G`I) (\_. 0))" - -lemma space_extend_measure: "G ` I \ Pow \ \ space (extend_measure \ I G \) = \" - unfolding extend_measure_def by simp - -lemma sets_extend_measure: "G ` I \ Pow \ \ sets (extend_measure \ I G \) = sigma_sets \ (G`I)" - unfolding extend_measure_def by simp - -lemma emeasure_extend_measure: - assumes M: "M = extend_measure \ I G \" - and eq: "\i. i \ I \ \' (G i) = \ i" - and ms: "G ` I \ Pow \" "positive (sets M) \'" "countably_additive (sets M) \'" - and "i \ I" - shows "emeasure M (G i) = \ i" -proof cases - assume *: "(\i\I. \ i = 0)" - with M have M_eq: "M = measure_of \ (G`I) (\_. 0)" - by (simp add: extend_measure_def) - from measure_space_0[OF ms(1)] ms `i\I` - have "emeasure M (G i) = 0" - by (intro emeasure_measure_of[OF M_eq]) (auto simp add: M measure_space_def sets_extend_measure) - with `i\I` * show ?thesis - by simp -next - def P \ "\\'. (\i\I. \' (G i) = \ i) \ measure_space \ (sigma_sets \ (G`I)) \'" - assume "\ (\i\I. \ i = 0)" - moreover - have "measure_space (space M) (sets M) \'" - using ms unfolding measure_space_def by auto default - with ms eq have "\\'. P \'" - unfolding P_def - by (intro exI[of _ \']) (auto simp add: M space_extend_measure sets_extend_measure) - ultimately have M_eq: "M = measure_of \ (G`I) (Eps P)" - by (simp add: M extend_measure_def P_def[symmetric]) - - from `\\'. P \'` have P: "P (Eps P)" by (rule someI_ex) - show "emeasure M (G i) = \ i" - proof (subst emeasure_measure_of[OF M_eq]) - have sets_M: "sets M = sigma_sets \ (G`I)" - using M_eq ms by (auto simp: sets_extend_measure) - then show "G i \ sets M" using `i \ I` by auto - show "positive (sets M) (Eps P)" "countably_additive (sets M) (Eps P)" "Eps P (G i) = \ i" - using P `i\I` by (auto simp add: sets_M measure_space_def P_def) - qed fact -qed - -lemma emeasure_extend_measure_Pair: - assumes M: "M = extend_measure \ {(i, j). I i j} (\(i, j). G i j) (\(i, j). \ i j)" - and eq: "\i j. I i j \ \' (G i j) = \ i j" - and ms: "\i j. I i j \ G i j \ Pow \" "positive (sets M) \'" "countably_additive (sets M) \'" - and "I i j" - shows "emeasure M (G i j) = \ i j" - using emeasure_extend_measure[OF M _ _ ms(2,3), of "(i,j)"] eq ms(1) `I i j` - by (auto simp: subset_eq) - -subsection {* Sigma algebra generated by function preimages *} - -definition - "vimage_algebra M S X = sigma S ((\A. X -` A \ S) ` sets M)" - -lemma sigma_algebra_preimages: - fixes f :: "'x \ 'a" - assumes "f \ S \ space M" - shows "sigma_algebra S ((\A. f -` A \ S) ` sets M)" - (is "sigma_algebra _ (?F ` sets M)") -proof (simp add: sigma_algebra_iff2, safe) - show "{} \ ?F ` sets M" by blast -next - fix A assume "A \ sets M" - moreover have "S - ?F A = ?F (space M - A)" - using assms by auto - ultimately show "S - ?F A \ ?F ` sets M" - by blast -next - fix A :: "nat \ 'x set" assume *: "range A \ ?F ` M" - have "\i. \b. b \ M \ A i = ?F b" - proof safe - fix i - have "A i \ ?F ` M" using * by auto - then show "\b. b \ M \ A i = ?F b" by auto - qed - from choice[OF this] obtain b where b: "range b \ M" "\i. A i = ?F (b i)" - by auto - then have "(\i. A i) = ?F (\i. b i)" by auto - then show "(\i. A i) \ ?F ` M" using b(1) by blast -qed - -lemma sets_vimage_algebra[simp]: - "f \ S \ space M \ sets (vimage_algebra M S f) = (\A. f -` A \ S) ` sets M" - using sigma_algebra.sets_measure_of_eq[OF sigma_algebra_preimages, of f S M] - by (simp add: vimage_algebra_def) - -lemma space_vimage_algebra[simp]: - "f \ S \ space M \ space (vimage_algebra M S f) = S" - using sigma_algebra.space_measure_of_eq[OF sigma_algebra_preimages, of f S M] - by (simp add: vimage_algebra_def) - -lemma in_vimage_algebra[simp]: - "f \ S \ space M \ A \ sets (vimage_algebra M S f) \ (\B\sets M. A = f -` B \ S)" - by (simp add: image_iff) - -lemma measurable_vimage_algebra: - fixes S :: "'c set" assumes "f \ S \ space M" - shows "f \ measurable (vimage_algebra M S f) M" - unfolding measurable_def using assms by force - -lemma measurable_vimage: - fixes g :: "'a \ 'c" and f :: "'d \ 'a" - assumes "g \ measurable M M2" "f \ S \ space M" - shows "(\x. g (f x)) \ measurable (vimage_algebra M S f) M2" -proof - - note measurable_vimage_algebra[OF assms(2)] - from measurable_comp[OF this assms(1)] - show ?thesis by (simp add: comp_def) -qed - -lemma sigma_sets_vimage: - assumes "f \ S' \ S" and "A \ Pow S" - shows "sigma_sets S' ((\X. f -` X \ S') ` A) = (\X. f -` X \ S') ` sigma_sets S A" -proof (intro set_eqI iffI) - let ?F = "\X. f -` X \ S'" - fix X assume "X \ sigma_sets S' (?F ` A)" - then show "X \ ?F ` sigma_sets S A" - proof induct - case (Basic X) then obtain X' where "X = ?F X'" "X' \ A" - by auto - then show ?case by auto - next - case Empty then show ?case - by (auto intro!: image_eqI[of _ _ "{}"] sigma_sets.Empty) - next - case (Compl X) then obtain X' where X: "X = ?F X'" and "X' \ sigma_sets S A" - by auto - then have "S - X' \ sigma_sets S A" - by (auto intro!: sigma_sets.Compl) - then show ?case - using X assms by (auto intro!: image_eqI[where x="S - X'"]) - next - case (Union F) - then have "\i. \F'. F' \ sigma_sets S A \ F i = f -` F' \ S'" - by (auto simp: image_iff Bex_def) - from choice[OF this] obtain F' where - "\i. F' i \ sigma_sets S A" and "\i. F i = f -` F' i \ S'" - by auto - then show ?case - by (auto intro!: sigma_sets.Union image_eqI[where x="\i. F' i"]) - qed -next - let ?F = "\X. f -` X \ S'" - fix X assume "X \ ?F ` sigma_sets S A" - then obtain X' where "X' \ sigma_sets S A" "X = ?F X'" by auto - then show "X \ sigma_sets S' (?F ` A)" - proof (induct arbitrary: X) - case Empty then show ?case by (auto intro: sigma_sets.Empty) - next - case (Compl X') - have "S' - (S' - X) \ sigma_sets S' (?F ` A)" - apply (rule sigma_sets.Compl) - using assms by (auto intro!: Compl.hyps simp: Compl.prems) - also have "S' - (S' - X) = X" - using assms Compl by auto - finally show ?case . - next - case (Union F) - have "(\i. f -` F i \ S') \ sigma_sets S' (?F ` A)" - by (intro sigma_sets.Union Union.hyps) simp - also have "(\i. f -` F i \ S') = X" - using assms Union by auto - finally show ?case . - qed auto -qed - -subsection {* Restricted Space Sigma Algebra *} - -definition "restrict_space M \ = measure_of \ ((op \ \) ` sets M) (\A. emeasure M (A \ \))" - -lemma space_restrict_space: "space (restrict_space M \) = \" - unfolding restrict_space_def by (subst space_measure_of) auto - -lemma sets_restrict_space: "\ \ space M \ sets (restrict_space M \) = (op \ \) ` sets M" - using sigma_sets_vimage[of "\x. x" \ "space M" "sets M"] - unfolding restrict_space_def - by (subst sets_measure_of) (auto simp: sets.sigma_sets_eq Int_commute[of _ \] sets.space_closed) - -lemma sets_restrict_space_iff: - "\ \ sets M \ A \ sets (restrict_space M \) \ (A \ \ \ A \ sets M)" - by (subst sets_restrict_space) (auto dest: sets.sets_into_space) - -lemma measurable_restrict_space1: - assumes \: "\ \ sets M" and f: "f \ measurable M N" shows "f \ measurable (restrict_space M \) N" - unfolding measurable_def -proof (intro CollectI conjI ballI) - show sp: "f \ space (restrict_space M \) \ space N" - using measurable_space[OF f] sets.sets_into_space[OF \] by (auto simp: space_restrict_space) - - fix A assume "A \ sets N" - have "f -` A \ space (restrict_space M \) = (f -` A \ space M) \ \" - using sets.sets_into_space[OF \] by (auto simp: space_restrict_space) - also have "\ \ sets (restrict_space M \)" - unfolding sets_restrict_space_iff[OF \] - using measurable_sets[OF f `A \ sets N`] \ by blast - finally show "f -` A \ space (restrict_space M \) \ sets (restrict_space M \)" . -qed - -lemma measurable_restrict_space2: - "\ \ sets N \ f \ space M \ \ \ f \ measurable M N \ f \ measurable M (restrict_space N \)" - by (simp add: measurable_def space_restrict_space sets_restrict_space_iff) - -subsection {* A Two-Element Series *} +subsubsection {* A Two-Element Series *} definition binaryset :: "'a set \ 'a set \ nat \ 'a set " where "binaryset A B = (\x. {})(0 := A, Suc 0 := B)" @@ -1830,7 +1076,7 @@ lemma UN_binaryset_eq: "(\i. binaryset A B i) = A \ B" by (simp add: SUP_def range_binaryset_eq) -section {* Closed CDI *} +subsubsection {* Closed CDI *} definition closed_cdi where "closed_cdi \ M \ @@ -2064,7 +1310,7 @@ by blast qed -subsection {* Dynkin systems *} +subsubsection {* Dynkin systems *} locale dynkin_system = subset_class + assumes space: "\ \ M" @@ -2126,7 +1372,7 @@ show ?thesis using sets_into_space by (fastforce intro!: dynkin_systemI) qed -subsection "Intersection stable algebras" +subsubsection "Intersection sets systems" definition "Int_stable M \ (\ a \ M. \ b \ M. a \ b \ M)" @@ -2162,7 +1408,7 @@ qed auto qed -subsection "Smallest Dynkin systems" +subsubsection "Smallest Dynkin systems" definition dynkin where "dynkin \ M = (\{D. dynkin_system \ D \ M \ D})" @@ -2309,6 +1555,11 @@ using assms by (auto simp: dynkin_def) qed +subsubsection {* Induction rule for intersection-stable generators *} + +text {* The reason to introduce Dynkin-systems is the following induction rules for $\sigma$-algebras +generated by a generator closed under intersection. *} + lemma sigma_sets_induct_disjoint[consumes 3, case_names basic empty compl union]: assumes "Int_stable G" and closed: "G \ Pow \" @@ -2330,4 +1581,756 @@ with A show ?thesis by auto qed +subsection {* Measure type *} + +definition positive :: "'a set set \ ('a set \ ereal) \ bool" where + "positive M \ \ \ {} = 0 \ (\A\M. 0 \ \ A)" + +definition countably_additive :: "'a set set \ ('a set \ ereal) \ bool" where + "countably_additive M f \ (\A. range A \ M \ disjoint_family A \ (\i. A i) \ M \ + (\i. f (A i)) = f (\i. A i))" + +definition measure_space :: "'a set \ 'a set set \ ('a set \ ereal) \ bool" where + "measure_space \ A \ \ sigma_algebra \ A \ positive A \ \ countably_additive A \" + +typedef 'a measure = "{(\::'a set, A, \). (\a\-A. \ a = 0) \ measure_space \ A \ }" +proof + have "sigma_algebra UNIV {{}, UNIV}" + by (auto simp: sigma_algebra_iff2) + then show "(UNIV, {{}, UNIV}, \A. 0) \ {(\, A, \). (\a\-A. \ a = 0) \ measure_space \ A \} " + by (auto simp: measure_space_def positive_def countably_additive_def) +qed + +definition space :: "'a measure \ 'a set" where + "space M = fst (Rep_measure M)" + +definition sets :: "'a measure \ 'a set set" where + "sets M = fst (snd (Rep_measure M))" + +definition emeasure :: "'a measure \ 'a set \ ereal" where + "emeasure M = snd (snd (Rep_measure M))" + +definition measure :: "'a measure \ 'a set \ real" where + "measure M A = real (emeasure M A)" + +declare [[coercion sets]] + +declare [[coercion measure]] + +declare [[coercion emeasure]] + +lemma measure_space: "measure_space (space M) (sets M) (emeasure M)" + by (cases M) (auto simp: space_def sets_def emeasure_def Abs_measure_inverse) + +interpretation sets!: sigma_algebra "space M" "sets M" for M :: "'a measure" + using measure_space[of M] by (auto simp: measure_space_def) + +definition measure_of :: "'a set \ 'a set set \ ('a set \ ereal) \ 'a measure" where + "measure_of \ A \ = Abs_measure (\, if A \ Pow \ then sigma_sets \ A else {{}, \}, + \a. if a \ sigma_sets \ A \ measure_space \ (sigma_sets \ A) \ then \ a else 0)" + +abbreviation "sigma \ A \ measure_of \ A (\x. 0)" + +lemma measure_space_0: "A \ Pow \ \ measure_space \ (sigma_sets \ A) (\x. 0)" + unfolding measure_space_def + by (auto intro!: sigma_algebra_sigma_sets simp: positive_def countably_additive_def) + +lemma sigma_algebra_trivial: "sigma_algebra \ {{}, \}" +by unfold_locales(fastforce intro: exI[where x="{{}}"] exI[where x="{\}"])+ + +lemma measure_space_0': "measure_space \ {{}, \} (\x. 0)" +by(simp add: measure_space_def positive_def countably_additive_def sigma_algebra_trivial) + +lemma measure_space_closed: + assumes "measure_space \ M \" + shows "M \ Pow \" +proof - + interpret sigma_algebra \ M using assms by(simp add: measure_space_def) + show ?thesis by(rule space_closed) +qed + +lemma (in ring_of_sets) positive_cong_eq: + "(\a. a \ M \ \' a = \ a) \ positive M \' = positive M \" + by (auto simp add: positive_def) + +lemma (in sigma_algebra) countably_additive_eq: + "(\a. a \ M \ \' a = \ a) \ countably_additive M \' = countably_additive M \" + unfolding countably_additive_def + by (intro arg_cong[where f=All] ext) (auto simp add: countably_additive_def subset_eq) + +lemma measure_space_eq: + assumes closed: "A \ Pow \" and eq: "\a. a \ sigma_sets \ A \ \ a = \' a" + shows "measure_space \ (sigma_sets \ A) \ = measure_space \ (sigma_sets \ A) \'" +proof - + interpret sigma_algebra \ "sigma_sets \ A" using closed by (rule sigma_algebra_sigma_sets) + from positive_cong_eq[OF eq, of "\i. i"] countably_additive_eq[OF eq, of "\i. i"] show ?thesis + by (auto simp: measure_space_def) +qed + +lemma measure_of_eq: + assumes closed: "A \ Pow \" and eq: "(\a. a \ sigma_sets \ A \ \ a = \' a)" + shows "measure_of \ A \ = measure_of \ A \'" +proof - + have "measure_space \ (sigma_sets \ A) \ = measure_space \ (sigma_sets \ A) \'" + using assms by (rule measure_space_eq) + with eq show ?thesis + by (auto simp add: measure_of_def intro!: arg_cong[where f=Abs_measure]) +qed + +lemma + shows space_measure_of_conv: "space (measure_of \ A \) = \" (is ?space) + and sets_measure_of_conv: + "sets (measure_of \ A \) = (if A \ Pow \ then sigma_sets \ A else {{}, \})" (is ?sets) + and emeasure_measure_of_conv: + "emeasure (measure_of \ A \) = + (\B. if B \ sigma_sets \ A \ measure_space \ (sigma_sets \ A) \ then \ B else 0)" (is ?emeasure) +proof - + have "?space \ ?sets \ ?emeasure" + proof(cases "measure_space \ (sigma_sets \ A) \") + case True + from measure_space_closed[OF this] sigma_sets_superset_generator[of A \] + have "A \ Pow \" by simp + hence "measure_space \ (sigma_sets \ A) \ = measure_space \ (sigma_sets \ A) + (\a. if a \ sigma_sets \ A then \ a else 0)" + by(rule measure_space_eq) auto + with True `A \ Pow \` show ?thesis + by(simp add: measure_of_def space_def sets_def emeasure_def Abs_measure_inverse) + next + case False thus ?thesis + by(cases "A \ Pow \")(simp_all add: Abs_measure_inverse measure_of_def sets_def space_def emeasure_def measure_space_0 measure_space_0') + qed + thus ?space ?sets ?emeasure by simp_all +qed + +lemma [simp]: + assumes A: "A \ Pow \" + shows sets_measure_of: "sets (measure_of \ A \) = sigma_sets \ A" + and space_measure_of: "space (measure_of \ A \) = \" +using assms +by(simp_all add: sets_measure_of_conv space_measure_of_conv) + +lemma (in sigma_algebra) sets_measure_of_eq[simp]: "sets (measure_of \ M \) = M" + using space_closed by (auto intro!: sigma_sets_eq) + +lemma (in sigma_algebra) space_measure_of_eq[simp]: "space (measure_of \ M \) = \" + by (rule space_measure_of_conv) + +lemma measure_of_subset: "M \ Pow \ \ M' \ M \ sets (measure_of \ M' \) \ sets (measure_of \ M \')" + by (auto intro!: sigma_sets_subseteq) + +lemma sigma_sets_mono'': + assumes "A \ sigma_sets C D" + assumes "B \ D" + assumes "D \ Pow C" + shows "sigma_sets A B \ sigma_sets C D" +proof + fix x assume "x \ sigma_sets A B" + thus "x \ sigma_sets C D" + proof induct + case (Basic a) with assms have "a \ D" by auto + thus ?case .. + next + case Empty show ?case by (rule sigma_sets.Empty) + next + from assms have "A \ sets (sigma C D)" by (subst sets_measure_of[OF `D \ Pow C`]) + moreover case (Compl a) hence "a \ sets (sigma C D)" by (subst sets_measure_of[OF `D \ Pow C`]) + ultimately have "A - a \ sets (sigma C D)" .. + thus ?case by (subst (asm) sets_measure_of[OF `D \ Pow C`]) + next + case (Union a) + thus ?case by (intro sigma_sets.Union) + qed +qed + +lemma in_measure_of[intro, simp]: "M \ Pow \ \ A \ M \ A \ sets (measure_of \ M \)" + by auto + +subsubsection {* Constructing simple @{typ "'a measure"} *} + +lemma emeasure_measure_of: + assumes M: "M = measure_of \ A \" + assumes ms: "A \ Pow \" "positive (sets M) \" "countably_additive (sets M) \" + assumes X: "X \ sets M" + shows "emeasure M X = \ X" +proof - + interpret sigma_algebra \ "sigma_sets \ A" by (rule sigma_algebra_sigma_sets) fact + have "measure_space \ (sigma_sets \ A) \" + using ms M by (simp add: measure_space_def sigma_algebra_sigma_sets) + thus ?thesis using X ms + by(simp add: M emeasure_measure_of_conv sets_measure_of_conv) +qed + +lemma emeasure_measure_of_sigma: + assumes ms: "sigma_algebra \ M" "positive M \" "countably_additive M \" + assumes A: "A \ M" + shows "emeasure (measure_of \ M \) A = \ A" +proof - + interpret sigma_algebra \ M by fact + have "measure_space \ (sigma_sets \ M) \" + using ms sigma_sets_eq by (simp add: measure_space_def) + thus ?thesis by(simp add: emeasure_measure_of_conv A) +qed + +lemma measure_cases[cases type: measure]: + obtains (measure) \ A \ where "x = Abs_measure (\, A, \)" "\a\-A. \ a = 0" "measure_space \ A \" + by atomize_elim (cases x, auto) + +lemma sets_eq_imp_space_eq: + "sets M = sets M' \ space M = space M'" + using sets.top[of M] sets.top[of M'] sets.space_closed[of M] sets.space_closed[of M'] + by blast + +lemma emeasure_notin_sets: "A \ sets M \ emeasure M A = 0" + by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def) + +lemma emeasure_neq_0_sets: "emeasure M A \ 0 \ A \ sets M" + using emeasure_notin_sets[of A M] by blast + +lemma measure_notin_sets: "A \ sets M \ measure M A = 0" + by (simp add: measure_def emeasure_notin_sets) + +lemma measure_eqI: + fixes M N :: "'a measure" + assumes "sets M = sets N" and eq: "\A. A \ sets M \ emeasure M A = emeasure N A" + shows "M = N" +proof (cases M N rule: measure_cases[case_product measure_cases]) + case (measure_measure \ A \ \' A' \') + interpret M: sigma_algebra \ A using measure_measure by (auto simp: measure_space_def) + interpret N: sigma_algebra \' A' using measure_measure by (auto simp: measure_space_def) + have "A = sets M" "A' = sets N" + using measure_measure by (simp_all add: sets_def Abs_measure_inverse) + with `sets M = sets N` have AA': "A = A'" by simp + moreover from M.top N.top M.space_closed N.space_closed AA' have "\ = \'" by auto + moreover { fix B have "\ B = \' B" + proof cases + assume "B \ A" + with eq `A = sets M` have "emeasure M B = emeasure N B" by simp + with measure_measure show "\ B = \' B" + by (simp add: emeasure_def Abs_measure_inverse) + next + assume "B \ A" + with `A = sets M` `A' = sets N` `A = A'` have "B \ sets M" "B \ sets N" + by auto + then have "emeasure M B = 0" "emeasure N B = 0" + by (simp_all add: emeasure_notin_sets) + with measure_measure show "\ B = \' B" + by (simp add: emeasure_def Abs_measure_inverse) + qed } + then have "\ = \'" by auto + ultimately show "M = N" + by (simp add: measure_measure) +qed + +lemma emeasure_sigma: "A \ Pow \ \ emeasure (sigma \ A) = (\_. 0)" + using measure_space_0[of A \] + by (simp add: measure_of_def emeasure_def Abs_measure_inverse) + +lemma sigma_eqI: + assumes [simp]: "M \ Pow \" "N \ Pow \" "sigma_sets \ M = sigma_sets \ N" + shows "sigma \ M = sigma \ N" + by (rule measure_eqI) (simp_all add: emeasure_sigma) + +subsubsection {* Measurable functions *} + +definition measurable :: "'a measure \ 'b measure \ ('a \ 'b) set" where + "measurable A B = {f \ space A -> space B. \y \ sets B. f -` y \ space A \ sets A}" + +lemma measurable_space: + "f \ measurable M A \ x \ space M \ f x \ space A" + unfolding measurable_def by auto + +lemma measurable_sets: + "f \ measurable M A \ S \ sets A \ f -` S \ space M \ sets M" + unfolding measurable_def by auto + +lemma measurable_sets_Collect: + assumes f: "f \ measurable M N" and P: "{x\space N. P x} \ sets N" shows "{x\space M. P (f x)} \ sets M" +proof - + have "f -` {x \ space N. P x} \ space M = {x\space M. P (f x)}" + using measurable_space[OF f] by auto + with measurable_sets[OF f P] show ?thesis + by simp +qed + +lemma measurable_sigma_sets: + assumes B: "sets N = sigma_sets \ A" "A \ Pow \" + and f: "f \ space M \ \" + and ba: "\y. y \ A \ (f -` y) \ space M \ sets M" + shows "f \ measurable M N" +proof - + interpret A: sigma_algebra \ "sigma_sets \ A" using B(2) by (rule sigma_algebra_sigma_sets) + from B sets.top[of N] A.top sets.space_closed[of N] A.space_closed have \: "\ = space N" by force + + { fix X assume "X \ sigma_sets \ A" + then have "f -` X \ space M \ sets M \ X \ \" + proof induct + case (Basic a) then show ?case + by (auto simp add: ba) (metis B(2) subsetD PowD) + next + case (Compl a) + have [simp]: "f -` \ \ space M = space M" + by (auto simp add: funcset_mem [OF f]) + then show ?case + by (auto simp add: vimage_Diff Diff_Int_distrib2 sets.compl_sets Compl) + next + case (Union a) + then show ?case + by (simp add: vimage_UN, simp only: UN_extend_simps(4)) blast + qed auto } + with f show ?thesis + by (auto simp add: measurable_def B \) +qed + +lemma measurable_measure_of: + assumes B: "N \ Pow \" + and f: "f \ space M \ \" + and ba: "\y. y \ N \ (f -` y) \ space M \ sets M" + shows "f \ measurable M (measure_of \ N \)" +proof - + have "sets (measure_of \ N \) = sigma_sets \ N" + using B by (rule sets_measure_of) + from this assms show ?thesis by (rule measurable_sigma_sets) +qed + +lemma measurable_iff_measure_of: + assumes "N \ Pow \" "f \ space M \ \" + shows "f \ measurable M (measure_of \ N \) \ (\A\N. f -` A \ space M \ sets M)" + by (metis assms in_measure_of measurable_measure_of assms measurable_sets) + +lemma measurable_cong_sets: + assumes sets: "sets M = sets M'" "sets N = sets N'" + shows "measurable M N = measurable M' N'" + using sets[THEN sets_eq_imp_space_eq] sets by (simp add: measurable_def) + +lemma measurable_cong: + assumes "\ w. w \ space M \ f w = g w" + shows "f \ measurable M M' \ g \ measurable M M'" + unfolding measurable_def using assms + by (simp cong: vimage_inter_cong Pi_cong) + +lemma measurable_cong_strong: + "M = N \ M' = N' \ (\w. w \ space M \ f w = g w) \ + f \ measurable M M' \ g \ measurable N N'" + by (metis measurable_cong) + +lemma measurable_eqI: + "\ space m1 = space m1' ; space m2 = space m2' ; + sets m1 = sets m1' ; sets m2 = sets m2' \ + \ measurable m1 m2 = measurable m1' m2'" + by (simp add: measurable_def sigma_algebra_iff2) + +lemma measurable_compose: + assumes f: "f \ measurable M N" and g: "g \ measurable N L" + shows "(\x. g (f x)) \ measurable M L" +proof - + have "\A. (\x. g (f x)) -` A \ space M = f -` (g -` A \ space N) \ space M" + using measurable_space[OF f] by auto + with measurable_space[OF f] measurable_space[OF g] show ?thesis + by (auto intro: measurable_sets[OF f] measurable_sets[OF g] + simp del: vimage_Int simp add: measurable_def) +qed + +lemma measurable_comp: + "f \ measurable M N \ g \ measurable N L \ g \ f \ measurable M L" + using measurable_compose[of f M N g L] by (simp add: comp_def) + +lemma measurable_const: + "c \ space M' \ (\x. c) \ measurable M M'" + by (auto simp add: measurable_def) + +lemma measurable_If: + assumes measure: "f \ measurable M M'" "g \ measurable M M'" + assumes P: "{x\space M. P x} \ sets M" + shows "(\x. if P x then f x else g x) \ measurable M M'" + unfolding measurable_def +proof safe + fix x assume "x \ space M" + thus "(if P x then f x else g x) \ space M'" + using measure unfolding measurable_def by auto +next + fix A assume "A \ sets M'" + hence *: "(\x. if P x then f x else g x) -` A \ space M = + ((f -` A \ space M) \ {x\space M. P x}) \ + ((g -` A \ space M) \ (space M - {x\space M. P x}))" + using measure unfolding measurable_def by (auto split: split_if_asm) + show "(\x. if P x then f x else g x) -` A \ space M \ sets M" + using `A \ sets M'` measure P unfolding * measurable_def + by (auto intro!: sets.Un) +qed + +lemma measurable_If_set: + assumes measure: "f \ measurable M M'" "g \ measurable M M'" + assumes P: "A \ space M \ sets M" + shows "(\x. if x \ A then f x else g x) \ measurable M M'" +proof (rule measurable_If[OF measure]) + have "{x \ space M. x \ A} = A \ space M" by auto + thus "{x \ space M. x \ A} \ sets M" using `A \ space M \ sets M` by auto +qed + +lemma measurable_ident: "id \ measurable M M" + by (auto simp add: measurable_def) + +lemma measurable_ident_sets: + assumes eq: "sets M = sets M'" shows "(\x. x) \ measurable M M'" + using measurable_ident[of M] + unfolding id_def measurable_def eq sets_eq_imp_space_eq[OF eq] . + +lemma sets_Least: + assumes meas: "\i::nat. {x\space M. P i x} \ M" + shows "(\x. LEAST j. P j x) -` A \ space M \ sets M" +proof - + { fix i have "(\x. LEAST j. P j x) -` {i} \ space M \ sets M" + proof cases + assume i: "(LEAST j. False) = i" + have "(\x. LEAST j. P j x) -` {i} \ space M = + {x\space M. P i x} \ (space M - (\jspace M. P j x})) \ (space M - (\i. {x\space M. P i x}))" + by (simp add: set_eq_iff, safe) + (insert i, auto dest: Least_le intro: LeastI intro!: Least_equality) + with meas show ?thesis + by (auto intro!: sets.Int) + next + assume i: "(LEAST j. False) \ i" + then have "(\x. LEAST j. P j x) -` {i} \ space M = + {x\space M. P i x} \ (space M - (\jspace M. P j x}))" + proof (simp add: set_eq_iff, safe) + fix x assume neq: "(LEAST j. False) \ (LEAST j. P j x)" + have "\j. P j x" + by (rule ccontr) (insert neq, auto) + then show "P (LEAST j. P j x) x" by (rule LeastI_ex) + qed (auto dest: Least_le intro!: Least_equality) + with meas show ?thesis + by auto + qed } + then have "(\i\A. (\x. LEAST j. P j x) -` {i} \ space M) \ sets M" + by (intro sets.countable_UN) auto + moreover have "(\i\A. (\x. LEAST j. P j x) -` {i} \ space M) = + (\x. LEAST j. P j x) -` A \ space M" by auto + ultimately show ?thesis by auto +qed + +lemma measurable_strong: + fixes f :: "'a \ 'b" and g :: "'b \ 'c" + assumes f: "f \ measurable a b" and g: "g \ space b \ space c" + and t: "f ` (space a) \ t" + and cb: "\s. s \ sets c \ (g -` s) \ t \ sets b" + shows "(g o f) \ measurable a c" +proof - + have fab: "f \ (space a -> space b)" + and ba: "\y. y \ sets b \ (f -` y) \ (space a) \ sets a" using f + by (auto simp add: measurable_def) + have eq: "\y. (g \ f) -` y \ space a = f -` (g -` y \ t) \ space a" using t + by force + show ?thesis + apply (auto simp add: measurable_def vimage_comp) + apply (metis funcset_mem fab g) + apply (subst eq) + apply (metis ba cb) + done +qed + +lemma measurable_mono1: + "M' \ Pow \ \ M \ M' \ + measurable (measure_of \ M \) N \ measurable (measure_of \ M' \') N" + using measure_of_subset[of M' \ M] by (auto simp add: measurable_def) + +subsubsection {* Counting space *} + +definition count_space :: "'a set \ 'a measure" where + "count_space \ = measure_of \ (Pow \) (\A. if finite A then ereal (card A) else \)" + +lemma + shows space_count_space[simp]: "space (count_space \) = \" + and sets_count_space[simp]: "sets (count_space \) = Pow \" + using sigma_sets_into_sp[of "Pow \" \] + by (auto simp: count_space_def) + +lemma measurable_count_space_eq1[simp]: + "f \ measurable (count_space A) M \ f \ A \ space M" + unfolding measurable_def by simp + +lemma measurable_count_space_eq2: + assumes "finite A" + shows "f \ measurable M (count_space A) \ (f \ space M \ A \ (\a\A. f -` {a} \ space M \ sets M))" +proof - + { fix X assume "X \ A" "f \ space M \ A" + with `finite A` have "f -` X \ space M = (\a\X. f -` {a} \ space M)" "finite X" + by (auto dest: finite_subset) + moreover assume "\a\A. f -` {a} \ space M \ sets M" + ultimately have "f -` X \ space M \ sets M" + using `X \ A` by (auto intro!: sets.finite_UN simp del: UN_simps) } + then show ?thesis + unfolding measurable_def by auto +qed + +lemma measurable_count_space_eq2_countable: + fixes f :: "'a => 'c::countable" + shows "f \ measurable M (count_space A) \ (f \ space M \ A \ (\a\A. f -` {a} \ space M \ sets M))" +proof - + { fix X assume "X \ A" "f \ space M \ A" + assume *: "\a. a\A \ f -` {a} \ space M \ sets M" + have "f -` X \ space M = (\a\X. f -` {a} \ space M)" + by auto + also have "\ \ sets M" + using * `X \ A` by (intro sets.countable_UN) auto + finally have "f -` X \ space M \ sets M" . } + then show ?thesis + unfolding measurable_def by auto +qed + +lemma measurable_compose_countable: + assumes f: "\i::'i::countable. (\x. f i x) \ measurable M N" and g: "g \ measurable M (count_space UNIV)" + shows "(\x. f (g x) x) \ measurable M N" + unfolding measurable_def +proof safe + fix x assume "x \ space M" then show "f (g x) x \ space N" + using f[THEN measurable_space] g[THEN measurable_space] by auto +next + fix A assume A: "A \ sets N" + have "(\x. f (g x) x) -` A \ space M = (\i. (g -` {i} \ space M) \ (f i -` A \ space M))" + by auto + also have "\ \ sets M" using f[THEN measurable_sets, OF A] g[THEN measurable_sets] + by (auto intro!: sets.countable_UN measurable_sets) + finally show "(\x. f (g x) x) -` A \ space M \ sets M" . +qed + +lemma measurable_count_space_const: + "(\x. c) \ measurable M (count_space UNIV)" + by (simp add: measurable_const) + +lemma measurable_count_space: + "f \ measurable (count_space A) (count_space UNIV)" + by simp + +lemma measurable_compose_rev: + assumes f: "f \ measurable L N" and g: "g \ measurable M L" + shows "(\x. f (g x)) \ measurable M N" + using measurable_compose[OF g f] . + +lemma measurable_count_space_eq_countable: + assumes "countable A" + shows "f \ measurable M (count_space A) \ (f \ space M \ A \ (\a\A. f -` {a} \ space M \ sets M))" +proof - + { fix X assume "X \ A" "f \ space M \ A" + with `countable A` have "f -` X \ space M = (\a\X. f -` {a} \ space M)" "countable X" + by (auto dest: countable_subset) + moreover assume "\a\A. f -` {a} \ space M \ sets M" + ultimately have "f -` X \ space M \ sets M" + using `X \ A` by (auto intro!: sets.countable_UN' simp del: UN_simps) } + then show ?thesis + unfolding measurable_def by auto +qed + +subsubsection {* Extend measure *} + +definition "extend_measure \ I G \ = + (if (\\'. (\i\I. \' (G i) = \ i) \ measure_space \ (sigma_sets \ (G`I)) \') \ \ (\i\I. \ i = 0) + then measure_of \ (G`I) (SOME \'. (\i\I. \' (G i) = \ i) \ measure_space \ (sigma_sets \ (G`I)) \') + else measure_of \ (G`I) (\_. 0))" + +lemma space_extend_measure: "G ` I \ Pow \ \ space (extend_measure \ I G \) = \" + unfolding extend_measure_def by simp + +lemma sets_extend_measure: "G ` I \ Pow \ \ sets (extend_measure \ I G \) = sigma_sets \ (G`I)" + unfolding extend_measure_def by simp + +lemma emeasure_extend_measure: + assumes M: "M = extend_measure \ I G \" + and eq: "\i. i \ I \ \' (G i) = \ i" + and ms: "G ` I \ Pow \" "positive (sets M) \'" "countably_additive (sets M) \'" + and "i \ I" + shows "emeasure M (G i) = \ i" +proof cases + assume *: "(\i\I. \ i = 0)" + with M have M_eq: "M = measure_of \ (G`I) (\_. 0)" + by (simp add: extend_measure_def) + from measure_space_0[OF ms(1)] ms `i\I` + have "emeasure M (G i) = 0" + by (intro emeasure_measure_of[OF M_eq]) (auto simp add: M measure_space_def sets_extend_measure) + with `i\I` * show ?thesis + by simp +next + def P \ "\\'. (\i\I. \' (G i) = \ i) \ measure_space \ (sigma_sets \ (G`I)) \'" + assume "\ (\i\I. \ i = 0)" + moreover + have "measure_space (space M) (sets M) \'" + using ms unfolding measure_space_def by auto default + with ms eq have "\\'. P \'" + unfolding P_def + by (intro exI[of _ \']) (auto simp add: M space_extend_measure sets_extend_measure) + ultimately have M_eq: "M = measure_of \ (G`I) (Eps P)" + by (simp add: M extend_measure_def P_def[symmetric]) + + from `\\'. P \'` have P: "P (Eps P)" by (rule someI_ex) + show "emeasure M (G i) = \ i" + proof (subst emeasure_measure_of[OF M_eq]) + have sets_M: "sets M = sigma_sets \ (G`I)" + using M_eq ms by (auto simp: sets_extend_measure) + then show "G i \ sets M" using `i \ I` by auto + show "positive (sets M) (Eps P)" "countably_additive (sets M) (Eps P)" "Eps P (G i) = \ i" + using P `i\I` by (auto simp add: sets_M measure_space_def P_def) + qed fact +qed + +lemma emeasure_extend_measure_Pair: + assumes M: "M = extend_measure \ {(i, j). I i j} (\(i, j). G i j) (\(i, j). \ i j)" + and eq: "\i j. I i j \ \' (G i j) = \ i j" + and ms: "\i j. I i j \ G i j \ Pow \" "positive (sets M) \'" "countably_additive (sets M) \'" + and "I i j" + shows "emeasure M (G i j) = \ i j" + using emeasure_extend_measure[OF M _ _ ms(2,3), of "(i,j)"] eq ms(1) `I i j` + by (auto simp: subset_eq) + +subsubsection {* Sigma algebra generated by function preimages *} + +definition + "vimage_algebra M S X = sigma S ((\A. X -` A \ S) ` sets M)" + +lemma sigma_algebra_preimages: + fixes f :: "'x \ 'a" + assumes "f \ S \ space M" + shows "sigma_algebra S ((\A. f -` A \ S) ` sets M)" + (is "sigma_algebra _ (?F ` sets M)") +proof (simp add: sigma_algebra_iff2, safe) + show "{} \ ?F ` sets M" by blast +next + fix A assume "A \ sets M" + moreover have "S - ?F A = ?F (space M - A)" + using assms by auto + ultimately show "S - ?F A \ ?F ` sets M" + by blast +next + fix A :: "nat \ 'x set" assume *: "range A \ ?F ` M" + have "\i. \b. b \ M \ A i = ?F b" + proof safe + fix i + have "A i \ ?F ` M" using * by auto + then show "\b. b \ M \ A i = ?F b" by auto + qed + from choice[OF this] obtain b where b: "range b \ M" "\i. A i = ?F (b i)" + by auto + then have "(\i. A i) = ?F (\i. b i)" by auto + then show "(\i. A i) \ ?F ` M" using b(1) by blast +qed + +lemma sets_vimage_algebra[simp]: + "f \ S \ space M \ sets (vimage_algebra M S f) = (\A. f -` A \ S) ` sets M" + using sigma_algebra.sets_measure_of_eq[OF sigma_algebra_preimages, of f S M] + by (simp add: vimage_algebra_def) + +lemma space_vimage_algebra[simp]: + "f \ S \ space M \ space (vimage_algebra M S f) = S" + using sigma_algebra.space_measure_of_eq[OF sigma_algebra_preimages, of f S M] + by (simp add: vimage_algebra_def) + +lemma in_vimage_algebra[simp]: + "f \ S \ space M \ A \ sets (vimage_algebra M S f) \ (\B\sets M. A = f -` B \ S)" + by (simp add: image_iff) + +lemma measurable_vimage_algebra: + fixes S :: "'c set" assumes "f \ S \ space M" + shows "f \ measurable (vimage_algebra M S f) M" + unfolding measurable_def using assms by force + +lemma measurable_vimage: + fixes g :: "'a \ 'c" and f :: "'d \ 'a" + assumes "g \ measurable M M2" "f \ S \ space M" + shows "(\x. g (f x)) \ measurable (vimage_algebra M S f) M2" +proof - + note measurable_vimage_algebra[OF assms(2)] + from measurable_comp[OF this assms(1)] + show ?thesis by (simp add: comp_def) +qed + +lemma sigma_sets_vimage: + assumes "f \ S' \ S" and "A \ Pow S" + shows "sigma_sets S' ((\X. f -` X \ S') ` A) = (\X. f -` X \ S') ` sigma_sets S A" +proof (intro set_eqI iffI) + let ?F = "\X. f -` X \ S'" + fix X assume "X \ sigma_sets S' (?F ` A)" + then show "X \ ?F ` sigma_sets S A" + proof induct + case (Basic X) then obtain X' where "X = ?F X'" "X' \ A" + by auto + then show ?case by auto + next + case Empty then show ?case + by (auto intro!: image_eqI[of _ _ "{}"] sigma_sets.Empty) + next + case (Compl X) then obtain X' where X: "X = ?F X'" and "X' \ sigma_sets S A" + by auto + then have "S - X' \ sigma_sets S A" + by (auto intro!: sigma_sets.Compl) + then show ?case + using X assms by (auto intro!: image_eqI[where x="S - X'"]) + next + case (Union F) + then have "\i. \F'. F' \ sigma_sets S A \ F i = f -` F' \ S'" + by (auto simp: image_iff Bex_def) + from choice[OF this] obtain F' where + "\i. F' i \ sigma_sets S A" and "\i. F i = f -` F' i \ S'" + by auto + then show ?case + by (auto intro!: sigma_sets.Union image_eqI[where x="\i. F' i"]) + qed +next + let ?F = "\X. f -` X \ S'" + fix X assume "X \ ?F ` sigma_sets S A" + then obtain X' where "X' \ sigma_sets S A" "X = ?F X'" by auto + then show "X \ sigma_sets S' (?F ` A)" + proof (induct arbitrary: X) + case Empty then show ?case by (auto intro: sigma_sets.Empty) + next + case (Compl X') + have "S' - (S' - X) \ sigma_sets S' (?F ` A)" + apply (rule sigma_sets.Compl) + using assms by (auto intro!: Compl.hyps simp: Compl.prems) + also have "S' - (S' - X) = X" + using assms Compl by auto + finally show ?case . + next + case (Union F) + have "(\i. f -` F i \ S') \ sigma_sets S' (?F ` A)" + by (intro sigma_sets.Union Union.hyps) simp + also have "(\i. f -` F i \ S') = X" + using assms Union by auto + finally show ?case . + qed auto +qed + +subsubsection {* Restricted Space Sigma Algebra *} + +definition "restrict_space M \ = measure_of \ ((op \ \) ` sets M) (\A. emeasure M (A \ \))" + +lemma space_restrict_space: "space (restrict_space M \) = \" + unfolding restrict_space_def by (subst space_measure_of) auto + +lemma sets_restrict_space: "\ \ space M \ sets (restrict_space M \) = (op \ \) ` sets M" + using sigma_sets_vimage[of "\x. x" \ "space M" "sets M"] + unfolding restrict_space_def + by (subst sets_measure_of) (auto simp: sets.sigma_sets_eq Int_commute[of _ \] sets.space_closed) + +lemma sets_restrict_space_iff: + "\ \ sets M \ A \ sets (restrict_space M \) \ (A \ \ \ A \ sets M)" + by (subst sets_restrict_space) (auto dest: sets.sets_into_space) + +lemma measurable_restrict_space1: + assumes \: "\ \ sets M" and f: "f \ measurable M N" shows "f \ measurable (restrict_space M \) N" + unfolding measurable_def +proof (intro CollectI conjI ballI) + show sp: "f \ space (restrict_space M \) \ space N" + using measurable_space[OF f] sets.sets_into_space[OF \] by (auto simp: space_restrict_space) + + fix A assume "A \ sets N" + have "f -` A \ space (restrict_space M \) = (f -` A \ space M) \ \" + using sets.sets_into_space[OF \] by (auto simp: space_restrict_space) + also have "\ \ sets (restrict_space M \)" + unfolding sets_restrict_space_iff[OF \] + using measurable_sets[OF f `A \ sets N`] \ by blast + finally show "f -` A \ space (restrict_space M \) \ sets (restrict_space M \)" . +qed + +lemma measurable_restrict_space2: + "\ \ sets N \ f \ space M \ \ \ f \ measurable M N \ f \ measurable M (restrict_space N \)" + by (simp add: measurable_def space_restrict_space sets_restrict_space_iff) + end diff -r e5366291d6aa -r 8d5e5ec1cac3 src/HOL/Probability/document/root.tex --- a/src/HOL/Probability/document/root.tex Mon May 19 12:04:45 2014 +0200 +++ b/src/HOL/Probability/document/root.tex Mon May 19 13:44:13 2014 +0200 @@ -13,7 +13,7 @@ \begin{document} -\title{Multivariate Analysis} +\title{Measure and Probability Theory} \maketitle \tableofcontents diff -r e5366291d6aa -r 8d5e5ec1cac3 src/HOL/Probability/ex/Dining_Cryptographers.thy --- a/src/HOL/Probability/ex/Dining_Cryptographers.thy Mon May 19 12:04:45 2014 +0200 +++ b/src/HOL/Probability/ex/Dining_Cryptographers.thy Mon May 19 13:44:13 2014 +0200 @@ -8,7 +8,7 @@ lemma Ex1_eq: "\!x. P x \ P x \ P y \ x = y" by auto -section "Define the state space" +subsection {* Define the state space *} text {* diff -r e5366291d6aa -r 8d5e5ec1cac3 src/HOL/ROOT --- a/src/HOL/ROOT Mon May 19 12:04:45 2014 +0200 +++ b/src/HOL/ROOT Mon May 19 13:44:13 2014 +0200 @@ -685,8 +685,9 @@ options [document_graph] theories [document = false] "~~/src/HOL/Library/Countable" - "~~/src/HOL/Multivariate_Analysis/Extended_Real_Limits" "~~/src/HOL/Library/Permutation" + "~~/src/HOL/Library/Order_Continuity" + "~~/src/HOL/Library/Diagonal_Subsequence" theories Probability "ex/Dining_Cryptographers"