# HG changeset patch # User bulwahn # Date 1307625261 -7200 # Node ID b017cfb10df4ffbde9f21bb84e076136094ec990 # Parent e83695ea0e0a74535cc45a9795b4ed7c0d31cb07# Parent 2929f96d3ae7a34f264bf8a8dc554ddc6eaf1810 merged diff -r e83695ea0e0a -r b017cfb10df4 src/HOL/Library/Convex.thy --- a/src/HOL/Library/Convex.thy Thu Jun 09 14:16:12 2011 +0200 +++ b/src/HOL/Library/Convex.thy Thu Jun 09 15:14:21 2011 +0200 @@ -465,6 +465,25 @@ thus "convex_on C f" unfolding convex_on_def by auto qed +lemma convex_on_diff: + fixes f :: "real \ real" + assumes f: "convex_on I f" and I: "x\I" "y\I" and t: "x < t" "t < y" + shows "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" "(f x - f y) / (x - y) \ (f t - f y) / (t - y)" +proof - + def a \ "(t - y) / (x - y)" + with t have "0 \ a" "0 \ 1 - a" by (auto simp: field_simps) + with f `x \ I` `y \ I` have cvx: "f (a * x + (1 - a) * y) \ a * f x + (1 - a) * f y" + by (auto simp: convex_on_def) + have "a * x + (1 - a) * y = a * (x - y) + y" by (simp add: field_simps) + also have "\ = t" unfolding a_def using `x < t` `t < y` by simp + finally have "f t \ a * f x + (1 - a) * f y" using cvx by simp + also have "\ = a * (f x - f y) + f y" by (simp add: field_simps) + finally have "f t - f y \ a * (f x - f y)" by simp + with t show "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" + by (simp add: times_divide_eq le_divide_eq divide_le_eq field_simps a_def) + with t show "(f x - f y) / (x - y) \ (f t - f y) / (t - y)" + by (simp add: times_divide_eq le_divide_eq divide_le_eq field_simps) +qed lemma pos_convex_function: fixes f :: "real \ real" diff -r e83695ea0e0a -r b017cfb10df4 src/HOL/Ln.thy --- a/src/HOL/Ln.thy Thu Jun 09 14:16:12 2011 +0200 +++ b/src/HOL/Ln.thy Thu Jun 09 15:14:21 2011 +0200 @@ -387,4 +387,47 @@ finally show ?thesis using b by (simp add: field_simps) qed +lemma ln_le_minus_one: + "0 < x \ ln x \ x - 1" + using exp_ge_add_one_self[of "ln x"] by simp + +lemma ln_eq_minus_one: + assumes "0 < x" "ln x = x - 1" shows "x = 1" +proof - + let "?l y" = "ln y - y + 1" + have D: "\x. 0 < x \ DERIV ?l x :> (1 / x - 1)" + by (auto intro!: DERIV_intros) + + show ?thesis + proof (cases rule: linorder_cases) + assume "x < 1" + from dense[OF `x < 1`] obtain a where "x < a" "a < 1" by blast + from `x < a` have "?l x < ?l a" + proof (rule DERIV_pos_imp_increasing, safe) + fix y assume "x \ y" "y \ a" + with `0 < x` `a < 1` have "0 < 1 / y - 1" "0 < y" + by (auto simp: field_simps) + with D show "\z. DERIV ?l y :> z \ 0 < z" + by auto + qed + also have "\ \ 0" + using ln_le_minus_one `0 < x` `x < a` by (auto simp: field_simps) + finally show "x = 1" using assms by auto + next + assume "1 < x" + from dense[OF `1 < x`] obtain a where "1 < a" "a < x" by blast + from `a < x` have "?l x < ?l a" + proof (rule DERIV_neg_imp_decreasing, safe) + fix y assume "a \ y" "y \ x" + with `1 < a` have "1 / y - 1 < 0" "0 < y" + by (auto simp: field_simps) + with D show "\z. DERIV ?l y :> z \ z < 0" + by blast + qed + also have "\ \ 0" + using ln_le_minus_one `1 < a` by (auto simp: field_simps) + finally show "x = 1" using assms by auto + qed simp +qed + end diff -r e83695ea0e0a -r b017cfb10df4 src/HOL/Multivariate_Analysis/Derivative.thy --- a/src/HOL/Multivariate_Analysis/Derivative.thy Thu Jun 09 14:16:12 2011 +0200 +++ b/src/HOL/Multivariate_Analysis/Derivative.thy Thu Jun 09 15:14:21 2011 +0200 @@ -105,6 +105,22 @@ "a \ s \ open s \ ((f has_derivative f') (at a within s) \ (f has_derivative f') (at a))" by (simp only: at_within_interior interior_open) +lemma has_derivative_right: + fixes f :: "real \ real" and y :: "real" + shows "(f has_derivative (op * y)) (at x within ({x <..} \ I)) \ + ((\t. (f x - f t) / (x - t)) ---> y) (at x within ({x <..} \ I))" +proof - + have "((\t. (f t - (f x + y * (t - x))) / \t - x\) ---> 0) (at x within ({x<..} \ I)) \ + ((\t. (f t - f x) / (t - x) - y) ---> 0) (at x within ({x<..} \ I))" + by (intro Lim_cong_within) (auto simp add: divide.diff divide.add) + also have "\ \ ((\t. (f t - f x) / (t - x)) ---> y) (at x within ({x<..} \ I))" + by (simp add: Lim_null[symmetric]) + also have "\ \ ((\t. (f x - f t) / (x - t)) ---> y) (at x within ({x<..} \ I))" + by (intro Lim_cong_within) (simp_all add: times_divide_eq field_simps) + finally show ?thesis + by (simp add: mult.bounded_linear_right has_derivative_within) +qed + lemma bounded_linear_imp_linear: "bounded_linear f \ linear f" (* TODO: move elsewhere *) proof - assume "bounded_linear f" diff -r e83695ea0e0a -r b017cfb10df4 src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Thu Jun 09 14:16:12 2011 +0200 +++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy Thu Jun 09 15:14:21 2011 +0200 @@ -1122,6 +1122,100 @@ thus ?lhs by (rule Lim_at_within) qed +lemma Lim_within_LIMSEQ: + fixes a :: real and L :: "'a::metric_space" + assumes "\S. (\n. S n \ a \ S n \ T) \ S ----> a \ (\n. X (S n)) ----> L" + shows "(X ---> L) (at a within T)" +proof (rule ccontr) + assume "\ (X ---> L) (at a within T)" + hence "\r>0. \s>0. \x\T. x \ a \ \x - a\ < s \ r \ dist (X x) L" + unfolding tendsto_iff eventually_within dist_norm by (simp add: not_less[symmetric]) + then obtain r where r: "r > 0" "\s. s > 0 \ \x\T-{a}. \x - a\ < s \ dist (X x) L \ r" by blast + + let ?F = "\n::nat. SOME x. x \ T \ x \ a \ \x - a\ < inverse (real (Suc n)) \ dist (X x) L \ r" + have "\n. \x. x \ T \ x \ a \ \x - a\ < inverse (real (Suc n)) \ dist (X x) L \ r" + using r by (simp add: Bex_def) + hence F: "\n. ?F n \ T \ ?F n \ a \ \?F n - a\ < inverse (real (Suc n)) \ dist (X (?F n)) L \ r" + by (rule someI_ex) + hence F1: "\n. ?F n \ T \ ?F n \ a" + and F2: "\n. \?F n - a\ < inverse (real (Suc n))" + and F3: "\n. dist (X (?F n)) L \ r" + by fast+ + + have "?F ----> a" + proof (rule LIMSEQ_I, unfold real_norm_def) + fix e::real + assume "0 < e" + (* choose no such that inverse (real (Suc n)) < e *) + then have "\no. inverse (real (Suc no)) < e" by (rule reals_Archimedean) + then obtain m where nodef: "inverse (real (Suc m)) < e" by auto + show "\no. \n. no \ n --> \?F n - a\ < e" + proof (intro exI allI impI) + fix n + assume mlen: "m \ n" + have "\?F n - a\ < inverse (real (Suc n))" + by (rule F2) + also have "inverse (real (Suc n)) \ inverse (real (Suc m))" + using mlen by auto + also from nodef have + "inverse (real (Suc m)) < e" . + finally show "\?F n - a\ < e" . + qed + qed + moreover note `\S. (\n. S n \ a \ S n \ T) \ S ----> a \ (\n. X (S n)) ----> L` + ultimately have "(\n. X (?F n)) ----> L" using F1 by simp + + moreover have "\ ((\n. X (?F n)) ----> L)" + proof - + { + fix no::nat + obtain n where "n = no + 1" by simp + then have nolen: "no \ n" by simp + (* We prove this by showing that for any m there is an n\m such that |X (?F n) - L| \ r *) + have "dist (X (?F n)) L \ r" + by (rule F3) + with nolen have "\n. no \ n \ dist (X (?F n)) L \ r" by fast + } + then have "(\no. \n. no \ n \ dist (X (?F n)) L \ r)" by simp + with r have "\e>0. (\no. \n. no \ n \ dist (X (?F n)) L \ e)" by auto + thus ?thesis by (unfold LIMSEQ_def, auto simp add: not_less) + qed + ultimately show False by simp +qed + +lemma Lim_right_bound: + fixes f :: "real \ real" + assumes mono: "\a b. a \ I \ b \ I \ x < a \ a \ b \ f a \ f b" + assumes bnd: "\a. a \ I \ x < a \ K \ f a" + shows "(f ---> Inf (f ` ({x<..} \ I))) (at x within ({x<..} \ I))" +proof cases + assume "{x<..} \ I = {}" then show ?thesis by (simp add: Lim_within_empty) +next + assume [simp]: "{x<..} \ I \ {}" + show ?thesis + proof (rule Lim_within_LIMSEQ, safe) + fix S assume S: "\n. S n \ x \ S n \ {x <..} \ I" "S ----> x" + + show "(\n. f (S n)) ----> Inf (f ` ({x<..} \ I))" + proof (rule LIMSEQ_I, rule ccontr) + fix r :: real assume "0 < r" + with Inf_close[of "f ` ({x<..} \ I)" r] + obtain y where y: "x < y" "y \ I" "f y < Inf (f ` ({x <..} \ I)) + r" by auto + from `x < y` have "0 < y - x" by auto + from S(2)[THEN LIMSEQ_D, OF this] + obtain N where N: "\n. N \ n \ \S n - x\ < y - x" by auto + + assume "\ (\N. \n\N. norm (f (S n) - Inf (f ` ({x<..} \ I))) < r)" + moreover have "\n. Inf (f ` ({x<..} \ I)) \ f (S n)" + using S bnd by (intro Inf_lower[where z=K]) auto + ultimately obtain n where n: "N \ n" "r + Inf (f ` ({x<..} \ I)) \ f (S n)" + by (auto simp: not_less field_simps) + with N[OF n(1)] mono[OF _ `y \ I`, of "S n"] S(1)[THEN spec, of n] y + show False by auto + qed + qed +qed + text{* Another limit point characterization. *} lemma islimpt_sequential: @@ -1505,14 +1599,16 @@ (* FIXME: Only one congruence rule for tendsto can be used at a time! *) lemma Lim_cong_within(*[cong add]*): - assumes "\x. x \ a \ f x = g x" - shows "((\x. f x) ---> l) (at a within S) \ ((g ---> l) (at a within S))" + assumes "a = b" "x = y" "S = T" + assumes "\x. x \ b \ x \ T \ f x = g x" + shows "(f ---> x) (at a within S) \ (g ---> y) (at b within T)" unfolding tendsto_def Limits.eventually_within eventually_at_topological using assms by simp lemma Lim_cong_at(*[cong add]*): + assumes "a = b" "x = y" assumes "\x. x \ a \ f x = g x" - shows "((\x. f x) ---> l) (at a) \ ((g ---> l) (at a))" + shows "((\x. f x) ---> x) (at a) \ ((g ---> y) (at a))" unfolding tendsto_def eventually_at_topological using assms by simp diff -r e83695ea0e0a -r b017cfb10df4 src/HOL/Probability/Independent_Family.thy --- a/src/HOL/Probability/Independent_Family.thy Thu Jun 09 14:16:12 2011 +0200 +++ b/src/HOL/Probability/Independent_Family.thy Thu Jun 09 15:14:21 2011 +0200 @@ -117,6 +117,16 @@ using indep[unfolded indep_set_def, THEN indep_setsD, of UNIV "bool_case a b"] ev by (simp add: ac_simps UNIV_bool) +lemma (in prob_space) indep_var_eq: + "indep_var S X T Y \ + (random_variable S X \ random_variable T Y) \ + indep_set + (sigma_sets (space M) { X -` A \ space M | A. A \ sets S}) + (sigma_sets (space M) { Y -` A \ space M | A. A \ sets T})" + unfolding indep_var_def indep_vars_def indep_set_def UNIV_bool + by (intro arg_cong2[where f="op \"] arg_cong2[where f=indep_sets] ext) + (auto split: bool.split) + lemma (in prob_space) assumes indep: "indep_set A B" shows indep_setD_ev1: "A \ events" @@ -491,7 +501,7 @@ proof (simp add: sigma_algebra_iff2, safe) let ?A = "(\n. sigma_sets (space M) (UNION {n..} A))" interpret A: sigma_algebra "\space = space M, sets = A i\" for i by fact - { fix X x assume "X \ ?A" "x \ X" + { fix X x assume "X \ ?A" "x \ X" then have "\n. X \ sigma_sets (space M) (UNION {n..} A)" by auto from this[of 0] have "X \ sigma_sets (space M) (UNION UNIV A)" by simp then have "X \ space M" @@ -572,7 +582,7 @@ show "Int_stable \space = space M, sets = A m\" unfolding Int_stable_def using A.Int by auto qed - also have "(\b. sigma_sets (space M) (\m\bool_case {..n} {Suc n..} b. A m)) = + also have "(\b. sigma_sets (space M) (\m\bool_case {..n} {Suc n..} b. A m)) = bool_case (sigma_sets (space M) (\m\{..n}. A m)) (sigma_sets (space M) (\m\{Suc n..}. A m))" by (auto intro!: ext split: bool.split) finally have indep: "indep_set (sigma_sets (space M) (\m\{..n}. A m)) (sigma_sets (space M) (\m\{Suc n..}. A m))" @@ -732,7 +742,7 @@ by (auto simp del: vimage_Int intro!: exI[of _ "A \ B"] dest: Int_stableD) qed } note indep_sets_sigma_sets_iff[OF this, simp] - + { fix i assume "i \ I" { fix A assume "A \ sets (M' i)" then have "A \ sets (sigma (M' i))" by (auto simp: sets_sigma intro: sigma_sets.Basic) @@ -745,7 +755,7 @@ "space M \ {X i -` A \ space M |A. A \ sets (M' i)}" by (auto intro!: exI[of _ "space (M' i)"]) } note indep_sets_finite[OF I this, simp] - + have "(\A\\ i\I. {X i -` A \ space M |A. A \ sets (M' i)}. prob (INTER I A) = (\j\I. prob (A j))) = (\A\\ i\I. sets (M' i). prob ((\j\I. X j -` A j) \ space M) = (\x\I. prob (X x -` A x \ space M)))" (is "?L = ?R") @@ -847,7 +857,7 @@ by (simp_all add: product_algebra_def) show "A \ sets (sigma P.G)" using `A \ sets P.P` by (simp add: product_algebra_def) - + fix E assume E: "E \ sets P.G" then have "E \ sets P.P" by (simp add: sets_sigma sigma_sets.Basic product_algebra_def) @@ -915,10 +925,67 @@ finally show ?thesis . qed +lemma (in prob_space) + assumes "indep_var S X T Y" + shows indep_var_rv1: "random_variable S X" + and indep_var_rv2: "random_variable T Y" +proof - + have "\i\UNIV. random_variable (bool_case S T i) (bool_case X Y i)" + using assms unfolding indep_var_def indep_vars_def by auto + then show "random_variable S X" "random_variable T Y" + unfolding UNIV_bool by auto +qed + lemma (in prob_space) indep_var_distributionD: - assumes "indep_var Ma A Mb B" - assumes "Xa \ sets Ma" "Xb \ sets Mb" - shows "joint_distribution A B (Xa \ Xb) = distribution A Xa * distribution B Xb" - unfolding distribution_def using assms by (rule indep_varD) + assumes indep: "indep_var S X T Y" + defines "P \ S\measure := extreal\distribution X\ \\<^isub>M T\measure := extreal\distribution Y\" + assumes "A \ sets P" + shows "joint_distribution X Y A = finite_measure.\' P A" +proof - + from indep have rvs: "random_variable S X" "random_variable T Y" + by (blast dest: indep_var_rv1 indep_var_rv2)+ + + let ?S = "S\measure := extreal\distribution X\" + let ?T = "T\measure := extreal\distribution Y\" + interpret X: prob_space ?S by (rule distribution_prob_space) fact + interpret Y: prob_space ?T by (rule distribution_prob_space) fact + interpret XY: pair_prob_space ?S ?T by default + + let ?J = "XY.P\ measure := extreal \ joint_distribution X Y \" + interpret J: prob_space ?J + by (rule joint_distribution_prob_space) (simp_all add: rvs) + + have "finite_measure.\' (XY.P\ measure := extreal \ joint_distribution X Y \) A = XY.\' A" + proof (rule prob_space_unique_Int_stable) + show "Int_stable (pair_measure_generator ?S ?T)" (is "Int_stable ?P") + by fact + show "space ?P \ sets ?P" + unfolding space_pair_measure[simplified pair_measure_def space_sigma] + using X.top Y.top by (auto intro!: pair_measure_generatorI) + + show "prob_space ?J" by default + show "space ?J = space ?P" + by (simp add: pair_measure_generator_def space_pair_measure) + show "sets ?J = sets (sigma ?P)" + by (simp add: pair_measure_def) + + show "prob_space XY.P" by default + show "space XY.P = space ?P" "sets XY.P = sets (sigma ?P)" + by (simp_all add: pair_measure_generator_def pair_measure_def) + + show "A \ sets (sigma ?P)" + using `A \ sets P` unfolding P_def pair_measure_def by simp + + fix X assume "X \ sets ?P" + then obtain A B where "A \ sets S" "B \ sets T" "X = A \ B" + by (auto simp: sets_pair_measure_generator) + then show "J.\' X = XY.\' X" + unfolding J.\'_def XY.\'_def using indep + by (simp add: XY.pair_measure_times) + (simp add: distribution_def indep_varD) + qed + then show ?thesis + using `A \ sets P` unfolding P_def J.\'_def XY.\'_def by simp +qed end diff -r e83695ea0e0a -r b017cfb10df4 src/HOL/Probability/Information.thy --- a/src/HOL/Probability/Information.thy Thu Jun 09 14:16:12 2011 +0200 +++ b/src/HOL/Probability/Information.thy Thu Jun 09 15:14:21 2011 +0200 @@ -7,14 +7,10 @@ theory Information imports - Probability_Measure + Independent_Family "~~/src/HOL/Library/Convex" begin -lemma (in prob_space) not_zero_less_distribution[simp]: - "(\ 0 < distribution X A) \ distribution X A = 0" - using distribution_positive[of X A] by arith - lemma log_le: "1 < a \ 0 < x \ x \ y \ log a x \ log a y" by (subst log_le_cancel_iff) auto @@ -175,7 +171,211 @@ Kullback$-$Leibler distance. *} definition - "KL_divergence b M \ = \x. log b (real (RN_deriv M \ x)) \M\measure := \\" + "entropy_density b M \ = log b \ real \ RN_deriv M \" + +definition + "KL_divergence b M \ = integral\<^isup>L (M\measure := \\) (entropy_density b M \)" + +lemma (in information_space) measurable_entropy_density: + assumes ps: "prob_space (M\measure := \\)" + assumes ac: "absolutely_continuous \" + shows "entropy_density b M \ \ borel_measurable M" +proof - + interpret \: prob_space "M\measure := \\" by fact + have "measure_space (M\measure := \\)" by fact + from RN_deriv[OF this ac] b_gt_1 show ?thesis + unfolding entropy_density_def + by (intro measurable_comp) auto +qed + +lemma (in information_space) KL_gt_0: + assumes ps: "prob_space (M\measure := \\)" + assumes ac: "absolutely_continuous \" + assumes int: "integrable (M\ measure := \ \) (entropy_density b M \)" + assumes A: "A \ sets M" "\ A \ \ A" + shows "0 < KL_divergence b M \" +proof - + interpret \: prob_space "M\measure := \\" by fact + have ms: "measure_space (M\measure := \\)" by default + have fms: "finite_measure (M\measure := \\)" by default + note RN = RN_deriv[OF ms ac] + + from real_RN_deriv[OF fms ac] guess D . note D = this + with absolutely_continuous_AE[OF ms] ac + have D\: "AE x in M\measure := \\. RN_deriv M \ x = extreal (D x)" + by auto + + def f \ "\x. if D x = 0 then 1 else 1 / D x" + with D have f_borel: "f \ borel_measurable M" + by (auto intro!: measurable_If) + + have "KL_divergence b M \ = 1 / ln b * (\ x. ln b * entropy_density b M \ x \M\measure := \\)" + unfolding KL_divergence_def using int b_gt_1 + by (simp add: integral_cmult) + + { fix A assume "A \ sets M" + with RN D have "\.\ A = (\\<^isup>+ x. extreal (D x) * indicator A x \M)" + by (auto intro!: positive_integral_cong_AE) } + note D_density = this + + have ln_entropy: "(\x. ln b * entropy_density b M \ x) \ borel_measurable M" + using measurable_entropy_density[OF ps ac] by auto + + have "integrable (M\measure := \\) (\x. ln b * entropy_density b M \ x)" + using int by auto + moreover have "integrable (M\measure := \\) (\x. ln b * entropy_density b M \ x) \ + integrable M (\x. D x * (ln b * entropy_density b M \ x))" + using D D_density ln_entropy + by (intro integral_translated_density) auto + ultimately have M_int: "integrable M (\x. D x * (ln b * entropy_density b M \ x))" + by simp + + have D_neg: "(\\<^isup>+ x. extreal (- D x) \M) = 0" + using D by (subst positive_integral_0_iff_AE) auto + + have "(\\<^isup>+ x. extreal (D x) \M) = \ (space M)" + using RN D by (auto intro!: positive_integral_cong_AE) + then have D_pos: "(\\<^isup>+ x. extreal (D x) \M) = 1" + using \.measure_space_1 by simp + + have "integrable M D" + using D_pos D_neg D by (auto simp: integrable_def) + + have "integral\<^isup>L M D = 1" + using D_pos D_neg by (auto simp: lebesgue_integral_def) + + let ?D_set = "{x\space M. D x \ 0}" + have [simp, intro]: "?D_set \ sets M" + using D by (auto intro: sets_Collect) + + have "0 \ 1 - \' ?D_set" + using prob_le_1 by (auto simp: field_simps) + also have "\ = (\ x. D x - indicator ?D_set x \M)" + using `integrable M D` `integral\<^isup>L M D = 1` + by (simp add: \'_def) + also have "\ < (\ x. D x * (ln b * entropy_density b M \ x) \M)" + proof (rule integral_less_AE) + show "integrable M (\x. D x - indicator ?D_set x)" + using `integrable M D` + by (intro integral_diff integral_indicator) auto + next + show "integrable M (\x. D x * (ln b * entropy_density b M \ x))" + by fact + next + show "\ {x\space M. D x \ 1 \ D x \ 0} \ 0" + proof + assume eq_0: "\ {x\space M. D x \ 1 \ D x \ 0} = 0" + then have disj: "AE x. D x = 1 \ D x = 0" + using D(1) by (auto intro!: AE_I[OF subset_refl] sets_Collect) + + have "\ {x\space M. D x = 1} = (\\<^isup>+ x. indicator {x\space M. D x = 1} x \M)" + using D(1) by auto + also have "\ = (\\<^isup>+ x. extreal (D x) * indicator {x\space M. D x \ 0} x \M)" + using disj by (auto intro!: positive_integral_cong_AE simp: indicator_def one_extreal_def) + also have "\ = \ {x\space M. D x \ 0}" + using D(1) D_density by auto + also have "\ = \ (space M)" + using D_density D(1) by (auto intro!: positive_integral_cong simp: indicator_def) + finally have "AE x. D x = 1" + using D(1) \.measure_space_1 by (intro AE_I_eq_1) auto + then have "(\\<^isup>+x. indicator A x\M) = (\\<^isup>+x. extreal (D x) * indicator A x\M)" + by (intro positive_integral_cong_AE) (auto simp: one_extreal_def[symmetric]) + also have "\ = \ A" + using `A \ sets M` D_density by simp + finally show False using `A \ sets M` `\ A \ \ A` by simp + qed + show "{x\space M. D x \ 1 \ D x \ 0} \ sets M" + using D(1) by (auto intro: sets_Collect) + + show "AE t. t \ {x\space M. D x \ 1 \ D x \ 0} \ + D t - indicator ?D_set t \ D t * (ln b * entropy_density b M \ t)" + using D(2) + proof (elim AE_mp, safe intro!: AE_I2) + fix t assume Dt: "t \ space M" "D t \ 1" "D t \ 0" + and RN: "RN_deriv M \ t = extreal (D t)" + and eq: "D t - indicator ?D_set t = D t * (ln b * entropy_density b M \ t)" + + have "D t - 1 = D t - indicator ?D_set t" + using Dt by simp + also note eq + also have "D t * (ln b * entropy_density b M \ t) = - D t * ln (1 / D t)" + using RN b_gt_1 `D t \ 0` `0 \ D t` + by (simp add: entropy_density_def log_def ln_div less_le) + finally have "ln (1 / D t) = 1 / D t - 1" + using `D t \ 0` by (auto simp: field_simps) + from ln_eq_minus_one[OF _ this] `D t \ 0` `0 \ D t` `D t \ 1` + show False by auto + qed + + show "AE t. D t - indicator ?D_set t \ D t * (ln b * entropy_density b M \ t)" + using D(2) + proof (elim AE_mp, intro AE_I2 impI) + fix t assume "t \ space M" and RN: "RN_deriv M \ t = extreal (D t)" + show "D t - indicator ?D_set t \ D t * (ln b * entropy_density b M \ t)" + proof cases + assume asm: "D t \ 0" + then have "0 < D t" using `0 \ D t` by auto + then have "0 < 1 / D t" by auto + have "D t - indicator ?D_set t \ - D t * (1 / D t - 1)" + using asm `t \ space M` by (simp add: field_simps) + also have "- D t * (1 / D t - 1) \ - D t * ln (1 / D t)" + using ln_le_minus_one `0 < 1 / D t` by (intro mult_left_mono_neg) auto + also have "\ = D t * (ln b * entropy_density b M \ t)" + using `0 < D t` RN b_gt_1 + by (simp_all add: log_def ln_div entropy_density_def) + finally show ?thesis by simp + qed simp + qed + qed + also have "\ = (\ x. ln b * entropy_density b M \ x \M\measure := \\)" + using D D_density ln_entropy + by (intro integral_translated_density[symmetric]) auto + also have "\ = ln b * (\ x. entropy_density b M \ x \M\measure := \\)" + using int by (rule \.integral_cmult) + finally show "0 < KL_divergence b M \" + using b_gt_1 by (auto simp: KL_divergence_def zero_less_mult_iff) +qed + +lemma (in sigma_finite_measure) KL_eq_0: + assumes eq: "\A\sets M. \ A = measure M A" + shows "KL_divergence b M \ = 0" +proof - + have "AE x. 1 = RN_deriv M \ x" + proof (rule RN_deriv_unique) + show "measure_space (M\measure := \\)" + using eq by (intro measure_space_cong) auto + show "absolutely_continuous \" + unfolding absolutely_continuous_def using eq by auto + show "(\x. 1) \ borel_measurable M" "AE x. 0 \ (1 :: extreal)" by auto + fix A assume "A \ sets M" + with eq show "\ A = \\<^isup>+ x. 1 * indicator A x \M" by simp + qed + then have "AE x. log b (real (RN_deriv M \ x)) = 0" + by (elim AE_mp) simp + from integral_cong_AE[OF this] + have "integral\<^isup>L M (entropy_density b M \) = 0" + by (simp add: entropy_density_def comp_def) + with eq show "KL_divergence b M \ = 0" + unfolding KL_divergence_def + by (subst integral_cong_measure) auto +qed + +lemma (in information_space) KL_eq_0_imp: + assumes ps: "prob_space (M\measure := \\)" + assumes ac: "absolutely_continuous \" + assumes int: "integrable (M\ measure := \ \) (entropy_density b M \)" + assumes KL: "KL_divergence b M \ = 0" + shows "\A\sets M. \ A = \ A" + by (metis less_imp_neq KL_gt_0 assms) + +lemma (in information_space) KL_ge_0: + assumes ps: "prob_space (M\measure := \\)" + assumes ac: "absolutely_continuous \" + assumes int: "integrable (M\ measure := \ \) (entropy_density b M \)" + shows "0 \ KL_divergence b M \" + using KL_eq_0 KL_gt_0[OF ps ac int] + by (cases "\A\sets M. \ A = measure M A") (auto simp: le_less) + lemma (in sigma_finite_measure) KL_divergence_vimage: assumes T: "T \ measure_preserving M M'" @@ -209,7 +409,7 @@ have AE: "AE x. RN_deriv M' \' (T x) = RN_deriv M \ x" by (rule RN_deriv_vimage[OF T T' inv \']) show ?thesis - unfolding KL_divergence_def + unfolding KL_divergence_def entropy_density_def comp_def proof (subst \'.integral_vimage[OF sa T']) show "(\x. log b (real (RN_deriv M \ x))) \ borel_measurable (M\measure := \\)" by (auto intro!: RN_deriv[OF M ac] borel_measurable_log[OF _ `1 < b`]) @@ -233,9 +433,9 @@ proof - interpret \: measure_space ?\ by fact have "KL_divergence b M \ = \x. log b (real (RN_deriv N \' x)) \?\" - by (simp cong: RN_deriv_cong \.integral_cong add: KL_divergence_def) + by (simp cong: RN_deriv_cong \.integral_cong add: KL_divergence_def entropy_density_def) also have "\ = KL_divergence b N \'" - by (auto intro!: \.integral_cong_measure[symmetric] simp: KL_divergence_def) + by (auto intro!: \.integral_cong_measure[symmetric] simp: KL_divergence_def entropy_density_def comp_def) finally show ?thesis . qed @@ -243,7 +443,7 @@ assumes v: "finite_measure_space (M\measure := \\)" assumes ac: "absolutely_continuous \" shows "KL_divergence b M \ = (\x\space M. real (\ {x}) * log b (real (\ {x}) / real (\ {x})))" (is "_ = ?sum") -proof (simp add: KL_divergence_def finite_measure_space.integral_finite_singleton[OF v]) +proof (simp add: KL_divergence_def finite_measure_space.integral_finite_singleton[OF v] entropy_density_def) interpret v: finite_measure_space "M\measure := \\" by fact have ms: "measure_space (M\measure := \\)" by default show "(\x \ space M. log b (real (RN_deriv M \ x)) * real (\ {x})) = ?sum" @@ -257,27 +457,10 @@ and "1 < b" shows "0 \ KL_divergence b M \" proof - + interpret information_space M by default fact interpret v: finite_prob_space "M\measure := \\" by fact - have ms: "finite_measure_space (M\measure := \\)" by default - - have "- (KL_divergence b M \) \ log b (\x\space M. real (\ {x}))" - proof (subst KL_divergence_eq_finite[OF ms ac], safe intro!: log_setsum_divide not_empty) - show "finite (space M)" using finite_space by simp - show "1 < b" by fact - show "(\x\space M. real (\ {x})) = 1" - using v.finite_sum_over_space_eq_1 by (simp add: v.\'_def) - - fix x assume "x \ space M" - then have x: "{x} \ sets M" unfolding sets_eq_Pow by auto - { assume "0 < real (\ {x})" - then have "\ {x} \ 0" by auto - then have "\ {x} \ 0" - using ac[unfolded absolutely_continuous_def, THEN bspec, of "{x}"] x by auto - thus "0 < real (\ {x})" using real_measure[OF x] by auto } - show "0 \ real (\ {x})" "0 \ real (\ {x})" - using real_measure[OF x] v.real_measure[of "{x}"] x by auto - qed - thus "0 \ KL_divergence b M \" using finite_sum_over_space_eq_1 by (simp add: \'_def) + have ps: "prob_space (M\measure := \\)" by default + from KL_ge_0[OF this ac v.integral_finite_singleton(1)] show ?thesis . qed subsection {* Mutual Information *} @@ -287,6 +470,163 @@ KL_divergence b (S\measure := extreal\distribution X\ \\<^isub>M T\measure := extreal\distribution Y\) (extreal\joint_distribution X Y)" +lemma (in information_space) + fixes S T X Y + defines "P \ S\measure := extreal\distribution X\ \\<^isub>M T\measure := extreal\distribution Y\" + shows "indep_var S X T Y \ + (random_variable S X \ random_variable T Y \ + measure_space.absolutely_continuous P (extreal\joint_distribution X Y) \ + integrable (P\measure := (extreal\joint_distribution X Y)\) + (entropy_density b P (extreal\joint_distribution X Y)) \ + mutual_information b S T X Y = 0)" +proof safe + assume indep: "indep_var S X T Y" + then have "random_variable S X" "random_variable T Y" + by (blast dest: indep_var_rv1 indep_var_rv2)+ + then show "sigma_algebra S" "X \ measurable M S" "sigma_algebra T" "Y \ measurable M T" + by blast+ + + interpret X: prob_space "S\measure := extreal\distribution X\" + by (rule distribution_prob_space) fact + interpret Y: prob_space "T\measure := extreal\distribution Y\" + by (rule distribution_prob_space) fact + interpret XY: pair_prob_space "S\measure := extreal\distribution X\" "T\measure := extreal\distribution Y\" by default + interpret XY: information_space XY.P b by default (rule b_gt_1) + + let ?J = "XY.P\ measure := (extreal\joint_distribution X Y) \" + { fix A assume "A \ sets XY.P" + then have "extreal (joint_distribution X Y A) = XY.\ A" + using indep_var_distributionD[OF indep] + by (simp add: XY.P.finite_measure_eq) } + note j_eq = this + + interpret J: prob_space ?J + using j_eq by (intro XY.prob_space_cong) auto + + have ac: "XY.absolutely_continuous (extreal\joint_distribution X Y)" + by (simp add: XY.absolutely_continuous_def j_eq) + then show "measure_space.absolutely_continuous P (extreal\joint_distribution X Y)" + unfolding P_def . + + have ed: "entropy_density b XY.P (extreal\joint_distribution X Y) \ borel_measurable XY.P" + by (rule XY.measurable_entropy_density) (default | fact)+ + + have "AE x in XY.P. 1 = RN_deriv XY.P (extreal\joint_distribution X Y) x" + proof (rule XY.RN_deriv_unique[OF _ ac]) + show "measure_space ?J" by default + fix A assume "A \ sets XY.P" + then show "(extreal\joint_distribution X Y) A = (\\<^isup>+ x. 1 * indicator A x \XY.P)" + by (simp add: j_eq) + qed (insert XY.measurable_const[of 1 borel], auto) + then have ae_XY: "AE x in XY.P. entropy_density b XY.P (extreal\joint_distribution X Y) x = 0" + by (elim XY.AE_mp) (simp add: entropy_density_def) + have ae_J: "AE x in ?J. entropy_density b XY.P (extreal\joint_distribution X Y) x = 0" + proof (rule XY.absolutely_continuous_AE) + show "measure_space ?J" by default + show "XY.absolutely_continuous (measure ?J)" + using ac by simp + qed (insert ae_XY, simp_all) + then show "integrable (P\measure := (extreal\joint_distribution X Y)\) + (entropy_density b P (extreal\joint_distribution X Y))" + unfolding P_def + using ed XY.measurable_const[of 0 borel] + by (subst J.integrable_cong_AE) auto + + show "mutual_information b S T X Y = 0" + unfolding mutual_information_def KL_divergence_def P_def + by (subst J.integral_cong_AE[OF ae_J]) simp +next + assume "sigma_algebra S" "X \ measurable M S" "sigma_algebra T" "Y \ measurable M T" + then have rvs: "random_variable S X" "random_variable T Y" by blast+ + + interpret X: prob_space "S\measure := extreal\distribution X\" + by (rule distribution_prob_space) fact + interpret Y: prob_space "T\measure := extreal\distribution Y\" + by (rule distribution_prob_space) fact + interpret XY: pair_prob_space "S\measure := extreal\distribution X\" "T\measure := extreal\distribution Y\" by default + interpret XY: information_space XY.P b by default (rule b_gt_1) + + let ?J = "XY.P\ measure := (extreal\joint_distribution X Y) \" + interpret J: prob_space ?J + using rvs by (intro joint_distribution_prob_space) auto + + assume ac: "measure_space.absolutely_continuous P (extreal\joint_distribution X Y)" + assume int: "integrable (P\measure := (extreal\joint_distribution X Y)\) + (entropy_density b P (extreal\joint_distribution X Y))" + assume I_eq_0: "mutual_information b S T X Y = 0" + + have eq: "\A\sets XY.P. (extreal \ joint_distribution X Y) A = XY.\ A" + proof (rule XY.KL_eq_0_imp) + show "prob_space ?J" by default + show "XY.absolutely_continuous (extreal\joint_distribution X Y)" + using ac by (simp add: P_def) + show "integrable ?J (entropy_density b XY.P (extreal\joint_distribution X Y))" + using int by (simp add: P_def) + show "KL_divergence b XY.P (extreal\joint_distribution X Y) = 0" + using I_eq_0 unfolding mutual_information_def by (simp add: P_def) + qed + + { fix S X assume "sigma_algebra S" + interpret S: sigma_algebra S by fact + have "Int_stable \space = space M, sets = {X -` A \ space M |A. A \ sets S}\" + proof (safe intro!: Int_stableI) + fix A B assume "A \ sets S" "B \ sets S" + then show "\C. (X -` A \ space M) \ (X -` B \ space M) = (X -` C \ space M) \ C \ sets S" + by (intro exI[of _ "A \ B"]) auto + qed } + note Int_stable = this + + show "indep_var S X T Y" unfolding indep_var_eq + proof (intro conjI indep_set_sigma_sets Int_stable) + show "indep_set {X -` A \ space M |A. A \ sets S} {Y -` A \ space M |A. A \ sets T}" + proof (safe intro!: indep_setI) + { fix A assume "A \ sets S" then show "X -` A \ space M \ sets M" + using `X \ measurable M S` by (auto intro: measurable_sets) } + { fix A assume "A \ sets T" then show "Y -` A \ space M \ sets M" + using `Y \ measurable M T` by (auto intro: measurable_sets) } + next + fix A B assume ab: "A \ sets S" "B \ sets T" + have "extreal (prob ((X -` A \ space M) \ (Y -` B \ space M))) = + extreal (joint_distribution X Y (A \ B))" + unfolding distribution_def + by (intro arg_cong[where f="\C. extreal (prob C)"]) auto + also have "\ = XY.\ (A \ B)" + using ab eq by (auto simp: XY.finite_measure_eq) + also have "\ = extreal (distribution X A) * extreal (distribution Y B)" + using ab by (simp add: XY.pair_measure_times) + finally show "prob ((X -` A \ space M) \ (Y -` B \ space M)) = + prob (X -` A \ space M) * prob (Y -` B \ space M)" + unfolding distribution_def by simp + qed + qed fact+ +qed + +lemma (in information_space) mutual_information_commute_generic: + assumes X: "random_variable S X" and Y: "random_variable T Y" + assumes ac: "measure_space.absolutely_continuous + (S\measure := extreal\distribution X\ \\<^isub>M T\measure := extreal\distribution Y\) (extreal\joint_distribution X Y)" + shows "mutual_information b S T X Y = mutual_information b T S Y X" +proof - + let ?S = "S\measure := extreal\distribution X\" and ?T = "T\measure := extreal\distribution Y\" + interpret S: prob_space ?S using X by (rule distribution_prob_space) + interpret T: prob_space ?T using Y by (rule distribution_prob_space) + interpret P: pair_prob_space ?S ?T .. + interpret Q: pair_prob_space ?T ?S .. + show ?thesis + unfolding mutual_information_def + proof (intro Q.KL_divergence_vimage[OF Q.measure_preserving_swap _ _ _ _ ac b_gt_1]) + show "(\(x,y). (y,x)) \ measure_preserving + (P.P \ measure := extreal\joint_distribution X Y\) (Q.P \ measure := extreal\joint_distribution Y X\)" + using X Y unfolding measurable_def + unfolding measure_preserving_def using P.pair_sigma_algebra_swap_measurable + by (auto simp add: space_pair_measure distribution_def intro!: arg_cong[where f=\']) + have "prob_space (P.P\ measure := extreal\joint_distribution X Y\)" + using X Y by (auto intro!: distribution_prob_space random_variable_pairI) + then show "measure_space (P.P\ measure := extreal\joint_distribution X Y\)" + unfolding prob_space_def by simp + qed auto +qed + definition (in prob_space) "entropy b s X = mutual_information b s s X X" @@ -356,32 +696,6 @@ unfolding mutual_information_def . qed -lemma (in information_space) mutual_information_commute_generic: - assumes X: "random_variable S X" and Y: "random_variable T Y" - assumes ac: "measure_space.absolutely_continuous - (S\measure := extreal\distribution X\ \\<^isub>M T\measure := extreal\distribution Y\) (extreal\joint_distribution X Y)" - shows "mutual_information b S T X Y = mutual_information b T S Y X" -proof - - let ?S = "S\measure := extreal\distribution X\" and ?T = "T\measure := extreal\distribution Y\" - interpret S: prob_space ?S using X by (rule distribution_prob_space) - interpret T: prob_space ?T using Y by (rule distribution_prob_space) - interpret P: pair_prob_space ?S ?T .. - interpret Q: pair_prob_space ?T ?S .. - show ?thesis - unfolding mutual_information_def - proof (intro Q.KL_divergence_vimage[OF Q.measure_preserving_swap _ _ _ _ ac b_gt_1]) - show "(\(x,y). (y,x)) \ measure_preserving - (P.P \ measure := extreal\joint_distribution X Y\) (Q.P \ measure := extreal\joint_distribution Y X\)" - using X Y unfolding measurable_def - unfolding measure_preserving_def using P.pair_sigma_algebra_swap_measurable - by (auto simp add: space_pair_measure distribution_def intro!: arg_cong[where f=\']) - have "prob_space (P.P\ measure := extreal\joint_distribution X Y\)" - using X Y by (auto intro!: distribution_prob_space random_variable_pairI) - then show "measure_space (P.P\ measure := extreal\joint_distribution X Y\)" - unfolding prob_space_def by simp - qed auto -qed - lemma (in information_space) mutual_information_commute: assumes X: "finite_random_variable S X" and Y: "finite_random_variable T Y" shows "mutual_information b S T X Y = mutual_information b T S Y X" diff -r e83695ea0e0a -r b017cfb10df4 src/HOL/Probability/Lebesgue_Integration.thy --- a/src/HOL/Probability/Lebesgue_Integration.thy Thu Jun 09 14:16:12 2011 +0200 +++ b/src/HOL/Probability/Lebesgue_Integration.thy Thu Jun 09 15:14:21 2011 +0200 @@ -1712,6 +1712,12 @@ shows "integral\<^isup>L N f = integral\<^isup>L M f" by (simp add: positive_integral_cong_measure[OF assms] lebesgue_integral_def) +lemma (in measure_space) integrable_cong_measure: + assumes "\A. A \ sets M \ measure N A = \ A" "sets N = sets M" "space N = space M" + shows "integrable N f \ integrable M f" + using assms + by (simp add: positive_integral_cong_measure[OF assms] integrable_def measurable_def) + lemma (in measure_space) integral_cong_AE: assumes cong: "AE x. f x = g x" shows "integral\<^isup>L M f = integral\<^isup>L M g" @@ -1722,6 +1728,18 @@ unfolding *[THEN positive_integral_cong_AE] lebesgue_integral_def .. qed +lemma (in measure_space) integrable_cong_AE: + assumes borel: "f \ borel_measurable M" "g \ borel_measurable M" + assumes "AE x. f x = g x" + shows "integrable M f = integrable M g" +proof - + have "(\\<^isup>+ x. extreal (f x) \M) = (\\<^isup>+ x. extreal (g x) \M)" + "(\\<^isup>+ x. extreal (- f x) \M) = (\\<^isup>+ x. extreal (- g x) \M)" + using assms by (auto intro!: positive_integral_cong_AE) + with assms show ?thesis + by (auto simp: integrable_def) +qed + lemma (in measure_space) integrable_cong: "(\x. x \ space M \ f x = g x) \ integrable M f \ integrable M g" by (simp cong: positive_integral_cong measurable_cong add: integrable_def) @@ -1767,6 +1785,48 @@ by (auto simp: borel[THEN positive_integral_vimage[OF T]]) qed +lemma (in measure_space) integral_translated_density: + assumes f: "f \ borel_measurable M" "AE x. 0 \ f x" + and g: "g \ borel_measurable M" + and N: "space N = space M" "sets N = sets M" + and density: "\A. A \ sets M \ measure N A = (\\<^isup>+ x. f x * indicator A x \M)" + (is "\A. _ \ _ = ?d A") + shows "integral\<^isup>L N g = (\ x. f x * g x \M)" (is ?integral) + and "integrable N g = integrable M (\x. f x * g x)" (is ?integrable) +proof - + from f have ms: "measure_space (M\measure := ?d\)" + by (intro measure_space_density[where u="\x. extreal (f x)"]) auto + + from ms density N have "(\\<^isup>+ x. g x \N) = (\\<^isup>+ x. max 0 (extreal (g x)) \M\measure := ?d\)" + unfolding positive_integral_max_0 + by (intro measure_space.positive_integral_cong_measure) auto + also have "\ = (\\<^isup>+ x. extreal (f x) * max 0 (extreal (g x)) \M)" + using f g by (intro positive_integral_translated_density) auto + also have "\ = (\\<^isup>+ x. max 0 (extreal (f x * g x)) \M)" + using f by (intro positive_integral_cong_AE) + (auto simp: extreal_max_0 zero_le_mult_iff split: split_max) + finally have pos: "(\\<^isup>+ x. g x \N) = (\\<^isup>+ x. f x * g x \M)" + by (simp add: positive_integral_max_0) + + from ms density N have "(\\<^isup>+ x. - (g x) \N) = (\\<^isup>+ x. max 0 (extreal (- g x)) \M\measure := ?d\)" + unfolding positive_integral_max_0 + by (intro measure_space.positive_integral_cong_measure) auto + also have "\ = (\\<^isup>+ x. extreal (f x) * max 0 (extreal (- g x)) \M)" + using f g by (intro positive_integral_translated_density) auto + also have "\ = (\\<^isup>+ x. max 0 (extreal (- f x * g x)) \M)" + using f by (intro positive_integral_cong_AE) + (auto simp: extreal_max_0 mult_le_0_iff split: split_max) + finally have neg: "(\\<^isup>+ x. - g x \N) = (\\<^isup>+ x. - (f x * g x) \M)" + by (simp add: positive_integral_max_0) + + have g_N: "g \ borel_measurable N" + using g N unfolding measurable_def by simp + + show ?integral ?integrable + unfolding lebesgue_integral_def integrable_def + using pos neg f g g_N by auto +qed + lemma (in measure_space) integral_minus[intro, simp]: assumes "integrable M f" shows "integrable M (\x. - f x)" "(\x. - f x \M) = - integral\<^isup>L M f" @@ -2221,9 +2281,14 @@ by (simp add: \'_def lebesgue_integral_def positive_integral_const_If) qed +lemma indicator_less[simp]: + "indicator A x \ (indicator B x::extreal) \ (x \ A \ x \ B)" + by (simp add: indicator_def not_le) + lemma (in finite_measure) integral_less_AE: assumes int: "integrable M X" "integrable M Y" - assumes gt: "AE x. X x < Y x" and neq0: "\ (space M) \ 0" + assumes A: "\ A \ 0" "A \ sets M" "AE x. x \ A \ X x \ Y x" + assumes gt: "AE x. X x \ Y x" shows "integral\<^isup>L M X < integral\<^isup>L M Y" proof - have "integral\<^isup>L M X \ integral\<^isup>L M Y" @@ -2231,24 +2296,30 @@ moreover have "integral\<^isup>L M X \ integral\<^isup>L M Y" proof - from gt have "AE x. Y x - X x \ 0" - by auto - then have \: "\ {x\space M. Y x - X x \ 0} = \ (space M)" - using int - by (intro AE_measure borel_measurable_neq) (auto simp add: integrable_def) - assume eq: "integral\<^isup>L M X = integral\<^isup>L M Y" have "integral\<^isup>L M (\x. \Y x - X x\) = integral\<^isup>L M (\x. Y x - X x)" using gt by (intro integral_cong_AE) auto also have "\ = 0" using eq int by simp - finally show False - using \ int neq0 - by (subst (asm) integral_0_iff) auto + finally have "\ {x \ space M. Y x - X x \ 0} = 0" + using int by (simp add: integral_0_iff) + moreover + have "(\\<^isup>+x. indicator A x \M) \ (\\<^isup>+x. indicator {x \ space M. Y x - X x \ 0} x \M)" + using A by (intro positive_integral_mono_AE) auto + then have "\ A \ \ {x \ space M. Y x - X x \ 0}" + using int A by (simp add: integrable_def) + moreover note `\ A \ 0` positive_measure[OF `A \ sets M`] + ultimately show False by auto qed ultimately show ?thesis by auto qed +lemma (in finite_measure) integral_less_AE_space: + assumes int: "integrable M X" "integrable M Y" + assumes gt: "AE x. X x < Y x" "\ (space M) \ 0" + shows "integral\<^isup>L M X < integral\<^isup>L M Y" + using gt by (intro integral_less_AE[OF int, where A="space M"]) auto + lemma (in measure_space) integral_dominated_convergence: assumes u: "\i. integrable M (u i)" and bound: "\x j. x\space M \ \u j x\ \ w x" and w: "integrable M w" diff -r e83695ea0e0a -r b017cfb10df4 src/HOL/Probability/Probability_Measure.thy --- a/src/HOL/Probability/Probability_Measure.thy Thu Jun 09 14:16:12 2011 +0200 +++ b/src/HOL/Probability/Probability_Measure.thy Thu Jun 09 15:14:21 2011 +0200 @@ -28,6 +28,14 @@ abbreviation (in prob_space) "joint_distribution X Y \ distribution (\x. (X x, Y x))" +lemma (in prob_space) prob_space_cong: + assumes "\A. A \ sets M \ measure N A = \ A" "space N = space M" "sets N = sets M" + shows "prob_space N" +proof - + interpret N: measure_space N by (intro measure_space_cong assms) + show ?thesis by default (insert assms measure_space_1, simp) +qed + lemma (in prob_space) distribution_cong: assumes "\x. x \ space M \ X x = Y x" shows "distribution X = distribution Y" @@ -54,15 +62,30 @@ lemma (in prob_space) distribution_positive[simp, intro]: "0 \ distribution X A" unfolding distribution_def by auto +lemma (in prob_space) not_zero_less_distribution[simp]: + "(\ 0 < distribution X A) \ distribution X A = 0" + using distribution_positive[of X A] by arith + 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 prob_space) not_empty: "space M \ {}" + using prob_space empty_measure' by auto + lemma (in prob_space) measure_le_1: "X \ sets M \ \ X \ 1" unfolding measure_space_1[symmetric] using sets_into_space by (intro measure_mono) auto +lemma (in prob_space) AE_I_eq_1: + assumes "\ {x\space M. P x} = 1" "{x\space M. P x} \ sets M" + shows "AE x. P x" +proof (rule AE_I) + show "\ (space M - {x \ space M. P x}) = 0" + using assms measure_space_1 by (simp add: measure_compl) +qed (insert assms, auto) + lemma (in prob_space) distribution_1: "distribution X A \ 1" unfolding distribution_def by simp @@ -245,6 +268,146 @@ using finite_measure_eq[of "X -` A \ space M"] by (auto simp: measurable_sets distribution_def) +lemma (in prob_space) expectation_less: + assumes [simp]: "integrable M X" + assumes gt: "\x\space M. X x < b" + shows "expectation X < b" +proof - + have "expectation X < expectation (\x. b)" + using gt measure_space_1 + by (intro integral_less_AE_space) auto + then show ?thesis using prob_space by simp +qed + +lemma (in prob_space) expectation_greater: + assumes [simp]: "integrable M X" + assumes gt: "\x\space M. a < X x" + shows "a < expectation X" +proof - + have "expectation (\x. a) < expectation X" + using gt measure_space_1 + by (intro integral_less_AE_space) auto + then show ?thesis using prob_space by simp +qed + +lemma convex_le_Inf_differential: + fixes f :: "real \ real" + assumes "convex_on I f" + assumes "x \ interior I" "y \ I" + shows "f y \ f x + Inf ((\t. (f x - f t) / (x - t)) ` ({x<..} \ I)) * (y - x)" + (is "_ \ _ + Inf (?F x) * (y - x)") +proof - + show ?thesis + proof (cases rule: linorder_cases) + assume "x < y" + moreover + have "open (interior I)" by auto + from openE[OF this `x \ interior I`] guess e . note e = this + moreover def t \ "min (x + e / 2) ((x + y) / 2)" + ultimately have "x < t" "t < y" "t \ ball x e" + by (auto simp: mem_ball dist_real_def field_simps split: split_min) + with `x \ interior I` e interior_subset[of I] have "t \ I" "x \ I" by auto + + have "open (interior I)" by auto + from openE[OF this `x \ interior I`] guess e . + moreover def K \ "x - e / 2" + with `0 < e` have "K \ ball x e" "K < x" by (auto simp: mem_ball dist_real_def) + ultimately have "K \ I" "K < x" "x \ I" + using interior_subset[of I] `x \ interior I` by auto + + have "Inf (?F x) \ (f x - f y) / (x - y)" + proof (rule Inf_lower2) + show "(f x - f t) / (x - t) \ ?F x" + using `t \ I` `x < t` by auto + show "(f x - f t) / (x - t) \ (f x - f y) / (x - y)" + using `convex_on I f` `x \ I` `y \ I` `x < t` `t < y` by (rule convex_on_diff) + next + fix y assume "y \ ?F x" + with order_trans[OF convex_on_diff[OF `convex_on I f` `K \ I` _ `K < x` _]] + show "(f K - f x) / (K - x) \ y" by auto + qed + then show ?thesis + using `x < y` by (simp add: field_simps) + next + assume "y < x" + moreover + have "open (interior I)" by auto + from openE[OF this `x \ interior I`] guess e . note e = this + moreover def t \ "x + e / 2" + ultimately have "x < t" "t \ ball x e" + by (auto simp: mem_ball dist_real_def field_simps) + with `x \ interior I` e interior_subset[of I] have "t \ I" "x \ I" by auto + + have "(f x - f y) / (x - y) \ Inf (?F x)" + proof (rule Inf_greatest) + have "(f x - f y) / (x - y) = (f y - f x) / (y - x)" + using `y < x` by (auto simp: field_simps) + also + fix z assume "z \ ?F x" + with order_trans[OF convex_on_diff[OF `convex_on I f` `y \ I` _ `y < x`]] + have "(f y - f x) / (y - x) \ z" by auto + finally show "(f x - f y) / (x - y) \ z" . + next + have "open (interior I)" by auto + from openE[OF this `x \ interior I`] guess e . note e = this + then have "x + e / 2 \ ball x e" by (auto simp: mem_ball dist_real_def) + with e interior_subset[of I] have "x + e / 2 \ {x<..} \ I" by auto + then show "?F x \ {}" by blast + qed + then show ?thesis + using `y < x` by (simp add: field_simps) + qed simp +qed + +lemma (in prob_space) jensens_inequality: + fixes a b :: real + assumes X: "integrable M X" "X ` space M \ I" + assumes I: "I = {a <..< b} \ I = {a <..} \ I = {..< b} \ I = UNIV" + assumes q: "integrable M (\x. q (X x))" "convex_on I q" + shows "q (expectation X) \ expectation (\x. q (X x))" +proof - + let "?F x" = "Inf ((\t. (q x - q t) / (x - t)) ` ({x<..} \ I))" + from not_empty X(2) have "I \ {}" by auto + + from I have "open I" by auto + + note I + moreover + { assume "I \ {a <..}" + with X have "a < expectation X" + by (intro expectation_greater) auto } + moreover + { assume "I \ {..< b}" + with X have "expectation X < b" + by (intro expectation_less) auto } + ultimately have "expectation X \ I" + by (elim disjE) (auto simp: subset_eq) + moreover + { fix y assume y: "y \ I" + with q(2) `open I` have "Sup ((\x. q x + ?F x * (y - x)) ` I) = q y" + by (auto intro!: Sup_eq_maximum convex_le_Inf_differential image_eqI[OF _ y] simp: interior_open) } + ultimately have "q (expectation X) = Sup ((\x. q x + ?F x * (expectation X - x)) ` I)" + by simp + also have "\ \ expectation (\w. q (X w))" + proof (rule Sup_least) + show "(\x. q x + ?F x * (expectation X - x)) ` I \ {}" + using `I \ {}` by auto + next + fix k assume "k \ (\x. q x + ?F x * (expectation X - x)) ` I" + then guess x .. note x = this + have "q x + ?F x * (expectation X - x) = expectation (\w. q x + ?F x * (X w - x))" + using prob_space + by (simp add: integral_add integral_cmult integral_diff lebesgue_integral_const X) + also have "\ \ expectation (\w. q (X w))" + using `x \ I` `open I` X(2) + by (intro integral_mono integral_add integral_cmult integral_diff + lebesgue_integral_const X q convex_le_Inf_differential) + (auto simp: interior_open) + finally show "k \ expectation (\w. q (X w))" using x by auto + qed + finally show "q (expectation X) \ expectation (\x. q (X x))" . +qed + lemma (in prob_space) distribution_eq_translated_integral: assumes "random_variable S X" "A \ sets S" shows "distribution X A = integral\<^isup>P (S\measure := extreal \ distribution X\) (indicator A)" @@ -722,9 +885,6 @@ unfolding finite_prob_space_def finite_measure_space_def prob_space_def prob_space_axioms_def by auto -lemma (in prob_space) not_empty: "space M \ {}" - using prob_space empty_measure' by auto - lemma (in finite_prob_space) sum_over_space_eq_1: "(\x\space M. \ {x}) = 1" using measure_space_1 sum_over_space by simp @@ -829,7 +989,7 @@ 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 + 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) diff -r e83695ea0e0a -r b017cfb10df4 src/HOL/Probability/Radon_Nikodym.thy --- a/src/HOL/Probability/Radon_Nikodym.thy Thu Jun 09 14:16:12 2011 +0200 +++ b/src/HOL/Probability/Radon_Nikodym.thy Thu Jun 09 15:14:21 2011 +0200 @@ -1311,6 +1311,59 @@ by (auto simp: RN_deriv_positive_integral[OF ms \(2)]) qed +lemma (in sigma_finite_measure) real_RN_deriv: + assumes "finite_measure (M\measure := \\)" (is "finite_measure ?\") + assumes ac: "absolutely_continuous \" + obtains D where "D \ borel_measurable M" + and "AE x. RN_deriv M \ x = extreal (D x)" + and "AE x in M\measure := \\. 0 < D x" + and "\x. 0 \ D x" +proof + interpret \: finite_measure ?\ by fact + have ms: "measure_space ?\" by default + note RN = RN_deriv[OF ms ac] + + let ?RN = "\t. {x \ space M. RN_deriv M \ x = t}" + + show "(\x. real (RN_deriv M \ x)) \ borel_measurable M" + using RN by auto + + have "\ (?RN \) = (\\<^isup>+ x. RN_deriv M \ x * indicator (?RN \) x \M)" + using RN by auto + also have "\ = (\\<^isup>+ x. \ * indicator (?RN \) x \M)" + by (intro positive_integral_cong) (auto simp: indicator_def) + also have "\ = \ * \ (?RN \)" + using RN by (intro positive_integral_cmult_indicator) auto + finally have eq: "\ (?RN \) = \ * \ (?RN \)" . + moreover + have "\ (?RN \) = 0" + proof (rule ccontr) + assume "\ {x \ space M. RN_deriv M \ x = \} \ 0" + moreover from RN have "0 \ \ {x \ space M. RN_deriv M \ x = \}" by auto + ultimately have "0 < \ {x \ space M. RN_deriv M \ x = \}" by auto + with eq have "\ (?RN \) = \" by simp + with \.finite_measure[of "?RN \"] RN show False by auto + qed + ultimately have "AE x. RN_deriv M \ x < \" + using RN by (intro AE_iff_measurable[THEN iffD2]) auto + then show "AE x. RN_deriv M \ x = extreal (real (RN_deriv M \ x))" + using RN(3) by (auto simp: extreal_real) + then have eq: "AE x in (M\measure := \\) . RN_deriv M \ x = extreal (real (RN_deriv M \ x))" + using ac absolutely_continuous_AE[OF ms] by auto + + show "\x. 0 \ real (RN_deriv M \ x)" + using RN by (auto intro: real_of_extreal_pos) + + have "\ (?RN 0) = (\\<^isup>+ x. RN_deriv M \ x * indicator (?RN 0) x \M)" + using RN by auto + also have "\ = (\\<^isup>+ x. 0 \M)" + by (intro positive_integral_cong) (auto simp: indicator_def) + finally have "AE x in (M\measure := \\). RN_deriv M \ x \ 0" + using RN by (intro \.AE_iff_measurable[THEN iffD2]) auto + with RN(3) eq show "AE x in (M\measure := \\). 0 < real (RN_deriv M \ x)" + by (auto simp: zero_less_real_of_extreal le_less) +qed + lemma (in sigma_finite_measure) RN_deriv_singleton: assumes "measure_space (M\measure := \\)" and ac: "absolutely_continuous \" diff -r e83695ea0e0a -r b017cfb10df4 src/HOL/Quickcheck_Narrowing.thy --- a/src/HOL/Quickcheck_Narrowing.thy Thu Jun 09 14:16:12 2011 +0200 +++ b/src/HOL/Quickcheck_Narrowing.thy Thu Jun 09 15:14:21 2011 +0200 @@ -26,7 +26,7 @@ code_reserved Haskell_Quickcheck Typerep -subsubsection {* Type @{text "code_int"} for Haskell_Quickcheck's Int type *} +subsubsection {* Type @{text "code_int"} for Haskell Quickcheck's Int type *} typedef (open) code_int = "UNIV \ int set" morphisms int_of of_int by rule @@ -218,7 +218,7 @@ datatype narrowing_term = Var "code_int list" narrowing_type | Ctr code_int "narrowing_term list" datatype 'a cons = C narrowing_type "(narrowing_term list => 'a) list" -subsubsection {* From narrowing's deep representation of terms to Code_Evaluation's terms *} +subsubsection {* From narrowing's deep representation of terms to @{theory Code_Evaluation}'s terms *} class partial_term_of = typerep + fixes partial_term_of :: "'a itself => narrowing_term => Code_Evaluation.term" diff -r e83695ea0e0a -r b017cfb10df4 src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy Thu Jun 09 14:16:12 2011 +0200 +++ b/src/HOL/Transcendental.thy Thu Jun 09 15:14:21 2011 +0200 @@ -1366,7 +1366,7 @@ declare DERIV_exp[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros] - DERIV_ln[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros] + DERIV_ln_divide[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros] DERIV_sin[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros] DERIV_cos[THEN DERIV_chain2, THEN DERIV_cong, DERIV_intros]