# HG changeset patch # User hoelzl # Date 1349863954 -7200 # Node ID 15ea98537c766fc885d0ac30037f30ff1a9dbc05 # Parent 8d5668f73c428bc6ce2fde75482a1e3ae5dfe782 strong nonnegativ (instead of ae nn) for induction rule diff -r 8d5668f73c42 -r 15ea98537c76 src/HOL/Probability/Lebesgue_Integration.thy --- a/src/HOL/Probability/Lebesgue_Integration.thy Wed Oct 10 12:12:33 2012 +0200 +++ b/src/HOL/Probability/Lebesgue_Integration.thy Wed Oct 10 12:12:34 2012 +0200 @@ -388,53 +388,40 @@ lemma simple_function_induct_nn[consumes 2, case_names cong set mult add]: fixes u :: "'a \ ereal" - assumes u: "simple_function M u" and nn: "AE x in M. 0 \ u x" - assumes cong: "\f g. simple_function M f \ simple_function M g \ (AE x in M. f x = g x) \ P f \ P g" + assumes u: "simple_function M u" and nn: "\x. 0 \ u x" + assumes cong: "\f g. simple_function M f \ simple_function M g \ (\x. x \ space M \ f x = g x) \ P f \ P g" assumes set: "\A. A \ sets M \ P (indicator A)" assumes mult: "\u c. 0 \ c \ simple_function M u \ (\x. 0 \ u x) \ P u \ P (\x. c * u x)" assumes add: "\u v. simple_function M u \ (\x. 0 \ u x) \ P u \ simple_function M v \ (\x. 0 \ v x) \ P v \ P (\x. v x + u x)" shows "P u" proof - - def v \ "max 0 \ u" - have v: "simple_function M v" - unfolding v_def using u by (rule simple_function_compose) - have v_nn: "\x. 0 \ v x" - unfolding v_def by (auto simp: max_def) - show ?thesis proof (rule cong) - have "AE x in M. u x = v x" - unfolding v_def using nn by eventually_elim (auto simp: max_def) - with AE_space - show "AE x in M. (\y\v ` space M. y * indicator (v -` {y} \ space M) x) = u x" - proof eventually_elim - fix x assume x: "x \ space M" and "u x = v x" - from simple_function_indicator_representation[OF v x] - show "(\y\v ` space M. y * indicator (v -` {y} \ space M) x) = u x" - unfolding `u x = v x` .. - qed + fix x assume x: "x \ space M" + from simple_function_indicator_representation[OF u x] + show "(\y\u ` space M. y * indicator (u -` {y} \ space M) x) = u x" .. next - show "simple_function M (\x. (\y\v ` space M. y * indicator (v -` {y} \ space M) x))" + show "simple_function M (\x. (\y\u ` space M. y * indicator (u -` {y} \ space M) x))" apply (subst simple_function_cong) apply (rule simple_function_indicator_representation[symmetric]) - apply (auto intro: v) + apply (auto intro: u) done next - from v v_nn have "finite (v ` space M)" "\x. x \ v ` space M \ 0 \ x" + from u nn have "finite (u ` space M)" "\x. x \ u ` space M \ 0 \ x" unfolding simple_function_def by auto - then show "P (\x. \y\v ` space M. y * indicator (v -` {y} \ space M) x)" + then show "P (\x. \y\u ` space M. y * indicator (u -` {y} \ space M) x)" proof induct case empty show ?case using set[of "{}"] by (simp add: indicator_def[abs_def]) - qed (auto intro!: add mult set simple_functionD v setsum_nonneg + qed (auto intro!: add mult set simple_functionD u setsum_nonneg simple_function_setsum) qed fact qed lemma borel_measurable_induct[consumes 2, case_names cong set mult add seq, induct set: borel_measurable]: fixes u :: "'a \ ereal" - assumes u: "u \ borel_measurable M" "AE x in M. 0 \ u x" - assumes cong: "\f g. f \ borel_measurable M \ g \ borel_measurable M \ (AE x in M. f x = g x) \ P f \ P g" + assumes u: "u \ borel_measurable M" "\x. 0 \ u x" + assumes cong: "\f g. f \ borel_measurable M \ g \ borel_measurable M \ (\x. x \ space M \ f x = g x) \ P g \ P f" assumes set: "\A. A \ sets M \ P (indicator A)" assumes mult: "\u c. 0 \ c \ u \ borel_measurable M \ (\x. 0 \ u x) \ P u \ P (\x. c * u x)" assumes add: "\u v. u \ borel_measurable M \ (\x. 0 \ u x) \ P u \ v \ borel_measurable M \ (\x. 0 \ v x) \ P v \ P (\x. v x + u x)" @@ -444,26 +431,20 @@ proof (induct rule: borel_measurable_implies_simple_function_sequence') fix U assume U: "\i. simple_function M (U i)" "incseq U" "\i. \ \ range (U i)" and sup: "\x. (SUP i. U i x) = max 0 (u x)" and nn: "\i x. 0 \ U i x" - have u_eq: "max 0 \ u = (SUP i. U i)" and "\i. AE x in M. 0 \ U i x" + have u_eq: "u = (SUP i. U i)" using nn u sup by (auto simp: max_def) from U have "\i. U i \ borel_measurable M" by (simp add: borel_measurable_simple_function) - have "max 0 \ u \ borel_measurable M" "u \ borel_measurable M" - by (auto intro!: measurable_comp u borel_measurable_ereal_max simp: id_def[symmetric]) - moreover have "AE x in M. (max 0 \ u) x = u x" - using u(2) by eventually_elim (auto simp: max_def) - moreover have "P (max 0 \ u)" + show "P u" unfolding u_eq proof (rule seq) fix i show "P (U i)" - using `simple_function M (U i)` `AE x in M. 0 \ U i x` + using `simple_function M (U i)` nn by (induct rule: simple_function_induct_nn) (auto intro: set mult add cong dest!: borel_measurable_simple_function) qed fact+ - ultimately show "P u" - by fact qed lemma simple_function_If_set: @@ -1427,23 +1408,21 @@ by (auto intro!: positive_integral_0_iff_AE[THEN iffD2]) lemma positive_integral_subalgebra: - assumes f: "f \ borel_measurable N" "AE x in N. 0 \ f x" + assumes f: "f \ borel_measurable N" "\x. 0 \ f x" and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ emeasure N A = emeasure M A" shows "integral\<^isup>P N f = integral\<^isup>P M f" proof - - from borel_measurable_implies_simple_function_sequence'[OF f(1)] guess fs . note fs = this - note sf = simple_function_subalgebra[OF fs(1) N(1,2)] - from positive_integral_monotone_convergence_simple[OF fs(2,5,1), symmetric] - have "integral\<^isup>P N f = (SUP i. \x\fs i ` space M. x * emeasure N (fs i -` {x} \ space M))" - unfolding fs(4) positive_integral_max_0 - unfolding simple_integral_def `space N = space M` by simp - also have "\ = (SUP i. \x\fs i ` space M. x * (emeasure M) (fs i -` {x} \ space M))" - using N simple_functionD(2)[OF fs(1)] unfolding `space N = space M` by auto - also have "\ = integral\<^isup>P M f" - using positive_integral_monotone_convergence_simple[OF fs(2,5) sf, symmetric] - unfolding fs(4) positive_integral_max_0 - unfolding simple_integral_def `space N = space M` by simp - finally show ?thesis . + have [simp]: "\f :: 'a \ ereal. f \ borel_measurable N \ f \ borel_measurable M" + using N by (auto simp: measurable_def) + have [simp]: "\P. (AE x in N. P x) \ (AE x in M. P x)" + using N by (auto simp add: eventually_ae_filter null_sets_def) + have [simp]: "\A. A \ sets N \ A \ sets M" + using N by auto + from f show ?thesis + apply induct + apply (simp_all add: positive_integral_add positive_integral_cmult positive_integral_monotone_convergence_SUP N) + apply (auto intro!: positive_integral_cong cong: positive_integral_cong simp: N(2)[symmetric]) + done qed section "Lebesgue Integral" @@ -2310,17 +2289,17 @@ lemma positive_integral_distr': assumes T: "T \ measurable M M'" - and f: "f \ borel_measurable (distr M M' T)" "AE x in (distr M M' T). 0 \ f x" + and f: "f \ borel_measurable (distr M M' T)" "\x. 0 \ f x" shows "integral\<^isup>P (distr M M' T) f = (\\<^isup>+ x. f (T x) \M)" using f proof induct case (cong f g) - then have eq: "AE x in (distr M M' T). g x = f x" "AE x in M. g (T x) = f (T x)" - by (auto dest: AE_distrD[OF T]) - show ?case - apply (subst eq(1)[THEN positive_integral_cong_AE]) - apply (subst eq(2)[THEN positive_integral_cong_AE]) - apply fact + with T show ?case + apply (subst positive_integral_cong[of _ f g]) + apply simp + apply (subst positive_integral_cong[of _ "\x. f (T x)" "\x. g (T x)"]) + apply (simp add: measurable_def Pi_iff) + apply simp done next case (set A) @@ -2529,16 +2508,13 @@ lemma positive_integral_density': assumes f: "f \ borel_measurable M" "AE x in M. 0 \ f x" - assumes g: "g \ borel_measurable M" "AE x in M. 0 \ g x" + assumes g: "g \ borel_measurable M" "\x. 0 \ g x" shows "integral\<^isup>P (density M f) g = (\\<^isup>+ x. f x * g x \M)" using g proof induct case (cong u v) - then have eq: "AE x in density M f. v x = u x" "AE x in M. f x * v x = f x * u x" - by (auto simp: AE_density f) - show ?case - apply (subst positive_integral_cong_AE[OF eq(1)]) - apply (subst positive_integral_cong_AE[OF eq(2)]) - apply fact + then show ?case + apply (subst positive_integral_cong[OF cong(3)]) + apply (simp_all cong: positive_integral_cong) done next case (set A) then show ?case