# HG changeset patch # User hoelzl # Date 1300109869 -3600 # Node ID cdf7693bbe08f54fb586203402d63ebb0a146b94 # Parent 28b51effc5ed3ffe196a09dcf4ef594edc4844e5 reworked Probability theory: measures are not type restricted to positive extended reals diff -r 28b51effc5ed -r cdf7693bbe08 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Mon Mar 14 14:37:47 2011 +0100 +++ b/src/HOL/IsaMakefile Mon Mar 14 14:37:49 2011 +0100 @@ -1168,9 +1168,10 @@ Multivariate_Analysis/Topology_Euclidean_Space.thy \ Multivariate_Analysis/document/root.tex \ Multivariate_Analysis/normarith.ML Library/Glbs.thy \ - Library/Indicator_Function.thy Library/Inner_Product.thy \ - Library/Numeral_Type.thy Library/Convex.thy Library/FrechetDeriv.thy \ - Library/Product_Vector.thy Library/Product_plus.thy + Library/Extended_Reals.thy Library/Indicator_Function.thy \ + Library/Inner_Product.thy Library/Numeral_Type.thy Library/Convex.thy \ + Library/FrechetDeriv.thy Library/Product_Vector.thy \ + Library/Product_plus.thy @cd Multivariate_Analysis; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL HOL-Multivariate_Analysis @@ -1185,7 +1186,6 @@ Probability/ex/Koepf_Duermuth_Countermeasure.thy \ Probability/Information.thy Probability/Lebesgue_Integration.thy \ Probability/Lebesgue_Measure.thy Probability/Measure.thy \ - Probability/Positive_Extended_Real.thy \ Probability/Probability_Space.thy Probability/Probability.thy \ Probability/Product_Measure.thy Probability/Radon_Nikodym.thy \ Probability/ROOT.ML Probability/Sigma_Algebra.thy \ diff -r 28b51effc5ed -r cdf7693bbe08 src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy --- a/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy Mon Mar 14 14:37:47 2011 +0100 +++ b/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy Mon Mar 14 14:37:49 2011 +0100 @@ -1028,7 +1028,7 @@ qed simp qed -lemma setsum_of_pextreal: +lemma setsum_real_of_extreal: assumes "\x. x \ S \ \f x\ \ \" shows "(\x\S. real (f x)) = real (setsum f S)" proof - @@ -1062,17 +1062,6 @@ by induct (auto simp: extreal_right_distrib setsum_nonneg) qed simp -lemma setsum_real_of_extreal: - assumes "\x. x \ A \ \f x\ \ \" - shows "real (\x\A. f x) = (\x\A. real (f x))" -proof cases - assume "finite A" from this assms show ?thesis - proof induct - case (insert a A) then show ?case - by (simp add: real_of_extreal_add setsum_Inf) - qed simp -qed simp - lemma sums_extreal_positive: fixes f :: "nat \ extreal" assumes "\i. 0 \ f i" shows "f sums (SUP n. \i (x::real) \ (\i. max (u i) 0) ----> max x 0" - by (fastsimp intro!: LIMSEQ_I dest!: LIMSEQ_D) - section "Generic Borel spaces" definition "borel = sigma \ space = UNIV::'a::topological_space set, sets = open\" @@ -112,26 +108,8 @@ ultimately show "?I \ borel_measurable M" by auto qed -lemma borel_measurable_translate: - assumes "A \ sets borel" and trans: "\B. open B \ f -` B \ sets borel" - shows "f -` A \ sets borel" -proof - - have "A \ sigma_sets UNIV open" using assms - by (simp add: borel_def sigma_def) - thus ?thesis - proof (induct rule: sigma_sets.induct) - case (Basic a) thus ?case using trans[of a] by (simp add: mem_def) - next - case (Compl a) - moreover have "UNIV \ sets borel" - using borel.top by simp - ultimately show ?case - by (auto simp: vimage_Diff borel.Diff) - qed (auto simp add: vimage_UN) -qed - lemma (in sigma_algebra) borel_measurable_restricted: - fixes f :: "'a \ 'x\{topological_space, semiring_1}" assumes "A \ sets M" + fixes f :: "'a \ extreal" assumes "A \ sets M" shows "f \ borel_measurable (restricted_space A) \ (\x. f x * indicator A x) \ borel_measurable M" (is "f \ borel_measurable ?R \ ?f \ borel_measurable M") @@ -142,7 +120,7 @@ show ?thesis unfolding * unfolding in_borel_measurable_borel proof (simp, safe) - fix S :: "'x set" assume "S \ sets borel" + fix S :: "extreal set" assume "S \ sets borel" "\S\sets borel. ?f -` S \ A \ op \ A ` sets M" then have "?f -` S \ A \ op \ A ` sets M" by auto then have f: "?f -` S \ A \ sets M" @@ -161,7 +139,7 @@ then show ?thesis using f by auto qed next - fix S :: "'x set" assume "S \ sets borel" + fix S :: "extreal set" assume "S \ sets borel" "\S\sets borel. ?f -` S \ space M \ sets M" then have f: "?f -` S \ space M \ sets M" by auto then show "?f -` S \ A \ op \ A ` sets M" @@ -1024,103 +1002,6 @@ using borel_measurable_euclidean_component unfolding nth_conv_component by auto -section "Borel space over the real line with infinity" - -lemma borel_Real_measurable: - "A \ sets borel \ Real -` A \ sets borel" -proof (rule borel_measurable_translate) - fix B :: "pextreal set" assume "open B" - then obtain T x where T: "open T" "Real ` (T \ {0..}) = B - {\}" and - x: "\ \ B \ 0 \ x" "\ \ B \ {Real x <..} \ B" - unfolding open_pextreal_def by blast - have "Real -` B = Real -` (B - {\})" by auto - also have "\ = Real -` (Real ` (T \ {0..}))" using T by simp - also have "\ = (if 0 \ T then T \ {.. 0} else T \ {0..})" - apply (auto simp add: Real_eq_Real image_iff) - apply (rule_tac x="max 0 x" in bexI) - by (auto simp: max_def) - finally show "Real -` B \ sets borel" - using `open T` by auto -qed simp - -lemma borel_real_measurable: - "A \ sets borel \ (real -` A :: pextreal set) \ sets borel" -proof (rule borel_measurable_translate) - fix B :: "real set" assume "open B" - { fix x have "0 < real x \ (\r>0. x = Real r)" by (cases x) auto } - note Ex_less_real = this - have *: "real -` B = (if 0 \ B then real -` (B \ {0 <..}) \ {0, \} else real -` (B \ {0 <..}))" - by (force simp: Ex_less_real) - - have "open (real -` (B \ {0 <..}) :: pextreal set)" - unfolding open_pextreal_def using `open B` - by (auto intro!: open_Int exI[of _ "B \ {0 <..}"] simp: image_iff Ex_less_real) - then show "(real -` B :: pextreal set) \ sets borel" unfolding * by auto -qed simp - -lemma (in sigma_algebra) borel_measurable_Real[intro, simp]: - assumes "f \ borel_measurable M" - shows "(\x. Real (f x)) \ borel_measurable M" - unfolding in_borel_measurable_borel -proof safe - fix S :: "pextreal set" assume "S \ sets borel" - from borel_Real_measurable[OF this] - have "(Real \ f) -` S \ space M \ sets M" - using assms - unfolding vimage_compose in_borel_measurable_borel - by auto - thus "(\x. Real (f x)) -` S \ space M \ sets M" by (simp add: comp_def) -qed - -lemma (in sigma_algebra) borel_measurable_real[intro, simp]: - fixes f :: "'a \ pextreal" - assumes "f \ borel_measurable M" - shows "(\x. real (f x)) \ borel_measurable M" - unfolding in_borel_measurable_borel -proof safe - fix S :: "real set" assume "S \ sets borel" - from borel_real_measurable[OF this] - have "(real \ f) -` S \ space M \ sets M" - using assms - unfolding vimage_compose in_borel_measurable_borel - by auto - thus "(\x. real (f x)) -` S \ space M \ sets M" by (simp add: comp_def) -qed - -lemma (in sigma_algebra) borel_measurable_Real_eq: - assumes "\x. x \ space M \ 0 \ f x" - shows "(\x. Real (f x)) \ borel_measurable M \ f \ borel_measurable M" -proof - have [simp]: "(\x. Real (f x)) -` {\} \ space M = {}" - by auto - assume "(\x. Real (f x)) \ borel_measurable M" - hence "(\x. real (Real (f x))) \ borel_measurable M" - by (rule borel_measurable_real) - moreover have "\x. x \ space M \ real (Real (f x)) = f x" - using assms by auto - ultimately show "f \ borel_measurable M" - by (simp cong: measurable_cong) -qed auto - -lemma (in sigma_algebra) borel_measurable_pextreal_eq_real: - "f \ borel_measurable M \ - ((\x. real (f x)) \ borel_measurable M \ f -` {\} \ space M \ sets M)" -proof safe - assume "f \ borel_measurable M" - then show "(\x. real (f x)) \ borel_measurable M" "f -` {\} \ space M \ sets M" - by (auto intro: borel_measurable_vimage borel_measurable_real) -next - assume *: "(\x. real (f x)) \ borel_measurable M" "f -` {\} \ space M \ sets M" - have "f -` {\} \ space M = {x\space M. f x = \}" by auto - with * have **: "{x\space M. f x = \} \ sets M" by simp - have f: "f = (\x. if f x = \ then \ else Real (real (f x)))" - by (simp add: fun_eq_iff Real_real) - show "f \ borel_measurable M" - apply (subst f) - apply (rule measurable_If) - using * ** by auto -qed - lemma borel_measurable_continuous_on1: fixes f :: "'a::topological_space \ 'b::t1_space" assumes "continuous_on UNIV f" @@ -1187,206 +1068,213 @@ using measurable_comp[OF f borel_measurable_borel_log[OF `1 < b`]] by (simp add: comp_def) +subsection "Borel space on the extended reals" + +lemma borel_measurable_extreal_borel: + "extreal \ borel_measurable borel" + unfolding borel_def[where 'a=extreal] +proof (rule borel.measurable_sigma) + fix X :: "extreal set" assume "X \ sets \space = UNIV, sets = open \" + then have "open X" by (auto simp: mem_def) + then have "open (extreal -` X \ space borel)" + by (simp add: open_extreal_vimage) + then show "extreal -` X \ space borel \ sets borel" by auto +qed auto + +lemma (in sigma_algebra) borel_measurable_extreal[simp, intro]: + assumes f: "f \ borel_measurable M" shows "(\x. extreal (f x)) \ borel_measurable M" + using measurable_comp[OF f borel_measurable_extreal_borel] unfolding comp_def . + +lemma borel_measurable_real_of_extreal_borel: + "(real :: extreal \ real) \ borel_measurable borel" + unfolding borel_def[where 'a=real] +proof (rule borel.measurable_sigma) + fix B :: "real set" assume "B \ sets \space = UNIV, sets = open \" + then have "open B" by (auto simp: mem_def) + have *: "extreal -` real -` (B - {0}) = B - {0}" by auto + have open_real: "open (real -` (B - {0}) :: extreal set)" + unfolding open_extreal_def * using `open B` by auto + show "(real -` B \ space borel :: extreal set) \ sets borel" + proof cases + assume "0 \ B" + then have *: "real -` B = real -` (B - {0}) \ {-\, \, 0}" + by (auto simp add: real_of_extreal_eq_0) + then show "(real -` B :: extreal set) \ space borel \ sets borel" + using open_real by auto + next + assume "0 \ B" + then have *: "(real -` B :: extreal set) = real -` (B - {0})" + by (auto simp add: real_of_extreal_eq_0) + then show "(real -` B :: extreal set) \ space borel \ sets borel" + using open_real by auto + qed +qed auto + +lemma (in sigma_algebra) borel_measurable_real_of_extreal[simp, intro]: + assumes f: "f \ borel_measurable M" shows "(\x. real (f x :: extreal)) \ borel_measurable M" + using measurable_comp[OF f borel_measurable_real_of_extreal_borel] unfolding comp_def . + +lemma (in sigma_algebra) borel_measurable_extreal_iff: + shows "(\x. extreal (f x)) \ borel_measurable M \ f \ borel_measurable M" +proof + assume "(\x. extreal (f x)) \ borel_measurable M" + from borel_measurable_real_of_extreal[OF this] + show "f \ borel_measurable M" by auto +qed auto + +lemma (in sigma_algebra) borel_measurable_extreal_iff_real: + "f \ borel_measurable M \ + ((\x. real (f x)) \ borel_measurable M \ f -` {\} \ space M \ sets M \ f -` {-\} \ space M \ sets M)" +proof safe + assume *: "(\x. real (f x)) \ borel_measurable M" "f -` {\} \ space M \ sets M" "f -` {-\} \ space M \ sets M" + have "f -` {\} \ space M = {x\space M. f x = \}" "f -` {-\} \ space M = {x\space M. f x = -\}" by auto + with * have **: "{x\space M. f x = \} \ sets M" "{x\space M. f x = -\} \ sets M" by simp_all + let "?f x" = "if f x = \ then \ else if f x = -\ then -\ else extreal (real (f x))" + have "?f \ borel_measurable M" using * ** by (intro measurable_If) auto + also have "?f = f" by (auto simp: fun_eq_iff extreal_real) + finally show "f \ borel_measurable M" . +qed (auto intro: measurable_sets borel_measurable_real_of_extreal) lemma (in sigma_algebra) less_eq_ge_measurable: fixes f :: "'a \ 'c::linorder" - shows "{x\space M. a < f x} \ sets M \ {x\space M. f x \ a} \ sets M" + shows "f -` {a <..} \ space M \ sets M \ f -` {..a} \ space M \ sets M" proof - assume "{x\space M. f x \ a} \ sets M" - moreover have "{x\space M. a < f x} = space M - {x\space M. f x \ a}" by auto - ultimately show "{x\space M. a < f x} \ sets M" by auto + assume "f -` {a <..} \ space M \ sets M" + moreover have "f -` {..a} \ space M = space M - f -` {a <..} \ space M" by auto + ultimately show "f -` {..a} \ space M \ sets M" by auto next - assume "{x\space M. a < f x} \ sets M" - moreover have "{x\space M. f x \ a} = space M - {x\space M. a < f x}" by auto - ultimately show "{x\space M. f x \ a} \ sets M" by auto + assume "f -` {..a} \ space M \ sets M" + moreover have "f -` {a <..} \ space M = space M - f -` {..a} \ space M" by auto + ultimately show "f -` {a <..} \ space M \ sets M" by auto qed lemma (in sigma_algebra) greater_eq_le_measurable: fixes f :: "'a \ 'c::linorder" - shows "{x\space M. f x < a} \ sets M \ {x\space M. a \ f x} \ sets M" + shows "f -` {..< a} \ space M \ sets M \ f -` {a ..} \ space M \ sets M" proof - assume "{x\space M. a \ f x} \ sets M" - moreover have "{x\space M. f x < a} = space M - {x\space M. a \ f x}" by auto - ultimately show "{x\space M. f x < a} \ sets M" by auto + assume "f -` {a ..} \ space M \ sets M" + moreover have "f -` {..< a} \ space M = space M - f -` {a ..} \ space M" by auto + ultimately show "f -` {..< a} \ space M \ sets M" by auto next - assume "{x\space M. f x < a} \ sets M" - moreover have "{x\space M. a \ f x} = space M - {x\space M. f x < a}" by auto - ultimately show "{x\space M. a \ f x} \ sets M" by auto + assume "f -` {..< a} \ space M \ sets M" + moreover have "f -` {a ..} \ space M = space M - f -` {..< a} \ space M" by auto + ultimately show "f -` {a ..} \ space M \ sets M" by auto qed -lemma (in sigma_algebra) less_eq_le_pextreal_measurable: - fixes f :: "'a \ pextreal" - shows "(\a. {x\space M. a < f x} \ sets M) \ (\a. {x\space M. a \ f x} \ sets M)" +lemma (in sigma_algebra) borel_measurable_uminus_borel_extreal: + "(uminus :: extreal \ extreal) \ borel_measurable borel" +proof (subst borel_def, rule borel.measurable_sigma) + fix X :: "extreal set" assume "X \ sets \space = UNIV, sets = open\" + then have "open X" by (simp add: mem_def) + have "uminus -` X = uminus ` X" by (force simp: image_iff) + then have "open (uminus -` X)" using `open X` extreal_open_uminus by auto + then show "uminus -` X \ space borel \ sets borel" by auto +qed auto + +lemma (in sigma_algebra) borel_measurable_uminus_extreal[intro]: + assumes "f \ borel_measurable M" + shows "(\x. - f x :: extreal) \ borel_measurable M" + using measurable_comp[OF assms borel_measurable_uminus_borel_extreal] by (simp add: comp_def) + +lemma (in sigma_algebra) borel_measurable_uminus_eq_extreal[simp]: + "(\x. - f x :: extreal) \ borel_measurable M \ f \ borel_measurable M" (is "?l = ?r") proof - assume a: "\a. {x\space M. a \ f x} \ sets M" - show "\a. {x \ space M. a < f x} \ sets M" - proof - fix a show "{x \ space M. a < f x} \ sets M" - proof (cases a) - case (preal r) - have "{x\space M. a < f x} = (\i. {x\space M. a + inverse (of_nat (Suc i)) \ f x})" - proof safe - fix x assume "a < f x" and [simp]: "x \ space M" - with ex_pextreal_inverse_of_nat_Suc_less[of "f x - a"] - obtain n where "a + inverse (of_nat (Suc n)) < f x" - by (cases "f x", auto simp: pextreal_minus_order) - then have "a + inverse (of_nat (Suc n)) \ f x" by simp - then show "x \ (\i. {x \ space M. a + inverse (of_nat (Suc i)) \ f x})" - by auto - next - fix i x assume [simp]: "x \ space M" - have "a < a + inverse (of_nat (Suc i))" using preal by auto - also assume "a + inverse (of_nat (Suc i)) \ f x" - finally show "a < f x" . - qed - with a show ?thesis by auto - qed simp + assume ?l from borel_measurable_uminus_extreal[OF this] show ?r by simp +qed auto + +lemma (in sigma_algebra) borel_measurable_eq_atMost_extreal: + "(f::'a \ extreal) \ borel_measurable M \ (\a. f -` {..a} \ space M \ sets M)" +proof (intro iffI allI) + assume pos[rule_format]: "\a. f -` {..a} \ space M \ sets M" + show "f \ borel_measurable M" + unfolding borel_measurable_extreal_iff_real borel_measurable_iff_le + proof (intro conjI allI) + fix a :: real + { fix x :: extreal assume *: "\i::nat. real i < x" + have "x = \" + proof (rule extreal_top) + fix B from real_arch_lt[of B] guess n .. + then have "extreal B < real n" by auto + with * show "B \ x" by (metis less_trans less_imp_le) + qed } + then have "f -` {\} \ space M = space M - (\i::nat. f -` {.. real i} \ space M)" + by (auto simp: not_le) + then show "f -` {\} \ space M \ sets M" using pos by (auto simp del: UN_simps intro!: Diff) + moreover + have "{-\} = {..-\}" by auto + then show "f -` {-\} \ space M \ sets M" using pos by auto + moreover have "{x\space M. f x \ extreal a} \ sets M" + using pos[of "extreal a"] by (simp add: vimage_def Int_def conj_commute) + moreover have "{w \ space M. real (f w) \ a} = + (if a < 0 then {w \ space M. f w \ extreal a} - f -` {-\} \ space M + else {w \ space M. f w \ extreal a} \ (f -` {\} \ space M) \ (f -` {-\} \ space M))" (is "?l = ?r") + proof (intro set_eqI) fix x show "x \ ?l \ x \ ?r" by (cases "f x") auto qed + ultimately show "{w \ space M. real (f w) \ a} \ sets M" by auto qed -next - assume a': "\a. {x \ space M. a < f x} \ sets M" - then have a: "\a. {x \ space M. f x \ a} \ sets M" unfolding less_eq_ge_measurable . - show "\a. {x \ space M. a \ f x} \ sets M" unfolding greater_eq_le_measurable[symmetric] - proof - fix a show "{x \ space M. f x < a} \ sets M" - proof (cases a) - case (preal r) - show ?thesis - proof cases - assume "a = 0" then show ?thesis by simp - next - assume "a \ 0" - have "{x\space M. f x < a} = (\i. {x\space M. f x \ a - inverse (of_nat (Suc i))})" - proof safe - fix x assume "f x < a" and [simp]: "x \ space M" - with ex_pextreal_inverse_of_nat_Suc_less[of "a - f x"] - obtain n where "inverse (of_nat (Suc n)) < a - f x" - using preal by (cases "f x") auto - then have "f x \ a - inverse (of_nat (Suc n)) " - using preal by (cases "f x") (auto split: split_if_asm) - then show "x \ (\i. {x \ space M. f x \ a - inverse (of_nat (Suc i))})" - by auto - next - fix i x assume [simp]: "x \ space M" - assume "f x \ a - inverse (of_nat (Suc i))" - also have "\ < a" using `a \ 0` preal by auto - finally show "f x < a" . - qed - with a show ?thesis by auto - qed - next - case infinite - have "f -` {\} \ space M = (\n. {x\space M. of_nat n < f x})" - proof (safe, simp_all, safe) - fix x assume *: "\n::nat. Real (real n) < f x" - show "f x = \" proof (rule ccontr) - assume "f x \ \" - with real_arch_lt[of "real (f x)"] obtain n where "f x < of_nat n" - by (auto simp: pextreal_noteq_omega_Ex) - with *[THEN spec, of n] show False by auto - qed - qed - with a' have \: "f -` {\} \ space M \ sets M" by auto - moreover have "{x \ space M. f x < a} = space M - f -` {\} \ space M" - using infinite by auto - ultimately show ?thesis by auto - qed - qed -qed +qed (simp add: measurable_sets) -lemma (in sigma_algebra) borel_measurable_pextreal_iff_greater: - "(f::'a \ pextreal) \ borel_measurable M \ (\a. {x\space M. a < f x} \ sets M)" -proof safe - fix a assume f: "f \ borel_measurable M" - have "{x\space M. a < f x} = f -` {a <..} \ space M" by auto - with f show "{x\space M. a < f x} \ sets M" - by (auto intro!: measurable_sets) -next - assume *: "\a. {x\space M. a < f x} \ sets M" - hence **: "\a. {x\space M. f x < a} \ sets M" - unfolding less_eq_le_pextreal_measurable - unfolding greater_eq_le_measurable . - show "f \ borel_measurable M" unfolding borel_measurable_pextreal_eq_real borel_measurable_iff_greater - proof safe - have "f -` {\} \ space M = space M - {x\space M. f x < \}" by auto - then show \: "f -` {\} \ space M \ sets M" using ** by auto - fix a - have "{w \ space M. a < real (f w)} = - (if 0 \ a then {w\space M. Real a < f w} - (f -` {\} \ space M) else space M)" - proof (split split_if, safe del: notI) - fix x assume "0 \ a" - { assume "a < real (f x)" then show "Real a < f x" "x \ f -` {\} \ space M" - using `0 \ a` by (cases "f x", auto) } - { assume "Real a < f x" "x \ f -` {\}" then show "a < real (f x)" - using `0 \ a` by (cases "f x", auto) } - next - fix x assume "\ 0 \ a" then show "a < real (f x)" by (cases "f x") auto - qed - then show "{w \ space M. a < real (f w)} \ sets M" - using \ * by (auto intro!: Diff) - qed -qed +lemma (in sigma_algebra) borel_measurable_eq_atLeast_extreal: + "(f::'a \ extreal) \ borel_measurable M \ (\a. f -` {a..} \ space M \ sets M)" +proof + assume pos: "\a. f -` {a..} \ space M \ sets M" + moreover have "\a. (\x. - f x) -` {..a} = f -` {-a ..}" + by (auto simp: extreal_uminus_le_reorder) + ultimately have "(\x. - f x) \ borel_measurable M" + unfolding borel_measurable_eq_atMost_extreal by auto + then show "f \ borel_measurable M" by simp +qed (simp add: measurable_sets) -lemma (in sigma_algebra) borel_measurable_pextreal_iff_less: - "(f::'a \ pextreal) \ borel_measurable M \ (\a. {x\space M. f x < a} \ sets M)" - using borel_measurable_pextreal_iff_greater unfolding less_eq_le_pextreal_measurable greater_eq_le_measurable . +lemma (in sigma_algebra) borel_measurable_extreal_iff_less: + "(f::'a \ extreal) \ borel_measurable M \ (\a. f -` {..< a} \ space M \ sets M)" + unfolding borel_measurable_eq_atLeast_extreal greater_eq_le_measurable .. -lemma (in sigma_algebra) borel_measurable_pextreal_iff_le: - "(f::'a \ pextreal) \ borel_measurable M \ (\a. {x\space M. f x \ a} \ sets M)" - using borel_measurable_pextreal_iff_greater unfolding less_eq_ge_measurable . +lemma (in sigma_algebra) borel_measurable_extreal_iff_ge: + "(f::'a \ extreal) \ borel_measurable M \ (\a. f -` {a <..} \ space M \ sets M)" + unfolding borel_measurable_eq_atMost_extreal less_eq_ge_measurable .. -lemma (in sigma_algebra) borel_measurable_pextreal_iff_ge: - "(f::'a \ pextreal) \ borel_measurable M \ (\a. {x\space M. a \ f x} \ sets M)" - using borel_measurable_pextreal_iff_greater unfolding less_eq_le_pextreal_measurable . - -lemma (in sigma_algebra) borel_measurable_pextreal_eq_const: - fixes f :: "'a \ pextreal" assumes "f \ borel_measurable M" +lemma (in sigma_algebra) borel_measurable_extreal_eq_const: + fixes f :: "'a \ extreal" assumes "f \ borel_measurable M" shows "{x\space M. f x = c} \ sets M" proof - have "{x\space M. f x = c} = (f -` {c} \ space M)" by auto then show ?thesis using assms by (auto intro!: measurable_sets) qed -lemma (in sigma_algebra) borel_measurable_pextreal_neq_const: - fixes f :: "'a \ pextreal" - assumes "f \ borel_measurable M" +lemma (in sigma_algebra) borel_measurable_extreal_neq_const: + fixes f :: "'a \ extreal" assumes "f \ borel_measurable M" shows "{x\space M. f x \ c} \ sets M" proof - have "{x\space M. f x \ c} = space M - (f -` {c} \ space M)" by auto then show ?thesis using assms by (auto intro!: measurable_sets) qed -lemma (in sigma_algebra) borel_measurable_pextreal_less[intro,simp]: - fixes f g :: "'a \ pextreal" +lemma (in sigma_algebra) borel_measurable_extreal_le[intro,simp]: + fixes f g :: "'a \ extreal" + assumes f: "f \ borel_measurable M" + assumes g: "g \ borel_measurable M" + shows "{x \ space M. f x \ g x} \ sets M" +proof - + have "{x \ space M. f x \ g x} = + {x \ space M. real (f x) \ real (g x)} - (f -` {\, -\} \ space M \ g -` {\, -\} \ space M) \ + f -` {-\} \ space M \ g -` {\} \ space M" (is "?l = ?r") + proof (intro set_eqI) + fix x show "x \ ?l \ x \ ?r" by (cases rule: extreal2_cases[of "f x" "g x"]) auto + qed + with f g show ?thesis by (auto intro!: Un simp: measurable_sets) +qed + +lemma (in sigma_algebra) borel_measurable_extreal_less[intro,simp]: + fixes f :: "'a \ extreal" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" shows "{x \ space M. f x < g x} \ sets M" proof - - have "(\x. real (f x)) \ borel_measurable M" - "(\x. real (g x)) \ borel_measurable M" - using assms by (auto intro!: borel_measurable_real) - from borel_measurable_less[OF this] - have "{x \ space M. real (f x) < real (g x)} \ sets M" . - moreover have "{x \ space M. f x \ \} \ sets M" using f by (rule borel_measurable_pextreal_neq_const) - moreover have "{x \ space M. g x = \} \ sets M" using g by (rule borel_measurable_pextreal_eq_const) - moreover have "{x \ space M. g x \ \} \ sets M" using g by (rule borel_measurable_pextreal_neq_const) - moreover have "{x \ space M. f x < g x} = ({x \ space M. g x = \} \ {x \ space M. f x \ \}) \ - ({x \ space M. g x \ \} \ {x \ space M. f x \ \} \ {x \ space M. real (f x) < real (g x)})" - by (auto simp: real_of_pextreal_strict_mono_iff) - ultimately show ?thesis by auto -qed - -lemma (in sigma_algebra) borel_measurable_pextreal_le[intro,simp]: - fixes f :: "'a \ pextreal" - assumes f: "f \ borel_measurable M" - assumes g: "g \ borel_measurable M" - shows "{x \ space M. f x \ g x} \ sets M" -proof - - have "{x \ space M. f x \ g x} = space M - {x \ space M. g x < f x}" by auto + have "{x \ space M. f x < g x} = space M - {x \ space M. g x \ f x}" by auto then show ?thesis using g f by auto qed -lemma (in sigma_algebra) borel_measurable_pextreal_eq[intro,simp]: - fixes f :: "'a \ pextreal" +lemma (in sigma_algebra) borel_measurable_extreal_eq[intro,simp]: + fixes f :: "'a \ extreal" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" shows "{w \ space M. f w = g w} \ sets M" @@ -1395,8 +1283,8 @@ then show ?thesis using g f by auto qed -lemma (in sigma_algebra) borel_measurable_pextreal_neq[intro,simp]: - fixes f :: "'a \ pextreal" +lemma (in sigma_algebra) borel_measurable_extreal_neq[intro,simp]: + fixes f :: "'a \ extreal" assumes f: "f \ borel_measurable M" assumes g: "g \ borel_measurable M" shows "{w \ space M. f w \ g w} \ sets M" @@ -1405,20 +1293,28 @@ thus ?thesis using f g by auto qed -lemma (in sigma_algebra) borel_measurable_pextreal_add[intro, simp]: - fixes f :: "'a \ pextreal" +lemma (in sigma_algebra) split_sets: + "{x\space M. P x \ Q x} = {x\space M. P x} \ {x\space M. Q x}" + "{x\space M. P x \ Q x} = {x\space M. P x} \ {x\space M. Q x}" + by auto + +lemma (in sigma_algebra) borel_measurable_extreal_add[intro, simp]: + fixes f :: "'a \ extreal" assumes "f \ borel_measurable M" "g \ borel_measurable M" shows "(\x. f x + g x) \ borel_measurable M" proof - - have *: "(\x. f x + g x) = - (\x. if f x = \ then \ else if g x = \ then \ else Real (real (f x) + real (g x)))" - by (auto simp: fun_eq_iff pextreal_noteq_omega_Ex) - show ?thesis using assms unfolding * - by (auto intro!: measurable_If) + { fix x assume "x \ space M" then have "f x + g x = + (if f x = \ \ g x = \ then \ + else if f x = -\ \ g x = -\ then -\ + else extreal (real (f x) + real (g x)))" + by (cases rule: extreal2_cases[of "f x" "g x"]) auto } + with assms show ?thesis + by (auto cong: measurable_cong simp: split_sets + intro!: Un measurable_If measurable_sets) qed -lemma (in sigma_algebra) borel_measurable_pextreal_setsum[simp, intro]: - fixes f :: "'c \ 'a \ pextreal" +lemma (in sigma_algebra) borel_measurable_extreal_setsum[simp, intro]: + fixes f :: "'c \ 'a \ extreal" assumes "\i. i \ S \ f i \ borel_measurable M" shows "(\x. \i\S. f i x) \ borel_measurable M" proof cases @@ -1427,20 +1323,49 @@ by induct auto qed (simp add: borel_measurable_const) -lemma (in sigma_algebra) borel_measurable_pextreal_times[intro, simp]: - fixes f :: "'a \ pextreal" assumes "f \ borel_measurable M" "g \ borel_measurable M" +lemma abs_extreal_ge0[simp]: "0 \ x \ \x :: extreal\ = x" + by (cases x) auto + +lemma abs_extreal_less0[simp]: "x < 0 \ \x :: extreal\ = -x" + by (cases x) auto + +lemma abs_extreal_pos[simp]: "0 \ \x :: extreal\" + by (cases x) auto + +lemma (in sigma_algebra) borel_measurable_extreal_abs[intro, simp]: + fixes f :: "'a \ extreal" assumes "f \ borel_measurable M" + shows "(\x. \f x\) \ borel_measurable M" +proof - + { fix x have "\f x\ = (if 0 \ f x then f x else - f x)" by auto } + then show ?thesis using assms by (auto intro!: measurable_If) +qed + +lemma (in sigma_algebra) borel_measurable_extreal_times[intro, simp]: + fixes f :: "'a \ extreal" assumes "f \ borel_measurable M" "g \ borel_measurable M" shows "(\x. f x * g x) \ borel_measurable M" proof - + { fix f g :: "'a \ extreal" + assume b: "f \ borel_measurable M" "g \ borel_measurable M" + and pos: "\x. 0 \ f x" "\x. 0 \ g x" + { fix x have *: "f x * g x = (if f x = 0 \ g x = 0 then 0 + else if f x = \ \ g x = \ then \ + else extreal (real (f x) * real (g x)))" + apply (cases rule: extreal2_cases[of "f x" "g x"]) + using pos[of x] by auto } + with b have "(\x. f x * g x) \ borel_measurable M" + by (auto cong: measurable_cong simp: split_sets + intro!: Un measurable_If measurable_sets) } + note pos_times = this have *: "(\x. f x * g x) = - (\x. if f x = 0 then 0 else if g x = 0 then 0 else if f x = \ then \ else if g x = \ then \ else - Real (real (f x) * real (g x)))" - by (auto simp: fun_eq_iff pextreal_noteq_omega_Ex) + (\x. if 0 \ f x \ 0 \ g x \ f x < 0 \ g x < 0 then \f x\ * \g x\ else - (\f x\ * \g x\))" + by (auto simp: fun_eq_iff) show ?thesis using assms unfolding * - by (auto intro!: measurable_If) + by (intro measurable_If pos_times borel_measurable_uminus_extreal) + (auto simp: split_sets intro!: Int) qed -lemma (in sigma_algebra) borel_measurable_pextreal_setprod[simp, intro]: - fixes f :: "'c \ 'a \ pextreal" +lemma (in sigma_algebra) borel_measurable_extreal_setprod[simp, intro]: + fixes f :: "'c \ 'a \ extreal" assumes "\i. i \ S \ f i \ borel_measurable M" shows "(\x. \i\S. f i x) \ borel_measurable M" proof cases @@ -1448,64 +1373,73 @@ thus ?thesis using assms by induct auto qed simp -lemma (in sigma_algebra) borel_measurable_pextreal_min[simp, intro]: - fixes f g :: "'a \ pextreal" +lemma (in sigma_algebra) borel_measurable_extreal_min[simp, intro]: + fixes f g :: "'a \ extreal" assumes "f \ borel_measurable M" assumes "g \ borel_measurable M" shows "(\x. min (g x) (f x)) \ borel_measurable M" using assms unfolding min_def by (auto intro!: measurable_If) -lemma (in sigma_algebra) borel_measurable_pextreal_max[simp, intro]: - fixes f g :: "'a \ pextreal" +lemma (in sigma_algebra) borel_measurable_extreal_max[simp, intro]: + fixes f g :: "'a \ extreal" assumes "f \ borel_measurable M" and "g \ borel_measurable M" shows "(\x. max (g x) (f x)) \ borel_measurable M" using assms unfolding max_def by (auto intro!: measurable_If) lemma (in sigma_algebra) borel_measurable_SUP[simp, intro]: - fixes f :: "'d\countable \ 'a \ pextreal" + fixes f :: "'d\countable \ 'a \ extreal" assumes "\i. i \ A \ f i \ borel_measurable M" shows "(\x. SUP i : A. f i x) \ borel_measurable M" (is "?sup \ borel_measurable M") - unfolding borel_measurable_pextreal_iff_greater -proof safe + unfolding borel_measurable_extreal_iff_ge +proof fix a - have "{x\space M. a < ?sup x} = (\i\A. {x\space M. a < f i x})" + have "?sup -` {a<..} \ space M = (\i\A. {x\space M. a < f i x})" by (auto simp: less_SUP_iff SUPR_apply) - then show "{x\space M. a < ?sup x} \ sets M" + then show "?sup -` {a<..} \ space M \ sets M" using assms by auto qed lemma (in sigma_algebra) borel_measurable_INF[simp, intro]: - fixes f :: "'d :: countable \ 'a \ pextreal" + fixes f :: "'d :: countable \ 'a \ extreal" assumes "\i. i \ A \ f i \ borel_measurable M" shows "(\x. INF i : A. f i x) \ borel_measurable M" (is "?inf \ borel_measurable M") - unfolding borel_measurable_pextreal_iff_less -proof safe + unfolding borel_measurable_extreal_iff_less +proof fix a - have "{x\space M. ?inf x < a} = (\i\A. {x\space M. f i x < a})" + have "?inf -` {.. space M = (\i\A. {x\space M. f i x < a})" by (auto simp: INF_less_iff INFI_apply) - then show "{x\space M. ?inf x < a} \ sets M" + then show "?inf -` {.. space M \ sets M" using assms by auto qed -lemma (in sigma_algebra) borel_measurable_pextreal_diff[simp, intro]: - fixes f g :: "'a \ pextreal" +lemma (in sigma_algebra) borel_measurable_liminf[simp, intro]: + fixes f :: "nat \ 'a \ extreal" + assumes "\i. f i \ borel_measurable M" + shows "(\x. liminf (\i. f i x)) \ borel_measurable M" + unfolding liminf_SUPR_INFI using assms by auto + +lemma (in sigma_algebra) borel_measurable_limsup[simp, intro]: + fixes f :: "nat \ 'a \ extreal" + assumes "\i. f i \ borel_measurable M" + shows "(\x. limsup (\i. f i x)) \ borel_measurable M" + unfolding limsup_INFI_SUPR using assms by auto + +lemma (in sigma_algebra) borel_measurable_extreal_diff[simp, intro]: + fixes f g :: "'a \ extreal" assumes "f \ borel_measurable M" assumes "g \ borel_measurable M" shows "(\x. f x - g x) \ borel_measurable M" - unfolding borel_measurable_pextreal_iff_greater -proof safe - fix a - have "{x \ space M. a < f x - g x} = {x \ space M. g x + a < f x}" - by (simp add: pextreal_less_minus_iff) - then show "{x \ space M. a < f x - g x} \ sets M" - using assms by auto -qed + unfolding minus_extreal_def using assms by auto lemma (in sigma_algebra) borel_measurable_psuminf[simp, intro]: - assumes "\i. f i \ borel_measurable M" - shows "(\x. (\\<^isub>\ i. f i x)) \ borel_measurable M" - using assms unfolding psuminf_def by auto + fixes f :: "nat \ 'a \ extreal" + assumes "\i. f i \ borel_measurable M" and pos: "\i x. x \ space M \ 0 \ f i x" + shows "(\x. (\i. f i x)) \ borel_measurable M" + apply (subst measurable_cong) + apply (subst suminf_extreal_eq_SUPR) + apply (rule pos) + using assms by auto section "LIMSEQ is borel measurable" @@ -1515,28 +1449,11 @@ and u: "\i. u i \ borel_measurable M" shows "u' \ borel_measurable M" proof - - let "?pu x i" = "max (u i x) 0" - let "?nu x i" = "max (- u i x) 0" - { fix x assume x: "x \ space M" - have "(?pu x) ----> max (u' x) 0" - "(?nu x) ----> max (- u' x) 0" - using u'[OF x] by (auto intro!: LIMSEQ_max LIMSEQ_minus) - from LIMSEQ_imp_lim_INF[OF _ this(1)] LIMSEQ_imp_lim_INF[OF _ this(2)] - have "(SUP n. INF m. Real (u (n + m) x)) = Real (u' x)" - "(SUP n. INF m. Real (- u (n + m) x)) = Real (- u' x)" - by (simp_all add: Real_max'[symmetric]) } - note eq = this - have *: "\x. real (Real (u' x)) - real (Real (- u' x)) = u' x" + have "\x. x \ space M \ liminf (\n. extreal (u n x)) = extreal (u' x)" + using u' by (simp add: lim_imp_Liminf trivial_limit_sequentially lim_extreal) + moreover from u have "(\x. liminf (\n. extreal (u n x))) \ borel_measurable M" by auto - have "(\x. SUP n. INF m. Real (u (n + m) x)) \ borel_measurable M" - "(\x. SUP n. INF m. Real (- u (n + m) x)) \ borel_measurable M" - using u by auto - with eq[THEN measurable_cong, of M "\x. x" borel] - have "(\x. Real (u' x)) \ borel_measurable M" - "(\x. Real (- u' x)) \ borel_measurable M" by auto - note this[THEN borel_measurable_real] - from borel_measurable_diff[OF this] - show ?thesis unfolding * . + ultimately show ?thesis by (simp cong: measurable_cong add: borel_measurable_extreal_iff) qed end diff -r 28b51effc5ed -r cdf7693bbe08 src/HOL/Probability/Caratheodory.thy --- a/src/HOL/Probability/Caratheodory.thy Mon Mar 14 14:37:47 2011 +0100 +++ b/src/HOL/Probability/Caratheodory.thy Mon Mar 14 14:37:49 2011 +0100 @@ -1,36 +1,66 @@ header {*Caratheodory Extension Theorem*} theory Caratheodory - imports Sigma_Algebra Positive_Extended_Real + imports Sigma_Algebra Extended_Real_Limits begin +lemma suminf_extreal_2dimen: + fixes f:: "nat \ nat \ extreal" + assumes pos: "\p. 0 \ f p" + assumes "\m. g m = (\n. f (m,n))" + shows "(\i. f (prod_decode i)) = suminf g" +proof - + have g_def: "g = (\m. (\n. f (m,n)))" + using assms by (simp add: fun_eq_iff) + have reindex: "\B. (\x\B. f (prod_decode x)) = setsum f (prod_decode ` B)" + by (simp add: setsum_reindex[OF inj_prod_decode] comp_def) + { fix n + let ?M = "\f. Suc (Max (f ` prod_decode ` {.. setsum f ({.. {..a b. setsum f (prod_decode ` {.. setsum f ({.. {.. ?M" + by (auto intro!: Max_ge le_imp_less_Suc image_eqI[where x="prod_encode (a', b')"]) } + then have "setsum f ({.. {.. setsum f ?M" + by (auto intro!: setsum_mono3 simp: pos) } + ultimately + show ?thesis unfolding g_def using pos + by (auto intro!: SUPR_eq simp: setsum_cartesian_product reindex le_SUPI2 + setsum_nonneg suminf_extreal_eq_SUPR SUPR_pair + SUPR_extreal_setsum[symmetric] incseq_setsumI setsum_nonneg) +qed + text{*From the Hurd/Coble measure theory development, translated by Lawrence Paulson.*} subsection {* Measure Spaces *} record 'a measure_space = "'a algebra" + - measure :: "'a set \ pextreal" + measure :: "'a set \ extreal" -definition positive where "positive M f \ f {} = (0::pextreal)" - -- "Positive is enforced by the type" +definition positive where "positive M f \ f {} = (0::extreal) \ (\A\sets M. 0 \ f A)" definition additive where "additive M f \ (\x \ sets M. \y \ sets M. x \ y = {} \ f (x \ y) = f x + f y)" -definition countably_additive where "countably_additive M f \ - (\A. range A \ sets M \ disjoint_family A \ (\i. A i) \ sets M \ - (\\<^isub>\ n. f (A n)) = f (\i. A i))" +definition countably_additive :: "('a, 'b) algebra_scheme \ ('a set \ extreal) \ bool" where + "countably_additive M f \ (\A. range A \ sets M \ disjoint_family A \ (\i. A i) \ sets M \ + (\i. f (A i)) = f (\i. A i))" definition increasing where "increasing M f \ (\x \ sets M. \y \ sets M. x \ y \ f x \ f y)" definition subadditive where "subadditive M f \ - (\x \ sets M. \y \ sets M. x \ y = {} \ - f (x \ y) \ f x + f y)" + (\x \ sets M. \y \ sets M. x \ y = {} \ f (x \ y) \ f x + f y)" definition countably_subadditive where "countably_subadditive M f \ (\A. range A \ sets M \ disjoint_family A \ (\i. A i) \ sets M \ - f (\i. A i) \ (\\<^isub>\ n. f (A n)))" + (f (\i. A i) \ (\i. f (A i))))" definition lambda_system where "lambda_system M f = {l \ sets M. \x \ sets M. f (l \ x) + f ((space M - l) \ x) = f x}" @@ -39,14 +69,19 @@ positive M f \ increasing M f \ countably_subadditive M f" definition measure_set where "measure_set M f X = {r. - \A. range A \ sets M \ disjoint_family A \ X \ (\i. A i) \ (\\<^isub>\ i. f (A i)) = r}" + \A. range A \ sets M \ disjoint_family A \ X \ (\i. A i) \ (\i. f (A i)) = r}" locale measure_space = sigma_algebra M for M :: "('a, 'b) measure_space_scheme" + - assumes empty_measure [simp]: "measure M {} = 0" + assumes measure_positive: "positive M (measure M)" and ca: "countably_additive M (measure M)" abbreviation (in measure_space) "\ \ measure M" +lemma (in measure_space) + shows empty_measure[simp, intro]: "\ {} = 0" + and positive_measure[simp, intro!]: "\A. A \ sets M \ 0 \ \ A" + using measure_positive unfolding positive_def by auto + lemma increasingD: "increasing M f \ x \ y \ x\sets M \ y\sets M \ f x \ f y" by (auto simp add: increasing_def) @@ -61,39 +96,30 @@ \ f (x \ y) = f x + f y" by (auto simp add: additive_def) -lemma countably_additiveD: - "countably_additive M f \ range A \ sets M \ disjoint_family A - \ (\i. A i) \ sets M \ (\\<^isub>\ n. f (A n)) = f (\i. A i)" - by (simp add: countably_additive_def) - -lemma countably_subadditiveD: - "countably_subadditive M f \ range A \ sets M \ disjoint_family A \ - (\i. A i) \ sets M \ f (\i. A i) \ psuminf (f o A)" - by (auto simp add: countably_subadditive_def o_def) - lemma countably_additiveI: - "(\A. range A \ sets M \ disjoint_family A \ (\i. A i) \ sets M - \ (\\<^isub>\ n. f (A n)) = f (\i. A i)) \ countably_additive M f" - by (simp add: countably_additive_def) + assumes "\A x. range A \ sets M \ disjoint_family A \ (\i. A i) \ sets M + \ (\i. f (A i)) = f (\i. A i)" + shows "countably_additive M f" + using assms by (simp add: countably_additive_def) section "Extend binary sets" lemma LIMSEQ_binaryset: assumes f: "f {} = 0" - shows "(\n. \i = 0.. f A + f B" + shows "(\n. \i f A + f B" proof - - have "(\n. \i = 0..< Suc (Suc n). f (binaryset A B i)) = (\n. f A + f B)" + have "(\n. \i < Suc (Suc n). f (binaryset A B i)) = (\n. f A + f B)" proof fix n - show "(\i\nat = 0\nat..i < Suc (Suc n). f (binaryset A B i)) = f A + f B" by (induct n) (auto simp add: binaryset_def f) qed moreover have "... ----> f A + f B" by (rule LIMSEQ_const) ultimately - have "(\n. \i = 0..< Suc (Suc n). f (binaryset A B i)) ----> f A + f B" + have "(\n. \i< Suc (Suc n). f (binaryset A B i)) ----> f A + f B" by metis - hence "(\n. \i = 0..< n+2. f (binaryset A B i)) ----> f A + f B" + hence "(\n. \i< n+2. f (binaryset A B i)) ----> f A + f B" by simp thus ?thesis by (rule LIMSEQ_offset [where k=2]) qed @@ -101,28 +127,13 @@ lemma binaryset_sums: assumes f: "f {} = 0" shows "(\n. f (binaryset A B n)) sums (f A + f B)" - by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f]) + by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f] atLeast0LessThan) lemma suminf_binaryset_eq: - fixes f :: "'a set \ real" + fixes f :: "'a set \ 'b::{comm_monoid_add, t2_space}" shows "f {} = 0 \ (\n. f (binaryset A B n)) = f A + f B" by (metis binaryset_sums sums_unique) -lemma binaryset_psuminf: - assumes "f {} = 0" - shows "(\\<^isub>\ n. f (binaryset A B n)) = f A + f B" (is "?suminf = ?sum") -proof - - have *: "{..<2} = {0, 1::nat}" by auto - have "\n\2. f (binaryset A B n) = 0" - unfolding binaryset_def - using assms by auto - hence "?suminf = (\N<2. f (binaryset A B N))" - by (rule psuminf_finite) - also have "... = ?sum" unfolding * binaryset_def - by simp - finally show ?thesis . -qed - subsection {* Lambda Systems *} lemma (in algebra) lambda_system_eq: @@ -144,7 +155,7 @@ by (simp add: lambda_system_def) lemma (in algebra) lambda_system_Compl: - fixes f:: "'a set \ pextreal" + fixes f:: "'a set \ extreal" assumes x: "x \ lambda_system M f" shows "space M - x \ lambda_system M f" proof - @@ -157,7 +168,7 @@ qed lemma (in algebra) lambda_system_Int: - fixes f:: "'a set \ pextreal" + fixes f:: "'a set \ extreal" assumes xl: "x \ lambda_system M f" and yl: "y \ lambda_system M f" shows "x \ y \ lambda_system M f" proof - @@ -191,7 +202,7 @@ qed lemma (in algebra) lambda_system_Un: - fixes f:: "'a set \ pextreal" + fixes f:: "'a set \ extreal" assumes xl: "x \ lambda_system M f" and yl: "y \ lambda_system M f" shows "x \ y \ lambda_system M f" proof - @@ -250,53 +261,54 @@ by (auto simp add: disjoint_family_on_def binaryset_def) hence "range (binaryset x y) \ sets M \ (\i. binaryset x y i) \ sets M \ - f (\i. binaryset x y i) \ (\\<^isub>\ n. f (binaryset x y n))" - using cs by (simp add: countably_subadditive_def) + f (\i. binaryset x y i) \ (\ n. f (binaryset x y n))" + using cs by (auto simp add: countably_subadditive_def) hence "{x,y,{}} \ sets M \ x \ y \ sets M \ - f (x \ y) \ (\\<^isub>\ n. f (binaryset x y n))" + f (x \ y) \ (\ n. f (binaryset x y n))" by (simp add: range_binaryset_eq UN_binaryset_eq) thus "f (x \ y) \ f x + f y" using f x y - by (auto simp add: Un o_def binaryset_psuminf positive_def) + by (auto simp add: Un o_def suminf_binaryset_eq positive_def) qed lemma (in algebra) additive_sum: fixes A:: "nat \ 'a set" - assumes f: "positive M f" and ad: "additive M f" + assumes f: "positive M f" and ad: "additive M f" and "finite S" and A: "range A \ sets M" - and disj: "disjoint_family A" - shows "setsum (f \ A) {0..i\{0..i\S. f (A i)) = f (\i\S. A i)" +using `finite S` disj proof induct + case empty show ?case using f by (simp add: positive_def) next - case (Suc n) - have "A n \ (\i\{0.. (\i\S. A i) = {}" + by (auto simp add: disjoint_family_on_def neq_iff) moreover - have "A n \ sets M" using A by blast - moreover have "(\i\{0.. sets M" - by (metis A UNION_in_sets atLeast0LessThan) + have "A s \ sets M" using A by blast + moreover have "(\i\S. A i) \ sets M" + using A `finite S` by auto moreover - ultimately have "f (A n \ (\i\{0..i\{0.. (\i\S. A i)) = f (A s) + f(\i\S. A i)" using ad UNION_in_sets A by (auto simp add: additive_def) - with Suc.hyps show ?case using ad - by (auto simp add: atLeastLessThanSuc additive_def) + with insert show ?case using ad disjoint_family_on_mono[of S "insert s S" A] + by (auto simp add: additive_def subset_insertI) qed lemma (in algebra) increasing_additive_bound: - fixes A:: "nat \ 'a set" and f :: "'a set \ pextreal" + fixes A:: "nat \ 'a set" and f :: "'a set \ extreal" assumes f: "positive M f" and ad: "additive M f" and inc: "increasing M f" and A: "range A \ sets M" and disj: "disjoint_family A" - shows "psuminf (f \ A) \ f (space M)" -proof (safe intro!: psuminf_bound) + shows "(\i. f (A i)) \ f (space M)" +proof (safe intro!: suminf_bound) fix N - have "setsum (f \ A) {0..i\{0..ii\{.. f (space M)" using space_closed A - by (blast intro: increasingD [OF inc] UNION_in_sets top) - finally show "setsum (f \ A) {.. f (space M)" by (simp add: atLeast0LessThan) -qed + by (intro increasingD[OF inc] finite_UN) auto + finally show "(\i f (space M)" by simp +qed (insert f A, auto simp: positive_def) lemma lambda_system_increasing: "increasing M f \ increasing (M (|sets := lambda_system M f|)) f" @@ -307,7 +319,7 @@ by (simp add: positive_def lambda_system_def) lemma (in algebra) lambda_system_strong_sum: - fixes A:: "nat \ 'a set" and f :: "'a set \ pextreal" + fixes A:: "nat \ 'a set" and f :: "'a set \ extreal" assumes f: "positive M f" and a: "a \ sets M" and A: "range A \ lambda_system M f" and disj: "disjoint_family A" @@ -331,7 +343,7 @@ assumes oms: "outer_measure_space M f" and A: "range A \ lambda_system M f" and disj: "disjoint_family A" - shows "(\i. A i) \ lambda_system M f \ psuminf (f \ A) = f (\i. A i)" + shows "(\i. A i) \ lambda_system M f \ (\i. f (A i)) = f (\i. A i)" proof - have pos: "positive M f" and inc: "increasing M f" and csa: "countably_subadditive M f" @@ -347,15 +359,16 @@ have U_in: "(\i. A i) \ sets M" by (metis A'' countable_UN) - have U_eq: "f (\i. A i) = psuminf (f o A)" + have U_eq: "f (\i. A i) = (\i. f (A i))" proof (rule antisym) - show "f (\i. A i) \ psuminf (f \ A)" - by (rule countably_subadditiveD [OF csa A'' disj U_in]) - show "psuminf (f \ A) \ f (\i. A i)" - by (rule psuminf_bound, unfold atLeast0LessThan[symmetric]) - (metis algebra.additive_sum [OF alg_ls] pos disj UN_Un Un_UNIV_right - lambda_system_positive lambda_system_additive - subset_Un_eq increasingD [OF inc] A' A'' UNION_in_sets U_in) + show "f (\i. A i) \ (\i. f (A i))" + using csa[unfolded countably_subadditive_def] A'' disj U_in by auto + have *: "\i. 0 \ f (A i)" using pos A'' unfolding positive_def by auto + have dis: "\N. disjoint_family_on A {..i. f (A i)) \ f (\i. A i)" + using algebra.additive_sum [OF alg_ls lambda_system_positive[OF pos] lambda_system_additive _ A' dis] + using A'' + by (intro suminf_bound[OF _ *]) (auto intro!: increasingD[OF inc] allI countable_UN) qed { fix a @@ -373,15 +386,15 @@ have "a \ (\i. A i) \ sets M" by (metis Int U_in a) ultimately - have "f (a \ (\i. A i)) \ psuminf (f \ (\i. a \ i) \ A)" - using countably_subadditiveD [OF csa, of "(\i. a \ A i)"] + have "f (a \ (\i. A i)) \ (\i. f (a \ A i))" + using csa[unfolded countably_subadditive_def, rule_format, of "(\i. a \ A i)"] by (simp add: o_def) hence "f (a \ (\i. A i)) + f (a - (\i. A i)) \ - psuminf (f \ (\i. a \ i) \ A) + f (a - (\i. A i))" + (\i. f (a \ A i)) + f (a - (\i. A i))" by (rule add_right_mono) moreover - have "psuminf (f \ (\i. a \ i) \ A) + f (a - (\i. A i)) \ f a" - proof (safe intro!: psuminf_bound_add) + have "(\i. f (a \ A i)) + f (a - (\i. A i)) \ f a" + proof (intro suminf_bound_add allI) fix n have UNION_in: "(\i\{0.. sets M" by (metis A'' UNION_in_sets) @@ -395,8 +408,14 @@ have "f (a - (\i. A i)) \ f (a - (\i\{0.. (\i. a \ i) \ A) {..i. A i)) \ f a" + thus "(\i A i)) + f (a - (\i. A i)) \ f a" by (simp add: lambda_system_strong_sum pos A disj eq_fa add_left_mono atLeast0LessThan[symmetric]) + next + have "\i. a \ A i \ sets M" using A'' by auto + then show "\i. 0 \ f (a \ A i)" using pos[unfolded positive_def] by auto + have "\i. a - (\i. A i) \ sets M" using A'' by auto + then have "\i. 0 \ f (a - (\i. A i))" using pos[unfolded positive_def] by auto + then show "f (a - (\i. A i)) \ -\" by auto qed ultimately show "f (a \ (\i. A i)) + f (a - (\i. A i)) \ f a" by (rule order_trans) @@ -443,12 +462,14 @@ proof (auto simp add: increasing_def) fix x y assume xy: "x \ sets M" "y \ sets M" "x \ y" - have "f x \ f x + f (y-x)" .. + then have "y - x \ sets M" by auto + then have "0 \ f (y-x)" using posf[unfolded positive_def] by auto + then have "f x + 0 \ f x + f (y-x)" by (intro add_left_mono) auto also have "... = f (x \ (y-x))" using addf by (auto simp add: additive_def) (metis Diff_disjoint Un_Diff_cancel Diff xy(1,2)) also have "... = f y" by (metis Un_Diff_cancel Un_absorb1 xy(3)) - finally show "f x \ f y" . + finally show "f x \ f y" by simp qed lemma (in algebra) countably_additive_additive: @@ -461,27 +482,27 @@ by (auto simp add: disjoint_family_on_def binaryset_def) hence "range (binaryset x y) \ sets M \ (\i. binaryset x y i) \ sets M \ - f (\i. binaryset x y i) = (\\<^isub>\ n. f (binaryset x y n))" + f (\i. binaryset x y i) = (\ n. f (binaryset x y n))" using ca by (simp add: countably_additive_def) hence "{x,y,{}} \ sets M \ x \ y \ sets M \ - f (x \ y) = (\\<^isub>\ n. f (binaryset x y n))" + f (x \ y) = (\n. f (binaryset x y n))" by (simp add: range_binaryset_eq UN_binaryset_eq) thus "f (x \ y) = f x + f y" using posf x y - by (auto simp add: Un binaryset_psuminf positive_def) + by (auto simp add: Un suminf_binaryset_eq positive_def) qed lemma inf_measure_nonempty: assumes f: "positive M f" and b: "b \ sets M" and a: "a \ b" "{} \ sets M" shows "f b \ measure_set M f a" proof - - have "psuminf (f \ (\i. {})(0 := b)) = setsum (f \ (\i. {})(0 := b)) {..<1::nat}" - by (rule psuminf_finite) (simp add: f[unfolded positive_def]) + let ?A = "\i::nat. (if i = 0 then b else {})" + have "(\i. f (?A i)) = (\i<1::nat. f (?A i))" + by (rule suminf_finite) (simp add: f[unfolded positive_def]) also have "... = f b" by simp - finally have "psuminf (f \ (\i. {})(0 := b)) = f b" . - thus ?thesis using assms - by (auto intro!: exI [of _ "(\i. {})(0 := b)"] + finally show ?thesis using assms + by (auto intro!: exI [of _ ?A] simp: measure_set_def disjoint_family_on_def split_if_mem2 comp_def) qed @@ -489,19 +510,19 @@ assumes posf: "positive M f" and ca: "countably_additive M f" and s: "s \ sets M" shows "Inf (measure_set M f s) = f s" - unfolding Inf_pextreal_def + unfolding Inf_extreal_def proof (safe intro!: Greatest_equality) fix z assume z: "z \ measure_set M f s" from this obtain A where A: "range A \ sets M" and disj: "disjoint_family A" - and "s \ (\x. A x)" and si: "psuminf (f \ A) = z" + and "s \ (\x. A x)" and si: "(\i. f (A i)) = z" by (auto simp add: measure_set_def comp_def) hence seq: "s = (\i. A i \ s)" by blast have inc: "increasing M f" by (metis additive_increasing ca countably_additive_additive posf) - have sums: "psuminf (\i. f (A i \ s)) = f (\i. A i \ s)" - proof (rule countably_additiveD [OF ca]) + have sums: "(\i. f (A i \ s)) = f (\i. A i \ s)" + proof (rule ca[unfolded countably_additive_def, rule_format]) show "range (\n. A n \ s) \ sets M" using A s by blast show "disjoint_family (\n. A n \ s)" using disj @@ -509,12 +530,14 @@ show "(\i. A i \ s) \ sets M" using A s by (metis UN_extend_simps(4) s seq) qed - hence "f s = psuminf (\i. f (A i \ s))" + hence "f s = (\i. f (A i \ s))" using seq [symmetric] by (simp add: sums_iff) - also have "... \ psuminf (f \ A)" - proof (rule psuminf_le) - fix n show "f (A n \ s) \ (f \ A) n" using A s + also have "... \ (\i. f (A i))" + proof (rule suminf_le_pos) + fix n show "f (A n \ s) \ f (A n)" using A s by (force intro: increasingD [OF inc]) + fix N have "A N \ s \ sets M" using A s by auto + then show "0 \ f (A N \ s)" using posf unfolding positive_def by auto qed also have "... = z" by (rule si) finally show "f s \ z" . @@ -525,18 +548,40 @@ by (blast intro: inf_measure_nonempty [of _ f, OF posf s subset_refl]) qed +lemma measure_set_pos: + assumes posf: "positive M f" "r \ measure_set M f X" + shows "0 \ r" +proof - + obtain A where "range A \ sets M" and r: "r = (\i. f (A i))" + using `r \ measure_set M f X` unfolding measure_set_def by auto + then show "0 \ r" using posf unfolding r positive_def + by (intro suminf_0_le) auto +qed + +lemma inf_measure_pos: + assumes posf: "positive M f" + shows "0 \ Inf (measure_set M f X)" +proof (rule complete_lattice_class.Inf_greatest) + fix r assume "r \ measure_set M f X" with posf show "0 \ r" + by (rule measure_set_pos) +qed + lemma inf_measure_empty: - assumes posf: "positive M f" "{} \ sets M" + assumes posf: "positive M f" and "{} \ sets M" shows "Inf (measure_set M f {}) = 0" proof (rule antisym) show "Inf (measure_set M f {}) \ 0" by (metis complete_lattice_class.Inf_lower `{} \ sets M` inf_measure_nonempty[OF posf] subset_refl posf[unfolded positive_def]) -qed simp +qed (rule inf_measure_pos[OF posf]) lemma (in algebra) inf_measure_positive: - "positive M f \ positive M (\x. Inf (measure_set M f x))" - by (simp add: positive_def inf_measure_empty) + assumes p: "positive M f" and "{} \ sets M" + shows "positive M (\x. Inf (measure_set M f x))" +proof (unfold positive_def, intro conjI ballI) + show "Inf (measure_set M f {}) = 0" using inf_measure_empty[OF assms] by auto + fix A assume "A \ sets M" +qed (rule inf_measure_pos[OF p]) lemma (in algebra) inf_measure_increasing: assumes posf: "positive M f" @@ -548,25 +593,25 @@ apply (clarsimp simp add: measure_set_def, rule_tac x=A in exI, blast) done - lemma (in algebra) inf_measure_le: assumes posf: "positive M f" and inc: "increasing M f" - and x: "x \ {r . \A. range A \ sets M \ s \ (\i. A i) \ psuminf (f \ A) = r}" + and x: "x \ {r . \A. range A \ sets M \ s \ (\i. A i) \ (\i. f (A i)) = r}" shows "Inf (measure_set M f s) \ x" proof - - from x obtain A where A: "range A \ sets M" and ss: "s \ (\i. A i)" - and xeq: "psuminf (f \ A) = x" - by auto + and xeq: "(\i. f (A i)) = x" + using x by auto have dA: "range (disjointed A) \ sets M" by (metis A range_disjointed_sets) - have "\n.(f o disjointed A) n \ (f \ A) n" unfolding comp_def + have "\n. f (disjointed A n) \ f (A n)" by (metis increasingD [OF inc] UNIV_I dA image_subset_iff disjointed_subset A comp_def) - hence sda: "psuminf (f o disjointed A) \ psuminf (f \ A)" - by (blast intro: psuminf_le) - hence ley: "psuminf (f o disjointed A) \ x" + moreover have "\i. 0 \ f (disjointed A i)" + using posf dA unfolding positive_def by auto + ultimately have sda: "(\i. f (disjointed A i)) \ (\i. f (A i))" + by (blast intro!: suminf_le_pos) + hence ley: "(\i. f (disjointed A i)) \ x" by (metis xeq) - hence y: "psuminf (f o disjointed A) \ measure_set M f s" + hence y: "(\i. f (disjointed A i)) \ measure_set M f s" apply (auto simp add: measure_set_def) apply (rule_tac x="disjointed A" in exI) apply (simp add: disjoint_family_disjointed UN_disjointed_eq ss dA comp_def) @@ -576,13 +621,16 @@ qed lemma (in algebra) inf_measure_close: + fixes e :: extreal assumes posf: "positive M f" and e: "0 < e" and ss: "s \ (space M)" shows "\A. range A \ sets M \ disjoint_family A \ s \ (\i. A i) \ - psuminf (f \ A) \ Inf (measure_set M f s) + e" -proof (cases "Inf (measure_set M f s) = \") + (\i. f (A i)) \ Inf (measure_set M f s) + e" +proof (cases "Inf (measure_set M f s) = \") case False + then have fin: "\Inf (measure_set M f s)\ \ \" + using inf_measure_pos[OF posf, of s] by auto obtain l where "l \ measure_set M f s" "l \ Inf (measure_set M f s) + e" - using Inf_close[OF False e] by auto + using Inf_extreal_close[OF fin e] by auto thus ?thesis by (auto intro!: exI[of _ l] simp: measure_set_def comp_def) next @@ -600,9 +648,8 @@ shows "countably_subadditive (| space = space M, sets = Pow (space M) |) (\x. Inf (measure_set M f x))" unfolding countably_subadditive_def o_def -proof (safe, simp, rule pextreal_le_epsilon) - fix A :: "nat \ 'a set" and e :: pextreal - +proof (safe, simp, rule extreal_le_epsilon, safe) + fix A :: "nat \ 'a set" and e :: extreal let "?outer n" = "Inf (measure_set M f (A n))" assume A: "range A \ Pow (space M)" and disj: "disjoint_family A" @@ -610,21 +657,27 @@ and e: "0 < e" hence "\BB. \n. range (BB n) \ sets M \ disjoint_family (BB n) \ A n \ (\i. BB n i) \ - psuminf (f o BB n) \ ?outer n + e * (1/2)^(Suc n)" - apply (safe intro!: choice inf_measure_close [of f, OF posf _]) - using e sb by (cases e, auto simp add: not_le mult_pos_pos) + (\i. f (BB n i)) \ ?outer n + e * (1/2)^(Suc n)" + apply (safe intro!: choice inf_measure_close [of f, OF posf]) + using e sb by (cases e) (auto simp add: not_le mult_pos_pos one_extreal_def) then obtain BB where BB: "\n. (range (BB n) \ sets M)" and disjBB: "\n. disjoint_family (BB n)" and sbBB: "\n. A n \ (\i. BB n i)" - and BBle: "\n. psuminf (f o BB n) \ ?outer n + e * (1/2)^(Suc n)" + and BBle: "\n. (\i. f (BB n i)) \ ?outer n + e * (1/2)^(Suc n)" by auto blast - have sll: "(\\<^isub>\ n. psuminf (f o BB n)) \ psuminf ?outer + e" + have sll: "(\n. \i. (f (BB n i))) \ suminf ?outer + e" proof - - have "(\\<^isub>\ n. psuminf (f o BB n)) \ (\\<^isub>\ n. ?outer n + e*(1/2) ^ Suc n)" - by (rule psuminf_le[OF BBle]) - also have "... = psuminf ?outer + e" - using psuminf_half_series by simp + have sum_eq_1: "(\n. e*(1/2) ^ Suc n) = e" + using suminf_half_series_extreal e + by (simp add: extreal_zero_le_0_iff zero_le_divide_extreal suminf_cmult_extreal) + have "\n i. 0 \ f (BB n i)" using posf[unfolded positive_def] BB by auto + then have "\n. 0 \ (\i. f (BB n i))" by (rule suminf_0_le) + then have "(\n. \i. (f (BB n i))) \ (\n. ?outer n + e*(1/2) ^ Suc n)" + by (rule suminf_le_pos[OF BBle]) + also have "... = suminf ?outer + e" + using sum_eq_1 inf_measure_pos[OF posf] e + by (subst suminf_add_extreal) (auto simp add: extreal_zero_le_0_iff) finally show ?thesis . qed def C \ "(split BB) o prod_decode" @@ -644,23 +697,25 @@ qed have "(f \ C) = (f \ (\(x, y). BB x y)) \ prod_decode" by (rule ext) (auto simp add: C_def) - moreover have "psuminf ... = (\\<^isub>\ n. psuminf (f o BB n))" using BBle - by (force intro!: psuminf_2dimen simp: o_def) - ultimately have Csums: "psuminf (f \ C) = (\\<^isub>\ n. psuminf (f o BB n))" by simp - have "Inf (measure_set M f (\i. A i)) \ (\\<^isub>\ n. psuminf (f o BB n))" + moreover have "suminf ... = (\n. \i. f (BB n i))" using BBle + using BB posf[unfolded positive_def] + by (force intro!: suminf_extreal_2dimen simp: o_def) + ultimately have Csums: "(\i. f (C i)) = (\n. \i. f (BB n i))" by (simp add: o_def) + have "Inf (measure_set M f (\i. A i)) \ (\n. \i. f (BB n i))" apply (rule inf_measure_le [OF posf(1) inc], auto) apply (rule_tac x="C" in exI) apply (auto simp add: C sbC Csums) done - also have "... \ (\\<^isub>\n. Inf (measure_set M f (A n))) + e" using sll + also have "... \ (\n. Inf (measure_set M f (A n))) + e" using sll by blast - finally show "Inf (measure_set M f (\i. A i)) \ psuminf ?outer + e" . + finally show "Inf (measure_set M f (\i. A i)) \ suminf ?outer + e" . qed lemma (in algebra) inf_measure_outer: "\ positive M f ; increasing M f \ \ outer_measure_space \ space = space M, sets = Pow (space M) \ (\x. Inf (measure_set M f x))" + using inf_measure_pos[of M f] by (simp add: outer_measure_space_def inf_measure_empty inf_measure_increasing inf_measure_countably_subadditive positive_def) @@ -680,13 +735,13 @@ by blast have "Inf (measure_set M f (s\x)) + Inf (measure_set M f (s-x)) \ Inf (measure_set M f s)" - proof (rule pextreal_le_epsilon) - fix e :: pextreal + proof (rule extreal_le_epsilon, intro allI impI) + fix e :: extreal assume e: "0 < e" from inf_measure_close [of f, OF posf e s] obtain A where A: "range A \ sets M" and disj: "disjoint_family A" and sUN: "s \ (\i. A i)" - and l: "psuminf (f \ A) \ Inf (measure_set M f s) + e" + and l: "(\i. f (A i)) \ Inf (measure_set M f s) + e" by auto have [simp]: "!!x. x \ sets M \ (f o (\z. z \ (space M - x)) o A) = (f o (\z. z - x) o A)" @@ -698,9 +753,9 @@ assume u: "u \ sets M" have [simp]: "\n. f (A n \ u) \ f (A n)" by (simp add: increasingD [OF inc] u Int range_subsetD [OF A]) - have 2: "Inf (measure_set M f (s \ u)) \ psuminf (f \ (\z. z \ u) \ A)" + have 2: "Inf (measure_set M f (s \ u)) \ (\i. f (A i \ u))" proof (rule complete_lattice_class.Inf_lower) - show "psuminf (f \ (\z. z \ u) \ A) \ measure_set M f (s \ u)" + show "(\i. f (A i \ u)) \ measure_set M f (s \ u)" apply (simp add: measure_set_def) apply (rule_tac x="(\z. z \ u) o A" in exI) apply (auto simp add: disjoint_family_subset [OF disj] o_def) @@ -709,15 +764,16 @@ done qed } note lesum = this - have inf1: "Inf (measure_set M f (s\x)) \ psuminf (f o (\z. z\x) o A)" + have [simp]: "\i. A i \ (space M - x) = A i - x" using A sets_into_space by auto + have inf1: "Inf (measure_set M f (s\x)) \ (\i. f (A i \ x))" and inf2: "Inf (measure_set M f (s \ (space M - x))) - \ psuminf (f o (\z. z \ (space M - x)) o A)" + \ (\i. f (A i \ (space M - x)))" by (metis Diff lesum top x)+ hence "Inf (measure_set M f (s\x)) + Inf (measure_set M f (s-x)) - \ psuminf (f o (\s. s\x) o A) + psuminf (f o (\s. s-x) o A)" - by (simp add: x add_mono) - also have "... \ psuminf (f o A)" - by (simp add: x psuminf_add[symmetric] o_def) + \ (\i. f (A i \ x)) + (\i. f (A i - x))" + by (simp add: add_mono x) + also have "... \ (\i. f (A i))" using posf[unfolded positive_def] A x + by (subst suminf_add_extreal[symmetric]) auto also have "... \ Inf (measure_set M f s) + e" by (rule l) finally show "Inf (measure_set M f (s\x)) + Inf (measure_set M f (s-x)) @@ -732,7 +788,7 @@ also have "... \ Inf (measure_set M f (s\x)) + Inf (measure_set M f (s-x))" apply (rule subadditiveD) apply (rule algebra.countably_subadditive_subadditive[OF algebra_Pow]) - apply (simp add: positive_def inf_measure_empty[OF posf]) + apply (simp add: positive_def inf_measure_empty[OF posf] inf_measure_pos[OF posf]) apply (rule inf_measure_countably_subadditive) using s by (auto intro!: posf inc) finally show ?thesis . @@ -751,7 +807,7 @@ theorem (in algebra) caratheodory: assumes posf: "positive M f" and ca: "countably_additive M f" - shows "\\ :: 'a set \ pextreal. (\s \ sets M. \ s = f s) \ + shows "\\ :: 'a set \ extreal. (\s \ sets M. \ s = f s) \ measure_space \ space = space M, sets = sets (sigma M), measure = \ \" proof - have inc: "increasing M f" diff -r 28b51effc5ed -r cdf7693bbe08 src/HOL/Probability/Complete_Measure.thy --- a/src/HOL/Probability/Complete_Measure.thy Mon Mar 14 14:37:47 2011 +0100 +++ b/src/HOL/Probability/Complete_Measure.thy Mon Mar 14 14:37:49 2011 +0100 @@ -1,7 +1,6 @@ -(* Title: HOL/Probability/Complete_Measure.thy +(* Title: Complete_Measure.thy Author: Robert Himmelmann, Johannes Hoelzl, TU Muenchen *) - theory Complete_Measure imports Product_Measure begin @@ -177,7 +176,8 @@ proof - show "measure_space completion" proof - show "measure completion {} = 0" by (auto simp: completion_def) + show "positive completion (measure completion)" + by (auto simp: completion_def positive_def) next show "countably_additive completion (measure completion)" proof (intro countably_additiveI) @@ -189,9 +189,9 @@ using A by (subst (1 2) main_part_null_part_Un) auto then show "main_part (A n) \ main_part (A m) = {}" by auto qed - then have "(\\<^isub>\n. measure completion (A n)) = \ (\i. main_part (A i))" + then have "(\n. measure completion (A n)) = \ (\i. main_part (A i))" unfolding completion_def using A by (auto intro!: measure_countably_additive) - then show "(\\<^isub>\n. measure completion (A n)) = measure completion (UNION UNIV A)" + then show "(\n. measure completion (A n)) = measure completion (UNION UNIV A)" by (simp add: completion_def \_main_part_UN[OF A(1)]) qed qed @@ -251,30 +251,52 @@ qed qed -lemma (in completeable_measure_space) completion_ex_borel_measurable: - fixes g :: "'a \ pextreal" - assumes g: "g \ borel_measurable completion" +lemma (in completeable_measure_space) completion_ex_borel_measurable_pos: + fixes g :: "'a \ extreal" + assumes g: "g \ borel_measurable completion" and "\x. 0 \ g x" shows "\g'\borel_measurable M. (AE x. g x = g' x)" proof - - from g[THEN completion.borel_measurable_implies_simple_function_sequence] - obtain f where "\i. simple_function completion (f i)" "f \ g" by auto - then have "\i. \f'. simple_function M f' \ (AE x. f i x = f' x)" - using completion_ex_simple_function by auto + from g[THEN completion.borel_measurable_implies_simple_function_sequence'] guess f . note f = this + from this(1)[THEN completion_ex_simple_function] + have "\i. \f'. simple_function M f' \ (AE x. f i x = f' x)" .. from this[THEN choice] obtain f' where sf: "\i. simple_function M (f' i)" and AE: "\i. AE x. f i x = f' i x" by auto show ?thesis proof (intro bexI) - from AE[unfolded all_AE_countable] + from AE[unfolded AE_all_countable[symmetric]] show "AE x. g x = (SUP i. f' i x)" (is "AE x. g x = ?f x") proof (elim AE_mp, safe intro!: AE_I2) fix x assume eq: "\i. f i x = f' i x" - moreover have "g = SUPR UNIV f" using `f \ g` unfolding isoton_def by simp - ultimately show "g x = ?f x" by (simp add: SUPR_apply) + moreover have "g x = (SUP i. f i x)" + unfolding f using `0 \ g x` by (auto split: split_max) + ultimately show "g x = ?f x" by auto qed show "?f \ borel_measurable M" using sf by (auto intro: borel_measurable_simple_function) qed qed +lemma (in completeable_measure_space) completion_ex_borel_measurable: + fixes g :: "'a \ extreal" + assumes g: "g \ borel_measurable completion" + shows "\g'\borel_measurable M. (AE x. g x = g' x)" +proof - + have "(\x. max 0 (g x)) \ borel_measurable completion" "\x. 0 \ max 0 (g x)" using g by auto + from completion_ex_borel_measurable_pos[OF this] guess g_pos .. + moreover + have "(\x. max 0 (- g x)) \ borel_measurable completion" "\x. 0 \ max 0 (- g x)" using g by auto + from completion_ex_borel_measurable_pos[OF this] guess g_neg .. + ultimately + show ?thesis + proof (safe intro!: bexI[of _ "\x. g_pos x - g_neg x"]) + show "AE x. max 0 (- g x) = g_neg x \ max 0 (g x) = g_pos x \ g x = g_pos x - g_neg x" + proof (intro AE_I2 impI) + fix x assume g: "max 0 (- g x) = g_neg x" "max 0 (g x) = g_pos x" + show "g x = g_pos x - g_neg x" unfolding g[symmetric] + by (cases "g x") (auto split: split_max) + qed + qed auto +qed + end diff -r 28b51effc5ed -r cdf7693bbe08 src/HOL/Probability/Information.thy --- a/src/HOL/Probability/Information.thy Mon Mar 14 14:37:47 2011 +0100 +++ b/src/HOL/Probability/Information.thy Mon Mar 14 14:37:49 2011 +0100 @@ -2,9 +2,12 @@ imports Probability_Space "~~/src/HOL/Library/Convex" - Lebesgue_Measure 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 @@ -238,7 +241,7 @@ have ms: "measure_space (M\measure := \\)" by default show "(\x \ space M. log b (real (RN_deriv M \ x)) * real (\ {x})) = ?sum" using RN_deriv_finite_measure[OF ms ac] - by (auto intro!: setsum_cong simp: field_simps real_of_pextreal_mult[symmetric]) + by (auto intro!: setsum_cong simp: field_simps) qed lemma (in finite_prob_space) KL_divergence_positive_finite: @@ -254,7 +257,8 @@ 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 + 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 @@ -262,17 +266,19 @@ 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 < prob {x}" using finite_measure[of "{x}"] x by auto } - qed auto - thus "0 \ KL_divergence b M \" using finite_sum_over_space_eq_1 by simp + 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) qed subsection {* Mutual Information *} definition (in prob_space) "mutual_information b S T X Y = - KL_divergence b (S\measure := distribution X\ \\<^isub>M T\measure := distribution Y\) - (joint_distribution X Y)" + KL_divergence b (S\measure := extreal\distribution X\ \\<^isub>M T\measure := extreal\distribution Y\) + (extreal\joint_distribution X Y)" definition (in prob_space) "entropy b s X = mutual_information b s s X X" @@ -280,38 +286,33 @@ abbreviation (in information_space) mutual_information_Pow ("\'(_ ; _')") where "\(X ; Y) \ mutual_information b - \ space = X`space M, sets = Pow (X`space M), measure = distribution X \ - \ space = Y`space M, sets = Pow (Y`space M), measure = distribution Y \ X Y" + \ space = X`space M, sets = Pow (X`space M), measure = extreal\distribution X \ + \ space = Y`space M, sets = Pow (Y`space M), measure = extreal\distribution Y \ X Y" lemma (in prob_space) finite_variables_absolutely_continuous: assumes X: "finite_random_variable S X" and Y: "finite_random_variable T Y" shows "measure_space.absolutely_continuous - (S\measure := distribution X\ \\<^isub>M T\measure := distribution Y\) - (joint_distribution X Y)" + (S\measure := extreal\distribution X\ \\<^isub>M T\measure := extreal\distribution Y\) + (extreal\joint_distribution X Y)" proof - - interpret X: finite_prob_space "S\measure := distribution X\" + interpret X: finite_prob_space "S\measure := extreal\distribution X\" using X by (rule distribution_finite_prob_space) - interpret Y: finite_prob_space "T\measure := distribution Y\" + interpret Y: finite_prob_space "T\measure := extreal\distribution Y\" using Y by (rule distribution_finite_prob_space) interpret XY: pair_finite_prob_space - "S\measure := distribution X\" "T\ measure := distribution Y\" by default - interpret P: finite_prob_space "XY.P\ measure := joint_distribution X Y\" + "S\measure := extreal\distribution X\" "T\ measure := extreal\distribution Y\" by default + interpret P: finite_prob_space "XY.P\ measure := extreal\joint_distribution X Y\" using assms by (auto intro!: joint_distribution_finite_prob_space) note rv = assms[THEN finite_random_variableD] - show "XY.absolutely_continuous (joint_distribution X Y)" + show "XY.absolutely_continuous (extreal\joint_distribution X Y)" proof (rule XY.absolutely_continuousI) - show "finite_measure_space (XY.P\ measure := joint_distribution X Y\)" by default + show "finite_measure_space (XY.P\ measure := extreal\joint_distribution X Y\)" by default fix x assume "x \ space XY.P" and "XY.\ {x} = 0" - then obtain a b where "(a, b) = x" and "a \ space S" "b \ space T" - and distr: "distribution X {a} * distribution Y {b} = 0" + then obtain a b where "x = (a, b)" + and "distribution X {a} = 0 \ distribution Y {b} = 0" by (cases x) (auto simp: space_pair_measure) - with X.sets_eq_Pow Y.sets_eq_Pow - joint_distribution_Times_le_fst[OF rv, of "{a}" "{b}"] - joint_distribution_Times_le_snd[OF rv, of "{a}" "{b}"] - have "joint_distribution X Y {x} \ distribution Y {b}" - "joint_distribution X Y {x} \ distribution X {a}" - by (auto simp del: X.sets_eq_Pow Y.sets_eq_Pow) - with distr show "joint_distribution X Y {x} = 0" by auto + with finite_distribution_order(5,6)[OF X Y] + show "(extreal \ joint_distribution X Y) {x} = 0" by auto qed qed @@ -320,28 +321,28 @@ assumes MY: "finite_random_variable MY Y" shows mutual_information_generic_eq: "mutual_information b MX MY X Y = (\ (x,y) \ space MX \ space MY. - real (joint_distribution X Y {(x,y)}) * - log b (real (joint_distribution X Y {(x,y)}) / - (real (distribution X {x}) * real (distribution Y {y}))))" + joint_distribution X Y {(x,y)} * + log b (joint_distribution X Y {(x,y)} / + (distribution X {x} * distribution Y {y})))" (is ?sum) and mutual_information_positive_generic: "0 \ mutual_information b MX MY X Y" (is ?positive) proof - - interpret X: finite_prob_space "MX\measure := distribution X\" + interpret X: finite_prob_space "MX\measure := extreal\distribution X\" using MX by (rule distribution_finite_prob_space) - interpret Y: finite_prob_space "MY\measure := distribution Y\" + interpret Y: finite_prob_space "MY\measure := extreal\distribution Y\" using MY by (rule distribution_finite_prob_space) - interpret XY: pair_finite_prob_space "MX\measure := distribution X\" "MY\measure := distribution Y\" by default - interpret P: finite_prob_space "XY.P\measure := joint_distribution X Y\" + interpret XY: pair_finite_prob_space "MX\measure := extreal\distribution X\" "MY\measure := extreal\distribution Y\" by default + interpret P: finite_prob_space "XY.P\measure := extreal\joint_distribution X Y\" using assms by (auto intro!: joint_distribution_finite_prob_space) - have P_ms: "finite_measure_space (XY.P\measure :=joint_distribution X Y\)" by default - have P_ps: "finite_prob_space (XY.P\measure := joint_distribution X Y\)" by default + have P_ms: "finite_measure_space (XY.P\measure := extreal\joint_distribution X Y\)" by default + have P_ps: "finite_prob_space (XY.P\measure := extreal\joint_distribution X Y\)" by default show ?sum unfolding Let_def mutual_information_def by (subst XY.KL_divergence_eq_finite[OF P_ms finite_variables_absolutely_continuous[OF MX MY]]) - (auto simp add: space_pair_measure setsum_cartesian_product' real_of_pextreal_mult[symmetric]) + (auto simp add: space_pair_measure setsum_cartesian_product') show ?positive using XY.KL_divergence_positive_finite[OF P_ps finite_variables_absolutely_continuous[OF MX MY] b_gt_1] @@ -351,10 +352,10 @@ 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 := distribution X\ \\<^isub>M T\measure := distribution Y\) (joint_distribution X Y)" + (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 := distribution X\" and ?T = "T\measure := distribution Y\" + 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 .. @@ -363,13 +364,13 @@ 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 := joint_distribution X Y\) (Q.P \ measure := joint_distribution Y X\)" + (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 := joint_distribution X Y\)" + 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 := joint_distribution X Y\)" + then show "measure_space (P.P\ measure := extreal\joint_distribution X Y\)" unfolding prob_space_def by simp qed auto qed @@ -389,8 +390,8 @@ lemma (in information_space) mutual_information_eq: assumes "simple_function M X" "simple_function M Y" shows "\(X;Y) = (\ (x,y) \ X ` space M \ Y ` space M. - real (distribution (\x. (X x, Y x)) {(x,y)}) * log b (real (distribution (\x. (X x, Y x)) {(x,y)}) / - (real (distribution X {x}) * real (distribution Y {y}))))" + distribution (\x. (X x, Y x)) {(x,y)} * log b (distribution (\x. (X x, Y x)) {(x,y)} / + (distribution X {x} * distribution Y {y})))" using assms by (simp add: mutual_information_generic_eq) lemma (in information_space) mutual_information_generic_cong: @@ -416,22 +417,27 @@ abbreviation (in information_space) entropy_Pow ("\'(_')") where - "\(X) \ entropy b \ space = X`space M, sets = Pow (X`space M), measure = distribution X \ X" + "\(X) \ entropy b \ space = X`space M, sets = Pow (X`space M), measure = extreal\distribution X \ X" lemma (in information_space) entropy_generic_eq: + fixes X :: "'a \ 'c" assumes MX: "finite_random_variable MX X" - shows "entropy b MX X = -(\ x \ space MX. real (distribution X {x}) * log b (real (distribution X {x})))" + shows "entropy b MX X = -(\ x \ space MX. distribution X {x} * log b (distribution X {x}))" proof - - interpret MX: finite_prob_space "MX\measure := distribution X\" + interpret MX: finite_prob_space "MX\measure := extreal\distribution X\" using MX by (rule distribution_finite_prob_space) - let "?X x" = "real (distribution X {x})" - let "?XX x y" = "real (joint_distribution X X {(x, y)})" - { fix x y - have "(\x. (X x, X x)) -` {(x, y)} = (if x = y then X -` {x} else {})" by auto + let "?X x" = "distribution X {x}" + let "?XX x y" = "joint_distribution X X {(x, y)}" + + { fix x y :: 'c + { assume "x \ y" + then have "(\x. (X x, X x)) -` {(x,y)} \ space M = {}" by auto + then have "joint_distribution X X {(x, y)} = 0" by (simp add: distribution_def) } then have "?XX x y * log b (?XX x y / (?X x * ?X y)) = (if x = y then - ?X y * log b (?X y) else 0)" - unfolding distribution_def by (auto simp: log_simps zero_less_mult_iff) } + by (auto simp: log_simps zero_less_mult_iff) } note remove_XX = this + show ?thesis unfolding entropy_def mutual_information_generic_eq[OF MX MX] unfolding setsum_cartesian_product[symmetric] setsum_negf[symmetric] remove_XX @@ -440,7 +446,7 @@ lemma (in information_space) entropy_eq: assumes "simple_function M X" - shows "\(X) = -(\ x \ X ` space M. real (distribution X {x}) * log b (real (distribution X {x})))" + shows "\(X) = -(\ x \ X ` space M. distribution X {x} * log b (distribution X {x}))" using assms by (simp add: entropy_generic_eq) lemma (in information_space) entropy_positive: @@ -448,63 +454,77 @@ unfolding entropy_def by (simp add: mutual_information_positive) lemma (in information_space) entropy_certainty_eq_0: - assumes "simple_function M X" and "x \ X ` space M" and "distribution X {x} = 1" + assumes X: "simple_function M X" and "x \ X ` space M" and "distribution X {x} = 1" shows "\(X) = 0" proof - - let ?X = "\ space = X ` space M, sets = Pow (X ` space M), measure = distribution X\" + let ?X = "\ space = X ` space M, sets = Pow (X ` space M), measure = extreal\distribution X\" note simple_function_imp_finite_random_variable[OF `simple_function M X`] - from distribution_finite_prob_space[OF this, of "\ measure = distribution X \"] + from distribution_finite_prob_space[OF this, of "\ measure = extreal\distribution X \"] interpret X: finite_prob_space ?X by simp have "distribution X (X ` space M - {x}) = distribution X (X ` space M) - distribution X {x}" using X.measure_compl[of "{x}"] assms by auto also have "\ = 0" using X.prob_space assms by auto finally have X0: "distribution X (X ` space M - {x}) = 0" by auto - { fix y assume asm: "y \ x" "y \ X ` space M" - hence "{y} \ X ` space M - {x}" by auto - from X.measure_mono[OF this] X0 asm - have "distribution X {y} = 0" by auto } - hence fi: "\ y. y \ X ` space M \ real (distribution X {y}) = (if x = y then 1 else 0)" - using assms by auto + { fix y assume *: "y \ X ` space M" + { assume asm: "y \ x" + with * have "{y} \ X ` space M - {x}" by auto + from X.measure_mono[OF this] X0 asm * + have "distribution X {y} = 0" by (auto intro: antisym) } + then have "distribution X {y} = (if x = y then 1 else 0)" + using assms by auto } + note fi = this have y: "\y. (if x = y then 1 else 0) * log b (if x = y then 1 else 0) = 0" by simp show ?thesis unfolding entropy_eq[OF `simple_function M X`] by (auto simp: y fi) qed lemma (in information_space) entropy_le_card_not_0: - assumes "simple_function M X" - shows "\(X) \ log b (real (card (X ` space M \ {x . distribution X {x} \ 0})))" + assumes X: "simple_function M X" + shows "\(X) \ log b (card (X ` space M \ {x. distribution X {x} \ 0}))" proof - - let "?d x" = "distribution X {x}" - let "?p x" = "real (?d x)" + let "?p x" = "distribution X {x}" have "\(X) = (\x\X`space M. ?p x * log b (1 / ?p x))" - by (auto intro!: setsum_cong simp: entropy_eq[OF `simple_function M X`] setsum_negf[symmetric] log_simps not_less) + unfolding entropy_eq[OF X] setsum_negf[symmetric] + by (auto intro!: setsum_cong simp: log_simps) also have "\ \ log b (\x\X`space M. ?p x * (1 / ?p x))" - apply (rule log_setsum') - using not_empty b_gt_1 `simple_function M X` sum_over_space_real_distribution - by (auto simp: simple_function_def) - also have "\ = log b (\x\X`space M. if ?d x \ 0 then 1 else 0)" - using distribution_finite[OF `simple_function M X`[THEN simple_function_imp_random_variable], simplified] - by (intro arg_cong[where f="\X. log b X"] setsum_cong) (auto simp: real_of_pextreal_eq_0) + using not_empty b_gt_1 `simple_function M X` sum_over_space_real_distribution[OF X] + by (intro log_setsum') (auto simp: simple_function_def) + also have "\ = log b (\x\X`space M. if ?p x \ 0 then 1 else 0)" + by (intro arg_cong[where f="\X. log b X"] setsum_cong) auto finally show ?thesis using `simple_function M X` by (auto simp: setsum_cases real_eq_of_nat simple_function_def) qed +lemma (in prob_space) measure'_translate: + assumes X: "random_variable S X" and A: "A \ sets S" + shows "finite_measure.\' (S\ measure := extreal\distribution X \) A = distribution X A" +proof - + interpret S: prob_space "S\ measure := extreal\distribution X \" + using distribution_prob_space[OF X] . + from A show "S.\' A = distribution X A" + unfolding S.\'_def by (simp add: distribution_def_raw \'_def) +qed + lemma (in information_space) entropy_uniform_max: - assumes "simple_function M X" + assumes X: "simple_function M X" assumes "\x y. \ x \ X ` space M ; y \ X ` space M \ \ distribution X {x} = distribution X {y}" shows "\(X) = log b (real (card (X ` space M)))" proof - - let ?X = "\ space = X ` space M, sets = Pow (X ` space M), measure = distribution X\" - note simple_function_imp_finite_random_variable[OF `simple_function M X`] - from distribution_finite_prob_space[OF this, of "\ measure = distribution X \"] + let ?X = "\ space = X ` space M, sets = Pow (X ` space M), measure = undefined\\ measure := extreal\distribution X\" + note frv = simple_function_imp_finite_random_variable[OF X] + from distribution_finite_prob_space[OF this, of "\ measure = extreal\distribution X \"] interpret X: finite_prob_space ?X by simp + note rv = finite_random_variableD[OF frv] have card_gt0: "0 < card (X ` space M)" unfolding card_gt_0_iff using `simple_function M X` not_empty by (auto simp: simple_function_def) - { fix x assume "x \ X ` space M" - hence "real (distribution X {x}) = 1 / real (card (X ` space M))" - proof (rule X.uniform_prob[simplified]) - fix x y assume "x \ X`space M" "y \ X`space M" - from assms(2)[OF this] show "real (distribution X {x}) = real (distribution X {y})" by simp - qed } + { fix x assume "x \ space ?X" + moreover then have "X.\' {x} = 1 / card (space ?X)" + proof (rule X.uniform_prob) + fix x y assume "x \ space ?X" "y \ space ?X" + with assms(2)[of x y] show "X.\' {x} = X.\' {y}" + by (subst (1 2) measure'_translate[OF rv]) auto + qed + ultimately have "distribution X {x} = 1 / card (space ?X)" + by (subst (asm) measure'_translate[OF rv]) auto } thus ?thesis using not_empty X.finite_space b_gt_1 card_gt0 by (simp add: entropy_eq[OF `simple_function M X`] real_eq_of_nat[symmetric] log_simps) @@ -552,8 +572,7 @@ lemma (in information_space) entropy_eq_cartesian_product: assumes "simple_function M X" "simple_function M Y" shows "\(\x. (X x, Y x)) = -(\x\X`space M. \y\Y`space M. - real (joint_distribution X Y {(x,y)}) * - log b (real (joint_distribution X Y {(x,y)})))" + joint_distribution X Y {(x,y)} * log b (joint_distribution X Y {(x,y)}))" proof - have sf: "simple_function M (\x. (X x, Y x))" using assms by (auto intro: simple_function_Pair) @@ -576,9 +595,9 @@ abbreviation (in information_space) conditional_mutual_information_Pow ("\'( _ ; _ | _ ')") where "\(X ; Y | Z) \ conditional_mutual_information b - \ space = X`space M, sets = Pow (X`space M), measure = distribution X \ - \ space = Y`space M, sets = Pow (Y`space M), measure = distribution Y \ - \ space = Z`space M, sets = Pow (Z`space M), measure = distribution Z \ + \ space = X`space M, sets = Pow (X`space M), measure = extreal\distribution X \ + \ space = Y`space M, sets = Pow (Y`space M), measure = extreal\distribution Y \ + \ space = Z`space M, sets = Pow (Z`space M), measure = extreal\distribution Z \ X Y Z" lemma (in information_space) conditional_mutual_information_generic_eq: @@ -586,58 +605,44 @@ and MY: "finite_random_variable MY Y" and MZ: "finite_random_variable MZ Z" shows "conditional_mutual_information b MX MY MZ X Y Z = (\(x, y, z) \ space MX \ space MY \ space MZ. - real (distribution (\x. (X x, Y x, Z x)) {(x, y, z)}) * - log b (real (distribution (\x. (X x, Y x, Z x)) {(x, y, z)}) / - (real (joint_distribution X Z {(x, z)}) * real (joint_distribution Y Z {(y,z)} / distribution Z {z}))))" - (is "_ = (\(x, y, z)\?S. ?XYZ x y z * log b (?XYZ x y z / (?XZ x z * ?YZdZ y z)))") + distribution (\x. (X x, Y x, Z x)) {(x, y, z)} * + log b (distribution (\x. (X x, Y x, Z x)) {(x, y, z)} / + (joint_distribution X Z {(x, z)} * (joint_distribution Y Z {(y,z)} / distribution Z {z}))))" + (is "_ = (\(x, y, z)\?S. ?XYZ x y z * log b (?XYZ x y z / (?XZ x z * (?YZ y z / ?Z z))))") proof - - let ?YZ = "\y z. real (joint_distribution Y Z {(y, z)})" - let ?X = "\x. real (distribution X {x})" - let ?Z = "\z. real (distribution Z {z})" - - txt {* This proof is actually quiet easy, however we need to show that the - distributions are finite and the joint distributions are zero when one of - the variables distribution is also zero. *} - + let ?X = "\x. distribution X {x}" note finite_var = MX MY MZ - note random_var = finite_var[THEN finite_random_variableD] - - note space_simps = space_pair_measure space_sigma algebra.simps - note YZ = finite_random_variable_pairI[OF finite_var(2,3)] + note XYZ = finite_random_variable_pairI[OF MX YZ] note XZ = finite_random_variable_pairI[OF finite_var(1,3)] note ZX = finite_random_variable_pairI[OF finite_var(3,1)] note YZX = finite_random_variable_pairI[OF finite_var(2) ZX] note order1 = - finite_distribution_order(5,6)[OF finite_var(1) YZ, simplified space_simps] - finite_distribution_order(5,6)[OF finite_var(1,3), simplified space_simps] + finite_distribution_order(5,6)[OF finite_var(1) YZ] + finite_distribution_order(5,6)[OF finite_var(1,3)] + note random_var = finite_var[THEN finite_random_variableD] note finite = finite_var(1) YZ finite_var(3) XZ YZX - note finite[THEN finite_distribution_finite, simplified space_simps, simp] have order2: "\x y z. \x \ space MX; y \ space MY; z \ space MZ; joint_distribution X Z {(x, z)} = 0\ \ joint_distribution X (\x. (Y x, Z x)) {(x, y, z)} = 0" unfolding joint_distribution_commute_singleton[of X] unfolding joint_distribution_assoc_singleton[symmetric] using finite_distribution_order(6)[OF finite_var(2) ZX] - by (auto simp: space_simps) + by auto - have "(\(x, y, z)\?S. ?XYZ x y z * log b (?XYZ x y z / (?XZ x z * ?YZdZ y z))) = + have "(\(x, y, z)\?S. ?XYZ x y z * log b (?XYZ x y z / (?XZ x z * (?YZ y z / ?Z z)))) = (\(x, y, z)\?S. ?XYZ x y z * (log b (?XYZ x y z / (?X x * ?YZ y z)) - log b (?XZ x z / (?X x * ?Z z))))" (is "(\(x, y, z)\?S. ?L x y z) = (\(x, y, z)\?S. ?R x y z)") proof (safe intro!: setsum_cong) fix x y z assume space: "x \ space MX" "y \ space MY" "z \ space MZ" - then have *: "?XYZ x y z / (?XZ x z * ?YZdZ y z) = - (?XYZ x y z / (?X x * ?YZ y z)) / (?XZ x z / (?X x * ?Z z))" - using order1(3) - by (auto simp: real_of_pextreal_mult[symmetric] real_of_pextreal_eq_0) show "?L x y z = ?R x y z" proof cases assume "?XYZ x y z \ 0" - with space b_gt_1 order1 order2 show ?thesis unfolding * - by (subst log_divide) - (auto simp: zero_less_divide_iff zero_less_real_of_pextreal - real_of_pextreal_eq_0 zero_less_mult_iff) + with space have "0 < ?X x" "0 < ?Z z" "0 < ?XZ x z" "0 < ?YZ y z" "0 < ?XYZ x y z" + using order1 order2 by (auto simp: less_le) + with b_gt_1 show ?thesis + by (simp add: log_mult log_divide zero_less_mult_iff zero_less_divide_iff) qed simp qed also have "\ = (\(x, y, z)\?S. ?XYZ x y z * log b (?XYZ x y z / (?X x * ?YZ y z))) - @@ -649,8 +654,8 @@ setsum_left_distrib[symmetric] unfolding joint_distribution_commute_singleton[of X] unfolding joint_distribution_assoc_singleton[symmetric] - using setsum_real_joint_distribution_singleton[OF finite_var(2) ZX, unfolded space_simps] - by (intro setsum_cong refl) simp + using setsum_joint_distribution_singleton[OF finite_var(2) ZX] + by (intro setsum_cong refl) (simp add: space_pair_measure) also have "(\(x, y, z)\?S. ?XYZ x y z * log b (?XYZ x y z / (?X x * ?YZ y z))) - (\(x, z)\space MX \ space MZ. ?XZ x z * log b (?XZ x z / (?X x * ?Z z))) = conditional_mutual_information b MX MY MZ X Y Z" @@ -664,11 +669,11 @@ lemma (in information_space) conditional_mutual_information_eq: assumes "simple_function M X" "simple_function M Y" "simple_function M Z" shows "\(X;Y|Z) = (\(x, y, z) \ X`space M \ Y`space M \ Z`space M. - real (distribution (\x. (X x, Y x, Z x)) {(x, y, z)}) * - log b (real (distribution (\x. (X x, Y x, Z x)) {(x, y, z)}) / - (real (joint_distribution X Z {(x, z)}) * real (joint_distribution Y Z {(y,z)} / distribution Z {z}))))" - using conditional_mutual_information_generic_eq[OF assms[THEN simple_function_imp_finite_random_variable]] - by simp + distribution (\x. (X x, Y x, Z x)) {(x, y, z)} * + log b (distribution (\x. (X x, Y x, Z x)) {(x, y, z)} / + (joint_distribution X Z {(x, z)} * joint_distribution Y Z {(y,z)} / distribution Z {z})))" + by (subst conditional_mutual_information_generic_eq[OF assms[THEN simple_function_imp_finite_random_variable]]) + simp lemma (in information_space) conditional_mutual_information_eq_mutual_information: assumes X: "simple_function M X" and Y: "simple_function M Y" @@ -683,10 +688,10 @@ qed lemma (in prob_space) distribution_unit[simp]: "distribution (\x. ()) {()} = 1" - unfolding distribution_def using measure_space_1 by auto + unfolding distribution_def using prob_space by auto lemma (in prob_space) joint_distribution_unit[simp]: "distribution (\x. (X x, ())) {(a, ())} = distribution X {a}" - unfolding distribution_def by (auto intro!: arg_cong[where f=\]) + unfolding distribution_def by (auto intro!: arg_cong[where f=\']) lemma (in prob_space) setsum_distribution: assumes X: "finite_random_variable MX X" shows "(\a\space MX. distribution X {a}) = 1" @@ -695,12 +700,13 @@ lemma (in prob_space) setsum_real_distribution: fixes MX :: "('c, 'd) measure_space_scheme" - assumes X: "finite_random_variable MX X" shows "(\a\space MX. real (distribution X {a})) = 1" - using setsum_real_joint_distribution[OF assms, of "\ space = UNIV, sets = Pow UNIV, measure = undefined \" "\x. ()" "{()}"] - using sigma_algebra_Pow[of "UNIV::unit set" "\ measure = undefined \"] by simp + assumes X: "finite_random_variable MX X" shows "(\a\space MX. distribution X {a}) = 1" + using setsum_joint_distribution[OF assms, of "\ space = UNIV, sets = Pow UNIV, measure = undefined \" "\x. ()" "{()}"] + using sigma_algebra_Pow[of "UNIV::unit set" "\ measure = undefined \"] + by auto lemma (in information_space) conditional_mutual_information_generic_positive: - assumes "finite_random_variable MX X" and "finite_random_variable MY Y" and "finite_random_variable MZ Z" + assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y" and Z: "finite_random_variable MZ Z" shows "0 \ conditional_mutual_information b MX MY MZ X Y Z" proof (cases "space MX \ space MY \ space MZ = {}") case True show ?thesis @@ -708,43 +714,35 @@ by simp next case False - let "?dXYZ A" = "real (distribution (\x. (X x, Y x, Z x)) A)" - let "?dXZ A" = "real (joint_distribution X Z A)" - let "?dYZ A" = "real (joint_distribution Y Z A)" - let "?dX A" = "real (distribution X A)" - let "?dZ A" = "real (distribution Z A)" + let ?dXYZ = "distribution (\x. (X x, Y x, Z x))" + let ?dXZ = "joint_distribution X Z" + let ?dYZ = "joint_distribution Y Z" + let ?dX = "distribution X" + let ?dZ = "distribution Z" let ?M = "space MX \ space MY \ space MZ" - have split_beta: "\f. split f = (\x. f (fst x) (snd x))" by (simp add: fun_eq_iff) - - note space_simps = space_pair_measure space_sigma algebra.simps - - note finite_var = assms - note YZ = finite_random_variable_pairI[OF finite_var(2,3)] - note XZ = finite_random_variable_pairI[OF finite_var(1,3)] - note ZX = finite_random_variable_pairI[OF finite_var(3,1)] - note YZ = finite_random_variable_pairI[OF finite_var(2,3)] - note XYZ = finite_random_variable_pairI[OF finite_var(1) YZ] - note finite = finite_var(3) YZ XZ XYZ - note finite = finite[THEN finite_distribution_finite, simplified space_simps] - + note YZ = finite_random_variable_pairI[OF Y Z] + note XZ = finite_random_variable_pairI[OF X Z] + note ZX = finite_random_variable_pairI[OF Z X] + note YZ = finite_random_variable_pairI[OF Y Z] + note XYZ = finite_random_variable_pairI[OF X YZ] + note finite = Z YZ XZ XYZ have order: "\x y z. \x \ space MX; y \ space MY; z \ space MZ; joint_distribution X Z {(x, z)} = 0\ \ joint_distribution X (\x. (Y x, Z x)) {(x, y, z)} = 0" unfolding joint_distribution_commute_singleton[of X] unfolding joint_distribution_assoc_singleton[symmetric] - using finite_distribution_order(6)[OF finite_var(2) ZX] - by (auto simp: space_simps) + using finite_distribution_order(6)[OF Y ZX] + by auto note order = order - finite_distribution_order(5,6)[OF finite_var(1) YZ, simplified space_simps] - finite_distribution_order(5,6)[OF finite_var(2,3), simplified space_simps] + finite_distribution_order(5,6)[OF X YZ] + finite_distribution_order(5,6)[OF Y Z] have "- conditional_mutual_information b MX MY MZ X Y Z = - (\(x, y, z) \ ?M. ?dXYZ {(x, y, z)} * log b (?dXYZ {(x, y, z)} / (?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z})))" - unfolding conditional_mutual_information_generic_eq[OF assms] neg_equal_iff_equal - by (intro setsum_cong) (auto intro!: arg_cong[where f="log b"] simp: real_of_pextreal_mult[symmetric]) + unfolding conditional_mutual_information_generic_eq[OF assms] neg_equal_iff_equal by auto also have "\ \ log b (\(x, y, z) \ ?M. ?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z})" - unfolding split_beta + unfolding split_beta' proof (rule log_setsum_divide) show "?M \ {}" using False by simp show "1 < b" using b_gt_1 . @@ -757,33 +755,31 @@ unfolding setsum_commute[of _ "space MY"] unfolding setsum_commute[of _ "space MZ"] by (simp_all add: space_pair_measure - setsum_real_joint_distribution_singleton[OF `finite_random_variable MX X` YZ] - setsum_real_joint_distribution_singleton[OF `finite_random_variable MY Y` finite_var(3)] - setsum_real_distribution[OF `finite_random_variable MZ Z`]) + setsum_joint_distribution_singleton[OF X YZ] + setsum_joint_distribution_singleton[OF Y Z] + setsum_distribution[OF Z]) fix x assume "x \ ?M" let ?x = "(fst x, fst (snd x), snd (snd x))" - show "0 \ ?dXYZ {?x}" using real_pextreal_nonneg . - show "0 \ ?dXZ {(fst x, snd (snd x))} * ?dYZ {(fst (snd x), snd (snd x))} / ?dZ {snd (snd x)}" - by (simp add: real_pextreal_nonneg mult_nonneg_nonneg divide_nonneg_nonneg) + show "0 \ ?dXYZ {?x}" + "0 \ ?dXZ {(fst x, snd (snd x))} * ?dYZ {(fst (snd x), snd (snd x))} / ?dZ {snd (snd x)}" + by (simp_all add: mult_nonneg_nonneg divide_nonneg_nonneg) assume *: "0 < ?dXYZ {?x}" - with `x \ ?M` show "0 < ?dXZ {(fst x, snd (snd x))} * ?dYZ {(fst (snd x), snd (snd x))} / ?dZ {snd (snd x)}" - using finite order - by (cases x) - (auto simp add: zero_less_real_of_pextreal zero_less_mult_iff zero_less_divide_iff) + with `x \ ?M` finite order show "0 < ?dXZ {(fst x, snd (snd x))} * ?dYZ {(fst (snd x), snd (snd x))} / ?dZ {snd (snd x)}" + by (cases x) (auto simp add: zero_le_mult_iff zero_le_divide_iff less_le) qed also have "(\(x, y, z) \ ?M. ?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z}) = (\z\space MZ. ?dZ {z})" apply (simp add: setsum_cartesian_product') apply (subst setsum_commute) apply (subst (2) setsum_commute) by (auto simp: setsum_divide_distrib[symmetric] setsum_product[symmetric] - setsum_real_joint_distribution_singleton[OF finite_var(1,3)] - setsum_real_joint_distribution_singleton[OF finite_var(2,3)] + setsum_joint_distribution_singleton[OF X Z] + setsum_joint_distribution_singleton[OF Y Z] intro!: setsum_cong) also have "log b (\z\space MZ. ?dZ {z}) = 0" - unfolding setsum_real_distribution[OF finite_var(3)] by simp + unfolding setsum_real_distribution[OF Z] by simp finally show ?thesis by simp qed @@ -800,57 +796,52 @@ abbreviation (in information_space) conditional_entropy_Pow ("\'(_ | _')") where "\(X | Y) \ conditional_entropy b - \ space = X`space M, sets = Pow (X`space M), measure = distribution X \ - \ space = Y`space M, sets = Pow (Y`space M), measure = distribution Y \ X Y" + \ space = X`space M, sets = Pow (X`space M), measure = extreal\distribution X \ + \ space = Y`space M, sets = Pow (Y`space M), measure = extreal\distribution Y \ X Y" lemma (in information_space) conditional_entropy_positive: "simple_function M X \ simple_function M Y \ 0 \ \(X | Y)" unfolding conditional_entropy_def by (auto intro!: conditional_mutual_information_positive) -lemma (in measure_space) empty_measureI: "A = {} \ \ A = 0" by simp - lemma (in information_space) conditional_entropy_generic_eq: fixes MX :: "('c, 'd) measure_space_scheme" and MY :: "('e, 'f) measure_space_scheme" assumes MX: "finite_random_variable MX X" assumes MZ: "finite_random_variable MZ Z" shows "conditional_entropy b MX MZ X Z = - (\(x, z)\space MX \ space MZ. - real (joint_distribution X Z {(x, z)}) * - log b (real (joint_distribution X Z {(x, z)}) / real (distribution Z {z})))" + joint_distribution X Z {(x, z)} * log b (joint_distribution X Z {(x, z)} / distribution Z {z}))" proof - interpret MX: finite_sigma_algebra MX using MX by simp interpret MZ: finite_sigma_algebra MZ using MZ by simp let "?XXZ x y z" = "joint_distribution X (\x. (X x, Z x)) {(x, y, z)}" let "?XZ x z" = "joint_distribution X Z {(x, z)}" let "?Z z" = "distribution Z {z}" - let "?f x y z" = "log b (real (?XXZ x y z) / (real (?XZ x z) * real (?XZ y z / ?Z z)))" + let "?f x y z" = "log b (?XXZ x y z * ?Z z / (?XZ x z * ?XZ y z))" { fix x z have "?XXZ x x z = ?XZ x z" - unfolding distribution_def by (auto intro!: arg_cong[where f=\]) } + unfolding distribution_def by (auto intro!: arg_cong[where f=\']) } note this[simp] { fix x x' :: 'c and z assume "x' \ x" then have "?XXZ x x' z = 0" - by (auto simp: distribution_def intro!: arg_cong[where f=\] empty_measureI) } + by (auto simp: distribution_def empty_measure'[symmetric] + simp del: empty_measure' intro!: arg_cong[where f=\']) } note this[simp] { fix x x' z assume *: "x \ space MX" "z \ space MZ" - then have "(\x'\space MX. real (?XXZ x x' z) * ?f x x' z) - = (\x'\space MX. if x = x' then real (?XZ x z) * ?f x x z else 0)" + then have "(\x'\space MX. ?XXZ x x' z * ?f x x' z) + = (\x'\space MX. if x = x' then ?XZ x z * ?f x x z else 0)" by (auto intro!: setsum_cong) - also have "\ = real (?XZ x z) * ?f x x z" + also have "\ = ?XZ x z * ?f x x z" using `x \ space MX` by (simp add: setsum_cases[OF MX.finite_space]) - also have "\ = real (?XZ x z) * log b (real (?Z z) / real (?XZ x z))" - by (auto simp: real_of_pextreal_mult[symmetric]) - also have "\ = - real (?XZ x z) * log b (real (?XZ x z) / real (?Z z))" - using assms[THEN finite_distribution_finite] + also have "\ = ?XZ x z * log b (?Z z / ?XZ x z)" by auto + also have "\ = - ?XZ x z * log b (?XZ x z / ?Z z)" using finite_distribution_order(6)[OF MX MZ] - by (auto simp: log_simps field_simps zero_less_mult_iff zero_less_real_of_pextreal real_of_pextreal_eq_0) - finally have "(\x'\space MX. real (?XXZ x x' z) * ?f x x' z) = - - real (?XZ x z) * log b (real (?XZ x z) / real (?Z z))" . } + by (auto simp: log_simps field_simps zero_less_mult_iff) + finally have "(\x'\space MX. ?XXZ x x' z * ?f x x' z) = - ?XZ x z * log b (?XZ x z / ?Z z)" . } note * = this show ?thesis unfolding conditional_entropy_def unfolding conditional_mutual_information_generic_eq[OF MX MX MZ] by (auto simp: setsum_cartesian_product' setsum_negf[symmetric] - setsum_commute[of _ "space MZ"] * simp del: divide_pextreal_def + setsum_commute[of _ "space MZ"] * intro!: setsum_cong) qed @@ -858,29 +849,27 @@ assumes "simple_function M X" "simple_function M Z" shows "\(X | Z) = - (\(x, z)\X ` space M \ Z ` space M. - real (joint_distribution X Z {(x, z)}) * - log b (real (joint_distribution X Z {(x, z)}) / real (distribution Z {z})))" - using conditional_entropy_generic_eq[OF assms[THEN simple_function_imp_finite_random_variable]] - by simp + joint_distribution X Z {(x, z)} * + log b (joint_distribution X Z {(x, z)} / distribution Z {z}))" + by (subst conditional_entropy_generic_eq[OF assms[THEN simple_function_imp_finite_random_variable]]) + simp lemma (in information_space) conditional_entropy_eq_ce_with_hypothesis: assumes X: "simple_function M X" and Y: "simple_function M Y" shows "\(X | Y) = - -(\y\Y`space M. real (distribution Y {y}) * - (\x\X`space M. real (joint_distribution X Y {(x,y)}) / real (distribution Y {(y)}) * - log b (real (joint_distribution X Y {(x,y)}) / real (distribution Y {(y)}))))" + -(\y\Y`space M. distribution Y {y} * + (\x\X`space M. joint_distribution X Y {(x,y)} / distribution Y {(y)} * + log b (joint_distribution X Y {(x,y)} / distribution Y {(y)})))" unfolding conditional_entropy_eq[OF assms] - using finite_distribution_finite[OF finite_random_variable_pairI[OF assms[THEN simple_function_imp_finite_random_variable]]] using finite_distribution_order(5,6)[OF assms[THEN simple_function_imp_finite_random_variable]] - using finite_distribution_finite[OF Y[THEN simple_function_imp_finite_random_variable]] - by (auto simp: setsum_cartesian_product' setsum_commute[of _ "Y`space M"] setsum_right_distrib real_of_pextreal_eq_0 + by (auto simp: setsum_cartesian_product' setsum_commute[of _ "Y`space M"] setsum_right_distrib intro!: setsum_cong) lemma (in information_space) conditional_entropy_eq_cartesian_product: assumes "simple_function M X" "simple_function M Y" shows "\(X | Y) = -(\x\X`space M. \y\Y`space M. - real (joint_distribution X Y {(x,y)}) * - log b (real (joint_distribution X Y {(x,y)}) / real (distribution Y {y})))" + joint_distribution X Y {(x,y)} * + log b (joint_distribution X Y {(x,y)} / distribution Y {y}))" unfolding conditional_entropy_eq[OF assms] by (auto intro!: setsum_cong simp: setsum_cartesian_product') @@ -890,24 +879,22 @@ assumes X: "simple_function M X" and Z: "simple_function M Z" shows "\(X ; Z) = \(X) - \(X | Z)" proof - - let "?XZ x z" = "real (joint_distribution X Z {(x, z)})" - let "?Z z" = "real (distribution Z {z})" - let "?X x" = "real (distribution X {x})" + let "?XZ x z" = "joint_distribution X Z {(x, z)}" + let "?Z z" = "distribution Z {z}" + let "?X x" = "distribution X {x}" note fX = X[THEN simple_function_imp_finite_random_variable] note fZ = Z[THEN simple_function_imp_finite_random_variable] - note fX[THEN finite_distribution_finite, simp] and fZ[THEN finite_distribution_finite, simp] note finite_distribution_order[OF fX fZ, simp] { fix x z assume "x \ X`space M" "z \ Z`space M" have "?XZ x z * log b (?XZ x z / (?X x * ?Z z)) = ?XZ x z * log b (?XZ x z / ?Z z) - ?XZ x z * log b (?X x)" - by (auto simp: log_simps real_of_pextreal_mult[symmetric] zero_less_mult_iff - zero_less_real_of_pextreal field_simps real_of_pextreal_eq_0 abs_mult) } + by (auto simp: log_simps zero_le_mult_iff field_simps less_le) } note * = this show ?thesis unfolding entropy_eq[OF X] conditional_entropy_eq[OF X Z] mutual_information_eq[OF X Z] - using setsum_real_joint_distribution_singleton[OF fZ fX, unfolded joint_distribution_commute_singleton[of Z X]] + using setsum_joint_distribution_singleton[OF fZ fX, unfolded joint_distribution_commute_singleton[of Z X]] by (simp add: * setsum_cartesian_product' setsum_subtractf setsum_left_distrib[symmetric] - setsum_real_distribution) + setsum_distribution) qed lemma (in information_space) conditional_entropy_less_eq_entropy: @@ -923,21 +910,19 @@ assumes X: "simple_function M X" and Y: "simple_function M Y" shows "\(\x. (X x, Y x)) = \(X) + \(Y|X)" proof - - let "?XY x y" = "real (joint_distribution X Y {(x, y)})" - let "?Y y" = "real (distribution Y {y})" - let "?X x" = "real (distribution X {x})" + let "?XY x y" = "joint_distribution X Y {(x, y)}" + let "?Y y" = "distribution Y {y}" + let "?X x" = "distribution X {x}" note fX = X[THEN simple_function_imp_finite_random_variable] note fY = Y[THEN simple_function_imp_finite_random_variable] - note fX[THEN finite_distribution_finite, simp] and fY[THEN finite_distribution_finite, simp] note finite_distribution_order[OF fX fY, simp] { fix x y assume "x \ X`space M" "y \ Y`space M" have "?XY x y * log b (?XY x y / ?X x) = ?XY x y * log b (?XY x y) - ?XY x y * log b (?X x)" - by (auto simp: log_simps real_of_pextreal_mult[symmetric] zero_less_mult_iff - zero_less_real_of_pextreal field_simps real_of_pextreal_eq_0 abs_mult) } + by (auto simp: log_simps zero_le_mult_iff field_simps less_le) } note * = this show ?thesis - using setsum_real_joint_distribution_singleton[OF fY fX] + using setsum_joint_distribution_singleton[OF fY fX] unfolding entropy_eq[OF X] conditional_entropy_eq_cartesian_product[OF Y X] entropy_eq_cartesian_product[OF X Y] unfolding joint_distribution_commute_singleton[of Y X] setsum_commute[of _ "X`space M"] by (simp add: * setsum_subtractf setsum_left_distrib[symmetric]) @@ -1063,23 +1048,21 @@ assumes svi: "subvimage (space M) X P" shows "\(X) = \(P) + \(X|P)" proof - - let "?XP x p" = "real (joint_distribution X P {(x, p)})" - let "?X x" = "real (distribution X {x})" - let "?P p" = "real (distribution P {p})" + let "?XP x p" = "joint_distribution X P {(x, p)}" + let "?X x" = "distribution X {x}" + let "?P p" = "distribution P {p}" note fX = sf(1)[THEN simple_function_imp_finite_random_variable] note fP = sf(2)[THEN simple_function_imp_finite_random_variable] - note fX[THEN finite_distribution_finite, simp] and fP[THEN finite_distribution_finite, simp] note finite_distribution_order[OF fX fP, simp] - have "(\x\X ` space M. real (distribution X {x}) * log b (real (distribution X {x}))) = - (\y\P `space M. \x\X ` space M. - real (joint_distribution X P {(x, y)}) * log b (real (joint_distribution X P {(x, y)})))" + have "(\x\X ` space M. ?X x * log b (?X x)) = + (\y\P `space M. \x\X ` space M. ?XP x y * log b (?XP x y))" proof (subst setsum_image_split[OF svi], safe intro!: setsum_mono_zero_cong_left imageI) show "finite (X ` space M)" "finite (X ` space M)" "finite (P ` space M)" using sf unfolding simple_function_def by auto next fix p x assume in_space: "p \ space M" "x \ space M" - assume "real (joint_distribution X P {(X x, P p)}) * log b (real (joint_distribution X P {(X x, P p)})) \ 0" + assume "?XP (X x) (P p) * log b (?XP (X x) (P p)) \ 0" hence "(\x. (X x, P x)) -` {(X x, P p)} \ space M \ {}" by (auto simp: distribution_def) with svi[unfolded subvimage_def, rule_format, OF `x \ space M`] show "x \ P -` {P p}" by auto @@ -1091,20 +1074,16 @@ by auto hence "(\x. (X x, P x)) -` {(X x, P p)} \ space M = X -` {X x} \ space M" by auto - thus "real (distribution X {X x}) * log b (real (distribution X {X x})) = - real (joint_distribution X P {(X x, P p)}) * - log b (real (joint_distribution X P {(X x, P p)}))" + thus "?X (X x) * log b (?X (X x)) = ?XP (X x) (P p) * log b (?XP (X x) (P p))" by (auto simp: distribution_def) qed - moreover have "\x y. real (joint_distribution X P {(x, y)}) * - log b (real (joint_distribution X P {(x, y)}) / real (distribution P {y})) = - real (joint_distribution X P {(x, y)}) * log b (real (joint_distribution X P {(x, y)})) - - real (joint_distribution X P {(x, y)}) * log b (real (distribution P {y}))" + moreover have "\x y. ?XP x y * log b (?XP x y / ?P y) = + ?XP x y * log b (?XP x y) - ?XP x y * log b (?P y)" by (auto simp add: log_simps zero_less_mult_iff field_simps) ultimately show ?thesis unfolding sf[THEN entropy_eq] conditional_entropy_eq[OF sf] - using setsum_real_joint_distribution_singleton[OF fX fP] - by (simp add: setsum_cartesian_product' setsum_subtractf setsum_real_distribution + using setsum_joint_distribution_singleton[OF fX fP] + by (simp add: setsum_cartesian_product' setsum_subtractf setsum_distribution setsum_left_distrib[symmetric] setsum_commute[where B="P`space M"]) qed diff -r 28b51effc5ed -r cdf7693bbe08 src/HOL/Probability/Lebesgue_Integration.thy --- a/src/HOL/Probability/Lebesgue_Integration.thy Mon Mar 14 14:37:47 2011 +0100 +++ b/src/HOL/Probability/Lebesgue_Integration.thy Mon Mar 14 14:37:49 2011 +0100 @@ -6,6 +6,88 @@ imports Measure Borel_Space begin +lemma extreal_indicator_pos[simp,intro]: "0 \ (indicator A x ::extreal)" + unfolding indicator_def by auto + +lemma tendsto_real_max: + fixes x y :: real + assumes "(X ---> x) net" + assumes "(Y ---> y) net" + shows "((\x. max (X x) (Y x)) ---> max x y) net" +proof - + have *: "\x y :: real. max x y = y + ((x - y) + norm (x - y)) / 2" + by (auto split: split_max simp: field_simps) + show ?thesis + unfolding * + by (intro tendsto_add assms tendsto_divide tendsto_norm tendsto_diff) auto +qed + +lemma (in measure_space) measure_Union: + assumes "finite S" "S \ sets M" "\A B. A \ S \ B \ S \ A \ B \ A \ B = {}" + shows "setsum \ S = \ (\S)" +proof - + have "setsum \ S = \ (\i\S. i)" + using assms by (intro measure_setsum[OF `finite S`]) (auto simp: disjoint_family_on_def) + also have "\ = \ (\S)" by (auto intro!: arg_cong[where f=\]) + finally show ?thesis . +qed + +lemma (in sigma_algebra) measurable_sets2[intro]: + assumes "f \ measurable M M'" "g \ measurable M M''" + and "A \ sets M'" "B \ sets M''" + shows "f -` A \ g -` B \ space M \ sets M" +proof - + have "f -` A \ g -` B \ space M = (f -` A \ space M) \ (g -` B \ space M)" + by auto + then show ?thesis using assms by (auto intro: measurable_sets) +qed + +lemma incseq_extreal: "incseq f \ incseq (\x. extreal (f x))" + unfolding incseq_def by auto + +lemma incseq_Suc_iff: "incseq f \ (\n. f n \ f (Suc n))" +proof + assume "\n. f n \ f (Suc n)" then show "incseq f" by (auto intro!: incseq_SucI) +qed (auto simp: incseq_def) + +lemma borel_measurable_real_floor: + "(\x::real. real \x\) \ borel_measurable borel" + unfolding borel.borel_measurable_iff_ge +proof (intro allI) + fix a :: real + { fix x have "a \ real \x\ \ real \a\ \ x" + using le_floor_eq[of "\a\" x] ceiling_le_iff[of a "\x\"] + unfolding real_eq_of_int by simp } + then have "{w::real \ space borel. a \ real \w\} = {real \a\..}" by auto + then show "{w::real \ space borel. a \ real \w\} \ sets borel" by auto +qed + +lemma measure_preservingD2: + "f \ measure_preserving A B \ f \ measurable A B" + unfolding measure_preserving_def by auto + +lemma measure_preservingD3: + "f \ measure_preserving A B \ f \ space A \ space B" + unfolding measure_preserving_def measurable_def by auto + +lemma measure_preservingD: + "T \ measure_preserving A B \ X \ sets B \ measure A (T -` X \ space A) = measure B X" + unfolding measure_preserving_def by auto + +lemma (in sigma_algebra) borel_measurable_real_natfloor[intro, simp]: + assumes "f \ borel_measurable M" + shows "(\x. real (natfloor (f x))) \ borel_measurable M" +proof - + have "\x. real (natfloor (f x)) = max 0 (real \f x\)" + by (auto simp: max_def natfloor_def) + with borel_measurable_max[OF measurable_comp[OF assms borel_measurable_real_floor] borel_measurable_const] + show ?thesis by (simp add: comp_def) +qed + +lemma (in measure_space) AE_not_in: + assumes N: "N \ null_sets" shows "AE x. x \ N" + using N by (rule AE_I') auto + lemma sums_If_finite: fixes f :: "nat \ 'a::real_normed_vector" assumes finite: "finite {r. P r}" @@ -55,8 +137,17 @@ by (auto intro!: finite_UN simp del: UN_simps simp: simple_function_def) qed +lemma (in sigma_algebra) simple_function_measurable2[intro]: + assumes "simple_function M f" "simple_function M g" + shows "f -` A \ g -` B \ space M \ sets M" +proof - + have "f -` A \ g -` B \ space M = (f -` A \ space M) \ (g -` B \ space M)" + by auto + then show ?thesis using assms[THEN simple_functionD(2)] by auto +qed + lemma (in sigma_algebra) simple_function_indicator_representation: - fixes f ::"'a \ pextreal" + fixes f ::"'a \ extreal" assumes f: "simple_function M f" and x: "x \ space M" shows "f x = (\y \ f ` space M. y * indicator (f -` {y} \ space M) x)" (is "?l = ?r") @@ -71,7 +162,7 @@ qed lemma (in measure_space) simple_function_notspace: - "simple_function M (\x. h x * indicator (- space M) x::pextreal)" (is "simple_function M ?h") + "simple_function M (\x. h x * indicator (- space M) x::extreal)" (is "simple_function M ?h") proof - have "?h ` space M \ {0}" unfolding indicator_def by auto hence [simp, intro]: "finite (?h ` space M)" by (auto intro: finite_subset) @@ -111,16 +202,22 @@ qed lemma (in sigma_algebra) simple_function_borel_measurable: - fixes f :: "'a \ 'x::t2_space" + fixes f :: "'a \ 'x::{t2_space}" assumes "f \ borel_measurable M" and "finite (f ` space M)" shows "simple_function M f" using assms unfolding simple_function_def by (auto intro: borel_measurable_vimage) +lemma (in sigma_algebra) simple_function_eq_borel_measurable: + fixes f :: "'a \ extreal" + shows "simple_function M f \ finite (f`space M) \ f \ borel_measurable M" + using simple_function_borel_measurable[of f] + borel_measurable_simple_function[of f] + by (fastsimp simp: simple_function_def) + lemma (in sigma_algebra) simple_function_const[intro, simp]: "simple_function M (\x. c)" by (auto intro: finite_subset simp: simple_function_def) - lemma (in sigma_algebra) simple_function_compose[intro, simp]: assumes "simple_function M f" shows "simple_function M (g \ f)" @@ -189,6 +286,7 @@ and simple_function_mult[intro, simp] = simple_function_compose2[where h="op *"] and simple_function_div[intro, simp] = simple_function_compose2[where h="op /"] and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"] + and simple_function_max[intro, simp] = simple_function_compose2[where h=max] lemma (in sigma_algebra) simple_function_setsum[intro, simp]: assumes "\i. i \ P \ simple_function M (f i)" @@ -197,247 +295,168 @@ assume "finite P" from this assms show ?thesis by induct auto qed auto -lemma (in sigma_algebra) simple_function_le_measurable: - assumes "simple_function M f" "simple_function M g" - shows "{x \ space M. f x \ g x} \ sets M" -proof - - have *: "{x \ space M. f x \ g x} = - (\(F, G)\f`space M \ g`space M. - if F \ G then (f -` {F} \ space M) \ (g -` {G} \ space M) else {})" - apply (auto split: split_if_asm) - apply (rule_tac x=x in bexI) - apply (rule_tac x=x in bexI) - by simp_all - have **: "\x y. x \ space M \ y \ space M \ - (f -` {f x} \ space M) \ (g -` {g y} \ space M) \ sets M" - using assms unfolding simple_function_def by auto - have "finite (f`space M \ g`space M)" - using assms unfolding simple_function_def by auto - thus ?thesis unfolding * - apply (rule finite_UN) - using assms unfolding simple_function_def - by (auto intro!: **) -qed +lemma (in sigma_algebra) + fixes f g :: "'a \ real" assumes sf: "simple_function M f" + shows simple_function_extreal[intro, simp]: "simple_function M (\x. extreal (f x))" + by (auto intro!: simple_function_compose1[OF sf]) + +lemma (in sigma_algebra) + fixes f g :: "'a \ nat" assumes sf: "simple_function M f" + shows simple_function_real_of_nat[intro, simp]: "simple_function M (\x. real (f x))" + by (auto intro!: simple_function_compose1[OF sf]) lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence: - fixes u :: "'a \ pextreal" + fixes u :: "'a \ extreal" assumes u: "u \ borel_measurable M" - shows "\f. (\i. simple_function M (f i) \ (\x\space M. f i x \ \)) \ f \ u" + shows "\f. incseq f \ (\i. \ \ range (f i) \ simple_function M (f i)) \ + (\x. (SUP i. f i x) = max 0 (u x)) \ (\i x. 0 \ f i x)" proof - - have "\f. \x j. (of_nat j \ u x \ f x j = j*2^j) \ - (u x < of_nat j \ of_nat (f x j) \ u x * 2^j \ u x * 2^j < of_nat (Suc (f x j)))" - (is "\f. \x j. ?P x j (f x j)") - proof(rule choice, rule, rule choice, rule) - fix x j show "\n. ?P x j n" - proof cases - assume *: "u x < of_nat j" - then obtain r where r: "u x = Real r" "0 \ r" by (cases "u x") auto - from reals_Archimedean6a[of "r * 2^j"] - obtain n :: nat where "real n \ r * 2 ^ j" "r * 2 ^ j < real (Suc n)" - using `0 \ r` by (auto simp: zero_le_power zero_le_mult_iff) - thus ?thesis using r * by (auto intro!: exI[of _ n]) - qed auto - qed - then obtain f where top: "\j x. of_nat j \ u x \ f x j = j*2^j" and - upper: "\j x. u x < of_nat j \ of_nat (f x j) \ u x * 2^j" and - lower: "\j x. u x < of_nat j \ u x * 2^j < of_nat (Suc (f x j))" by blast - - { fix j x P - assume 1: "of_nat j \ u x \ P (j * 2^j)" - assume 2: "\k. \ u x < of_nat j ; of_nat k \ u x * 2^j ; u x * 2^j < of_nat (Suc k) \ \ P k" - have "P (f x j)" - proof cases - assume "of_nat j \ u x" thus "P (f x j)" - using top[of j x] 1 by auto - next - assume "\ of_nat j \ u x" - hence "u x < of_nat j" "of_nat (f x j) \ u x * 2^j" "u x * 2^j < of_nat (Suc (f x j))" - using upper lower by auto - from 2[OF this] show "P (f x j)" . - qed } - note fI = this - - { fix j x - have "f x j = j * 2 ^ j \ of_nat j \ u x" - by (rule fI, simp, cases "u x") (auto split: split_if_asm) } - note f_eq = this - - { fix j x - have "f x j \ j * 2 ^ j" - proof (rule fI) - fix k assume *: "u x < of_nat j" - assume "of_nat k \ u x * 2 ^ j" - also have "\ \ of_nat (j * 2^j)" - using * by (cases "u x") (auto simp: zero_le_mult_iff) - finally show "k \ j*2^j" by (auto simp del: real_of_nat_mult) - qed simp } + def f \ "\x i. if real i \ u x then i * 2 ^ i else natfloor (real (u x) * 2 ^ i)" + { fix x j have "f x j \ j * 2 ^ j" unfolding f_def + proof (split split_if, intro conjI impI) + assume "\ real j \ u x" + then have "natfloor (real (u x) * 2 ^ j) \ natfloor (j * 2 ^ j)" + by (cases "u x") (auto intro!: natfloor_mono simp: mult_nonneg_nonneg) + moreover have "real (natfloor (j * 2 ^ j)) \ j * 2^j" + by (intro real_natfloor_le) (auto simp: mult_nonneg_nonneg) + ultimately show "natfloor (real (u x) * 2 ^ j) \ j * 2 ^ j" + unfolding real_of_nat_le_iff by auto + qed auto } note f_upper = this - let "?g j x" = "of_nat (f x j) / 2^j :: pextreal" - show ?thesis unfolding simple_function_def isoton_fun_expand unfolding isoton_def - proof (safe intro!: exI[of _ ?g]) - fix j - have *: "?g j ` space M \ (\x. of_nat x / 2^j) ` {..j * 2^j}" - using f_upper by auto - thus "finite (?g j ` space M)" by (rule finite_subset) auto - next - fix j t assume "t \ space M" - have **: "?g j -` {?g j t} \ space M = {x \ space M. f t j = f x j}" - by (auto simp: power_le_zero_eq Real_eq_Real mult_le_0_iff) + have real_f: + "\i x. real (f x i) = (if real i \ u x then i * 2 ^ i else real (natfloor (real (u x) * 2 ^ i)))" + unfolding f_def by auto - show "?g j -` {?g j t} \ space M \ sets M" - proof cases - assume "of_nat j \ u t" - hence "?g j -` {?g j t} \ space M = {x\space M. of_nat j \ u x}" - unfolding ** f_eq[symmetric] by auto - thus "?g j -` {?g j t} \ space M \ sets M" - using u by auto - next - assume not_t: "\ of_nat j \ u t" - hence less: "f t j < j*2^j" using f_eq[symmetric] `f t j \ j*2^j` by auto - have split_vimage: "?g j -` {?g j t} \ space M = - {x\space M. of_nat (f t j)/2^j \ u x} \ {x\space M. u x < of_nat (Suc (f t j))/2^j}" - unfolding ** - proof safe - fix x assume [simp]: "f t j = f x j" - have *: "\ of_nat j \ u x" using not_t f_eq[symmetric] by simp - hence "of_nat (f t j) \ u x * 2^j \ u x * 2^j < of_nat (Suc (f t j))" - using upper lower by auto - hence "of_nat (f t j) / 2^j \ u x \ u x < of_nat (Suc (f t j))/2^j" using * - by (cases "u x") (auto simp: zero_le_mult_iff power_le_zero_eq divide_real_def[symmetric] field_simps) - thus "of_nat (f t j) / 2^j \ u x" "u x < of_nat (Suc (f t j))/2^j" by auto + let "?g j x" = "real (f x j) / 2^j :: extreal" + show ?thesis + proof (intro exI[of _ ?g] conjI allI ballI) + fix i + have "simple_function M (\x. real (f x i))" + proof (intro simple_function_borel_measurable) + show "(\x. real (f x i)) \ borel_measurable M" + using u by (auto intro!: measurable_If simp: real_f) + have "(\x. real (f x i))`space M \ real`{..i*2^i}" + using f_upper[of _ i] by auto + then show "finite ((\x. real (f x i))`space M)" + by (rule finite_subset) auto + qed + then show "simple_function M (?g i)" + by (auto intro: simple_function_extreal simple_function_div) + next + show "incseq ?g" + proof (intro incseq_extreal incseq_SucI le_funI) + fix x and i :: nat + have "f x i * 2 \ f x (Suc i)" unfolding f_def + proof ((split split_if)+, intro conjI impI) + assume "extreal (real i) \ u x" "\ extreal (real (Suc i)) \ u x" + then show "i * 2 ^ i * 2 \ natfloor (real (u x) * 2 ^ Suc i)" + by (cases "u x") (auto intro!: le_natfloor) next - fix x - assume "of_nat (f t j) / 2^j \ u x" "u x < of_nat (Suc (f t j))/2^j" - hence "of_nat (f t j) \ u x * 2 ^ j \ u x * 2 ^ j < of_nat (Suc (f t j))" - by (cases "u x") (auto simp: zero_le_mult_iff power_le_zero_eq divide_real_def[symmetric] field_simps) - hence 1: "of_nat (f t j) \ u x * 2 ^ j" and 2: "u x * 2 ^ j < of_nat (Suc (f t j))" by auto - note 2 - also have "\ \ of_nat (j*2^j)" - using less by (auto simp: zero_le_mult_iff simp del: real_of_nat_mult) - finally have bound_ux: "u x < of_nat j" - by (cases "u x") (auto simp: zero_le_mult_iff power_le_zero_eq) - show "f t j = f x j" - proof (rule antisym) - from 1 lower[OF bound_ux] - show "f t j \ f x j" by (cases "u x") (auto split: split_if_asm) - from upper[OF bound_ux] 2 - show "f x j \ f t j" by (cases "u x") (auto split: split_if_asm) + assume "\ extreal (real i) \ u x" "extreal (real (Suc i)) \ u x" + then show "natfloor (real (u x) * 2 ^ i) * 2 \ Suc i * 2 ^ Suc i" + by (cases "u x") auto + next + assume "\ extreal (real i) \ u x" "\ extreal (real (Suc i)) \ u x" + have "natfloor (real (u x) * 2 ^ i) * 2 = natfloor (real (u x) * 2 ^ i) * natfloor 2" + by simp + also have "\ \ natfloor (real (u x) * 2 ^ i * 2)" + proof cases + assume "0 \ u x" then show ?thesis + by (intro le_mult_natfloor) (cases "u x", auto intro!: mult_nonneg_nonneg) + next + assume "\ 0 \ u x" then show ?thesis + by (cases "u x") (auto simp: natfloor_neg mult_nonpos_nonneg) qed - qed - show ?thesis unfolding split_vimage using u by auto + also have "\ = natfloor (real (u x) * 2 ^ Suc i)" + by (simp add: ac_simps) + finally show "natfloor (real (u x) * 2 ^ i) * 2 \ natfloor (real (u x) * 2 ^ Suc i)" . + qed simp + then show "?g i x \ ?g (Suc i) x" + by (auto simp: field_simps) qed next - fix j t assume "?g j t = \" thus False by (auto simp: power_le_zero_eq) - next - fix t - { fix i - have "f t i * 2 \ f t (Suc i)" - proof (rule fI) - assume "of_nat (Suc i) \ u t" - hence "of_nat i \ u t" by (cases "u t") auto - thus "f t i * 2 \ Suc i * 2 ^ Suc i" unfolding f_eq[symmetric] by simp - next - fix k - assume *: "u t * 2 ^ Suc i < of_nat (Suc k)" - show "f t i * 2 \ k" - proof (rule fI) - assume "of_nat i \ u t" - hence "of_nat (i * 2^Suc i) \ u t * 2^Suc i" - by (cases "u t") (auto simp: zero_le_mult_iff power_le_zero_eq) - also have "\ < of_nat (Suc k)" using * by auto - finally show "i * 2 ^ i * 2 \ k" - by (auto simp del: real_of_nat_mult) - next - fix j assume "of_nat j \ u t * 2 ^ i" - with * show "j * 2 \ k" by (cases "u t") (auto simp: zero_le_mult_iff power_le_zero_eq) + fix x show "(SUP i. ?g i x) = max 0 (u x)" + proof (rule extreal_SUPI) + fix i show "?g i x \ max 0 (u x)" unfolding max_def f_def + by (cases "u x") (auto simp: field_simps real_natfloor_le natfloor_neg + mult_nonpos_nonneg mult_nonneg_nonneg) + next + fix y assume *: "\i. i \ UNIV \ ?g i x \ y" + have "\i. 0 \ ?g i x" by (auto simp: divide_nonneg_pos) + from order_trans[OF this *] have "0 \ y" by simp + show "max 0 (u x) \ y" + proof (cases y) + case (real r) + with * have *: "\i. f x i \ r * 2^i" by (auto simp: divide_le_eq) + from real_arch_lt[of r] * have "u x \ \" by (auto simp: f_def) (metis less_le_not_le) + then have "\p. max 0 (u x) = extreal p \ 0 \ p" by (cases "u x") (auto simp: max_def) + then guess p .. note ux = this + obtain m :: nat where m: "p < real m" using real_arch_lt .. + have "p \ r" + proof (rule ccontr) + assume "\ p \ r" + with LIMSEQ_inverse_realpow_zero[unfolded LIMSEQ_iff, rule_format, of 2 "p - r"] + obtain N where "\n\N. r * 2^n < p * 2^n - 1" by (auto simp: inverse_eq_divide field_simps) + then have "r * 2^max N m < p * 2^max N m - 1" by simp + moreover + have "real (natfloor (p * 2 ^ max N m)) \ r * 2 ^ max N m" + using *[of "max N m"] m unfolding real_f using ux + by (cases "0 \ u x") (simp_all add: max_def mult_nonneg_nonneg split: split_if_asm) + then have "p * 2 ^ max N m - 1 < r * 2 ^ max N m" + by (metis real_natfloor_gt_diff_one less_le_trans) + ultimately show False by auto qed - qed - thus "?g i t \ ?g (Suc i) t" - by (auto simp: zero_le_mult_iff power_le_zero_eq divide_real_def[symmetric] field_simps) } - hence mono: "mono (\i. ?g i t)" unfolding mono_iff_le_Suc by auto + then show "max 0 (u x) \ y" using real ux by simp + qed (insert `0 \ y`, auto) + qed + qed (auto simp: divide_nonneg_pos) +qed - show "(SUP j. of_nat (f t j) / 2 ^ j) = u t" - proof (rule pextreal_SUPI) - fix j show "of_nat (f t j) / 2 ^ j \ u t" - proof (rule fI) - assume "of_nat j \ u t" thus "of_nat (j * 2 ^ j) / 2 ^ j \ u t" - by (cases "u t") (auto simp: power_le_zero_eq divide_real_def[symmetric] field_simps) - next - fix k assume "of_nat k \ u t * 2 ^ j" - thus "of_nat k / 2 ^ j \ u t" - by (cases "u t") - (auto simp: power_le_zero_eq divide_real_def[symmetric] field_simps zero_le_mult_iff) - qed - next - fix y :: pextreal assume *: "\j. j \ UNIV \ of_nat (f t j) / 2 ^ j \ y" - show "u t \ y" - proof (cases "u t") - case (preal r) - show ?thesis - proof (rule ccontr) - assume "\ u t \ y" - then obtain p where p: "y = Real p" "0 \ p" "p < r" using preal by (cases y) auto - with LIMSEQ_inverse_realpow_zero[of 2, unfolded LIMSEQ_iff, rule_format, of "r - p"] - obtain n where n: "\N. n \ N \ inverse (2^N) < r - p" by auto - let ?N = "max n (natfloor r + 1)" - have "u t < of_nat ?N" "n \ ?N" - using ge_natfloor_plus_one_imp_gt[of r n] preal - using real_natfloor_add_one_gt - by (auto simp: max_def real_of_nat_Suc) - from lower[OF this(1)] - have "r < (real (f t ?N) + 1) / 2 ^ ?N" unfolding less_divide_eq - using preal by (auto simp add: power_le_zero_eq zero_le_mult_iff real_of_nat_Suc split: split_if_asm) - hence "u t < of_nat (f t ?N) / 2 ^ ?N + 1 / 2 ^ ?N" - using preal by (auto simp: field_simps divide_real_def[symmetric]) - with n[OF `n \ ?N`] p preal *[of ?N] - show False - by (cases "f t ?N") (auto simp: power_le_zero_eq split: split_if_asm) - qed - next - case infinite - { fix j have "f t j = j*2^j" using top[of j t] infinite by simp - hence "of_nat j \ y" using *[of j] - by (cases y) (auto simp: power_le_zero_eq zero_le_mult_iff field_simps) } - note all_less_y = this - show ?thesis unfolding infinite - proof (rule ccontr) - assume "\ \ \ y" then obtain r where r: "y = Real r" "0 \ r" by (cases y) auto - moreover obtain n ::nat where "r < real n" using ex_less_of_nat by (auto simp: real_eq_of_nat) - with all_less_y[of n] r show False by auto - qed - qed - qed +lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence': + fixes u :: "'a \ extreal" + assumes u: "u \ borel_measurable M" + obtains f where "\i. simple_function M (f i)" "incseq f" "\i. \ \ range (f i)" + "\x. (SUP i. f i x) = max 0 (u x)" "\i x. 0 \ f i x" + using borel_measurable_implies_simple_function_sequence[OF u] by auto + +lemma (in sigma_algebra) simple_function_If_set: + assumes sf: "simple_function M f" "simple_function M g" and A: "A \ space M \ sets M" + shows "simple_function M (\x. if x \ A then f x else g x)" (is "simple_function M ?IF") +proof - + def F \ "\x. f -` {x} \ space M" and G \ "\x. g -` {x} \ space M" + show ?thesis unfolding simple_function_def + proof safe + have "?IF ` space M \ f ` space M \ g ` space M" by auto + from finite_subset[OF this] assms + show "finite (?IF ` space M)" unfolding simple_function_def by auto + next + fix x assume "x \ space M" + then have *: "?IF -` {?IF x} \ space M = (if x \ A + then ((F (f x) \ (A \ space M)) \ (G (f x) - (G (f x) \ (A \ space M)))) + else ((F (g x) \ (A \ space M)) \ (G (g x) - (G (g x) \ (A \ space M)))))" + using sets_into_space[OF A] by (auto split: split_if_asm simp: G_def F_def) + have [intro]: "\x. F x \ sets M" "\x. G x \ sets M" + unfolding F_def G_def using sf[THEN simple_functionD(2)] by auto + show "?IF -` {?IF x} \ space M \ sets M" unfolding * using A by auto qed qed -lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence': - fixes u :: "'a \ pextreal" - assumes "u \ borel_measurable M" - obtains (x) f where "f \ u" "\i. simple_function M (f i)" "\i. \\f i`space M" +lemma (in sigma_algebra) simple_function_If: + assumes sf: "simple_function M f" "simple_function M g" and P: "{x\space M. P x} \ sets M" + shows "simple_function M (\x. if P x then f x else g x)" proof - - from borel_measurable_implies_simple_function_sequence[OF assms] - obtain f where x: "\i. simple_function M (f i)" "f \ u" - and fin: "\i. \x. x\space M \ f i x \ \" by auto - { fix i from fin[of _ i] have "\ \ f i`space M" by fastsimp } - with x show thesis by (auto intro!: that[of f]) + have "{x\space M. P x} = {x. P x} \ space M" by auto + with simple_function_If_set[OF sf, of "{x. P x}"] P show ?thesis by simp qed -lemma (in sigma_algebra) simple_function_eq_borel_measurable: - fixes f :: "'a \ pextreal" - shows "simple_function M f \ - finite (f`space M) \ f \ borel_measurable M" - using simple_function_borel_measurable[of f] - borel_measurable_simple_function[of f] - by (fastsimp simp: simple_function_def) - lemma (in measure_space) simple_function_restricted: - fixes f :: "'a \ pextreal" assumes "A \ sets M" + fixes f :: "'a \ extreal" assumes "A \ sets M" shows "simple_function (restricted_space A) f \ simple_function M (\x. f x * indicator A x)" (is "simple_function ?R f \ simple_function M ?f") proof - interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \ sets M`]) - have "finite (f`A) \ finite (?f`space M)" + have f: "finite (f`A) \ finite (?f`space M)" proof cases assume "A = space M" then have "f`A = ?f`space M" by (fastsimp simp: image_iff) @@ -456,7 +475,7 @@ using `A \ sets M` sets_into_space by (auto intro!: bexI[of _ x]) next fix x - assume "indicator A x \ (0::pextreal)" + assume "indicator A x \ (0::extreal)" then have "x \ A" by (auto simp: indicator_def split: split_if_asm) moreover assume "x \ space M" "\y\A. ?f x \ f y" ultimately show "f x = 0" by auto @@ -467,7 +486,8 @@ unfolding simple_function_eq_borel_measurable R.simple_function_eq_borel_measurable unfolding borel_measurable_restricted[OF `A \ sets M`] - by auto + using assms(1)[THEN sets_into_space] + by (auto simp: indicator_def) qed lemma (in sigma_algebra) simple_function_subalgebra: @@ -504,7 +524,7 @@ "integral\<^isup>S M f = (\x \ f ` space M. x * measure M (f -` {x} \ space M))" syntax - "_simple_integral" :: "'a \ ('a \ pextreal) \ ('a, 'b) measure_space_scheme \ pextreal" ("\\<^isup>S _. _ \_" [60,61] 110) + "_simple_integral" :: "'a \ ('a \ extreal) \ ('a, 'b) measure_space_scheme \ extreal" ("\\<^isup>S _. _ \_" [60,61] 110) translations "\\<^isup>S x. f \M" == "CONST integral\<^isup>S M (%x. f)" @@ -540,7 +560,7 @@ qed lemma (in measure_space) simple_function_partition: - assumes "simple_function M f" and "simple_function M g" + assumes f: "simple_function M f" and g: "simple_function M g" shows "integral\<^isup>S M f = (\A\(\x. f -` {f x} \ g -` {g x} \ space M) ` space M. the_elem (f`A) * \ A)" (is "_ = setsum _ (?p ` space M)") proof- @@ -559,23 +579,16 @@ hence "finite (?p ` (A \ space M))" by (rule finite_subset) auto } note this[intro, simp] + note sets = simple_function_measurable2[OF f g] { fix x assume "x \ space M" have "\(?sub (f x)) = (f -` {f x} \ space M)" by auto - moreover { - fix x y - have *: "f -` {f x} \ g -` {g x} \ space M - = (f -` {f x} \ space M) \ (g -` {g x} \ space M)" by auto - assume "x \ space M" "y \ space M" - hence "f -` {f x} \ g -` {g x} \ space M \ sets M" - using assms unfolding simple_function_def * by auto } - ultimately - have "\ (f -` {f x} \ space M) = setsum (\) (?sub (f x))" - by (subst measure_finitely_additive) auto } + with sets have "\ (f -` {f x} \ space M) = setsum \ (?sub (f x))" + by (subst measure_Union) auto } hence "integral\<^isup>S M f = (\(x,A)\?SIGMA. x * \ A)" - unfolding simple_integral_def - by (subst setsum_Sigma[symmetric], - auto intro!: setsum_cong simp: setsum_right_distrib[symmetric]) + unfolding simple_integral_def using f sets + by (subst setsum_Sigma[symmetric]) + (auto intro!: setsum_cong setsum_extreal_right_distrib) also have "\ = (\A\?p ` space M. the_elem (f`A) * \ A)" proof - have [simp]: "\x. x \ space M \ f ` ?p x = {f x}" by (auto intro!: imageI) @@ -595,7 +608,7 @@ qed lemma (in measure_space) simple_integral_add[simp]: - assumes "simple_function M f" and "simple_function M g" + assumes f: "simple_function M f" and "\x. 0 \ f x" and g: "simple_function M g" and "\x. 0 \ g x" shows "(\\<^isup>Sx. f x + g x \M) = integral\<^isup>S M f + integral\<^isup>S M g" proof - { fix x let ?S = "g -` {g x} \ f -` {f x} \ space M" @@ -603,63 +616,43 @@ hence "(\a. f a + g a) ` ?S = {f x + g x}" "f ` ?S = {f x}" "g ` ?S = {g x}" "(\x. (f x, g x)) -` {(f x, g x)} \ space M = ?S" by auto } - thus ?thesis + with assms show ?thesis unfolding - simple_function_partition[OF simple_function_add[OF assms] simple_function_Pair[OF assms]] - simple_function_partition[OF `simple_function M f` `simple_function M g`] - simple_function_partition[OF `simple_function M g` `simple_function M f`] - apply (subst (3) Int_commute) - by (auto simp add: field_simps setsum_addf[symmetric] intro!: setsum_cong) + simple_function_partition[OF simple_function_add[OF f g] simple_function_Pair[OF f g]] + simple_function_partition[OF f g] + simple_function_partition[OF g f] + by (subst (3) Int_commute) + (auto simp add: extreal_left_distrib setsum_addf[symmetric] intro!: setsum_cong) qed lemma (in measure_space) simple_integral_setsum[simp]: + assumes "\i x. i \ P \ 0 \ f i x" assumes "\i. i \ P \ simple_function M (f i)" shows "(\\<^isup>Sx. (\i\P. f i x) \M) = (\i\P. integral\<^isup>S M (f i))" proof cases assume "finite P" from this assms show ?thesis - by induct (auto simp: simple_function_setsum simple_integral_add) + by induct (auto simp: simple_function_setsum simple_integral_add setsum_nonneg) qed auto lemma (in measure_space) simple_integral_mult[simp]: - assumes "simple_function M f" + assumes f: "simple_function M f" "\x. 0 \ f x" shows "(\\<^isup>Sx. c * f x \M) = c * integral\<^isup>S M f" proof - - note mult = simple_function_mult[OF simple_function_const[of c] assms] + note mult = simple_function_mult[OF simple_function_const[of c] f(1)] { fix x let ?S = "f -` {f x} \ (\x. c * f x) -` {c * f x} \ space M" assume "x \ space M" hence "(\x. c * f x) ` ?S = {c * f x}" "f ` ?S = {f x}" by auto } - thus ?thesis - unfolding simple_function_partition[OF mult assms] - simple_function_partition[OF assms mult] - by (auto simp: setsum_right_distrib field_simps intro!: setsum_cong) -qed - -lemma (in sigma_algebra) simple_function_If: - assumes sf: "simple_function M f" "simple_function M g" and A: "A \ sets M" - shows "simple_function M (\x. if x \ A then f x else g x)" (is "simple_function M ?IF") -proof - - def F \ "\x. f -` {x} \ space M" and G \ "\x. g -` {x} \ space M" - show ?thesis unfolding simple_function_def - proof safe - have "?IF ` space M \ f ` space M \ g ` space M" by auto - from finite_subset[OF this] assms - show "finite (?IF ` space M)" unfolding simple_function_def by auto - next - fix x assume "x \ space M" - then have *: "?IF -` {?IF x} \ space M = (if x \ A - then ((F (f x) \ A) \ (G (f x) - (G (f x) \ A))) - else ((F (g x) \ A) \ (G (g x) - (G (g x) \ A))))" - using sets_into_space[OF A] by (auto split: split_if_asm simp: G_def F_def) - have [intro]: "\x. F x \ sets M" "\x. G x \ sets M" - unfolding F_def G_def using sf[THEN simple_functionD(2)] by auto - show "?IF -` {?IF x} \ space M \ sets M" unfolding * using A by auto - qed + with assms show ?thesis + unfolding simple_function_partition[OF mult f(1)] + simple_function_partition[OF f(1) mult] + by (subst setsum_extreal_right_distrib) + (auto intro!: extreal_0_le_mult setsum_cong simp: mult_assoc) qed lemma (in measure_space) simple_integral_mono_AE: - assumes "simple_function M f" and "simple_function M g" + assumes f: "simple_function M f" and g: "simple_function M g" and mono: "AE x. f x \ g x" shows "integral\<^isup>S M f \ integral\<^isup>S M g" proof - @@ -668,14 +661,16 @@ "\x. f -` {f x} \ g -` {g x} \ space M = ?S x" by auto show ?thesis unfolding * - simple_function_partition[OF `simple_function M f` `simple_function M g`] - simple_function_partition[OF `simple_function M g` `simple_function M f`] + simple_function_partition[OF f g] + simple_function_partition[OF g f] proof (safe intro!: setsum_mono) fix x assume "x \ space M" then have *: "f ` ?S x = {f x}" "g ` ?S x = {g x}" by auto show "the_elem (f`?S x) * \ (?S x) \ the_elem (g`?S x) * \ (?S x)" proof (cases "f x \ g x") - case True then show ?thesis using * by (auto intro!: mult_right_mono) + case True then show ?thesis + using * assms(1,2)[THEN simple_functionD(2)] + by (auto intro!: extreal_mult_right_mono) next case False obtain N where N: "{x\space M. \ f x \ g x} \ N" "N \ sets M" "\ N = 0" @@ -685,7 +680,10 @@ by (rule_tac Int) (auto intro!: simple_functionD) ultimately have "\ (?S x) \ \ N" using `N \ sets M` by (auto intro!: measure_mono) - then show ?thesis using `\ N = 0` by auto + moreover have "0 \ \ (?S x)" + using assms(1,2)[THEN simple_functionD(2)] by auto + ultimately have "\ (?S x) = 0" using `\ N = 0` by auto + then show ?thesis by simp qed qed qed @@ -697,7 +695,8 @@ using assms by (intro simple_integral_mono_AE) auto lemma (in measure_space) simple_integral_cong_AE: - assumes "simple_function M f" "simple_function M g" and "AE x. f x = g x" + assumes "simple_function M f" and "simple_function M g" + and "AE x. f x = g x" shows "integral\<^isup>S M f = integral\<^isup>S M g" using assms by (auto simp: eq_iff intro!: simple_integral_mono_AE) @@ -765,7 +764,7 @@ assume "space M = {}" hence "A = {}" using sets_into_space[OF assms] by auto thus ?thesis unfolding simple_integral_def using `space M = {}` by auto next - assume "space M \ {}" hence "(\x. 1) ` space M = {1::pextreal}" by auto + assume "space M \ {}" hence "(\x. 1) ` space M = {1::extreal}" by auto thus ?thesis using simple_integral_indicator[OF assms simple_function_const[of 1]] using sets_into_space[OF assms] @@ -773,13 +772,13 @@ qed lemma (in measure_space) simple_integral_null_set: - assumes "simple_function M u" "N \ null_sets" + assumes "simple_function M u" "\x. 0 \ u x" and "N \ null_sets" shows "(\\<^isup>Sx. u x * indicator N x \M) = 0" proof - - have "AE x. indicator N x = (0 :: pextreal)" + have "AE x. indicator N x = (0 :: extreal)" using `N \ null_sets` by (auto simp: indicator_def intro!: AE_I[of _ N]) then have "(\\<^isup>Sx. u x * indicator N x \M) = (\\<^isup>Sx. 0 \M)" - using assms by (intro simple_integral_cong_AE) auto + using assms apply (intro simple_integral_cong_AE) by auto then show ?thesis by simp qed @@ -813,7 +812,7 @@ by (auto simp: indicator_def split: split_if_asm) then show "f x * \ (f -` {f x} \ A) = f x * \ (?f -` {f x} \ space M)" - unfolding pextreal_mult_cancel_left by auto + unfolding extreal_mult_cancel_left by auto qed lemma (in measure_space) simple_integral_subalgebra: @@ -821,10 +820,6 @@ shows "integral\<^isup>S N = integral\<^isup>S M" unfolding simple_integral_def_raw by simp -lemma measure_preservingD: - "T \ measure_preserving A B \ X \ sets B \ measure A (T -` X \ space A) = measure B X" - unfolding measure_preserving_def by auto - lemma (in measure_space) simple_integral_vimage: assumes T: "sigma_algebra M'" "T \ measure_preserving M M'" and f: "simple_function M' f" @@ -853,196 +848,164 @@ qed qed +lemma (in measure_space) simple_integral_cmult_indicator: + assumes A: "A \ sets M" + shows "(\\<^isup>Sx. c * indicator A x \M) = c * \ A" + using simple_integral_mult[OF simple_function_indicator[OF A]] + unfolding simple_integral_indicator_only[OF A] by simp + +lemma (in measure_space) simple_integral_positive: + assumes f: "simple_function M f" and ae: "AE x. 0 \ f x" + shows "0 \ integral\<^isup>S M f" +proof - + have "integral\<^isup>S M (\x. 0) \ integral\<^isup>S M f" + using simple_integral_mono_AE[OF _ f ae] by auto + then show ?thesis by simp +qed + section "Continuous positive integration" definition positive_integral_def: - "integral\<^isup>P M f = (SUP g : {g. simple_function M g \ g \ f}. integral\<^isup>S M g)" + "integral\<^isup>P M f = (SUP g : {g. simple_function M g \ g \ max 0 \ f}. integral\<^isup>S M g)" syntax - "_positive_integral" :: "'a \ ('a \ pextreal) \ ('a, 'b) measure_space_scheme \ pextreal" ("\\<^isup>+ _. _ \_" [60,61] 110) + "_positive_integral" :: "'a \ ('a \ extreal) \ ('a, 'b) measure_space_scheme \ extreal" ("\\<^isup>+ _. _ \_" [60,61] 110) translations "\\<^isup>+ x. f \M" == "CONST integral\<^isup>P M (%x. f)" -lemma (in measure_space) positive_integral_alt: "integral\<^isup>P M f = - (SUP g : {g. simple_function M g \ g \ f \ \ \ g`space M}. integral\<^isup>S M g)" - (is "_ = ?alt") -proof (rule antisym SUP_leI) - show "integral\<^isup>P M f \ ?alt" unfolding positive_integral_def - proof (safe intro!: SUP_leI) - fix g assume g: "simple_function M g" "g \ f" - let ?G = "g -` {\} \ space M" - show "integral\<^isup>S M g \ - (SUP h : {i. simple_function M i \ i \ f \ \ \ i ` space M}. integral\<^isup>S M h)" - (is "integral\<^isup>S M g \ SUPR ?A _") - proof cases - let ?g = "\x. indicator (space M - ?G) x * g x" - have g': "simple_function M ?g" - using g by (auto intro: simple_functionD) - moreover - assume "\ ?G = 0" - then have "AE x. g x = ?g x" using g - by (intro AE_I[where N="?G"]) - (auto intro: simple_functionD simp: indicator_def) - with g(1) g' have "integral\<^isup>S M g = integral\<^isup>S M ?g" - by (rule simple_integral_cong_AE) - moreover have "?g \ g" by (auto simp: le_fun_def indicator_def) - from this `g \ f` have "?g \ f" by (rule order_trans) - moreover have "\ \ ?g ` space M" - by (auto simp: indicator_def split: split_if_asm) - ultimately show ?thesis by (auto intro!: le_SUPI) - next - assume "\ ?G \ 0" - then have "?G \ {}" by auto - then have "\ \ g`space M" by force - then have "space M \ {}" by auto - have "SUPR ?A (integral\<^isup>S M) = \" - proof (intro SUP_\[THEN iffD2] allI impI) - fix x assume "x < \" - then guess n unfolding less_\_Ex_of_nat .. note n = this - then have "0 < n" by (intro neq0_conv[THEN iffD1] notI) simp - let ?g = "\x. (of_nat n / (if \ ?G = \ then 1 else \ ?G)) * indicator ?G x" - show "\i\?A. x < integral\<^isup>S M i" - proof (intro bexI impI CollectI conjI) - show "simple_function M ?g" using g - by (auto intro!: simple_functionD simple_function_add) - have "?g \ g" by (auto simp: le_fun_def indicator_def) - from this g(2) show "?g \ f" by (rule order_trans) - show "\ \ ?g ` space M" - using `\ ?G \ 0` by (auto simp: indicator_def split: split_if_asm) - have "x < (of_nat n / (if \ ?G = \ then 1 else \ ?G)) * \ ?G" - using n `\ ?G \ 0` `0 < n` - by (auto simp: pextreal_noteq_omega_Ex field_simps) - also have "\ = integral\<^isup>S M ?g" using g `space M \ {}` - by (subst simple_integral_indicator) - (auto simp: image_constant ac_simps dest: simple_functionD) - finally show "x < integral\<^isup>S M ?g" . - qed - qed - then show ?thesis by simp - qed - qed -qed (auto intro!: SUP_subset simp: positive_integral_def) - lemma (in measure_space) positive_integral_cong_measure: assumes "\A. A \ sets M \ measure N A = \ A" "sets N = sets M" "space N = space M" shows "integral\<^isup>P N f = integral\<^isup>P M f" -proof - - interpret v: measure_space N - by (rule measure_space_cong) fact+ - with assms show ?thesis - unfolding positive_integral_def SUPR_def - by (auto intro!: arg_cong[where f=Sup] image_cong - simp: simple_integral_cong_measure[OF assms] - simple_function_cong_algebra[OF assms(2,3)]) -qed + unfolding positive_integral_def + unfolding simple_function_cong_algebra[OF assms(2,3), symmetric] + using AE_cong_measure[OF assms] + using simple_integral_cong_measure[OF assms] + by (auto intro!: SUP_cong) + +lemma (in measure_space) positive_integral_positive: + "0 \ integral\<^isup>P M f" + by (auto intro!: le_SUPI2[of "\x. 0"] simp: positive_integral_def le_fun_def) -lemma (in measure_space) positive_integral_alt1: - "integral\<^isup>P M f = - (SUP g : {g. simple_function M g \ (\x\space M. g x \ f x \ g x \ \)}. integral\<^isup>S M g)" - unfolding positive_integral_alt SUPR_def -proof (safe intro!: arg_cong[where f=Sup]) - fix g let ?g = "\x. if x \ space M then g x else f x" - assume "simple_function M g" "\x\space M. g x \ f x \ g x \ \" - hence "?g \ f" "simple_function M ?g" "integral\<^isup>S M g = integral\<^isup>S M ?g" - "\ \ g`space M" - unfolding le_fun_def by (auto cong: simple_function_cong simple_integral_cong) - thus "integral\<^isup>S M g \ integral\<^isup>S M ` {g. simple_function M g \ g \ f \ \ \ g`space M}" - by auto -next - fix g assume "simple_function M g" "g \ f" "\ \ g`space M" - hence "simple_function M g" "\x\space M. g x \ f x \ g x \ \" - by (auto simp add: le_fun_def image_iff) - thus "integral\<^isup>S M g \ integral\<^isup>S M ` {g. simple_function M g \ (\x\space M. g x \ f x \ g x \ \)}" - by auto -qed +lemma (in measure_space) positive_integral_def_finite: + "integral\<^isup>P M f = (SUP g : {g. simple_function M g \ g \ max 0 \ f \ range g \ {0 ..< \}}. integral\<^isup>S M g)" + (is "_ = SUPR ?A ?f") + unfolding positive_integral_def +proof (safe intro!: antisym SUP_leI) + fix g assume g: "simple_function M g" "g \ max 0 \ f" + let ?G = "{x \ space M. \ g x \ \}" + note gM = g(1)[THEN borel_measurable_simple_function] + have \G_pos: "0 \ \ ?G" using gM by auto + let "?g y x" = "if g x = \ then y else max 0 (g x)" + from g gM have g_in_A: "\y. 0 \ y \ y \ \ \ ?g y \ ?A" + apply (safe intro!: simple_function_max simple_function_If) + apply (force simp: max_def le_fun_def split: split_if_asm)+ + done + show "integral\<^isup>S M g \ SUPR ?A ?f" + proof cases + have g0: "?g 0 \ ?A" by (intro g_in_A) auto + assume "\ ?G = 0" + with gM have "AE x. x \ ?G" by (simp add: AE_iff_null_set) + with gM g show ?thesis + by (intro le_SUPI2[OF g0] simple_integral_mono_AE) + (auto simp: max_def intro!: simple_function_If) + next + assume \G: "\ ?G \ 0" + have "SUPR ?A (integral\<^isup>S M) = \" + proof (intro SUP_PInfty) + fix n :: nat + let ?y = "extreal (real n) / (if \ ?G = \ then 1 else \ ?G)" + have "0 \ ?y" "?y \ \" using \G \G_pos by (auto simp: extreal_divide_eq) + then have "?g ?y \ ?A" by (rule g_in_A) + have "real n \ ?y * \ ?G" + using \G \G_pos by (cases "\ ?G") (auto simp: field_simps) + also have "\ = (\\<^isup>Sx. ?y * indicator ?G x \M)" + using `0 \ ?y` `?g ?y \ ?A` gM + by (subst simple_integral_cmult_indicator) auto + also have "\ \ integral\<^isup>S M (?g ?y)" using `?g ?y \ ?A` gM + by (intro simple_integral_mono) auto + finally show "\i\?A. real n \ integral\<^isup>S M i" + using `?g ?y \ ?A` by blast + qed + then show ?thesis by simp + qed +qed (auto intro: le_SUPI) -lemma (in measure_space) positive_integral_cong: - assumes "\x. x \ space M \ f x = g x" - shows "integral\<^isup>P M f = integral\<^isup>P M g" -proof - - have "\h. (\x\space M. h x \ f x \ h x \ \) = (\x\space M. h x \ g x \ h x \ \)" - using assms by auto - thus ?thesis unfolding positive_integral_alt1 by auto +lemma (in measure_space) positive_integral_mono_AE: + assumes ae: "AE x. u x \ v x" shows "integral\<^isup>P M u \ integral\<^isup>P M v" + unfolding positive_integral_def +proof (safe intro!: SUP_mono) + fix n assume n: "simple_function M n" "n \ max 0 \ u" + from ae[THEN AE_E] guess N . note N = this + then have ae_N: "AE x. x \ N" by (auto intro: AE_not_in) + let "?n x" = "n x * indicator (space M - N) x" + have "AE x. n x \ ?n x" "simple_function M ?n" + using n N ae_N by auto + moreover + { fix x have "?n x \ max 0 (v x)" + proof cases + assume x: "x \ space M - N" + with N have "u x \ v x" by auto + with n(2)[THEN le_funD, of x] x show ?thesis + by (auto simp: max_def split: split_if_asm) + qed simp } + then have "?n \ max 0 \ v" by (auto simp: le_funI) + moreover have "integral\<^isup>S M n \ integral\<^isup>S M ?n" + using ae_N N n by (auto intro!: simple_integral_mono_AE) + ultimately show "\m\{g. simple_function M g \ g \ max 0 \ v}. integral\<^isup>S M n \ integral\<^isup>S M m" + by force qed -lemma (in measure_space) positive_integral_eq_simple_integral: - assumes "simple_function M f" - shows "integral\<^isup>P M f = integral\<^isup>S M f" - unfolding positive_integral_def -proof (safe intro!: pextreal_SUPI) - fix g assume "simple_function M g" "g \ f" - with assms show "integral\<^isup>S M g \ integral\<^isup>S M f" - by (auto intro!: simple_integral_mono simp: le_fun_def) -next - fix y assume "\x. x\{g. simple_function M g \ g \ f} \ integral\<^isup>S M x \ y" - with assms show "integral\<^isup>S M f \ y" by auto -qed - -lemma (in measure_space) positive_integral_mono_AE: - assumes ae: "AE x. u x \ v x" - shows "integral\<^isup>P M u \ integral\<^isup>P M v" - unfolding positive_integral_alt1 -proof (safe intro!: SUPR_mono) - fix a assume a: "simple_function M a" and mono: "\x\space M. a x \ u x \ a x \ \" - from ae obtain N where N: "{x\space M. \ u x \ v x} \ N" "N \ sets M" "\ N = 0" - by (auto elim!: AE_E) - have "simple_function M (\x. a x * indicator (space M - N) x)" - using `N \ sets M` a by auto - with a show "\b\{g. simple_function M g \ (\x\space M. g x \ v x \ g x \ \)}. - integral\<^isup>S M a \ integral\<^isup>S M b" - proof (safe intro!: bexI[of _ "\x. a x * indicator (space M - N) x"] - simple_integral_mono_AE) - show "AE x. a x \ a x * indicator (space M - N) x" - proof (rule AE_I, rule subset_refl) - have *: "{x \ space M. \ a x \ a x * indicator (space M - N) x} = - N \ {x \ space M. a x \ 0}" (is "?N = _") - using `N \ sets M`[THEN sets_into_space] by (auto simp: indicator_def) - then show "?N \ sets M" - using `N \ sets M` `simple_function M a`[THEN borel_measurable_simple_function] - by (auto intro!: measure_mono Int) - then have "\ ?N \ \ N" - unfolding * using `N \ sets M` by (auto intro!: measure_mono) - then show "\ ?N = 0" using `\ N = 0` by auto - qed - next - fix x assume "x \ space M" - show "a x * indicator (space M - N) x \ v x" - proof (cases "x \ N") - case True then show ?thesis by simp - next - case False - with N mono have "a x \ u x" "u x \ v x" using `x \ space M` by auto - with False `x \ space M` show "a x * indicator (space M - N) x \ v x" by auto - qed - assume "a x * indicator (space M - N) x = \" - with mono `x \ space M` show False - by (simp split: split_if_asm add: indicator_def) - qed -qed +lemma (in measure_space) positive_integral_mono: + "(\x. x \ space M \ u x \ v x) \ integral\<^isup>P M u \ integral\<^isup>P M v" + by (auto intro: positive_integral_mono_AE) lemma (in measure_space) positive_integral_cong_AE: "AE x. u x = v x \ integral\<^isup>P M u = integral\<^isup>P M v" by (auto simp: eq_iff intro!: positive_integral_mono_AE) -lemma (in measure_space) positive_integral_mono: - "(\x. x \ space M \ u x \ v x) \ integral\<^isup>P M u \ integral\<^isup>P M v" - by (auto intro: positive_integral_mono_AE) +lemma (in measure_space) positive_integral_cong: + "(\x. x \ space M \ u x = v x) \ integral\<^isup>P M u = integral\<^isup>P M v" + by (auto intro: positive_integral_cong_AE) -lemma image_set_cong: - assumes A: "\x. x \ A \ \y\B. f x = g y" - assumes B: "\y. y \ B \ \x\A. g y = f x" - shows "f ` A = g ` B" - using assms by blast +lemma (in measure_space) positive_integral_eq_simple_integral: + assumes f: "simple_function M f" "\x. 0 \ f x" shows "integral\<^isup>P M f = integral\<^isup>S M f" +proof - + let "?f x" = "f x * indicator (space M) x" + have f': "simple_function M ?f" using f by auto + with f(2) have [simp]: "max 0 \ ?f = ?f" + by (auto simp: fun_eq_iff max_def split: split_indicator) + have "integral\<^isup>P M ?f \ integral\<^isup>S M ?f" using f' + by (force intro!: SUP_leI simple_integral_mono simp: le_fun_def positive_integral_def) + moreover have "integral\<^isup>S M ?f \ integral\<^isup>P M ?f" + unfolding positive_integral_def + using f' by (auto intro!: le_SUPI) + ultimately show ?thesis + by (simp cong: positive_integral_cong simple_integral_cong) +qed + +lemma (in measure_space) positive_integral_eq_simple_integral_AE: + assumes f: "simple_function M f" "AE x. 0 \ f x" shows "integral\<^isup>P M f = integral\<^isup>S M f" +proof - + have "AE x. f x = max 0 (f x)" using f by (auto split: split_max) + with f have "integral\<^isup>P M f = integral\<^isup>S M (\x. max 0 (f x))" + by (simp cong: positive_integral_cong_AE simple_integral_cong_AE + add: positive_integral_eq_simple_integral) + with assms show ?thesis + by (auto intro!: simple_integral_cong_AE split: split_max) +qed lemma (in measure_space) positive_integral_SUP_approx: - assumes "f \ s" - and f: "\i. f i \ borel_measurable M" - and "simple_function M u" - and le: "u \ s" and real: "\ \ u`space M" + assumes f: "incseq f" "\i. f i \ borel_measurable M" "\i x. 0 \ f i x" + and u: "simple_function M u" "u \ (SUP i. f i)" "u`space M \ {0..<\}" shows "integral\<^isup>S M u \ (SUP i. integral\<^isup>P M (f i))" (is "_ \ ?S") -proof (rule pextreal_le_mult_one_interval) - fix a :: pextreal assume "0 < a" "a < 1" +proof (rule extreal_le_mult_one_interval) + have "0 \ (SUP i. integral\<^isup>P M (f i))" + using f(3) by (auto intro!: le_SUPI2 positive_integral_positive) + then show "(SUP i. integral\<^isup>P M (f i)) \ -\" by auto + have u_range: "\x. x \ space M \ 0 \ u x \ u x \ \" + using u(3) by auto + fix a :: extreal assume "0 < a" "a < 1" hence "a \ 0" by auto let "?B i" = "{x \ space M. a * u x \ f i x}" have B: "\i. ?B i \ sets M" @@ -1054,203 +1017,269 @@ proof safe fix i x assume "a * u x \ f i x" also have "\ \ f (Suc i) x" - using `f \ s` unfolding isoton_def le_fun_def by auto + using `incseq f`[THEN incseq_SucD] unfolding le_fun_def by auto finally show "a * u x \ f (Suc i) x" . qed } note B_mono = this - have u: "\x. x \ space M \ u -` {u x} \ space M \ sets M" - using `simple_function M u` by (auto simp add: simple_function_def) + note B_u = Int[OF u(1)[THEN simple_functionD(2)] B] - have "\i. (\n. (u -` {i} \ space M) \ ?B n) \ (u -` {i} \ space M)" using B_mono unfolding isoton_def - proof safe - fix x i assume "x \ space M" - show "x \ (\i. (u -` {u x} \ space M) \ ?B i)" - proof cases - assume "u x = 0" thus ?thesis using `x \ space M` by simp - next - assume "u x \ 0" - with `a < 1` real `x \ space M` - have "a * u x < 1 * u x" by (rule_tac pextreal_mult_strict_right_mono) (auto simp: image_iff) - also have "\ \ (SUP i. f i x)" using le `f \ s` - unfolding isoton_fun_expand by (auto simp: isoton_def le_fun_def) - finally obtain i where "a * u x < f i x" unfolding SUPR_def - by (auto simp add: less_Sup_iff) - hence "a * u x \ f i x" by auto - thus ?thesis using `x \ space M` by auto + let "?B' i n" = "(u -` {i} \ space M) \ ?B n" + have measure_conv: "\i. \ (u -` {i} \ space M) = (SUP n. \ (?B' i n))" + proof - + fix i + have 1: "range (?B' i) \ sets M" using B_u by auto + have 2: "incseq (?B' i)" using B_mono by (auto intro!: incseq_SucI) + have "(\n. ?B' i n) = u -` {i} \ space M" + proof safe + fix x i assume x: "x \ space M" + show "x \ (\i. ?B' (u x) i)" + proof cases + assume "u x = 0" thus ?thesis using `x \ space M` f(3) by simp + next + assume "u x \ 0" + with `a < 1` u_range[OF `x \ space M`] + have "a * u x < 1 * u x" + by (intro extreal_mult_strict_right_mono) (auto simp: image_iff) + also have "\ \ (SUP i. f i x)" using u(2) by (auto simp: le_fun_def SUPR_apply) + finally obtain i where "a * u x < f i x" unfolding SUPR_def + by (auto simp add: less_Sup_iff) + hence "a * u x \ f i x" by auto + thus ?thesis using `x \ space M` by auto + qed qed - qed auto - note measure_conv = measure_up[OF Int[OF u B] this] + then show "?thesis i" using continuity_from_below[OF 1 2] by simp + qed have "integral\<^isup>S M u = (SUP i. integral\<^isup>S M (?uB i))" unfolding simple_integral_indicator[OF B `simple_function M u`] - proof (subst SUPR_pextreal_setsum, safe) + proof (subst SUPR_extreal_setsum, safe) fix x n assume "x \ space M" - have "\ (u -` {u x} \ space M \ {x \ space M. a * u x \ f n x}) - \ \ (u -` {u x} \ space M \ {x \ space M. a * u x \ f (Suc n) x})" - using B_mono Int[OF u[OF `x \ space M`] B] by (auto intro!: measure_mono) - thus "u x * \ (u -` {u x} \ space M \ ?B n) - \ u x * \ (u -` {u x} \ space M \ ?B (Suc n))" - by (auto intro: mult_left_mono) + with u_range show "incseq (\i. u x * \ (?B' (u x) i))" "\i. 0 \ u x * \ (?B' (u x) i)" + using B_mono B_u by (auto intro!: measure_mono extreal_mult_left_mono incseq_SucI simp: extreal_zero_le_0_iff) next - show "integral\<^isup>S M u = - (\i\u ` space M. SUP n. i * \ (u -` {i} \ space M \ ?B n))" - using measure_conv unfolding simple_integral_def isoton_def - by (auto intro!: setsum_cong simp: pextreal_SUP_cmult) + show "integral\<^isup>S M u = (\i\u ` space M. SUP n. i * \ (?B' i n))" + using measure_conv u_range B_u unfolding simple_integral_def + by (auto intro!: setsum_cong SUPR_extreal_cmult[symmetric]) qed moreover have "a * (SUP i. integral\<^isup>S M (?uB i)) \ ?S" - unfolding pextreal_SUP_cmult[symmetric] + apply (subst SUPR_extreal_cmult[symmetric]) proof (safe intro!: SUP_mono bexI) fix i have "a * integral\<^isup>S M (?uB i) = (\\<^isup>Sx. a * ?uB i x \M)" - using B `simple_function M u` - by (subst simple_integral_mult[symmetric]) (auto simp: field_simps) + using B `simple_function M u` u_range + by (subst simple_integral_mult) (auto split: split_indicator) also have "\ \ integral\<^isup>P M (f i)" proof - - have "\x. a * ?uB i x \ f i x" unfolding indicator_def by auto - hence *: "simple_function M (\x. a * ?uB i x)" using B assms(3) - by (auto intro!: simple_integral_mono) - show ?thesis unfolding positive_integral_eq_simple_integral[OF *, symmetric] - by (auto intro!: positive_integral_mono simp: indicator_def) + have *: "simple_function M (\x. a * ?uB i x)" using B `0 < a` u(1) by auto + show ?thesis using f(3) * u_range `0 < a` + by (subst positive_integral_eq_simple_integral[symmetric]) + (auto intro!: positive_integral_mono split: split_indicator) qed finally show "a * integral\<^isup>S M (?uB i) \ integral\<^isup>P M (f i)" by auto - qed simp + next + fix i show "0 \ \\<^isup>S x. ?uB i x \M" using B `0 < a` u(1) u_range + by (intro simple_integral_positive) (auto split: split_indicator) + qed (insert `0 < a`, auto) ultimately show "a * integral\<^isup>S M u \ ?S" by simp qed +lemma (in measure_space) incseq_positive_integral: + assumes "incseq f" shows "incseq (\i. integral\<^isup>P M (f i))" +proof - + have "\i x. f i x \ f (Suc i) x" + using assms by (auto dest!: incseq_SucD simp: le_fun_def) + then show ?thesis + by (auto intro!: incseq_SucI positive_integral_mono) +qed + text {* Beppo-Levi monotone convergence theorem *} -lemma (in measure_space) positive_integral_isoton: - assumes "f \ u" "\i. f i \ borel_measurable M" - shows "(\i. integral\<^isup>P M (f i)) \ integral\<^isup>P M u" - unfolding isoton_def -proof safe - fix i show "integral\<^isup>P M (f i) \ integral\<^isup>P M (f (Suc i))" - apply (rule positive_integral_mono) - using `f \ u` unfolding isoton_def le_fun_def by auto +lemma (in measure_space) positive_integral_monotone_convergence_SUP: + assumes f: "incseq f" "\i. f i \ borel_measurable M" "\i x. 0 \ f i x" + shows "(\\<^isup>+ x. (SUP i. f i x) \M) = (SUP i. integral\<^isup>P M (f i))" +proof (rule antisym) + show "(SUP j. integral\<^isup>P M (f j)) \ (\\<^isup>+ x. (SUP i. f i x) \M)" + by (auto intro!: SUP_leI le_SUPI positive_integral_mono) next - have u: "u = (SUP i. f i)" using `f \ u` unfolding isoton_def by auto - show "(SUP i. integral\<^isup>P M (f i)) = integral\<^isup>P M u" - proof (rule antisym) - from `f \ u`[THEN isoton_Sup, unfolded le_fun_def] - show "(SUP j. integral\<^isup>P M (f j)) \ integral\<^isup>P M u" - by (auto intro!: SUP_leI positive_integral_mono) - next - show "integral\<^isup>P M u \ (SUP i. integral\<^isup>P M (f i))" - unfolding positive_integral_alt[of u] - by (auto intro!: SUP_leI positive_integral_SUP_approx[OF assms]) + show "(\\<^isup>+ x. (SUP i. f i x) \M) \ (SUP j. integral\<^isup>P M (f j))" + unfolding positive_integral_def_finite[of "\x. SUP i. f i x"] + proof (safe intro!: SUP_leI) + fix g assume g: "simple_function M g" + and "g \ max 0 \ (\x. SUP i. f i x)" "range g \ {0..<\}" + moreover then have "\x. 0 \ (SUP i. f i x)" and g': "g`space M \ {0..<\}" + using f by (auto intro!: le_SUPI2) + ultimately show "integral\<^isup>S M g \ (SUP j. integral\<^isup>P M (f j))" + by (intro positive_integral_SUP_approx[OF f g _ g']) + (auto simp: le_fun_def max_def SUPR_apply) qed qed -lemma (in measure_space) positive_integral_monotone_convergence_SUP: - assumes "\i x. x \ space M \ f i x \ f (Suc i) x" - assumes "\i. f i \ borel_measurable M" - shows "(SUP i. integral\<^isup>P M (f i)) = (\\<^isup>+ x. (SUP i. f i x) \M)" - (is "_ = integral\<^isup>P M ?u") +lemma (in measure_space) positive_integral_monotone_convergence_SUP_AE: + assumes f: "\i. AE x. f i x \ f (Suc i) x \ 0 \ f i x" "\i. f i \ borel_measurable M" + shows "(\\<^isup>+ x. (SUP i. f i x) \M) = (SUP i. integral\<^isup>P M (f i))" proof - - show ?thesis - proof (rule antisym) - show "(SUP j. integral\<^isup>P M (f j)) \ integral\<^isup>P M ?u" - by (auto intro!: SUP_leI positive_integral_mono le_SUPI) - next - def rf \ "\i. \x\space M. f i x" and ru \ "\x\space M. ?u x" - have "\i. rf i \ borel_measurable M" unfolding rf_def - using assms by (simp cong: measurable_cong) - moreover have iso: "rf \ ru" using assms unfolding rf_def ru_def - unfolding isoton_def le_fun_def fun_eq_iff SUPR_apply - using SUP_const[OF UNIV_not_empty] - by (auto simp: restrict_def le_fun_def fun_eq_iff) - ultimately have "integral\<^isup>P M ru \ (SUP i. integral\<^isup>P M (rf i))" - unfolding positive_integral_alt[of ru] - by (auto simp: le_fun_def intro!: SUP_leI positive_integral_SUP_approx) - then show "integral\<^isup>P M ?u \ (SUP i. integral\<^isup>P M (f i))" - unfolding ru_def rf_def by (simp cong: positive_integral_cong) + from f have "AE x. \i. f i x \ f (Suc i) x \ 0 \ f i x" + by (simp add: AE_all_countable) + from this[THEN AE_E] guess N . note N = this + let "?f i x" = "if x \ space M - N then f i x else 0" + have f_eq: "AE x. \i. ?f i x = f i x" using N by (auto intro!: AE_I[of _ N]) + then have "(\\<^isup>+ x. (SUP i. f i x) \M) = (\\<^isup>+ x. (SUP i. ?f i x) \M)" + by (auto intro!: positive_integral_cong_AE) + also have "\ = (SUP i. (\\<^isup>+ x. ?f i x \M))" + proof (rule positive_integral_monotone_convergence_SUP) + show "incseq ?f" using N(1) by (force intro!: incseq_SucI le_funI) + { fix i show "(\x. if x \ space M - N then f i x else 0) \ borel_measurable M" + using f N(3) by (intro measurable_If_set) auto + fix x show "0 \ ?f i x" + using N(1) by auto } qed + also have "\ = (SUP i. (\\<^isup>+ x. f i x \M))" + using f_eq by (force intro!: arg_cong[where f="SUPR UNIV"] positive_integral_cong_AE ext) + finally show ?thesis . +qed + +lemma (in measure_space) positive_integral_monotone_convergence_SUP_AE_incseq: + assumes f: "incseq f" "\i. AE x. 0 \ f i x" and borel: "\i. f i \ borel_measurable M" + shows "(\\<^isup>+ x. (SUP i. f i x) \M) = (SUP i. integral\<^isup>P M (f i))" + using f[unfolded incseq_Suc_iff le_fun_def] + by (intro positive_integral_monotone_convergence_SUP_AE[OF _ borel]) + auto + +lemma (in measure_space) positive_integral_monotone_convergence_simple: + assumes f: "incseq f" "\i x. 0 \ f i x" "\i. simple_function M (f i)" + shows "(SUP i. integral\<^isup>S M (f i)) = (\\<^isup>+x. (SUP i. f i x) \M)" + using assms unfolding positive_integral_monotone_convergence_SUP[OF f(1) + f(3)[THEN borel_measurable_simple_function] f(2)] + by (auto intro!: positive_integral_eq_simple_integral[symmetric] arg_cong[where f="SUPR UNIV"] ext) + +lemma positive_integral_max_0: + "(\\<^isup>+x. max 0 (f x) \M) = integral\<^isup>P M f" + by (simp add: le_fun_def positive_integral_def) + +lemma (in measure_space) positive_integral_cong_pos: + assumes "\x. x \ space M \ f x \ 0 \ g x \ 0 \ f x = g x" + shows "integral\<^isup>P M f = integral\<^isup>P M g" +proof - + have "integral\<^isup>P M (\x. max 0 (f x)) = integral\<^isup>P M (\x. max 0 (g x))" + proof (intro positive_integral_cong) + fix x assume "x \ space M" + from assms[OF this] show "max 0 (f x) = max 0 (g x)" + by (auto split: split_max) + qed + then show ?thesis by (simp add: positive_integral_max_0) qed lemma (in measure_space) SUP_simple_integral_sequences: - assumes f: "f \ u" "\i. simple_function M (f i)" - and g: "g \ u" "\i. simple_function M (g i)" + assumes f: "incseq f" "\i x. 0 \ f i x" "\i. simple_function M (f i)" + and g: "incseq g" "\i x. 0 \ g i x" "\i. simple_function M (g i)" + and eq: "AE x. (SUP i. f i x) = (SUP i. g i x)" shows "(SUP i. integral\<^isup>S M (f i)) = (SUP i. integral\<^isup>S M (g i))" (is "SUPR _ ?F = SUPR _ ?G") proof - - have "(SUP i. ?F i) = (SUP i. integral\<^isup>P M (f i))" - using assms by (simp add: positive_integral_eq_simple_integral) - also have "\ = integral\<^isup>P M u" - using positive_integral_isoton[OF f(1) borel_measurable_simple_function[OF f(2)]] - unfolding isoton_def by simp - also have "\ = (SUP i. integral\<^isup>P M (g i))" - using positive_integral_isoton[OF g(1) borel_measurable_simple_function[OF g(2)]] - unfolding isoton_def by simp + have "(SUP i. integral\<^isup>S M (f i)) = (\\<^isup>+x. (SUP i. f i x) \M)" + using f by (rule positive_integral_monotone_convergence_simple) + also have "\ = (\\<^isup>+x. (SUP i. g i x) \M)" + unfolding eq[THEN positive_integral_cong_AE] .. also have "\ = (SUP i. ?G i)" - using assms by (simp add: positive_integral_eq_simple_integral) - finally show ?thesis . + using g by (rule positive_integral_monotone_convergence_simple[symmetric]) + finally show ?thesis by simp qed lemma (in measure_space) positive_integral_const[simp]: - "(\\<^isup>+ x. c \M) = c * \ (space M)" + "0 \ c \ (\\<^isup>+ x. c \M) = c * \ (space M)" by (subst positive_integral_eq_simple_integral) auto -lemma (in measure_space) positive_integral_isoton_simple: - assumes "f \ u" and e: "\i. simple_function M (f i)" - shows "(\i. integral\<^isup>S M (f i)) \ integral\<^isup>P M u" - using positive_integral_isoton[OF `f \ u` e(1)[THEN borel_measurable_simple_function]] - unfolding positive_integral_eq_simple_integral[OF e] . - -lemma measure_preservingD2: - "f \ measure_preserving A B \ f \ measurable A B" - unfolding measure_preserving_def by auto - lemma (in measure_space) positive_integral_vimage: - assumes T: "sigma_algebra M'" "T \ measure_preserving M M'" and f: "f \ borel_measurable M'" + assumes T: "sigma_algebra M'" "T \ measure_preserving M M'" + and f: "f \ borel_measurable M'" shows "integral\<^isup>P M' f = (\\<^isup>+ x. f (T x) \M)" proof - interpret T: measure_space M' by (rule measure_space_vimage[OF T]) - obtain f' where f': "f' \ f" "\i. simple_function M' (f' i)" - using T.borel_measurable_implies_simple_function_sequence[OF f] by blast - then have f: "(\i x. f' i (T x)) \ (\x. f (T x))" "\i. simple_function M (\x. f' i (T x))" - using simple_function_vimage[OF T(1) measure_preservingD2[OF T(2)]] unfolding isoton_fun_expand by auto + from T.borel_measurable_implies_simple_function_sequence'[OF f] + guess f' . note f' = this + let "?f i x" = "f' i (T x)" + have inc: "incseq ?f" using f' by (force simp: le_fun_def incseq_def) + have sup: "\x. (SUP i. ?f i x) = max 0 (f (T x))" + using f'(4) . + have sf: "\i. simple_function M (\x. f' i (T x))" + using simple_function_vimage[OF T(1) measure_preservingD2[OF T(2)] f'(1)] . show "integral\<^isup>P M' f = (\\<^isup>+ x. f (T x) \M)" - using positive_integral_isoton_simple[OF f] - using T.positive_integral_isoton_simple[OF f'] - by (simp add: simple_integral_vimage[OF T f'(2)] isoton_def) + using + T.positive_integral_monotone_convergence_simple[OF f'(2,5,1)] + positive_integral_monotone_convergence_simple[OF inc f'(5) sf] + by (simp add: positive_integral_max_0 simple_integral_vimage[OF T f'(1)] f') qed lemma (in measure_space) positive_integral_linear: - assumes f: "f \ borel_measurable M" - and g: "g \ borel_measurable M" + assumes f: "f \ borel_measurable M" "\x. 0 \ f x" and "0 \ a" + and g: "g \ borel_measurable M" "\x. 0 \ g x" shows "(\\<^isup>+ x. a * f x + g x \M) = a * integral\<^isup>P M f + integral\<^isup>P M g" (is "integral\<^isup>P M ?L = _") proof - - from borel_measurable_implies_simple_function_sequence'[OF f] guess u . - note u = this positive_integral_isoton_simple[OF this(1-2)] - from borel_measurable_implies_simple_function_sequence'[OF g] guess v . - note v = this positive_integral_isoton_simple[OF this(1-2)] + from borel_measurable_implies_simple_function_sequence'[OF f(1)] guess u . + note u = positive_integral_monotone_convergence_simple[OF this(2,5,1)] this + from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess v . + note v = positive_integral_monotone_convergence_simple[OF this(2,5,1)] this let "?L' i x" = "a * u i x + v i x" - have "?L \ borel_measurable M" - using assms by simp + have "?L \ borel_measurable M" using assms by auto from borel_measurable_implies_simple_function_sequence'[OF this] guess l . - note positive_integral_isoton_simple[OF this(1-2)] and l = this - moreover have "(SUP i. integral\<^isup>S M (l i)) = (SUP i. integral\<^isup>S M (?L' i))" - proof (rule SUP_simple_integral_sequences[OF l(1-2)]) - show "?L' \ ?L" "\i. simple_function M (?L' i)" - using u v by (auto simp: isoton_fun_expand isoton_add isoton_cmult_right) + note l = positive_integral_monotone_convergence_simple[OF this(2,5,1)] this + + have inc: "incseq (\i. a * integral\<^isup>S M (u i))" "incseq (\i. integral\<^isup>S M (v i))" + using u v `0 \ a` + by (auto simp: incseq_Suc_iff le_fun_def + intro!: add_mono extreal_mult_left_mono simple_integral_mono) + have pos: "\i. 0 \ integral\<^isup>S M (u i)" "\i. 0 \ integral\<^isup>S M (v i)" "\i. 0 \ a * integral\<^isup>S M (u i)" + using u v `0 \ a` by (auto simp: simple_integral_positive) + { fix i from pos[of i] have "a * integral\<^isup>S M (u i) \ -\" "integral\<^isup>S M (v i) \ -\" + by (auto split: split_if_asm) } + note not_MInf = this + + have l': "(SUP i. integral\<^isup>S M (l i)) = (SUP i. integral\<^isup>S M (?L' i))" + proof (rule SUP_simple_integral_sequences[OF l(3,6,2)]) + show "incseq ?L'" "\i x. 0 \ ?L' i x" "\i. simple_function M (?L' i)" + using u v `0 \ a` unfolding incseq_Suc_iff le_fun_def + by (auto intro!: add_mono extreal_mult_left_mono extreal_add_nonneg_nonneg) + { fix x + { fix i have "a * u i x \ -\" "v i x \ -\" "u i x \ -\" using `0 \ a` u(6)[of i x] v(6)[of i x] + by auto } + then have "(SUP i. a * u i x + v i x) = a * (SUP i. u i x) + (SUP i. v i x)" + using `0 \ a` u(3) v(3) u(6)[of _ x] v(6)[of _ x] + by (subst SUPR_extreal_cmult[symmetric, OF u(6) `0 \ a`]) + (auto intro!: SUPR_extreal_add + simp: incseq_Suc_iff le_fun_def add_mono extreal_mult_left_mono extreal_add_nonneg_nonneg) } + then show "AE x. (SUP i. l i x) = (SUP i. ?L' i x)" + unfolding l(5) using `0 \ a` u(5) v(5) l(5) f(2) g(2) + by (intro AE_I2) (auto split: split_max simp add: extreal_add_nonneg_nonneg) qed - moreover from u v have L'_isoton: - "(\i. integral\<^isup>S M (?L' i)) \ a * integral\<^isup>P M f + integral\<^isup>P M g" - by (simp add: isoton_add isoton_cmult_right) - ultimately show ?thesis by (simp add: isoton_def) + also have "\ = (SUP i. a * integral\<^isup>S M (u i) + integral\<^isup>S M (v i))" + using u(2, 6) v(2, 6) `0 \ a` by (auto intro!: arg_cong[where f="SUPR UNIV"] ext) + finally have "(\\<^isup>+ x. max 0 (a * f x + g x) \M) = a * (\\<^isup>+x. max 0 (f x) \M) + (\\<^isup>+x. max 0 (g x) \M)" + unfolding l(5)[symmetric] u(5)[symmetric] v(5)[symmetric] + unfolding l(1)[symmetric] u(1)[symmetric] v(1)[symmetric] + apply (subst SUPR_extreal_cmult[symmetric, OF pos(1) `0 \ a`]) + apply (subst SUPR_extreal_add[symmetric, OF inc not_MInf]) . + then show ?thesis by (simp add: positive_integral_max_0) qed lemma (in measure_space) positive_integral_cmult: - assumes "f \ borel_measurable M" + assumes f: "f \ borel_measurable M" "AE x. 0 \ f x" "0 \ c" shows "(\\<^isup>+ x. c * f x \M) = c * integral\<^isup>P M f" - using positive_integral_linear[OF assms, of "\x. 0" c] by auto +proof - + have [simp]: "\x. c * max 0 (f x) = max 0 (c * f x)" using `0 \ c` + by (auto split: split_max simp: extreal_zero_le_0_iff) + have "(\\<^isup>+ x. c * f x \M) = (\\<^isup>+ x. c * max 0 (f x) \M)" + by (simp add: positive_integral_max_0) + then show ?thesis + using positive_integral_linear[OF _ _ `0 \ c`, of "\x. max 0 (f x)" "\x. 0"] f + by (auto simp: positive_integral_max_0) +qed lemma (in measure_space) positive_integral_multc: - assumes "f \ borel_measurable M" + assumes "f \ borel_measurable M" "AE x. 0 \ f x" "0 \ c" shows "(\\<^isup>+ x. f x * c \M) = integral\<^isup>P M f * c" unfolding mult_commute[of _ c] positive_integral_cmult[OF assms] by simp @@ -1260,143 +1289,172 @@ (auto simp: simple_function_indicator simple_integral_indicator) lemma (in measure_space) positive_integral_cmult_indicator: - "A \ sets M \ (\\<^isup>+ x. c * indicator A x \M) = c * \ A" + "0 \ c \ A \ sets M \ (\\<^isup>+ x. c * indicator A x \M) = c * \ A" by (subst positive_integral_eq_simple_integral) (auto simp: simple_function_indicator simple_integral_indicator) lemma (in measure_space) positive_integral_add: - assumes "f \ borel_measurable M" "g \ borel_measurable M" + assumes f: "f \ borel_measurable M" "AE x. 0 \ f x" + and g: "g \ borel_measurable M" "AE x. 0 \ g x" shows "(\\<^isup>+ x. f x + g x \M) = integral\<^isup>P M f + integral\<^isup>P M g" - using positive_integral_linear[OF assms, of 1] by simp +proof - + have ae: "AE x. max 0 (f x) + max 0 (g x) = max 0 (f x + g x)" + using assms by (auto split: split_max simp: extreal_add_nonneg_nonneg) + have "(\\<^isup>+ x. f x + g x \M) = (\\<^isup>+ x. max 0 (f x + g x) \M)" + by (simp add: positive_integral_max_0) + also have "\ = (\\<^isup>+ x. max 0 (f x) + max 0 (g x) \M)" + unfolding ae[THEN positive_integral_cong_AE] .. + also have "\ = (\\<^isup>+ x. max 0 (f x) \M) + (\\<^isup>+ x. max 0 (g x) \M)" + using positive_integral_linear[of "\x. max 0 (f x)" 1 "\x. max 0 (g x)"] f g + by auto + finally show ?thesis + by (simp add: positive_integral_max_0) +qed lemma (in measure_space) positive_integral_setsum: - assumes "\i. i\P \ f i \ borel_measurable M" + assumes "\i. i\P \ f i \ borel_measurable M" "\i. i\P \ AE x. 0 \ f i x" shows "(\\<^isup>+ x. (\i\P. f i x) \M) = (\i\P. integral\<^isup>P M (f i))" proof cases - assume "finite P" - from this assms show ?thesis + assume f: "finite P" + from assms have "AE x. \i\P. 0 \ f i x" unfolding AE_finite_all[OF f] by auto + from f this assms(1) show ?thesis proof induct case (insert i P) - have "f i \ borel_measurable M" - "(\x. \i\P. f i x) \ borel_measurable M" - using insert by (auto intro!: borel_measurable_pextreal_setsum) + then have "f i \ borel_measurable M" "AE x. 0 \ f i x" + "(\x. \i\P. f i x) \ borel_measurable M" "AE x. 0 \ (\i\P. f i x)" + by (auto intro!: borel_measurable_extreal_setsum setsum_nonneg) from positive_integral_add[OF this] show ?case using insert by auto qed simp qed simp +lemma (in measure_space) positive_integral_Markov_inequality: + assumes u: "u \ borel_measurable M" "AE x. 0 \ u x" and "A \ sets M" and c: "0 \ c" "c \ \" + shows "\ ({x\space M. 1 \ c * u x} \ A) \ c * (\\<^isup>+ x. u x * indicator A x \M)" + (is "\ ?A \ _ * ?PI") +proof - + have "?A \ sets M" + using `A \ sets M` u by auto + hence "\ ?A = (\\<^isup>+ x. indicator ?A x \M)" + using positive_integral_indicator by simp + also have "\ \ (\\<^isup>+ x. c * (u x * indicator A x) \M)" using u c + by (auto intro!: positive_integral_mono_AE + simp: indicator_def extreal_zero_le_0_iff) + also have "\ = c * (\\<^isup>+ x. u x * indicator A x \M)" + using assms + by (auto intro!: positive_integral_cmult borel_measurable_indicator simp: extreal_zero_le_0_iff) + finally show ?thesis . +qed + +lemma (in measure_space) positive_integral_noteq_infinite: + assumes g: "g \ borel_measurable M" "AE x. 0 \ g x" + and "integral\<^isup>P M g \ \" + shows "AE x. g x \ \" +proof (rule ccontr) + assume c: "\ (AE x. g x \ \)" + have "\ {x\space M. g x = \} \ 0" + using c g by (simp add: AE_iff_null_set) + moreover have "0 \ \ {x\space M. g x = \}" using g by (auto intro: measurable_sets) + ultimately have "0 < \ {x\space M. g x = \}" by auto + then have "\ = \ * \ {x\space M. g x = \}" by auto + also have "\ \ (\\<^isup>+x. \ * indicator {x\space M. g x = \} x \M)" + using g by (subst positive_integral_cmult_indicator) auto + also have "\ \ integral\<^isup>P M g" + using assms by (auto intro!: positive_integral_mono_AE simp: indicator_def) + finally show False using `integral\<^isup>P M g \ \` by auto +qed + lemma (in measure_space) positive_integral_diff: - assumes f: "f \ borel_measurable M" and g: "g \ borel_measurable M" - and fin: "integral\<^isup>P M g \ \" - and mono: "\x. x \ space M \ g x \ f x" + assumes f: "f \ borel_measurable M" + and g: "g \ borel_measurable M" "AE x. 0 \ g x" + and fin: "integral\<^isup>P M g \ \" + and mono: "AE x. g x \ f x" shows "(\\<^isup>+ x. f x - g x \M) = integral\<^isup>P M f - integral\<^isup>P M g" proof - - have borel: "(\x. f x - g x) \ borel_measurable M" - using f g by (rule borel_measurable_pextreal_diff) - have "(\\<^isup>+x. f x - g x \M) + integral\<^isup>P M g = integral\<^isup>P M f" - unfolding positive_integral_add[OF borel g, symmetric] - proof (rule positive_integral_cong) - fix x assume "x \ space M" - from mono[OF this] show "f x - g x + g x = f x" - by (cases "f x", cases "g x", simp, simp, cases "g x", auto) - qed - with mono show ?thesis - by (subst minus_pextreal_eq2[OF _ fin]) (auto intro!: positive_integral_mono) + have diff: "(\x. f x - g x) \ borel_measurable M" "AE x. 0 \ f x - g x" + using assms by (auto intro: extreal_diff_positive) + have pos_f: "AE x. 0 \ f x" using mono g by auto + { fix a b :: extreal assume "0 \ a" "a \ \" "0 \ b" "a \ b" then have "b - a + a = b" + by (cases rule: extreal2_cases[of a b]) auto } + note * = this + then have "AE x. f x = f x - g x + g x" + using mono positive_integral_noteq_infinite[OF g fin] assms by auto + then have **: "integral\<^isup>P M f = (\\<^isup>+x. f x - g x \M) + integral\<^isup>P M g" + unfolding positive_integral_add[OF diff g, symmetric] + by (rule positive_integral_cong_AE) + show ?thesis unfolding ** + using fin positive_integral_positive[of g] + by (cases rule: extreal2_cases[of "\\<^isup>+ x. f x - g x \M" "integral\<^isup>P M g"]) auto qed -lemma (in measure_space) positive_integral_psuminf: - assumes "\i. f i \ borel_measurable M" - shows "(\\<^isup>+ x. (\\<^isub>\ i. f i x) \M) = (\\<^isub>\ i. integral\<^isup>P M (f i))" +lemma (in measure_space) positive_integral_suminf: + assumes f: "\i. f i \ borel_measurable M" "\i. AE x. 0 \ f i x" + shows "(\\<^isup>+ x. (\i. f i x) \M) = (\i. integral\<^isup>P M (f i))" proof - - have "(\i. (\\<^isup>+x. (\iM)) \ (\\<^isup>+x. (\\<^isub>\i. f i x) \M)" - by (rule positive_integral_isoton) - (auto intro!: borel_measurable_pextreal_setsum assms positive_integral_mono - arg_cong[where f=Sup] - simp: isoton_def le_fun_def psuminf_def fun_eq_iff SUPR_def Sup_fun_def) - thus ?thesis - by (auto simp: isoton_def psuminf_def positive_integral_setsum[OF assms]) + have all_pos: "AE x. \i. 0 \ f i x" + using assms by (auto simp: AE_all_countable) + have "(\i. integral\<^isup>P M (f i)) = (SUP n. \iP M (f i))" + using positive_integral_positive by (rule suminf_extreal_eq_SUPR) + also have "\ = (SUP n. \\<^isup>+x. (\iM)" + unfolding positive_integral_setsum[OF f] .. + also have "\ = \\<^isup>+x. (SUP n. \iM" using f all_pos + by (intro positive_integral_monotone_convergence_SUP_AE[symmetric]) + (elim AE_mp, auto simp: setsum_nonneg simp del: setsum_lessThan_Suc intro!: AE_I2 setsum_mono3) + also have "\ = \\<^isup>+x. (\i. f i x) \M" using all_pos + by (intro positive_integral_cong_AE) (auto simp: suminf_extreal_eq_SUPR) + finally show ?thesis by simp qed text {* Fatou's lemma: convergence theorem on limes inferior *} lemma (in measure_space) positive_integral_lim_INF: - fixes u :: "nat \ 'a \ pextreal" - assumes "\i. u i \ borel_measurable M" - shows "(\\<^isup>+ x. (SUP n. INF m. u (m + n) x) \M) \ - (SUP n. INF m. integral\<^isup>P M (u (m + n)))" + fixes u :: "nat \ 'a \ extreal" + assumes u: "\i. u i \ borel_measurable M" "\i. AE x. 0 \ u i x" + shows "(\\<^isup>+ x. liminf (\n. u n x) \M) \ liminf (\n. integral\<^isup>P M (u n))" proof - - have "(\\<^isup>+x. (SUP n. INF m. u (m + n) x) \M) - = (SUP n. (\\<^isup>+x. (INF m. u (m + n) x) \M))" - using assms - by (intro positive_integral_monotone_convergence_SUP[symmetric] INF_mono) - (auto simp del: add_Suc simp add: add_Suc[symmetric]) - also have "\ \ (SUP n. INF m. integral\<^isup>P M (u (m + n)))" - by (auto intro!: SUP_mono bexI le_INFI positive_integral_mono INF_leI) + have pos: "AE x. \i. 0 \ u i x" using u by (auto simp: AE_all_countable) + have "(\\<^isup>+ x. liminf (\n. u n x) \M) = + (SUP n. \\<^isup>+ x. (INF i:{n..}. u i x) \M)" + unfolding liminf_SUPR_INFI using pos u + by (intro positive_integral_monotone_convergence_SUP_AE) + (elim AE_mp, auto intro!: AE_I2 intro: le_INFI INF_subset) + also have "\ \ liminf (\n. integral\<^isup>P M (u n))" + unfolding liminf_SUPR_INFI + by (auto intro!: SUP_mono exI le_INFI positive_integral_mono INF_leI) finally show ?thesis . qed lemma (in measure_space) measure_space_density: - assumes borel: "u \ borel_measurable M" + assumes u: "u \ borel_measurable M" "AE x. 0 \ u x" and M'[simp]: "M' = (M\measure := \A. (\\<^isup>+ x. u x * indicator A x \M)\)" shows "measure_space M'" proof - interpret M': sigma_algebra M' by (intro sigma_algebra_cong) auto show ?thesis proof - show "measure M' {} = 0" unfolding M' by simp + have pos: "\A. AE x. 0 \ u x * indicator A x" + using u by (auto simp: extreal_zero_le_0_iff) + then show "positive M' (measure M')" unfolding M' + using u(1) by (auto simp: positive_def intro!: positive_integral_positive) show "countably_additive M' (measure M')" proof (intro countably_additiveI) fix A :: "nat \ 'a set" assume "range A \ sets M'" - then have "\i. (\x. u x * indicator (A i) x) \ borel_measurable M" - using borel by (auto intro: borel_measurable_indicator) - moreover assume "disjoint_family A" - note psuminf_indicator[OF this] - ultimately show "(\\<^isub>\n. measure M' (A n)) = measure M' (\x. A x)" - by (simp add: positive_integral_psuminf[symmetric]) + then have *: "\i. (\x. u x * indicator (A i) x) \ borel_measurable M" + using u by (auto intro: borel_measurable_indicator) + assume disj: "disjoint_family A" + have "(\n. measure M' (A n)) = (\\<^isup>+ x. (\n. u x * indicator (A n) x) \M)" + unfolding M' using u(1) * + by (simp add: positive_integral_suminf[OF _ pos, symmetric]) + also have "\ = (\\<^isup>+ x. u x * (\n. indicator (A n) x) \M)" using u + by (intro positive_integral_cong_AE) + (elim AE_mp, auto intro!: AE_I2 suminf_cmult_extreal) + also have "\ = (\\<^isup>+ x. u x * indicator (\n. A n) x \M)" + unfolding suminf_indicator[OF disj] .. + finally show "(\n. measure M' (A n)) = measure M' (\x. A x)" + unfolding M' by simp qed qed qed -lemma (in measure_space) positive_integral_translated_density: - assumes "f \ borel_measurable M" "g \ borel_measurable M" - and M': "M' = (M\ measure := \A. (\\<^isup>+ x. f x * indicator A x \M)\)" - shows "integral\<^isup>P M' g = (\\<^isup>+ x. f x * g x \M)" -proof - - from measure_space_density[OF assms(1) M'] - interpret T: measure_space M' . - have borel[simp]: - "borel_measurable M' = borel_measurable M" - "simple_function M' = simple_function M" - unfolding measurable_def simple_function_def_raw by (auto simp: M') - from borel_measurable_implies_simple_function_sequence[OF assms(2)] - obtain G where G: "\i. simple_function M (G i)" "G \ g" by blast - note G_borel = borel_measurable_simple_function[OF this(1)] - from T.positive_integral_isoton[unfolded borel, OF `G \ g` G_borel] - have *: "(\i. integral\<^isup>P M' (G i)) \ integral\<^isup>P M' g" . - { fix i - have [simp]: "finite (G i ` space M)" - using G(1) unfolding simple_function_def by auto - have "integral\<^isup>P M' (G i) = integral\<^isup>S M' (G i)" - using G T.positive_integral_eq_simple_integral by simp - also have "\ = (\\<^isup>+x. f x * (\y\G i`space M. y * indicator (G i -` {y} \ space M) x) \M)" - apply (simp add: simple_integral_def M') - apply (subst positive_integral_cmult[symmetric]) - using G_borel assms(1) apply (fastsimp intro: borel_measurable_vimage) - apply (subst positive_integral_setsum[symmetric]) - using G_borel assms(1) apply (fastsimp intro: borel_measurable_vimage) - by (simp add: setsum_right_distrib field_simps) - also have "\ = (\\<^isup>+x. f x * G i x \M)" - by (auto intro!: positive_integral_cong - simp: indicator_def if_distrib setsum_cases) - finally have "integral\<^isup>P M' (G i) = (\\<^isup>+x. f x * G i x \M)" . } - with * have eq_Tg: "(\i. (\\<^isup>+x. f x * G i x \M)) \ integral\<^isup>P M' g" by simp - from G(2) have "(\i x. f x * G i x) \ (\x. f x * g x)" - unfolding isoton_fun_expand by (auto intro!: isoton_cmult_right) - then have "(\i. (\\<^isup>+x. f x * G i x \M)) \ (\\<^isup>+x. f x * g x \M)" - using assms(1) G_borel by (auto intro!: positive_integral_isoton borel_measurable_pextreal_times) - with eq_Tg show "integral\<^isup>P M' g = (\\<^isup>+x. f x * g x \M)" - unfolding isoton_def by simp -qed - lemma (in measure_space) positive_integral_null_set: assumes "N \ null_sets" shows "(\\<^isup>+ x. u x * indicator N x \M) = 0" proof - @@ -1410,144 +1468,199 @@ then show ?thesis by simp qed -lemma (in measure_space) positive_integral_Markov_inequality: - assumes borel: "u \ borel_measurable M" and "A \ sets M" and c: "c \ \" - shows "\ ({x\space M. 1 \ c * u x} \ A) \ c * (\\<^isup>+ x. u x * indicator A x \M)" - (is "\ ?A \ _ * ?PI") +lemma (in measure_space) positive_integral_translated_density: + assumes f: "f \ borel_measurable M" "AE x. 0 \ f x" + assumes g: "g \ borel_measurable M" "AE x. 0 \ g x" + and M': "M' = (M\ measure := \A. (\\<^isup>+ x. f x * indicator A x \M)\)" + shows "integral\<^isup>P M' g = (\\<^isup>+ x. f x * g x \M)" proof - - have "?A \ sets M" - using `A \ sets M` borel by auto - hence "\ ?A = (\\<^isup>+ x. indicator ?A x \M)" - using positive_integral_indicator by simp - also have "\ \ (\\<^isup>+ x. c * (u x * indicator A x) \M)" - proof (rule positive_integral_mono) - fix x assume "x \ space M" - show "indicator ?A x \ c * (u x * indicator A x)" - by (cases "x \ ?A") auto - qed - also have "\ = c * (\\<^isup>+ x. u x * indicator A x \M)" - using assms - by (auto intro!: positive_integral_cmult borel_measurable_indicator) - finally show ?thesis . + from measure_space_density[OF f M'] + interpret T: measure_space M' . + have borel[simp]: + "borel_measurable M' = borel_measurable M" + "simple_function M' = simple_function M" + unfolding measurable_def simple_function_def_raw by (auto simp: M') + from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess G . note G = this + note G' = borel_measurable_simple_function[OF this(1)] simple_functionD[OF G(1)] + note G'(2)[simp] + { fix P have "AE x. P x \ AE x in M'. P x" + using positive_integral_null_set[of _ f] + unfolding T.almost_everywhere_def almost_everywhere_def + by (auto simp: M') } + note ac = this + from G(4) g(2) have G_M': "AE x in M'. (SUP i. G i x) = g x" + by (auto intro!: ac split: split_max) + { fix i + let "?I y x" = "indicator (G i -` {y} \ space M) x" + { fix x assume *: "x \ space M" "0 \ f x" "0 \ g x" + then have [simp]: "G i ` space M \ {y. G i x = y \ x \ space M} = {G i x}" by auto + from * G' G have "(\y\G i`space M. y * (f x * ?I y x)) = f x * (\y\G i`space M. (y * ?I y x))" + by (subst setsum_extreal_right_distrib) (auto simp: ac_simps) + also have "\ = f x * G i x" + by (simp add: indicator_def if_distrib setsum_cases) + finally have "(\y\G i`space M. y * (f x * ?I y x)) = f x * G i x" . } + note to_singleton = this + have "integral\<^isup>P M' (G i) = integral\<^isup>S M' (G i)" + using G T.positive_integral_eq_simple_integral by simp + also have "\ = (\y\G i`space M. y * (\\<^isup>+x. f x * ?I y x \M))" + unfolding simple_integral_def M' by simp + also have "\ = (\y\G i`space M. (\\<^isup>+x. y * (f x * ?I y x) \M))" + using f G' G by (auto intro!: setsum_cong positive_integral_cmult[symmetric]) + also have "\ = (\\<^isup>+x. (\y\G i`space M. y * (f x * ?I y x)) \M)" + using f G' G by (auto intro!: positive_integral_setsum[symmetric]) + finally have "integral\<^isup>P M' (G i) = (\\<^isup>+x. f x * G i x \M)" + using f g G' to_singleton by (auto intro!: positive_integral_cong_AE) } + note [simp] = this + have "integral\<^isup>P M' g = (SUP i. integral\<^isup>P M' (G i))" using G'(1) G_M'(1) G + using T.positive_integral_monotone_convergence_SUP[symmetric, OF `incseq G`] + by (simp cong: T.positive_integral_cong_AE) + also have "\ = (SUP i. (\\<^isup>+x. f x * G i x \M))" by simp + also have "\ = (\\<^isup>+x. (SUP i. f x * G i x) \M)" + using f G' G(2)[THEN incseq_SucD] G + by (intro positive_integral_monotone_convergence_SUP_AE[symmetric]) + (auto simp: extreal_mult_left_mono le_fun_def extreal_zero_le_0_iff) + also have "\ = (\\<^isup>+x. f x * g x \M)" using f G' G g + by (intro positive_integral_cong_AE) + (auto simp add: SUPR_extreal_cmult split: split_max) + finally show "integral\<^isup>P M' g = (\\<^isup>+x. f x * g x \M)" . qed lemma (in measure_space) positive_integral_0_iff: - assumes borel: "u \ borel_measurable M" + assumes u: "u \ borel_measurable M" and pos: "AE x. 0 \ u x" shows "integral\<^isup>P M u = 0 \ \ {x\space M. u x \ 0} = 0" (is "_ \ \ ?A = 0") proof - - have A: "?A \ sets M" using borel by auto - have u: "(\\<^isup>+ x. u x * indicator ?A x \M) = integral\<^isup>P M u" + have u_eq: "(\\<^isup>+ x. u x * indicator ?A x \M) = integral\<^isup>P M u" by (auto intro!: positive_integral_cong simp: indicator_def) - show ?thesis proof assume "\ ?A = 0" - hence "?A \ null_sets" using `?A \ sets M` by auto - from positive_integral_null_set[OF this] - have "0 = (\\<^isup>+ x. u x * indicator ?A x \M)" by simp - thus "integral\<^isup>P M u = 0" unfolding u by simp + with positive_integral_null_set[of ?A u] u + show "integral\<^isup>P M u = 0" by (simp add: u_eq) next + { fix r :: extreal and n :: nat assume gt_1: "1 \ real n * r" + then have "0 < real n * r" by (cases r) (auto split: split_if_asm simp: one_extreal_def) + then have "0 \ r" by (auto simp add: extreal_zero_less_0_iff) } + note gt_1 = this assume *: "integral\<^isup>P M u = 0" - let "?M n" = "{x \ space M. 1 \ of_nat n * u x}" + let "?M n" = "{x \ space M. 1 \ real (n::nat) * u x}" have "0 = (SUP n. \ (?M n \ ?A))" proof - - { fix n - from positive_integral_Markov_inequality[OF borel `?A \ sets M`, of "of_nat n"] - have "\ (?M n \ ?A) = 0" unfolding * u by simp } + { fix n :: nat + from positive_integral_Markov_inequality[OF u pos, of ?A "extreal (real n)"] + have "\ (?M n \ ?A) \ 0" unfolding u_eq * using u by simp + moreover have "0 \ \ (?M n \ ?A)" using u by auto + ultimately have "\ (?M n \ ?A) = 0" by auto } thus ?thesis by simp qed also have "\ = \ (\n. ?M n \ ?A)" proof (safe intro!: continuity_from_below) fix n show "?M n \ ?A \ sets M" - using borel by (auto intro!: Int) + using u by (auto intro!: Int) next - fix n x assume "1 \ of_nat n * u x" - also have "\ \ of_nat (Suc n) * u x" - by (subst (1 2) mult_commute) (auto intro!: pextreal_mult_cancel) - finally show "1 \ of_nat (Suc n) * u x" . + show "incseq (\n. {x \ space M. 1 \ real n * u x} \ {x \ space M. u x \ 0})" + proof (safe intro!: incseq_SucI) + fix n :: nat and x + assume *: "1 \ real n * u x" + also from gt_1[OF this] have "real n * u x \ real (Suc n) * u x" + using `0 \ u x` by (auto intro!: extreal_mult_right_mono) + finally show "1 \ real (Suc n) * u x" by auto + qed qed - also have "\ = \ ?A" - proof (safe intro!: arg_cong[where f="\"]) - fix x assume "u x \ 0" and [simp, intro]: "x \ space M" + also have "\ = \ {x\space M. 0 < u x}" + proof (safe intro!: arg_cong[where f="\"] dest!: gt_1) + fix x assume "0 < u x" and [simp, intro]: "x \ space M" show "x \ (\n. ?M n \ ?A)" proof (cases "u x") - case (preal r) - obtain j where "1 / r \ of_nat j" using ex_le_of_nat .. - hence "1 / r * r \ of_nat j * r" using preal unfolding mult_le_cancel_right by auto - hence "1 \ of_nat j * r" using preal `u x \ 0` by auto - thus ?thesis using `u x \ 0` preal by (auto simp: real_of_nat_def[symmetric]) - qed auto - qed - finally show "\ ?A = 0" by simp + case (real r) with `0 < u x` have "0 < r" by auto + obtain j :: nat where "1 / r \ real j" using real_arch_simple .. + hence "1 / r * r \ real j * r" unfolding mult_le_cancel_right using `0 < r` by auto + hence "1 \ real j * r" using real `0 < r` by auto + thus ?thesis using `0 < r` real by (auto simp: one_extreal_def) + qed (insert `0 < u x`, auto) + qed auto + finally have "\ {x\space M. 0 < u x} = 0" by simp + moreover + from pos have "AE x. \ (u x < 0)" by auto + then have "\ {x\space M. u x < 0} = 0" + using AE_iff_null_set u by auto + moreover have "\ {x\space M. u x \ 0} = \ {x\space M. u x < 0} + \ {x\space M. 0 < u x}" + using u by (subst measure_additive) (auto intro!: arg_cong[where f=\]) + ultimately show "\ ?A = 0" by simp qed qed lemma (in measure_space) positive_integral_0_iff_AE: assumes u: "u \ borel_measurable M" - shows "integral\<^isup>P M u = 0 \ (AE x. u x = 0)" + shows "integral\<^isup>P M u = 0 \ (AE x. u x \ 0)" proof - - have sets: "{x\space M. u x \ 0} \ sets M" + have sets: "{x\space M. max 0 (u x) \ 0} \ sets M" using u by auto - then show ?thesis - using positive_integral_0_iff[OF u] AE_iff_null_set[OF sets] by auto + from positive_integral_0_iff[of "\x. max 0 (u x)"] + have "integral\<^isup>P M u = 0 \ (AE x. max 0 (u x) = 0)" + unfolding positive_integral_max_0 + using AE_iff_null_set[OF sets] u by auto + also have "\ \ (AE x. u x \ 0)" by (auto split: split_max) + finally show ?thesis . qed lemma (in measure_space) positive_integral_restricted: - assumes "A \ sets M" + assumes A: "A \ sets M" shows "integral\<^isup>P (restricted_space A) f = (\\<^isup>+ x. f x * indicator A x \M)" (is "integral\<^isup>P ?R f = integral\<^isup>P M ?f") proof - - have msR: "measure_space ?R" by (rule restricted_measure_space[OF `A \ sets M`]) - then interpret R: measure_space ?R . - have saR: "sigma_algebra ?R" by fact - have *: "integral\<^isup>P ?R f = integral\<^isup>P ?R ?f" - by (intro R.positive_integral_cong) auto + interpret R: measure_space ?R + by (rule restricted_measure_space) fact + let "?I g x" = "g x * indicator A x :: extreal" show ?thesis - unfolding * positive_integral_def - unfolding simple_function_restricted[OF `A \ sets M`] - apply (simp add: SUPR_def) - apply (rule arg_cong[where f=Sup]) - proof (auto simp add: image_iff simple_integral_restricted[OF `A \ sets M`]) - fix g assume "simple_function M (\x. g x * indicator A x)" - "g \ f" - then show "\x. simple_function M x \ x \ (\x. f x * indicator A x) \ - (\\<^isup>Sx. g x * indicator A x \M) = integral\<^isup>S M x" - apply (rule_tac exI[of _ "\x. g x * indicator A x"]) - by (auto simp: indicator_def le_fun_def) + unfolding positive_integral_def + unfolding simple_function_restricted[OF A] + unfolding AE_restricted[OF A] + proof (safe intro!: SUPR_eq) + fix g assume g: "simple_function M (?I g)" and le: "g \ max 0 \ f" + show "\j\{g. simple_function M g \ g \ max 0 \ ?I f}. + integral\<^isup>S (restricted_space A) g \ integral\<^isup>S M j" + proof (safe intro!: bexI[of _ "?I g"]) + show "integral\<^isup>S (restricted_space A) g \ integral\<^isup>S M (?I g)" + using g A by (simp add: simple_integral_restricted) + show "?I g \ max 0 \ ?I f" + using le by (auto simp: le_fun_def max_def indicator_def split: split_if_asm) + qed fact next - fix g assume g: "simple_function M g" "g \ (\x. f x * indicator A x)" - then have *: "(\x. g x * indicator A x) = g" - "\x. g x * indicator A x = g x" - "\x. g x \ f x" - by (auto simp: le_fun_def fun_eq_iff indicator_def split: split_if_asm) - from g show "\x. simple_function M (\xa. x xa * indicator A xa) \ x \ f \ - integral\<^isup>S M g = integral\<^isup>S M (\xa. x xa * indicator A xa)" - using `A \ sets M`[THEN sets_into_space] - apply (rule_tac exI[of _ "\x. g x * indicator A x"]) - by (fastsimp simp: le_fun_def *) + fix g assume g: "simple_function M g" and le: "g \ max 0 \ ?I f" + show "\i\{g. simple_function M (?I g) \ g \ max 0 \ f}. + integral\<^isup>S M g \ integral\<^isup>S (restricted_space A) i" + proof (safe intro!: bexI[of _ "?I g"]) + show "?I g \ max 0 \ f" + using le by (auto simp: le_fun_def max_def indicator_def split: split_if_asm) + from le have "\x. g x \ ?I (?I g) x" + by (auto simp: le_fun_def max_def indicator_def split: split_if_asm) + then show "integral\<^isup>S M g \ integral\<^isup>S (restricted_space A) (?I g)" + using A g by (auto intro!: simple_integral_mono simp: simple_integral_restricted) + show "simple_function M (?I (?I g))" using g A by auto + qed qed qed lemma (in measure_space) positive_integral_subalgebra: - assumes borel: "f \ borel_measurable N" + assumes f: "f \ borel_measurable N" "AE x in N. 0 \ f x" and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ measure N A = \ A" and sa: "sigma_algebra N" shows "integral\<^isup>P N f = integral\<^isup>P M f" proof - interpret N: measure_space N using measure_space_subalgebra[OF sa N] . - from N.borel_measurable_implies_simple_function_sequence[OF borel] - obtain fs where Nsf: "\i. simple_function N (fs i)" and "fs \ f" by blast - then have sf: "\i. simple_function M (fs i)" - using simple_function_subalgebra[OF _ N(1,2)] by blast - from N.positive_integral_isoton_simple[OF `fs \ f` Nsf] + from N.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 N.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 * N.\ (fs i -` {x} \ space M))" - unfolding isoton_def simple_integral_def `space N = space M` by simp + 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 * \ (fs i -` {x} \ space M))" - using N N.simple_functionD(2)[OF Nsf] unfolding `space N = space M` by auto + using N N.simple_functionD(2)[OF fs(1)] unfolding `space N = space M` by auto also have "\ = integral\<^isup>P M f" - using positive_integral_isoton_simple[OF `fs \ f` sf] - unfolding isoton_def simple_integral_def `space N = space M` by simp + 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 . qed @@ -1555,16 +1668,15 @@ definition integrable where "integrable M f \ f \ borel_measurable M \ - (\\<^isup>+ x. Real (f x) \M) \ \ \ - (\\<^isup>+ x. Real (- f x) \M) \ \" + (\\<^isup>+ x. extreal (f x) \M) \ \ \ (\\<^isup>+ x. extreal (- f x) \M) \ \" lemma integrableD[dest]: assumes "integrable M f" - shows "f \ borel_measurable M" "(\\<^isup>+ x. Real (f x) \M) \ \" "(\\<^isup>+ x. Real (- f x) \M) \ \" + shows "f \ borel_measurable M" "(\\<^isup>+ x. extreal (f x) \M) \ \" "(\\<^isup>+ x. extreal (- f x) \M) \ \" using assms unfolding integrable_def by auto definition lebesgue_integral_def: - "integral\<^isup>L M f = real ((\\<^isup>+ x. Real (f x) \M)) - real ((\\<^isup>+ x. Real (- f x) \M))" + "integral\<^isup>L M f = real ((\\<^isup>+ x. extreal (f x) \M)) - real ((\\<^isup>+ x. extreal (- f x) \M))" syntax "_lebesgue_integral" :: "'a \ ('a \ real) \ ('a, 'b) measure_space_scheme \ real" ("\ _. _ \_" [60,61] 110) @@ -1572,6 +1684,17 @@ translations "\ x. f \M" == "CONST integral\<^isup>L M (%x. f)" +lemma (in measure_space) integrableE: + assumes "integrable M f" + obtains r q where + "(\\<^isup>+x. extreal (f x)\M) = extreal r" + "(\\<^isup>+x. extreal (-f x)\M) = extreal q" + "f \ borel_measurable M" "integral\<^isup>L M f = r - q" + using assms unfolding integrable_def lebesgue_integral_def + using positive_integral_positive[of "\x. extreal (f x)"] + using positive_integral_positive[of "\x. extreal (-f x)"] + by (cases rule: extreal2_cases[of "(\\<^isup>+x. extreal (-f x)\M)" "(\\<^isup>+x. extreal (f x)\M)"]) auto + lemma (in measure_space) integral_cong: assumes "\x. x \ space M \ f x = g x" shows "integral\<^isup>L M f = integral\<^isup>L M g" @@ -1580,21 +1703,16 @@ lemma (in measure_space) integral_cong_measure: assumes "\A. A \ sets M \ measure N A = \ A" "sets N = sets M" "space N = space M" shows "integral\<^isup>L N f = integral\<^isup>L M f" -proof - - interpret v: measure_space N - by (rule measure_space_cong) fact+ - show ?thesis - by (simp add: positive_integral_cong_measure[OF assms] lebesgue_integral_def) -qed + by (simp add: positive_integral_cong_measure[OF assms] lebesgue_integral_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" proof - - have "AE x. Real (f x) = Real (g x)" using cong by auto - moreover have "AE x. Real (- f x) = Real (- g x)" using cong by auto - ultimately show ?thesis - by (simp cong: positive_integral_cong_AE add: lebesgue_integral_def) + have *: "AE x. extreal (f x) = extreal (g x)" + "AE x. extreal (- f x) = extreal (- g x)" using cong by auto + show ?thesis + unfolding *[THEN positive_integral_cong_AE] lebesgue_integral_def .. qed lemma (in measure_space) integrable_cong: @@ -1602,11 +1720,14 @@ by (simp cong: positive_integral_cong measurable_cong add: integrable_def) lemma (in measure_space) integral_eq_positive_integral: - assumes "\x. 0 \ f x" - shows "integral\<^isup>L M f = real (\\<^isup>+ x. Real (f x) \M)" + assumes f: "\x. 0 \ f x" + shows "integral\<^isup>L M f = real (\\<^isup>+ x. extreal (f x) \M)" proof - - have "\x. Real (- f x) = 0" using assms by simp - thus ?thesis by (simp del: Real_eq_0 add: lebesgue_integral_def) + { fix x have "max 0 (extreal (- f x)) = 0" using f[of x] by (simp split: split_max) } + then have "0 = (\\<^isup>+ x. max 0 (extreal (- f x)) \M)" by simp + also have "\ = (\\<^isup>+ x. extreal (- f x) \M)" unfolding positive_integral_max_0 .. + finally show ?thesis + unfolding lebesgue_integral_def by simp qed lemma (in measure_space) integral_vimage: @@ -1616,7 +1737,7 @@ proof - interpret T: measure_space M' by (rule measure_space_vimage[OF T]) from measurable_comp[OF measure_preservingD2[OF T(2)], of f borel] - have borel: "(\x. Real (f x)) \ borel_measurable M'" "(\x. Real (- f x)) \ borel_measurable M'" + have borel: "(\x. extreal (f x)) \ borel_measurable M'" "(\x. extreal (- f x)) \ borel_measurable M'" and "(\x. f (T x)) \ borel_measurable M" using f by (auto simp: comp_def) then show ?thesis @@ -1631,7 +1752,7 @@ proof - interpret T: measure_space M' by (rule measure_space_vimage[OF T]) from measurable_comp[OF measure_preservingD2[OF T(2)], of f borel] - have borel: "(\x. Real (f x)) \ borel_measurable M'" "(\x. Real (- f x)) \ borel_measurable M'" + have borel: "(\x. extreal (f x)) \ borel_measurable M'" "(\x. extreal (- f x)) \ borel_measurable M'" and "(\x. f (T x)) \ borel_measurable M" using f by (auto simp: comp_def) then show ?thesis @@ -1649,10 +1770,10 @@ and f_def: "\x. f x = u x - v x" and pos: "\x. 0 \ u x" "\x. 0 \ v x" shows "integrable M f" and "integral\<^isup>L M f = integral\<^isup>L M u - integral\<^isup>L M v" proof - - let "?f x" = "Real (f x)" - let "?mf x" = "Real (- f x)" - let "?u x" = "Real (u x)" - let "?v x" = "Real (v x)" + let "?f x" = "max 0 (extreal (f x))" + let "?mf x" = "max 0 (extreal (- f x))" + let "?u x" = "max 0 (extreal (u x))" + let "?v x" = "max 0 (extreal (v x))" from borel_measurable_diff[of u v] integrable have f_borel: "?f \ borel_measurable M" and @@ -1662,73 +1783,62 @@ "f \ borel_measurable M" by (auto simp: f_def[symmetric] integrable_def) - have "(\\<^isup>+ x. Real (u x - v x) \M) \ integral\<^isup>P M ?u" - using pos by (auto intro!: positive_integral_mono) - moreover have "(\\<^isup>+ x. Real (v x - u x) \M) \ integral\<^isup>P M ?v" - using pos by (auto intro!: positive_integral_mono) + have "(\\<^isup>+ x. extreal (u x - v x) \M) \ integral\<^isup>P M ?u" + using pos by (auto intro!: positive_integral_mono simp: positive_integral_max_0) + moreover have "(\\<^isup>+ x. extreal (v x - u x) \M) \ integral\<^isup>P M ?v" + using pos by (auto intro!: positive_integral_mono simp: positive_integral_max_0) ultimately show f: "integrable M f" using `integrable M u` `integrable M v` `f \ borel_measurable M` - by (auto simp: integrable_def f_def) - hence mf: "integrable M (\x. - f x)" .. - - have *: "\x. Real (- v x) = 0" "\x. Real (- u x) = 0" - using pos by auto + by (auto simp: integrable_def f_def positive_integral_max_0) have "\x. ?u x + ?mf x = ?v x + ?f x" - unfolding f_def using pos by simp - hence "(\\<^isup>+ x. ?u x + ?mf x \M) = (\\<^isup>+ x. ?v x + ?f x \M)" by simp - hence "real (integral\<^isup>P M ?u + integral\<^isup>P M ?mf) = + unfolding f_def using pos by (simp split: split_max) + then have "(\\<^isup>+ x. ?u x + ?mf x \M) = (\\<^isup>+ x. ?v x + ?f x \M)" by simp + then have "real (integral\<^isup>P M ?u + integral\<^isup>P M ?mf) = real (integral\<^isup>P M ?v + integral\<^isup>P M ?f)" - using positive_integral_add[OF u_borel mf_borel] - using positive_integral_add[OF v_borel f_borel] + using positive_integral_add[OF u_borel _ mf_borel] + using positive_integral_add[OF v_borel _ f_borel] by auto then show "integral\<^isup>L M f = integral\<^isup>L M u - integral\<^isup>L M v" - using f mf `integrable M u` `integrable M v` - unfolding lebesgue_integral_def integrable_def * - by (cases "integral\<^isup>P M ?f", cases "integral\<^isup>P M ?mf", cases "integral\<^isup>P M ?v", cases "integral\<^isup>P M ?u") - (auto simp add: field_simps) + unfolding positive_integral_max_0 + unfolding pos[THEN integral_eq_positive_integral] + using integrable f by (auto elim!: integrableE) qed lemma (in measure_space) integral_linear: assumes "integrable M f" "integrable M g" and "0 \ a" shows "integrable M (\t. a * f t + g t)" - and "(\ t. a * f t + g t \M) = a * integral\<^isup>L M f + integral\<^isup>L M g" + and "(\ t. a * f t + g t \M) = a * integral\<^isup>L M f + integral\<^isup>L M g" (is ?EQ) proof - - let ?PI = "integral\<^isup>P M" - let "?f x" = "Real (f x)" - let "?g x" = "Real (g x)" - let "?mf x" = "Real (- f x)" - let "?mg x" = "Real (- g x)" + let "?f x" = "max 0 (extreal (f x))" + let "?g x" = "max 0 (extreal (g x))" + let "?mf x" = "max 0 (extreal (- f x))" + let "?mg x" = "max 0 (extreal (- g x))" let "?p t" = "max 0 (a * f t) + max 0 (g t)" let "?n t" = "max 0 (- (a * f t)) + max 0 (- g t)" - have pos: "?f \ borel_measurable M" "?g \ borel_measurable M" - and neg: "?mf \ borel_measurable M" "?mg \ borel_measurable M" - and p: "?p \ borel_measurable M" - and n: "?n \ borel_measurable M" - using assms by (simp_all add: integrable_def) + from assms have linear: + "(\\<^isup>+ x. extreal a * ?f x + ?g x \M) = extreal a * integral\<^isup>P M ?f + integral\<^isup>P M ?g" + "(\\<^isup>+ x. extreal a * ?mf x + ?mg x \M) = extreal a * integral\<^isup>P M ?mf + integral\<^isup>P M ?mg" + by (auto intro!: positive_integral_linear simp: integrable_def) - have *: "\x. Real (?p x) = Real a * ?f x + ?g x" - "\x. Real (?n x) = Real a * ?mf x + ?mg x" - "\x. Real (- ?p x) = 0" - "\x. Real (- ?n x) = 0" - using `0 \ a` by (auto simp: max_def min_def zero_le_mult_iff mult_le_0_iff add_nonpos_nonpos) - - note linear = - positive_integral_linear[OF pos] - positive_integral_linear[OF neg] + have *: "(\\<^isup>+x. extreal (- ?p x) \M) = 0" "(\\<^isup>+x. extreal (- ?n x) \M) = 0" + using `0 \ a` assms by (auto simp: positive_integral_0_iff_AE integrable_def) + have **: "\x. extreal a * ?f x + ?g x = max 0 (extreal (?p x))" + "\x. extreal a * ?mf x + ?mg x = max 0 (extreal (?n x))" + using `0 \ a` by (auto split: split_max simp: zero_le_mult_iff mult_le_0_iff) have "integrable M ?p" "integrable M ?n" "\t. a * f t + g t = ?p t - ?n t" "\t. 0 \ ?p t" "\t. 0 \ ?n t" - using assms p n unfolding integrable_def * linear by auto + using linear assms unfolding integrable_def ** * + by (auto simp: positive_integral_max_0) note diff = integral_of_positive_diff[OF this] show "integrable M (\t. a * f t + g t)" by (rule diff) - - from assms show "(\ t. a * f t + g t \M) = a * integral\<^isup>L M f + integral\<^isup>L M g" - unfolding diff(2) unfolding lebesgue_integral_def * linear integrable_def - by (cases "?PI ?f", cases "?PI ?mf", cases "?PI ?g", cases "?PI ?mg") - (auto simp add: field_simps zero_le_mult_iff) + from assms linear show ?EQ + unfolding diff(2) ** positive_integral_max_0 + unfolding lebesgue_integral_def * + by (auto elim!: integrableE simp: field_simps) qed lemma (in measure_space) integral_add[simp, intro]: @@ -1772,13 +1882,13 @@ and mono: "AE t. f t \ g t" shows "integral\<^isup>L M f \ integral\<^isup>L M g" proof - - have "AE x. Real (f x) \ Real (g x)" + have "AE x. extreal (f x) \ extreal (g x)" using mono by auto - moreover have "AE x. Real (- g x) \ Real (- f x)" + moreover have "AE x. extreal (- g x) \ extreal (- f x)" using mono by auto ultimately show ?thesis using fg - by (auto simp: lebesgue_integral_def integrable_def diff_minus - intro!: add_mono real_of_pextreal_mono positive_integral_mono_AE) + by (auto intro!: add_mono positive_integral_mono_AE real_of_extreal_positive_mono + simp: positive_integral_positive lebesgue_integral_def diff_minus) qed lemma (in measure_space) integral_mono: @@ -1795,20 +1905,21 @@ by auto lemma (in measure_space) integral_indicator[simp, intro]: - assumes "a \ sets M" and "\ a \ \" - shows "integral\<^isup>L M (indicator a) = real (\ a)" (is ?int) - and "integrable M (indicator a)" (is ?able) + assumes "A \ sets M" and "\ A \ \" + shows "integral\<^isup>L M (indicator A) = real (\ A)" (is ?int) + and "integrable M (indicator A)" (is ?able) proof - - have *: - "\A x. Real (indicator A x) = indicator A x" - "\A x. Real (- indicator A x) = 0" unfolding indicator_def by auto + from `A \ sets M` have *: + "\x. extreal (indicator A x) = indicator A x" + "(\\<^isup>+x. extreal (- indicator A x) \M) = 0" + by (auto split: split_indicator simp: positive_integral_0_iff_AE one_extreal_def) show ?int ?able using assms unfolding lebesgue_integral_def integrable_def by (auto simp: * positive_integral_indicator borel_measurable_indicator) qed lemma (in measure_space) integral_cmul_indicator: - assumes "A \ sets M" and "c \ 0 \ \ A \ \" + assumes "A \ sets M" and "c \ 0 \ \ A \ \" shows "integrable M (\x. c * indicator A x)" (is ?P) and "(\x. c * indicator A x \M) = c * real (\ A)" (is ?I) proof - @@ -1840,15 +1951,11 @@ assumes "integrable M f" shows "integrable M (\ x. \f x\)" proof - - have *: - "\x. Real \f x\ = Real (f x) + Real (- f x)" - "\x. Real (- \f x\) = 0" by auto - have abs: "(\x. \f x\) \ borel_measurable M" and - f: "(\x. Real (f x)) \ borel_measurable M" - "(\x. Real (- f x)) \ borel_measurable M" - using assms unfolding integrable_def by auto - from abs assms show ?thesis unfolding integrable_def * - using positive_integral_linear[OF f, of 1] by simp + from assms have *: "(\\<^isup>+x. extreal (- \f x\)\M) = 0" + "\x. extreal \f x\ = max 0 (extreal (f x)) + max 0 (extreal (- f x))" + by (auto simp: integrable_def positive_integral_0_iff_AE split: split_max) + with assms show ?thesis + by (simp add: positive_integral_add positive_integral_max_0 integrable_def) qed lemma (in measure_space) integral_subalgebra: @@ -1858,12 +1965,13 @@ and "integral\<^isup>L N f = integral\<^isup>L M f" (is ?I) proof - interpret N: measure_space N using measure_space_subalgebra[OF sa N] . - have "(\x. Real (f x)) \ borel_measurable N" "(\x. Real (- f x)) \ borel_measurable N" - using borel by auto - note * = this[THEN positive_integral_subalgebra[OF _ N sa]] - have "f \ borel_measurable M \ f \ borel_measurable N" + have "(\\<^isup>+ x. max 0 (extreal (f x)) \N) = (\\<^isup>+ x. max 0 (extreal (f x)) \M)" + "(\\<^isup>+ x. max 0 (extreal (- f x)) \N) = (\\<^isup>+ x. max 0 (extreal (- f x)) \M)" + using borel by (auto intro!: positive_integral_subalgebra N sa) + moreover have "f \ borel_measurable M \ f \ borel_measurable N" using assms unfolding measurable_def by auto - then show ?P ?I by (auto simp: * integrable_def lebesgue_integral_def) + ultimately show ?P ?I + by (auto simp: integrable_def lebesgue_integral_def positive_integral_max_0) qed lemma (in measure_space) integrable_bound: @@ -1873,21 +1981,21 @@ assumes borel: "g \ borel_measurable M" shows "integrable M g" proof - - have "(\\<^isup>+ x. Real (g x) \M) \ (\\<^isup>+ x. Real \g x\ \M)" + have "(\\<^isup>+ x. extreal (g x) \M) \ (\\<^isup>+ x. extreal \g x\ \M)" by (auto intro!: positive_integral_mono) - also have "\ \ (\\<^isup>+ x. Real (f x) \M)" + also have "\ \ (\\<^isup>+ x. extreal (f x) \M)" using f by (auto intro!: positive_integral_mono) - also have "\ < \" - using `integrable M f` unfolding integrable_def by (auto simp: pextreal_less_\) - finally have pos: "(\\<^isup>+ x. Real (g x) \M) < \" . + also have "\ < \" + using `integrable M f` unfolding integrable_def by auto + finally have pos: "(\\<^isup>+ x. extreal (g x) \M) < \" . - have "(\\<^isup>+ x. Real (- g x) \M) \ (\\<^isup>+ x. Real (\g x\) \M)" + have "(\\<^isup>+ x. extreal (- g x) \M) \ (\\<^isup>+ x. extreal (\g x\) \M)" by (auto intro!: positive_integral_mono) - also have "\ \ (\\<^isup>+ x. Real (f x) \M)" + also have "\ \ (\\<^isup>+ x. extreal (f x) \M)" using f by (auto intro!: positive_integral_mono) - also have "\ < \" - using `integrable M f` unfolding integrable_def by (auto simp: pextreal_less_\) - finally have neg: "(\\<^isup>+ x. Real (- g x) \M) < \" . + also have "\ < \" + using `integrable M f` unfolding integrable_def by auto + finally have neg: "(\\<^isup>+ x. extreal (- g x) \M) < \" . from neg pos borel show ?thesis unfolding integrable_def by auto @@ -1959,41 +2067,34 @@ by (simp add: mono_def incseq_def) } note pos_u = this - hence [simp]: "\i x. Real (- f i x) = 0" "\x. Real (- u x) = 0" - using pos by auto + have SUP_F: "\x. (SUP n. extreal (f n x)) = extreal (u x)" + unfolding SUP_eq_LIMSEQ[OF mono] by (rule lim) - have SUP_F: "\x. (SUP n. Real (f n x)) = Real (u x)" - using mono pos pos_u lim by (rule SUP_eq_LIMSEQ[THEN iffD2]) - - have borel_f: "\i. (\x. Real (f i x)) \ borel_measurable M" + have borel_f: "\i. (\x. extreal (f i x)) \ borel_measurable M" using i unfolding integrable_def by auto - hence "(\x. SUP i. Real (f i x)) \ borel_measurable M" + hence "(\x. SUP i. extreal (f i x)) \ borel_measurable M" by auto hence borel_u: "u \ borel_measurable M" - using pos_u by (auto simp: borel_measurable_Real_eq SUP_F) + by (auto simp: borel_measurable_extreal_iff SUP_F) - have integral_eq: "\n. (\\<^isup>+ x. Real (f n x) \M) = Real (integral\<^isup>L M (f n))" - using i unfolding lebesgue_integral_def integrable_def by (auto simp: Real_real) + hence [simp]: "\i. (\\<^isup>+x. extreal (- f i x) \M) = 0" "(\\<^isup>+x. extreal (- u x) \M) = 0" + using i borel_u pos pos_u by (auto simp: positive_integral_0_iff_AE integrable_def) + + have integral_eq: "\n. (\\<^isup>+ x. extreal (f n x) \M) = extreal (integral\<^isup>L M (f n))" + using i positive_integral_positive by (auto simp: extreal_real lebesgue_integral_def integrable_def) have pos_integral: "\n. 0 \ integral\<^isup>L M (f n)" using pos i by (auto simp: integral_positive) hence "0 \ x" using LIMSEQ_le_const[OF ilim, of 0] by auto - have "(\i. (\\<^isup>+ x. Real (f i x) \M)) \ (\\<^isup>+ x. Real (u x) \M)" - proof (rule positive_integral_isoton) - from SUP_F mono pos - show "(\i x. Real (f i x)) \ (\x. Real (u x))" - unfolding isoton_fun_expand by (auto simp: isoton_def mono_def) - qed (rule borel_f) - hence pI: "(\\<^isup>+ x. Real (u x) \M) = (SUP n. (\\<^isup>+ x. Real (f n x) \M))" - unfolding isoton_def by simp - also have "\ = Real x" unfolding integral_eq + from mono pos i have pI: "(\\<^isup>+ x. extreal (u x) \M) = (SUP n. (\\<^isup>+ x. extreal (f n x) \M))" + by (auto intro!: positive_integral_monotone_convergence_SUP + simp: integrable_def incseq_mono incseq_Suc_iff le_fun_def SUP_F[symmetric]) + also have "\ = extreal x" unfolding integral_eq proof (rule SUP_eq_LIMSEQ[THEN iffD2]) show "mono (\n. integral\<^isup>L M (f n))" using mono i by (auto simp: mono_def intro!: integral_mono) - show "\n. 0 \ integral\<^isup>L M (f n)" using pos_integral . - show "0 \ x" using `0 \ x` . show "(\n. integral\<^isup>L M (f n)) ----> x" using ilim . qed finally show "integrable M u" "integral\<^isup>L M u = x" using borel_u `0 \ x` @@ -2028,61 +2129,72 @@ assumes "integrable M f" shows "(\x. \f x\ \M) = 0 \ \ {x\space M. f x \ 0} = 0" proof - - have *: "\x. Real (- \f x\) = 0" by auto + have *: "(\\<^isup>+x. extreal (- \f x\) \M) = 0" + using assms by (auto simp: positive_integral_0_iff_AE integrable_def) have "integrable M (\x. \f x\)" using assms by (rule integrable_abs) - hence "(\x. Real (\f x\)) \ borel_measurable M" - "(\\<^isup>+ x. Real \f x\ \M) \ \" unfolding integrable_def by auto + hence "(\x. extreal (\f x\)) \ borel_measurable M" + "(\\<^isup>+ x. extreal \f x\ \M) \ \" unfolding integrable_def by auto from positive_integral_0_iff[OF this(1)] this(2) show ?thesis unfolding lebesgue_integral_def * - by (simp add: real_of_pextreal_eq_0) + using positive_integral_positive[of "\x. extreal \f x\"] + by (auto simp add: real_of_extreal_eq_0) qed -lemma (in measure_space) positive_integral_omega: - assumes "f \ borel_measurable M" - and "integral\<^isup>P M f \ \" - shows "\ (f -` {\} \ space M) = 0" +lemma (in measure_space) positive_integral_PInf: + assumes f: "f \ borel_measurable M" + and not_Inf: "integral\<^isup>P M f \ \" + shows "\ (f -` {\} \ space M) = 0" proof - - have "\ * \ (f -` {\} \ space M) = (\\<^isup>+ x. \ * indicator (f -` {\} \ space M) x \M)" - using positive_integral_cmult_indicator[OF borel_measurable_vimage, OF assms(1), of \ \] by simp - also have "\ \ integral\<^isup>P M f" - by (auto intro!: positive_integral_mono simp: indicator_def) - finally show ?thesis - using assms(2) by (cases ?thesis) auto + have "\ * \ (f -` {\} \ space M) = (\\<^isup>+ x. \ * indicator (f -` {\} \ space M) x \M)" + using f by (subst positive_integral_cmult_indicator) (auto simp: measurable_sets) + also have "\ \ integral\<^isup>P M (\x. max 0 (f x))" + by (auto intro!: positive_integral_mono simp: indicator_def max_def) + finally have "\ * \ (f -` {\} \ space M) \ integral\<^isup>P M f" + by (simp add: positive_integral_max_0) + moreover have "0 \ \ (f -` {\} \ space M)" + using f by (simp add: measurable_sets) + ultimately show ?thesis + using assms by (auto split: split_if_asm) qed -lemma (in measure_space) positive_integral_omega_AE: - assumes "f \ borel_measurable M" "integral\<^isup>P M f \ \" shows "AE x. f x \ \" +lemma (in measure_space) positive_integral_PInf_AE: + assumes "f \ borel_measurable M" "integral\<^isup>P M f \ \" shows "AE x. f x \ \" proof (rule AE_I) - show "\ (f -` {\} \ space M) = 0" - by (rule positive_integral_omega[OF assms]) - show "f -` {\} \ space M \ sets M" + show "\ (f -` {\} \ space M) = 0" + by (rule positive_integral_PInf[OF assms]) + show "f -` {\} \ space M \ sets M" using assms by (auto intro: borel_measurable_vimage) qed auto -lemma (in measure_space) simple_integral_omega: - assumes "simple_function M f" - and "integral\<^isup>S M f \ \" - shows "\ (f -` {\} \ space M) = 0" -proof (rule positive_integral_omega) +lemma (in measure_space) simple_integral_PInf: + assumes "simple_function M f" "\x. 0 \ f x" + and "integral\<^isup>S M f \ \" + shows "\ (f -` {\} \ space M) = 0" +proof (rule positive_integral_PInf) show "f \ borel_measurable M" using assms by (auto intro: borel_measurable_simple_function) - show "integral\<^isup>P M f \ \" + show "integral\<^isup>P M f \ \" using assms by (simp add: positive_integral_eq_simple_integral) qed lemma (in measure_space) integral_real: - fixes f :: "'a \ pextreal" - assumes [simp]: "AE x. f x \ \" - shows "(\x. real (f x) \M) = real (integral\<^isup>P M f)" (is ?plus) - and "(\x. - real (f x) \M) = - real (integral\<^isup>P M f)" (is ?minus) -proof - - have "(\\<^isup>+ x. Real (real (f x)) \M) = integral\<^isup>P M f" - by (auto intro!: positive_integral_cong_AE simp: Real_real) - moreover - have "(\\<^isup>+ x. Real (- real (f x)) \M) = (\\<^isup>+ x. 0 \M)" - by (intro positive_integral_cong) auto - ultimately show ?plus ?minus - by (auto simp: lebesgue_integral_def integrable_def) -qed + "AE x. \f x\ \ \ \ (\x. real (f x) \M) = real (integral\<^isup>P M f) - real (integral\<^isup>P M (\x. - f x))" + using assms unfolding lebesgue_integral_def + by (subst (1 2) positive_integral_cong_AE) (auto simp add: extreal_real) + +lemma liminf_extreal_cminus: + fixes f :: "nat \ extreal" assumes "c \ -\" + shows "liminf (\x. c - f x) = c - limsup f" +proof (cases c) + case PInf then show ?thesis by (simp add: Liminf_const) +next + case (real r) then show ?thesis + unfolding liminf_SUPR_INFI limsup_INFI_SUPR + apply (subst INFI_extreal_cminus) + apply auto + apply (subst SUPR_extreal_cminus) + apply auto + done +qed (insert `c \ -\`, simp) 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" @@ -2129,61 +2241,76 @@ finally have "\u j x - u' x\ \ 2 * w x" by simp } note diff_less_2w = this - have PI_diff: "\m n. (\\<^isup>+ x. Real (?diff (m + n) x) \M) = - (\\<^isup>+ x. Real (2 * w x) \M) - (\\<^isup>+ x. Real \u (m + n) x - u' x\ \M)" + have PI_diff: "\n. (\\<^isup>+ x. extreal (?diff n x) \M) = + (\\<^isup>+ x. extreal (2 * w x) \M) - (\\<^isup>+ x. extreal \u n x - u' x\ \M)" using diff w diff_less_2w w_pos by (subst positive_integral_diff[symmetric]) (auto simp: integrable_def intro!: positive_integral_cong) have "integrable M (\x. 2 * w x)" using w by (auto intro: integral_cmult) - hence I2w_fin: "(\\<^isup>+ x. Real (2 * w x) \M) \ \" and - borel_2w: "(\x. Real (2 * w x)) \ borel_measurable M" + hence I2w_fin: "(\\<^isup>+ x. extreal (2 * w x) \M) \ \" and + borel_2w: "(\x. extreal (2 * w x)) \ borel_measurable M" unfolding integrable_def by auto - have "(INF n. SUP m. (\\<^isup>+ x. Real \u (m + n) x - u' x\ \M)) = 0" (is "?lim_SUP = 0") + have "limsup (\n. \\<^isup>+ x. extreal \u n x - u' x\ \M) = 0" (is "limsup ?f = 0") proof cases - assume eq_0: "(\\<^isup>+ x. Real (2 * w x) \M) = 0" - have "\i. (\\<^isup>+ x. Real \u i x - u' x\ \M) \ (\\<^isup>+ x. Real (2 * w x) \M)" - proof (rule positive_integral_mono) - fix i x assume "x \ space M" from diff_less_2w[OF this, of i] - show "Real \u i x - u' x\ \ Real (2 * w x)" by auto - qed - hence "\i. (\\<^isup>+ x. Real \u i x - u' x\ \M) = 0" using eq_0 by auto - thus ?thesis by simp + assume eq_0: "(\\<^isup>+ x. max 0 (extreal (2 * w x)) \M) = 0" (is "?wx = 0") + { fix n + have "?f n \ ?wx" (is "integral\<^isup>P M ?f' \ _") + using diff_less_2w[of _ n] unfolding positive_integral_max_0 + by (intro positive_integral_mono) auto + then have "?f n = 0" + using positive_integral_positive[of ?f'] eq_0 by auto } + then show ?thesis by (simp add: Limsup_const) next - assume neq_0: "(\\<^isup>+ x. Real (2 * w x) \M) \ 0" - have "(\\<^isup>+ x. Real (2 * w x) \M) = (\\<^isup>+ x. (SUP n. INF m. Real (?diff (m + n) x)) \M)" - proof (rule positive_integral_cong, subst add_commute) + assume neq_0: "(\\<^isup>+ x. max 0 (extreal (2 * w x)) \M) \ 0" (is "?wx \ 0") + have "0 = limsup (\n. 0 :: extreal)" by (simp add: Limsup_const) + also have "\ \ limsup (\n. \\<^isup>+ x. extreal \u n x - u' x\ \M)" + by (intro limsup_mono positive_integral_positive) + finally have pos: "0 \ limsup (\n. \\<^isup>+ x. extreal \u n x - u' x\ \M)" . + have "?wx = (\\<^isup>+ x. liminf (\n. max 0 (extreal (?diff n x))) \M)" + proof (rule positive_integral_cong) fix x assume x: "x \ space M" - show "Real (2 * w x) = (SUP n. INF m. Real (?diff (n + m) x))" - proof (rule LIMSEQ_imp_lim_INF[symmetric]) - fix j show "0 \ ?diff j x" using diff_less_2w[OF x, of j] by simp - next + show "max 0 (extreal (2 * w x)) = liminf (\n. max 0 (extreal (?diff n x)))" + unfolding extreal_max_0 + proof (rule lim_imp_Liminf[symmetric], unfold lim_extreal) have "(\i. ?diff i x) ----> 2 * w x - \u' x - u' x\" using u'[OF x] by (safe intro!: LIMSEQ_diff LIMSEQ_const LIMSEQ_imp_rabs) - thus "(\i. ?diff i x) ----> 2 * w x" by simp - qed + then show "(\i. max 0 (?diff i x)) ----> max 0 (2 * w x)" + by (auto intro!: tendsto_real_max simp add: lim_extreal) + qed (rule trivial_limit_sequentially) qed - also have "\ \ (SUP n. INF m. (\\<^isup>+ x. Real (?diff (m + n) x) \M))" + also have "\ \ liminf (\n. \\<^isup>+ x. max 0 (extreal (?diff n x)) \M)" using u'_borel w u unfolding integrable_def - by (auto intro!: positive_integral_lim_INF) - also have "\ = (\\<^isup>+ x. Real (2 * w x) \M) - - (INF n. SUP m. \\<^isup>+ x. Real \u (m + n) x - u' x\ \M)" - unfolding PI_diff pextreal_INF_minus[OF I2w_fin] pextreal_SUP_minus .. - finally show ?thesis using neq_0 I2w_fin by (rule pextreal_le_minus_imp_0) + by (intro positive_integral_lim_INF) (auto intro!: positive_integral_lim_INF) + also have "\ = (\\<^isup>+ x. extreal (2 * w x) \M) - + limsup (\n. \\<^isup>+ x. extreal \u n x - u' x\ \M)" + unfolding PI_diff positive_integral_max_0 + using positive_integral_positive[of "\x. extreal (2 * w x)"] + by (subst liminf_extreal_cminus) auto + finally show ?thesis + using neq_0 I2w_fin positive_integral_positive[of "\x. extreal (2 * w x)"] pos + unfolding positive_integral_max_0 + by (cases rule: extreal2_cases[of "\\<^isup>+ x. extreal (2 * w x) \M" "limsup (\n. \\<^isup>+ x. extreal \u n x - u' x\ \M)"]) + auto qed - have [simp]: "\n m x. Real (- \u (m + n) x - u' x\) = 0" by auto - - have [simp]: "\n m. (\\<^isup>+ x. Real \u (m + n) x - u' x\ \M) = - Real ((\x. \u (n + m) x - u' x\ \M))" - using diff by (subst add_commute) (simp add: lebesgue_integral_def integrable_def Real_real) - - have "(SUP n. INF m. (\\<^isup>+ x. Real \u (m + n) x - u' x\ \M)) \ ?lim_SUP" - (is "?lim_INF \ _") by (subst (1 2) add_commute) (rule lim_INF_le_lim_SUP) - hence "?lim_INF = Real 0" "?lim_SUP = Real 0" using `?lim_SUP = 0` by auto - thus ?lim_diff using diff by (auto intro!: integral_positive lim_INF_eq_lim_SUP) + have "liminf ?f \ limsup ?f" + by (intro extreal_Liminf_le_Limsup trivial_limit_sequentially) + moreover + { have "0 = liminf (\n. 0 :: extreal)" by (simp add: Liminf_const) + also have "\ \ liminf ?f" + by (intro liminf_mono positive_integral_positive) + finally have "0 \ liminf ?f" . } + ultimately have liminf_limsup_eq: "liminf ?f = extreal 0" "limsup ?f = extreal 0" + using `limsup ?f = 0` by auto + have "\n. (\\<^isup>+ x. extreal \u n x - u' x\ \M) = extreal (\x. \u n x - u' x\ \M)" + using diff positive_integral_positive + by (subst integral_eq_positive_integral) (auto simp: extreal_real integrable_def) + then show ?lim_diff + using extreal_Liminf_eq_Limsup[OF trivial_limit_sequentially liminf_limsup_eq] + by (simp add: lim_extreal) show ?lim proof (rule LIMSEQ_I) @@ -2266,7 +2393,7 @@ assumes f: "f \ borel_measurable M" and bij: "bij_betw enum S (f ` space M)" and enum_zero: "enum ` (-S) \ {0}" - and fin: "\x. x \ 0 \ \ (f -` {x} \ space M) \ \" + and fin: "\x. x \ 0 \ \ (f -` {x} \ space M) \ \" and abs_summable: "summable (\r. \enum r * real (\ (f -` {enum r} \ space M))\)" shows "integrable M f" and "(\r. enum r * real (\ (f -` {enum r} \ space M))) sums integral\<^isup>L M f" (is ?sums) @@ -2303,7 +2430,7 @@ also have "\ = \enum r\ * real (\ (?A r))" using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator) finally have "(\x. \?F r x\ \M) = \enum r * real (\ (?A r))\" - by (simp add: abs_mult_pos real_pextreal_pos) } + using f by (subst (2) abs_mult_pos[symmetric]) (auto intro!: real_of_extreal_pos measurable_sets) } note int_abs_F = this have 1: "\i. integrable M (\x. ?F i x)" @@ -2323,7 +2450,7 @@ lemma (in measure_space) integral_on_finite: assumes f: "f \ borel_measurable M" and finite: "finite (f`space M)" - and fin: "\x. x \ 0 \ \ (f -` {x} \ space M) \ \" + and fin: "\x. x \ 0 \ \ (f -` {x} \ space M) \ \" shows "integrable M f" and "(\x. f x \M) = (\ r \ f`space M. r * real (\ (f -` {r} \ space M)))" (is "?integral") @@ -2353,11 +2480,12 @@ by (auto intro: borel_measurable_simple_function) lemma (in finite_measure_space) positive_integral_finite_eq_setsum: - "integral\<^isup>P M f = (\x \ space M. f x * \ {x})" + assumes pos: "\x. x \ space M \ 0 \ f x" + shows "integral\<^isup>P M f = (\x \ space M. f x * \ {x})" proof - have *: "integral\<^isup>P M f = (\\<^isup>+ x. (\y\space M. f y * indicator {y} x) \M)" by (auto intro!: positive_integral_cong simp add: indicator_def if_distrib setsum_cases[OF finite_space]) - show ?thesis unfolding * using borel_measurable_finite[of f] + show ?thesis unfolding * using borel_measurable_finite[of f] pos by (simp add: positive_integral_setsum positive_integral_cmult_indicator) qed @@ -2365,16 +2493,20 @@ shows "integrable M f" and "integral\<^isup>L M f = (\x \ space M. f x * real (\ {x}))" (is ?I) proof - - have [simp]: - "(\\<^isup>+ x. Real (f x) \M) = (\x \ space M. Real (f x) * \ {x})" - "(\\<^isup>+ x. Real (- f x) \M) = (\x \ space M. Real (- f x) * \ {x})" - unfolding positive_integral_finite_eq_setsum by auto - show "integrable M f" using finite_space finite_measure - by (simp add: setsum_\ integrable_def) - show ?I using finite_measure - apply (simp add: lebesgue_integral_def real_of_pextreal_setsum[symmetric] - real_of_pextreal_mult[symmetric] setsum_subtractf[symmetric]) - by (rule setsum_cong) (simp_all split: split_if) + have *: + "(\\<^isup>+ x. max 0 (extreal (f x)) \M) = (\x \ space M. max 0 (extreal (f x)) * \ {x})" + "(\\<^isup>+ x. max 0 (extreal (- f x)) \M) = (\x \ space M. max 0 (extreal (- f x)) * \ {x})" + by (simp_all add: positive_integral_finite_eq_setsum) + then show "integrable M f" using finite_space finite_measure + by (simp add: setsum_Pinfty integrable_def positive_integral_max_0 + split: split_max) + show ?I using finite_measure * + apply (simp add: positive_integral_max_0 lebesgue_integral_def) + apply (subst (1 2) setsum_real_of_extreal[symmetric]) + apply (simp_all split: split_max add: setsum_subtractf[symmetric]) + apply (intro setsum_cong[OF refl]) + apply (simp split: split_max) + done qed end diff -r 28b51effc5ed -r cdf7693bbe08 src/HOL/Probability/Lebesgue_Measure.thy --- a/src/HOL/Probability/Lebesgue_Measure.thy Mon Mar 14 14:37:47 2011 +0100 +++ b/src/HOL/Probability/Lebesgue_Measure.thy Mon Mar 14 14:37:49 2011 +0100 @@ -48,12 +48,12 @@ lemma Pi_iff: "f \ Pi I X \ (\i\I. f i \ X i)" unfolding Pi_def by auto -subsection {* Lebesgue measure *} +subsection {* Lebesgue measure *} definition lebesgue :: "'a::ordered_euclidean_space measure_space" where "lebesgue = \ space = UNIV, sets = {A. \n. (indicator A :: 'a \ real) integrable_on cube n}, - measure = \A. SUP n. Real (integral (cube n) (indicator A)) \" + measure = \A. SUP n. extreal (integral (cube n) (indicator A)) \" lemma space_lebesgue[simp]: "space lebesgue = UNIV" unfolding lebesgue_def by simp @@ -114,10 +114,33 @@ qed (auto intro: LIMSEQ_indicator_UN simp: cube_def) qed simp +lemma suminf_SUP_eq: + fixes f :: "nat \ nat \ extreal" + assumes "\i. incseq (\n. f n i)" "\n i. 0 \ f n i" + shows "(\i. SUP n. f n i) = (SUP n. \i. f n i)" +proof - + { fix n :: nat + have "(\iix. 0 :: real)" by (simp add: fun_eq_iff) - show "measure lebesgue {} = 0" by (simp add: integral_0 * lebesgue_def) + show "positive lebesgue (measure lebesgue)" + proof (unfold positive_def, safe) + show "measure lebesgue {} = 0" by (simp add: integral_0 * lebesgue_def) + fix A assume "A \ sets lebesgue" + then show "0 \ measure lebesgue A" + unfolding lebesgue_def + by (auto intro!: le_SUPI2 integral_nonneg) + qed next show "countably_additive lebesgue (measure lebesgue)" proof (intro countably_additive_def[THEN iffD2] allI impI) @@ -130,23 +153,17 @@ assume "(\i. A i) \ sets lebesgue" then have UN_A[simp, intro]: "\i n. (indicator (\i. A i) :: _ \ real) integrable_on cube n" by (auto dest: lebesgueD) - show "(\\<^isub>\n. measure lebesgue (A n)) = measure lebesgue (\i. A i)" - proof (simp add: lebesgue_def, subst psuminf_SUP_eq) - fix n i show "Real (?m n i) \ Real (?m (Suc n) i)" - using cube_subset[of n "Suc n"] by (auto intro!: integral_subset_le) + show "(\n. measure lebesgue (A n)) = measure lebesgue (\i. A i)" + proof (simp add: lebesgue_def, subst suminf_SUP_eq, safe intro!: incseq_SucI) + fix i n show "extreal (?m n i) \ extreal (?m (Suc n) i)" + using cube_subset[of n "Suc n"] by (auto intro!: integral_subset_le incseq_SucI) next - show "(SUP n. \\<^isub>\i. Real (?m n i)) = (SUP n. Real (?M n UNIV))" - unfolding psuminf_def - proof (subst setsum_Real, (intro arg_cong[where f="SUPR UNIV"] ext ballI nn SUP_eq_LIMSEQ[THEN iffD2])+) - fix n :: nat show "mono (\m. \xx (\x ?M n UNIV" - using UN_A by (auto intro!: integral_nonneg) - fix m show "0 \ (\x extreal (?m n i)" + using rA unfolding lebesgue_def + by (auto intro!: le_SUPI2 integral_nonneg) + next + show "(SUP n. \i. extreal (?m n i)) = (SUP n. extreal (?M n UNIV))" + proof (intro arg_cong[where f="SUPR UNIV"] ext sums_unique[symmetric] sums_extreal[THEN iffD2] sums_def[THEN iffD2]) fix n have "\m. (UNION {.. sets lebesgue" using rA by auto from lebesgueD[OF this] @@ -171,8 +188,8 @@ ultimately show ?case using Suc A by (simp add: integral_add[symmetric]) qed auto } - ultimately show "(\m. \x ?M n UNIV" - by simp + ultimately show "(\m. \x = 0.. ?M n UNIV" + by (simp add: atLeast0LessThan) qed qed qed @@ -232,13 +249,11 @@ lemma lmeasure_iff_LIMSEQ: assumes "A \ sets lebesgue" "0 \ m" - shows "lebesgue.\ A = Real m \ (\n. integral (cube n) (indicator A :: _ \ real)) ----> m" + shows "lebesgue.\ A = extreal m \ (\n. integral (cube n) (indicator A :: _ \ real)) ----> m" proof (simp add: lebesgue_def, intro SUP_eq_LIMSEQ) show "mono (\n. integral (cube n) (indicator A::_=>real))" using cube_subset assms by (intro monoI integral_subset_le) (auto dest!: lebesgueD) - fix n show "0 \ integral (cube n) (indicator A::_=>real)" - using assms by (auto dest!: lebesgueD intro!: integral_nonneg) -qed fact +qed lemma has_integral_indicator_UNIV: fixes s A :: "'a::ordered_euclidean_space set" and x :: real @@ -260,7 +275,7 @@ lemma lmeasure_finite_has_integral: fixes s :: "'a::ordered_euclidean_space set" - assumes "s \ sets lebesgue" "lebesgue.\ s = Real m" "0 \ m" + assumes "s \ sets lebesgue" "lebesgue.\ s = extreal m" "0 \ m" shows "(indicator s has_integral m) UNIV" proof - let ?I = "indicator :: 'a set \ 'a \ real" @@ -302,12 +317,14 @@ unfolding m by (intro integrable_integral **) qed -lemma lmeasure_finite_integrable: assumes "s \ sets lebesgue" "lebesgue.\ s \ \" +lemma lmeasure_finite_integrable: assumes s: "s \ sets lebesgue" and "lebesgue.\ s \ \" shows "(indicator s :: _ \ real) integrable_on UNIV" proof (cases "lebesgue.\ s") - case (preal m) from lmeasure_finite_has_integral[OF `s \ sets lebesgue` this] + case (real m) + with lmeasure_finite_has_integral[OF `s \ sets lebesgue` this] + lebesgue.positive_measure[OF s] show ?thesis unfolding integrable_on_def by auto -qed (insert assms, auto) +qed (insert assms lebesgue.positive_measure[OF s], auto) lemma has_integral_lebesgue: assumes "((indicator s :: _\real) has_integral m) UNIV" shows "s \ sets lebesgue" @@ -321,7 +338,7 @@ qed lemma has_integral_lmeasure: assumes "((indicator s :: _\real) has_integral m) UNIV" - shows "lebesgue.\ s = Real m" + shows "lebesgue.\ s = extreal m" proof (intro lmeasure_iff_LIMSEQ[THEN iffD2]) let ?I = "indicator :: 'a set \ 'a \ real" show "s \ sets lebesgue" using has_integral_lebesgue[OF assms] . @@ -346,28 +363,28 @@ qed lemma has_integral_iff_lmeasure: - "(indicator A has_integral m) UNIV \ (A \ sets lebesgue \ 0 \ m \ lebesgue.\ A = Real m)" + "(indicator A has_integral m) UNIV \ (A \ sets lebesgue \ 0 \ m \ lebesgue.\ A = extreal m)" proof assume "(indicator A has_integral m) UNIV" with has_integral_lmeasure[OF this] has_integral_lebesgue[OF this] - show "A \ sets lebesgue \ 0 \ m \ lebesgue.\ A = Real m" + show "A \ sets lebesgue \ 0 \ m \ lebesgue.\ A = extreal m" by (auto intro: has_integral_nonneg) next - assume "A \ sets lebesgue \ 0 \ m \ lebesgue.\ A = Real m" + assume "A \ sets lebesgue \ 0 \ m \ lebesgue.\ A = extreal m" then show "(indicator A has_integral m) UNIV" by (intro lmeasure_finite_has_integral) auto qed lemma lmeasure_eq_integral: assumes "(indicator s::_\real) integrable_on UNIV" - shows "lebesgue.\ s = Real (integral UNIV (indicator s))" + shows "lebesgue.\ s = extreal (integral UNIV (indicator s))" using assms unfolding integrable_on_def proof safe fix y :: real assume "(indicator s has_integral y) UNIV" from this[unfolded has_integral_iff_lmeasure] integral_unique[OF this] - show "lebesgue.\ s = Real (integral UNIV (indicator s))" by simp + show "lebesgue.\ s = extreal (integral UNIV (indicator s))" by simp qed lemma lebesgue_simple_function_indicator: - fixes f::"'a::ordered_euclidean_space \ pextreal" + fixes f::"'a::ordered_euclidean_space \ extreal" assumes f:"simple_function lebesgue f" shows "f = (\x. (\y \ f ` UNIV. y * indicator (f -` {y}) x))" by (rule, subst lebesgue.simple_function_indicator_representation[OF f]) auto @@ -376,7 +393,7 @@ "(indicator s::_\real) integrable_on UNIV \ integral UNIV (indicator s) = real (lebesgue.\ s)" by (subst lmeasure_eq_integral) (auto intro!: integral_nonneg) -lemma lmeasure_finite: assumes "(indicator s::_\real) integrable_on UNIV" shows "lebesgue.\ s \ \" +lemma lmeasure_finite: assumes "(indicator s::_\real) integrable_on UNIV" shows "lebesgue.\ s \ \" using lmeasure_eq_integral[OF assms] by auto lemma negligible_iff_lebesgue_null_sets: @@ -409,37 +426,29 @@ shows "integral {a .. b} (\x. c) = content {a .. b} *\<^sub>R c" by (rule integral_unique) (rule has_integral_const) -lemma lmeasure_UNIV[intro]: "lebesgue.\ (UNIV::'a::ordered_euclidean_space set) = \" -proof (simp add: lebesgue_def SUP_\, intro allI impI) - fix x assume "x < \" - then obtain r where r: "x = Real r" "0 \ r" by (cases x) auto - then obtain n where n: "r < of_nat n" using ex_less_of_nat by auto - show "\i. x < Real (integral (cube i) (indicator UNIV::'a\real))" - proof (intro exI[of _ n]) - have [simp]: "indicator UNIV = (\x. 1)" by (auto simp: fun_eq_iff) - { fix m :: nat assume "0 < m" then have "real n \ (\x 0" then have "real n \ (\x Real (integral (cube n) (indicator UNIV::'a\real))" - by (auto simp: real_eq_of_nat[symmetric] cube_def content_closed_interval_cases) - finally show "x < Real (integral (cube n) (indicator UNIV::'a\real))" . - qed -qed +lemma lmeasure_UNIV[intro]: "lebesgue.\ (UNIV::'a::ordered_euclidean_space set) = \" +proof (simp add: lebesgue_def, intro SUP_PInfty bexI) + fix n :: nat + have "indicator UNIV = (\x::'a. 1 :: real)" by auto + moreover + { have "real n \ (2 * real n) ^ DIM('a)" + proof (cases n) + case 0 then show ?thesis by auto + next + case (Suc n') + have "real n \ (2 * real n)^1" by auto + also have "(2 * real n)^1 \ (2 * real n) ^ DIM('a)" + using Suc DIM_positive[where 'a='a] by (intro power_increasing) (auto simp: real_of_nat_Suc) + finally show ?thesis . + qed } + ultimately show "extreal (real n) \ extreal (integral (cube n) (indicator UNIV::'a\real))" + using integral_const DIM_positive[where 'a='a] + by (auto simp: cube_def content_closed_interval_cases setprod_constant) +qed simp lemma fixes a b ::"'a::ordered_euclidean_space" - shows lmeasure_atLeastAtMost[simp]: "lebesgue.\ {a..b} = Real (content {a..b})" + shows lmeasure_atLeastAtMost[simp]: "lebesgue.\ {a..b} = extreal (content {a..b})" proof - have "(indicator (UNIV \ {a..b})::_\real) integrable_on UNIV" unfolding integrable_indicator_UNIV by (simp add: integrable_const indicator_def_raw) @@ -467,7 +476,7 @@ lemma fixes a b :: real shows lmeasure_real_greaterThanAtMost[simp]: - "lebesgue.\ {a <.. b} = Real (if a \ b then b - a else 0)" + "lebesgue.\ {a <.. b} = extreal (if a \ b then b - a else 0)" proof cases assume "a < b" then have "lebesgue.\ {a <.. b} = lebesgue.\ {a .. b} - lebesgue.\ {a}" @@ -479,7 +488,7 @@ lemma fixes a b :: real shows lmeasure_real_atLeastLessThan[simp]: - "lebesgue.\ {a ..< b} = Real (if a \ b then b - a else 0)" + "lebesgue.\ {a ..< b} = extreal (if a \ b then b - a else 0)" proof cases assume "a < b" then have "lebesgue.\ {a ..< b} = lebesgue.\ {a .. b} - lebesgue.\ {b}" @@ -491,7 +500,7 @@ lemma fixes a b :: real shows lmeasure_real_greaterThanLessThan[simp]: - "lebesgue.\ {a <..< b} = Real (if a \ b then b - a else 0)" + "lebesgue.\ {a <..< b} = extreal (if a \ b then b - a else 0)" proof cases assume "a < b" then have "lebesgue.\ {a <..< b} = lebesgue.\ {a <.. b} - lebesgue.\ {b}" @@ -511,19 +520,16 @@ and measurable_lborel[simp]: "measurable lborel = measurable borel" by (simp_all add: measurable_def_raw lborel_def) -interpretation lborel: measure_space lborel +interpretation lborel: measure_space "lborel :: ('a::ordered_euclidean_space) measure_space" where "space lborel = UNIV" and "sets lborel = sets borel" and "measure lborel = lebesgue.\" and "measurable lborel = measurable borel" -proof - - show "measure_space lborel" - proof - show "countably_additive lborel (measure lborel)" - using lebesgue.ca unfolding countably_additive_def lborel_def - apply safe apply (erule_tac x=A in allE) by auto - qed (auto simp: lborel_def) -qed simp_all +proof (rule lebesgue.measure_space_subalgebra) + have "sigma_algebra (lborel::'a measure_space) \ sigma_algebra (borel::'a algebra)" + unfolding sigma_algebra_iff2 lborel_def by simp + then show "sigma_algebra (lborel::'a measure_space)" by simp default +qed auto interpretation lborel: sigma_finite_measure lborel where "space lborel = UNIV" @@ -536,7 +542,7 @@ show "range cube \ sets lborel" by (auto intro: borel_closed) { fix x have "\n. x\cube n" using mem_big_cube by auto } thus "(\i. cube i) = space lborel" by auto - show "\i. measure lborel (cube i) \ \" by (simp add: cube_def) + show "\i. measure lborel (cube i) \ \" by (simp add: cube_def) qed qed simp_all @@ -544,171 +550,221 @@ proof from lborel.sigma_finite guess A .. moreover then have "range A \ sets lebesgue" using lebesgueI_borel by blast - ultimately show "\A::nat \ 'b set. range A \ sets lebesgue \ (\i. A i) = space lebesgue \ (\i. lebesgue.\ (A i) \ \)" + ultimately show "\A::nat \ 'b set. range A \ sets lebesgue \ (\i. A i) = space lebesgue \ (\i. lebesgue.\ (A i) \ \)" by auto qed subsection {* Lebesgue integrable implies Gauge integrable *} +lemma positive_not_Inf: + "0 \ x \ x \ \ \ \x\ \ \" + by (cases x) auto + +lemma has_integral_cmult_real: + fixes c :: real + assumes "c \ 0 \ (f has_integral x) A" + shows "((\x. c * f x) has_integral c * x) A" +proof cases + assume "c \ 0" + from has_integral_cmul[OF assms[OF this], of c] show ?thesis + unfolding real_scaleR_def . +qed simp + lemma simple_function_has_integral: - fixes f::"'a::ordered_euclidean_space \ pextreal" + fixes f::"'a::ordered_euclidean_space \ extreal" assumes f:"simple_function lebesgue f" - and f':"\x. f x \ \" - and om:"\x\range f. lebesgue.\ (f -` {x} \ UNIV) = \ \ x = 0" + and f':"range f \ {0..<\}" + and om:"\x. x \ range f \ lebesgue.\ (f -` {x} \ UNIV) = \ \ x = 0" shows "((\x. real (f x)) has_integral (real (integral\<^isup>S lebesgue f))) UNIV" - unfolding simple_integral_def - apply(subst lebesgue_simple_function_indicator[OF f]) -proof - - case goal1 - have *:"\x. \y\range f. y * indicator (f -` {y}) x \ \" - "\x\range f. x * lebesgue.\ (f -` {x} \ UNIV) \ \" - using f' om unfolding indicator_def by auto - show ?case unfolding space_lebesgue real_of_pextreal_setsum'[OF *(1),THEN sym] - unfolding real_of_pextreal_setsum'[OF *(2),THEN sym] - unfolding real_of_pextreal_setsum space_lebesgue - apply(rule has_integral_setsum) - proof safe show "finite (range f)" using f by (auto dest: lebesgue.simple_functionD) - fix y::'a show "((\x. real (f y * indicator (f -` {f y}) x)) has_integral - real (f y * lebesgue.\ (f -` {f y} \ UNIV))) UNIV" - proof(cases "f y = 0") case False - have mea:"(indicator (f -` {f y}) ::_\real) integrable_on UNIV" - apply(rule lmeasure_finite_integrable) - using assms unfolding simple_function_def using False by auto - have *:"\x. real (indicator (f -` {f y}) x::pextreal) = (indicator (f -` {f y}) x)" - by (auto simp: indicator_def) - show ?thesis unfolding real_of_pextreal_mult[THEN sym] - apply(rule has_integral_cmul[where 'b=real, unfolded real_scaleR_def]) - unfolding Int_UNIV_right lmeasure_eq_integral[OF mea,THEN sym] - unfolding integral_eq_lmeasure[OF mea, symmetric] * - apply(rule integrable_integral) using mea . - qed auto + unfolding simple_integral_def space_lebesgue +proof (subst lebesgue_simple_function_indicator) + let "?M x" = "lebesgue.\ (f -` {x} \ UNIV)" + let "?F x" = "indicator (f -` {x})" + { fix x y assume "y \ range f" + from subsetD[OF f' this] have "y * ?F y x = extreal (real y * ?F y x)" + by (cases rule: extreal2_cases[of y "?F y x"]) + (auto simp: indicator_def one_extreal_def split: split_if_asm) } + moreover + { fix x assume x: "x\range f" + have "x * ?M x = real x * real (?M x)" + proof cases + assume "x \ 0" with om[OF x] have "?M x \ \" by auto + with subsetD[OF f' x] f[THEN lebesgue.simple_functionD(2)] show ?thesis + by (cases rule: extreal2_cases[of x "?M x"]) auto + qed simp } + ultimately + have "((\x. real (\y\range f. y * ?F y x)) has_integral real (\x\range f. x * ?M x)) UNIV \ + ((\x. \y\range f. real y * ?F y x) has_integral (\x\range f. real x * real (?M x))) UNIV" + by simp + also have \ + proof (intro has_integral_setsum has_integral_cmult_real lmeasure_finite_has_integral + real_of_extreal_pos lebesgue.positive_measure ballI) + show *: "finite (range f)" "\y. f -` {y} \ sets lebesgue" "\y. f -` {y} \ UNIV \ sets lebesgue" + using lebesgue.simple_functionD[OF f] by auto + fix y assume "real y \ 0" "y \ range f" + with * om[OF this(2)] show "lebesgue.\ (f -` {y}) = extreal (real (?M y))" + by (auto simp: extreal_real) qed -qed + finally show "((\x. real (\y\range f. y * ?F y x)) has_integral real (\x\range f. x * ?M x)) UNIV" . +qed fact lemma bounded_realI: assumes "\x\s. abs (x::real) \ B" shows "bounded s" unfolding bounded_def dist_real_def apply(rule_tac x=0 in exI) using assms by auto lemma simple_function_has_integral': - fixes f::"'a::ordered_euclidean_space \ pextreal" - assumes f:"simple_function lebesgue f" - and i: "integral\<^isup>S lebesgue f \ \" + fixes f::"'a::ordered_euclidean_space \ extreal" + assumes f: "simple_function lebesgue f" "\x. 0 \ f x" + and i: "integral\<^isup>S lebesgue f \ \" shows "((\x. real (f x)) has_integral (real (integral\<^isup>S lebesgue f))) UNIV" -proof- let ?f = "\x. if f x = \ then 0 else f x" - { fix x have "real (f x) = real (?f x)" by (cases "f x") auto } note * = this - have **:"{x. f x \ ?f x} = f -` {\}" by auto - have **:"lebesgue.\ {x\space lebesgue. f x \ ?f x} = 0" - using lebesgue.simple_integral_omega[OF assms] by(auto simp add:**) - show ?thesis apply(subst lebesgue.simple_integral_cong'[OF f _ **]) - apply(rule lebesgue.simple_function_compose1[OF f]) - unfolding * defer apply(rule simple_function_has_integral) - proof- - show "simple_function lebesgue ?f" - using lebesgue.simple_function_compose1[OF f] . - show "\x. ?f x \ \" by auto - show "\x\range ?f. lebesgue.\ (?f -` {x} \ UNIV) = \ \ x = 0" - proof (safe, simp, safe, rule ccontr) - fix y assume "f y \ \" "f y \ 0" - hence "(\x. if f x = \ then 0 else f x) -` {if f y = \ then 0 else f y} = f -` {f y}" - by (auto split: split_if_asm) - moreover assume "lebesgue.\ ((\x. if f x = \ then 0 else f x) -` {if f y = \ then 0 else f y}) = \" - ultimately have "lebesgue.\ (f -` {f y}) = \" by simp - moreover - have "f y * lebesgue.\ (f -` {f y}) \ \" using i f - unfolding simple_integral_def setsum_\ simple_function_def - by auto - ultimately have "f y = 0" by (auto split: split_if_asm) - then show False using `f y \ 0` by simp - qed +proof - + let ?f = "\x. if x \ f -` {\} then 0 else f x" + note f(1)[THEN lebesgue.simple_functionD(2)] + then have [simp, intro]: "\X. f -` X \ sets lebesgue" by auto + have f': "simple_function lebesgue ?f" + using f by (intro lebesgue.simple_function_If_set) auto + have rng: "range ?f \ {0..<\}" using f by auto + have "AE x in lebesgue. f x = ?f x" + using lebesgue.simple_integral_PInf[OF f i] + by (intro lebesgue.AE_I[where N="f -` {\} \ space lebesgue"]) auto + from f(1) f' this have eq: "integral\<^isup>S lebesgue f = integral\<^isup>S lebesgue ?f" + by (rule lebesgue.simple_integral_cong_AE) + have real_eq: "\x. real (f x) = real (?f x)" by auto + + show ?thesis + unfolding eq real_eq + proof (rule simple_function_has_integral[OF f' rng]) + fix x assume x: "x \ range ?f" and inf: "lebesgue.\ (?f -` {x} \ UNIV) = \" + have "x * lebesgue.\ (?f -` {x} \ UNIV) = (\\<^isup>S y. x * indicator (?f -` {x}) y \lebesgue)" + using f'[THEN lebesgue.simple_functionD(2)] + by (simp add: lebesgue.simple_integral_cmult_indicator) + also have "\ \ integral\<^isup>S lebesgue f" + using f'[THEN lebesgue.simple_functionD(2)] f + by (intro lebesgue.simple_integral_mono lebesgue.simple_function_mult lebesgue.simple_function_indicator) + (auto split: split_indicator) + finally show "x = 0" unfolding inf using i subsetD[OF rng x] by (auto split: split_if_asm) qed qed -lemma (in measure_space) positive_integral_monotone_convergence: - fixes f::"nat \ 'a \ pextreal" - assumes i: "\i. f i \ borel_measurable M" and mono: "\x. mono (\n. f n x)" - and lim: "\x. (\i. f i x) ----> u x" - shows "u \ borel_measurable M" - and "(\i. integral\<^isup>P M (f i)) ----> integral\<^isup>P M u" (is ?ilim) -proof - - from positive_integral_isoton[unfolded isoton_fun_expand isoton_iff_Lim_mono, of f u] - show ?ilim using mono lim i by auto - have "\x. (SUP i. f i x) = u x" using mono lim SUP_Lim_pextreal - unfolding fun_eq_iff mono_def by auto - moreover have "(\x. SUP i. f i x) \ borel_measurable M" - using i by auto - ultimately show "u \ borel_measurable M" by simp -qed +lemma real_of_extreal_positive_mono: + "\0 \ x; x \ y; y \ \\ \ real x \ real y" + by (cases rule: extreal2_cases[of x y]) auto lemma positive_integral_has_integral: - fixes f::"'a::ordered_euclidean_space => pextreal" - assumes f:"f \ borel_measurable lebesgue" - and int_om:"integral\<^isup>P lebesgue f \ \" - and f_om:"\x. f x \ \" (* TODO: remove this *) + fixes f :: "'a::ordered_euclidean_space \ extreal" + assumes f: "f \ borel_measurable lebesgue" "range f \ {0..<\}" "integral\<^isup>P lebesgue f \ \" shows "((\x. real (f x)) has_integral (real (integral\<^isup>P lebesgue f))) UNIV" -proof- let ?i = "integral\<^isup>P lebesgue f" - from lebesgue.borel_measurable_implies_simple_function_sequence[OF f] - guess u .. note conjunctD2[OF this,rule_format] note u = conjunctD2[OF this(1)] this(2) - let ?u = "\i x. real (u i x)" and ?f = "\x. real (f x)" - have u_simple:"\k. integral\<^isup>S lebesgue (u k) = integral\<^isup>P lebesgue (u k)" - apply(subst lebesgue.positive_integral_eq_simple_integral[THEN sym,OF u(1)]) .. - have int_u_le:"\k. integral\<^isup>S lebesgue (u k) \ integral\<^isup>P lebesgue f" - unfolding u_simple apply(rule lebesgue.positive_integral_mono) - using isoton_Sup[OF u(3)] unfolding le_fun_def by auto - have u_int_om:"\i. integral\<^isup>S lebesgue (u i) \ \" - proof- case goal1 thus ?case using int_u_le[of i] int_om by auto qed +proof - + from lebesgue.borel_measurable_implies_simple_function_sequence'[OF f(1)] + guess u . note u = this + have SUP_eq: "\x. (SUP i. u i x) = f x" + using u(4) f(2)[THEN subsetD] by (auto split: split_max) + let "?u i x" = "real (u i x)" + note u_eq = lebesgue.positive_integral_eq_simple_integral[OF u(1,5), symmetric] + { fix i + note u_eq + also have "integral\<^isup>P lebesgue (u i) \ (\\<^isup>+x. max 0 (f x) \lebesgue)" + by (intro lebesgue.positive_integral_mono) (auto intro: le_SUPI simp: u(4)[symmetric]) + finally have "integral\<^isup>S lebesgue (u i) \ \" + unfolding positive_integral_max_0 using f by auto } + note u_fin = this + then have u_int: "\i. (?u i has_integral real (integral\<^isup>S lebesgue (u i))) UNIV" + by (rule simple_function_has_integral'[OF u(1,5)]) + have "\x. \r\0. f x = extreal r" + proof + fix x from f(2) have "0 \ f x" "f x \ \" by (auto simp: subset_eq) + then show "\r\0. f x = extreal r" by (cases "f x") auto + qed + from choice[OF this] obtain f' where f': "f = (\x. extreal (f' x))" "\x. 0 \ f' x" by auto + + have "\i. \r. \x. 0 \ r x \ u i x = extreal (r x)" + proof + fix i show "\r. \x. 0 \ r x \ u i x = extreal (r x)" + proof (intro choice allI) + fix i x have "u i x \ \" using u(3)[of i] by (auto simp: image_iff) metis + then show "\r\0. u i x = extreal r" using u(5)[of i x] by (cases "u i x") auto + qed + qed + from choice[OF this] obtain u' where + u': "u = (\i x. extreal (u' i x))" "\i x. 0 \ u' i x" by (auto simp: fun_eq_iff) - note u_int = simple_function_has_integral'[OF u(1) this] - have "(\x. real (f x)) integrable_on UNIV \ - (\k. Integration.integral UNIV (\x. real (u k x))) ----> Integration.integral UNIV (\x. real (f x))" - apply(rule monotone_convergence_increasing) apply(rule,rule,rule u_int) - proof safe case goal1 show ?case apply(rule real_of_pextreal_mono) using u(2,3) by auto - next case goal2 show ?case using u(3) apply(subst lim_Real[THEN sym]) - prefer 3 apply(subst Real_real') defer apply(subst Real_real') - using isotone_Lim[OF u(3)[unfolded isoton_fun_expand, THEN spec]] using f_om u by auto - next case goal3 - show ?case apply(rule bounded_realI[where B="real (integral\<^isup>P lebesgue f)"]) - apply safe apply(subst abs_of_nonneg) apply(rule integral_nonneg,rule) apply(rule u_int) - unfolding integral_unique[OF u_int] defer apply(rule real_of_pextreal_mono[OF _ int_u_le]) - using u int_om by auto - qed note int = conjunctD2[OF this] + have convergent: "f' integrable_on UNIV \ + (\k. integral UNIV (u' k)) ----> integral UNIV f'" + proof (intro monotone_convergence_increasing allI ballI) + show int: "\k. (u' k) integrable_on UNIV" + using u_int unfolding integrable_on_def u' by auto + show "\k x. u' k x \ u' (Suc k) x" using u(2,3,5) + by (auto simp: incseq_Suc_iff le_fun_def image_iff u' intro!: real_of_extreal_positive_mono) + show "\x. (\k. u' k x) ----> f' x" + using SUP_eq u(2) + by (intro SUP_eq_LIMSEQ[THEN iffD1]) (auto simp: u' f' incseq_mono incseq_Suc_iff le_fun_def) + show "bounded {integral UNIV (u' k)|k. True}" + proof (safe intro!: bounded_realI) + fix k + have "\integral UNIV (u' k)\ = integral UNIV (u' k)" + by (intro abs_of_nonneg integral_nonneg int ballI u') + also have "\ = real (integral\<^isup>S lebesgue (u k))" + using u_int[THEN integral_unique] by (simp add: u') + also have "\ = real (integral\<^isup>P lebesgue (u k))" + using lebesgue.positive_integral_eq_simple_integral[OF u(1,5)] by simp + also have "\ \ real (integral\<^isup>P lebesgue f)" using f + by (auto intro!: real_of_extreal_positive_mono lebesgue.positive_integral_positive + lebesgue.positive_integral_mono le_SUPI simp: SUP_eq[symmetric]) + finally show "\integral UNIV (u' k)\ \ real (integral\<^isup>P lebesgue f)" . + qed + qed - have "(\i. integral\<^isup>S lebesgue (u i)) ----> ?i" unfolding u_simple - apply(rule lebesgue.positive_integral_monotone_convergence(2)) - apply(rule lebesgue.borel_measurable_simple_function[OF u(1)]) - using isotone_Lim[OF u(3)[unfolded isoton_fun_expand, THEN spec]] by auto - hence "(\i. real (integral\<^isup>S lebesgue (u i))) ----> real ?i" apply- - apply(subst lim_Real[THEN sym]) prefer 3 - apply(subst Real_real') defer apply(subst Real_real') - using u f_om int_om u_int_om by auto - note * = LIMSEQ_unique[OF this int(2)[unfolded integral_unique[OF u_int]]] - show ?thesis unfolding * by(rule integrable_integral[OF int(1)]) + have "integral\<^isup>P lebesgue f = extreal (integral UNIV f')" + proof (rule tendsto_unique[OF trivial_limit_sequentially]) + have "(\i. integral\<^isup>S lebesgue (u i)) ----> (SUP i. integral\<^isup>P lebesgue (u i))" + unfolding u_eq by (intro LIMSEQ_extreal_SUPR lebesgue.incseq_positive_integral u) + also note lebesgue.positive_integral_monotone_convergence_SUP + [OF u(2) lebesgue.borel_measurable_simple_function[OF u(1)] u(5), symmetric] + finally show "(\k. integral\<^isup>S lebesgue (u k)) ----> integral\<^isup>P lebesgue f" + unfolding SUP_eq . + + { fix k + have "0 \ integral\<^isup>S lebesgue (u k)" + using u by (auto intro!: lebesgue.simple_integral_positive) + then have "integral\<^isup>S lebesgue (u k) = extreal (real (integral\<^isup>S lebesgue (u k)))" + using u_fin by (auto simp: extreal_real) } + note * = this + show "(\k. integral\<^isup>S lebesgue (u k)) ----> extreal (integral UNIV f')" + using convergent using u_int[THEN integral_unique, symmetric] + by (subst *) (simp add: lim_extreal u') + qed + then show ?thesis using convergent by (simp add: f' integrable_integral) qed lemma lebesgue_integral_has_integral: - fixes f::"'a::ordered_euclidean_space => real" - assumes f:"integrable lebesgue f" + fixes f :: "'a::ordered_euclidean_space \ real" + assumes f: "integrable lebesgue f" shows "(f has_integral (integral\<^isup>L lebesgue f)) UNIV" -proof- let ?n = "\x. - min (f x) 0" and ?p = "\x. max (f x) 0" - have *:"f = (\x. ?p x - ?n x)" apply rule by auto - note f = integrableD[OF f] - show ?thesis unfolding lebesgue_integral_def apply(subst *) - proof(rule has_integral_sub) case goal1 - have *:"\x. Real (f x) \ \" by auto - note lebesgue.borel_measurable_Real[OF f(1)] - from positive_integral_has_integral[OF this f(2) *] - show ?case unfolding real_Real_max . - next case goal2 - have *:"\x. Real (- f x) \ \" by auto - note lebesgue.borel_measurable_uminus[OF f(1)] - note lebesgue.borel_measurable_Real[OF this] - from positive_integral_has_integral[OF this f(3) *] - show ?case unfolding real_Real_max minus_min_eq_max by auto - qed +proof - + let ?n = "\x. real (extreal (max 0 (- f x)))" and ?p = "\x. real (extreal (max 0 (f x)))" + have *: "f = (\x. ?p x - ?n x)" by (auto simp del: extreal_max) + { fix f have "(\\<^isup>+ x. extreal (f x) \lebesgue) = (\\<^isup>+ x. extreal (max 0 (f x)) \lebesgue)" + by (intro lebesgue.positive_integral_cong_pos) (auto split: split_max) } + note eq = this + show ?thesis + unfolding lebesgue_integral_def + apply (subst *) + apply (rule has_integral_sub) + unfolding eq[of f] eq[of "\x. - f x"] + apply (safe intro!: positive_integral_has_integral) + using integrableD[OF f] + by (auto simp: zero_extreal_def[symmetric] positive_integral_max_0 split: split_max + intro!: lebesgue.measurable_If lebesgue.borel_measurable_extreal) qed lemma lebesgue_positive_integral_eq_borel: - "f \ borel_measurable borel \ integral\<^isup>P lebesgue f = integral\<^isup>P lborel f" - by (auto intro!: lebesgue.positive_integral_subalgebra[symmetric]) default + assumes f: "f \ borel_measurable borel" + shows "integral\<^isup>P lebesgue f = integral\<^isup>P lborel f" +proof - + from f have "integral\<^isup>P lebesgue (\x. max 0 (f x)) = integral\<^isup>P lborel (\x. max 0 (f x))" + by (auto intro!: lebesgue.positive_integral_subalgebra[symmetric]) default + then show ?thesis unfolding positive_integral_max_0 . +qed lemma lebesgue_integral_eq_borel: assumes "f \ borel_measurable borel" @@ -771,7 +827,7 @@ have "sets ?G = sets (\\<^isub>M i\I. sigma \ space = UNIV::real set, sets = range lessThan, measure = lebesgue.\ \)" by (subst sigma_product_algebra_sigma_eq[of I "\_ i. {.. sets ?E" unfolding cube_def_raw by auto + show "incseq cube" using cube_subset_Suc by (auto intro!: incseq_SucI) { fix x have "\n. x \ cube n" using mem_big_cube[of x] by fastsimp } - then show "cube \ space ?E" by (intro isotoneI cube_subset_Suc) auto - { fix i show "lborel.\ (cube i) \ \" unfolding cube_def by auto } + then show "(\i. cube i) = space ?E" by auto + { fix i show "lborel.\ (cube i) \ \" unfolding cube_def by auto } show "A \ sets (sigma ?E)" "sets (sigma ?E) = sets lborel" "space ?E = space lborel" using assms by (simp_all add: borel_eq_atLeastAtMost) @@ -857,7 +914,7 @@ by (simp add: interval_ne_empty eucl_le[where 'a='a]) then have "lborel.\ X = (\x {a $$ x .. b $$ x})" by (auto simp: content_closed_interval eucl_le[where 'a='a] - intro!: Real_setprod ) + intro!: setprod_extreal[symmetric]) also have "\ = measure ?P (?T X)" unfolding * by (subst lborel_space.measure_times) auto finally show ?thesis . @@ -882,7 +939,7 @@ using lborel_eq_lborel_space[OF A] by simp lemma borel_fubini_positiv_integral: - fixes f :: "'a::ordered_euclidean_space \ pextreal" + fixes f :: "'a::ordered_euclidean_space \ extreal" assumes f: "f \ borel_measurable borel" shows "integral\<^isup>P lborel f = \\<^isup>+x. f (p2e x) \(lborel_space.P DIM('a))" proof (rule lborel_space.positive_integral_vimage[OF _ measure_preserving_p2e _]) diff -r 28b51effc5ed -r cdf7693bbe08 src/HOL/Probability/Measure.thy --- a/src/HOL/Probability/Measure.thy Mon Mar 14 14:37:47 2011 +0100 +++ b/src/HOL/Probability/Measure.thy Mon Mar 14 14:37:49 2011 +0100 @@ -76,7 +76,7 @@ lemma (in measure_space) measure_countably_additive: assumes "range A \ sets M" "disjoint_family A" - shows "psuminf (\i. \ (A i)) = \ (\i. A i)" + shows "(\i. \ (A i)) = \ (\i. A i)" proof - have "(\ i. A i) \ sets M" using assms(1) by (rule countable_UN) with ca assms show ?thesis by (simp add: countably_additive_def) @@ -94,13 +94,13 @@ interpret N: sigma_algebra N by (intro sigma_algebra_cong assms) show ?thesis proof - show "measure N {} = 0" using assms by auto + show "positive N (measure N)" using assms by (auto simp: positive_def) show "countably_additive N (measure N)" unfolding countably_additive_def proof safe fix A :: "nat \ 'a set" assume A: "range A \ sets N" "disjoint_family A" then have "\i. A i \ sets M" "(UNION UNIV A) \ sets M" unfolding assms by auto from measure_countably_additive[of A] A this[THEN assms(1)] - show "(\\<^isub>\n. measure N (A n)) = measure N (UNION UNIV A)" + show "(\n. measure N (A n)) = measure N (UNION UNIV A)" unfolding assms by simp qed qed @@ -124,51 +124,51 @@ have "b = a \ (b - a)" using assms by auto moreover have "{} = a \ (b - a)" by auto ultimately have "\ b = \ a + \ (b - a)" - using measure_additive[of a "b - a"] local.Diff[of b a] assms by auto - moreover have "\ (b - a) \ 0" using assms by auto + using measure_additive[of a "b - a"] Diff[of b a] assms by auto + moreover have "\ a + 0 \ \ a + \ (b - a)" using assms by (intro add_mono) auto ultimately show "\ a \ \ b" by auto qed lemma (in measure_space) measure_compl: - assumes s: "s \ sets M" and fin: "\ s \ \" + assumes s: "s \ sets M" and fin: "\ s \ \" shows "\ (space M - s) = \ (space M) - \ s" proof - have s_less_space: "\ s \ \ (space M)" using s by (auto intro!: measure_mono sets_into_space) - + from s have "0 \ \ s" by auto have "\ (space M) = \ (s \ (space M - s))" using s by (metis Un_Diff_cancel Un_absorb1 s sets_into_space) also have "... = \ s + \ (space M - s)" by (rule additiveD [OF additive]) (auto simp add: s) finally have "\ (space M) = \ s + \ (space M - s)" . - thus ?thesis - unfolding minus_pextreal_eq2[OF s_less_space fin] - by (simp add: ac_simps) + then show ?thesis + using fin `0 \ \ s` + unfolding extreal_eq_minus_iff by (auto simp: ac_simps) qed lemma (in measure_space) measure_Diff: - assumes finite: "\ B \ \" + assumes finite: "\ B \ \" and measurable: "A \ sets M" "B \ sets M" "B \ A" shows "\ (A - B) = \ A - \ B" proof - - have *: "(A - B) \ B = A" using `B \ A` by auto - - have "\ ((A - B) \ B) = \ (A - B) + \ B" - using measurable by (rule_tac measure_additive[symmetric]) auto - thus ?thesis unfolding * using `\ B \ \` - by (simp add: pextreal_cancel_plus_minus) + have "0 \ \ B" using assms by auto + have "(A - B) \ B = A" using `B \ A` by auto + then have "\ A = \ ((A - B) \ B)" by simp + also have "\ = \ (A - B) + \ B" + using measurable by (subst measure_additive[symmetric]) auto + finally show "\ (A - B) = \ A - \ B" + unfolding extreal_eq_minus_iff + using finite `0 \ \ B` by auto qed lemma (in measure_space) measure_countable_increasing: assumes A: "range A \ sets M" and A0: "A 0 = {}" - and ASuc: "\n. A n \ A (Suc n)" + and ASuc: "\n. A n \ A (Suc n)" shows "(SUP n. \ (A n)) = \ (\i. A i)" proof - - { - fix n - have "\ (A n) = - setsum (\ \ (\i. A (Suc i) - A i)) {.. (A n) = (\i (A (Suc i) - A i))" proof (induct n) case 0 thus ?case by (auto simp add: A0) next @@ -199,92 +199,83 @@ by (metis A Diff range_subsetD) have A2: "(\i. A (Suc i) - A i) \ sets M" by (blast intro: range_subsetD [OF A]) - have "psuminf ( (\i. \ (A (Suc i) - A i))) = \ (\i. A (Suc i) - A i)" + have "(SUP n. \i (A (Suc i) - A i)) = (\i. \ (A (Suc i) - A i))" + using A by (auto intro!: suminf_extreal_eq_SUPR[symmetric]) + also have "\ = \ (\i. A (Suc i) - A i)" by (rule measure_countably_additive) (auto simp add: disjoint_family_Suc ASuc A1 A2) also have "... = \ (\i. A i)" by (simp add: Aeq) - finally have "psuminf (\i. \ (A (Suc i) - A i)) = \ (\i. A i)" . - thus ?thesis - by (auto simp add: Meq psuminf_def) + finally have "(SUP n. \i (A (Suc i) - A i)) = \ (\i. A i)" . + then show ?thesis by (auto simp add: Meq) qed lemma (in measure_space) continuity_from_below: - assumes A: "range A \ sets M" - and ASuc: "!!n. A n \ A (Suc n)" + assumes A: "range A \ sets M" and "incseq A" shows "(SUP n. \ (A n)) = \ (\i. A i)" proof - have *: "(SUP n. \ (nat_case {} A (Suc n))) = (SUP n. \ (nat_case {} A n))" - apply (rule Sup_mono_offset_Suc) - apply (rule measure_mono) - using assms by (auto split: nat.split) - + using A by (auto intro!: SUPR_eq exI split: nat.split) have ueq: "(\i. nat_case {} A i) = (\i. A i)" by (auto simp add: split: nat.splits) have meq: "\n. \ (A n) = (\ \ nat_case {} A) (Suc n)" by simp have "(SUP n. \ (nat_case {} A n)) = \ (\i. nat_case {} A i)" - by (rule measure_countable_increasing) - (auto simp add: range_subsetD [OF A] subsetD [OF ASuc] split: nat.splits) + using range_subsetD[OF A] incseq_SucD[OF `incseq A`] + by (force split: nat.splits intro!: measure_countable_increasing) also have "\ (\i. nat_case {} A i) = \ (\i. A i)" by (simp add: ueq) finally have "(SUP n. \ (nat_case {} A n)) = \ (\i. A i)" . thus ?thesis unfolding meq * comp_def . qed -lemma (in measure_space) measure_up: - assumes "\i. B i \ sets M" "B \ P" - shows "(\i. \ (B i)) \ \ P" - using assms unfolding isoton_def - by (auto intro!: measure_mono continuity_from_below) +lemma (in measure_space) measure_incseq: + assumes "range B \ sets M" "incseq B" + shows "incseq (\i. \ (B i))" + using assms by (auto simp: incseq_def intro!: measure_mono) -lemma (in measure_space) continuity_from_below': - assumes A: "range A \ sets M" "\i. A i \ A (Suc i)" - shows "(\i. (\ (A i))) ----> (\ (\i. A i))" -proof- let ?A = "\i. A i" - have " (\i. \ (A i)) \ \ ?A" apply(rule measure_up) - using assms unfolding complete_lattice_class.isoton_def by auto - thus ?thesis by(rule isotone_Lim(1)) -qed +lemma (in measure_space) continuity_from_below_Lim: + assumes A: "range A \ sets M" "incseq A" + shows "(\i. (\ (A i))) ----> \ (\i. A i)" + using LIMSEQ_extreal_SUPR[OF measure_incseq, OF A] + continuity_from_below[OF A] by simp + +lemma (in measure_space) measure_decseq: + assumes "range B \ sets M" "decseq B" + shows "decseq (\i. \ (B i))" + using assms by (auto simp: decseq_def intro!: measure_mono) lemma (in measure_space) continuity_from_above: - assumes A: "range A \ sets M" - and mono_Suc: "\n. A (Suc n) \ A n" - and finite: "\i. \ (A i) \ \" + assumes A: "range A \ sets M" and "decseq A" + and finite: "\i. \ (A i) \ \" shows "(INF n. \ (A n)) = \ (\i. A i)" proof - - { fix n have "A n \ A 0" using mono_Suc by (induct n) auto } - note mono = this - have le_MI: "\ (\i. A i) \ \ (A 0)" using A by (auto intro!: measure_mono) - hence *: "\ (\i. A i) \ \" using finite[of 0] by auto + hence *: "\ (\i. A i) \ \" using finite[of 0] by auto + + have A0: "0 \ \ (A 0)" using A by auto - have le_IM: "(INF n. \ (A n)) \ \ (A 0)" - by (rule INF_leI) simp - - have "\ (A 0) - (INF n. \ (A n)) = (SUP n. \ (A 0 - A n))" - unfolding pextreal_SUP_minus[symmetric] - using mono A finite by (subst measure_Diff) auto + have "\ (A 0) - (INF n. \ (A n)) = \ (A 0) + (SUP n. - \ (A n))" + by (simp add: extreal_SUPR_uminus minus_extreal_def) + also have "\ = (SUP n. \ (A 0) - \ (A n))" + unfolding minus_extreal_def using A0 assms + by (subst SUPR_extreal_add) (auto simp add: measure_decseq) + also have "\ = (SUP n. \ (A 0 - A n))" + using A finite `decseq A`[unfolded decseq_def] by (subst measure_Diff) auto also have "\ = \ (\i. A 0 - A i)" proof (rule continuity_from_below) show "range (\n. A 0 - A n) \ sets M" using A by auto - show "\n. A 0 - A n \ A 0 - A (Suc n)" - using mono_Suc by auto + show "incseq (\n. A 0 - A n)" + using `decseq A` by (auto simp add: incseq_def decseq_def) qed also have "\ = \ (A 0) - \ (\i. A i)" - using mono A finite * by (simp, subst measure_Diff) auto + using A finite * by (simp, subst measure_Diff) auto finally show ?thesis - by (rule pextreal_diff_eq_diff_imp_eq[OF finite[of 0] le_IM le_MI]) + unfolding extreal_minus_eq_minus_iff using finite A0 by auto qed -lemma (in measure_space) measure_down: - assumes "\i. B i \ sets M" "B \ P" - and finite: "\i. \ (B i) \ \" - shows "(\i. \ (B i)) \ \ P" - using assms unfolding antiton_def - by (auto intro!: measure_mono continuity_from_above) lemma (in measure_space) measure_insert: assumes sets: "{x} \ sets M" "A \ sets M" and "x \ A" shows "\ (insert x A) = \ {x} + \ A" @@ -293,109 +284,26 @@ from measure_additive[OF sets this] show ?thesis by simp qed -lemma (in measure_space) measure_finite_singleton: - assumes fin: "finite S" - and ssets: "\x. x \ S \ {x} \ sets M" - shows "\ S = (\x\S. \ {x})" +lemma (in measure_space) measure_setsum: + assumes "finite S" and "\i. i \ S \ A i \ sets M" + assumes disj: "disjoint_family_on A S" + shows "(\i\S. \ (A i)) = \ (\i\S. A i)" using assms proof induct - case (insert x S) - have *: "\ S = (\x\S. \ {x})" "{x} \ sets M" - using insert.prems by (blast intro!: insert.hyps(3))+ - - have "(\x\S. {x}) \ sets M" - using insert.prems `finite S` by (blast intro!: finite_UN) - hence "S \ sets M" by auto - from measure_insert[OF `{x} \ sets M` this `x \ S`] - show ?case using `x \ S` `finite S` * by simp + case (insert i S) + then have "(\i\S. \ (A i)) = \ (\a\S. A a)" + by (auto intro: disjoint_family_on_mono) + moreover have "A i \ (\a\S. A a) = {}" + using `disjoint_family_on A (insert i S)` `i \ S` + by (auto simp: disjoint_family_on_def) + ultimately show ?case using insert + by (auto simp: measure_additive finite_UN) qed simp -lemma (in measure_space) measure_finitely_additive': - assumes "f \ ({..< n :: nat} \ sets M)" - assumes "\ a b. \a < n ; b < n ; a \ b\ \ f a \ f b = {}" - assumes "s = \ (f ` {..< n})" - shows "(\i \ f) i) = \ s" -proof - - def f' == "\ i. (if i < n then f i else {})" - have rf: "range f' \ sets M" unfolding f'_def - using assms empty_sets by auto - have df: "disjoint_family f'" unfolding f'_def disjoint_family_on_def - using assms by simp - have "\ range f' = (\ i \ {..< n}. f i)" - unfolding f'_def by auto - then have "\ s = \ (\ range f')" - using assms by simp - then have part1: "(\\<^isub>\ i. \ (f' i)) = \ s" - using df rf ca[unfolded countably_additive_def, rule_format, of f'] - by auto - - have "(\\<^isub>\ i. \ (f' i)) = (\ i< n. \ (f' i))" - by (rule psuminf_finite) (simp add: f'_def) - also have "\ = (\i (f i))" - unfolding f'_def by auto - finally have part2: "(\\<^isub>\ i. \ (f' i)) = (\i (f i))" by simp - show ?thesis using part1 part2 by auto -qed - - -lemma (in measure_space) measure_finitely_additive: - assumes "finite r" - assumes "r \ sets M" - assumes d: "\ a b. \a \ r ; b \ r ; a \ b\ \ a \ b = {}" - shows "(\ i \ r. \ i) = \ (\ r)" -using assms -proof - - (* counting the elements in r is enough *) - let ?R = "{.. ?R \ sets M" using assms by auto - have disj: "\ a b. \a \ ?R ; b \ ?R ; a \ b\ \ f a \ f b = {}" - proof - - fix a b assume asm: "a \ ?R" "b \ ?R" "a \ b" - hence neq: "f a \ f b" using f[unfolded inj_on_def, rule_format] by blast - from asm have "f a \ r" "f b \ r" using f by auto - thus "f a \ f b = {}" using d[of "f a" "f b"] f using neq by auto - qed - have "(\ r) = (\ i \ ?R. f i)" - using f by auto - hence "\ (\ r)= \ (\ i \ ?R. f i)" by simp - also have "\ = (\ i \ ?R. \ (f i))" - using measure_finitely_additive'[OF f_into_sets disj] by simp - also have "\ = (\ a \ r. \ a)" - using f[rule_format] setsum_reindex[of f ?R "\ a. \ a"] by auto - finally show ?thesis by simp -qed - -lemma (in measure_space) measure_finitely_additive'': - assumes "finite s" - assumes "\ i. i \ s \ a i \ sets M" - assumes d: "disjoint_family_on a s" - shows "(\ i \ s. \ (a i)) = \ (\ i \ s. a i)" -using assms -proof - - (* counting the elements in s is enough *) - let ?R = "{..< card s}" - obtain f where f: "f ` ?R = s" "inj_on f ?R" - using ex_bij_betw_nat_finite[unfolded bij_betw_def, OF `finite s`] - unfolding atLeast0LessThan by auto - hence f_into_sets: "a \ f \ ?R \ sets M" using assms unfolding o_def by auto - have disj: "\ i j. \i \ ?R ; j \ ?R ; i \ j\ \ (a \ f) i \ (a \ f) j = {}" - proof - - fix i j assume asm: "i \ ?R" "j \ ?R" "i \ j" - hence neq: "f i \ f j" using f[unfolded inj_on_def, rule_format] by blast - from asm have "f i \ s" "f j \ s" using f by auto - thus "(a \ f) i \ (a \ f) j = {}" - using d f neq unfolding disjoint_family_on_def by auto - qed - have "(\ i \ s. a i) = (\ i \ f ` ?R. a i)" using f by auto - hence "(\ i \ s. a i) = (\ i \ ?R. a (f i))" by auto - hence "\ (\ i \ s. a i) = (\ i \ ?R. \ (a (f i)))" - using measure_finitely_additive'[OF f_into_sets disj] by simp - also have "\ = (\ i \ s. \ (a i))" - using f[rule_format] setsum_reindex[of f ?R "\ i. \ (a i)"] by auto - finally show ?thesis by simp -qed +lemma (in measure_space) measure_finite_singleton: + assumes "finite S" "\x. x \ S \ {x} \ sets M" + shows "\ S = (\x\S. \ {x})" + using measure_setsum[of S "\x. {x}", OF assms] + by (auto simp: disjoint_family_on_def) lemma finite_additivity_sufficient: assumes "sigma_algebra M" @@ -405,7 +313,7 @@ interpret sigma_algebra M by fact show ?thesis proof - show [simp]: "measure M {} = 0" using pos by (simp add: positive_def) + show [simp]: "positive M (measure M)" using pos by (simp add: positive_def) show "countably_additive M (measure M)" proof (auto simp add: countably_additive_def) fix A :: "nat \ 'a set" @@ -434,12 +342,12 @@ by blast qed then obtain N where N: "\m\N. A m = {}" by blast - then have "\m\N. measure M (A m) = 0" by simp - then have "(\\<^isub>\ n. measure M (A n)) = setsum (\m. measure M (A m)) {..m\N. measure M (A m) = 0" using pos[unfolded positive_def] by simp + then have "(\n. measure M (A n)) = (\mi (\ x i\<^isub>\ n. measure M (A n)) = measure M (\i. A i)" . + finally show "(\n. measure M (A n)) = measure M (\i. A i)" . qed qed qed lemma (in measure_space) measure_setsum_split: - assumes "finite r" and "a \ sets M" and br_in_M: "b ` r \ sets M" - assumes "(\i \ r. b i) = space M" - assumes "disjoint_family_on b r" - shows "\ a = (\ i \ r. \ (a \ (b i)))" + assumes "finite S" and "A \ sets M" and br_in_M: "B ` S \ sets M" + assumes "(\i\S. B i) = space M" + assumes "disjoint_family_on B S" + shows "\ A = (\i\S. \ (A \ (B i)))" proof - - have *: "\ a = \ (\i \ r. a \ b i)" + have *: "\ A = \ (\i\S. A \ B i)" using assms by auto show ?thesis unfolding * - proof (rule measure_finitely_additive''[symmetric]) - show "finite r" using `finite r` by auto - { fix i assume "i \ r" - hence "b i \ sets M" using br_in_M by auto - thus "a \ b i \ sets M" using `a \ sets M` by auto - } - show "disjoint_family_on (\i. a \ b i) r" - using `disjoint_family_on b r` + proof (rule measure_setsum[symmetric]) + show "disjoint_family_on (\i. A \ B i) S" + using `disjoint_family_on B S` unfolding disjoint_family_on_def by auto - qed + qed (insert assms, auto) qed lemma (in measure_space) measure_subadditive: @@ -506,7 +409,7 @@ lemma (in measure_space) measure_eq_0: assumes "N \ sets M" and "\ N = 0" and "K \ N" and "K \ sets M" shows "\ K = 0" -using measure_mono[OF assms(3,4,1)] assms(2) by auto + using measure_mono[OF assms(3,4,1)] assms(2) positive_measure[OF assms(4)] by auto lemma (in measure_space) measure_finitely_subadditive: assumes "finite I" "A ` I \ sets M" @@ -523,35 +426,38 @@ lemma (in measure_space) measure_countably_subadditive: assumes "range f \ sets M" - shows "\ (\i. f i) \ (\\<^isub>\ i. \ (f i))" + shows "\ (\i. f i) \ (\i. \ (f i))" proof - have "\ (\i. f i) = \ (\i. disjointed f i)" unfolding UN_disjointed_eq .. - also have "\ = (\\<^isub>\ i. \ (disjointed f i))" + also have "\ = (\i. \ (disjointed f i))" using range_disjointed_sets[OF assms] measure_countably_additive by (simp add: disjoint_family_disjointed comp_def) - also have "\ \ (\\<^isub>\ i. \ (f i))" - proof (rule psuminf_le, rule measure_mono) - fix i show "disjointed f i \ f i" by (rule disjointed_subset) - show "f i \ sets M" "disjointed f i \ sets M" - using assms range_disjointed_sets[OF assms] by auto - qed + also have "\ \ (\i. \ (f i))" + using range_disjointed_sets[OF assms] assms + by (auto intro!: suminf_le_pos measure_mono positive_measure disjointed_subset) finally show ?thesis . qed lemma (in measure_space) measure_UN_eq_0: - assumes "\ i :: nat. \ (N i) = 0" and "range N \ sets M" + assumes "\i::nat. \ (N i) = 0" and "range N \ sets M" shows "\ (\ i. N i) = 0" -using measure_countably_subadditive[OF assms(2)] assms(1) by auto +proof - + have "0 \ \ (\ i. N i)" using assms by auto + moreover have "\ (\ i. N i) \ 0" + using measure_countably_subadditive[OF assms(2)] assms(1) by simp + ultimately show ?thesis by simp +qed lemma (in measure_space) measure_inter_full_set: - assumes "S \ sets M" "T \ sets M" and not_\: "\ (T - S) \ \" + assumes "S \ sets M" "T \ sets M" and fin: "\ (T - S) \ \" assumes T: "\ T = \ (space M)" shows "\ (S \ T) = \ S" proof (rule antisym) show " \ (S \ T) \ \ S" using assms by (auto intro!: measure_mono) + have pos: "0 \ \ (T - S)" using assms by auto show "\ S \ \ (S \ T)" proof (rule ccontr) assume contr: "\ ?thesis" @@ -560,7 +466,7 @@ also have "\ \ \ (T - S) + \ (S \ T)" using assms by (auto intro!: measure_subadditive) also have "\ < \ (T - S) + \ S" - by (rule pextreal_less_add[OF not_\]) (insert contr, auto) + using fin contr pos by (intro extreal_less_add) auto also have "\ = \ (T \ S)" using assms by (subst measure_additive) auto also have "\ \ \ (space M)" @@ -572,11 +478,11 @@ lemma measure_unique_Int_stable: fixes E :: "('a, 'b) algebra_scheme" and A :: "nat \ 'a set" assumes "Int_stable E" - and A: "range A \ sets E" "A \ space E" + and A: "range A \ sets E" "incseq A" "(\i. A i) = space E" and M: "measure_space \space = space E, sets = sets (sigma E), measure = \\" (is "measure_space ?M") and N: "measure_space \space = space E, sets = sets (sigma E), measure = \\" (is "measure_space ?N") and eq: "\X. X \ sets E \ \ X = \ X" - and finite: "\i. \ (A i) \ \" + and finite: "\i. \ (A i) \ \" assumes "X \ sets (sigma E)" shows "\ X = \ X" proof - @@ -585,9 +491,9 @@ where "space ?M = space E" and "sets ?M = sets (sigma E)" and "measure ?M = \" by (simp_all add: M) interpret N: measure_space ?N where "space ?N = space E" and "sets ?N = sets (sigma E)" and "measure ?N = \" by (simp_all add: N) - { fix F assume "F \ sets E" and "\ F \ \" + { fix F assume "F \ sets E" and "\ F \ \" then have [intro]: "F \ sets (sigma E)" by auto - have "\ F \ \" using `\ F \ \` `F \ sets E` eq by simp + have "\ F \ \" using `\ F \ \` `F \ sets E` eq by simp interpret D: dynkin_system "\space=space E, sets=?D F\" proof (rule dynkin_systemI, simp_all) fix A assume "A \ sets (sigma E) \ \ (F \ A) = \ (F \ A)" @@ -602,14 +508,14 @@ and [intro]: "F \ A \ sets (sigma E)" using `F \ sets E` M.sets_into_space by auto have "\ (F \ A) \ \ F" by (auto intro!: N.measure_mono) - then have "\ (F \ A) \ \" using `\ F \ \` by auto + then have "\ (F \ A) \ \" using `\ F \ \` by auto have "\ (F \ A) \ \ F" by (auto intro!: M.measure_mono) - then have "\ (F \ A) \ \" using `\ F \ \` by auto + then have "\ (F \ A) \ \" using `\ F \ \` by auto then have "\ (F \ (space (sigma E) - A)) = \ F - \ (F \ A)" unfolding ** using `F \ A \ sets (sigma E)` by (auto intro!: M.measure_Diff) also have "\ = \ F - \ (F \ A)" using eq `F \ sets E` * by simp also have "\ = \ (F \ (space (sigma E) - A))" unfolding ** - using `F \ A \ sets (sigma E)` `\ (F \ A) \ \` by (auto intro!: N.measure_Diff[symmetric]) + using `F \ A \ sets (sigma E)` `\ (F \ A) \ \` by (auto intro!: N.measure_Diff[symmetric]) finally show "space E - A \ sets (sigma E) \ \ (F \ (space E - A)) = \ (F \ (space E - A))" using * by auto next @@ -630,15 +536,13 @@ have "\D. D \ sets (sigma E) \ \ (F \ D) = \ (F \ D)" by (subst (asm) *) auto } note * = this - { fix i have "\ (A i \ X) = \ (A i \ X)" + let "?A i" = "A i \ X" + have A': "range ?A \ sets (sigma E)" "incseq ?A" + using A(1,2) `X \ sets (sigma E)` by (auto simp: incseq_def) + { fix i have "\ (?A i) = \ (?A i)" using *[of "A i" X] `X \ sets (sigma E)` A finite by auto } - moreover - have "(\i. A i \ X) \ X" - using `X \ sets (sigma E)` M.sets_into_space A - by (auto simp: isoton_def) - then have "(\i. \ (A i \ X)) \ \ X" "(\i. \ (A i \ X)) \ \ X" - using `X \ sets (sigma E)` A by (auto intro!: M.measure_up N.measure_up M.Int simp: subset_eq) - ultimately show ?thesis by (simp add: isoton_def) + with M.continuity_from_below[OF A'] N.continuity_from_below[OF A'] + show ?thesis using A(3) `X \ sets (sigma E)` by auto qed section "@{text \}-null sets" @@ -650,10 +554,10 @@ shows "N \ N' \ null_sets" proof (intro conjI CollectI) show "N \ N' \ sets M" using assms by auto - have "\ (N \ N') \ \ N + \ N'" + then have "0 \ \ (N \ N')" by simp + moreover have "\ (N \ N') \ \ N + \ N'" using assms by (intro measure_subadditive) auto - then show "\ (N \ N') = 0" - using assms by auto + ultimately show "\ (N \ N') = 0" using assms by auto qed lemma UN_from_nat: "(\i. N i) = (\i. N (Countable.from_nat i))" @@ -669,11 +573,11 @@ shows "(\i. N i) \ null_sets" proof (intro conjI CollectI) show "(\i. N i) \ sets M" using assms by auto - have "\ (\i. N i) \ (\\<^isub>\ n. \ (N (Countable.from_nat n)))" + then have "0 \ \ (\i. N i)" by simp + moreover have "\ (\i. N i) \ (\n. \ (N (Countable.from_nat n)))" unfolding UN_from_nat[of N] using assms by (intro measure_countably_subadditive) auto - then show "\ (\i. N i) = 0" - using assms by auto + ultimately show "\ (\i. N i) = 0" using assms by auto qed lemma (in measure_space) null_sets_finite_UN: @@ -681,10 +585,10 @@ shows "(\i\S. A i) \ null_sets" proof (intro CollectI conjI) show "(\i\S. A i) \ sets M" using assms by (intro finite_UN) auto - have "\ (\i\S. A i) \ (\i\S. \ (A i))" + then have "0 \ \ (\i\S. A i)" by simp + moreover have "\ (\i\S. A i) \ (\i\S. \ (A i))" using assms by (intro measure_finitely_subadditive) auto - then show "\ (\i\S. A i) = 0" - using assms by auto + ultimately show "\ (\i\S. A i) = 0" using assms by auto qed lemma (in measure_space) null_set_Int1: @@ -731,6 +635,23 @@ almost_everywhere :: "('a \ bool) \ bool" (binder "AE " 10) where "almost_everywhere P \ (\N\null_sets. { x \ space M. \ P x } \ N)" +syntax + "_almost_everywhere" :: "'a \ ('a, 'b) measure_space_scheme \ ('a \ bool) \ bool" ("AE _ in _. _" [0,0,10] 10) + +translations + "AE x in M. P" == "CONST measure_space.almost_everywhere M (%x. P)" + +lemma (in measure_space) AE_cong_measure: + assumes "\A. A \ sets M \ measure N A = \ A" "sets N = sets M" "space N = space M" + shows "(AE x in N. P x) \ (AE x. P x)" +proof - + interpret N: measure_space N + by (rule measure_space_cong) fact+ + show ?thesis + unfolding N.almost_everywhere_def almost_everywhere_def + by (auto simp: assms) +qed + lemma (in measure_space) AE_I': "N \ null_sets \ {x\space M. \ P x} \ N \ (AE x. P x)" unfolding almost_everywhere_def by auto @@ -741,13 +662,19 @@ proof assume "AE x. P x" then obtain N where N: "N \ sets M" "?P \ N" "\ N = 0" unfolding almost_everywhere_def by auto + have "0 \ \ ?P" using assms by simp moreover have "\ ?P \ \ N" using assms N(1,2) by (auto intro: measure_mono) - ultimately show "?P \ null_sets" using assms by auto + ultimately have "\ ?P = 0" unfolding `\ N = 0` by auto + then show "?P \ null_sets" using assms by simp next assume "?P \ null_sets" with assms show "AE x. P x" by (auto intro: AE_I') qed +lemma (in measure_space) AE_iff_measurable: + "N \ sets M \ {x\space M. \ P x} = N \ (AE x. P x) \ \ N = 0" + using AE_iff_null_set[of P] by simp + lemma (in measure_space) AE_True[intro, simp]: "AE x. True" unfolding almost_everywhere_def by auto @@ -760,13 +687,9 @@ assumes "AE x. P x" "{x\space M. P x} \ sets M" shows "\ {x\space M. \ P x} = 0" (is "\ ?P = 0") proof - - obtain A where A: "?P \ A" "A \ sets M" "\ A = 0" - using assms by (auto elim!: AE_E) - have "?P = space M - {x\space M. P x}" by auto - then have "?P \ sets M" using assms by auto - with assms `?P \ A` `A \ sets M` have "\ ?P \ \ A" - by (auto intro!: measure_mono) - then show "\ ?P = 0" using A by simp + have "{x\space M. \ P x} = space M - {x\space M. P x}" + by auto + with AE_iff_null_set[of P] assms show ?thesis by auto qed lemma (in measure_space) AE_I: @@ -788,8 +711,10 @@ show ?thesis proof (intro AE_I) - show "A \ B \ sets M" "\ (A \ B) = 0" using A B - using measure_subadditive[of A B] by auto + have "0 \ \ (A \ B)" using A B by auto + moreover have "\ (A \ B) \ 0" + using measure_subadditive[of A B] A B by auto + ultimately show "A \ B \ sets M" "\ (A \ B) = 0" using A B by auto show "{x\space M. \ Q x} \ A \ B" using P imp by auto qed @@ -818,8 +743,8 @@ "(\x. x \ space M \ P x \ Q x) \ (AE x. P x) \ (AE x. Q x)" by auto -lemma (in measure_space) all_AE_countable: - "(\i::'i::countable. AE x. P i x) \ (AE x. \i. P i x)" +lemma (in measure_space) AE_all_countable: + "(AE x. \i. P i x) \ (\i::'i::countable. AE x. P i x)" proof assume "\i. AE x. P i x" from this[unfolded almost_everywhere_def Bex_def, THEN choice] @@ -833,6 +758,10 @@ unfolding almost_everywhere_def by auto qed auto +lemma (in measure_space) AE_finite_all: + assumes f: "finite S" shows "(AE x. \i\S. P i x) \ (\i\S. AE x. P i x)" + using f by induct auto + lemma (in measure_space) restricted_measure_space: assumes "S \ sets M" shows "measure_space (restricted_space S)" @@ -840,7 +769,7 @@ unfolding measure_space_def measure_space_axioms_def proof safe show "sigma_algebra ?r" using restricted_sigma_algebra[OF assms] . - show "measure ?r {} = 0" by simp + show "positive ?r (measure ?r)" using `S \ sets M` by (auto simp: positive_def) show "countably_additive ?r (measure ?r)" unfolding countably_additive_def @@ -848,13 +777,38 @@ fix A :: "nat \ 'a set" assume *: "range A \ sets ?r" and **: "disjoint_family A" from restriction_in_sets[OF assms *[simplified]] ** - show "(\\<^isub>\ n. measure ?r (A n)) = measure ?r (\i. A i)" + show "(\n. measure ?r (A n)) = measure ?r (\i. A i)" using measure_countably_additive by simp qed qed +lemma (in measure_space) AE_restricted: + assumes "A \ sets M" + shows "(AE x in restricted_space A. P x) \ (AE x. x \ A \ P x)" +proof - + interpret R: measure_space "restricted_space A" + by (rule restricted_measure_space[OF `A \ sets M`]) + show ?thesis + proof + assume "AE x in restricted_space A. P x" + from this[THEN R.AE_E] guess N' . + then obtain N where "{x \ A. \ P x} \ A \ N" "\ (A \ N) = 0" "N \ sets M" + by auto + moreover then have "{x \ space M. \ (x \ A \ P x)} \ A \ N" + using `A \ sets M` sets_into_space by auto + ultimately show "AE x. x \ A \ P x" + using `A \ sets M` by (auto intro!: AE_I[where N="A \ N"]) + next + assume "AE x. x \ A \ P x" + from this[THEN AE_E] guess N . + then show "AE x in restricted_space A. P x" + using null_set_Int1[OF _ `A \ sets M`] `A \ sets M`[THEN sets_into_space] + by (auto intro!: R.AE_I[where N="A \ N"] simp: subset_eq) + qed +qed + lemma (in measure_space) measure_space_subalgebra: - assumes "sigma_algebra N" and [simp]: "sets N \ sets M" "space N = space M" + assumes "sigma_algebra N" and "sets N \ sets M" "space N = space M" and measure[simp]: "\X. X \ sets N \ measure N X = measure M X" shows "measure_space N" proof - @@ -864,13 +818,26 @@ from `sets N \ sets M` have "\A. range A \ sets N \ range A \ sets M" by blast then show "countably_additive N (measure N)" by (auto intro!: measure_countably_additive simp: countably_additive_def subset_eq) - qed simp + show "positive N (measure_space.measure N)" + using assms(2) by (auto simp add: positive_def) + qed +qed + +lemma (in measure_space) AE_subalgebra: + assumes ae: "AE x in N. P x" + and N: "sets N \ sets M" "space N = space M" "\A. A \ sets N \ measure N A = \ A" + and sa: "sigma_algebra N" + shows "AE x. P x" +proof - + interpret N: measure_space N using measure_space_subalgebra[OF sa N] . + from ae[THEN N.AE_E] guess N . + with N show ?thesis unfolding almost_everywhere_def by auto qed section "@{text \}-finite Measures" locale sigma_finite_measure = measure_space + - assumes sigma_finite: "\A::nat \ 'a set. range A \ sets M \ (\i. A i) = space M \ (\i. \ (A i) \ \)" + assumes sigma_finite: "\A::nat \ 'a set. range A \ sets M \ (\i. A i) = space M \ (\i. \ (A i) \ \)" lemma (in sigma_finite_measure) restricted_sigma_finite_measure: assumes "S \ sets M" @@ -881,9 +848,9 @@ show "measure_space ?r" using restricted_measure_space[OF assms] . next obtain A :: "nat \ 'a set" where - "range A \ sets M" "(\i. A i) = space M" "\i. \ (A i) \ \" + "range A \ sets M" "(\i. A i) = space M" "\i. \ (A i) \ \" using sigma_finite by auto - show "\A::nat \ 'a set. range A \ sets ?r \ (\i. A i) = space ?r \ (\i. measure ?r (A i) \ \)" + show "\A::nat \ 'a set. range A \ sets ?r \ (\i. A i) = space ?r \ (\i. measure ?r (A i) \ \)" proof (safe intro!: exI[of _ "\i. A i \ S"] del: notI) fix i show "A i \ S \ sets ?r" @@ -897,8 +864,7 @@ fix i have "\ (A i \ S) \ \ (A i)" using `range A \ sets M` `S \ sets M` by (auto intro!: measure_mono) - also have "\ < \" using `\ (A i) \ \` by (auto simp: pextreal_less_\) - finally show "measure ?r (A i \ S) \ \" by (auto simp: pextreal_less_\) + then show "measure ?r (A i \ S) \ \" using `\ (A i) \ \` by auto qed qed @@ -909,7 +875,7 @@ interpret M': measure_space M' by (intro measure_space_cong cong) from sigma_finite guess A .. note A = this then have "\i. A i \ sets M" by auto - with A have fin: "(\i. measure M' (A i) \ \)" using cong by auto + with A have fin: "\i. measure M' (A i) \ \" using cong by auto show ?thesis apply default using A fin cong by auto @@ -917,30 +883,30 @@ lemma (in sigma_finite_measure) disjoint_sigma_finite: "\A::nat\'a set. range A \ sets M \ (\i. A i) = space M \ - (\i. \ (A i) \ \) \ disjoint_family A" + (\i. \ (A i) \ \) \ disjoint_family A" proof - obtain A :: "nat \ 'a set" where range: "range A \ sets M" and space: "(\i. A i) = space M" and - measure: "\i. \ (A i) \ \" + measure: "\i. \ (A i) \ \" using sigma_finite by auto note range' = range_disjointed_sets[OF range] range { fix i have "\ (disjointed A i) \ \ (A i)" using range' disjointed_subset[of A i] by (auto intro!: measure_mono) - then have "\ (disjointed A i) \ \" + then have "\ (disjointed A i) \ \" using measure[of i] by auto } with disjoint_family_disjointed UN_disjointed_eq[of A] space range' show ?thesis by (auto intro!: exI[of _ "disjointed A"]) qed lemma (in sigma_finite_measure) sigma_finite_up: - "\F. range F \ sets M \ F \ space M \ (\i. \ (F i) \ \)" + "\F. range F \ sets M \ incseq F \ (\i. F i) = space M \ (\i. \ (F i) \ \)" proof - obtain F :: "nat \ 'a set" where - F: "range F \ sets M" "(\i. F i) = space M" "\i. \ (F i) \ \" + F: "range F \ sets M" "(\i. F i) = space M" "\i. \ (F i) \ \" using sigma_finite by auto - then show ?thesis unfolding isoton_def + then show ?thesis proof (intro exI[of _ "\n. \i\n. F i"] conjI allI) from F have "\x. x \ space M \ \i. x \ F i" by auto then show "(\n. \ i\n. F i) = space M" @@ -949,16 +915,16 @@ fix n have "\ (\ i\n. F i) \ (\i\n. \ (F i))" using F by (auto intro!: measure_finitely_subadditive) - also have "\ < \" - using F by (auto simp: setsum_\) - finally show "\ (\ i\n. F i) \ \" by simp - qed force+ + also have "\ < \" + using F by (auto simp: setsum_Pinfty) + finally show "\ (\ i\n. F i) \ \" by simp + qed (force simp: incseq_def)+ qed section {* Measure preserving *} definition "measure_preserving A B = - {f \ measurable A B. (\y \ sets B. measure A (f -` y \ space A) = measure B y)}" + {f \ measurable A B. (\y \ sets B. measure B y = measure A (f -` y \ space A))}" lemma measure_preservingI[intro?]: assumes "f \ measurable A B" @@ -974,7 +940,8 @@ interpret M': sigma_algebra M' by fact show ?thesis proof - show "measure M' {} = 0" using T by (force simp: measure_preserving_def) + show "positive M' (measure M')" using T + by (auto simp: measure_preserving_def positive_def measurable_sets) show "countably_additive M' (measure M')" proof (intro countably_additiveI) @@ -986,7 +953,7 @@ using * by blast moreover have **: "disjoint_family (\i. T -` A i \ space M)" using `disjoint_family A` by (auto simp: disjoint_family_on_def) - ultimately show "(\\<^isub>\ i. measure M' (A i)) = measure M' (\i. A i)" + ultimately show "(\i. measure M' (A i)) = measure M' (\i. A i)" using measure_countably_additive[OF _ **] A T by (auto simp: comp_def vimage_UN measure_preserving_def) qed @@ -1009,13 +976,13 @@ lemma measure_unique_Int_stable_vimage: fixes A :: "nat \ 'a set" assumes E: "Int_stable E" - and A: "range A \ sets E" "A \ space E" "\i. measure M (A i) \ \" + and A: "range A \ sets E" "incseq A" "(\i. A i) = space E" "\i. measure M (A i) \ \" and N: "measure_space N" "T \ measurable N M" and M: "measure_space M" "sets (sigma E) = sets M" "space E = space M" and eq: "\X. X \ sets E \ measure M X = measure N (T -` X \ space N)" assumes X: "X \ sets (sigma E)" shows "measure M X = measure N (T -` X \ space N)" -proof (rule measure_unique_Int_stable[OF E A(1,2) _ _ eq _ X]) +proof (rule measure_unique_Int_stable[OF E A(1,2,3) _ _ eq _ X]) interpret M: measure_space M by fact interpret N: measure_space N by fact let "?T X" = "T -` X \ space N" @@ -1028,12 +995,12 @@ show "T \ measure_preserving N ?E" using `T \ measurable N M` by (auto simp: M measurable_def measure_preserving_def) qed - show "\i. M.\ (A i) \ \" by fact + show "\i. M.\ (A i) \ \" by fact qed lemma (in measure_space) measure_preserving_Int_stable: fixes A :: "nat \ 'a set" - assumes E: "Int_stable E" "range A \ sets E" "A \ space E" "\i. measure E (A i) \ \" + assumes E: "Int_stable E" "range A \ sets E" "incseq A" "(\i. A i) = space E" "\i. measure E (A i) \ \" and N: "measure_space (sigma E)" and T: "T \ measure_preserving M E" shows "T \ measure_preserving M (sigma E)" @@ -1046,7 +1013,7 @@ show "\ (T -` X \ space M) = E.\ X" proof (rule measure_unique_Int_stable_vimage[symmetric]) show "sets (sigma E) = sets (sigma E)" "space E = space (sigma E)" - "\i. E.\ (A i) \ \" using E by auto + "\i. E.\ (A i) \ \" using E by auto show "measure_space M" by default next fix X assume "X \ sets E" then show "E.\ X = \ (T -` X \ space M)" @@ -1057,28 +1024,28 @@ section "Real measure values" lemma (in measure_space) real_measure_Union: - assumes finite: "\ A \ \" "\ B \ \" + assumes finite: "\ A \ \" "\ B \ \" and measurable: "A \ sets M" "B \ sets M" "A \ B = {}" shows "real (\ (A \ B)) = real (\ A) + real (\ B)" unfolding measure_additive[symmetric, OF measurable] - using finite by (auto simp: real_of_pextreal_add) + using measurable(1,2)[THEN positive_measure] + using finite by (cases rule: extreal2_cases[of "\ A" "\ B"]) auto lemma (in measure_space) real_measure_finite_Union: assumes measurable: "finite S" "\i. i \ S \ A i \ sets M" "disjoint_family_on A S" - assumes finite: "\i. i \ S \ \ (A i) \ \" + assumes finite: "\i. i \ S \ \ (A i) \ \" shows "real (\ (\i\S. A i)) = (\i\S. real (\ (A i)))" - using real_of_pextreal_setsum[of S, OF finite] - measure_finitely_additive''[symmetric, OF measurable] - by simp + using finite measurable(2)[THEN positive_measure] + by (force intro!: setsum_real_of_extreal[symmetric] + simp: measure_setsum[OF measurable, symmetric]) lemma (in measure_space) real_measure_Diff: - assumes finite: "\ A \ \" + assumes finite: "\ A \ \" and measurable: "A \ sets M" "B \ sets M" "B \ A" shows "real (\ (A - B)) = real (\ A) - real (\ B)" proof - - have "\ (A - B) \ \ A" - "\ B \ \ A" + have "\ (A - B) \ \ A" "\ B \ \ A" using measurable by (auto intro!: measure_mono) hence "real (\ ((A - B) \ B)) = real (\ (A - B)) + real (\ B)" using measurable finite by (rule_tac real_measure_Union) auto @@ -1087,117 +1054,155 @@ lemma (in measure_space) real_measure_UNION: assumes measurable: "range A \ sets M" "disjoint_family A" - assumes finite: "\ (\i. A i) \ \" + assumes finite: "\ (\i. A i) \ \" shows "(\i. real (\ (A i))) sums (real (\ (\i. A i)))" proof - - have *: "(\\<^isub>\ i. \ (A i)) = \ (\i. A i)" - using measure_countably_additive[OF measurable] by (simp add: comp_def) - hence "(\\<^isub>\ i. \ (A i)) \ \" using finite by simp - from psuminf_imp_suminf[OF this] - show ?thesis using * by simp + have "\i. 0 \ \ (A i)" using measurable by auto + with summable_sums[OF summable_extreal_pos, of "\i. \ (A i)"] + measure_countably_additive[OF measurable] + have "(\i. \ (A i)) sums (\ (\i. A i))" by simp + moreover + { fix i + have "\ (A i) \ \ (\i. A i)" + using measurable by (auto intro!: measure_mono) + moreover have "0 \ \ (A i)" using measurable by auto + ultimately have "\ (A i) = extreal (real (\ (A i)))" + using finite by (cases "\ (A i)") auto } + moreover + have "0 \ \ (\i. A i)" using measurable by auto + then have "\ (\i. A i) = extreal (real (\ (\i. A i)))" + using finite by (cases "\ (\i. A i)") auto + ultimately show ?thesis + unfolding sums_extreal[symmetric] by simp qed lemma (in measure_space) real_measure_subadditive: assumes measurable: "A \ sets M" "B \ sets M" - and fin: "\ A \ \" "\ B \ \" + and fin: "\ A \ \" "\ B \ \" shows "real (\ (A \ B)) \ real (\ A) + real (\ B)" proof - - have "real (\ (A \ B)) \ real (\ A + \ B)" - using measure_subadditive[OF measurable] fin by (auto intro!: real_of_pextreal_mono) - also have "\ = real (\ A) + real (\ B)" - using fin by (simp add: real_of_pextreal_add) - finally show ?thesis . -qed - -lemma (in measure_space) real_measure_countably_subadditive: - assumes "range f \ sets M" and "(\\<^isub>\ i. \ (f i)) \ \" - shows "real (\ (\i. f i)) \ (\ i. real (\ (f i)))" -proof - - have "real (\ (\i. f i)) \ real (\\<^isub>\ i. \ (f i))" - using assms by (auto intro!: real_of_pextreal_mono measure_countably_subadditive) - also have "\ = (\ i. real (\ (f i)))" - using assms by (auto intro!: sums_unique psuminf_imp_suminf) - finally show ?thesis . + have "0 \ \ (A \ B)" using measurable by auto + then show "real (\ (A \ B)) \ real (\ A) + real (\ B)" + using measure_subadditive[OF measurable] fin + by (cases rule: extreal3_cases[of "\ (A \ B)" "\ A" "\ B"]) auto qed lemma (in measure_space) real_measure_setsum_singleton: assumes S: "finite S" "\x. x \ S \ {x} \ sets M" - and fin: "\x. x \ S \ \ {x} \ \" + and fin: "\x. x \ S \ \ {x} \ \" shows "real (\ S) = (\x\S. real (\ {x}))" - using measure_finite_singleton[OF S] fin by (simp add: real_of_pextreal_setsum) + using measure_finite_singleton[OF S] fin + using positive_measure[OF S(2)] + by (force intro!: setsum_real_of_extreal[symmetric]) lemma (in measure_space) real_continuity_from_below: - assumes A: "range A \ sets M" "\i. A i \ A (Suc i)" and "\ (\i. A i) \ \" + assumes A: "range A \ sets M" "incseq A" and fin: "\ (\i. A i) \ \" shows "(\i. real (\ (A i))) ----> real (\ (\i. A i))" -proof (rule SUP_eq_LIMSEQ[THEN iffD1]) - { fix n have "\ (A n) \ \ (\i. A i)" - using A by (auto intro!: measure_mono) - hence "\ (A n) \ \" using assms by auto } - note this[simp] +proof - + have "0 \ \ (\i. A i)" using A by auto + then have "extreal (real (\ (\i. A i))) = \ (\i. A i)" + using fin by (auto intro: extreal_real') + then show ?thesis + using continuity_from_below_Lim[OF A] + by (intro lim_real_of_extreal) simp +qed - show "mono (\i. real (\ (A i)))" unfolding mono_iff_le_Suc using A - by (auto intro!: real_of_pextreal_mono measure_mono) - - show "(SUP n. Real (real (\ (A n)))) = Real (real (\ (\i. A i)))" - using continuity_from_below[OF A(1), OF A(2)] - using assms by (simp add: Real_real) -qed simp_all +lemma (in measure_space) continuity_from_above_Lim: + assumes A: "range A \ sets M" "decseq A" and fin: "\i. \ (A i) \ \" + shows "(\i. (\ (A i))) ----> \ (\i. A i)" + using LIMSEQ_extreal_INFI[OF measure_decseq, OF A] + using continuity_from_above[OF A fin] by simp lemma (in measure_space) real_continuity_from_above: - assumes A: "range A \ sets M" - and mono_Suc: "\n. A (Suc n) \ A n" - and finite: "\i. \ (A i) \ \" + assumes A: "range A \ sets M" "decseq A" and fin: "\i. \ (A i) \ \" shows "(\n. real (\ (A n))) ----> real (\ (\i. A i))" -proof (rule INF_eq_LIMSEQ[THEN iffD1]) - { fix n have "\ (\i. A i) \ \ (A n)" - using A by (auto intro!: measure_mono) - hence "\ (\i. A i) \ \" using assms by auto } - note this[simp] +proof - + have "0 \ \ (\i. A i)" using A by auto + moreover + have "\ (\i. A i) \ \ (A 0)" + using A by (auto intro!: measure_mono) + ultimately have "extreal (real (\ (\i. A i))) = \ (\i. A i)" + using fin by (auto intro: extreal_real') + then show ?thesis + using continuity_from_above_Lim[OF A fin] + by (intro lim_real_of_extreal) simp +qed - show "mono (\i. - real (\ (A i)))" unfolding mono_iff_le_Suc using assms - by (auto intro!: real_of_pextreal_mono measure_mono) - - show "(INF n. Real (real (\ (A n)))) = - Real (real (\ (\i. A i)))" - using continuity_from_above[OF A, OF mono_Suc finite] - using assms by (simp add: Real_real) -qed simp_all +lemma (in measure_space) real_measure_countably_subadditive: + assumes A: "range A \ sets M" and fin: "(\i. \ (A i)) \ \" + shows "real (\ (\i. A i)) \ (\i. real (\ (A i)))" +proof - + { fix i + have "0 \ \ (A i)" using A by auto + moreover have "\ (A i) \ \" using A by (intro suminf_PInfty[OF _ fin]) auto + ultimately have "\\ (A i)\ \ \" by auto } + moreover have "0 \ \ (\i. A i)" using A by auto + ultimately have "extreal (real (\ (\i. A i))) \ (\i. extreal (real (\ (A i))))" + using measure_countably_subadditive[OF A] by (auto simp: extreal_real) + also have "\ = extreal (\i. real (\ (A i)))" + using A + by (auto intro!: sums_unique[symmetric] sums_extreal[THEN iffD2] summable_sums summable_real_of_extreal fin) + finally show ?thesis by simp +qed locale finite_measure = measure_space + - assumes finite_measure_of_space: "\ (space M) \ \" + assumes finite_measure_of_space: "\ (space M) \ \" sublocale finite_measure < sigma_finite_measure proof - show "\A. range A \ sets M \ (\i. A i) = space M \ (\i. \ (A i) \ \)" + show "\A. range A \ sets M \ (\i. A i) = space M \ (\i. \ (A i) \ \)" using finite_measure_of_space by (auto intro!: exI[of _ "\x. space M"]) qed lemma (in finite_measure) finite_measure[simp, intro]: assumes "A \ sets M" - shows "\ A \ \" + shows "\ A \ \" proof - from `A \ sets M` have "A \ space M" using sets_into_space by blast - hence "\ A \ \ (space M)" + then have "\ A \ \ (space M)" using assms top by (rule measure_mono) - also have "\ < \" - using finite_measure_of_space unfolding pextreal_less_\ . - finally show ?thesis unfolding pextreal_less_\ . + then show ?thesis + using finite_measure_of_space by auto qed +lemma (in finite_measure) measure_not_inf: + assumes A: "A \ sets M" + shows "\\ A\ \ \" + using finite_measure[OF A] positive_measure[OF A] by auto + +definition (in finite_measure) + "\' A = (if A \ sets M then real (\ A) else 0)" + +lemma (in finite_measure) finite_measure_eq: "A \ sets M \ \ A = extreal (\' A)" + using measure_not_inf[of A] by (auto simp: \'_def) + +lemma (in finite_measure) positive_measure': "0 \ \' A" + unfolding \'_def by (auto simp: real_of_extreal_pos) + +lemma (in finite_measure) bounded_measure: "\' A \ \' (space M)" +proof cases + assume "A \ sets M" + moreover then have "\ A \ \ (space M)" + using sets_into_space by (auto intro!: measure_mono) + ultimately show ?thesis + using measure_not_inf[of A] measure_not_inf[of "space M"] + by (auto simp: \'_def) +qed (simp add: \'_def real_of_extreal_pos) + lemma (in finite_measure) restricted_finite_measure: assumes "S \ sets M" shows "finite_measure (restricted_space S)" (is "finite_measure ?r") unfolding finite_measure_def finite_measure_axioms_def -proof (safe del: notI) +proof (intro conjI) show "measure_space ?r" using restricted_measure_space[OF assms] . next - show "measure ?r (space ?r) \ \" using finite_measure[OF `S \ sets M`] by auto + show "measure ?r (space ?r) \ \" using finite_measure[OF `S \ sets M`] by auto qed lemma (in measure_space) restricted_to_finite_measure: - assumes "S \ sets M" "\ S \ \" + assumes "S \ sets M" "\ S \ \" shows "finite_measure (restricted_space S)" proof - have "measure_space (restricted_space S)" @@ -1207,202 +1212,128 @@ using assms by auto qed -lemma (in finite_measure) real_finite_measure_Diff: - assumes "A \ sets M" "B \ sets M" "B \ A" - shows "real (\ (A - B)) = real (\ A) - real (\ B)" - using finite_measure[OF `A \ sets M`] by (rule real_measure_Diff[OF _ assms]) - -lemma (in finite_measure) real_finite_measure_Union: - assumes sets: "A \ sets M" "B \ sets M" and "A \ B = {}" - shows "real (\ (A \ B)) = real (\ A) + real (\ B)" - using sets by (auto intro!: real_measure_Union[OF _ _ assms] finite_measure) +lemma (in finite_measure) finite_measure_Diff: + assumes sets: "A \ sets M" "B \ sets M" and "B \ A" + shows "\' (A - B) = \' A - \' B" + using sets[THEN finite_measure_eq] + using Diff[OF sets, THEN finite_measure_eq] + using measure_Diff[OF _ assms] by simp -lemma (in finite_measure) real_finite_measure_finite_Union: - assumes "finite S" and dis: "disjoint_family_on A S" - and *: "\i. i \ S \ A i \ sets M" - shows "real (\ (\i\S. A i)) = (\i\S. real (\ (A i)))" -proof (rule real_measure_finite_Union[OF `finite S` _ dis]) - fix i assume "i \ S" from *[OF this] show "A i \ sets M" . - from finite_measure[OF this] show "\ (A i) \ \" . -qed - -lemma (in finite_measure) real_finite_measure_UNION: - assumes measurable: "range A \ sets M" "disjoint_family A" - shows "(\i. real (\ (A i))) sums (real (\ (\i. A i)))" -proof (rule real_measure_UNION[OF assms]) - have "(\i. A i) \ sets M" using measurable(1) by (rule countable_UN) - thus "\ (\i. A i) \ \" by (rule finite_measure) -qed - -lemma (in finite_measure) real_measure_mono: - "A \ sets M \ B \ sets M \ A \ B \ real (\ A) \ real (\ B)" - by (auto intro!: measure_mono real_of_pextreal_mono finite_measure) - -lemma (in finite_measure) real_finite_measure_subadditive: - assumes measurable: "A \ sets M" "B \ sets M" - shows "real (\ (A \ B)) \ real (\ A) + real (\ B)" - using measurable measurable[THEN finite_measure] by (rule real_measure_subadditive) +lemma (in finite_measure) finite_measure_Union: + assumes sets: "A \ sets M" "B \ sets M" and "A \ B = {}" + shows "\' (A \ B) = \' A + \' B" + using measure_additive[OF assms] + using sets[THEN finite_measure_eq] + using Un[OF sets, THEN finite_measure_eq] + by simp -lemma (in finite_measure) real_finite_measure_countably_subadditive: - assumes "range f \ sets M" and "summable (\i. real (\ (f i)))" - shows "real (\ (\i. f i)) \ (\ i. real (\ (f i)))" -proof (rule real_measure_countably_subadditive[OF assms(1)]) - have *: "\i. f i \ sets M" using assms by auto - have "(\i. real (\ (f i))) sums (\i. real (\ (f i)))" - using assms(2) by (rule summable_sums) - from suminf_imp_psuminf[OF this] - have "(\\<^isub>\i. \ (f i)) = Real (\i. real (\ (f i)))" - using * by (simp add: Real_real finite_measure) - thus "(\\<^isub>\i. \ (f i)) \ \" by simp -qed - -lemma (in finite_measure) real_finite_measure_finite_singelton: - assumes "finite S" and *: "\x. x \ S \ {x} \ sets M" - shows "real (\ S) = (\x\S. real (\ {x}))" -proof (rule real_measure_setsum_singleton[OF `finite S`]) - fix x assume "x \ S" thus "{x} \ sets M" by (rule *) - with finite_measure show "\ {x} \ \" . -qed +lemma (in finite_measure) finite_measure_finite_Union: + assumes S: "finite S" "\i. i \ S \ A i \ sets M" + and dis: "disjoint_family_on A S" + shows "\' (\i\S. A i) = (\i\S. \' (A i))" + using measure_setsum[OF assms] + using finite_UN[of S A, OF S, THEN finite_measure_eq] + using S(2)[THEN finite_measure_eq] + by simp -lemma (in finite_measure) real_finite_continuity_from_below: - assumes "range A \ sets M" "\i. A i \ A (Suc i)" - shows "(\i. real (\ (A i))) ----> real (\ (\i. A i))" - using real_continuity_from_below[OF assms(1), OF assms(2) finite_measure] assms by auto - -lemma (in finite_measure) real_finite_continuity_from_above: - assumes A: "range A \ sets M" - and mono_Suc: "\n. A (Suc n) \ A n" - shows "(\n. real (\ (A n))) ----> real (\ (\i. A i))" - using real_continuity_from_above[OF A, OF mono_Suc finite_measure] A - by auto - -lemma (in finite_measure) real_finite_measure_finite_Union': - assumes "finite S" "A`S \ sets M" "disjoint_family_on A S" - shows "real (\ (\i\S. A i)) = (\i\S. real (\ (A i)))" - using assms real_finite_measure_finite_Union[of S A] by auto - -lemma (in finite_measure) finite_measure_compl: - assumes s: "s \ sets M" - shows "\ (space M - s) = \ (space M) - \ s" - using measure_compl[OF s, OF finite_measure, OF s] . - -lemma (in finite_measure) finite_measure_inter_full_set: - assumes "S \ sets M" "T \ sets M" - assumes T: "\ T = \ (space M)" - shows "\ (S \ T) = \ S" - using measure_inter_full_set[OF assms(1,2) finite_measure assms(3)] assms +lemma (in finite_measure) finite_measure_UNION: + assumes A: "range A \ sets M" "disjoint_family A" + shows "(\i. \' (A i)) sums (\' (\i. A i))" + using real_measure_UNION[OF A] + using countable_UN[OF A(1), THEN finite_measure_eq] + using A(1)[THEN subsetD, THEN finite_measure_eq] by auto -lemma (in finite_measure) measure_preserving_lift: - fixes f :: "'a \ 'c" and A :: "('c, 'd) measure_space_scheme" - assumes "algebra A" "finite_measure (sigma A)" (is "finite_measure ?sA") - assumes fmp: "f \ measure_preserving M A" - shows "f \ measure_preserving M (sigma A)" +lemma (in finite_measure) finite_measure_mono: + assumes B: "B \ sets M" and "A \ B" shows "\' A \ \' B" +proof cases + assume "A \ sets M" + from this[THEN finite_measure_eq] B[THEN finite_measure_eq] + show ?thesis using measure_mono[OF `A \ B` `A \ sets M` `B \ sets M`] by simp +next + assume "A \ sets M" then show ?thesis + using positive_measure'[of B] unfolding \'_def by auto +qed + +lemma (in finite_measure) finite_measure_subadditive: + assumes m: "A \ sets M" "B \ sets M" + shows "\' (A \ B) \ \' A + \' B" + using measure_subadditive[OF m] + using m[THEN finite_measure_eq] Un[OF m, THEN finite_measure_eq] by simp + +lemma (in finite_measure) finite_measure_countably_subadditive: + assumes A: "range A \ sets M" and sum: "summable (\i. \' (A i))" + shows "\' (\i. A i) \ (\i. \' (A i))" proof - - interpret sA: finite_measure ?sA by fact - interpret A: algebra A by fact - show ?thesis using fmp - proof (clarsimp simp add: measure_preserving_def) - assume fm: "f \ measurable M A" - and "\y\sets A. \ (f -` y \ space M) = measure A y" - then have meq: "\y\sets A. \ (f -` y \ space M) = sA.\ y" - by simp - have f12: "f \ measurable M ?sA" - using measurable_subset[OF A.sets_into_space] fm by auto - hence ffn: "f \ space M \ space A" - by (simp add: measurable_def) - have "space M \ f -` (space A)" - by auto (metis PiE ffn) - hence fveq [simp]: "(f -` (space A)) \ space M = space M" - by blast - { - fix y - assume y: "y \ sets ?sA" - have "sets ?sA = sigma_sets (space A) (sets A)" (is "_ = ?A") by (auto simp: sigma_def) - also have "\ \ {s . \ ((f -` s) \ space M) = sA.\ s}" - proof (rule A.sigma_property_disjoint, safe) - fix x assume "x \ sets A" then show "\ (f -` x \ space M) = sA.\ x" by (simp add: meq) - next - fix s - assume eq: "\ (f -` s \ space M) = sA.\ s" and s: "s \ ?A" - then have s': "s \ sets ?sA" by (simp add: sigma_def) - show "\ (f -` (space A - s) \ space M) = measure ?sA (space A - s)" - using sA.finite_measure_compl[OF s'] - using measurable_sets[OF f12 s'] meq[THEN bspec, OF A.top] - by (simp add: vimage_Diff Diff_Int_distrib2 finite_measure_compl eq) - next - fix S - assume S0: "S 0 = {}" - and SSuc: "\n. S n \ S (Suc n)" - and rS1: "range S \ {s. \ (f -` s \ space M) = sA.\ s} \ ?A" - hence rS2: "range S \ sets ?sA" by (simp add: sigma_def) - have eq1: "\i. \ (f -` S i \ space M) = sA.\ (S i)" - using rS1 by blast - have *: "(\n. sA.\ (S n)) = (\n. \ (f -` S n \ space M))" - by (simp add: eq1) - have "(SUP n. ... n) = \ (\i. f -` S i \ space M)" - proof (rule measure_countable_increasing) - show "range (\i. f -` S i \ space M) \ sets M" - using f12 rS2 by (auto simp add: measurable_def) - show "f -` S 0 \ space M = {}" using S0 - by blast - show "\n. f -` S n \ space M \ f -` S (Suc n) \ space M" - using SSuc by auto - qed - also have "\ (\i. f -` S i \ space M) = \ (f -` (\i. S i) \ space M)" - by (simp add: vimage_UN) - finally have "(SUP n. sA.\ (S n)) = \ (f -` (\i. S i) \ space M)" unfolding * . - moreover - have "(SUP n. sA.\ (S n)) = sA.\ (\i. S i)" - by (rule sA.measure_countable_increasing[OF rS2, OF S0 SSuc]) - ultimately - show "\ (f -` (\i. S i) \ space M) = sA.\ (\i. S i)" by simp - next - fix S :: "nat \ 'c set" - assume dS: "disjoint_family S" - and rS1: "range S \ {s. \ (f -` s \ space M) = sA.\ s} \ ?A" - hence rS2: "range S \ sets ?sA" by (simp add: sigma_def) - have "\i. \ (f -` S i \ space M) = sA.\ (S i)" - using rS1 by blast - hence *: "(\i. sA.\ (S i)) = (\n. \ (f -` S n \ space M))" - by simp - have "psuminf ... = \ (\i. f -` S i \ space M)" - proof (rule measure_countably_additive) - show "range (\i. f -` S i \ space M) \ sets M" - using f12 rS2 by (auto simp add: measurable_def) - show "disjoint_family (\i. f -` S i \ space M)" using dS - by (auto simp add: disjoint_family_on_def) - qed - hence "(\\<^isub>\ i. sA.\ (S i)) = \ (\i. f -` S i \ space M)" unfolding * . - with sA.measure_countably_additive [OF rS2 dS] - have "\ (\i. f -` S i \ space M) = sA.\ (\i. S i)" - by simp - moreover have "\ (f -` (\i. S i) \ space M) = \ (\i. f -` S i \ space M)" - by (simp add: vimage_UN) - ultimately show "\ (f -` (\i. S i) \ space M) = sA.\ (\i. S i)" - by metis - qed - finally have "sets ?sA \ {s . \ ((f -` s) \ space M) = sA.\ s}" . - hence "\ (f -` y \ space M) = sA.\ y" using y by force - } - thus "f \ measurable M ?sA \ (\y\sets ?sA. \ (f -` y \ space M) = measure A y)" - by simp_all (blast intro: f12) - qed + note A[THEN subsetD, THEN finite_measure_eq, simp] + note countable_UN[OF A, THEN finite_measure_eq, simp] + from `summable (\i. \' (A i))` + have "(\i. extreal (\' (A i))) sums extreal (\i. \' (A i))" + by (simp add: sums_extreal) (rule summable_sums) + from sums_unique[OF this, symmetric] + measure_countably_subadditive[OF A] + show ?thesis by simp qed +lemma (in finite_measure) finite_measure_finite_singleton: + assumes "finite S" and *: "\x. x \ S \ {x} \ sets M" + shows "\' S = (\x\S. \' {x})" + using real_measure_setsum_singleton[OF assms] + using *[THEN finite_measure_eq] + using finite_UN[of S "\x. {x}", OF assms, THEN finite_measure_eq] + by simp + +lemma (in finite_measure) finite_continuity_from_below: + assumes A: "range A \ sets M" and "incseq A" + shows "(\i. \' (A i)) ----> \' (\i. A i)" + using real_continuity_from_below[OF A, OF `incseq A` finite_measure] assms + using A[THEN subsetD, THEN finite_measure_eq] + using countable_UN[OF A, THEN finite_measure_eq] + by auto + +lemma (in finite_measure) finite_continuity_from_above: + assumes A: "range A \ sets M" and "decseq A" + shows "(\n. \' (A n)) ----> \' (\i. A i)" + using real_continuity_from_above[OF A, OF `decseq A` finite_measure] assms + using A[THEN subsetD, THEN finite_measure_eq] + using countable_INT[OF A, THEN finite_measure_eq] + by auto + +lemma (in finite_measure) finite_measure_compl: + assumes S: "S \ sets M" + shows "\' (space M - S) = \' (space M) - \' S" + using measure_compl[OF S, OF finite_measure, OF S] + using S[THEN finite_measure_eq] + using compl_sets[OF S, THEN finite_measure_eq] + using top[THEN finite_measure_eq] + by simp + +lemma (in finite_measure) finite_measure_inter_full_set: + assumes S: "S \ sets M" "T \ sets M" + assumes T: "\' T = \' (space M)" + shows "\' (S \ T) = \' S" + using measure_inter_full_set[OF S finite_measure] + using T Diff[OF S(2,1)] Diff[OF S, THEN finite_measure_eq] + using Int[OF S, THEN finite_measure_eq] + using S[THEN finite_measure_eq] top[THEN finite_measure_eq] + by simp + +lemma (in finite_measure) empty_measure'[simp]: "\' {} = 0" + unfolding \'_def by simp + section "Finite spaces" locale finite_measure_space = measure_space + finite_sigma_algebra + - assumes finite_single_measure[simp]: "\x. x \ space M \ \ {x} \ \" + assumes finite_single_measure[simp]: "\x. x \ space M \ \ {x} \ \" lemma (in finite_measure_space) sum_over_space: "(\x\space M. \ {x}) = \ (space M)" - using measure_finitely_additive''[of "space M" "\i. {i}"] + using measure_setsum[of "space M" "\i. {i}"] by (simp add: sets_eq_Pow disjoint_family_on_def finite_space) lemma finite_measure_spaceI: - assumes "finite (space M)" "sets M = Pow(space M)" and space: "measure M (space M) \ \" + assumes "finite (space M)" "sets M = Pow(space M)" and space: "measure M (space M) \ \" and add: "\A B. A\space M \ B\space M \ A \ B = {} \ measure M (A \ B) = measure M A + measure M B" - and "measure M {} = 0" + and "measure M {} = 0" "\A. A \ space M \ 0 \ measure M A" shows "finite_measure_space M" unfolding finite_measure_space_def finite_measure_space_axioms_def proof (intro allI impI conjI) @@ -1414,7 +1345,7 @@ using sigma_algebra_Pow[of "space M" "algebra.more M"] unfolding * . show "finite (space M)" by fact - show "positive M (measure M)" unfolding positive_def by fact + show "positive M (measure M)" unfolding positive_def using assms by auto show "additive M (measure M)" unfolding additive_def using assms by simp qed then interpret measure_space M . @@ -1425,19 +1356,20 @@ qed { fix x assume *: "x \ space M" with add[of "{x}" "space M - {x}"] space - show "\ {x} \ \" by (auto simp: insert_absorb[OF *] Diff_subset) } + show "\ {x} \ \" by (auto simp: insert_absorb[OF *] Diff_subset) } qed sublocale finite_measure_space \ finite_measure proof - show "\ (space M) \ \" - unfolding sum_over_space[symmetric] setsum_\ + show "\ (space M) \ \" + unfolding sum_over_space[symmetric] setsum_Pinfty using finite_space finite_single_measure by auto qed lemma finite_measure_space_iff: "finite_measure_space M \ - finite (space M) \ sets M = Pow(space M) \ measure M (space M) \ \ \ measure M {} = 0 \ + finite (space M) \ sets M = Pow(space M) \ measure M (space M) \ \ \ + measure M {} = 0 \ (\A\space M. 0 \ measure M A) \ (\A\space M. \B\space M. A \ B = {} \ measure M (A \ B) = measure M A + measure M B)" (is "_ = ?rhs") proof (intro iffI) @@ -1451,29 +1383,31 @@ by (auto intro!: finite_measure_spaceI) qed -lemma psuminf_cmult_indicator: - assumes "disjoint_family A" "x \ A i" - shows "(\\<^isub>\ n. f n * indicator (A n) x) = f i" +lemma suminf_cmult_indicator: + fixes f :: "nat \ extreal" + assumes "disjoint_family A" "x \ A i" "\i. 0 \ f i" + shows "(\n. f n * indicator (A n) x) = f i" proof - - have **: "\n. f n * indicator (A n) x = (if n = i then f n else 0 :: pextreal)" + have **: "\n. f n * indicator (A n) x = (if n = i then f n else 0 :: extreal)" using `x \ A i` assms unfolding disjoint_family_on_def indicator_def by auto - then have "\n. (\jn. (\jn. n \ UNIV \ (if i < n then f i else 0) \ y" + moreover have "(SUP n. if i < n then f i else 0) = (f i :: extreal)" + proof (rule extreal_SUPI) + fix y :: extreal assume "\n. n \ UNIV \ (if i < n then f i else 0) \ y" from this[of "Suc i"] show "f i \ y" by auto - qed simp - ultimately show ?thesis using `x \ A i` unfolding psuminf_def by auto + qed (insert assms, simp) + ultimately show ?thesis using assms + by (subst suminf_extreal_eq_SUPR) (auto simp: indicator_def) qed -lemma psuminf_indicator: +lemma suminf_indicator: assumes "disjoint_family A" - shows "(\\<^isub>\ n. indicator (A n) x) = indicator (\i. A i) x" + shows "(\n. indicator (A n) x :: extreal) = indicator (\i. A i) x" proof cases assume *: "x \ (\i. A i)" then obtain i where "x \ A i" by auto - from psuminf_cmult_indicator[OF assms, OF this, of "\i. 1"] + from suminf_cmult_indicator[OF assms(1), OF `x \ A i`, of "\k. 1"] show ?thesis using * by simp qed simp diff -r 28b51effc5ed -r cdf7693bbe08 src/HOL/Probability/Positive_Extended_Real.thy --- a/src/HOL/Probability/Positive_Extended_Real.thy Mon Mar 14 14:37:47 2011 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,2908 +0,0 @@ -(* Author: Johannes Hoelzl, TU Muenchen *) - -header {* A type for positive real numbers with infinity *} - -theory Positive_Extended_Real - imports Complex_Main "~~/src/HOL/Library/Nat_Bijection" Multivariate_Analysis -begin - -lemma (in complete_lattice) Sup_start: - assumes *: "\x. f x \ f 0" - shows "(SUP n. f n) = f 0" -proof (rule antisym) - show "f 0 \ (SUP n. f n)" by (rule le_SUPI) auto - show "(SUP n. f n) \ f 0" by (rule SUP_leI[OF *]) -qed - -lemma (in complete_lattice) Inf_start: - assumes *: "\x. f 0 \ f x" - shows "(INF n. f n) = f 0" -proof (rule antisym) - show "(INF n. f n) \ f 0" by (rule INF_leI) simp - show "f 0 \ (INF n. f n)" by (rule le_INFI[OF *]) -qed - -lemma (in complete_lattice) Sup_mono_offset: - fixes f :: "'b :: {ordered_ab_semigroup_add,monoid_add} \ 'a" - assumes *: "\x y. x \ y \ f x \ f y" and "0 \ k" - shows "(SUP n . f (k + n)) = (SUP n. f n)" -proof (rule antisym) - show "(SUP n. f (k + n)) \ (SUP n. f n)" - by (auto intro!: Sup_mono simp: SUPR_def) - { fix n :: 'b - have "0 + n \ k + n" using `0 \ k` by (rule add_right_mono) - with * have "f n \ f (k + n)" by simp } - thus "(SUP n. f n) \ (SUP n. f (k + n))" - by (auto intro!: Sup_mono exI simp: SUPR_def) -qed - -lemma (in complete_lattice) Sup_mono_offset_Suc: - assumes *: "\x. f x \ f (Suc x)" - shows "(SUP n . f (Suc n)) = (SUP n. f n)" - unfolding Suc_eq_plus1 - apply (subst add_commute) - apply (rule Sup_mono_offset) - by (auto intro!: order.lift_Suc_mono_le[of "op \" "op <" f, OF _ *]) default - -lemma (in complete_lattice) Inf_mono_offset: - fixes f :: "'b :: {ordered_ab_semigroup_add,monoid_add} \ 'a" - assumes *: "\x y. x \ y \ f y \ f x" and "0 \ k" - shows "(INF n . f (k + n)) = (INF n. f n)" -proof (rule antisym) - show "(INF n. f n) \ (INF n. f (k + n))" - by (auto intro!: Inf_mono simp: INFI_def) - { fix n :: 'b - have "0 + n \ k + n" using `0 \ k` by (rule add_right_mono) - with * have "f (k + n) \ f n" by simp } - thus "(INF n. f (k + n)) \ (INF n. f n)" - by (auto intro!: Inf_mono exI simp: INFI_def) -qed - -lemma (in complete_lattice) isotone_converge: - fixes f :: "nat \ 'a" assumes "\x y. x \ y \ f x \ f y " - shows "(INF n. SUP m. f (n + m)) = (SUP n. INF m. f (n + m))" -proof - - have "\n. (SUP m. f (n + m)) = (SUP n. f n)" - apply (rule Sup_mono_offset) - apply (rule assms) - by simp_all - moreover - { fix n have "(INF m. f (n + m)) = f n" - using Inf_start[of "\m. f (n + m)"] assms by simp } - ultimately show ?thesis by simp -qed - -lemma (in complete_lattice) antitone_converges: - fixes f :: "nat \ 'a" assumes "\x y. x \ y \ f y \ f x" - shows "(INF n. SUP m. f (n + m)) = (SUP n. INF m. f (n + m))" -proof - - have "\n. (INF m. f (n + m)) = (INF n. f n)" - apply (rule Inf_mono_offset) - apply (rule assms) - by simp_all - moreover - { fix n have "(SUP m. f (n + m)) = f n" - using Sup_start[of "\m. f (n + m)"] assms by simp } - ultimately show ?thesis by simp -qed - -lemma (in complete_lattice) lim_INF_le_lim_SUP: - fixes f :: "nat \ 'a" - shows "(SUP n. INF m. f (n + m)) \ (INF n. SUP m. f (n + m))" -proof (rule SUP_leI, rule le_INFI) - fix i j show "(INF m. f (i + m)) \ (SUP m. f (j + m))" - proof (cases rule: le_cases) - assume "i \ j" - have "(INF m. f (i + m)) \ f (i + (j - i))" by (rule INF_leI) simp - also have "\ = f (j + 0)" using `i \ j` by auto - also have "\ \ (SUP m. f (j + m))" by (rule le_SUPI) simp - finally show ?thesis . - next - assume "j \ i" - have "(INF m. f (i + m)) \ f (i + 0)" by (rule INF_leI) simp - also have "\ = f (j + (i - j))" using `j \ i` by auto - also have "\ \ (SUP m. f (j + m))" by (rule le_SUPI) simp - finally show ?thesis . - qed -qed - -text {* - -We introduce the the positive real numbers as needed for measure theory. - -*} - -typedef pextreal = "(Some ` {0::real..}) \ {None}" - by (rule exI[of _ None]) simp - -subsection "Introduce @{typ pextreal} similar to a datatype" - -definition "Real x = Abs_pextreal (Some (sup 0 x))" -definition "\ = Abs_pextreal None" - -definition "pextreal_case f i x = (if x = \ then i else f (THE r. 0 \ r \ x = Real r))" - -definition "of_pextreal = pextreal_case (\x. x) 0" - -defs (overloaded) - real_of_pextreal_def [code_unfold]: "real == of_pextreal" - -lemma pextreal_Some[simp]: "0 \ x \ Some x \ pextreal" - unfolding pextreal_def by simp - -lemma pextreal_Some_sup[simp]: "Some (sup 0 x) \ pextreal" - by (simp add: sup_ge1) - -lemma pextreal_None[simp]: "None \ pextreal" - unfolding pextreal_def by simp - -lemma Real_inj[simp]: - assumes "0 \ x" and "0 \ y" - shows "Real x = Real y \ x = y" - unfolding Real_def assms[THEN sup_absorb2] - using assms by (simp add: Abs_pextreal_inject) - -lemma Real_neq_\[simp]: - "Real x = \ \ False" - "\ = Real x \ False" - by (simp_all add: Abs_pextreal_inject \_def Real_def) - -lemma Real_neg: "x < 0 \ Real x = Real 0" - unfolding Real_def by (auto simp add: Abs_pextreal_inject intro!: sup_absorb1) - -lemma pextreal_cases[case_names preal infinite, cases type: pextreal]: - assumes preal: "\r. x = Real r \ 0 \ r \ P" and inf: "x = \ \ P" - shows P -proof (cases x rule: pextreal.Abs_pextreal_cases) - case (Abs_pextreal y) - hence "y = None \ (\x \ 0. y = Some x)" - unfolding pextreal_def by auto - thus P - proof (rule disjE) - assume "\x\0. y = Some x" then guess x .. - thus P by (simp add: preal[of x] Real_def Abs_pextreal(1) sup_absorb2) - qed (simp add: \_def Abs_pextreal(1) inf) -qed - -lemma pextreal_case_\[simp]: "pextreal_case f i \ = i" - unfolding pextreal_case_def by simp - -lemma pextreal_case_Real[simp]: "pextreal_case f i (Real x) = (if 0 \ x then f x else f 0)" -proof (cases "0 \ x") - case True thus ?thesis unfolding pextreal_case_def by (auto intro: theI2) -next - case False - moreover have "(THE r. 0 \ r \ Real 0 = Real r) = 0" - by (auto intro!: the_equality) - ultimately show ?thesis unfolding pextreal_case_def by (simp add: Real_neg) -qed - -lemma pextreal_case_cancel[simp]: "pextreal_case (\c. i) i x = i" - by (cases x) simp_all - -lemma pextreal_case_split: - "P (pextreal_case f i x) = ((x = \ \ P i) \ (\r\0. x = Real r \ P (f r)))" - by (cases x) simp_all - -lemma pextreal_case_split_asm: - "P (pextreal_case f i x) = (\ (x = \ \ \ P i \ (\r. r \ 0 \ x = Real r \ \ P (f r))))" - by (cases x) auto - -lemma pextreal_case_cong[cong]: - assumes eq: "x = x'" "i = i'" and cong: "\r. 0 \ r \ f r = f' r" - shows "pextreal_case f i x = pextreal_case f' i' x'" - unfolding eq using cong by (cases x') simp_all - -lemma real_Real[simp]: "real (Real x) = (if 0 \ x then x else 0)" - unfolding real_of_pextreal_def of_pextreal_def by simp - -lemma Real_real_image: - assumes "\ \ A" shows "Real ` real ` A = A" -proof safe - fix x assume "x \ A" - hence *: "x = Real (real x)" - using `\ \ A` by (cases x) auto - show "x \ Real ` real ` A" - using `x \ A` by (subst *) (auto intro!: imageI) -next - fix x assume "x \ A" - thus "Real (real x) \ A" - using `\ \ A` by (cases x) auto -qed - -lemma real_pextreal_nonneg[simp, intro]: "0 \ real (x :: pextreal)" - unfolding real_of_pextreal_def of_pextreal_def - by (cases x) auto - -lemma real_\[simp]: "real \ = 0" - unfolding real_of_pextreal_def of_pextreal_def by simp - -lemma pextreal_noteq_omega_Ex: "X \ \ \ (\r\0. X = Real r)" by (cases X) auto - -subsection "@{typ pextreal} is a monoid for addition" - -instantiation pextreal :: comm_monoid_add -begin - -definition "0 = Real 0" -definition "x + y = pextreal_case (\r. pextreal_case (\p. Real (r + p)) \ y) \ x" - -lemma pextreal_plus[simp]: - "Real r + Real p = (if 0 \ r then if 0 \ p then Real (r + p) else Real r else Real p)" - "x + 0 = x" - "0 + x = x" - "x + \ = \" - "\ + x = \" - by (simp_all add: plus_pextreal_def Real_neg zero_pextreal_def split: pextreal_case_split) - -lemma \_neq_0[simp]: - "\ = 0 \ False" - "0 = \ \ False" - by (simp_all add: zero_pextreal_def) - -lemma Real_eq_0[simp]: - "Real r = 0 \ r \ 0" - "0 = Real r \ r \ 0" - by (auto simp add: Abs_pextreal_inject zero_pextreal_def Real_def sup_real_def) - -lemma Real_0[simp]: "Real 0 = 0" by (simp add: zero_pextreal_def) - -instance -proof - fix a :: pextreal - show "0 + a = a" by (cases a) simp_all - - fix b show "a + b = b + a" - by (cases a, cases b) simp_all - - fix c show "a + b + c = a + (b + c)" - by (cases a, cases b, cases c) simp_all -qed -end - -lemma Real_minus_abs[simp]: "Real (- \x\) = 0" - by simp - -lemma pextreal_plus_eq_\[simp]: "(a :: pextreal) + b = \ \ a = \ \ b = \" - by (cases a, cases b) auto - -lemma pextreal_add_cancel_left: - "a + b = a + c \ (a = \ \ b = c)" - by (cases a, cases b, cases c, simp_all, cases c, simp_all) - -lemma pextreal_add_cancel_right: - "b + a = c + a \ (a = \ \ b = c)" - by (cases a, cases b, cases c, simp_all, cases c, simp_all) - -lemma Real_eq_Real: - "Real a = Real b \ (a = b \ (a \ 0 \ b \ 0))" -proof (cases "a \ 0 \ b \ 0") - case False with Real_inj[of a b] show ?thesis by auto -next - case True - thus ?thesis - proof - assume "a \ 0" - hence *: "Real a = 0" by simp - show ?thesis using `a \ 0` unfolding * by auto - next - assume "b \ 0" - hence *: "Real b = 0" by simp - show ?thesis using `b \ 0` unfolding * by auto - qed -qed - -lemma real_pextreal_0[simp]: "real (0 :: pextreal) = 0" - unfolding zero_pextreal_def real_Real by simp - -lemma real_of_pextreal_eq_0: "real X = 0 \ (X = 0 \ X = \)" - by (cases X) auto - -lemma real_of_pextreal_eq: "real X = real Y \ - (X = Y \ (X = 0 \ Y = \) \ (Y = 0 \ X = \))" - by (cases X, cases Y) (auto simp add: real_of_pextreal_eq_0) - -lemma real_of_pextreal_add: "real X + real Y = - (if X = \ then real Y else if Y = \ then real X else real (X + Y))" - by (auto simp: pextreal_noteq_omega_Ex) - -subsection "@{typ pextreal} is a monoid for multiplication" - -instantiation pextreal :: comm_monoid_mult -begin - -definition "1 = Real 1" -definition "x * y = (if x = 0 \ y = 0 then 0 else - pextreal_case (\r. pextreal_case (\p. Real (r * p)) \ y) \ x)" - -lemma pextreal_times[simp]: - "Real r * Real p = (if 0 \ r \ 0 \ p then Real (r * p) else 0)" - "\ * x = (if x = 0 then 0 else \)" - "x * \ = (if x = 0 then 0 else \)" - "0 * x = 0" - "x * 0 = 0" - "1 = \ \ False" - "\ = 1 \ False" - by (auto simp add: times_pextreal_def one_pextreal_def) - -lemma pextreal_one_mult[simp]: - "Real x + 1 = (if 0 \ x then Real (x + 1) else 1)" - "1 + Real x = (if 0 \ x then Real (1 + x) else 1)" - unfolding one_pextreal_def by simp_all - -instance -proof - fix a :: pextreal show "1 * a = a" - by (cases a) (simp_all add: one_pextreal_def) - - fix b show "a * b = b * a" - by (cases a, cases b) (simp_all add: mult_nonneg_nonneg) - - fix c show "a * b * c = a * (b * c)" - apply (cases a, cases b, cases c) - apply (simp_all add: mult_nonneg_nonneg not_le mult_pos_pos) - apply (cases b, cases c) - apply (simp_all add: mult_nonneg_nonneg not_le mult_pos_pos) - done -qed -end - -lemma pextreal_mult_cancel_left: - "a * b = a * c \ (a = 0 \ b = c \ (a = \ \ b \ 0 \ c \ 0))" - by (cases a, cases b, cases c, auto simp: Real_eq_Real mult_le_0_iff, cases c, auto) - -lemma pextreal_mult_cancel_right: - "b * a = c * a \ (a = 0 \ b = c \ (a = \ \ b \ 0 \ c \ 0))" - by (cases a, cases b, cases c, auto simp: Real_eq_Real mult_le_0_iff, cases c, auto) - -lemma Real_1[simp]: "Real 1 = 1" by (simp add: one_pextreal_def) - -lemma real_pextreal_1[simp]: "real (1 :: pextreal) = 1" - unfolding one_pextreal_def real_Real by simp - -lemma real_of_pextreal_mult: "real X * real Y = real (X * Y :: pextreal)" - by (cases X, cases Y) (auto simp: zero_le_mult_iff) - -lemma Real_mult_nonneg: assumes "x \ 0" "y \ 0" - shows "Real (x * y) = Real x * Real y" using assms by auto - -lemma Real_setprod: assumes "\x\A. f x \ 0" shows "Real (setprod f A) = setprod (\x. Real (f x)) A" -proof(cases "finite A") - case True thus ?thesis using assms - proof(induct A) case (insert x A) - have "0 \ setprod f A" apply(rule setprod_nonneg) using insert by auto - thus ?case unfolding setprod_insert[OF insert(1-2)] apply- - apply(subst Real_mult_nonneg) prefer 3 apply(subst insert(3)[THEN sym]) - using insert by auto - qed auto -qed auto - -subsection "@{typ pextreal} is a linear order" - -instantiation pextreal :: linorder -begin - -definition "x < y \ pextreal_case (\i. pextreal_case (\j. i < j) True y) False x" -definition "x \ y \ pextreal_case (\j. pextreal_case (\i. i \ j) False x) True y" - -lemma pextreal_less[simp]: - "Real r < \" - "Real r < Real p \ (if 0 \ r \ 0 \ p then r < p else 0 < p)" - "\ < x \ False" - "0 < \" - "0 < Real r \ 0 < r" - "x < 0 \ False" - "0 < (1::pextreal)" - by (simp_all add: less_pextreal_def zero_pextreal_def one_pextreal_def del: Real_0 Real_1) - -lemma pextreal_less_eq[simp]: - "x \ \" - "Real r \ Real p \ (if 0 \ r \ 0 \ p then r \ p else r \ 0)" - "0 \ x" - by (simp_all add: less_eq_pextreal_def zero_pextreal_def del: Real_0) - -lemma pextreal_\_less_eq[simp]: - "\ \ x \ x = \" - by (cases x) (simp_all add: not_le less_eq_pextreal_def) - -lemma pextreal_less_eq_zero[simp]: - "(x::pextreal) \ 0 \ x = 0" - by (cases x) (simp_all add: zero_pextreal_def del: Real_0) - -instance -proof - fix x :: pextreal - show "x \ x" by (cases x) simp_all - fix y - show "(x < y) = (x \ y \ \ y \ x)" - by (cases x, cases y) auto - show "x \ y \ y \ x " - by (cases x, cases y) auto - { assume "x \ y" "y \ x" thus "x = y" - by (cases x, cases y) auto } - { fix z assume "x \ y" "y \ z" - thus "x \ z" by (cases x, cases y, cases z) auto } -qed -end - -lemma pextreal_zero_lessI[intro]: - "(a :: pextreal) \ 0 \ 0 < a" - by (cases a) auto - -lemma pextreal_less_omegaI[intro, simp]: - "a \ \ \ a < \" - by (cases a) auto - -lemma pextreal_plus_eq_0[simp]: "(a :: pextreal) + b = 0 \ a = 0 \ b = 0" - by (cases a, cases b) auto - -lemma pextreal_le_add1[simp, intro]: "n \ n + (m::pextreal)" - by (cases n, cases m) simp_all - -lemma pextreal_le_add2: "(n::pextreal) + m \ k \ m \ k" - by (cases n, cases m, cases k) simp_all - -lemma pextreal_le_add3: "(n::pextreal) + m \ k \ n \ k" - by (cases n, cases m, cases k) simp_all - -lemma pextreal_less_\: "x < \ \ x \ \" - by (cases x) auto - -lemma pextreal_0_less_mult_iff[simp]: - fixes x y :: pextreal shows "0 < x * y \ 0 < x \ 0 < y" - by (cases x, cases y) (auto simp: zero_less_mult_iff) - -lemma pextreal_ord_one[simp]: - "Real p < 1 \ p < 1" - "Real p \ 1 \ p \ 1" - "1 < Real p \ 1 < p" - "1 \ Real p \ 1 \ p" - by (simp_all add: one_pextreal_def del: Real_1) - -subsection {* @{text "x - y"} on @{typ pextreal} *} - -instantiation pextreal :: minus -begin -definition "x - y = (if y < x then THE d. x = y + d else 0 :: pextreal)" - -lemma minus_pextreal_eq: - "(x - y = (z :: pextreal)) \ (if y < x then x = y + z else z = 0)" - (is "?diff \ ?if") -proof - assume ?diff - thus ?if - proof (cases "y < x") - case True - then obtain p where p: "y = Real p" "0 \ p" by (cases y) auto - - show ?thesis unfolding `?diff`[symmetric] if_P[OF True] minus_pextreal_def - proof (rule theI2[where Q="\d. x = y + d"]) - show "x = y + pextreal_case (\r. Real (r - real y)) \ x" (is "x = y + ?d") - using `y < x` p by (cases x) simp_all - - fix d assume "x = y + d" - thus "d = ?d" using `y < x` p by (cases d, cases x) simp_all - qed simp - qed (simp add: minus_pextreal_def) -next - assume ?if - thus ?diff - proof (cases "y < x") - case True - then obtain p where p: "y = Real p" "0 \ p" by (cases y) auto - - from True `?if` have "x = y + z" by simp - - show ?thesis unfolding minus_pextreal_def if_P[OF True] unfolding `x = y + z` - proof (rule the_equality) - fix d :: pextreal assume "y + z = y + d" - thus "d = z" using `y < x` p - by (cases d, cases z) simp_all - qed simp - qed (simp add: minus_pextreal_def) -qed - -instance .. -end - -lemma pextreal_minus[simp]: - "Real r - Real p = (if 0 \ r \ p < r then if 0 \ p then Real (r - p) else Real r else 0)" - "(A::pextreal) - A = 0" - "\ - Real r = \" - "Real r - \ = 0" - "A - 0 = A" - "0 - A = 0" - by (auto simp: minus_pextreal_eq not_less) - -lemma pextreal_le_epsilon: - fixes x y :: pextreal - assumes "\e. 0 < e \ x \ y + e" - shows "x \ y" -proof (cases y) - case (preal r) - then obtain p where x: "x = Real p" "0 \ p" - using assms[of 1] by (cases x) auto - { fix e have "0 < e \ p \ r + e" - using assms[of "Real e"] preal x by auto } - hence "p \ r" by (rule field_le_epsilon) - thus ?thesis using preal x by auto -qed simp - -instance pextreal :: "{ordered_comm_semiring, comm_semiring_1}" -proof - show "0 \ (1::pextreal)" unfolding zero_pextreal_def one_pextreal_def - by (simp del: Real_1 Real_0) - - fix a :: pextreal - show "0 * a = 0" "a * 0 = 0" by simp_all - - fix b c :: pextreal - show "(a + b) * c = a * c + b * c" - by (cases c, cases a, cases b) - (auto intro!: arg_cong[where f=Real] simp: field_simps not_le mult_le_0_iff mult_less_0_iff) - - { assume "a \ b" thus "c + a \ c + b" - by (cases c, cases a, cases b) auto } - - assume "a \ b" "0 \ c" - thus "c * a \ c * b" - apply (cases c, cases a, cases b) - by (auto simp: mult_left_mono mult_le_0_iff mult_less_0_iff not_le) -qed - -lemma mult_\[simp]: "x * y = \ \ (x = \ \ y = \) \ x \ 0 \ y \ 0" - by (cases x, cases y) auto - -lemma \_mult[simp]: "(\ = x * y) = ((x = \ \ y = \) \ x \ 0 \ y \ 0)" - by (cases x, cases y) auto - -lemma pextreal_mult_0[simp]: "x * y = 0 \ x = 0 \ (y::pextreal) = 0" - by (cases x, cases y) (auto simp: mult_le_0_iff) - -lemma pextreal_mult_cancel: - fixes x y z :: pextreal - assumes "y \ z" - shows "x * y \ x * z" - using assms - by (cases x, cases y, cases z) - (auto simp: mult_le_cancel_left mult_le_0_iff mult_less_0_iff not_le) - -lemma Real_power[simp]: - "Real x ^ n = (if x \ 0 then (if n = 0 then 1 else 0) else Real (x ^ n))" - by (induct n) auto - -lemma Real_power_\[simp]: - "\ ^ n = (if n = 0 then 1 else \)" - by (induct n) auto - -lemma pextreal_of_nat[simp]: "of_nat m = Real (real m)" - by (induct m) (auto simp: real_of_nat_Suc one_pextreal_def simp del: Real_1) - -lemma less_\_Ex_of_nat: "x < \ \ (\n. x < of_nat n)" -proof safe - assume "x < \" - then obtain r where "0 \ r" "x = Real r" by (cases x) auto - moreover obtain n where "r < of_nat n" using ex_less_of_nat by auto - ultimately show "\n. x < of_nat n" by (auto simp: real_eq_of_nat) -qed auto - -lemma real_of_pextreal_mono: - fixes a b :: pextreal - assumes "b \ \" "a \ b" - shows "real a \ real b" -using assms by (cases b, cases a) auto - -lemma setprod_pextreal_0: - "(\i\I. f i) = (0::pextreal) \ finite I \ (\i\I. f i = 0)" -proof cases - assume "finite I" then show ?thesis - proof (induct I) - case (insert i I) - then show ?case by simp - qed simp -qed simp - -lemma setprod_\: - "(\i\I. f i) = \ \ finite I \ (\i\I. f i = \) \ (\i\I. f i \ 0)" -proof cases - assume "finite I" then show ?thesis - proof (induct I) - case (insert i I) then show ?case - by (auto simp: setprod_pextreal_0) - qed simp -qed simp - -instance pextreal :: "semiring_char_0" -proof - fix m n - show "inj (of_nat::nat\pextreal)" by (auto intro!: inj_onI) -qed - -subsection "@{typ pextreal} is a complete lattice" - -instantiation pextreal :: lattice -begin -definition [simp]: "sup x y = (max x y :: pextreal)" -definition [simp]: "inf x y = (min x y :: pextreal)" -instance proof qed simp_all -end - -instantiation pextreal :: complete_lattice -begin - -definition "bot = Real 0" -definition "top = \" - -definition "Sup S = (LEAST z. \x\S. x \ z :: pextreal)" -definition "Inf S = (GREATEST z. \x\S. z \ x :: pextreal)" - -lemma pextreal_complete_Sup: - fixes S :: "pextreal set" assumes "S \ {}" - shows "\x. (\y\S. y \ x) \ (\z. (\y\S. y \ z) \ x \ z)" -proof (cases "\x\0. \a\S. a \ Real x") - case False - hence *: "\x. x\0 \ \a\S. \a \ Real x" by simp - show ?thesis - proof (safe intro!: exI[of _ \]) - fix y assume **: "\z\S. z \ y" - show "\ \ y" unfolding pextreal_\_less_eq - proof (rule ccontr) - assume "y \ \" - then obtain x where [simp]: "y = Real x" and "0 \ x" by (cases y) auto - from *[OF `0 \ x`] show False using ** by auto - qed - qed simp -next - case True then obtain y where y: "\a. a\S \ a \ Real y" and "0 \ y" by auto - from y[of \] have "\ \ S" by auto - - with `S \ {}` obtain x where "x \ S" and "x \ \" by auto - - have bound: "\x\real ` S. x \ y" - proof - fix z assume "z \ real ` S" then guess a .. - with y[of a] `\ \ S` `0 \ y` show "z \ y" by (cases a) auto - qed - with reals_complete2[of "real ` S"] `x \ S` - obtain s where s: "\y\S. real y \ s" "\z. ((\y\S. real y \ z) \ s \ z)" - by auto - - show ?thesis - proof (safe intro!: exI[of _ "Real s"]) - fix z assume "z \ S" thus "z \ Real s" - using s `\ \ S` by (cases z) auto - next - fix z assume *: "\y\S. y \ z" - show "Real s \ z" - proof (cases z) - case (preal u) - { fix v assume "v \ S" - hence "v \ Real u" using * preal by auto - hence "real v \ u" using `\ \ S` `0 \ u` by (cases v) auto } - hence "s \ u" using s(2) by auto - thus "Real s \ z" using preal by simp - qed simp - qed -qed - -lemma pextreal_complete_Inf: - fixes S :: "pextreal set" assumes "S \ {}" - shows "\x. (\y\S. x \ y) \ (\z. (\y\S. z \ y) \ z \ x)" -proof (cases "S = {\}") - case True thus ?thesis by (auto intro!: exI[of _ \]) -next - case False with `S \ {}` have "S - {\} \ {}" by auto - hence not_empty: "\x. x \ uminus ` real ` (S - {\})" by auto - have bounds: "\x. \y\uminus ` real ` (S - {\}). y \ x" by (auto intro!: exI[of _ 0]) - from reals_complete2[OF not_empty bounds] - obtain s where s: "\y. y\S - {\} \ - real y \ s" "\z. ((\y\S - {\}. - real y \ z) \ s \ z)" - by auto - - show ?thesis - proof (safe intro!: exI[of _ "Real (-s)"]) - fix z assume "z \ S" - show "Real (-s) \ z" - proof (cases z) - case (preal r) - with s `z \ S` have "z \ S - {\}" by simp - hence "- r \ s" using preal s(1)[of z] by auto - hence "- s \ r" by (subst neg_le_iff_le[symmetric]) simp - thus ?thesis using preal by simp - qed simp - next - fix z assume *: "\y\S. z \ y" - show "z \ Real (-s)" - proof (cases z) - case (preal u) - { fix v assume "v \ S-{\}" - hence "Real u \ v" using * preal by auto - hence "- real v \ - u" using `0 \ u` `v \ S - {\}` by (cases v) auto } - hence "u \ - s" using s(2) by (subst neg_le_iff_le[symmetric]) auto - thus "z \ Real (-s)" using preal by simp - next - case infinite - with * have "S = {\}" using `S \ {}` by auto - with `S - {\} \ {}` show ?thesis by simp - qed - qed -qed - -instance -proof - fix x :: pextreal and A - - show "bot \ x" by (cases x) (simp_all add: bot_pextreal_def) - show "x \ top" by (simp add: top_pextreal_def) - - { assume "x \ A" - with pextreal_complete_Sup[of A] - obtain s where s: "\y\A. y \ s" "\z. (\y\A. y \ z) \ s \ z" by auto - hence "x \ s" using `x \ A` by auto - also have "... = Sup A" using s unfolding Sup_pextreal_def - by (auto intro!: Least_equality[symmetric]) - finally show "x \ Sup A" . } - - { assume "x \ A" - with pextreal_complete_Inf[of A] - obtain i where i: "\y\A. i \ y" "\z. (\y\A. z \ y) \ z \ i" by auto - hence "Inf A = i" unfolding Inf_pextreal_def - by (auto intro!: Greatest_equality) - also have "i \ x" using i `x \ A` by auto - finally show "Inf A \ x" . } - - { assume *: "\z. z \ A \ z \ x" - show "Sup A \ x" - proof (cases "A = {}") - case True - hence "Sup A = 0" unfolding Sup_pextreal_def - by (auto intro!: Least_equality) - thus "Sup A \ x" by simp - next - case False - with pextreal_complete_Sup[of A] - obtain s where s: "\y\A. y \ s" "\z. (\y\A. y \ z) \ s \ z" by auto - hence "Sup A = s" - unfolding Sup_pextreal_def by (auto intro!: Least_equality) - also have "s \ x" using * s by auto - finally show "Sup A \ x" . - qed } - - { assume *: "\z. z \ A \ x \ z" - show "x \ Inf A" - proof (cases "A = {}") - case True - hence "Inf A = \" unfolding Inf_pextreal_def - by (auto intro!: Greatest_equality) - thus "x \ Inf A" by simp - next - case False - with pextreal_complete_Inf[of A] - obtain i where i: "\y\A. i \ y" "\z. (\y\A. z \ y) \ z \ i" by auto - have "x \ i" using * i by auto - also have "i = Inf A" using i - unfolding Inf_pextreal_def by (auto intro!: Greatest_equality[symmetric]) - finally show "x \ Inf A" . - qed } -qed -end - -lemma Inf_pextreal_iff: - fixes z :: pextreal - shows "(\x. x \ X \ z \ x) \ (\x\X. x Inf X < y" - by (metis complete_lattice_class.Inf_greatest complete_lattice_class.Inf_lower less_le_not_le linear - order_less_le_trans) - -lemma Inf_greater: - fixes z :: pextreal assumes "Inf X < z" - shows "\x \ X. x < z" -proof - - have "X \ {}" using assms by (auto simp: Inf_empty top_pextreal_def) - with assms show ?thesis - by (metis Inf_pextreal_iff mem_def not_leE) -qed - -lemma Inf_close: - fixes e :: pextreal assumes "Inf X \ \" "0 < e" - shows "\x \ X. x < Inf X + e" -proof (rule Inf_greater) - show "Inf X < Inf X + e" using assms - by (cases "Inf X", cases e) auto -qed - -lemma pextreal_SUPI: - fixes x :: pextreal - assumes "\i. i \ A \ f i \ x" - assumes "\y. (\i. i \ A \ f i \ y) \ x \ y" - shows "(SUP i:A. f i) = x" - unfolding SUPR_def Sup_pextreal_def - using assms by (auto intro!: Least_equality) - -lemma Sup_pextreal_iff: - fixes z :: pextreal - shows "(\x. x \ X \ x \ z) \ (\x\X. y y < Sup X" - by (metis complete_lattice_class.Sup_least complete_lattice_class.Sup_upper less_le_not_le linear - order_less_le_trans) - -lemma Sup_lesser: - fixes z :: pextreal assumes "z < Sup X" - shows "\x \ X. z < x" -proof - - have "X \ {}" using assms by (auto simp: Sup_empty bot_pextreal_def) - with assms show ?thesis - by (metis Sup_pextreal_iff mem_def not_leE) -qed - -lemma Sup_eq_\: "\ \ S \ Sup S = \" - unfolding Sup_pextreal_def - by (auto intro!: Least_equality) - -lemma Sup_close: - assumes "0 < e" and S: "Sup S \ \" "S \ {}" - shows "\X\S. Sup S < X + e" -proof cases - assume "Sup S = 0" - moreover obtain X where "X \ S" using `S \ {}` by auto - ultimately show ?thesis using `0 < e` by (auto intro!: bexI[OF _ `X\S`]) -next - assume "Sup S \ 0" - have "\X\S. Sup S - e < X" - proof (rule Sup_lesser) - show "Sup S - e < Sup S" using `0 < e` `Sup S \ 0` `Sup S \ \` - by (cases e) (auto simp: pextreal_noteq_omega_Ex) - qed - then guess X .. note X = this - with `Sup S \ \` Sup_eq_\ have "X \ \" by auto - thus ?thesis using `Sup S \ \` X unfolding pextreal_noteq_omega_Ex - by (cases e) (auto intro!: bexI[OF _ `X\S`] simp: split: split_if_asm) -qed - -lemma Sup_\: "(SUP i::nat. Real (real i)) = \" -proof (rule pextreal_SUPI) - fix y assume *: "\i::nat. i \ UNIV \ Real (real i) \ y" - thus "\ \ y" - proof (cases y) - case (preal r) - then obtain k :: nat where "r < real k" - using ex_less_of_nat by (auto simp: real_eq_of_nat) - with *[of k] preal show ?thesis by auto - qed simp -qed simp - -lemma SUP_\: "(SUP i:A. f i) = \ \ (\x<\. \i\A. x < f i)" -proof - assume *: "(SUP i:A. f i) = \" - show "(\x<\. \i\A. x < f i)" unfolding *[symmetric] - proof (intro allI impI) - fix x assume "x < SUPR A f" then show "\i\A. x < f i" - unfolding less_SUP_iff by auto - qed -next - assume *: "\x<\. \i\A. x < f i" - show "(SUP i:A. f i) = \" - proof (rule pextreal_SUPI) - fix y assume **: "\i. i \ A \ f i \ y" - show "\ \ y" - proof cases - assume "y < \" - from *[THEN spec, THEN mp, OF this] - obtain i where "i \ A" "\ (f i \ y)" by auto - with ** show ?thesis by auto - qed auto - qed auto -qed - -subsubsection {* Equivalence between @{text "f ----> x"} and @{text SUP} on @{typ pextreal} *} - -lemma monoseq_monoI: "mono f \ monoseq f" - unfolding mono_def monoseq_def by auto - -lemma incseq_mono: "mono f \ incseq f" - unfolding mono_def incseq_def by auto - -lemma SUP_eq_LIMSEQ: - assumes "mono f" and "\n. 0 \ f n" and "0 \ x" - shows "(SUP n. Real (f n)) = Real x \ f ----> x" -proof - assume x: "(SUP n. Real (f n)) = Real x" - { fix n - have "Real (f n) \ Real x" using x[symmetric] by (auto intro: le_SUPI) - hence "f n \ x" using assms by simp } - show "f ----> x" - proof (rule LIMSEQ_I) - fix r :: real assume "0 < r" - show "\no. \n\no. norm (f n - x) < r" - proof (rule ccontr) - assume *: "\ ?thesis" - { fix N - from * obtain n where "N \ n" "r \ x - f n" - using `\n. f n \ x` by (auto simp: not_less) - hence "f N \ f n" using `mono f` by (auto dest: monoD) - hence "f N \ x - r" using `r \ x - f n` by auto - hence "Real (f N) \ Real (x - r)" and "r \ x" using `0 \ f N` by auto } - hence "(SUP n. Real (f n)) \ Real (x - r)" - and "Real (x - r) < Real x" using `0 < r` by (auto intro: SUP_leI) - hence "(SUP n. Real (f n)) < Real x" by (rule le_less_trans) - thus False using x by auto - qed - qed -next - assume "f ----> x" - show "(SUP n. Real (f n)) = Real x" - proof (rule pextreal_SUPI) - fix n - from incseq_le[of f x] `mono f` `f ----> x` - show "Real (f n) \ Real x" using assms incseq_mono by auto - next - fix y assume *: "\n. n\UNIV \ Real (f n) \ y" - show "Real x \ y" - proof (cases y) - case (preal r) - with * have "\N. \n\N. f n \ r" using assms by fastsimp - from LIMSEQ_le_const2[OF `f ----> x` this] - show "Real x \ y" using `0 \ x` preal by auto - qed simp - qed -qed - -lemma SUPR_bound: - assumes "\N. f N \ x" - shows "(SUP n. f n) \ x" - using assms by (simp add: SUPR_def Sup_le_iff) - -lemma pextreal_less_eq_diff_eq_sum: - fixes x y z :: pextreal - assumes "y \ x" and "x \ \" - shows "z \ x - y \ z + y \ x" - using assms - apply (cases z, cases y, cases x) - by (simp_all add: field_simps minus_pextreal_eq) - -lemma Real_diff_less_omega: "Real r - x < \" by (cases x) auto - -subsubsection {* Numbers on @{typ pextreal} *} - -instantiation pextreal :: number -begin -definition [simp]: "number_of x = Real (number_of x)" -instance proof qed -end - -subsubsection {* Division on @{typ pextreal} *} - -instantiation pextreal :: inverse -begin - -definition "inverse x = pextreal_case (\x. if x = 0 then \ else Real (inverse x)) 0 x" -definition [simp]: "x / y = x * inverse (y :: pextreal)" - -instance proof qed -end - -lemma pextreal_inverse[simp]: - "inverse 0 = \" - "inverse (Real x) = (if x \ 0 then \ else Real (inverse x))" - "inverse \ = 0" - "inverse (1::pextreal) = 1" - "inverse (inverse x) = x" - by (simp_all add: inverse_pextreal_def one_pextreal_def split: pextreal_case_split del: Real_1) - -lemma pextreal_inverse_le_eq: - assumes "x \ 0" "x \ \" - shows "y \ z / x \ x * y \ (z :: pextreal)" -proof - - from assms obtain r where r: "x = Real r" "0 < r" by (cases x) auto - { fix p q :: real assume "0 \ p" "0 \ q" - have "p \ q * inverse r \ p \ q / r" by (simp add: divide_inverse) - also have "... \ p * r \ q" using `0 < r` by (auto simp: field_simps) - finally have "p \ q * inverse r \ p * r \ q" . } - with r show ?thesis - by (cases y, cases z, auto simp: zero_le_mult_iff field_simps) -qed - -lemma inverse_antimono_strict: - fixes x y :: pextreal - assumes "x < y" shows "inverse y < inverse x" - using assms by (cases x, cases y) auto - -lemma inverse_antimono: - fixes x y :: pextreal - assumes "x \ y" shows "inverse y \ inverse x" - using assms by (cases x, cases y) auto - -lemma pextreal_inverse_\_iff[simp]: "inverse x = \ \ x = 0" - by (cases x) auto - -subsection "Infinite sum over @{typ pextreal}" - -text {* - -The infinite sum over @{typ pextreal} has the nice property that it is always -defined. - -*} - -definition psuminf :: "(nat \ pextreal) \ pextreal" (binder "\\<^isub>\" 10) where - "(\\<^isub>\ x. f x) = (SUP n. \i n. f n"} and @{text "\\<^isub>\ n. f n"} *} - -lemma setsum_Real: - assumes "\x\A. 0 \ f x" - shows "(\x\A. Real (f x)) = Real (\x\A. f x)" -proof (cases "finite A") - case True - thus ?thesis using assms - proof induct case (insert x s) - hence "0 \ setsum f s" apply-apply(rule setsum_nonneg) by auto - thus ?case using insert by auto - qed auto -qed simp - -lemma setsum_Real': - assumes "\x. 0 \ f x" - shows "(\x\A. Real (f x)) = Real (\x\A. f x)" - apply(rule setsum_Real) using assms by auto - -lemma setsum_\: - "(\x\P. f x) = \ \ (finite P \ (\i\P. f i = \))" -proof safe - assume *: "setsum f P = \" - show "finite P" - proof (rule ccontr) assume "infinite P" with * show False by auto qed - show "\i\P. f i = \" - proof (rule ccontr) - assume "\ ?thesis" hence "\i. i\P \ f i \ \" by auto - from `finite P` this have "setsum f P \ \" - by induct auto - with * show False by auto - qed -next - fix i assume "finite P" "i \ P" "f i = \" - thus "setsum f P = \" - proof induct - case (insert x A) - show ?case using insert by (cases "x = i") auto - qed simp -qed - -lemma real_of_pextreal_setsum: - assumes "\x. x \ S \ f x \ \" - shows "(\x\S. real (f x)) = real (setsum f S)" -proof cases - assume "finite S" - from this assms show ?thesis - by induct (simp_all add: real_of_pextreal_add setsum_\) -qed simp - -lemma setsum_0: - fixes f :: "'a \ pextreal" assumes "finite A" - shows "(\x\A. f x) = 0 \ (\i\A. f i = 0)" - using assms by induct auto - -lemma suminf_imp_psuminf: - assumes "f sums x" and "\n. 0 \ f n" - shows "(\\<^isub>\ x. Real (f x)) = Real x" - unfolding psuminf_def setsum_Real'[OF assms(2)] -proof (rule SUP_eq_LIMSEQ[THEN iffD2]) - show "mono (\n. setsum f {.. ?S n" - using setsum_nonneg[of "{.. x" "?S ----> x" - using `f sums x` LIMSEQ_le_const - by (auto simp: atLeast0LessThan sums_def) -qed - -lemma psuminf_equality: - assumes "\n. setsum f {.. x" - and "\y. y \ \ \ (\n. setsum f {.. y) \ x \ y" - shows "psuminf f = x" - unfolding psuminf_def -proof (safe intro!: pextreal_SUPI) - fix n show "setsum f {.. x" using assms(1) . -next - fix y assume *: "\n. n \ UNIV \ setsum f {.. y" - show "x \ y" - proof (cases "y = \") - assume "y \ \" from assms(2)[OF this] * - show "x \ y" by auto - qed simp -qed - -lemma psuminf_\: - assumes "f i = \" - shows "(\\<^isub>\ x. f x) = \" -proof (rule psuminf_equality) - fix y assume *: "\n. setsum f {.. y" - have "setsum f {.." - using assms by (simp add: setsum_\) - thus "\ \ y" using *[of "Suc i"] by auto -qed simp - -lemma psuminf_imp_suminf: - assumes "(\\<^isub>\ x. f x) \ \" - shows "(\x. real (f x)) sums real (\\<^isub>\ x. f x)" -proof - - have "\i. \r. f i = Real r \ 0 \ r" - proof - fix i show "\r. f i = Real r \ 0 \ r" using psuminf_\ assms by (cases "f i") auto - qed - from choice[OF this] obtain r where f: "f = (\i. Real (r i))" - and pos: "\i. 0 \ r i" - by (auto simp: fun_eq_iff) - hence [simp]: "\i. real (f i) = r i" by auto - - have "mono (\n. setsum r {.. ?S n" - using setsum_nonneg[of "{..\<^isub>\ x. f x) = Real p" and "0 \ p" - by (cases "(\\<^isub>\ x. f x)") auto - show ?thesis unfolding * using * pos `0 \ p` SUP_eq_LIMSEQ[OF `mono ?S` `\n. 0 \ ?S n` `0 \ p`] - by (simp add: f atLeast0LessThan sums_def psuminf_def setsum_Real'[OF pos] f) -qed - -lemma psuminf_bound: - assumes "\N. (\n x" - shows "(\\<^isub>\ n. f n) \ x" - using assms by (simp add: psuminf_def SUPR_def Sup_le_iff) - -lemma psuminf_bound_add: - assumes "\N. (\n x" - shows "(\\<^isub>\ n. f n) + y \ x" -proof (cases "x = \") - have "y \ x" using assms by (auto intro: pextreal_le_add2) - assume "x \ \" - note move_y = pextreal_less_eq_diff_eq_sum[OF `y \ x` this] - - have "\N. (\n x - y" using assms by (simp add: move_y) - hence "(\\<^isub>\ n. f n) \ x - y" by (rule psuminf_bound) - thus ?thesis by (simp add: move_y) -qed simp - -lemma psuminf_finite: - assumes "\N\n. f N = 0" - shows "(\\<^isub>\ n. f n) = (\N setsum f {.. {n..n (\\<^isub>\ n. f n)" - unfolding psuminf_def SUPR_def - by (auto intro: complete_lattice_class.Sup_upper image_eqI) - -lemma psuminf_le: - assumes "\N. f N \ g N" - shows "psuminf f \ psuminf g" -proof (safe intro!: psuminf_bound) - fix n - have "setsum f {.. setsum g {.. psuminf g" by (rule psuminf_upper) - finally show "setsum f {.. psuminf g" . -qed - -lemma psuminf_const[simp]: "psuminf (\n. c) = (if c = 0 then 0 else \)" (is "_ = ?if") -proof (rule psuminf_equality) - fix y assume *: "\n :: nat. (\n y" and "y \ \" - then obtain r p where - y: "y = Real r" "0 \ r" and - c: "c = Real p" "0 \ p" - using *[of 1] by (cases c, cases y) auto - show "(if c = 0 then 0 else \) \ y" - proof (cases "p = 0") - assume "p = 0" with c show ?thesis by simp - next - assume "p \ 0" - with * c y have **: "\n :: nat. real n \ r / p" - by (auto simp: zero_le_mult_iff field_simps) - from ex_less_of_nat[of "r / p"] guess n .. - with **[of n] show ?thesis by (simp add: real_eq_of_nat) - qed -qed (cases "c = 0", simp_all) - -lemma psuminf_add[simp]: "psuminf (\n. f n + g n) = psuminf f + psuminf g" -proof (rule psuminf_equality) - fix n - from psuminf_upper[of f n] psuminf_upper[of g n] - show "(\n psuminf f + psuminf g" - by (auto simp add: setsum_addf intro!: add_mono) -next - fix y assume *: "\n. (\n y" and "y \ \" - { fix n m - have **: "(\nn y" - proof (cases rule: linorder_le_cases) - assume "n \ m" - hence "(\nn (\nn y" - using * by (simp add: setsum_addf) - finally show ?thesis . - next - assume "m \ n" - hence "(\nn (\nn y" - using * by (simp add: setsum_addf) - finally show ?thesis . - qed } - hence "\m. \n. (\nn y" by simp - from psuminf_bound_add[OF this] - have "\m. (\n y" by (simp add: ac_simps) - from psuminf_bound_add[OF this] - show "psuminf f + psuminf g \ y" by (simp add: ac_simps) -qed - -lemma psuminf_0: "psuminf f = 0 \ (\i. f i = 0)" -proof safe - assume "\i. f i = 0" - hence "f = (\i. 0)" by (simp add: fun_eq_iff) - thus "psuminf f = 0" using psuminf_const by simp -next - fix i assume "psuminf f = 0" - hence "(\nn. c * f n) = c * psuminf f" -proof (rule psuminf_equality) - fix n show "(\n c * psuminf f" - by (auto simp: setsum_right_distrib[symmetric] intro: mult_left_mono psuminf_upper) -next - fix y - assume "\n. (\n y" - hence *: "\n. c * (\n y" by (auto simp add: setsum_right_distrib) - thus "c * psuminf f \ y" - proof (cases "c = \ \ c = 0") - assume "c = \ \ c = 0" - thus ?thesis - using * by (fastsimp simp add: psuminf_0 setsum_0 split: split_if_asm) - next - assume "\ (c = \ \ c = 0)" - hence "c \ 0" "c \ \" by auto - note rewrite_div = pextreal_inverse_le_eq[OF this, of _ y] - hence "\n. (\n y / c" using * by simp - hence "psuminf f \ y / c" by (rule psuminf_bound) - thus ?thesis using rewrite_div by simp - qed -qed - -lemma psuminf_cmult_left[simp]: "psuminf (\n. f n * c) = psuminf f * c" - using psuminf_cmult_right[of c f] by (simp add: ac_simps) - -lemma psuminf_half_series: "psuminf (\n. (1/2)^Suc n) = 1" - using suminf_imp_psuminf[OF power_half_series] by auto - -lemma setsum_pinfsum: "(\\<^isub>\ n. \m\A. f n m) = (\m\A. (\\<^isub>\ n. f n m))" -proof (cases "finite A") - assume "finite A" - thus ?thesis by induct simp_all -qed simp - -lemma psuminf_reindex: - fixes f:: "nat \ nat" assumes "bij f" - shows "psuminf (g \ f) = psuminf g" -proof - - have [intro, simp]: "\A. inj_on f A" using `bij f` unfolding bij_def by (auto intro: subset_inj_on) - have f[intro, simp]: "\x. f (inv f x) = x" - using `bij f` unfolding bij_def by (auto intro: surj_f_inv_f) - show ?thesis - proof (rule psuminf_equality) - fix n - have "setsum (g \ f) {.. \ setsum g {..Max (f ` {.. \ psuminf g" unfolding lessThan_Suc_atMost[symmetric] by (rule psuminf_upper) - finally show "setsum (g \ f) {.. psuminf g" . - next - fix y assume *: "\n. setsum (g \ f) {.. y" - show "psuminf g \ y" - proof (safe intro!: psuminf_bound) - fix N - have "setsum g {.. setsum g (f ` {..Max (inv f ` {.. = setsum (g \ f) {..Max (inv f ` {.. \ y" unfolding lessThan_Suc_atMost[symmetric] by (rule *) - finally show "setsum g {.. y" . - qed - qed -qed - -lemma pextreal_mult_less_right: - assumes "b * a < c * a" "0 < a" "a < \" - shows "b < c" - using assms - by (cases a, cases b, cases c) (auto split: split_if_asm simp: zero_less_mult_iff zero_le_mult_iff) - -lemma pextreal_\_eq_plus[simp]: "\ = a + b \ (a = \ \ b = \)" - by (cases a, cases b) auto - -lemma pextreal_of_nat_le_iff: - "(of_nat k :: pextreal) \ of_nat m \ k \ m" by auto - -lemma pextreal_of_nat_less_iff: - "(of_nat k :: pextreal) < of_nat m \ k < m" by auto - -lemma pextreal_bound_add: - assumes "\N. f N + y \ (x::pextreal)" - shows "(SUP n. f n) + y \ x" -proof (cases "x = \") - have "y \ x" using assms by (auto intro: pextreal_le_add2) - assume "x \ \" - note move_y = pextreal_less_eq_diff_eq_sum[OF `y \ x` this] - - have "\N. f N \ x - y" using assms by (simp add: move_y) - hence "(SUP n. f n) \ x - y" by (rule SUPR_bound) - thus ?thesis by (simp add: move_y) -qed simp - -lemma SUPR_pextreal_add: - fixes f g :: "nat \ pextreal" - assumes f: "\n. f n \ f (Suc n)" and g: "\n. g n \ g (Suc n)" - shows "(SUP n. f n + g n) = (SUP n. f n) + (SUP n. g n)" -proof (rule pextreal_SUPI) - fix n :: nat from le_SUPI[of n UNIV f] le_SUPI[of n UNIV g] - show "f n + g n \ (SUP n. f n) + (SUP n. g n)" - by (auto intro!: add_mono) -next - fix y assume *: "\n. n \ UNIV \ f n + g n \ y" - { fix n m - have "f n + g m \ y" - proof (cases rule: linorder_le_cases) - assume "n \ m" - hence "f n + g m \ f m + g m" - using f lift_Suc_mono_le by (auto intro!: add_right_mono) - also have "\ \ y" using * by simp - finally show ?thesis . - next - assume "m \ n" - hence "f n + g m \ f n + g n" - using g lift_Suc_mono_le by (auto intro!: add_left_mono) - also have "\ \ y" using * by simp - finally show ?thesis . - qed } - hence "\m. \n. f n + g m \ y" by simp - from pextreal_bound_add[OF this] - have "\m. (g m) + (SUP n. f n) \ y" by (simp add: ac_simps) - from pextreal_bound_add[OF this] - show "SUPR UNIV f + SUPR UNIV g \ y" by (simp add: ac_simps) -qed - -lemma SUPR_pextreal_setsum: - fixes f :: "'x \ nat \ pextreal" - assumes "\i. i \ P \ \n. f i n \ f i (Suc n)" - shows "(SUP n. \i\P. f i n) = (\i\P. SUP n. f i n)" -proof cases - assume "finite P" from this assms show ?thesis - proof induct - case (insert i P) - thus ?case - apply simp - apply (subst SUPR_pextreal_add) - by (auto intro!: setsum_mono) - qed simp -qed simp - -lemma psuminf_SUP_eq: - assumes "\n i. f n i \ f (Suc n) i" - shows "(\\<^isub>\ i. SUP n::nat. f n i) = (SUP n::nat. \\<^isub>\ i. f n i)" -proof - - { fix n :: nat - have "(\ii\<^isub>\ i j. f i j) = (\\<^isub>\ j i. f i j)" -proof - - have "(SUP n. \ i < n. SUP m. \ j < m. f i j) = (SUP n. SUP m. \ i < n. \ j < m. f i j)" - apply (subst SUPR_pextreal_setsum) - by auto - also have "\ = (SUP m n. \ j < m. \ i < n. f i j)" - apply (subst SUP_commute) - apply (subst setsum_commute) - by auto - also have "\ = (SUP m. \ j < m. SUP n. \ i < n. f i j)" - apply (subst SUPR_pextreal_setsum) - by auto - finally show ?thesis - unfolding psuminf_def by auto -qed - -lemma psuminf_2dimen: - fixes f:: "nat * nat \ pextreal" - assumes fsums: "\m. g m = (\\<^isub>\ n. f (m,n))" - shows "psuminf (f \ prod_decode) = psuminf g" -proof (rule psuminf_equality) - fix n :: nat - let ?P = "prod_decode ` {.. prod_decode) {.. \ setsum f ({..Max (fst ` ?P)} \ {..Max (snd ` ?P)})" - proof (safe intro!: setsum_mono3 Max_ge image_eqI) - fix a b x assume "(a, b) = prod_decode x" - from this[symmetric] show "a = fst (prod_decode x)" "b = snd (prod_decode x)" - by simp_all - qed simp_all - also have "\ = (\m\Max (fst ` ?P). (\n\Max (snd ` ?P). f (m,n)))" - unfolding setsum_cartesian_product by simp - also have "\ \ (\m\Max (fst ` ?P). g m)" - by (auto intro!: setsum_mono psuminf_upper simp del: setsum_lessThan_Suc - simp: fsums lessThan_Suc_atMost[symmetric]) - also have "\ \ psuminf g" - by (auto intro!: psuminf_upper simp del: setsum_lessThan_Suc - simp: lessThan_Suc_atMost[symmetric]) - finally show "setsum (f \ prod_decode) {.. psuminf g" . -next - fix y assume *: "\n. setsum (f \ prod_decode) {.. y" - have g: "g = (\m. \\<^isub>\ n. f (m,n))" unfolding fsums[symmetric] .. - show "psuminf g \ y" unfolding g - proof (rule psuminf_bound, unfold setsum_pinfsum[symmetric], safe intro!: psuminf_bound) - fix N M :: nat - let ?P = "{.. {..nm (\(m, n)\?P. f (m, n))" - unfolding setsum_commute[of _ _ "{.. \ (\(m,n)\(prod_decode ` {..?M}). f (m, n))" - by (auto intro!: setsum_mono3 image_eqI[where f=prod_decode, OF prod_encode_inverse[symmetric]]) - also have "\ \ y" using *[of "Suc ?M"] - by (simp add: lessThan_Suc_atMost[symmetric] setsum_reindex - inj_prod_decode del: setsum_lessThan_Suc) - finally show "(\nm y" . - qed -qed - -lemma Real_max: - assumes "x \ 0" "y \ 0" - shows "Real (max x y) = max (Real x) (Real y)" - using assms unfolding max_def by (auto simp add:not_le) - -lemma Real_real: "Real (real x) = (if x = \ then 0 else x)" - using assms by (cases x) auto - -lemma inj_on_real: "inj_on real (UNIV - {\})" -proof (rule inj_onI) - fix x y assume mem: "x \ UNIV - {\}" "y \ UNIV - {\}" and "real x = real y" - thus "x = y" by (cases x, cases y) auto -qed - -lemma inj_on_Real: "inj_on Real {0..}" - by (auto intro!: inj_onI) - -lemma range_Real[simp]: "range Real = UNIV - {\}" -proof safe - fix x assume "x \ range Real" - thus "x = \" by (cases x) auto -qed auto - -lemma image_Real[simp]: "Real ` {0..} = UNIV - {\}" -proof safe - fix x assume "x \ Real ` {0..}" - thus "x = \" by (cases x) auto -qed auto - -lemma pextreal_SUP_cmult: - fixes f :: "'a \ pextreal" - shows "(SUP i : R. z * f i) = z * (SUP i : R. f i)" -proof (rule pextreal_SUPI) - fix i assume "i \ R" - from le_SUPI[OF this] - show "z * f i \ z * (SUP i:R. f i)" by (rule pextreal_mult_cancel) -next - fix y assume "\i. i\R \ z * f i \ y" - hence *: "\i. i\R \ z * f i \ y" by auto - show "z * (SUP i:R. f i) \ y" - proof (cases "\i\R. f i = 0") - case True - show ?thesis - proof cases - assume "R \ {}" hence "f ` R = {0}" using True by auto - thus ?thesis by (simp add: SUPR_def) - qed (simp add: SUPR_def Sup_empty bot_pextreal_def) - next - case False then obtain i where i: "i \ R" and f0: "f i \ 0" by auto - show ?thesis - proof (cases "z = 0 \ z = \") - case True with f0 *[OF i] show ?thesis by auto - next - case False hence z: "z \ 0" "z \ \" by auto - note div = pextreal_inverse_le_eq[OF this, symmetric] - hence "\i. i\R \ f i \ y / z" using * by auto - thus ?thesis unfolding div SUP_le_iff by simp - qed - qed -qed - -instantiation pextreal :: topological_space -begin - -definition "open A \ - (\T. open T \ (Real ` (T\{0..}) = A - {\})) \ (\ \ A \ (\x\0. {Real x <..} \ A))" - -lemma open_omega: "open A \ \ \ A \ (\x\0. {Real x<..} \ A)" - unfolding open_pextreal_def by auto - -lemma open_omegaD: assumes "open A" "\ \ A" obtains x where "x\0" "{Real x<..} \ A" - using open_omega[OF assms] by auto - -lemma pextreal_openE: assumes "open A" obtains A' x where - "open A'" "Real ` (A' \ {0..}) = A - {\}" - "x \ 0" "\ \ A \ {Real x<..} \ A" - using assms open_pextreal_def by auto - -instance -proof - let ?U = "UNIV::pextreal set" - show "open ?U" unfolding open_pextreal_def - by (auto intro!: exI[of _ "UNIV"] exI[of _ 0]) -next - fix S T::"pextreal set" assume "open S" and "open T" - from `open S`[THEN pextreal_openE] guess S' xS . note S' = this - from `open T`[THEN pextreal_openE] guess T' xT . note T' = this - - from S'(1-3) T'(1-3) - show "open (S \ T)" unfolding open_pextreal_def - proof (safe intro!: exI[of _ "S' \ T'"] exI[of _ "max xS xT"]) - fix x assume *: "Real (max xS xT) < x" and "\ \ S" "\ \ T" - from `\ \ S`[THEN S'(4)] * show "x \ S" - by (cases x, auto simp: max_def split: split_if_asm) - from `\ \ T`[THEN T'(4)] * show "x \ T" - by (cases x, auto simp: max_def split: split_if_asm) - next - fix x assume x: "x \ Real ` (S' \ T' \ {0..})" - have *: "S' \ T' \ {0..} = (S' \ {0..}) \ (T' \ {0..})" by auto - assume "x \ T" "x \ S" - with S'(2) T'(2) show "x = \" - using x[unfolded *] inj_on_image_Int[OF inj_on_Real] by auto - qed auto -next - fix K assume openK: "\S \ K. open (S:: pextreal set)" - hence "\S\K. \T. open T \ Real ` (T \ {0..}) = S - {\}" by (auto simp: open_pextreal_def) - from bchoice[OF this] guess T .. note T = this[rule_format] - - show "open (\K)" unfolding open_pextreal_def - proof (safe intro!: exI[of _ "\(T ` K)"]) - fix x S assume "0 \ x" "x \ T S" "S \ K" - with T[OF `S \ K`] show "Real x \ \K" by auto - next - fix x S assume x: "x \ Real ` (\T ` K \ {0..})" "S \ K" "x \ S" - hence "x \ Real ` (T S \ {0..})" - by (auto simp: image_UN UN_simps[symmetric] simp del: UN_simps) - thus "x = \" using T[OF `S \ K`] `x \ S` by auto - next - fix S assume "\ \ S" "S \ K" - from openK[rule_format, OF `S \ K`, THEN pextreal_openE] guess S' x . - from this(3, 4) `\ \ S` - show "\x\0. {Real x<..} \ \K" - by (auto intro!: exI[of _ x] bexI[OF _ `S \ K`]) - next - from T[THEN conjunct1] show "open (\T`K)" by auto - qed auto -qed -end - -lemma minus_omega[simp]: "x - \ = 0" by (cases x) auto - -lemma open_pextreal_alt: "open A \ - (\x\A. \e>0. {x-e <..< x+e} \ A) \ (\ \ A \ (\x\0. {Real x <..} \ A))" -proof - - have *: "(\T. open T \ (Real ` (T\{0..}) = A - {\})) \ - open (real ` (A - {\}) \ {..<0})" - proof safe - fix T assume "open T" and A: "Real ` (T \ {0..}) = A - {\}" - have *: "(\x. real (Real x)) ` (T \ {0..}) = T \ {0..}" - by auto - have **: "T \ {0..} \ {..<0} = T \ {..<0}" by auto - show "open (real ` (A - {\}) \ {..<0})" - unfolding A[symmetric] image_image * ** using `open T` by auto - next - assume "open (real ` (A - {\}) \ {..<0})" - moreover have "Real ` ((real ` (A - {\}) \ {..<0}) \ {0..}) = A - {\}" - apply auto - apply (case_tac xb) - apply auto - apply (case_tac x) - apply (auto simp: image_iff) - apply (erule_tac x="Real r" in ballE) - apply auto - done - ultimately show "\T. open T \ Real ` (T \ {0..}) = A - {\}" by auto - qed - also have "\ \ (\x\A. \e>0. {x-e <..< x+e} \ A)" - proof (intro iffI ballI open_subopen[THEN iffD2]) - fix x assume *: "\x\A. \e>0. {x - e<.. A" and x: "x \ real ` (A - {\}) \ {..<0}" - show "\T. open T \ x \ T \ T \ real ` (A - {\}) \ {..<0}" - proof (cases rule: linorder_cases) - assume "x < 0" then show ?thesis by (intro exI[of _ "{..<0}"]) auto - next - assume "x = 0" with x - have "0 \ A" - apply auto by (case_tac x) auto - with * obtain e where "e > 0" "{0 - e<..<0 + e} \ A" by auto - then have "{.. A" using `0 \ A` - apply auto - apply (case_tac "x = 0") - by (auto dest!: pextreal_zero_lessI) - then have *: "{.. A - {\}" by auto - def e' \ "if e = \ then 1 else real e" - then have "0 < e'" using `e > 0` by (cases e) auto - have "{0.. real ` (A - {\})" - proof (cases e) - case infinite - then have "{..}" by auto - then have A: "A - {\} = UNIV - {\}" using * by auto - show ?thesis unfolding e'_def infinite A - apply safe - apply (rule_tac x="Real x" in image_eqI) - apply auto - done - next - case (preal r) - then show ?thesis unfolding e'_def using * - apply safe - apply (rule_tac x="Real x" in image_eqI) - by (auto simp: subset_eq) - qed - then have "{0.. {..<0} \ real ` (A - {\}) \ {..<0}" by auto - moreover have "{0.. {..<0} = {.. real ` (A - {\}) \ {..<0}" by simp - then show ?thesis using `0 < e'` `x = 0` by auto - next - assume "0 < x" - with x have "Real x \ A" apply auto by (case_tac x) auto - with * obtain e where "0 < e" and e: "{Real x - e<.. A" by auto - show ?thesis - proof cases - assume "e < Real x" - with `0 < e` obtain r where r: "e = Real r" "0 < r" by (cases e) auto - then have "r < x" using `e < Real x` `0 < e` by (auto split: split_if_asm) - then have "{x - r <..< x + r} \ real ` (A - {\}) \ {..<0}" - using e unfolding r - apply (auto simp: subset_eq) - apply (rule_tac x="Real xa" in image_eqI) - by auto - then show ?thesis using `0 < r` by (intro exI[of _ "{x - r <..< x + r}"]) auto - next - assume "\ e < Real x" - moreover then have "Real x - e = 0" by (cases e) auto - moreover have "\z. 0 < z \ z * 2 < 3 * x \ Real z < Real x + e" - using `\ e < Real x` by (cases e) auto - ultimately have "{0 <..< x + x / 2} \ real ` (A - {\}) \ {..<0}" - using e - apply (auto simp: subset_eq) - apply (erule_tac x="Real xa" in ballE) - apply (auto simp: not_less) - apply (rule_tac x="Real xa" in image_eqI) - apply auto - done - moreover have "x \ {0 <..< x + x / 2}" using `0 < x` by auto - ultimately show ?thesis by (intro exI[of _ "{0 <..< x + x / 2}"]) auto - qed - qed - next - fix x assume x: "x \ A" "open (real ` (A - {\}) \ {..<0})" - then show "\e>0. {x - e<.. A" - proof (cases x) - case infinite then show ?thesis by (intro exI[of _ 2]) auto - next - case (preal r) - with `x \ A` have r: "r \ real ` (A - {\}) \ {..<0}" by force - from x(2)[unfolded open_real, THEN bspec, OF r] - obtain e where e: "0 < e" "\x'. \x' - r\ < e \ x' \ real ` (A - {\}) \ {..<0}" - by auto - show ?thesis using `0 < e` preal - proof (auto intro!: exI[of _ "Real e"] simp: greaterThanLessThan_iff not_less - split: split_if_asm) - fix z assume *: "Real (r - e) < z" "z < Real (r + e)" - then obtain q where [simp]: "z = Real q" "0 \ q" by (cases z) auto - with * have "r - e < q" "q < r + e" by (auto split: split_if_asm) - with e(2)[of q] have "q \ real ` (A - {\}) \ {..<0}" by auto - then show "z \ A" using `0 \ q` apply auto apply (case_tac x) apply auto done - next - fix z assume *: "0 < z" "z < Real (r + e)" "r \ e" - then obtain q where [simp]: "z = Real q" and "0 < q" by (cases z) auto - with * have "q < r + e" by (auto split: split_if_asm) - moreover have "r - e < q" using `r \ e` `0 < q` by auto - ultimately have "q \ real ` (A - {\}) \ {..<0}" using e(2)[of q] by auto - then show "z \ A" using `0 < q` apply auto apply (case_tac x) apply auto done - qed - qed - qed - finally show ?thesis unfolding open_pextreal_def by simp -qed - -lemma open_pextreal_lessThan[simp]: - "open {..< a :: pextreal}" -proof (cases a) - case (preal x) thus ?thesis unfolding open_pextreal_def - proof (safe intro!: exI[of _ "{..< x}"]) - fix y assume "y < Real x" - moreover assume "y \ Real ` ({.. {0..})" - ultimately have "y \ Real (real y)" using preal by (cases y) auto - thus "y = \" by (auto simp: Real_real split: split_if_asm) - qed auto -next - case infinite thus ?thesis - unfolding open_pextreal_def by (auto intro!: exI[of _ UNIV]) -qed - -lemma open_pextreal_greaterThan[simp]: - "open {a :: pextreal <..}" -proof (cases a) - case (preal x) thus ?thesis unfolding open_pextreal_def - proof (safe intro!: exI[of _ "{x <..}"]) - fix y assume "Real x < y" - moreover assume "y \ Real ` ({x<..} \ {0..})" - ultimately have "y \ Real (real y)" using preal by (cases y) auto - thus "y = \" by (auto simp: Real_real split: split_if_asm) - qed auto -next - case infinite thus ?thesis - unfolding open_pextreal_def by (auto intro!: exI[of _ "{}"]) -qed - -lemma pextreal_open_greaterThanLessThan[simp]: "open {a::pextreal <..< b}" - unfolding greaterThanLessThan_def by auto - -lemma closed_pextreal_atLeast[simp, intro]: "closed {a :: pextreal ..}" -proof - - have "- {a ..} = {..< a}" by auto - then show "closed {a ..}" - unfolding closed_def using open_pextreal_lessThan by auto -qed - -lemma closed_pextreal_atMost[simp, intro]: "closed {.. b :: pextreal}" -proof - - have "- {.. b} = {b <..}" by auto - then show "closed {.. b}" - unfolding closed_def using open_pextreal_greaterThan by auto -qed - -lemma closed_pextreal_atLeastAtMost[simp, intro]: - shows "closed {a :: pextreal .. b}" - unfolding atLeastAtMost_def by auto - -lemma pextreal_dense: - fixes x y :: pextreal assumes "x < y" - shows "\z. x < z \ z < y" -proof - - from `x < y` obtain p where p: "x = Real p" "0 \ p" by (cases x) auto - show ?thesis - proof (cases y) - case (preal r) with p `x < y` have "p < r" by auto - with dense obtain z where "p < z" "z < r" by auto - thus ?thesis using preal p by (auto intro!: exI[of _ "Real z"]) - next - case infinite thus ?thesis using `x < y` p - by (auto intro!: exI[of _ "Real p + 1"]) - qed -qed - -instance pextreal :: t2_space -proof - fix x y :: pextreal assume "x \ y" - let "?P x (y::pextreal)" = "\ U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" - - { fix x y :: pextreal assume "x < y" - from pextreal_dense[OF this] obtain z where z: "x < z" "z < y" by auto - have "?P x y" - apply (rule exI[of _ "{.. y` - show "\U V. open U \ open V \ x \ U \ y \ V \ U \ V = {}" - proof (cases rule: linorder_cases) - assume "x = y" with `x \ y` show ?thesis by simp - next assume "x < y" from *[OF this] show ?thesis by auto - next assume "y < x" from *[OF this] show ?thesis by auto - qed -qed - -definition (in complete_lattice) isoton :: "(nat \ 'a) \ 'a \ bool" (infix "\" 50) where - "A \ X \ (\i. A i \ A (Suc i)) \ (SUP i. A i) = X" - -definition (in complete_lattice) antiton (infix "\" 50) where - "A \ X \ (\i. A i \ A (Suc i)) \ (INF i. A i) = X" - -lemma isotoneI[intro?]: "\ \i. f i \ f (Suc i) ; (SUP i. f i) = F \ \ f \ F" - unfolding isoton_def by auto - -lemma (in complete_lattice) isotonD[dest]: - assumes "A \ X" shows "A i \ A (Suc i)" "(SUP i. A i) = X" - using assms unfolding isoton_def by auto - -lemma isotonD'[dest]: - assumes "(A::_=>_) \ X" shows "A i x \ A (Suc i) x" "(SUP i. A i) = X" - using assms unfolding isoton_def le_fun_def by auto - -lemma isoton_mono_le: - assumes "f \ x" "i \ j" - shows "f i \ f j" - using `f \ x`[THEN isotonD(1)] lift_Suc_mono_le[of f, OF _ `i \ j`] by auto - -lemma isoton_const: - shows "(\ i. c) \ c" -unfolding isoton_def by auto - -lemma isoton_cmult_right: - assumes "f \ (x::pextreal)" - shows "(\i. c * f i) \ (c * x)" - using assms unfolding isoton_def pextreal_SUP_cmult - by (auto intro: pextreal_mult_cancel) - -lemma isoton_cmult_left: - "f \ (x::pextreal) \ (\i. f i * c) \ (x * c)" - by (subst (1 2) mult_commute) (rule isoton_cmult_right) - -lemma isoton_add: - assumes "f \ (x::pextreal)" and "g \ y" - shows "(\i. f i + g i) \ (x + y)" - using assms unfolding isoton_def - by (auto intro: pextreal_mult_cancel add_mono simp: SUPR_pextreal_add) - -lemma isoton_fun_expand: - "f \ x \ (\i. (\j. f j i) \ (x i))" -proof - - have "\i. {y. \f'\range f. y = f' i} = range (\j. f j i)" - by auto - with assms show ?thesis - by (auto simp add: isoton_def le_fun_def Sup_fun_def SUPR_def) -qed - -lemma isoton_indicator: - assumes "f \ g" - shows "(\i x. f i x * indicator A x) \ (\x. g x * indicator A x :: pextreal)" - using assms unfolding isoton_fun_expand by (auto intro!: isoton_cmult_left) - -lemma isoton_setsum: - fixes f :: "'a \ nat \ pextreal" - assumes "finite A" "A \ {}" - assumes "\ x. x \ A \ f x \ y x" - shows "(\ i. (\ x \ A. f x i)) \ (\ x \ A. y x)" -using assms -proof (induct A rule:finite_ne_induct) - case singleton thus ?case by auto -next - case (insert a A) note asms = this - hence *: "(\ i. \ x \ A. f x i) \ (\ x \ A. y x)" by auto - have **: "(\ i. f a i) \ y a" using asms by simp - have "(\ i. f a i + (\ x \ A. f x i)) \ (y a + (\ x \ A. y x))" - using * ** isoton_add by auto - thus "(\ i. \ x \ insert a A. f x i) \ (\ x \ insert a A. y x)" - using asms by fastsimp -qed - -lemma isoton_Sup: - assumes "f \ u" - shows "f i \ u" - using le_SUPI[of i UNIV f] assms - unfolding isoton_def by auto - -lemma isoton_mono: - assumes iso: "x \ a" "y \ b" and *: "\n. x n \ y (N n)" - shows "a \ b" -proof - - from iso have "a = (SUP n. x n)" "b = (SUP n. y n)" - unfolding isoton_def by auto - with * show ?thesis by (auto intro!: SUP_mono) -qed - -lemma pextreal_le_mult_one_interval: - fixes x y :: pextreal - assumes "\z. \ 0 < z ; z < 1 \ \ z * x \ y" - shows "x \ y" -proof (cases x, cases y) - assume "x = \" - with assms[of "1 / 2"] - show "x \ y" by simp -next - fix r p assume *: "y = Real p" "x = Real r" and **: "0 \ r" "0 \ p" - have "r \ p" - proof (rule field_le_mult_one_interval) - fix z :: real assume "0 < z" and "z < 1" - with assms[of "Real z"] - show "z * r \ p" using ** * by (auto simp: zero_le_mult_iff) - qed - thus "x \ y" using ** * by simp -qed simp - -lemma pextreal_greater_0[intro]: - fixes a :: pextreal - assumes "a \ 0" - shows "a > 0" -using assms apply (cases a) by auto - -lemma pextreal_mult_strict_right_mono: - assumes "a < b" and "0 < c" "c < \" - shows "a * c < b * c" - using assms - by (cases a, cases b, cases c) - (auto simp: zero_le_mult_iff pextreal_less_\) - -lemma minus_pextreal_eq2: - fixes x y z :: pextreal - assumes "y \ x" and "y \ \" shows "z = x - y \ z + y = x" - using assms - apply (subst eq_commute) - apply (subst minus_pextreal_eq) - by (cases x, cases z, auto simp add: ac_simps not_less) - -lemma pextreal_diff_eq_diff_imp_eq: - assumes "a \ \" "b \ a" "c \ a" - assumes "a - b = a - c" - shows "b = c" - using assms - by (cases a, cases b, cases c) (auto split: split_if_asm) - -lemma pextreal_inverse_eq_0: "inverse x = 0 \ x = \" - by (cases x) auto - -lemma pextreal_mult_inverse: - "\ x \ \ ; x \ 0 \ \ x * inverse x = 1" - by (cases x) auto - -lemma pextreal_zero_less_diff_iff: - fixes a b :: pextreal shows "0 < a - b \ b < a" - apply (cases a, cases b) - apply (auto simp: pextreal_noteq_omega_Ex pextreal_less_\) - apply (cases b) - by auto - -lemma pextreal_less_Real_Ex: - fixes a b :: pextreal shows "x < Real r \ (\p\0. p < r \ x = Real p)" - by (cases x) auto - -lemma open_Real: assumes "open S" shows "open (Real ` ({0..} \ S))" - unfolding open_pextreal_def apply(rule,rule,rule,rule assms) by auto - -lemma pextreal_zero_le_diff: - fixes a b :: pextreal shows "a - b = 0 \ a \ b" - by (cases a, cases b, simp_all, cases b, auto) - -lemma lim_Real[simp]: assumes "\n. f n \ 0" "m\0" - shows "(\n. Real (f n)) ----> Real m \ (\n. f n) ----> m" (is "?l = ?r") -proof assume ?l show ?r unfolding Lim_sequentially - proof safe fix e::real assume e:"e>0" - note open_ball[of m e] note open_Real[OF this] - note * = `?l`[unfolded tendsto_def,rule_format,OF this] - have "eventually (\x. Real (f x) \ Real ` ({0..} \ ball m e)) sequentially" - apply(rule *) unfolding image_iff using assms(2) e by auto - thus "\N. \n\N. dist (f n) m < e" unfolding eventually_sequentially - apply safe apply(rule_tac x=N in exI,safe) apply(erule_tac x=n in allE,safe) - proof- fix n x assume "Real (f n) = Real x" "0 \ x" - hence *:"f n = x" using assms(1) by auto - assume "x \ ball m e" thus "dist (f n) m < e" unfolding * - by (auto simp add:dist_commute) - qed qed -next assume ?r show ?l unfolding tendsto_def eventually_sequentially - proof safe fix S assume S:"open S" "Real m \ S" - guess T y using S(1) apply-apply(erule pextreal_openE) . note T=this - have "m\real ` (S - {\})" unfolding image_iff - apply(rule_tac x="Real m" in bexI) using assms(2) S(2) by auto - hence "m \ T" unfolding T(2)[THEN sym] by auto - from `?r`[unfolded tendsto_def eventually_sequentially,rule_format,OF T(1) this] - guess N .. note N=this[rule_format] - show "\N. \n\N. Real (f n) \ S" apply(rule_tac x=N in exI) - proof safe fix n assume n:"N\n" - have "f n \ real ` (S - {\})" using N[OF n] assms unfolding T(2)[THEN sym] - unfolding image_iff apply-apply(rule_tac x="Real (f n)" in bexI) - unfolding real_Real by auto - then guess x unfolding image_iff .. note x=this - show "Real (f n) \ S" unfolding x apply(subst Real_real) using x by auto - qed - qed -qed - -lemma pextreal_INFI: - fixes x :: pextreal - assumes "\i. i \ A \ x \ f i" - assumes "\y. (\i. i \ A \ y \ f i) \ y \ x" - shows "(INF i:A. f i) = x" - unfolding INFI_def Inf_pextreal_def - using assms by (auto intro!: Greatest_equality) - -lemma real_of_pextreal_less:"x < y \ y\\ \ real x < real y" -proof- case goal1 - have *:"y = Real (real y)" "x = Real (real x)" using goal1 Real_real by auto - show ?case using goal1 apply- apply(subst(asm) *(1))apply(subst(asm) *(2)) - unfolding pextreal_less by auto -qed - -lemma not_less_omega[simp]:"\ x < \ \ x = \" - by (metis antisym_conv3 pextreal_less(3)) - -lemma Real_real': assumes "x\\" shows "Real (real x) = x" -proof- have *:"(THE r. 0 \ r \ x = Real r) = real x" - apply(rule the_equality) using assms unfolding Real_real by auto - have "Real (THE r. 0 \ r \ x = Real r) = x" unfolding * - using assms unfolding Real_real by auto - thus ?thesis unfolding real_of_pextreal_def of_pextreal_def - unfolding pextreal_case_def using assms by auto -qed - -lemma Real_less_plus_one:"Real x < Real (max (x + 1) 1)" - unfolding pextreal_less by auto - -lemma Lim_omega: "f ----> \ \ (\B. \N. \n\N. f n \ Real B)" (is "?l = ?r") -proof assume ?r show ?l apply(rule topological_tendstoI) - unfolding eventually_sequentially - proof- fix S assume "open S" "\ \ S" - from open_omega[OF this] guess B .. note B=this - from `?r`[rule_format,of "(max B 0)+1"] guess N .. note N=this - show "\N. \n\N. f n \ S" apply(rule_tac x=N in exI) - proof safe case goal1 - have "Real B < Real ((max B 0) + 1)" by auto - also have "... \ f n" using goal1 N by auto - finally show ?case using B by fastsimp - qed - qed -next assume ?l show ?r - proof fix B::real have "open {Real B<..}" "\ \ {Real B<..}" by auto - from topological_tendstoD[OF `?l` this,unfolded eventually_sequentially] - guess N .. note N=this - show "\N. \n\N. Real B \ f n" apply(rule_tac x=N in exI) using N by auto - qed -qed - -lemma Lim_bounded_omgea: assumes lim:"f ----> l" and "\n. f n \ Real B" shows "l \ \" -proof(rule ccontr,unfold not_not) let ?B = "max (B + 1) 1" assume as:"l=\" - from lim[unfolded this Lim_omega,rule_format,of "?B"] - guess N .. note N=this[rule_format,OF le_refl] - hence "Real ?B \ Real B" using assms(2)[of N] by(rule order_trans) - hence "Real ?B < Real ?B" using Real_less_plus_one[of B] by(rule le_less_trans) - thus False by auto -qed - -lemma incseq_le_pextreal: assumes inc: "\n m. n\m \ X n \ X m" - and lim: "X ----> (L::pextreal)" shows "X n \ L" -proof(cases "L = \") - case False have "\n. X n \ \" - proof(rule ccontr,unfold not_all not_not,safe) - case goal1 hence "\n\x. X n = \" using inc[of x] by auto - hence "X ----> \" unfolding tendsto_def eventually_sequentially - apply safe apply(rule_tac x=x in exI) by auto - note Lim_unique[OF trivial_limit_sequentially this lim] - with False show False by auto - qed note * =this[rule_format] - - have **:"\m n. m \ n \ Real (real (X m)) \ Real (real (X n))" - unfolding Real_real using * inc by auto - have "real (X n) \ real L" apply-apply(rule incseq_le) defer - apply(subst lim_Real[THEN sym]) apply(rule,rule,rule) - unfolding Real_real'[OF *] Real_real'[OF False] - unfolding incseq_def using ** lim by auto - hence "Real (real (X n)) \ Real (real L)" by auto - thus ?thesis unfolding Real_real using * False by auto -qed auto - -lemma SUP_Lim_pextreal: assumes "\n m. n\m \ f n \ f m" "f ----> l" - shows "(SUP n. f n) = (l::pextreal)" unfolding SUPR_def Sup_pextreal_def -proof (safe intro!: Least_equality) - fix n::nat show "f n \ l" apply(rule incseq_le_pextreal) - using assms by auto -next fix y assume y:"\x\range f. x \ y" show "l \ y" - proof(rule ccontr,cases "y=\",unfold not_le) - case False assume as:"y < l" - have l:"l \ \" apply(rule Lim_bounded_omgea[OF assms(2), of "real y"]) - using False y unfolding Real_real by auto - - have yl:"real y < real l" using as apply- - apply(subst(asm) Real_real'[THEN sym,OF `y\\`]) - apply(subst(asm) Real_real'[THEN sym,OF `l\\`]) - unfolding pextreal_less apply(subst(asm) if_P) by auto - hence "y + (y - l) * Real (1 / 2) < l" apply- - apply(subst Real_real'[THEN sym,OF `y\\`]) apply(subst(2) Real_real'[THEN sym,OF `y\\`]) - apply(subst Real_real'[THEN sym,OF `l\\`]) apply(subst(2) Real_real'[THEN sym,OF `l\\`]) by auto - hence *:"l \ {y + (y - l) / 2<..}" by auto - have "open {y + (y-l)/2 <..}" by auto - note topological_tendstoD[OF assms(2) this *] - from this[unfolded eventually_sequentially] guess N .. note this[rule_format, of N] - hence "y + (y - l) * Real (1 / 2) < y" using y[rule_format,of "f N"] by auto - hence "Real (real y) + (Real (real y) - Real (real l)) * Real (1 / 2) < Real (real y)" - unfolding Real_real using `y\\` `l\\` by auto - thus False using yl by auto - qed auto -qed - -lemma Real_max':"Real x = Real (max x 0)" -proof(cases "x < 0") case True - hence *:"max x 0 = 0" by auto - show ?thesis unfolding * using True by auto -qed auto - -lemma lim_pextreal_increasing: assumes "\n m. n\m \ f n \ f m" - obtains l where "f ----> (l::pextreal)" -proof(cases "\B. \n. f n < Real B") - case False thus thesis apply- apply(rule that[of \]) unfolding Lim_omega not_ex not_all - apply safe apply(erule_tac x=B in allE,safe) apply(rule_tac x=x in exI,safe) - apply(rule order_trans[OF _ assms[rule_format]]) by auto -next case True then guess B .. note B = this[rule_format] - hence *:"\n. f n < \" apply-apply(rule less_le_trans,assumption) by auto - have *:"\n. f n \ \" proof- case goal1 show ?case using *[of n] by auto qed - have B':"\n. real (f n) \ max 0 B" proof- case goal1 thus ?case - using B[of n] apply-apply(subst(asm) Real_real'[THEN sym]) defer - apply(subst(asm)(2) Real_max') unfolding pextreal_less apply(subst(asm) if_P) using *[of n] by auto - qed - have "\l. (\n. real (f n)) ----> l" apply(rule Topology_Euclidean_Space.bounded_increasing_convergent) - proof safe show "bounded {real (f n) |n. True}" - unfolding bounded_def apply(rule_tac x=0 in exI,rule_tac x="max 0 B" in exI) - using B' unfolding dist_norm by auto - fix n::nat have "Real (real (f n)) \ Real (real (f (Suc n)))" - using assms[rule_format,of n "Suc n"] apply(subst Real_real)+ - using *[of n] *[of "Suc n"] by fastsimp - thus "real (f n) \ real (f (Suc n))" by auto - qed then guess l .. note l=this - have "0 \ l" apply(rule LIMSEQ_le_const[OF l]) - by(rule_tac x=0 in exI,auto) - - thus ?thesis apply-apply(rule that[of "Real l"]) - using l apply-apply(subst(asm) lim_Real[THEN sym]) prefer 3 - unfolding Real_real using * by auto -qed - -lemma setsum_neq_omega: assumes "finite s" "\x. x \ s \ f x \ \" - shows "setsum f s \ \" using assms -proof induct case (insert x s) - show ?case unfolding setsum.insert[OF insert(1-2)] - using insert by auto -qed auto - - -lemma real_Real': "0 \ x \ real (Real x) = x" - unfolding real_Real by auto - -lemma real_pextreal_pos[intro]: - assumes "x \ 0" "x \ \" - shows "real x > 0" - apply(subst real_Real'[THEN sym,of 0]) defer - apply(rule real_of_pextreal_less) using assms by auto - -lemma Lim_omega_gt: "f ----> \ \ (\B. \N. \n\N. f n > Real B)" (is "?l = ?r") -proof assume ?l thus ?r unfolding Lim_omega apply safe - apply(erule_tac x="max B 0 +1" in allE,safe) - apply(rule_tac x=N in exI,safe) apply(erule_tac x=n in allE,safe) - apply(rule_tac y="Real (max B 0 + 1)" in less_le_trans) by auto -next assume ?r thus ?l unfolding Lim_omega apply safe - apply(erule_tac x=B in allE,safe) apply(rule_tac x=N in exI,safe) by auto -qed - -lemma pextreal_minus_le_cancel: - fixes a b c :: pextreal - assumes "b \ a" - shows "c - a \ c - b" - using assms by (cases a, cases b, cases c, simp, simp, simp, cases b, cases c, simp_all) - -lemma pextreal_minus_\[simp]: "x - \ = 0" by (cases x) simp_all - -lemma pextreal_minus_mono[intro]: "a - x \ (a::pextreal)" -proof- have "a - x \ a - 0" - apply(rule pextreal_minus_le_cancel) by auto - thus ?thesis by auto -qed - -lemma pextreal_minus_eq_\[simp]: "x - y = \ \ (x = \ \ y \ \)" - by (cases x, cases y) (auto, cases y, auto) - -lemma pextreal_less_minus_iff: - fixes a b c :: pextreal - shows "a < b - c \ c + a < b" - by (cases c, cases a, cases b, auto) - -lemma pextreal_minus_less_iff: - fixes a b c :: pextreal shows "a - c < b \ (0 < b \ (c \ \ \ a < b + c))" - by (cases c, cases a, cases b, auto) - -lemma pextreal_le_minus_iff: - fixes a b c :: pextreal - shows "a \ c - b \ ((c \ b \ a = 0) \ (b < c \ a + b \ c))" - by (cases a, cases c, cases b, auto simp: pextreal_noteq_omega_Ex) - -lemma pextreal_minus_le_iff: - fixes a b c :: pextreal - shows "a - c \ b \ (c \ a \ a \ b + c)" - by (cases a, cases c, cases b, auto simp: pextreal_noteq_omega_Ex) - -lemmas pextreal_minus_order = pextreal_minus_le_iff pextreal_minus_less_iff pextreal_le_minus_iff pextreal_less_minus_iff - -lemma pextreal_minus_strict_mono: - assumes "a > 0" "x > 0" "a\\" - shows "a - x < (a::pextreal)" - using assms by(cases x, cases a, auto) - -lemma pextreal_minus': - "Real r - Real p = (if 0 \ r \ p \ r then if 0 \ p then Real (r - p) else Real r else 0)" - by (auto simp: minus_pextreal_eq not_less) - -lemma pextreal_minus_plus: - "x \ (a::pextreal) \ a - x + x = a" - by (cases a, cases x) auto - -lemma pextreal_cancel_plus_minus: "b \ \ \ a + b - b = a" - by (cases a, cases b) auto - -lemma pextreal_minus_le_cancel_right: - fixes a b c :: pextreal - assumes "a \ b" "c \ a" - shows "a - c \ b - c" - using assms by (cases a, cases b, cases c, auto, cases c, auto) - -lemma real_of_pextreal_setsum': - assumes "\x \ S. f x \ \" - shows "(\x\S. real (f x)) = real (setsum f S)" -proof cases - assume "finite S" - from this assms show ?thesis - by induct (simp_all add: real_of_pextreal_add setsum_\) -qed simp - -lemma Lim_omega_pos: "f ----> \ \ (\B>0. \N. \n\N. f n \ Real B)" (is "?l = ?r") - unfolding Lim_omega apply safe defer - apply(erule_tac x="max 1 B" in allE) apply safe defer - apply(rule_tac x=N in exI,safe) apply(erule_tac x=n in allE,safe) - apply(rule_tac y="Real (max 1 B)" in order_trans) by auto - -lemma pextreal_LimI_finite: - assumes "x \ \" "\r. 0 < r \ \N. \n\N. u n < x + r \ x < u n + r" - shows "u ----> x" -proof (rule topological_tendstoI, unfold eventually_sequentially) - fix S assume "open S" "x \ S" - then obtain A where "open A" and A_eq: "Real ` (A \ {0..}) = S - {\}" by (auto elim!: pextreal_openE) - then have "x \ Real ` (A \ {0..})" using `x \ S` `x \ \` by auto - then have "real x \ A" by auto - then obtain r where "0 < r" and dist: "\y. dist y (real x) < r \ y \ A" - using `open A` unfolding open_real_def by auto - then obtain n where - upper: "\N. n \ N \ u N < x + Real r" and - lower: "\N. n \ N \ x < u N + Real r" using assms(2)[of "Real r"] by auto - show "\N. \n\N. u n \ S" - proof (safe intro!: exI[of _ n]) - fix N assume "n \ N" - from upper[OF this] `x \ \` `0 < r` - have "u N \ \" by (force simp: pextreal_noteq_omega_Ex) - with `x \ \` `0 < r` lower[OF `n \ N`] upper[OF `n \ N`] - have "dist (real (u N)) (real x) < r" "u N \ \" - by (auto simp: pextreal_noteq_omega_Ex dist_real_def abs_diff_less_iff field_simps) - from dist[OF this(1)] - have "u N \ Real ` (A \ {0..})" using `u N \ \` - by (auto intro!: image_eqI[of _ _ "real (u N)"] simp: pextreal_noteq_omega_Ex Real_real) - thus "u N \ S" using A_eq by simp - qed -qed - -lemma real_Real_max:"real (Real x) = max x 0" - unfolding real_Real by auto - -lemma Sup_lim: - assumes "\n. b n \ s" "b ----> (a::pextreal)" - shows "a \ Sup s" -proof(rule ccontr,unfold not_le) - assume as:"Sup s < a" hence om:"Sup s \ \" by auto - have s:"s \ {}" using assms by auto - { presume *:"\n. b n < a \ False" - show False apply(cases,rule *,assumption,unfold not_all not_less) - proof- case goal1 then guess n .. note n=this - thus False using complete_lattice_class.Sup_upper[OF assms(1)[rule_format,of n]] - using as by auto - qed - } assume b:"\n. b n < a" - show False - proof(cases "a = \") - case False have *:"a - Sup s > 0" - using False as by(auto simp: pextreal_zero_le_diff) - have "(a - Sup s) / 2 \ a / 2" unfolding divide_pextreal_def - apply(rule mult_right_mono) by auto - also have "... = Real (real (a / 2))" apply(rule Real_real'[THEN sym]) - using False by auto - also have "... < Real (real a)" unfolding pextreal_less using as False - by(auto simp add: real_of_pextreal_mult[THEN sym]) - also have "... = a" apply(rule Real_real') using False by auto - finally have asup:"a > (a - Sup s) / 2" . - have "\n. a - b n < (a - Sup s) / 2" - proof(rule ccontr,unfold not_ex not_less) - case goal1 - have "(a - Sup s) * Real (1 / 2) > 0" - using * by auto - hence "a - (a - Sup s) * Real (1 / 2) < a" - apply-apply(rule pextreal_minus_strict_mono) - using False * by auto - hence *:"a \ {a - (a - Sup s) / 2<..}"using asup by auto - note topological_tendstoD[OF assms(2) open_pextreal_greaterThan,OF *] - from this[unfolded eventually_sequentially] guess n .. - note n = this[rule_format,of n] - have "b n + (a - Sup s) / 2 \ a" - using add_right_mono[OF goal1[rule_format,of n],of "b n"] - unfolding pextreal_minus_plus[OF less_imp_le[OF b[rule_format]]] - by(auto simp: add_commute) - hence "b n \ a - (a - Sup s) / 2" unfolding pextreal_le_minus_iff - using asup by auto - hence "b n \ {a - (a - Sup s) / 2<..}" by auto - thus False using n by auto - qed - then guess n .. note n = this - have "Sup s < a - (a - Sup s) / 2" - using False as om by (cases a) (auto simp: pextreal_noteq_omega_Ex field_simps) - also have "... \ b n" - proof- note add_right_mono[OF less_imp_le[OF n],of "b n"] - note this[unfolded pextreal_minus_plus[OF less_imp_le[OF b[rule_format]]]] - hence "a - (a - Sup s) / 2 \ (a - Sup s) / 2 + b n - (a - Sup s) / 2" - apply(rule pextreal_minus_le_cancel_right) using asup by auto - also have "... = b n + (a - Sup s) / 2 - (a - Sup s) / 2" - by(auto simp add: add_commute) - also have "... = b n" apply(subst pextreal_cancel_plus_minus) - proof(rule ccontr,unfold not_not) case goal1 - show ?case using asup unfolding goal1 by auto - qed auto - finally show ?thesis . - qed - finally show False - using complete_lattice_class.Sup_upper[OF assms(1)[rule_format,of n]] by auto - next case True - from assms(2)[unfolded True Lim_omega_gt,rule_format,of "real (Sup s)"] - guess N .. note N = this[rule_format,of N] - thus False using complete_lattice_class.Sup_upper[OF assms(1)[rule_format,of N]] - unfolding Real_real using om by auto - qed qed - -lemma Sup_mono_lim: - assumes "\a\A. \b. \n. b n \ B \ b ----> (a::pextreal)" - shows "Sup A \ Sup B" - unfolding Sup_le_iff apply(rule) apply(drule assms[rule_format]) apply safe - apply(rule_tac b=b in Sup_lim) by auto - -lemma pextreal_less_add: - assumes "x \ \" "a < b" - shows "x + a < x + b" - using assms by (cases a, cases b, cases x) auto - -lemma SUPR_lim: - assumes "\n. b n \ B" "(\n. f (b n)) ----> (f a::pextreal)" - shows "f a \ SUPR B f" - unfolding SUPR_def apply(rule Sup_lim[of "\n. f (b n)"]) - using assms by auto - -lemma SUP_\_imp: - assumes "(SUP i. f i) = \" - shows "\i. Real x < f i" -proof (rule ccontr) - assume "\ ?thesis" hence "\i. f i \ Real x" by (simp add: not_less) - hence "(SUP i. f i) \ Real x" unfolding SUP_le_iff by auto - with assms show False by auto -qed - -lemma SUPR_mono_lim: - assumes "\a\A. \b. \n. b n \ B \ (\n. f (b n)) ----> (f a::pextreal)" - shows "SUPR A f \ SUPR B f" - unfolding SUPR_def apply(rule Sup_mono_lim) - apply safe apply(drule assms[rule_format],safe) - apply(rule_tac x="\n. f (b n)" in exI) by auto - -lemma real_0_imp_eq_0: - assumes "x \ \" "real x = 0" - shows "x = 0" -using assms by (cases x) auto - -lemma SUPR_mono: - assumes "\a\A. \b\B. f b \ f a" - shows "SUPR A f \ SUPR B f" - unfolding SUPR_def apply(rule Sup_mono) - using assms by auto - -lemma less_add_Real: - fixes x :: real - fixes a b :: pextreal - assumes "x \ 0" "a < b" - shows "a + Real x < b + Real x" -using assms by (cases a, cases b) auto - -lemma le_add_Real: - fixes x :: real - fixes a b :: pextreal - assumes "x \ 0" "a \ b" - shows "a + Real x \ b + Real x" -using assms by (cases a, cases b) auto - -lemma le_imp_less_pextreal: - fixes x :: pextreal - assumes "x > 0" "a + x \ b" "a \ \" - shows "a < b" -using assms by (cases x, cases a, cases b) auto - -lemma pextreal_INF_minus: - fixes f :: "nat \ pextreal" - assumes "c \ \" - shows "(INF i. c - f i) = c - (SUP i. f i)" -proof (cases "SUP i. f i") - case infinite - from `c \ \` obtain x where [simp]: "c = Real x" by (cases c) auto - from SUP_\_imp[OF infinite] obtain i where "Real x < f i" by auto - have "(INF i. c - f i) \ c - f i" - by (auto intro!: complete_lattice_class.INF_leI) - also have "\ = 0" using `Real x < f i` by (auto simp: minus_pextreal_eq) - finally show ?thesis using infinite by auto -next - case (preal r) - from `c \ \` obtain x where c: "c = Real x" by (cases c) auto - - show ?thesis unfolding c - proof (rule pextreal_INFI) - fix i have "f i \ (SUP i. f i)" by (rule le_SUPI) simp - thus "Real x - (SUP i. f i) \ Real x - f i" by (rule pextreal_minus_le_cancel) - next - fix y assume *: "\i. i \ UNIV \ y \ Real x - f i" - from this[of 0] obtain p where p: "y = Real p" "0 \ p" - by (cases "f 0", cases y, auto split: split_if_asm) - hence "\i. Real p \ Real x - f i" using * by auto - hence *: "\i. Real x \ f i \ Real p = 0" - "\i. f i < Real x \ Real p + f i \ Real x" - unfolding pextreal_le_minus_iff by auto - show "y \ Real x - (SUP i. f i)" unfolding p pextreal_le_minus_iff - proof safe - assume x_less: "Real x \ (SUP i. f i)" - show "Real p = 0" - proof (rule ccontr) - assume "Real p \ 0" - hence "0 < Real p" by auto - from Sup_close[OF this, of "range f"] - obtain i where e: "(SUP i. f i) < f i + Real p" - using preal unfolding SUPR_def by auto - hence "Real x \ f i + Real p" using x_less by auto - show False - proof cases - assume "\i. f i < Real x" - hence "Real p + f i \ Real x" using * by auto - hence "f i + Real p \ (SUP i. f i)" using x_less by (auto simp: field_simps) - thus False using e by auto - next - assume "\ (\i. f i < Real x)" - then obtain i where "Real x \ f i" by (auto simp: not_less) - from *(1)[OF this] show False using `Real p \ 0` by auto - qed - qed - next - have "\i. f i \ (SUP i. f i)" by (rule complete_lattice_class.le_SUPI) auto - also assume "(SUP i. f i) < Real x" - finally have "\i. f i < Real x" by auto - hence *: "\i. Real p + f i \ Real x" using * by auto - have "Real p \ Real x" using *[of 0] by (cases "f 0") (auto split: split_if_asm) - - have SUP_eq: "(SUP i. f i) \ Real x - Real p" - proof (rule SUP_leI) - fix i show "f i \ Real x - Real p" unfolding pextreal_le_minus_iff - proof safe - assume "Real x \ Real p" - with *[of i] show "f i = 0" - by (cases "f i") (auto split: split_if_asm) - next - assume "Real p < Real x" - show "f i + Real p \ Real x" using * by (auto simp: field_simps) - qed - qed - - show "Real p + (SUP i. f i) \ Real x" - proof cases - assume "Real x \ Real p" - with `Real p \ Real x` have [simp]: "Real p = Real x" by (rule antisym) - { fix i have "f i = 0" using *[of i] by (cases "f i") (auto split: split_if_asm) } - hence "(SUP i. f i) \ 0" by (auto intro!: SUP_leI) - thus ?thesis by simp - next - assume "\ Real x \ Real p" hence "Real p < Real x" unfolding not_le . - with SUP_eq show ?thesis unfolding pextreal_le_minus_iff by (auto simp: field_simps) - qed - qed - qed -qed - -lemma pextreal_SUP_minus: - fixes f :: "nat \ pextreal" - shows "(SUP i. c - f i) = c - (INF i. f i)" -proof (rule pextreal_SUPI) - fix i have "(INF i. f i) \ f i" by (rule INF_leI) simp - thus "c - f i \ c - (INF i. f i)" by (rule pextreal_minus_le_cancel) -next - fix y assume *: "\i. i \ UNIV \ c - f i \ y" - show "c - (INF i. f i) \ y" - proof (cases y) - case (preal p) - - show ?thesis unfolding pextreal_minus_le_iff preal - proof safe - assume INF_le_x: "(INF i. f i) \ c" - from * have *: "\i. f i \ c \ c \ Real p + f i" - unfolding pextreal_minus_le_iff preal by auto - - have INF_eq: "c - Real p \ (INF i. f i)" - proof (rule le_INFI) - fix i show "c - Real p \ f i" unfolding pextreal_minus_le_iff - proof safe - assume "Real p \ c" - show "c \ f i + Real p" - proof cases - assume "f i \ c" from *[OF this] - show ?thesis by (simp add: field_simps) - next - assume "\ f i \ c" - hence "c \ f i" by auto - also have "\ \ f i + Real p" by auto - finally show ?thesis . - qed - qed - qed - - show "c \ Real p + (INF i. f i)" - proof cases - assume "Real p \ c" - with INF_eq show ?thesis unfolding pextreal_minus_le_iff by (auto simp: field_simps) - next - assume "\ Real p \ c" - hence "c \ Real p" by auto - also have "Real p \ Real p + (INF i. f i)" by auto - finally show ?thesis . - qed - qed - qed simp -qed - -lemma pextreal_le_minus_imp_0: - fixes a b :: pextreal - shows "a \ a - b \ a \ 0 \ a \ \ \ b = 0" - by (cases a, cases b, auto split: split_if_asm) - -lemma lim_INF_eq_lim_SUP: - fixes X :: "nat \ real" - assumes "\i. 0 \ X i" and "0 \ x" - and lim_INF: "(SUP n. INF m. Real (X (n + m))) = Real x" (is "(SUP n. ?INF n) = _") - and lim_SUP: "(INF n. SUP m. Real (X (n + m))) = Real x" (is "(INF n. ?SUP n) = _") - shows "X ----> x" -proof (rule LIMSEQ_I) - fix r :: real assume "0 < r" - hence "0 \ r" by auto - from Sup_close[of "Real r" "range ?INF"] - obtain n where inf: "Real x < ?INF n + Real r" - unfolding SUPR_def lim_INF[unfolded SUPR_def] using `0 < r` by auto - - from Inf_close[of "range ?SUP" "Real r"] - obtain n' where sup: "?SUP n' < Real x + Real r" - unfolding INFI_def lim_SUP[unfolded INFI_def] using `0 < r` by auto - - show "\N. \n\N. norm (X n - x) < r" - proof (safe intro!: exI[of _ "max n n'"]) - fix m assume "max n n' \ m" hence "n \ m" "n' \ m" by auto - - note inf - also have "?INF n + Real r \ Real (X (n + (m - n))) + Real r" - by (rule le_add_Real, auto simp: `0 \ r` intro: INF_leI) - finally have up: "x < X m + r" - using `0 \ X m` `0 \ x` `0 \ r` `n \ m` by auto - - have "Real (X (n' + (m - n'))) \ ?SUP n'" - by (auto simp: `0 \ r` intro: le_SUPI) - also note sup - finally have down: "X m < x + r" - using `0 \ X m` `0 \ x` `0 \ r` `n' \ m` by auto - - show "norm (X m - x) < r" using up down by auto - qed -qed - -lemma Sup_countable_SUPR: - assumes "Sup A \ \" "A \ {}" - shows "\ f::nat \ pextreal. range f \ A \ Sup A = SUPR UNIV f" -proof - - have "\n. 0 < 1 / (of_nat n :: pextreal)" by auto - from Sup_close[OF this assms] - have "\n. \x. x \ A \ Sup A < x + 1 / of_nat n" by blast - from choice[OF this] obtain f where "range f \ A" and - epsilon: "\n. Sup A < f n + 1 / of_nat n" by blast - have "SUPR UNIV f = Sup A" - proof (rule pextreal_SUPI) - fix i show "f i \ Sup A" using `range f \ A` - by (auto intro!: complete_lattice_class.Sup_upper) - next - fix y assume bound: "\i. i \ UNIV \ f i \ y" - show "Sup A \ y" - proof (rule pextreal_le_epsilon) - fix e :: pextreal assume "0 < e" - show "Sup A \ y + e" - proof (cases e) - case (preal r) - hence "0 < r" using `0 < e` by auto - then obtain n where *: "inverse (of_nat n) < r" "0 < n" - using ex_inverse_of_nat_less by auto - have "Sup A \ f n + 1 / of_nat n" using epsilon[of n] by auto - also have "1 / of_nat n \ e" using preal * by (auto simp: real_eq_of_nat) - with bound have "f n + 1 / of_nat n \ y + e" by (rule add_mono) simp - finally show "Sup A \ y + e" . - qed simp - qed - qed - with `range f \ A` show ?thesis by (auto intro!: exI[of _ f]) -qed - -lemma SUPR_countable_SUPR: - assumes "SUPR A g \ \" "A \ {}" - shows "\ f::nat \ pextreal. range f \ g`A \ SUPR A g = SUPR UNIV f" -proof - - have "Sup (g`A) \ \" "g`A \ {}" using assms unfolding SUPR_def by auto - from Sup_countable_SUPR[OF this] - show ?thesis unfolding SUPR_def . -qed - -lemma pextreal_setsum_subtractf: - assumes "\i. i\A \ g i \ f i" and "\i. i\A \ f i \ \" - shows "(\i\A. f i - g i) = (\i\A. f i) - (\i\A. g i)" -proof cases - assume "finite A" from this assms show ?thesis - proof induct - case (insert x A) - hence hyp: "(\i\A. f i - g i) = (\i\A. f i) - (\i\A. g i)" - by auto - { fix i assume *: "i \ insert x A" - hence "g i \ f i" using insert by simp - also have "f i < \" using * insert by (simp add: pextreal_less_\) - finally have "g i \ \" by (simp add: pextreal_less_\) } - hence "setsum g A \ \" "g x \ \" by (auto simp: setsum_\) - moreover have "setsum f A \ \" "f x \ \" using insert by (auto simp: setsum_\) - moreover have "g x \ f x" using insert by auto - moreover have "(\i\A. g i) \ (\i\A. f i)" using insert by (auto intro!: setsum_mono) - ultimately show ?case using `finite A` `x \ A` hyp - by (auto simp: pextreal_noteq_omega_Ex) - qed simp -qed simp - -lemma real_of_pextreal_diff: - "y \ x \ x \ \ \ real x - real y = real (x - y)" - by (cases x, cases y) auto - -lemma psuminf_minus: - assumes ord: "\i. g i \ f i" and fin: "psuminf g \ \" "psuminf f \ \" - shows "(\\<^isub>\ i. f i - g i) = psuminf f - psuminf g" -proof - - have [simp]: "\i. f i \ \" using fin by (auto intro: psuminf_\) - from fin have "(\x. real (f x)) sums real (\\<^isub>\x. f x)" - and "(\x. real (g x)) sums real (\\<^isub>\x. g x)" - by (auto intro: psuminf_imp_suminf) - from sums_diff[OF this] - have "(\n. real (f n - g n)) sums (real ((\\<^isub>\x. f x) - (\\<^isub>\x. g x)))" using fin ord - by (subst (asm) (1 2) real_of_pextreal_diff) (auto simp: psuminf_\ psuminf_le) - hence "(\\<^isub>\ i. Real (real (f i - g i))) = Real (real ((\\<^isub>\x. f x) - (\\<^isub>\x. g x)))" - by (rule suminf_imp_psuminf) simp - thus ?thesis using fin by (simp add: Real_real psuminf_\) -qed - -lemma INF_eq_LIMSEQ: - assumes "mono (\i. - f i)" and "\n. 0 \ f n" and "0 \ x" - shows "(INF n. Real (f n)) = Real x \ f ----> x" -proof - assume x: "(INF n. Real (f n)) = Real x" - { fix n - have "Real x \ Real (f n)" using x[symmetric] by (auto intro: INF_leI) - hence "x \ f n" using assms by simp - hence "\f n - x\ = f n - x" by auto } - note abs_eq = this - show "f ----> x" - proof (rule LIMSEQ_I) - fix r :: real assume "0 < r" - show "\no. \n\no. norm (f n - x) < r" - proof (rule ccontr) - assume *: "\ ?thesis" - { fix N - from * obtain n where *: "N \ n" "r \ f n - x" - using abs_eq by (auto simp: not_less) - hence "x + r \ f n" by auto - also have "f n \ f N" using `mono (\i. - f i)` * by (auto dest: monoD) - finally have "Real (x + r) \ Real (f N)" using `0 \ f N` by auto } - hence "Real x < Real (x + r)" - and "Real (x + r) \ (INF n. Real (f n))" using `0 < r` `0 \ x` by (auto intro: le_INFI) - hence "Real x < (INF n. Real (f n))" by (rule less_le_trans) - thus False using x by auto - qed - qed -next - assume "f ----> x" - show "(INF n. Real (f n)) = Real x" - proof (rule pextreal_INFI) - fix n - from decseq_le[OF _ `f ----> x`] assms - show "Real x \ Real (f n)" unfolding decseq_eq_incseq incseq_mono by auto - next - fix y assume *: "\n. n\UNIV \ y \ Real (f n)" - thus "y \ Real x" - proof (cases y) - case (preal r) - with * have "\N. \n\N. r \ f n" using assms by fastsimp - from LIMSEQ_le_const[OF `f ----> x` this] - show "y \ Real x" using `0 \ x` preal by auto - qed simp - qed -qed - -lemma INFI_bound: - assumes "\N. x \ f N" - shows "x \ (INF n. f n)" - using assms by (simp add: INFI_def le_Inf_iff) - -lemma LIMSEQ_imp_lim_INF: - assumes pos: "\i. 0 \ X i" and lim: "X ----> x" - shows "(SUP n. INF m. Real (X (n + m))) = Real x" -proof - - have "0 \ x" using assms by (auto intro!: LIMSEQ_le_const) - - have "\n. (INF m. Real (X (n + m))) \ Real (X (n + 0))" by (rule INF_leI) simp - also have "\n. Real (X (n + 0)) < \" by simp - finally have "\n. \r\0. (INF m. Real (X (n + m))) = Real r" - by (auto simp: pextreal_less_\ pextreal_noteq_omega_Ex) - from choice[OF this] obtain r where r: "\n. (INF m. Real (X (n + m))) = Real (r n)" "\n. 0 \ r n" - by auto - - show ?thesis unfolding r - proof (subst SUP_eq_LIMSEQ) - show "mono r" unfolding mono_def - proof safe - fix x y :: nat assume "x \ y" - have "Real (r x) \ Real (r y)" unfolding r(1)[symmetric] using pos - proof (safe intro!: INF_mono bexI) - fix m have "x + (m + y - x) = y + m" - using `x \ y` by auto - thus "Real (X (x + (m + y - x))) \ Real (X (y + m))" by simp - qed simp - thus "r x \ r y" using r by auto - qed - show "\n. 0 \ r n" by fact - show "0 \ x" by fact - show "r ----> x" - proof (rule LIMSEQ_I) - fix e :: real assume "0 < e" - hence "0 < e/2" by auto - from LIMSEQ_D[OF lim this] obtain N where *: "\n. N \ n \ \X n - x\ < e/2" - by auto - show "\N. \n\N. norm (r n - x) < e" - proof (safe intro!: exI[of _ N]) - fix n assume "N \ n" - show "norm (r n - x) < e" - proof cases - assume "r n < x" - have "x - r n \ e/2" - proof cases - assume e: "e/2 \ x" - have "Real (x - e/2) \ Real (r n)" unfolding r(1)[symmetric] - proof (rule le_INFI) - fix m show "Real (x - e / 2) \ Real (X (n + m))" - using *[of "n + m"] `N \ n` - using pos by (auto simp: field_simps abs_real_def split: split_if_asm) - qed - with e show ?thesis using pos `0 \ x` r(2) by auto - next - assume "\ e/2 \ x" hence "x - e/2 < 0" by auto - with `0 \ r n` show ?thesis by auto - qed - with `r n < x` show ?thesis by simp - next - assume e: "\ r n < x" - have "Real (r n) \ Real (X (n + 0))" unfolding r(1)[symmetric] - by (rule INF_leI) simp - hence "r n - x \ X n - x" using r pos by auto - also have "\ < e/2" using *[OF `N \ n`] by (auto simp: field_simps abs_real_def split: split_if_asm) - finally have "r n - x < e" using `0 < e` by auto - with e show ?thesis by auto - qed - qed - qed - qed -qed - -lemma real_of_pextreal_strict_mono_iff: - "real a < real b \ (b \ \ \ ((a = \ \ 0 < b) \ (a < b)))" -proof (cases a) - case infinite thus ?thesis by (cases b) auto -next - case preal thus ?thesis by (cases b) auto -qed - -lemma real_of_pextreal_mono_iff: - "real a \ real b \ (a = \ \ (b \ \ \ a \ b) \ (b = \ \ a = 0))" -proof (cases a) - case infinite thus ?thesis by (cases b) auto -next - case preal thus ?thesis by (cases b) auto -qed - -lemma ex_pextreal_inverse_of_nat_Suc_less: - fixes e :: pextreal assumes "0 < e" shows "\n. inverse (of_nat (Suc n)) < e" -proof (cases e) - case (preal r) - with `0 < e` ex_inverse_of_nat_Suc_less[of r] - obtain n where "inverse (of_nat (Suc n)) < r" by auto - with preal show ?thesis - by (auto simp: real_eq_of_nat[symmetric]) -qed auto - -lemma Lim_eq_Sup_mono: - fixes u :: "nat \ pextreal" assumes "mono u" - shows "u ----> (SUP i. u i)" -proof - - from lim_pextreal_increasing[of u] `mono u` - obtain l where l: "u ----> l" unfolding mono_def by auto - from SUP_Lim_pextreal[OF _ this] `mono u` - have "(SUP i. u i) = l" unfolding mono_def by auto - with l show ?thesis by simp -qed - -lemma isotone_Lim: - fixes x :: pextreal assumes "u \ x" - shows "u ----> x" (is ?lim) and "mono u" (is ?mono) -proof - - show ?mono using assms unfolding mono_iff_le_Suc isoton_def by auto - from Lim_eq_Sup_mono[OF this] `u \ x` - show ?lim unfolding isoton_def by simp -qed - -lemma isoton_iff_Lim_mono: - fixes u :: "nat \ pextreal" - shows "u \ x \ (mono u \ u ----> x)" -proof safe - assume "mono u" and x: "u ----> x" - with SUP_Lim_pextreal[OF _ x] - show "u \ x" unfolding isoton_def - using `mono u`[unfolded mono_def] - using `mono u`[unfolded mono_iff_le_Suc] - by auto -qed (auto dest: isotone_Lim) - -lemma pextreal_inverse_inverse[simp]: - fixes x :: pextreal - shows "inverse (inverse x) = x" - by (cases x) auto - -lemma atLeastAtMost_omega_eq_atLeast: - shows "{a .. \} = {a ..}" -by auto - -lemma atLeast0AtMost_eq_atMost: "{0 :: pextreal .. a} = {.. a}" by auto - -lemma greaterThan_omega_Empty: "{\ <..} = {}" by auto - -lemma lessThan_0_Empty: "{..< 0 :: pextreal} = {}" by auto - -lemma real_of_pextreal_inverse[simp]: - fixes X :: pextreal - shows "real (inverse X) = 1 / real X" - by (cases X) (auto simp: inverse_eq_divide) - -lemma real_of_pextreal_le_0[simp]: "real (X :: pextreal) \ 0 \ (X = 0 \ X = \)" - by (cases X) auto - -lemma real_of_pextreal_less_0[simp]: "\ (real (X :: pextreal) < 0)" - by (cases X) auto - -lemma abs_real_of_pextreal[simp]: "\real (X :: pextreal)\ = real X" - by simp - -lemma zero_less_real_of_pextreal: "0 < real (X :: pextreal) \ X \ 0 \ X \ \" - by (cases X) auto - -end diff -r 28b51effc5ed -r cdf7693bbe08 src/HOL/Probability/Probability.thy --- a/src/HOL/Probability/Probability.thy Mon Mar 14 14:37:47 2011 +0100 +++ b/src/HOL/Probability/Probability.thy Mon Mar 14 14:37:49 2011 +0100 @@ -1,6 +1,7 @@ theory Probability imports Complete_Measure + Lebesgue_Measure Information "ex/Dining_Cryptographers" "ex/Koepf_Duermuth_Countermeasure" diff -r 28b51effc5ed -r cdf7693bbe08 src/HOL/Probability/Probability_Space.thy --- a/src/HOL/Probability/Probability_Space.thy Mon Mar 14 14:37:47 2011 +0100 +++ b/src/HOL/Probability/Probability_Space.thy Mon Mar 14 14:37:49 2011 +0100 @@ -2,33 +2,33 @@ imports Lebesgue_Integration Radon_Nikodym Product_Measure begin -lemma real_of_pextreal_inverse[simp]: - fixes X :: pextreal +lemma real_of_extreal_inverse[simp]: + fixes X :: extreal shows "real (inverse X) = 1 / real X" by (cases X) (auto simp: inverse_eq_divide) -lemma real_of_pextreal_le_0[simp]: "real (X :: pextreal) \ 0 \ (X = 0 \ X = \)" +lemma real_of_extreal_le_0[simp]: "real (X :: extreal) \ 0 \ (X \ 0 \ X = \)" + by (cases X) auto + +lemma abs_real_of_extreal[simp]: "\real (X :: extreal)\ = real \X\" by (cases X) auto -lemma real_of_pextreal_less_0[simp]: "\ (real (X :: pextreal) < 0)" +lemma zero_less_real_of_extreal: "0 < real X \ (0 < X \ X \ \)" by (cases X) auto +lemma real_of_extreal_le_1: fixes X :: extreal shows "X \ 1 \ real X \ 1" + by (cases X) (auto simp: one_extreal_def) + locale prob_space = measure_space + assumes measure_space_1: "measure M (space M) = 1" -lemma abs_real_of_pextreal[simp]: "\real (X :: pextreal)\ = real X" - by simp - -lemma zero_less_real_of_pextreal: "0 < real (X :: pextreal) \ X \ 0 \ X \ \" - by (cases X) auto - sublocale prob_space < finite_measure proof - from measure_space_1 show "\ (space M) \ \" by simp + from measure_space_1 show "\ (space M) \ \" by simp qed abbreviation (in prob_space) "events \ sets M" -abbreviation (in prob_space) "prob \ \A. real (\ A)" +abbreviation (in prob_space) "prob \ \'" abbreviation (in prob_space) "prob_preserving \ measure_preserving" abbreviation (in prob_space) "random_variable M' X \ sigma_algebra M' \ X \ measurable M M'" abbreviation (in prob_space) "expectation \ integral\<^isup>L M" @@ -40,53 +40,57 @@ "indep_families F G \ (\ A \ F. \ B \ G. indep A B)" definition (in prob_space) - "distribution X = (\s. \ ((X -` s) \ (space M)))" + "distribution X A = \' (X -` A \ space M)" abbreviation (in prob_space) "joint_distribution X Y \ distribution (\x. (X x, Y x))" +declare (in finite_measure) positive_measure'[intro, simp] + lemma (in prob_space) distribution_cong: assumes "\x. x \ space M \ X x = Y x" shows "distribution X = distribution Y" unfolding distribution_def fun_eq_iff - using assms by (auto intro!: arg_cong[where f="\"]) + using assms by (auto intro!: arg_cong[where f="\'"]) lemma (in prob_space) joint_distribution_cong: assumes "\x. x \ space M \ X x = X' x" assumes "\x. x \ space M \ Y x = Y' x" shows "joint_distribution X Y = joint_distribution X' Y'" unfolding distribution_def fun_eq_iff - using assms by (auto intro!: arg_cong[where f="\"]) + using assms by (auto intro!: arg_cong[where f="\'"]) lemma (in prob_space) distribution_id[simp]: - assumes "N \ sets M" shows "distribution (\x. x) N = \ N" - using assms by (auto simp: distribution_def intro!: arg_cong[where f=\]) + "N \ events \ distribution (\x. x) N = prob N" + by (auto simp: distribution_def intro!: arg_cong[where f=prob]) lemma (in prob_space) prob_space: "prob (space M) = 1" - unfolding measure_space_1 by simp + using measure_space_1 unfolding \'_def by (simp add: one_extreal_def) + +lemma (in prob_space) prob_le_1[simp, intro]: "prob A \ 1" + using bounded_measure[of A] by (simp add: prob_space) + +lemma (in prob_space) distribution_positive[simp, intro]: + "0 \ distribution X A" unfolding distribution_def by auto -lemma (in prob_space) measure_le_1[simp, intro]: - assumes "A \ events" shows "\ A \ 1" -proof - - have "\ A \ \ (space M)" - using assms sets_into_space by(auto intro!: measure_mono) - also note measure_space_1 - finally show ?thesis . -qed +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) distribution_1: + "distribution X A \ 1" + unfolding distribution_def by simp lemma (in prob_space) prob_compl: - assumes "A \ events" + assumes A: "A \ events" shows "prob (space M - A) = 1 - prob A" - using `A \ events`[THEN sets_into_space] `A \ events` measure_space_1 - by (subst real_finite_measure_Diff) auto + using finite_measure_compl[OF A] by (simp add: prob_space) -lemma (in prob_space) indep_space: - assumes "s \ events" - shows "indep (space M) s" - using assms prob_space by (simp add: indep_def) +lemma (in prob_space) indep_space: "s \ events \ indep (space M) s" + by (simp add: indep_def prob_space) lemma (in prob_space) prob_space_increasing: "increasing M prob" - by (auto intro!: real_measure_mono simp: increasing_def) + by (auto intro!: finite_measure_mono simp: increasing_def) lemma (in prob_space) prob_zero_union: assumes "s \ events" "t \ events" "prob t = 0" @@ -94,9 +98,9 @@ using assms proof - have "prob (s \ t) \ prob s" - using real_finite_measure_subadditive[of s t] assms by auto + using finite_measure_subadditive[of s t] assms by auto moreover have "prob (s \ t) \ prob s" - using assms by (blast intro: real_measure_mono) + using assms by (blast intro: finite_measure_mono) ultimately show ?thesis by simp qed @@ -127,9 +131,9 @@ using assms proof - have a: "(\ i. prob (f i)) sums (prob (\ i. f i))" - by (rule real_finite_measure_UNION[OF assms(1,3)]) + by (rule finite_measure_UNION[OF assms(1,3)]) have b: "(\ i. prob (g i)) sums (prob (\ i. g i))" - by (rule real_finite_measure_UNION[OF assms(2,4)]) + by (rule finite_measure_UNION[OF assms(2,4)]) show ?thesis using sums_unique[OF b] sums_unique[OF a] assms by simp qed @@ -139,10 +143,9 @@ shows "prob (\ i :: nat. c i) = 0" proof (rule antisym) show "prob (\ i :: nat. c i) \ 0" - using real_finite_measure_countably_subadditive[OF assms(1)] + using finite_measure_countably_subadditive[OF assms(1)] by (simp add: assms(2) suminf_zero summable_zero) - show "0 \ prob (\ i :: nat. c i)" by (rule real_pextreal_nonneg) -qed +qed simp lemma (in prob_space) indep_sym: "indep a b \ indep b a" @@ -163,7 +166,7 @@ from someI_ex[OF this] assms have prob_some: "\ x. x \ s \ prob {x} = prob {SOME y. y \ s}" by blast have "prob s = (\ x \ s. prob {x})" - using real_finite_measure_finite_singelton[OF s_finite] by simp + using finite_measure_finite_singleton[OF s_finite] by simp also have "\ = (\ x \ s. prob {SOME y. y \ s})" using prob_some by auto also have "\ = real (card s) * prob {(SOME x. x \ s)}" using setsum_constant assms by (simp add: real_eq_of_nat) @@ -182,7 +185,7 @@ using `e \ events` sets_into_space upper by blast hence "prob e = prob (\ i \ s. e \ f i)" by simp also have "\ = (\ x \ s. prob (e \ f x))" - proof (rule real_finite_measure_finite_Union) + proof (rule finite_measure_finite_Union) show "finite s" by fact show "\i. i \ s \ e \ f i \ events" by fact show "disjoint_family_on (\i. e \ f i) s" @@ -193,64 +196,62 @@ lemma (in prob_space) distribution_prob_space: assumes "random_variable S X" - shows "prob_space (S\measure := distribution X\)" + shows "prob_space (S\measure := extreal \ distribution X\)" proof - - interpret S: measure_space "S\measure := distribution X\" - unfolding distribution_def using assms - by (intro measure_space_vimage) - (auto intro!: sigma_algebra.sigma_algebra_cong[of S] simp: measure_preserving_def) + interpret S: measure_space "S\measure := extreal \ distribution X\" + proof (rule measure_space.measure_space_cong) + show "measure_space (S\ measure := \A. \ (X -` A \ space M) \)" + using assms by (auto intro!: measure_space_vimage simp: measure_preserving_def) + qed (insert assms, auto simp add: finite_measure_eq distribution_def measurable_sets) show ?thesis proof (default, simp) have "X -` space S \ space M = space M" using `random_variable S X` by (auto simp: measurable_def) - then show "distribution X (space S) = 1" - using measure_space_1 by (simp add: distribution_def) + then show "extreal (distribution X (space S)) = 1" + by (simp add: distribution_def one_extreal_def prob_space) qed qed lemma (in prob_space) AE_distribution: - assumes X: "random_variable MX X" and "measure_space.almost_everywhere (MX\measure := distribution X\) (\x. Q x)" + assumes X: "random_variable MX X" and "AE x in MX\measure := extreal \ distribution X\. Q x" shows "AE x. Q (X x)" proof - - interpret X: prob_space "MX\measure := distribution X\" using X by (rule distribution_prob_space) + interpret X: prob_space "MX\measure := extreal \ distribution X\" using X by (rule distribution_prob_space) obtain N where N: "N \ sets MX" "distribution X N = 0" "{x\space MX. \ Q x} \ N" using assms unfolding X.almost_everywhere_def by auto - show "AE x. Q (X x)" - using X[unfolded measurable_def] N unfolding distribution_def - by (intro AE_I'[where N="X -` N \ space M"]) auto + from X[unfolded measurable_def] N show "AE x. Q (X x)" + by (intro AE_I'[where N="X -` N \ space M"]) + (auto simp: finite_measure_eq distribution_def measurable_sets) qed -lemma (in prob_space) distribution_lebesgue_thm1: - assumes "random_variable s X" - assumes "A \ sets s" - shows "real (distribution X A) = expectation (indicator (X -` A \ space M))" -unfolding distribution_def -using assms unfolding measurable_def -using integral_indicator by auto +lemma (in prob_space) distribution_eq_integral: + "random_variable S X \ A \ sets S \ distribution X A = expectation (indicator (X -` A \ space M))" + using finite_measure_eq[of "X -` A \ space M"] + by (auto simp: measurable_sets distribution_def) -lemma (in prob_space) distribution_lebesgue_thm2: - assumes "random_variable S X" and "A \ sets S" - shows "distribution X A = integral\<^isup>P (S\measure := distribution X\) (indicator A)" +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)" proof - - interpret S: prob_space "S\measure := distribution X\" + interpret S: prob_space "S\measure := extreal \ distribution X\" using assms(1) by (rule distribution_prob_space) show ?thesis - using S.positive_integral_indicator(1) - using assms unfolding distribution_def by auto + using S.positive_integral_indicator(1)[of A] assms by simp qed lemma (in prob_space) finite_expectation1: assumes f: "finite (X`space M)" and rv: "random_variable borel X" - shows "expectation X = (\ r \ X ` space M. r * prob (X -` {r} \ space M))" -proof (rule integral_on_finite(2)[OF rv[THEN conjunct2] f]) - fix x have "X -` {x} \ space M \ sets M" - using rv unfolding measurable_def by auto - thus "\ (X -` {x} \ space M) \ \" using finite_measure by simp + shows "expectation X = (\r \ X ` space M. r * prob (X -` {r} \ space M))" (is "_ = ?r") +proof (subst integral_on_finite) + show "X \ borel_measurable M" "finite (X`space M)" using assms by auto + show "(\ r \ X ` space M. r * real (\ (X -` {r} \ space M))) = ?r" + "\x. \ (X -` {x} \ space M) \ \" + using finite_measure_eq[OF borel_measurable_vimage, of X] rv by auto qed lemma (in prob_space) finite_expectation: assumes "finite (X`space M)" "random_variable borel X" - shows "expectation X = (\ r \ X ` (space M). r * real (distribution X {r}))" + shows "expectation X = (\ r \ X ` (space M). r * distribution X {r})" using assms unfolding distribution_def using finite_expectation1 by auto lemma (in prob_space) prob_x_eq_1_imp_prob_y_eq_0: @@ -267,23 +268,18 @@ lemma (in prob_space) distribution_space[simp]: "distribution X (X ` space M) = 1" proof - have "X -` X ` space M \ space M = space M" by auto - thus ?thesis unfolding distribution_def by (simp add: measure_space_1) + thus ?thesis unfolding distribution_def by (simp add: prob_space) qed lemma (in prob_space) distribution_one: assumes "random_variable M' X" and "A \ sets M'" shows "distribution X A \ 1" proof - - have "distribution X A \ \ (space M)" unfolding distribution_def - using assms[unfolded measurable_def] by (auto intro!: measure_mono) - thus ?thesis by (simp add: measure_space_1) + have "distribution X A \ \' (space M)" unfolding distribution_def + using assms[unfolded measurable_def] by (auto intro!: finite_measure_mono) + thus ?thesis by (simp add: prob_space) qed -lemma (in prob_space) distribution_finite: - assumes "random_variable M' X" and "A \ sets M'" - shows "distribution X A \ \" - using distribution_one[OF assms] by auto - lemma (in prob_space) distribution_x_eq_1_imp_distribution_y_eq_0: assumes X: "random_variable \space = X ` (space M), sets = Pow (X ` (space M))\ X" (is "random_variable ?S X") @@ -302,8 +298,8 @@ have "X -` {x} \ space M - X -` {y} \ space M = X -` {x} \ space M" "X -` {y} \ space M \ (X -` {x} \ space M) = {}" using `y \ x` by auto - with measure_inter_full_set[OF single single, of x y] assms(2) - show ?thesis unfolding distribution_def measure_space_1 by auto + with finite_measure_inter_full_set[OF single single, of x y] assms(2) + show ?thesis by (auto simp: distribution_def prob_space) next assume "{y} \ sets ?S" then have "X -` {y} \ space M = {}" by auto @@ -315,20 +311,17 @@ and A: "A \ sets MX" and B: "B \ sets MY" shows "joint_distribution X Y (A \ B) \ distribution X A" unfolding distribution_def -proof (intro measure_mono) +proof (intro finite_measure_mono) show "(\x. (X x, Y x)) -` (A \ B) \ space M \ X -` A \ space M" by force show "X -` A \ space M \ events" using X A unfolding measurable_def by simp have *: "(\x. (X x, Y x)) -` (A \ B) \ space M = (X -` A \ space M) \ (Y -` B \ space M)" by auto - show "(\x. (X x, Y x)) -` (A \ B) \ space M \ events" - unfolding * apply (rule Int) - using assms unfolding measurable_def by auto qed lemma (in prob_space) joint_distribution_commute: "joint_distribution X Y x = joint_distribution Y X ((\(x,y). (y,x))`x)" - unfolding distribution_def by (auto intro!: arg_cong[where f=\]) + unfolding distribution_def by (auto intro!: arg_cong[where f=\']) lemma (in prob_space) joint_distribution_Times_le_snd: assumes X: "random_variable MX X" and Y: "random_variable MY Y" @@ -352,27 +345,14 @@ unfolding P.measurable_pair_iff[OF sa] using assms by (simp add: comp_def) qed -lemma (in prob_space) distribution_order: - assumes "random_variable MX X" "random_variable MY Y" - assumes "{x} \ sets MX" "{y} \ sets MY" - shows "r \ joint_distribution X Y {(x, y)} \ r \ distribution X {x}" - and "r \ joint_distribution X Y {(x, y)} \ r \ distribution Y {y}" - and "r < joint_distribution X Y {(x, y)} \ r < distribution X {x}" - and "r < joint_distribution X Y {(x, y)} \ r < distribution Y {y}" - and "distribution X {x} = 0 \ joint_distribution X Y {(x, y)} = 0" - and "distribution Y {y} = 0 \ joint_distribution X Y {(x, y)} = 0" - using joint_distribution_Times_le_snd[OF assms] - using joint_distribution_Times_le_fst[OF assms] - by auto - lemma (in prob_space) joint_distribution_commute_singleton: "joint_distribution X Y {(x, y)} = joint_distribution Y X {(y, x)}" - unfolding distribution_def by (auto intro!: arg_cong[where f=\]) + unfolding distribution_def by (auto intro!: arg_cong[where f=\']) lemma (in prob_space) joint_distribution_assoc_singleton: "joint_distribution X (\x. (Y x, Z x)) {(x, y, z)} = joint_distribution (\x. (X x, Y x)) Z {((x, y), z)}" - unfolding distribution_def by (auto intro!: arg_cong[where f=\]) + unfolding distribution_def by (auto intro!: arg_cong[where f=\']) locale pair_prob_space = M1: prob_space M1 + M2: prob_space M2 for M1 M2 @@ -383,13 +363,13 @@ lemma countably_additiveI[case_names countably]: assumes "\A. \ range A \ sets M ; disjoint_family A ; (\i. A i) \ sets M\ \ - (\\<^isub>\n. \ (A n)) = \ (\i. A i)" + (\n. \ (A n)) = \ (\i. A i)" shows "countably_additive M \" using assms unfolding countably_additive_def by auto lemma (in prob_space) joint_distribution_prob_space: assumes "random_variable MX X" "random_variable MY Y" - shows "prob_space ((MX \\<^isub>M MY) \ measure := joint_distribution X Y\)" + shows "prob_space ((MX \\<^isub>M MY) \ measure := extreal \ joint_distribution X Y\)" using random_variable_pairI[OF assms] by (rule distribution_prob_space) section "Probability spaces on finite sets" @@ -407,20 +387,13 @@ lemma (in prob_space) distribution_finite_prob_space: assumes "finite_random_variable MX X" - shows "finite_prob_space (MX\measure := distribution X\)" + shows "finite_prob_space (MX\measure := extreal \ distribution X\)" proof - - interpret X: prob_space "MX\measure := distribution X\" + interpret X: prob_space "MX\measure := extreal \ distribution X\" using assms[THEN finite_random_variableD] by (rule distribution_prob_space) interpret MX: finite_sigma_algebra MX using assms by auto - show ?thesis - proof (default, simp_all) - fix x assume "x \ space MX" - then have "X -` {x} \ space M \ sets M" - using assms unfolding measurable_def by simp - then show "distribution X {x} \ \" - unfolding distribution_def by simp - qed (rule MX.finite_space) + show ?thesis by default (simp_all add: MX.finite_space) qed lemma (in prob_space) simple_function_imp_finite_random_variable[simp, intro]: @@ -451,9 +424,9 @@ by (auto dest!: finite_random_variableD) lemma (in prob_space) sum_over_space_real_distribution: - "simple_function M X \ (\x\X`space M. real (distribution X {x})) = 1" + "simple_function M X \ (\x\X`space M. distribution X {x}) = 1" unfolding distribution_def prob_space[symmetric] - by (subst real_finite_measure_finite_Union[symmetric]) + by (subst finite_measure_finite_Union[symmetric]) (auto simp add: disjoint_family_on_def simple_function_def intro!: arg_cong[where f=prob]) @@ -475,7 +448,7 @@ "finite_random_variable MX X \ x \ space MX \ {x} \ sets MX" unfolding finite_sigma_algebra_def finite_sigma_algebra_axioms_def by simp -lemma (in prob_space) finite_random_variable_vimage: +lemma (in prob_space) finite_random_variable_measurable: assumes X: "finite_random_variable MX X" shows "X -` A \ space M \ events" proof - interpret X: finite_sigma_algebra MX using X by simp @@ -492,15 +465,12 @@ assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y" shows "joint_distribution X Y (A \ B) \ distribution X A" unfolding distribution_def -proof (intro measure_mono) +proof (intro finite_measure_mono) show "(\x. (X x, Y x)) -` (A \ B) \ space M \ X -` A \ space M" by force show "X -` A \ space M \ events" - using finite_random_variable_vimage[OF X] . + using finite_random_variable_measurable[OF X] . have *: "(\x. (X x, Y x)) -` (A \ B) \ space M = (X -` A \ space M) \ (Y -` B \ space M)" by auto - show "(\x. (X x, Y x)) -` (A \ B) \ space M \ events" - unfolding * apply (rule Int) - using assms[THEN finite_random_variable_vimage] by auto qed lemma (in prob_space) joint_distribution_finite_Times_le_snd: @@ -511,6 +481,7 @@ (simp add: swap_product joint_distribution_finite_Times_le_fst) lemma (in prob_space) finite_distribution_order: + fixes MX :: "('c, 'd) measure_space_scheme" and MY :: "('e, 'f) measure_space_scheme" assumes "finite_random_variable MX X" "finite_random_variable MY Y" shows "r \ joint_distribution X Y {(x, y)} \ r \ distribution X {x}" and "r \ joint_distribution X Y {(x, y)} \ r \ distribution Y {y}" @@ -520,26 +491,14 @@ and "distribution Y {y} = 0 \ joint_distribution X Y {(x, y)} = 0" using joint_distribution_finite_Times_le_snd[OF assms, of "{x}" "{y}"] using joint_distribution_finite_Times_le_fst[OF assms, of "{x}" "{y}"] - by auto - -lemma (in prob_space) finite_distribution_finite: - assumes "finite_random_variable M' X" - shows "distribution X {x} \ \" -proof - - have "distribution X {x} \ \ (space M)" - unfolding distribution_def - using finite_random_variable_vimage[OF assms] - by (intro measure_mono) auto - then show ?thesis - by auto -qed + by (auto intro: antisym) lemma (in prob_space) setsum_joint_distribution: assumes X: "finite_random_variable MX X" assumes Y: "random_variable MY Y" "B \ sets MY" shows "(\a\space MX. joint_distribution X Y ({a} \ B)) = distribution Y B" unfolding distribution_def -proof (subst measure_finitely_additive'') +proof (subst finite_measure_finite_Union[symmetric]) interpret MX: finite_sigma_algebra MX using X by auto show "finite (space MX)" using MX.finite_space . let "?d i" = "(\x. (X x, Y x)) -` ({i} \ B) \ space M" @@ -549,9 +508,8 @@ using measurable_sets[of X M MX] measurable_sets[of Y M MY B] X Y using MX.sets_eq_Pow by auto } show "disjoint_family_on ?d (space MX)" by (auto simp: disjoint_family_on_def) - show "\ (\i\space MX. ?d i) = \ (Y -` B \ space M)" - using X[unfolded measurable_def] - by (auto intro!: arg_cong[where f=\]) + show "\' (\i\space MX. ?d i) = \' (Y -` B \ space M)" + using X[unfolded measurable_def] by (auto intro!: arg_cong[where f=\']) qed lemma (in prob_space) setsum_joint_distribution_singleton: @@ -562,28 +520,6 @@ finite_random_variableD[OF Y(1)] finite_random_variable_imp_sets[OF Y]] by simp -lemma (in prob_space) setsum_real_joint_distribution: - fixes MX :: "('c, 'x) measure_space_scheme" and MY :: "('d, 'y) measure_space_scheme" - assumes X: "finite_random_variable MX X" - assumes Y: "random_variable MY Y" "B \ sets MY" - shows "(\a\space MX. real (joint_distribution X Y ({a} \ B))) = real (distribution Y B)" -proof - - interpret MX: finite_sigma_algebra MX using X by auto - show ?thesis - unfolding setsum_joint_distribution[OF assms, symmetric] - using distribution_finite[OF random_variable_pairI[OF finite_random_variableD[OF X] Y(1)]] Y(2) - by (simp add: space_pair_measure real_of_pextreal_setsum) -qed - -lemma (in prob_space) setsum_real_joint_distribution_singleton: - fixes MX :: "('c, 'x) measure_space_scheme" and MY :: "('d, 'y) measure_space_scheme" - assumes X: "finite_random_variable MX X" - assumes Y: "finite_random_variable MY Y" "b \ space MY" - shows "(\a\space MX. real (joint_distribution X Y {(a,b)})) = real (distribution Y {b})" - using setsum_real_joint_distribution[OF X - finite_random_variableD[OF Y(1)] - finite_random_variable_imp_sets[OF Y]] by simp - locale pair_finite_prob_space = M1: finite_prob_space M1 + M2: finite_prob_space M2 for M1 M2 sublocale pair_finite_prob_space \ pair_prob_space M1 M2 by default @@ -593,7 +529,7 @@ lemma (in prob_space) joint_distribution_finite_prob_space: assumes X: "finite_random_variable MX X" assumes Y: "finite_random_variable MY Y" - shows "finite_prob_space ((MX \\<^isub>M MY)\ measure := joint_distribution X Y\)" + shows "finite_prob_space ((MX \\<^isub>M MY)\ measure := extreal \ joint_distribution X Y\)" by (intro distribution_finite_prob_space finite_random_variable_pairI X Y) lemma finite_prob_space_eq: @@ -602,18 +538,15 @@ by auto lemma (in prob_space) not_empty: "space M \ {}" - using prob_space empty_measure by auto + 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 -lemma (in finite_prob_space) positive_distribution: "0 \ distribution X x" - unfolding distribution_def by simp - lemma (in finite_prob_space) joint_distribution_restriction_fst: "joint_distribution X Y A \ distribution X (fst ` A)" unfolding distribution_def -proof (safe intro!: measure_mono) +proof (safe intro!: finite_measure_mono) fix x assume "x \ space M" and *: "(X x, Y x) \ A" show "x \ X -` fst ` A" by (auto intro!: image_eqI[OF _ *]) @@ -622,7 +555,7 @@ lemma (in finite_prob_space) joint_distribution_restriction_snd: "joint_distribution X Y A \ distribution Y (snd ` A)" unfolding distribution_def -proof (safe intro!: measure_mono) +proof (safe intro!: finite_measure_mono) fix x assume "x \ space M" and *: "(X x, Y x) \ A" show "x \ Y -` snd ` A" by (auto intro!: image_eqI[OF _ *]) @@ -637,17 +570,16 @@ and "r < joint_distribution X Y {(x, y)} \ r < distribution Y {y}" and "distribution X {x} = 0 \ joint_distribution X Y {(x, y)} = 0" and "distribution Y {y} = 0 \ joint_distribution X Y {(x, y)} = 0" - using positive_distribution[of X x'] - positive_distribution[of "\x. (X x, Y x)" "{(x, y)}"] + using joint_distribution_restriction_fst[of X Y "{(x, y)}"] joint_distribution_restriction_snd[of X Y "{(x, y)}"] - by auto + by (auto intro: antisym) lemma (in finite_prob_space) distribution_mono: assumes "\t. \ t \ space M ; X t \ x \ \ Y t \ y" shows "distribution X x \ distribution Y y" unfolding distribution_def - using assms by (auto simp: sets_eq_Pow intro!: measure_mono) + using assms by (auto simp: sets_eq_Pow intro!: finite_measure_mono) lemma (in finite_prob_space) distribution_mono_gt_0: assumes gt_0: "0 < distribution X x" @@ -657,49 +589,21 @@ lemma (in finite_prob_space) sum_over_space_distrib: "(\x\X`space M. distribution X {x}) = 1" - unfolding distribution_def measure_space_1[symmetric] using finite_space - by (subst measure_finitely_additive'') - (auto simp add: disjoint_family_on_def sets_eq_Pow intro!: arg_cong[where f=\]) + unfolding distribution_def prob_space[symmetric] using finite_space + by (subst finite_measure_finite_Union[symmetric]) + (auto simp add: disjoint_family_on_def sets_eq_Pow + intro!: arg_cong[where f=\']) lemma (in finite_prob_space) sum_over_space_real_distribution: - "(\x\X`space M. real (distribution X {x})) = 1" + "(\x\X`space M. distribution X {x}) = 1" unfolding distribution_def prob_space[symmetric] using finite_space - by (subst real_finite_measure_finite_Union[symmetric]) + by (subst finite_measure_finite_Union[symmetric]) (auto simp add: disjoint_family_on_def sets_eq_Pow intro!: arg_cong[where f=prob]) lemma (in finite_prob_space) finite_sum_over_space_eq_1: - "(\x\space M. real (\ {x})) = 1" - using sum_over_space_eq_1 finite_measure by (simp add: real_of_pextreal_setsum sets_eq_Pow) - -lemma (in finite_prob_space) distribution_finite: - "distribution X A \ \" - using finite_measure[of "X -` A \ space M"] - unfolding distribution_def sets_eq_Pow by auto - -lemma (in finite_prob_space) real_distribution_gt_0[simp]: - "0 < real (distribution Y y) \ 0 < distribution Y y" - using assms by (auto intro!: real_pextreal_pos distribution_finite) - -lemma (in finite_prob_space) real_distribution_mult_pos_pos: - assumes "0 < distribution Y y" - and "0 < distribution X x" - shows "0 < real (distribution Y y * distribution X x)" - unfolding real_of_pextreal_mult[symmetric] - using assms by (auto intro!: mult_pos_pos) - -lemma (in finite_prob_space) real_distribution_divide_pos_pos: - assumes "0 < distribution Y y" - and "0 < distribution X x" - shows "0 < real (distribution Y y / distribution X x)" - unfolding divide_pextreal_def real_of_pextreal_mult[symmetric] - using assms distribution_finite[of X x] by (cases "distribution X x") (auto intro!: mult_pos_pos) - -lemma (in finite_prob_space) real_distribution_mult_inverse_pos_pos: - assumes "0 < distribution Y y" - and "0 < distribution X x" - shows "0 < real (distribution Y y * inverse (distribution X x))" - unfolding divide_pextreal_def real_of_pextreal_mult[symmetric] - using assms distribution_finite[of X x] by (cases "distribution X x") (auto intro!: mult_pos_pos) + "(\x\space M. prob {x}) = 1" + using prob_space finite_space + by (subst (asm) finite_measure_finite_singleton) auto lemma (in prob_space) distribution_remove_const: shows "joint_distribution X (\x. ()) {(x, ())} = distribution X {x}" @@ -707,8 +611,7 @@ and "joint_distribution X (\x. (Y x, ())) {(x, y, ())} = joint_distribution X Y {(x, y)}" and "joint_distribution X (\x. ((), Y x)) {(x, (), y)} = joint_distribution X Y {(x, y)}" and "distribution (\x. ()) {()} = 1" - unfolding measure_space_1[symmetric] - by (auto intro!: arg_cong[where f="\"] simp: distribution_def) + by (auto intro!: arg_cong[where f=\'] simp: distribution_def prob_space[symmetric]) lemma (in finite_prob_space) setsum_distribution_gen: assumes "Z -` {c} \ space M = (\x \ X`space M. Y -` {f x}) \ space M" @@ -716,7 +619,7 @@ shows "(\x \ X`space M. distribution Y {f x}) = distribution Z {c}" unfolding distribution_def assms using finite_space assms - by (subst measure_finitely_additive'') + by (subst finite_measure_finite_Union[symmetric]) (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def intro!: arg_cong[where f=prob]) @@ -728,61 +631,17 @@ "(\z \ Z`space M. joint_distribution X (\x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Y {(x, y)}" by (auto intro!: inj_onI setsum_distribution_gen) -lemma (in finite_prob_space) setsum_real_distribution_gen: - assumes "Z -` {c} \ space M = (\x \ X`space M. Y -` {f x}) \ space M" - and "inj_on f (X`space M)" - shows "(\x \ X`space M. real (distribution Y {f x})) = real (distribution Z {c})" - unfolding distribution_def assms - using finite_space assms - by (subst real_finite_measure_finite_Union[symmetric]) - (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def - intro!: arg_cong[where f=prob]) - -lemma (in finite_prob_space) setsum_real_distribution: - "(\x \ X`space M. real (joint_distribution X Y {(x, y)})) = real (distribution Y {y})" - "(\y \ Y`space M. real (joint_distribution X Y {(x, y)})) = real (distribution X {x})" - "(\x \ X`space M. real (joint_distribution X (\x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution Y Z {(y, z)})" - "(\y \ Y`space M. real (joint_distribution X (\x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution X Z {(x, z)})" - "(\z \ Z`space M. real (joint_distribution X (\x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution X Y {(x, y)})" - by (auto intro!: inj_onI setsum_real_distribution_gen) - -lemma (in finite_prob_space) real_distribution_order: - shows "r \ real (joint_distribution X Y {(x, y)}) \ r \ real (distribution X {x})" - and "r \ real (joint_distribution X Y {(x, y)}) \ r \ real (distribution Y {y})" - and "r < real (joint_distribution X Y {(x, y)}) \ r < real (distribution X {x})" - and "r < real (joint_distribution X Y {(x, y)}) \ r < real (distribution Y {y})" - and "distribution X {x} = 0 \ real (joint_distribution X Y {(x, y)}) = 0" - and "distribution Y {y} = 0 \ real (joint_distribution X Y {(x, y)}) = 0" - using real_of_pextreal_mono[OF distribution_finite joint_distribution_restriction_fst, of X Y "{(x, y)}"] - using real_of_pextreal_mono[OF distribution_finite joint_distribution_restriction_snd, of X Y "{(x, y)}"] - using real_pextreal_nonneg[of "joint_distribution X Y {(x, y)}"] - by auto - -lemma (in prob_space) joint_distribution_remove[simp]: - "joint_distribution X X {(x, x)} = distribution X {x}" - unfolding distribution_def by (auto intro!: arg_cong[where f="\"]) - -lemma (in finite_prob_space) distribution_1: - "distribution X A \ 1" - unfolding distribution_def measure_space_1[symmetric] - by (auto intro!: measure_mono simp: sets_eq_Pow) - -lemma (in finite_prob_space) real_distribution_1: - "real (distribution X A) \ 1" - unfolding real_pextreal_1[symmetric] - by (rule real_of_pextreal_mono[OF _ distribution_1]) simp - lemma (in finite_prob_space) uniform_prob: assumes "x \ space M" assumes "\ x y. \x \ space M ; y \ space M\ \ prob {x} = prob {y}" - shows "prob {x} = 1 / real (card (space M))" + shows "prob {x} = 1 / card (space M)" proof - have prob_x: "\ y. y \ space M \ prob {y} = prob {x}" using assms(2)[OF _ `x \ space M`] by blast have "1 = prob (space M)" using prob_space by auto also have "\ = (\ x \ space M. prob {x})" - using real_finite_measure_finite_Union[of "space M" "\ x. {x}", simplified] + using finite_measure_finite_Union[of "space M" "\ x. {x}", simplified] sets_eq_Pow inj_singleton[unfolded inj_on_def, rule_format] finite_space unfolding disjoint_family_on_def prob_space[symmetric] by (auto simp add:setsum_restrict_set) @@ -809,7 +668,7 @@ qed lemma (in prob_space) prob_space_of_restricted_space: - assumes "\ A \ 0" "\ A \ \" "A \ sets M" + assumes "\ A \ 0" "A \ sets M" shows "prob_space (restricted_space A \measure := \S. \ S / \ A\)" (is "prob_space ?P") proof - @@ -820,29 +679,48 @@ show "prob_space ?P" proof show "measure ?P (space ?P) = 1" - using `\ A \ 0` `\ A \ \` by (auto simp: pextreal_noteq_omega_Ex) - show "measure ?P {} = 0" by auto - have "\S. \ S / \ A = inverse (\ A) * \ S" by (simp add: mult_commute) - then show "countably_additive ?P (measure ?P)" - unfolding countably_additive_def psuminf_cmult_right - using A.measure_countably_additive by auto + using real_measure[OF `A \ events`] `\ A \ 0` by auto + show "positive ?P (measure ?P)" + proof (simp add: positive_def, safe) + show "0 / \ A = 0" using `\ A \ 0` by (cases "\ A") (auto simp: zero_extreal_def) + fix B assume "B \ events" + with real_measure[of "A \ B"] real_measure[OF `A \ events`] `A \ sets M` + show "0 \ \ (A \ B) / \ A" by (auto simp: Int) + qed + show "countably_additive ?P (measure ?P)" + proof (simp add: countably_additive_def, safe) + fix B and F :: "nat \ 'a set" + assume F: "range F \ op \ A ` events" "disjoint_family F" + { fix i + from F have "F i \ op \ A ` events" by auto + with `A \ events` have "F i \ events" by auto } + moreover then have "range F \ events" by auto + moreover have "\S. \ S / \ A = inverse (\ A) * \ S" + by (simp add: mult_commute divide_extreal_def) + moreover have "0 \ inverse (\ A)" + using real_measure[OF `A \ events`] by auto + ultimately show "(\i. \ (F i) / \ A) = \ (\i. F i) / \ A" + using measure_countably_additive[of F] F + by (auto simp: suminf_cmult_extreal) + qed qed qed lemma finite_prob_spaceI: - assumes "finite (space M)" "sets M = Pow(space M)" "measure M (space M) = 1" "measure M {} = 0" + assumes "finite (space M)" "sets M = Pow(space M)" + and "measure M (space M) = 1" "measure M {} = 0" "\A. A \ space M \ 0 \ measure M A" and "\A B. A\space M \ B\space M \ A \ B = {} \ measure M (A \ B) = measure M A + measure M B" shows "finite_prob_space M" unfolding finite_prob_space_eq proof show "finite_measure_space M" using assms - by (auto intro!: finite_measure_spaceI) + by (auto intro!: finite_measure_spaceI) show "measure M (space M) = 1" by fact qed lemma (in finite_prob_space) finite_measure_space: fixes X :: "'a \ 'x" - shows "finite_measure_space \space = X ` space M, sets = Pow (X ` space M), measure = distribution X\" + shows "finite_measure_space \space = X ` space M, sets = Pow (X ` space M), measure = extreal \ distribution X\" (is "finite_measure_space ?S") proof (rule finite_measure_spaceI, simp_all) show "finite (X ` space M)" using finite_space by simp @@ -850,51 +728,41 @@ fix A B :: "'x set" assume "A \ B = {}" then show "distribution X (A \ B) = distribution X A + distribution X B" unfolding distribution_def - by (subst measure_additive) - (auto intro!: arg_cong[where f=\] simp: sets_eq_Pow) + by (subst finite_measure_Union[symmetric]) + (auto intro!: arg_cong[where f=\'] simp: sets_eq_Pow) qed lemma (in finite_prob_space) finite_prob_space_of_images: - "finite_prob_space \ space = X ` space M, sets = Pow (X ` space M), measure = distribution X \" - by (simp add: finite_prob_space_eq finite_measure_space) - -lemma (in finite_prob_space) real_distribution_order': - shows "real (distribution X {x}) = 0 \ real (joint_distribution X Y {(x, y)}) = 0" - and "real (distribution Y {y}) = 0 \ real (joint_distribution X Y {(x, y)}) = 0" - using real_of_pextreal_mono[OF distribution_finite joint_distribution_restriction_fst, of X Y "{(x, y)}"] - using real_of_pextreal_mono[OF distribution_finite joint_distribution_restriction_snd, of X Y "{(x, y)}"] - using real_pextreal_nonneg[of "joint_distribution X Y {(x, y)}"] - by auto + "finite_prob_space \ space = X ` space M, sets = Pow (X ` space M), measure = extreal \ distribution X \" + by (simp add: finite_prob_space_eq finite_measure_space measure_space_1 one_extreal_def) lemma (in finite_prob_space) finite_product_measure_space: fixes X :: "'a \ 'x" and Y :: "'a \ 'y" assumes "finite s1" "finite s2" - shows "finite_measure_space \ space = s1 \ s2, sets = Pow (s1 \ s2), measure = joint_distribution X Y\" + shows "finite_measure_space \ space = s1 \ s2, sets = Pow (s1 \ s2), measure = extreal \ joint_distribution X Y\" (is "finite_measure_space ?M") proof (rule finite_measure_spaceI, simp_all) show "finite (s1 \ s2)" using assms by auto - show "joint_distribution X Y (s1\s2) \ \" - using distribution_finite . next fix A B :: "('x*'y) set" assume "A \ B = {}" then show "joint_distribution X Y (A \ B) = joint_distribution X Y A + joint_distribution X Y B" unfolding distribution_def - by (subst measure_additive) - (auto intro!: arg_cong[where f=\] simp: sets_eq_Pow) + by (subst finite_measure_Union[symmetric]) + (auto intro!: arg_cong[where f=\'] simp: sets_eq_Pow) qed lemma (in finite_prob_space) finite_product_measure_space_of_images: shows "finite_measure_space \ space = X ` space M \ Y ` space M, sets = Pow (X ` space M \ Y ` space M), - measure = joint_distribution X Y \" + measure = extreal \ joint_distribution X Y \" using finite_space by (auto intro!: finite_product_measure_space) lemma (in finite_prob_space) finite_product_prob_space_of_images: "finite_prob_space \ space = X ` space M \ Y ` space M, sets = Pow (X ` space M \ Y ` space M), - measure = joint_distribution X Y \" + measure = extreal \ joint_distribution X Y \" (is "finite_prob_space ?S") -proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images) +proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images one_extreal_def) have "X -` X ` space M \ Y -` Y ` space M \ space M = space M" by auto thus "joint_distribution X Y (X ` space M \ Y ` space M) = 1" by (simp add: distribution_def prob_space vimage_Times comp_def measure_space_1) @@ -903,11 +771,11 @@ section "Conditional Expectation and Probability" lemma (in prob_space) conditional_expectation_exists: - fixes X :: "'a \ pextreal" and N :: "('a, 'b) measure_space_scheme" - assumes borel: "X \ borel_measurable M" + fixes X :: "'a \ extreal" and N :: "('a, 'b) measure_space_scheme" + assumes borel: "X \ borel_measurable M" "AE x. 0 \ X x" and N: "sigma_algebra N" "sets N \ sets M" "space N = space M" "\A. measure N A = \ A" - shows "\Y\borel_measurable N. \C\sets N. - (\\<^isup>+x. Y x * indicator C x \M) = (\\<^isup>+x. X x * indicator C x \M)" + shows "\Y\borel_measurable N. (\x. 0 \ Y x) \ (\C\sets N. + (\\<^isup>+x. Y x * indicator C x \M) = (\\<^isup>+x. X x * indicator C x \M))" proof - note N(4)[simp] interpret P: prob_space N @@ -926,34 +794,29 @@ unfolding P.absolutely_continuous_def proof safe fix A assume "A \ sets N" "P.\ A = 0" - moreover then have f_borel: "?f A \ borel_measurable M" - using borel N by (auto intro: borel_measurable_indicator) - moreover have "{x\space M. ?f A x \ 0} = (?f A -` {0<..} \ space M) \ A" - by (auto simp: indicator_def) - moreover have "\ \ \ \ A" - using `A \ sets N` N f_borel - by (auto intro!: measure_mono Int[of _ A] measurable_sets) - ultimately show "?Q A = 0" - by (simp add: positive_integral_0_iff) + then have f_borel: "?f A \ borel_measurable M" "AE x. x \ A" + using borel N by (auto intro!: borel_measurable_indicator AE_not_in) + then show "?Q A = 0" + by (auto simp add: positive_integral_0_iff_AE) qed from P.Radon_Nikodym[OF Q this] - obtain Y where Y: "Y \ borel_measurable N" + obtain Y where Y: "Y \ borel_measurable N" "\x. 0 \ Y x" "\A. A \ sets N \ ?Q A =(\\<^isup>+x. Y x * indicator A x \N)" by blast with N(2) show ?thesis - by (auto intro!: bexI[OF _ Y(1)] simp: positive_integral_subalgebra[OF _ N(2,3,4,1)]) + by (auto intro!: bexI[OF _ Y(1)] simp: positive_integral_subalgebra[OF _ _ N(2,3,4,1)]) qed definition (in prob_space) - "conditional_expectation N X = (SOME Y. Y\borel_measurable N + "conditional_expectation N X = (SOME Y. Y\borel_measurable N \ (\x. 0 \ Y x) \ (\C\sets N. (\\<^isup>+x. Y x * indicator C x\M) = (\\<^isup>+x. X x * indicator C x\M)))" abbreviation (in prob_space) "conditional_prob N A \ conditional_expectation N (indicator A)" lemma (in prob_space) - fixes X :: "'a \ pextreal" and N :: "('a, 'b) measure_space_scheme" - assumes borel: "X \ borel_measurable M" + fixes X :: "'a \ extreal" and N :: "('a, 'b) measure_space_scheme" + assumes borel: "X \ borel_measurable M" "AE x. 0 \ X x" and N: "sigma_algebra N" "sets N \ sets M" "space N = space M" "\A. measure N A = \ A" shows borel_measurable_conditional_expectation: "conditional_expectation N X \ borel_measurable N" @@ -970,36 +833,23 @@ unfolding conditional_expectation_def by (rule someI2_ex) blast qed -lemma (in sigma_algebra) factorize_measurable_function: - fixes Z :: "'a \ pextreal" and Y :: "'a \ 'c" +lemma (in sigma_algebra) factorize_measurable_function_pos: + fixes Z :: "'a \ extreal" and Y :: "'a \ 'c" assumes "sigma_algebra M'" and "Y \ measurable M M'" "Z \ borel_measurable M" - shows "Z \ borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y) - \ (\g\borel_measurable M'. \x\space M. Z x = g (Y x))" -proof safe + assumes Z: "Z \ borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y)" + shows "\g\borel_measurable M'. \x\space M. max 0 (Z x) = g (Y x)" +proof - interpret M': sigma_algebra M' by fact have Y: "Y \ space M \ space M'" using assms unfolding measurable_def by auto from M'.sigma_algebra_vimage[OF this] interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" . - { fix g :: "'c \ pextreal" assume "g \ borel_measurable M'" - with M'.measurable_vimage_algebra[OF Y] - have "g \ Y \ borel_measurable (M'.vimage_algebra (space M) Y)" - by (rule measurable_comp) - moreover assume "\x\space M. Z x = g (Y x)" - then have "Z \ borel_measurable (M'.vimage_algebra (space M) Y) \ - g \ Y \ borel_measurable (M'.vimage_algebra (space M) Y)" - by (auto intro!: measurable_cong) - ultimately show "Z \ borel_measurable (M'.vimage_algebra (space M) Y)" - by simp } - - assume "Z \ borel_measurable (M'.vimage_algebra (space M) Y)" - from va.borel_measurable_implies_simple_function_sequence[OF this] - obtain f where f: "\i. simple_function (M'.vimage_algebra (space M) Y) (f i)" and "f \ Z" by blast + from va.borel_measurable_implies_simple_function_sequence'[OF Z] guess f . note f = this have "\i. \g. simple_function M' g \ (\x\space M. f i x = g (Y x))" proof fix i - from f[of i] have "finite (f i`space M)" and B_ex: + from f(1)[of i] have "finite (f i`space M)" and B_ex: "\z\(f i)`space M. \B. B \ sets M' \ (f i) -` {z} \ space M = Y -` B \ space M" unfolding simple_function_def by auto from B_ex[THEN bchoice] guess B .. note B = this @@ -1011,23 +861,67 @@ show "simple_function M' ?g" using B by auto fix x assume "x \ space M" - then have "\z. z \ f i`space M \ indicator (B z) (Y x) = (indicator (f i -` {z} \ space M) x::pextreal)" + then have "\z. z \ f i`space M \ indicator (B z) (Y x) = (indicator (f i -` {z} \ space M) x::extreal)" unfolding indicator_def using B by auto - then show "f i x = ?g (Y x)" using `x \ space M` f[of i] + then show "f i x = ?g (Y x)" using `x \ space M` f(1)[of i] by (subst va.simple_function_indicator_representation) auto qed qed from choice[OF this] guess g .. note g = this - show "\g\borel_measurable M'. \x\space M. Z x = g (Y x)" + show ?thesis proof (intro ballI bexI) show "(\x. SUP i. g i x) \ borel_measurable M'" using g by (auto intro: M'.borel_measurable_simple_function) fix x assume "x \ space M" - have "Z x = (SUP i. f i) x" using `f \ Z` unfolding isoton_def by simp - also have "\ = (SUP i. g i (Y x))" unfolding SUPR_apply + have "max 0 (Z x) = (SUP i. f i x)" using f by simp + also have "\ = (SUP i. g i (Y x))" using g `x \ space M` by simp - finally show "Z x = (SUP i. g i (Y x))" . + finally show "max 0 (Z x) = (SUP i. g i (Y x))" . + qed +qed + +lemma extreal_0_le_iff_le_0[simp]: + fixes a :: extreal shows "0 \ -a \ a \ 0" + by (cases rule: extreal2_cases[of a]) auto + +lemma (in sigma_algebra) factorize_measurable_function: + fixes Z :: "'a \ extreal" and Y :: "'a \ 'c" + assumes "sigma_algebra M'" and "Y \ measurable M M'" "Z \ borel_measurable M" + shows "Z \ borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y) + \ (\g\borel_measurable M'. \x\space M. Z x = g (Y x))" +proof safe + interpret M': sigma_algebra M' by fact + have Y: "Y \ space M \ space M'" using assms unfolding measurable_def by auto + from M'.sigma_algebra_vimage[OF this] + interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" . + + { fix g :: "'c \ extreal" assume "g \ borel_measurable M'" + with M'.measurable_vimage_algebra[OF Y] + have "g \ Y \ borel_measurable (M'.vimage_algebra (space M) Y)" + by (rule measurable_comp) + moreover assume "\x\space M. Z x = g (Y x)" + then have "Z \ borel_measurable (M'.vimage_algebra (space M) Y) \ + g \ Y \ borel_measurable (M'.vimage_algebra (space M) Y)" + by (auto intro!: measurable_cong) + ultimately show "Z \ borel_measurable (M'.vimage_algebra (space M) Y)" + by simp } + + assume Z: "Z \ borel_measurable (M'.vimage_algebra (space M) Y)" + with assms have "(\x. - Z x) \ borel_measurable M" + "(\x. - Z x) \ borel_measurable (M'.vimage_algebra (space M) Y)" + by auto + from factorize_measurable_function_pos[OF assms(1,2) this] guess n .. note n = this + from factorize_measurable_function_pos[OF assms Z] guess p .. note p = this + let "?g x" = "p x - n x" + show "\g\borel_measurable M'. \x\space M. Z x = g (Y x)" + proof (intro bexI ballI) + show "?g \ borel_measurable M'" using p n by auto + fix x assume "x \ space M" + then have "p (Y x) = max 0 (Z x)" "n (Y x) = max 0 (- Z x)" + using p n by auto + then show "Z x = ?g (Y x)" + by (auto split: split_max) qed qed diff -r 28b51effc5ed -r cdf7693bbe08 src/HOL/Probability/Product_Measure.thy --- a/src/HOL/Probability/Product_Measure.thy Mon Mar 14 14:37:47 2011 +0100 +++ b/src/HOL/Probability/Product_Measure.thy Mon Mar 14 14:37:49 2011 +0100 @@ -416,7 +416,7 @@ show "?f -` A \ space ?P \ sets ?P" by simp qed -lemma (in pair_sigma_algebra) measurable_cut_fst: +lemma (in pair_sigma_algebra) measurable_cut_fst[simp,intro]: assumes "Q \ sets P" shows "Pair x -` Q \ sets M2" proof - let ?Q' = "{Q. Q \ space P \ Pair x -` Q \ sets M2}" @@ -504,15 +504,14 @@ fix A assume "A \ sets ?D" with sets_into_space have "\x. measure M2 (Pair x -` (space M1 \ space M2 - A)) = (if x \ space M1 then measure M2 (space M2) - ?s A x else 0)" - by (auto intro!: M2.finite_measure_compl measurable_cut_fst - simp: vimage_Diff) + by (auto intro!: M2.measure_compl simp: vimage_Diff) with `A \ sets ?D` top show "space ?D - A \ sets ?D" - by (auto intro!: Diff M1.measurable_If M1.borel_measurable_pextreal_diff) + by (auto intro!: Diff M1.measurable_If M1.borel_measurable_extreal_diff) next fix F :: "nat \ ('a\'b) set" assume "disjoint_family F" "range F \ sets ?D" - moreover then have "\x. measure M2 (\i. Pair x -` F i) = (\\<^isub>\ i. ?s (F i) x)" + moreover then have "\x. measure M2 (\i. Pair x -` F i) = (\i. ?s (F i) x)" by (intro M2.measure_countably_additive[symmetric]) - (auto intro!: measurable_cut_fst simp: disjoint_family_on_def) + (auto simp: disjoint_family_on_def) ultimately show "(\i. F i) \ sets ?D" by (auto simp: vimage_UN intro!: M1.borel_measurable_psuminf) qed @@ -546,7 +545,7 @@ have [intro]: "sigma_algebra M1" and [intro]: "sigma_algebra M2" by default+ have M1: "sigma_finite_measure M1" by default from M2.disjoint_sigma_finite guess F .. note F = this - then have "\i. F i \ sets M2" by auto + then have F_sets: "\i. F i \ sets M2" by auto let "?C x i" = "F i \ Pair x -` Q" { fix i let ?R = "M2.restricted_space (F i)" @@ -578,10 +577,10 @@ by simp } moreover { fix x - have "(\\<^isub>\i. measure M2 (?C x i)) = measure M2 (\i. ?C x i)" + have "(\i. measure M2 (?C x i)) = measure M2 (\i. ?C x i)" proof (intro M2.measure_countably_additive) show "range (?C x) \ sets M2" - using F `Q \ sets P` by (auto intro!: M2.Int measurable_cut_fst) + using F `Q \ sets P` by (auto intro!: M2.Int) have "disjoint_family F" using F by auto show "disjoint_family (?C x)" by (rule disjoint_family_on_bisimulation[OF `disjoint_family F`]) auto @@ -589,10 +588,10 @@ also have "(\i. ?C x i) = Pair x -` Q" using F sets_into_space `Q \ sets P` by (auto simp: space_pair_measure) - finally have "measure M2 (Pair x -` Q) = (\\<^isub>\i. measure M2 (?C x i))" + finally have "measure M2 (Pair x -` Q) = (\i. measure M2 (?C x i))" by simp } - ultimately show ?thesis - by (auto intro!: M1.borel_measurable_psuminf) + ultimately show ?thesis using `Q \ sets P` F_sets + by (auto intro!: M1.borel_measurable_psuminf M2.Int) qed lemma (in pair_sigma_finite) measure_cut_measurable_snd: @@ -620,7 +619,7 @@ apply (simp add: pair_measure_def pair_measure_generator_def) proof (rule M1.positive_integral_cong) fix x assume "x \ space M1" - have *: "\y. indicator A (x, y) = (indicator (Pair x -` A) y :: pextreal)" + have *: "\y. indicator A (x, y) = (indicator (Pair x -` A) y :: extreal)" unfolding indicator_def by auto show "(\\<^isup>+ y. indicator A (x, y) \M2) = measure M2 (Pair x -` A)" unfolding * @@ -639,18 +638,21 @@ by (simp add: M1.positive_integral_cmult_indicator ac_simps) qed +lemma (in measure_space) measure_not_negative[simp,intro]: + assumes A: "A \ sets M" shows "\ A \ - \" + using positive_measure[OF A] by auto + lemma (in pair_sigma_finite) sigma_finite_up_in_pair_measure_generator: - "\F::nat \ ('a \ 'b) set. range F \ sets E \ F \ space E \ - (\i. measure (M1 \\<^isub>M M2) (F i) \ \)" + "\F::nat \ ('a \ 'b) set. range F \ sets E \ incseq F \ (\i. F i) = space E \ + (\i. measure (M1 \\<^isub>M M2) (F i) \ \)" proof - obtain F1 :: "nat \ 'a set" and F2 :: "nat \ 'b set" where - F1: "range F1 \ sets M1" "F1 \ space M1" "\i. M1.\ (F1 i) \ \" and - F2: "range F2 \ sets M2" "F2 \ space M2" "\i. M2.\ (F2 i) \ \" + F1: "range F1 \ sets M1" "incseq F1" "(\i. F1 i) = space M1" "\i. M1.\ (F1 i) \ \" and + F2: "range F2 \ sets M2" "incseq F2" "(\i. F2 i) = space M2" "\i. M2.\ (F2 i) \ \" using M1.sigma_finite_up M2.sigma_finite_up by auto - then have space: "space M1 = (\i. F1 i)" "space M2 = (\i. F2 i)" - unfolding isoton_def by auto + then have space: "space M1 = (\i. F1 i)" "space M2 = (\i. F2 i)" by auto let ?F = "\i. F1 i \ F2 i" - show ?thesis unfolding isoton_def space_pair_measure + show ?thesis unfolding space_pair_measure proof (intro exI[of _ ?F] conjI allI) show "range ?F \ sets E" using F1 F2 by (fastsimp intro!: pair_measure_generatorI) @@ -661,8 +663,8 @@ then obtain i j where "fst x \ F1 i" "snd x \ F2 j" by (auto simp: space) then have "fst x \ F1 (max i j)" "snd x \ F2 (max j i)" - using `F1 \ space M1` `F2 \ space M2` - by (auto simp: max_def dest: isoton_mono_le) + using `incseq F1` `incseq F2` unfolding incseq_def + by (force split: split_max)+ then have "(fst x, snd x) \ F1 (max i j) \ F2 (max i j)" by (intro SigmaI) (auto simp add: min_max.sup_commute) then show "x \ (\i. ?F i)" by auto @@ -670,21 +672,22 @@ then show "(\i. ?F i) = space E" using space by (auto simp: space pair_measure_generator_def) next - fix i show "F1 i \ F2 i \ F1 (Suc i) \ F2 (Suc i)" - using `F1 \ space M1` `F2 \ space M2` unfolding isoton_def - by auto + fix i show "incseq (\i. F1 i \ F2 i)" + using `incseq F1` `incseq F2` unfolding incseq_Suc_iff by auto next fix i from F1 F2 have "F1 i \ sets M1" "F2 i \ sets M2" by auto - with F1 F2 show "measure P (F1 i \ F2 i) \ \" + with F1 F2 M1.positive_measure[OF this(1)] M2.positive_measure[OF this(2)] + show "measure P (F1 i \ F2 i) \ \" by (simp add: pair_measure_times) qed qed sublocale pair_sigma_finite \ sigma_finite_measure P proof - show "measure P {} = 0" - unfolding pair_measure_def pair_measure_generator_def sigma_def by auto + show "positive P (measure P)" + unfolding pair_measure_def pair_measure_generator_def sigma_def positive_def + by (auto intro: M1.positive_integral_positive M2.positive_integral_positive) show "countably_additive P (measure P)" unfolding countably_additive_def @@ -697,20 +700,20 @@ moreover have "\x. disjoint_family (\i. Pair x -` F i)" by (intro disjoint_family_on_bisimulation[OF F(2)]) auto moreover have "\x. x \ space M1 \ range (\i. Pair x -` F i) \ sets M2" - using F by (auto intro!: measurable_cut_fst) - ultimately show "(\\<^isub>\n. measure P (F n)) = measure P (\i. F i)" - by (simp add: pair_measure_alt vimage_UN M1.positive_integral_psuminf[symmetric] + using F by auto + ultimately show "(\n. measure P (F n)) = measure P (\i. F i)" + by (simp add: pair_measure_alt vimage_UN M1.positive_integral_suminf[symmetric] M2.measure_countably_additive cong: M1.positive_integral_cong) qed from sigma_finite_up_in_pair_measure_generator guess F :: "nat \ ('a \ 'b) set" .. note F = this - show "\F::nat \ ('a \ 'b) set. range F \ sets P \ (\i. F i) = space P \ (\i. measure P (F i) \ \)" + show "\F::nat \ ('a \ 'b) set. range F \ sets P \ (\i. F i) = space P \ (\i. measure P (F i) \ \)" proof (rule exI[of _ F], intro conjI) show "range F \ sets P" using F by (auto simp: pair_measure_def) show "(\i. F i) = space P" - using F by (auto simp: pair_measure_def pair_measure_generator_def isoton_def) - show "\i. measure P (F i) \ \" using F by auto + using F by (auto simp: pair_measure_def pair_measure_generator_def) + show "\i. measure P (F i) \ \" using F by auto qed qed @@ -741,7 +744,7 @@ show "(\(y, x). (x, y)) \ measurable Q.P P" by (rule Q.pair_sigma_algebra_swap_measurable) show "sets (sigma E) = sets P" "space E = space P" "A \ sets (sigma E)" using assms unfolding pair_measure_def by auto - show "range F \ sets E" "F \ space E" "\i. \ (F i) \ \" + show "range F \ sets E" "incseq F" "(\i. F i) = space E" "\i. \ (F i) \ \" using F `A \ sets P` by (auto simp: pair_measure_def) fix X assume "X \ sets E" then obtain A B where X[simp]: "X = A \ B" and AB: "A \ sets M1" "B \ sets M2" @@ -758,8 +761,8 @@ qed lemma pair_sigma_algebra_sigma: - assumes 1: "S1 \ (space E1)" "range S1 \ sets E1" and E1: "sets E1 \ Pow (space E1)" - assumes 2: "S2 \ (space E2)" "range S2 \ sets E2" and E2: "sets E2 \ Pow (space E2)" + assumes 1: "incseq S1" "(\i. S1 i) = space E1" "range S1 \ sets E1" and E1: "sets E1 \ Pow (space E1)" + assumes 2: "decseq S2" "(\i. S2 i) = space E2" "range S2 \ sets E2" and E2: "sets E2 \ Pow (space E2)" shows "sets (sigma (pair_measure_generator (sigma E1) (sigma E2))) = sets (sigma (pair_measure_generator E1 E2))" (is "sets ?S = sets ?E") proof - @@ -771,22 +774,22 @@ using E1 E2 by (intro sigma_algebra_sigma) auto { fix A assume "A \ sets E1" then have "fst -` A \ space ?E = A \ (\i. S2 i)" - using E1 2 unfolding isoton_def pair_measure_generator_def by auto + using E1 2 unfolding pair_measure_generator_def by auto also have "\ = (\i. A \ S2 i)" by auto also have "\ \ sets ?E" unfolding pair_measure_generator_def sets_sigma using 2 `A \ sets E1` by (intro sigma_sets.Union) - (auto simp: image_subset_iff intro!: sigma_sets.Basic) + (force simp: image_subset_iff intro!: sigma_sets.Basic) finally have "fst -` A \ space ?E \ sets ?E" . } moreover { fix B assume "B \ sets E2" then have "snd -` B \ space ?E = (\i. S1 i) \ B" - using E2 1 unfolding isoton_def pair_measure_generator_def by auto + using E2 1 unfolding pair_measure_generator_def by auto also have "\ = (\i. S1 i \ B)" by auto also have "\ \ sets ?E" using 1 `B \ sets E2` unfolding pair_measure_generator_def sets_sigma by (intro sigma_sets.Union) - (auto simp: image_subset_iff intro!: sigma_sets.Basic) + (force simp: image_subset_iff intro!: sigma_sets.Basic) finally have "snd -` B \ space ?E \ sets ?E" . } ultimately have proj: "fst \ measurable ?E (sigma E1) \ snd \ measurable ?E (sigma E2)" @@ -819,12 +822,12 @@ section "Fubinis theorem" lemma (in pair_sigma_finite) simple_function_cut: - assumes f: "simple_function P f" + assumes f: "simple_function P f" "\x. 0 \ f x" shows "(\x. \\<^isup>+y. f (x, y) \M2) \ borel_measurable M1" and "(\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1) = integral\<^isup>P P f" proof - have f_borel: "f \ borel_measurable P" - using f by (rule borel_measurable_simple_function) + using f(1) by (rule borel_measurable_simple_function) let "?F z" = "f -` {z} \ space P" let "?F' x z" = "Pair x -` ?F z" { fix x assume "x \ space M1" @@ -836,15 +839,15 @@ by (intro borel_measurable_vimage measurable_cut_fst) ultimately have "simple_function M2 (\ y. f (x, y))" apply (rule_tac M2.simple_function_cong[THEN iffD2, OF _]) - apply (rule simple_function_indicator_representation[OF f]) + apply (rule simple_function_indicator_representation[OF f(1)]) using `x \ space M1` by (auto simp del: space_sigma) } note M2_sf = this { fix x assume x: "x \ space M1" then have "(\\<^isup>+y. f (x, y) \M2) = (\z\f ` space P. z * M2.\ (?F' x z))" - unfolding M2.positive_integral_eq_simple_integral[OF M2_sf[OF x]] + unfolding M2.positive_integral_eq_simple_integral[OF M2_sf[OF x] f(2)] unfolding simple_integral_def proof (safe intro!: setsum_mono_zero_cong_left) - from f show "finite (f ` space P)" by (rule simple_functionD) + from f(1) show "finite (f ` space P)" by (rule simple_functionD) next fix y assume "y \ space M2" then show "f (x, y) \ f ` space P" using `x \ space M1` by (auto simp: space_pair_measure) @@ -862,11 +865,16 @@ by (auto intro!: f_borel borel_measurable_vimage simp del: space_sigma) moreover then have "\z. (\x. M2.\ (?F' x z)) \ borel_measurable M1" by (auto intro!: measure_cut_measurable_fst simp del: vimage_Int) + moreover have *: "\i x. 0 \ M2.\ (Pair x -` (f -` {i} \ space P))" + using f(1)[THEN simple_functionD(2)] f(2) by (intro M2.positive_measure measurable_cut_fst) + moreover { fix i assume "i \ f`space P" + with * have "\x. 0 \ i * M2.\ (Pair x -` (f -` {i} \ space P))" + using f(2) by auto } ultimately show "(\x. \\<^isup>+y. f (x, y) \M2) \ borel_measurable M1" - and "(\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1) = integral\<^isup>P P f" + and "(\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1) = integral\<^isup>P P f" using f(2) by (auto simp del: vimage_Int cong: measurable_cong - intro!: M1.borel_measurable_pextreal_setsum + intro!: M1.borel_measurable_extreal_setsum setsum_cong simp add: M1.positive_integral_setsum simple_integral_def M1.positive_integral_cmult M1.positive_integral_cong[OF eq] @@ -880,42 +888,38 @@ (is "?C f \ borel_measurable M1") and "(\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1) = integral\<^isup>P P f" proof - - from borel_measurable_implies_simple_function_sequence[OF f] - obtain F where F: "\i. simple_function P (F i)" "F \ f" by auto + from borel_measurable_implies_simple_function_sequence'[OF f] guess F . note F = this then have F_borel: "\i. F i \ borel_measurable P" - and F_mono: "\i x. F i x \ F (Suc i) x" - and F_SUPR: "\x. (SUP i. F i x) = f x" - unfolding isoton_fun_expand unfolding isoton_def le_fun_def by (auto intro: borel_measurable_simple_function) - note sf = simple_function_cut[OF F(1)] + note sf = simple_function_cut[OF F(1,5)] then have "(\x. SUP i. ?C (F i) x) \ borel_measurable M1" using F(1) by auto moreover { fix x assume "x \ space M1" - have isotone: "(\ i y. F i (x, y)) \ (\y. f (x, y))" - using `F \ f` unfolding isoton_fun_expand - by (auto simp: isoton_def) - note measurable_pair_image_snd[OF F_borel`x \ space M1`] - from M2.positive_integral_isoton[OF isotone this] - have "(SUP i. ?C (F i) x) = ?C f x" - by (simp add: isoton_def) } + from F measurable_pair_image_snd[OF F_borel`x \ space M1`] + have "(\\<^isup>+y. (SUP i. F i (x, y)) \M2) = (SUP i. ?C (F i) x)" + by (intro M2.positive_integral_monotone_convergence_SUP) + (auto simp: incseq_Suc_iff le_fun_def) + then have "(SUP i. ?C (F i) x) = ?C f x" + unfolding F(4) positive_integral_max_0 by simp } note SUPR_C = this ultimately show "?C f \ borel_measurable M1" by (simp cong: measurable_cong) have "(\\<^isup>+x. (SUP i. F i x) \P) = (SUP i. integral\<^isup>P P (F i))" - using F_borel F_mono - by (auto intro!: positive_integral_monotone_convergence_SUP[symmetric]) + using F_borel F + by (intro positive_integral_monotone_convergence_SUP) auto also have "(SUP i. integral\<^isup>P P (F i)) = (SUP i. \\<^isup>+ x. (\\<^isup>+ y. F i (x, y) \M2) \M1)" unfolding sf(2) by simp - also have "\ = \\<^isup>+ x. (SUP i. \\<^isup>+ y. F i (x, y) \M2) \M1" - by (auto intro!: M1.positive_integral_monotone_convergence_SUP[OF _ sf(1)] - M2.positive_integral_mono F_mono) + also have "\ = \\<^isup>+ x. (SUP i. \\<^isup>+ y. F i (x, y) \M2) \M1" using F sf(1) + by (intro M1.positive_integral_monotone_convergence_SUP[symmetric]) + (auto intro!: M2.positive_integral_mono M2.positive_integral_positive + simp: incseq_Suc_iff le_fun_def) also have "\ = \\<^isup>+ x. (\\<^isup>+ y. (SUP i. F i (x, y)) \M2) \M1" - using F_borel F_mono - by (auto intro!: M2.positive_integral_monotone_convergence_SUP - M1.positive_integral_cong measurable_pair_image_snd) + using F_borel F(2,5) + by (auto intro!: M1.positive_integral_cong M2.positive_integral_monotone_convergence_SUP[symmetric] + simp: incseq_Suc_iff le_fun_def measurable_pair_image_snd) finally show "(\\<^isup>+ x. (\\<^isup>+ y. f (x, y) \M2) \M1) = integral\<^isup>P P f" - unfolding F_SUPR by simp + using F by (simp add: positive_integral_max_0) qed lemma (in pair_sigma_finite) measure_preserving_swap: @@ -963,8 +967,8 @@ unfolding positive_integral_fst_measurable[OF assms] .. lemma (in pair_sigma_finite) AE_pair: - assumes "almost_everywhere (\x. Q x)" - shows "M1.almost_everywhere (\x. M2.almost_everywhere (\y. Q (x, y)))" + assumes "AE x in P. Q x" + shows "AE x in M1. (AE y in M2. Q (x, y))" proof - obtain N where N: "N \ sets P" "\ N = 0" "{x\space P. \ Q x} \ N" using assms unfolding almost_everywhere_def by auto @@ -972,9 +976,9 @@ proof (rule M1.AE_I) from N measure_cut_measurable_fst[OF `N \ sets P`] show "M1.\ {x\space M1. M2.\ (Pair x -` N) \ 0} = 0" - by (simp add: M1.positive_integral_0_iff pair_measure_alt vimage_def) + by (auto simp: pair_measure_alt M1.positive_integral_0_iff) show "{x \ space M1. M2.\ (Pair x -` N) \ 0} \ sets M1" - by (intro M1.borel_measurable_pextreal_neq_const measure_cut_measurable_fst N) + by (intro M1.borel_measurable_extreal_neq_const measure_cut_measurable_fst N) { fix x assume "x \ space M1" "M2.\ (Pair x -` N) = 0" have "M2.almost_everywhere (\y. Q (x, y))" proof (rule M2.AE_I) @@ -1036,41 +1040,52 @@ shows "M1.almost_everywhere (\x. integrable M2 (\ y. f (x, y)))" (is "?AE") and "(\x. (\y. f (x, y) \M2) \M1) = integral\<^isup>L P f" (is "?INT") proof - - let "?pf x" = "Real (f x)" and "?nf x" = "Real (- f x)" + let "?pf x" = "extreal (f x)" and "?nf x" = "extreal (- f x)" have borel: "?nf \ borel_measurable P""?pf \ borel_measurable P" and - int: "integral\<^isup>P P ?nf \ \" "integral\<^isup>P P ?pf \ \" + int: "integral\<^isup>P P ?nf \ \" "integral\<^isup>P P ?pf \ \" using assms by auto - have "(\\<^isup>+x. (\\<^isup>+y. Real (f (x, y)) \M2) \M1) \ \" - "(\\<^isup>+x. (\\<^isup>+y. Real (- f (x, y)) \M2) \M1) \ \" + have "(\\<^isup>+x. (\\<^isup>+y. extreal (f (x, y)) \M2) \M1) \ \" + "(\\<^isup>+x. (\\<^isup>+y. extreal (- f (x, y)) \M2) \M1) \ \" using borel[THEN positive_integral_fst_measurable(1)] int unfolding borel[THEN positive_integral_fst_measurable(2)] by simp_all with borel[THEN positive_integral_fst_measurable(1)] - have AE: "M1.almost_everywhere (\x. (\\<^isup>+y. Real (f (x, y)) \M2) \ \)" - "M1.almost_everywhere (\x. (\\<^isup>+y. Real (- f (x, y)) \M2) \ \)" - by (auto intro!: M1.positive_integral_omega_AE) - then show ?AE using assms + have AE_pos: "AE x in M1. (\\<^isup>+y. extreal (f (x, y)) \M2) \ \" + "AE x in M1. (\\<^isup>+y. extreal (- f (x, y)) \M2) \ \" + by (auto intro!: M1.positive_integral_PInf_AE ) + then have AE: "AE x in M1. \\\<^isup>+y. extreal (f (x, y)) \M2\ \ \" + "AE x in M1. \\\<^isup>+y. extreal (- f (x, y)) \M2\ \ \" + by (auto simp: M2.positive_integral_positive) + from AE_pos show ?AE using assms by (simp add: measurable_pair_image_snd integrable_def) - { fix f assume borel: "(\x. Real (f x)) \ borel_measurable P" - and int: "integral\<^isup>P P (\x. Real (f x)) \ \" - and AE: "M1.almost_everywhere (\x. (\\<^isup>+y. Real (f (x, y)) \M2) \ \)" - have "integrable M1 (\x. real (\\<^isup>+y. Real (f (x, y)) \M2))" (is "integrable M1 ?f") + { fix f have "(\\<^isup>+ x. - \\<^isup>+ y. extreal (f x y) \M2 \M1) = (\\<^isup>+x. 0 \M1)" + using M2.positive_integral_positive + by (intro M1.positive_integral_cong_pos) (auto simp: extreal_uminus_le_reorder) + then have "(\\<^isup>+ x. - \\<^isup>+ y. extreal (f x y) \M2 \M1) = 0" by simp } + note this[simp] + { fix f assume borel: "(\x. extreal (f x)) \ borel_measurable P" + and int: "integral\<^isup>P P (\x. extreal (f x)) \ \" + and AE: "M1.almost_everywhere (\x. (\\<^isup>+y. extreal (f (x, y)) \M2) \ \)" + have "integrable M1 (\x. real (\\<^isup>+y. extreal (f (x, y)) \M2))" (is "integrable M1 ?f") proof (intro integrable_def[THEN iffD2] conjI) show "?f \ borel_measurable M1" - using borel by (auto intro!: M1.borel_measurable_real positive_integral_fst_measurable) - have "(\\<^isup>+x. Real (?f x) \M1) = (\\<^isup>+x. (\\<^isup>+y. Real (f (x, y)) \M2) \M1)" - using AE by (auto intro!: M1.positive_integral_cong_AE simp: Real_real) - then show "(\\<^isup>+x. Real (?f x) \M1) \ \" + using borel by (auto intro!: M1.borel_measurable_real_of_extreal positive_integral_fst_measurable) + have "(\\<^isup>+x. extreal (?f x) \M1) = (\\<^isup>+x. (\\<^isup>+y. extreal (f (x, y)) \M2) \M1)" + using AE M2.positive_integral_positive + by (auto intro!: M1.positive_integral_cong_AE simp: extreal_real) + then show "(\\<^isup>+x. extreal (?f x) \M1) \ \" using positive_integral_fst_measurable[OF borel] int by simp - have "(\\<^isup>+x. Real (- ?f x) \M1) = (\\<^isup>+x. 0 \M1)" - by (intro M1.positive_integral_cong) simp - then show "(\\<^isup>+x. Real (- ?f x) \M1) \ \" by simp + have "(\\<^isup>+x. extreal (- ?f x) \M1) = (\\<^isup>+x. 0 \M1)" + by (intro M1.positive_integral_cong_pos) + (simp add: M2.positive_integral_positive real_of_extreal_pos) + then show "(\\<^isup>+x. extreal (- ?f x) \M1) \ \" by simp qed } - from this[OF borel(1) int(1) AE(2)] this[OF borel(2) int(2) AE(1)] + with this[OF borel(1) int(1) AE_pos(2)] this[OF borel(2) int(2) AE_pos(1)] show ?INT unfolding lebesgue_integral_def[of P] lebesgue_integral_def[of M2] borel[THEN positive_integral_fst_measurable(2), symmetric] - using AE by (simp add: M1.integral_real) + using AE[THEN M1.integral_real] + by simp qed lemma (in pair_sigma_finite) integrable_snd_measurable: @@ -1195,7 +1210,8 @@ lemma sigma_product_algebra_sigma_eq: assumes "finite I" - assumes isotone: "\i. i \ I \ (S i) \ (space (E i))" + assumes mono: "\i. i \ I \ incseq (S i)" + assumes union: "\i. i \ I \ (\j. S i j) = space (E i)" assumes sets_into: "\i. i \ I \ range (S i) \ sets (E i)" and E: "\i. sets (E i) \ Pow (space (E i))" shows "sets (\\<^isub>M i\I. sigma (E i)) = sets (\\<^isub>M i\I. E i)" @@ -1214,13 +1230,13 @@ by (intro sigma_algebra.sigma_algebra_cong[OF sigma_algebra_sigma]) (auto dest: Pi_mem) { fix A i assume "i \ I" and A: "A \ sets (E i)" then have "(\x. x i) -` A \ space ?E = (\\<^isub>E j\I. if j = i then A else \n. S j n) \ space ?E" - using isotone unfolding isoton_def space_product_algebra + using mono union unfolding incseq_Suc_iff space_product_algebra by (auto dest: Pi_mem) also have "\ = (\n. (\\<^isub>E j\I. if j = i then A else S j n))" unfolding space_product_algebra apply simp apply (subst Pi_UN[OF `finite I`]) - using isotone[THEN isoton_mono_le] apply simp + using mono[THEN incseqD] apply simp apply (simp add: PiE_Int) apply (intro PiE_cong) using A sets_into by (auto intro!: into_space) @@ -1324,7 +1340,7 @@ obtain E where E: "A = Pi\<^isub>E (I \ J) E" "E \ (\ i\I \ J. sets (M i))" . then have *: "?f -` A \ space (?I \\<^isub>M ?J) = Pi\<^isub>E I E \ Pi\<^isub>E J E" using sets_into_space `I \ J = {}` - by (auto simp add: space_pair_measure) blast+ + by (auto simp add: space_pair_measure) fast+ then show "?f -` A \ space (?I \\<^isub>M ?J) \ sets (?I \\<^isub>M ?J)" using E unfolding * by (auto intro!: pair_measureI in_sigma product_algebra_generatorI simp: product_algebra_def) @@ -1360,6 +1376,44 @@ sublocale finite_product_sigma_finite \ finite_product_sigma_algebra by default (fact finite_index') +lemma setprod_extreal_0: + fixes f :: "'a \ extreal" + shows "(\i\A. f i) = 0 \ (finite A \ (\i\A. f i = 0))" +proof cases + assume "finite A" + then show ?thesis by (induct A) auto +qed auto + +lemma setprod_extreal_pos: + fixes f :: "'a \ extreal" assumes pos: "\i. i \ I \ 0 \ f i" shows "0 \ (\i\I. f i)" +proof cases + assume "finite I" from this pos show ?thesis by induct auto +qed simp + +lemma setprod_PInf: + assumes "\i. i \ I \ 0 \ f i" + shows "(\i\I. f i) = \ \ finite I \ (\i\I. f i = \) \ (\i\I. f i \ 0)" +proof cases + assume "finite I" from this assms show ?thesis + proof (induct I) + case (insert i I) + then have pos: "0 \ f i" "0 \ setprod f I" by (auto intro!: setprod_extreal_pos) + from insert have "(\j\insert i I. f j) = \ \ setprod f I * f i = \" by auto + also have "\ \ (setprod f I = \ \ f i = \) \ f i \ 0 \ setprod f I \ 0" + using setprod_extreal_pos[of I f] pos + by (cases rule: extreal2_cases[of "f i" "setprod f I"]) auto + also have "\ \ finite (insert i I) \ (\j\insert i I. f j = \) \ (\j\insert i I. f j \ 0)" + using insert by (auto simp: setprod_extreal_0) + finally show ?case . + qed simp +qed simp + +lemma setprod_extreal: "(\i\A. extreal (f i)) = extreal (setprod f A)" +proof cases + assume "finite A" then show ?thesis + by induct (auto simp: one_extreal_def) +qed (simp add: one_extreal_def) + lemma (in finite_product_sigma_finite) product_algebra_generator_measure: assumes "Pi\<^isub>E I F \ sets G" shows "measure G (Pi\<^isub>E I F) = (\i\I. M.\ i (F i))" @@ -1373,13 +1427,13 @@ next assume empty: "\ (\j\I. F j \ {})" then have "(\j\I. M.\ j (F j)) = 0" - by (auto simp: setprod_pextreal_0 intro!: bexI) + by (auto simp: setprod_extreal_0 intro!: bexI) moreover have "\j\I. (SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j = {}" by (rule someI2[where P="\F'. Pi\<^isub>E I F = Pi\<^isub>E I F'"]) (insert empty, auto simp: Pi_eq_empty_iff[symmetric]) then have "(\j\I. M.\ j ((SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j)) = 0" - by (auto simp: setprod_pextreal_0 intro!: bexI) + by (auto simp: setprod_extreal_0 intro!: bexI) ultimately show ?thesis unfolding product_algebra_generator_def by simp qed @@ -1387,34 +1441,32 @@ lemma (in finite_product_sigma_finite) sigma_finite_pairs: "\F::'i \ nat \ 'a set. (\i\I. range (F i) \ sets (M i)) \ - (\k. \i\I. \ i (F i k) \ \) \ - (\k. \\<^isub>E i\I. F i k) \ space G" + (\k. \i\I. \ i (F i k) \ \) \ incseq (\k. \\<^isub>E i\I. F i k) \ + (\k. \\<^isub>E i\I. F i k) = space G" proof - - have "\i::'i. \F::nat \ 'a set. range F \ sets (M i) \ F \ space (M i) \ (\k. \ i (F k) \ \)" + have "\i::'i. \F::nat \ 'a set. range F \ sets (M i) \ incseq F \ (\i. F i) = space (M i) \ (\k. \ i (F k) \ \)" using M.sigma_finite_up by simp from choice[OF this] guess F :: "'i \ nat \ 'a set" .. - then have "\i. range (F i) \ sets (M i)" "\i. F i \ space (M i)" "\i k. \ i (F i k) \ \" + then have F: "\i. range (F i) \ sets (M i)" "\i. incseq (F i)" "\i. (\j. F i j) = space (M i)" "\i k. \ i (F i k) \ \" by auto let ?F = "\k. \\<^isub>E i\I. F i k" note space_product_algebra[simp] show ?thesis - proof (intro exI[of _ F] conjI allI isotoneI set_eqI iffI ballI) + proof (intro exI[of _ F] conjI allI incseq_SucI set_eqI iffI ballI) fix i show "range (F i) \ sets (M i)" by fact next - fix i k show "\ i (F i k) \ \" by fact + fix i k show "\ i (F i k) \ \" by fact next fix A assume "A \ (\i. ?F i)" then show "A \ space G" using `\i. range (F i) \ sets (M i)` M.sets_into_space by (force simp: image_subset_iff) next fix f assume "f \ space G" - with Pi_UN[OF finite_index, of "\k i. F i k"] - `\i. F i \ space (M i)`[THEN isotonD(2)] - `\i. F i \ space (M i)`[THEN isoton_mono_le] - show "f \ (\i. ?F i)" by auto + with Pi_UN[OF finite_index, of "\k i. F i k"] F + show "f \ (\i. ?F i)" by (auto simp: incseq_def) next fix i show "?F i \ ?F (Suc i)" - using `\i. F i \ space (M i)`[THEN isotonD(1)] by auto + using `\i. incseq (F i)`[THEN incseq_SucD] by auto qed qed @@ -1438,7 +1490,7 @@ using `finite I` proof induct case empty interpret finite_product_sigma_finite M "{}" by default simp - let ?\ = "(\A. if A = {} then 0 else 1) :: 'd set \ pextreal" + let ?\ = "(\A. if A = {} then 0 else 1) :: 'd set \ extreal" show ?case proof (intro exI conjI ballI) have "sigma_algebra (sigma G \measure := ?\\)" @@ -1488,22 +1540,23 @@ proof (unfold *, default, simp) from I'.sigma_finite_pairs guess F :: "'i \ nat \ 'a set" .. then have F: "\j. j \ insert i I \ range (F j) \ sets (M j)" - "(\k. \\<^isub>E j \ insert i I. F j k) \ space I'.G" - "\k. \j. j \ insert i I \ \ j (F j k) \ \" + "incseq (\k. \\<^isub>E j \ insert i I. F j k)" + "(\k. \\<^isub>E j \ insert i I. F j k) = space I'.G" + "\k. \j. j \ insert i I \ \ j (F j k) \ \" by blast+ let "?F k" = "\\<^isub>E j \ insert i I. F j k" show "\F::nat \ ('i \ 'a) set. range F \ sets I'.P \ - (\i. F i) = (\\<^isub>E i\insert i I. space (M i)) \ (\i. ?m (F i) \ \)" + (\i. F i) = (\\<^isub>E i\insert i I. space (M i)) \ (\i. ?m (F i) \ \)" proof (intro exI[of _ ?F] conjI allI) show "range ?F \ sets I'.P" using F(1) by auto next - from F(2)[THEN isotonD(2)] - show "(\i. ?F i) = (\\<^isub>E i\insert i I. space (M i))" by simp + from F(3) show "(\i. ?F i) = (\\<^isub>E i\insert i I. space (M i))" by simp next fix j - show "?\ (?F j) \ \" - using F `finite I` - by (subst product) (auto simp: setprod_\) + have "\k. k \ insert i I \ 0 \ measure (M k) (F k j)" + using F(1) by auto + with F `finite I` setprod_PInf[of "insert i I", OF this] show "?\ (?F j) \ \" + by (subst product) auto qed qed qed @@ -1542,7 +1595,8 @@ by (simp_all add: sigma_def image_constant) lemma (in product_sigma_finite) positive_integral_empty: - "integral\<^isup>P (Pi\<^isub>M {} M) f = f (\k. undefined)" + assumes pos: "0 \ f (\k. undefined)" + shows "integral\<^isup>P (Pi\<^isub>M {} M) f = f (\k. undefined)" proof - interpret finite_product_sigma_finite M "{}" by default (fact finite.emptyI) have "\A. measure (Pi\<^isub>M {} M) (Pi\<^isub>E {} A) = 1" @@ -1550,10 +1604,10 @@ then show ?thesis unfolding positive_integral_def simple_function_def simple_integral_def_raw proof (simp add: P_empty, intro antisym) - show "f (\k. undefined) \ (SUP f:{g. g \ f}. f (\k. undefined))" - by (intro le_SUPI) auto - show "(SUP f:{g. g \ f}. f (\k. undefined)) \ f (\k. undefined)" - by (intro SUP_leI) (auto simp: le_fun_def) + show "f (\k. undefined) \ (SUP f:{g. g \ max 0 \ f}. f (\k. undefined))" + by (intro le_SUPI) (auto simp: le_fun_def split: split_max) + show "(SUP f:{g. g \ max 0 \ f}. f (\k. undefined)) \ f (\k. undefined)" using pos + by (intro SUP_leI) (auto simp: le_fun_def simp: max_def split: split_if_asm) qed qed @@ -1572,8 +1626,9 @@ let "?X S" = "?g -` S \ space P.P" from IJ.sigma_finite_pairs obtain F where F: "\i. i\ I \ J \ range (F i) \ sets (M i)" - "(\k. \\<^isub>E i\I \ J. F i k) \ space IJ.G" - "\k. \i\I\J. \ i (F i k) \ \" + "incseq (\k. \\<^isub>E i\I \ J. F i k)" + "(\k. \\<^isub>E i\I \ J. F i k) = space IJ.G" + "\k. \i\I\J. \ i (F i k) \ \" by auto let ?F = "\k. \\<^isub>E i\I \ J. F i k" show "IJ.\ A = P.\ (?X A)" @@ -1589,12 +1644,13 @@ show "range ?F \ sets IJ.G" using F by (simp add: image_subset_iff product_algebra_def product_algebra_generator_def) - show "?F \ space IJ.G " using F(2) by simp - show "\k. IJ.\ (?F k) \ \" - using `finite I` F - by (subst IJ.measure_times) (auto simp add: setprod_\) - show "?g \ measurable P.P IJ.P" - using IJ by (rule measurable_merge) + show "incseq ?F" "(\i. ?F i) = space IJ.G " by fact+ + next + fix k + have "\j. j \ I \ J \ 0 \ measure (M j) (F j k)" + using F(1) by auto + with F `finite I` setprod_PInf[of "I \ J", OF this] + show "IJ.\ (?F k) \ \" by (subst IJ.measure_times) auto next fix A assume "A \ sets IJ.G" then obtain F where A: "A = Pi\<^isub>E (I \ J) F" @@ -1611,7 +1667,7 @@ using `finite J` `finite I` F unfolding A by (intro IJ.measure_times[symmetric]) auto finally show "IJ.\ A = P.\ (?X A)" by (rule sym) - qed + qed (rule measurable_merge[OF IJ]) qed lemma (in product_sigma_finite) measure_preserving_merge: @@ -1692,8 +1748,9 @@ qed lemma (in product_sigma_finite) product_positive_integral_setprod: - fixes f :: "'i \ 'a \ pextreal" + fixes f :: "'i \ 'a \ extreal" assumes "finite I" and borel: "\i. i \ I \ f i \ borel_measurable (M i)" + and pos: "\i x. i \ I \ 0 \ f i x" shows "(\\<^isup>+ x. (\i\I. f i (x i)) \Pi\<^isub>M I M) = (\i\I. integral\<^isup>P (M i) (f i))" using assms proof induct case empty @@ -1707,12 +1764,15 @@ using insert by (auto intro!: setprod_cong) have prod: "\J. J \ insert i I \ (\x. (\i\J. f i (x i))) \ borel_measurable (Pi\<^isub>M J M)" using sets_into_space insert - by (intro sigma_algebra.borel_measurable_pextreal_setprod sigma_algebra_product_algebra + by (intro sigma_algebra.borel_measurable_extreal_setprod sigma_algebra_product_algebra measurable_comp[OF measurable_component_singleton, unfolded comp_def]) auto - show ?case - by (simp add: product_positive_integral_insert[OF insert(1,2) prod]) - (simp add: insert I.positive_integral_cmult M.positive_integral_multc * prod subset_insertI) + then show ?case + apply (simp add: product_positive_integral_insert[OF insert(1,2) prod]) + apply (simp add: insert * pos borel setprod_extreal_pos M.positive_integral_multc) + apply (subst I.positive_integral_cmult) + apply (auto simp add: pos borel insert setprod_extreal_pos positive_integral_positive) + done qed lemma (in product_sigma_finite) product_integral_singleton: @@ -1720,8 +1780,8 @@ shows "(\x. f (x i) \Pi\<^isub>M {i} M) = integral\<^isup>L (M i) f" proof - interpret I: finite_product_sigma_finite M "{i}" by default simp - have *: "(\x. Real (f x)) \ borel_measurable (M i)" - "(\x. Real (- f x)) \ borel_measurable (M i)" + have *: "(\x. extreal (f x)) \ borel_measurable (M i)" + "(\x. extreal (- f x)) \ borel_measurable (M i)" using assms by auto show ?thesis unfolding lebesgue_integral_def *[THEN product_positive_integral_singleton] .. @@ -1795,15 +1855,17 @@ proof (unfold integrable_def, intro conjI) show "(\x. abs (?f x)) \ borel_measurable P" using borel by auto - have "(\\<^isup>+x. Real (abs (?f x)) \P) = (\\<^isup>+x. (\i\I. Real (abs (f i (x i)))) \P)" - by (simp add: Real_setprod abs_setprod) - also have "\ = (\i\I. (\\<^isup>+x. Real (abs (f i x)) \M i))" + have "(\\<^isup>+x. extreal (abs (?f x)) \P) = (\\<^isup>+x. (\i\I. extreal (abs (f i (x i)))) \P)" + by (simp add: setprod_extreal abs_setprod) + also have "\ = (\i\I. (\\<^isup>+x. extreal (abs (f i x)) \M i))" using f by (subst product_positive_integral_setprod) auto - also have "\ < \" + also have "\ < \" using integrable[THEN M.integrable_abs] - unfolding pextreal_less_\ setprod_\ integrable_def by simp - finally show "(\\<^isup>+x. Real (abs (?f x)) \P) \ \" by auto - show "(\\<^isup>+x. Real (- abs (?f x)) \P) \ \" by simp + by (simp add: setprod_PInf integrable_def M.positive_integral_positive) + finally show "(\\<^isup>+x. extreal (abs (?f x)) \P) \ \" by auto + have "(\\<^isup>+x. extreal (- abs (?f x)) \P) = (\\<^isup>+x. 0 \P)" + by (intro positive_integral_cong_pos) auto + then show "(\\<^isup>+x. extreal (- abs (?f x)) \P) \ \" by simp qed ultimately show ?thesis by (rule integrable_abs_iff[THEN iffD1]) diff -r 28b51effc5ed -r cdf7693bbe08 src/HOL/Probability/Radon_Nikodym.thy --- a/src/HOL/Probability/Radon_Nikodym.thy Mon Mar 14 14:37:47 2011 +0100 +++ b/src/HOL/Probability/Radon_Nikodym.thy Mon Mar 14 14:37:49 2011 +0100 @@ -2,87 +2,81 @@ imports Lebesgue_Integration begin -lemma less_\_Ex_of_nat: "x < \ \ (\n. x < of_nat n)" -proof safe - assume "x < \" - then obtain r where "0 \ r" "x = Real r" by (cases x) auto - moreover obtain n where "r < of_nat n" using ex_less_of_nat by auto - ultimately show "\n. x < of_nat n" by (auto simp: real_eq_of_nat) -qed auto - lemma (in sigma_finite_measure) Ex_finite_integrable_function: - shows "\h\borel_measurable M. integral\<^isup>P M h \ \ \ (\x\space M. 0 < h x \ h x < \)" + shows "\h\borel_measurable M. integral\<^isup>P M h \ \ \ (\x\space M. 0 < h x \ h x < \) \ (\x. 0 \ h x)" proof - obtain A :: "nat \ 'a set" where range: "range A \ sets M" and space: "(\i. A i) = space M" and - measure: "\i. \ (A i) \ \" and + measure: "\i. \ (A i) \ \" and disjoint: "disjoint_family A" using disjoint_sigma_finite by auto let "?B i" = "2^Suc i * \ (A i)" have "\i. \x. 0 < x \ x < inverse (?B i)" proof - fix i show "\x. 0 < x \ x < inverse (?B i)" - proof cases - assume "\ (A i) = 0" - then show ?thesis by (auto intro!: exI[of _ 1]) - next - assume not_0: "\ (A i) \ 0" - then have "?B i \ \" using measure[of i] by auto - then have "inverse (?B i) \ 0" unfolding pextreal_inverse_eq_0 by simp - then show ?thesis using measure[of i] not_0 - by (auto intro!: exI[of _ "inverse (?B i) / 2"] - simp: pextreal_noteq_omega_Ex zero_le_mult_iff zero_less_mult_iff mult_le_0_iff power_le_zero_eq) - qed + fix i have Ai: "A i \ sets M" using range by auto + from measure positive_measure[OF this] + show "\x. 0 < x \ x < inverse (?B i)" + by (auto intro!: extreal_dense simp: extreal_0_gt_inverse) qed from choice[OF this] obtain n where n: "\i. 0 < n i" "\i. n i < inverse (2^Suc i * \ (A i))" by auto - let "?h x" = "\\<^isub>\ i. n i * indicator (A i) x" + { fix i have "0 \ n i" using n(1)[of i] by auto } note pos = this + let "?h x" = "\i. n i * indicator (A i) x" show ?thesis proof (safe intro!: bexI[of _ ?h] del: notI) have "\i. A i \ sets M" using range by fastsimp+ - then have "integral\<^isup>P M ?h = (\\<^isub>\ i. n i * \ (A i))" - by (simp add: positive_integral_psuminf positive_integral_cmult_indicator) - also have "\ \ (\\<^isub>\ i. Real ((1 / 2)^Suc i))" - proof (rule psuminf_le) - fix N show "n N * \ (A N) \ Real ((1 / 2) ^ Suc N)" + then have "integral\<^isup>P M ?h = (\i. n i * \ (A i))" using pos + by (simp add: positive_integral_suminf positive_integral_cmult_indicator) + also have "\ \ (\i. (1 / 2)^Suc i)" + proof (rule suminf_le_pos) + fix N + have "n N * \ (A N) \ inverse (2^Suc N * \ (A N)) * \ (A N)" + using positive_measure[OF `A N \ sets M`] n[of N] + by (intro extreal_mult_right_mono) auto + also have "\ \ (1 / 2) ^ Suc N" using measure[of N] n[of N] - by (cases "n N") - (auto simp: pextreal_noteq_omega_Ex field_simps zero_le_mult_iff - mult_le_0_iff mult_less_0_iff power_less_zero_eq - power_le_zero_eq inverse_eq_divide less_divide_eq - power_divide split: split_if_asm) + by (cases rule: extreal2_cases[of "n N" "\ (A N)"]) + (simp_all add: inverse_eq_divide power_divide one_extreal_def extreal_power_divide) + finally show "n N * \ (A N) \ (1 / 2) ^ Suc N" . + show "0 \ n N * \ (A N)" using n[of N] `A N \ sets M` by simp qed - also have "\ = Real 1" - by (rule suminf_imp_psuminf, rule power_half_series, auto) - finally show "integral\<^isup>P M ?h \ \" by auto + finally show "integral\<^isup>P M ?h \ \" unfolding suminf_half_series_extreal by auto next - fix x assume "x \ space M" - then obtain i where "x \ A i" using space[symmetric] by auto - from psuminf_cmult_indicator[OF disjoint, OF this] - have "?h x = n i" by simp - then show "0 < ?h x" and "?h x < \" using n[of i] by auto + { fix x assume "x \ space M" + then obtain i where "x \ A i" using space[symmetric] by auto + with disjoint n have "?h x = n i" + by (auto intro!: suminf_cmult_indicator intro: less_imp_le) + then show "0 < ?h x" and "?h x < \" using n[of i] by auto } + note pos = this + fix x show "0 \ ?h x" + proof cases + assume "x \ space M" then show "0 \ ?h x" using pos by (auto intro: less_imp_le) + next + assume "x \ space M" then have "\i. x \ A i" using space by auto + then show "0 \ ?h x" by auto + qed next - show "?h \ borel_measurable M" using range - by (auto intro!: borel_measurable_psuminf borel_measurable_pextreal_times) + show "?h \ borel_measurable M" using range n + by (auto intro!: borel_measurable_psuminf borel_measurable_extreal_times extreal_0_le_mult intro: less_imp_le) qed qed subsection "Absolutely continuous" definition (in measure_space) - "absolutely_continuous \ = (\N\null_sets. \ N = (0 :: pextreal))" + "absolutely_continuous \ = (\N\null_sets. \ N = (0 :: extreal))" lemma (in measure_space) absolutely_continuous_AE: assumes "measure_space M'" and [simp]: "sets M' = sets M" "space M' = space M" and "absolutely_continuous (measure M')" "AE x. P x" - shows "measure_space.almost_everywhere M' P" + shows "AE x in M'. P x" proof - interpret \: measure_space M' by fact from `AE x. P x` obtain N where N: "N \ null_sets" and "{x\space M. \ P x} \ N" unfolding almost_everywhere_def by auto - show "\.almost_everywhere P" + show "AE x in M'. P x" proof (rule \.AE_I') show "{x\space M'. \ P x} \ N" by simp fact from `absolutely_continuous (measure M')` show "N \ \.null_sets" @@ -99,7 +93,7 @@ interpret v: finite_measure_space ?\ by fact have "\ N = measure ?\ (\x\N. {x})" by simp also have "\ = (\x\N. measure ?\ {x})" - proof (rule v.measure_finitely_additive''[symmetric]) + proof (rule v.measure_setsum[symmetric]) show "finite N" using `N \ space M` finite_space by (auto intro: finite_subset) show "disjoint_family_on (\i. {i}) N" unfolding disjoint_family_on_def by auto fix x assume "x \ N" thus "{x} \ sets ?\" using `N \ space M` sets_eq_Pow by auto @@ -107,8 +101,10 @@ also have "\ = 0" proof (safe intro!: setsum_0') fix x assume "x \ N" - hence "\ {x} \ \ N" using sets_eq_Pow `N \ space M` by (auto intro!: measure_mono) - hence "\ {x} = 0" using `\ N = 0` by simp + hence "\ {x} \ \ N" "0 \ \ {x}" + using sets_eq_Pow `N \ space M` positive_measure[of "{x}"] + by (auto intro!: measure_mono) + then have "\ {x} = 0" using `\ N = 0` by simp thus "measure ?\ {x} = 0" using v[of x] `x \ N` `N \ space M` by auto qed finally show "\ N = 0" by simp @@ -125,12 +121,12 @@ lemma (in finite_measure) Radon_Nikodym_aux_epsilon: fixes e :: real assumes "0 < e" assumes "finite_measure (M\measure := \\)" - shows "\A\sets M. real (\ (space M)) - real (\ (space M)) \ - real (\ A) - real (\ A) \ - (\B\sets M. B \ A \ - e < real (\ B) - real (\ B))" + shows "\A\sets M. \' (space M) - finite_measure.\' (M\measure := \\) (space M) \ + \' A - finite_measure.\' (M\measure := \\) A \ + (\B\sets M. B \ A \ - e < \' B - finite_measure.\' (M\measure := \\) B)" proof - - let "?d A" = "real (\ A) - real (\ A)" interpret M': finite_measure "M\measure := \\" by fact + let "?d A" = "\' A - M'.\' A" let "?A A" = "if (\B\sets M. B \ space M - A \ -e < ?d B) then {} else (SOME B. B \ sets M \ B \ space M - A \ ?d B \ -e)" @@ -157,7 +153,7 @@ fix B assume "B \ sets M \ B \ space M - A n \ ?d B \ - e" hence "A n \ B = {}" "B \ sets M" and dB: "?d B \ -e" by auto hence "?d (A n \ B) = ?d (A n) + ?d B" - using `A n \ sets M` real_finite_measure_Union M'.real_finite_measure_Union by simp + using `A n \ sets M` finite_measure_Union M'.finite_measure_Union by simp also have "\ \ ?d (A n) - e" using dB by simp finally show "?d (A n \ B) \ ?d (A n) - e" . qed } @@ -186,11 +182,7 @@ fix n assume "?d (space M) \ ?d (space M - A n)" also have "\ \ ?d (space M - A (Suc n))" using A_in_sets sets_into_space dA_mono[of n] - real_finite_measure_Diff[of "space M"] - real_finite_measure_Diff[of "space M"] - M'.real_finite_measure_Diff[of "space M"] - M'.real_finite_measure_Diff[of "space M"] - by (simp del: A_simps) + by (simp del: A_simps add: finite_measure_Diff M'.finite_measure_Diff) finally show "?d (space M) \ ?d (space M - A (Suc n))" . qed simp qed @@ -200,13 +192,16 @@ { fix n have "?d (A n) \ - real n * e" proof (induct n) case (Suc n) with dA_epsilon[of n, OF B] show ?case by (simp del: A_simps add: real_of_nat_Suc field_simps) - qed simp } note dA_less = this + next + case 0 with M'.empty_measure show ?case by (simp add: zero_extreal_def) + qed } note dA_less = this have decseq: "decseq (\n. ?d (A n))" unfolding decseq_eq_incseq proof (rule incseq_SucI) fix n show "- ?d (A n) \ - ?d (A (Suc n))" using dA_mono[of n] by auto qed - from real_finite_continuity_from_below[of A] `range A \ sets M` - M'.real_finite_continuity_from_below[of A] + have A: "incseq A" by (auto intro!: incseq_SucI) + from finite_continuity_from_below[OF _ A] `range A \ sets M` + M'.finite_continuity_from_below[OF _ A] have convergent: "(\i. ?d (A i)) ----> ?d (\i. A i)" by (auto intro!: LIMSEQ_diff) obtain n :: nat where "- ?d (\i. A i) / e < real n" using reals_Archimedean2 by auto @@ -216,33 +211,55 @@ qed qed +lemma (in finite_measure) restricted_measure_subset: + assumes S: "S \ sets M" and X: "X \ S" + shows "finite_measure.\' (restricted_space S) X = \' X" +proof cases + note r = restricted_finite_measure[OF S] + { assume "X \ sets M" with S X show ?thesis + unfolding finite_measure.\'_def[OF r] \'_def by auto } + { assume "X \ sets M" + moreover then have "S \ X \ sets M" + using X by (simp add: Int_absorb1) + ultimately show ?thesis + unfolding finite_measure.\'_def[OF r] \'_def using S by auto } +qed + +lemma (in finite_measure) restricted_measure: + assumes X: "S \ sets M" "X \ sets (restricted_space S)" + shows "finite_measure.\' (restricted_space S) X = \' X" +proof - + from X have "S \ sets M" "X \ S" by auto + from restricted_measure_subset[OF this] show ?thesis . +qed + lemma (in finite_measure) Radon_Nikodym_aux: assumes "finite_measure (M\measure := \\)" (is "finite_measure ?M'") - shows "\A\sets M. real (\ (space M)) - real (\ (space M)) \ - real (\ A) - real (\ A) \ - (\B\sets M. B \ A \ 0 \ real (\ B) - real (\ B))" + shows "\A\sets M. \' (space M) - finite_measure.\' (M\measure := \\) (space M) \ + \' A - finite_measure.\' (M\measure := \\) A \ + (\B\sets M. B \ A \ 0 \ \' B - finite_measure.\' (M\measure := \\) B)" proof - - let "?d A" = "real (\ A) - real (\ A)" - let "?P A B n" = "A \ sets M \ A \ B \ ?d B \ ?d A \ (\C\sets M. C \ A \ - 1 / real (Suc n) < ?d C)" interpret M': finite_measure ?M' where "space ?M' = space M" and "sets ?M' = sets M" and "measure ?M' = \" by fact auto + let "?d A" = "\' A - M'.\' A" + let "?P A B n" = "A \ sets M \ A \ B \ ?d B \ ?d A \ (\C\sets M. C \ A \ - 1 / real (Suc n) < ?d C)" let "?r S" = "restricted_space S" - { fix S n - assume S: "S \ sets M" - hence **: "\X. X \ op \ S ` sets M \ X \ sets M \ X \ S" by auto - have [simp]: "(restricted_space S\measure := \\) = M'.restricted_space S" - by (cases M) simp - from M'.restricted_finite_measure[of S] restricted_finite_measure[OF S] S - have "finite_measure (?r S)" "0 < 1 / real (Suc n)" + { fix S n assume S: "S \ sets M" + note r = M'.restricted_finite_measure[of S] restricted_finite_measure[OF S] S + then have "finite_measure (?r S)" "0 < 1 / real (Suc n)" "finite_measure (?r S\measure := \\)" by auto - from finite_measure.Radon_Nikodym_aux_epsilon[OF this] guess X .. - hence "?P X S n" - proof (simp add: **, safe) - fix C assume C: "C \ sets M" "C \ X" "X \ S" and - *: "\B\sets M. S \ B \ X \ - (1 / real (Suc n)) < ?d (S \ B)" - hence "C \ S" "C \ X" "S \ C = C" by auto - with *[THEN bspec, OF `C \ sets M`] - show "- (1 / real (Suc n)) < ?d C" by auto + from finite_measure.Radon_Nikodym_aux_epsilon[OF this] guess X .. note X = this + have "?P X S n" + proof (intro conjI ballI impI) + show "X \ sets M" "X \ S" using X(1) `S \ sets M` by auto + have "S \ op \ S ` sets M" using `S \ sets M` by auto + then show "?d S \ ?d X" + using S X restricted_measure[OF S] M'.restricted_measure[OF S] by simp + fix C assume "C \ sets M" "C \ X" + then have "C \ sets (restricted_space S)" "C \ X" + using `S \ sets M` `X \ S` by auto + with X(2) show "- 1 / real (Suc n) < ?d C" + using S X restricted_measure[OF S] M'.restricted_measure[OF S] by auto qed hence "\A. ?P A S n" by auto } note Ex_P = this @@ -268,10 +285,11 @@ show ?thesis proof (safe intro!: bexI[of _ "\i. A i"]) show "(\i. A i) \ sets M" using A_in_sets by auto - from `range A \ sets M` A_mono - real_finite_continuity_from_above[of A] - M'.real_finite_continuity_from_above[of A] - have "(\i. ?d (A i)) ----> ?d (\i. A i)" by (auto intro!: LIMSEQ_diff) + have A: "decseq A" using A_mono by (auto intro!: decseq_SucI) + from + finite_continuity_from_above[OF `range A \ sets M` A] + M'.finite_continuity_from_above[OF `range A \ sets M` A] + have "(\i. ?d (A i)) ----> ?d (\i. A i)" by (intro LIMSEQ_diff) thus "?d (space M) \ ?d (\i. A i)" using mono_dA[THEN monoD, of 0 _] by (rule_tac LIMSEQ_le_const) (auto intro!: exI) next @@ -290,6 +308,10 @@ qed qed +lemma (in finite_measure) real_measure: + assumes A: "A \ sets M" shows "\r. 0 \ r \ \ A = extreal r" + using finite_measure[OF A] positive_measure[OF A] by (cases "\ A") auto + lemma (in finite_measure) Radon_Nikodym_finite_measure: assumes "finite_measure (M\ measure := \\)" (is "finite_measure ?M'") assumes "absolutely_continuous \" @@ -298,7 +320,7 @@ interpret M': finite_measure ?M' where "space ?M' = space M" and "sets ?M' = sets M" and "measure ?M' = \" using assms(1) by auto - def G \ "{g \ borel_measurable M. \A\sets M. (\\<^isup>+x. g x * indicator A x \M) \ \ A}" + def G \ "{g \ borel_measurable M. (\x. 0 \ g x) \ (\A\sets M. (\\<^isup>+x. g x * indicator A x \M) \ \ A)}" have "(\x. 0) \ G" unfolding G_def by auto hence "G \ {}" by auto { fix f g assume f: "f \ G" and g: "g \ G" @@ -324,24 +346,28 @@ also have "\ = \ A" using M'.measure_additive[OF sets] union by auto finally show "(\\<^isup>+x. max (g x) (f x) * indicator A x \M) \ \ A" . + next + fix x show "0 \ max (g x) (f x)" using f g by (auto simp: G_def split: split_max) qed } note max_in_G = this - { fix f g assume "f \ g" and f: "\i. f i \ G" - have "g \ G" unfolding G_def + { fix f assume "incseq f" and f: "\i. f i \ G" + have "(\x. SUP i. f i x) \ G" unfolding G_def proof safe - from `f \ g` have [simp]: "g = (\x. SUP i. f i x)" - unfolding isoton_def fun_eq_iff SUPR_apply by simp - have f_borel: "\i. f i \ borel_measurable M" using f unfolding G_def by simp - thus "g \ borel_measurable M" by auto + show "(\x. SUP i. f i x) \ borel_measurable M" + using f by (auto simp: G_def) + { fix x show "0 \ (SUP i. f i x)" + using f by (auto simp: G_def intro: le_SUPI2) } + next fix A assume "A \ sets M" - hence "\i. (\x. f i x * indicator A x) \ borel_measurable M" - using f_borel by (auto intro!: borel_measurable_indicator) - from positive_integral_isoton[OF isoton_indicator[OF `f \ g`] this] - have SUP: "(\\<^isup>+x. g x * indicator A x \M) = - (SUP i. (\\<^isup>+x. f i x * indicator A x \M))" - unfolding isoton_def by simp - show "(\\<^isup>+x. g x * indicator A x \M) \ \ A" unfolding SUP - using f `A \ sets M` unfolding G_def by (auto intro!: SUP_leI) + have "(\\<^isup>+x. (SUP i. f i x) * indicator A x \M) = + (\\<^isup>+x. (SUP i. f i x * indicator A x) \M)" + by (intro positive_integral_cong) (simp split: split_indicator) + also have "\ = (SUP i. (\\<^isup>+x. f i x * indicator A x \M))" + using `incseq f` f `A \ sets M` + by (intro positive_integral_monotone_convergence_SUP) + (auto simp: G_def incseq_Suc_iff le_fun_def split: split_indicator) + finally show "(\\<^isup>+x. (SUP i. f i x) * indicator A x \M) \ \ A" + using f `A \ sets M` by (auto intro!: SUP_leI simp: G_def) qed } note SUP_in_G = this let ?y = "SUP g : G. integral\<^isup>P M g" @@ -351,9 +377,8 @@ from this[THEN bspec, OF top] show "integral\<^isup>P M g \ \ (space M)" by (simp cong: positive_integral_cong) qed - hence "?y \ \" using M'.finite_measure_of_space by auto - from SUPR_countable_SUPR[OF this `G \ {}`] guess ys .. note ys = this - hence "\n. \g. g\G \ integral\<^isup>P M g = ys n" + from SUPR_countable_SUPR[OF `G \ {}`, of "integral\<^isup>P M"] guess ys .. note ys = this + then have "\n. \g. g\G \ integral\<^isup>P M g = ys n" proof safe fix n assume "range ys \ integral\<^isup>P M ` G" hence "ys n \ integral\<^isup>P M ` G" by auto @@ -362,8 +387,9 @@ from choice[OF this] obtain gs where "\i. gs i \ G" "\n. integral\<^isup>P M (gs n) = ys n" by auto hence y_eq: "?y = (SUP i. integral\<^isup>P M (gs i))" using ys by auto let "?g i x" = "Max ((\n. gs n x) ` {..i})" - def f \ "SUP i. ?g i" - have gs_not_empty: "\i. (\n. gs n x) ` {..i} \ {}" by auto + def f \ "\x. SUP i. ?g i x" + let "?F A x" = "f x * indicator A x" + have gs_not_empty: "\i x. (\n. gs n x) ` {..i} \ {}" by auto { fix i have "?g i \ G" proof (induct i) case 0 thus ?case by simp fact @@ -373,15 +399,13 @@ by (auto simp add: atMost_Suc intro!: max_in_G) qed } note g_in_G = this - have "\x. \i. ?g i x \ ?g (Suc i) x" - using gs_not_empty by (simp add: atMost_Suc) - hence isoton_g: "?g \ f" by (simp add: isoton_def le_fun_def f_def) - from SUP_in_G[OF this g_in_G] have "f \ G" . - hence [simp, intro]: "f \ borel_measurable M" unfolding G_def by auto - have "(\i. integral\<^isup>P M (?g i)) \ integral\<^isup>P M f" - using isoton_g g_in_G by (auto intro!: positive_integral_isoton simp: G_def f_def) - hence "integral\<^isup>P M f = (SUP i. integral\<^isup>P M (?g i))" - unfolding isoton_def by simp + have "incseq ?g" using gs_not_empty + by (auto intro!: incseq_SucI le_funI simp add: atMost_Suc) + from SUP_in_G[OF this g_in_G] have "f \ G" unfolding f_def . + then have [simp, intro]: "f \ borel_measurable M" unfolding G_def by auto + have "integral\<^isup>P M f = (SUP i. integral\<^isup>P M (?g i))" unfolding f_def + using g_in_G `incseq ?g` + by (auto intro!: positive_integral_monotone_convergence_SUP simp: G_def) also have "\ = ?y" proof (rule antisym) show "(SUP i. integral\<^isup>P M (?g i)) \ ?y" @@ -390,42 +414,57 @@ by (auto intro!: SUP_mono positive_integral_mono Max_ge) qed finally have int_f_eq_y: "integral\<^isup>P M f = ?y" . - let "?t A" = "\ A - (\\<^isup>+x. f x * indicator A x \M)" + have "\x. 0 \ f x" + unfolding f_def using `\i. gs i \ G` + by (auto intro!: le_SUPI2 Max_ge_iff[THEN iffD2] simp: G_def) + let "?t A" = "\ A - (\\<^isup>+x. ?F A x \M)" let ?M = "M\ measure := ?t\" interpret M: sigma_algebra ?M by (intro sigma_algebra_cong) auto + have f_le_\: "\A. A \ sets M \ (\\<^isup>+x. ?F A x \M) \ \ A" + using `f \ G` unfolding G_def by auto have fmM: "finite_measure ?M" - proof (default, simp_all add: countably_additive_def, safe) + proof (default, simp_all add: countably_additive_def positive_def, safe del: notI) fix A :: "nat \ 'a set" assume A: "range A \ sets M" "disjoint_family A" - have "(\\<^isub>\ n. (\\<^isup>+x. f x * indicator (A n) x \M)) - = (\\<^isup>+x. (\\<^isub>\n. f x * indicator (A n) x) \M)" - using `range A \ sets M` - by (rule_tac positive_integral_psuminf[symmetric]) (auto intro!: borel_measurable_indicator) - also have "\ = (\\<^isup>+x. f x * indicator (\n. A n) x \M)" - apply (rule positive_integral_cong) - apply (subst psuminf_cmult_right) - unfolding psuminf_indicator[OF `disjoint_family A`] .. - finally have "(\\<^isub>\ n. (\\<^isup>+x. f x * indicator (A n) x \M)) - = (\\<^isup>+x. f x * indicator (\n. A n) x \M)" . - moreover have "(\\<^isub>\n. \ (A n)) = \ (\n. A n)" + have "(\n. (\\<^isup>+x. ?F (A n) x \M)) = (\\<^isup>+x. (\n. ?F (A n) x) \M)" + using `range A \ sets M` `\x. 0 \ f x` + by (intro positive_integral_suminf[symmetric]) auto + also have "\ = (\\<^isup>+x. ?F (\n. A n) x \M)" + using `\x. 0 \ f x` + by (intro positive_integral_cong) (simp add: suminf_cmult_extreal suminf_indicator[OF `disjoint_family A`]) + finally have "(\n. (\\<^isup>+x. ?F (A n) x \M)) = (\\<^isup>+x. ?F (\n. A n) x \M)" . + moreover have "(\n. \ (A n)) = \ (\n. A n)" using M'.measure_countably_additive A by (simp add: comp_def) - moreover have "\i. (\\<^isup>+x. f x * indicator (A i) x \M) \ \ (A i)" - using A `f \ G` unfolding G_def by auto - moreover have v_fin: "\ (\i. A i) \ \" using M'.finite_measure A by (simp add: countable_UN) + moreover have v_fin: "\ (\i. A i) \ \" using M'.finite_measure A by (simp add: countable_UN) moreover { - have "(\\<^isup>+x. f x * indicator (\i. A i) x \M) \ \ (\i. A i)" + have "(\\<^isup>+x. ?F (\i. A i) x \M) \ \ (\i. A i)" using A `f \ G` unfolding G_def by (auto simp: countable_UN) - also have "\ (\i. A i) < \" using v_fin by (simp add: pextreal_less_\) - finally have "(\\<^isup>+x. f x * indicator (\i. A i) x \M) \ \" - by (simp add: pextreal_less_\) } + also have "\ (\i. A i) < \" using v_fin by simp + finally have "(\\<^isup>+x. ?F (\i. A i) x \M) \ \" by simp } + moreover have "\i. (\\<^isup>+x. ?F (A i) x \M) \ \ (A i)" + using A by (intro f_le_\) auto ultimately - show "(\\<^isub>\ n. ?t (A n)) = ?t (\i. A i)" - apply (subst psuminf_minus) by simp_all + show "(\n. ?t (A n)) = ?t (\i. A i)" + by (subst suminf_extreal_minus) (simp_all add: positive_integral_positive) + next + fix A assume A: "A \ sets M" show "0 \ \ A - \\<^isup>+ x. ?F A x \M" + using f_le_\[OF A] `f \ G` M'.finite_measure[OF A] by (auto simp: G_def extreal_le_minus_iff) + next + show "\ (space M) - (\\<^isup>+ x. ?F (space M) x \M) \ \" (is "?a - ?b \ _") + using positive_integral_positive[of "?F (space M)"] + by (cases rule: extreal2_cases[of ?a ?b]) auto qed then interpret M: finite_measure ?M where "space ?M = space M" and "sets ?M = sets M" and "measure ?M = ?t" by (simp_all add: fmM) - have ac: "absolutely_continuous ?t" using `absolutely_continuous \` unfolding absolutely_continuous_def by auto + have ac: "absolutely_continuous ?t" unfolding absolutely_continuous_def + proof + fix N assume N: "N \ null_sets" + with `absolutely_continuous \` have "\ N = 0" unfolding absolutely_continuous_def by auto + moreover with N have "(\\<^isup>+ x. ?F N x \M) \ \ N" using `f \ G` by (auto simp: G_def) + ultimately show "\ N - (\\<^isup>+ x. ?F N x \M) = 0" + using positive_integral_positive by (auto intro!: antisym) + qed have upper_bound: "\A\sets M. ?t A \ 0" proof (rule ccontr) assume "\ ?thesis" @@ -436,43 +475,54 @@ using M.measure_mono[of A "space M"] A sets_into_space by simp finally have pos_t: "0 < ?t (space M)" by simp moreover - hence pos_M: "0 < \ (space M)" - using ac top unfolding absolutely_continuous_def by auto + then have "\ (space M) \ 0" + using ac unfolding absolutely_continuous_def by auto + then have pos_M: "0 < \ (space M)" + using positive_measure[OF top] by (simp add: le_less) moreover have "(\\<^isup>+x. f x * indicator (space M) x \M) \ \ (space M)" using `f \ G` unfolding G_def by auto - hence "(\\<^isup>+x. f x * indicator (space M) x \M) \ \" + hence "(\\<^isup>+x. f x * indicator (space M) x \M) \ \" using M'.finite_measure_of_space by auto moreover def b \ "?t (space M) / \ (space M) / 2" - ultimately have b: "b \ 0" "b \ \" - using M'.finite_measure_of_space - by (auto simp: pextreal_inverse_eq_0 finite_measure_of_space) + ultimately have b: "b \ 0 \ 0 \ b \ b \ \" + using M'.finite_measure_of_space positive_integral_positive[of "?F (space M)"] + by (cases rule: extreal3_cases[of "integral\<^isup>P M (?F (space M))" "\ (space M)" "\ (space M)"]) + (simp_all add: field_simps) + then have b: "b \ 0" "0 \ b" "0 < b" "b \ \" by auto let ?Mb = "?M\measure := \A. b * \ A\" interpret b: sigma_algebra ?Mb by (intro sigma_algebra_cong) auto - have "finite_measure ?Mb" - by default - (insert finite_measure_of_space b measure_countably_additive, - auto simp: psuminf_cmult_right countably_additive_def) + have Mb: "finite_measure ?Mb" + proof + show "positive ?Mb (measure ?Mb)" + using `0 \ b` by (auto simp: positive_def) + show "countably_additive ?Mb (measure ?Mb)" + using `0 \ b` measure_countably_additive + by (auto simp: countably_additive_def suminf_cmult_extreal subset_eq) + show "measure ?Mb (space ?Mb) \ \" + using b by auto + qed from M.Radon_Nikodym_aux[OF this] obtain A0 where "A0 \ sets M" and space_less_A0: "real (?t (space M)) - real (b * \ (space M)) \ real (?t A0) - real (b * \ A0)" and - *: "\B. \ B \ sets M ; B \ A0 \ \ 0 \ real (?t B) - real (b * \ B)" by auto - { fix B assume "B \ sets M" "B \ A0" + *: "\B. \ B \ sets M ; B \ A0 \ \ 0 \ real (?t B) - real (b * \ B)" + unfolding M.\'_def finite_measure.\'_def[OF Mb] by auto + { fix B assume B: "B \ sets M" "B \ A0" with *[OF this] have "b * \ B \ ?t B" - using M'.finite_measure b finite_measure - by (cases "b * \ B", cases "?t B") (auto simp: field_simps) } + using M'.finite_measure b finite_measure M.positive_measure[OF B(1)] + by (cases rule: extreal2_cases[of "?t B" "b * \ B"]) auto } note bM_le_t = this let "?f0 x" = "f x + b * indicator A0 x" { fix A assume A: "A \ sets M" hence "A \ A0 \ sets M" using `A0 \ sets M` by auto have "(\\<^isup>+x. ?f0 x * indicator A x \M) = (\\<^isup>+x. f x * indicator A x + b * indicator (A \ A0) x \M)" - by (auto intro!: positive_integral_cong simp: field_simps indicator_inter_arith) + by (auto intro!: positive_integral_cong split: split_indicator) hence "(\\<^isup>+x. ?f0 x * indicator A x \M) = (\\<^isup>+x. f x * indicator A x \M) + b * \ (A \ A0)" - using `A0 \ sets M` `A \ A0 \ sets M` A - by (simp add: borel_measurable_indicator positive_integral_add positive_integral_cmult_indicator) } + using `A0 \ sets M` `A \ A0 \ sets M` A b `f \ G` + by (simp add: G_def positive_integral_add positive_integral_cmult_indicator) } note f0_eq = this { fix A assume A: "A \ sets M" hence "A \ A0 \ sets M" using `A0 \ sets M` by auto @@ -487,39 +537,57 @@ using M.measure_mono[simplified, OF _ `A \ A0 \ sets M` `A \ sets M`] by (auto intro!: add_left_mono) also have "\ \ \ A" - using f_le_v M'.finite_measure[simplified, OF `A \ sets M`] - by (cases "(\\<^isup>+x. f x * indicator A x \M)", cases "\ A", auto) + using f_le_v M'.finite_measure[simplified, OF `A \ sets M`] positive_integral_positive[of "?F A"] + by (cases "\\<^isup>+x. ?F A x \M", cases "\ A") auto finally have "(\\<^isup>+x. ?f0 x * indicator A x \M) \ \ A" . } - hence "?f0 \ G" using `A0 \ sets M` unfolding G_def - by (auto intro!: borel_measurable_indicator borel_measurable_pextreal_add borel_measurable_pextreal_times) - have real: "?t (space M) \ \" "?t A0 \ \" - "b * \ (space M) \ \" "b * \ A0 \ \" + hence "?f0 \ G" using `A0 \ sets M` b `f \ G` unfolding G_def + by (auto intro!: borel_measurable_indicator borel_measurable_extreal_add + borel_measurable_extreal_times extreal_add_nonneg_nonneg) + have real: "?t (space M) \ \" "?t A0 \ \" + "b * \ (space M) \ \" "b * \ A0 \ \" using `A0 \ sets M` b finite_measure[of A0] M.finite_measure[of A0] finite_measure_of_space M.finite_measure_of_space by auto - have int_f_finite: "integral\<^isup>P M f \ \" - using M'.finite_measure_of_space pos_t unfolding pextreal_zero_less_diff_iff + have int_f_finite: "integral\<^isup>P M f \ \" + using M'.finite_measure_of_space pos_t unfolding extreal_less_minus_iff by (auto cong: positive_integral_cong) - have "?t (space M) > b * \ (space M)" unfolding b_def - apply (simp add: field_simps) - apply (subst mult_assoc[symmetric]) - apply (subst pextreal_mult_inverse) + have "0 < ?t (space M) - b * \ (space M)" unfolding b_def using finite_measure_of_space M'.finite_measure_of_space pos_t pos_M - using pextreal_mult_strict_right_mono[of "Real (1/2)" 1 "?t (space M)"] - by simp_all - hence "0 < ?t (space M) - b * \ (space M)" - by (simp add: pextreal_zero_less_diff_iff) + using positive_integral_positive[of "?F (space M)"] + by (cases rule: extreal3_cases[of "\ (space M)" "\ (space M)" "integral\<^isup>P M (?F (space M))"]) + (auto simp: field_simps mult_less_cancel_left) also have "\ \ ?t A0 - b * \ A0" - using space_less_A0 pos_M pos_t b real[unfolded pextreal_noteq_omega_Ex] by auto - finally have "b * \ A0 < ?t A0" unfolding pextreal_zero_less_diff_iff . - hence "0 < ?t A0" by auto - hence "0 < \ A0" using ac unfolding absolutely_continuous_def + using space_less_A0 b + using + `A0 \ sets M`[THEN M.real_measure] + top[THEN M.real_measure] + apply safe + apply simp + using + `A0 \ sets M`[THEN real_measure] + `A0 \ sets M`[THEN M'.real_measure] + top[THEN real_measure] + top[THEN M'.real_measure] + by (cases b) auto + finally have 1: "b * \ A0 < ?t A0" + using + `A0 \ sets M`[THEN M.real_measure] + apply safe + apply simp + using + `A0 \ sets M`[THEN real_measure] + `A0 \ sets M`[THEN M'.real_measure] + by (cases b) auto + have "0 < ?t A0" + using b `A0 \ sets M` by (auto intro!: le_less_trans[OF _ 1]) + then have "\ A0 \ 0" using ac unfolding absolutely_continuous_def using `A0 \ sets M` by auto - hence "0 < b * \ A0" using b by auto - from int_f_finite this - have "?y + 0 < integral\<^isup>P M f + b * \ A0" unfolding int_f_eq_y - by (rule pextreal_less_add) + then have "0 < \ A0" using positive_measure[OF `A0 \ sets M`] by auto + hence "0 < b * \ A0" using b by (auto simp: extreal_zero_less_0_iff) + with int_f_finite have "?y + 0 < integral\<^isup>P M f + b * \ A0" unfolding int_f_eq_y + using `f \ G` + by (intro extreal_add_strict_mono) (auto intro!: le_SUPI2 positive_integral_positive) also have "\ = integral\<^isup>P M ?f0" using f0_eq[OF top] `A0 \ sets M` sets_into_space by (simp cong: positive_integral_cong) finally have "?y < integral\<^isup>P M ?f0" by simp @@ -528,14 +596,15 @@ qed show ?thesis proof (safe intro!: bexI[of _ f]) - fix A assume "A\sets M" + fix A assume A: "A\sets M" show "\ A = (\\<^isup>+x. f x * indicator A x \M)" proof (rule antisym) show "(\\<^isup>+x. f x * indicator A x \M) \ \ A" using `f \ G` `A \ sets M` unfolding G_def by auto show "\ A \ (\\<^isup>+x. f x * indicator A x \M)" using upper_bound[THEN bspec, OF `A \ sets M`] - by (simp add: pextreal_zero_le_diff) + using M'.real_measure[OF A] + by (cases "integral\<^isup>P M (?F A)") auto qed qed simp qed @@ -543,22 +612,22 @@ lemma (in finite_measure) split_space_into_finite_sets_and_rest: assumes "measure_space (M\measure := \\)" (is "measure_space ?N") assumes ac: "absolutely_continuous \" - shows "\\0\sets M. \\::nat\'a set. disjoint_family \ \ range \ \ sets M \ \0 = space M - (\i. \ i) \ - (\A\sets M. A \ \0 \ (\ A = 0 \ \ A = 0) \ (\ A > 0 \ \ A = \)) \ - (\i. \ (\ i) \ \)" + shows "\A0\sets M. \B::nat\'a set. disjoint_family B \ range B \ sets M \ A0 = space M - (\i. B i) \ + (\A\sets M. A \ A0 \ (\ A = 0 \ \ A = 0) \ (\ A > 0 \ \ A = \)) \ + (\i. \ (B i) \ \)" proof - interpret v: measure_space ?N where "space ?N = space M" and "sets ?N = sets M" and "measure ?N = \" by fact auto - let ?Q = "{Q\sets M. \ Q \ \}" + let ?Q = "{Q\sets M. \ Q \ \}" let ?a = "SUP Q:?Q. \ Q" have "{} \ ?Q" using v.empty_measure by auto then have Q_not_empty: "?Q \ {}" by blast have "?a \ \ (space M)" using sets_into_space by (auto intro!: SUP_leI measure_mono top) - then have "?a \ \" using finite_measure_of_space + then have "?a \ \" using finite_measure_of_space by auto - from SUPR_countable_SUPR[OF this Q_not_empty] + from SUPR_countable_SUPR[OF Q_not_empty, of \] obtain Q'' where "range Q'' \ \ ` ?Q" and a: "?a = (SUP i::nat. Q'' i)" by auto then have "\i. \Q'. Q'' i = \ Q' \ Q' \ ?Q" by auto @@ -569,7 +638,7 @@ have Union: "(SUP i. \ (?O i)) = \ (\i. ?O i)" proof (rule continuity_from_below[of ?O]) show "range ?O \ sets M" using Q' by (auto intro!: finite_UN) - show "\i. ?O i \ ?O (Suc i)" by fastsimp + show "incseq ?O" by (fastsimp intro!: incseq_SucI) qed have Q'_sets: "\i. Q' i \ sets M" using Q' by auto have O_sets: "\i. ?O i \ sets M" @@ -580,8 +649,8 @@ using Q' by (auto intro: finite_UN) with v.measure_finitely_subadditive[of "{.. i}" Q'] have "\ (?O i) \ (\i\i. \ (Q' i))" by auto - also have "\ < \" unfolding setsum_\ pextreal_less_\ using Q' by auto - finally show "\ (?O i) \ \" unfolding pextreal_less_\ by auto + also have "\ < \" using Q' by (simp add: setsum_Pinfty) + finally show "\ (?O i) \ \" by simp qed auto have O_mono: "\n. ?O n \ ?O (Suc n)" by fastsimp have a_eq: "?a = \ (\i. ?O i)" unfolding Union[symmetric] @@ -592,7 +661,7 @@ proof (safe intro!: Sup_mono, unfold bex_simps) fix i have *: "(\Q' ` {..i}) = ?O i" by auto - then show "\x. (x \ sets M \ \ x \ \) \ + then show "\x. (x \ sets M \ \ x \ \) \ \ (\Q' ` {..i}) \ \ x" using O_in_G[of i] by (auto intro!: exI[of _ "?O i"]) qed @@ -610,50 +679,52 @@ show "range Q \ sets M" using Q_sets by auto { fix A assume A: "A \ sets M" "A \ space M - ?O_0" - show "\ A = 0 \ \ A = 0 \ 0 < \ A \ \ A = \" + show "\ A = 0 \ \ A = 0 \ 0 < \ A \ \ A = \" proof (rule disjCI, simp) - assume *: "0 < \ A \ \ A \ \" + assume *: "0 < \ A \ \ A \ \" show "\ A = 0 \ \ A = 0" proof cases assume "\ A = 0" moreover with ac A have "\ A = 0" unfolding absolutely_continuous_def by auto ultimately show ?thesis by simp next - assume "\ A \ 0" with * have "\ A \ \" by auto + assume "\ A \ 0" with * have "\ A \ \" using positive_measure[OF A(1)] by auto with A have "\ ?O_0 + \ A = \ (?O_0 \ A)" using Q' by (auto intro!: measure_additive countable_UN) also have "\ = (SUP i. \ (?O i \ A))" proof (rule continuity_from_below[of "\i. ?O i \ A", symmetric, simplified]) show "range (\i. ?O i \ A) \ sets M" - using `\ A \ \` O_sets A by auto - qed fastsimp + using `\ A \ \` O_sets A by auto + qed (fastsimp intro!: incseq_SucI) also have "\ \ ?a" - proof (safe intro!: SUPR_bound) + proof (safe intro!: SUP_leI) fix i have "?O i \ A \ ?Q" proof (safe del: notI) show "?O i \ A \ sets M" using O_sets A by auto from O_in_G[of i] moreover have "\ (?O i \ A) \ \ (?O i) + \ A" using v.measure_subadditive[of "?O i" A] A O_sets by auto - ultimately show "\ (?O i \ A) \ \" - using `\ A \ \` by auto + ultimately show "\ (?O i \ A) \ \" + using `\ A \ \` by auto qed then show "\ (?O i \ A) \ ?a" by (rule le_SUPI) qed - finally have "\ A = 0" unfolding a_eq using finite_measure[OF `?O_0 \ sets M`] - by (cases "\ A") (auto simp: pextreal_noteq_omega_Ex) + finally have "\ A = 0" + unfolding a_eq using real_measure[OF `?O_0 \ sets M`] real_measure[OF A(1)] by auto with `\ A \ 0` show ?thesis by auto qed qed } - { fix i show "\ (Q i) \ \" + { fix i show "\ (Q i) \ \" proof (cases i) case 0 then show ?thesis unfolding Q_def using Q'[of 0] by simp next case (Suc n) then show ?thesis unfolding Q_def - using `?O n \ ?Q` `?O (Suc n) \ ?Q` O_mono - using v.measure_Diff[of "?O n" "?O (Suc n)"] by auto + using `?O n \ ?Q` `?O (Suc n) \ ?Q` + using v.measure_mono[OF O_mono, of n] v.positive_measure[of "?O n"] v.positive_measure[of "?O (Suc n)"] + using v.measure_Diff[of "?O n" "?O (Suc n)", OF _ _ _ O_mono] + by (cases rule: extreal2_cases[of "\ (\ x\Suc n. Q' x)" "\ (\ i\n. Q' i)"]) auto qed } show "space M - ?O_0 \ sets M" using Q'_sets by auto { fix j have "(\i\j. ?O i) = (\i\j. Q i)" @@ -675,7 +746,7 @@ lemma (in finite_measure) Radon_Nikodym_finite_measure_infinite: assumes "measure_space (M\measure := \\)" (is "measure_space ?N") assumes "absolutely_continuous \" - shows "\f \ borel_measurable M. \A\sets M. \ A = (\\<^isup>+x. f x * indicator A x \M)" + shows "\f \ borel_measurable M. (\x. 0 \ f x) \ (\A\sets M. \ A = (\\<^isup>+x. f x * indicator A x \M))" proof - interpret v: measure_space ?N where "space ?N = space M" and "sets ?N = sets M" and "measure ?N = \" @@ -684,14 +755,14 @@ obtain Q0 and Q :: "nat \ 'a set" where Q: "disjoint_family Q" "range Q \ sets M" and Q0: "Q0 \ sets M" "Q0 = space M - (\i. Q i)" - and in_Q0: "\A. A \ sets M \ A \ Q0 \ \ A = 0 \ \ A = 0 \ 0 < \ A \ \ A = \" - and Q_fin: "\i. \ (Q i) \ \" by force + and in_Q0: "\A. A \ sets M \ A \ Q0 \ \ A = 0 \ \ A = 0 \ 0 < \ A \ \ A = \" + and Q_fin: "\i. \ (Q i) \ \" by force from Q have Q_sets: "\i. Q i \ sets M" by auto - have "\i. \f. f\borel_measurable M \ (\A\sets M. + have "\i. \f. f\borel_measurable M \ (\x. 0 \ f x) \ (\A\sets M. \ (Q i \ A) = (\\<^isup>+x. f x * indicator (Q i \ A) x \M))" proof fix i - have indicator_eq: "\f x A. (f x :: pextreal) * indicator (Q i \ A) x * indicator (Q i) x + have indicator_eq: "\f x A. (f x :: extreal) * indicator (Q i \ A) x * indicator (Q i) x = (f x * indicator (Q i) x) * indicator A x" unfolding indicator_def by auto have fm: "finite_measure (restricted_space (Q i))" @@ -702,7 +773,7 @@ proof show "measure_space ?Q" using v.restricted_measure_space Q_sets[of i] by auto - show "measure ?Q (space ?Q) \ \" using Q_fin by simp + show "measure ?Q (space ?Q) \ \" using Q_fin by simp qed have "R.absolutely_continuous \" using `absolutely_continuous \` `Q i \ sets M` @@ -712,48 +783,40 @@ and f_int: "\A. A\sets M \ \ (Q i \ A) = (\\<^isup>+x. (f x * indicator (Q i) x) * indicator A x \M)" unfolding Bex_def borel_measurable_restricted[OF `Q i \ sets M`] positive_integral_restricted[OF `Q i \ sets M`] by (auto simp: indicator_eq) - then show "\f. f\borel_measurable M \ (\A\sets M. + then show "\f. f\borel_measurable M \ (\x. 0 \ f x) \ (\A\sets M. \ (Q i \ A) = (\\<^isup>+x. f x * indicator (Q i \ A) x \M))" - by (fastsimp intro!: exI[of _ "\x. f x * indicator (Q i) x"] positive_integral_cong - simp: indicator_def) + by (auto intro!: exI[of _ "\x. max 0 (f x * indicator (Q i) x)"] positive_integral_cong_pos + split: split_indicator split_if_asm simp: max_def) qed - from choice[OF this] obtain f where borel: "\i. f i \ borel_measurable M" + from choice[OF this] obtain f where borel: "\i. f i \ borel_measurable M" "\i x. 0 \ f i x" and f: "\A i. A \ sets M \ \ (Q i \ A) = (\\<^isup>+x. f i x * indicator (Q i \ A) x \M)" by auto - let "?f x" = - "(\\<^isub>\ i. f i x * indicator (Q i) x) + \ * indicator Q0 x" + let "?f x" = "(\i. f i x * indicator (Q i) x) + \ * indicator Q0 x" show ?thesis proof (safe intro!: bexI[of _ ?f]) - show "?f \ borel_measurable M" - by (safe intro!: borel_measurable_psuminf borel_measurable_pextreal_times - borel_measurable_pextreal_add borel_measurable_indicator - borel_measurable_const borel Q_sets Q0 Diff countable_UN) + show "?f \ borel_measurable M" using Q0 borel Q_sets + by (auto intro!: measurable_If) + show "\x. 0 \ ?f x" using borel by (auto intro!: suminf_0_le simp: indicator_def) fix A assume "A \ sets M" - have *: - "\x i. indicator A x * (f i x * indicator (Q i) x) = - f i x * indicator (Q i \ A) x" - "\x i. (indicator A x * indicator Q0 x :: pextreal) = - indicator (Q0 \ A) x" by (auto simp: indicator_def) - have "(\\<^isup>+x. ?f x * indicator A x \M) = - (\\<^isub>\ i. \ (Q i \ A)) + \ * \ (Q0 \ A)" - unfolding f[OF `A \ sets M`] - apply (simp del: pextreal_times(2) add: field_simps *) - apply (subst positive_integral_add) - apply (fastsimp intro: Q0 `A \ sets M`) - apply (fastsimp intro: Q_sets `A \ sets M` borel_measurable_psuminf borel) - apply (subst positive_integral_cmult_indicator) - apply (fastsimp intro: Q0 `A \ sets M`) - unfolding psuminf_cmult_right[symmetric] - apply (subst positive_integral_psuminf) - apply (fastsimp intro: `A \ sets M` Q_sets borel) - apply (simp add: *) - done - moreover have "(\\<^isub>\i. \ (Q i \ A)) = \ ((\i. Q i) \ A)" + have Qi: "\i. Q i \ sets M" using Q by auto + have [intro,simp]: "\i. (\x. f i x * indicator (Q i \ A) x) \ borel_measurable M" + "\i. AE x. 0 \ f i x * indicator (Q i \ A) x" + using borel Qi Q0(1) `A \ sets M` by (auto intro!: borel_measurable_extreal_times) + have "(\\<^isup>+x. ?f x * indicator A x \M) = (\\<^isup>+x. (\i. f i x * indicator (Q i \ A) x) + \ * indicator (Q0 \ A) x \M)" + using borel by (intro positive_integral_cong) (auto simp: indicator_def) + also have "\ = (\\<^isup>+x. (\i. f i x * indicator (Q i \ A) x) \M) + \ * \ (Q0 \ A)" + using borel Qi Q0(1) `A \ sets M` + by (subst positive_integral_add) (auto simp del: extreal_infty_mult + simp add: positive_integral_cmult_indicator Int intro!: suminf_0_le) + also have "\ = (\i. \ (Q i \ A)) + \ * \ (Q0 \ A)" + by (subst f[OF `A \ sets M`], subst positive_integral_suminf) auto + finally have "(\\<^isup>+x. ?f x * indicator A x \M) = (\i. \ (Q i \ A)) + \ * \ (Q0 \ A)" . + moreover have "(\i. \ (Q i \ A)) = \ ((\i. Q i) \ A)" using Q Q_sets `A \ sets M` by (intro v.measure_countably_additive[of "\i. Q i \ A", unfolded comp_def, simplified]) (auto simp: disjoint_family_on_def) - moreover have "\ * \ (Q0 \ A) = \ (Q0 \ A)" + moreover have "\ * \ (Q0 \ A) = \ (Q0 \ A)" proof - have "Q0 \ A \ sets M" using Q0(1) `A \ sets M` by blast from in_Q0[OF this] show ?thesis by auto @@ -770,40 +833,43 @@ lemma (in sigma_finite_measure) Radon_Nikodym: assumes "measure_space (M\measure := \\)" (is "measure_space ?N") - assumes "absolutely_continuous \" - shows "\f \ borel_measurable M. \A\sets M. \ A = (\\<^isup>+x. f x * indicator A x \M)" + assumes ac: "absolutely_continuous \" + shows "\f \ borel_measurable M. (\x. 0 \ f x) \ (\A\sets M. \ A = (\\<^isup>+x. f x * indicator A x \M))" proof - from Ex_finite_integrable_function - obtain h where finite: "integral\<^isup>P M h \ \" and + obtain h where finite: "integral\<^isup>P M h \ \" and borel: "h \ borel_measurable M" and + nn: "\x. 0 \ h x" and pos: "\x. x \ space M \ 0 < h x" and - "\x. x \ space M \ h x < \" by auto + "\x. x \ space M \ h x < \" by auto let "?T A" = "(\\<^isup>+x. h x * indicator A x \M)" let ?MT = "M\ measure := ?T \" - from measure_space_density[OF borel] finite interpret T: finite_measure ?MT where "space ?MT = space M" and "sets ?MT = sets M" and "measure ?MT = ?T" - unfolding finite_measure_def finite_measure_axioms_def - by (simp_all cong: positive_integral_cong) - have "\N. N \ sets M \ {x \ space M. h x \ 0 \ indicator N x \ (0::pextreal)} = N" - using sets_into_space pos by (force simp: indicator_def) - then have "T.absolutely_continuous \" using assms(2) borel - unfolding T.absolutely_continuous_def absolutely_continuous_def - by (fastsimp simp: borel_measurable_indicator positive_integral_0_iff) + unfolding finite_measure_def finite_measure_axioms_def using borel finite nn + by (auto intro!: measure_space_density cong: positive_integral_cong) + have "T.absolutely_continuous \" + proof (unfold T.absolutely_continuous_def, safe) + fix N assume "N \ sets M" "(\\<^isup>+x. h x * indicator N x \M) = 0" + with borel ac pos have "AE x. x \ N" + by (subst (asm) positive_integral_0_iff_AE) (auto split: split_indicator simp: not_le) + then have "N \ null_sets" using `N \ sets M` sets_into_space + by (subst (asm) AE_iff_measurable[OF `N \ sets M`]) auto + then show "\ N = 0" using ac by (auto simp: absolutely_continuous_def) + qed from T.Radon_Nikodym_finite_measure_infinite[simplified, OF assms(1) this] - obtain f where f_borel: "f \ borel_measurable M" and + obtain f where f_borel: "f \ borel_measurable M" "\x. 0 \ f x" and fT: "\A. A \ sets M \ \ A = (\\<^isup>+ x. f x * indicator A x \?MT)" by (auto simp: measurable_def) show ?thesis proof (safe intro!: bexI[of _ "\x. h x * f x"]) show "(\x. h x * f x) \ borel_measurable M" - using borel f_borel by (auto intro: borel_measurable_pextreal_times) + using borel f_borel by (auto intro: borel_measurable_extreal_times) + show "\x. 0 \ h x * f x" using nn f_borel by auto fix A assume "A \ sets M" - then have "(\x. f x * indicator A x) \ borel_measurable M" - using f_borel by (auto intro: borel_measurable_pextreal_times borel_measurable_indicator) - from positive_integral_translated_density[OF borel this] - show "\ A = (\\<^isup>+x. h x * f x * indicator A x \M)" - unfolding fT[OF `A \ sets M`] by (simp add: field_simps) + then show "\ A = (\\<^isup>+x. h x * f x * indicator A x \M)" + unfolding fT[OF `A \ sets M`] mult_assoc using nn borel f_borel + by (intro positive_integral_translated_density) auto qed qed @@ -811,7 +877,8 @@ lemma (in measure_space) finite_density_unique: assumes borel: "f \ borel_measurable M" "g \ borel_measurable M" - and fin: "integral\<^isup>P M f < \" + assumes pos: "AE x. 0 \ f x" "AE x. 0 \ g x" + and fin: "integral\<^isup>P M f \ \" shows "(\A\sets M. (\\<^isup>+x. f x * indicator A x \M) = (\\<^isup>+x. g x * indicator A x \M)) \ (AE x. f x = g x)" (is "(\A\sets M. ?P f A = ?P g A) \ _") @@ -822,42 +889,38 @@ next assume eq: "\A\sets M. ?P f A = ?P g A" from this[THEN bspec, OF top] fin - have g_fin: "integral\<^isup>P M g < \" by (simp cong: positive_integral_cong) + have g_fin: "integral\<^isup>P M g \ \" by (simp cong: positive_integral_cong) { fix f g assume borel: "f \ borel_measurable M" "g \ borel_measurable M" - and g_fin: "integral\<^isup>P M g < \" and eq: "\A\sets M. ?P f A = ?P g A" + and pos: "AE x. 0 \ f x" "AE x. 0 \ g x" + and g_fin: "integral\<^isup>P M g \ \" and eq: "\A\sets M. ?P f A = ?P g A" let ?N = "{x\space M. g x < f x}" have N: "?N \ sets M" using borel by simp + have "?P g ?N \ integral\<^isup>P M g" using pos + by (intro positive_integral_mono_AE) (auto split: split_indicator) + then have Pg_fin: "?P g ?N \ \" using g_fin by auto have "?P (\x. (f x - g x)) ?N = (\\<^isup>+x. f x * indicator ?N x - g x * indicator ?N x \M)" by (auto intro!: positive_integral_cong simp: indicator_def) also have "\ = ?P f ?N - ?P g ?N" proof (rule positive_integral_diff) show "(\x. f x * indicator ?N x) \ borel_measurable M" "(\x. g x * indicator ?N x) \ borel_measurable M" using borel N by auto - have "?P g ?N \ integral\<^isup>P M g" - by (auto intro!: positive_integral_mono simp: indicator_def) - then show "?P g ?N \ \" using g_fin by auto - fix x assume "x \ space M" - show "g x * indicator ?N x \ f x * indicator ?N x" - by (auto simp: indicator_def) - qed + show "AE x. g x * indicator ?N x \ f x * indicator ?N x" + "AE x. 0 \ g x * indicator ?N x" + using pos by (auto split: split_indicator) + qed fact also have "\ = 0" - using eq[THEN bspec, OF N] by simp - finally have "\ {x\space M. (f x - g x) * indicator ?N x \ 0} = 0" - using borel N by (subst (asm) positive_integral_0_iff) auto - moreover have "{x\space M. (f x - g x) * indicator ?N x \ 0} = ?N" - by (auto simp: pextreal_zero_le_diff) - ultimately have "?N \ null_sets" using N by simp } - from this[OF borel g_fin eq] this[OF borel(2,1) fin] - have "{x\space M. g x < f x} \ {x\space M. f x < g x} \ null_sets" - using eq by (intro null_sets_Un) auto - also have "{x\space M. g x < f x} \ {x\space M. f x < g x} = {x\space M. f x \ g x}" - by auto - finally show "AE x. f x = g x" - unfolding almost_everywhere_def by auto + unfolding eq[THEN bspec, OF N] using positive_integral_positive Pg_fin by auto + finally have "AE x. f x \ g x" + using pos borel positive_integral_PInf_AE[OF borel(2) g_fin] + by (subst (asm) positive_integral_0_iff_AE) + (auto split: split_indicator simp: not_less extreal_minus_le_iff) } + from this[OF borel pos g_fin eq] this[OF borel(2,1) pos(2,1) fin] eq + show "AE x. f x = g x" by auto qed lemma (in finite_measure) density_unique_finite_measure: assumes borel: "f \ borel_measurable M" "f' \ borel_measurable M" + assumes pos: "AE x. 0 \ f x" "AE x. 0 \ f' x" assumes f: "\A. A \ sets M \ (\\<^isup>+x. f x * indicator A x \M) = (\\<^isup>+x. f' x * indicator A x \M)" (is "\A. A \ sets M \ ?P f A = ?P f' A") shows "AE x. f x = f' x" @@ -865,26 +928,26 @@ let "?\ A" = "?P f A" and "?\' A" = "?P f' A" let "?f A x" = "f x * indicator A x" and "?f' A x" = "f' x * indicator A x" interpret M: measure_space "M\ measure := ?\\" - using borel(1) by (rule measure_space_density) simp + using borel(1) pos(1) by (rule measure_space_density) simp have ac: "absolutely_continuous ?\" using f by (rule density_is_absolutely_continuous) from split_space_into_finite_sets_and_rest[OF `measure_space (M\ measure := ?\\)` ac] obtain Q0 and Q :: "nat \ 'a set" where Q: "disjoint_family Q" "range Q \ sets M" and Q0: "Q0 \ sets M" "Q0 = space M - (\i. Q i)" - and in_Q0: "\A. A \ sets M \ A \ Q0 \ \ A = 0 \ ?\ A = 0 \ 0 < \ A \ ?\ A = \" - and Q_fin: "\i. ?\ (Q i) \ \" by force + and in_Q0: "\A. A \ sets M \ A \ Q0 \ \ A = 0 \ ?\ A = 0 \ 0 < \ A \ ?\ A = \" + and Q_fin: "\i. ?\ (Q i) \ \" by force from Q have Q_sets: "\i. Q i \ sets M" by auto let ?N = "{x\space M. f x \ f' x}" have "?N \ sets M" using borel by auto - have *: "\i x A. \y::pextreal. y * indicator (Q i) x * indicator A x = y * indicator (Q i \ A) x" + have *: "\i x A. \y::extreal. y * indicator (Q i) x * indicator A x = y * indicator (Q i \ A) x" unfolding indicator_def by auto - have "\i. AE x. ?f (Q i) x = ?f' (Q i) x" using borel Q_fin Q + have "\i. AE x. ?f (Q i) x = ?f' (Q i) x" using borel Q_fin Q pos by (intro finite_density_unique[THEN iffD1] allI) - (auto intro!: borel_measurable_pextreal_times f Int simp: *) + (auto intro!: borel_measurable_extreal_times f Int simp: *) moreover have "AE x. ?f Q0 x = ?f' Q0 x" proof (rule AE_I') - { fix f :: "'a \ pextreal" assume borel: "f \ borel_measurable M" + { fix f :: "'a \ extreal" assume borel: "f \ borel_measurable M" and eq: "\A. A \ sets M \ ?\ A = (\\<^isup>+x. f x * indicator A x \M)" let "?A i" = "Q0 \ {x \ space M. f x < of_nat i}" have "(\i. ?A i) \ null_sets" @@ -896,69 +959,74 @@ by (auto intro!: positive_integral_mono simp: indicator_def) also have "\ = of_nat i * \ (?A i)" using `?A i \ sets M` by (auto intro!: positive_integral_cmult_indicator) - also have "\ < \" + also have "\ < \" using `?A i \ sets M`[THEN finite_measure] by auto - finally have "?\ (?A i) \ \" by simp + finally have "?\ (?A i) \ \" by simp then show "?A i \ null_sets" using in_Q0[OF `?A i \ sets M`] `?A i \ sets M` by auto qed - also have "(\i. ?A i) = Q0 \ {x\space M. f x < \}" - by (auto simp: less_\_Ex_of_nat) - finally have "Q0 \ {x\space M. f x \ \} \ null_sets" by (simp add: pextreal_less_\) } + also have "(\i. ?A i) = Q0 \ {x\space M. f x \ \}" + by (auto simp: less_PInf_Ex_of_nat real_eq_of_nat) + finally have "Q0 \ {x\space M. f x \ \} \ null_sets" by simp } from this[OF borel(1) refl] this[OF borel(2) f] - have "Q0 \ {x\space M. f x \ \} \ null_sets" "Q0 \ {x\space M. f' x \ \} \ null_sets" by simp_all - then show "(Q0 \ {x\space M. f x \ \}) \ (Q0 \ {x\space M. f' x \ \}) \ null_sets" by (rule null_sets_Un) + have "Q0 \ {x\space M. f x \ \} \ null_sets" "Q0 \ {x\space M. f' x \ \} \ null_sets" by simp_all + then show "(Q0 \ {x\space M. f x \ \}) \ (Q0 \ {x\space M. f' x \ \}) \ null_sets" by (rule null_sets_Un) show "{x \ space M. ?f Q0 x \ ?f' Q0 x} \ - (Q0 \ {x\space M. f x \ \}) \ (Q0 \ {x\space M. f' x \ \})" by (auto simp: indicator_def) + (Q0 \ {x\space M. f x \ \}) \ (Q0 \ {x\space M. f' x \ \})" by (auto simp: indicator_def) qed moreover have "\x. (?f Q0 x = ?f' Q0 x) \ (\i. ?f (Q i) x = ?f' (Q i) x) \ ?f (space M) x = ?f' (space M) x" by (auto simp: indicator_def Q0) ultimately have "AE x. ?f (space M) x = ?f' (space M) x" - by (auto simp: all_AE_countable) + by (auto simp: AE_all_countable[symmetric]) then show "AE x. f x = f' x" by auto qed lemma (in sigma_finite_measure) density_unique: - assumes borel: "f \ borel_measurable M" "f' \ borel_measurable M" - assumes f: "\A. A \ sets M \ (\\<^isup>+x. f x * indicator A x \M) = (\\<^isup>+x. f' x * indicator A x \M)" + assumes f: "f \ borel_measurable M" "AE x. 0 \ f x" + assumes f': "f' \ borel_measurable M" "AE x. 0 \ f' x" + assumes eq: "\A. A \ sets M \ (\\<^isup>+x. f x * indicator A x \M) = (\\<^isup>+x. f' x * indicator A x \M)" (is "\A. A \ sets M \ ?P f A = ?P f' A") shows "AE x. f x = f' x" proof - obtain h where h_borel: "h \ borel_measurable M" - and fin: "integral\<^isup>P M h \ \" and pos: "\x. x \ space M \ 0 < h x \ h x < \" + and fin: "integral\<^isup>P M h \ \" and pos: "\x. x \ space M \ 0 < h x \ h x < \" "\x. 0 \ h x" using Ex_finite_integrable_function by auto - interpret h: measure_space "M\ measure := \A. (\\<^isup>+x. h x * indicator A x \M) \" - using h_borel by (rule measure_space_density) simp + then have h_nn: "AE x. 0 \ h x" by auto + let ?H = "M\ measure := \A. (\\<^isup>+x. h x * indicator A x \M) \" + have H: "measure_space ?H" + using h_borel h_nn by (rule measure_space_density) simp + then interpret h: measure_space ?H . interpret h: finite_measure "M\ measure := \A. (\\<^isup>+x. h x * indicator A x \M) \" by default (simp cong: positive_integral_cong add: fin) let ?fM = "M\measure := \A. (\\<^isup>+x. f x * indicator A x \M)\" interpret f: measure_space ?fM - using borel(1) by (rule measure_space_density) simp + using f by (rule measure_space_density) simp let ?f'M = "M\measure := \A. (\\<^isup>+x. f' x * indicator A x \M)\" interpret f': measure_space ?f'M - using borel(2) by (rule measure_space_density) simp + using f' by (rule measure_space_density) simp { fix A assume "A \ sets M" - then have " {x \ space M. h x \ 0 \ indicator A x \ (0::pextreal)} = A" - using pos sets_into_space by (force simp: indicator_def) + then have "{x \ space M. h x * indicator A x \ 0} = A" + using pos(1) sets_into_space by (force simp: indicator_def) then have "(\\<^isup>+x. h x * indicator A x \M) = 0 \ A \ null_sets" - using h_borel `A \ sets M` by (simp add: positive_integral_0_iff) } + using h_borel `A \ sets M` h_nn by (subst positive_integral_0_iff) auto } note h_null_sets = this { fix A assume "A \ sets M" - have "(\\<^isup>+x. h x * (f x * indicator A x) \M) = (\\<^isup>+x. h x * indicator A x \?fM)" - using `A \ sets M` h_borel borel - by (simp add: positive_integral_translated_density ac_simps cong: positive_integral_cong) + have "(\\<^isup>+x. f x * (h x * indicator A x) \M) = (\\<^isup>+x. h x * indicator A x \?fM)" + using `A \ sets M` h_borel h_nn f f' + by (intro positive_integral_translated_density[symmetric]) auto also have "\ = (\\<^isup>+x. h x * indicator A x \?f'M)" - by (rule f'.positive_integral_cong_measure) (simp_all add: f) - also have "\ = (\\<^isup>+x. h x * (f' x * indicator A x) \M)" - using `A \ sets M` h_borel borel - by (simp add: positive_integral_translated_density ac_simps cong: positive_integral_cong) - finally have "(\\<^isup>+x. h x * (f x * indicator A x) \M) = (\\<^isup>+x. h x * (f' x * indicator A x) \M)" . } - then have "h.almost_everywhere (\x. f x = f' x)" - using h_borel borel - apply (intro h.density_unique_finite_measure) - apply (simp add: measurable_def) - apply (simp add: measurable_def) - by (simp add: positive_integral_translated_density) + by (rule f'.positive_integral_cong_measure) (simp_all add: eq) + also have "\ = (\\<^isup>+x. f' x * (h x * indicator A x) \M)" + using `A \ sets M` h_borel h_nn f f' + by (intro positive_integral_translated_density) auto + finally have "(\\<^isup>+x. h x * (f x * indicator A x) \M) = (\\<^isup>+x. h x * (f' x * indicator A x) \M)" + by (simp add: ac_simps) + then have "(\\<^isup>+x. (f x * indicator A x) \?H) = (\\<^isup>+x. (f' x * indicator A x) \?H)" + using `A \ sets M` h_borel h_nn f f' + by (subst (asm) (1 2) positive_integral_translated_density[symmetric]) auto } + then have "AE x in ?H. f x = f' x" using h_borel h_nn f f' + by (intro h.density_unique_finite_measure absolutely_continuous_AE[OF H] density_is_absolutely_continuous) + simp_all then show "AE x. f x = f' x" unfolding h.almost_everywhere_def almost_everywhere_def by (auto simp add: h_null_sets) @@ -966,41 +1034,42 @@ lemma (in sigma_finite_measure) sigma_finite_iff_density_finite: assumes \: "measure_space (M\measure := \\)" (is "measure_space ?N") - and f: "f \ borel_measurable M" + and f: "f \ borel_measurable M" "AE x. 0 \ f x" and eq: "\A. A \ sets M \ \ A = (\\<^isup>+x. f x * indicator A x \M)" - shows "sigma_finite_measure (M\measure := \\) \ (AE x. f x \ \)" + shows "sigma_finite_measure (M\measure := \\) \ (AE x. f x \ \)" proof assume "sigma_finite_measure ?N" then interpret \: sigma_finite_measure ?N where "borel_measurable ?N = borel_measurable M" and "measure ?N = \" and "sets ?N = sets M" and "space ?N = space M" by (auto simp: measurable_def) from \.Ex_finite_integrable_function obtain h where - h: "h \ borel_measurable M" "integral\<^isup>P ?N h \ \" - and fin: "\x\space M. 0 < h x \ h x < \" by auto - have "AE x. f x * h x \ \" + h: "h \ borel_measurable M" "integral\<^isup>P ?N h \ \" and + h_nn: "\x. 0 \ h x" and + fin: "\x\space M. 0 < h x \ h x < \" by auto + have "AE x. f x * h x \ \" proof (rule AE_I') - have "integral\<^isup>P ?N h = (\\<^isup>+x. f x * h x \M)" using f h + have "integral\<^isup>P ?N h = (\\<^isup>+x. f x * h x \M)" using f h h_nn by (subst \.positive_integral_cong_measure[symmetric, of "M\ measure := \ A. \\<^isup>+x. f x * indicator A x \M \"]) (auto intro!: positive_integral_translated_density simp: eq) - then have "(\\<^isup>+x. f x * h x \M) \ \" + then have "(\\<^isup>+x. f x * h x \M) \ \" using h(2) by simp - then show "(\x. f x * h x) -` {\} \ space M \ null_sets" - using f h(1) by (auto intro!: positive_integral_omega borel_measurable_vimage) + then show "(\x. f x * h x) -` {\} \ space M \ null_sets" + using f h(1) by (auto intro!: positive_integral_PInf borel_measurable_vimage) qed auto - then show "AE x. f x \ \" + then show "AE x. f x \ \" using fin by (auto elim!: AE_Ball_mp) next - assume AE: "AE x. f x \ \" + assume AE: "AE x. f x \ \" from sigma_finite guess Q .. note Q = this interpret \: measure_space ?N where "borel_measurable ?N = borel_measurable M" and "measure ?N = \" and "sets ?N = sets M" and "space ?N = space M" using \ by (auto simp: measurable_def) - def A \ "\i. f -` (case i of 0 \ {\} | Suc n \ {.. of_nat (Suc n)}) \ space M" + def A \ "\i. f -` (case i of 0 \ {\} | Suc n \ {.. of_nat (Suc n)}) \ space M" { fix i j have "A i \ Q j \ sets M" unfolding A_def using f Q apply (rule_tac Int) - by (cases i) (auto intro: measurable_sets[OF f]) } + by (cases i) (auto intro: measurable_sets[OF f(1)]) } note A_in_sets = this let "?A n" = "case prod_decode n of (i,j) \ A i \ Q j" show "sigma_finite_measure ?N" @@ -1021,18 +1090,21 @@ fix x assume x: "x \ space M" show "x \ (\i. A i)" proof (cases "f x") - case infinite then show ?thesis using x unfolding A_def by (auto intro: exI[of _ 0]) + case PInf with x show ?thesis unfolding A_def by (auto intro: exI[of _ 0]) next - case (preal r) - with less_\_Ex_of_nat[of "f x"] obtain n where "f x < of_nat n" by auto - then show ?thesis using x preal unfolding A_def by (auto intro!: exI[of _ "Suc n"]) + case (real r) + with less_PInf_Ex_of_nat[of "f x"] obtain n where "f x < of_nat n" by (auto simp: real_eq_of_nat) + then show ?thesis using x real unfolding A_def by (auto intro!: exI[of _ "Suc n"]) + next + case MInf with x show ?thesis + unfolding A_def by (auto intro!: exI[of _ "Suc 0"]) qed qed (auto simp: A_def) finally show "(\i. ?A i) = space ?N" by simp next fix n obtain i j where [simp]: "prod_decode n = (i, j)" by (cases "prod_decode n") auto - have "(\\<^isup>+x. f x * indicator (A i \ Q j) x \M) \ \" + have "(\\<^isup>+x. f x * indicator (A i \ Q j) x \M) \ \" proof (cases i) case 0 have "AE x. f x * indicator (A i \ Q j) x = 0" @@ -1045,11 +1117,11 @@ by (auto intro!: positive_integral_mono simp: indicator_def A_def) also have "\ = of_nat (Suc n) * \ (Q j)" using Q by (auto intro!: positive_integral_cmult_indicator) - also have "\ < \" - using Q by auto + also have "\ < \" + using Q by (auto simp: real_eq_of_nat[symmetric]) finally show ?thesis by simp qed - then show "measure ?N (?A n) \ \" + then show "measure ?N (?A n) \ \" using A_in_sets Q eq by auto qed qed @@ -1057,7 +1129,7 @@ section "Radon-Nikodym derivative" definition - "RN_deriv M \ \ SOME f. f \ borel_measurable M \ + "RN_deriv M \ \ SOME f. f \ borel_measurable M \ (\x. 0 \ f x) \ (\A \ sets M. \ A = (\\<^isup>+x. f x * indicator A x \M))" lemma (in sigma_finite_measure) RN_deriv_cong: @@ -1078,9 +1150,12 @@ shows "RN_deriv M \ \ borel_measurable M" (is ?borel) and "\A. A \ sets M \ \ A = (\\<^isup>+x. RN_deriv M \ x * indicator A x \M)" (is "\A. _ \ ?int A") + and "0 \ RN_deriv M \ x" proof - note Ex = Radon_Nikodym[OF assms, unfolded Bex_def] - thus ?borel unfolding RN_deriv_def by (rule someI2_ex) auto + then show ?borel unfolding RN_deriv_def by (rule someI2_ex) auto + from Ex show "0 \ RN_deriv M \ x" unfolding RN_deriv_def + by (rule someI2_ex) simp fix A assume "A \ sets M" from Ex show "?int A" unfolding RN_deriv_def by (rule someI2_ex) (simp add: `A \ sets M`) @@ -1092,22 +1167,28 @@ shows "integral\<^isup>P (M\measure := \\) f = (\\<^isup>+x. RN_deriv M \ x * f x \M)" proof - interpret \: measure_space "M\measure := \\" by fact - have "integral\<^isup>P (M\measure := \\) f = - integral\<^isup>P (M\measure := \A. (\\<^isup>+x. RN_deriv M \ x * indicator A x \M)\) f" - by (intro \.positive_integral_cong_measure[symmetric]) - (simp_all add: RN_deriv(2)[OF \, symmetric]) + note RN = RN_deriv[OF \] + have "integral\<^isup>P (M\measure := \\) f = (\\<^isup>+x. max 0 (f x) \M\measure := \\)" + unfolding positive_integral_max_0 .. + also have "(\\<^isup>+x. max 0 (f x) \M\measure := \\) = + (\\<^isup>+x. max 0 (f x) \M\measure := \A. (\\<^isup>+x. RN_deriv M \ x * indicator A x \M)\)" + by (intro \.positive_integral_cong_measure[symmetric]) (simp_all add: RN(2)) + also have "\ = (\\<^isup>+x. RN_deriv M \ x * max 0 (f x) \M)" + by (intro positive_integral_translated_density) (auto simp add: RN f) also have "\ = (\\<^isup>+x. RN_deriv M \ x * f x \M)" - by (intro positive_integral_translated_density) - (simp_all add: RN_deriv[OF \] f) + using RN_deriv(3)[OF \] + by (auto intro!: positive_integral_cong_pos split: split_if_asm + simp: max_def extreal_mult_le_0_iff) finally show ?thesis . qed lemma (in sigma_finite_measure) RN_deriv_unique: assumes \: "measure_space (M\measure := \\)" "absolutely_continuous \" - and f: "f \ borel_measurable M" + and f: "f \ borel_measurable M" "AE x. 0 \ f x" and eq: "\A. A \ sets M \ \ A = (\\<^isup>+x. f x * indicator A x \M)" shows "AE x. f x = RN_deriv M \ x" proof (rule density_unique[OF f RN_deriv(1)[OF \]]) + show "AE x. 0 \ RN_deriv M \ x" using RN_deriv[OF \] by auto fix A assume A: "A \ sets M" show "(\\<^isup>+x. f x * indicator A x \M) = (\\<^isup>+x. RN_deriv M \ x * indicator A x \M)" unfolding eq[OF A, symmetric] RN_deriv(2)[OF \ A, symmetric] .. @@ -1143,7 +1224,7 @@ interpret M': sigma_finite_measure M' proof from sigma_finite guess F .. note F = this - show "\A::nat \ 'c set. range A \ sets M' \ (\i. A i) = space M' \ (\i. M'.\ (A i) \ \)" + show "\A::nat \ 'c set. range A \ sets M' \ (\i. A i) = space M' \ (\i. M'.\ (A i) \ \)" proof (intro exI conjI allI) show *: "range (\i. T' -` F i \ space M') \ sets M'" using F T' by (auto simp: measurable_def measure_preserving_def) @@ -1157,7 +1238,7 @@ then have "T -` (T' -` F i \ space M') \ space M = F i" using T inv sets_into_space[OF Fi] by (auto simp: measurable_def measure_preserving_def) - ultimately show "measure M' (T' -` F i \ space M') \ \" + ultimately show "measure M' (T' -` F i \ space M') \ \" using F by simp qed qed @@ -1165,6 +1246,7 @@ by (intro measurable_comp[where b=M'] M'.RN_deriv(1) measure_preservingD2[OF T]) fact+ then show "(\x. RN_deriv M' \' (T x)) \ borel_measurable M" by (simp add: comp_def) + show "AE x. 0 \ RN_deriv M' \' (T x)" using M'.RN_deriv(3)[OF \'] by auto fix A let ?A = "T' -` A \ space M'" assume A: "A \ sets M" then have A': "?A \ sets M'" using T' unfolding measurable_def measure_preserving_def @@ -1185,12 +1267,12 @@ lemma (in sigma_finite_measure) RN_deriv_finite: assumes sfm: "sigma_finite_measure (M\measure := \\)" and ac: "absolutely_continuous \" - shows "AE x. RN_deriv M \ x \ \" + shows "AE x. RN_deriv M \ x \ \" proof - interpret \: sigma_finite_measure "M\measure := \\" by fact have \: "measure_space (M\measure := \\)" by default from sfm show ?thesis - using sigma_finite_iff_density_finite[OF \ RN_deriv[OF \ ac]] by simp + using sigma_finite_iff_density_finite[OF \ RN_deriv(1)[OF \ ac]] RN_deriv(2,3)[OF \ ac] by simp qed lemma (in sigma_finite_measure) @@ -1203,22 +1285,24 @@ proof - interpret \: sigma_finite_measure "M\measure := \\" by fact have ms: "measure_space (M\measure := \\)" by default - have minus_cong: "\A B A' B'::pextreal. A = A' \ B = B' \ real A - real B = real A' - real B'" by simp + have minus_cong: "\A B A' B'::extreal. A = A' \ B = B' \ real A - real B = real A' - real B'" by simp have f': "(\x. - f x) \ borel_measurable M" using f by auto - have Nf: "f \ borel_measurable (M\measure := \\)" using f unfolding measurable_def by auto + have Nf: "f \ borel_measurable (M\measure := \\)" using f by simp { fix f :: "'a \ real" - { fix x assume *: "RN_deriv M \ x \ \" - have "Real (real (RN_deriv M \ x)) * Real (f x) = Real (real (RN_deriv M \ x) * f x)" + { fix x assume *: "RN_deriv M \ x \ \" + have "extreal (real (RN_deriv M \ x)) * extreal (f x) = extreal (real (RN_deriv M \ x) * f x)" by (simp add: mult_le_0_iff) - then have "RN_deriv M \ x * Real (f x) = Real (real (RN_deriv M \ x) * f x)" - using * by (simp add: Real_real) } - then have "(\\<^isup>+x. RN_deriv M \ x * Real (f x) \M) = (\\<^isup>+x. Real (real (RN_deriv M \ x) * f x) \M)" - using RN_deriv_finite[OF \] by (auto intro: positive_integral_cong_AE) } - with this this f f' Nf + then have "RN_deriv M \ x * extreal (f x) = extreal (real (RN_deriv M \ x) * f x)" + using RN_deriv(3)[OF ms \(2)] * by (auto simp add: extreal_real split: split_if_asm) } + then have "(\\<^isup>+x. extreal (real (RN_deriv M \ x) * f x) \M) = (\\<^isup>+x. RN_deriv M \ x * extreal (f x) \M)" + "(\\<^isup>+x. extreal (- (real (RN_deriv M \ x) * f x)) \M) = (\\<^isup>+x. RN_deriv M \ x * extreal (- f x) \M)" + using RN_deriv_finite[OF \] unfolding extreal_mult_minus_right uminus_extreal.simps(1)[symmetric] + by (auto intro!: positive_integral_cong_AE) } + note * = this show ?integral ?integrable - unfolding lebesgue_integral_def integrable_def - by (auto intro!: RN_deriv(1)[OF ms \(2)] minus_cong - simp: RN_deriv_positive_integral[OF ms \(2)]) + unfolding lebesgue_integral_def integrable_def * + using f RN_deriv(1)[OF ms \(2)] + by (auto simp: RN_deriv_positive_integral[OF ms \(2)]) qed lemma (in sigma_finite_measure) RN_deriv_singleton: @@ -1231,7 +1315,7 @@ from deriv(2)[OF `{x} \ sets M`] have "\ {x} = (\\<^isup>+w. RN_deriv M \ x * indicator {x} w \M)" by (auto simp: indicator_def intro!: positive_integral_cong) - thus ?thesis using positive_integral_cmult_indicator[OF `{x} \ sets M`] + thus ?thesis using positive_integral_cmult_indicator[OF _ `{x} \ sets M`] deriv(3) by auto qed diff -r 28b51effc5ed -r cdf7693bbe08 src/HOL/Probability/Sigma_Algebra.thy --- a/src/HOL/Probability/Sigma_Algebra.thy Mon Mar 14 14:37:47 2011 +0100 +++ b/src/HOL/Probability/Sigma_Algebra.thy Mon Mar 14 14:37:49 2011 +0100 @@ -1,7 +1,7 @@ -(* Title: HOL/Probability/Sigma_Algebra.thy - Author: Stefan Richter - Author: Markus Wenzel - Author: Lawrence Paulson +(* Title: Sigma_Algebra.thy + Author: Stefan Richter, Markus Wenzel, TU Muenchen + Plus material from the Hurd/Coble measure theory development, + translated by Lawrence Paulson. *) header {* Sigma Algebras *} @@ -70,6 +70,16 @@ "finite X \ X \ sets M \ Union X \ sets M" by (induct set: finite) (auto simp add: Un) +lemma (in algebra) finite_UN[intro]: + assumes "finite I" and "\i. i \ I \ A i \ sets M" + shows "(\i\I. A i) \ sets M" + using assms by induct auto + +lemma (in algebra) finite_INT[intro]: + assumes "finite I" "I \ {}" "\i. i \ I \ A i \ sets M" + shows "(\i\I. A i) \ sets M" + using assms by (induct rule: finite_ne_induct) auto + lemma algebra_iff_Int: "algebra M \ sets M \ Pow (space M) & {} \ sets M & @@ -149,11 +159,6 @@ ultimately show ?thesis by simp qed -lemma (in sigma_algebra) finite_UN: - assumes "finite I" and "\i. i \ I \ A i \ sets M" - shows "(\i\I. A i) \ sets M" - using assms by induct auto - lemma (in sigma_algebra) countable_INT [intro]: fixes A :: "'i::countable \ 'a set" assumes A: "A`X \ sets M" "X \ {}" @@ -167,11 +172,6 @@ ultimately show ?thesis by metis qed -lemma (in sigma_algebra) finite_INT: - assumes "finite I" "I \ {}" "\i. i \ I \ A i \ sets M" - shows "(\i\I. A i) \ sets M" - using assms by (induct rule: finite_ne_induct) auto - lemma algebra_Pow: "algebra \ space = sp, sets = Pow sp, \ = X \" by (auto simp add: algebra_def) diff -r 28b51effc5ed -r cdf7693bbe08 src/HOL/Probability/ex/Dining_Cryptographers.thy --- a/src/HOL/Probability/ex/Dining_Cryptographers.thy Mon Mar 14 14:37:47 2011 +0100 +++ b/src/HOL/Probability/ex/Dining_Cryptographers.thy Mon Mar 14 14:37:49 2011 +0100 @@ -1,5 +1,5 @@ theory Dining_Cryptographers -imports Information +imports "~~/src/HOL/Probability/Information" begin locale finite_space = @@ -8,7 +8,7 @@ and not_empty[simp]: "S \ {}" definition (in finite_space) "M = \ space = S, sets = Pow S, - measure = (\A. of_nat (card A) / of_nat (card S) :: pextreal) \" + measure = \A. extreal (card A / card S) \" lemma (in finite_space) shows space_M[simp]: "space M = S" @@ -19,13 +19,14 @@ proof (rule finite_measure_spaceI) fix A B :: "'a set" assume "A \ B = {}" "A \ space M" "B \ space M" then show "measure M (A \ B) = measure M A + measure M B" - by (simp add: inverse_eq_divide field_simps Real_real M_def - divide_le_0_iff zero_le_divide_iff - card_Un_disjoint finite_subset[OF _ finite]) -qed (auto simp: M_def) + by (simp add: M_def card_Un_disjoint finite_subset[OF _ finite] field_simps) +qed (auto simp: M_def divide_nonneg_nonneg) sublocale finite_space \ information_space M 2 - by default (simp_all add: M_def) + by default (simp_all add: M_def one_extreal_def) + +lemma (in finite_space) \'_eq[simp]: "\' A = (if A \ S then card A / card S else 0)" + unfolding \'_def by (auto simp: M_def) lemma set_of_list_extend: "{xs. length xs = Suc n \ (\x\set xs. x \ A)} = @@ -491,7 +492,7 @@ let ?dI = "distribution inversion" { have "\(inversion|payer) = - - (\x\inversion`dc_crypto. (\z\Some ` {0..x\inversion`dc_crypto. (\z\Some ` {0.. dc_crypto = {z} \ {xs. length xs = n}" by (auto simp: dc_crypto payer_def) hence "card (payer -` {z} \ dc_crypto) = 2^n" using card_list_length[where A="UNIV::bool set"] by (simp add: card_cartesian_product_singleton) - hence "real (?dP {z}) = 1 / real n" unfolding distribution_def - by (simp add: card_dc_crypto field_simps inverse_eq_divide M_def - mult_le_0_iff zero_le_mult_iff power_le_zero_eq) + hence "?dP {z} = 1 / real n" + by (simp add: distribution_def card_dc_crypto) ultimately - show "real (?dIP {(x,z)}) * log 2 (real (?dIP {(x,z)}) / real (?dP {z})) = + show "?dIP {(x,z)} * log 2 (?dIP {(x,z)} / ?dP {z}) = 2 / (real n * 2^n) * (1 - real n)" - by (simp add: field_simps log_divide log_nat_power[of 2]) + by (simp add: log_divide log_nat_power[of 2]) qed also have "... = real n - 1" using n finite_space by (simp add: card_image_inversion card_image[OF inj_Some] field_simps real_eq_of_nat[symmetric]) finally have "\(inversion|payer) = real n - 1" . } moreover - { have "\(inversion) = - (\x \ inversion`dc_crypto. real (?dI {x}) * log 2 (real (?dI {x})))" + { have "\(inversion) = - (\x \ inversion`dc_crypto. ?dI {x} * log 2 (?dI {x}))" unfolding entropy_eq[OF simple_function_finite] by simp also have "... = - (\x \ inversion`dc_crypto. 2 * (1 - real n) / 2^n)" unfolding neg_equal_iff_equal @@ -536,10 +534,9 @@ fix x assume x_inv: "x \ inversion ` dc_crypto" hence "length x = n" by (auto simp: inversion_def_raw dc_crypto) moreover have "inversion -` {x} \ dc_crypto = {dc \ dc_crypto. inversion dc = x}" by auto - ultimately have "?dI {x} = 2 / 2^n" using `0 < n` unfolding distribution_def - by (simp add: card_inversion[OF x_inv] card_dc_crypto M_def - mult_le_0_iff zero_le_mult_iff power_le_zero_eq) - thus "real (?dI {x}) * log 2 (real (?dI {x})) = 2 * (1 - real n) / 2^n" + ultimately have "?dI {x} = 2 / 2^n" using `0 < n` + by (auto simp: card_inversion[OF x_inv] card_dc_crypto distribution_def) + thus "?dI {x} * log 2 (?dI {x}) = 2 * (1 - real n) / 2^n" by (simp add: log_divide log_nat_power power_le_zero_eq inverse_eq_divide) qed also have "... = real n - 1" diff -r 28b51effc5ed -r cdf7693bbe08 src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy --- a/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy Mon Mar 14 14:37:47 2011 +0100 +++ b/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy Mon Mar 14 14:37:49 2011 +0100 @@ -3,7 +3,7 @@ header {* Formalization of a Countermeasure by Koepf \& Duermuth 2009 *} theory Koepf_Duermuth_Countermeasure - imports Information "~~/src/HOL/Library/Permutation" + imports "~~/src/HOL/Probability/Information" "~~/src/HOL/Library/Permutation" begin lemma @@ -201,14 +201,17 @@ lemma (in finite_information) positive_p_sum[simp]: "0 \ setsum p X" by (auto intro!: setsum_nonneg) -sublocale finite_information \ finite_measure_space "\ space = \, sets = Pow \, measure = \S. Real (setsum p S)\" +sublocale finite_information \ finite_measure_space "\ space = \, sets = Pow \, measure = \S. extreal (setsum p S)\" by (rule finite_measure_spaceI) (simp_all add: setsum_Un_disjoint finite_subset) -sublocale finite_information \ finite_prob_space "\ space = \, sets = Pow \, measure = \S. Real (setsum p S)\" +sublocale finite_information \ finite_prob_space "\ space = \, sets = Pow \, measure = \S. extreal (setsum p S)\" + by default (simp add: one_extreal_def) + +sublocale finite_information \ information_space "\ space = \, sets = Pow \, measure = \S. extreal (setsum p S)\" b by default simp -sublocale finite_information \ information_space "\ space = \, sets = Pow \, measure = \S. Real (setsum p S)\" b - by default simp +lemma (in finite_information) \'_eq: "A \ \ \ \' A = setsum p A" + unfolding \'_def by auto locale koepf_duermuth = K: finite_information keys K b + M: finite_information messages M b for b :: real @@ -259,16 +262,6 @@ "snd ` (SIGMA x:f`X. f -` {x} \ X) = X" by (auto simp: image_iff) -lemma zero_le_eq_True: "0 \ (x::pextreal) \ True" by simp - -lemma Real_setprod: - assumes"\i. i\X \ 0 \ f i" - shows "(\i\X. Real (f i)) = Real (\i\X. f i)" -proof cases - assume "finite X" - from this assms show ?thesis by induct (auto simp: mult_le_0_iff) -qed simp - lemma inj_Cons[simp]: "\X. inj_on (\(xs, x). x#xs) X" by (auto intro!: inj_onI) lemma setprod_setsum_distrib_lists: @@ -323,10 +316,10 @@ "p A \ setsum P A" abbreviation probability ("\

'(_') _") where - "\

(X) x \ real (distribution X x)" + "\

(X) x \ distribution X x" abbreviation joint_probability ("\

'(_, _') _") where - "\

(X, Y) x \ real (joint_distribution X Y x)" + "\

(X, Y) x \ joint_distribution X Y x" abbreviation conditional_probability ("\

'(_|_') _") where "\

(X|Y) x \ \

(X, Y) x / \

(Y) (snd`x)" @@ -355,11 +348,12 @@ from assms have *: "fst -` {k} \ msgs = {k}\{ms. length ms = n \ (\M\set ms. M \ messages)}" unfolding msgs_def by auto - show "\

(fst) {k} = K k" unfolding distribution_def - apply (simp add: *) unfolding P_def + show "\

(fst) {k} = K k" + apply (simp add: \'_eq distribution_def) + apply (simp add: * P_def) apply (simp add: setsum_cartesian_product') - using setprod_setsum_distrib_lists[OF M.finite_space, of M n "\x x. True"] - by (simp add: setsum_right_distrib[symmetric] subset_eq setprod_1) + using setprod_setsum_distrib_lists[OF M.finite_space, of M n "\x x. True"] `k \ keys` + by (auto simp add: setsum_right_distrib[symmetric] subset_eq setprod_1) qed lemma fst_image_msgs[simp]: "fst`msgs = keys" @@ -390,7 +384,7 @@ have "\

(OB, fst) {(obs, k)} / K k = p ({k}\{ms. (k,ms) \ msgs \ OB (k,ms) = obs}) / K k" - unfolding distribution_def by (auto intro!: arg_cong[where f=p]) + apply (simp add: distribution_def \'_eq) by (auto intro!: arg_cong[where f=p]) also have "\ = (\im\{m\messages. observe k m = obs ! i}. M m)" unfolding P_def using `K k \ 0` `k \ keys` @@ -415,7 +409,7 @@ unfolding disjoint_family_on_def by auto have "\

(t\OB, fst) {(t obs, k)} = (\obs'\?S obs. \

(OB, fst) {(obs', k)})" unfolding distribution_def comp_def - using real_finite_measure_finite_Union[OF _ df] + using finite_measure_finite_Union[OF _ _ df] by (force simp add: * intro!: setsum_nonneg) also have "(\obs'\?S obs. \

(OB, fst) {(obs', k)}) = real (card (?S obs)) * \

(OB, fst) {(obs, k)}" by (simp add: t_eq_imp[OF `k \ keys` `K k \ 0` obs] real_eq_of_nat) @@ -433,11 +427,11 @@ then have "real (card ?S) \ 0" by auto have "\

(fst | t\OB) {(k, t obs)} = \

(t\OB | fst) {(t obs, k)} * \

(fst) {k} / \

(t\OB) {t obs}" - using real_distribution_order'[of fst k "t\OB" "t obs"] + using distribution_order(7,8)[where X=fst and x=k and Y="t\OB" and y="t obs"] by (subst joint_distribution_commute) auto also have "\

(t\OB) {t obs} = (\k'\keys. \

(t\OB|fst) {(t obs, k')} * \

(fst) {k'})" - using setsum_real_distribution(2)[of "t\OB" fst "t obs", symmetric] - using real_distribution_order'[of fst _ "t\OB" "t obs"] by (auto intro!: setsum_cong) + using setsum_distribution(2)[of "t\OB" fst "t obs", symmetric] + by (auto intro!: setsum_cong distribution_order(8)) also have "\

(t\OB | fst) {(t obs, k)} * \

(fst) {k} / (\k'\keys. \

(t\OB|fst) {(t obs, k')} * \

(fst) {k'}) = \

(OB | fst) {(obs, k)} * \

(fst) {k} / (\k'\keys. \

(OB|fst) {(obs, k')} * \

(fst) {k'})" using CP_t_K[OF `k\keys` obs] CP_t_K[OF _ obs] `real (card ?S) \ 0` @@ -445,11 +439,10 @@ mult_divide_mult_cancel_left[OF `real (card ?S) \ 0`] cong: setsum_cong) also have "(\k'\keys. \

(OB|fst) {(obs, k')} * \

(fst) {k'}) = \

(OB) {obs}" - using setsum_real_distribution(2)[of OB fst obs, symmetric] - using real_distribution_order'[of fst _ OB obs] by (auto intro!: setsum_cong) + using setsum_distribution(2)[of OB fst obs, symmetric] + by (auto intro!: setsum_cong distribution_order(8)) also have "\

(OB | fst) {(obs, k)} * \

(fst) {k} / \

(OB) {obs} = \

(fst | OB) {(k, obs)}" - using real_distribution_order'[of fst k OB obs] - by (subst joint_distribution_commute) auto + by (subst joint_distribution_commute) (auto intro!: distribution_order(8)) finally have "\

(fst | t\OB) {(k, t obs)} = \

(fst | OB) {(k, obs)}" . } note CP_T_eq_CP_O = this @@ -472,7 +465,7 @@ unfolding disjoint_family_on_def by auto have "\

(t\OB) {t (OB x)} = (\obs\?S (OB x). \

(OB) {obs})" unfolding distribution_def comp_def - using real_finite_measure_finite_Union[OF _ df] + using finite_measure_finite_Union[OF _ _ df] by (force simp add: * intro!: setsum_nonneg) } note P_t_sum_P_O = this