diff -r 596a499318ab -r b0b9a10e4bf4 src/HOL/Probability/Information.thy --- a/src/HOL/Probability/Information.thy Wed Jun 11 13:39:38 2014 +0200 +++ b/src/HOL/Probability/Information.thy Thu Jun 12 15:47:36 2014 +0200 @@ -8,7 +8,7 @@ theory Information imports Independent_Family - Radon_Nikodym + Distributions "~~/src/HOL/Library/Convex" begin @@ -400,27 +400,6 @@ done qed -lemma distributed_integrable: - "distributed M N X f \ g \ borel_measurable N \ - integrable N (\x. f x * g x) \ integrable M (\x. g (X x))" - by (auto simp: distributed_real_AE - distributed_distr_eq_density[symmetric] integrable_real_density[symmetric] integrable_distr_eq) - -lemma distributed_transform_integrable: - assumes Px: "distributed M N X Px" - assumes "distributed M P Y Py" - assumes Y: "Y = (\x. T (X x))" and T: "T \ measurable N P" and f: "f \ borel_measurable P" - shows "integrable P (\x. Py x * f x) \ integrable N (\x. Px x * f (T x))" -proof - - have "integrable P (\x. Py x * f x) \ integrable M (\x. f (Y x))" - by (rule distributed_integrable) fact+ - also have "\ \ integrable M (\x. f (T (X x)))" - using Y by simp - also have "\ \ integrable N (\x. Px x * f (T x))" - using measurable_comp[OF T f] Px by (intro distributed_integrable[symmetric]) (auto simp: comp_def) - finally show ?thesis . -qed - lemma integrable_cong_AE_imp: "integrable M g \ f \ borel_measurable M \ (AE x in M. g x = f x) \ integrable M f" using integrable_cong_AE[of f M g] by (auto simp: eq_commute) @@ -937,6 +916,35 @@ by (subst integral_mult_right) (auto simp: measure_def) qed +lemma (in information_space) entropy_exponential: + assumes D: "distributed M lborel X (exponential_density l)" + shows "entropy b lborel X = log b (exp 1 / l)" +proof - + have l[simp, arith]: "0 < l" by (rule exponential_distributed_params[OF D]) + + have [simp]: "integrable lborel (exponential_density l)" + using distributed_integrable[OF D, of "\_. 1"] by simp + + have [simp]: "integral\<^sup>L lborel (exponential_density l) = 1" + using distributed_integral[OF D, of "\_. 1"] by (simp add: prob_space) + + have [simp]: "integrable lborel (\x. exponential_density l x * x)" + using erlang_ith_moment_integrable[OF l D, of 1] distributed_integrable[OF D, of "\x. x"] by simp + + have [simp]: "integral\<^sup>L lborel (\x. exponential_density l x * x) = 1 / l" + using erlang_ith_moment[OF l D, of 1] distributed_integral[OF D, of "\x. x"] by simp + + have "entropy b lborel X = - (\ x. exponential_density l x * log b (exponential_density l x) \lborel)" + using D by (rule entropy_distr) + also have "(\ x. exponential_density l x * log b (exponential_density l x) \lborel) = + (\ x. (ln l * exponential_density l x - l * (exponential_density l x * x)) / ln b \lborel)" + by (intro integral_cong) (auto simp: log_def ln_mult exponential_density_def field_simps) + also have "\ = (ln l - 1) / ln b" + by simp + finally show ?thesis + by (simp add: log_def divide_simps ln_div) +qed + lemma (in information_space) entropy_simple_distributed: "simple_distributed M X f \ \(X) = - (\x\X`space M. f x * log b (f x))" by (subst entropy_distr[OF simple_distributed])