--- 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 \<Rightarrow> real"
+ assumes f: "convex_on I f" and I: "x\<in>I" "y\<in>I" and t: "x < t" "t < y"
+ shows "(f x - f t) / (x - t) \<le> (f x - f y) / (x - y)" "(f x - f y) / (x - y) \<le> (f t - f y) / (t - y)"
+proof -
+ def a \<equiv> "(t - y) / (x - y)"
+ with t have "0 \<le> a" "0 \<le> 1 - a" by (auto simp: field_simps)
+ with f `x \<in> I` `y \<in> I` have cvx: "f (a * x + (1 - a) * y) \<le> 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 "\<dots> = t" unfolding a_def using `x < t` `t < y` by simp
+ finally have "f t \<le> a * f x + (1 - a) * f y" using cvx by simp
+ also have "\<dots> = a * (f x - f y) + f y" by (simp add: field_simps)
+ finally have "f t - f y \<le> a * (f x - f y)" by simp
+ with t show "(f x - f t) / (x - t) \<le> (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) \<le> (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 \<Rightarrow> real"
--- 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 \<Longrightarrow> ln x \<le> 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: "\<And>x. 0 < x \<Longrightarrow> 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 \<le> y" "y \<le> a"
+ with `0 < x` `a < 1` have "0 < 1 / y - 1" "0 < y"
+ by (auto simp: field_simps)
+ with D show "\<exists>z. DERIV ?l y :> z \<and> 0 < z"
+ by auto
+ qed
+ also have "\<dots> \<le> 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 \<le> y" "y \<le> x"
+ with `1 < a` have "1 / y - 1 < 0" "0 < y"
+ by (auto simp: field_simps)
+ with D show "\<exists>z. DERIV ?l y :> z \<and> z < 0"
+ by blast
+ qed
+ also have "\<dots> \<le> 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
--- 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 \<in> s \<Longrightarrow> open s \<Longrightarrow> ((f has_derivative f') (at a within s) \<longleftrightarrow> (f has_derivative f') (at a))"
by (simp only: at_within_interior interior_open)
+lemma has_derivative_right:
+ fixes f :: "real \<Rightarrow> real" and y :: "real"
+ shows "(f has_derivative (op * y)) (at x within ({x <..} \<inter> I)) \<longleftrightarrow>
+ ((\<lambda>t. (f x - f t) / (x - t)) ---> y) (at x within ({x <..} \<inter> I))"
+proof -
+ have "((\<lambda>t. (f t - (f x + y * (t - x))) / \<bar>t - x\<bar>) ---> 0) (at x within ({x<..} \<inter> I)) \<longleftrightarrow>
+ ((\<lambda>t. (f t - f x) / (t - x) - y) ---> 0) (at x within ({x<..} \<inter> I))"
+ by (intro Lim_cong_within) (auto simp add: divide.diff divide.add)
+ also have "\<dots> \<longleftrightarrow> ((\<lambda>t. (f t - f x) / (t - x)) ---> y) (at x within ({x<..} \<inter> I))"
+ by (simp add: Lim_null[symmetric])
+ also have "\<dots> \<longleftrightarrow> ((\<lambda>t. (f x - f t) / (x - t)) ---> y) (at x within ({x<..} \<inter> 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 \<Longrightarrow> linear f" (* TODO: move elsewhere *)
proof -
assume "bounded_linear f"
--- 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 "\<forall>S. (\<forall>n. S n \<noteq> a \<and> S n \<in> T) \<and> S ----> a \<longrightarrow> (\<lambda>n. X (S n)) ----> L"
+ shows "(X ---> L) (at a within T)"
+proof (rule ccontr)
+ assume "\<not> (X ---> L) (at a within T)"
+ hence "\<exists>r>0. \<forall>s>0. \<exists>x\<in>T. x \<noteq> a \<and> \<bar>x - a\<bar> < s \<and> r \<le> dist (X x) L"
+ unfolding tendsto_iff eventually_within dist_norm by (simp add: not_less[symmetric])
+ then obtain r where r: "r > 0" "\<And>s. s > 0 \<Longrightarrow> \<exists>x\<in>T-{a}. \<bar>x - a\<bar> < s \<and> dist (X x) L \<ge> r" by blast
+
+ let ?F = "\<lambda>n::nat. SOME x. x \<in> T \<and> x \<noteq> a \<and> \<bar>x - a\<bar> < inverse (real (Suc n)) \<and> dist (X x) L \<ge> r"
+ have "\<And>n. \<exists>x. x \<in> T \<and> x \<noteq> a \<and> \<bar>x - a\<bar> < inverse (real (Suc n)) \<and> dist (X x) L \<ge> r"
+ using r by (simp add: Bex_def)
+ hence F: "\<And>n. ?F n \<in> T \<and> ?F n \<noteq> a \<and> \<bar>?F n - a\<bar> < inverse (real (Suc n)) \<and> dist (X (?F n)) L \<ge> r"
+ by (rule someI_ex)
+ hence F1: "\<And>n. ?F n \<in> T \<and> ?F n \<noteq> a"
+ and F2: "\<And>n. \<bar>?F n - a\<bar> < inverse (real (Suc n))"
+ and F3: "\<And>n. dist (X (?F n)) L \<ge> 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 "\<exists>no. inverse (real (Suc no)) < e" by (rule reals_Archimedean)
+ then obtain m where nodef: "inverse (real (Suc m)) < e" by auto
+ show "\<exists>no. \<forall>n. no \<le> n --> \<bar>?F n - a\<bar> < e"
+ proof (intro exI allI impI)
+ fix n
+ assume mlen: "m \<le> n"
+ have "\<bar>?F n - a\<bar> < inverse (real (Suc n))"
+ by (rule F2)
+ also have "inverse (real (Suc n)) \<le> inverse (real (Suc m))"
+ using mlen by auto
+ also from nodef have
+ "inverse (real (Suc m)) < e" .
+ finally show "\<bar>?F n - a\<bar> < e" .
+ qed
+ qed
+ moreover note `\<forall>S. (\<forall>n. S n \<noteq> a \<and> S n \<in> T) \<and> S ----> a \<longrightarrow> (\<lambda>n. X (S n)) ----> L`
+ ultimately have "(\<lambda>n. X (?F n)) ----> L" using F1 by simp
+
+ moreover have "\<not> ((\<lambda>n. X (?F n)) ----> L)"
+ proof -
+ {
+ fix no::nat
+ obtain n where "n = no + 1" by simp
+ then have nolen: "no \<le> n" by simp
+ (* We prove this by showing that for any m there is an n\<ge>m such that |X (?F n) - L| \<ge> r *)
+ have "dist (X (?F n)) L \<ge> r"
+ by (rule F3)
+ with nolen have "\<exists>n. no \<le> n \<and> dist (X (?F n)) L \<ge> r" by fast
+ }
+ then have "(\<forall>no. \<exists>n. no \<le> n \<and> dist (X (?F n)) L \<ge> r)" by simp
+ with r have "\<exists>e>0. (\<forall>no. \<exists>n. no \<le> n \<and> dist (X (?F n)) L \<ge> 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 \<Rightarrow> real"
+ assumes mono: "\<And>a b. a \<in> I \<Longrightarrow> b \<in> I \<Longrightarrow> x < a \<Longrightarrow> a \<le> b \<Longrightarrow> f a \<le> f b"
+ assumes bnd: "\<And>a. a \<in> I \<Longrightarrow> x < a \<Longrightarrow> K \<le> f a"
+ shows "(f ---> Inf (f ` ({x<..} \<inter> I))) (at x within ({x<..} \<inter> I))"
+proof cases
+ assume "{x<..} \<inter> I = {}" then show ?thesis by (simp add: Lim_within_empty)
+next
+ assume [simp]: "{x<..} \<inter> I \<noteq> {}"
+ show ?thesis
+ proof (rule Lim_within_LIMSEQ, safe)
+ fix S assume S: "\<forall>n. S n \<noteq> x \<and> S n \<in> {x <..} \<inter> I" "S ----> x"
+
+ show "(\<lambda>n. f (S n)) ----> Inf (f ` ({x<..} \<inter> I))"
+ proof (rule LIMSEQ_I, rule ccontr)
+ fix r :: real assume "0 < r"
+ with Inf_close[of "f ` ({x<..} \<inter> I)" r]
+ obtain y where y: "x < y" "y \<in> I" "f y < Inf (f ` ({x <..} \<inter> 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: "\<And>n. N \<le> n \<Longrightarrow> \<bar>S n - x\<bar> < y - x" by auto
+
+ assume "\<not> (\<exists>N. \<forall>n\<ge>N. norm (f (S n) - Inf (f ` ({x<..} \<inter> I))) < r)"
+ moreover have "\<And>n. Inf (f ` ({x<..} \<inter> I)) \<le> f (S n)"
+ using S bnd by (intro Inf_lower[where z=K]) auto
+ ultimately obtain n where n: "N \<le> n" "r + Inf (f ` ({x<..} \<inter> I)) \<le> f (S n)"
+ by (auto simp: not_less field_simps)
+ with N[OF n(1)] mono[OF _ `y \<in> 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 "\<And>x. x \<noteq> a \<Longrightarrow> f x = g x"
- shows "((\<lambda>x. f x) ---> l) (at a within S) \<longleftrightarrow> ((g ---> l) (at a within S))"
+ assumes "a = b" "x = y" "S = T"
+ assumes "\<And>x. x \<noteq> b \<Longrightarrow> x \<in> T \<Longrightarrow> f x = g x"
+ shows "(f ---> x) (at a within S) \<longleftrightarrow> (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 "\<And>x. x \<noteq> a \<Longrightarrow> f x = g x"
- shows "((\<lambda>x. f x) ---> l) (at a) \<longleftrightarrow> ((g ---> l) (at a))"
+ shows "((\<lambda>x. f x) ---> x) (at a) \<longleftrightarrow> ((g ---> y) (at a))"
unfolding tendsto_def eventually_at_topological
using assms by simp
--- 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 \<longleftrightarrow>
+ (random_variable S X \<and> random_variable T Y) \<and>
+ indep_set
+ (sigma_sets (space M) { X -` A \<inter> space M | A. A \<in> sets S})
+ (sigma_sets (space M) { Y -` A \<inter> space M | A. A \<in> sets T})"
+ unfolding indep_var_def indep_vars_def indep_set_def UNIV_bool
+ by (intro arg_cong2[where f="op \<and>"] 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 \<subseteq> events"
@@ -491,7 +501,7 @@
proof (simp add: sigma_algebra_iff2, safe)
let ?A = "(\<Inter>n. sigma_sets (space M) (UNION {n..} A))"
interpret A: sigma_algebra "\<lparr>space = space M, sets = A i\<rparr>" for i by fact
- { fix X x assume "X \<in> ?A" "x \<in> X"
+ { fix X x assume "X \<in> ?A" "x \<in> X"
then have "\<And>n. X \<in> sigma_sets (space M) (UNION {n..} A)" by auto
from this[of 0] have "X \<in> sigma_sets (space M) (UNION UNIV A)" by simp
then have "X \<subseteq> space M"
@@ -572,7 +582,7 @@
show "Int_stable \<lparr>space = space M, sets = A m\<rparr>"
unfolding Int_stable_def using A.Int by auto
qed
- also have "(\<lambda>b. sigma_sets (space M) (\<Union>m\<in>bool_case {..n} {Suc n..} b. A m)) =
+ also have "(\<lambda>b. sigma_sets (space M) (\<Union>m\<in>bool_case {..n} {Suc n..} b. A m)) =
bool_case (sigma_sets (space M) (\<Union>m\<in>{..n}. A m)) (sigma_sets (space M) (\<Union>m\<in>{Suc n..}. A m))"
by (auto intro!: ext split: bool.split)
finally have indep: "indep_set (sigma_sets (space M) (\<Union>m\<in>{..n}. A m)) (sigma_sets (space M) (\<Union>m\<in>{Suc n..}. A m))"
@@ -732,7 +742,7 @@
by (auto simp del: vimage_Int intro!: exI[of _ "A \<inter> B"] dest: Int_stableD)
qed }
note indep_sets_sigma_sets_iff[OF this, simp]
-
+
{ fix i assume "i \<in> I"
{ fix A assume "A \<in> sets (M' i)"
then have "A \<in> sets (sigma (M' i))" by (auto simp: sets_sigma intro: sigma_sets.Basic)
@@ -745,7 +755,7 @@
"space M \<in> {X i -` A \<inter> space M |A. A \<in> sets (M' i)}"
by (auto intro!: exI[of _ "space (M' i)"]) }
note indep_sets_finite[OF I this, simp]
-
+
have "(\<forall>A\<in>\<Pi> i\<in>I. {X i -` A \<inter> space M |A. A \<in> sets (M' i)}. prob (INTER I A) = (\<Prod>j\<in>I. prob (A j))) =
(\<forall>A\<in>\<Pi> i\<in>I. sets (M' i). prob ((\<Inter>j\<in>I. X j -` A j) \<inter> space M) = (\<Prod>x\<in>I. prob (X x -` A x \<inter> space M)))"
(is "?L = ?R")
@@ -847,7 +857,7 @@
by (simp_all add: product_algebra_def)
show "A \<in> sets (sigma P.G)"
using `A \<in> sets P.P` by (simp add: product_algebra_def)
-
+
fix E assume E: "E \<in> sets P.G"
then have "E \<in> 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 "\<forall>i\<in>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 \<in> sets Ma" "Xb \<in> sets Mb"
- shows "joint_distribution A B (Xa \<times> 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 \<equiv> S\<lparr>measure := extreal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
+ assumes "A \<in> sets P"
+ shows "joint_distribution X Y A = finite_measure.\<mu>' 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\<lparr>measure := extreal\<circ>distribution X\<rparr>"
+ let ?T = "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
+ 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\<lparr> measure := extreal \<circ> joint_distribution X Y \<rparr>"
+ interpret J: prob_space ?J
+ by (rule joint_distribution_prob_space) (simp_all add: rvs)
+
+ have "finite_measure.\<mu>' (XY.P\<lparr> measure := extreal \<circ> joint_distribution X Y \<rparr>) A = XY.\<mu>' 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 \<in> 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 \<in> sets (sigma ?P)"
+ using `A \<in> sets P` unfolding P_def pair_measure_def by simp
+
+ fix X assume "X \<in> sets ?P"
+ then obtain A B where "A \<in> sets S" "B \<in> sets T" "X = A \<times> B"
+ by (auto simp: sets_pair_measure_generator)
+ then show "J.\<mu>' X = XY.\<mu>' X"
+ unfolding J.\<mu>'_def XY.\<mu>'_def using indep
+ by (simp add: XY.pair_measure_times)
+ (simp add: distribution_def indep_varD)
+ qed
+ then show ?thesis
+ using `A \<in> sets P` unfolding P_def J.\<mu>'_def XY.\<mu>'_def by simp
+qed
end
--- 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]:
- "(\<not> 0 < distribution X A) \<longleftrightarrow> distribution X A = 0"
- using distribution_positive[of X A] by arith
-
lemma log_le: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> x \<le> y \<Longrightarrow> log a x \<le> log a y"
by (subst log_le_cancel_iff) auto
@@ -175,7 +171,211 @@
Kullback$-$Leibler distance. *}
definition
- "KL_divergence b M \<nu> = \<integral>x. log b (real (RN_deriv M \<nu> x)) \<partial>M\<lparr>measure := \<nu>\<rparr>"
+ "entropy_density b M \<nu> = log b \<circ> real \<circ> RN_deriv M \<nu>"
+
+definition
+ "KL_divergence b M \<nu> = integral\<^isup>L (M\<lparr>measure := \<nu>\<rparr>) (entropy_density b M \<nu>)"
+
+lemma (in information_space) measurable_entropy_density:
+ assumes ps: "prob_space (M\<lparr>measure := \<nu>\<rparr>)"
+ assumes ac: "absolutely_continuous \<nu>"
+ shows "entropy_density b M \<nu> \<in> borel_measurable M"
+proof -
+ interpret \<nu>: prob_space "M\<lparr>measure := \<nu>\<rparr>" by fact
+ have "measure_space (M\<lparr>measure := \<nu>\<rparr>)" 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\<lparr>measure := \<nu>\<rparr>)"
+ assumes ac: "absolutely_continuous \<nu>"
+ assumes int: "integrable (M\<lparr> measure := \<nu> \<rparr>) (entropy_density b M \<nu>)"
+ assumes A: "A \<in> sets M" "\<nu> A \<noteq> \<mu> A"
+ shows "0 < KL_divergence b M \<nu>"
+proof -
+ interpret \<nu>: prob_space "M\<lparr>measure := \<nu>\<rparr>" by fact
+ have ms: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" by default
+ have fms: "finite_measure (M\<lparr>measure := \<nu>\<rparr>)" 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\<nu>: "AE x in M\<lparr>measure := \<nu>\<rparr>. RN_deriv M \<nu> x = extreal (D x)"
+ by auto
+
+ def f \<equiv> "\<lambda>x. if D x = 0 then 1 else 1 / D x"
+ with D have f_borel: "f \<in> borel_measurable M"
+ by (auto intro!: measurable_If)
+
+ have "KL_divergence b M \<nu> = 1 / ln b * (\<integral> x. ln b * entropy_density b M \<nu> x \<partial>M\<lparr>measure := \<nu>\<rparr>)"
+ unfolding KL_divergence_def using int b_gt_1
+ by (simp add: integral_cmult)
+
+ { fix A assume "A \<in> sets M"
+ with RN D have "\<nu>.\<mu> A = (\<integral>\<^isup>+ x. extreal (D x) * indicator A x \<partial>M)"
+ by (auto intro!: positive_integral_cong_AE) }
+ note D_density = this
+
+ have ln_entropy: "(\<lambda>x. ln b * entropy_density b M \<nu> x) \<in> borel_measurable M"
+ using measurable_entropy_density[OF ps ac] by auto
+
+ have "integrable (M\<lparr>measure := \<nu>\<rparr>) (\<lambda>x. ln b * entropy_density b M \<nu> x)"
+ using int by auto
+ moreover have "integrable (M\<lparr>measure := \<nu>\<rparr>) (\<lambda>x. ln b * entropy_density b M \<nu> x) \<longleftrightarrow>
+ integrable M (\<lambda>x. D x * (ln b * entropy_density b M \<nu> x))"
+ using D D_density ln_entropy
+ by (intro integral_translated_density) auto
+ ultimately have M_int: "integrable M (\<lambda>x. D x * (ln b * entropy_density b M \<nu> x))"
+ by simp
+
+ have D_neg: "(\<integral>\<^isup>+ x. extreal (- D x) \<partial>M) = 0"
+ using D by (subst positive_integral_0_iff_AE) auto
+
+ have "(\<integral>\<^isup>+ x. extreal (D x) \<partial>M) = \<nu> (space M)"
+ using RN D by (auto intro!: positive_integral_cong_AE)
+ then have D_pos: "(\<integral>\<^isup>+ x. extreal (D x) \<partial>M) = 1"
+ using \<nu>.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\<in>space M. D x \<noteq> 0}"
+ have [simp, intro]: "?D_set \<in> sets M"
+ using D by (auto intro: sets_Collect)
+
+ have "0 \<le> 1 - \<mu>' ?D_set"
+ using prob_le_1 by (auto simp: field_simps)
+ also have "\<dots> = (\<integral> x. D x - indicator ?D_set x \<partial>M)"
+ using `integrable M D` `integral\<^isup>L M D = 1`
+ by (simp add: \<mu>'_def)
+ also have "\<dots> < (\<integral> x. D x * (ln b * entropy_density b M \<nu> x) \<partial>M)"
+ proof (rule integral_less_AE)
+ show "integrable M (\<lambda>x. D x - indicator ?D_set x)"
+ using `integrable M D`
+ by (intro integral_diff integral_indicator) auto
+ next
+ show "integrable M (\<lambda>x. D x * (ln b * entropy_density b M \<nu> x))"
+ by fact
+ next
+ show "\<mu> {x\<in>space M. D x \<noteq> 1 \<and> D x \<noteq> 0} \<noteq> 0"
+ proof
+ assume eq_0: "\<mu> {x\<in>space M. D x \<noteq> 1 \<and> D x \<noteq> 0} = 0"
+ then have disj: "AE x. D x = 1 \<or> D x = 0"
+ using D(1) by (auto intro!: AE_I[OF subset_refl] sets_Collect)
+
+ have "\<mu> {x\<in>space M. D x = 1} = (\<integral>\<^isup>+ x. indicator {x\<in>space M. D x = 1} x \<partial>M)"
+ using D(1) by auto
+ also have "\<dots> = (\<integral>\<^isup>+ x. extreal (D x) * indicator {x\<in>space M. D x \<noteq> 0} x \<partial>M)"
+ using disj by (auto intro!: positive_integral_cong_AE simp: indicator_def one_extreal_def)
+ also have "\<dots> = \<nu> {x\<in>space M. D x \<noteq> 0}"
+ using D(1) D_density by auto
+ also have "\<dots> = \<nu> (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) \<nu>.measure_space_1 by (intro AE_I_eq_1) auto
+ then have "(\<integral>\<^isup>+x. indicator A x\<partial>M) = (\<integral>\<^isup>+x. extreal (D x) * indicator A x\<partial>M)"
+ by (intro positive_integral_cong_AE) (auto simp: one_extreal_def[symmetric])
+ also have "\<dots> = \<nu> A"
+ using `A \<in> sets M` D_density by simp
+ finally show False using `A \<in> sets M` `\<nu> A \<noteq> \<mu> A` by simp
+ qed
+ show "{x\<in>space M. D x \<noteq> 1 \<and> D x \<noteq> 0} \<in> sets M"
+ using D(1) by (auto intro: sets_Collect)
+
+ show "AE t. t \<in> {x\<in>space M. D x \<noteq> 1 \<and> D x \<noteq> 0} \<longrightarrow>
+ D t - indicator ?D_set t \<noteq> D t * (ln b * entropy_density b M \<nu> t)"
+ using D(2)
+ proof (elim AE_mp, safe intro!: AE_I2)
+ fix t assume Dt: "t \<in> space M" "D t \<noteq> 1" "D t \<noteq> 0"
+ and RN: "RN_deriv M \<nu> t = extreal (D t)"
+ and eq: "D t - indicator ?D_set t = D t * (ln b * entropy_density b M \<nu> 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 \<nu> t) = - D t * ln (1 / D t)"
+ using RN b_gt_1 `D t \<noteq> 0` `0 \<le> 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 \<noteq> 0` by (auto simp: field_simps)
+ from ln_eq_minus_one[OF _ this] `D t \<noteq> 0` `0 \<le> D t` `D t \<noteq> 1`
+ show False by auto
+ qed
+
+ show "AE t. D t - indicator ?D_set t \<le> D t * (ln b * entropy_density b M \<nu> t)"
+ using D(2)
+ proof (elim AE_mp, intro AE_I2 impI)
+ fix t assume "t \<in> space M" and RN: "RN_deriv M \<nu> t = extreal (D t)"
+ show "D t - indicator ?D_set t \<le> D t * (ln b * entropy_density b M \<nu> t)"
+ proof cases
+ assume asm: "D t \<noteq> 0"
+ then have "0 < D t" using `0 \<le> D t` by auto
+ then have "0 < 1 / D t" by auto
+ have "D t - indicator ?D_set t \<le> - D t * (1 / D t - 1)"
+ using asm `t \<in> space M` by (simp add: field_simps)
+ also have "- D t * (1 / D t - 1) \<le> - D t * ln (1 / D t)"
+ using ln_le_minus_one `0 < 1 / D t` by (intro mult_left_mono_neg) auto
+ also have "\<dots> = D t * (ln b * entropy_density b M \<nu> 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 "\<dots> = (\<integral> x. ln b * entropy_density b M \<nu> x \<partial>M\<lparr>measure := \<nu>\<rparr>)"
+ using D D_density ln_entropy
+ by (intro integral_translated_density[symmetric]) auto
+ also have "\<dots> = ln b * (\<integral> x. entropy_density b M \<nu> x \<partial>M\<lparr>measure := \<nu>\<rparr>)"
+ using int by (rule \<nu>.integral_cmult)
+ finally show "0 < KL_divergence b M \<nu>"
+ 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: "\<forall>A\<in>sets M. \<nu> A = measure M A"
+ shows "KL_divergence b M \<nu> = 0"
+proof -
+ have "AE x. 1 = RN_deriv M \<nu> x"
+ proof (rule RN_deriv_unique)
+ show "measure_space (M\<lparr>measure := \<nu>\<rparr>)"
+ using eq by (intro measure_space_cong) auto
+ show "absolutely_continuous \<nu>"
+ unfolding absolutely_continuous_def using eq by auto
+ show "(\<lambda>x. 1) \<in> borel_measurable M" "AE x. 0 \<le> (1 :: extreal)" by auto
+ fix A assume "A \<in> sets M"
+ with eq show "\<nu> A = \<integral>\<^isup>+ x. 1 * indicator A x \<partial>M" by simp
+ qed
+ then have "AE x. log b (real (RN_deriv M \<nu> x)) = 0"
+ by (elim AE_mp) simp
+ from integral_cong_AE[OF this]
+ have "integral\<^isup>L M (entropy_density b M \<nu>) = 0"
+ by (simp add: entropy_density_def comp_def)
+ with eq show "KL_divergence b M \<nu> = 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\<lparr>measure := \<nu>\<rparr>)"
+ assumes ac: "absolutely_continuous \<nu>"
+ assumes int: "integrable (M\<lparr> measure := \<nu> \<rparr>) (entropy_density b M \<nu>)"
+ assumes KL: "KL_divergence b M \<nu> = 0"
+ shows "\<forall>A\<in>sets M. \<nu> A = \<mu> A"
+ by (metis less_imp_neq KL_gt_0 assms)
+
+lemma (in information_space) KL_ge_0:
+ assumes ps: "prob_space (M\<lparr>measure := \<nu>\<rparr>)"
+ assumes ac: "absolutely_continuous \<nu>"
+ assumes int: "integrable (M\<lparr> measure := \<nu> \<rparr>) (entropy_density b M \<nu>)"
+ shows "0 \<le> KL_divergence b M \<nu>"
+ using KL_eq_0 KL_gt_0[OF ps ac int]
+ by (cases "\<forall>A\<in>sets M. \<nu> A = measure M A") (auto simp: le_less)
+
lemma (in sigma_finite_measure) KL_divergence_vimage:
assumes T: "T \<in> measure_preserving M M'"
@@ -209,7 +409,7 @@
have AE: "AE x. RN_deriv M' \<nu>' (T x) = RN_deriv M \<nu> x"
by (rule RN_deriv_vimage[OF T T' inv \<nu>'])
show ?thesis
- unfolding KL_divergence_def
+ unfolding KL_divergence_def entropy_density_def comp_def
proof (subst \<nu>'.integral_vimage[OF sa T'])
show "(\<lambda>x. log b (real (RN_deriv M \<nu> x))) \<in> borel_measurable (M\<lparr>measure := \<nu>\<rparr>)"
by (auto intro!: RN_deriv[OF M ac] borel_measurable_log[OF _ `1 < b`])
@@ -233,9 +433,9 @@
proof -
interpret \<nu>: measure_space ?\<nu> by fact
have "KL_divergence b M \<nu> = \<integral>x. log b (real (RN_deriv N \<nu>' x)) \<partial>?\<nu>"
- by (simp cong: RN_deriv_cong \<nu>.integral_cong add: KL_divergence_def)
+ by (simp cong: RN_deriv_cong \<nu>.integral_cong add: KL_divergence_def entropy_density_def)
also have "\<dots> = KL_divergence b N \<nu>'"
- by (auto intro!: \<nu>.integral_cong_measure[symmetric] simp: KL_divergence_def)
+ by (auto intro!: \<nu>.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\<lparr>measure := \<nu>\<rparr>)"
assumes ac: "absolutely_continuous \<nu>"
shows "KL_divergence b M \<nu> = (\<Sum>x\<in>space M. real (\<nu> {x}) * log b (real (\<nu> {x}) / real (\<mu> {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\<lparr>measure := \<nu>\<rparr>" by fact
have ms: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" by default
show "(\<Sum>x \<in> space M. log b (real (RN_deriv M \<nu> x)) * real (\<nu> {x})) = ?sum"
@@ -257,27 +457,10 @@
and "1 < b"
shows "0 \<le> KL_divergence b M \<nu>"
proof -
+ interpret information_space M by default fact
interpret v: finite_prob_space "M\<lparr>measure := \<nu>\<rparr>" by fact
- have ms: "finite_measure_space (M\<lparr>measure := \<nu>\<rparr>)" by default
-
- have "- (KL_divergence b M \<nu>) \<le> log b (\<Sum>x\<in>space M. real (\<mu> {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 "(\<Sum>x\<in>space M. real (\<nu> {x})) = 1"
- using v.finite_sum_over_space_eq_1 by (simp add: v.\<mu>'_def)
-
- fix x assume "x \<in> space M"
- then have x: "{x} \<in> sets M" unfolding sets_eq_Pow by auto
- { assume "0 < real (\<nu> {x})"
- then have "\<nu> {x} \<noteq> 0" by auto
- then have "\<mu> {x} \<noteq> 0"
- using ac[unfolded absolutely_continuous_def, THEN bspec, of "{x}"] x by auto
- thus "0 < real (\<mu> {x})" using real_measure[OF x] by auto }
- show "0 \<le> real (\<mu> {x})" "0 \<le> real (\<nu> {x})"
- using real_measure[OF x] v.real_measure[of "{x}"] x by auto
- qed
- thus "0 \<le> KL_divergence b M \<nu>" using finite_sum_over_space_eq_1 by (simp add: \<mu>'_def)
+ have ps: "prob_space (M\<lparr>measure := \<nu>\<rparr>)" 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\<lparr>measure := extreal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := extreal\<circ>distribution Y\<rparr>)
(extreal\<circ>joint_distribution X Y)"
+lemma (in information_space)
+ fixes S T X Y
+ defines "P \<equiv> S\<lparr>measure := extreal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
+ shows "indep_var S X T Y \<longleftrightarrow>
+ (random_variable S X \<and> random_variable T Y \<and>
+ measure_space.absolutely_continuous P (extreal\<circ>joint_distribution X Y) \<and>
+ integrable (P\<lparr>measure := (extreal\<circ>joint_distribution X Y)\<rparr>)
+ (entropy_density b P (extreal\<circ>joint_distribution X Y)) \<and>
+ 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 \<in> measurable M S" "sigma_algebra T" "Y \<in> measurable M T"
+ by blast+
+
+ interpret X: prob_space "S\<lparr>measure := extreal\<circ>distribution X\<rparr>"
+ by (rule distribution_prob_space) fact
+ interpret Y: prob_space "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
+ by (rule distribution_prob_space) fact
+ interpret XY: pair_prob_space "S\<lparr>measure := extreal\<circ>distribution X\<rparr>" "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>" by default
+ interpret XY: information_space XY.P b by default (rule b_gt_1)
+
+ let ?J = "XY.P\<lparr> measure := (extreal\<circ>joint_distribution X Y) \<rparr>"
+ { fix A assume "A \<in> sets XY.P"
+ then have "extreal (joint_distribution X Y A) = XY.\<mu> 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\<circ>joint_distribution X Y)"
+ by (simp add: XY.absolutely_continuous_def j_eq)
+ then show "measure_space.absolutely_continuous P (extreal\<circ>joint_distribution X Y)"
+ unfolding P_def .
+
+ have ed: "entropy_density b XY.P (extreal\<circ>joint_distribution X Y) \<in> borel_measurable XY.P"
+ by (rule XY.measurable_entropy_density) (default | fact)+
+
+ have "AE x in XY.P. 1 = RN_deriv XY.P (extreal\<circ>joint_distribution X Y) x"
+ proof (rule XY.RN_deriv_unique[OF _ ac])
+ show "measure_space ?J" by default
+ fix A assume "A \<in> sets XY.P"
+ then show "(extreal\<circ>joint_distribution X Y) A = (\<integral>\<^isup>+ x. 1 * indicator A x \<partial>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\<circ>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\<circ>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\<lparr>measure := (extreal\<circ>joint_distribution X Y)\<rparr>)
+ (entropy_density b P (extreal\<circ>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 \<in> measurable M S" "sigma_algebra T" "Y \<in> measurable M T"
+ then have rvs: "random_variable S X" "random_variable T Y" by blast+
+
+ interpret X: prob_space "S\<lparr>measure := extreal\<circ>distribution X\<rparr>"
+ by (rule distribution_prob_space) fact
+ interpret Y: prob_space "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
+ by (rule distribution_prob_space) fact
+ interpret XY: pair_prob_space "S\<lparr>measure := extreal\<circ>distribution X\<rparr>" "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>" by default
+ interpret XY: information_space XY.P b by default (rule b_gt_1)
+
+ let ?J = "XY.P\<lparr> measure := (extreal\<circ>joint_distribution X Y) \<rparr>"
+ interpret J: prob_space ?J
+ using rvs by (intro joint_distribution_prob_space) auto
+
+ assume ac: "measure_space.absolutely_continuous P (extreal\<circ>joint_distribution X Y)"
+ assume int: "integrable (P\<lparr>measure := (extreal\<circ>joint_distribution X Y)\<rparr>)
+ (entropy_density b P (extreal\<circ>joint_distribution X Y))"
+ assume I_eq_0: "mutual_information b S T X Y = 0"
+
+ have eq: "\<forall>A\<in>sets XY.P. (extreal \<circ> joint_distribution X Y) A = XY.\<mu> A"
+ proof (rule XY.KL_eq_0_imp)
+ show "prob_space ?J" by default
+ show "XY.absolutely_continuous (extreal\<circ>joint_distribution X Y)"
+ using ac by (simp add: P_def)
+ show "integrable ?J (entropy_density b XY.P (extreal\<circ>joint_distribution X Y))"
+ using int by (simp add: P_def)
+ show "KL_divergence b XY.P (extreal\<circ>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 \<lparr>space = space M, sets = {X -` A \<inter> space M |A. A \<in> sets S}\<rparr>"
+ proof (safe intro!: Int_stableI)
+ fix A B assume "A \<in> sets S" "B \<in> sets S"
+ then show "\<exists>C. (X -` A \<inter> space M) \<inter> (X -` B \<inter> space M) = (X -` C \<inter> space M) \<and> C \<in> sets S"
+ by (intro exI[of _ "A \<inter> 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 \<inter> space M |A. A \<in> sets S} {Y -` A \<inter> space M |A. A \<in> sets T}"
+ proof (safe intro!: indep_setI)
+ { fix A assume "A \<in> sets S" then show "X -` A \<inter> space M \<in> sets M"
+ using `X \<in> measurable M S` by (auto intro: measurable_sets) }
+ { fix A assume "A \<in> sets T" then show "Y -` A \<inter> space M \<in> sets M"
+ using `Y \<in> measurable M T` by (auto intro: measurable_sets) }
+ next
+ fix A B assume ab: "A \<in> sets S" "B \<in> sets T"
+ have "extreal (prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M))) =
+ extreal (joint_distribution X Y (A \<times> B))"
+ unfolding distribution_def
+ by (intro arg_cong[where f="\<lambda>C. extreal (prob C)"]) auto
+ also have "\<dots> = XY.\<mu> (A \<times> B)"
+ using ab eq by (auto simp: XY.finite_measure_eq)
+ also have "\<dots> = extreal (distribution X A) * extreal (distribution Y B)"
+ using ab by (simp add: XY.pair_measure_times)
+ finally show "prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)) =
+ prob (X -` A \<inter> space M) * prob (Y -` B \<inter> 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\<lparr>measure := extreal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := extreal\<circ>distribution Y\<rparr>) (extreal\<circ>joint_distribution X Y)"
+ shows "mutual_information b S T X Y = mutual_information b T S Y X"
+proof -
+ let ?S = "S\<lparr>measure := extreal\<circ>distribution X\<rparr>" and ?T = "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
+ 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 "(\<lambda>(x,y). (y,x)) \<in> measure_preserving
+ (P.P \<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>) (Q.P \<lparr> measure := extreal\<circ>joint_distribution Y X\<rparr>)"
+ 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=\<mu>'])
+ have "prob_space (P.P\<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>)"
+ using X Y by (auto intro!: distribution_prob_space random_variable_pairI)
+ then show "measure_space (P.P\<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>)"
+ 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\<lparr>measure := extreal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := extreal\<circ>distribution Y\<rparr>) (extreal\<circ>joint_distribution X Y)"
- shows "mutual_information b S T X Y = mutual_information b T S Y X"
-proof -
- let ?S = "S\<lparr>measure := extreal\<circ>distribution X\<rparr>" and ?T = "T\<lparr>measure := extreal\<circ>distribution Y\<rparr>"
- 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 "(\<lambda>(x,y). (y,x)) \<in> measure_preserving
- (P.P \<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>) (Q.P \<lparr> measure := extreal\<circ>joint_distribution Y X\<rparr>)"
- 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=\<mu>'])
- have "prob_space (P.P\<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>)"
- using X Y by (auto intro!: distribution_prob_space random_variable_pairI)
- then show "measure_space (P.P\<lparr> measure := extreal\<circ>joint_distribution X Y\<rparr>)"
- 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"
--- 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 "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "sets N = sets M" "space N = space M"
+ shows "integrable N f \<longleftrightarrow> 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 \<in> borel_measurable M" "g \<in> borel_measurable M"
+ assumes "AE x. f x = g x"
+ shows "integrable M f = integrable M g"
+proof -
+ have "(\<integral>\<^isup>+ x. extreal (f x) \<partial>M) = (\<integral>\<^isup>+ x. extreal (g x) \<partial>M)"
+ "(\<integral>\<^isup>+ x. extreal (- f x) \<partial>M) = (\<integral>\<^isup>+ x. extreal (- g x) \<partial>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:
"(\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> integrable M f \<longleftrightarrow> 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 \<in> borel_measurable M" "AE x. 0 \<le> f x"
+ and g: "g \<in> borel_measurable M"
+ and N: "space N = space M" "sets N = sets M"
+ and density: "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = (\<integral>\<^isup>+ x. f x * indicator A x \<partial>M)"
+ (is "\<And>A. _ \<Longrightarrow> _ = ?d A")
+ shows "integral\<^isup>L N g = (\<integral> x. f x * g x \<partial>M)" (is ?integral)
+ and "integrable N g = integrable M (\<lambda>x. f x * g x)" (is ?integrable)
+proof -
+ from f have ms: "measure_space (M\<lparr>measure := ?d\<rparr>)"
+ by (intro measure_space_density[where u="\<lambda>x. extreal (f x)"]) auto
+
+ from ms density N have "(\<integral>\<^isup>+ x. g x \<partial>N) = (\<integral>\<^isup>+ x. max 0 (extreal (g x)) \<partial>M\<lparr>measure := ?d\<rparr>)"
+ unfolding positive_integral_max_0
+ by (intro measure_space.positive_integral_cong_measure) auto
+ also have "\<dots> = (\<integral>\<^isup>+ x. extreal (f x) * max 0 (extreal (g x)) \<partial>M)"
+ using f g by (intro positive_integral_translated_density) auto
+ also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (extreal (f x * g x)) \<partial>M)"
+ using f by (intro positive_integral_cong_AE)
+ (auto simp: extreal_max_0 zero_le_mult_iff split: split_max)
+ finally have pos: "(\<integral>\<^isup>+ x. g x \<partial>N) = (\<integral>\<^isup>+ x. f x * g x \<partial>M)"
+ by (simp add: positive_integral_max_0)
+
+ from ms density N have "(\<integral>\<^isup>+ x. - (g x) \<partial>N) = (\<integral>\<^isup>+ x. max 0 (extreal (- g x)) \<partial>M\<lparr>measure := ?d\<rparr>)"
+ unfolding positive_integral_max_0
+ by (intro measure_space.positive_integral_cong_measure) auto
+ also have "\<dots> = (\<integral>\<^isup>+ x. extreal (f x) * max 0 (extreal (- g x)) \<partial>M)"
+ using f g by (intro positive_integral_translated_density) auto
+ also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (extreal (- f x * g x)) \<partial>M)"
+ using f by (intro positive_integral_cong_AE)
+ (auto simp: extreal_max_0 mult_le_0_iff split: split_max)
+ finally have neg: "(\<integral>\<^isup>+ x. - g x \<partial>N) = (\<integral>\<^isup>+ x. - (f x * g x) \<partial>M)"
+ by (simp add: positive_integral_max_0)
+
+ have g_N: "g \<in> 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 (\<lambda>x. - f x)" "(\<integral>x. - f x \<partial>M) = - integral\<^isup>L M f"
@@ -2221,9 +2281,14 @@
by (simp add: \<mu>'_def lebesgue_integral_def positive_integral_const_If)
qed
+lemma indicator_less[simp]:
+ "indicator A x \<le> (indicator B x::extreal) \<longleftrightarrow> (x \<in> A \<longrightarrow> x \<in> 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: "\<mu> (space M) \<noteq> 0"
+ assumes A: "\<mu> A \<noteq> 0" "A \<in> sets M" "AE x. x \<in> A \<longrightarrow> X x \<noteq> Y x"
+ assumes gt: "AE x. X x \<le> Y x"
shows "integral\<^isup>L M X < integral\<^isup>L M Y"
proof -
have "integral\<^isup>L M X \<le> integral\<^isup>L M Y"
@@ -2231,24 +2296,30 @@
moreover
have "integral\<^isup>L M X \<noteq> integral\<^isup>L M Y"
proof
- from gt have "AE x. Y x - X x \<noteq> 0"
- by auto
- then have \<mu>: "\<mu> {x\<in>space M. Y x - X x \<noteq> 0} = \<mu> (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 (\<lambda>x. \<bar>Y x - X x\<bar>) = integral\<^isup>L M (\<lambda>x. Y x - X x)"
using gt by (intro integral_cong_AE) auto
also have "\<dots> = 0"
using eq int by simp
- finally show False
- using \<mu> int neq0
- by (subst (asm) integral_0_iff) auto
+ finally have "\<mu> {x \<in> space M. Y x - X x \<noteq> 0} = 0"
+ using int by (simp add: integral_0_iff)
+ moreover
+ have "(\<integral>\<^isup>+x. indicator A x \<partial>M) \<le> (\<integral>\<^isup>+x. indicator {x \<in> space M. Y x - X x \<noteq> 0} x \<partial>M)"
+ using A by (intro positive_integral_mono_AE) auto
+ then have "\<mu> A \<le> \<mu> {x \<in> space M. Y x - X x \<noteq> 0}"
+ using int A by (simp add: integrable_def)
+ moreover note `\<mu> A \<noteq> 0` positive_measure[OF `A \<in> 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" "\<mu> (space M) \<noteq> 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: "\<And>i. integrable M (u i)" and bound: "\<And>x j. x\<in>space M \<Longrightarrow> \<bar>u j x\<bar> \<le> w x"
and w: "integrable M w"
--- 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 \<equiv> distribution (\<lambda>x. (X x, Y x))"
+lemma (in prob_space) prob_space_cong:
+ assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> 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 "\<And>x. x \<in> space M \<Longrightarrow> X x = Y x"
shows "distribution X = distribution Y"
@@ -54,15 +62,30 @@
lemma (in prob_space) distribution_positive[simp, intro]:
"0 \<le> distribution X A" unfolding distribution_def by auto
+lemma (in prob_space) not_zero_less_distribution[simp]:
+ "(\<not> 0 < distribution X A) \<longleftrightarrow> 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=\<mu>'])
+lemma (in prob_space) not_empty: "space M \<noteq> {}"
+ using prob_space empty_measure' by auto
+
lemma (in prob_space) measure_le_1: "X \<in> sets M \<Longrightarrow> \<mu> X \<le> 1"
unfolding measure_space_1[symmetric]
using sets_into_space
by (intro measure_mono) auto
+lemma (in prob_space) AE_I_eq_1:
+ assumes "\<mu> {x\<in>space M. P x} = 1" "{x\<in>space M. P x} \<in> sets M"
+ shows "AE x. P x"
+proof (rule AE_I)
+ show "\<mu> (space M - {x \<in> 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 \<le> 1"
unfolding distribution_def by simp
@@ -245,6 +268,146 @@
using finite_measure_eq[of "X -` A \<inter> space M"]
by (auto simp: measurable_sets distribution_def)
+lemma (in prob_space) expectation_less:
+ assumes [simp]: "integrable M X"
+ assumes gt: "\<forall>x\<in>space M. X x < b"
+ shows "expectation X < b"
+proof -
+ have "expectation X < expectation (\<lambda>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: "\<forall>x\<in>space M. a < X x"
+ shows "a < expectation X"
+proof -
+ have "expectation (\<lambda>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 \<Rightarrow> real"
+ assumes "convex_on I f"
+ assumes "x \<in> interior I" "y \<in> I"
+ shows "f y \<ge> f x + Inf ((\<lambda>t. (f x - f t) / (x - t)) ` ({x<..} \<inter> I)) * (y - x)"
+ (is "_ \<ge> _ + 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 \<in> interior I`] guess e . note e = this
+ moreover def t \<equiv> "min (x + e / 2) ((x + y) / 2)"
+ ultimately have "x < t" "t < y" "t \<in> ball x e"
+ by (auto simp: mem_ball dist_real_def field_simps split: split_min)
+ with `x \<in> interior I` e interior_subset[of I] have "t \<in> I" "x \<in> I" by auto
+
+ have "open (interior I)" by auto
+ from openE[OF this `x \<in> interior I`] guess e .
+ moreover def K \<equiv> "x - e / 2"
+ with `0 < e` have "K \<in> ball x e" "K < x" by (auto simp: mem_ball dist_real_def)
+ ultimately have "K \<in> I" "K < x" "x \<in> I"
+ using interior_subset[of I] `x \<in> interior I` by auto
+
+ have "Inf (?F x) \<le> (f x - f y) / (x - y)"
+ proof (rule Inf_lower2)
+ show "(f x - f t) / (x - t) \<in> ?F x"
+ using `t \<in> I` `x < t` by auto
+ show "(f x - f t) / (x - t) \<le> (f x - f y) / (x - y)"
+ using `convex_on I f` `x \<in> I` `y \<in> I` `x < t` `t < y` by (rule convex_on_diff)
+ next
+ fix y assume "y \<in> ?F x"
+ with order_trans[OF convex_on_diff[OF `convex_on I f` `K \<in> I` _ `K < x` _]]
+ show "(f K - f x) / (K - x) \<le> 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 \<in> interior I`] guess e . note e = this
+ moreover def t \<equiv> "x + e / 2"
+ ultimately have "x < t" "t \<in> ball x e"
+ by (auto simp: mem_ball dist_real_def field_simps)
+ with `x \<in> interior I` e interior_subset[of I] have "t \<in> I" "x \<in> I" by auto
+
+ have "(f x - f y) / (x - y) \<le> 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 \<in> ?F x"
+ with order_trans[OF convex_on_diff[OF `convex_on I f` `y \<in> I` _ `y < x`]]
+ have "(f y - f x) / (y - x) \<le> z" by auto
+ finally show "(f x - f y) / (x - y) \<le> z" .
+ next
+ have "open (interior I)" by auto
+ from openE[OF this `x \<in> interior I`] guess e . note e = this
+ then have "x + e / 2 \<in> ball x e" by (auto simp: mem_ball dist_real_def)
+ with e interior_subset[of I] have "x + e / 2 \<in> {x<..} \<inter> I" by auto
+ then show "?F x \<noteq> {}" 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 \<subseteq> I"
+ assumes I: "I = {a <..< b} \<or> I = {a <..} \<or> I = {..< b} \<or> I = UNIV"
+ assumes q: "integrable M (\<lambda>x. q (X x))" "convex_on I q"
+ shows "q (expectation X) \<le> expectation (\<lambda>x. q (X x))"
+proof -
+ let "?F x" = "Inf ((\<lambda>t. (q x - q t) / (x - t)) ` ({x<..} \<inter> I))"
+ from not_empty X(2) have "I \<noteq> {}" by auto
+
+ from I have "open I" by auto
+
+ note I
+ moreover
+ { assume "I \<subseteq> {a <..}"
+ with X have "a < expectation X"
+ by (intro expectation_greater) auto }
+ moreover
+ { assume "I \<subseteq> {..< b}"
+ with X have "expectation X < b"
+ by (intro expectation_less) auto }
+ ultimately have "expectation X \<in> I"
+ by (elim disjE) (auto simp: subset_eq)
+ moreover
+ { fix y assume y: "y \<in> I"
+ with q(2) `open I` have "Sup ((\<lambda>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 ((\<lambda>x. q x + ?F x * (expectation X - x)) ` I)"
+ by simp
+ also have "\<dots> \<le> expectation (\<lambda>w. q (X w))"
+ proof (rule Sup_least)
+ show "(\<lambda>x. q x + ?F x * (expectation X - x)) ` I \<noteq> {}"
+ using `I \<noteq> {}` by auto
+ next
+ fix k assume "k \<in> (\<lambda>x. q x + ?F x * (expectation X - x)) ` I"
+ then guess x .. note x = this
+ have "q x + ?F x * (expectation X - x) = expectation (\<lambda>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 "\<dots> \<le> expectation (\<lambda>w. q (X w))"
+ using `x \<in> 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 \<le> expectation (\<lambda>w. q (X w))" using x by auto
+ qed
+ finally show "q (expectation X) \<le> expectation (\<lambda>x. q (X x))" .
+qed
+
lemma (in prob_space) distribution_eq_translated_integral:
assumes "random_variable S X" "A \<in> sets S"
shows "distribution X A = integral\<^isup>P (S\<lparr>measure := extreal \<circ> distribution X\<rparr>) (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 \<noteq> {}"
- using prob_space empty_measure' by auto
-
lemma (in finite_prob_space) sum_over_space_eq_1: "(\<Sum>x\<in>space M. \<mu> {x}) = 1"
using measure_space_1 sum_over_space by simp
@@ -829,7 +989,7 @@
also have "\<dots> = 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)) \<noteq> 0" by fastsimp
+ hence two: "real (card (space M)) \<noteq> 0" by fastsimp
from one have three: "prob {x} \<noteq> 0" by fastsimp
thus ?thesis using one two three divide_cancel_right
by (auto simp:field_simps)
--- 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 \<nu>(2)])
qed
+lemma (in sigma_finite_measure) real_RN_deriv:
+ assumes "finite_measure (M\<lparr>measure := \<nu>\<rparr>)" (is "finite_measure ?\<nu>")
+ assumes ac: "absolutely_continuous \<nu>"
+ obtains D where "D \<in> borel_measurable M"
+ and "AE x. RN_deriv M \<nu> x = extreal (D x)"
+ and "AE x in M\<lparr>measure := \<nu>\<rparr>. 0 < D x"
+ and "\<And>x. 0 \<le> D x"
+proof
+ interpret \<nu>: finite_measure ?\<nu> by fact
+ have ms: "measure_space ?\<nu>" by default
+ note RN = RN_deriv[OF ms ac]
+
+ let ?RN = "\<lambda>t. {x \<in> space M. RN_deriv M \<nu> x = t}"
+
+ show "(\<lambda>x. real (RN_deriv M \<nu> x)) \<in> borel_measurable M"
+ using RN by auto
+
+ have "\<nu> (?RN \<infinity>) = (\<integral>\<^isup>+ x. RN_deriv M \<nu> x * indicator (?RN \<infinity>) x \<partial>M)"
+ using RN by auto
+ also have "\<dots> = (\<integral>\<^isup>+ x. \<infinity> * indicator (?RN \<infinity>) x \<partial>M)"
+ by (intro positive_integral_cong) (auto simp: indicator_def)
+ also have "\<dots> = \<infinity> * \<mu> (?RN \<infinity>)"
+ using RN by (intro positive_integral_cmult_indicator) auto
+ finally have eq: "\<nu> (?RN \<infinity>) = \<infinity> * \<mu> (?RN \<infinity>)" .
+ moreover
+ have "\<mu> (?RN \<infinity>) = 0"
+ proof (rule ccontr)
+ assume "\<mu> {x \<in> space M. RN_deriv M \<nu> x = \<infinity>} \<noteq> 0"
+ moreover from RN have "0 \<le> \<mu> {x \<in> space M. RN_deriv M \<nu> x = \<infinity>}" by auto
+ ultimately have "0 < \<mu> {x \<in> space M. RN_deriv M \<nu> x = \<infinity>}" by auto
+ with eq have "\<nu> (?RN \<infinity>) = \<infinity>" by simp
+ with \<nu>.finite_measure[of "?RN \<infinity>"] RN show False by auto
+ qed
+ ultimately have "AE x. RN_deriv M \<nu> x < \<infinity>"
+ using RN by (intro AE_iff_measurable[THEN iffD2]) auto
+ then show "AE x. RN_deriv M \<nu> x = extreal (real (RN_deriv M \<nu> x))"
+ using RN(3) by (auto simp: extreal_real)
+ then have eq: "AE x in (M\<lparr>measure := \<nu>\<rparr>) . RN_deriv M \<nu> x = extreal (real (RN_deriv M \<nu> x))"
+ using ac absolutely_continuous_AE[OF ms] by auto
+
+ show "\<And>x. 0 \<le> real (RN_deriv M \<nu> x)"
+ using RN by (auto intro: real_of_extreal_pos)
+
+ have "\<nu> (?RN 0) = (\<integral>\<^isup>+ x. RN_deriv M \<nu> x * indicator (?RN 0) x \<partial>M)"
+ using RN by auto
+ also have "\<dots> = (\<integral>\<^isup>+ x. 0 \<partial>M)"
+ by (intro positive_integral_cong) (auto simp: indicator_def)
+ finally have "AE x in (M\<lparr>measure := \<nu>\<rparr>). RN_deriv M \<nu> x \<noteq> 0"
+ using RN by (intro \<nu>.AE_iff_measurable[THEN iffD2]) auto
+ with RN(3) eq show "AE x in (M\<lparr>measure := \<nu>\<rparr>). 0 < real (RN_deriv M \<nu> x)"
+ by (auto simp: zero_less_real_of_extreal le_less)
+qed
+
lemma (in sigma_finite_measure) RN_deriv_singleton:
assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)"
and ac: "absolutely_continuous \<nu>"
--- 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 \<Colon> 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"
--- 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]