# HG changeset patch # User hoelzl # Date 1283453073 -7200 # Node ID e9467adb8b525232bd56541787373c88e2926a9f # Parent bddc3d3f6e537222900e7a66fb44c57b9e517fc7# Parent 5c714fd55b1e165b58bc5151e4a6c4717765c056 merged diff -r bddc3d3f6e53 -r e9467adb8b52 src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Thu Sep 02 18:45:23 2010 +0200 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Thu Sep 02 20:44:33 2010 +0200 @@ -5027,7 +5027,7 @@ (\a\s. \b\s. \x. (\i x$$i \ x$$i \ b$$i) \ (b$$i \ x$$i \ x$$i \ a$$i))) \ x \ s)" lemma is_interval_interval: "is_interval {a .. b::'a::ordered_euclidean_space}" (is ?th1) - "is_interval {a<..x y z::real. x < y \ y < z \ x < z" by auto show ?th1 ?th2 unfolding is_interval_def mem_interval Ball_def atLeastAtMost_iff by(meson order_trans le_less_trans less_le_trans *)+ qed @@ -5051,6 +5051,9 @@ lemma continuous_at_inner: "continuous (at x) (inner a)" unfolding continuous_at by (intro tendsto_intros) +lemma continuous_at_euclidean_component[intro!, simp]: "continuous (at x) (\x. x $$ i)" + unfolding euclidean_component_def by (rule continuous_at_inner) + lemma continuous_on_inner: fixes s :: "'a::real_inner set" shows "continuous_on s (inner a)" @@ -5159,6 +5162,9 @@ by (simp add: closed_def open_halfspace_component_lt) qed +lemma open_vimage_euclidean_component: "open S \ open ((\x. x $$ i) -` S)" + by (auto intro!: continuous_open_vimage) + text{* This gives a simple derivation of limit component bounds. *} lemma Lim_component_le: fixes f :: "'a \ 'b::euclidean_space" diff -r bddc3d3f6e53 -r e9467adb8b52 src/HOL/Probability/Borel.thy --- a/src/HOL/Probability/Borel.thy Thu Sep 02 18:45:23 2010 +0200 +++ b/src/HOL/Probability/Borel.thy Thu Sep 02 20:44:33 2010 +0200 @@ -6,6 +6,10 @@ imports Sigma_Algebra Positive_Infinite_Real Multivariate_Analysis begin +lemma LIMSEQ_max: + "u ----> (x::real) \ (\i. max (u i) 0) ----> max x 0" + by (fastsimp intro!: LIMSEQ_I dest!: LIMSEQ_D) + section "Generic Borel spaces" definition "borel_space = sigma (UNIV::'a::topological_space set) open" @@ -81,7 +85,7 @@ "(\x. c) \ borel_measurable M" by (auto intro!: measurable_const) -lemma (in sigma_algebra) borel_measurable_indicator: +lemma (in sigma_algebra) borel_measurable_indicator[simp, intro!]: assumes A: "A \ sets M" shows "indicator A \ borel_measurable M" unfolding indicator_def_raw using A @@ -105,6 +109,53 @@ qed (auto simp add: vimage_UN) qed +lemma (in sigma_algebra) borel_measurable_restricted: + fixes f :: "'a \ 'x\{topological_space, semiring_1}" assumes "A \ sets M" + shows "f \ borel_measurable (restricted_space A) \ + (\x. f x * indicator A x) \ borel_measurable M" + (is "f \ borel_measurable ?R \ ?f \ borel_measurable M") +proof - + interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \ sets M`]) + have *: "f \ borel_measurable ?R \ ?f \ borel_measurable ?R" + by (auto intro!: measurable_cong) + show ?thesis unfolding * + unfolding in_borel_measurable_borel_space + proof (simp, safe) + fix S :: "'x set" assume "S \ sets borel_space" + "\S\sets borel_space. ?f -` S \ A \ op \ A ` sets M" + then have "?f -` S \ A \ op \ A ` sets M" by auto + then have f: "?f -` S \ A \ sets M" + using `A \ sets M` sets_into_space by fastsimp + show "?f -` S \ space M \ sets M" + proof cases + assume "0 \ S" + then have "?f -` S \ space M = ?f -` S \ A \ (space M - A)" + using `A \ sets M` sets_into_space by auto + then show ?thesis using f `A \ sets M` by (auto intro!: Un Diff) + next + assume "0 \ S" + then have "?f -` S \ space M = ?f -` S \ A" + using `A \ sets M` sets_into_space + by (auto simp: indicator_def split: split_if_asm) + then show ?thesis using f by auto + qed + next + fix S :: "'x set" assume "S \ sets borel_space" + "\S\sets borel_space. ?f -` S \ space M \ sets M" + then have f: "?f -` S \ space M \ sets M" by auto + then show "?f -` S \ A \ op \ A ` sets M" + using `A \ sets M` sets_into_space + apply (simp add: image_iff) + apply (rule bexI[OF _ f]) + by auto + qed +qed + +lemma (in sigma_algebra) borel_measurable_subalgebra: + assumes "N \ sets M" "f \ borel_measurable (M\sets:=N\)" + shows "f \ borel_measurable M" + using assms unfolding measurable_def by auto + section "Borel spaces on euclidean spaces" lemma lessThan_borel[simp, intro]: @@ -658,6 +709,30 @@ "(f::'a \ real) \ borel_measurable M = (\a. {w \ space M. a < f w} \ sets M)" using borel_measurable_iff_halfspace_greater[where 'c=real] by simp +lemma borel_measureable_euclidean_component: + "(\x::'a::euclidean_space. x $$ i) \ borel_measurable borel_space" + unfolding borel_space_def[where 'a=real] +proof (rule borel_space.measurable_sigma) + fix S::"real set" assume "S \ open" then have "open S" unfolding mem_def . + from open_vimage_euclidean_component[OF this] + show "(\x. x $$ i) -` S \ space borel_space \ sets borel_space" + by (auto intro: borel_space_open) +qed auto + +lemma (in sigma_algebra) borel_measureable_euclidean_space: + fixes f :: "'a \ 'c::ordered_euclidean_space" + shows "f \ borel_measurable M \ (\ix. f x $$ i) \ borel_measurable M)" +proof safe + fix i assume "f \ borel_measurable M" + then show "(\x. f x $$ i) \ borel_measurable M" + using measurable_comp[of f _ _ "\x. x $$ i", unfolded comp_def] + by (auto intro: borel_measureable_euclidean_component) +next + assume f: "\ix. f x $$ i) \ borel_measurable M" + then show "f \ borel_measurable M" + unfolding borel_measurable_iff_halfspace_le by auto +qed + subsection "Borel measurable operators" lemma (in sigma_algebra) affine_borel_measurable_vector: @@ -1270,4 +1345,46 @@ using assms by auto qed +lemma (in sigma_algebra) borel_measurable_psuminf: + assumes "\i. f i \ borel_measurable M" + shows "(\x. (\\<^isub>\ i. f i x)) \ borel_measurable M" + using assms unfolding psuminf_def + by (auto intro!: borel_measurable_SUP[unfolded SUPR_fun_expand]) + +section "LIMSEQ is borel measurable" + +lemma (in sigma_algebra) borel_measurable_LIMSEQ: + fixes u :: "nat \ 'a \ real" + assumes u': "\x. x \ space M \ (\i. u i x) ----> u' x" + and u: "\i. u i \ borel_measurable M" + shows "u' \ borel_measurable M" +proof - + let "?pu x i" = "max (u i x) 0" + let "?nu x i" = "max (- u i x) 0" + + { fix x assume x: "x \ space M" + have "(?pu x) ----> max (u' x) 0" + "(?nu x) ----> max (- u' x) 0" + using u'[OF x] by (auto intro!: LIMSEQ_max LIMSEQ_minus) + from LIMSEQ_imp_lim_INF[OF _ this(1)] LIMSEQ_imp_lim_INF[OF _ this(2)] + have "(SUP n. INF m. Real (u (n + m) x)) = Real (u' x)" + "(SUP n. INF m. Real (- u (n + m) x)) = Real (- u' x)" + by (simp_all add: Real_max'[symmetric]) } + note eq = this + + have *: "\x. real (Real (u' x)) - real (Real (- u' x)) = u' x" + by auto + + have "(SUP n. INF m. (\x. Real (u (n + m) x))) \ borel_measurable M" + "(SUP n. INF m. (\x. Real (- u (n + m) x))) \ borel_measurable M" + using u by (auto intro: borel_measurable_SUP borel_measurable_INF borel_measurable_Real) + with eq[THEN measurable_cong, of M "\x. x" borel_space] + have "(\x. Real (u' x)) \ borel_measurable M" + "(\x. Real (- u' x)) \ borel_measurable M" + unfolding SUPR_fun_expand INFI_fun_expand by auto + note this[THEN borel_measurable_real] + from borel_measurable_diff[OF this] + show ?thesis unfolding * . +qed + end diff -r bddc3d3f6e53 -r e9467adb8b52 src/HOL/Probability/Caratheodory.thy --- a/src/HOL/Probability/Caratheodory.thy Thu Sep 02 18:45:23 2010 +0200 +++ b/src/HOL/Probability/Caratheodory.thy Thu Sep 02 20:44:33 2010 +0200 @@ -445,21 +445,6 @@ by intro_locales (auto simp add: sigma_algebra_def) qed - -lemma (in algebra) inf_measure_nonempty: - assumes f: "positive f" and b: "b \ sets M" and a: "a \ b" - shows "f b \ measure_set M f a" -proof - - have "psuminf (f \ (\i. {})(0 := b)) = setsum (f \ (\i. {})(0 := b)) {..<1::nat}" - by (rule psuminf_finite) (simp add: f[unfolded positive_def]) - also have "... = f b" - by simp - finally have "psuminf (f \ (\i. {})(0 := b)) = f b" . - thus ?thesis using a b - by (auto intro!: exI [of _ "(\i. {})(0 := b)"] - simp: measure_set_def disjoint_family_on_def split_if_mem2 comp_def) -qed - lemma (in algebra) additive_increasing: assumes posf: "positive f" and addf: "additive M f" shows "increasing M f" @@ -494,6 +479,20 @@ by (auto simp add: Un binaryset_psuminf positive_def) qed +lemma inf_measure_nonempty: + assumes f: "positive f" and b: "b \ sets M" and a: "a \ b" "{} \ sets M" + shows "f b \ measure_set M f a" +proof - + have "psuminf (f \ (\i. {})(0 := b)) = setsum (f \ (\i. {})(0 := b)) {..<1::nat}" + by (rule psuminf_finite) (simp add: f[unfolded positive_def]) + also have "... = f b" + by simp + finally have "psuminf (f \ (\i. {})(0 := b)) = f b" . + thus ?thesis using assms + by (auto intro!: exI [of _ "(\i. {})(0 := b)"] + simp: measure_set_def disjoint_family_on_def split_if_mem2 comp_def) +qed + lemma (in algebra) inf_measure_agrees: assumes posf: "positive f" and ca: "countably_additive M f" and s: "s \ sets M" @@ -535,11 +534,11 @@ qed lemma (in algebra) inf_measure_empty: - assumes posf: "positive f" + assumes posf: "positive f" "{} \ sets M" shows "Inf (measure_set M f {}) = 0" proof (rule antisym) show "Inf (measure_set M f {}) \ 0" - by (metis complete_lattice_class.Inf_lower empty_sets inf_measure_nonempty[OF posf] subset_refl posf[unfolded positive_def]) + by (metis complete_lattice_class.Inf_lower `{} \ sets M` inf_measure_nonempty[OF posf] subset_refl posf[unfolded positive_def]) qed simp lemma (in algebra) inf_measure_positive: @@ -597,7 +596,7 @@ next case True have "measure_set M f s \ {}" - by (metis emptyE ss inf_measure_nonempty [of f, OF posf top]) + by (metis emptyE ss inf_measure_nonempty [of f, OF posf top _ empty_sets]) then obtain l where "l \ measure_set M f s" by auto moreover from True have "l \ Inf (measure_set M f s) + e" by simp ultimately show ?thesis diff -r bddc3d3f6e53 -r e9467adb8b52 src/HOL/Probability/Information.thy --- a/src/HOL/Probability/Information.thy Thu Sep 02 18:45:23 2010 +0200 +++ b/src/HOL/Probability/Information.thy Thu Sep 02 20:44:33 2010 +0200 @@ -2,11 +2,53 @@ imports Probability_Space Product_Measure Convex Radon_Nikodym begin +lemma log_le: "1 < a \ 0 < x \ x \ y \ log a x \ log a y" + by (subst log_le_cancel_iff) auto + +lemma log_less: "1 < a \ 0 < x \ x < y \ log a x < log a y" + by (subst log_less_cancel_iff) auto + +lemma setsum_cartesian_product': + "(\x\A \ B. f x) = (\x\A. setsum (\y. f (x, y)) B)" + unfolding setsum_cartesian_product by simp + lemma real_of_pinfreal_inverse[simp]: fixes X :: pinfreal shows "real (inverse X) = 1 / real X" by (cases X) (auto simp: inverse_eq_divide) +lemma (in finite_prob_space) finite_product_prob_space_of_images: + "finite_prob_space \ space = X ` space M \ Y ` space M, sets = Pow (X ` space M \ Y ` space M)\ + (joint_distribution X Y)" + (is "finite_prob_space ?S _") +proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images) + have "X -` X ` space M \ Y -` Y ` space M \ space M = space M" by auto + thus "joint_distribution X Y (X ` space M \ Y ` space M) = 1" + by (simp add: distribution_def prob_space vimage_Times comp_def measure_space_1) +qed + +lemma (in finite_prob_space) finite_measure_space_prod: + assumes X: "finite_measure_space MX (distribution X)" + assumes Y: "finite_measure_space MY (distribution Y)" + shows "finite_measure_space (prod_measure_space MX MY) (joint_distribution X Y)" + (is "finite_measure_space ?M ?D") +proof (intro finite_measure_spaceI) + interpret X: finite_measure_space MX "distribution X" by fact + interpret Y: finite_measure_space MY "distribution Y" by fact + note finite_measure_space.finite_prod_measure_space[OF X Y, simp] + show "finite (space ?M)" using X.finite_space Y.finite_space by auto + show "joint_distribution X Y {} = 0" by simp + show "sets ?M = Pow (space ?M)" by simp + { fix x show "?D (space ?M) \ \" by (rule distribution_finite) } + { fix A B assume "A \ space ?M" "B \ space ?M" "A \ B = {}" + have *: "(\t. (X t, Y t)) -` (A \ B) \ space M = + (\t. (X t, Y t)) -` A \ space M \ (\t. (X t, Y t)) -` B \ space M" + by auto + show "?D (A \ B) = ?D A + ?D B" unfolding distribution_def * + apply (rule measure_additive[symmetric]) + using `A \ B = {}` by (auto simp: sets_eq_Pow) } +qed + section "Convex theory" lemma log_setsum: @@ -105,51 +147,19 @@ finally show ?thesis . qed -lemma (in finite_prob_space) sum_over_space_distrib: - "(\x\X`space M. distribution X {x}) = 1" - unfolding distribution_def measure_space_1[symmetric] using finite_space - by (subst measure_finitely_additive'') - (auto simp add: disjoint_family_on_def sets_eq_Pow intro!: arg_cong[where f=\]) - -lemma (in finite_prob_space) sum_over_space_real_distribution: - "(\x\X`space M. real (distribution X {x})) = 1" - unfolding distribution_def prob_space[symmetric] using finite_space - by (subst real_finite_measure_finite_Union[symmetric]) - (auto simp add: disjoint_family_on_def sets_eq_Pow intro!: arg_cong[where f=prob]) +lemma split_pairs: + shows + "((A, B) = X) \ (fst X = A \ snd X = B)" and + "(X = (A, B)) \ (fst X = A \ snd X = B)" by auto section "Information theory" -definition - "KL_divergence b M \ \ = - measure_space.integral M \ (\x. log b (real (sigma_finite_measure.RN_deriv M \ \ x)))" - locale finite_information_space = finite_prob_space + fixes b :: real assumes b_gt_1: "1 < b" -lemma (in finite_prob_space) distribution_mono: - assumes "\t. \ t \ space M ; X t \ x \ \ Y t \ y" - shows "distribution X x \ distribution Y y" - unfolding distribution_def - using assms by (auto simp: sets_eq_Pow intro!: measure_mono) - -lemma (in prob_space) distribution_remove_const: - shows "joint_distribution X (\x. ()) {(x, ())} = distribution X {x}" - and "joint_distribution (\x. ()) X {((), x)} = distribution X {x}" - and "joint_distribution X (\x. (Y x, ())) {(x, y, ())} = joint_distribution X Y {(x, y)}" - and "joint_distribution X (\x. ((), Y x)) {(x, (), y)} = joint_distribution X Y {(x, y)}" - and "distribution (\x. ()) {()} = 1" - unfolding measure_space_1[symmetric] - by (auto intro!: arg_cong[where f="\"] simp: distribution_def) - context finite_information_space begin -lemma distribution_mono_gt_0: - assumes gt_0: "0 < distribution X x" - assumes *: "\t. \ t \ space M ; X t \ x \ \ Y t \ y" - shows "0 < distribution Y y" - by (rule less_le_trans[OF gt_0 distribution_mono]) (rule *) - lemma assumes "0 \ A" and pos: "0 < A \ 0 < B" "0 < A \ 0 < C" shows mult_log_mult: "A * log b (B * C) = A * log b B + A * log b C" (is "?mult") @@ -165,41 +175,6 @@ thus ?mult and ?div by auto qed -lemma split_pairs: - shows - "((A, B) = X) \ (fst X = A \ snd X = B)" and - "(X = (A, B)) \ (fst X = A \ snd X = B)" by auto - -lemma (in finite_information_space) distribution_finite: - "distribution X A \ \" - using measure_finite[of "X -` A \ space M"] - unfolding distribution_def sets_eq_Pow by auto - -lemma (in finite_information_space) real_distribution_gt_0[simp]: - "0 < real (distribution Y y) \ 0 < distribution Y y" - using assms by (auto intro!: real_pinfreal_pos distribution_finite) - -lemma real_distribution_mult_pos_pos: - assumes "0 < distribution Y y" - and "0 < distribution X x" - shows "0 < real (distribution Y y * distribution X x)" - unfolding real_of_pinfreal_mult[symmetric] - using assms by (auto intro!: mult_pos_pos) - -lemma real_distribution_divide_pos_pos: - assumes "0 < distribution Y y" - and "0 < distribution X x" - shows "0 < real (distribution Y y / distribution X x)" - unfolding divide_pinfreal_def real_of_pinfreal_mult[symmetric] - using assms distribution_finite[of X x] by (cases "distribution X x") (auto intro!: mult_pos_pos) - -lemma real_distribution_mult_inverse_pos_pos: - assumes "0 < distribution Y y" - and "0 < distribution X x" - shows "0 < real (distribution Y y * inverse (distribution X x))" - unfolding divide_pinfreal_def real_of_pinfreal_mult[symmetric] - using assms distribution_finite[of X x] by (cases "distribution X x") (auto intro!: mult_pos_pos) - ML {* (* tactic to solve equations of the form @{term "W * log b (X / (Y * Z)) = W * log b X - W * log b (Y * Z)"} @@ -252,31 +227,14 @@ end -lemma (in finite_measure_space) absolutely_continuousI: - assumes "finite_measure_space M \" - assumes v: "\x. \ x \ space M ; \ {x} = 0 \ \ \ {x} = 0" - shows "absolutely_continuous \" -proof (unfold absolutely_continuous_def sets_eq_Pow, safe) - fix N assume "\ N = 0" "N \ space M" - - interpret v: finite_measure_space M \ by fact +subsection "Kullback$-$Leibler divergence" - have "\ N = \ (\x\N. {x})" by simp - also have "\ = (\x\N. \ {x})" - proof (rule v.measure_finitely_additive''[symmetric]) - show "finite N" using `N \ space M` finite_space by (auto intro: finite_subset) - show "disjoint_family_on (\i. {i}) N" unfolding disjoint_family_on_def by auto - fix x assume "x \ N" thus "{x} \ sets M" using `N \ space M` sets_eq_Pow by auto - qed - also have "\ = 0" - proof (safe intro!: setsum_0') - fix x assume "x \ N" - hence "\ {x} \ \ N" using sets_eq_Pow `N \ space M` by (auto intro!: measure_mono) - hence "\ {x} = 0" using `\ N = 0` by simp - thus "\ {x} = 0" using v[of x] `x \ N` `N \ space M` by auto - qed - finally show "\ N = 0" . -qed +text {* The Kullback$-$Leibler divergence is also known as relative entropy or +Kullback$-$Leibler distance. *} + +definition + "KL_divergence b M \ \ = + measure_space.integral M \ (\x. log b (real (sigma_finite_measure.RN_deriv M \ \ x)))" lemma (in finite_measure_space) KL_divergence_eq_finite: assumes v: "finite_measure_space M \" @@ -285,19 +243,13 @@ proof (simp add: KL_divergence_def finite_measure_space.integral_finite_singleton[OF v]) interpret v: finite_measure_space M \ by fact have ms: "measure_space M \" by fact - have ac: "absolutely_continuous \" using ac by (auto intro!: absolutely_continuousI[OF v]) - show "(\x \ space M. log b (real (RN_deriv \ x)) * real (\ {x})) = ?sum" using RN_deriv_finite_measure[OF ms ac] by (auto intro!: setsum_cong simp: field_simps real_of_pinfreal_mult[symmetric]) qed -lemma (in finite_prob_space) finite_sum_over_space_eq_1: - "(\x\space M. real (\ {x})) = 1" - using sum_over_space_eq_1 finite_measure by (simp add: real_of_pinfreal_setsum sets_eq_Pow) - lemma (in finite_prob_space) KL_divergence_positive_finite: assumes v: "finite_prob_space M \" assumes ac: "\x. \ x \ space M ; \ {x} = 0 \ \ \ {x} = 0" @@ -322,13 +274,15 @@ fix x assume x: "x \ space M" { assume "0 < real (\ {x})" hence "\ {x} \ 0" using ac[OF x] by auto - thus "0 < prob {x}" using measure_finite[of "{x}"] sets_eq_Pow x + thus "0 < prob {x}" using finite_measure[of "{x}"] sets_eq_Pow x by (cases "\ {x}") simp_all } qed auto qed thus "0 \ KL_divergence b M \ \" using finite_sum_over_space_eq_1 by simp qed +subsection {* Mutual Information *} + definition (in prob_space) "mutual_information b S T X Y = KL_divergence b (prod_measure_space S T) @@ -341,24 +295,48 @@ \ space = X`space M, sets = Pow (X`space M) \ \ space = Y`space M, sets = Pow (Y`space M) \ X Y" -lemma prod_measure_times_finite: - assumes fms: "finite_measure_space M \" "finite_measure_space N \" and a: "a \ space M \ space N" - shows "prod_measure M \ N \ {a} = \ {fst a} * \ {snd a}" -proof (cases a) - case (Pair b c) - hence a_eq: "{a} = {b} \ {c}" by simp - - interpret M: finite_measure_space M \ by fact - interpret N: finite_measure_space N \ by fact - - from finite_measure_space.finite_prod_measure_times[OF fms, of "{b}" "{c}"] M.sets_eq_Pow N.sets_eq_Pow a Pair - show ?thesis unfolding a_eq by simp +lemma (in finite_information_space) mutual_information_generic_eq: + assumes MX: "finite_measure_space MX (distribution X)" + assumes MY: "finite_measure_space MY (distribution Y)" + shows "mutual_information b MX MY X Y = (\ (x,y) \ space MX \ space MY. + real (joint_distribution X Y {(x,y)}) * + log b (real (joint_distribution X Y {(x,y)}) / + (real (distribution X {x}) * real (distribution Y {y}))))" +proof - + let ?P = "prod_measure_space MX MY" + let ?\ = "prod_measure MX (distribution X) MY (distribution Y)" + let ?\ = "joint_distribution X Y" + interpret X: finite_measure_space MX "distribution X" by fact + moreover interpret Y: finite_measure_space MY "distribution Y" by fact + have fms: "finite_measure_space MX (distribution X)" + "finite_measure_space MY (distribution Y)" by fact+ + have fms_P: "finite_measure_space ?P ?\" + by (rule X.finite_measure_space_finite_prod_measure) fact + then interpret P: finite_measure_space ?P ?\ . + have fms_P': "finite_measure_space ?P ?\" + using finite_product_measure_space[of "space MX" "space MY"] + X.finite_space Y.finite_space sigma_prod_sets_finite[OF X.finite_space Y.finite_space] + X.sets_eq_Pow Y.sets_eq_Pow + by (simp add: prod_measure_space_def sigma_def) + then interpret P': finite_measure_space ?P ?\ . + { fix x assume "x \ space ?P" + hence in_MX: "{fst x} \ sets MX" "{snd x} \ sets MY" using X.sets_eq_Pow Y.sets_eq_Pow + by (auto simp: prod_measure_space_def) + assume "?\ {x} = 0" + with X.finite_prod_measure_times[OF fms(2), of "{fst x}" "{snd x}"] in_MX + have "distribution X {fst x} = 0 \ distribution Y {snd x} = 0" + by (simp add: prod_measure_space_def) + hence "joint_distribution X Y {x} = 0" + by (cases x) (auto simp: distribution_order) } + note measure_0 = this + show ?thesis + unfolding Let_def mutual_information_def + using measure_0 fms_P fms_P' MX MY P.absolutely_continuous_def + by (subst P.KL_divergence_eq_finite) + (auto simp add: prod_measure_space_def prod_measure_times_finite + finite_prob_space_eq setsum_cartesian_product' real_of_pinfreal_mult[symmetric]) qed -lemma setsum_cartesian_product': - "(\x\A \ B. f x) = (\x\A. setsum (\y. f (x, y)) B)" - unfolding setsum_cartesian_product by simp - lemma (in finite_information_space) assumes MX: "finite_prob_space MX (distribution X)" assumes MY: "finite_prob_space MY (distribution Y)" @@ -436,9 +414,26 @@ (real (distribution X {x}) * real (distribution Y {y}))))" by (subst mutual_information_eq_generic) (simp_all add: finite_prob_space_of_images) +lemma (in finite_information_space) mutual_information_cong: + assumes X: "\x. x \ space M \ X x = X' x" + assumes Y: "\x. x \ space M \ Y x = Y' x" + shows "\(X ; Y) = \(X' ; Y')" +proof - + have "X ` space M = X' ` space M" using X by (auto intro!: image_eqI) + moreover have "Y ` space M = Y' ` space M" using Y by (auto intro!: image_eqI) + ultimately show ?thesis + unfolding mutual_information_eq + using + assms[THEN distribution_cong] + joint_distribution_cong[OF assms] + by (auto intro!: setsum_cong) +qed + lemma (in finite_information_space) mutual_information_positive: "0 \ \(X;Y)" by (subst mutual_information_positive_generic) (simp_all add: finite_prob_space_of_images) +subsection {* Entropy *} + definition (in prob_space) "entropy b s X = mutual_information b s s X X" @@ -446,32 +441,146 @@ finite_entropy ("\'(_')") where "\(X) \ entropy b \ space = X`space M, sets = Pow (X`space M) \ X" -lemma (in finite_information_space) joint_distribution_remove[simp]: - "joint_distribution X X {(x, x)} = distribution X {x}" - unfolding distribution_def by (auto intro!: arg_cong[where f="\"]) +lemma (in finite_information_space) entropy_generic_eq: + assumes MX: "finite_measure_space MX (distribution X)" + shows "entropy b MX X = -(\ x \ space MX. real (distribution X {x}) * log b (real (distribution X {x})))" +proof - + let "?X x" = "real (distribution X {x})" + let "?XX x y" = "real (joint_distribution X X {(x, y)})" + interpret MX: finite_measure_space MX "distribution X" by fact + { fix x y + have "(\x. (X x, X x)) -` {(x, y)} = (if x = y then X -` {x} else {})" by auto + then have "?XX x y * log b (?XX x y / (?X x * ?X y)) = + (if x = y then - ?X y * log b (?X y) else 0)" + unfolding distribution_def by (auto simp: mult_log_divide) } + note remove_XX = this + show ?thesis + unfolding entropy_def mutual_information_generic_eq[OF MX MX] + unfolding setsum_cartesian_product[symmetric] setsum_negf[symmetric] remove_XX + by (auto simp: setsum_cases MX.finite_space) +qed lemma (in finite_information_space) entropy_eq: "\(X) = -(\ x \ X ` space M. real (distribution X {x}) * log b (real (distribution X {x})))" -proof - - { fix f - { fix x y - have "(\x. (X x, X x)) -` {(x, y)} = (if x = y then X -` {x} else {})" by auto - hence "real (distribution (\x. (X x, X x)) {(x,y)}) * f x y = - (if x = y then real (distribution X {x}) * f x y else 0)" - unfolding distribution_def by auto } - hence "(\(x, y) \ X ` space M \ X ` space M. real (joint_distribution X X {(x, y)}) * f x y) = - (\x \ X ` space M. real (distribution X {x}) * f x x)" - unfolding setsum_cartesian_product' by (simp add: setsum_cases finite_space) } - note remove_cartesian_product = this - - show ?thesis - unfolding entropy_def mutual_information_eq setsum_negf[symmetric] remove_cartesian_product - by (auto intro!: setsum_cong) -qed + by (simp add: finite_measure_space entropy_generic_eq) lemma (in finite_information_space) entropy_positive: "0 \ \(X)" unfolding entropy_def using mutual_information_positive . +lemma (in finite_information_space) entropy_certainty_eq_0: + assumes "x \ X ` space M" and "distribution X {x} = 1" + shows "\(X) = 0" +proof - + interpret X: finite_prob_space "\ space = X ` space M, sets = Pow (X ` space M) \" "distribution X" + by (rule finite_prob_space_of_images) + + have "distribution X (X ` space M - {x}) = distribution X (X ` space M) - distribution X {x}" + using X.measure_compl[of "{x}"] assms by auto + also have "\ = 0" using X.prob_space assms by auto + finally have X0: "distribution X (X ` space M - {x}) = 0" by auto + + { fix y assume asm: "y \ x" "y \ X ` space M" + hence "{y} \ X ` space M - {x}" by auto + from X.measure_mono[OF this] X0 asm + have "distribution X {y} = 0" by auto } + + hence fi: "\ y. y \ X ` space M \ real (distribution X {y}) = (if x = y then 1 else 0)" + using assms by auto + + have y: "\y. (if x = y then 1 else 0) * log b (if x = y then 1 else 0) = 0" by simp + + show ?thesis unfolding entropy_eq by (auto simp: y fi) +qed + +lemma (in finite_information_space) entropy_le_card_not_0: + "\(X) \ log b (real (card (X ` space M \ {x . distribution X {x} \ 0})))" +proof - + let "?d x" = "distribution X {x}" + let "?p x" = "real (?d x)" + have "\(X) = (\x\X`space M. ?p x * log b (1 / ?p x))" + by (auto intro!: setsum_cong simp: entropy_eq setsum_negf[symmetric]) + also have "\ \ log b (\x\X`space M. ?p x * (1 / ?p x))" + apply (rule log_setsum') + using not_empty b_gt_1 finite_space sum_over_space_real_distribution + by auto + also have "\ = log b (\x\X`space M. if ?d x \ 0 then 1 else 0)" + apply (rule arg_cong[where f="\f. log b (\x\X`space M. f x)"]) + using distribution_finite[of X] by (auto simp: expand_fun_eq real_of_pinfreal_eq_0) + finally show ?thesis + using finite_space by (auto simp: setsum_cases real_eq_of_nat) +qed + +lemma (in finite_information_space) entropy_uniform_max: + assumes "\x y. \ x \ X ` space M ; y \ X ` space M \ \ distribution X {x} = distribution X {y}" + shows "\(X) = log b (real (card (X ` space M)))" +proof - + note uniform = + finite_prob_space_of_images[of X, THEN finite_prob_space.uniform_prob, simplified] + + have card_gt0: "0 < card (X ` space M)" unfolding card_gt_0_iff + using finite_space not_empty by auto + + { fix x assume "x \ X ` space M" + hence "real (distribution X {x}) = 1 / real (card (X ` space M))" + proof (rule uniform) + fix x y assume "x \ X`space M" "y \ X`space M" + from assms[OF this] show "real (distribution X {x}) = real (distribution X {y})" by simp + qed } + thus ?thesis + using not_empty finite_space b_gt_1 card_gt0 + by (simp add: entropy_eq real_eq_of_nat[symmetric] log_divide) +qed + +lemma (in finite_information_space) entropy_le_card: + "\(X) \ log b (real (card (X ` space M)))" +proof cases + assume "X ` space M \ {x. distribution X {x} \ 0} = {}" + then have "\x. x\X`space M \ distribution X {x} = 0" by auto + moreover + have "0 < card (X`space M)" + using finite_space not_empty unfolding card_gt_0_iff by auto + then have "log b 1 \ log b (real (card (X`space M)))" + using b_gt_1 by (intro log_le) auto + ultimately show ?thesis unfolding entropy_eq by simp +next + assume False: "X ` space M \ {x. distribution X {x} \ 0} \ {}" + have "card (X ` space M \ {x. distribution X {x} \ 0}) \ card (X ` space M)" + (is "?A \ ?B") using finite_space not_empty by (auto intro!: card_mono) + note entropy_le_card_not_0 + also have "log b (real ?A) \ log b (real ?B)" + using b_gt_1 False finite_space not_empty `?A \ ?B` + by (auto intro!: log_le simp: card_gt_0_iff) + finally show ?thesis . +qed + +lemma (in finite_information_space) entropy_commute: + "\(\x. (X x, Y x)) = \(\x. (Y x, X x))" +proof - + have *: "(\x. (Y x, X x))`space M = (\(a,b). (b,a))`(\x. (X x, Y x))`space M" + by auto + have inj: "\X. inj_on (\(a,b). (b,a)) X" + by (auto intro!: inj_onI) + show ?thesis + unfolding entropy_eq unfolding * setsum_reindex[OF inj] + by (simp add: joint_distribution_commute[of Y X] split_beta) +qed + +lemma (in finite_information_space) entropy_eq_cartesian_sum: + "\(\x. (X x, Y x)) = -(\x\X`space M. \y\Y`space M. + real (joint_distribution X Y {(x,y)}) * + log b (real (joint_distribution X Y {(x,y)})))" +proof - + { fix x assume "x\(\x. (X x, Y x))`space M" + then have "(\x. (X x, Y x)) -` {x} \ space M = {}" by auto + then have "joint_distribution X Y {x} = 0" + unfolding distribution_def by auto } + then show ?thesis using finite_space + unfolding entropy_eq neg_equal_iff_equal setsum_cartesian_product + by (auto intro!: setsum_mono_zero_cong_left) +qed + +subsection {* Conditional Mutual Information *} + definition (in prob_space) "conditional_mutual_information b M1 M2 M3 X Y Z \ mutual_information b M1 (prod_measure_space M2 M3) X (\x. (Y x, Z x)) - @@ -485,87 +594,32 @@ \ space = Z`space M, sets = Pow (Z`space M) \ X Y Z" -lemma (in finite_information_space) setsum_distribution_gen: - assumes "Z -` {c} \ space M = (\x \ X`space M. Y -` {f x}) \ space M" - and "inj_on f (X`space M)" - shows "(\x \ X`space M. distribution Y {f x}) = distribution Z {c}" - unfolding distribution_def assms - using finite_space assms - by (subst measure_finitely_additive'') - (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def - intro!: arg_cong[where f=prob]) - -lemma (in finite_information_space) setsum_distribution: - "(\x \ X`space M. joint_distribution X Y {(x, y)}) = distribution Y {y}" - "(\y \ Y`space M. joint_distribution X Y {(x, y)}) = distribution X {x}" - "(\x \ X`space M. joint_distribution X (\x. (Y x, Z x)) {(x, y, z)}) = joint_distribution Y Z {(y, z)}" - "(\y \ Y`space M. joint_distribution X (\x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Z {(x, z)}" - "(\z \ Z`space M. joint_distribution X (\x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Y {(x, y)}" - by (auto intro!: inj_onI setsum_distribution_gen) - -lemma (in finite_information_space) setsum_real_distribution_gen: - assumes "Z -` {c} \ space M = (\x \ X`space M. Y -` {f x}) \ space M" - and "inj_on f (X`space M)" - shows "(\x \ X`space M. real (distribution Y {f x})) = real (distribution Z {c})" - unfolding distribution_def assms - using finite_space assms - by (subst real_finite_measure_finite_Union[symmetric]) - (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def - intro!: arg_cong[where f=prob]) - -lemma (in finite_information_space) setsum_real_distribution: - "(\x \ X`space M. real (joint_distribution X Y {(x, y)})) = real (distribution Y {y})" - "(\y \ Y`space M. real (joint_distribution X Y {(x, y)})) = real (distribution X {x})" - "(\x \ X`space M. real (joint_distribution X (\x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution Y Z {(y, z)})" - "(\y \ Y`space M. real (joint_distribution X (\x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution X Z {(x, z)})" - "(\z \ Z`space M. real (joint_distribution X (\x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution X Y {(x, y)})" - by (auto intro!: inj_onI setsum_real_distribution_gen) +lemma (in finite_information_space) conditional_mutual_information_generic_eq: + assumes MX: "finite_measure_space MX (distribution X)" + assumes MY: "finite_measure_space MY (distribution Y)" + assumes MZ: "finite_measure_space MZ (distribution Z)" + shows "conditional_mutual_information b MX MY MZ X Y Z = + (\(x, y, z)\space MX \ space MY \ space MZ. + real (joint_distribution X (\x. (Y x, Z x)) {(x, y, z)}) * + log b (real (joint_distribution X (\x. (Y x, Z x)) {(x, y, z)}) / + (real (distribution X {x}) * real (joint_distribution Y Z {(y, z)})))) - + (\(x, y)\space MX \ space MZ. + real (joint_distribution X Z {(x, y)}) * + log b (real (joint_distribution X Z {(x, y)}) / (real (distribution X {x}) * real (distribution Z {y}))))" + using assms finite_measure_space_prod[OF MY MZ] + unfolding conditional_mutual_information_def + by (subst (1 2) mutual_information_generic_eq) + (simp_all add: setsum_cartesian_product' finite_measure_space.finite_prod_measure_space) -lemma (in finite_information_space) conditional_mutual_information_eq_sum: - "\(X ; Y | Z) = - (\(x, y, z)\X ` space M \ (\x. (Y x, Z x)) ` space M. - real (distribution (\x. (X x, Y x, Z x)) {(x, y, z)}) * - log b (real (distribution (\x. (X x, Y x, Z x)) {(x, y, z)})/ - real (distribution (\x. (Y x, Z x)) {(y, z)}))) - - (\(x, z)\X ` space M \ Z ` space M. - real (distribution (\x. (X x, Z x)) {(x,z)}) * log b (real (distribution (\x. (X x, Z x)) {(x,z)}) / real (distribution Z {z})))" - (is "_ = ?rhs") -proof - - have setsum_product: - "\f x. (\v\(\x. (Y x, Z x)) ` space M. real (joint_distribution X (\x. (Y x, Z x)) {(x,v)}) * f v) - = (\v\Y ` space M \ Z ` space M. real (joint_distribution X (\x. (Y x, Z x)) {(x,v)}) * f v)" - proof (safe intro!: setsum_mono_zero_cong_left imageI) - fix x y z f - assume *: "(Y y, Z z) \ (\x. (Y x, Z x)) ` space M" and "y \ space M" "z \ space M" - hence "(\x. (X x, Y x, Z x)) -` {(x, Y y, Z z)} \ space M = {}" - proof safe - fix x' assume x': "x' \ space M" and eq: "Y x' = Y y" "Z x' = Z z" - have "(Y y, Z z) \ (\x. (Y x, Z x)) ` space M" using eq[symmetric] x' by auto - thus "x' \ {}" using * by auto - qed - thus "real (joint_distribution X (\x. (Y x, Z x)) {(x, Y y, Z z)}) * f (Y y) (Z z) = 0" - unfolding distribution_def by simp - qed (simp add: finite_space) - - thus ?thesis - unfolding conditional_mutual_information_def Let_def mutual_information_eq - by (subst mutual_information_eq_generic) - (auto simp: prod_measure_space_def sigma_prod_sets_finite finite_space sigma_def - finite_prob_space_of_images finite_product_prob_space_of_images - setsum_cartesian_product' setsum_product setsum_subtractf setsum_addf - setsum_left_distrib[symmetric] setsum_real_distribution - cong: setsum_cong) -qed lemma (in finite_information_space) conditional_mutual_information_eq: "\(X ; Y | Z) = (\(x, y, z) \ X ` space M \ Y ` space M \ Z ` space M. real (distribution (\x. (X x, Y x, Z x)) {(x, y, z)}) * log b (real (distribution (\x. (X x, Y x, Z x)) {(x, y, z)}) / (real (joint_distribution X Z {(x, z)}) * real (joint_distribution Y Z {(y,z)} / distribution Z {z}))))" - unfolding conditional_mutual_information_def Let_def mutual_information_eq - by (subst mutual_information_eq_generic) + by (subst conditional_mutual_information_generic_eq) (auto simp add: prod_measure_space_def sigma_prod_sets_finite finite_space - finite_prob_space_of_images finite_product_prob_space_of_images sigma_def + finite_measure_space finite_product_prob_space_of_images sigma_def setsum_cartesian_product' setsum_product setsum_subtractf setsum_addf setsum_left_distrib[symmetric] setsum_real_distribution setsum_commute[where A="Y`space M"] real_of_pinfreal_mult[symmetric] @@ -581,22 +635,6 @@ by (simp add: setsum_cartesian_product' distribution_remove_const) qed -lemma (in finite_prob_space) distribution_finite: - "distribution X A \ \" - by (auto simp: sets_eq_Pow distribution_def intro!: measure_finite) - -lemma (in finite_prob_space) real_distribution_order: - shows "r \ real (joint_distribution X Y {(x, y)}) \ r \ real (distribution X {x})" - and "r \ real (joint_distribution X Y {(x, y)}) \ r \ real (distribution Y {y})" - and "r < real (joint_distribution X Y {(x, y)}) \ r < real (distribution X {x})" - and "r < real (joint_distribution X Y {(x, y)}) \ r < real (distribution Y {y})" - and "distribution X {x} = 0 \ real (joint_distribution X Y {(x, y)}) = 0" - and "distribution Y {y} = 0 \ real (joint_distribution X Y {(x, y)}) = 0" - using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_fst, of X Y "{(x, y)}"] - using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_snd, of X Y "{(x, y)}"] - using real_pinfreal_nonneg[of "joint_distribution X Y {(x, y)}"] - by auto - lemma (in finite_information_space) conditional_mutual_information_positive: "0 \ \(X ; Y | Z)" proof - @@ -640,6 +678,8 @@ by (simp add: real_of_pinfreal_mult[symmetric]) qed +subsection {* Conditional Entropy *} + definition (in prob_space) "conditional_entropy b S T X Y = conditional_mutual_information b S S T X X Y" @@ -652,19 +692,69 @@ lemma (in finite_information_space) conditional_entropy_positive: "0 \ \(X | Y)" unfolding conditional_entropy_def using conditional_mutual_information_positive . +lemma (in finite_information_space) conditional_entropy_generic_eq: + assumes MX: "finite_measure_space MX (distribution X)" + assumes MY: "finite_measure_space MZ (distribution Z)" + shows "conditional_entropy b MX MZ X Z = + - (\(x, z)\space MX \ space MZ. + real (joint_distribution X Z {(x, z)}) * + log b (real (joint_distribution X Z {(x, z)}) / real (distribution Z {z})))" + unfolding conditional_entropy_def using assms + apply (simp add: conditional_mutual_information_generic_eq + setsum_cartesian_product' setsum_commute[of _ "space MZ"] + setsum_negf[symmetric] setsum_subtractf[symmetric]) +proof (safe intro!: setsum_cong, simp) + fix z x assume "z \ space MZ" "x \ space MX" + let "?XXZ x'" = "real (joint_distribution X (\x. (X x, Z x)) {(x, x', z)})" + let "?XZ x'" = "real (joint_distribution X Z {(x', z)})" + let "?X" = "real (distribution X {x})" + interpret MX: finite_measure_space MX "distribution X" by fact + have *: "\A. A = {} \ prob A = 0" by simp + have XXZ: "\x'. ?XXZ x' = (if x' = x then ?XZ x else 0)" + by (auto simp: distribution_def intro!: arg_cong[where f=prob] *) + have "(\x'\space MX. ?XXZ x' * log b (?XXZ x') - (?XXZ x' * log b ?X + ?XXZ x' * log b (?XZ x'))) = + (\x'\{x}. ?XZ x' * log b (?XZ x') - (?XZ x' * log b ?X + ?XZ x' * log b (?XZ x')))" + using `x \ space MX` MX.finite_space + by (safe intro!: setsum_mono_zero_cong_right) + (auto split: split_if_asm simp: XXZ) + then show "(\x'\space MX. ?XXZ x' * log b (?XXZ x') - (?XXZ x' * log b ?X + ?XXZ x' * log b (?XZ x'))) + + ?XZ x * log b ?X = 0" by simp +qed + lemma (in finite_information_space) conditional_entropy_eq: "\(X | Z) = - (\(x, z)\X ` space M \ Z ` space M. real (joint_distribution X Z {(x, z)}) * log b (real (joint_distribution X Z {(x, z)}) / real (distribution Z {z})))" + by (simp add: finite_measure_space conditional_entropy_generic_eq) + +lemma (in finite_information_space) conditional_entropy_eq_ce_with_hypothesis: + "\(X | Y) = + -(\y\Y`space M. real (distribution Y {y}) * + (\x\X`space M. real (joint_distribution X Y {(x,y)}) / real (distribution Y {(y)}) * + log b (real (joint_distribution X Y {(x,y)}) / real (distribution Y {(y)}))))" + unfolding conditional_entropy_eq neg_equal_iff_equal + apply (simp add: setsum_commute[of _ "Y`space M"] setsum_cartesian_product' setsum_divide_distrib[symmetric]) + apply (safe intro!: setsum_cong) + using real_distribution_order'[of Y _ X _] + by (auto simp add: setsum_subtractf[of _ _ "X`space M"]) + +lemma (in finite_information_space) conditional_entropy_eq_cartesian_sum: + "\(X | Y) = -(\x\X`space M. \y\Y`space M. + real (joint_distribution X Y {(x,y)}) * + log b (real (joint_distribution X Y {(x,y)}) / real (distribution Y {y})))" proof - - have *: "\x y z. (\x. (X x, X x, Z x)) -` {(x, y, z)} = (if x = y then (\x. (X x, Z x)) -` {(x, z)} else {})" by auto - show ?thesis - unfolding conditional_mutual_information_eq_sum - conditional_entropy_def distribution_def * - by (auto intro!: setsum_0') + { fix x assume "x\(\x. (X x, Y x))`space M" + then have "(\x. (X x, Y x)) -` {x} \ space M = {}" by auto + then have "joint_distribution X Y {x} = 0" + unfolding distribution_def by auto } + then show ?thesis using finite_space + unfolding conditional_entropy_eq neg_equal_iff_equal setsum_cartesian_product + by (auto intro!: setsum_mono_zero_cong_left) qed +subsection {* Equalities *} + lemma (in finite_information_space) mutual_information_eq_entropy_conditional_entropy: "\(X ; Z) = \(X) - \(X | Z)" unfolding mutual_information_eq entropy_eq conditional_entropy_eq @@ -680,109 +770,15 @@ show ?thesis by auto qed -(* -------------Entropy of a RV with a certain event is zero---------------- *) - -lemma (in finite_information_space) finite_entropy_certainty_eq_0: - assumes "x \ X ` space M" and "distribution X {x} = 1" - shows "\(X) = 0" -proof - - interpret X: finite_prob_space "\ space = X ` space M, sets = Pow (X ` space M) \" "distribution X" - by (rule finite_prob_space_of_images) - - have "distribution X (X ` space M - {x}) = distribution X (X ` space M) - distribution X {x}" - using X.measure_compl[of "{x}"] assms by auto - also have "\ = 0" using X.prob_space assms by auto - finally have X0: "distribution X (X ` space M - {x}) = 0" by auto - - { fix y assume asm: "y \ x" "y \ X ` space M" - hence "{y} \ X ` space M - {x}" by auto - from X.measure_mono[OF this] X0 asm - have "distribution X {y} = 0" by auto } - - hence fi: "\ y. y \ X ` space M \ real (distribution X {y}) = (if x = y then 1 else 0)" - using assms by auto - - have y: "\y. (if x = y then 1 else 0) * log b (if x = y then 1 else 0) = 0" by simp - - show ?thesis unfolding entropy_eq by (auto simp: y fi) -qed -(* --------------- upper bound on entropy for a rv ------------------------- *) - -lemma (in finite_prob_space) distribution_1: - "distribution X A \ 1" - unfolding distribution_def measure_space_1[symmetric] - by (auto intro!: measure_mono simp: sets_eq_Pow) - -lemma (in finite_prob_space) real_distribution_1: - "real (distribution X A) \ 1" - unfolding real_pinfreal_1[symmetric] - by (rule real_of_pinfreal_mono[OF _ distribution_1]) simp +lemma (in finite_information_space) entropy_chain_rule: + "\(\x. (X x, Y x)) = \(X) + \(Y|X)" + unfolding entropy_eq[of X] entropy_eq_cartesian_sum conditional_entropy_eq_cartesian_sum + unfolding setsum_commute[of _ "X`space M"] setsum_negf[symmetric] setsum_addf[symmetric] + by (rule setsum_cong) + (simp_all add: setsum_negf setsum_addf setsum_subtractf setsum_real_distribution + setsum_left_distrib[symmetric] joint_distribution_commute[of X Y]) -lemma (in finite_information_space) finite_entropy_le_card: - "\(X) \ log b (real (card (X ` space M \ {x . distribution X {x} \ 0})))" -proof - - let "?d x" = "distribution X {x}" - let "?p x" = "real (?d x)" - have "\(X) = (\x\X`space M. ?p x * log b (1 / ?p x))" - by (auto intro!: setsum_cong simp: entropy_eq setsum_negf[symmetric]) - also have "\ \ log b (\x\X`space M. ?p x * (1 / ?p x))" - apply (rule log_setsum') - using not_empty b_gt_1 finite_space sum_over_space_real_distribution - by auto - also have "\ = log b (\x\X`space M. if ?d x \ 0 then 1 else 0)" - apply (rule arg_cong[where f="\f. log b (\x\X`space M. f x)"]) - using distribution_finite[of X] by (auto simp: expand_fun_eq real_of_pinfreal_eq_0) - finally show ?thesis - using finite_space by (auto simp: setsum_cases real_eq_of_nat) -qed - -(* --------------- entropy is maximal for a uniform rv --------------------- *) - -lemma (in finite_prob_space) uniform_prob: - assumes "x \ space M" - assumes "\ x y. \x \ space M ; y \ space M\ \ prob {x} = prob {y}" - shows "prob {x} = 1 / real (card (space M))" -proof - - have prob_x: "\ y. y \ space M \ prob {y} = prob {x}" - using assms(2)[OF _ `x \ space M`] by blast - have "1 = prob (space M)" - using prob_space by auto - also have "\ = (\ x \ space M. prob {x})" - using real_finite_measure_finite_Union[of "space M" "\ x. {x}", simplified] - sets_eq_Pow inj_singleton[unfolded inj_on_def, rule_format] - finite_space unfolding disjoint_family_on_def prob_space[symmetric] - by (auto simp add:setsum_restrict_set) - also have "\ = (\ y \ space M. prob {x})" - using prob_x by auto - also have "\ = real_of_nat (card (space M)) * prob {x}" by simp - finally have one: "1 = real (card (space M)) * prob {x}" - using real_eq_of_nat by auto - hence two: "real (card (space M)) \ 0" by fastsimp - from one have three: "prob {x} \ 0" by fastsimp - thus ?thesis using one two three divide_cancel_right - by (auto simp:field_simps) -qed - -lemma (in finite_information_space) finite_entropy_uniform_max: - assumes "\x y. \ x \ X ` space M ; y \ X ` space M \ \ distribution X {x} = distribution X {y}" - shows "\(X) = log b (real (card (X ` space M)))" -proof - - note uniform = - finite_prob_space_of_images[of X, THEN finite_prob_space.uniform_prob, simplified] - - have card_gt0: "0 < card (X ` space M)" unfolding card_gt_0_iff - using finite_space not_empty by auto - - { fix x assume "x \ X ` space M" - hence "real (distribution X {x}) = 1 / real (card (X ` space M))" - proof (rule uniform) - fix x y assume "x \ X`space M" "y \ X`space M" - from assms[OF this] show "real (distribution X {x}) = real (distribution X {y})" by simp - qed } - thus ?thesis - using not_empty finite_space b_gt_1 card_gt0 - by (simp add: entropy_eq real_eq_of_nat[symmetric] log_divide) -qed +section {* Partitioning *} definition "subvimage A f g \ (\x \ A. f -` {f x} \ A \ g -` {g x} \ A)" @@ -934,38 +930,6 @@ "\(f \ X) \ \(X)" by (subst (2) entropy_partition[of _ "f \ X"]) (auto intro: conditional_entropy_positive) -lemma (in prob_space) distribution_cong: - assumes "\x. x \ space M \ X x = Y x" - shows "distribution X = distribution Y" - unfolding distribution_def expand_fun_eq - using assms by (auto intro!: arg_cong[where f="\"]) - -lemma (in prob_space) joint_distribution_cong: - assumes "\x. x \ space M \ X x = X' x" - assumes "\x. x \ space M \ Y x = Y' x" - shows "joint_distribution X Y = joint_distribution X' Y'" - unfolding distribution_def expand_fun_eq - using assms by (auto intro!: arg_cong[where f="\"]) - -lemma image_cong: - "\ \x. x \ S \ X x = X' x \ \ X ` S = X' ` S" - by (auto intro!: image_eqI) - -lemma (in finite_information_space) mutual_information_cong: - assumes X: "\x. x \ space M \ X x = X' x" - assumes Y: "\x. x \ space M \ Y x = Y' x" - shows "\(X ; Y) = \(X' ; Y')" -proof - - have "X ` space M = X' ` space M" using X by (rule image_cong) - moreover have "Y ` space M = Y' ` space M" using Y by (rule image_cong) - ultimately show ?thesis - unfolding mutual_information_eq - using - assms[THEN distribution_cong] - joint_distribution_cong[OF assms] - by (auto intro!: setsum_cong) -qed - corollary (in finite_information_space) entropy_of_inj: assumes "inj_on f (X`space M)" shows "\(f \ X) = \(X)" diff -r bddc3d3f6e53 -r e9467adb8b52 src/HOL/Probability/Lebesgue_Integration.thy --- a/src/HOL/Probability/Lebesgue_Integration.thy Thu Sep 02 18:45:23 2010 +0200 +++ b/src/HOL/Probability/Lebesgue_Integration.thy Thu Sep 02 20:44:33 2010 +0200 @@ -209,19 +209,6 @@ by (auto intro!: **) qed -lemma setsum_indicator_disjoint_family: - fixes f :: "'d \ 'e::semiring_1" - assumes d: "disjoint_family_on A P" and "x \ A j" and "finite P" and "j \ P" - shows "(\i\P. f i * indicator (A i) x) = f j" -proof - - have "P \ {i. x \ A i} = {j}" - using d `x \ A j` `j \ P` unfolding disjoint_family_on_def - by auto - thus ?thesis - unfolding indicator_def - by (simp add: if_distrib setsum_cases[OF `finite P`]) -qed - lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence: fixes u :: "'a \ pinfreal" assumes u: "u \ borel_measurable M" @@ -426,6 +413,62 @@ with x show thesis by (auto intro!: that[of f]) qed +lemma (in sigma_algebra) simple_function_eq_borel_measurable: + fixes f :: "'a \ pinfreal" + shows "simple_function f \ + finite (f`space M) \ f \ borel_measurable M" + using simple_function_borel_measurable[of f] + borel_measurable_simple_function[of f] + by (fastsimp simp: simple_function_def) + +lemma (in measure_space) simple_function_restricted: + fixes f :: "'a \ pinfreal" assumes "A \ sets M" + shows "sigma_algebra.simple_function (restricted_space A) f \ simple_function (\x. f x * indicator A x)" + (is "sigma_algebra.simple_function ?R f \ simple_function ?f") +proof - + interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \ sets M`]) + have "finite (f`A) \ finite (?f`space M)" + proof cases + assume "A = space M" + then have "f`A = ?f`space M" by (fastsimp simp: image_iff) + then show ?thesis by simp + next + assume "A \ space M" + then obtain x where x: "x \ space M" "x \ A" + using sets_into_space `A \ sets M` by auto + have *: "?f`space M = f`A \ {0}" + proof (auto simp add: image_iff) + show "\x\space M. f x = 0 \ indicator A x = 0" + using x by (auto intro!: bexI[of _ x]) + next + fix x assume "x \ A" + then show "\y\space M. f x = f y * indicator A y" + using `A \ sets M` sets_into_space by (auto intro!: bexI[of _ x]) + next + fix x + assume "indicator A x \ (0::pinfreal)" + then have "x \ A" by (auto simp: indicator_def split: split_if_asm) + moreover assume "x \ space M" "\y\A. ?f x \ f y" + ultimately show "f x = 0" by auto + qed + then show ?thesis by auto + qed + then show ?thesis + unfolding simple_function_eq_borel_measurable + R.simple_function_eq_borel_measurable + unfolding borel_measurable_restricted[OF `A \ sets M`] + by auto +qed + +lemma (in sigma_algebra) simple_function_subalgebra: + assumes "sigma_algebra.simple_function (M\sets:=N\) f" + and N_subalgebra: "N \ sets M" "sigma_algebra (M\sets:=N\)" + shows "simple_function f" + using assms + unfolding simple_function_def + unfolding sigma_algebra.simple_function_def[OF N_subalgebra(2)] + by auto + section "Simple integral" definition (in measure_space) @@ -668,6 +711,41 @@ qed qed +lemma (in measure_space) simple_integral_restricted: + assumes "A \ sets M" + assumes sf: "simple_function (\x. f x * indicator A x)" + shows "measure_space.simple_integral (restricted_space A) \ f = simple_integral (\x. f x * indicator A x)" + (is "_ = simple_integral ?f") + unfolding measure_space.simple_integral_def[OF restricted_measure_space[OF `A \ sets M`]] + unfolding simple_integral_def +proof (simp, safe intro!: setsum_mono_zero_cong_left) + from sf show "finite (?f ` space M)" + unfolding simple_function_def by auto +next + fix x assume "x \ A" + then show "f x \ ?f ` space M" + using sets_into_space `A \ sets M` by (auto intro!: image_eqI[of _ _ x]) +next + fix x assume "x \ space M" "?f x \ f`A" + then have "x \ A" by (auto simp: image_iff) + then show "?f x * \ (?f -` {?f x} \ space M) = 0" by simp +next + fix x assume "x \ A" + then have "f x \ 0 \ + f -` {f x} \ A = ?f -` {f x} \ space M" + using `A \ sets M` sets_into_space + by (auto simp: indicator_def split: split_if_asm) + then show "f x * \ (f -` {f x} \ A) = + f x * \ (?f -` {f x} \ space M)" + unfolding pinfreal_mult_cancel_left by auto +qed + +lemma (in measure_space) simple_integral_subalgebra[simp]: + assumes "measure_space (M\sets := N\) \" + shows "measure_space.simple_integral (M\sets := N\) \ = simple_integral" + unfolding simple_integral_def_raw + unfolding measure_space.simple_integral_def_raw[OF assms] by simp + section "Continuous posititve integration" definition (in measure_space) @@ -1077,6 +1155,43 @@ qed qed +lemma (in measure_space) positive_integral_translated_density: + assumes "f \ borel_measurable M" "g \ borel_measurable M" + shows "measure_space.positive_integral M (\A. positive_integral (\x. f x * indicator A x)) g = + positive_integral (\x. f x * g x)" (is "measure_space.positive_integral M ?T _ = _") +proof - + from measure_space_density[OF assms(1)] + interpret T: measure_space M ?T . + from borel_measurable_implies_simple_function_sequence[OF assms(2)] + obtain G where G: "\i. simple_function (G i)" "G \ g" by blast + note G_borel = borel_measurable_simple_function[OF this(1)] + from T.positive_integral_isoton[OF `G \ g` G_borel] + have *: "(\i. T.positive_integral (G i)) \ T.positive_integral g" . + { fix i + have [simp]: "finite (G i ` space M)" + using G(1) unfolding simple_function_def by auto + have "T.positive_integral (G i) = T.simple_integral (G i)" + using G T.positive_integral_eq_simple_integral by simp + also have "\ = positive_integral (\x. f x * (\y\G i`space M. y * indicator (G i -` {y} \ space M) x))" + apply (simp add: T.simple_integral_def) + apply (subst positive_integral_cmult[symmetric]) + using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage) + apply (subst positive_integral_setsum[symmetric]) + using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage) + by (simp add: setsum_right_distrib field_simps) + also have "\ = positive_integral (\x. f x * G i x)" + by (auto intro!: positive_integral_cong + simp: indicator_def if_distrib setsum_cases) + finally have "T.positive_integral (G i) = positive_integral (\x. f x * G i x)" . } + with * have eq_Tg: "(\i. positive_integral (\x. f x * G i x)) \ T.positive_integral g" by simp + from G(2) have "(\i x. f x * G i x) \ (\x. f x * g x)" + unfolding isoton_fun_expand by (auto intro!: isoton_cmult_right) + then have "(\i. positive_integral (\x. f x * G i x)) \ positive_integral (\x. f x * g x)" + using assms(1) G_borel by (auto intro!: positive_integral_isoton borel_measurable_pinfreal_times) + with eq_Tg show "T.positive_integral g = positive_integral (\x. f x * g x)" + unfolding isoton_def by simp +qed + lemma (in measure_space) positive_integral_null_set: assumes borel: "u \ borel_measurable M" and "N \ null_sets" shows "positive_integral (\x. u x * indicator N x) = 0" (is "?I = 0") @@ -1222,6 +1337,58 @@ finally show ?thesis by simp qed +lemma (in measure_space) positive_integral_restricted: + assumes "A \ sets M" + shows "measure_space.positive_integral (restricted_space A) \ f = positive_integral (\x. f x * indicator A x)" + (is "measure_space.positive_integral ?R \ f = positive_integral ?f") +proof - + have msR: "measure_space ?R \" by (rule restricted_measure_space[OF `A \ sets M`]) + then interpret R: measure_space ?R \ . + have saR: "sigma_algebra ?R" by fact + have *: "R.positive_integral f = R.positive_integral ?f" + by (auto intro!: R.positive_integral_cong) + show ?thesis + unfolding * R.positive_integral_def positive_integral_def + unfolding simple_function_restricted[OF `A \ sets M`] + apply (simp add: SUPR_def) + apply (rule arg_cong[where f=Sup]) + proof (auto simp: image_iff simple_integral_restricted[OF `A \ sets M`]) + fix g assume "simple_function (\x. g x * indicator A x)" + "g \ f" "\x\A. \ \ g x" + then show "\x. simple_function x \ x \ (\x. f x * indicator A x) \ (\y\space M. \ \ x y) \ + simple_integral (\x. g x * indicator A x) = simple_integral x" + apply (rule_tac exI[of _ "\x. g x * indicator A x"]) + by (auto simp: indicator_def le_fun_def) + next + fix g assume g: "simple_function g" "g \ (\x. f x * indicator A x)" + "\x\space M. \ \ g x" + then have *: "(\x. g x * indicator A x) = g" + "\x. g x * indicator A x = g x" + "\x. g x \ f x" + by (auto simp: le_fun_def expand_fun_eq indicator_def split: split_if_asm) + from g show "\x. simple_function (\xa. x xa * indicator A xa) \ x \ f \ (\xa\A. \ \ x xa) \ + simple_integral g = simple_integral (\xa. x xa * indicator A xa)" + using `A \ sets M`[THEN sets_into_space] + apply (rule_tac exI[of _ "\x. g x * indicator A x"]) + by (fastsimp simp: le_fun_def *) + qed +qed + +lemma (in measure_space) positive_integral_subalgebra[simp]: + assumes borel: "f \ borel_measurable (M\sets := N\)" + and N_subalgebra: "N \ sets M" "sigma_algebra (M\sets := N\)" + shows "measure_space.positive_integral (M\sets := N\) \ f = positive_integral f" +proof - + note msN = measure_space_subalgebra[OF N_subalgebra] + then interpret N: measure_space "M\sets:=N\" \ . + from N.borel_measurable_implies_simple_function_sequence[OF borel] + obtain fs where Nsf: "\i. N.simple_function (fs i)" and "fs \ f" by blast + then have sf: "\i. simple_function (fs i)" + using simple_function_subalgebra[OF _ N_subalgebra] by blast + from positive_integral_isoton_simple[OF `fs \ f` sf] N.positive_integral_isoton_simple[OF `fs \ f` Nsf] + show ?thesis unfolding simple_integral_subalgebra[OF msN] isoton_def by simp +qed + section "Lebesgue Integral" definition (in measure_space) integrable where @@ -1629,44 +1796,6 @@ by (simp add: real_of_pinfreal_eq_0) qed -lemma LIMSEQ_max: - "u ----> (x::real) \ (\i. max (u i) 0) ----> max x 0" - by (fastsimp intro!: LIMSEQ_I dest!: LIMSEQ_D) - -lemma (in sigma_algebra) borel_measurable_LIMSEQ: - fixes u :: "nat \ 'a \ real" - assumes u': "\x. x \ space M \ (\i. u i x) ----> u' x" - and u: "\i. u i \ borel_measurable M" - shows "u' \ borel_measurable M" -proof - - let "?pu x i" = "max (u i x) 0" - let "?nu x i" = "max (- u i x) 0" - - { fix x assume x: "x \ space M" - have "(?pu x) ----> max (u' x) 0" - "(?nu x) ----> max (- u' x) 0" - using u'[OF x] by (auto intro!: LIMSEQ_max LIMSEQ_minus) - from LIMSEQ_imp_lim_INF[OF _ this(1)] LIMSEQ_imp_lim_INF[OF _ this(2)] - have "(SUP n. INF m. Real (u (n + m) x)) = Real (u' x)" - "(SUP n. INF m. Real (- u (n + m) x)) = Real (- u' x)" - by (simp_all add: Real_max'[symmetric]) } - note eq = this - - have *: "\x. real (Real (u' x)) - real (Real (- u' x)) = u' x" - by auto - - have "(SUP n. INF m. (\x. Real (u (n + m) x))) \ borel_measurable M" - "(SUP n. INF m. (\x. Real (- u (n + m) x))) \ borel_measurable M" - using u by (auto intro: borel_measurable_SUP borel_measurable_INF borel_measurable_Real) - with eq[THEN measurable_cong, of M "\x. x" borel_space] - have "(\x. Real (u' x)) \ borel_measurable M" - "(\x. Real (- u' x)) \ borel_measurable M" - unfolding SUPR_fun_expand INFI_fun_expand by auto - note this[THEN borel_measurable_real] - from borel_measurable_diff[OF this] - show ?thesis unfolding * . -qed - lemma (in measure_space) integral_dominated_convergence: assumes u: "\i. integrable (u i)" and bound: "\x j. x\space M \ \u j x\ \ w x" and w: "integrable w" "\x. x \ space M \ 0 \ w x" @@ -1926,41 +2055,11 @@ by (simp_all add: integral_cmul_indicator borel_measurable_vimage) qed -lemma sigma_algebra_cong: - fixes M :: "('a, 'b) algebra_scheme" and M' :: "('a, 'c) algebra_scheme" - assumes *: "sigma_algebra M" - and cong: "space M = space M'" "sets M = sets M'" - shows "sigma_algebra M'" -using * unfolding sigma_algebra_def algebra_def sigma_algebra_axioms_def unfolding cong . - -lemma finite_Pow_additivity_sufficient: - assumes "finite (space M)" and "sets M = Pow (space M)" - and "positive \" and "additive M \" - and "\x. x \ space M \ \ {x} \ \" - shows "finite_measure_space M \" -proof - - have "sigma_algebra M" - using assms by (auto intro!: sigma_algebra_cong[OF sigma_algebra_Pow]) - - have "measure_space M \" - by (rule sigma_algebra.finite_additivity_sufficient) (fact+) - thus ?thesis - unfolding finite_measure_space_def finite_measure_space_axioms_def - using assms by simp -qed - -lemma finite_measure_spaceI: - assumes "measure_space M \" and "finite (space M)" and "sets M = Pow (space M)" - and "\x. x \ space M \ \ {x} \ \" - shows "finite_measure_space M \" - unfolding finite_measure_space_def finite_measure_space_axioms_def - using assms by simp +lemma (in finite_measure_space) simple_function_finite[simp, intro]: "simple_function f" + unfolding simple_function_def sets_eq_Pow using finite_space by auto lemma (in finite_measure_space) borel_measurable_finite[intro, simp]: "f \ borel_measurable M" - unfolding measurable_def sets_eq_Pow by auto - -lemma (in finite_measure_space) simple_function_finite: "simple_function f" - unfolding simple_function_def sets_eq_Pow using finite_space by auto + by (auto intro: borel_measurable_simple_function) lemma (in finite_measure_space) positive_integral_finite_eq_setsum: "positive_integral f = (\x \ space M. f x * \ {x})" @@ -1979,10 +2078,8 @@ "positive_integral (\x. Real (f x)) = (\x \ space M. Real (f x) * \ {x})" "positive_integral (\x. Real (- f x)) = (\x \ space M. Real (- f x) * \ {x})" unfolding positive_integral_finite_eq_setsum by auto - show "integrable f" using finite_space finite_measure by (simp add: setsum_\ integrable_def sets_eq_Pow) - show ?I using finite_measure apply (simp add: integral_def sets_eq_Pow real_of_pinfreal_setsum[symmetric] real_of_pinfreal_mult[symmetric] setsum_subtractf[symmetric]) diff -r bddc3d3f6e53 -r e9467adb8b52 src/HOL/Probability/Measure.thy --- a/src/HOL/Probability/Measure.thy Thu Sep 02 18:45:23 2010 +0200 +++ b/src/HOL/Probability/Measure.thy Thu Sep 02 20:44:33 2010 +0200 @@ -414,6 +414,19 @@ finally show ?thesis . qed +lemma (in measure_space) measure_finitely_subadditive: + assumes "finite I" "A ` I \ sets M" + shows "\ (\i\I. A i) \ (\i\I. \ (A i))" +using assms proof induct + case (insert i I) + then have "(\i\I. A i) \ sets M" by (auto intro: finite_UN) + then have "\ (\i\insert i I. A i) \ \ (A i) + \ (\i\I. A i)" + using insert by (simp add: measure_subadditive) + also have "\ \ (\i\insert i I. \ (A i))" + using insert by (auto intro!: add_left_mono) + finally show ?case . +qed simp + lemma (in measure_space) measurable_countably_subadditive: assumes "range f \ sets M" shows "\ (\i. f i) \ (\\<^isub>\ i. \ (f i))" @@ -432,9 +445,34 @@ finally show ?thesis . qed +lemma (in measure_space) measure_inter_full_set: + assumes "S \ sets M" "T \ sets M" and not_\: "\ (T - S) \ \" + assumes T: "\ T = \ (space M)" + shows "\ (S \ T) = \ S" +proof (rule antisym) + show " \ (S \ T) \ \ S" + using assms by (auto intro!: measure_mono) + + show "\ S \ \ (S \ T)" + proof (rule ccontr) + assume contr: "\ ?thesis" + have "\ (space M) = \ ((T - S) \ (S \ T))" + unfolding T[symmetric] by (auto intro!: arg_cong[where f="\"]) + also have "\ \ \ (T - S) + \ (S \ T)" + using assms by (auto intro!: measure_subadditive) + also have "\ < \ (T - S) + \ S" + by (rule pinfreal_less_add[OF not_\]) (insert contr, auto) + also have "\ = \ (T \ S)" + using assms by (subst measure_additive) auto + also have "\ \ \ (space M)" + using assms sets_into_space by (auto intro!: measure_mono) + finally show False .. + qed +qed + lemma (in measure_space) restricted_measure_space: assumes "S \ sets M" - shows "measure_space (M\ space := S, sets := (\A. S \ A) ` sets M \) \" + shows "measure_space (restricted_space S) \" (is "measure_space ?r \") unfolding measure_space_def measure_space_axioms_def proof safe @@ -451,6 +489,46 @@ qed qed +lemma (in measure_space) measure_space_vimage: + assumes "f \ measurable M M'" + and "sigma_algebra M'" + shows "measure_space M' (\A. \ (f -` A \ space M))" (is "measure_space M' ?T") +proof - + interpret M': sigma_algebra M' by fact + + show ?thesis + proof + show "?T {} = 0" by simp + + show "countably_additive M' ?T" + proof (unfold countably_additive_def, safe) + fix A :: "nat \ 'c set" assume "range A \ sets M'" "disjoint_family A" + hence *: "\i. f -` (A i) \ space M \ sets M" + using `f \ measurable M M'` by (auto simp: measurable_def) + moreover have "(\i. f -` A i \ space M) \ sets M" + using * by blast + moreover have **: "disjoint_family (\i. f -` A i \ space M)" + using `disjoint_family A` by (auto simp: disjoint_family_on_def) + ultimately show "(\\<^isub>\ i. ?T (A i)) = ?T (\i. A i)" + using measure_countably_additive[OF _ **] by (auto simp: comp_def vimage_UN) + qed + qed +qed + +lemma (in measure_space) measure_space_subalgebra: + assumes "N \ sets M" "sigma_algebra (M\ sets := N \)" + shows "measure_space (M\ sets := N \) \" +proof - + interpret N: sigma_algebra "M\ sets := N \" by fact + show ?thesis + proof + show "countably_additive (M\sets := N\) \" + using `N \ sets M` + by (auto simp add: countably_additive_def + intro!: measure_countably_additive) + qed simp +qed + section "@{text \}-finite Measures" locale sigma_finite_measure = measure_space + @@ -458,7 +536,7 @@ lemma (in sigma_finite_measure) restricted_sigma_finite_measure: assumes "S \ sets M" - shows "sigma_finite_measure (M\ space := S, sets := (\A. S \ A) ` sets M \) \" + shows "sigma_finite_measure (restricted_space S) \" (is "sigma_finite_measure ?r _") unfolding sigma_finite_measure_def sigma_finite_measure_axioms_def proof safe @@ -486,6 +564,25 @@ qed qed +lemma (in sigma_finite_measure) disjoint_sigma_finite: + "\A::nat\'a set. range A \ sets M \ (\i. A i) = space M \ + (\i. \ (A i) \ \) \ disjoint_family A" +proof - + obtain A :: "nat \ 'a set" where + range: "range A \ sets M" and + space: "(\i. A i) = space M" and + measure: "\i. \ (A i) \ \" + using sigma_finite by auto + note range' = range_disjointed_sets[OF range] range + { fix i + have "\ (disjointed A i) \ \ (A i)" + using range' disjointed_subset[of A i] by (auto intro!: measure_mono) + then have "\ (disjointed A i) \ \" + using measure[of i] by auto } + with disjoint_family_disjointed UN_disjointed_eq[of A] space range' + show ?thesis by (auto intro!: exI[of _ "disjointed A"]) +qed + section "Real measure values" lemma (in measure_space) real_measure_Union: @@ -604,7 +701,7 @@ using finite_measure_of_space by (auto intro!: exI[of _ "\x. space M"]) qed -lemma (in finite_measure) finite_measure: +lemma (in finite_measure) finite_measure[simp, intro]: assumes "A \ sets M" shows "\ A \ \" proof - @@ -619,7 +716,7 @@ lemma (in finite_measure) restricted_finite_measure: assumes "S \ sets M" - shows "finite_measure (M\ space := S, sets := (\A. S \ A) ` sets M \) \" + shows "finite_measure (restricted_space S) \" (is "finite_measure ?r _") unfolding finite_measure_def finite_measure_axioms_def proof (safe del: notI) @@ -707,6 +804,13 @@ shows "\ (space M - s) = \ (space M) - \ s" using measure_compl[OF s, OF finite_measure, OF s] . +lemma (in finite_measure) finite_measure_inter_full_set: + assumes "S \ sets M" "T \ sets M" + assumes T: "\ T = \ (space M)" + shows "\ (S \ T) = \ S" + using measure_inter_full_set[OF assms(1,2) finite_measure assms(3)] assms + by auto + section {* Measure preserving *} definition "measure_preserving A \ B \ = @@ -817,10 +921,51 @@ and sets_eq_Pow: "sets M = Pow (space M)" and finite_single_measure: "\x. x \ space M \ \ {x} \ \" +lemma (in finite_measure_space) sets_image_space_eq_Pow: + "sets (image_space X) = Pow (space (image_space X))" +proof safe + fix x S assume "S \ sets (image_space X)" "x \ S" + then show "x \ space (image_space X)" + using sets_into_space by (auto intro!: imageI simp: image_space_def) +next + fix S assume "S \ space (image_space X)" + then obtain S' where "S = X`S'" "S'\sets M" + by (auto simp: subset_image_iff sets_eq_Pow image_space_def) + then show "S \ sets (image_space X)" + by (auto simp: image_space_def) +qed + lemma (in finite_measure_space) sum_over_space: "(\x\space M. \ {x}) = \ (space M)" using measure_finitely_additive''[of "space M" "\i. {i}"] by (simp add: sets_eq_Pow disjoint_family_on_def finite_space) +lemma finite_measure_spaceI: + assumes "finite (space M)" "sets M = Pow(space M)" and space: "\ (space M) \ \" + and add: "\A B. A\space M \ B\space M \ A \ B = {} \ \ (A \ B) = \ A + \ B" + and "\ {} = 0" + shows "finite_measure_space M \" + unfolding finite_measure_space_def finite_measure_space_axioms_def +proof (safe del: notI) + show "measure_space M \" + proof (rule sigma_algebra.finite_additivity_sufficient) + show "sigma_algebra M" + apply (rule sigma_algebra_cong) + apply (rule sigma_algebra_Pow[of "space M"]) + using assms by simp_all + show "finite (space M)" by fact + show "positive \" unfolding positive_def by fact + show "additive M \" unfolding additive_def using assms by simp + qed + show "finite (space M)" by fact + { fix A x assume "A \ sets M" "x \ A" then show "x \ space M" + using assms by auto } + { fix A assume "A \ space M" then show "A \ sets M" + using assms by auto } + { fix x assume *: "x \ space M" + with add[of "{x}" "space M - {x}"] space + show "\ {x} \ \" by (auto simp: insert_absorb[OF *] Diff_subset) } +qed + sublocale finite_measure_space < finite_measure proof show "\ (space M) \ \" @@ -828,6 +973,22 @@ using finite_space finite_single_measure by auto qed +lemma finite_measure_space_iff: + "finite_measure_space M \ \ + finite (space M) \ sets M = Pow(space M) \ \ (space M) \ \ \ \ {} = 0 \ + (\A\space M. \B\space M. A \ B = {} \ \ (A \ B) = \ A + \ B)" + (is "_ = ?rhs") +proof (intro iffI) + assume "finite_measure_space M \" + then interpret finite_measure_space M \ . + show ?rhs + using finite_space sets_eq_Pow measure_additive empty_measure finite_measure + by auto +next + assume ?rhs then show "finite_measure_space M \" + by (auto intro!: finite_measure_spaceI) +qed + lemma psuminf_cmult_indicator: assumes "disjoint_family A" "x \ A i" shows "(\\<^isub>\ n. f n * indicator (A n) x) = f i" diff -r bddc3d3f6e53 -r e9467adb8b52 src/HOL/Probability/Positive_Infinite_Real.thy --- a/src/HOL/Probability/Positive_Infinite_Real.thy Thu Sep 02 18:45:23 2010 +0200 +++ b/src/HOL/Probability/Positive_Infinite_Real.thy Thu Sep 02 20:44:33 2010 +0200 @@ -411,6 +411,10 @@ lemma pinfreal_less_\: "x < \ \ x \ \" by (cases x) auto +lemma pinfreal_0_less_mult_iff[simp]: + fixes x y :: pinfreal shows "0 < x * y \ 0 < x \ 0 < y" + by (cases x, cases y) (auto simp: zero_less_mult_iff) + subsection {* @{text "x - y"} on @{typ pinfreal} *} instantiation pinfreal :: minus diff -r bddc3d3f6e53 -r e9467adb8b52 src/HOL/Probability/Probability_Space.thy --- a/src/HOL/Probability/Probability_Space.thy Thu Sep 02 18:45:23 2010 +0200 +++ b/src/HOL/Probability/Probability_Space.thy Thu Sep 02 20:44:33 2010 +0200 @@ -1,39 +1,7 @@ theory Probability_Space -imports Lebesgue_Integration +imports Lebesgue_Integration Radon_Nikodym begin -lemma (in measure_space) measure_inter_full_set: - assumes "S \ sets M" "T \ sets M" and not_\: "\ (T - S) \ \" - assumes T: "\ T = \ (space M)" - shows "\ (S \ T) = \ S" -proof (rule antisym) - show " \ (S \ T) \ \ S" - using assms by (auto intro!: measure_mono) - - show "\ S \ \ (S \ T)" - proof (rule ccontr) - assume contr: "\ ?thesis" - have "\ (space M) = \ ((T - S) \ (S \ T))" - unfolding T[symmetric] by (auto intro!: arg_cong[where f="\"]) - also have "\ \ \ (T - S) + \ (S \ T)" - using assms by (auto intro!: measure_subadditive) - also have "\ < \ (T - S) + \ S" - by (rule pinfreal_less_add[OF not_\]) (insert contr, auto) - also have "\ = \ (T \ S)" - using assms by (subst measure_additive) auto - also have "\ \ \ (space M)" - using assms sets_into_space by (auto intro!: measure_mono) - finally show False .. - qed -qed - -lemma (in finite_measure) finite_measure_inter_full_set: - assumes "S \ sets M" "T \ sets M" - assumes T: "\ T = \ (space M)" - shows "\ (S \ T) = \ S" - using measure_inter_full_set[OF assms(1,2) finite_measure assms(3)] assms - by auto - locale prob_space = measure_space + assumes measure_space_1: "\ (space M) = 1" @@ -63,6 +31,19 @@ abbreviation "joint_distribution X Y \ distribution (\x. (X x, Y x))" +lemma (in prob_space) distribution_cong: + assumes "\x. x \ space M \ X x = Y x" + shows "distribution X = distribution Y" + unfolding distribution_def expand_fun_eq + using assms by (auto intro!: arg_cong[where f="\"]) + +lemma (in prob_space) joint_distribution_cong: + assumes "\x. x \ space M \ X x = X' x" + assumes "\x. x \ space M \ Y x = Y' x" + shows "joint_distribution X Y = joint_distribution X' Y'" + unfolding distribution_def expand_fun_eq + using assms by (auto intro!: arg_cong[where f="\"]) + lemma prob_space: "prob (space M) = 1" unfolding measure_space_1 by simp @@ -75,10 +56,6 @@ finally show ?thesis . qed -lemma measure_finite[simp, intro]: - assumes "A \ events" shows "\ A \ \" - using measure_le_1[OF assms] by auto - lemma prob_compl: assumes "A \ events" shows "prob (space M - A) = 1 - prob A" @@ -197,35 +174,17 @@ qed lemma distribution_prob_space: - fixes S :: "('c, 'd) algebra_scheme" - assumes "sigma_algebra S" "random_variable S X" + assumes S: "sigma_algebra S" "random_variable S X" shows "prob_space S (distribution X)" proof - - interpret S: sigma_algebra S by fact + interpret S: measure_space S "distribution X" + using measure_space_vimage[OF S(2,1)] unfolding distribution_def . show ?thesis proof - show "distribution X {} = 0" unfolding distribution_def by simp have "X -` space S \ space M = space M" using `random_variable S X` by (auto simp: measurable_def) - then show "distribution X (space S) = 1" using measure_space_1 by (simp add: distribution_def) - - show "countably_additive S (distribution X)" - proof (unfold countably_additive_def, safe) - fix A :: "nat \ 'c set" assume "range A \ sets S" "disjoint_family A" - hence *: "\i. X -` A i \ space M \ sets M" - using `random_variable S X` by (auto simp: measurable_def) - moreover hence "\i. \ (X -` A i \ space M) \ \" - using finite_measure by auto - moreover have "(\i. X -` A i \ space M) \ sets M" - using * by blast - moreover hence "\ (\i. X -` A i \ space M) \ \" - using finite_measure by auto - moreover have **: "disjoint_family (\i. X -` A i \ space M)" - using `disjoint_family A` by (auto simp: disjoint_family_on_def) - ultimately show "(\\<^isub>\ i. distribution X (A i)) = distribution X (\i. A i)" - using measure_countably_additive[OF _ **] - by (auto simp: distribution_def Real_real comp_def vimage_UN) - qed + then show "distribution X (space S) = 1" + using measure_space_1 by (simp add: distribution_def) qed qed @@ -379,39 +338,246 @@ joint_distribution_restriction_snd[of X Y "{(x, y)}"] by auto +lemma (in finite_prob_space) distribution_mono: + assumes "\t. \ t \ space M ; X t \ x \ \ Y t \ y" + shows "distribution X x \ distribution Y y" + unfolding distribution_def + using assms by (auto simp: sets_eq_Pow intro!: measure_mono) + +lemma (in finite_prob_space) distribution_mono_gt_0: + assumes gt_0: "0 < distribution X x" + assumes *: "\t. \ t \ space M ; X t \ x \ \ Y t \ y" + shows "0 < distribution Y y" + by (rule less_le_trans[OF gt_0 distribution_mono]) (rule *) + +lemma (in finite_prob_space) sum_over_space_distrib: + "(\x\X`space M. distribution X {x}) = 1" + unfolding distribution_def measure_space_1[symmetric] using finite_space + by (subst measure_finitely_additive'') + (auto simp add: disjoint_family_on_def sets_eq_Pow intro!: arg_cong[where f=\]) + +lemma (in finite_prob_space) sum_over_space_real_distribution: + "(\x\X`space M. real (distribution X {x})) = 1" + unfolding distribution_def prob_space[symmetric] using finite_space + by (subst real_finite_measure_finite_Union[symmetric]) + (auto simp add: disjoint_family_on_def sets_eq_Pow intro!: arg_cong[where f=prob]) + +lemma (in finite_prob_space) finite_sum_over_space_eq_1: + "(\x\space M. real (\ {x})) = 1" + using sum_over_space_eq_1 finite_measure by (simp add: real_of_pinfreal_setsum sets_eq_Pow) + +lemma (in finite_prob_space) distribution_finite: + "distribution X A \ \" + using finite_measure[of "X -` A \ space M"] + unfolding distribution_def sets_eq_Pow by auto + +lemma (in finite_prob_space) real_distribution_gt_0[simp]: + "0 < real (distribution Y y) \ 0 < distribution Y y" + using assms by (auto intro!: real_pinfreal_pos distribution_finite) + +lemma (in finite_prob_space) real_distribution_mult_pos_pos: + assumes "0 < distribution Y y" + and "0 < distribution X x" + shows "0 < real (distribution Y y * distribution X x)" + unfolding real_of_pinfreal_mult[symmetric] + using assms by (auto intro!: mult_pos_pos) + +lemma (in finite_prob_space) real_distribution_divide_pos_pos: + assumes "0 < distribution Y y" + and "0 < distribution X x" + shows "0 < real (distribution Y y / distribution X x)" + unfolding divide_pinfreal_def real_of_pinfreal_mult[symmetric] + using assms distribution_finite[of X x] by (cases "distribution X x") (auto intro!: mult_pos_pos) + +lemma (in finite_prob_space) real_distribution_mult_inverse_pos_pos: + assumes "0 < distribution Y y" + and "0 < distribution X x" + shows "0 < real (distribution Y y * inverse (distribution X x))" + unfolding divide_pinfreal_def real_of_pinfreal_mult[symmetric] + using assms distribution_finite[of X x] by (cases "distribution X x") (auto intro!: mult_pos_pos) + +lemma (in prob_space) distribution_remove_const: + shows "joint_distribution X (\x. ()) {(x, ())} = distribution X {x}" + and "joint_distribution (\x. ()) X {((), x)} = distribution X {x}" + and "joint_distribution X (\x. (Y x, ())) {(x, y, ())} = joint_distribution X Y {(x, y)}" + and "joint_distribution X (\x. ((), Y x)) {(x, (), y)} = joint_distribution X Y {(x, y)}" + and "distribution (\x. ()) {()} = 1" + unfolding measure_space_1[symmetric] + by (auto intro!: arg_cong[where f="\"] simp: distribution_def) + +lemma (in finite_prob_space) setsum_distribution_gen: + assumes "Z -` {c} \ space M = (\x \ X`space M. Y -` {f x}) \ space M" + and "inj_on f (X`space M)" + shows "(\x \ X`space M. distribution Y {f x}) = distribution Z {c}" + unfolding distribution_def assms + using finite_space assms + by (subst measure_finitely_additive'') + (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def + intro!: arg_cong[where f=prob]) + +lemma (in finite_prob_space) setsum_distribution: + "(\x \ X`space M. joint_distribution X Y {(x, y)}) = distribution Y {y}" + "(\y \ Y`space M. joint_distribution X Y {(x, y)}) = distribution X {x}" + "(\x \ X`space M. joint_distribution X (\x. (Y x, Z x)) {(x, y, z)}) = joint_distribution Y Z {(y, z)}" + "(\y \ Y`space M. joint_distribution X (\x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Z {(x, z)}" + "(\z \ Z`space M. joint_distribution X (\x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Y {(x, y)}" + by (auto intro!: inj_onI setsum_distribution_gen) + +lemma (in finite_prob_space) setsum_real_distribution_gen: + assumes "Z -` {c} \ space M = (\x \ X`space M. Y -` {f x}) \ space M" + and "inj_on f (X`space M)" + shows "(\x \ X`space M. real (distribution Y {f x})) = real (distribution Z {c})" + unfolding distribution_def assms + using finite_space assms + by (subst real_finite_measure_finite_Union[symmetric]) + (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def + intro!: arg_cong[where f=prob]) + +lemma (in finite_prob_space) setsum_real_distribution: + "(\x \ X`space M. real (joint_distribution X Y {(x, y)})) = real (distribution Y {y})" + "(\y \ Y`space M. real (joint_distribution X Y {(x, y)})) = real (distribution X {x})" + "(\x \ X`space M. real (joint_distribution X (\x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution Y Z {(y, z)})" + "(\y \ Y`space M. real (joint_distribution X (\x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution X Z {(x, z)})" + "(\z \ Z`space M. real (joint_distribution X (\x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution X Y {(x, y)})" + by (auto intro!: inj_onI setsum_real_distribution_gen) + +lemma (in finite_prob_space) real_distribution_order: + shows "r \ real (joint_distribution X Y {(x, y)}) \ r \ real (distribution X {x})" + and "r \ real (joint_distribution X Y {(x, y)}) \ r \ real (distribution Y {y})" + and "r < real (joint_distribution X Y {(x, y)}) \ r < real (distribution X {x})" + and "r < real (joint_distribution X Y {(x, y)}) \ r < real (distribution Y {y})" + and "distribution X {x} = 0 \ real (joint_distribution X Y {(x, y)}) = 0" + and "distribution Y {y} = 0 \ real (joint_distribution X Y {(x, y)}) = 0" + using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_fst, of X Y "{(x, y)}"] + using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_snd, of X Y "{(x, y)}"] + using real_pinfreal_nonneg[of "joint_distribution X Y {(x, y)}"] + by auto + +lemma (in prob_space) joint_distribution_remove[simp]: + "joint_distribution X X {(x, x)} = distribution X {x}" + unfolding distribution_def by (auto intro!: arg_cong[where f="\"]) + +lemma (in finite_prob_space) distribution_1: + "distribution X A \ 1" + unfolding distribution_def measure_space_1[symmetric] + by (auto intro!: measure_mono simp: sets_eq_Pow) + +lemma (in finite_prob_space) real_distribution_1: + "real (distribution X A) \ 1" + unfolding real_pinfreal_1[symmetric] + by (rule real_of_pinfreal_mono[OF _ distribution_1]) simp + +lemma (in finite_prob_space) uniform_prob: + assumes "x \ space M" + assumes "\ x y. \x \ space M ; y \ space M\ \ prob {x} = prob {y}" + shows "prob {x} = 1 / real (card (space M))" +proof - + have prob_x: "\ y. y \ space M \ prob {y} = prob {x}" + using assms(2)[OF _ `x \ space M`] by blast + have "1 = prob (space M)" + using prob_space by auto + also have "\ = (\ x \ space M. prob {x})" + using real_finite_measure_finite_Union[of "space M" "\ x. {x}", simplified] + sets_eq_Pow inj_singleton[unfolded inj_on_def, rule_format] + finite_space unfolding disjoint_family_on_def prob_space[symmetric] + by (auto simp add:setsum_restrict_set) + also have "\ = (\ y \ space M. prob {x})" + using prob_x by auto + also have "\ = real_of_nat (card (space M)) * prob {x}" by simp + finally have one: "1 = real (card (space M)) * prob {x}" + using real_eq_of_nat by auto + hence two: "real (card (space M)) \ 0" by fastsimp + from one have three: "prob {x} \ 0" by fastsimp + thus ?thesis using one two three divide_cancel_right + by (auto simp:field_simps) +qed + +lemma (in prob_space) prob_space_subalgebra: + assumes "N \ sets M" "sigma_algebra (M\ sets := N \)" + shows "prob_space (M\ sets := N \) \" +proof - + interpret N: measure_space "M\ sets := N \" \ + using measure_space_subalgebra[OF assms] . + show ?thesis + proof qed (simp add: measure_space_1) +qed + +lemma (in prob_space) prob_space_of_restricted_space: + assumes "\ A \ 0" "\ A \ \" "A \ sets M" + shows "prob_space (restricted_space A) (\S. \ S / \ A)" + unfolding prob_space_def prob_space_axioms_def +proof + show "\ (space (restricted_space A)) / \ A = 1" + using `\ A \ 0` `\ A \ \` by (auto simp: pinfreal_noteq_omega_Ex) + have *: "\S. \ S / \ A = inverse (\ A) * \ S" by (simp add: mult_commute) + interpret A: measure_space "restricted_space A" \ + using `A \ sets M` by (rule restricted_measure_space) + show "measure_space (restricted_space A) (\S. \ S / \ A)" + proof + show "\ {} / \ A = 0" by auto + show "countably_additive (restricted_space A) (\S. \ S / \ A)" + unfolding countably_additive_def psuminf_cmult_right * + using A.measure_countably_additive by auto + qed +qed + +lemma finite_prob_spaceI: + assumes "finite (space M)" "sets M = Pow(space M)" "\ (space M) = 1" "\ {} = 0" + and "\A B. A\space M \ B\space M \ A \ B = {} \ \ (A \ B) = \ A + \ B" + shows "finite_prob_space M \" + unfolding finite_prob_space_eq +proof + show "finite_measure_space M \" using assms + by (auto intro!: finite_measure_spaceI) + show "\ (space M) = 1" by fact +qed + +lemma (in finite_prob_space) finite_measure_space: + fixes X :: "'a \ 'x" + shows "finite_measure_space \space = X ` space M, sets = Pow (X ` space M)\ (distribution X)" + (is "finite_measure_space ?S _") +proof (rule finite_measure_spaceI, simp_all) + show "finite (X ` space M)" using finite_space by simp +next + fix A B :: "'x set" assume "A \ B = {}" + then show "distribution X (A \ B) = distribution X A + distribution X B" + unfolding distribution_def + by (subst measure_additive) + (auto intro!: arg_cong[where f=\] simp: sets_eq_Pow) +qed + +lemma (in finite_prob_space) finite_prob_space_of_images: + "finite_prob_space \ space = X ` space M, sets = Pow (X ` space M)\ (distribution X)" + by (simp add: finite_prob_space_eq finite_measure_space) + +lemma (in prob_space) joint_distribution_commute: + "joint_distribution X Y x = joint_distribution Y X ((\(x,y). (y,x))`x)" + unfolding distribution_def by (auto intro!: arg_cong[where f=\]) + +lemma (in finite_prob_space) real_distribution_order': + shows "real (distribution X {x}) = 0 \ real (joint_distribution X Y {(x, y)}) = 0" + and "real (distribution Y {y}) = 0 \ real (joint_distribution X Y {(x, y)}) = 0" + using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_fst, of X Y "{(x, y)}"] + using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_snd, of X Y "{(x, y)}"] + using real_pinfreal_nonneg[of "joint_distribution X Y {(x, y)}"] + by auto + lemma (in finite_prob_space) finite_product_measure_space: + fixes X :: "'a \ 'x" and Y :: "'a \ 'y" assumes "finite s1" "finite s2" shows "finite_measure_space \ space = s1 \ s2, sets = Pow (s1 \ s2)\ (joint_distribution X Y)" (is "finite_measure_space ?M ?D") -proof (rule finite_Pow_additivity_sufficient) - show "positive ?D" - unfolding positive_def using assms sets_eq_Pow - by (simp add: distribution_def) - - show "additive ?M ?D" unfolding additive_def - proof safe - fix x y - have A: "((\x. (X x, Y x)) -` x) \ space M \ sets M" using assms sets_eq_Pow by auto - have B: "((\x. (X x, Y x)) -` y) \ space M \ sets M" using assms sets_eq_Pow by auto - assume "x \ y = {}" - hence "(\x. (X x, Y x)) -` x \ space M \ ((\x. (X x, Y x)) -` y \ space M) = {}" - by auto - from additive[unfolded additive_def, rule_format, OF A B] this - finite_measure[OF A] finite_measure[OF B] - show "?D (x \ y) = ?D x + ?D y" - apply (simp add: distribution_def) - apply (subst Int_Un_distrib2) - by (auto simp: real_of_pinfreal_add) - qed - - show "finite (space ?M)" +proof (rule finite_measure_spaceI, simp_all) + show "finite (s1 \ s2)" using assms by auto - - show "sets ?M = Pow (space ?M)" - by simp - - { fix x assume "x \ space ?M" thus "?D {x} \ \" - unfolding distribution_def by (auto intro!: finite_measure simp: sets_eq_Pow) } + show "joint_distribution X Y (s1\s2) \ \" + using distribution_finite . +next + fix A B :: "('x*'y) set" assume "A \ B = {}" + then show "joint_distribution X Y (A \ B) = joint_distribution X Y A + joint_distribution X Y B" + unfolding distribution_def + by (subst measure_additive) + (auto intro!: arg_cong[where f=\] simp: sets_eq_Pow) qed lemma (in finite_prob_space) finite_product_measure_space_of_images: @@ -420,47 +586,133 @@ (joint_distribution X Y)" using finite_space by (auto intro!: finite_product_measure_space) -lemma (in finite_prob_space) finite_measure_space: - shows "finite_measure_space \space = X ` space M, sets = Pow (X ` space M)\ (distribution X)" - (is "finite_measure_space ?S _") -proof (rule finite_Pow_additivity_sufficient, simp_all) - show "finite (X ` space M)" using finite_space by simp +section "Conditional Expectation and Probability" - show "positive (distribution X)" - unfolding distribution_def positive_def using sets_eq_Pow by auto +lemma (in prob_space) conditional_expectation_exists: + fixes X :: "'a \ pinfreal" + assumes borel: "X \ borel_measurable M" + and N_subalgebra: "N \ sets M" "sigma_algebra (M\ sets := N \)" + shows "\Y\borel_measurable (M\ sets := N \). \C\N. + positive_integral (\x. Y x * indicator C x) = positive_integral (\x. X x * indicator C x)" +proof - + interpret P: prob_space "M\ sets := N \" \ + using prob_space_subalgebra[OF N_subalgebra] . + + let "?f A" = "\x. X x * indicator A x" + let "?Q A" = "positive_integral (?f A)" + + from measure_space_density[OF borel] + have Q: "measure_space (M\ sets := N \) ?Q" + by (rule measure_space.measure_space_subalgebra[OF _ N_subalgebra]) + then interpret Q: measure_space "M\ sets := N \" ?Q . - show "additive ?S (distribution X)" unfolding additive_def distribution_def - proof (simp, safe) - fix x y - have x: "(X -` x) \ space M \ sets M" - and y: "(X -` y) \ space M \ sets M" using sets_eq_Pow by auto - assume "x \ y = {}" - hence "X -` x \ space M \ (X -` y \ space M) = {}" by auto - from additive[unfolded additive_def, rule_format, OF x y] this - finite_measure[OF x] finite_measure[OF y] - have "\ (((X -` x) \ (X -` y)) \ space M) = - \ ((X -` x) \ space M) + \ ((X -` y) \ space M)" - by (subst Int_Un_distrib2) auto - thus "\ ((X -` x \ X -` y) \ space M) = \ (X -` x \ space M) + \ (X -` y \ space M)" - by auto + have "P.absolutely_continuous ?Q" + unfolding P.absolutely_continuous_def + proof (safe, simp) + fix A assume "A \ N" "\ A = 0" + moreover then have f_borel: "?f A \ borel_measurable M" + using borel N_subalgebra by (auto intro: borel_measurable_indicator) + moreover have "{x\space M. ?f A x \ 0} = (?f A -` {0<..} \ space M) \ A" + by (auto simp: indicator_def) + moreover have "\ \ \ \ A" + using `A \ N` N_subalgebra f_borel + by (auto intro!: measure_mono Int[of _ A] measurable_sets) + ultimately show "?Q A = 0" + by (simp add: positive_integral_0_iff) qed - - { fix x assume "x \ X ` space M" thus "distribution X {x} \ \" - unfolding distribution_def by (auto intro!: finite_measure simp: sets_eq_Pow) } + from P.Radon_Nikodym[OF Q this] + obtain Y where Y: "Y \ borel_measurable (M\sets := N\)" + "\A. A \ sets (M\sets:=N\) \ ?Q A = P.positive_integral (\x. Y x * indicator A x)" + by blast + with N_subalgebra show ?thesis + by (auto intro!: bexI[OF _ Y(1)]) qed -lemma (in finite_prob_space) finite_prob_space_of_images: - "finite_prob_space \ space = X ` space M, sets = Pow (X ` space M)\ (distribution X)" - by (simp add: finite_prob_space_eq finite_measure_space) +definition (in prob_space) + "conditional_expectation N X = (SOME Y. Y\borel_measurable (M\sets:=N\) + \ (\C\N. positive_integral (\x. Y x * indicator C x) = positive_integral (\x. X x * indicator C x)))" + +abbreviation (in prob_space) + "conditional_prob N A \ conditional_expectation N (indicator A)" + +lemma (in prob_space) + fixes X :: "'a \ pinfreal" + assumes borel: "X \ borel_measurable M" + and N_subalgebra: "N \ sets M" "sigma_algebra (M\ sets := N \)" + shows borel_measurable_conditional_expectation: + "conditional_expectation N X \ borel_measurable (M\ sets := N \)" + and conditional_expectation: "\C. C \ N \ + positive_integral (\x. conditional_expectation N X x * indicator C x) = + positive_integral (\x. X x * indicator C x)" + (is "\C. C \ N \ ?eq C") +proof - + note CE = conditional_expectation_exists[OF assms, unfolded Bex_def] + then show "conditional_expectation N X \ borel_measurable (M\ sets := N \)" + unfolding conditional_expectation_def by (rule someI2_ex) blast + + from CE show "\C. C\N \ ?eq C" + unfolding conditional_expectation_def by (rule someI2_ex) blast +qed + +lemma (in sigma_algebra) factorize_measurable_function: + fixes Z :: "'a \ pinfreal" and Y :: "'a \ 'c" + assumes "sigma_algebra M'" and "Y \ measurable M M'" "Z \ borel_measurable M" + shows "Z \ borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y) + \ (\g\borel_measurable M'. \x\space M. Z x = g (Y x))" +proof safe + interpret M': sigma_algebra M' by fact + have Y: "Y \ space M \ space M'" using assms unfolding measurable_def by auto + from M'.sigma_algebra_vimage[OF this] + interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" . -lemma (in finite_prob_space) finite_product_prob_space_of_images: - "finite_prob_space \ space = X ` space M \ Y ` space M, sets = Pow (X ` space M \ Y ` space M)\ - (joint_distribution X Y)" - (is "finite_prob_space ?S _") -proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images) - have "X -` X ` space M \ Y -` Y ` space M \ space M = space M" by auto - thus "joint_distribution X Y (X ` space M \ Y ` space M) = 1" - by (simp add: distribution_def prob_space vimage_Times comp_def measure_space_1) + { fix g :: "'c \ pinfreal" assume "g \ borel_measurable M'" + with M'.measurable_vimage_algebra[OF Y] + have "g \ Y \ borel_measurable (M'.vimage_algebra (space M) Y)" + by (rule measurable_comp) + moreover assume "\x\space M. Z x = g (Y x)" + then have "Z \ borel_measurable (M'.vimage_algebra (space M) Y) \ + g \ Y \ borel_measurable (M'.vimage_algebra (space M) Y)" + by (auto intro!: measurable_cong) + ultimately show "Z \ borel_measurable (M'.vimage_algebra (space M) Y)" + by simp } + + assume "Z \ borel_measurable (M'.vimage_algebra (space M) Y)" + from va.borel_measurable_implies_simple_function_sequence[OF this] + obtain f where f: "\i. va.simple_function (f i)" and "f \ Z" by blast + + have "\i. \g. M'.simple_function g \ (\x\space M. f i x = g (Y x))" + proof + fix i + from f[of i] have "finite (f i`space M)" and B_ex: + "\z\(f i)`space M. \B. B \ sets M' \ (f i) -` {z} \ space M = Y -` B \ space M" + unfolding va.simple_function_def by auto + from B_ex[THEN bchoice] guess B .. note B = this + + let ?g = "\x. \z\f i`space M. z * indicator (B z) x" + + show "\g. M'.simple_function g \ (\x\space M. f i x = g (Y x))" + proof (intro exI[of _ ?g] conjI ballI) + show "M'.simple_function ?g" using B by auto + + fix x assume "x \ space M" + then have "\z. z \ f i`space M \ indicator (B z) (Y x) = (indicator (f i -` {z} \ space M) x::pinfreal)" + unfolding indicator_def using B by auto + then show "f i x = ?g (Y x)" using `x \ space M` f[of i] + by (subst va.simple_function_indicator_representation) auto + qed + qed + from choice[OF this] guess g .. note g = this + + show "\g\borel_measurable M'. \x\space M. Z x = g (Y x)" + proof (intro ballI bexI) + show "(SUP i. g i) \ borel_measurable M'" + using g by (auto intro: M'.borel_measurable_simple_function) + fix x assume "x \ space M" + have "Z x = (SUP i. f i) x" using `f \ Z` unfolding isoton_def by simp + also have "\ = (SUP i. g i) (Y x)" unfolding SUPR_fun_expand + using g `x \ space M` by simp + finally show "Z x = (SUP i. g i) (Y x)" . + qed qed end diff -r bddc3d3f6e53 -r e9467adb8b52 src/HOL/Probability/Product_Measure.thy --- a/src/HOL/Probability/Product_Measure.thy Thu Sep 02 18:45:23 2010 +0200 +++ b/src/HOL/Probability/Product_Measure.thy Thu Sep 02 20:44:33 2010 +0200 @@ -2,14 +2,412 @@ imports Lebesgue_Integration begin +definition "dynkin M \ + space M \ sets M \ + (\ A \ sets M. A \ space M) \ + (\ a \ sets M. \ b \ sets M. a \ b \ b - a \ sets M) \ + (\A. disjoint_family A \ range A \ sets M \ (\i::nat. A i) \ sets M)" + +lemma dynkinI: + assumes "\ A. A \ sets M \ A \ space M" + assumes "space M \ sets M" and "\ b \ sets M. \ a \ sets M. a \ b \ b - a \ sets M" + assumes "\ a. (\ i j :: nat. i \ j \ a i \ a j = {}) + \ (\ i :: nat. a i \ sets M) \ UNION UNIV a \ sets M" + shows "dynkin M" +using assms by (auto simp: dynkin_def disjoint_family_on_def image_subset_iff) + +lemma dynkin_subset: + assumes "dynkin M" + shows "\ A. A \ sets M \ A \ space M" +using assms unfolding dynkin_def by auto + +lemma dynkin_space: + assumes "dynkin M" + shows "space M \ sets M" +using assms unfolding dynkin_def by auto + +lemma dynkin_diff: + assumes "dynkin M" + shows "\ a b. \ a \ sets M ; b \ sets M ; a \ b \ \ b - a \ sets M" +using assms unfolding dynkin_def by auto + +lemma dynkin_empty: + assumes "dynkin M" + shows "{} \ sets M" +using dynkin_diff[OF assms dynkin_space[OF assms] dynkin_space[OF assms]] by auto + +lemma dynkin_UN: + assumes "dynkin M" + assumes "\ i j :: nat. i \ j \ a i \ a j = {}" + assumes "\ i :: nat. a i \ sets M" + shows "UNION UNIV a \ sets M" +using assms by (auto simp: dynkin_def disjoint_family_on_def image_subset_iff) + +definition "Int_stable M \ (\ a \ sets M. \ b \ sets M. a \ b \ sets M)" + +lemma dynkin_trivial: + shows "dynkin \ space = A, sets = Pow A \" +by (rule dynkinI) auto + +lemma dynkin_lemma: + fixes D :: "'a algebra" + assumes stab: "Int_stable E" + and spac: "space E = space D" + and subsED: "sets E \ sets D" + and subsDE: "sets D \ sigma_sets (space E) (sets E)" + and dyn: "dynkin D" + shows "sigma (space E) (sets E) = D" +proof - + def sets_\E == "\ {sets d | d :: 'a algebra. dynkin d \ space d = space E \ sets E \ sets d}" + def \E == "\ space = space E, sets = sets_\E \" + have "\ space = space E, sets = Pow (space E) \ \ {d | d. dynkin d \ space d = space E \ sets E \ sets d}" + using dynkin_trivial spac subsED dynkin_subset[OF dyn] by fastsimp + hence not_empty: "{sets (d :: 'a algebra) | d. dynkin d \ space d = space E \ sets E \ sets d} \ {}" + using exI[of "\ x. space x = space E \ dynkin x \ sets E \ sets x" "\ space = space E, sets = Pow (space E) \", simplified] + by auto + have \E_D: "sets_\E \ sets D" + unfolding sets_\E_def using assms by auto + have \ynkin: "dynkin \E" + proof (rule dynkinI, safe) + fix A x assume asm: "A \ sets \E" "x \ A" + { fix d :: "('a, 'b) algebra_scheme" assume "A \ sets d" "dynkin d \ space d = space E" + hence "A \ space d" using dynkin_subset by auto } + show "x \ space \E" using asm unfolding \E_def sets_\E_def using not_empty + by simp (metis dynkin_subset in_mono mem_def) + next + show "space \E \ sets \E" + unfolding \E_def sets_\E_def + using dynkin_space by fastsimp + next + fix a b assume "a \ sets \E" "b \ sets \E" "a \ b" + thus "b - a \ sets \E" + unfolding \E_def sets_\E_def by (auto intro:dynkin_diff) + next + fix a assume asm: "\i j :: nat. i \ j \ a i \ a j = {}" "\i. a i \ sets \E" + thus "UNION UNIV a \ sets \E" + unfolding \E_def sets_\E_def apply (auto intro!:dynkin_UN[OF _ asm(1)]) + by blast + qed + + def Dy == "\ d. {A | A. A \ sets_\E \ A \ d \ sets_\E}" + { fix d assume dasm: "d \ sets_\E" + have "dynkin \ space = space E, sets = Dy d \" + proof (rule dynkinI, safe, simp_all) + fix A x assume "A \ Dy d" "x \ A" + thus "x \ space E" unfolding Dy_def sets_\E_def using not_empty + by simp (metis dynkin_subset in_mono mem_def) + next + show "space E \ Dy d" + unfolding Dy_def \E_def sets_\E_def + proof auto + fix d assume asm: "dynkin d" "space d = space E" "sets E \ sets d" + hence "space d \ sets d" using dynkin_space[OF asm(1)] by auto + thus "space E \ sets d" using asm by auto + next + fix da :: "'a algebra" assume asm: "dynkin da" "space da = space E" "sets E \ sets da" + have d: "d = space E \ d" + using dasm dynkin_subset[OF asm(1)] asm(2) dynkin_subset[OF \ynkin] + unfolding \E_def by auto + hence "space E \ d \ sets \E" unfolding \E_def + using dasm by auto + have "sets \E \ sets da" unfolding \E_def sets_\E_def using asm + by auto + thus "space E \ d \ sets da" using dasm asm d dynkin_subset[OF \ynkin] + unfolding \E_def by auto + qed + next + fix a b assume absm: "a \ Dy d" "b \ Dy d" "a \ b" + hence "a \ sets \E" "b \ sets \E" + unfolding Dy_def \E_def by auto + hence *: "b - a \ sets \E" + using dynkin_diff[OF \ynkin] `a \ b` by auto + have "a \ d \ sets \E" "b \ d \ sets \E" + using absm unfolding Dy_def \E_def by auto + hence "(b \ d) - (a \ d) \ sets \E" + using dynkin_diff[OF \ynkin] `a \ b` by auto + hence **: "(b - a) \ d \ sets \E" by (auto simp add:Diff_Int_distrib2) + thus "b - a \ Dy d" + using * ** unfolding Dy_def \E_def by auto + next + fix a assume aasm: "\i j :: nat. i \ j \ a i \ a j = {}" "\i. a i \ Dy d" + hence "\ i. a i \ sets \E" + unfolding Dy_def \E_def by auto + from dynkin_UN[OF \ynkin aasm(1) this] + have *: "UNION UNIV a \ sets \E" by auto + from aasm + have aE: "\ i. a i \ d \ sets \E" + unfolding Dy_def \E_def by auto + from aasm + have "\i j :: nat. i \ j \ (a i \ d) \ (a j \ d) = {}" by auto + from dynkin_UN[OF \ynkin this] + have "UNION UNIV (\ i. a i \ d) \ sets \E" + using aE by auto + hence **: "UNION UNIV a \ d \ sets \E" by auto + from * ** show "UNION UNIV a \ Dy d" unfolding Dy_def \E_def by auto + qed } note Dy_nkin = this + have E_\E: "sets E \ sets \E" + unfolding \E_def sets_\E_def by auto + { fix d assume dasm: "d \ sets \E" + { fix e assume easm: "e \ sets E" + hence deasm: "e \ sets \E" + unfolding \E_def sets_\E_def by auto + have subset: "Dy e \ sets \E" + unfolding Dy_def \E_def by auto + { fix e' assume e'asm: "e' \ sets E" + have "e' \ e \ sets E" + using easm e'asm stab unfolding Int_stable_def by auto + hence "e' \ e \ sets \E" + unfolding \E_def sets_\E_def by auto + hence "e' \ Dy e" using e'asm unfolding Dy_def \E_def sets_\E_def by auto } + hence E_Dy: "sets E \ Dy e" by auto + have "\ space = space E, sets = Dy e \ \ {d | d. dynkin d \ space d = space E \ sets E \ sets d}" + using Dy_nkin[OF deasm[unfolded \E_def, simplified]] E_\E E_Dy by auto + hence "sets_\E \ Dy e" + unfolding sets_\E_def by auto (metis E_Dy simps(1) simps(2) spac) + hence "sets \E = Dy e" using subset unfolding \E_def by auto + hence "d \ e \ sets \E" + using dasm easm deasm unfolding Dy_def \E_def by auto + hence "e \ Dy d" using deasm + unfolding Dy_def \E_def + by (auto simp add:Int_commute) } + hence "sets E \ Dy d" by auto + hence "sets \E \ Dy d" using Dy_nkin[OF dasm[unfolded \E_def, simplified]] + unfolding \E_def sets_\E_def + by auto (metis `sets E <= Dy d` simps(1) simps(2) spac) + hence *: "sets \E = Dy d" + unfolding Dy_def \E_def by auto + fix a assume aasm: "a \ sets \E" + hence "a \ d \ sets \E" + using * dasm unfolding Dy_def \E_def by auto } note \E_stab = this + { fix A :: "nat \ 'a set" assume Asm: "range A \ sets \E" "\A. A \ sets \E \ A \ space \E" + "\a. a \ sets \E \ space \E - a \ sets \E" + "{} \ sets \E" "space \E \ sets \E" + let "?A i" = "A i \ (\ j \ {..< i}. space \E - A j)" + { fix i :: nat + have *: "(\ j \ {..< i}. space \E - A j) \ space \E \ sets \E" + apply (induct i) + using lessThan_Suc Asm \E_stab apply fastsimp + apply (subst lessThan_Suc) + apply (subst INT_insert) + apply (subst Int_assoc) + apply (subst \E_stab) + using lessThan_Suc Asm \E_stab Asm + apply (fastsimp simp add:Int_assoc dynkin_diff[OF \ynkin]) + prefer 2 apply simp + apply (rule dynkin_diff[OF \ynkin, of _ "space \E", OF _ dynkin_space[OF \ynkin]]) + using Asm by auto + have **: "\ i. A i \ space \E" using Asm by auto + have "(\ j \ {..< i}. space \E - A j) \ space \E \ (\ j \ {..< i}. A j) = UNIV \ i = 0" + apply (cases i) + using Asm ** dynkin_subset[OF \ynkin, of "A (i - 1)"] + by auto + hence Aisets: "?A i \ sets \E" + apply (cases i) + using Asm * apply fastsimp + apply (rule \E_stab) + using Asm * ** + by (auto simp add:Int_absorb2) + have "?A i = disjointed A i" unfolding disjointed_def + atLeast0LessThan using Asm by auto + hence "?A i = disjointed A i" "?A i \ sets \E" + using Aisets by auto + } note Ai = this + from dynkin_UN[OF \ynkin _ this(2)] this disjoint_family_disjointed[of A] + have "(\ i. ?A i) \ sets \E" + by (auto simp add:disjoint_family_on_def disjointed_def) + hence "(\ i. A i) \ sets \E" + using Ai(1) UN_disjointed_eq[of A] by auto } note \E_UN = this + { fix a b assume asm: "a \ sets \E" "b \ sets \E" + let ?ab = "\ i. if (i::nat) = 0 then a else if i = 1 then b else {}" + have *: "(\ i. ?ab i) \ sets \E" + apply (rule \E_UN) + using asm \E_UN dynkin_empty[OF \ynkin] + dynkin_subset[OF \ynkin] + dynkin_space[OF \ynkin] + dynkin_diff[OF \ynkin] by auto + have "(\ i. ?ab i) = a \ b" apply auto + apply (case_tac "i = 0") + apply auto + apply (case_tac "i = 1") + by auto + hence "a \ b \ sets \E" using * by auto} note \E_Un = this + have "sigma_algebra \E" + apply unfold_locales + using dynkin_subset[OF \ynkin] + using dynkin_diff[OF \ynkin, of _ "space \E", OF _ dynkin_space[OF \ynkin]] + using dynkin_diff[OF \ynkin, of "space \E" "space \E", OF dynkin_space[OF \ynkin] dynkin_space[OF \ynkin]] + using dynkin_space[OF \ynkin] + using \E_UN \E_Un + by auto + from sigma_algebra.sigma_subset[OF this E_\E] \E_D subsDE spac + show ?thesis by (auto simp add:\E_def sigma_def) +qed + +lemma measure_eq: + assumes fin: "\ (space M) = \ (space M)" "\ (space M) < \" + assumes E: "M = sigma (space E) (sets E)" "Int_stable E" + assumes eq: "\ e. e \ sets E \ \ e = \ e" + assumes ms: "measure_space M \" "measure_space M \" + assumes A: "A \ sets M" + shows "\ A = \ A" +proof - + interpret M: measure_space M \ + using ms by simp + interpret M': measure_space M \ + using ms by simp + + let ?D_sets = "{A. A \ sets M \ \ A = \ A}" + have \: "dynkin \ space = space M , sets = ?D_sets \" + proof (rule dynkinI, safe, simp_all) + fix A x assume "A \ sets M \ \ A = \ A" "x \ A" + thus "x \ space M" using assms M.sets_into_space by auto + next + show "\ (space M) = \ (space M)" + using fin by auto + next + fix a b + assume asm: "a \ sets M \ \ a = \ a" + "b \ sets M \ \ b = \ b" "a \ b" + hence "a \ space M" + using M.sets_into_space by auto + from M.measure_mono[OF this] + have "\ a \ \ (space M)" + using asm by auto + hence afin: "\ a < \" + using fin by auto + have *: "b = b - a \ a" using asm by auto + have **: "(b - a) \ a = {}" using asm by auto + have iv: "\ (b - a) + \ a = \ b" + using M.measure_additive[of "b - a" a] + conjunct1[OF asm(1)] conjunct1[OF asm(2)] * ** + by auto + have v: "\ (b - a) + \ a = \ b" + using M'.measure_additive[of "b - a" a] + conjunct1[OF asm(1)] conjunct1[OF asm(2)] * ** + by auto + from iv v have "\ (b - a) = \ (b - a)" using asm afin + pinfreal_add_cancel_right[of "\ (b - a)" "\ a" "\ (b - a)"] + by auto + thus "b - a \ sets M \ \ (b - a) = \ (b - a)" + using asm by auto + next + fix a assume "\i j :: nat. i \ j \ a i \ a j = {}" + "\i. a i \ sets M \ \ (a i) = \ (a i)" + thus "(\x. a x) \ sets M \ \ (\x. a x) = \ (\x. a x)" + using M.measure_countably_additive + M'.measure_countably_additive + M.countable_UN + apply (auto simp add:disjoint_family_on_def image_def) + apply (subst M.measure_countably_additive[symmetric]) + apply (auto simp add:disjoint_family_on_def) + apply (subst M'.measure_countably_additive[symmetric]) + by (auto simp add:disjoint_family_on_def) + qed + have *: "sets E \ ?D_sets" + using eq E sigma_sets.Basic[of _ "sets E"] + by (auto simp add:sigma_def) + have **: "?D_sets \ sets M" by auto + have "M = \ space = space M , sets = ?D_sets \" + unfolding E(1) + apply (rule dynkin_lemma[OF E(2)]) + using eq E space_sigma \ sigma_sets.Basic + by (auto simp add:sigma_def) + from subst[OF this, of "\ M. A \ sets M", OF A] + show ?thesis by auto +qed +(* +lemma + assumes sfin: "range A \ sets M" "(\i. A i) = space M" "\ i :: nat. \ (A i) < \" + assumes A: "\ i. \ (A i) = \ (A i)" "\ i. A i \ A (Suc i)" + assumes E: "M = sigma (space E) (sets E)" "Int_stable E" + assumes eq: "\ e. e \ sets E \ \ e = \ e" + assumes ms: "measure_space (M :: 'a algebra) \" "measure_space M \" + assumes B: "B \ sets M" + shows "\ B = \ B" +proof - + interpret M: measure_space M \ by (rule ms) + interpret M': measure_space M \ by (rule ms) + have *: "M = \ space = space M, sets = sets M \" by auto + { fix i :: nat + have **: "M\ space := A i, sets := op \ (A i) ` sets M \ = + \ space = A i, sets = op \ (A i) ` sets M \" + by auto + have mu_i: "measure_space \ space = A i, sets = op \ (A i) ` sets M \ \" + using M.restricted_measure_space[of "A i", simplified **] + sfin by auto + have nu_i: "measure_space \ space = A i, sets = op \ (A i) ` sets M \ \" + using M'.restricted_measure_space[of "A i", simplified **] + sfin by auto + let ?M = "\ space = A i, sets = op \ (A i) ` sets M \" + have "\ (A i \ B) = \ (A i \ B)" + apply (rule measure_eq[of \ ?M \ "\ space = space E \ A i, sets = op \ (A i) ` sets E\" "A i \ B", simplified]) + using assms nu_i mu_i + apply (auto simp add:image_def) (* TODO *) sorry + show ?thesis sorry +qed +*) definition prod_sets where "prod_sets A B = {z. \x \ A. \y \ B. z = x \ y}" definition - "prod_measure M \ N \ = (\A. measure_space.positive_integral M \ (\s0. \ ((\s1. (s0, s1)) -` A)))" + "prod_measure_space M1 M2 = sigma (space M1 \ space M2) (prod_sets (sets M1) (sets M2))" + +lemma + fixes M1 :: "'a algebra" and M2 :: "'b algebra" + assumes "algebra M1" "algebra M2" + shows measureable_fst[intro!, simp]: + "fst \ measurable (prod_measure_space M1 M2) M1" (is ?fst) + and measureable_snd[intro!, simp]: + "snd \ measurable (prod_measure_space M1 M2) M2" (is ?snd) +proof - + interpret M1: algebra M1 by fact + interpret M2: algebra M2 by fact + + { fix X assume "X \ sets M1" + then have "\X1\sets M1. \X2\sets M2. fst -` X \ space M1 \ space M2 = X1 \ X2" + apply - apply (rule bexI[of _ X]) apply (rule bexI[of _ "space M2"]) + using M1.sets_into_space by force+ } + moreover + { fix X assume "X \ sets M2" + then have "\X1\sets M1. \X2\sets M2. snd -` X \ space M1 \ space M2 = X1 \ X2" + apply - apply (rule bexI[of _ "space M1"]) apply (rule bexI[of _ X]) + using M2.sets_into_space by force+ } + ultimately show ?fst ?snd + by (force intro!: sigma_sets.Basic + simp: measurable_def prod_measure_space_def prod_sets_def sets_sigma)+ +qed + +lemma (in sigma_algebra) measureable_prod: + fixes M1 :: "'a algebra" and M2 :: "'b algebra" + assumes "algebra M1" "algebra M2" + shows "f \ measurable M (prod_measure_space M1 M2) \ + (fst \ f) \ measurable M M1 \ (snd \ f) \ measurable M M2" +using assms proof (safe intro!: measurable_comp[where b="prod_measure_space M1 M2"]) + interpret M1: algebra M1 by fact + interpret M2: algebra M2 by fact + assume f: "(fst \ f) \ measurable M M1" and s: "(snd \ f) \ measurable M M2" + + show "f \ measurable M (prod_measure_space M1 M2)" unfolding prod_measure_space_def + proof (rule measurable_sigma) + show "prod_sets (sets M1) (sets M2) \ Pow (space M1 \ space M2)" + unfolding prod_sets_def using M1.sets_into_space M2.sets_into_space by auto + show "f \ space M \ space M1 \ space M2" + using f s by (auto simp: mem_Times_iff measurable_def comp_def) + fix A assume "A \ prod_sets (sets M1) (sets M2)" + then obtain B C where "B \ sets M1" "C \ sets M2" "A = B \ C" + unfolding prod_sets_def by auto + moreover have "(fst \ f) -` B \ space M \ sets M" + using f `B \ sets M1` unfolding measurable_def by auto + moreover have "(snd \ f) -` C \ space M \ sets M" + using s `C \ sets M2` unfolding measurable_def by auto + moreover have "f -` A \ space M = ((fst \ f) -` B \ space M) \ ((snd \ f) -` C \ space M)" + unfolding `A = B \ C` by (auto simp: vimage_Times) + ultimately show "f -` A \ space M \ sets M" by auto + qed +qed definition - "prod_measure_space M1 M2 = sigma (space M1 \ space M2) (prod_sets (sets M1) (sets M2))" + "prod_measure M \ N \ = (\A. measure_space.positive_integral M \ (\s0. \ ((\s1. (s0, s1)) -` A)))" lemma prod_setsI: "x \ A \ y \ B \ (x \ y) \ prod_sets A B" by (auto simp add: prod_sets_def) @@ -114,36 +512,25 @@ qed lemma (in finite_measure_space) finite_measure_space_finite_prod_measure: - assumes "finite_measure_space N \" + fixes N :: "('c, 'd) algebra_scheme" + assumes N: "finite_measure_space N \" shows "finite_measure_space (prod_measure_space M N) (prod_measure M \ N \)" unfolding finite_prod_measure_space[OF assms] -proof (rule finite_measure_spaceI) +proof (rule finite_measure_spaceI, simp_all) interpret N: finite_measure_space N \ by fact - - let ?P = "\space = space M \ space N, sets = Pow (space M \ space N)\" - show "measure_space ?P (prod_measure M \ N \)" - proof (rule sigma_algebra.finite_additivity_sufficient) - show "sigma_algebra ?P" by (rule sigma_algebra_Pow) - show "finite (space ?P)" using finite_space N.finite_space by auto - from finite_prod_measure_times[OF assms, of "{}" "{}"] - show "positive (prod_measure M \ N \)" - unfolding positive_def by simp + show "finite (space M \ space N)" using finite_space N.finite_space by auto + show "prod_measure M \ N \ (space M \ space N) \ \" + using finite_prod_measure_times[OF N top N.top] by simp + show "prod_measure M \ N \ {} = 0" + using finite_prod_measure_times[OF N empty_sets N.empty_sets] by simp - show "additive ?P (prod_measure M \ N \)" - unfolding additive_def - apply (auto simp add: sets_eq_Pow prod_measure_def positive_integral_add[symmetric] - intro!: positive_integral_cong) - apply (subst N.measure_additive[symmetric]) - by (auto simp: N.sets_eq_Pow sets_eq_Pow) - qed - show "finite (space ?P)" using finite_space N.finite_space by auto - show "sets ?P = Pow (space ?P)" by simp - - fix x assume "x \ space ?P" - with finite_prod_measure_times[OF assms, of "{fst x}" "{snd x}"] - finite_measure[of "{fst x}"] N.finite_measure[of "{snd x}"] - show "prod_measure M \ N \ {x} \ \" - by (auto simp add: sets_eq_Pow N.sets_eq_Pow elim!: SigmaE) + fix A B :: "('a * 'c) set" assume "A \ B = {}" "A \ space M \ space N" "B \ space M \ space N" + then show "prod_measure M \ N \ (A \ B) = prod_measure M \ N \ A + prod_measure M \ N \ B" + apply (auto simp add: sets_eq_Pow prod_measure_def positive_integral_add[symmetric] + intro!: positive_integral_cong) + apply (subst N.measure_additive) + apply (auto intro!: arg_cong[where f=\] simp: N.sets_eq_Pow sets_eq_Pow) + done qed lemma (in finite_measure_space) finite_measure_space_finite_prod_measure_alterantive: @@ -153,4 +540,18 @@ unfolding finite_prod_measure_space[OF N, symmetric] using finite_measure_space_finite_prod_measure[OF N] . -end \ No newline at end of file +lemma prod_measure_times_finite: + assumes fms: "finite_measure_space M \" "finite_measure_space N \" and a: "a \ space M \ space N" + shows "prod_measure M \ N \ {a} = \ {fst a} * \ {snd a}" +proof (cases a) + case (Pair b c) + hence a_eq: "{a} = {b} \ {c}" by simp + + interpret M: finite_measure_space M \ by fact + interpret N: finite_measure_space N \ by fact + + from finite_measure_space.finite_prod_measure_times[OF fms, of "{b}" "{c}"] M.sets_eq_Pow N.sets_eq_Pow a Pair + show ?thesis unfolding a_eq by simp +qed + +end diff -r bddc3d3f6e53 -r e9467adb8b52 src/HOL/Probability/Radon_Nikodym.thy --- a/src/HOL/Probability/Radon_Nikodym.thy Thu Sep 02 18:45:23 2010 +0200 +++ b/src/HOL/Probability/Radon_Nikodym.thy Thu Sep 02 20:44:33 2010 +0200 @@ -2,201 +2,6 @@ imports Lebesgue_Integration begin -lemma (in measure_space) measure_finitely_subadditive: - assumes "finite I" "A ` I \ sets M" - shows "\ (\i\I. A i) \ (\i\I. \ (A i))" -using assms proof induct - case (insert i I) - then have "(\i\I. A i) \ sets M" by (auto intro: finite_UN) - then have "\ (\i\insert i I. A i) \ \ (A i) + \ (\i\I. A i)" - using insert by (simp add: measure_subadditive) - also have "\ \ (\i\insert i I. \ (A i))" - using insert by (auto intro!: add_left_mono) - finally show ?case . -qed simp - -lemma (in sigma_algebra) borel_measurable_restricted: - fixes f :: "'a \ pinfreal" assumes "A \ sets M" - shows "f \ borel_measurable (M\ space := A, sets := op \ A ` sets M \) \ - (\x. f x * indicator A x) \ borel_measurable M" - (is "f \ borel_measurable ?R \ ?f \ borel_measurable M") -proof - - interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \ sets M`]) - have *: "f \ borel_measurable ?R \ ?f \ borel_measurable ?R" - by (auto intro!: measurable_cong) - show ?thesis unfolding * - unfolding in_borel_measurable_borel_space - proof (simp, safe) - fix S :: "pinfreal set" assume "S \ sets borel_space" - "\S\sets borel_space. ?f -` S \ A \ op \ A ` sets M" - then have "?f -` S \ A \ op \ A ` sets M" by auto - then have f: "?f -` S \ A \ sets M" - using `A \ sets M` sets_into_space by fastsimp - show "?f -` S \ space M \ sets M" - proof cases - assume "0 \ S" - then have "?f -` S \ space M = ?f -` S \ A \ (space M - A)" - using `A \ sets M` sets_into_space by auto - then show ?thesis using f `A \ sets M` by (auto intro!: Un Diff) - next - assume "0 \ S" - then have "?f -` S \ space M = ?f -` S \ A" - using `A \ sets M` sets_into_space - by (auto simp: indicator_def split: split_if_asm) - then show ?thesis using f by auto - qed - next - fix S :: "pinfreal set" assume "S \ sets borel_space" - "\S\sets borel_space. ?f -` S \ space M \ sets M" - then have f: "?f -` S \ space M \ sets M" by auto - then show "?f -` S \ A \ op \ A ` sets M" - using `A \ sets M` sets_into_space - apply (simp add: image_iff) - apply (rule bexI[OF _ f]) - by auto - qed -qed - -lemma (in sigma_algebra) simple_function_eq_borel_measurable: - fixes f :: "'a \ pinfreal" - shows "simple_function f \ - finite (f`space M) \ f \ borel_measurable M" - using simple_function_borel_measurable[of f] - borel_measurable_simple_function[of f] - by (fastsimp simp: simple_function_def) - -lemma (in measure_space) simple_function_restricted: - fixes f :: "'a \ pinfreal" assumes "A \ sets M" - shows "sigma_algebra.simple_function (M\ space := A, sets := op \ A ` sets M \) f \ simple_function (\x. f x * indicator A x)" - (is "sigma_algebra.simple_function ?R f \ simple_function ?f") -proof - - interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \ sets M`]) - have "finite (f`A) \ finite (?f`space M)" - proof cases - assume "A = space M" - then have "f`A = ?f`space M" by (fastsimp simp: image_iff) - then show ?thesis by simp - next - assume "A \ space M" - then obtain x where x: "x \ space M" "x \ A" - using sets_into_space `A \ sets M` by auto - have *: "?f`space M = f`A \ {0}" - proof (auto simp add: image_iff) - show "\x\space M. f x = 0 \ indicator A x = 0" - using x by (auto intro!: bexI[of _ x]) - next - fix x assume "x \ A" - then show "\y\space M. f x = f y * indicator A y" - using `A \ sets M` sets_into_space by (auto intro!: bexI[of _ x]) - next - fix x - assume "indicator A x \ (0::pinfreal)" - then have "x \ A" by (auto simp: indicator_def split: split_if_asm) - moreover assume "x \ space M" "\y\A. ?f x \ f y" - ultimately show "f x = 0" by auto - qed - then show ?thesis by auto - qed - then show ?thesis - unfolding simple_function_eq_borel_measurable - R.simple_function_eq_borel_measurable - unfolding borel_measurable_restricted[OF `A \ sets M`] - by auto -qed - -lemma (in measure_space) simple_integral_restricted: - assumes "A \ sets M" - assumes sf: "simple_function (\x. f x * indicator A x)" - shows "measure_space.simple_integral (M\ space := A, sets := op \ A ` sets M \) \ f = simple_integral (\x. f x * indicator A x)" - (is "_ = simple_integral ?f") - unfolding measure_space.simple_integral_def[OF restricted_measure_space[OF `A \ sets M`]] - unfolding simple_integral_def -proof (simp, safe intro!: setsum_mono_zero_cong_left) - from sf show "finite (?f ` space M)" - unfolding simple_function_def by auto -next - fix x assume "x \ A" - then show "f x \ ?f ` space M" - using sets_into_space `A \ sets M` by (auto intro!: image_eqI[of _ _ x]) -next - fix x assume "x \ space M" "?f x \ f`A" - then have "x \ A" by (auto simp: image_iff) - then show "?f x * \ (?f -` {?f x} \ space M) = 0" by simp -next - fix x assume "x \ A" - then have "f x \ 0 \ - f -` {f x} \ A = ?f -` {f x} \ space M" - using `A \ sets M` sets_into_space - by (auto simp: indicator_def split: split_if_asm) - then show "f x * \ (f -` {f x} \ A) = - f x * \ (?f -` {f x} \ space M)" - unfolding pinfreal_mult_cancel_left by auto -qed - -lemma (in measure_space) positive_integral_restricted: - assumes "A \ sets M" - shows "measure_space.positive_integral (M\ space := A, sets := op \ A ` sets M \) \ f = positive_integral (\x. f x * indicator A x)" - (is "measure_space.positive_integral ?R \ f = positive_integral ?f") -proof - - have msR: "measure_space ?R \" by (rule restricted_measure_space[OF `A \ sets M`]) - then interpret R: measure_space ?R \ . - have saR: "sigma_algebra ?R" by fact - have *: "R.positive_integral f = R.positive_integral ?f" - by (auto intro!: R.positive_integral_cong) - show ?thesis - unfolding * R.positive_integral_def positive_integral_def - unfolding simple_function_restricted[OF `A \ sets M`] - apply (simp add: SUPR_def) - apply (rule arg_cong[where f=Sup]) - proof (auto simp: image_iff simple_integral_restricted[OF `A \ sets M`]) - fix g assume "simple_function (\x. g x * indicator A x)" - "g \ f" "\x\A. \ \ g x" - then show "\x. simple_function x \ x \ (\x. f x * indicator A x) \ (\y\space M. \ \ x y) \ - simple_integral (\x. g x * indicator A x) = simple_integral x" - apply (rule_tac exI[of _ "\x. g x * indicator A x"]) - by (auto simp: indicator_def le_fun_def) - next - fix g assume g: "simple_function g" "g \ (\x. f x * indicator A x)" - "\x\space M. \ \ g x" - then have *: "(\x. g x * indicator A x) = g" - "\x. g x * indicator A x = g x" - "\x. g x \ f x" - by (auto simp: le_fun_def expand_fun_eq indicator_def split: split_if_asm) - from g show "\x. simple_function (\xa. x xa * indicator A xa) \ x \ f \ (\xa\A. \ \ x xa) \ - simple_integral g = simple_integral (\xa. x xa * indicator A xa)" - using `A \ sets M`[THEN sets_into_space] - apply (rule_tac exI[of _ "\x. g x * indicator A x"]) - by (fastsimp simp: le_fun_def *) - qed -qed - -lemma (in sigma_algebra) borel_measurable_psuminf: - assumes "\i. f i \ borel_measurable M" - shows "(\x. (\\<^isub>\ i. f i x)) \ borel_measurable M" - using assms unfolding psuminf_def - by (auto intro!: borel_measurable_SUP[unfolded SUPR_fun_expand]) - -lemma (in sigma_finite_measure) disjoint_sigma_finite: - "\A::nat\'a set. range A \ sets M \ (\i. A i) = space M \ - (\i. \ (A i) \ \) \ disjoint_family A" -proof - - obtain A :: "nat \ 'a set" where - range: "range A \ sets M" and - space: "(\i. A i) = space M" and - measure: "\i. \ (A i) \ \" - using sigma_finite by auto - - note range' = range_disjointed_sets[OF range] range - - { fix i - have "\ (disjointed A i) \ \ (A i)" - using range' disjointed_subset[of A i] by (auto intro!: measure_mono) - then have "\ (disjointed A i) \ \" - using measure[of i] by auto } - with disjoint_family_disjointed UN_disjointed_eq[of A] space range' - show ?thesis by (auto intro!: exI[of _ "disjointed A"]) -qed - lemma (in sigma_finite_measure) Ex_finite_integrable_function: shows "\h\borel_measurable M. positive_integral h \ \ \ (\x\space M. 0 < h x \ h x < \)" proof - @@ -206,7 +11,6 @@ measure: "\i. \ (A i) \ \" and disjoint: "disjoint_family A" using disjoint_sigma_finite by auto - let "?B i" = "2^Suc i * \ (A i)" have "\i. \x. 0 < x \ x < inverse (?B i)" proof @@ -225,20 +29,22 @@ qed from choice[OF this] obtain n where n: "\i. 0 < n i" "\i. n i < inverse (2^Suc i * \ (A i))" by auto - let "?h x" = "\\<^isub>\ i. n i * indicator (A i) x" show ?thesis proof (safe intro!: bexI[of _ ?h] del: notI) - have "positive_integral ?h = (\\<^isub>\ i. n i * \ (A i))" - apply (subst positive_integral_psuminf) - using range apply (fastsimp intro!: borel_measurable_pinfreal_times borel_measurable_const borel_measurable_indicator) - apply (subst positive_integral_cmult_indicator) - using range by auto + have "\i. A i \ sets M" + using range by fastsimp+ + then have "positive_integral ?h = (\\<^isub>\ i. n i * \ (A i))" + by (simp add: positive_integral_psuminf positive_integral_cmult_indicator) also have "\ \ (\\<^isub>\ i. Real ((1 / 2)^Suc i))" proof (rule psuminf_le) fix N show "n N * \ (A N) \ Real ((1 / 2) ^ Suc N)" using measure[of N] n[of N] - by (cases "n N") (auto simp: pinfreal_noteq_omega_Ex field_simps zero_le_mult_iff mult_le_0_iff mult_less_0_iff power_less_zero_eq power_le_zero_eq inverse_eq_divide less_divide_eq power_divide split: split_if_asm) + by (cases "n N") + (auto simp: pinfreal_noteq_omega_Ex field_simps zero_le_mult_iff + mult_le_0_iff mult_less_0_iff power_less_zero_eq + power_le_zero_eq inverse_eq_divide less_divide_eq + power_divide split: split_if_asm) qed also have "\ = Real 1" by (rule suminf_imp_psuminf, rule power_half_series, auto) @@ -251,13 +57,37 @@ then show "0 < ?h x" and "?h x < \" using n[of i] by auto next show "?h \ borel_measurable M" using range - by (auto intro!: borel_measurable_psuminf borel_measurable_pinfreal_times borel_measurable_indicator) + by (auto intro!: borel_measurable_psuminf borel_measurable_pinfreal_times) qed qed definition (in measure_space) "absolutely_continuous \ = (\N\null_sets. \ N = (0 :: pinfreal))" +lemma (in finite_measure_space) absolutely_continuousI: + assumes "finite_measure_space M \" + assumes v: "\x. \ x \ space M ; \ {x} = 0 \ \ \ {x} = 0" + shows "absolutely_continuous \" +proof (unfold absolutely_continuous_def sets_eq_Pow, safe) + fix N assume "\ N = 0" "N \ space M" + interpret v: finite_measure_space M \ by fact + have "\ N = \ (\x\N. {x})" by simp + also have "\ = (\x\N. \ {x})" + proof (rule v.measure_finitely_additive''[symmetric]) + show "finite N" using `N \ space M` finite_space by (auto intro: finite_subset) + show "disjoint_family_on (\i. {i}) N" unfolding disjoint_family_on_def by auto + fix x assume "x \ N" thus "{x} \ sets M" using `N \ space M` sets_eq_Pow by auto + qed + also have "\ = 0" + proof (safe intro!: setsum_0') + fix x assume "x \ N" + hence "\ {x} \ \ N" using sets_eq_Pow `N \ space M` by (auto intro!: measure_mono) + hence "\ {x} = 0" using `\ N = 0` by simp + thus "\ {x} = 0" using v[of x] `x \ N` `N \ space M` by auto + qed + finally show "\ N = 0" . +qed + lemma (in finite_measure) Radon_Nikodym_aux_epsilon: fixes e :: real assumes "0 < e" assumes "finite_measure M s" @@ -370,7 +200,7 @@ interpret M': finite_measure M s by fact - let "?r S" = "M\ space := S, sets := (\C. S \ C)`sets M\" + let "?r S" = "restricted_space S" { fix S n assume S: "S \ sets M" @@ -838,7 +668,7 @@ = (f x * indicator (Q i) x) * indicator A x" unfolding indicator_def by auto - have fm: "finite_measure (M\space := Q i, sets := op \ (Q i) ` sets M\) \" + have fm: "finite_measure (restricted_space (Q i)) \" (is "finite_measure ?R \") by (rule restricted_finite_measure[OF Q_sets[of i]]) then interpret R: finite_measure ?R . have fmv: "finite_measure ?R \" @@ -935,47 +765,6 @@ qed qed -lemma (in measure_space) positive_integral_translated_density: - assumes "f \ borel_measurable M" "g \ borel_measurable M" - shows "measure_space.positive_integral M (\A. positive_integral (\x. f x * indicator A x)) g = - positive_integral (\x. f x * g x)" (is "measure_space.positive_integral M ?T _ = _") -proof - - from measure_space_density[OF assms(1)] - interpret T: measure_space M ?T . - - from borel_measurable_implies_simple_function_sequence[OF assms(2)] - obtain G where G: "\i. simple_function (G i)" "G \ g" by blast - note G_borel = borel_measurable_simple_function[OF this(1)] - - from T.positive_integral_isoton[OF `G \ g` G_borel] - have *: "(\i. T.positive_integral (G i)) \ T.positive_integral g" . - - { fix i - have [simp]: "finite (G i ` space M)" - using G(1) unfolding simple_function_def by auto - have "T.positive_integral (G i) = T.simple_integral (G i)" - using G T.positive_integral_eq_simple_integral by simp - also have "\ = positive_integral (\x. f x * (\y\G i`space M. y * indicator (G i -` {y} \ space M) x))" - apply (simp add: T.simple_integral_def) - apply (subst positive_integral_cmult[symmetric]) - using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage) - apply (subst positive_integral_setsum[symmetric]) - using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage) - by (simp add: setsum_right_distrib field_simps) - also have "\ = positive_integral (\x. f x * G i x)" - by (auto intro!: positive_integral_cong - simp: indicator_def if_distrib setsum_cases) - finally have "T.positive_integral (G i) = positive_integral (\x. f x * G i x)" . } - with * have eq_Tg: "(\i. positive_integral (\x. f x * G i x)) \ T.positive_integral g" by simp - - from G(2) have "(\i x. f x * G i x) \ (\x. f x * g x)" - unfolding isoton_fun_expand by (auto intro!: isoton_cmult_right) - then have "(\i. positive_integral (\x. f x * G i x)) \ positive_integral (\x. f x * g x)" - using assms(1) G_borel by (auto intro!: positive_integral_isoton borel_measurable_pinfreal_times) - with eq_Tg show "T.positive_integral g = positive_integral (\x. f x * g x)" - unfolding isoton_def by simp -qed - lemma (in sigma_finite_measure) Radon_Nikodym: assumes "measure_space M \" assumes "absolutely_continuous \" diff -r bddc3d3f6e53 -r e9467adb8b52 src/HOL/Probability/Sigma_Algebra.thy --- a/src/HOL/Probability/Sigma_Algebra.thy Thu Sep 02 18:45:23 2010 +0200 +++ b/src/HOL/Probability/Sigma_Algebra.thy Thu Sep 02 20:44:33 2010 +0200 @@ -6,7 +6,7 @@ header {* Sigma Algebras *} -theory Sigma_Algebra imports Main Countable FuncSet begin +theory Sigma_Algebra imports Main Countable FuncSet Indicator_Function begin text {* Sigma algebras are an elementary concept in measure theory. To measure --- that is to integrate --- functions, we first have @@ -95,10 +95,13 @@ lemma (in algebra) Int_space_eq2 [simp]: "x \ sets M \ x \ space M = x" by (metis Int_absorb2 sets_into_space) +section {* Restricted algebras *} + +abbreviation (in algebra) + "restricted_space A \ \ space = A, sets = (\S. (A \ S)) ` sets M \" + lemma (in algebra) restricted_algebra: - assumes "S \ sets M" - shows "algebra (M\ space := S, sets := (\A. S \ A) ` sets M \)" - (is "algebra ?r") + assumes "A \ sets M" shows "algebra (restricted_space A)" using assms by unfold_locales auto subsection {* Sigma Algebras *} @@ -107,6 +110,13 @@ assumes countable_nat_UN [intro]: "!!A. range A \ sets M \ (\i::nat. A i) \ sets M" +lemma sigma_algebra_cong: + fixes M :: "('a, 'b) algebra_scheme" and M' :: "('a, 'c) algebra_scheme" + assumes *: "sigma_algebra M" + and cong: "space M = space M'" "sets M = sets M'" + shows "sigma_algebra M'" +using * unfolding sigma_algebra_def algebra_def sigma_algebra_axioms_def unfolding cong . + lemma countable_UN_eq: fixes A :: "'i::countable \ 'a set" shows "(range A \ sets M \ (\i. A i) \ sets M) \ @@ -320,15 +330,14 @@ lemma (in sigma_algebra) restricted_sigma_algebra: assumes "S \ sets M" - shows "sigma_algebra (M\ space := S, sets := (\A. S \ A) ` sets M \)" - (is "sigma_algebra ?r") + shows "sigma_algebra (restricted_space S)" unfolding sigma_algebra_def sigma_algebra_axioms_def proof safe - show "algebra ?r" using restricted_algebra[OF assms] . + show "algebra (restricted_space S)" using restricted_algebra[OF assms] . next - fix A :: "nat \ 'a set" assume "range A \ sets ?r" + fix A :: "nat \ 'a set" assume "range A \ sets (restricted_space S)" from restriction_in_sets[OF assms this[simplified]] - show "(\i. A i) \ sets ?r" by simp + show "(\i. A i) \ sets (restricted_space S)" by simp qed section {* Measurable functions *} @@ -560,6 +569,19 @@ (metis insert_absorb insert_subset le_SucE le_antisym not_leE) qed +lemma setsum_indicator_disjoint_family: + fixes f :: "'d \ 'e::semiring_1" + assumes d: "disjoint_family_on A P" and "x \ A j" and "finite P" and "j \ P" + shows "(\i\P. f i * indicator (A i) x) = f j" +proof - + have "P \ {i. x \ A i} = {j}" + using d `x \ A j` `j \ P` unfolding disjoint_family_on_def + by auto + thus ?thesis + unfolding indicator_def + by (simp add: if_distrib setsum_cases[OF `finite P`]) +qed + definition disjointed :: "(nat \ 'a set) \ nat \ 'a set " where "disjointed A n = A n - (\i\{0..i::nat. A i) \ sets M" by (simp add: UN_disjointed_eq) qed +subsection {* Sigma algebra generated by function preimages *} + +definition (in sigma_algebra) + "vimage_algebra S f = \ space = S, sets = (\A. f -` A \ S) ` sets M \" + +lemma (in sigma_algebra) in_vimage_algebra[simp]: + "A \ sets (vimage_algebra S f) \ (\B\sets M. A = f -` B \ S)" + by (simp add: vimage_algebra_def image_iff) + +lemma (in sigma_algebra) space_vimage_algebra[simp]: + "space (vimage_algebra S f) = S" + by (simp add: vimage_algebra_def) + +lemma (in sigma_algebra) sigma_algebra_vimage: + fixes S :: "'c set" assumes "f \ S \ space M" + shows "sigma_algebra (vimage_algebra S f)" +proof + fix A assume "A \ sets (vimage_algebra S f)" + then guess B unfolding in_vimage_algebra .. + then show "space (vimage_algebra S f) - A \ sets (vimage_algebra S f)" + using sets_into_space assms + by (auto intro!: bexI[of _ "space M - B"]) +next + fix A assume "A \ sets (vimage_algebra S f)" + then guess A' unfolding in_vimage_algebra .. note A' = this + fix B assume "B \ sets (vimage_algebra S f)" + then guess B' unfolding in_vimage_algebra .. note B' = this + then show "A \ B \ sets (vimage_algebra S f)" + using sets_into_space assms A' B' + by (auto intro!: bexI[of _ "A' \ B'"]) +next + fix A::"nat \ 'c set" assume "range A \ sets (vimage_algebra S f)" + then have "\i. \B. A i = f -` B \ S \ B \ sets M" + by (simp add: subset_eq) blast + from this[THEN choice] obtain B + where B: "\i. A i = f -` B i \ S" "range B \ sets M" by auto + show "(\i. A i) \ sets (vimage_algebra S f)" + using B by (auto intro!: bexI[of _ "\i. B i"]) +qed auto + +lemma (in sigma_algebra) measurable_vimage_algebra: + fixes S :: "'c set" assumes "f \ S \ space M" + shows "f \ measurable (vimage_algebra S f) M" + unfolding measurable_def using assms by force + +section {* Conditional space *} + +definition (in algebra) + "image_space X = \ space = X`space M, sets = (\S. X`S) ` sets M \" + +definition (in algebra) + "conditional_space X A = algebra.image_space (restricted_space A) X" + +lemma (in algebra) space_conditional_space: + assumes "A \ sets M" shows "space (conditional_space X A) = X`A" +proof - + interpret r: algebra "restricted_space A" using assms by (rule restricted_algebra) + show ?thesis unfolding conditional_space_def r.image_space_def + by simp +qed + subsection {* A Two-Element Series *} definition binaryset :: "'a set \ 'a set \ nat \ 'a set " diff -r bddc3d3f6e53 -r e9467adb8b52 src/HOL/Probability/ex/Dining_Cryptographers.thy --- a/src/HOL/Probability/ex/Dining_Cryptographers.thy Thu Sep 02 18:45:23 2010 +0200 +++ b/src/HOL/Probability/ex/Dining_Cryptographers.thy Thu Sep 02 20:44:33 2010 +0200 @@ -26,14 +26,13 @@ let ?measure = "\s::'a set. real (card s) / real (card S)" show "finite_measure_space M \" - proof (rule finite_Pow_additivity_sufficient, simp_all) - show "positive \" by (simp add: positive_def) - - show "additive M \" - by (simp add: additive_def inverse_eq_divide field_simps Real_real + proof (rule finite_measure_spaceI) + fix A B :: "'a set" assume "A \ B = {}" "A \ space M" "B \ space M" + then show "\ (A \ B) = \ A + \ B" + by (simp add: inverse_eq_divide field_simps Real_real divide_le_0_iff zero_le_divide_iff card_Un_disjoint finite_subset[OF _ finite]) - qed + qed auto qed simp_all lemma set_of_list_extend: